From 4da7d053210483b56ef07e2e7447345577012592 Mon Sep 17 00:00:00 2001 From: Axel Date: Wed, 12 Feb 2020 23:43:03 +0100 Subject: [PATCH 01/30] First update --- Cubical/Foundations/HLevels.agda | 10 + Cubical/Foundations/Pointed/Base.agda | 5 + Cubical/Foundations/Prelude.agda | 1 + .../PropositionalTruncation/Properties.agda | 9 + Cubical/HITs/SetTruncation/Properties.agda | 20 ++ Cubical/HITs/Truncation/Base.agda | 1 + Cubical/HITs/Truncation/Properties.agda | 188 +++++++++++++++++- Cubical/ZCohomology/Base.agda | 51 +++++ Cubical/ZCohomology/Everything.agda | 7 + Cubical/ZCohomology/Properties.agda | 88 ++++++++ Cubical/ZCohomology/S1/S1.agda | 57 ++++++ 11 files changed, 435 insertions(+), 2 deletions(-) create mode 100644 Cubical/ZCohomology/Base.agda create mode 100644 Cubical/ZCohomology/Everything.agda create mode 100644 Cubical/ZCohomology/Properties.agda create mode 100644 Cubical/ZCohomology/S1/S1.agda diff --git a/Cubical/Foundations/HLevels.agda b/Cubical/Foundations/HLevels.agda index 0be65a054..b7e2c3324 100644 --- a/Cubical/Foundations/HLevels.agda +++ b/Cubical/Foundations/HLevels.agda @@ -246,6 +246,16 @@ hLevelSuc 0 A = isContr→isProp hLevelSuc 1 A = isProp→isSet hLevelSuc (suc (suc n)) A h a b = hLevelSuc (suc n) (a ≡ b) (h a b) +hLevelAdd : (n : ℕ) (m : ℕ) (A : Type ℓ) → isOfHLevel n A → isOfHLevel (n + m) A +hLevelAdd n zero A h = transport (λ i → isOfHLevel (helper n i) A) h where + helper : (n : ℕ) → n ≡ n + zero + helper zero = refl + helper (suc n) = cong (λ x → (suc x)) (helper n) +hLevelAdd n (suc m) A h = transport (λ i → (isOfHLevel (helper n m i) A )) (hLevelSuc (n + m) A (hLevelAdd n m A h)) where + helper : (n m : ℕ) → (suc (n + m)) ≡ (n + (suc m)) + helper zero m = refl + helper (suc n) m = cong (λ x → (suc x)) (helper n m) + hLevelPath : (n : ℕ) → isOfHLevel n A → (x y : A) → isOfHLevel n (x ≡ y) hLevelPath 0 h x y = isContrPath h x y hLevelPath 1 h x y = hLevelSuc zero _ (isProp→isContr≡ h x y) diff --git a/Cubical/Foundations/Pointed/Base.agda b/Cubical/Foundations/Pointed/Base.agda index 83ac03040..2c2d7e6d4 100644 --- a/Cubical/Foundations/Pointed/Base.agda +++ b/Cubical/Foundations/Pointed/Base.agda @@ -11,3 +11,8 @@ typ = fst pt : ∀ {ℓ} (A∙ : Pointed ℓ) → typ A∙ pt = snd + +{- Pointed functions -} +_→*_ : ∀{ℓ ℓ'} → (A : Pointed ℓ) → (B : Pointed ℓ') → Type (ℓ-max ℓ ℓ') +_→*_ A B = Σ (typ A → typ B) λ f → f (pt A) ≡ pt B + diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index b35b3af20..31086087a 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -131,6 +131,7 @@ transport p a = transp (λ i → p i) i0 a transportRefl : (x : A) → transport refl x ≡ x transportRefl {A = A} x i = transp (λ _ → A) i x + -- We want B to be explicit in subst subst : (B : A → Type ℓ') (p : x ≡ y) → B x → B y subst B p pa = transport (λ i → B (p i)) pa diff --git a/Cubical/HITs/PropositionalTruncation/Properties.agda b/Cubical/HITs/PropositionalTruncation/Properties.agda index 26c2b0dfe..a22739907 100644 --- a/Cubical/HITs/PropositionalTruncation/Properties.agda +++ b/Cubical/HITs/PropositionalTruncation/Properties.agda @@ -40,6 +40,15 @@ elimPropTrunc {A = A} {P = P} Pprop f (squash x y i) = PpropOver : {a b : ∥ A ∥} → (sq : a ≡ b) → ∀ x y → PathP (λ i → P (sq i)) x y PpropOver {a} = J (λ b (sq : a ≡ b) → ∀ x y → PathP (λ i → P (sq i)) x y) (Pprop a) + +propId : isProp A → ∥ A ∥ ≡ A +propId {A = A} hA = isoToPath (iso (recPropTrunc hA (idfun A)) (λ x → ∣ x ∣) (λ _ → refl) rinv) + where + rinv : ∀ (x : ∥ A ∥) → ∣ recPropTrunc hA (idfun A) x ∣ ≡ x + rinv x = elimPropTrunc {P = λ x → ∣ recPropTrunc hA (idfun A) x ∣ ≡ x} + (λ _ → isProp→isSet propTruncIsProp _ _) + (λ _ → refl) x + -- We could also define the eliminator using the recursor elimPropTrunc' : ∀ {P : ∥ A ∥ → Type ℓ} → ((a : ∥ A ∥) → isProp (P a)) → ((x : A) → P ∣ x ∣) → (a : ∥ A ∥) → P a diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index 6a9f364c7..ebe21aa20 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -56,3 +56,23 @@ elimSetTrunc3 : {B : (x y z : ∥ A ∥₀) → Type ℓ} (x y z : ∥ A ∥₀) → B x y z elimSetTrunc3 Bset g = elimSetTrunc2 (λ _ _ → hLevelPi 2 λ _ → Bset _ _ _) (λ a b → elimSetTrunc (λ _ → Bset _ _ _) (λ c → g a b c)) + + +setTruncIsSet : isSet ∥ A ∥₀ +setTruncIsSet a b p q = squash₀ a b p q + +setId : isSet A → A ≡ ∥ A ∥₀ +setId {A = A} isset = isoToPath (iso (λ x → ∣ x ∣₀) + (elimSetTrunc {A = A} + (λ _ → isset) + (λ x → x)) + (λ a → idLemma a) + (λ _ → refl)) + where + idLemma : ∀ (b : ∥ A ∥₀) → ∣ elimSetTrunc (λ x → isset) (λ x → x) b ∣₀ ≡ b + idLemma b = elimSetTrunc {B = (λ x → ∣ elimSetTrunc (λ _ → isset) (λ x → x) x ∣₀ ≡ x)} + (λ x → (hLevelSuc 2 ∥ A ∥₀ (setTruncIsSet {A = A})) + ∣ elimSetTrunc (λ _ → isset) (λ x₁ → x₁) x ∣₀ + x) + (λ _ → refl) + b diff --git a/Cubical/HITs/Truncation/Base.agda b/Cubical/HITs/Truncation/Base.agda index ef1993225..f69ddacb9 100644 --- a/Cubical/HITs/Truncation/Base.agda +++ b/Cubical/HITs/Truncation/Base.agda @@ -22,3 +22,4 @@ open import Cubical.HITs.Sn ∥_∥_ : ∀ {ℓ} → Type ℓ → ℕ₋₂ → Type ℓ ∥ A ∥ n = Null (S (1+ n)) A + diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index 231f19698..d720fe948 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -13,9 +13,9 @@ open Modality open import Cubical.Data.Empty open import Cubical.Data.Nat -open import Cubical.Data.NatMinusOne using (ℕ₋₁; neg1; suc; ℕ→ℕ₋₁) +open import Cubical.Data.NatMinusOne using (ℕ₋₁; neg1; suc; ℕ→ℕ₋₁) renaming (-1+_ to -1+₋₁_ ; 1+_ to 1+₋₁_) import Cubical.Data.NatMinusOne as ℕ₋₁ -open import Cubical.Data.NatMinusTwo +open import Cubical.Data.NatMinusTwo open import Cubical.HITs.Sn open import Cubical.HITs.Susp open import Cubical.HITs.Nullification @@ -124,10 +124,20 @@ isSnNull→isOfHLevel {n = suc n} nA = isSphereFilled→isOfHLevelSuc (λ f → isOfHLevel∥∥ : (n : ℕ₋₂) → isOfHLevel (2+ n) (∥ A ∥ n) isOfHLevel∥∥ neg2 = hub ⊥-elim , λ _ → ≡hub ⊥-elim isOfHLevel∥∥ (suc n) = isSphereFilled→isOfHLevelSuc isSphereFilled∥∥ + + + -- isOfHLevel∥∥ n = isSnNull→isOfHLevel isNull-Null -- ∥_∥ n is a modality +rec : {n : ℕ₋₂} + {B : Type ℓ'} → + (isOfHLevel (2+ n) B) → + (g : (a : A) → B) → + (∥ A ∥ n → B) +rec {B = B} h = Null-ind {B = λ _ → B} λ x → isOfHLevel→isSnNull h + ind : {n : ℕ₋₂} {B : ∥ A ∥ n → Type ℓ'} (hB : (x : ∥ A ∥ n) → isOfHLevel (2+ n) (B x)) @@ -145,6 +155,9 @@ ind2 : {n : ℕ₋₂} ind2 {n = n} hB g = ind (λ _ → hLevelPi (2+ n) (λ _ → hB _ _)) λ a → ind (λ _ → hB _ _) (λ b → g a b) + + + ind3 : {n : ℕ₋₂} {B : (x y z : ∥ A ∥ n) → Type ℓ'} (hB : ((x y z : ∥ A ∥ n) → isOfHLevel (2+ n) (B x y z))) @@ -204,3 +217,174 @@ groupoidTrunc≃Trunc1 = (ind (λ _ → squash₂) ∣_∣₂) (ind (λ _ → hLevelPath 4 (isOfHLevel∥∥ 2) _ _) (λ _ → refl)) (g2TruncElim _ _ (λ _ → hLevelPath 4 squash₂ _ _) (λ _ → refl))) + + + +---- ∥ Ω A ∥ ₙ ≡ Ω ∥ A ∥ₙ₊₁ ---- + + + {- Proofs of Theorem 7.3.12. and Corollary 7.3.13. in the HoTT book -} + +private + {- We define the fibration P to show a more general result -} + P : ∀ {ℓ} {B : Type ℓ}{n : ℕ₋₂} → ∥ B ∥ (suc n) → ∥ B ∥ (suc n) → Type ℓ + P x y = fst (P₁ x y) + + where + P₁ : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} → ∥ B ∥ (suc n) → ∥ B ∥ (suc n) → (HLevel ℓ (2+ n)) + P₁ {ℓ} {n = n} x y = ind2 (λ _ _ → hLevelHLevel (2+ n)) + (λ a b → (∥ a ≡ b ∥ n , isOfHLevel∥∥ n )) + x + y + + {- We will need P to be of hLevel n + 3 -} + hLevelP : ∀{ℓ} {n : ℕ₋₂} {B : Type ℓ} + (a b : ∥ B ∥ (suc n)) → + isOfHLevel (2+ (suc n)) (P a b ) + hLevelP {n = n} {B = B} = ind2 {A = B} + {n = (suc n)} + {B = λ x y → isOfHLevel (2+ (suc n)) (P x y)} + (λ x y → hLevelAdd 1 (2+ n) (isOfHLevel (2+ suc n) (P x y)) (isPropIsOfHLevel (2+ suc n) (P x y)) ) + λ a b → ( hLevelSuc (2+ n) (P ∣ a ∣ ∣ b ∣) ) + (isOfHLevel∥∥ {A = a ≡ b} n) + + {- decode function from P x y to x ≡ y -} + decode-fun : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x y : ∥ B ∥ (suc n)) → + P x y → + _≡_ {A = ∥ B ∥ (suc n)} x y + + decode-fun {B = B} {n = n} x y = ind2 {B = λ u v → P u v → _≡_ {A = ∥ B ∥ (suc n)} u v } + (λ u v → hLevelPi {A = P u v} {B = λ _ → u ≡ v} + (2+ suc n) + λ _ → (((hLevelSuc (2+ suc n) (∥ B ∥ suc n) (isOfHLevel∥∥ {A = B} (suc n))) u v)) ) + decode* x y + where + decode* : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂}(u v : B) → + (P {n = n} + (∣_∣ {S = (S (1+ (suc n)))} u) + (∣_∣ {S = (S (1+ (suc n)))} v)) → + _≡_ {A = ∥ B ∥ (suc n) } + ∣ u ∣ + ∣ v ∣ + decode* {B = B} {n = neg2} u v = rec {A = u ≡ v} {n = neg2} + {B = _≡_ {A = ∥ B ∥ (suc (neg2)) } (∣ u ∣) (∣ v ∣)} + ((isOfHLevel∥∥ {A = B} (suc neg2) ∣ u ∣ ∣ v ∣) , + λ y → (hLevelSuc (2+ suc neg2) (∥ B ∥ suc neg2) + (isOfHLevel∥∥ {A = B} (suc neg2)) ∣ u ∣ ∣ v ∣) + (isOfHLevel∥∥ {A = B} (suc neg2) ∣ u ∣ ∣ v ∣) y ) + (λ p → cong (λ z → ∣ z ∣) p) + decode* {B = B} {n = suc n} u v = rec {A = u ≡ v} {n = suc n} + {B = _≡_ {A = ∥ B ∥ (suc (suc n)) } (∣ u ∣) (∣ v ∣)} + (isOfHLevel∥∥ {A = B} (suc (suc n)) ∣ u ∣ ∣ v ∣) + (λ p → cong (λ z → ∣ z ∣) p) + + + {- auxilliary function r used to define encode -} + r : ∀ {ℓ} {B : Type ℓ} {m : ℕ₋₂} (u : ∥ B ∥ (suc m)) → P u u + r {m = m} = ind {B = (λ u → P u u)} + (λ x → hLevelP x x) + (λ a → ∣ refl {x = a} ∣) + + {- encode function from x ≡ y to P x y -} + encode-fun : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x y : ∥ B ∥ (suc n)) → + _≡_ {A = ∥ B ∥ (suc n)} x y → + P x y + encode-fun x y p = transport (λ i → P x (p i )) (r x) + + {- We need the following two lemmas on the functions behaviour for refl -} + dec-refl : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x : ∥ B ∥ (suc n)) → + decode-fun x x (r x) ≡ refl {x = x} + dec-refl {B = B} {n = neg2} x = ind {A = B} {n = suc neg2} + {B = λ x → decode-fun x x (r x) ≡ refl {x = x} } + (λ x → (hLevelSuc (2+ (suc neg2)) (x ≡ x) + (hLevelSuc (2+ (suc neg2)) + (∥ B ∥ suc neg2) + (isOfHLevel∥∥ {A = B} (suc neg2)) x x)) + (decode-fun x x (r x)) refl) + (λ a → refl) x + dec-refl {B = B} {n = suc n} = ind {A = B} {n = suc (suc n)} + {B = λ x → decode-fun x x (r x) ≡ refl {x = x} } + (λ x → hLevelSuc (2+ suc n) (decode-fun x x (r x) ≡ refl) + (hLevelSuc (2+ suc n) (x ≡ x) + (isOfHLevel∥∥ {A = B} (suc (suc n)) x x ) + (decode-fun x x (r x)) refl)) + λ c → refl + + + enc-refl : ∀ {ℓ} {B : Type ℓ} + {n : ℕ₋₂} + (x : ∥ B ∥ (suc n)) → + encode-fun x x (refl {x = x}) ≡ r x + enc-refl x j = transp (λ i → P x (refl {x = x} i)) j (r x) + + + + {- decode-fun is a right-inverse -} + P-rinv : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (u v : ∥ B ∥ (suc n)) → + (x : _≡_ {A = ∥ B ∥ (suc n)} u v) → + decode-fun u v (encode-fun u v x) ≡ x + P-rinv {ℓ = ℓ} {B = B} {n = n} u v = J {ℓ} { ∥ B ∥ (suc n)} {u} {ℓ} + (λ y p → decode-fun u y (encode-fun u y p) ≡ p) + ((λ i → (decode-fun u u (enc-refl u i))) ∙ dec-refl u) + {v} + + {- decode-fun is a left-inverse -} + P-linv : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (u v : ∥ B ∥ (suc n )) → + (x : P u v) → + encode-fun u v (decode-fun u v x) ≡ x + P-linv {ℓ = ℓ} {B = B} {n = n} u v = ind2 {A = B} + {n = (suc n)} + {B = λ u v → (x : P u v) → encode-fun u v (decode-fun u v x) ≡ x} + (λ x y → hLevelPi {A = P x y} + (2+ suc n) + λ z → hLevelSuc (2+ suc n) (P x y) + (hLevelP {n = n} x y) (encode-fun x y (decode-fun x y z)) z) + helper u v + + + where + helper : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} + (a b : B) + (x : P {n = n} ∣ a ∣ ∣ b ∣) → + encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x) ≡ x + helper {ℓ = ℓ} {B = B} {n = neg2} a b = ind {A = (a ≡ b)} + {n = neg2} + {B = λ x → encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x) ≡ x} + (λ x → (sym ((snd (isOfHLevel∥∥ {A = a ≡ b} neg2)) + (encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x))) + ∙ ((snd (isOfHLevel∥∥ {A = a ≡ b} neg2)) x)) + , + λ y → hLevelSuc (2+ (suc neg2)) (∥ a ≡ b ∥ neg2) + (hLevelSuc (2+ neg2) (∥ a ≡ b ∥ neg2) (isOfHLevel∥∥ {A = a ≡ b} (neg2))) + (encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x)) x + ((sym ((snd (isOfHLevel∥∥ {A = a ≡ b} neg2)) + (encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x))) + ∙ ((snd (isOfHLevel∥∥ {A = a ≡ b} neg2)) x))) + y) + (J {ℓ}{B}{a}{ℓ} + ((λ y p → (encode-fun {n = neg2}) ∣ a ∣ ∣ y ∣ + ((decode-fun ∣ a ∣ ∣ y ∣) ∣ p ∣) ≡ ∣ p ∣)) + (enc-refl {n = neg2} ∣ a ∣ ) + {b}) + helper {ℓ = ℓ} {B = B} {n = suc n} a b = ind {A = (a ≡ b)} + {n = suc n} + {B = λ x → encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x) ≡ x} + (λ x → (hLevelP {n = suc n} ∣ a ∣ ∣ b ∣ (encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x)) x) ) + (J {ℓ}{B}{a}{ℓ} ((λ y p → (encode-fun {n = suc n}) ∣ a ∣ ∣ y ∣ ((decode-fun ∣ a ∣ ∣ y ∣) ∣ p ∣) ≡ ∣ p ∣)) + (enc-refl {n = suc n} ∣ a ∣ ) + {b}) + + {- The final Iso established -} + IsoFinal : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x y : ∥ B ∥ (suc n)) → Iso (x ≡ y) (P x y) + IsoFinal x y = iso (encode-fun x y ) (decode-fun x y) (P-linv x y) (P-rinv x y) + +IsoIdTrunc : {a b : A} (n : ℕ₋₂) → Iso (_≡_ {A = ∥ A ∥ (suc n)} ∣ a ∣ ∣ b ∣) (∥ a ≡ b ∥ n) +IsoIdTrunc {a = a} {b = b} n = IsoFinal {n = n} ∣ a ∣ ∣ b ∣ + +IsoΩ : {a : A} (n : ℕ₋₂) → Iso (_≡_ {A = ∥ A ∥ (suc n)} ∣ a ∣ ∣ a ∣) (∥ a ≡ a ∥ n) +IsoΩ {a = a} n = IsoIdTrunc {a = a} {b = a} n + + + + + diff --git a/Cubical/ZCohomology/Base.agda b/Cubical/ZCohomology/Base.agda new file mode 100644 index 000000000..8732a2aae --- /dev/null +++ b/Cubical/ZCohomology/Base.agda @@ -0,0 +1,51 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.Base where + + +open import Cubical.Data.Int.Base +open import Cubical.Data.Nat.Base +open import Cubical.Data.NatMinusTwo.Base +open import Cubical.Data.Sigma.Base + +open import Cubical.Foundations.Pointed.Base + +open import Cubical.HITs.Nullification.Base +open import Cubical.HITs.SetTruncation.Base +open import Cubical.HITs.Sn.Base +open import Cubical.HITs.Susp.Base +open import Cubical.HITs.Truncation.Base + +private + variable + ℓ : Level + A : Type ℓ + +{- +Definition of cohomology with cocefficients in the integers. For now, this definition uses ℕ₋₂, for instance defining H⁰ as +coHom neg2, H¹ as coHom (suc neg2), and so on. This has so far been easier to work with than a definition relying on ℕ→ℕ₋₂. +-} + + +--- Cohomology --- + +{- Types Kₙ from Brunerie 2016 -} +coHomK : (n : ℕ₋₂) → Type₀ +coHomK neg2 = Int +coHomK (suc n) = ∥ S₊ (2+ suc n) ∥ (suc (suc (suc n))) + +{- Cohomology -} +coHom : (n : ℕ₋₂) → Type ℓ → Type ℓ +coHom n A = ∥ (A → coHomK n) ∥₀ + + +--- Reduced cohomology --- + +{- Pointed version of Kₙ -} +coHomK-ptd : (n : ℕ₋₂) → Pointed (ℓ-zero) +coHomK-ptd neg2 = coHomK neg2 , (pos 0) +coHomK-ptd (suc n) = (coHomK (suc n) , ∣ north ∣) + +{- Reduced cohomology -} +coHomRed : ∀ {ℓ} → (n : ℕ₋₂) → (A : Pointed ℓ) → Type ℓ +coHomRed neg2 A = ∥ (A →* (coHomK-ptd neg2)) ∥₀ +coHomRed (suc n) A = ∥ (A →* (coHomK-ptd (suc n))) ∥₀ diff --git a/Cubical/ZCohomology/Everything.agda b/Cubical/ZCohomology/Everything.agda new file mode 100644 index 000000000..160104dce --- /dev/null +++ b/Cubical/ZCohomology/Everything.agda @@ -0,0 +1,7 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.Everything where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties + +open import Cubical.ZCohomology.S1.S1 diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda new file mode 100644 index 000000000..b4a2583d1 --- /dev/null +++ b/Cubical/ZCohomology/Properties.agda @@ -0,0 +1,88 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.Properties where + + +open import Cubical.ZCohomology.Base +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.NatMinusTwo.Base renaming (-1+_ to -1+₋₂_ ; 1+_ to 1+₋₂_) +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation +open import Cubical.HITs.Nullification +open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup +private + variable + ℓ : Level + A : Type ℓ + + + + + +{- Equivalence between cohomology and the reduced version -} +coHom↔coHomRed : ∀ {ℓ} → (n : ℕ₋₂) → (A : Set ℓ) → (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) +coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ + module ClaireCaron where + helplemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →* C))) + helplemma {C = C} = isoToPath (iso map1 map2 (λ b → linvPf b) (λ a → refl)) where + map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →* C)) + map1 f = map1' , refl module helpmap where + map1' : A ⊎ Unit → fst C + map1' (inl x) = f x + map1' (inr x) = pt C + + map2 : ((((A ⊎ Unit) , inr (tt)) →* C)) → (A → typ C) + map2 (g , pf) x = g (inl x) + linvPf : (b :((((A ⊎ Unit) , inr (tt)) →* C))) → map1 (map2 b) ≡ b + linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) where + helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x + helper (inl x) = refl + helper (inr tt) = sym snd + +coHom↔coHomRed (suc n) A i = ∥ ClaireCaron.helplemma A i {C = ((K (suc n)) , ∣ north ∣)} i ∥₀ + + + + +{- Unfinished proof that Kₙ ≡ Ω Kₙ₊₁ -} + +invPathCancel : {a b : A} → (p : a ≡ b) → p ∙ (sym p) ≡ refl +invPathCancel {a = a} {b = b} p = J {x = a} (λ y p → p ∙ (sym p) ≡ refl ) (sym (lUnit refl)) p + +φ : (a : A) → A → (north {A = A} ≡ north {A = A}) +φ x = (λ a → ((merid a) ∙ sym ((merid x)))) + +φ* : (A : Pointed ℓ) → A →* Ω ((Susp (typ A)) , north) +φ* A = (φ (pt A)) , invPathCancel (merid (pt A)) + +σ : (n : ℕ₋₂) → (typ (K-ptd n)) → (typ (Ω (K-ptd (suc n)))) +σ neg2 k = looper k ( cong {B = λ _ → (∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1)} + (λ x → ∣ x ∣) + ((merid north) ∙ (sym (merid south))) ) + where + looper : ∀ {ℓ} → {A : Type ℓ} → {a : A} → (n : Int) → (a ≡ a) → (a ≡ a) -- defining loopⁿ + looper {A = A} {a = a} (pos zero) p = refl + looper {A = A} {a = a} (pos (suc n)) p = (looper (pos n) p) ∙ p + looper {A = A} {a = a} (negsuc zero) p = sym p + looper {A = A} {a = a} (negsuc (suc n)) p = (sym p) ∙ (looper (negsuc n) p) +σ (suc n) x = rec {n = suc (suc (suc n))} + {B = (typ (Ω (K-ptd (suc (suc n)))))} + (isOfHLevel∥∥ {A = S₊ (2+ suc (suc n))} (suc (suc (suc (suc n)))) ∣ north ∣ ∣ north ∣) + (λ y → cong {B = λ _ → Null (S (1+₋₂ (suc (suc (suc (suc n)))))) (S₊ (2+ (suc (suc n))))} + (λ z → ∣ z ∣) + (φ north y)) + x + + + + diff --git a/Cubical/ZCohomology/S1/S1.agda b/Cubical/ZCohomology/S1/S1.agda new file mode 100644 index 000000000..605253014 --- /dev/null +++ b/Cubical/ZCohomology/S1/S1.agda @@ -0,0 +1,57 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.S1.S1 where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.Data.Sigma +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.Data.Empty.Base +open import Cubical.Data.Unit.Base +open import Cubical.Foundations.Everything +open import Cubical.Data.NatMinusTwo.Base renaming (-1+_ to -1+₋₂_ ; 1+_ to 1+₋₂_) +open import Cubical.Foundations.Equiv +open import Cubical.Data.Prod +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation +open import Cubical.HITs.Nullification +open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.Unit.Base +open import Cubical.Data.HomotopyGroup +private + variable + ℓ : Level + A : Type ℓ + + + +---- H⁰(S¹) ---- + +coHom0-S1 : (coHom neg2 S¹) ≡ Int +coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ sym (setId isSetInt) + where + helpLemma : (S¹ → Int) ≡ Int + helpLemma = isoToPath (iso fun funinv (λ b → refl) λ f → funExt (λ x → rinvLemma f x)) where + fun : (S¹ → Int) → Int + fun f = f base + + funinv : Int → (S¹ → Int) + funinv a base = a + funinv a (loop i) = refl {x = a} i + + rinvLemma : ( f : (S¹ → Int)) → (x : S¹) → funinv (fun f) x ≡ f x + rinvLemma f base = refl + rinvLemma f (loop i) = sym (helper f i) where + helper : (f : (S¹ → Int)) → (i : I) → (f (loop i) ≡ f base) + helper f i = (λ l → ((isSetInt (f base) (f base) (λ k → (f (loop k))) refl) l) i) + +------------------------- + + + + + From 56a5e0f5b5e0c271a482a03673ea354acb693807 Mon Sep 17 00:00:00 2001 From: Axel Date: Thu, 13 Feb 2020 00:06:24 +0100 Subject: [PATCH 02/30] Minor fixes --- Cubical/ZCohomology/Properties.agda | 19 ++++++++++--------- Cubical/ZCohomology/S1/S1.agda | 4 ++-- 2 files changed, 12 insertions(+), 11 deletions(-) diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index b4a2583d1..de90ec943 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -13,7 +13,7 @@ open import Cubical.Data.NatMinusTwo.Base renaming (-1+_ to -1+₋₂_ ; 1+_ to open import Cubical.HITs.Susp open import Cubical.HITs.SetTruncation open import Cubical.HITs.Nullification -open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Int open import Cubical.Data.Nat open import Cubical.HITs.Truncation @@ -27,12 +27,10 @@ private - - {- Equivalence between cohomology and the reduced version -} coHom↔coHomRed : ∀ {ℓ} → (n : ℕ₋₂) → (A : Set ℓ) → (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ - module ClaireCaron where + module coHom↔coHomRed where helplemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →* C))) helplemma {C = C} = isoToPath (iso map1 map2 (λ b → linvPf b) (λ a → refl)) where map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →* C)) @@ -49,12 +47,12 @@ coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ helper (inl x) = refl helper (inr tt) = sym snd -coHom↔coHomRed (suc n) A i = ∥ ClaireCaron.helplemma A i {C = ((K (suc n)) , ∣ north ∣)} i ∥₀ +coHom↔coHomRed (suc n) A i = ∥ coHom↔coHomRed.helplemma A i {C = ((coHomK (suc n)) , ∣ north ∣)} i ∥₀ -{- Unfinished proof that Kₙ ≡ Ω Kₙ₊₁ -} +{- TODO: Prove Kₙ ≡ Ω Kₙ₊₁ -} invPathCancel : {a b : A} → (p : a ≡ b) → p ∙ (sym p) ≡ refl invPathCancel {a = a} {b = b} p = J {x = a} (λ y p → p ∙ (sym p) ≡ refl ) (sym (lUnit refl)) p @@ -65,18 +63,21 @@ invPathCancel {a = a} {b = b} p = J {x = a} (λ y p → p ∙ (sym p) ≡ refl ) φ* : (A : Pointed ℓ) → A →* Ω ((Susp (typ A)) , north) φ* A = (φ (pt A)) , invPathCancel (merid (pt A)) -σ : (n : ℕ₋₂) → (typ (K-ptd n)) → (typ (Ω (K-ptd (suc n)))) +σ : (n : ℕ₋₂) → (typ (coHomK-ptd n)) → (typ (Ω (coHomK-ptd (suc n)))) σ neg2 k = looper k ( cong {B = λ _ → (∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1)} (λ x → ∣ x ∣) ((merid north) ∙ (sym (merid south))) ) where - looper : ∀ {ℓ} → {A : Type ℓ} → {a : A} → (n : Int) → (a ≡ a) → (a ≡ a) -- defining loopⁿ + looper : ∀ {ℓ}{A : Type ℓ}{a : A} -- defining loopⁿ + (n : Int) → + (a ≡ a) → + (a ≡ a) looper {A = A} {a = a} (pos zero) p = refl looper {A = A} {a = a} (pos (suc n)) p = (looper (pos n) p) ∙ p looper {A = A} {a = a} (negsuc zero) p = sym p looper {A = A} {a = a} (negsuc (suc n)) p = (sym p) ∙ (looper (negsuc n) p) σ (suc n) x = rec {n = suc (suc (suc n))} - {B = (typ (Ω (K-ptd (suc (suc n)))))} + {B = (typ (Ω (coHomK-ptd (suc (suc n)))))} (isOfHLevel∥∥ {A = S₊ (2+ suc (suc n))} (suc (suc (suc (suc n)))) ∣ north ∣ ∣ north ∣) (λ y → cong {B = λ _ → Null (S (1+₋₂ (suc (suc (suc (suc n)))))) (S₊ (2+ (suc (suc n))))} (λ z → ∣ z ∣) diff --git a/Cubical/ZCohomology/S1/S1.agda b/Cubical/ZCohomology/S1/S1.agda index 605253014..c349d04ed 100644 --- a/Cubical/ZCohomology/S1/S1.agda +++ b/Cubical/ZCohomology/S1/S1.agda @@ -15,7 +15,7 @@ open import Cubical.Data.Prod open import Cubical.HITs.Susp open import Cubical.HITs.SetTruncation open import Cubical.HITs.Nullification -open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Int open import Cubical.Data.Nat open import Cubical.HITs.Truncation open import Cubical.HITs.Pushout @@ -51,7 +51,7 @@ coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ sym (setId isSetInt) ------------------------- - +{- TODO : give Hᵏ(S¹) for all k -} From 098153c949c8731c6948f4c03756d32d934e37b5 Mon Sep 17 00:00:00 2001 From: Axel Date: Thu, 13 Feb 2020 00:14:40 +0100 Subject: [PATCH 03/30] Cleanup --- Cubical/ZCohomology/Base.agda | 2 +- Cubical/ZCohomology/Properties.agda | 22 +++++++++++++++------- 2 files changed, 16 insertions(+), 8 deletions(-) diff --git a/Cubical/ZCohomology/Base.agda b/Cubical/ZCohomology/Base.agda index 8732a2aae..1c02e2057 100644 --- a/Cubical/ZCohomology/Base.agda +++ b/Cubical/ZCohomology/Base.agda @@ -46,6 +46,6 @@ coHomK-ptd neg2 = coHomK neg2 , (pos 0) coHomK-ptd (suc n) = (coHomK (suc n) , ∣ north ∣) {- Reduced cohomology -} -coHomRed : ∀ {ℓ} → (n : ℕ₋₂) → (A : Pointed ℓ) → Type ℓ +coHomRed : (n : ℕ₋₂) → (A : Pointed ℓ) → Type ℓ coHomRed neg2 A = ∥ (A →* (coHomK-ptd neg2)) ∥₀ coHomRed (suc n) A = ∥ (A →* (coHomK-ptd (suc n))) ∥₀ diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index de90ec943..139514488 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -28,11 +28,17 @@ private {- Equivalence between cohomology and the reduced version -} -coHom↔coHomRed : ∀ {ℓ} → (n : ℕ₋₂) → (A : Set ℓ) → (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) +coHom↔coHomRed : (n : ℕ₋₂) → + (A : Set ℓ) → + (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ module coHom↔coHomRed where helplemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →* C))) - helplemma {C = C} = isoToPath (iso map1 map2 (λ b → linvPf b) (λ a → refl)) where + helplemma {C = C} = isoToPath (iso map1 + map2 + (λ b → linvPf b) + (λ _ → refl)) + where map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →* C)) map1 f = map1' , refl module helpmap where map1' : A ⊎ Unit → fst C @@ -41,8 +47,10 @@ coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ map2 : ((((A ⊎ Unit) , inr (tt)) →* C)) → (A → typ C) map2 (g , pf) x = g (inl x) + linvPf : (b :((((A ⊎ Unit) , inr (tt)) →* C))) → map1 (map2 b) ≡ b - linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) where + linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) + where helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x helper (inl x) = refl helper (inr tt) = sym snd @@ -68,10 +76,10 @@ invPathCancel {a = a} {b = b} p = J {x = a} (λ y p → p ∙ (sym p) ≡ refl ) (λ x → ∣ x ∣) ((merid north) ∙ (sym (merid south))) ) where - looper : ∀ {ℓ}{A : Type ℓ}{a : A} -- defining loopⁿ - (n : Int) → - (a ≡ a) → - (a ≡ a) + looper : {a : A} -- defining loopⁿ + (n : Int) → + (a ≡ a) → + (a ≡ a) looper {A = A} {a = a} (pos zero) p = refl looper {A = A} {a = a} (pos (suc n)) p = (looper (pos n) p) ∙ p looper {A = A} {a = a} (negsuc zero) p = sym p From 67a450d928db1674957a1596a50074ac425dccdf Mon Sep 17 00:00:00 2001 From: Axel Date: Thu, 13 Feb 2020 00:19:05 +0100 Subject: [PATCH 04/30] Cleanup --- Cubical/ZCohomology/S1/S1.agda | 14 +++++++++++--- 1 file changed, 11 insertions(+), 3 deletions(-) diff --git a/Cubical/ZCohomology/S1/S1.agda b/Cubical/ZCohomology/S1/S1.agda index c349d04ed..311fc0590 100644 --- a/Cubical/ZCohomology/S1/S1.agda +++ b/Cubical/ZCohomology/S1/S1.agda @@ -35,7 +35,11 @@ coHom0-S1 : (coHom neg2 S¹) ≡ Int coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ sym (setId isSetInt) where helpLemma : (S¹ → Int) ≡ Int - helpLemma = isoToPath (iso fun funinv (λ b → refl) λ f → funExt (λ x → rinvLemma f x)) where + helpLemma = isoToPath (iso fun + funinv + (λ b → refl) + λ f → funExt (λ x → rinvLemma f x)) + where fun : (S¹ → Int) → Int fun f = f base @@ -43,10 +47,14 @@ coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ sym (setId isSetInt) funinv a base = a funinv a (loop i) = refl {x = a} i - rinvLemma : ( f : (S¹ → Int)) → (x : S¹) → funinv (fun f) x ≡ f x + rinvLemma : ( f : (S¹ → Int)) → + (x : S¹) → + funinv (fun f) x ≡ f x rinvLemma f base = refl rinvLemma f (loop i) = sym (helper f i) where - helper : (f : (S¹ → Int)) → (i : I) → (f (loop i) ≡ f base) + helper : (f : (S¹ → Int)) → + (i : I) → + (f (loop i) ≡ f base) helper f i = (λ l → ((isSetInt (f base) (f base) (λ k → (f (loop k))) refl) l) i) ------------------------- From 1df167b1d7c5c02474d213ec9129b2a8ff811897 Mon Sep 17 00:00:00 2001 From: Axel Date: Mon, 24 Feb 2020 16:19:36 +0100 Subject: [PATCH 05/30] fixes --- Cubical/Foundations/HLevels.agda | 10 ---- Cubical/Foundations/Pointed/Base.agda | 4 +- Cubical/Foundations/Prelude.agda | 1 - Cubical/HITs/SetTruncation/Properties.agda | 10 ++-- Cubical/HITs/Truncation/Base.agda | 1 - Cubical/HITs/Truncation/Properties.agda | 17 ++---- Cubical/ZCohomology/Base.agda | 24 +++----- Cubical/ZCohomology/Properties.agda | 65 ++++++---------------- Cubical/ZCohomology/S1/S1.agda | 24 ++------ 9 files changed, 44 insertions(+), 112 deletions(-) diff --git a/Cubical/Foundations/HLevels.agda b/Cubical/Foundations/HLevels.agda index b7e2c3324..0be65a054 100644 --- a/Cubical/Foundations/HLevels.agda +++ b/Cubical/Foundations/HLevels.agda @@ -246,16 +246,6 @@ hLevelSuc 0 A = isContr→isProp hLevelSuc 1 A = isProp→isSet hLevelSuc (suc (suc n)) A h a b = hLevelSuc (suc n) (a ≡ b) (h a b) -hLevelAdd : (n : ℕ) (m : ℕ) (A : Type ℓ) → isOfHLevel n A → isOfHLevel (n + m) A -hLevelAdd n zero A h = transport (λ i → isOfHLevel (helper n i) A) h where - helper : (n : ℕ) → n ≡ n + zero - helper zero = refl - helper (suc n) = cong (λ x → (suc x)) (helper n) -hLevelAdd n (suc m) A h = transport (λ i → (isOfHLevel (helper n m i) A )) (hLevelSuc (n + m) A (hLevelAdd n m A h)) where - helper : (n m : ℕ) → (suc (n + m)) ≡ (n + (suc m)) - helper zero m = refl - helper (suc n) m = cong (λ x → (suc x)) (helper n m) - hLevelPath : (n : ℕ) → isOfHLevel n A → (x y : A) → isOfHLevel n (x ≡ y) hLevelPath 0 h x y = isContrPath h x y hLevelPath 1 h x y = hLevelSuc zero _ (isProp→isContr≡ h x y) diff --git a/Cubical/Foundations/Pointed/Base.agda b/Cubical/Foundations/Pointed/Base.agda index 2c2d7e6d4..d796dee28 100644 --- a/Cubical/Foundations/Pointed/Base.agda +++ b/Cubical/Foundations/Pointed/Base.agda @@ -13,6 +13,8 @@ pt : ∀ {ℓ} (A∙ : Pointed ℓ) → typ A∙ pt = snd {- Pointed functions -} -_→*_ : ∀{ℓ ℓ'} → (A : Pointed ℓ) → (B : Pointed ℓ') → Type (ℓ-max ℓ ℓ') +_→*_ : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → Type (ℓ-max ℓ ℓ') _→*_ A B = Σ (typ A → typ B) λ f → f (pt A) ≡ pt B +_→*_/_ : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → (A →* B) → Pointed (ℓ-max ℓ ℓ') +A →* B / f = (A →* B) , f diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 31086087a..b35b3af20 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -131,7 +131,6 @@ transport p a = transp (λ i → p i) i0 a transportRefl : (x : A) → transport refl x ≡ x transportRefl {A = A} x i = transp (λ _ → A) i x - -- We want B to be explicit in subst subst : (B : A → Type ℓ') (p : x ≡ y) → B x → B y subst B p pa = transport (λ i → B (p i)) pa diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index ebe21aa20..9c7393674 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -61,13 +61,13 @@ elimSetTrunc3 Bset g = elimSetTrunc2 (λ _ _ → hLevelPi 2 λ _ → Bset _ _ _) setTruncIsSet : isSet ∥ A ∥₀ setTruncIsSet a b p q = squash₀ a b p q -setId : isSet A → A ≡ ∥ A ∥₀ -setId {A = A} isset = isoToPath (iso (λ x → ∣ x ∣₀) - (elimSetTrunc {A = A} +setId : isSet A → ∥ A ∥₀ ≡ A +setId {A = A} isset = isoToPath (iso (elimSetTrunc {A = A} (λ _ → isset) (λ x → x)) - (λ a → idLemma a) - (λ _ → refl)) + (λ x → ∣ x ∣₀) + (λ b → refl) + λ b → idLemma b) where idLemma : ∀ (b : ∥ A ∥₀) → ∣ elimSetTrunc (λ x → isset) (λ x → x) b ∣₀ ≡ b idLemma b = elimSetTrunc {B = (λ x → ∣ elimSetTrunc (λ _ → isset) (λ x → x) x ∣₀ ≡ x)} diff --git a/Cubical/HITs/Truncation/Base.agda b/Cubical/HITs/Truncation/Base.agda index f69ddacb9..ef1993225 100644 --- a/Cubical/HITs/Truncation/Base.agda +++ b/Cubical/HITs/Truncation/Base.agda @@ -22,4 +22,3 @@ open import Cubical.HITs.Sn ∥_∥_ : ∀ {ℓ} → Type ℓ → ℕ₋₂ → Type ℓ ∥ A ∥ n = Null (S (1+ n)) A - diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index d720fe948..aa9a1f7f5 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -125,7 +125,6 @@ isOfHLevel∥∥ : (n : ℕ₋₂) → isOfHLevel (2+ n) (∥ A ∥ n) isOfHLevel∥∥ neg2 = hub ⊥-elim , λ _ → ≡hub ⊥-elim isOfHLevel∥∥ (suc n) = isSphereFilled→isOfHLevelSuc isSphereFilled∥∥ - -- isOfHLevel∥∥ n = isSnNull→isOfHLevel isNull-Null @@ -244,7 +243,8 @@ private hLevelP {n = n} {B = B} = ind2 {A = B} {n = (suc n)} {B = λ x y → isOfHLevel (2+ (suc n)) (P x y)} - (λ x y → hLevelAdd 1 (2+ n) (isOfHLevel (2+ suc n) (P x y)) (isPropIsOfHLevel (2+ suc n) (P x y)) ) + (λ x y → transport (λ i → isOfHLevel (+-comm (2+ n) 1 i) + (isOfHLevel (2+ suc n) (P x y))) (hLevelLift (2+ n) (isPropIsOfHLevel (2+ suc n) (P x y))) ) λ a b → ( hLevelSuc (2+ n) (P ∣ a ∣ ∣ b ∣) ) (isOfHLevel∥∥ {A = a ≡ b} n) @@ -378,13 +378,8 @@ private IsoFinal : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x y : ∥ B ∥ (suc n)) → Iso (x ≡ y) (P x y) IsoFinal x y = iso (encode-fun x y ) (decode-fun x y) (P-linv x y) (P-rinv x y) -IsoIdTrunc : {a b : A} (n : ℕ₋₂) → Iso (_≡_ {A = ∥ A ∥ (suc n)} ∣ a ∣ ∣ b ∣) (∥ a ≡ b ∥ n) -IsoIdTrunc {a = a} {b = b} n = IsoFinal {n = n} ∣ a ∣ ∣ b ∣ - -IsoΩ : {a : A} (n : ℕ₋₂) → Iso (_≡_ {A = ∥ A ∥ (suc n)} ∣ a ∣ ∣ a ∣) (∥ a ≡ a ∥ n) -IsoΩ {a = a} n = IsoIdTrunc {a = a} {b = a} n - - - - +PathIdTrunc : {a b : A} (n : ℕ₋₂) → (_≡_ {A = ∥ A ∥ (suc n)} ∣ a ∣ ∣ b ∣) ≡ (∥ a ≡ b ∥ n) +PathIdTrunc {a = a} {b = b} n = isoToPath (IsoFinal {n = n} ∣ a ∣ ∣ b ∣) +PathΩ : {a : A} (n : ℕ₋₂) → (_≡_ {A = ∥ A ∥ (suc n)} ∣ a ∣ ∣ a ∣) ≡ (∥ a ≡ a ∥ n) +PathΩ {a = a} n = PathIdTrunc {a = a} {b = a} n diff --git a/Cubical/ZCohomology/Base.agda b/Cubical/ZCohomology/Base.agda index 1c02e2057..1b9fc98ea 100644 --- a/Cubical/ZCohomology/Base.agda +++ b/Cubical/ZCohomology/Base.agda @@ -5,7 +5,7 @@ module Cubical.ZCohomology.Base where open import Cubical.Data.Int.Base open import Cubical.Data.Nat.Base open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Sigma.Base +open import Cubical.Data.Sigma open import Cubical.Foundations.Pointed.Base @@ -20,32 +20,26 @@ private ℓ : Level A : Type ℓ -{- -Definition of cohomology with cocefficients in the integers. For now, this definition uses ℕ₋₂, for instance defining H⁰ as -coHom neg2, H¹ as coHom (suc neg2), and so on. This has so far been easier to work with than a definition relying on ℕ→ℕ₋₂. --} - --- Cohomology --- {- Types Kₙ from Brunerie 2016 -} -coHomK : (n : ℕ₋₂) → Type₀ -coHomK neg2 = Int -coHomK (suc n) = ∥ S₊ (2+ suc n) ∥ (suc (suc (suc n))) +coHomK : (n : ℕ) → Type₀ +coHomK zero = Int +coHomK (suc n) = ∥ S₊ (suc n) ∥ (suc (ℕ→ℕ₋₂ n)) {- Cohomology -} -coHom : (n : ℕ₋₂) → Type ℓ → Type ℓ +coHom : (n : ℕ) → Type ℓ → Type ℓ coHom n A = ∥ (A → coHomK n) ∥₀ --- Reduced cohomology --- {- Pointed version of Kₙ -} -coHomK-ptd : (n : ℕ₋₂) → Pointed (ℓ-zero) -coHomK-ptd neg2 = coHomK neg2 , (pos 0) +coHomK-ptd : (n : ℕ) → Pointed (ℓ-zero) +coHomK-ptd zero = coHomK zero , (pos 0) coHomK-ptd (suc n) = (coHomK (suc n) , ∣ north ∣) {- Reduced cohomology -} -coHomRed : (n : ℕ₋₂) → (A : Pointed ℓ) → Type ℓ -coHomRed neg2 A = ∥ (A →* (coHomK-ptd neg2)) ∥₀ -coHomRed (suc n) A = ∥ (A →* (coHomK-ptd (suc n))) ∥₀ +coHomRed : (n : ℕ) → (A : Pointed ℓ) → Type ℓ +coHomRed n A = ∥ (A →* (coHomK-ptd n)) ∥₀ diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 139514488..04ded167c 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -5,11 +5,17 @@ module Cubical.ZCohomology.Properties where open import Cubical.ZCohomology.Base open import Cubical.HITs.Sn open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Everything open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws open import Cubical.Data.NatMinusTwo.Base renaming (-1+_ to -1+₋₂_ ; 1+_ to 1+₋₂_) +open import Cubical.Data.NatMinusOne.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma +open import Cubical.Data.Prod.Base open import Cubical.HITs.Susp open import Cubical.HITs.SetTruncation open import Cubical.HITs.Nullification @@ -20,21 +26,21 @@ open import Cubical.HITs.Truncation open import Cubical.HITs.Pushout open import Cubical.Data.Sum.Base open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Bool private variable - ℓ : Level + ℓ ℓ' : Level A : Type ℓ + B : Type ℓ' - - -{- Equivalence between cohomology and the reduced version -} -coHom↔coHomRed : (n : ℕ₋₂) → +{- Equivalence between cohomology of A and reduced cohomology of (A + 1) -} +coHomRed+1Equiv : (n : ℕ) → (A : Set ℓ) → (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) -coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ - module coHom↔coHomRed where - helplemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →* C))) - helplemma {C = C} = isoToPath (iso map1 +coHomRed+1Equiv zero A i = ∥ helpLemma {C = (Int , pos 0)} i ∥₀ + module coHomRed+1 where + helpLemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →* C))) + helpLemma {C = C} = isoToPath (iso map1 map2 (λ b → linvPf b) (λ _ → refl)) @@ -55,43 +61,4 @@ coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ helper (inl x) = refl helper (inr tt) = sym snd -coHom↔coHomRed (suc n) A i = ∥ coHom↔coHomRed.helplemma A i {C = ((coHomK (suc n)) , ∣ north ∣)} i ∥₀ - - - - -{- TODO: Prove Kₙ ≡ Ω Kₙ₊₁ -} - -invPathCancel : {a b : A} → (p : a ≡ b) → p ∙ (sym p) ≡ refl -invPathCancel {a = a} {b = b} p = J {x = a} (λ y p → p ∙ (sym p) ≡ refl ) (sym (lUnit refl)) p - -φ : (a : A) → A → (north {A = A} ≡ north {A = A}) -φ x = (λ a → ((merid a) ∙ sym ((merid x)))) - -φ* : (A : Pointed ℓ) → A →* Ω ((Susp (typ A)) , north) -φ* A = (φ (pt A)) , invPathCancel (merid (pt A)) - -σ : (n : ℕ₋₂) → (typ (coHomK-ptd n)) → (typ (Ω (coHomK-ptd (suc n)))) -σ neg2 k = looper k ( cong {B = λ _ → (∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1)} - (λ x → ∣ x ∣) - ((merid north) ∙ (sym (merid south))) ) - where - looper : {a : A} -- defining loopⁿ - (n : Int) → - (a ≡ a) → - (a ≡ a) - looper {A = A} {a = a} (pos zero) p = refl - looper {A = A} {a = a} (pos (suc n)) p = (looper (pos n) p) ∙ p - looper {A = A} {a = a} (negsuc zero) p = sym p - looper {A = A} {a = a} (negsuc (suc n)) p = (sym p) ∙ (looper (negsuc n) p) -σ (suc n) x = rec {n = suc (suc (suc n))} - {B = (typ (Ω (coHomK-ptd (suc (suc n)))))} - (isOfHLevel∥∥ {A = S₊ (2+ suc (suc n))} (suc (suc (suc (suc n)))) ∣ north ∣ ∣ north ∣) - (λ y → cong {B = λ _ → Null (S (1+₋₂ (suc (suc (suc (suc n)))))) (S₊ (2+ (suc (suc n))))} - (λ z → ∣ z ∣) - (φ north y)) - x - - - - +coHomRed+1Equiv (suc n) A i = ∥ coHomRed+1.helpLemma A i {C = ((coHomK (suc n)) , ∣ north ∣)} i ∥₀ diff --git a/Cubical/ZCohomology/S1/S1.agda b/Cubical/ZCohomology/S1/S1.agda index 311fc0590..0516dbdce 100644 --- a/Cubical/ZCohomology/S1/S1.agda +++ b/Cubical/ZCohomology/S1/S1.agda @@ -3,36 +3,25 @@ module Cubical.ZCohomology.S1.S1 where open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.Properties -open import Cubical.Data.Sigma open import Cubical.HITs.S1 -open import Cubical.HITs.Sn -open import Cubical.Data.Empty.Base -open import Cubical.Data.Unit.Base -open import Cubical.Foundations.Everything -open import Cubical.Data.NatMinusTwo.Base renaming (-1+_ to -1+₋₂_ ; 1+_ to 1+₋₂_) -open import Cubical.Foundations.Equiv -open import Cubical.Data.Prod -open import Cubical.HITs.Susp +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism open import Cubical.HITs.SetTruncation open import Cubical.HITs.Nullification open import Cubical.Data.Int open import Cubical.Data.Nat open import Cubical.HITs.Truncation -open import Cubical.HITs.Pushout -open import Cubical.Data.Sum.Base -open import Cubical.Data.Unit.Base -open import Cubical.Data.HomotopyGroup + private variable ℓ : Level A : Type ℓ - ---- H⁰(S¹) ---- -coHom0-S1 : (coHom neg2 S¹) ≡ Int -coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ sym (setId isSetInt) +coHom0-S1 : (coHom zero S¹) ≡ Int +coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ (setId isSetInt) where helpLemma : (S¹ → Int) ≡ Int helpLemma = isoToPath (iso fun @@ -60,6 +49,3 @@ coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ sym (setId isSetInt) ------------------------- {- TODO : give Hᵏ(S¹) for all k -} - - - From 95d1c5bca2babd2b9c235a33e03604811f78fd66 Mon Sep 17 00:00:00 2001 From: Axel Date: Mon, 24 Feb 2020 16:24:55 +0100 Subject: [PATCH 06/30] fixes1 --- Cubical/AndersFile.agda | 41 ++ Cubical/Data/NatMinusTwo/Base.agda | 17 + Cubical/ZCohomology/FunLoopSpace.agda | 205 +++++++ Cubical/ZCohomology/FunLoopSpace2.agda | 34 ++ Cubical/ZCohomology/PropertiesTrash.agda | 660 +++++++++++++++++++++++ 5 files changed, 957 insertions(+) create mode 100644 Cubical/AndersFile.agda create mode 100644 Cubical/ZCohomology/FunLoopSpace.agda create mode 100644 Cubical/ZCohomology/FunLoopSpace2.agda create mode 100644 Cubical/ZCohomology/PropertiesTrash.agda diff --git a/Cubical/AndersFile.agda b/Cubical/AndersFile.agda new file mode 100644 index 000000000..59285462e --- /dev/null +++ b/Cubical/AndersFile.agda @@ -0,0 +1,41 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.AndersFile where + + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.Nat +open import Cubical.Data.HomotopyGroup + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + + +FunLR : (n : ℕ) {f : A → B} → + ((typ ((Ω^ (suc n)) ((A → B) , f))) → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) +LR-id : (n : ℕ) {f : A → B} → FunLR n {f} refl ≡ (λ x → refl) +FunLR n {f} = {!!} +LR-id n {f} = {!!} + + +FunRL : (n : ℕ) {f : A → B} → + ((x : A) → (typ ((Ω^ (suc n)) (B , f x)))) → + ((typ ((Ω^ (suc n)) ((A → B) , f)))) +RL-id : (n : ℕ) {f : A → B} → + FunRL n (λ x → pt (((Ω^ (suc n)) (B , f x)))) ≡ pt ((Ω^ (suc n)) ((A → B) , f)) +FunRL n {f} = {!!} +RL-id n {f} = {!!} + + +wantToShow : (n : ℕ) (f : A → B) → + ((x : A) → (typ ((Ω^ n) (B , f x)))) ≡ (typ ((Ω^ n) ((A → B) , f))) +wantToShow n f = {!!} + + diff --git a/Cubical/Data/NatMinusTwo/Base.agda b/Cubical/Data/NatMinusTwo/Base.agda index 35e5b3ac8..4ca66c900 100644 --- a/Cubical/Data/NatMinusTwo/Base.agda +++ b/Cubical/Data/NatMinusTwo/Base.agda @@ -3,6 +3,7 @@ module Cubical.Data.NatMinusTwo.Base where open import Cubical.Core.Primitives open import Cubical.Data.Nat +open import Cubical.Data.Bool open import Cubical.Data.Empty import Cubical.Data.NatMinusOne as ℕ₋₁ @@ -32,6 +33,22 @@ data ℕ₋₂ : Set where ℕ→ℕ₋₂ : ℕ → ℕ₋₂ ℕ→ℕ₋₂ n = ℕ₋₁→ℕ₋₂ (ℕ→ℕ₋₁ n) +_<₋₂_ : ℕ₋₂ → ℕ₋₂ → Bool +neg2 <₋₂ m = true +suc n <₋₂ neg2 = false +suc n <₋₂ suc m = n <₋₂ m + +_+₋₂_ : ℕ₋₂ → ℕ₋₂ → ℕ₋₂ +neg2 +₋₂ m = m +suc n +₋₂ neg2 = suc n +suc n +₋₂ suc m = suc (suc ( n +₋₂ m )) + +_-₋₂_ : ℕ₋₂ → ℕ₋₂ → ℕ₋₂ +neg2 -₋₂ m = neg2 +suc n -₋₂ neg2 = suc n +suc n -₋₂ suc m = n -₋₂ m + + -- Natural number and negative integer literals for ℕ₋₂ instance diff --git a/Cubical/ZCohomology/FunLoopSpace.agda b/Cubical/ZCohomology/FunLoopSpace.agda new file mode 100644 index 000000000..215e12a0b --- /dev/null +++ b/Cubical/ZCohomology/FunLoopSpace.agda @@ -0,0 +1,205 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.FunLoopSpace where + + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.FunExtEquiv +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.Nat +open import Cubical.Data.HomotopyGroup + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + + + + + +{- Beware: Use of J -} +cancelReflMid : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) (q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q +cancelReflMid {ℓ = ℓ}{A = A} {a = a} {b = b} p q = J {ℓ} {A} {a} {ℓ} (λ b p → ((q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q)) (λ r → sym (lUnit (refl ∙ r ))) {b} p q + +--- + + +FunLR : (n : ℕ) {f : A → B} → + ((typ ((Ω^ (suc n)) ((A → B) , f))) → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) +LR-id : (n : ℕ) {f : A → B} → FunLR n {f} refl ≡ (λ x → refl) +FunLR zero {f} = funExt⁻ +FunLR (suc n) {f} p x = cong (λ y → y x) (sym (LR-id n)) ∙ (funExt⁻ (cong (FunLR n) p ) x) ∙ cong (λ y → y x) (LR-id n) +LR-id zero {f} = refl +LR-id (suc n) {f} = funExt λ y → (cancelReflMid (λ i → (LR-id n) (~ i) y) + (λ i → (LR-id n {f}) i y) ∙ + rCancel (λ i → (LR-id n {f}) (~ i) y)) + +{- +lemma2 : (n : ℕ) {f : A → B} {x : A} → PathP (λ i → LR-id n {f} i x ≡ λ _ → pt ((Ω^ n) (B , f x))) (cong (λ y → y x) ((LR-id n {f}))) refl +lemma2 zero = refl +lemma2 (suc n) {f} {x} j = hcomp {!!} {!(cong (λ y → y x) (funExt (λ y → cancelReflMid (λ i → LR-id n {f} (~ (i)) y) (λ i → LR-id n {f} (i ∨ j) y) ∙ rCancel (λ i → LR-id n ((~ i)) y)))) !} +-} + + +FunRL : (n : ℕ) {f : A → B} → + ((x : A) → (typ ((Ω^ (suc n)) (B , f x)))) → + ((typ ((Ω^ (suc n)) ((A → B) , f)))) +RL-id : (n : ℕ) {f : A → B} → + FunRL n (λ x → pt (((Ω^ (suc n)) (B , f x)))) ≡ pt ((Ω^ (suc n)) ((A → B) , f)) +FunRL zero = funExt +FunRL (suc n) {f} = λ g → (sym (RL-id n {f}) ∙ cong (FunRL n {f}) (funExt g) ∙ (RL-id n {f})) +RL-id zero = refl +RL-id (suc n) {f} = cancelReflMid (λ i → (RL-id n {f}) (~ i)) (RL-id n) ∙ + rCancel (λ i → (RL-id n) (~ i)) + + + +---- + + + +wantToShow : (n : ℕ) (f : A → B) → + ((x : A) → (typ ((Ω^ n) (B , f x)))) ≡ (typ ((Ω^ n) ((A → B) , f))) +wantToShow zero f = refl +wantToShow (suc zero) f = ua funExtEquiv +wantToShow (suc (suc n)) f = isoToPath (iso (FunRL (suc n)) (FunLR (suc n)) (λ b → {!!}) {!!}) + where + hauptLemma : (n : ℕ) (f : A → B) (b : typ (Ω ((Ω^ n) ((A → B) , f)))) → FunRL n (FunLR n b ) ≡ b + hauptLemma zero f b = refl + hauptLemma (suc n) f b = {!!} + where + + oi : {!λ j → ((λ i → RL-id n {f} (~ i)) ∙ + (λ i → (hauptLemma n f (b i) (i))) ∙ + (λ i → RL-id n {f} i))!} + oi = {!!} + + + firstTransp : PathP (λ j → RL-id n (~ j) ≡ RL-id n (~ j)) + (FunRL (suc n) (FunLR (suc n) b )) + ((λ i → FunRL n {f} (funExt (λ x → (λ i₁ → LR-id n {f} (~ i₁) x) ∙ (λ i₁ → FunLR n {f} (b i₁) x) ∙ (λ i₁ → LR-id n {f} i₁ x)) i))) + firstTransp j = hcomp (λ k → λ{(j = i0) → FunRL (suc n) (FunLR (suc n) b ) ; + (j = i1) → (sym (lUnit ((λ i → FunRL n (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ + (λ i₁ → FunLR n {f} (b i₁) x) ∙ (λ i₁ → LR-id n {f} i₁ x)) i)) ∙ + (λ i → RL-id n {f} (i0)))) ∙ + sym (rUnit (λ i → FunRL n (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ + (λ i₁ → FunLR n {f} (b i₁) x) ∙ + (λ i₁ → LR-id n {f} i₁ x)) i)))) k }) + ((λ i → RL-id n {f} (~ i ∧ ( ~ j))) ∙ + (λ i → FunRL n (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ + (λ i₁ → FunLR n {f} (b i₁) x) ∙ + (λ i₁ → LR-id n {f} i₁ x)) i)) ∙ + (λ i → RL-id n {f} (i ∧ (~ j)))) + + SNDtransp : PathP (λ j → (FunRL n {f} (λ x → LR-id n {f} (~ j) x)) ≡ (FunRL n {f} (λ x → LR-id n {f} (~ j) x))) + ((λ i → FunRL n {f} (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ + (λ i₁ → FunLR n {f} (b i₁) x) ∙ + (λ i₁ → LR-id n {f} i₁ x)) i))) λ i → FunRL n {f} (FunLR n {f} (b i)) + SNDtransp j = hcomp (λ k → λ{ (j = i0) → + ((λ i → FunRL n (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ + (λ i₁ → FunLR n {f} (b i₁) x) ∙ + (λ i₁ → LR-id n {f} i₁ x)) i))) + ; (j = i1) → {!(FunRL n {f} (λ x → LR-id n {f} (~ i1) x))!}}) + (λ i → FunRL n {f} (funExt (λ x → (λ i₁ → LR-id n {f} (~ i₁ ∧ (~ j)) x) ∙ + (λ i₁ → FunLR n {f} (b i₁) x) ∙ + (λ i₁ → LR-id n {f} (i₁ ∧ (~ j)) x)) i)) + + 3transp : PathP (λ j → (hauptLemma n f refl j) ≡ (hauptLemma n f refl j)) (λ i → FunRL n {f} (FunLR n {f} (b i))) b + 3transp j = (λ i → (hauptLemma n f (b i) j)) + + test : (λ j₂ → RL-id n {f} (~ j₂) ≡ RL-id n {f} (~ j₂)) ∙ (λ i → FunRL n {f} (LR-id n {f} (~ i)) ≡ FunRL n {f} (LR-id n {f} (~ i))) ∙ (λ i → hauptLemma n f refl i ≡ hauptLemma n f refl i) ≡ refl + test j = {!((λ j₂ → RL-id n {f} ((~ j₂) ∨ j) ≡ RL-id n {f} ((~ j₂) ∨ j)) ∙ (λ i → FunRL n {f} (LR-id n {f} (?)) ≡ FunRL n {f} (LR-id n {f} (?))))!} + where + + + + postulate + wowiecomp : (i : I) → FunRL (n) (funExt (λ x → (λ i₁ → LR-id (n) (~ i₁) x) ∙ + (λ i₁ → FunLR (n) (b i₁) x) ∙ + (λ i₁ → LR-id (n) {f} i₁ x)) i) ≡ b i + + + + needToShow : ((λ j₁ → ((λ j₂ → RL-id n {f} (~ j₂) ≡ RL-id n {f} (~ j₂)) ∙ (λ i → FunRL n {f} (LR-id n (~ i)) ≡ FunRL n {f} (LR-id n {f} (~ i)))) + j₁) ∙ (λ i → hauptLemma n f refl i ≡ hauptLemma n f refl i)) ≡ refl + needToShow j k = {!((λ j₁ → ((λ j₂ → RL-id n {f} (~ j₂) ≡ RL-id n {f} (~ j₂)) ∙ (λ i → FunRL n {f} (LR-id n (~ i)) ≡ FunRL n {f} (LR-id n {f} (~ i))))!} + + +{- +where + + linvLemma : (n : ℕ) {f : A → B} {b : typ ((Ω^ (suc (suc (suc n)))) ((A → B) , f))} → + (λ i → RL-id (suc n) (~ i)) ∙ + (λ i → FunRL (suc n) (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x) ∙ + (λ i₁ → FunLR (suc n) (b i₁) x) ∙ + (λ i₁ → LR-id (suc n) {f} i₁ x)) i)) ∙ + RL-id (suc n) {f} + ≡ b + linvLemma n {f} {b} j = {!hcomp ? ( (λ i → RL-id (suc n) (~ i ∧ j)) ∙ + (λ i → FunRL (suc n) (funExt (λ x → (λ i₁ → LR-id (suc n) ((~ i₁)) x) ∙ + (λ i₁ → FunLR (suc n) (b i₁) x) ∙ + (λ i₁ → LR-id (suc n) {f} (i₁ ) x)) (i))) ∙ + (λ i → (RL-id (suc n) {f} (i ∧ j)))) !} + where + cunt : (i : I) → (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x) ∙ (λ i₁ → FunLR (suc n) {f} (b i₁) x) ∙ (λ i₁ → LR-id (suc n) {f} i₁ x )) i ) + ≡ (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x)) ∙ funExt (λ x → (λ i₁ → FunLR (suc n) {f} (b i₁) x)) ∙ funExt (λ x → (λ i₁ → LR-id (suc n) {f} i₁ x ))) i + cunt i = refl + + cunt2 : (i : I) → (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x) ∙ (λ i₁ → FunLR (suc n) {f} (b i₁) x) ∙ (λ i₁ → LR-id (suc n) {f} i₁ x )) i ) + ≡ (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x)) ∙ (λ j → FunLR (suc n) {f} (b j)) ∙ funExt (λ x → (λ i₁ → LR-id (suc n) {f} i₁ x ))) i + cunt2 i = refl +{- + cunt3 : (i : I) → FunRL n (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ (λ i₁ → FunLR n (b i₁) x) ∙ (λ i₁ → LR-id n {f} i₁ x)) i) ≡ FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i) + cunt3 i = refl + + cunt4 : (i : I) → FunRL (suc n) (((λ i₁ → LR-id (suc n) {f} (~ i₁)) ∙ (λ j → FunLR (suc n) {f} (b j)) ∙ (λ i₁ → LR-id (suc n) {f} i₁)) i) ≡ FunRL n {f} ( LR-id n {f} (~ i)) ∙ FunRL n (FunLR n {f} (b i)) ∙ FunRL n {f} ( LR-id n {f} i) + cunt4 i = {!!} + + +{- + main2 : (λ i → RL-id n (~ i)) ∙ (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) ∙ (λ i → RL-id n i) ≡ {!!} + main2 j i = {!FunRL n (((λ i₁ → LR-id n {f} (~ i₁ ∧ i0)) ∙ (λ k → FunLR n {f} (b (k))) ∙ (λ i₁ → LR-id n {f} (i₁ ∧ i0))) i0)!} + + main3 : PathP (λ j → (RL-id n {f} j ≡ RL-id n j)) (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) ((λ i → RL-id n (~ i)) ∙ (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) ∙ (λ i → RL-id n {f} i)) + main3 j = hcomp (λ k → λ{(j = i0) → {!((λ i → RL-id n ((~ i) ∧ j))!} ; (j = i1) → {!!}}) ((λ i → RL-id n ((~ i) ∧ j)) ∙ (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) ∙ (λ i → RL-id n (i ∧ j))) + + main4 : PathP (λ j → FunRL n (LR-id n j) ≡ FunRL n (LR-id n j)) {!LR-id n {f} i0!} (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) + main4 j = hcomp (λ k → {! LR-id n {f} (i0)!}) (λ i → FunRL n {f} (((λ i₁ → LR-id n {f} (~ i₁ ∧ j)) ∙ (λ k → FunLR n {f} (b (k))) ∙ (λ i₁ → LR-id n {f} (i₁ ∧ j))) i)) +-} + + wowie : (i : I) → ( FunRL (suc n) (((λ i₁ → LR-id (suc n) {f} (~ i₁)) ∙ (λ j → FunLR (suc n) {f} (b j)) ∙ (λ i₁ → LR-id (suc n) {f} i₁)) i)) ≡ (FunRL (suc n) (((λ i₁ → (FunLR (suc n) refl)) ∙ (λ j → FunLR (suc n) {f} (b j)) ∙ (λ i₁ → (FunLR (suc n) refl))) i)) + wowie i = sym (λ j → FunRL (suc n) (((λ i₁ → LR-id (suc n) {f} (~ i₁ ∧ j)) ∙ (λ k → FunLR (suc n) {f} (b (k))) ∙ (λ i₁ → LR-id (suc n) {f} (i₁ ∧ j))) i)) + + wowie2 : (i : I) → (FunRL (suc n) (((λ i₁ → (FunLR (suc n) refl)) ∙ (λ j → FunLR (suc n) {f} (b j)) ∙ (λ i₁ → (FunLR (suc n) refl))) i)) ≡ (FunRL (suc n) ((( FunLR (suc n) {f} (b i))) )) + wowie2 i = (λ j → FunRL (suc n) {f} ((sym (lUnit ((λ j → FunLR (suc n) {f} (b j)) ∙ refl)) j) i) ) ∙ ((λ j → FunRL (suc n) {f} ((sym (rUnit ((λ j → FunLR (suc n) {f} (b j)))) j) i))) + + main4 : PathP (λ j → FunRL n (LR-id n j) ≡ FunRL n (LR-id n j)) {!LR-id n {f} i0!} (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) + main4 j = hcomp (λ k → {! LR-id n {f} (i0)!}) (λ i → FunRL n {f} (((λ i₁ → LR-id n {f} (~ i₁ ∧ j)) ∙ (λ k → FunLR n {f} (b (k))) ∙ (λ i₁ → LR-id n {f} (i₁ ∧ j))) i)) + + postulate + fixLater : (i : I) → (FunRL (suc n) ((( FunLR (suc n) {f} (b i))) )) ≡ b i + + wowie3 : + (λ i → FunRL (suc n) (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x) ∙ + (λ i₁ → FunLR (suc n) (b i₁) x) ∙ + (λ i₁ → LR-id (suc n) {f} i₁ x)) i)) + ≡ + {!b!} + wowie3 j = {!FunRL!} + + wowiecomp : (i : I) → FunRL (suc n) (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x) ∙ + (λ i₁ → FunLR (suc n) (b i₁) x) ∙ + (λ i₁ → LR-id (suc n) {f} i₁ x)) i) ≡ b i + wowiecomp i = wowie i ∙ wowie2 i ∙ fixLater {!wowiecomp i0!} + + +-} + +-} diff --git a/Cubical/ZCohomology/FunLoopSpace2.agda b/Cubical/ZCohomology/FunLoopSpace2.agda new file mode 100644 index 000000000..ae60e32b6 --- /dev/null +++ b/Cubical/ZCohomology/FunLoopSpace2.agda @@ -0,0 +1,34 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.FunLoopSpace2 where + + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.FunExtEquiv +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.Nat +open import Cubical.Data.HomotopyGroup +open import Cubical.HITs.Nullification +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Truncation +open import Cubical.Data.NatMinusTwo + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + + + + + +{- Beware: Use of J -} +cancelReflMid : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) (q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q +cancelReflMid {ℓ = ℓ}{A = A} {a = a} {b = b} p q = J {ℓ} {A} {a} {ℓ} (λ b p → ((q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q)) (λ r → sym (lUnit (refl ∙ r ))) {b} p q + diff --git a/Cubical/ZCohomology/PropertiesTrash.agda b/Cubical/ZCohomology/PropertiesTrash.agda new file mode 100644 index 000000000..5a9005989 --- /dev/null +++ b/Cubical/ZCohomology/PropertiesTrash.agda @@ -0,0 +1,660 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.Properties where + + +open import Cubical.ZCohomology.Base +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.NatMinusTwo.Base renaming (-1+_ to -1+₋₂_ ; 1+_ to 1+₋₂_) +open import Cubical.Data.NatMinusOne.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma +open import Cubical.Data.Prod.Base +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation +open import Cubical.HITs.Nullification +open import Cubical.Data.Int +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Bool +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + + +{- Equivalence between cohomology and the reduced version -} +coHom↔coHomRed : (n : ℕ₋₂) → + (A : Set ℓ) → + (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) +coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ + module coHom↔coHomRed where + helplemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →* C))) + helplemma {C = C} = isoToPath (iso map1 + map2 + (λ b → linvPf b) + (λ _ → refl)) + where + map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →* C)) + map1 f = map1' , refl module helpmap where + map1' : A ⊎ Unit → fst C + map1' (inl x) = f x + map1' (inr x) = pt C + + map2 : ((((A ⊎ Unit) , inr (tt)) →* C)) → (A → typ C) + map2 (g , pf) x = g (inl x) + + linvPf : (b :((((A ⊎ Unit) , inr (tt)) →* C))) → map1 (map2 b) ≡ b + linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) + where + helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x + helper (inl x) = refl + helper (inr tt) = sym snd + +coHom↔coHomRed (suc n) A i = ∥ coHom↔coHomRed.helplemma A i {C = ((coHomK (suc n)) , ∣ north ∣)} i ∥₀ + + + +{- Test -} +SigmaEqHelper : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) (b : B) (p1 p2 : Σ A (λ a → (f a) ≡ b)) → (fst p1 ≡ fst p2) → p1 ≡ p2 +SigmaEqHelper {ℓ = ℓ} {ℓ' = ℓ'} {A = A} {B = B} f b p1 p2 path = J {ℓ'} {B} {f (fst p1)} (λ b pf → ( (pair2 : Σ A (λ a → (f a) ≡ b)) → (fst p1) ≡ (fst pair2) → ((fst p1 , pf) ≡ pair2) ) ) (λ pr id → ΣPathP (id , {!transport (λ i → f (id i) ≡ ((snd pr) i))!})) {y = b} (snd p1) (p2) path {- J {ℓ} {A} {fst p1} {ℓ-max ℓ ℓ'} (λ y pf → ( (s : (f y ≡ b)) → (p1 ≡ (y , s)) )) + (λ s → ΣPathP (refl , {!J {ℓ'} {B} {x = (f (fst p1))} {ℓ'} (λ c → )!})) + {fst p2} path (snd p2) -} + + +{- TODO: Prove Kₙ ≡ Ω Kₙ₊₁ -} + +invPathCancel : {a b : A} → (p : a ≡ b) → p ∙ (sym p) ≡ refl +invPathCancel {a = a} {b = b} p = J {x = a} (λ y p → p ∙ (sym p) ≡ refl ) (sym (lUnit refl)) p +invPathCancel2 : {a b : A} → (p : a ≡ b) → (sym p) ∙ p ≡ refl +invPathCancel2 {a = a} {b = b} p = J {x = a} (λ y p → (sym p) ∙ p ≡ refl ) (sym (lUnit refl)) p + +cancelReflMid : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) (q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q +cancelReflMid {ℓ = ℓ}{A = A} {a = a} {b = b} p q = J {ℓ} {A} {a} {ℓ} (λ b p → ((q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q)) (λ r → sym (lUnit (refl ∙ r ))) {b} p q + +φ' : (a : A) → A → typ (Ω (Susp A , north )) +φ' x = (λ a → ((merid a) ∙ sym ((merid x)))) + +φ'' : (b : Susp A) → b ≡ south → A → north ≡ b +φ'' b p a = (merid a) ∙ (sym p) + +--φ''' : (a : A) (b : Susp A) → A → north ≡ b +--φ''' (a : A) b p = (merid a) ∙ (sym p) + + + + +looper : {a : A} -- defining loopⁿ + (n : Int) → + (a ≡ a) → + (a ≡ a) +looper {A = A} {a = a} (pos zero) p = refl +looper {A = A} {a = a} (pos (suc n)) p = (looper (pos n) p) ∙ p +looper {A = A} {a = a} (negsuc zero) p = sym p +looper {A = A} {a = a} (negsuc (suc n)) p = (sym p) ∙ (looper (negsuc n) p) + + + +φ : (a : A) → A → (north {A = A} ≡ north {A = A}) +φ x = (λ a → ((merid a) ∙ sym ((merid x)))) + +φ* : (A : Pointed ℓ) → A →* Ω ((Susp (typ A)) , north) +φ* A = (φ (pt A)) , invPathCancel (merid (pt A)) + +σ : (n : ℕ₋₂) → (typ (coHomK-ptd n)) → (typ (Ω (coHomK-ptd (suc n)))) +σ neg2 k = looper k ( cong {B = λ _ → (∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1)} + (λ x → ∣ x ∣) + ((merid north) ∙ (sym (merid south))) ) +σ (suc n) x = rec {n = suc (suc (suc n))} + {B = (typ (Ω (coHomK-ptd (suc (suc n)))))} + (isOfHLevel∥∥ {A = S₊ (2+ suc (suc n))} (suc (suc (suc (suc n)))) ∣ north ∣ ∣ north ∣) + (λ y → cong {B = λ _ → Null (S (1+₋₂ (suc (suc (suc (suc n)))))) (S₊ (2+ (suc (suc n))))} + (λ z → ∣ z ∣) + (φ north y)) + x + +testing : (n : ℕ) → (a : typ (Ω (S₊ (suc n) , north))) → isContr (∥ (Σ (S₊ n) λ k → (φ' {A = (S₊ n)} north k ) ≡ a) ∥ (ℕ→ℕ₋₂ n) ) +testing zero a = {!!} +testing (suc zero) a = {!!} +testing (suc (suc n)) a = ∣ {!!} , {!a!} ∣ , λ y → {!!} + + +-- φ : (a : A) → A → (north {A = A} ≡ north {A = A}) +-- φ x = (λ a → ((merid a) ∙ sym ((merid x)))) + +isConnectedT : (n : ℕ₋₂) (A : Type ℓ) → Type ℓ +isConnectedT n A = isContr (∥ A ∥ n) + +isConnectedF : ∀{ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ₋₂) → (A → B) → Type ((ℓ-max ℓ ℓ')) +isConnectedF {A = A} {B = B} n f = ((b : B) → isContr (∥ fiber f b ∥ n)) + +inducedFun : ∀{ℓ''} {k : ℕ₋₂} + (f : A → B) → + (P : (B → HLevel ℓ'' (2+ k))) → + (((b : B) → fst (P b)) → (a : A) → fst (P (f a)) ) +inducedFun f P = λ g a → g (f a) + +Lemma861 : ∀{ℓ''} (n k : ℕ₋₂) → + Σ ℕ₋₂ (λ num → (suc (suc n)) +₋₂ num ≡ (suc (suc k)) ) → + (f : A → B) → + (isConnectedF (suc (suc n)) f) → + (P : (B → HLevel ℓ'' (2+ (suc (suc k))))) → + isConnectedF (((suc (suc n)) -₋₂ (suc (suc k))) -₋₂ (suc (suc neg2))) (inducedFun {k = (suc (suc k))} f P) +Lemma861 n k (neg2 , pf) f conf P b = {!fiber!} +Lemma861 {A = A} {B = B} n k (suc j , pf) f conf P g = {!!} where + helper : (l : (x : A) → (fst (P (f x)) )) → + (fiber (inducedFun {k = (suc (suc k))} f P) l) ≡ Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a) + helper l = isoToPath (iso (fun) finv (λ b → refl) λ b → refl) where + fun : (fiber (inducedFun {k = (suc (suc k))} f P) l) → Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a) + fun (fs , sn) = fs , λ a → cong (λ x → x a) sn + + finv : (Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → (fiber (inducedFun {k = (suc (suc k))} f P) l) + finv (fs , sn) = fs , (funExt sn) + + helper'3 : (l : (x : A) → (fst (P (f x)) )) + (gp hq : Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → + (gp ≡ hq) + ≡ Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) + helper'3 l gp hq = isoToPath (iso fun finv (λ b i → ({!!})) {!!}) where + fun : (gp ≡ hq) → Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) + fun path = (λ x i → fst (path i) x ) , funExt λ x → pathFinal {x} where + pathHelper : {x : A} → PathP (λ i → fst gp (f x) ≡ fst (path i) (f x)) refl (snd gp x ∙ sym (snd hq x)) + pathHelper {x = x} i = hcomp (λ j → λ{ (i = i0) → rCancel (snd gp x) j ; (i = i1) → (snd gp x ∙ sym (snd hq x)) } ) (snd gp x ∙ sym (snd (path i) x)) + + pathHelper2 : {x : A} → PathP (λ i → fst gp (f x) ≡ fst (path (~ i)) (f x)) (λ i → fst (path i) (f x)) refl + pathHelper2 {x = x} i = λ i₁ → fst (path (i₁ ∧ (~ i))) (f x) + + pathFinal : {x : A} → (λ i → fst (path i) (f x)) ≡ snd gp x ∙ sym (snd hq x) + pathFinal {x = x} = transport (λ i → PathP (λ j₁ → (lCancel (λ i → fst gp (f x) ≡ fst (path i) (f x))) i j₁) + (λ i → fst (path i) (f x)) + (snd gp x ∙ sym (snd hq x))) + (compPathP pathHelper2 pathHelper) + + finv : Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) → (gp ≡ hq) + finv (fs , sn) = ΣPathP {x = gp} {y = hq} (funExt (λ b → fs b) , {! !}) -- (λ b → (fs b) i) , λ a → hcomp (λ k → λ{(i = i0) → (snd gp a) ; (i = i1) → {!!}}) (transp (λ j → fs (f a) (i ∧ j) ≡ l a) (~ i) ((snd gp) a)) + where + potTerm : {a : A} → PathP (λ i → (a : A) → funExt (λ b → fs b) i (f a) ≡ l a) (snd gp) (snd hq) + potTerm {a = a} j = λ a → hcomp {!(transp (λ i → funExt (λ b → fs b) (i) (f a) ≡ l a) i0 ((snd gp a))) !} (transp (λ i → funExt (λ b → fs b) (i ∧ j) (f a) ≡ l a) i0 ((snd gp a))) + + helper2 : {a : A} → (fs (f a)) ∙ (snd hq a) ≡ snd gp a + helper2 {a = a} = (cong (λ x → x ∙ (snd hq a)) (cong (λ x → x a) sn)) ∙ sym (assoc (snd gp a) (sym (snd hq a)) (snd hq a)) ∙ (λ j → (snd gp a) ∙ (lCancel (snd hq a)) j) ∙ sym (rUnit (snd gp a)) + + + -- fst gp (f x) ≡ fst hq (f x) + -- fst gp (f x) ≡ fst gp (f x) + + + + + helper2 : (l : (x : A) → (fst (P (f x)) )) + (gp hq : Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → + Σ (fst gp ≡ fst hq) (λ a≡ → PathP (λ i → (a : A) → a≡ i (f a) ≡ l a) (snd gp) (snd hq)) + ≡ Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) + helper2 l gp hq = isoToPath (iso fun finv (λ b i → {!fst gp x!} , {!!}) {!!}) where + + + finv : Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) (λ r → (inducedFun {k = (suc (suc k))} f (λ b → (fst gp b ≡ fst hq b) , hLevelSuc (suc (suc (1+ (1+₋₂ k)))) (fst (P b)) (snd (P b)) (fst gp b) (fst hq b) ) r ≡ (λ (x : A) → (snd gp x) ∙ sym (snd hq x)))) → Σ (fst gp ≡ fst hq) (λ a≡ → PathP (λ i → (a : A) → a≡ i (f a) ≡ l a) (snd gp) (snd hq)) + finv (fs , sn) = {!!} {- (funExt fs) , {!(snd hq) !}-} where + propLemma : (a : A) → isProp (fst hq (f a) ≡ l a) + propLemma a f g = {!!} + + + fun : Σ (fst gp ≡ fst hq) (λ a≡ → PathP (λ i → (a : A) → a≡ i (f a) ≡ l a) (snd gp) (snd hq)) → Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) (λ r → (inducedFun {k = (suc (suc k))} f (λ b → (fst gp b ≡ fst hq b) , hLevelSuc (suc (suc (1+ (1+₋₂ k)))) (fst (P b)) (snd (P b)) (fst gp b) (fst hq b) ) r ≡ (λ (x : A) → (snd gp x) ∙ sym (snd hq x)))) + fun (fs , sn) = (funExt⁻ fs) , (funExt (λ x → test5)) where + + test5 : {x : A} → PathP (λ _ → fst gp (f x) ≡ fst hq (f x) ) (λ i → (fs i) (f x)) (snd gp x ∙ sym (snd hq x)) + test5 {x = x} i j = hcomp ((λ k → λ{ (i = i0) → fs j (f x) ; (i = i1) → test3 {x = x} k j ; (j = i0) → fst gp (f x) ; (j = i1) → fs (~ i ∨ k) (f x) })) (test4 {x = x} i j) where + test3 : {x : A} → PathP (λ i → fst gp (f x) ≡ fs i (f x)) (snd gp x ∙ sym (snd gp x)) (snd gp x ∙ sym (snd hq x)) + test3 {x = x} i = (snd gp x) ∙ (sym (sn i x)) + + test4 : {x : A} → PathP (λ i → fst gp (f x) ≡ fs (~ i) (f x)) (λ i → (fs i) (f x)) (snd gp x ∙ (sym (snd gp x))) + test4 {x = x} i = hcomp (λ j → λ { (i = i0) → (λ i → (fs i) (f x)) ; (i = i1) → (sym (invPathCancel (snd gp x)) j) } ) (test i) where + test : {x : A} → PathP (λ i → fst gp (f x) ≡ fs (~ i) (f x)) (λ i → (fs i) (f x)) (refl {x = (fst gp (f x))}) + test {x = x} j i = fs (i ∧ ( ~ j)) (f x) + + helper3 : (l : (x : A) → (fst (P (f x)) )) (gp hq : Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → (gp ≡ hq) ≡ Σ (fst gp ≡ fst hq) + (λ a≡ → PathP (λ i → (a : A) → a≡ i (f a) ≡ l a) (snd gp) (snd hq)) + helper3 l gp hq = sym (ua (Σ≡ {A = ((a : B) → fst (P a))} {B = (λ x → (a : A) → x (f a) ≡ l a)} {x = gp} {y = hq})) + + + + +------------------ +{- Sn ≡ Ω(Sn+1) -} + + + + + + + + +-------------------- Guillaume Test -------------------- + + + + + + + + + + +sphereSuc : (n : ℕ) → (S₊ n) → (S₊ (suc n)) +sphereSuc n north = north +sphereSuc n south = south +sphereSuc (suc n) (merid a i) = merid (sphereSuc n a) i + +helplemma : {A : Type ℓ} {a : A} {p q : a ≡ a} {P : p ≡ q} → (C : (a ≡ a) → Type ℓ) → (a1 a2 : C (P i0)) → (b1 b2 : C (P (i1))) → PathP (λ i → C (P i)) a1 b2 → (a1 ≡ a2) → (b1 ≡ b2) → PathP (λ i → C (P i)) a2 b2 +helplemma {ℓ = ℓ} {a = a}{p = p} {q = q} {P = P} C a1 a2 b1 b2 depP a1a2 b1b2 = {!J {ℓ} {?} !} + +pathExtender : (B : A → Type ℓ) → {!!} +pathExtender = {!!} + +promote : (n : ℕ) → (S₊ n) → typ (Ω ((S₊ (suc n)) , north) ) +promote zero north = refl {x = north} +promote zero south = refl {x = north} +promote (suc n) north = refl {x = north} +promote (suc n) south = refl {x = north} +promote (suc n) (merid a i) = (( sym (rCancel (merid (merid a i0))) ∙ (λ i → (merid (merid a i) ∙ (sym (merid (merid north i)))))) ∙ rCancel (merid (merid a i1)) ) i + +decode' : (n : ℕ) → ∥ (S₊ n) ∥ (ℕ→ℕ₋₂ n) → ∥ typ (Ω (S₊ (suc n) , north)) ∥ (ℕ→ℕ₋₂ n) +decode' n = rec {A = (S₊ n)} {B = ∥ typ (Ω (S₊ (suc n) , north)) ∥ (ℕ→ℕ₋₂ n) } (isOfHLevel∥∥ {A = typ (Ω (S₊ (suc n) , north))} (ℕ→ℕ₋₂ n)) λ x → ∣ promote n x ∣ + +CODES : (n : ℕ) → S₊ (suc n) → Type₀ +CODES n north = ∥ S₊ n ∥ (ℕ→ℕ₋₂ n) +CODES n south = ∥ S₊ n ∥ (ℕ→ℕ₋₂ n) +CODES n (merid a i) = {!(typ ((Ω^ (suc n)) (Type₀ , ∥ S₊ n ∥ (ℕ→ℕ₋₂ n))))!} + +encode' : (n : ℕ) → ∥ typ (Ω ((S₊ (suc n)) , north)) ∥ (ℕ→ℕ₋₂ n) → ∥ S₊ n ∥ (ℕ→ℕ₋₂ n) +encode' n = rec {A = typ (Ω ((S₊ (suc n)) , north))} + {B = ∥ S₊ n ∥ (ℕ→ℕ₋₂ n)} + (isOfHLevel∥∥ {A = (S₊ n )} (ℕ→ℕ₋₂ n)) + λ x → {!(cong (CODES n) x) ?!} + +{- +functions : (n : ℕ) (f : A → B) → + (typ ((Ω^ n) ((A → B) , f))) ≡ + ((x : A) → (typ ((Ω^ n) (B , f x)))) +functions zero = {!!} +functions {A = A} {B = B}(suc zero) f = isoToPath (iso (λ g x i → (g i) x) (λ g → funExt {f = f} {g = f} g ) (λ b → refl) λ b → refl) +functions (suc (suc zero)) f = isoToPath (iso {!!} {!!} {!!} {!!}) + where + helper' : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + typ (Ω (Ω ((A → B) , f))) ≡ ( _≡_ {A = ( _≡_ f f)} (refl {x = f}) (refl {x = f})) + helper' f = refl + helper : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + (typ (Ω (Ω ((A → B) , f)))) ≡ ( _≡_ {A = ((x : A) → (f x ≡ f x) )} + (λ x _ → f x) + (λ x _ → f x)) + helper f = isoToPath (iso (λ g i x j → (g i j) x) (λ x k j y → ((x k) y) j) (λ b → refl) λ b → refl) + + helper2 :{A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + ( _≡_ {A = ((x : A) → (f x ≡ f x) )} (λ x _ → f x) (λ x _ → f x)) + ≡ + ((x : A) → (_≡_ {A = f x ≡ f x} (λ _ → f x) (λ _ → f x ))) + helper2 = {!!} + + +functions (suc (suc (suc n))) f = {!!} -- isoToPath (iso {!!} (λ x → {!!}) {!!} {!!}) + where + helper : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + typ (Ω (Ω (Ω ((Ω^ n) ((A → B) , f))))) ≡ ( _≡_ {A = typ (Ω (Ω ((Ω^ n) ((A → B) , f))))} + (λ _ → pt (Ω ((Ω^ n) ((A → B) , f)))) + (λ _ → pt (Ω ((Ω^ n) ((A → B) , f))))) + helper f = refl + + where + Ω^n : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + typ (Ω (Ω ((Ω^ n) ((A → B) , f)))) ≡ ((x : A) → typ (Ω (Ω ((Ω^ n) (B , f x))))) + Ω^n f = functions (suc (suc n)) f + Ω^n' : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + typ (Ω (Ω (Ω ((Ω^ n) ((A → B) , f))))) ≡ typ (Ω (((x : A) → typ (Ω (Ω ((Ω^ n) (B , f x))))) , (λ _ → refl))) + Ω^n' f i = typ (Ω (( Ω^n f i) , {!!})) + + +-} + +Fun1000 : (n : ℕ) {f : A → B} → + ((typ ((Ω^ (suc n)) ((A → B) , f))) → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) +-- Fun1001 : (n : ℕ) {f : A → B} → Fun1000 n {f} refl ≡ (λ x → refl) +Fun1000 zero p x i = p i x +Fun1000 (suc n) p x = {!Fun1000 n (p j) x i!} +-- Fun1001 = {!!} + + +lrFun2* : (n : ℕ) {f : A → B} → + Σ ((typ ((Ω^ (suc n)) ((A → B) , f))) → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) + (λ g → (g refl) ≡ (λ x → refl)) +lrFun2* zero {f = f} = funExt⁻ , refl +lrFun2* (suc n) {f = f} = (λ p x → {!(funExt⁻ (cong (fst (lrFun2* n)) p) x)!}) + , + funExt (λ y → cancelReflMid (λ i → snd (lrFun2* n) (~ i) y) + (λ i → snd (lrFun2* n {f = f}) i y) ∙ + invPathCancel (λ i → snd (lrFun2* n) (~ i) y)) + +lrFun : (n : ℕ) {f : A → B} → + ((typ ((Ω^ (suc n)) ((A → B) , f))) → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) +lrFun n = fst (lrFun2* n) + + + where + + +rlFun* : (n : ℕ) {f : A → B} → + Σ ((((x : A) → (typ ((Ω^ (suc n)) (B , f x))))) → + ((typ ((Ω^ (suc n)) ((A → B) , f)))) ) + (λ g → (g (λ x → snd (((Ω^ (suc n)) (B , f x)))) ≡ pt ((Ω^ (suc n)) ((A → B) , f)))) + +rlFun* zero {f = f} = (λ g → funExt g) , refl +rlFun* {A = A} {B = B}(suc n) {f = f} = (λ g → (sym (snd (rlFun* n)) ∙ cong (fst (rlFun* n)) (funExt g) ∙ snd (rlFun* n))) + , + (cancelReflMid (λ i → snd (rlFun* n) (~ i)) (snd (rlFun* n))) ∙ + invPathCancel (λ i → snd (rlFun* n) (~ i)) + +rlFun : (n : ℕ) {f : A → B} → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x))))) → + ((typ ((Ω^ (suc n)) ((A → B) , f)))) +rlFun n {f = f} = fst (rlFun* n {f = f}) + where + +functions2 : (n : ℕ) (f : A → B) → + ((x : A) → (typ ((Ω^ n) (B , f x)))) ≃ (typ ((Ω^ n) ((A → B) , f))) + +functions2 {A = A} {B = B} zero f = idEquiv (A → B) +functions2 {A = A} {B = B} (suc n) f = rlFun n , record { equiv-proof = λ y → ((fst (lrFun2* n)) y , {!!}) , {!!} } where + + helper : (n : ℕ) (y : (typ ((Ω^ (suc n)) ((A → B) , f)))) → (rlFun n ((fst (lrFun2* n)) y) ≡ y) + helper zero y = refl + helper (suc n) y = (λ i → (sym (snd (rlFun* n))) ∙ ((helper5 (funExt (λ x → sym (cong (λ g → g x ) (snd (lrFun2* n))))) (cong (fst (lrFun2* n)) y) (funExt (λ x → (cong (λ g → g x ) (snd (lrFun2* n))))) (fst (rlFun* n))) i ) ∙ (snd (rlFun* n))) ∙ {!helper !} + where + + + idTest : (n : ℕ) → (rlFun n {f} (λ x → snd (lrFun2* n {f}) i1 x) ≡ rlFun n (λ x → snd (lrFun2* n {f}) i0 x)) ≡ ((fst (rlFun* n) (λ x → snd (Ω ((Ω^ n) (B , f x))))) ≡ (pt (Ω ((Ω^ n) ((A → B) , f))))) + idTest n i = refl {x = rlFun n {f} (λ x → snd (lrFun2* n {f}) i1 x)} i ≡ idTest1 i where + idTest1 : rlFun n (λ x → snd (lrFun2* n {f}) i0 x) ≡ (pt (Ω ((Ω^ n) ((A → B) , f)))) + idTest1 = helper n (pt (Ω ((Ω^ n) ((A → B) , f)))) + + + test : (n : ℕ) → PathP (λ i → {!!}) ( (cong (rlFun n {f}) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))) (snd (rlFun* n {f})) + test n i = {!!} where + + + maybe2 : {n : ℕ} (i : I) → snd (rlFun* n) (~ i) ≡ (pt (Ω ((Ω^ n) ((A → B) , f)))) + maybe2 {n} i j = snd (rlFun* n) ((~ i) ∨ j) + + maybe3 : (n : ℕ) → (pt (Ω ((Ω^ n) ((A → B) , f))) ≡ rlFun n (λ x → snd (lrFun2* n {f}) i0 x)) + maybe3 n = sym (snd (rlFun* n {f})) ∙ λ i → rlFun n (hihih ( ~ i)) where + hihih : (λ (x : A) → snd (lrFun2* n {f}) i0 x) ≡ (λ x → snd (Ω ((Ω^ n) (B , f x)))) + hihih = funExt λ x → cong (λ g → g x) (snd (lrFun2* n)) + + maybe : (n : ℕ) (y : (typ (Ω ((Ω^ (suc n)) ((A → B) , f))))) → ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x)))))) ∙ (cong (rlFun n) (funExt (λ x → (λ i → fst (lrFun2* n) (y i) x)))) ∙ ((sym (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x))))))) ≡ (snd (rlFun* n)) ∙ y ∙ (sym (snd (rlFun* n))) + maybe n y i = test n i ∙ (cong (λ x → (helper n x) i) y ) ∙ sym (test n i) + -- (λ j → (snd (rlFun* n)) (~ j ∨ i)) ∙ {!!} + + -- ∙ {!λ i → (cong (λ x → helper n x i ) y)!} ∙ {!!} -- (cong (λ x → helper n x i ) ?) + + compTest : (n : ℕ) (y : (typ (Ω ((Ω^ (suc n)) ((A → B) , f))))) → + (sym (snd (rlFun* n))) ∙ + ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x)))))) ∙ + (cong (rlFun n) (funExt (λ x → (λ i → fst (lrFun2* n) (y i) x)))) ∙ + ((sym (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x))))))) ∙ + (snd (rlFun* n)) ≡ + refl ∙ {!!} ∙ y ∙ {!!} ∙ refl + compTest n y i = (λ i₁ → snd (rlFun* n {f}) (~ i₁ ∨ i)) ∙ + {!snd (rlFun* n {f}) (~ i1 ∨ i)!} {- hcomp (λ k → λ{(i = i0) → {!maxZeuner i0!} ; (i = i1) → (maxZeuner i1)}) (maxZeuner i) -} ∙ + (cong (λ x → (helper n x) i) y ) ∙ + {!snd (rlFun* n {f}) (~ i1 ∨ i)!} + ∙ λ i₁ → snd (rlFun* n {f}) (i₁ ∨ i) + + where + + maxZeuner : (i : I) → snd (rlFun* n) i ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i + maxZeuner i j = hcomp (λ k → λ{(j = i0) → snd (rlFun* n) i ; (j = i1) → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i ∨ (~ k)) }) (snd (rlFun* n {f}) (i ∨ j)) + maxZeunerHelper : (i : I) → maxZeuner i0 ≡ {!snd (rlFun* n) i0 ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i0!} + maxZeunerHelper = {!!} + + gimmeMore : (i : I) → rlFun n (λ x → snd (lrFun2* n {f}) i0 x) ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i + gimmeMore i = (λ j → rlFun n (λ x → snd (lrFun2* n {f}) j x)) ∙ (snd (rlFun* n {f})) ∙ λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i ∨ (~ j)) + + gimmeLeft : (i : I) → rlFun n (λ x → snd (lrFun2* n {f}) i1 x) ≡ snd (rlFun* n) i + gimmeLeft i = λ j → snd (rlFun* n {f}) (j ∧ i) + + gimmei0 : ((sym (gimmeLeft i0)) ∙ ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))) ∙ (gimmeMore i0)) ≡ ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))) + gimmei0 = {!((sym (gimmeLeft i0)))!} where + + gimmei0' : ((sym (gimmeLeft i0))) ≡ refl + gimmei0' j = refl + + + + + gimmei1 : ((sym (gimmeLeft i1)) ∙ ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))) ∙ (gimmeMore i1)) ≡ refl + gimmei1 = gimmei1' ∙ more!! where + + gimmei1' : ((sym (gimmeLeft i1)) ∙ ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))) ∙ (gimmeMore i1)) ≡ (sym (gimmeLeft i1)) ∙ (snd (rlFun* n {f})) ∙ λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j)) + gimmei1' = ∙-assoc (sym (gimmeLeft i1)) (( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x))))))) ((snd (rlFun* n {f})) ∙ λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j))) where + + ∙-assoc : {A : Type ℓ}{a b c d : A} (p : a ≡ b) (q : b ≡ c) (t : b ≡ d) → p ∙ q ∙ (sym q ∙ t) ≡ p ∙ t + ∙-assoc {ℓ = ℓ} {A = A} {a = a} {b = b} {c = c} {d = d} p q t = cong (λ x → p ∙ x) {!q ∙ (sym q ∙ t) ≡ t!} + + + more!! : (sym (gimmeLeft i1)) ∙ (snd (rlFun* n {f})) ∙ (λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j))) ≡ refl + more!! = (assoc (sym (gimmeLeft i1)) (snd (rlFun* n {f})) (λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j)))) ∙ (λ j → (invPathCancel (sym (gimmeLeft i1)) j) ∙ (λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j)))) ∙ lCancel (λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j))) + + more!!! : (λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j))) ≡ refl + more!!! = refl + + + + hello : (i : I) → snd (rlFun* n) i ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i + hello i = (λ j → snd (rlFun* n {f}) (i ∨ j)) ∙ λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i ∨ (~ j)) + + kittyHelper : PathP (λ i → (snd (rlFun* n {f})) i ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i) (hello i0) refl + kittyHelper j = hcomp (λ k → λ{ (j = i0) → hello i0 ; + (j = i1) → invPathCancel refl k }) + ((λ i → snd (rlFun* n {f}) (j ∨ i)) ∙ λ i → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (j ∨ (~ i))) +{- + kittyHelperB : PathP (λ i → (snd (rlFun* n {f})) i ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i) ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x)))))) refl + kittyHelperB i j = hcomp (λ k → λ{ (j = i0) → {!!} + ; (j = i1) → {!(cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x)))))!}} ) + {!( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ j) x))))))!} +-} + + -- helloKitty : hello i0 ≡ ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x)))))) + -- helloKitty k = {!!} + + -- helloKitty2 : hello i1 ≡ refl + -- helloKitty2 = {!!} + + mini : (n : ℕ) → rlFun n (λ x → snd (lrFun2* n {f}) i0 x) ≡ pt (Ω ((Ω^ n) ((A → B) , f))) + mini n = helper n refl + + mini2 : (n : ℕ) → rlFun n (λ x → snd (lrFun2* n {f}) i1 x) ≡ pt (Ω ((Ω^ n) ((A → B) , f))) + mini2 n = (λ i → rlFun n (λ x → snd (lrFun2* n {f}) (~ i) x)) ∙ helper n refl + + leftMost : PathP (λ i → (pt (Ω ((Ω^ n) ((A → B) , f))) ≡ (snd (rlFun* n)) i)) (sym (snd (rlFun* n))) refl + leftMost i = λ i₁ → snd (rlFun* n) (~ i₁ ∨ i) + + + + notLeftMost : (n : ℕ) → PathP (λ i → mini2 n i ≡ mini n i ) (( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x))))))) refl + notLeftMost zero i j = {!fst (rlFun* (suc n)) (λ x → snd (lrFun2* (suc n) {f}) (i1) x)!} + notLeftMost (suc n) i = {!(( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))))!} + where + ffs2 : (n : ℕ) (i : I) → (fst (rlFun* n {f}) (λ x → snd (lrFun2* n {f}) (~ i1 ∨ i) x)) ≡ mini n i + ffs2 n i = ((λ j → fst (rlFun* n) (λ x → snd (lrFun2* n {f}) (i ∨ j) x)) ∙ snd (rlFun* n {f}) ) ∙ λ j → mini n (i ∨ (~ j)) + + ffs : (n : ℕ) (i : I) → (fst (rlFun* n {f}) (λ x → snd (lrFun2* n {f}) (~ i0 ∨ i) x)) ≡ mini2 n i + ffs n i = snd (rlFun* n {f}) ∙ {!mini2 i1!} + + {- + + tester : ((sym (ffs i0)) ∙ (λ j → fst (rlFun* n) (snd (lrFun2* n {f}) (i0 ∨ (~ j)))) ∙ (ffs2 i0)) ≡ (( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x))))))) + tester = (({!((sym (ffs i1)) ∙ (λ j → fst (rlFun* n) (snd (lrFun2* n {f}) (i1 ∨ (~ j)))) ∙ (ffs2 i1))!}) ∙ sym (rUnit (refl ∙ (λ j → fst (rlFun* n) (snd (lrFun2* n {f}) (~ j)))))) ∙ sym (lUnit (λ j → fst (rlFun* n) (snd (lrFun2* n {f}) (i0 ∨ (~ j))))) + where + helpi : (ffs2 i0) ≡ refl + helpi = {!!} + + snLeftMost : {!(λ j → fst (rlFun* n {f}) (snd (lrFun2* n) (~ j)))!} + snLeftMost = {!!} + + -} + + + helper2 : (n : ℕ) (y : typ (Ω (Ω ((Ω^ n) ((A → B) , f))))) → funExt ((λ (x : A) → (λ i → snd (lrFun2* n) (~ i) x) ∙ (λ i → fst (lrFun2* n) (y i) x) ∙ (λ i → snd (lrFun2* n {f = f}) i x))) ≡ (funExt (λ (x : A) → (λ i → snd (lrFun2* n) (~ i) x))) ∙ funExt (λ (x : A) → (λ i → fst (lrFun2* n) (y i) x)) ∙ funExt λ (x : A) → (λ i → snd (lrFun2* n {f = f}) i x) + helper2 zero y = refl + helper2 (suc n) y = refl + + helper4 : {A : Type ℓ} {B : Type ℓ} {a b c : A} (p : a ≡ b) (q : b ≡ c) (g : A → B) → cong g (p ∙ q) ≡ ((cong g p) ∙ (cong g q)) + helper4 {ℓ = ℓ} {A} {B} {a} {b} {c} p q g = J {ℓ} {A} {a} {ℓ} (λ b1 p → ((q1 : b1 ≡ c) → cong g (p ∙ q1) ≡ ((cong g p) ∙ (cong g q1)))) + (λ q1 → J {ℓ} {A} {a} {ℓ} (λ c2 q2 → (λ i → g ((refl ∙ q2) i)) ≡ (λ i → g (refl i)) ∙ (λ i → g (q2 i))) ( (λ j → ((λ i → g (((rUnit (λ _ → a)) (~ j) ) i)))) ∙ lUnit ((λ i → g (refl i)))) {c} q1 ) + p q + + helper5 : {A : Type ℓ} {B : Type ℓ} {a b c d : A} (p : a ≡ b) (q : b ≡ c) (r : c ≡ d) (g : A → B) → cong g (p ∙ q ∙ r) ≡ ((cong g p) ∙ (cong g q) ∙ (cong g r)) + helper5 {ℓ = ℓ} {A} {B} {a} {b} {c} {d} p q r g = helper4 p (q ∙ r) g ∙ ((cong (λ x → (cong g p) ∙ x )) (helper4 q r g)) + + helper3 : (n : ℕ) (y : typ (Ω (Ω ((Ω^ n) ((A → B) , f))))) → (cong (rlFun n) (funExt (λ x → (λ i₁ → fst (lrFun2* n) (y i₁) x)))) ≡ {!!} + helper3 zero y = {!!} + helper3 (suc n) y = {!(cong (rlFun (suc n)) (funExt (λ x → (λ i₁ → fst (lrFun2* (suc n)) (y i₁) x))))!} + +{- +functions' {A = A} {B = B}(suc zero) f = isoToPath (iso (λ g x i → (g i) x) (λ g → funExt {f = f} {g = f} g ) (λ b → refl) λ b → refl) + where + helper : transport (λ i → (sym (functions' (suc zero) f)) i) (λ x → refl) ≡ (pt ((Ω ) ((A → B) , f))) + helper = {!!} +functions' (suc (suc zero)) f = isoToPath (iso {!!} {!!} {!!} {!!}) + where + helper' : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + typ (Ω (Ω ((A → B) , f))) ≡ ( _≡_ {A = ( _≡_ f f)} (refl {x = f}) (refl {x = f})) + helper' f = refl + helper : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + (typ (Ω (Ω ((A → B) , f)))) ≡ ( _≡_ {A = ((x : A) → (f x ≡ f x) )} + (λ x _ → f x) + (λ x _ → f x)) + helper f = isoToPath (iso (λ g i x j → (g i j) x) (λ x k j y → ((x k) y) j) (λ b → refl) λ b → refl) + + helper2 :{A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + ( _≡_ {A = ((x : A) → (f x ≡ f x) )} (λ x _ → f x) (λ x _ → f x)) + ≡ + ((x : A) → (_≡_ {A = f x ≡ f x} (λ _ → f x) (λ _ → f x ))) + helper2 = {!!} +functions' (suc (suc (suc n))) f = {!!} + where + helper : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + typ (Ω (Ω (Ω ((Ω^ n) ((A → B) , f))))) ≡ ( _≡_ {A = typ (Ω (Ω ((Ω^ n) ((A → B) , f))))} + (λ _ → pt (Ω ((Ω^ n) ((A → B) , f)))) + (λ _ → pt (Ω ((Ω^ n) ((A → B) , f))))) + helper f = refl + + where + Ω^n : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + typ (Ω (Ω ((Ω^ n) ((A → B) , f)))) ≡ ((x : A) → typ (Ω (Ω ((Ω^ n) (B , f x))))) + Ω^n f = functions (suc (suc n)) f + Ω^n' : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + (Ω (Ω (Ω ((Ω^ n) ((A → B) , f))))) ≡ (Ω (((x : A) → typ (Ω (Ω ((Ω^ n) (B , f x))))) , (λ _ → refl))) + Ω^n' f i = (Ω (( Ω^n f i) , {!(transport (λ j → ( Ω^n f j)) (refl i)) !})) where + lemma : transport (λ i → Ω^n f i) refl ≡ λ _ → refl + lemma = funExt {f = transport (λ i → Ω^n f i) refl} {g = λ _ → refl} λ x → {!!} + + + + + ---------- + +lrFun3* : (n : ℕ) {f : A → B} → + Σ ((typ ((Ω^ (suc n)) ((A → B) , f))) → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) (λ g → (g refl) ≡ (λ x → refl)) +lrFun3* zero {f = f} = funExt⁻ , refl +lrFun3* {A = A} {B = B} (suc n) {f = f} = (λ p x → transport (λ i → (cong (λ g → g x) (snd (lrFun3* n {f = f}))) i ≡ (cong (λ g → g x) (snd (lrFun3* n {f = f}))) i) (funExt⁻ (cong (λ x → (fst (lrFun3* n) x )) p) x)) , funExt (λ x → transpLemma x) + module lr where + helpLemma : (x : A) → ((fst (lrFun3* n) (snd (Ω ((Ω^ n) ((A → B) , f)))) x ≡ (fst (lrFun3* n) (snd (Ω ((Ω^ n) ((A → B) , f)))) x))) ≡ (typ (Ω (Ω ((Ω^ n) (B , f x))))) + helpLemma x i = helper x i ≡ helper x i where + + helper : (x : A) → fst (lrFun3* n) (snd (Ω ((Ω^ n) ((A → B) , f)))) x ≡ pt ((Ω ((Ω^ n) (B , f x)))) + helper x = cong (λ g → g x) (snd (lrFun3* n {f = f})) + + transpLemma : (x : A) → transp (λ i → helpLemma x i) i0 refl ≡ refl + transpLemma x j = transp (λ i → helpLemma x (i ∨ j)) j refl + +lrFun3 : (n : ℕ) {f : A → B} → ((typ ((Ω^ (suc n)) ((A → B) , f))) → (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) +lrFun3 n {f = f} = (fst (lrFun3* n {f = f})) + + +rlFun3* : (n : ℕ) {f : A → B} → + Σ ((((x : A) → (typ ((Ω^ (suc n)) (B , f x))))) → + ((typ ((Ω^ (suc n)) ((A → B) , f)))) ) + (λ g → (g (λ x → snd (((Ω^ (suc n)) (B , f x)))) ≡ pt ((Ω^ (suc n)) ((A → B) , f)))) + + + + +rlFun3* zero {f = f} = (λ g → funExt g) , refl +rlFun3* {A = A} {B = B}(suc n) {f = f} = (λ g → transport (λ i → (snd (rlFun3* n {f})) i ≡ (snd (rlFun3* n {f})) i) (cong (fst (rlFun3* n)) (funExt g))) , transpLemma -- (cancelReflMid (λ i → snd (rlFun* n) (~ i)) (snd (rlFun* n))) ∙ invPathCancel (λ i → snd (rlFun* n) (~ i)) + module rl where + helpLemma : (fst (rlFun3* n) (λ x → snd (Ω ((Ω^ n) (B , f x)))) ≡ fst (rlFun3* n) (λ x → snd (Ω ((Ω^ n) (B , f x))))) ≡ typ (Ω (Ω ((Ω^ n) ((A → B) , f)))) + helpLemma i = (snd (rlFun3* n {f})) i ≡ (snd (rlFun3* n {f})) i + + transpLemma : transport (λ i → helpLemma i) refl ≡ refl + transpLemma j = transp (λ i → helpLemma (i ∨ j)) j refl -- transp (λ i → helpLemma (i ∨ j)) j refl + +rlFun3 : (n : ℕ) {f : A → B} → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x))))) → + ((typ ((Ω^ (suc n)) ((A → B) , f)))) +rlFun3 n {f = f} = fst (rlFun3* n {f = f}) + where + +functions3 : (n : ℕ) (f : A → B) → + ((x : A) → (typ ((Ω^ n) (B , f x)))) ≃ (typ ((Ω^ n) ((A → B) , f))) + +functions3 {A = A} {B = B} zero f = idEquiv (A → B) +functions3 {A = A} {B = B} (suc n) f = rlFun3 n , record { equiv-proof = λ y → (( (fst (lrFun3* n)) y) , λ i → {!!}) , {!!} } where + linv : (n : ℕ) (y : typ (Ω ((Ω^ n) ((A → B) , f)))) → rlFun3 n (fst (lrFun3* n {f = f}) y) ≡ y + linv zero y = refl + linv (suc n) y = {!!} where + maybe : rlFun3 (suc n) (fst (lrFun3* (suc n) {f = f}) y) ≡ λ j x → {!((cong (λ x → (fst (lrFun3* n) x )) y) (~ j)) x!} + maybe = {!!} + + +inv1 : (n : ℕ) {f : A → B} (g : ((typ ((Ω^ (suc n)) ((A → B) , f))) )) → (rlFun3 n) ((lrFun3 n) g ) ≡ g +inv1 zero {f = f} g = refl +inv1 {A = A} {B = B}(suc n) {f = f} g j = {! -- transp (λ i → (snd (rlFun3* n {f})) (i ∨ j) ≡ (snd (rlFun3* n {f})) (i ∨ j)) i0 + (cong (fst (rlFun3* n)) (funExt ( + (λ x → transp (λ i → (cong (λ g → g x) (snd (lrFun3* n {f = f}))) (i ∨ j) + ≡ (cong (λ g → g x) (snd (lrFun3* n {f = f}))) (i ∨ j)) (~ j) + (funExt⁻ (cong (λ x → (fst (lrFun3* n) x )) g) x)) ) ) )!} + + where + maybe : PathP (λ i → {! !} ≡ {!!} ) (cong (fst (rlFun3* n)) (funExt ( + (λ x → transport (λ i → (cong (λ g → g x) (snd (lrFun3* n {f = f}))) i + ≡ (cong (λ g → g x) (snd (lrFun3* n {f = f}))) i) + (funExt⁻ (cong (λ x → (fst (lrFun3* n {f}) x )) g) x)) ) ) ) + {!(cong (fst (rlFun3* n)) (funExt ( (λ x → (funExt⁻ (cong (λ x → (fst (lrFun3* n {f}) x )) g) x))))) !} + maybe j = {!fst (rlFun3* n) (λ x → snd (lrFun3* n {f}) i1 x)!} +-} From abb8a5246bd3cf47e3d67419f363c6619a34f2d5 Mon Sep 17 00:00:00 2001 From: Axel Date: Tue, 25 Feb 2020 10:57:25 +0100 Subject: [PATCH 07/30] Pointed funs change --- Cubical/Foundations/Pointed/Base.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Cubical/Foundations/Pointed/Base.agda b/Cubical/Foundations/Pointed/Base.agda index d796dee28..79ce3184b 100644 --- a/Cubical/Foundations/Pointed/Base.agda +++ b/Cubical/Foundations/Pointed/Base.agda @@ -16,5 +16,5 @@ pt = snd _→*_ : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → Type (ℓ-max ℓ ℓ') _→*_ A B = Σ (typ A → typ B) λ f → f (pt A) ≡ pt B -_→*_/_ : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → (A →* B) → Pointed (ℓ-max ℓ ℓ') -A →* B / f = (A →* B) , f +_→*_* : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +A →* B * = (A →* B) , (λ x → pt B) , refl From c4b83e3407e4940c23115073ae187a1fa6e37e94 Mon Sep 17 00:00:00 2001 From: Axel Date: Fri, 28 Feb 2020 12:55:34 +0100 Subject: [PATCH 08/30] backup --- Cubical/HITs/SetTruncation/Properties.agda | 3 +-- Cubical/HITs/Truncation/Properties.agda | 20 ++++---------- Cubical/ZCohomology/Properties.agda | 9 +++---- Cubical/ZCohomology/S1/S1.agda | 31 ++++++---------------- GNUmakefile | 1 + 5 files changed, 19 insertions(+), 45 deletions(-) diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index 9c7393674..38806271f 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -57,11 +57,10 @@ elimSetTrunc3 : {B : (x y z : ∥ A ∥₀) → Type ℓ} elimSetTrunc3 Bset g = elimSetTrunc2 (λ _ _ → hLevelPi 2 λ _ → Bset _ _ _) (λ a b → elimSetTrunc (λ _ → Bset _ _ _) (λ c → g a b c)) - setTruncIsSet : isSet ∥ A ∥₀ setTruncIsSet a b p q = squash₀ a b p q -setId : isSet A → ∥ A ∥₀ ≡ A +setId : isSet A → ∥ A ∥₀ ≡ A setId {A = A} isset = isoToPath (iso (elimSetTrunc {A = A} (λ _ → isset) (λ x → x)) diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index aa9a1f7f5..c63cfdb18 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -125,13 +125,12 @@ isOfHLevel∥∥ : (n : ℕ₋₂) → isOfHLevel (2+ n) (∥ A ∥ n) isOfHLevel∥∥ neg2 = hub ⊥-elim , λ _ → ≡hub ⊥-elim isOfHLevel∥∥ (suc n) = isSphereFilled→isOfHLevelSuc isSphereFilled∥∥ - -- isOfHLevel∥∥ n = isSnNull→isOfHLevel isNull-Null -- ∥_∥ n is a modality rec : {n : ℕ₋₂} - {B : Type ℓ'} → + {B : Type ℓ'} → (isOfHLevel (2+ n) B) → (g : (a : A) → B) → (∥ A ∥ n → B) @@ -217,17 +216,14 @@ groupoidTrunc≃Trunc1 = (ind (λ _ → hLevelPath 4 (isOfHLevel∥∥ 2) _ _) (λ _ → refl)) (g2TruncElim _ _ (λ _ → hLevelPath 4 squash₂ _ _) (λ _ → refl))) - - ---- ∥ Ω A ∥ ₙ ≡ Ω ∥ A ∥ₙ₊₁ ---- - {- Proofs of Theorem 7.3.12. and Corollary 7.3.13. in the HoTT book -} private {- We define the fibration P to show a more general result -} P : ∀ {ℓ} {B : Type ℓ}{n : ℕ₋₂} → ∥ B ∥ (suc n) → ∥ B ∥ (suc n) → Type ℓ - P x y = fst (P₁ x y) + P x y = fst (P₁ x y) where P₁ : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} → ∥ B ∥ (suc n) → ∥ B ∥ (suc n) → (HLevel ℓ (2+ n)) @@ -278,7 +274,6 @@ private (isOfHLevel∥∥ {A = B} (suc (suc n)) ∣ u ∣ ∣ v ∣) (λ p → cong (λ z → ∣ z ∣) p) - {- auxilliary function r used to define encode -} r : ∀ {ℓ} {B : Type ℓ} {m : ℕ₋₂} (u : ∥ B ∥ (suc m)) → P u u r {m = m} = ind {B = (λ u → P u u)} @@ -287,13 +282,13 @@ private {- encode function from x ≡ y to P x y -} encode-fun : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x y : ∥ B ∥ (suc n)) → - _≡_ {A = ∥ B ∥ (suc n)} x y → - P x y + _≡_ {A = ∥ B ∥ (suc n)} x y → + P x y encode-fun x y p = transport (λ i → P x (p i )) (r x) {- We need the following two lemmas on the functions behaviour for refl -} dec-refl : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x : ∥ B ∥ (suc n)) → - decode-fun x x (r x) ≡ refl {x = x} + decode-fun x x (r x) ≡ refl {x = x} dec-refl {B = B} {n = neg2} x = ind {A = B} {n = suc neg2} {B = λ x → decode-fun x x (r x) ≡ refl {x = x} } (λ x → (hLevelSuc (2+ (suc neg2)) (x ≡ x) @@ -310,15 +305,12 @@ private (decode-fun x x (r x)) refl)) λ c → refl - enc-refl : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x : ∥ B ∥ (suc n)) → encode-fun x x (refl {x = x}) ≡ r x enc-refl x j = transp (λ i → P x (refl {x = x} i)) j (r x) - - {- decode-fun is a right-inverse -} P-rinv : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (u v : ∥ B ∥ (suc n)) → (x : _≡_ {A = ∥ B ∥ (suc n)} u v) → @@ -340,8 +332,6 @@ private λ z → hLevelSuc (2+ suc n) (P x y) (hLevelP {n = n} x y) (encode-fun x y (decode-fun x y z)) z) helper u v - - where helper : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (a b : B) diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 04ded167c..5529101fb 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -1,7 +1,6 @@ {-# OPTIONS --cubical --safe #-} module Cubical.ZCohomology.Properties where - open import Cubical.ZCohomology.Base open import Cubical.HITs.Sn open import Cubical.Foundations.HLevels @@ -17,7 +16,7 @@ open import Cubical.Data.Empty open import Cubical.Data.Sigma open import Cubical.Data.Prod.Base open import Cubical.HITs.Susp -open import Cubical.HITs.SetTruncation +open import Cubical.HITs.SetTruncation open import Cubical.HITs.Nullification open import Cubical.Data.Int open import Cubical.Data.Nat @@ -53,12 +52,12 @@ coHomRed+1Equiv zero A i = ∥ helpLemma {C = (Int , pos 0)} i ∥₀ map2 : ((((A ⊎ Unit) , inr (tt)) →* C)) → (A → typ C) map2 (g , pf) x = g (inl x) - + linvPf : (b :((((A ⊎ Unit) , inr (tt)) →* C))) → map1 (map2 b) ≡ b linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) where - helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x + helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x helper (inl x) = refl helper (inr tt) = sym snd - + coHomRed+1Equiv (suc n) A i = ∥ coHomRed+1.helpLemma A i {C = ((coHomK (suc n)) , ∣ north ∣)} i ∥₀ diff --git a/Cubical/ZCohomology/S1/S1.agda b/Cubical/ZCohomology/S1/S1.agda index 0516dbdce..029a6eeed 100644 --- a/Cubical/ZCohomology/S1/S1.agda +++ b/Cubical/ZCohomology/S1/S1.agda @@ -6,45 +6,30 @@ open import Cubical.ZCohomology.Properties open import Cubical.HITs.S1 open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism -open import Cubical.HITs.SetTruncation +open import Cubical.HITs.SetTruncation open import Cubical.HITs.Nullification open import Cubical.Data.Int open import Cubical.Data.Nat open import Cubical.HITs.Truncation -private - variable - ℓ : Level - A : Type ℓ +---- H⁰(S¹) = ℤ ---- - ----- H⁰(S¹) ---- - -coHom0-S1 : (coHom zero S¹) ≡ Int -coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ (setId isSetInt) +coHom0-S1 : coHom zero S¹ ≡ Int +coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ setId isSetInt where helpLemma : (S¹ → Int) ≡ Int - helpLemma = isoToPath (iso fun - funinv - (λ b → refl) - λ f → funExt (λ x → rinvLemma f x)) + helpLemma = isoToPath (iso fun funinv (λ _ → refl) (λ f → funExt (rinvLemma f))) where fun : (S¹ → Int) → Int fun f = f base funinv : Int → (S¹ → Int) funinv a base = a - funinv a (loop i) = refl {x = a} i + funinv a (loop i) = a - rinvLemma : ( f : (S¹ → Int)) → - (x : S¹) → - funinv (fun f) x ≡ f x + rinvLemma : (f : S¹ → Int) → (x : S¹) → funinv (fun f) x ≡ f x rinvLemma f base = refl - rinvLemma f (loop i) = sym (helper f i) where - helper : (f : (S¹ → Int)) → - (i : I) → - (f (loop i) ≡ f base) - helper f i = (λ l → ((isSetInt (f base) (f base) (λ k → (f (loop k))) refl) l) i) + rinvLemma f (loop i) j = isSetInt (f base) (f base) (λ k → f (loop k)) refl (~ j) i ------------------------- diff --git a/GNUmakefile b/GNUmakefile index b859a1242..38669713e 100644 --- a/GNUmakefile +++ b/GNUmakefile @@ -28,6 +28,7 @@ check: $(wildcard Cubical/**/*.agda) $(AGDA) Cubical/Modalities/Everything.agda $(AGDA) Cubical/WithK.agda $(AGDA) Cubical/Experiments/Everything.agda + $(AGDA) Cubical/ZCohomology/Everything.agda .PHONY: listings listings: $(wildcard Cubical/**/*.agda) From c2220abe1971f63e11d1a6267ff77a1086aeddc3 Mon Sep 17 00:00:00 2001 From: Axel Date: Mon, 2 Mar 2020 18:58:27 +0100 Subject: [PATCH 09/30] nice --- Cubical/ZCohomology/PropertiesTrash.agda | 518 +++++++++++++++++------ 1 file changed, 398 insertions(+), 120 deletions(-) diff --git a/Cubical/ZCohomology/PropertiesTrash.agda b/Cubical/ZCohomology/PropertiesTrash.agda index 5a9005989..3319aaae9 100644 --- a/Cubical/ZCohomology/PropertiesTrash.agda +++ b/Cubical/ZCohomology/PropertiesTrash.agda @@ -1,5 +1,5 @@ {-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.Properties where +module Cubical.ZCohomology.PropertiesTrash where open import Cubical.ZCohomology.Base @@ -19,8 +19,9 @@ open import Cubical.Data.Prod.Base open import Cubical.HITs.Susp open import Cubical.HITs.SetTruncation open import Cubical.HITs.Nullification -open import Cubical.Data.Int +open import Cubical.Data.Int hiding (_+_ ; _-_) open import Cubical.Data.Nat +open import Cubical.Data.Unit open import Cubical.HITs.Truncation open import Cubical.HITs.Pushout @@ -33,44 +34,59 @@ private A : Type ℓ B : Type ℓ' +is-_-Truncated : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + (n : ℕ₋₂) → (f : A → B) → Type (ℓ-max ℓ ℓ') +is-_-Truncated {A = A} {B = B} n f = (b : B) → isOfHLevel (2+ n) (fiber f b) -{- Equivalence between cohomology and the reduced version -} -coHom↔coHomRed : (n : ℕ₋₂) → - (A : Set ℓ) → - (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) -coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ - module coHom↔coHomRed where - helplemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →* C))) - helplemma {C = C} = isoToPath (iso map1 - map2 - (λ b → linvPf b) - (λ _ → refl)) - where - map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →* C)) - map1 f = map1' , refl module helpmap where - map1' : A ⊎ Unit → fst C - map1' (inl x) = f x - map1' (inr x) = pt C - - map2 : ((((A ⊎ Unit) , inr (tt)) →* C)) → (A → typ C) - map2 (g , pf) x = g (inl x) - - linvPf : (b :((((A ⊎ Unit) , inr (tt)) →* C))) → map1 (map2 b) ≡ b - linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) - where - helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x - helper (inl x) = refl - helper (inr tt) = sym snd - -coHom↔coHomRed (suc n) A i = ∥ coHom↔coHomRed.helplemma A i {C = ((coHomK (suc n)) , ∣ north ∣)} i ∥₀ +_-_ : ℕ → ℕ → ℕ +zero - m = zero +suc n - zero = suc n +suc n - suc m = n - m + +canceller : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (p q : a ≡ b) (r : b ≡ c) → p ∙ r ≡ q ∙ r → p ≡ q +canceller {ℓ} {A} {a} {b} {c} p q r id = (rUnit p ∙ (λ j → (p ∙ (rCancel r (~ j)))) ∙ assoc p r (sym r)) ∙ (cong (λ x → x ∙ (sym r)) id) ∙ sym (assoc q r (sym r)) ∙ (λ j → q ∙ (rCancel r j)) ∙ sym (rUnit q) + +switcher : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (p : a ≡ b) (q r : a ≡ c) → p ∙ (sym p) ≡ q ∙ (sym r) → r ≡ q +switcher {ℓ} {A} {a} {b} {c} p q r id = (lUnit r) ∙ ((λ j → ((rCancel p (~ j))) ∙ r)) ∙ (cong (λ x → x ∙ r) id) ∙ sym (assoc q (sym r) r) ∙ (λ j → q ∙ (lCancel r j)) ∙ sym (rUnit q) + + +Lemma296b : ∀{ℓ ℓ' ℓ''} {X : Type ℓ} {x y : X} {A : X → Type ℓ'} {B : X → Type ℓ''} + (p : x ≡ y) + (f : (A x) → (B x)) + (g : (A y) → (B y)) → + (transport (λ i → (A (p i)) → (B (p i))) f ≡ g) ≡ ((a : A y) → transport (λ i → B (p (~ i) )) (g a) ≡ f (transport (λ i → A (p (~ i))) a)) +Lemma296b {ℓ = ℓ} {X = X} {x = x} {y = y} {A = A} {B = B} = J {ℓ} {X} {x} (λ y p → (f : (A x) → (B x)) (g : (A y) → (B y)) → + (transport (λ i → (A (p i)) → (B (p i))) f ≡ g) ≡ + ((a : A y) → transport (λ i → B (p (~ i))) (g a) ≡ f (transport (λ i → A (p (~ i))) a))) + λ f g → transport (λ i → helper1 f g (~ i)) (helper2 f g) + where + helper1 : (f : (A x) → (B x)) (g : (A x) → (B x)) → + ((transport (λ i → A (refl {x = x} i) → B (refl {x = x} i)) f ≡ g) ≡ ((a : A x) → transport (λ i → B (refl {x = x} (~ i))) (g a) ≡ f (transport (λ i → A (refl {x = x} (~ i))) a))) + ≡ ((f ≡ g) ≡ ((a : A x) → g a ≡ f a)) + helper1 f g i = ((transportRefl f i ≡ g) ≡ ((a : A x) → transportRefl (g a) i ≡ f (transportRefl a i))) + + helper2 : (f : (A x) → (B x)) (g : (A x) → (B x)) → ((f ≡ g) ≡ ((a : A x) → g a ≡ f a)) + helper2 f g = isoToPath (iso (λ p a → sym (funExt⁻ p a)) (λ p → sym (funExt p) ) (λ x → refl) λ x → refl) + +Lemma296 : ∀{ℓ ℓ' ℓ''} {X : Type ℓ} {x y : X} {A : X → Type ℓ'} {B : X → Type ℓ''} + (p : x ≡ y) + (f : (A x) → (B x)) + (g : (A y) → (B y)) → + (transport (λ i → (A (p i)) → (B (p i))) f ≡ g) ≡ ((a : A x) → transport (λ i → B (p i)) (f a) ≡ g (transport (λ i → A (p i)) a)) +Lemma296 {ℓ = ℓ} {X = X} {x = x} {y = y} {A = A} {B = B} = J {ℓ} {X} {x} + (λ y p → (f : (A x) → (B x)) (g : (A y) → (B y)) → + (transport (λ i → (A (p i)) → (B (p i))) f ≡ g) ≡ + ((a : A x) → transport (λ i → B (p i)) (f a) ≡ g (transport (λ i → A (p i)) a))) + (λ f g → transport (λ i → reflTransport f g (~ i)) (reflCase f g)) + where + reflTransport : (f g : A x → B x) → ((transport (λ i → A (refl {x = x} i) → B (refl {x = x} i)) f ≡ g) ≡ ((a : A x) → transport (λ i → B (refl {x = x} i)) (f a) ≡ g (transport (λ i → A (refl {x = x} i)) a))) ≡ ((f ≡ g) ≡ ((a : A x) → f a ≡ g a)) + reflTransport f g j = (transportRefl f j ≡ g) ≡ ((a : A x) → transportRefl (f a) j ≡ g (transportRefl a j)) + + reflCase : (f g : A x → B x) → ((f ≡ g) ≡ ((a : A x) → f a ≡ g a)) + reflCase f g = isoToPath (iso (λ p → funExt⁻ p) (λ p → funExt p) (λ x → refl) λ x → refl) -{- Test -} -SigmaEqHelper : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) (b : B) (p1 p2 : Σ A (λ a → (f a) ≡ b)) → (fst p1 ≡ fst p2) → p1 ≡ p2 -SigmaEqHelper {ℓ = ℓ} {ℓ' = ℓ'} {A = A} {B = B} f b p1 p2 path = J {ℓ'} {B} {f (fst p1)} (λ b pf → ( (pair2 : Σ A (λ a → (f a) ≡ b)) → (fst p1) ≡ (fst pair2) → ((fst p1 , pf) ≡ pair2) ) ) (λ pr id → ΣPathP (id , {!transport (λ i → f (id i) ≡ ((snd pr) i))!})) {y = b} (snd p1) (p2) path {- J {ℓ} {A} {fst p1} {ℓ-max ℓ ℓ'} (λ y pf → ( (s : (f y ≡ b)) → (p1 ≡ (y , s)) )) - (λ s → ΣPathP (refl , {!J {ℓ'} {B} {x = (f (fst p1))} {ℓ'} (λ c → )!})) - {fst p2} path (snd p2) -} {- TODO: Prove Kₙ ≡ Ω Kₙ₊₁ -} @@ -83,119 +99,255 @@ invPathCancel2 {a = a} {b = b} p = J {x = a} (λ y p → (sym p) ∙ p ≡ refl cancelReflMid : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) (q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q cancelReflMid {ℓ = ℓ}{A = A} {a = a} {b = b} p q = J {ℓ} {A} {a} {ℓ} (λ b p → ((q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q)) (λ r → sym (lUnit (refl ∙ r ))) {b} p q -φ' : (a : A) → A → typ (Ω (Susp A , north )) -φ' x = (λ a → ((merid a) ∙ sym ((merid x)))) -φ'' : (b : Susp A) → b ≡ south → A → north ≡ b -φ'' b p a = (merid a) ∙ (sym p) ---φ''' : (a : A) (b : Susp A) → A → north ≡ b ---φ''' (a : A) b p = (merid a) ∙ (sym p) +inducedFun : ∀{ℓ''} {k : ℕ₋₂} + (f : A → B) → + (P : (B → HLevel ℓ'' (2+ k))) → + (((b : B) → fst (P b)) → (a : A) → fst (P (f a)) ) +inducedFun f P = λ g a → g (f a) +Lemma757i→ii : ∀ {ℓ} (f : A → B) (n : ℕ₋₂) → + is- n -Connected f → + (P : B → HLevel ℓ (2+ n)) → + isEquiv (inducedFun f P) +Lemma757i→ii f n isCon P = {!!} + +Lemma757i→iii : ∀ {ℓ} (f : A → B) (n : ℕ₋₂) → + is- n -Connected f → + (P : B → HLevel ℓ (2+ n)) → + hasSection (inducedFun f P) +Lemma757i→iii f n isCon P = {!!} , (λ b → {!!}) + +Lemma861* : ∀{ℓ} (n k : ℕ₋₂) (f : A → B) → + (n ≡ neg2 → ⊥) → (k ≡ neg2 → ⊥) → + (n ≡ (suc neg2) → ⊥) → (k ≡ (suc neg2) → ⊥) → + (is- n -Connected f) → + (P : B → HLevel ℓ (2+ k)) → + Σ ℕ₋₂ (λ a → (a +₋₂ n ≡ k) × ((a ≡ neg2 → ⊥) × (a ≡ (suc neg2) → ⊥))) → + is- ((k -₋₂ n) -₋₂ (suc (suc (suc (suc neg2))))) -Truncated (inducedFun {k = k } f P) +Lemma861* n neg2 f n2 k2 n1 k1 isCon P pair = ⊥-elim (k2 refl) +Lemma861* n (suc neg2) f n2 k2 n1 k1 isCon P pair = ⊥-elim (k1 refl) +Lemma861* neg2 (suc (suc k)) f n2 k2 n1 k1 isCon P pair = ⊥-elim (n2 refl) +Lemma861* (suc neg2) (suc (suc k)) f n2 k2 n1 k1 isCon P pair = ⊥-elim (n1 refl) +Lemma861* (suc (suc n)) (suc (suc k)) f n2 k2 n1 k1 isCon P (neg2 , (x , (x₁ , x₂))) = ⊥-elim (x₁ refl) +Lemma861* (suc (suc n)) (suc (suc k)) f n2 k2 n1 k1 isCon P (suc neg2 , (x , (x₁ , x₂))) = ⊥-elim (x₂ refl) +Lemma861* (suc (suc n)) (suc (suc k)) f n2 k2 n1 k1 isCon P (suc (suc neg2) , (pf , (x₁ , x₂))) = {!pf!} +Lemma861* (suc (suc n)) (suc (suc k)) f n2 k2 n1 k1 isCon P (suc (suc (suc dif)) , (pf , (x₁ , x₂))) = {!isCon!} + + +Lemma861 : ∀{ℓ} (n k : ℕ) (f : A → B) → + (is- (ℕ→ℕ₋₂ n) -Connected f) → + (P : B → HLevel ℓ (suc (suc k))) → + Σ ℕ (λ a → a + n ≡ k) → + is- (ℕ→ℕ₋₂ ((k - n) - 2)) -Truncated (inducedFun f P) +Lemma861 {A = A} {B = B} n k f isCon P (dif , pf) = {!!} where + + helper1 : (l : (a : A) → fst (P (f a))) (a b : Σ ((b : B) → fst (P b)) λ g → (a : A) → g (f a) ≡ l a ) → (a ≡ b) ≡ Σ (fst a ≡ fst b) λ r → PathP (λ i → (x : A) → (r i) (f x) ≡ l x) (snd a) (snd b) + helper1 l a b = sym (ua (Σ≡ )) + + helper2 : (l : (a : A) → fst (P (f a))) (a b : Σ ((b : B) → fst (P b)) λ g → (a : A) → g (f a) ≡ l a ) → (Σ (fst a ≡ fst b) λ r → PathP (λ i → (x : A) → (r i) (f x) ≡ l x) (snd a) (snd b)) ≡ Σ ((x : B) → (fst a) x ≡ (fst b) x) λ r → PathP (λ i → (x : A) → r (f x) i ≡ l x) (snd a) (snd b) + helper2 l (g , p) (h , q) = isoToPath (iso lrFun rlFun (λ b → refl) λ b → refl) where + lrFun : (Σ (g ≡ h) λ r → PathP (λ i → (x : A) → (r i) (f x) ≡ l x) p q) → Σ ((x : B) → g x ≡ h x) λ r → PathP (λ i → (x : A) → r (f x) i ≡ l x) p q + lrFun (a , b) = funExt⁻ a , b + rlFun : (Σ ((x : B) → g x ≡ h x) λ r → PathP (λ i → (x : A) → r (f x) i ≡ l x) p q) → (Σ (g ≡ h) λ r → PathP (λ i → (x : A) → (r i) (f x) ≡ l x) p q) + rlFun (a , b) = (funExt a) , b + abstract + helper3 : (l : (a : A) → fst (P (f a))) (a b : Σ ((b : B) → fst (P b)) λ g → (a : A) → g (f a) ≡ l a ) → (r : (x : B) → (fst a) x ≡ (fst b) x) → (PathP (λ i → (x : A) → r (f x) i ≡ l x) (snd a) (snd b)) ≡ ((x : A) → r (f x) ≡ ((snd a) x) ∙ sym ((snd b) x) ) + helper3 l (g , p) (h , q) r j = hcomp (λ k → λ{ (j = i0) → (PathP (λ i → (x : A) → r (f x) (i ∧ i1) ≡ l x) p λ x → (sym (lUnit (q x)) k )) + ; (j = i1) → throwAbout k}) + ((PathP (λ i → (x : A) → r (f x) (i ∧ ~ j) ≡ l x) p λ x → (λ i → r (f x) (~ j ∨ i)) ∙ (q x))) + where + postulate + throwAbout : (p ≡ λ x → r (f x) ∙ q x) ≡ ((x : A) → (((r (f x))) ≡ (p x) ∙ (sym (q x)))) + {- Very heavy proof + throwAbout : (p ≡ λ x → r (f x) ∙ q x) ≡ ((x : A) → (((r (f x))) ≡ (p x) ∙ (sym (q x)))) + throwAbout = isoToPath (iso (λ g x → transport (λ i → throwAboutHelper (p x) (r (f x)) (sym (q x)) i ) (funExt⁻ g x) ) + (λ g → funExt λ x → transport (λ i → throwAboutHelper (p x) (r (f x)) (sym (q x)) (~ i) ) (g x)) + (λ b → funExt λ x → (λ j → transport (λ i → throwAboutHelper (p x) (r (f x)) (λ i₁ → q x (~ i₁)) (i ∨ j)) (transport (λ i₁ → throwAboutHelper (p x) (r (f x)) (λ i₂ → q x (~ i₂)) ((~ i₁) ∨ ( j))) (b x))) ∙ λ j → transportRefl (transportRefl (b x) j) j ) + λ b → (λ j → funExt (λ x → transport (λ i → throwAboutHelper (p x) (r (f x)) (λ i₁ → q x (~ i₁)) (~ i ∧ (~ j))) + (transport (λ i → throwAboutHelper (p x) (r (f x)) (λ i₁ → q x (~ i₁)) (i ∧ (~ j))) + (λ i → b i x)))) + ∙ (λ j → funExt (λ x → transportRefl ((transportRefl (λ i → b i x)) j) j ) )) where + throwAboutHelper : ∀{ℓ} {A : Type ℓ} {a b c : A} → (p : a ≡ b) (q : a ≡ c) (r : b ≡ c) → (p ≡ q ∙ (sym r) ) ≡ (q ≡ p ∙ r) + throwAboutHelper {ℓ} {A} {a} {b} {c} = J {ℓ} {A} {a} (λ b p → (q : a ≡ c) (r : b ≡ c) → (p ≡ q ∙ (sym r) ) ≡ (q ≡ p ∙ r)) + (J {ℓ} {A} {a} (λ c q → (r : a ≡ c) → (refl ≡ q ∙ (sym r) ) ≡ (q ≡ refl ∙ r) ) + λ r → (λ i → refl ≡ lUnit (sym r) (~ i)) ∙ isoToPath (iso (λ id → cong (λ x → sym x) id) (λ id → cong (λ x → sym x) id ) (λ b → refl) λ b → refl) ∙ λ i → (refl ≡ lUnit r i) ) + -} -looper : {a : A} -- defining loopⁿ - (n : Int) → - (a ≡ a) → - (a ≡ a) -looper {A = A} {a = a} (pos zero) p = refl -looper {A = A} {a = a} (pos (suc n)) p = (looper (pos n) p) ∙ p -looper {A = A} {a = a} (negsuc zero) p = sym p -looper {A = A} {a = a} (negsuc (suc n)) p = (sym p) ∙ (looper (negsuc n) p) + helper4 : (l : (a : A) → fst (P (f a))) (a b : Σ ((b : B) → fst (P b)) λ g → (a : A) → g (f a) ≡ l a ) → (r : (x : B) → (fst a) x ≡ (fst b) x) → ((x : A) → r (f x) ≡ ((snd a) x) ∙ sym ((snd b) x) ) ≡ ((λ (x : A) → r (f x)) ≡ (λ (x : A) → ((snd a) x) ∙ sym ((snd b) x))) + helper4 l (h , q) (g , p) r = isoToPath (iso (λ p → funExt p) (λ p → funExt⁻ p) (λ b → refl) λ b → refl) + FINAL : (l : (a : A) → fst (P (f a))) (a b : Σ ((b : B) → fst (P b)) λ g → (a : A) → g (f a) ≡ l a ) → (a ≡ b) ≡ Σ ((x : B) → (fst a) x ≡ (fst b) x) (λ r → ((λ (x : A) → r (f x)) ≡ (λ (x : A) → ((snd a) x) ∙ sym ((snd b) x)))) + FINAL l (g , p) (h , q) = helper1 _ _ _ ∙ helper2 _ _ _ ∙ (λ j → Σ ((x : B) → g x ≡ h x) (λ r → helper3 l (g , p) (h , q) r j)) ∙ λ j → Σ ((x : B) → g x ≡ h x) (λ r → helper4 l (g , p) (h , q) r j) -φ : (a : A) → A → (north {A = A} ≡ north {A = A}) -φ x = (λ a → ((merid a) ∙ sym ((merid x)))) + + +{- +Lemma862 : (A : Pointed ℓ) (B : Pointed ℓ') (n m : ℕ₋₂) → + (is- (suc (suc n)) -ConnectedType (typ A)) → + (is- (suc (suc m)) -ConnectedType (typ B)) → + (P : (typ A) → (typ B) → HLevel (ℓ-max ℓ ℓ') (2+ ((suc (suc n)) +₋₂ (suc (suc m))))) → + (f : ((a : (typ A)) → fst (P a (pt B)))) → + (g : ((b : (typ B)) → fst (P (pt A) b))) → + (p : f (pt A) ≡ g (pt B)) → + Σ ((a : typ A) → (b : typ B) → fst (P a b)) λ h → Σ (((a : typ A) → h a (pt B) ≡ f a) × ((b : typ B) → h (pt A) b ≡ g b) ) λ pair → p ≡ sym ((proj₁ pair) (pt A) ) ∙ (proj₂ pair) (pt B) +Lemma862 {ℓ = ℓ} {ℓ' = ℓ'} (A , a₀) (B , b₀) n m conA conB P f g p = (λ a b → {!!}) , ({!!} , {!!}) where + Q : A → Type (ℓ-max ℓ ℓ') + Q a = Σ ((b : B) → fst (P a b )) λ k → f a ≡ k (b₀) -φ* : (A : Pointed ℓ) → A →* Ω ((Susp (typ A)) , north) -φ* A = (φ (pt A)) , invPathCancel (merid (pt A)) + a₀-map : Unit → A + a₀-map x = a₀ -σ : (n : ℕ₋₂) → (typ (coHomK-ptd n)) → (typ (Ω (coHomK-ptd (suc n)))) -σ neg2 k = looper k ( cong {B = λ _ → (∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1)} - (λ x → ∣ x ∣) - ((merid north) ∙ (sym (merid south))) ) -σ (suc n) x = rec {n = suc (suc (suc n))} - {B = (typ (Ω (coHomK-ptd (suc (suc n)))))} - (isOfHLevel∥∥ {A = S₊ (2+ suc (suc n))} (suc (suc (suc (suc n)))) ∣ north ∣ ∣ north ∣) - (λ y → cong {B = λ _ → Null (S (1+₋₂ (suc (suc (suc (suc n)))))) (S₊ (2+ (suc (suc n))))} - (λ z → ∣ z ∣) - (φ north y)) - x + conOfa₀-map : is- (suc n) -Connected a₀-map + conOfa₀-map = {!!} -testing : (n : ℕ) → (a : typ (Ω (S₊ (suc n) , north))) → isContr (∥ (Σ (S₊ n) λ k → (φ' {A = (S₊ n)} north k ) ≡ a) ∥ (ℕ→ℕ₋₂ n) ) -testing zero a = {!!} -testing (suc zero) a = {!!} -testing (suc (suc n)) a = ∣ {!!} , {!a!} ∣ , λ y → {!!} + b₀-map : Unit → B + b₀-map x = b₀ + conOfb₀-map : is- (suc n) -Connected b₀-map + conOfb₀-map = {!!} --- φ : (a : A) → A → (north {A = A} ≡ north {A = A}) --- φ x = (λ a → ((merid a) ∙ sym ((merid x)))) + conOfQ : (a : A) → isOfHLevel (2+ (suc n)) (Q a) + conOfQ = {!!} -isConnectedT : (n : ℕ₋₂) (A : Type ℓ) → Type ℓ -isConnectedT n A = isContr (∥ A ∥ n) + typesQ : (a : A) → Σ ((a : A) → Q a) λ ℓ → ℓ a₀ ≡ (g , p) + typesQ a = (fst (Lemma757i→iii a₀-map (suc n) conOfa₀-map (λ a → (Q a , conOfQ a ))) (λ x → (g , p))) , + cong (λ f → f tt) (snd (Lemma757i→iii a₀-map (suc n) conOfa₀-map (λ a → (Q a , conOfQ a ))) (λ x → (g , p))) -isConnectedF : ∀{ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ₋₂) → (A → B) → Type ((ℓ-max ℓ ℓ')) -isConnectedF {A = A} {B = B} n f = ((b : B) → isContr (∥ fiber f b ∥ n)) + QisFib : (a : A) → Q a ≡ fiber ( λ x → inducedFun {A = Unit} {B = B} b₀-map (P a) x tt) (f a) + QisFib a = {!!} -inducedFun : ∀{ℓ''} {k : ℕ₋₂} - (f : A → B) → - (P : (B → HLevel ℓ'' (2+ k))) → - (((b : B) → fst (P b)) → (a : A) → fst (P (f a)) ) -inducedFun f P = λ g a → g (f a) + Ok : (a : A) → is- (suc n) -ConnectedType (fiber ( λ x → inducedFun {A = Unit} {B = B} b₀-map (P a) x tt) (f a)) + Ok a = {!!} -Lemma861 : ∀{ℓ''} (n k : ℕ₋₂) → - Σ ℕ₋₂ (λ num → (suc (suc n)) +₋₂ num ≡ (suc (suc k)) ) → - (f : A → B) → - (isConnectedF (suc (suc n)) f) → - (P : (B → HLevel ℓ'' (2+ (suc (suc k))))) → - isConnectedF (((suc (suc n)) -₋₂ (suc (suc k))) -₋₂ (suc (suc neg2))) (inducedFun {k = (suc (suc k))} f P) -Lemma861 n k (neg2 , pf) f conf P b = {!fiber!} -Lemma861 {A = A} {B = B} n k (suc j , pf) f conf P g = {!!} where - helper : (l : (x : A) → (fst (P (f x)) )) → - (fiber (inducedFun {k = (suc (suc k))} f P) l) ≡ Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a) - helper l = isoToPath (iso (fun) finv (λ b → refl) λ b → refl) where - fun : (fiber (inducedFun {k = (suc (suc k))} f P) l) → Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a) - fun (fs , sn) = fs , λ a → cong (λ x → x a) sn +σ : A → {a : A} → typ (Ω ((Susp A) , north)) +σ x {a} = merid x ∙ sym (merid a) - finv : (Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → (fiber (inducedFun {k = (suc (suc k))} f P) l) - finv (fs , sn) = fs , (funExt sn) - helper'3 : (l : (x : A) → (fst (P (f x)) )) - (gp hq : Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → - (gp ≡ hq) - ≡ Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) - helper'3 l gp hq = isoToPath (iso fun finv (λ b i → ({!!})) {!!}) where - fun : (gp ≡ hq) → Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) - fun path = (λ x i → fst (path i) x ) , funExt λ x → pathFinal {x} where - pathHelper : {x : A} → PathP (λ i → fst gp (f x) ≡ fst (path i) (f x)) refl (snd gp x ∙ sym (snd hq x)) - pathHelper {x = x} i = hcomp (λ j → λ{ (i = i0) → rCancel (snd gp x) j ; (i = i1) → (snd gp x ∙ sym (snd hq x)) } ) (snd gp x ∙ sym (snd (path i) x)) +code : ∀{ℓ} {A : Type ℓ} {a : A} (n : ℕ₋₂) → is- (suc (suc n)) -ConnectedType A → (y : Susp A) → (north ≡ y → Type ℓ) +code {A = A} {a = a} n iscon north p = (∥ fiber (λ y → σ y {a}) p ∥ ((suc (suc n)) +₋₂ (suc (suc n)))) +code {ℓ} {A = A} {a = a} n iscon south q = ∥ fiber merid q ∥ ((suc (suc n)) +₋₂ (suc (suc n))) +code {ℓ} {A = A} {a = a} n iscon (merid c i) = pathToConstruct i where + pathToConstruct : PathP (λ i → north ≡ merid c i → Type ℓ) (code {a = a} n iscon north) (code {a = a} n iscon south) + pathToConstruct = toPathP (transport (λ i → transpId (~ i)) λ q → sym (ua ((RlFun q) , (RlFunEquiv q)))) + where - pathHelper2 : {x : A} → PathP (λ i → fst gp (f x) ≡ fst (path (~ i)) (f x)) (λ i → fst (path i) (f x)) refl - pathHelper2 {x = x} i = λ i₁ → fst (path (i₁ ∧ (~ i))) (f x) + transpId : (transport (λ i₁ → north ≡ merid c i₁ → Type ℓ) (code {a = a} n iscon north) ≡ code {a = a} n iscon south) ≡ + ((q : north ≡ south) → (code {a = a} n iscon south q) ≡ (code {a = a} n iscon north (q ∙ sym (merid c)))) + transpId = (transport (λ i₁ → north ≡ merid c i₁ → Type ℓ) (code {a = a} n iscon north) ≡ code {a = a} n iscon south) + ≡⟨ Lemma296b {A = λ x → north ≡ x} {B = λ _ → Type ℓ} (merid c) (code {a = a} n iscon north) (code {a = a} n iscon south) ⟩ + ((q : north ≡ south) → code {a = a} n iscon south q ≡ code {a = a} n iscon north (transport (λ i → north ≡ (merid c (~ i))) q)) + ≡⟨ (λ i → ((q : north ≡ south) → code {a = a} n iscon south q ≡ code {a = a} n iscon north (helper q (sym (merid c)) i))) ⟩ + ((q : north ≡ south) → (code {a = a} n iscon south q) ≡ (code {a = a} n iscon north (q ∙ sym (merid c)))) ∎ + + where + helper : ∀{ℓ} {A : Type ℓ} {a b c : A} (q : a ≡ b) (p : b ≡ c) → + (transport (λ i → a ≡ (p i)) q) ≡ (q ∙ p) + helper {ℓ = ℓ} {A = A} {a = a} {b = b} {c = c} q = J {ℓ} {A} {b} + (λ c p → (transport (λ i → a ≡ (p i)) q) ≡ (q ∙ p)) + (transportRefl q ∙ rUnit q) - pathFinal : {x : A} → (λ i → fst (path i) (f x)) ≡ snd gp x ∙ sym (snd hq x) - pathFinal {x = x} = transport (λ i → PathP (λ j₁ → (lCancel (λ i → fst gp (f x) ≡ fst (path i) (f x))) i j₁) - (λ i → fst (path i) (f x)) - (snd gp x ∙ sym (snd hq x))) - (compPathP pathHelper2 pathHelper) + - finv : Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) → (gp ≡ hq) - finv (fs , sn) = ΣPathP {x = gp} {y = hq} (funExt (λ b → fs b) , {! !}) -- (λ b → (fs b) i) , λ a → hcomp (λ k → λ{(i = i0) → (snd gp a) ; (i = i1) → {!!}}) (transp (λ j → fs (f a) (i ∧ j) ≡ l a) (~ i) ((snd gp) a)) + RlFun : {c : A} (q : north ≡ south) → code {a = a} n iscon north (q ∙ sym (merid c)) → code {a = a} n iscon south q + RlFun {c = c} q x = (funHelper a) x where - potTerm : {a : A} → PathP (λ i → (a : A) → funExt (λ b → fs b) i (f a) ≡ l a) (snd gp) (snd hq) - potTerm {a = a} j = λ a → hcomp {!(transp (λ i → funExt (λ b → fs b) (i) (f a) ≡ l a) i0 ((snd gp a))) !} (transp (λ i → funExt (λ b → fs b) (i ∧ j) (f a) ≡ l a) i0 ((snd gp a))) - - helper2 : {a : A} → (fs (f a)) ∙ (snd hq a) ≡ snd gp a - helper2 {a = a} = (cong (λ x → x ∙ (snd hq a)) (cong (λ x → x a) sn)) ∙ sym (assoc (snd gp a) (sym (snd hq a)) (snd hq a)) ∙ (λ j → (snd gp a) ∙ (lCancel (snd hq a)) j) ∙ sym (rUnit (snd gp a)) - + funHelper : (a : A) → (∥ Σ A (λ x₁ → merid x₁ ∙ (λ i₁ → merid a (~ i₁)) ≡ q ∙ sym (merid c)) ∥ (suc (suc n) +₋₂ suc (suc n))) → ∥ Σ A (λ x → merid x ≡ q) ∥ ((suc (suc n) +₋₂ suc (suc n))) + funHelper a = ind (λ x → isOfHLevel∥∥ _) λ r → sufMap (fst r) (snd r) + where + + sufMap : (x₂ : A) → (merid x₂ ∙ sym (merid a)) ≡ (q ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ q) ∥ ((suc (suc n) +₋₂ suc (suc n))) + sufMap x₂ = fst (Lemma862 (A , a) (A , a) n n iscon iscon (λ x₂ c → (((merid x₂ ∙ sym (merid a)) ≡ (q ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ q) ∥ ((suc (suc n) +₋₂ suc (suc n))) ) , isOfHLevelPi (2+ ((suc (suc n) +₋₂ suc (suc n)))) λ _ → isOfHLevel∥∥ _)) firstFun secondFun (funExt (λ x → λ j → ∣ (refl j) , (maybe (merid a) q (canceller (merid a) q (sym (merid a)) x ) x j) ∣ ))) x₂ c + where + + + + firstFun : (a₁ : A) → merid a₁ ∙ (λ i₁ → merid a (~ i₁)) ≡ q ∙ (λ i₁ → merid (pt (A , a)) (~ i₁)) → ∥ Σ A (λ x₁ → merid x₁ ≡ q) ∥ (suc (suc n) +₋₂ suc (suc n)) + firstFun x r = ∣ x , canceller (merid x) q (sym (merid a)) r ∣ + + {-J {ℓ} {A} {b} (λ d r → ((p ∙ r) ≡ (q ∙ r)) ≡ (p ≡ q)) λ i → sym (rUnit p) i ≡ sym (rUnit q) i-} + secondFun : (b : typ (A , a)) → merid (pt (A , a)) ∙ (λ i₁ → merid a (~ i₁)) ≡ q ∙ (λ i₁ → merid b (~ i₁)) → ∥ Σ A (λ x₁ → merid x₁ ≡ q) ∥ (suc (suc n) +₋₂ suc (suc n)) + secondFun b s = ∣ b , switcher (merid a) q (merid b) s ∣ + + {- need eckmann hilton -} + + + maybe : ∀ {ℓ} {A : Type ℓ} {a b : A} → (p q : a ≡ b) → (p ≡ q) → (P : p ∙ (sym p) ≡ q ∙ (sym p)) → canceller p q (sym p) P ≡ switcher p q p P + maybe {ℓ} {A} {a} {b} p q = J {ℓ} {a ≡ b} {p} (λ q _ → (P : p ∙ (sym p) ≡ q ∙ (sym p)) → canceller p q (sym p) P ≡ switcher p q p P) (J {ℓ} {A} {a} (λ b p → (P : p ∙ (sym p) ≡ p ∙ (sym p)) → canceller p p (sym p) P ≡ switcher p p p P) (λ P → {!refl!}) p) + + + canceller2 : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (p q : a ≡ b) (r : b ≡ c) → p ∙ r ≡ q ∙ r → p ≡ q + canceller2 {ℓ} {A} {a} {b} {c} p q r id = (rUnit p ∙ (λ j → (p ∙ (rCancel r (~ j)))) ∙ assoc p r (sym r)) ∙ (cong (λ x → x ∙ (sym r)) id) ∙ sym (assoc q r (sym r)) ∙ (λ j → q ∙ (rCancel r j)) ∙ sym (rUnit q) + + RlFunEquiv : {c : A} → (q : north ≡ south) → isEquiv (RlFun {c = c} q) + RlFunEquiv {c = c} q = fst (Lemma757i→iii {A = Unit} {B = A} (λ x → a) ((suc (suc n))) (λ b → {!fst iscon!} , {!!}) (λ a → ((isEquiv (RlFun {c = a} q)) , {!!} ))) (λ t → baseCase) c where + baseCase : isEquiv (RlFun {c = a} q) + baseCase = snd (isoToEquiv (iso (RlFun {c = a} q) (ind (λ _ → isOfHLevel∥∥ _) λ a → {!!}) {!!} {!!})) + + claim : RlFun {c = a} q ≡ ind (λ _ → isOfHLevel∥∥ _) λ b → ∣ (fst b) , canceller2 _ _ _ (snd b) ∣ + claim = {!!} --cannot say before I have prev lemmas + + isCona₀ : is- (suc (suc n)) -Connected (λ (x₁ : Unit) → a) + isCona₀ = λ b → {!!} , {!!} + + +Lemma8610 : ∀{ℓ ℓ' ℓ''} {A : Type ℓ} {a1 a2 : A} {B : A → Type ℓ'} (C : (a : A) → (B a → Type ℓ'')) (p : a1 ≡ a2) (b : B a2) → transport (λ i → {!uncurry C ?!}) {!!} ≡ {!(ua ?)!} +Lemma8610 = {!!} + +Thm864 : (n : ℕ₋₂) (a : A) (iscon : is- (suc (suc n)) -ConnectedType A) {y : Susp A} → (p : north ≡ y) → isContr (code {a = a} n iscon y p) +Thm864 {ℓ} {A} n a iscon = J {ℓ} {Susp A} {north} (λ y p → (isContr (code {a = a} n iscon y p))) (∣ a , (rCancel _) ∣ , (λ y → {!refl!})) where + center : {!!} + center = {!!} + +Freudenthal : (n : ℕ₋₂ ) (A : Pointed ℓ) → is- (suc (suc n)) -ConnectedType (typ A) → ∥ typ A ∥ (((suc (suc n) +₋₂ suc (suc n))) ) ≡ ∥ typ (Ω (Susp (typ A) , north)) ∥ (((suc (suc n) +₋₂ suc (suc n))) ) +Freudenthal n A isCon = ua ({!!} , {!Lemma757i→ii ? ? ? ?!}) + + + +-} + + + + + + + + + + + + + + + + + + +{- J {ℓ} {A} {a} (λ b p → ((q : a ≡ b) (P : p ∙ (sym p) ≡ q ∙ (sym p)) → canceller p q (sym p) P ≡ switcher p q p P)) λ q id → (λ i → canceller refl q ((sym symRefl ) (~ i)) id) ∙ {!!} ∙ λ j → switcher refl q (symRefl (~ j)) id where + test : {q : a ≡ a} {id : refl ∙ (λ i₁ → refl (~ i₁)) ≡ q ∙ (λ i₁ → refl {x = a} (~ i₁))} → canceller refl q refl id ≡ (sym (rCancel refl)) ∙ (cong (λ x → x ∙ refl) ((lUnit refl) ∙ id)) ∙ (sym (assoc q refl refl)) ∙ (λ j → q ∙ (rCancel refl j)) ∙ (sym (rUnit q)) + test {q = q} {id = id} i = {!(sym (rCancel refl)) ∙ (cong (λ x → x ∙ (symRefl (~ i))) ((lUnit refl) ∙ id)) ∙ (sym (assoc q refl (symRefl (~ i)))) ∙ (λ j → q ∙ (rCancel refl j)) ∙ (sym (rUnit q))!} -} + +{- + + RlFun : (q : north ≡ south) → (∥ ((fiber (λ y → σ y) (q ∙ sym (merid c))) ≡ (Σ A (λ x₁ → merid x₁ ∙ (λ i₁ → merid a (~ i₁)) ≡ q ∙ sym (merid c)))) ∥ ((suc n) +₋₂ suc (suc n))) → (∥ (fiber merid q) ≡ (Σ A (λ x₁ → merid x₁ ≡ q)) ∥ ((suc n) +₋₂ suc (suc n))) + RlFun q x = ind (λ x → isOfHLevel∥∥ _) ({!cong funHelper!}) x where + funHelper : (a : A) → merid a ∙ (sym (merid c)) ≡ q ∙ (sym (merid c)) → ∥ Σ A (λ x → merid x ≡ q) ∥ ((suc (suc n) +₋₂ suc (suc n))) + funHelper a = {!!} + + LrFun : ((q : north ≡ south) → (code {a = a} n iscon south q) → (code {a = a} n iscon north (q ∙ sym (merid c)))) + LrFun = {!!} + +-} -- fst gp (f x) ≡ fst hq (f x) -- fst gp (f x) ≡ fst gp (f x) - +{- helper2 : (l : (x : A) → (fst (P (f x)) )) (gp hq : Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → @@ -226,8 +378,75 @@ Lemma861 {A = A} {B = B} n k (suc j , pf) f conf P g = {!!} where helper3 : (l : (x : A) → (fst (P (f x)) )) (gp hq : Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → (gp ≡ hq) ≡ Σ (fst gp ≡ fst hq) (λ a≡ → PathP (λ i → (a : A) → a≡ i (f a) ≡ l a) (snd gp) (snd hq)) helper3 l gp hq = sym (ua (Σ≡ {A = ((a : B) → fst (P a))} {B = (λ x → (a : A) → x (f a) ≡ l a)} {x = gp} {y = hq})) + + +-} + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + +{- + + + + + + + + + + + + + + + + + + + ------------------ @@ -658,3 +877,62 @@ inv1 {A = A} {B = B}(suc n) {f = f} g j = {! -- transp (λ i → (snd (rlFun3* n {!(cong (fst (rlFun3* n)) (funExt ( (λ x → (funExt⁻ (cong (λ x → (fst (lrFun3* n {f}) x )) g) x))))) !} maybe j = {!fst (rlFun3* n) (λ x → snd (lrFun3* n {f}) i1 x)!} -} +-} + + + + + + + + + +{- +Lemma861 : ∀{ℓ''} (n k : ℕ₋₂) → + Σ ℕ₋₂ (λ num → (suc (suc n)) +₋₂ num ≡ (suc (suc k)) ) → + (f : A → B) → + (is- (suc (suc n)) -Connected f) → + (P : (B → HLevel ℓ'' (2+ (suc (suc k))))) → + is- (((suc (suc n)) -₋₂ (suc (suc k))) -₋₂ (suc (suc neg2))) -Connected (inducedFun {k = (suc (suc k))} f P) +Lemma861 n k (neg2 , pf) f conf P b = {!!} + where + helper : (((suc (suc n)) -₋₂ (suc (suc k))) -₋₂ (suc (suc neg2))) ≡ neg2 + helper = {!!} +Lemma861 {A = A} {B = B} n k (suc j , pf) f conf P g = {!!} where + helper : (l : (x : A) → (fst (P (f x)) )) → + (fiber (inducedFun {k = (suc (suc k))} f P) l) ≡ Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a) + helper l = isoToPath (iso (fun) finv (λ b → refl) λ b → refl) where + fun : (fiber (inducedFun {k = (suc (suc k))} f P) l) → Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a) + fun (fs , sn) = fs , λ a → cong (λ x → x a) sn + + finv : (Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → (fiber (inducedFun {k = (suc (suc k))} f P) l) + finv (fs , sn) = fs , (funExt sn) + + helper'3 : (l : (x : A) → (fst (P (f x)) )) + (gp hq : Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → + (gp ≡ hq) + ≡ Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) + helper'3 l gp hq = isoToPath (iso fun finv (λ b i → ({!!})) {!!}) where + fun : (gp ≡ hq) → Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) + fun path = (λ x i → fst (path i) x ) , funExt λ x → pathFinal {x} where + pathHelper : {x : A} → PathP (λ i → fst gp (f x) ≡ fst (path i) (f x)) refl (snd gp x ∙ sym (snd hq x)) + pathHelper {x = x} i = hcomp (λ j → λ{ (i = i0) → rCancel (snd gp x) j ; (i = i1) → (snd gp x ∙ sym (snd hq x)) } ) (snd gp x ∙ sym (snd (path i) x)) + + pathHelper2 : {x : A} → PathP (λ i → fst gp (f x) ≡ fst (path (~ i)) (f x)) (λ i → fst (path i) (f x)) refl + pathHelper2 {x = x} i = λ i₁ → fst (path (i₁ ∧ (~ i))) (f x) + + pathFinal : {x : A} → (λ i → fst (path i) (f x)) ≡ snd gp x ∙ sym (snd hq x) + pathFinal {x = x} = transport (λ i → PathP (λ j₁ → (lCancel (λ i → fst gp (f x) ≡ fst (path i) (f x))) i j₁) + (λ i → fst (path i) (f x)) + (snd gp x ∙ sym (snd hq x))) + (compPathP pathHelper2 pathHelper) + + + finv : Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) → (gp ≡ hq) + finv (fs , sn) = ΣPathP {x = gp} {y = hq} (funExt (λ b → fs b) , {! !}) -- (λ b → (fs b) i) , λ a → hcomp (λ k → λ{(i = i0) → (snd gp a) ; (i = i1) → {!!}}) (transp (λ j → fs (f a) (i ∧ j) ≡ l a) (~ i) ((snd gp) a)) + where + potTerm : {a : A} → PathP (λ i → (a : A) → funExt (λ b → fs b) i (f a) ≡ l a) (snd gp) (snd hq) + potTerm {a = a} j = λ a → hcomp {!(transp (λ i → funExt (λ b → fs b) (i) (f a) ≡ l a) i0 ((snd gp a))) !} (transp (λ i → funExt (λ b → fs b) (i ∧ j) (f a) ≡ l a) i0 ((snd gp a))) + + helper2 : {a : A} → (fs (f a)) ∙ (snd hq a) ≡ snd gp a + helper2 {a = a} = (cong (λ x → x ∙ (snd hq a)) (cong (λ x → x a) sn)) ∙ sym (assoc (snd gp a) (sym (snd hq a)) (snd hq a)) ∙ (λ j → (snd gp a) ∙ (lCancel (snd hq a)) j) ∙ sym (rUnit (snd gp a)) -} From 37735ee7073d593d5a3ae9842b2c23e41c0d3a47 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 6 Jul 2021 16:33:28 +0200 Subject: [PATCH 10/30] stuff --- Cubical/Algebra/AbGroup.agda | 4 + Cubical/Algebra/AbGroup/Base.agda | 192 +++ Cubical/Algebra/Algebra.agda | 4 + Cubical/Algebra/Algebra/Base.agda | 305 +++++ Cubical/Algebra/CommAlgebra.agda | 4 + Cubical/Algebra/CommAlgebra/Base.agda | 175 +++ Cubical/Algebra/CommAlgebra/Localisation.agda | 217 ++++ Cubical/Algebra/CommAlgebra/Properties.agda | 113 ++ Cubical/Algebra/CommRing.agda | 5 + Cubical/Algebra/CommRing/Base.agda | 144 +++ Cubical/Algebra/CommRing/BinomialThm.agda | 132 ++ Cubical/Algebra/CommRing/FGIdeal.agda | 104 ++ Cubical/Algebra/CommRing/Ideal.agda | 79 ++ .../Algebra/CommRing/Instances/BiInvInt.agda | 26 + Cubical/Algebra/CommRing/Instances/Int.agda | 26 + .../Algebra/CommRing/Localisation/Base.agda | 294 +++++ .../Localisation/InvertingElements.agda | 606 ++++++++++ .../Localisation/UniversalProperty.agda | 417 +++++++ Cubical/Algebra/CommRing/Properties.agda | 225 ++++ Cubical/Algebra/CommRing/RadicalIdeal.agda | 122 ++ Cubical/Algebra/Everything.agda | 49 + Cubical/Algebra/FreeCommAlgebra.agda | 455 +++++++ Cubical/Algebra/Group.agda | 15 + Cubical/Algebra/Group/Base.agda | 186 +++ Cubical/Algebra/Group/DirProd.agda | 33 + Cubical/Algebra/Group/EilenbergMacLane1.agda | 151 +++ Cubical/Algebra/Group/GroupPath.agda | 129 ++ Cubical/Algebra/Group/Instances/Bool.agda | 94 ++ Cubical/Algebra/Group/Instances/Int.agda | 21 + Cubical/Algebra/Group/Instances/Unit.agda | 56 + .../Algebra/Group/IsomorphismTheorems.agda | 105 ++ Cubical/Algebra/Group/MorphismProperties.agda | 294 +++++ Cubical/Algebra/Group/Morphisms.agda | 103 ++ Cubical/Algebra/Group/Properties.agda | 97 ++ Cubical/Algebra/Group/QuotientGroup.agda | 105 ++ Cubical/Algebra/Group/Subgroup.agda | 196 +++ Cubical/Algebra/Matrix.agda | 249 ++++ Cubical/Algebra/Module.agda | 4 + Cubical/Algebra/Module/Base.agda | 165 +++ Cubical/Algebra/Monoid.agda | 4 + Cubical/Algebra/Monoid/Base.agda | 144 +++ Cubical/Algebra/Monoid/BigOp.agda | 54 + Cubical/Algebra/Ring.agda | 7 + Cubical/Algebra/Ring/Base.agda | 259 ++++ Cubical/Algebra/Ring/BigOps.agda | 115 ++ Cubical/Algebra/Ring/Ideal.agda | 81 ++ Cubical/Algebra/Ring/Kernel.agda | 56 + Cubical/Algebra/Ring/Properties.agda | 201 +++ Cubical/Algebra/Ring/QuotientRing.agda | 234 ++++ .../Algebra/RingSolver/AlgebraExpression.agda | 38 + Cubical/Algebra/RingSolver/AlmostRing.agda | 111 ++ .../RingSolver/CommRingAsAlmostRing.agda | 45 + .../Algebra/RingSolver/CommRingEvalHom.agda | 282 +++++ .../Algebra/RingSolver/CommRingExamples.agda | 81 ++ .../RingSolver/CommRingHornerForms.agda | 139 +++ .../Algebra/RingSolver/CommRingSolver.agda | 222 ++++ .../RingSolver/EvaluationHomomorphism.agda | 228 ++++ Cubical/Algebra/RingSolver/Examples.agda | 52 + Cubical/Algebra/RingSolver/HornerForms.agda | 118 ++ Cubical/Algebra/RingSolver/IntAsRawRing.agda | 15 + .../Algebra/RingSolver/NatAsAlmostRing.agda | 21 + Cubical/Algebra/RingSolver/NatExamples.agda | 205 ++++ Cubical/Algebra/RingSolver/README.md | 25 + Cubical/Algebra/RingSolver/RawAlgebra.agda | 179 +++ Cubical/Algebra/RingSolver/RawRing.agda | 35 + .../Algebra/RingSolver/ReflectionSolving.agda | 209 ++++ .../Algebra/RingSolver/RingExpression.agda | 39 + Cubical/Algebra/RingSolver/Solver.agda | 150 +++ Cubical/Algebra/Semigroup.agda | 4 + Cubical/Algebra/Semigroup/Base.agda | 100 ++ Cubical/Algebra/SymmetricGroup.agda | 26 + .../Algebra/ZariskiLattice/BasicOpens.agda | 236 ++++ Cubical/AndersFile.agda | 41 - Cubical/Categories/Adjoint.agda | 319 +++++ Cubical/Categories/Category.agda | 25 + Cubical/Categories/Category/Base.agda | 101 ++ Cubical/Categories/Category/Properties.agda | 93 ++ Cubical/Categories/Commutativity.agda | 33 + .../Categories/Constructions/Elements.agda | 137 +++ Cubical/Categories/Constructions/Slice.agda | 383 ++++++ Cubical/Categories/Equivalence.agda | 7 + Cubical/Categories/Equivalence/Base.agda | 31 + .../Categories/Equivalence/Properties.agda | 92 ++ Cubical/Categories/Everything.agda | 20 + Cubical/Categories/Functor.agda | 7 + Cubical/Categories/Functor/Base.agda | 89 ++ Cubical/Categories/Functor/Compose.agda | 47 + Cubical/Categories/Functor/Properties.agda | 116 ++ Cubical/Categories/Instances/Cospan.agda | 60 + Cubical/Categories/Instances/Functors.agda | 59 + Cubical/Categories/Instances/Sets.agda | 106 ++ Cubical/Categories/Limits.agda | 7 + Cubical/Categories/Limits/Base.agda | 160 +++ Cubical/Categories/Limits/Pullback.agda | 80 ++ Cubical/Categories/Limits/Terminal.agda | 25 + Cubical/Categories/Morphism.agda | 113 ++ Cubical/Categories/NaturalTransformation.agda | 6 + .../NaturalTransformation/Base.agda | 222 ++++ .../NaturalTransformation/Properties.agda | 146 +++ Cubical/Categories/Presheaf.agda | 8 + Cubical/Categories/Presheaf/Base.agda | 84 ++ Cubical/Categories/Presheaf/KanExtension.agda | 341 ++++++ Cubical/Categories/Presheaf/Properties.agda | 389 ++++++ Cubical/Categories/Sets.agda | 93 ++ .../TypesOfCategories/TypeCategory.agda | 187 +++ Cubical/Categories/Yoneda.agda | 192 +++ Cubical/Codata/Conat.agda | 3 +- Cubical/Codata/Conat/Base.agda | 2 +- Cubical/Codata/Conat/Properties.agda | 31 +- Cubical/Codata/Everything.agda | 18 +- Cubical/Codata/EverythingSafe.agda | 2 +- Cubical/Codata/M.agda | 17 +- Cubical/Codata/M/AsLimit/Coalg.agda | 5 + Cubical/Codata/M/AsLimit/Coalg/Base.agda | 39 + Cubical/Codata/M/AsLimit/Container.agda | 88 ++ Cubical/Codata/M/AsLimit/M.agda | 6 + Cubical/Codata/M/AsLimit/M/Base.agda | 194 +++ Cubical/Codata/M/AsLimit/M/Properties.agda | 61 + Cubical/Codata/M/AsLimit/helper.agda | 116 ++ Cubical/Codata/M/AsLimit/itree.agda | 75 ++ Cubical/Codata/M/AsLimit/stream.agda | 30 + Cubical/Codata/M/Bisimilarity.agda | 2 +- Cubical/Codata/Stream.agda | 3 +- Cubical/Codata/Stream/Base.agda | 2 +- Cubical/Codata/Stream/Properties.agda | 6 +- Cubical/Core/Everything.agda | 2 +- Cubical/Core/Glue.agda | 96 +- Cubical/Core/Id.agda | 2 +- Cubical/Core/Primitives.agda | 32 +- Cubical/Data/BinNat.agda | 2 +- Cubical/Data/BinNat/BinNat.agda | 16 +- Cubical/Data/Bool.agda | 2 +- Cubical/Data/Bool/Base.agda | 54 +- Cubical/Data/Bool/Properties.agda | 136 ++- Cubical/Data/Bool/SwitchStatement.agda | 42 + Cubical/Data/DescendingList.agda | 6 +- Cubical/Data/DescendingList/Base.agda | 2 +- Cubical/Data/DescendingList/Examples.agda | 21 +- Cubical/Data/DescendingList/Properties.agda | 32 +- Cubical/Data/DescendingList/Strict.agda | 2 +- .../DescendingList/Strict/Properties.agda | 79 +- Cubical/Data/DiffInt.agda | 5 - Cubical/Data/DiffInt/Base.agda | 18 - Cubical/Data/Empty.agda | 2 +- Cubical/Data/Empty/Base.agda | 18 +- Cubical/Data/Empty/Properties.agda | 21 +- Cubical/Data/Equality.agda | 4 +- Cubical/Data/Everything.agda | 68 +- Cubical/Data/Fin.agda | 5 +- Cubical/Data/Fin/Base.agda | 45 +- Cubical/Data/Fin/LehmerCode.agda | 222 ++++ Cubical/Data/Fin/Literals.agda | 18 + Cubical/Data/Fin/Properties.agda | 519 +++++++- Cubical/Data/Fin/Recursive.agda | 6 + Cubical/Data/Fin/Recursive/Base.agda | 34 + Cubical/Data/Fin/Recursive/Properties.agda | 266 ++++ Cubical/Data/FinData.agda | 5 + Cubical/Data/FinData/Base.agda | 80 ++ Cubical/Data/FinData/Properties.agda | 56 + Cubical/Data/FinInd.agda | 48 + Cubical/Data/FinSet.agda | 105 ++ Cubical/Data/FinSet/Binary/Large.agda | 147 +++ Cubical/Data/FinSet/Binary/Small.agda | 6 + Cubical/Data/FinSet/Binary/Small/Base.agda | 20 + .../Data/FinSet/Binary/Small/Properties.agda | 115 ++ Cubical/Data/Graph.agda | 2 +- Cubical/Data/Graph/Base.agda | 2 +- Cubical/Data/Graph/Examples.agda | 5 +- Cubical/Data/Group.agda | 4 - Cubical/Data/Group/Base.agda | 162 --- Cubical/Data/HomotopyGroup.agda | 2 +- Cubical/Data/HomotopyGroup/Base.agda | 60 +- Cubical/Data/InfNat.agda | 2 +- Cubical/Data/InfNat/Base.agda | 18 +- Cubical/Data/InfNat/Properties.agda | 2 +- Cubical/Data/Int.agda | 5 +- Cubical/Data/Int/Base.agda | 91 +- Cubical/Data/Int/MoreInts/BiInvInt.agda | 6 + .../Int/MoreInts}/BiInvInt/Base.agda | 126 +- .../Int/MoreInts/BiInvInt/Properties.agda | 243 ++++ Cubical/Data/Int/MoreInts/DeltaInt.agda | 6 + .../Int/MoreInts}/DeltaInt/Base.agda | 25 +- .../Int/MoreInts}/DeltaInt/Properties.agda | 20 +- Cubical/Data/Int/MoreInts/DiffInt.agda | 5 + Cubical/Data/Int/MoreInts/DiffInt/Base.agda | 84 ++ .../MoreInts}/DiffInt/Properties.agda | 11 +- Cubical/Data/Int/MoreInts/QuoInt.agda | 6 + Cubical/Data/Int/MoreInts/QuoInt/Base.agda | 215 ++++ .../Data/Int/MoreInts/QuoInt/Properties.agda | 235 ++++ Cubical/Data/Int/Properties.agda | 439 ++++--- Cubical/Data/List.agda | 2 +- Cubical/Data/List/Base.agda | 32 +- Cubical/Data/List/Properties.agda | 56 +- Cubical/Data/Maybe.agda | 2 +- Cubical/Data/Maybe/Base.agda | 6 +- Cubical/Data/Maybe/Properties.agda | 38 +- Cubical/Data/Nat.agda | 2 +- Cubical/Data/Nat/Algebra.agda | 21 +- Cubical/Data/Nat/Base.agda | 65 +- Cubical/Data/Nat/Coprime.agda | 95 ++ Cubical/Data/Nat/Divisibility.agda | 104 ++ Cubical/Data/Nat/GCD.agda | 161 +++ Cubical/Data/Nat/Literals.agda | 22 + Cubical/Data/Nat/Order.agda | 151 ++- Cubical/Data/Nat/Order/Recursive.agda | 104 ++ Cubical/Data/Nat/Properties.agda | 206 +++- Cubical/Data/NatMinusOne.agda | 4 +- Cubical/Data/NatMinusOne/Base.agda | 33 +- Cubical/Data/NatMinusOne/Properties.agda | 13 + Cubical/Data/NatMinusTwo.agda | 4 - Cubical/Data/NatMinusTwo/Base.agda | 61 - Cubical/Data/NatPlusOne.agda | 6 + Cubical/Data/NatPlusOne/Base.agda | 32 + Cubical/Data/NatPlusOne/Properties.agda | 38 + Cubical/Data/Prod.agda | 2 +- Cubical/Data/Prod/Base.agda | 30 +- Cubical/Data/Prod/Properties.agda | 64 +- Cubical/Data/Queue.agda | 7 + Cubical/Data/Queue/1List.agda | 87 ++ Cubical/Data/Queue/Base.agda | 6 + Cubical/Data/Queue/Finite.agda | 58 + Cubical/Data/Queue/Truncated2List.agda | 150 +++ Cubical/Data/Queue/Untruncated2List.agda | 166 +++ .../Data/Queue/Untruncated2ListInvariant.agda | 81 ++ Cubical/Data/Sigma.agda | 2 +- Cubical/Data/Sigma/Base.agda | 50 +- Cubical/Data/Sigma/Properties.agda | 403 ++++-- Cubical/Data/SubFinSet.agda | 49 + Cubical/Data/Sum.agda | 2 +- Cubical/Data/Sum/Base.agda | 29 +- Cubical/Data/Sum/Properties.agda | 98 +- Cubical/Data/SumFin.agda | 3 +- Cubical/Data/SumFin/Base.agda | 24 +- Cubical/Data/SumFin/Properties.agda | 48 + Cubical/Data/Unit.agda | 2 +- Cubical/Data/Unit/Base.agda | 15 +- Cubical/Data/Unit/Properties.agda | 55 +- Cubical/Data/Vec.agda | 6 + Cubical/Data/Vec/Base.agda | 67 + Cubical/Data/Vec/NAry.agda | 55 + Cubical/Data/Vec/Properties.agda | 60 + Cubical/Displayed/Auto.agda | 367 ++++++ Cubical/Displayed/Base.agda | 87 ++ Cubical/Displayed/Constant.agda | 39 + Cubical/Displayed/Everything.agda | 16 + Cubical/Displayed/Function.agda | 163 +++ Cubical/Displayed/Generic.agda | 36 + Cubical/Displayed/Morphism.agda | 67 + Cubical/Displayed/Prop.agda | 52 + Cubical/Displayed/Properties.agda | 143 +++ Cubical/Displayed/Record.agda | 231 ++++ Cubical/Displayed/Sigma.agda | 101 ++ Cubical/Displayed/Subst.agda | 60 + Cubical/Displayed/Unit.agda | 27 + Cubical/Displayed/Universe.agda | 32 + Cubical/Experiments/Brunerie.agda | 123 +- Cubical/Experiments/CohomologyGroups.agda | 140 +++ Cubical/Experiments/EscardoSIP.agda | 133 +- Cubical/Experiments/Everything.agda | 9 +- Cubical/Experiments/FunExtFromUA.agda | 21 +- Cubical/Experiments/Generic.agda | 22 +- Cubical/Experiments/HAEquivInt.agda | 5 + .../Ints => Experiments}/HAEquivInt/Base.agda | 8 +- Cubical/Experiments/HInt.agda | 1 - .../{Foundations => Experiments}/HoTT-UF.agda | 4 +- Cubical/Experiments/IsoInt.agda | 5 + .../Ints => Experiments}/IsoInt/Base.agda | 20 +- Cubical/Experiments/List.agda | 285 +++++ Cubical/Experiments/LocalisationDefs.agda | 146 +++ Cubical/Experiments/NatMinusTwo.agda | 21 + Cubical/Experiments/NatMinusTwo/Base.agda | 41 + .../Experiments/NatMinusTwo/Properties.agda | 11 + .../NatMinusTwo/ToNatMinusOne.agda | 24 + Cubical/Experiments/Problem.agda | 50 +- .../Experiments/ZCohomology/Benchmarks.agda | 385 ++++++ .../Experiments/ZCohomologyExperiments.agda | 79 ++ Cubical/Experiments/ZCohomologyOld/Base.agda | 46 + .../ZCohomologyOld/KcompPrelims.agda | 212 ++++ .../MayerVietorisUnreduced.agda | 384 ++++++ .../ZCohomologyOld/Properties.agda | 634 ++++++++++ Cubical/Foundations/CartesianKanOps.agda | 10 +- Cubical/Foundations/Embedding.agda | 130 -- Cubical/Foundations/Equiv.agda | 293 +++-- Cubical/Foundations/Equiv/Base.agda | 27 + .../BiInvertible.agda} | 14 +- Cubical/Foundations/Equiv/Fiberwise.agda | 19 +- Cubical/Foundations/Equiv/HalfAdjoint.agda | 166 +++ .../PathSplit.agda} | 65 +- Cubical/Foundations/Equiv/Properties.agda | 228 +++- Cubical/Foundations/Everything.agda | 25 +- Cubical/Foundations/FunExtEquiv.agda | 28 - Cubical/Foundations/Function.agda | 63 +- Cubical/Foundations/GroupoidLaws.agda | 422 +++++-- Cubical/Foundations/HAEquiv.agda | 147 --- Cubical/Foundations/HLevels.agda | 628 +++++++--- Cubical/Foundations/Id.agda | 56 +- Cubical/Foundations/Isomorphism.agda | 111 +- Cubical/Foundations/Path.agda | 331 +++-- Cubical/Foundations/Pointed.agda | 4 +- Cubical/Foundations/Pointed/Base.agda | 77 +- Cubical/Foundations/Pointed/FunExt.agda | 48 + Cubical/Foundations/Pointed/Homogeneous.agda | 75 +- Cubical/Foundations/Pointed/Homotopy.agda | 119 ++ Cubical/Foundations/Pointed/Properties.agda | 84 +- Cubical/Foundations/Powerset.agda | 65 + Cubical/Foundations/Prelude.agda | 370 +++++- Cubical/Foundations/RelationalStructure.agda | 269 +++++ Cubical/Foundations/SIP.agda | 568 ++------- Cubical/Foundations/Structure.agda | 44 + Cubical/Foundations/Surjection.agda | 53 - Cubical/Foundations/TotalFiber.agda | 65 - Cubical/Foundations/Transport.agda | 108 +- Cubical/Foundations/Univalence.agda | 203 +++- Cubical/Foundations/Univalence/Universe.agda | 94 +- Cubical/Foundations/UnivalenceId.agda | 39 - Cubical/Functions/Bundle.agda | 50 + Cubical/Functions/Embedding.agda | 427 +++++++ Cubical/Functions/Everything.agda | 13 + Cubical/Functions/Fibration.agda | 105 ++ Cubical/Functions/Fixpoint.agda | 43 + Cubical/Functions/FunExtEquiv.agda | 193 +++ Cubical/Functions/Implicit.agda | 17 + Cubical/Functions/Involution.agda | 42 + Cubical/{Foundations => Functions}/Logic.agda | 197 +-- Cubical/Functions/Morphism.agda | 84 ++ Cubical/Functions/Surjection.agda | 80 ++ Cubical/HITs/2GroupoidTruncation.agda | 2 +- Cubical/HITs/2GroupoidTruncation/Base.agda | 8 +- .../HITs/2GroupoidTruncation/Properties.agda | 117 +- Cubical/HITs/AssocList.agda | 6 + Cubical/HITs/AssocList/Base.agda | 69 ++ Cubical/HITs/AssocList/Properties.agda | 162 +++ Cubical/HITs/Colimit.agda | 2 +- Cubical/HITs/Colimit/Base.agda | 20 +- Cubical/HITs/Colimit/Examples.agda | 2 +- Cubical/HITs/Cost.agda | 4 + Cubical/HITs/Cost/Base.agda | 124 ++ Cubical/HITs/Cylinder.agda | 2 +- Cubical/HITs/Cylinder/Base.agda | 61 +- Cubical/HITs/Delooping/Two/Base.agda | 55 + Cubical/HITs/Delooping/Two/Properties.agda | 198 +++ Cubical/HITs/DunceCap.agda | 2 +- Cubical/HITs/DunceCap/Base.agda | 2 +- Cubical/HITs/DunceCap/Properties.agda | 4 +- Cubical/HITs/EilenbergMacLane1.agda | 5 + Cubical/HITs/EilenbergMacLane1/Base.agda | 43 + .../HITs/EilenbergMacLane1/Properties.agda | 100 ++ Cubical/HITs/Everything.agda | 79 +- Cubical/HITs/FiniteMultiset.agda | 2 +- Cubical/HITs/FiniteMultiset/Base.agda | 17 +- .../FiniteMultiset/CountExtensionality.agda | 191 +++ Cubical/HITs/FiniteMultiset/Properties.agda | 170 ++- Cubical/HITs/GroupoidQuotients.agda | 5 + Cubical/HITs/GroupoidQuotients/Base.agda | 43 + .../HITs/GroupoidQuotients/Properties.agda | 106 ++ Cubical/HITs/GroupoidTruncation.agda | 2 +- Cubical/HITs/GroupoidTruncation/Base.agda | 8 +- .../HITs/GroupoidTruncation/Properties.agda | 99 +- Cubical/HITs/Hopf.agda | 120 +- Cubical/HITs/InfNat.agda | 2 +- Cubical/HITs/InfNat/Base.agda | 2 +- Cubical/HITs/InfNat/Properties.agda | 2 +- Cubical/HITs/Interval.agda | 2 +- Cubical/HITs/Interval/Base.agda | 12 +- Cubical/HITs/Ints/BiInvInt.agda | 6 - Cubical/HITs/Ints/BiInvInt/Properties.agda | 148 --- Cubical/HITs/Ints/DeltaInt.agda | 6 - Cubical/HITs/Ints/HAEquivInt.agda | 5 - Cubical/HITs/Ints/IsoInt.agda | 5 - Cubical/HITs/Ints/QuoInt.agda | 6 - Cubical/HITs/Ints/QuoInt/Base.agda | 180 --- Cubical/HITs/Ints/QuoInt/Properties.agda | 68 -- Cubical/HITs/Join.agda | 2 +- Cubical/HITs/Join/Base.agda | 11 +- Cubical/HITs/Join/Properties.agda | 34 +- Cubical/HITs/KleinBottle.agda | 2 +- Cubical/HITs/KleinBottle/Base.agda | 4 +- Cubical/HITs/KleinBottle/Properties.agda | 38 +- Cubical/HITs/ListedFiniteSet.agda | 2 +- Cubical/HITs/ListedFiniteSet/Base.agda | 75 +- Cubical/HITs/ListedFiniteSet/Properties.agda | 127 +- Cubical/HITs/Localization.agda | 4 +- Cubical/HITs/Localization/Base.agda | 4 +- Cubical/HITs/Localization/Properties.agda | 4 +- Cubical/HITs/MappingCones.agda | 2 +- Cubical/HITs/MappingCones/Base.agda | 2 +- Cubical/HITs/MappingCones/Properties.agda | 2 +- Cubical/HITs/Modulo.agda | 4 +- Cubical/HITs/Modulo/Base.agda | 7 +- Cubical/HITs/Modulo/FinEquiv.agda | 2 +- Cubical/HITs/Modulo/Properties.agda | 2 +- Cubical/HITs/Nullification.agda | 4 +- Cubical/HITs/Nullification/Base.agda | 4 +- Cubical/HITs/Nullification/Properties.agda | 85 +- Cubical/HITs/PropositionalTruncation.agda | 2 +- .../HITs/PropositionalTruncation/Base.agda | 2 +- .../PropositionalTruncation/MagicTrick.agda | 5 +- .../HITs/PropositionalTruncation/Monad.agda | 33 + .../PropositionalTruncation/Properties.agda | 266 +++- Cubical/HITs/Pushout.agda | 3 +- Cubical/HITs/Pushout/Base.agda | 21 +- Cubical/HITs/Pushout/Flattening.agda | 91 ++ Cubical/HITs/Pushout/KrausVonRaumer.agda | 137 +++ Cubical/HITs/Pushout/Properties.agda | 55 +- Cubical/HITs/RPn.agda | 4 + Cubical/HITs/RPn/Base.agda | 320 +++++ Cubical/HITs/Rational.agda | 6 - Cubical/HITs/Rational/Base.agda | 18 - Cubical/HITs/Rationals/HITQ.agda | 6 + Cubical/HITs/Rationals/HITQ/Base.agda | 36 + Cubical/HITs/Rationals/QuoQ.agda | 6 + Cubical/HITs/Rationals/QuoQ/Base.agda | 76 ++ Cubical/HITs/Rationals/QuoQ/Properties.agda | 245 ++++ Cubical/HITs/Rationals/SigmaQ.agda | 6 + Cubical/HITs/Rationals/SigmaQ/Base.agda | 54 + Cubical/HITs/Rationals/SigmaQ/Properties.agda | 78 ++ Cubical/HITs/S1.agda | 2 +- Cubical/HITs/S1/Base.agda | 333 +++-- Cubical/HITs/S1/Properties.agda | 14 +- Cubical/HITs/S2.agda | 2 +- Cubical/HITs/S2/Base.agda | 11 +- Cubical/HITs/S3.agda | 2 +- Cubical/HITs/S3/Base.agda | 3 +- Cubical/HITs/SetQuotients.agda | 2 +- Cubical/HITs/SetQuotients/Base.agda | 2 +- Cubical/HITs/SetQuotients/Properties.agda | 299 +++-- Cubical/HITs/SetTruncation.agda | 2 +- Cubical/HITs/SetTruncation/Base.agda | 8 +- Cubical/HITs/SetTruncation/Properties.agda | 352 +++++- Cubical/HITs/SmashProduct.agda | 2 +- Cubical/HITs/SmashProduct/Base.agda | 215 +++- Cubical/HITs/Sn.agda | 3 +- Cubical/HITs/Sn/Base.agda | 20 +- Cubical/HITs/Sn/Properties.agda | 310 +++++ Cubical/HITs/Susp.agda | 4 +- Cubical/HITs/Susp/Base.agda | 103 +- Cubical/HITs/Susp/Properties.agda | 119 ++ Cubical/HITs/Torus.agda | 2 +- Cubical/HITs/Torus/Base.agda | 14 +- Cubical/HITs/Truncation.agda | 4 +- Cubical/HITs/Truncation/Base.agda | 54 +- Cubical/HITs/Truncation/FromNegOne.agda | 5 - Cubical/HITs/Truncation/FromNegOne/Base.agda | 17 - .../Truncation/FromNegOne/Properties.agda | 180 --- Cubical/HITs/Truncation/FromNegTwo.agda | 5 + Cubical/HITs/Truncation/FromNegTwo/Base.agda | 29 + .../Truncation/FromNegTwo/Properties.agda | 396 ++++++ Cubical/HITs/Truncation/Properties.agda | 813 ++++++++----- Cubical/HITs/TypeQuotients.agda | 5 + Cubical/HITs/TypeQuotients/Base.agda | 17 + Cubical/HITs/TypeQuotients/Properties.agda | 60 + Cubical/HITs/Wedge.agda | 4 + Cubical/HITs/Wedge/Base.agda | 26 + Cubical/Homotopy/Base.agda | 19 + Cubical/Homotopy/Connected.agda | 306 +++++ Cubical/Homotopy/EilenbergSteenrod.agda | 47 + Cubical/Homotopy/Everything.agda | 10 + Cubical/Homotopy/Freudenthal.agda | 138 +++ Cubical/Homotopy/Loopspace.agda | 81 ++ Cubical/Homotopy/MayerVietorisCofiber.agda | 175 +++ Cubical/Homotopy/WedgeConnectivity.agda | 59 + Cubical/Induction/Everything.agda | 4 +- Cubical/Induction/WellFounded.agda | 2 +- Cubical/Modalities/Everything.agda | 5 +- Cubical/Modalities/Lex.agda | 271 +++++ Cubical/Modalities/Modality.agda | 24 +- Cubical/Papers/Everything.agda | 6 + .../Papers/RepresentationIndependence.agda | 314 +++++ Cubical/Papers/Synthetic.agda | 214 ++++ Cubical/Papers/ZCohomology.agda | 522 ++++++++ Cubical/README.agda | 41 +- Cubical/Reflection/Base.agda | 59 + Cubical/Reflection/Everything.agda | 6 + Cubical/Reflection/RecordEquiv.agda | 196 +++ Cubical/Reflection/StrictEquiv.agda | 81 ++ Cubical/Relation/Binary.agda | 2 +- Cubical/Relation/Binary/Base.agda | 128 +- Cubical/Relation/Binary/Poset.agda | 321 +++++ Cubical/Relation/Binary/Properties.agda | 2 +- Cubical/Relation/Everything.agda | 12 +- Cubical/Relation/Nullary.agda | 61 +- Cubical/Relation/Nullary/Base.agda | 66 + Cubical/Relation/Nullary/DecidableEq.agda | 34 +- Cubical/Relation/Nullary/HLevels.agda | 35 + Cubical/Relation/Nullary/Properties.agda | 163 +++ .../ZigZag/Applications/MultiSet.agda | 358 ++++++ Cubical/Relation/ZigZag/Base.agda | 161 +++ Cubical/Structures/Auto.agda | 250 ++++ Cubical/Structures/Axioms.agda | 69 ++ Cubical/Structures/Constant.agda | 35 + Cubical/Structures/Everything.agda | 27 + Cubical/Structures/Function.agda | 82 ++ Cubical/Structures/LeftAction.agda | 17 + Cubical/Structures/Macro.agda | 157 +++ Cubical/Structures/Maybe.agda | 108 ++ Cubical/Structures/MultiSet.agda | 42 + Cubical/Structures/Parameterized.agda | 48 + Cubical/Structures/Pointed.agda | 41 + Cubical/Structures/Product.agda | 50 + Cubical/Structures/Queue.agda | 115 ++ Cubical/Structures/Record.agda | 453 +++++++ Cubical/Structures/Relational/Auto.agda | 239 ++++ Cubical/Structures/Relational/Constant.agda | 57 + Cubical/Structures/Relational/Equalizer.agda | 67 + Cubical/Structures/Relational/Function.agda | 191 +++ Cubical/Structures/Relational/Macro.agda | 161 +++ Cubical/Structures/Relational/Maybe.agda | 127 ++ .../Structures/Relational/Parameterized.agda | 70 ++ Cubical/Structures/Relational/Pointed.agda | 56 + Cubical/Structures/Relational/Product.agda | 140 +++ Cubical/Structures/Transfer.agda | 47 + Cubical/Structures/TypeEqvTo.agda | 41 + Cubical/Talks/EPA2020.agda | 341 ++++++ Cubical/Talks/Everything.agda | 4 + Cubical/WithK.agda | 8 +- Cubical/ZCohomology/#Benchmarks.agda# | 3 + Cubical/ZCohomology/Base.agda | 32 +- Cubical/ZCohomology/Benchmarks.agda | 4 + Cubical/ZCohomology/EilenbergSteenrodZ.agda | 379 ++++++ Cubical/ZCohomology/Everything.agda | 23 +- Cubical/ZCohomology/FunLoopSpace.agda | 205 ---- Cubical/ZCohomology/FunLoopSpace2.agda | 34 - Cubical/ZCohomology/GroupStructure.agda | 919 ++++++++++++++ Cubical/ZCohomology/Groups/CP2.agda | 125 ++ Cubical/ZCohomology/Groups/Connected.agda | 45 + Cubical/ZCohomology/Groups/KleinBottle.agda | 454 +++++++ Cubical/ZCohomology/Groups/Prelims.agda | 280 +++++ Cubical/ZCohomology/Groups/RP2.agda | 128 ++ Cubical/ZCohomology/Groups/Sn.agda | 330 +++++ Cubical/ZCohomology/Groups/Torus.agda | 309 +++++ Cubical/ZCohomology/Groups/Unit.agda | 85 ++ Cubical/ZCohomology/Groups/Wedge.agda | 243 ++++ .../ZCohomology/Groups/WedgeOfSpheres.agda | 134 ++ .../ZCohomology/MayerVietorisUnreduced.agda | 249 ++++ Cubical/ZCohomology/Properties.agda | 645 +++++++++- Cubical/ZCohomology/PropertiesTrash.agda | 938 -------------- .../ZCohomology/RingStructure/CupProduct.agda | 87 ++ .../RingStructure/GradedCommutativity.agda | 1075 +++++++++++++++++ .../ZCohomology/RingStructure/RingLaws.agda | 657 ++++++++++ Cubical/ZCohomology/S1/S1.agda | 36 - 540 files changed, 47029 insertions(+), 6436 deletions(-) create mode 100644 Cubical/Algebra/AbGroup.agda create mode 100644 Cubical/Algebra/AbGroup/Base.agda create mode 100644 Cubical/Algebra/Algebra.agda create mode 100644 Cubical/Algebra/Algebra/Base.agda create mode 100644 Cubical/Algebra/CommAlgebra.agda create mode 100644 Cubical/Algebra/CommAlgebra/Base.agda create mode 100644 Cubical/Algebra/CommAlgebra/Localisation.agda create mode 100644 Cubical/Algebra/CommAlgebra/Properties.agda create mode 100644 Cubical/Algebra/CommRing.agda create mode 100644 Cubical/Algebra/CommRing/Base.agda create mode 100644 Cubical/Algebra/CommRing/BinomialThm.agda create mode 100644 Cubical/Algebra/CommRing/FGIdeal.agda create mode 100644 Cubical/Algebra/CommRing/Ideal.agda create mode 100644 Cubical/Algebra/CommRing/Instances/BiInvInt.agda create mode 100644 Cubical/Algebra/CommRing/Instances/Int.agda create mode 100644 Cubical/Algebra/CommRing/Localisation/Base.agda create mode 100644 Cubical/Algebra/CommRing/Localisation/InvertingElements.agda create mode 100644 Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda create mode 100644 Cubical/Algebra/CommRing/Properties.agda create mode 100644 Cubical/Algebra/CommRing/RadicalIdeal.agda create mode 100644 Cubical/Algebra/Everything.agda create mode 100644 Cubical/Algebra/FreeCommAlgebra.agda create mode 100644 Cubical/Algebra/Group.agda create mode 100644 Cubical/Algebra/Group/Base.agda create mode 100644 Cubical/Algebra/Group/DirProd.agda create mode 100644 Cubical/Algebra/Group/EilenbergMacLane1.agda create mode 100644 Cubical/Algebra/Group/GroupPath.agda create mode 100644 Cubical/Algebra/Group/Instances/Bool.agda create mode 100644 Cubical/Algebra/Group/Instances/Int.agda create mode 100644 Cubical/Algebra/Group/Instances/Unit.agda create mode 100644 Cubical/Algebra/Group/IsomorphismTheorems.agda create mode 100644 Cubical/Algebra/Group/MorphismProperties.agda create mode 100644 Cubical/Algebra/Group/Morphisms.agda create mode 100644 Cubical/Algebra/Group/Properties.agda create mode 100644 Cubical/Algebra/Group/QuotientGroup.agda create mode 100644 Cubical/Algebra/Group/Subgroup.agda create mode 100644 Cubical/Algebra/Matrix.agda create mode 100644 Cubical/Algebra/Module.agda create mode 100644 Cubical/Algebra/Module/Base.agda create mode 100644 Cubical/Algebra/Monoid.agda create mode 100644 Cubical/Algebra/Monoid/Base.agda create mode 100644 Cubical/Algebra/Monoid/BigOp.agda create mode 100644 Cubical/Algebra/Ring.agda create mode 100644 Cubical/Algebra/Ring/Base.agda create mode 100644 Cubical/Algebra/Ring/BigOps.agda create mode 100644 Cubical/Algebra/Ring/Ideal.agda create mode 100644 Cubical/Algebra/Ring/Kernel.agda create mode 100644 Cubical/Algebra/Ring/Properties.agda create mode 100644 Cubical/Algebra/Ring/QuotientRing.agda create mode 100644 Cubical/Algebra/RingSolver/AlgebraExpression.agda create mode 100644 Cubical/Algebra/RingSolver/AlmostRing.agda create mode 100644 Cubical/Algebra/RingSolver/CommRingAsAlmostRing.agda create mode 100644 Cubical/Algebra/RingSolver/CommRingEvalHom.agda create mode 100644 Cubical/Algebra/RingSolver/CommRingExamples.agda create mode 100644 Cubical/Algebra/RingSolver/CommRingHornerForms.agda create mode 100644 Cubical/Algebra/RingSolver/CommRingSolver.agda create mode 100644 Cubical/Algebra/RingSolver/EvaluationHomomorphism.agda create mode 100644 Cubical/Algebra/RingSolver/Examples.agda create mode 100644 Cubical/Algebra/RingSolver/HornerForms.agda create mode 100644 Cubical/Algebra/RingSolver/IntAsRawRing.agda create mode 100644 Cubical/Algebra/RingSolver/NatAsAlmostRing.agda create mode 100644 Cubical/Algebra/RingSolver/NatExamples.agda create mode 100644 Cubical/Algebra/RingSolver/README.md create mode 100644 Cubical/Algebra/RingSolver/RawAlgebra.agda create mode 100644 Cubical/Algebra/RingSolver/RawRing.agda create mode 100644 Cubical/Algebra/RingSolver/ReflectionSolving.agda create mode 100644 Cubical/Algebra/RingSolver/RingExpression.agda create mode 100644 Cubical/Algebra/RingSolver/Solver.agda create mode 100644 Cubical/Algebra/Semigroup.agda create mode 100644 Cubical/Algebra/Semigroup/Base.agda create mode 100644 Cubical/Algebra/SymmetricGroup.agda create mode 100644 Cubical/Algebra/ZariskiLattice/BasicOpens.agda delete mode 100644 Cubical/AndersFile.agda create mode 100644 Cubical/Categories/Adjoint.agda create mode 100644 Cubical/Categories/Category.agda create mode 100644 Cubical/Categories/Category/Base.agda create mode 100644 Cubical/Categories/Category/Properties.agda create mode 100644 Cubical/Categories/Commutativity.agda create mode 100644 Cubical/Categories/Constructions/Elements.agda create mode 100644 Cubical/Categories/Constructions/Slice.agda create mode 100644 Cubical/Categories/Equivalence.agda create mode 100644 Cubical/Categories/Equivalence/Base.agda create mode 100644 Cubical/Categories/Equivalence/Properties.agda create mode 100644 Cubical/Categories/Everything.agda create mode 100644 Cubical/Categories/Functor.agda create mode 100644 Cubical/Categories/Functor/Base.agda create mode 100644 Cubical/Categories/Functor/Compose.agda create mode 100644 Cubical/Categories/Functor/Properties.agda create mode 100644 Cubical/Categories/Instances/Cospan.agda create mode 100644 Cubical/Categories/Instances/Functors.agda create mode 100644 Cubical/Categories/Instances/Sets.agda create mode 100644 Cubical/Categories/Limits.agda create mode 100644 Cubical/Categories/Limits/Base.agda create mode 100644 Cubical/Categories/Limits/Pullback.agda create mode 100644 Cubical/Categories/Limits/Terminal.agda create mode 100644 Cubical/Categories/Morphism.agda create mode 100644 Cubical/Categories/NaturalTransformation.agda create mode 100644 Cubical/Categories/NaturalTransformation/Base.agda create mode 100644 Cubical/Categories/NaturalTransformation/Properties.agda create mode 100644 Cubical/Categories/Presheaf.agda create mode 100644 Cubical/Categories/Presheaf/Base.agda create mode 100644 Cubical/Categories/Presheaf/KanExtension.agda create mode 100644 Cubical/Categories/Presheaf/Properties.agda create mode 100644 Cubical/Categories/Sets.agda create mode 100644 Cubical/Categories/TypesOfCategories/TypeCategory.agda create mode 100644 Cubical/Categories/Yoneda.agda create mode 100644 Cubical/Codata/M/AsLimit/Coalg.agda create mode 100644 Cubical/Codata/M/AsLimit/Coalg/Base.agda create mode 100644 Cubical/Codata/M/AsLimit/Container.agda create mode 100644 Cubical/Codata/M/AsLimit/M.agda create mode 100644 Cubical/Codata/M/AsLimit/M/Base.agda create mode 100644 Cubical/Codata/M/AsLimit/M/Properties.agda create mode 100644 Cubical/Codata/M/AsLimit/helper.agda create mode 100644 Cubical/Codata/M/AsLimit/itree.agda create mode 100644 Cubical/Codata/M/AsLimit/stream.agda create mode 100644 Cubical/Data/Bool/SwitchStatement.agda delete mode 100644 Cubical/Data/DiffInt.agda delete mode 100644 Cubical/Data/DiffInt/Base.agda create mode 100644 Cubical/Data/Fin/LehmerCode.agda create mode 100644 Cubical/Data/Fin/Literals.agda create mode 100644 Cubical/Data/Fin/Recursive.agda create mode 100644 Cubical/Data/Fin/Recursive/Base.agda create mode 100644 Cubical/Data/Fin/Recursive/Properties.agda create mode 100644 Cubical/Data/FinData.agda create mode 100644 Cubical/Data/FinData/Base.agda create mode 100644 Cubical/Data/FinData/Properties.agda create mode 100644 Cubical/Data/FinInd.agda create mode 100644 Cubical/Data/FinSet.agda create mode 100644 Cubical/Data/FinSet/Binary/Large.agda create mode 100644 Cubical/Data/FinSet/Binary/Small.agda create mode 100644 Cubical/Data/FinSet/Binary/Small/Base.agda create mode 100644 Cubical/Data/FinSet/Binary/Small/Properties.agda delete mode 100644 Cubical/Data/Group.agda delete mode 100644 Cubical/Data/Group/Base.agda create mode 100644 Cubical/Data/Int/MoreInts/BiInvInt.agda rename Cubical/{HITs/Ints => Data/Int/MoreInts}/BiInvInt/Base.agda (76%) create mode 100644 Cubical/Data/Int/MoreInts/BiInvInt/Properties.agda create mode 100644 Cubical/Data/Int/MoreInts/DeltaInt.agda rename Cubical/{HITs/Ints => Data/Int/MoreInts}/DeltaInt/Base.agda (64%) rename Cubical/{HITs/Ints => Data/Int/MoreInts}/DeltaInt/Properties.agda (81%) create mode 100644 Cubical/Data/Int/MoreInts/DiffInt.agda create mode 100644 Cubical/Data/Int/MoreInts/DiffInt/Base.agda rename Cubical/Data/{ => Int/MoreInts}/DiffInt/Properties.agda (87%) create mode 100644 Cubical/Data/Int/MoreInts/QuoInt.agda create mode 100644 Cubical/Data/Int/MoreInts/QuoInt/Base.agda create mode 100644 Cubical/Data/Int/MoreInts/QuoInt/Properties.agda create mode 100644 Cubical/Data/Nat/Coprime.agda create mode 100644 Cubical/Data/Nat/Divisibility.agda create mode 100644 Cubical/Data/Nat/GCD.agda create mode 100644 Cubical/Data/Nat/Literals.agda create mode 100644 Cubical/Data/Nat/Order/Recursive.agda create mode 100644 Cubical/Data/NatMinusOne/Properties.agda delete mode 100644 Cubical/Data/NatMinusTwo.agda delete mode 100644 Cubical/Data/NatMinusTwo/Base.agda create mode 100644 Cubical/Data/NatPlusOne.agda create mode 100644 Cubical/Data/NatPlusOne/Base.agda create mode 100644 Cubical/Data/NatPlusOne/Properties.agda create mode 100644 Cubical/Data/Queue.agda create mode 100644 Cubical/Data/Queue/1List.agda create mode 100644 Cubical/Data/Queue/Base.agda create mode 100644 Cubical/Data/Queue/Finite.agda create mode 100644 Cubical/Data/Queue/Truncated2List.agda create mode 100644 Cubical/Data/Queue/Untruncated2List.agda create mode 100644 Cubical/Data/Queue/Untruncated2ListInvariant.agda create mode 100644 Cubical/Data/SubFinSet.agda create mode 100644 Cubical/Data/SumFin/Properties.agda create mode 100644 Cubical/Data/Vec.agda create mode 100644 Cubical/Data/Vec/Base.agda create mode 100644 Cubical/Data/Vec/NAry.agda create mode 100644 Cubical/Data/Vec/Properties.agda create mode 100644 Cubical/Displayed/Auto.agda create mode 100644 Cubical/Displayed/Base.agda create mode 100644 Cubical/Displayed/Constant.agda create mode 100644 Cubical/Displayed/Everything.agda create mode 100644 Cubical/Displayed/Function.agda create mode 100644 Cubical/Displayed/Generic.agda create mode 100644 Cubical/Displayed/Morphism.agda create mode 100644 Cubical/Displayed/Prop.agda create mode 100644 Cubical/Displayed/Properties.agda create mode 100644 Cubical/Displayed/Record.agda create mode 100644 Cubical/Displayed/Sigma.agda create mode 100644 Cubical/Displayed/Subst.agda create mode 100644 Cubical/Displayed/Unit.agda create mode 100644 Cubical/Displayed/Universe.agda create mode 100644 Cubical/Experiments/CohomologyGroups.agda create mode 100644 Cubical/Experiments/HAEquivInt.agda rename Cubical/{HITs/Ints => Experiments}/HAEquivInt/Base.agda (67%) rename Cubical/{Foundations => Experiments}/HoTT-UF.agda (96%) create mode 100644 Cubical/Experiments/IsoInt.agda rename Cubical/{HITs/Ints => Experiments}/IsoInt/Base.agda (83%) create mode 100644 Cubical/Experiments/List.agda create mode 100644 Cubical/Experiments/LocalisationDefs.agda create mode 100644 Cubical/Experiments/NatMinusTwo.agda create mode 100644 Cubical/Experiments/NatMinusTwo/Base.agda create mode 100644 Cubical/Experiments/NatMinusTwo/Properties.agda create mode 100644 Cubical/Experiments/NatMinusTwo/ToNatMinusOne.agda create mode 100644 Cubical/Experiments/ZCohomology/Benchmarks.agda create mode 100644 Cubical/Experiments/ZCohomologyExperiments.agda create mode 100644 Cubical/Experiments/ZCohomologyOld/Base.agda create mode 100644 Cubical/Experiments/ZCohomologyOld/KcompPrelims.agda create mode 100644 Cubical/Experiments/ZCohomologyOld/MayerVietorisUnreduced.agda create mode 100644 Cubical/Experiments/ZCohomologyOld/Properties.agda delete mode 100644 Cubical/Foundations/Embedding.agda create mode 100644 Cubical/Foundations/Equiv/Base.agda rename Cubical/Foundations/{BiInvEquiv.agda => Equiv/BiInvertible.agda} (87%) create mode 100644 Cubical/Foundations/Equiv/HalfAdjoint.agda rename Cubical/Foundations/{PathSplitEquiv.agda => Equiv/PathSplit.agda} (73%) delete mode 100644 Cubical/Foundations/FunExtEquiv.agda delete mode 100644 Cubical/Foundations/HAEquiv.agda create mode 100644 Cubical/Foundations/Pointed/FunExt.agda create mode 100644 Cubical/Foundations/Pointed/Homotopy.agda create mode 100644 Cubical/Foundations/Powerset.agda create mode 100644 Cubical/Foundations/RelationalStructure.agda create mode 100644 Cubical/Foundations/Structure.agda delete mode 100644 Cubical/Foundations/Surjection.agda delete mode 100644 Cubical/Foundations/TotalFiber.agda delete mode 100644 Cubical/Foundations/UnivalenceId.agda create mode 100644 Cubical/Functions/Bundle.agda create mode 100644 Cubical/Functions/Embedding.agda create mode 100644 Cubical/Functions/Everything.agda create mode 100644 Cubical/Functions/Fibration.agda create mode 100644 Cubical/Functions/Fixpoint.agda create mode 100644 Cubical/Functions/FunExtEquiv.agda create mode 100644 Cubical/Functions/Implicit.agda create mode 100644 Cubical/Functions/Involution.agda rename Cubical/{Foundations => Functions}/Logic.agda (53%) create mode 100644 Cubical/Functions/Morphism.agda create mode 100644 Cubical/Functions/Surjection.agda create mode 100644 Cubical/HITs/AssocList.agda create mode 100644 Cubical/HITs/AssocList/Base.agda create mode 100644 Cubical/HITs/AssocList/Properties.agda create mode 100644 Cubical/HITs/Cost.agda create mode 100644 Cubical/HITs/Cost/Base.agda create mode 100644 Cubical/HITs/Delooping/Two/Base.agda create mode 100644 Cubical/HITs/Delooping/Two/Properties.agda create mode 100644 Cubical/HITs/EilenbergMacLane1.agda create mode 100644 Cubical/HITs/EilenbergMacLane1/Base.agda create mode 100644 Cubical/HITs/EilenbergMacLane1/Properties.agda create mode 100644 Cubical/HITs/FiniteMultiset/CountExtensionality.agda create mode 100644 Cubical/HITs/GroupoidQuotients.agda create mode 100644 Cubical/HITs/GroupoidQuotients/Base.agda create mode 100644 Cubical/HITs/GroupoidQuotients/Properties.agda delete mode 100644 Cubical/HITs/Ints/BiInvInt.agda delete mode 100644 Cubical/HITs/Ints/BiInvInt/Properties.agda delete mode 100644 Cubical/HITs/Ints/DeltaInt.agda delete mode 100644 Cubical/HITs/Ints/HAEquivInt.agda delete mode 100644 Cubical/HITs/Ints/IsoInt.agda delete mode 100644 Cubical/HITs/Ints/QuoInt.agda delete mode 100644 Cubical/HITs/Ints/QuoInt/Base.agda delete mode 100644 Cubical/HITs/Ints/QuoInt/Properties.agda create mode 100644 Cubical/HITs/PropositionalTruncation/Monad.agda create mode 100644 Cubical/HITs/Pushout/Flattening.agda create mode 100644 Cubical/HITs/Pushout/KrausVonRaumer.agda create mode 100644 Cubical/HITs/RPn.agda create mode 100644 Cubical/HITs/RPn/Base.agda delete mode 100644 Cubical/HITs/Rational.agda delete mode 100644 Cubical/HITs/Rational/Base.agda create mode 100644 Cubical/HITs/Rationals/HITQ.agda create mode 100644 Cubical/HITs/Rationals/HITQ/Base.agda create mode 100644 Cubical/HITs/Rationals/QuoQ.agda create mode 100644 Cubical/HITs/Rationals/QuoQ/Base.agda create mode 100644 Cubical/HITs/Rationals/QuoQ/Properties.agda create mode 100644 Cubical/HITs/Rationals/SigmaQ.agda create mode 100644 Cubical/HITs/Rationals/SigmaQ/Base.agda create mode 100644 Cubical/HITs/Rationals/SigmaQ/Properties.agda create mode 100644 Cubical/HITs/Sn/Properties.agda create mode 100644 Cubical/HITs/Susp/Properties.agda delete mode 100644 Cubical/HITs/Truncation/FromNegOne.agda delete mode 100644 Cubical/HITs/Truncation/FromNegOne/Base.agda delete mode 100644 Cubical/HITs/Truncation/FromNegOne/Properties.agda create mode 100644 Cubical/HITs/Truncation/FromNegTwo.agda create mode 100644 Cubical/HITs/Truncation/FromNegTwo/Base.agda create mode 100644 Cubical/HITs/Truncation/FromNegTwo/Properties.agda create mode 100644 Cubical/HITs/TypeQuotients.agda create mode 100644 Cubical/HITs/TypeQuotients/Base.agda create mode 100644 Cubical/HITs/TypeQuotients/Properties.agda create mode 100644 Cubical/HITs/Wedge.agda create mode 100644 Cubical/HITs/Wedge/Base.agda create mode 100644 Cubical/Homotopy/Base.agda create mode 100644 Cubical/Homotopy/Connected.agda create mode 100644 Cubical/Homotopy/EilenbergSteenrod.agda create mode 100644 Cubical/Homotopy/Everything.agda create mode 100644 Cubical/Homotopy/Freudenthal.agda create mode 100644 Cubical/Homotopy/Loopspace.agda create mode 100644 Cubical/Homotopy/MayerVietorisCofiber.agda create mode 100644 Cubical/Homotopy/WedgeConnectivity.agda create mode 100644 Cubical/Modalities/Lex.agda create mode 100644 Cubical/Papers/Everything.agda create mode 100644 Cubical/Papers/RepresentationIndependence.agda create mode 100644 Cubical/Papers/Synthetic.agda create mode 100644 Cubical/Papers/ZCohomology.agda create mode 100644 Cubical/Reflection/Base.agda create mode 100644 Cubical/Reflection/Everything.agda create mode 100644 Cubical/Reflection/RecordEquiv.agda create mode 100644 Cubical/Reflection/StrictEquiv.agda create mode 100644 Cubical/Relation/Binary/Poset.agda create mode 100644 Cubical/Relation/Nullary/Base.agda create mode 100644 Cubical/Relation/Nullary/HLevels.agda create mode 100644 Cubical/Relation/Nullary/Properties.agda create mode 100644 Cubical/Relation/ZigZag/Applications/MultiSet.agda create mode 100644 Cubical/Relation/ZigZag/Base.agda create mode 100644 Cubical/Structures/Auto.agda create mode 100644 Cubical/Structures/Axioms.agda create mode 100644 Cubical/Structures/Constant.agda create mode 100644 Cubical/Structures/Everything.agda create mode 100644 Cubical/Structures/Function.agda create mode 100644 Cubical/Structures/LeftAction.agda create mode 100644 Cubical/Structures/Macro.agda create mode 100644 Cubical/Structures/Maybe.agda create mode 100644 Cubical/Structures/MultiSet.agda create mode 100644 Cubical/Structures/Parameterized.agda create mode 100644 Cubical/Structures/Pointed.agda create mode 100644 Cubical/Structures/Product.agda create mode 100644 Cubical/Structures/Queue.agda create mode 100644 Cubical/Structures/Record.agda create mode 100644 Cubical/Structures/Relational/Auto.agda create mode 100644 Cubical/Structures/Relational/Constant.agda create mode 100644 Cubical/Structures/Relational/Equalizer.agda create mode 100644 Cubical/Structures/Relational/Function.agda create mode 100644 Cubical/Structures/Relational/Macro.agda create mode 100644 Cubical/Structures/Relational/Maybe.agda create mode 100644 Cubical/Structures/Relational/Parameterized.agda create mode 100644 Cubical/Structures/Relational/Pointed.agda create mode 100644 Cubical/Structures/Relational/Product.agda create mode 100644 Cubical/Structures/Transfer.agda create mode 100644 Cubical/Structures/TypeEqvTo.agda create mode 100644 Cubical/Talks/EPA2020.agda create mode 100644 Cubical/Talks/Everything.agda create mode 100644 Cubical/ZCohomology/#Benchmarks.agda# create mode 100644 Cubical/ZCohomology/Benchmarks.agda create mode 100644 Cubical/ZCohomology/EilenbergSteenrodZ.agda delete mode 100644 Cubical/ZCohomology/FunLoopSpace.agda delete mode 100644 Cubical/ZCohomology/FunLoopSpace2.agda create mode 100644 Cubical/ZCohomology/GroupStructure.agda create mode 100644 Cubical/ZCohomology/Groups/CP2.agda create mode 100644 Cubical/ZCohomology/Groups/Connected.agda create mode 100644 Cubical/ZCohomology/Groups/KleinBottle.agda create mode 100644 Cubical/ZCohomology/Groups/Prelims.agda create mode 100644 Cubical/ZCohomology/Groups/RP2.agda create mode 100644 Cubical/ZCohomology/Groups/Sn.agda create mode 100644 Cubical/ZCohomology/Groups/Torus.agda create mode 100644 Cubical/ZCohomology/Groups/Unit.agda create mode 100644 Cubical/ZCohomology/Groups/Wedge.agda create mode 100644 Cubical/ZCohomology/Groups/WedgeOfSpheres.agda create mode 100644 Cubical/ZCohomology/MayerVietorisUnreduced.agda delete mode 100644 Cubical/ZCohomology/PropertiesTrash.agda create mode 100644 Cubical/ZCohomology/RingStructure/CupProduct.agda create mode 100644 Cubical/ZCohomology/RingStructure/GradedCommutativity.agda create mode 100644 Cubical/ZCohomology/RingStructure/RingLaws.agda delete mode 100644 Cubical/ZCohomology/S1/S1.agda diff --git a/Cubical/Algebra/AbGroup.agda b/Cubical/Algebra/AbGroup.agda new file mode 100644 index 000000000..dd0ffe841 --- /dev/null +++ b/Cubical/Algebra/AbGroup.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.AbGroup where + +open import Cubical.Algebra.AbGroup.Base public diff --git a/Cubical/Algebra/AbGroup/Base.agda b/Cubical/Algebra/AbGroup/Base.agda new file mode 100644 index 000000000..84dc2c07b --- /dev/null +++ b/Cubical/Algebra/AbGroup/Base.agda @@ -0,0 +1,192 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.AbGroup.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP +open import Cubical.Data.Sigma +open import Cubical.Data.Unit +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.Group + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open Iso + +private + variable + ℓ ℓ' : Level + +record IsAbGroup {G : Type ℓ} + (0g : G) (_+_ : G → G → G) (-_ : G → G) : Type ℓ where + + constructor isabgroup + + field + isGroup : IsGroup 0g _+_ -_ + comm : (x y : G) → x + y ≡ y + x + + open IsGroup isGroup public + +record AbGroupStr (A : Type ℓ) : Type (ℓ-suc ℓ) where + + constructor abgroupstr + + field + 0g : A + _+_ : A → A → A + -_ : A → A + isAbGroup : IsAbGroup 0g _+_ -_ + + infix 8 -_ + infixr 7 _+_ + + open IsAbGroup isAbGroup public + +AbGroup : ∀ ℓ → Type (ℓ-suc ℓ) +AbGroup ℓ = TypeWithStr ℓ AbGroupStr + +makeIsAbGroup : {G : Type ℓ} {0g : G} {_+_ : G → G → G} { -_ : G → G} + (is-setG : isSet G) + (assoc : (x y z : G) → x + (y + z) ≡ (x + y) + z) + (rid : (x : G) → x + 0g ≡ x) + (rinv : (x : G) → x + (- x) ≡ 0g) + (comm : (x y : G) → x + y ≡ y + x) + → IsAbGroup 0g _+_ -_ +makeIsAbGroup is-setG assoc rid rinv comm = + isabgroup (makeIsGroup is-setG assoc rid (λ x → comm _ _ ∙ rid x) rinv (λ x → comm _ _ ∙ rinv x)) comm + +makeAbGroup : {G : Type ℓ} (0g : G) (_+_ : G → G → G) (-_ : G → G) + (is-setG : isSet G) + (assoc : (x y z : G) → x + (y + z) ≡ (x + y) + z) + (rid : (x : G) → x + 0g ≡ x) + (rinv : (x : G) → x + (- x) ≡ 0g) + (comm : (x y : G) → x + y ≡ y + x) + → AbGroup ℓ +makeAbGroup 0g _+_ -_ is-setG assoc rid rinv comm = + _ , abgroupstr 0g _+_ -_ (makeIsAbGroup is-setG assoc rid rinv comm) + +open GroupStr +open AbGroupStr +open IsAbGroup + +AbGroupStr→GroupStr : {G : Type ℓ} → AbGroupStr G → GroupStr G +AbGroupStr→GroupStr A .1g = A .0g +AbGroupStr→GroupStr A ._·_ = A ._+_ +AbGroupStr→GroupStr A .inv = A .-_ +AbGroupStr→GroupStr A .isGroup = A .isAbGroup .isGroup + +AbGroup→Group : AbGroup ℓ → Group ℓ +fst (AbGroup→Group A) = fst A +snd (AbGroup→Group A) = AbGroupStr→GroupStr (snd A) + +Group→AbGroup : (G : Group ℓ) → ((x y : fst G) → _·_ (snd G) x y ≡ _·_ (snd G) y x) → AbGroup ℓ +fst (Group→AbGroup G comm) = fst G +AbGroupStr.0g (snd (Group→AbGroup G comm)) = 1g (snd G) +AbGroupStr._+_ (snd (Group→AbGroup G comm)) = _·_ (snd G) +AbGroupStr.- snd (Group→AbGroup G comm) = inv (snd G) +IsAbGroup.isGroup (AbGroupStr.isAbGroup (snd (Group→AbGroup G comm))) = isGroup (snd G) +IsAbGroup.comm (AbGroupStr.isAbGroup (snd (Group→AbGroup G comm))) = comm + +isSetAbGroup : (A : AbGroup ℓ) → isSet ⟨ A ⟩ +isSetAbGroup A = isSetGroup (AbGroup→Group A) + +AbGroupHom : (G : AbGroup ℓ) (H : AbGroup ℓ') → Type (ℓ-max ℓ ℓ') +AbGroupHom G H = GroupHom (AbGroup→Group G) (AbGroup→Group H) + +IsAbGroupEquiv : {A : Type ℓ} {B : Type ℓ'} + (G : AbGroupStr A) (e : A ≃ B) (H : AbGroupStr B) → Type (ℓ-max ℓ ℓ') +IsAbGroupEquiv G e H = IsGroupHom (AbGroupStr→GroupStr G) (e .fst) (AbGroupStr→GroupStr H) + +AbGroupEquiv : (G : AbGroup ℓ) (H : AbGroup ℓ') → Type (ℓ-max ℓ ℓ') +AbGroupEquiv G H = Σ[ e ∈ (G .fst ≃ H .fst) ] IsAbGroupEquiv (G .snd) e (H .snd) + +isPropIsAbGroup : {G : Type ℓ} (0g : G) (_+_ : G → G → G) (- : G → G) + → isProp (IsAbGroup 0g _+_ -) +isPropIsAbGroup 0g _+_ -_ (isabgroup GG GC) (isabgroup HG HC) = + λ i → isabgroup (isPropIsGroup _ _ _ GG HG i) (isPropComm GC HC i) + where + isSetG : isSet _ + isSetG = GG .IsGroup.isMonoid .IsMonoid.isSemigroup .IsSemigroup.is-set + + isPropComm : isProp ((x y : _) → x + y ≡ y + x) + isPropComm = isPropΠ2 λ _ _ → isSetG _ _ + +𝒮ᴰ-AbGroup : DUARel (𝒮-Univ ℓ) AbGroupStr ℓ +𝒮ᴰ-AbGroup = + 𝒮ᴰ-Record (𝒮-Univ _) IsAbGroupEquiv + (fields: + data[ _+_ ∣ autoDUARel _ _ ∣ pres· ] + data[ 0g ∣ autoDUARel _ _ ∣ pres1 ] + data[ -_ ∣ autoDUARel _ _ ∣ presinv ] + prop[ isAbGroup ∣ (λ _ _ → isPropIsAbGroup _ _ _) ]) + where + open AbGroupStr + open IsGroupHom + +-- Extract the characterization of equality of groups +AbGroupPath : (G H : AbGroup ℓ) → (AbGroupEquiv G H) ≃ (G ≡ H) +AbGroupPath = ∫ 𝒮ᴰ-AbGroup .UARel.ua + +-- TODO: Induced structure results are temporarily inconvenient while we transition between algebra +-- representations +module _ (G : AbGroup ℓ) {A : Type ℓ} (m : A → A → A) + (e : ⟨ G ⟩ ≃ A) + (p· : ∀ x y → e .fst (G .snd ._+_ x y) ≡ m (e .fst x) (e .fst y)) + where + + private + module G = AbGroupStr (G .snd) + + FamilyΣ : Σ[ B ∈ Type ℓ ] (B → B → B) → Type ℓ + FamilyΣ (B , n) = + Σ[ e ∈ B ] + Σ[ i ∈ (B → B) ] + IsAbGroup e n i + + inducedΣ : FamilyΣ (A , m) + inducedΣ = + subst FamilyΣ + (UARel.≅→≡ (autoUARel (Σ[ B ∈ Type ℓ ] (B → B → B))) (e , p·)) + (G.0g , G.-_ , G.isAbGroup) + + InducedAbGroup : AbGroup ℓ + InducedAbGroup .fst = A + InducedAbGroup .snd ._+_ = m + InducedAbGroup .snd .0g = inducedΣ .fst + InducedAbGroup .snd .-_ = inducedΣ .snd .fst + InducedAbGroup .snd .isAbGroup = inducedΣ .snd .snd + + InducedAbGroupPath : G ≡ InducedAbGroup + InducedAbGroupPath = AbGroupPath _ _ .fst (e , makeIsGroupHom p·) + +open IsMonoid +open IsSemigroup +open IsGroup + +dirProdAb : AbGroup ℓ → AbGroup ℓ' → AbGroup (ℓ-max ℓ ℓ') +dirProdAb A B = + Group→AbGroup (DirProd (AbGroup→Group A) (AbGroup→Group B)) + λ p q → ΣPathP (comm (isAbGroup (snd A)) _ _ + , comm (isAbGroup (snd B)) _ _) + +trivialAbGroup : ∀ {ℓ} → AbGroup ℓ +fst trivialAbGroup = Unit* +0g (snd trivialAbGroup) = tt* +_+_ (snd trivialAbGroup) _ _ = tt* +(- snd trivialAbGroup) _ = tt* +is-set (isSemigroup (isMonoid (isGroup (isAbGroup (snd trivialAbGroup))))) = + isProp→isSet isPropUnit* +assoc (isSemigroup (isMonoid (isGroup (isAbGroup (snd trivialAbGroup))))) _ _ _ = refl +identity (isMonoid (isGroup (isAbGroup (snd trivialAbGroup)))) _ = refl , refl +inverse (isGroup (isAbGroup (snd trivialAbGroup))) _ = refl , refl +comm (isAbGroup (snd trivialAbGroup)) _ _ = refl diff --git a/Cubical/Algebra/Algebra.agda b/Cubical/Algebra/Algebra.agda new file mode 100644 index 000000000..2f9c7d908 --- /dev/null +++ b/Cubical/Algebra/Algebra.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Algebra where + +open import Cubical.Algebra.Algebra.Base public diff --git a/Cubical/Algebra/Algebra/Base.agda b/Cubical/Algebra/Algebra/Base.agda new file mode 100644 index 000000000..6a108996b --- /dev/null +++ b/Cubical/Algebra/Algebra/Base.agda @@ -0,0 +1,305 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Algebra.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open import Cubical.Reflection.RecordEquiv + +open import Cubical.Algebra.Module +open import Cubical.Algebra.Ring +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Group +open import Cubical.Algebra.Monoid + +open Iso + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + +record IsAlgebra (R : Ring ℓ) {A : Type ℓ'} + (0a 1a : A) (_+_ _·_ : A → A → A) (-_ : A → A) + (_⋆_ : ⟨ R ⟩ → A → A) : Type (ℓ-max ℓ ℓ') where + + constructor isalgebra + + open RingStr (snd R) using (1r) renaming (_+_ to _+r_; _·_ to _·r_) + + field + isLeftModule : IsLeftModule R 0a _+_ -_ _⋆_ + ·-isMonoid : IsMonoid 1a _·_ + dist : (x y z : A) → (x · (y + z) ≡ (x · y) + (x · z)) + × ((x + y) · z ≡ (x · z) + (y · z)) + ⋆-lassoc : (r : ⟨ R ⟩) (x y : A) → (r ⋆ x) · y ≡ r ⋆ (x · y) + ⋆-rassoc : (r : ⟨ R ⟩) (x y : A) → r ⋆ (x · y) ≡ x · (r ⋆ y) + + open IsLeftModule isLeftModule public + + isRing : IsRing _ _ _ _ _ + isRing = isring (IsLeftModule.+-isAbGroup isLeftModule) ·-isMonoid dist + open IsRing isRing public hiding (_-_; +Assoc; +Lid; +Linv; +Rid; +Rinv; +Comm) + +unquoteDecl IsAlgebraIsoΣ = declareRecordIsoΣ IsAlgebraIsoΣ (quote IsAlgebra) + +record AlgebraStr (R : Ring ℓ) (A : Type ℓ') : Type (ℓ-max ℓ ℓ') where + + constructor algebrastr + + field + 0a : A + 1a : A + _+_ : A → A → A + _·_ : A → A → A + -_ : A → A + _⋆_ : ⟨ R ⟩ → A → A + isAlgebra : IsAlgebra R 0a 1a _+_ _·_ -_ _⋆_ + + open IsAlgebra isAlgebra public + +Algebra : (R : Ring ℓ) → ∀ ℓ' → Type (ℓ-max ℓ (ℓ-suc ℓ')) +Algebra R ℓ' = Σ[ A ∈ Type ℓ' ] AlgebraStr R A + +module commonExtractors {R : Ring ℓ} where + + Algebra→Module : (A : Algebra R ℓ') → LeftModule R ℓ' + Algebra→Module (_ , algebrastr A _ _ _ _ _ (isalgebra isLeftModule _ _ _ _)) = + _ , leftmodulestr A _ _ _ isLeftModule + + Algebra→Ring : (A : Algebra R ℓ') → Ring ℓ' + Algebra→Ring (_ , str) = _ , ringstr _ _ _ _ _ (IsAlgebra.isRing (AlgebraStr.isAlgebra str)) + + Algebra→AbGroup : (A : Algebra R ℓ') → AbGroup ℓ' + Algebra→AbGroup A = LeftModule→AbGroup (Algebra→Module A) + + Algebra→Group : (A : Algebra R ℓ') → Group ℓ' + Algebra→Group A = Ring→Group (Algebra→Ring A) + + Algebra→AddMonoid : (A : Algebra R ℓ') → Monoid ℓ' + Algebra→AddMonoid A = Group→Monoid (Algebra→Group A) + + Algebra→MultMonoid : (A : Algebra R ℓ') → Monoid ℓ' + Algebra→MultMonoid A = Ring→MultMonoid (Algebra→Ring A) + + isSetAlgebra : (A : Algebra R ℓ') → isSet ⟨ A ⟩ + isSetAlgebra A = isSetAbGroup (Algebra→AbGroup A) + + open RingStr (snd R) using (1r; ·Ldist+) renaming (_+_ to _+r_; _·_ to _·s_) + + makeIsAlgebra : {A : Type ℓ'} {0a 1a : A} + {_+_ _·_ : A → A → A} { -_ : A → A} {_⋆_ : ⟨ R ⟩ → A → A} + (isSet-A : isSet A) + (+-assoc : (x y z : A) → x + (y + z) ≡ (x + y) + z) + (+-rid : (x : A) → x + 0a ≡ x) + (+-rinv : (x : A) → x + (- x) ≡ 0a) + (+-comm : (x y : A) → x + y ≡ y + x) + (·-assoc : (x y z : A) → x · (y · z) ≡ (x · y) · z) + (·-rid : (x : A) → x · 1a ≡ x) + (·-lid : (x : A) → 1a · x ≡ x) + (·-rdist-+ : (x y z : A) → x · (y + z) ≡ (x · y) + (x · z)) + (·-ldist-+ : (x y z : A) → (x + y) · z ≡ (x · z) + (y · z)) + (⋆-assoc : (r s : ⟨ R ⟩) (x : A) → (r ·s s) ⋆ x ≡ r ⋆ (s ⋆ x)) + (⋆-ldist : (r s : ⟨ R ⟩) (x : A) → (r +r s) ⋆ x ≡ (r ⋆ x) + (s ⋆ x)) + (⋆-rdist : (r : ⟨ R ⟩) (x y : A) → r ⋆ (x + y) ≡ (r ⋆ x) + (r ⋆ y)) + (⋆-lid : (x : A) → 1r ⋆ x ≡ x) + (⋆-lassoc : (r : ⟨ R ⟩) (x y : A) → (r ⋆ x) · y ≡ r ⋆ (x · y)) + (⋆-rassoc : (r : ⟨ R ⟩) (x y : A) → r ⋆ (x · y) ≡ x · (r ⋆ y)) + → IsAlgebra R 0a 1a _+_ _·_ -_ _⋆_ + makeIsAlgebra isSet-A + +-assoc +-rid +-rinv +-comm + ·-assoc ·-rid ·-lid ·-rdist-+ ·-ldist-+ + ⋆-assoc ⋆-ldist ⋆-rdist ⋆-lid ⋆-lassoc ⋆-rassoc = + isalgebra + (makeIsLeftModule isSet-A + +-assoc +-rid +-rinv +-comm + ⋆-assoc ⋆-ldist ⋆-rdist ⋆-lid) + (makeIsMonoid isSet-A ·-assoc ·-rid ·-lid) + (λ x y z → ·-rdist-+ x y z , ·-ldist-+ x y z) + ⋆-lassoc ⋆-rassoc + + +open commonExtractors public + +record IsAlgebraHom {R : Ring ℓ} {A : Type ℓ'} {B : Type ℓ''} + (M : AlgebraStr R A) (f : A → B) (N : AlgebraStr R B) + : Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) + where + + -- Shorter qualified names + private + module M = AlgebraStr M + module N = AlgebraStr N + + field + pres0 : f M.0a ≡ N.0a + pres1 : f M.1a ≡ N.1a + pres+ : (x y : A) → f (x M.+ y) ≡ f x N.+ f y + pres· : (x y : A) → f (x M.· y) ≡ f x N.· f y + pres- : (x : A) → f (M.- x) ≡ N.- (f x) + pres⋆ : (r : ⟨ R ⟩) (y : A) → f (r M.⋆ y) ≡ r N.⋆ f y + +unquoteDecl IsAlgebraHomIsoΣ = declareRecordIsoΣ IsAlgebraHomIsoΣ (quote IsAlgebraHom) +open IsAlgebraHom + +AlgebraHom : {R : Ring ℓ} (M : Algebra R ℓ') (N : Algebra R ℓ'') → Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) +AlgebraHom M N = Σ[ f ∈ (⟨ M ⟩ → ⟨ N ⟩) ] IsAlgebraHom (M .snd) f (N .snd) + +idAlgHom : {R : Ring ℓ} {A : Algebra R ℓ'} → AlgebraHom A A +fst idAlgHom x = x +pres0 (snd idAlgHom) = refl +pres1 (snd idAlgHom) = refl +pres+ (snd idAlgHom) x y = refl +pres· (snd idAlgHom) x y = refl +pres- (snd idAlgHom) x = refl +pres⋆ (snd idAlgHom) r x = refl + +IsAlgebraEquiv : {R : Ring ℓ} {A B : Type ℓ'} + (M : AlgebraStr R A) (e : A ≃ B) (N : AlgebraStr R B) + → Type (ℓ-max ℓ ℓ') +IsAlgebraEquiv M e N = IsAlgebraHom M (e .fst) N + +AlgebraEquiv : {R : Ring ℓ} (M N : Algebra R ℓ') → Type (ℓ-max ℓ ℓ') +AlgebraEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsAlgebraEquiv (M .snd) e (N .snd) + +_$a_ : {R : Ring ℓ} {A : Algebra R ℓ'} {B : Algebra R ℓ''} → AlgebraHom A B → ⟨ A ⟩ → ⟨ B ⟩ +f $a x = fst f x + +AlgebraEquiv→AlgebraHom : {R : Ring ℓ} {A B : Algebra R ℓ'} + → AlgebraEquiv A B → AlgebraHom A B +AlgebraEquiv→AlgebraHom (e , eIsHom) = e .fst , eIsHom + +isPropIsAlgebra : (R : Ring ℓ) {A : Type ℓ'} + (0a 1a : A) + (_+_ _·_ : A → A → A) + (-_ : A → A) + (_⋆_ : ⟨ R ⟩ → A → A) + → isProp (IsAlgebra R 0a 1a _+_ _·_ -_ _⋆_) +isPropIsAlgebra R _ _ _ _ _ _ = let open IsLeftModule in + isOfHLevelRetractFromIso 1 IsAlgebraIsoΣ + (isPropΣ + (isPropIsLeftModule _ _ _ _ _) + (λ mo → isProp×3 (isPropIsMonoid _ _) + (isPropΠ3 λ _ _ _ → isProp× (mo .is-set _ _) (mo .is-set _ _)) + (isPropΠ3 λ _ _ _ → mo .is-set _ _) + (isPropΠ3 λ _ _ _ → mo .is-set _ _) )) + + +isPropIsAlgebraHom : (R : Ring ℓ) {A : Type ℓ'} {B : Type ℓ''} + (AS : AlgebraStr R A) (f : A → B) (BS : AlgebraStr R B) + → isProp (IsAlgebraHom AS f BS) +isPropIsAlgebraHom R AS f BS = isOfHLevelRetractFromIso 1 IsAlgebraHomIsoΣ + (isProp×5 (isSetAlgebra (_ , BS) _ _) + (isSetAlgebra (_ , BS) _ _) + (isPropΠ2 λ _ _ → isSetAlgebra (_ , BS) _ _) + (isPropΠ2 λ _ _ → isSetAlgebra (_ , BS) _ _) + (isPropΠ λ _ → isSetAlgebra (_ , BS) _ _) + (isPropΠ2 λ _ _ → isSetAlgebra (_ , BS) _ _)) + +isSetAlgebraHom : {R : Ring ℓ} (M : Algebra R ℓ') (N : Algebra R ℓ'') + → isSet (AlgebraHom M N) +isSetAlgebraHom _ N = isSetΣ (isSetΠ (λ _ → isSetAlgebra N)) + λ _ → isProp→isSet (isPropIsAlgebraHom _ _ _ _) + + +isSetAlgebraEquiv : {R : Ring ℓ} (M N : Algebra R ℓ') + → isSet (AlgebraEquiv M N) +isSetAlgebraEquiv M N = isSetΣ (isOfHLevel≃ 2 (isSetAlgebra M) (isSetAlgebra N)) + λ _ → isProp→isSet (isPropIsAlgebraHom _ _ _ _) + +𝒮ᴰ-Algebra : (R : Ring ℓ) → DUARel (𝒮-Univ ℓ') (AlgebraStr R) (ℓ-max ℓ ℓ') +𝒮ᴰ-Algebra R = + 𝒮ᴰ-Record (𝒮-Univ _) (IsAlgebraEquiv {R = R}) + (fields: + data[ 0a ∣ nul ∣ pres0 ] + data[ 1a ∣ nul ∣ pres1 ] + data[ _+_ ∣ bin ∣ pres+ ] + data[ _·_ ∣ bin ∣ pres· ] + data[ -_ ∣ autoDUARel _ _ ∣ pres- ] + data[ _⋆_ ∣ autoDUARel _ _ ∣ pres⋆ ] + prop[ isAlgebra ∣ (λ _ _ → isPropIsAlgebra _ _ _ _ _ _ _) ]) + where + open AlgebraStr + + -- faster with some sharing + nul = autoDUARel (𝒮-Univ _) (λ A → A) + bin = autoDUARel (𝒮-Univ _) (λ A → A → A → A) + +AlgebraPath : {R : Ring ℓ} (A B : Algebra R ℓ') → (AlgebraEquiv A B) ≃ (A ≡ B) +AlgebraPath {R = R} = ∫ (𝒮ᴰ-Algebra R) .UARel.ua + +compIsAlgebraHom : {R : Ring ℓ} {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''} + {g : ⟨ B ⟩ → ⟨ C ⟩} {f : ⟨ A ⟩ → ⟨ B ⟩} + → IsAlgebraHom (B .snd) g (C .snd) + → IsAlgebraHom (A .snd) f (B .snd) + → IsAlgebraHom (A .snd) (g ∘ f) (C .snd) +compIsAlgebraHom {g = g} {f} gh fh .pres0 = cong g (fh .pres0) ∙ gh .pres0 +compIsAlgebraHom {g = g} {f} gh fh .pres1 = cong g (fh .pres1) ∙ gh .pres1 +compIsAlgebraHom {g = g} {f} gh fh .pres+ x y = cong g (fh .pres+ x y) ∙ gh .pres+ (f x) (f y) +compIsAlgebraHom {g = g} {f} gh fh .pres· x y = cong g (fh .pres· x y) ∙ gh .pres· (f x) (f y) +compIsAlgebraHom {g = g} {f} gh fh .pres- x = cong g (fh .pres- x) ∙ gh .pres- (f x) +compIsAlgebraHom {g = g} {f} gh fh .pres⋆ r x = cong g (fh .pres⋆ r x) ∙ gh .pres⋆ r (f x) + +_∘a_ : {R : Ring ℓ} {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''} + → AlgebraHom B C → AlgebraHom A B → AlgebraHom A C +_∘a_ g f .fst = g .fst ∘ f .fst +_∘a_ g f .snd = compIsAlgebraHom (g .snd) (f .snd) + +module AlgebraTheory (R : Ring ℓ) (A : Algebra R ℓ') where + open RingStr (snd R) renaming (_+_ to _+r_ ; _·_ to _·r_) + open AlgebraStr (A .snd) + + 0-actsNullifying : (x : ⟨ A ⟩) → 0r ⋆ x ≡ 0a + 0-actsNullifying x = + let idempotent-+ = 0r ⋆ x ≡⟨ cong (λ u → u ⋆ x) (sym (RingTheory.0Idempotent R)) ⟩ + (0r +r 0r) ⋆ x ≡⟨ ⋆-ldist 0r 0r x ⟩ + (0r ⋆ x) + (0r ⋆ x) ∎ + in RingTheory.+Idempotency→0 (Algebra→Ring A) (0r ⋆ x) idempotent-+ + + ⋆Dist· : (x y : ⟨ R ⟩) (a b : ⟨ A ⟩) → (x ·r y) ⋆ (a · b) ≡ (x ⋆ a) · (y ⋆ b) + ⋆Dist· x y a b = (x ·r y) ⋆ (a · b) ≡⟨ ⋆-rassoc _ _ _ ⟩ + a · ((x ·r y) ⋆ b) ≡⟨ cong (a ·_) (⋆-assoc _ _ _) ⟩ + a · (x ⋆ (y ⋆ b)) ≡⟨ sym (⋆-rassoc _ _ _) ⟩ + x ⋆ (a · (y ⋆ b)) ≡⟨ sym (⋆-lassoc _ _ _) ⟩ + (x ⋆ a) · (y ⋆ b) ∎ + + +-- Smart constructor for ring homomorphisms +-- that infers the other equations from pres1, pres+, and pres· + +module _ {R : Ring ℓ} {A : Algebra R ℓ} {B : Algebra R ℓ'} {f : ⟨ A ⟩ → ⟨ B ⟩} where + + private + module A = AlgebraStr (A .snd) + module B = AlgebraStr (B .snd) + + module _ + (p1 : f A.1a ≡ B.1a) + (p+ : (x y : ⟨ A ⟩) → f (x A.+ y) ≡ f x B.+ f y) + (p· : (x y : ⟨ A ⟩) → f (x A.· y) ≡ f x B.· f y) + (p⋆ : (r : ⟨ R ⟩) (x : ⟨ A ⟩) → f (r A.⋆ x) ≡ r B.⋆ f x) + where + + open IsAlgebraHom + private + isGHom : IsGroupHom (Algebra→Group A .snd) f (Algebra→Group B .snd) + isGHom = makeIsGroupHom p+ + + makeIsAlgebraHom : IsAlgebraHom (A .snd) f (B .snd) + makeIsAlgebraHom .pres0 = isGHom .IsGroupHom.pres1 + makeIsAlgebraHom .pres1 = p1 + makeIsAlgebraHom .pres+ = p+ + makeIsAlgebraHom .pres· = p· + makeIsAlgebraHom .pres- = isGHom .IsGroupHom.presinv + makeIsAlgebraHom .pres⋆ = p⋆ diff --git a/Cubical/Algebra/CommAlgebra.agda b/Cubical/Algebra/CommAlgebra.agda new file mode 100644 index 000000000..ba2d1ee87 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra where + +open import Cubical.Algebra.CommAlgebra.Base public diff --git a/Cubical/Algebra/CommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/Base.agda new file mode 100644 index 000000000..95b262eba --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/Base.agda @@ -0,0 +1,175 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Algebra + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open import Cubical.Reflection.RecordEquiv + +private + variable + ℓ ℓ' : Level + +record IsCommAlgebra (R : CommRing ℓ) {A : Type ℓ'} + (0a : A) (1a : A) + (_+_ : A → A → A) (_·_ : A → A → A) (-_ : A → A) + (_⋆_ : ⟨ R ⟩ → A → A) : Type (ℓ-max ℓ ℓ') where + + constructor iscommalgebra + + field + isAlgebra : IsAlgebra (CommRing→Ring R) 0a 1a _+_ _·_ -_ _⋆_ + ·-comm : (x y : A) → x · y ≡ y · x + + open IsAlgebra isAlgebra public + +unquoteDecl IsCommAlgebraIsoΣ = declareRecordIsoΣ IsCommAlgebraIsoΣ (quote IsCommAlgebra) + +record CommAlgebraStr (R : CommRing ℓ) (A : Type ℓ') : Type (ℓ-max ℓ ℓ') where + + constructor commalgebrastr + + field + 0a : A + 1a : A + _+_ : A → A → A + _·_ : A → A → A + -_ : A → A + _⋆_ : ⟨ R ⟩ → A → A + isCommAlgebra : IsCommAlgebra R 0a 1a _+_ _·_ -_ _⋆_ + + open IsCommAlgebra isCommAlgebra public + +CommAlgebra : (R : CommRing ℓ) → ∀ ℓ' → Type (ℓ-max ℓ (ℓ-suc ℓ')) +CommAlgebra R ℓ' = Σ[ A ∈ Type ℓ' ] CommAlgebraStr R A + +module _ {R : CommRing ℓ} where + open CommRingStr (snd R) using (1r) renaming (_+_ to _+r_; _·_ to _·s_) + + CommAlgebraStr→AlgebraStr : {A : Type ℓ'} → CommAlgebraStr R A → AlgebraStr (CommRing→Ring R) A + CommAlgebraStr→AlgebraStr (commalgebrastr _ _ _ _ _ _ (iscommalgebra isAlgebra ·-comm)) = + algebrastr _ _ _ _ _ _ isAlgebra + + CommAlgebra→Algebra : (A : CommAlgebra R ℓ') → Algebra (CommRing→Ring R) ℓ' + CommAlgebra→Algebra (_ , str) = (_ , CommAlgebraStr→AlgebraStr str) + + CommAlgebra→CommRing : (A : CommAlgebra R ℓ') → CommRing ℓ' + CommAlgebra→CommRing (_ , commalgebrastr _ _ _ _ _ _ (iscommalgebra isAlgebra ·-comm)) = + _ , commringstr _ _ _ _ _ (iscommring (IsAlgebra.isRing isAlgebra) ·-comm) + + isSetCommAlgebra : (A : CommAlgebra R ℓ') → isSet ⟨ A ⟩ + isSetCommAlgebra A = isSetAlgebra (CommAlgebra→Algebra A) + + makeIsCommAlgebra : {A : Type ℓ'} {0a 1a : A} + {_+_ _·_ : A → A → A} { -_ : A → A} {_⋆_ : ⟨ R ⟩ → A → A} + (isSet-A : isSet A) + (+-assoc : (x y z : A) → x + (y + z) ≡ (x + y) + z) + (+-rid : (x : A) → x + 0a ≡ x) + (+-rinv : (x : A) → x + (- x) ≡ 0a) + (+-comm : (x y : A) → x + y ≡ y + x) + (·-assoc : (x y z : A) → x · (y · z) ≡ (x · y) · z) + (·-lid : (x : A) → 1a · x ≡ x) + (·-ldist-+ : (x y z : A) → (x + y) · z ≡ (x · z) + (y · z)) + (·-comm : (x y : A) → x · y ≡ y · x) + (⋆-assoc : (r s : ⟨ R ⟩) (x : A) → (r ·s s) ⋆ x ≡ r ⋆ (s ⋆ x)) + (⋆-ldist : (r s : ⟨ R ⟩) (x : A) → (r +r s) ⋆ x ≡ (r ⋆ x) + (s ⋆ x)) + (⋆-rdist : (r : ⟨ R ⟩) (x y : A) → r ⋆ (x + y) ≡ (r ⋆ x) + (r ⋆ y)) + (⋆-lid : (x : A) → 1r ⋆ x ≡ x) + (⋆-lassoc : (r : ⟨ R ⟩) (x y : A) → (r ⋆ x) · y ≡ r ⋆ (x · y)) + → IsCommAlgebra R 0a 1a _+_ _·_ -_ _⋆_ + makeIsCommAlgebra {A = A} {0a} {1a} {_+_} {_·_} { -_} {_⋆_} isSet-A + +-assoc +-rid +-rinv +-comm + ·-assoc ·-lid ·-ldist-+ ·-comm + ⋆-assoc ⋆-ldist ⋆-rdist ⋆-lid ⋆-lassoc + = iscommalgebra + (makeIsAlgebra + isSet-A + +-assoc +-rid +-rinv +-comm + ·-assoc + (λ x → x · 1a ≡⟨ ·-comm _ _ ⟩ 1a · x ≡⟨ ·-lid _ ⟩ x ∎) + ·-lid + (λ x y z → x · (y + z) ≡⟨ ·-comm _ _ ⟩ + (y + z) · x ≡⟨ ·-ldist-+ _ _ _ ⟩ + (y · x) + (z · x) ≡⟨ cong (λ u → (y · x) + u) (·-comm _ _) ⟩ + (y · x) + (x · z) ≡⟨ cong (λ u → u + (x · z)) (·-comm _ _) ⟩ + (x · y) + (x · z) ∎) + ·-ldist-+ + ⋆-assoc + ⋆-ldist + ⋆-rdist + ⋆-lid + ⋆-lassoc + λ r x y → r ⋆ (x · y) ≡⟨ cong (λ u → r ⋆ u) (·-comm _ _) ⟩ + r ⋆ (y · x) ≡⟨ sym (⋆-lassoc _ _ _) ⟩ + (r ⋆ y) · x ≡⟨ ·-comm _ _ ⟩ + x · (r ⋆ y) ∎) + ·-comm + + IsCommAlgebraEquiv : {A B : Type ℓ'} + (M : CommAlgebraStr R A) (e : A ≃ B) (N : CommAlgebraStr R B) + → Type (ℓ-max ℓ ℓ') + IsCommAlgebraEquiv M e N = + IsAlgebraHom (CommAlgebraStr→AlgebraStr M) (e .fst) (CommAlgebraStr→AlgebraStr N) + + CommAlgebraEquiv : (M N : CommAlgebra R ℓ') → Type (ℓ-max ℓ ℓ') + CommAlgebraEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsCommAlgebraEquiv (M .snd) e (N .snd) + + IsCommAlgebraHom : {A B : Type ℓ'} + (M : CommAlgebraStr R A) (f : A → B) (N : CommAlgebraStr R B) + → Type (ℓ-max ℓ ℓ') + IsCommAlgebraHom M f N = + IsAlgebraHom (CommAlgebraStr→AlgebraStr M) f (CommAlgebraStr→AlgebraStr N) + + CommAlgebraHom : (M N : CommAlgebra R ℓ') → Type (ℓ-max ℓ ℓ') + CommAlgebraHom M N = Σ[ f ∈ (⟨ M ⟩ → ⟨ N ⟩) ] IsCommAlgebraHom (M .snd) f (N .snd) + +isPropIsCommAlgebra : (R : CommRing ℓ) {A : Type ℓ'} + (0a 1a : A) + (_+_ _·_ : A → A → A) + (-_ : A → A) + (_⋆_ : ⟨ R ⟩ → A → A) + → isProp (IsCommAlgebra R 0a 1a _+_ _·_ -_ _⋆_) +isPropIsCommAlgebra R _ _ _ _ _ _ = + isOfHLevelRetractFromIso 1 IsCommAlgebraIsoΣ + (isPropΣ (isPropIsAlgebra _ _ _ _ _ _ _) + (λ alg → isPropΠ2 λ _ _ → alg .IsAlgebra.is-set _ _)) + +𝒮ᴰ-CommAlgebra : (R : CommRing ℓ) → DUARel (𝒮-Univ ℓ') (CommAlgebraStr R) (ℓ-max ℓ ℓ') +𝒮ᴰ-CommAlgebra R = + 𝒮ᴰ-Record (𝒮-Univ _) (IsCommAlgebraEquiv {R = R}) + (fields: + data[ 0a ∣ nul ∣ pres0 ] + data[ 1a ∣ nul ∣ pres1 ] + data[ _+_ ∣ bin ∣ pres+ ] + data[ _·_ ∣ bin ∣ pres· ] + data[ -_ ∣ autoDUARel _ _ ∣ pres- ] + data[ _⋆_ ∣ autoDUARel _ _ ∣ pres⋆ ] + prop[ isCommAlgebra ∣ (λ _ _ → isPropIsCommAlgebra _ _ _ _ _ _ _) ]) + where + open CommAlgebraStr + open IsAlgebraHom + + -- faster with some sharing + nul = autoDUARel (𝒮-Univ _) (λ A → A) + bin = autoDUARel (𝒮-Univ _) (λ A → A → A → A) + +CommAlgebraPath : (R : CommRing ℓ) → (A B : CommAlgebra R ℓ') → (CommAlgebraEquiv A B) ≃ (A ≡ B) +CommAlgebraPath R = ∫ (𝒮ᴰ-CommAlgebra R) .UARel.ua + +isGroupoidCommAlgebra : {R : CommRing ℓ} → isGroupoid (CommAlgebra R ℓ') +isGroupoidCommAlgebra A B = isOfHLevelRespectEquiv 2 (CommAlgebraPath _ _ _) (isSetAlgebraEquiv _ _) diff --git a/Cubical/Algebra/CommAlgebra/Localisation.agda b/Cubical/Algebra/CommAlgebra/Localisation.agda new file mode 100644 index 000000000..cdbd82691 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/Localisation.agda @@ -0,0 +1,217 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.CommAlgebra.Localisation where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Powerset + +open import Cubical.Data.Sigma + +open import Cubical.Reflection.StrictEquiv + +open import Cubical.Structures.Axioms +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.CommRing.Base +open import Cubical.Algebra.CommRing.Properties +open import Cubical.Algebra.CommRing.Localisation.Base +open import Cubical.Algebra.CommRing.Localisation.UniversalProperty +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Algebra +open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebra.Properties + +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.HITs.PropositionalTruncation as PT + + +private + variable + ℓ ℓ′ : Level + + + +module AlgLoc (R' : CommRing ℓ) + (S' : ℙ (fst R')) (SMultClosedSubset : isMultClosedSubset R' S') where + open isMultClosedSubset + private R = fst R' + open CommAlgebraStr + open IsAlgebraHom + open CommRingStr (snd R') renaming (_+_ to _+r_ ; _·_ to _·r_ ; ·Rid to ·rRid) + open RingTheory (CommRing→Ring R') + open CommRingTheory R' + open Loc R' S' SMultClosedSubset + open S⁻¹RUniversalProp R' S' SMultClosedSubset + open CommAlgChar R' + + + S⁻¹RAsCommAlg : CommAlgebra R' ℓ + S⁻¹RAsCommAlg = toCommAlg (S⁻¹RAsCommRing , /1AsCommRingHom) + + + hasLocAlgUniversalProp : (A : CommAlgebra R' ℓ) + → (∀ s → s ∈ S' → _⋆_ (snd A) s (1a (snd A)) ∈ (CommAlgebra→CommRing A) ˣ) + → Type (ℓ-suc ℓ) + hasLocAlgUniversalProp A _ = (B : CommAlgebra R' ℓ) + → (∀ s → s ∈ S' → _⋆_ (snd B) s (1a (snd B)) ∈ (CommAlgebra→CommRing B) ˣ) + → isContr (CommAlgebraHom A B) + + S⋆1⊆S⁻¹Rˣ : ∀ s → s ∈ S' → _⋆_ (snd S⁻¹RAsCommAlg) s (1a (snd S⁻¹RAsCommAlg)) ∈ S⁻¹Rˣ + S⋆1⊆S⁻¹Rˣ s s∈S' = subst-∈ S⁻¹Rˣ + (cong [_] (≡-× (sym (·rRid s)) (Σ≡Prop (λ x → S' x .snd) (sym (·rRid _))))) + (S/1⊆S⁻¹Rˣ s s∈S') + + + S⁻¹RHasAlgUniversalProp : hasLocAlgUniversalProp S⁻¹RAsCommAlg S⋆1⊆S⁻¹Rˣ + S⁻¹RHasAlgUniversalProp B' S⋆1⊆Bˣ = χᴬ , χᴬuniqueness + where + B = fromCommAlg B' .fst + φ = fromCommAlg B' .snd + open CommRingStr (snd B) renaming (_·_ to _·b_ ; 1r to 1b ; ·Lid to ·bLid) + + χ : CommRingHom S⁻¹RAsCommRing B + χ = S⁻¹RHasUniversalProp B φ S⋆1⊆Bˣ .fst .fst + + χcomp : ∀ r → fst χ (r /1) ≡ fst φ r + χcomp = funExt⁻ (S⁻¹RHasUniversalProp B φ S⋆1⊆Bˣ .fst .snd) + + χᴬ : CommAlgebraHom S⁻¹RAsCommAlg B' + fst χᴬ = fst χ + pres0 (snd χᴬ) = IsRingHom.pres0 (snd χ) + pres1 (snd χᴬ) = IsRingHom.pres1 (snd χ) + pres+ (snd χᴬ) = IsRingHom.pres+ (snd χ) + pres· (snd χᴬ) = IsRingHom.pres· (snd χ) + pres- (snd χᴬ) = IsRingHom.pres- (snd χ) + pres⋆ (snd χᴬ) r x = path + where + path : fst χ ((r /1) ·ₗ x) ≡ _⋆_ (snd B') r (fst χ x) + path = fst χ ((r /1) ·ₗ x) ≡⟨ IsRingHom.pres· (snd χ) _ _ ⟩ + fst χ (r /1) ·b fst χ x ≡⟨ cong (_·b fst χ x) (χcomp r) ⟩ + fst φ r ·b fst χ x ≡⟨ refl ⟩ + _⋆_ (snd B') r 1b ·b fst χ x ≡⟨ ⋆-lassoc (snd B') _ _ _ ⟩ + _⋆_ (snd B') r (1b ·b fst χ x) ≡⟨ cong (_⋆_ (snd B') r) (·bLid _) ⟩ + _⋆_ (snd B') r (fst χ x) ∎ + + + χᴬuniqueness : (ψ : CommAlgebraHom S⁻¹RAsCommAlg B') → χᴬ ≡ ψ + χᴬuniqueness ψ = Σ≡Prop (λ _ → isPropIsAlgebraHom _ _ _ _) + (cong (fst ∘ fst) (χuniqueness (ψ' , funExt ψ'r/1≡φr))) + where + χuniqueness = S⁻¹RHasUniversalProp B φ S⋆1⊆Bˣ .snd + + ψ' : CommRingHom S⁻¹RAsCommRing B + fst ψ' = fst ψ + IsRingHom.pres0 (snd ψ') = pres0 (snd ψ) + IsRingHom.pres1 (snd ψ') = pres1 (snd ψ) + IsRingHom.pres+ (snd ψ') = pres+ (snd ψ) + IsRingHom.pres· (snd ψ') = pres· (snd ψ) + IsRingHom.pres- (snd ψ') = pres- (snd ψ) + + ψ'r/1≡φr : ∀ r → fst ψ (r /1) ≡ fst φ r + ψ'r/1≡φr r = + fst ψ (r /1) ≡⟨ cong (fst ψ) (sym (·ₗ-rid _)) ⟩ + fst ψ (_⋆_ (snd S⁻¹RAsCommAlg) r (1a (snd S⁻¹RAsCommAlg))) ≡⟨ pres⋆ (snd ψ) _ _ ⟩ + _⋆_ (snd B') r (fst ψ (1a (snd S⁻¹RAsCommAlg))) ≡⟨ cong (_⋆_ (snd B') r) (pres1 (snd ψ)) ⟩ + _⋆_ (snd B') r 1b ∎ + + + -- an immediate corrollary: + isContrHomS⁻¹RS⁻¹R : isContr (CommAlgebraHom S⁻¹RAsCommAlg S⁻¹RAsCommAlg) + isContrHomS⁻¹RS⁻¹R = S⁻¹RHasAlgUniversalProp S⁻¹RAsCommAlg S⋆1⊆S⁻¹Rˣ + + + + +module AlgLocTwoSubsets (R' : CommRing ℓ) + (S₁ : ℙ (fst R')) (S₁MultClosedSubset : isMultClosedSubset R' S₁) + (S₂ : ℙ (fst R')) (S₂MultClosedSubset : isMultClosedSubset R' S₂) where + open isMultClosedSubset + open CommRingStr (snd R') hiding (is-set) + open RingTheory (CommRing→Ring R') + open Loc R' S₁ S₁MultClosedSubset renaming (S⁻¹R to S₁⁻¹R ; + S⁻¹RAsCommRing to S₁⁻¹RAsCommRing) + open Loc R' S₂ S₂MultClosedSubset renaming (S⁻¹R to S₂⁻¹R ; + S⁻¹RAsCommRing to S₂⁻¹RAsCommRing) + open AlgLoc R' S₁ S₁MultClosedSubset renaming ( S⁻¹RAsCommAlg to S₁⁻¹RAsCommAlg + ; S⋆1⊆S⁻¹Rˣ to S₁⋆1⊆S₁⁻¹Rˣ + ; S⁻¹RHasAlgUniversalProp to S₁⁻¹RHasAlgUniversalProp + ; isContrHomS⁻¹RS⁻¹R to isContrHomS₁⁻¹RS₁⁻¹R) + open AlgLoc R' S₂ S₂MultClosedSubset renaming ( S⁻¹RAsCommAlg to S₂⁻¹RAsCommAlg + ; S⋆1⊆S⁻¹Rˣ to S₂⋆1⊆S₂⁻¹Rˣ + ; S⁻¹RHasAlgUniversalProp to S₂⁻¹RHasAlgUniversalProp + ; isContrHomS⁻¹RS⁻¹R to isContrHomS₂⁻¹RS₂⁻¹R) + + open IsAlgebraHom + open CommAlgebraStr ⦃...⦄ + + private + R = fst R' + S₁⁻¹Rˣ = S₁⁻¹RAsCommRing ˣ + S₂⁻¹Rˣ = S₂⁻¹RAsCommRing ˣ + instance + _ = snd S₁⁻¹RAsCommAlg + _ = snd S₂⁻¹RAsCommAlg + + + isContrS₁⁻¹R≡S₂⁻¹R : (∀ s₁ → s₁ ∈ S₁ → s₁ ⋆ 1a ∈ S₂⁻¹Rˣ) + → (∀ s₂ → s₂ ∈ S₂ → s₂ ⋆ 1a ∈ S₁⁻¹Rˣ) + → isContr (S₁⁻¹RAsCommAlg ≡ S₂⁻¹RAsCommAlg) + isContrS₁⁻¹R≡S₂⁻¹R S₁⊆S₂⁻¹Rˣ S₂⊆S₁⁻¹Rˣ = isOfHLevelRetractFromIso 0 + (equivToIso (invEquiv (CommAlgebraPath _ _ _))) + isContrS₁⁻¹R≅S₂⁻¹R + where + χ₁ : CommAlgebraHom S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg + χ₁ = S₁⁻¹RHasAlgUniversalProp S₂⁻¹RAsCommAlg S₁⊆S₂⁻¹Rˣ .fst + + χ₂ : CommAlgebraHom S₂⁻¹RAsCommAlg S₁⁻¹RAsCommAlg + χ₂ = S₂⁻¹RHasAlgUniversalProp S₁⁻¹RAsCommAlg S₂⊆S₁⁻¹Rˣ .fst + + χ₁∘χ₂≡id : χ₁ ∘a χ₂ ≡ idAlgHom + χ₁∘χ₂≡id = isContr→isProp isContrHomS₂⁻¹RS₂⁻¹R _ _ + + χ₂∘χ₁≡id : χ₂ ∘a χ₁ ≡ idAlgHom + χ₂∘χ₁≡id = isContr→isProp isContrHomS₁⁻¹RS₁⁻¹R _ _ + + IsoS₁⁻¹RS₂⁻¹R : Iso S₁⁻¹R S₂⁻¹R + Iso.fun IsoS₁⁻¹RS₂⁻¹R = fst χ₁ + Iso.inv IsoS₁⁻¹RS₂⁻¹R = fst χ₂ + Iso.rightInv IsoS₁⁻¹RS₂⁻¹R = funExt⁻ (cong fst χ₁∘χ₂≡id) + Iso.leftInv IsoS₁⁻¹RS₂⁻¹R = funExt⁻ (cong fst χ₂∘χ₁≡id) + + isContrS₁⁻¹R≅S₂⁻¹R : isContr (CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg) + isContrS₁⁻¹R≅S₂⁻¹R = center , uniqueness + where + center : CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg + fst center = isoToEquiv IsoS₁⁻¹RS₂⁻¹R + pres0 (snd center) = pres0 (snd χ₁) + pres1 (snd center) = pres1 (snd χ₁) + pres+ (snd center) = pres+ (snd χ₁) + pres· (snd center) = pres· (snd χ₁) + pres- (snd center) = pres- (snd χ₁) + pres⋆ (snd center) = pres⋆ (snd χ₁) + + uniqueness : (φ : CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg) → center ≡ φ + uniqueness φ = Σ≡Prop (λ _ → isPropIsAlgebraHom _ _ _ _) + (equivEq (cong fst + (S₁⁻¹RHasAlgUniversalProp S₂⁻¹RAsCommAlg S₁⊆S₂⁻¹Rˣ .snd + (AlgebraEquiv→AlgebraHom φ)))) + + + isPropS₁⁻¹R≡S₂⁻¹R : isProp (S₁⁻¹RAsCommAlg ≡ S₂⁻¹RAsCommAlg) + isPropS₁⁻¹R≡S₂⁻¹R S₁⁻¹R≡S₂⁻¹R = + isContr→isProp (isContrS₁⁻¹R≡S₂⁻¹R S₁⊆S₂⁻¹Rˣ S₂⊆S₁⁻¹Rˣ) S₁⁻¹R≡S₂⁻¹R + where + S₁⊆S₂⁻¹Rˣ : ∀ s₁ → s₁ ∈ S₁ → s₁ ⋆ 1a ∈ S₂⁻¹Rˣ + S₁⊆S₂⁻¹Rˣ s₁ s₁∈S₁ = + transport (λ i → _⋆_ ⦃ S₁⁻¹R≡S₂⁻¹R i .snd ⦄ s₁ (1a ⦃ S₁⁻¹R≡S₂⁻¹R i .snd ⦄) + ∈ (CommAlgebra→CommRing (S₁⁻¹R≡S₂⁻¹R i)) ˣ) (S₁⋆1⊆S₁⁻¹Rˣ s₁ s₁∈S₁) + + S₂⊆S₁⁻¹Rˣ : ∀ s₂ → s₂ ∈ S₂ → s₂ ⋆ 1a ∈ S₁⁻¹Rˣ + S₂⊆S₁⁻¹Rˣ s₂ s₂∈S₂ = + transport (λ i → _⋆_ ⦃ (sym S₁⁻¹R≡S₂⁻¹R) i .snd ⦄ s₂ (1a ⦃ (sym S₁⁻¹R≡S₂⁻¹R) i .snd ⦄) + ∈ (CommAlgebra→CommRing ((sym S₁⁻¹R≡S₂⁻¹R) i)) ˣ) (S₂⋆1⊆S₂⁻¹Rˣ s₂ s₂∈S₂) diff --git a/Cubical/Algebra/CommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebra/Properties.agda new file mode 100644 index 000000000..9641f3406 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/Properties.agda @@ -0,0 +1,113 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.CommAlgebra.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Reflection.StrictEquiv + +open import Cubical.Structures.Axioms +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Algebra +open import Cubical.Algebra.CommAlgebra.Base + +private + variable + ℓ ℓ′ : Level + + +-- An R-algebra is the same as a CommRing A with a CommRingHom φ : R → A +module CommAlgChar (R : CommRing ℓ) where + open Iso + open IsRingHom + open CommRingTheory + + + CommRingWithHom : Type (ℓ-suc ℓ) + CommRingWithHom = Σ[ A ∈ CommRing ℓ ] CommRingHom R A + + toCommAlg : CommRingWithHom → CommAlgebra R ℓ + toCommAlg (A , φ , φIsHom) = ⟨ A ⟩ , ACommAlgStr + where + open CommRingStr (snd A) + ACommAlgStr : CommAlgebraStr R ⟨ A ⟩ + CommAlgebraStr.0a ACommAlgStr = 0r + CommAlgebraStr.1a ACommAlgStr = 1r + CommAlgebraStr._+_ ACommAlgStr = _+_ + CommAlgebraStr._·_ ACommAlgStr = _·_ + CommAlgebraStr.- ACommAlgStr = -_ + CommAlgebraStr._⋆_ ACommAlgStr r a = (φ r) · a + CommAlgebraStr.isCommAlgebra ACommAlgStr = makeIsCommAlgebra + is-set +Assoc +Rid +Rinv +Comm ·Assoc ·Lid ·Ldist+ ·-comm + (λ _ _ x → cong (λ y → y · x) (pres· φIsHom _ _) ∙ sym (·Assoc _ _ _)) + (λ _ _ x → cong (λ y → y · x) (pres+ φIsHom _ _) ∙ ·Ldist+ _ _ _) + (λ _ _ _ → ·Rdist+ _ _ _) + (λ x → cong (λ y → y · x) (pres1 φIsHom) ∙ ·Lid x) + (λ _ _ _ → sym (·Assoc _ _ _)) + + + fromCommAlg : CommAlgebra R ℓ → CommRingWithHom + fromCommAlg A = (CommAlgebra→CommRing A) , φ , φIsHom + where + open CommRingStr (snd R) renaming (_·_ to _·r_) hiding (·Lid) + open CommAlgebraStr (snd A) + open AlgebraTheory (CommRing→Ring R) (CommAlgebra→Algebra A) + φ : ⟨ R ⟩ → ⟨ A ⟩ + φ r = r ⋆ 1a + φIsHom : IsRingHom (CommRing→Ring R .snd) φ (CommRing→Ring (CommAlgebra→CommRing A) .snd) + φIsHom = makeIsRingHom (⋆-lid _) (λ _ _ → ⋆-ldist _ _ _) + λ x y → cong (λ a → (x ·r y) ⋆ a) (sym (·Lid _)) ∙ ⋆Dist· _ _ _ _ + + + CommRingWithHomRoundTrip : (Aφ : CommRingWithHom) → fromCommAlg (toCommAlg Aφ) ≡ Aφ + CommRingWithHomRoundTrip (A , φ) = ΣPathP (APath , φPathP) + where + open CommRingStr + -- note that the proofs of the axioms might differ! + APath : fst (fromCommAlg (toCommAlg (A , φ))) ≡ A + fst (APath i) = ⟨ A ⟩ + 0r (snd (APath i)) = 0r (snd A) + 1r (snd (APath i)) = 1r (snd A) + _+_ (snd (APath i)) = _+_ (snd A) + _·_ (snd (APath i)) = _·_ (snd A) + -_ (snd (APath i)) = -_ (snd A) + isCommRing (snd (APath i)) = isProp→PathP (λ i → isPropIsCommRing _ _ _ _ _ ) + (isCommRing (snd (fst (fromCommAlg (toCommAlg (A , φ)))))) (isCommRing (snd A)) i + + -- this only works because fst (APath i) = fst A definitionally! + φPathP : PathP (λ i → CommRingHom R (APath i)) (snd (fromCommAlg (toCommAlg (A , φ)))) φ + φPathP = RingHomEqDep _ _ _ _ _ _ λ i x → ·Rid (snd A) (fst φ x) i + + + CommAlgRoundTrip : (A : CommAlgebra R ℓ) → toCommAlg (fromCommAlg A) ≡ A + CommAlgRoundTrip A = ΣPathP (refl , AlgStrPathP) + where + open CommAlgebraStr ⦃...⦄ + instance + _ = snd A + AlgStrPathP : PathP (λ i → CommAlgebraStr R ⟨ A ⟩) (snd (toCommAlg (fromCommAlg A))) (snd A) + CommAlgebraStr.0a (AlgStrPathP i) = 0a + CommAlgebraStr.1a (AlgStrPathP i) = 1a + CommAlgebraStr._+_ (AlgStrPathP i) = _+_ + CommAlgebraStr._·_ (AlgStrPathP i) = _·_ + CommAlgebraStr.-_ (AlgStrPathP i) = -_ + CommAlgebraStr._⋆_ (AlgStrPathP i) r x = (⋆-lassoc r 1a x ∙ cong (r ⋆_) (·Lid x)) i + CommAlgebraStr.isCommAlgebra (AlgStrPathP i) = isProp→PathP + (λ i → isPropIsCommAlgebra _ _ _ _ _ _ (CommAlgebraStr._⋆_ (AlgStrPathP i))) + (CommAlgebraStr.isCommAlgebra (snd (toCommAlg (fromCommAlg A)))) isCommAlgebra i + + + CommAlgIso : Iso (CommAlgebra R ℓ) CommRingWithHom + fun CommAlgIso = fromCommAlg + inv CommAlgIso = toCommAlg + rightInv CommAlgIso = CommRingWithHomRoundTrip + leftInv CommAlgIso = CommAlgRoundTrip diff --git a/Cubical/Algebra/CommRing.agda b/Cubical/Algebra/CommRing.agda new file mode 100644 index 000000000..0f159f20a --- /dev/null +++ b/Cubical/Algebra/CommRing.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommRing where + +open import Cubical.Algebra.CommRing.Base public +open import Cubical.Algebra.CommRing.Properties public diff --git a/Cubical/Algebra/CommRing/Base.agda b/Cubical/Algebra/CommRing/Base.agda new file mode 100644 index 000000000..311875263 --- /dev/null +++ b/Cubical/Algebra/CommRing/Base.agda @@ -0,0 +1,144 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommRing.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Ring.Base + +open Iso + +private + variable + ℓ ℓ' : Level + +record IsCommRing {R : Type ℓ} + (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) : Type ℓ where + + constructor iscommring + + field + isRing : IsRing 0r 1r _+_ _·_ -_ + ·-comm : (x y : R) → x · y ≡ y · x + + open IsRing isRing public + +record CommRingStr (A : Type ℓ) : Type (ℓ-suc ℓ) where + + constructor commringstr + + field + 0r : A + 1r : A + _+_ : A → A → A + _·_ : A → A → A + -_ : A → A + isCommRing : IsCommRing 0r 1r _+_ _·_ -_ + + infix 8 -_ + infixl 7 _·_ + infixl 6 _+_ + + open IsCommRing isCommRing public + +CommRing : ∀ ℓ → Type (ℓ-suc ℓ) +CommRing ℓ = TypeWithStr ℓ CommRingStr + + +makeIsCommRing : {R : Type ℓ} {0r 1r : R} {_+_ _·_ : R → R → R} { -_ : R → R} + (is-setR : isSet R) + (+-assoc : (x y z : R) → x + (y + z) ≡ (x + y) + z) + (+-rid : (x : R) → x + 0r ≡ x) + (+-rinv : (x : R) → x + (- x) ≡ 0r) + (+-comm : (x y : R) → x + y ≡ y + x) + (·-assoc : (x y z : R) → x · (y · z) ≡ (x · y) · z) + (·-rid : (x : R) → x · 1r ≡ x) + (·-rdist-+ : (x y z : R) → x · (y + z) ≡ (x · y) + (x · z)) + (·-comm : (x y : R) → x · y ≡ y · x) + → IsCommRing 0r 1r _+_ _·_ -_ +makeIsCommRing {_+_ = _+_} is-setR +-assoc +-rid +-rinv +-comm ·-assoc ·-rid ·-rdist-+ ·-comm = + iscommring (makeIsRing is-setR +-assoc +-rid +-rinv +-comm ·-assoc ·-rid + (λ x → ·-comm _ _ ∙ ·-rid x) ·-rdist-+ + (λ x y z → ·-comm _ _ ∙∙ ·-rdist-+ z x y ∙∙ λ i → (·-comm z x i) + (·-comm z y i))) ·-comm + +makeCommRing : {R : Type ℓ} (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) + (is-setR : isSet R) + (+-assoc : (x y z : R) → x + (y + z) ≡ (x + y) + z) + (+-rid : (x : R) → x + 0r ≡ x) + (+-rinv : (x : R) → x + (- x) ≡ 0r) + (+-comm : (x y : R) → x + y ≡ y + x) + (·-assoc : (x y z : R) → x · (y · z) ≡ (x · y) · z) + (·-rid : (x : R) → x · 1r ≡ x) + (·-rdist-+ : (x y z : R) → x · (y + z) ≡ (x · y) + (x · z)) + (·-comm : (x y : R) → x · y ≡ y · x) + → CommRing ℓ +makeCommRing 0r 1r _+_ _·_ -_ is-setR +-assoc +-rid +-rinv +-comm ·-assoc ·-rid ·-rdist-+ ·-comm = + _ , commringstr _ _ _ _ _ (makeIsCommRing is-setR +-assoc +-rid +-rinv +-comm ·-assoc ·-rid ·-rdist-+ ·-comm) + +CommRingStr→RingStr : {A : Type ℓ} → CommRingStr A → RingStr A +CommRingStr→RingStr (commringstr _ _ _ _ _ H) = ringstr _ _ _ _ _ (IsCommRing.isRing H) + +CommRing→Ring : CommRing ℓ → Ring ℓ +CommRing→Ring (_ , commringstr _ _ _ _ _ H) = _ , ringstr _ _ _ _ _ (IsCommRing.isRing H) + +CommRingHom : (R : CommRing ℓ) (S : CommRing ℓ') → Type (ℓ-max ℓ ℓ') +CommRingHom R S = RingHom (CommRing→Ring R) (CommRing→Ring S) + +IsCommRingEquiv : {A : Type ℓ} {B : Type ℓ'} + (R : CommRingStr A) (e : A ≃ B) (S : CommRingStr B) → Type (ℓ-max ℓ ℓ') +IsCommRingEquiv R e S = IsRingHom (CommRingStr→RingStr R) (e .fst) (CommRingStr→RingStr S) + +CommRingEquiv : (R : CommRing ℓ) (S : CommRing ℓ') → Type (ℓ-max ℓ ℓ') +CommRingEquiv R S = Σ[ e ∈ (R .fst ≃ S .fst) ] IsCommRingEquiv (R .snd) e (S .snd) + +isPropIsCommRing : {R : Type ℓ} (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) + → isProp (IsCommRing 0r 1r _+_ _·_ -_) +isPropIsCommRing 0r 1r _+_ _·_ -_ (iscommring RR RC) (iscommring SR SC) = + λ i → iscommring (isPropIsRing _ _ _ _ _ RR SR i) + (isPropComm RC SC i) + where + isSetR : isSet _ + isSetR = RR .IsRing.·IsMonoid .IsMonoid.isSemigroup .IsSemigroup.is-set + + isPropComm : isProp ((x y : _) → x · y ≡ y · x) + isPropComm = isPropΠ2 λ _ _ → isSetR _ _ + +𝒮ᴰ-CommRing : DUARel (𝒮-Univ ℓ) CommRingStr ℓ +𝒮ᴰ-CommRing = + 𝒮ᴰ-Record (𝒮-Univ _) IsCommRingEquiv + (fields: + data[ 0r ∣ null ∣ pres0 ] + data[ 1r ∣ null ∣ pres1 ] + data[ _+_ ∣ bin ∣ pres+ ] + data[ _·_ ∣ bin ∣ pres· ] + data[ -_ ∣ autoDUARel _ _ ∣ pres- ] + prop[ isCommRing ∣ (λ _ _ → isPropIsCommRing _ _ _ _ _) ]) + where + open CommRingStr + open IsRingHom + + -- faster with some sharing + null = autoDUARel (𝒮-Univ _) (λ A → A) + bin = autoDUARel (𝒮-Univ _) (λ A → A → A → A) + +CommRingPath : (R S : CommRing ℓ) → CommRingEquiv R S ≃ (R ≡ S) +CommRingPath = ∫ 𝒮ᴰ-CommRing .UARel.ua + +isSetCommRing : ((R , str) : CommRing ℓ) → isSet R +isSetCommRing (R , str) = str .CommRingStr.is-set diff --git a/Cubical/Algebra/CommRing/BinomialThm.agda b/Cubical/Algebra/CommRing/BinomialThm.agda new file mode 100644 index 000000000..b0bc88371 --- /dev/null +++ b/Cubical/Algebra/CommRing/BinomialThm.agda @@ -0,0 +1,132 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.CommRing.BinomialThm where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function + +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ + ; +-comm to +ℕ-comm + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm + ; _choose_ to _ℕchoose_ ; snotz to ℕsnotz) +open import Cubical.Data.Nat.Order +open import Cubical.Data.FinData +open import Cubical.Data.Empty as ⊥ + +open import Cubical.Algebra.Monoid.BigOp +open import Cubical.Algebra.Ring.BigOps +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.RingSolver.ReflectionSolving + +private + variable + ℓ : Level + +module BinomialThm (R' : CommRing ℓ) where + open CommRingStr (snd R') + open Exponentiation R' + open CommRingTheory R' + open Sum (CommRing→Ring R') + private R = fst R' + + _choose_ : ℕ → ℕ → R + n choose zero = 1r + zero choose suc k = 0r + suc n choose suc k = n choose (suc k) + n choose k + + n n≡m + + powersFormMultClosedSubset : (f : R) → isMultClosedSubset R' [ f ⁿ|n≥0] + powersFormMultClosedSubset f .containsOne = PT.∣ zero , refl ∣ + powersFormMultClosedSubset f .multClosed = + PT.map2 λ (m , p) (n , q) → (m +ℕ n) , (λ i → (p i) · (q i)) ∙ ·-of-^-is-^-of-+ f m n + + + R[1/_] : R → Type ℓ + R[1/ f ] = Loc.S⁻¹R R' [ f ⁿ|n≥0] (powersFormMultClosedSubset f) + + + R[1/_]AsCommRing : R → CommRing ℓ + R[1/ f ]AsCommRing = Loc.S⁻¹RAsCommRing R' [ f ⁿ|n≥0] (powersFormMultClosedSubset f) + + -- A useful lemma: (gⁿ/1)≡(g/1)ⁿ in R[1/f] + ^-respects-/1 : {f g : R} (n : ℕ) → [ (g ^ n) , 1r , PT.∣ 0 , (λ _ → 1r) ∣ ] ≡ + Exponentiation._^_ R[1/ f ]AsCommRing [ g , 1r , powersFormMultClosedSubset _ .containsOne ] n + ^-respects-/1 zero = refl + ^-respects-/1 {f} {g} (suc n) = eq/ _ _ ( (1r , powersFormMultClosedSubset f .containsOne) + , cong (1r · (g · (g ^ n)) ·_) (·Lid 1r)) + ∙ cong (CommRingStr._·_ (R[1/ f ]AsCommRing .snd) + [ g , 1r , powersFormMultClosedSubset f .containsOne ]) (^-respects-/1 n) + + -- A slight improvement for eliminating into propositions + InvElPropElim : {f : R} {P : R[1/ f ] → Type ℓ'} + → (∀ x → isProp (P x)) + → (∀ (r : R) (n : ℕ) → P [ r , (f ^ n) , PT.∣ n , refl ∣ ]) -- ∀ r n → P (r/fⁿ) + ---------------------------------------------------------- + → (∀ x → P x) + InvElPropElim {f = f} {P = P} PisProp base = elimProp (λ _ → PisProp _) []-case + where + S[f] = Loc.S R' [ f ⁿ|n≥0] (powersFormMultClosedSubset f) + []-case : (a : R × S[f]) → P [ a ] + []-case (r , s , s∈S[f]) = PT.rec (PisProp _) Σhelper s∈S[f] + where + Σhelper : Σ[ n ∈ ℕ ] s ≡ f ^ n → P [ r , s , s∈S[f] ] + Σhelper (n , p) = subst P (cong [_] (≡-× refl (Σ≡Prop (λ _ → isPropPropTrunc) (sym p)))) (base r n) + + -- For predicates over the set of powers + powersPropElim : {f : R} {P : R → Type ℓ'} + → (∀ x → isProp (P x)) + → (∀ n → P (f ^ n)) + ------------------------------ + → ∀ s → s ∈ [ f ⁿ|n≥0] → P s + powersPropElim {f = f} {P = P} PisProp base s = + PT.rec (PisProp s) λ (n , p) → subst P (sym p) (base n) + + + +module DoubleLoc (R' : CommRing ℓ) (f g : (fst R')) where + open isMultClosedSubset + open CommRingStr (snd R') + open CommRingTheory R' + open Exponentiation R' + open RingTheory (CommRing→Ring R') + open CommRingStr (snd (R[1/_]AsCommRing R' f)) renaming ( _·_ to _·ᶠ_ ; 1r to 1ᶠ + ; _+_ to _+ᶠ_ ; 0r to 0ᶠ + ; ·Lid to ·ᶠ-lid ; ·Rid to ·ᶠ-rid + ; ·Assoc to ·ᶠ-assoc ; ·-comm to ·ᶠ-comm) + open IsRingHom + + private + R = fst R' + R[1/f] = R[1/_] R' f + R[1/f]AsCommRing = R[1/_]AsCommRing R' f + R[1/fg] = R[1/_] R' (f · g) + R[1/fg]AsCommRing = R[1/_]AsCommRing R' (f · g) + R[1/f][1/g] = R[1/_] (R[1/_]AsCommRing R' f) + [ g , 1r , powersFormMultClosedSubset R' f .containsOne ] + R[1/f][1/g]AsCommRing = R[1/_]AsCommRing (R[1/_]AsCommRing R' f) + [ g , 1r , powersFormMultClosedSubset R' f .containsOne ] + R[1/f][1/g]ˣ = R[1/f][1/g]AsCommRing ˣ + + + _/1/1 : R → R[1/f][1/g] + r /1/1 = [ [ r , 1r , PT.∣ 0 , refl ∣ ] , 1ᶠ , PT.∣ 0 , refl ∣ ] + + /1/1AsCommRingHom : CommRingHom R' R[1/f][1/g]AsCommRing + fst /1/1AsCommRingHom = _/1/1 + snd /1/1AsCommRingHom = makeIsRingHom refl lem+ lem· + where + lem+ : _ + lem+ r r' = + cong [_] + (≡-× + (cong [_] + (≡-× + (cong₂ _+_ + (sym (·Rid _) ∙ (λ i → (·Rid r (~ i)) · (·Rid 1r (~ i)))) + (sym (·Rid _) ∙ (λ i → (·Rid r' (~ i)) · (·Rid 1r (~ i))))) + (Σ≡Prop (λ _ → isPropPropTrunc) + (sym (·Lid _) ∙ (λ i → (·Lid 1r (~ i)) · (·Lid 1r (~ i))))))) + (Σ≡Prop (λ _ → isPropPropTrunc) (sym (·ᶠ-lid 1ᶠ)))) + + lem· : _ + lem· r r' = + cong [_] + (≡-× + (cong [_] (≡-× refl (Σ≡Prop (λ _ → isPropPropTrunc) (sym (·Lid _))))) + (Σ≡Prop (λ _ → isPropPropTrunc) (sym (·ᶠ-lid 1ᶠ)))) + + -- this will give us a map R[1/fg] → R[1/f][1/g] by the universal property of localisation + fⁿgⁿ/1/1∈R[1/f][1/g]ˣ : (s : R) → s ∈ ([_ⁿ|n≥0] R' (f · g)) → s /1/1 ∈ R[1/f][1/g]ˣ + fⁿgⁿ/1/1∈R[1/f][1/g]ˣ = powersPropElim R' (λ s → R[1/f][1/g]ˣ (s /1/1) .snd) ℕcase + where + ℕcase : (n : ℕ) → ((f · g) ^ n) /1/1 ∈ R[1/f][1/g]ˣ + ℕcase n = [ [ 1r , (f ^ n) , PT.∣ n , refl ∣ ] + , [ (g ^ n) , 1r , PT.∣ 0 , refl ∣ ] --denominator + , PT.∣ n , ^-respects-/1 _ n ∣ ] + , eq/ _ _ ((1ᶠ , powersFormMultClosedSubset _ _ .containsOne) + , eq/ _ _ ((1r , powersFormMultClosedSubset _ _ .containsOne) , path)) + where + eq1 : ∀ x → 1r · (1r · (x · 1r) · 1r) · (1r · 1r · (1r · 1r)) ≡ x + eq1 = solve R' + + eq2 : ∀ x y → x · y ≡ 1r · (1r · 1r · (1r · y)) · (1r · (1r · x) · 1r) + eq2 = solve R' + + path : 1r · (1r · ((f · g) ^ n · 1r) · 1r) · (1r · 1r · (1r · 1r)) + ≡ 1r · (1r · 1r · (1r · g ^ n)) · (1r · (1r · f ^ n) · 1r) + path = 1r · (1r · ((f · g) ^ n · 1r) · 1r) · (1r · 1r · (1r · 1r)) ≡⟨ eq1 ((f · g) ^ n) ⟩ + (f · g) ^ n ≡⟨ ^-ldist-· _ _ _ ⟩ + f ^ n · g ^ n ≡⟨ eq2 (f ^ n) (g ^ n) ⟩ + 1r · (1r · 1r · (1r · g ^ n)) · (1r · (1r · f ^ n) · 1r) ∎ + + + -- the main result: localising at one element and then at another is + -- the same as localising at the product. + -- takes forever to compute without experimental lossy unification + R[1/fg]≡R[1/f][1/g] : R[1/fg]AsCommRing ≡ R[1/f][1/g]AsCommRing + R[1/fg]≡R[1/f][1/g] = S⁻¹RChar R' ([_ⁿ|n≥0] R' (f · g)) + (powersFormMultClosedSubset R' (f · g)) _ /1/1AsCommRingHom pathtoR[1/fg] + where + open PathToS⁻¹R + pathtoR[1/fg] : PathToS⁻¹R R' ([_ⁿ|n≥0] R' (f · g)) (powersFormMultClosedSubset R' (f · g)) + R[1/f][1/g]AsCommRing /1/1AsCommRingHom + φS⊆Aˣ pathtoR[1/fg] = fⁿgⁿ/1/1∈R[1/f][1/g]ˣ + + kerφ⊆annS pathtoR[1/fg] r p = toGoal helperR[1/f] + where + open RingTheory (CommRing→Ring R[1/f]AsCommRing) renaming ( 0RightAnnihilates to 0ᶠRightAnnihilates + ; 0LeftAnnihilates to 0ᶠ-leftNullifies) + open Exponentiation R[1/f]AsCommRing renaming (_^_ to _^ᶠ_) + hiding (·-of-^-is-^-of-+ ; ^-ldist-·) + + S[f] = Loc.S R' ([_ⁿ|n≥0] R' f) (powersFormMultClosedSubset R' f) + S[fg] = Loc.S R' ([_ⁿ|n≥0] R' (f · g)) (powersFormMultClosedSubset R' (f · g)) + g/1 : R[1/_] R' f + g/1 = [ g , 1r , powersFormMultClosedSubset R' f .containsOne ] + S[g/1] = Loc.S R[1/f]AsCommRing + ([_ⁿ|n≥0] R[1/f]AsCommRing g/1) + (powersFormMultClosedSubset R[1/f]AsCommRing g/1) + r/1 : R[1/_] R' f + r/1 = [ r , 1r , powersFormMultClosedSubset R' f .containsOne ] + + -- this is the crucial step, modulo truncation we can take p to be generated + -- by the quotienting relation of localisation. Note that we wouldn't be able + -- to prove our goal if kerφ⊆annS was formulated with a Σ instead of a ∃ + ∥r/1,1/1≈0/1,1/1∥ : ∃[ u ∈ S[g/1] ] fst u ·ᶠ r/1 ·ᶠ 1ᶠ ≡ fst u ·ᶠ 0ᶠ ·ᶠ 1ᶠ + ∥r/1,1/1≈0/1,1/1∥ = Iso.fun (SQ.isEquivRel→TruncIso (Loc.locIsEquivRel _ _ _) _ _) p + + helperR[1/f] : ∃[ n ∈ ℕ ] [ g ^ n · r , 1r , PT.∣ 0 , refl ∣ ] ≡ 0ᶠ + helperR[1/f] = PT.rec isPropPropTrunc + (uncurry (uncurry (powersPropElim R[1/f]AsCommRing + (λ _ → isPropΠ (λ _ → isPropPropTrunc)) baseCase))) + ∥r/1,1/1≈0/1,1/1∥ + where + baseCase : ∀ n → g/1 ^ᶠ n ·ᶠ r/1 ·ᶠ 1ᶠ ≡ g/1 ^ᶠ n ·ᶠ 0ᶠ ·ᶠ 1ᶠ + → ∃[ n ∈ ℕ ] [ g ^ n · r , 1r , PT.∣ 0 , refl ∣ ] ≡ 0ᶠ + baseCase n q = PT.∣ n , path ∣ + where + path : [ g ^ n · r , 1r , PT.∣ 0 , refl ∣ ] ≡ 0ᶠ + path = [ g ^ n · r , 1r , PT.∣ 0 , refl ∣ ] + + ≡⟨ cong [_] (≡-× refl (Σ≡Prop (λ _ → isPropPropTrunc) (sym (·Rid _)))) ⟩ + + [ g ^ n , 1r , PT.∣ 0 , refl ∣ ] ·ᶠ r/1 + + ≡⟨ cong (_·ᶠ r/1) (^-respects-/1 _ n) ⟩ + + g/1 ^ᶠ n ·ᶠ r/1 + + ≡⟨ sym (·ᶠ-rid _) ⟩ + + g/1 ^ᶠ n ·ᶠ r/1 ·ᶠ 1ᶠ + + ≡⟨ q ⟩ + + g/1 ^ᶠ n ·ᶠ 0ᶠ ·ᶠ 1ᶠ + + ≡⟨ cong (_·ᶠ 1ᶠ) (0ᶠRightAnnihilates _) ∙ 0ᶠ-leftNullifies 1ᶠ ⟩ + + 0ᶠ ∎ + + + toGoal : ∃[ n ∈ ℕ ] [ g ^ n · r , 1r , PT.∣ 0 , refl ∣ ] ≡ 0ᶠ + → ∃[ u ∈ S[fg] ] fst u · r ≡ 0r + toGoal = PT.rec isPropPropTrunc Σhelper + where + Σhelper : Σ[ n ∈ ℕ ] [ g ^ n · r , 1r , PT.∣ 0 , refl ∣ ] ≡ 0ᶠ + → ∃[ u ∈ S[fg] ] fst u · r ≡ 0r + Σhelper (n , q) = PT.map Σhelper2 helperR + where + -- now, repeat the above strategy with q + ∥gⁿr≈0∥ : ∃[ u ∈ S[f] ] fst u · (g ^ n · r) · 1r ≡ fst u · 0r · 1r + ∥gⁿr≈0∥ = Iso.fun (SQ.isEquivRel→TruncIso (Loc.locIsEquivRel _ _ _) _ _) q + + helperR : ∃[ m ∈ ℕ ] f ^ m · g ^ n · r ≡ 0r + helperR = PT.rec isPropPropTrunc + (uncurry (uncurry (powersPropElim R' + (λ _ → isPropΠ (λ _ → isPropPropTrunc)) baseCase))) + ∥gⁿr≈0∥ + where + baseCase : (m : ℕ) → f ^ m · (g ^ n · r) · 1r ≡ f ^ m · 0r · 1r + → ∃[ m ∈ ℕ ] f ^ m · g ^ n · r ≡ 0r + baseCase m q' = PT.∣ m , path ∣ + where + path : f ^ m · g ^ n · r ≡ 0r + path = (λ i → ·Rid (·Assoc (f ^ m) (g ^ n) r (~ i)) (~ i)) + ∙∙ q' ∙∙ (λ i → ·Rid (0RightAnnihilates (f ^ m) i) i) + + Σhelper2 : Σ[ m ∈ ℕ ] f ^ m · g ^ n · r ≡ 0r + → Σ[ u ∈ S[fg] ] fst u · r ≡ 0r + Σhelper2 (m , q') = (((f · g) ^ l) , PT.∣ l , refl ∣) , path + where + l = max m n + + path : (f · g) ^ l · r ≡ 0r + path = (f · g) ^ l · r + + ≡⟨ cong (_· r) (^-ldist-· _ _ _) ⟩ + + f ^ l · g ^ l · r + + ≡⟨ cong₂ (λ x y → f ^ x · g ^ y · r) (sym (≤-∸-+-cancel {m = m} left-≤-max)) + (sym (≤-∸-+-cancel {m = n} right-≤-max)) ⟩ + + f ^ (l ∸ m +ℕ m) · g ^ (l ∸ n +ℕ n) · r + + ≡⟨ cong₂ (λ x y → x · y · r) (sym (·-of-^-is-^-of-+ _ _ _)) + (sym (·-of-^-is-^-of-+ _ _ _)) ⟩ + + f ^ (l ∸ m) · f ^ m · (g ^ (l ∸ n) · g ^ n) · r + + ≡⟨ cong (_· r) (·-commAssocSwap _ _ _ _) ⟩ + + f ^ (l ∸ m) · g ^ (l ∸ n) · (f ^ m · g ^ n) · r + + ≡⟨ sym (·Assoc _ _ _) ⟩ + + f ^ (l ∸ m) · g ^ (l ∸ n) · (f ^ m · g ^ n · r) + + ≡⟨ cong (f ^ (l ∸ m) · g ^ (l ∸ n) ·_) q' ⟩ + + f ^ (l ∸ m) · g ^ (l ∸ n) · 0r + + ≡⟨ 0RightAnnihilates _ ⟩ + + 0r ∎ + + + surχ pathtoR[1/fg] = InvElPropElim _ (λ _ → isPropPropTrunc) toGoal + where + open Exponentiation R[1/f]AsCommRing renaming (_^_ to _^ᶠ_) + hiding (·-of-^-is-^-of-+ ; ^-ldist-·) + open CommRingStr (snd R[1/f][1/g]AsCommRing) renaming (_·_ to _·R[1/f][1/g]_) + hiding (1r ; ·Lid ; ·Rid ; ·Assoc) + open Units R[1/f][1/g]AsCommRing + g/1 : R[1/_] R' f + g/1 = [ g , 1r , powersFormMultClosedSubset R' f .containsOne ] + S[fg] = Loc.S R' ([_ⁿ|n≥0] R' (f · g)) (powersFormMultClosedSubset R' (f · g)) + + baseCase : (r : R) (m n : ℕ) → ∃[ x ∈ R × S[fg] ] (x .fst /1/1) + ≡ [ [ r , f ^ m , PT.∣ m , refl ∣ ] + , [ g ^ n , 1r , PT.∣ 0 , refl ∣ ] , PT.∣ n , ^-respects-/1 _ n ∣ ] + ·R[1/f][1/g] (x .snd .fst /1/1) + baseCase r m n = PT.∣ ((r · f ^ (l ∸ m) · g ^ (l ∸ n)) -- x .fst + , (f · g) ^ l , PT.∣ l , refl ∣) -- x .snd + , eq/ _ _ ((1ᶠ , PT.∣ 0 , refl ∣) , eq/ _ _ ((1r , PT.∣ 0 , refl ∣) , path)) ∣ + -- reduce equality of double fractions into equality in R + where + eq1 : ∀ r flm gln gn fm + → 1r · (1r · (r · flm · gln) · (gn · 1r)) · (1r · (fm · 1r) · 1r) + ≡ r · flm · (gln · gn) · fm + eq1 = solve R' + + eq2 : ∀ r flm gl fm → r · flm · gl · fm ≡ r · (flm · fm ) · gl + eq2 = solve R' + + eq3 : ∀ r fgl → r · fgl ≡ 1r · (1r · (r · fgl) · 1r) · (1r · 1r · (1r · 1r)) + eq3 = solve R' + + l = max m n + + path : 1r · (1r · (r · f ^ (l ∸ m) · g ^ (l ∸ n)) · (g ^ n · 1r)) · (1r · (f ^ m · 1r) · 1r) + ≡ 1r · (1r · (r · (f · g) ^ l) · 1r) · (1r · 1r · (1r · 1r)) + path = 1r · (1r · (r · f ^ (l ∸ m) · g ^ (l ∸ n)) · (g ^ n · 1r)) · (1r · (f ^ m · 1r) · 1r) + + ≡⟨ eq1 r (f ^ (l ∸ m)) (g ^ (l ∸ n)) (g ^ n) (f ^ m) ⟩ + + r · f ^ (l ∸ m) · (g ^ (l ∸ n) · g ^ n) · f ^ m + + ≡⟨ cong (λ x → r · f ^ (l ∸ m) · x · f ^ m) (·-of-^-is-^-of-+ _ _ _) ⟩ + + r · f ^ (l ∸ m) · g ^ (l ∸ n +ℕ n) · f ^ m + + ≡⟨ cong (λ x → r · f ^ (l ∸ m) · g ^ x · f ^ m) (≤-∸-+-cancel right-≤-max) ⟩ + + r · f ^ (l ∸ m) · g ^ l · f ^ m + + ≡⟨ eq2 r (f ^ (l ∸ m)) (g ^ l) (f ^ m) ⟩ + + r · (f ^ (l ∸ m) · f ^ m) · g ^ l + + ≡⟨ cong (λ x → r · x · g ^ l) (·-of-^-is-^-of-+ _ _ _) ⟩ + + r · f ^ (l ∸ m +ℕ m) · g ^ l + + ≡⟨ cong (λ x → r · f ^ x · g ^ l) (≤-∸-+-cancel left-≤-max) ⟩ + + r · f ^ l · g ^ l + + ≡⟨ sym (·Assoc _ _ _) ⟩ + + r · (f ^ l · g ^ l) + + ≡⟨ cong (r ·_) (sym (^-ldist-· _ _ _)) ⟩ + + r · (f · g) ^ l + + ≡⟨ eq3 r ((f · g) ^ l) ⟩ + + 1r · (1r · (r · (f · g) ^ l) · 1r) · (1r · 1r · (1r · 1r)) ∎ + + + base-^ᶠ-helper : (r : R) (m n : ℕ) → ∃[ x ∈ R × S[fg] ] (x .fst /1/1) + ≡ [ [ r , f ^ m , PT.∣ m , refl ∣ ] + , g/1 ^ᶠ n , PT.∣ n , refl ∣ ] ·R[1/f][1/g] (x .snd .fst /1/1) + base-^ᶠ-helper r m n = subst (λ y → ∃[ x ∈ R × S[fg] ] (x .fst /1/1) + ≡ [ [ r , f ^ m , PT.∣ m , refl ∣ ] , y ] ·R[1/f][1/g] (x .snd .fst /1/1)) + (Σ≡Prop (λ _ → isPropPropTrunc) (^-respects-/1 _ n)) (baseCase r m n) + + indStep : (r : R[1/_] R' f) (n : ℕ) → ∃[ x ∈ R × S[fg] ] + (x .fst /1/1) ≡ [ r , g/1 ^ᶠ n , PT.∣ n , refl ∣ ] ·R[1/f][1/g] (x .snd .fst /1/1) + indStep = InvElPropElim _ (λ _ → isPropΠ λ _ → isPropPropTrunc) base-^ᶠ-helper + + toGoal : (r : R[1/_] R' f) (n : ℕ) → ∃[ x ∈ R × S[fg] ] + (x .fst /1/1) ·R[1/f][1/g] + ((x .snd .fst /1/1) ⁻¹) ⦃ φS⊆Aˣ pathtoR[1/fg] (x .snd .fst) (x .snd .snd) ⦄ + ≡ [ r , g/1 ^ᶠ n , PT.∣ n , refl ∣ ] + toGoal r n = PT.map Σhelper (indStep r n) + where + Σhelper : Σ[ x ∈ R × S[fg] ] + (x .fst /1/1) ≡ [ r , g/1 ^ᶠ n , PT.∣ n , refl ∣ ] ·R[1/f][1/g] (x .snd .fst /1/1) + → Σ[ x ∈ R × S[fg] ] + (x .fst /1/1) ·R[1/f][1/g] ((x .snd .fst /1/1) ⁻¹) + ⦃ φS⊆Aˣ pathtoR[1/fg] (x .snd .fst) (x .snd .snd) ⦄ + ≡ [ r , g/1 ^ᶠ n , PT.∣ n , refl ∣ ] + Σhelper ((r' , s , s∈S[fg]) , p) = (r' , s , s∈S[fg]) + , ⁻¹-eq-elim ⦃ φS⊆Aˣ pathtoR[1/fg] s s∈S[fg] ⦄ p + + + -- In this module we construct the map R[1/fg]→R[1/f][1/g] directly + -- and show that it is equal (although not judgementally) to the map induced + -- by the universal property of localisation, i.e. transporting along the path + -- constructed above. Given that this is the easier direction, we can see that + -- it is pretty cumbersome to prove R[1/fg]≡R[1/f][1/g] directly, + -- which illustrates the usefulness of S⁻¹RChar quite well. + private + module check where + φ : R[1/fg] → R[1/f][1/g] + φ = SQ.rec squash/ ϕ ϕcoh + where + S[fg] = Loc.S R' ([_ⁿ|n≥0] R' (f · g)) (powersFormMultClosedSubset R' (f · g)) + + curriedϕΣ : (r s : R) → Σ[ n ∈ ℕ ] s ≡ (f · g) ^ n → R[1/f][1/g] + curriedϕΣ r s (n , s≡fg^n) = [ [ r , (f ^ n) , PT.∣ n , refl ∣ ] + , [ (g ^ n) , 1r , PT.∣ 0 , refl ∣ ] --denominator + , PT.∣ n , ^-respects-/1 R' n ∣ ] + + curriedϕ : (r s : R) → ∃[ n ∈ ℕ ] s ≡ (f · g) ^ n → R[1/f][1/g] + curriedϕ r s = elim→Set (λ _ → squash/) (curriedϕΣ r s) coh + where + coh : (x y : Σ[ n ∈ ℕ ] s ≡ (f · g) ^ n) → curriedϕΣ r s x ≡ curriedϕΣ r s y + coh (n , s≡fg^n) (m , s≡fg^m) = eq/ _ _ ((1ᶠ , PT.∣ 0 , refl ∣) , + eq/ _ _ ( (1r , powersFormMultClosedSubset R' f .containsOne) + , path)) + where + eq1 : ∀ r gm fm → 1r · (1r · r · gm) · (1r · fm · 1r) ≡ r · (gm · fm) + eq1 = solve R' + + path : 1r · (1r · r · (g ^ m)) · (1r · (f ^ m) · 1r) + ≡ 1r · (1r · r · (g ^ n)) · (1r · (f ^ n) · 1r) + path = 1r · (1r · r · (g ^ m)) · (1r · (f ^ m) · 1r) + + ≡⟨ eq1 r (g ^ m) (f ^ m) ⟩ + + r · (g ^ m · f ^ m) + + ≡⟨ cong (r ·_) (sym (^-ldist-· g f m)) ⟩ + + r · ((g · f) ^ m) + + ≡⟨ cong (λ x → r · (x ^ m)) (·-comm _ _) ⟩ + + r · ((f · g) ^ m) + + ≡⟨ cong (r ·_) ((sym s≡fg^m) ∙ s≡fg^n) ⟩ + + r · ((f · g) ^ n) + + ≡⟨ cong (λ x → r · (x ^ n)) (·-comm _ _) ⟩ + + r · ((g · f) ^ n) + + ≡⟨ cong (r ·_) (^-ldist-· g f n) ⟩ + + r · (g ^ n · f ^ n) + + ≡⟨ sym (eq1 r (g ^ n) (f ^ n)) ⟩ + + 1r · (1r · r · (g ^ n)) · (1r · (f ^ n) · 1r) ∎ + + + ϕ : R × S[fg] → R[1/f][1/g] + ϕ (r , s , |n,s≡fg^n|) = curriedϕ r s |n,s≡fg^n| + -- λ (r / (fg)ⁿ) → ((r / fⁿ) / gⁿ) + + curriedϕcohΣ : (r s r' s' u : R) → (p : u · r · s' ≡ u · r' · s) + → (α : Σ[ n ∈ ℕ ] s ≡ (f · g) ^ n) + → (β : Σ[ m ∈ ℕ ] s' ≡ (f · g) ^ m) + → (γ : Σ[ l ∈ ℕ ] u ≡ (f · g) ^ l) + → ϕ (r , s , PT.∣ α ∣) ≡ ϕ (r' , s' , PT.∣ β ∣) + curriedϕcohΣ r s r' s' u p (n , s≡fgⁿ) (m , s'≡fgᵐ) (l , u≡fgˡ) = + eq/ _ _ ( ( [ (g ^ l) , 1r , powersFormMultClosedSubset R' f .containsOne ] + , PT.∣ l , ^-respects-/1 R' l ∣) + , eq/ _ _ ((f ^ l , PT.∣ l , refl ∣) , path)) + where + eq1 : ∀ fl gl r gm fm + → fl · (gl · r · gm) · (1r · fm · 1r) ≡ fl · gl · r · (gm · fm) + eq1 = solve R' + + path : f ^ l · (g ^ l · transp (λ i → R) i0 r · transp (λ i → R) i0 (g ^ m)) + · (1r · transp (λ i → R) i0 (f ^ m) · transp (λ i → R) i0 1r) + ≡ f ^ l · (g ^ l · transp (λ i → R) i0 r' · transp (λ i → R) i0 (g ^ n)) + · (1r · transp (λ i → R) i0 (f ^ n) · transp (λ i → R) i0 1r) + path = f ^ l · (g ^ l · transp (λ i → R) i0 r · transp (λ i → R) i0 (g ^ m)) + · (1r · transp (λ i → R) i0 (f ^ m) · transp (λ i → R) i0 1r) + + ≡⟨ (λ i → f ^ l · (g ^ l · transportRefl r i · transportRefl (g ^ m) i) + · (1r · transportRefl (f ^ m) i · transportRefl 1r i)) ⟩ + + f ^ l · (g ^ l · r · g ^ m) · (1r · f ^ m · 1r) + + ≡⟨ eq1 (f ^ l) (g ^ l) r (g ^ m) (f ^ m) ⟩ + + f ^ l · g ^ l · r · (g ^ m · f ^ m) + + ≡⟨ (λ i → ^-ldist-· f g l (~ i) · r · ^-ldist-· g f m (~ i)) ⟩ + + (f · g) ^ l · r · (g · f) ^ m + + ≡⟨ cong (λ x → (f · g) ^ l · r · x ^ m) (·-comm _ _) ⟩ + + (f · g) ^ l · r · (f · g) ^ m + + ≡⟨ (λ i → u≡fgˡ (~ i) · r · s'≡fgᵐ (~ i)) ⟩ + + u · r · s' + + ≡⟨ p ⟩ + + u · r' · s + + ≡⟨ (λ i → u≡fgˡ i · r' · s≡fgⁿ i) ⟩ + + (f · g) ^ l · r' · (f · g) ^ n + + ≡⟨ cong (λ x → (f · g) ^ l · r' · x ^ n) (·-comm _ _) ⟩ + + (f · g) ^ l · r' · (g · f) ^ n + + ≡⟨ (λ i → ^-ldist-· f g l i · r' · ^-ldist-· g f n i) ⟩ + + f ^ l · g ^ l · r' · (g ^ n · f ^ n) + + ≡⟨ sym (eq1 (f ^ l) (g ^ l) r' (g ^ n) (f ^ n)) ⟩ + + f ^ l · (g ^ l · r' · g ^ n) · (1r · f ^ n · 1r) + + ≡⟨ (λ i → f ^ l · (g ^ l · transportRefl r' (~ i) · transportRefl (g ^ n) (~ i)) + · (1r · transportRefl (f ^ n) (~ i) · transportRefl 1r (~ i))) ⟩ + + f ^ l · (g ^ l · transp (λ i → R) i0 r' · transp (λ i → R) i0 (g ^ n)) + · (1r · transp (λ i → R) i0 (f ^ n) · transp (λ i → R) i0 1r) ∎ + + + curriedϕcoh : (r s r' s' u : R) → (p : u · r · s' ≡ u · r' · s) + → (α : ∃[ n ∈ ℕ ] s ≡ (f · g) ^ n) + → (β : ∃[ m ∈ ℕ ] s' ≡ (f · g) ^ m) + → (γ : ∃[ l ∈ ℕ ] u ≡ (f · g) ^ l) + → ϕ (r , s , α) ≡ ϕ (r' , s' , β) + curriedϕcoh r s r' s' u p = PT.elim (λ _ → isPropΠ2 (λ _ _ → squash/ _ _)) + λ α → PT.elim (λ _ → isPropΠ (λ _ → squash/ _ _)) + λ β → PT.rec (squash/ _ _) + λ γ → curriedϕcohΣ r s r' s' u p α β γ + + ϕcoh : (a b : R × S[fg]) + → Loc._≈_ R' ([_ⁿ|n≥0] R' (f · g)) (powersFormMultClosedSubset R' (f · g)) a b + → ϕ a ≡ ϕ b + ϕcoh (r , s , α) (r' , s' , β) ((u , γ) , p) = curriedϕcoh r s r' s' u p α β γ + + + + + -- the map induced by the universal property + open S⁻¹RUniversalProp R' ([_ⁿ|n≥0] R' (f · g)) (powersFormMultClosedSubset R' (f · g)) + χ : R[1/fg] → R[1/f][1/g] + χ = S⁻¹RHasUniversalProp R[1/f][1/g]AsCommRing /1/1AsCommRingHom fⁿgⁿ/1/1∈R[1/f][1/g]ˣ .fst .fst .fst + + -- the sanity check: + -- both maps send a fraction r/(fg)ⁿ to a double fraction, + -- where numerator and denominator are almost the same fraction respectively. + -- unfortunately the proofs that the denominators are powers are quite different for + -- the two maps, but of course we can ignore them. + -- The definition of χ introduces a lot of (1r ·_). Perhaps most surprisingly, + -- we have to give the path eq1 for the equality of the numerator of the numerator. + φ≡χ : ∀ r → φ r ≡ χ r + φ≡χ = InvElPropElim _ (λ _ → squash/ _ _) ℕcase + where + ℕcase : (r : R) (n : ℕ) + → φ [ r , (f · g) ^ n , PT.∣ n , refl ∣ ] ≡ χ [ r , (f · g) ^ n , PT.∣ n , refl ∣ ] + ℕcase r n = cong [_] (ΣPathP --look into the components of the double-fractions + ( cong [_] (ΣPathP (eq1 , Σ≡Prop (λ x → S'[f] x .snd) (sym (·Lid _)))) + , Σ≡Prop (λ x → S'[f][g] x .snd) --ignore proof that denominator is power of g/1 + ( cong [_] (ΣPathP (sym (·Lid _) , Σ≡Prop (λ x → S'[f] x .snd) (sym (·Lid _))))))) + where + S'[f] = ([_ⁿ|n≥0] R' f) + S'[f][g] = ([_ⁿ|n≥0] R[1/f]AsCommRing [ g , 1r , powersFormMultClosedSubset R' f .containsOne ]) + + eq1 : transp (λ i → fst R') i0 r ≡ r · transp (λ i → fst R') i0 1r + eq1 = transportRefl r ∙∙ sym (·Rid r) ∙∙ cong (r ·_) (sym (transportRefl 1r)) diff --git a/Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda b/Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda new file mode 100644 index 000000000..e889f16d3 --- /dev/null +++ b/Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda @@ -0,0 +1,417 @@ +-- We define what it means to satisfy the universal property +-- of localisation and show that the localisation in Base satisfies +-- it. We will also show that the localisation is uniquely determined +-- by the universal property, and give sufficient criteria for +-- satisfying the universal property. + +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommRing.Localisation.UniversalProperty where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Transport +open import Cubical.Functions.FunExtEquiv +open import Cubical.Functions.Surjection +open import Cubical.Functions.Embedding + +import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Bool +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ + ; +-comm to +ℕ-comm ; +-assoc to +ℕ-assoc + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm) +open import Cubical.Data.Vec +open import Cubical.Data.Sigma.Base +open import Cubical.Data.Sigma.Properties +open import Cubical.Data.FinData +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.RingSolver.ReflectionSolving +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Localisation.Base + +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.HITs.PropositionalTruncation as PT + +open Iso + +private + variable + ℓ ℓ' : Level + + +module _ (R' : CommRing ℓ) (S' : ℙ (fst R')) (SMultClosedSubset : isMultClosedSubset R' S') where + open isMultClosedSubset + private R = fst R' + open CommRingStr (snd R') hiding (is-set) + open RingTheory (CommRing→Ring R') + open IsRingHom + + + + hasLocUniversalProp : (A : CommRing ℓ) (φ : CommRingHom R' A) + → (∀ s → s ∈ S' → fst φ s ∈ A ˣ) + → Type (ℓ-suc ℓ) + hasLocUniversalProp A φ _ = (B : CommRing ℓ) (ψ : CommRingHom R' B) + → (∀ s → s ∈ S' → fst ψ s ∈ B ˣ) + → ∃![ χ ∈ CommRingHom A B ] (fst χ) ∘ (fst φ) ≡ (fst ψ) + + isPropUniversalProp : (A : CommRing ℓ) (φ : CommRingHom R' A) + → (φS⊆Aˣ : ∀ s → s ∈ S' → fst φ s ∈ A ˣ) + → isProp (hasLocUniversalProp A φ φS⊆Aˣ) + isPropUniversalProp A φ φS⊆Aˣ = isPropΠ3 (λ _ _ _ → isPropIsContr) + + -- S⁻¹R has the universal property + module S⁻¹RUniversalProp where + open Loc R' S' SMultClosedSubset + _/1 : R → S⁻¹R + r /1 = [ r , 1r , SMultClosedSubset .containsOne ] + + /1AsCommRingHom : CommRingHom R' S⁻¹RAsCommRing + fst /1AsCommRingHom = _/1 + snd /1AsCommRingHom = + makeIsRingHom + refl + (λ r r' → cong [_] (≡-× (cong₂ (_+_) (sym (·Rid r)) (sym (·Rid r'))) + (Σ≡Prop (λ x → S' x .snd) (sym (·Lid 1r))))) + (λ _ _ → cong [_] (≡-× refl (Σ≡Prop (λ x → S' x .snd) (sym (·Lid 1r))))) + + S⁻¹Rˣ = S⁻¹RAsCommRing ˣ + S/1⊆S⁻¹Rˣ : ∀ s → s ∈ S' → (s /1) ∈ S⁻¹Rˣ + S/1⊆S⁻¹Rˣ s s∈S' = [ 1r , s , s∈S' ] , eq/ _ _ ((1r , SMultClosedSubset .containsOne) , path s) + where + path : ∀ s → 1r · (s · 1r) · 1r ≡ 1r · 1r · (1r · s) + path = solve R' + + S⁻¹RHasUniversalProp : hasLocUniversalProp S⁻¹RAsCommRing /1AsCommRingHom S/1⊆S⁻¹Rˣ + S⁻¹RHasUniversalProp B' ψ ψS⊆Bˣ = (χ , funExt χcomp) , χunique + where + B = fst B' + open CommRingStr (snd B') renaming ( is-set to Bset ; _·_ to _·B_ ; 1r to 1b + ; _+_ to _+B_ + ; ·Assoc to ·B-assoc ; ·-comm to ·B-comm + ; ·Lid to ·B-lid ; ·Rid to ·B-rid + ; ·Ldist+ to ·B-ldist-+) + open Units B' renaming (Rˣ to Bˣ ; RˣMultClosed to BˣMultClosed ; RˣContainsOne to BˣContainsOne) + open RingTheory (CommRing→Ring B') renaming (·-assoc2 to ·B-assoc2) + open CommRingTheory B' renaming (·-commAssocl to ·B-commAssocl ; ·-commAssocSwap to ·B-commAssocSwap) + + ψ₀ = fst ψ + module ψ = IsRingHom (snd ψ) + + χ : CommRingHom S⁻¹RAsCommRing B' + fst χ = SQ.rec Bset fχ fχcoh + where + fχ : R × S → B + fχ (r , s , s∈S') = (fst ψ r) ·B ((fst ψ s) ⁻¹) ⦃ ψS⊆Bˣ s s∈S' ⦄ + fχcoh : (a b : R × S) → a ≈ b → fχ a ≡ fχ b + fχcoh (r , s , s∈S') (r' , s' , s'∈S') ((u , u∈S') , p) = instancepath + ⦃ ψS⊆Bˣ s s∈S' ⦄ ⦃ ψS⊆Bˣ s' s'∈S' ⦄ ⦃ ψS⊆Bˣ (u · s · s') + (SMultClosedSubset .multClosed (SMultClosedSubset .multClosed u∈S' s∈S') s'∈S') ⦄ + ⦃ BˣMultClosed _ _ ⦃ ψS⊆Bˣ (u · s) (SMultClosedSubset .multClosed u∈S' s∈S') ⦄ + ⦃ ψS⊆Bˣ s' s'∈S' ⦄ ⦄ + ⦃ ψS⊆Bˣ (u · s) (SMultClosedSubset .multClosed u∈S' s∈S') ⦄ + where + -- only a few indidividual steps can be solved by the ring solver yet + instancepath : ⦃ _ : ψ₀ s ∈ Bˣ ⦄ ⦃ _ : ψ₀ s' ∈ Bˣ ⦄ + ⦃ _ : ψ₀ (u · s · s') ∈ Bˣ ⦄ ⦃ _ : ψ₀ (u · s) ·B ψ₀ s' ∈ Bˣ ⦄ + ⦃ _ : ψ₀ (u · s) ∈ Bˣ ⦄ + → ψ₀ r ·B ψ₀ s ⁻¹ ≡ ψ₀ r' ·B ψ₀ s' ⁻¹ + instancepath = ψ₀ r ·B ψ₀ s ⁻¹ + + ≡⟨ sym (·B-rid _) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B 1b + + ≡⟨ cong (ψ₀ r ·B ψ₀ s ⁻¹ ·B_) (sym (·-rinv _)) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B (ψ₀ (u · s · s') ·B ψ₀ (u · s · s') ⁻¹) + + ≡⟨ ·B-assoc _ _ _ ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B ψ₀ (u · s · s') ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (λ x → ψ₀ r ·B ψ₀ s ⁻¹ ·B x ·B ψ₀ (u · s · s') ⁻¹) (ψ.pres· _ _) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B (ψ₀ (u · s) ·B ψ₀ s') ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (_·B ψ₀ (u · s · s') ⁻¹) (·B-assoc _ _ _) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B ψ₀ (u · s) ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (λ x → ψ₀ r ·B ψ₀ s ⁻¹ ·B x ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹) + (ψ.pres· _ _) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B (ψ₀ u ·B ψ₀ s) ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (λ x → x ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹) (·B-commAssocSwap _ _ _ _) ⟩ + + ψ₀ r ·B ψ₀ u ·B (ψ₀ s ⁻¹ ·B ψ₀ s) ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ (λ i → ·B-comm (ψ₀ r) (ψ₀ u) i ·B (·-linv (ψ₀ s) i) + ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹) ⟩ + + ψ₀ u ·B ψ₀ r ·B 1b ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ (λ i → (·B-rid (sym (ψ.pres· u r) i) i) ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹) ⟩ + + ψ₀ (u · r) ·B ψ₀ s' ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (_·B ψ₀ (u · s · s') ⁻¹) (sym (ψ.pres· _ _)) ⟩ + + ψ₀ (u · r · s') ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (λ x → ψ₀ x ·B ψ₀ (u · s · s') ⁻¹) p ⟩ + + ψ₀ (u · r' · s) ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (_·B ψ₀ (u · s · s') ⁻¹) (ψ.pres· _ _) ⟩ + + ψ₀ (u · r') ·B ψ₀ s ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (λ x → x ·B ψ₀ s ·B ψ₀ (u · s · s') ⁻¹) (ψ.pres· _ _) ⟩ + + ψ₀ u ·B ψ₀ r' ·B ψ₀ s ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (_·B ψ₀ (u · s · s') ⁻¹) (sym (·B-assoc _ _ _)) ⟩ + + ψ₀ u ·B (ψ₀ r' ·B ψ₀ s) ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (_·B ψ₀ (u · s · s') ⁻¹) (·B-commAssocl _ _ _) ⟩ + + ψ₀ r' ·B (ψ₀ u ·B ψ₀ s) ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (λ x → ψ₀ r' ·B x ·B ψ₀ (u · s · s') ⁻¹) (sym (ψ.pres· _ _)) ⟩ + + ψ₀ r' ·B ψ₀ (u · s) ·B ψ₀ (u · s · s') ⁻¹ + + ≡⟨ cong (ψ₀ r' ·B ψ₀ (u · s) ·B_) (unitCong (ψ.pres· _ _)) ⟩ + + ψ₀ r' ·B ψ₀ (u · s) ·B (ψ₀ (u · s) ·B ψ₀ s') ⁻¹ + + ≡⟨ cong (ψ₀ r' ·B ψ₀ (u · s) ·B_) (⁻¹-dist-· _ _) ⟩ + + ψ₀ r' ·B ψ₀ (u · s) ·B (ψ₀ (u · s) ⁻¹ ·B ψ₀ s' ⁻¹) + + ≡⟨ ·B-assoc2 _ _ _ _ ⟩ + + ψ₀ r' ·B (ψ₀ (u · s) ·B ψ₀ (u · s) ⁻¹) ·B ψ₀ s' ⁻¹ + + ≡⟨ cong (λ x → ψ₀ r' ·B x ·B ψ₀ s' ⁻¹) (·-rinv _) ⟩ + + ψ₀ r' ·B 1b ·B ψ₀ s' ⁻¹ + + ≡⟨ cong (_·B ψ₀ s' ⁻¹) (·B-rid _) ⟩ + + ψ₀ r' ·B ψ₀ s' ⁻¹ ∎ + + snd χ = + makeIsRingHom + (instancepres1χ ⦃ ψS⊆Bˣ 1r (SMultClosedSubset .containsOne) ⦄ ⦃ BˣContainsOne ⦄) + (elimProp2 (λ _ _ _ _ → Bset _ _ _ _) pres+[]) + (elimProp2 (λ _ _ _ _ → Bset _ _ _ _) pres·[]) + where + instancepres1χ : ⦃ _ : ψ₀ 1r ∈ Bˣ ⦄ ⦃ _ : 1b ∈ Bˣ ⦄ + → ψ₀ 1r ·B (ψ₀ 1r) ⁻¹ ≡ 1b + instancepres1χ = (λ i → (ψ.pres1 i) ·B (unitCong (ψ.pres1) i)) + ∙ (λ i → ·B-lid (1⁻¹≡1 i) i) + + pres+[] : (a b : R × S) → fst χ ([ a ] +ₗ [ b ]) ≡ (fst χ [ a ]) +B (fst χ [ b ]) + pres+[] (r , s , s∈S') (r' , s' , s'∈S') = instancepath + ⦃ ψS⊆Bˣ s s∈S' ⦄ ⦃ ψS⊆Bˣ s' s'∈S' ⦄ ⦃ ψS⊆Bˣ (s · s') (SMultClosedSubset .multClosed s∈S' s'∈S') ⦄ + ⦃ BˣMultClosed _ _ ⦃ ψS⊆Bˣ s s∈S' ⦄ ⦃ ψS⊆Bˣ s' s'∈S' ⦄ ⦄ + where + -- only a few indidividual steps can be solved by the ring solver yet + instancepath : ⦃ _ : ψ₀ s ∈ Bˣ ⦄ ⦃ _ : ψ₀ s' ∈ Bˣ ⦄ + ⦃ _ : ψ₀ (s · s') ∈ Bˣ ⦄ ⦃ _ : ψ₀ s ·B ψ₀ s' ∈ Bˣ ⦄ + → ψ₀ (r · s' + r' · s) ·B ψ₀ (s · s') ⁻¹ ≡ ψ₀ r ·B ψ₀ s ⁻¹ +B ψ₀ r' ·B ψ₀ s' ⁻¹ + instancepath = + ψ₀ (r · s' + r' · s) ·B ψ₀ (s · s') ⁻¹ + + ≡⟨ (λ i → ψ.pres+ (r · s') (r' · s) i ·B unitCong (ψ.pres· s s') i) ⟩ + + (ψ₀ (r · s') +B ψ₀ (r' · s)) ·B (ψ₀ s ·B ψ₀ s') ⁻¹ + + ≡⟨ (λ i → (ψ.pres· r s' i +B ψ.pres· r' s i) ·B ⁻¹-dist-· (ψ₀ s) (ψ₀ s') i) ⟩ + + (ψ₀ r ·B ψ₀ s' +B ψ₀ r' ·B ψ₀ s) ·B (ψ₀ s ⁻¹ ·B ψ₀ s' ⁻¹) + + ≡⟨ ·B-ldist-+ _ _ _ ⟩ + + ψ₀ r ·B ψ₀ s' ·B (ψ₀ s ⁻¹ ·B ψ₀ s' ⁻¹) +B ψ₀ r' ·B ψ₀ s ·B (ψ₀ s ⁻¹ ·B ψ₀ s' ⁻¹) + + ≡⟨ (λ i → ·B-commAssocSwap (ψ₀ r) (ψ₀ s') (ψ₀ s ⁻¹) (ψ₀ s' ⁻¹) i + +B ·B-assoc2 (ψ₀ r') (ψ₀ s) (ψ₀ s ⁻¹) (ψ₀ s' ⁻¹) i) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B (ψ₀ s' ·B ψ₀ s' ⁻¹) +B ψ₀ r' ·B (ψ₀ s ·B ψ₀ s ⁻¹) ·B ψ₀ s' ⁻¹ + + ≡⟨ (λ i → ψ₀ r ·B ψ₀ s ⁻¹ ·B (·-rinv (ψ₀ s') i) + +B ψ₀ r' ·B (·-rinv (ψ₀ s) i) ·B ψ₀ s' ⁻¹) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B 1b +B ψ₀ r' ·B 1b ·B ψ₀ s' ⁻¹ + + ≡⟨ (λ i → ·B-rid (ψ₀ r ·B ψ₀ s ⁻¹) i +B ·B-rid (ψ₀ r') i ·B ψ₀ s' ⁻¹) ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ +B ψ₀ r' ·B ψ₀ s' ⁻¹ ∎ + + pres·[] : (a b : R × S) → fst χ ([ a ] ·ₗ [ b ]) ≡ (fst χ [ a ]) ·B (fst χ [ b ]) + pres·[] (r , s , s∈S') (r' , s' , s'∈S') = instancepath + ⦃ ψS⊆Bˣ s s∈S' ⦄ ⦃ ψS⊆Bˣ s' s'∈S' ⦄ ⦃ ψS⊆Bˣ (s · s') (SMultClosedSubset .multClosed s∈S' s'∈S') ⦄ + ⦃ BˣMultClosed _ _ ⦃ ψS⊆Bˣ s s∈S' ⦄ ⦃ ψS⊆Bˣ s' s'∈S' ⦄ ⦄ + where + -- only a indidividual steps can be solved by the ring solver yet + instancepath : ⦃ _ : ψ₀ s ∈ Bˣ ⦄ ⦃ _ : ψ₀ s' ∈ Bˣ ⦄ + ⦃ _ : ψ₀ (s · s') ∈ Bˣ ⦄ ⦃ _ : ψ₀ s ·B ψ₀ s' ∈ Bˣ ⦄ + → ψ₀ (r · r') ·B ψ₀ (s · s') ⁻¹ ≡ (ψ₀ r ·B ψ₀ s ⁻¹) ·B (ψ₀ r' ·B ψ₀ s' ⁻¹) + instancepath = ψ₀ (r · r') ·B ψ₀ (s · s') ⁻¹ + + ≡⟨ (λ i → ψ.pres· r r' i ·B unitCong (ψ.pres· s s') i) ⟩ + + ψ₀ r ·B ψ₀ r' ·B (ψ₀ s ·B ψ₀ s') ⁻¹ + + ≡⟨ cong (ψ₀ r ·B ψ₀ r' ·B_) (⁻¹-dist-· _ _) ⟩ + + ψ₀ r ·B ψ₀ r' ·B (ψ₀ s ⁻¹ ·B ψ₀ s' ⁻¹) + + ≡⟨ ·B-commAssocSwap _ _ _ _ ⟩ + + ψ₀ r ·B ψ₀ s ⁻¹ ·B (ψ₀ r' ·B ψ₀ s' ⁻¹) ∎ + + + χcomp : (r : R) → fst χ (r /1) ≡ ψ₀ r + χcomp = instanceχcomp ⦃ ψS⊆Bˣ 1r (SMultClosedSubset .containsOne) ⦄ ⦃ Units.RˣContainsOne B' ⦄ + where + instanceχcomp : ⦃ _ : ψ₀ 1r ∈ Bˣ ⦄ ⦃ _ : 1b ∈ Bˣ ⦄ + (r : R) → ψ₀ r ·B (ψ₀ 1r) ⁻¹ ≡ ψ₀ r + instanceχcomp r = ψ₀ r ·B (ψ₀ 1r) ⁻¹ ≡⟨ cong (ψ₀ r ·B_) (unitCong (ψ.pres1)) ⟩ + ψ₀ r ·B 1b ⁻¹ ≡⟨ cong (ψ₀ r ·B_) 1⁻¹≡1 ⟩ + ψ₀ r ·B 1b ≡⟨ ·B-rid _ ⟩ + ψ₀ r ∎ + + + χunique : (y : Σ[ χ' ∈ CommRingHom S⁻¹RAsCommRing B' ] fst χ' ∘ _/1 ≡ ψ₀) + → (χ , funExt χcomp) ≡ y + χunique (χ' , χ'/1≡ψ) = Σ≡Prop (λ x → isSetΠ (λ _ → Bset) _ _) (RingHom≡f _ _ fχ≡fχ') + where + open RingHomRespUnits {A' = S⁻¹RAsCommRing} {B' = B'} χ' + renaming (φ[x⁻¹]≡φ[x]⁻¹ to χ'[x⁻¹]≡χ'[x]⁻¹) + + module χ' = IsRingHom (χ' .snd) + + []-path : (a : R × S) → fst χ [ a ] ≡ fst χ' [ a ] + []-path (r , s , s∈S') = instancepath ⦃ ψS⊆Bˣ s s∈S' ⦄ ⦃ S/1⊆S⁻¹Rˣ s s∈S' ⦄ + ⦃ RingHomRespInv _ ⦃ S/1⊆S⁻¹Rˣ s s∈S' ⦄ ⦄ + where + open Units S⁻¹RAsCommRing renaming (_⁻¹ to _⁻¹ˡ ; inverseUniqueness to S⁻¹RInverseUniqueness) + hiding (unitCong) + + s-inv : ⦃ s/1∈S⁻¹Rˣ : s /1 ∈ S⁻¹Rˣ ⦄ → s /1 ⁻¹ˡ ≡ [ 1r , s , s∈S' ] + s-inv ⦃ s/1∈S⁻¹Rˣ ⦄ = PathPΣ (S⁻¹RInverseUniqueness (s /1) s/1∈S⁻¹Rˣ + (_ , eq/ _ _ ((1r , SMultClosedSubset .containsOne) , path s))) .fst + where + path : ∀ s → 1r · (s · 1r) · 1r ≡ 1r · 1r · (1r · s) + path = solve R' + + ·ₗ-path : [ r , s , s∈S' ] ≡ [ r , 1r , SMultClosedSubset .containsOne ] + ·ₗ [ 1r , s , s∈S' ] + ·ₗ-path = cong [_] (≡-× (sym (·Rid r)) (Σ≡Prop (λ x → S' x .snd) (sym (·Lid s)))) + + instancepath : ⦃ _ : ψ₀ s ∈ Bˣ ⦄ ⦃ _ : s /1 ∈ S⁻¹Rˣ ⦄ ⦃ _ : fst χ' (s /1) ∈ Bˣ ⦄ + → ψ₀ r ·B ψ₀ s ⁻¹ ≡ fst χ' [ r , s , s∈S' ] + instancepath = ψ₀ r ·B ψ₀ s ⁻¹ + + ≡⟨ cong (ψ₀ r ·B_) (unitCong (cong (λ φ → φ s) (sym χ'/1≡ψ))) ⟩ + + ψ₀ r ·B fst χ' (s /1) ⁻¹ + + ≡⟨ cong (ψ₀ r ·B_) (sym (χ'[x⁻¹]≡χ'[x]⁻¹ _)) ⟩ + + ψ₀ r ·B fst χ' (s /1 ⁻¹ˡ) + + ≡⟨ cong (λ x → ψ₀ r ·B fst χ' x) s-inv ⟩ + + ψ₀ r ·B fst χ' [ 1r , s , s∈S' ] + + ≡⟨ cong (_·B fst χ' [ 1r , s , s∈S' ]) (cong (λ φ → φ r) (sym χ'/1≡ψ)) ⟩ + + fst χ' [ r , 1r , SMultClosedSubset .containsOne ] ·B fst χ' [ 1r , s , s∈S' ] + + ≡⟨ sym (χ'.pres· _ _) ⟩ + + fst χ' ([ r , 1r , SMultClosedSubset .containsOne ] ·ₗ [ 1r , s , s∈S' ]) + + ≡⟨ cong (fst χ') (sym ·ₗ-path) ⟩ + + fst χ' [ r , s , s∈S' ] ∎ + + fχ≡fχ' : fst χ ≡ fst χ' + fχ≡fχ' = funExt (SQ.elimProp (λ _ → Bset _ _) []-path) + + + -- sufficient conditions for having the universal property + -- used as API in the leanprover-community/mathlib + -- Corollary 3.2 in Atiyah-McDonald + open S⁻¹RUniversalProp + open Loc R' S' SMultClosedSubset + + record PathToS⁻¹R (A' : CommRing ℓ) (φ : CommRingHom R' A') : Type ℓ where + constructor + pathtoS⁻¹R + open Units A' renaming (Rˣ to Aˣ) + open CommRingStr (snd A') renaming (is-set to Aset ; 0r to 0a ; _·_ to _·A_) + field + φS⊆Aˣ : ∀ s → s ∈ S' → fst φ s ∈ Aˣ + kerφ⊆annS : ∀ r → fst φ r ≡ 0a → ∃[ s ∈ S ] (s .fst) · r ≡ 0r + surχ : ∀ a → ∃[ x ∈ R × S ] fst φ (x .fst) ·A (fst φ (x .snd .fst) ⁻¹) ⦃ φS⊆Aˣ _ (x .snd .snd) ⦄ ≡ a + + S⁻¹RChar : (A' : CommRing ℓ) (φ : CommRingHom R' A') + → PathToS⁻¹R A' φ + → S⁻¹RAsCommRing ≡ A' + S⁻¹RChar A' φ cond = CommRingPath S⁻¹RAsCommRing A' .fst + (S⁻¹R≃A , χ .snd) + where + open CommRingStr (snd A') renaming ( is-set to Aset ; 0r to 0a ; _·_ to _·A_ ; 1r to 1a + ; ·Rid to ·A-rid) + open Units A' renaming (Rˣ to Aˣ ; RˣInvClosed to AˣInvClosed) + open PathToS⁻¹R ⦃...⦄ + private + A = fst A' + instance + _ = cond + χ = (S⁻¹RHasUniversalProp A' φ φS⊆Aˣ .fst .fst) + open HomTheory χ + + S⁻¹R≃A : S⁻¹R ≃ A + S⁻¹R≃A = fst χ , isEmbedding×isSurjection→isEquiv (Embχ , Surχ) + where + Embχ : isEmbedding (fst χ) + Embχ = injEmbedding squash/ Aset (ker≡0→inj λ {x} → kerχ≡0 x) + where + kerχ≡0 : (r/s : S⁻¹R) → fst χ r/s ≡ 0a → r/s ≡ 0ₗ + kerχ≡0 = SQ.elimProp (λ _ → isPropΠ λ _ → squash/ _ _) kerχ≡[] + where + kerχ≡[] : (a : R × S) → fst χ [ a ] ≡ 0a → [ a ] ≡ 0ₗ + kerχ≡[] (r , s , s∈S') p = PT.rec (squash/ _ _) Σhelper + (kerφ⊆annS r (UnitsAreNotZeroDivisors _ _ p)) + where + instance + _ : fst φ s ∈ Aˣ + _ = φS⊆Aˣ s s∈S' + _ : fst φ s ⁻¹ ∈ Aˣ + _ = AˣInvClosed _ + + Σhelper : Σ[ s ∈ S ] (s .fst) · r ≡ 0r → [ r , s , s∈S' ] ≡ 0ₗ + Σhelper ((u , u∈S') , q) = eq/ _ _ ((u , u∈S') , path) + where + path : u · r · 1r ≡ u · 0r · s + path = (λ i → ·Rid (q i) i) ∙∙ sym (0LeftAnnihilates _) + ∙∙ cong (_· s) (sym (0RightAnnihilates _)) + + Surχ : isSurjection (fst χ) + Surχ a = PT.rec isPropPropTrunc (λ x → PT.∣ [ x .fst ] , x .snd ∣) (surχ a) diff --git a/Cubical/Algebra/CommRing/Properties.agda b/Cubical/Algebra/CommRing/Properties.agda new file mode 100644 index 000000000..8e204b565 --- /dev/null +++ b/Cubical/Algebra/CommRing/Properties.agda @@ -0,0 +1,225 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommRing.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Powerset + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm) + +open import Cubical.Structures.Axioms +open import Cubical.Structures.Auto +open import Cubical.Structures.Macro +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing.Base + +private + variable + ℓ : Level + +module Units (R' : CommRing ℓ) where + open CommRingStr (snd R') + open RingTheory (CommRing→Ring R') + private R = fst R' + + inverseUniqueness : (r : R) → isProp (Σ[ r' ∈ R ] r · r' ≡ 1r) + inverseUniqueness r (r' , rr'≡1) (r'' , rr''≡1) = Σ≡Prop (λ _ → is-set _ _) path + where + path : r' ≡ r'' + path = r' ≡⟨ sym (·Rid _) ⟩ + r' · 1r ≡⟨ cong (r' ·_) (sym rr''≡1) ⟩ + r' · (r · r'') ≡⟨ ·Assoc _ _ _ ⟩ + (r' · r) · r'' ≡⟨ cong (_· r'') (·-comm _ _) ⟩ + (r · r') · r'' ≡⟨ cong (_· r'') rr'≡1 ⟩ + 1r · r'' ≡⟨ ·Lid _ ⟩ + r'' ∎ + + + Rˣ : ℙ R + Rˣ r = (Σ[ r' ∈ R ] r · r' ≡ 1r) , inverseUniqueness r + + -- some notation using instance arguments + _⁻¹ : (r : R) → ⦃ r ∈ Rˣ ⦄ → R + _⁻¹ r ⦃ r∈Rˣ ⦄ = r∈Rˣ .fst + + infix 9 _⁻¹ + + -- some results about inverses + ·-rinv : (r : R) ⦃ r∈Rˣ : r ∈ Rˣ ⦄ → r · r ⁻¹ ≡ 1r + ·-rinv r ⦃ r∈Rˣ ⦄ = r∈Rˣ .snd + + ·-linv : (r : R) ⦃ r∈Rˣ : r ∈ Rˣ ⦄ → r ⁻¹ · r ≡ 1r + ·-linv r ⦃ r∈Rˣ ⦄ = ·-comm _ _ ∙ r∈Rˣ .snd + + + RˣMultClosed : (r r' : R) ⦃ r∈Rˣ : r ∈ Rˣ ⦄ ⦃ r'∈Rˣ : r' ∈ Rˣ ⦄ + → (r · r') ∈ Rˣ + RˣMultClosed r r' = (r ⁻¹ · r' ⁻¹) , path + where + path : r · r' · (r ⁻¹ · r' ⁻¹) ≡ 1r + path = r · r' · (r ⁻¹ · r' ⁻¹) ≡⟨ cong (_· (r ⁻¹ · r' ⁻¹)) (·-comm _ _) ⟩ + r' · r · (r ⁻¹ · r' ⁻¹) ≡⟨ ·Assoc _ _ _ ⟩ + r' · r · r ⁻¹ · r' ⁻¹ ≡⟨ cong (_· r' ⁻¹) (sym (·Assoc _ _ _)) ⟩ + r' · (r · r ⁻¹) · r' ⁻¹ ≡⟨ cong (λ x → r' · x · r' ⁻¹) (·-rinv _) ⟩ + r' · 1r · r' ⁻¹ ≡⟨ cong (_· r' ⁻¹) (·Rid _) ⟩ + r' · r' ⁻¹ ≡⟨ ·-rinv _ ⟩ + 1r ∎ + + RˣContainsOne : 1r ∈ Rˣ + RˣContainsOne = 1r , ·Lid _ + + RˣInvClosed : (r : R) ⦃ _ : r ∈ Rˣ ⦄ → r ⁻¹ ∈ Rˣ + RˣInvClosed r = r , ·-linv _ + + UnitsAreNotZeroDivisors : (r : R) ⦃ _ : r ∈ Rˣ ⦄ + → ∀ r' → r' · r ≡ 0r → r' ≡ 0r + UnitsAreNotZeroDivisors r r' p = r' ≡⟨ sym (·Rid _) ⟩ + r' · 1r ≡⟨ cong (r' ·_) (sym (·-rinv _)) ⟩ + r' · (r · r ⁻¹) ≡⟨ ·Assoc _ _ _ ⟩ + r' · r · r ⁻¹ ≡⟨ cong (_· r ⁻¹) p ⟩ + 0r · r ⁻¹ ≡⟨ 0LeftAnnihilates _ ⟩ + 0r ∎ + + + -- laws keeping the instance arguments + 1⁻¹≡1 : ⦃ 1∈Rˣ' : 1r ∈ Rˣ ⦄ → 1r ⁻¹ ≡ 1r + 1⁻¹≡1 ⦃ 1∈Rˣ' ⦄ = (sym (·Lid _)) ∙ 1∈Rˣ' .snd + + ⁻¹-dist-· : (r r' : R) ⦃ r∈Rˣ : r ∈ Rˣ ⦄ ⦃ r'∈Rˣ : r' ∈ Rˣ ⦄ ⦃ rr'∈Rˣ : (r · r') ∈ Rˣ ⦄ + → (r · r') ⁻¹ ≡ r ⁻¹ · r' ⁻¹ + ⁻¹-dist-· r r' ⦃ r∈Rˣ ⦄ ⦃ r'∈Rˣ ⦄ ⦃ rr'∈Rˣ ⦄ = + sym path ∙∙ cong (r ⁻¹ · r' ⁻¹ ·_) (rr'∈Rˣ .snd) ∙∙ (·Rid _) + where + path : r ⁻¹ · r' ⁻¹ · (r · r' · (r · r') ⁻¹) ≡ (r · r') ⁻¹ + path = r ⁻¹ · r' ⁻¹ · (r · r' · (r · r') ⁻¹) + ≡⟨ ·Assoc _ _ _ ⟩ + r ⁻¹ · r' ⁻¹ · (r · r') · (r · r') ⁻¹ + ≡⟨ cong (λ x → r ⁻¹ · r' ⁻¹ · x · (r · r') ⁻¹) (·-comm _ _) ⟩ + r ⁻¹ · r' ⁻¹ · (r' · r) · (r · r') ⁻¹ + ≡⟨ cong (_· (r · r') ⁻¹) (sym (·Assoc _ _ _)) ⟩ + r ⁻¹ · (r' ⁻¹ · (r' · r)) · (r · r') ⁻¹ + ≡⟨ cong (λ x → r ⁻¹ · x · (r · r') ⁻¹) (·Assoc _ _ _) ⟩ + r ⁻¹ · (r' ⁻¹ · r' · r) · (r · r') ⁻¹ + ≡⟨ cong (λ x → r ⁻¹ · (x · r) · (r · r') ⁻¹) (·-linv _) ⟩ + r ⁻¹ · (1r · r) · (r · r') ⁻¹ + ≡⟨ cong (λ x → r ⁻¹ · x · (r · r') ⁻¹) (·Lid _) ⟩ + r ⁻¹ · r · (r · r') ⁻¹ + ≡⟨ cong (_· (r · r') ⁻¹) (·-linv _) ⟩ + 1r · (r · r') ⁻¹ + ≡⟨ ·Lid _ ⟩ + (r · r') ⁻¹ ∎ + + unitCong : {r r' : R} → r ≡ r' → ⦃ r∈Rˣ : r ∈ Rˣ ⦄ ⦃ r'∈Rˣ : r' ∈ Rˣ ⦄ → r ⁻¹ ≡ r' ⁻¹ + unitCong {r = r} {r' = r'} p ⦃ r∈Rˣ ⦄ ⦃ r'∈Rˣ ⦄ = + PathPΣ (inverseUniqueness r' (r ⁻¹ , subst (λ x → x · r ⁻¹ ≡ 1r) p (r∈Rˣ .snd)) r'∈Rˣ) .fst + + ⁻¹-eq-elim : {r r' r'' : R} ⦃ r∈Rˣ : r ∈ Rˣ ⦄ → r' ≡ r'' · r → r' · r ⁻¹ ≡ r'' + ⁻¹-eq-elim {r = r} {r'' = r''} p = cong (_· r ⁻¹) p + ∙ sym (·Assoc _ _ _) + ∙ cong (r'' ·_) (·-rinv _) + ∙ ·Rid _ + + +-- some convenient notation +_ˣ : (R' : CommRing ℓ) → ℙ (R' .fst) +R' ˣ = Units.Rˣ R' + +module RingHomRespUnits {A' B' : CommRing ℓ} (φ : CommRingHom A' B') where + open Units A' renaming (Rˣ to Aˣ ; _⁻¹ to _⁻¹ᵃ ; ·-rinv to ·A-rinv ; ·-linv to ·A-linv) + private A = fst A' + open CommRingStr (snd A') renaming (_·_ to _·A_ ; 1r to 1a) + open Units B' renaming (Rˣ to Bˣ ; _⁻¹ to _⁻¹ᵇ ; ·-rinv to ·B-rinv) + open CommRingStr (snd B') renaming ( _·_ to _·B_ ; 1r to 1b + ; ·Lid to ·B-lid ; ·Rid to ·B-rid + ; ·Assoc to ·B-assoc) + + private + f = fst φ + open IsRingHom (φ .snd) + + RingHomRespInv : (r : A) ⦃ r∈Aˣ : r ∈ Aˣ ⦄ → f r ∈ Bˣ + RingHomRespInv r = f (r ⁻¹ᵃ) , (sym (pres· r (r ⁻¹ᵃ)) ∙∙ cong (f) (·A-rinv r) ∙∙ pres1) + + φ[x⁻¹]≡φ[x]⁻¹ : (r : A) ⦃ r∈Aˣ : r ∈ Aˣ ⦄ ⦃ φr∈Bˣ : f r ∈ Bˣ ⦄ + → f (r ⁻¹ᵃ) ≡ (f r) ⁻¹ᵇ + φ[x⁻¹]≡φ[x]⁻¹ r ⦃ r∈Aˣ ⦄ ⦃ φr∈Bˣ ⦄ = + f (r ⁻¹ᵃ) ≡⟨ sym (·B-rid _) ⟩ + f (r ⁻¹ᵃ) ·B 1b ≡⟨ cong (f (r ⁻¹ᵃ) ·B_) (sym (·B-rinv _)) ⟩ + f (r ⁻¹ᵃ) ·B ((f r) ·B (f r) ⁻¹ᵇ) ≡⟨ ·B-assoc _ _ _ ⟩ + f (r ⁻¹ᵃ) ·B (f r) ·B (f r) ⁻¹ᵇ ≡⟨ cong (_·B (f r) ⁻¹ᵇ) (sym (pres· _ _)) ⟩ + f (r ⁻¹ᵃ ·A r) ·B (f r) ⁻¹ᵇ ≡⟨ cong (λ x → f x ·B (f r) ⁻¹ᵇ) (·A-linv _) ⟩ + f 1a ·B (f r) ⁻¹ᵇ ≡⟨ cong (_·B (f r) ⁻¹ᵇ) (pres1) ⟩ + 1b ·B (f r) ⁻¹ᵇ ≡⟨ ·B-lid _ ⟩ + (f r) ⁻¹ᵇ ∎ + + +module Exponentiation (R' : CommRing ℓ) where + open CommRingStr (snd R') + private R = fst R' + + -- introduce exponentiation + _^_ : R → ℕ → R + f ^ zero = 1r + f ^ suc n = f · (f ^ n) + + infix 9 _^_ + + -- and prove some laws + 1ⁿ≡1 : (n : ℕ) → 1r ^ n ≡ 1r + 1ⁿ≡1 zero = refl + 1ⁿ≡1 (suc n) = ·Lid _ ∙ 1ⁿ≡1 n + + ·-of-^-is-^-of-+ : (f : R) (m n : ℕ) → (f ^ m) · (f ^ n) ≡ f ^ (m +ℕ n) + ·-of-^-is-^-of-+ f zero n = ·Lid _ + ·-of-^-is-^-of-+ f (suc m) n = sym (·Assoc _ _ _) ∙ cong (f ·_) (·-of-^-is-^-of-+ f m n) + + ^-ldist-· : (f g : R) (n : ℕ) → (f · g) ^ n ≡ (f ^ n) · (g ^ n) + ^-ldist-· f g zero = sym (·Lid 1r) + ^-ldist-· f g (suc n) = path + where + path : f · g · ((f · g) ^ n) ≡ f · (f ^ n) · (g · (g ^ n)) + path = f · g · ((f · g) ^ n) ≡⟨ cong (f · g ·_) (^-ldist-· f g n) ⟩ + f · g · ((f ^ n) · (g ^ n)) ≡⟨ ·Assoc _ _ _ ⟩ + f · g · (f ^ n) · (g ^ n) ≡⟨ cong (_· (g ^ n)) (sym (·Assoc _ _ _)) ⟩ + f · (g · (f ^ n)) · (g ^ n) ≡⟨ cong (λ r → (f · r) · (g ^ n)) (·-comm _ _) ⟩ + f · ((f ^ n) · g) · (g ^ n) ≡⟨ cong (_· (g ^ n)) (·Assoc _ _ _) ⟩ + f · (f ^ n) · g · (g ^ n) ≡⟨ sym (·Assoc _ _ _) ⟩ + f · (f ^ n) · (g · (g ^ n)) ∎ + + ^-rdist-·ℕ : (f : R) (n m : ℕ) → f ^ (n ·ℕ m) ≡ (f ^ n) ^ m + ^-rdist-·ℕ f zero m = sym (1ⁿ≡1 m) + ^-rdist-·ℕ f (suc n) m = sym (·-of-^-is-^-of-+ f m (n ·ℕ m)) + ∙∙ cong (f ^ m ·_) (^-rdist-·ℕ f n m) + ∙∙ sym (^-ldist-· f (f ^ n) m) + + +-- like in Ring.Properties we provide helpful lemmas here +module CommRingTheory (R' : CommRing ℓ) where + open CommRingStr (snd R') + private R = fst R' + + ·-commAssocl : (x y z : R) → x · (y · z) ≡ y · (x · z) + ·-commAssocl x y z = ·Assoc x y z ∙∙ cong (_· z) (·-comm x y) ∙∙ sym (·Assoc y x z) + + ·-commAssocr : (x y z : R) → x · y · z ≡ x · z · y + ·-commAssocr x y z = sym (·Assoc x y z) ∙∙ cong (x ·_) (·-comm y z) ∙∙ ·Assoc x z y + + + ·-commAssocr2 : (x y z : R) → x · y · z ≡ z · y · x + ·-commAssocr2 x y z = ·-commAssocr _ _ _ ∙∙ cong (_· y) (·-comm _ _) ∙∙ ·-commAssocr _ _ _ + + ·-commAssocSwap : (x y z w : R) → (x · y) · (z · w) ≡ (x · z) · (y · w) + ·-commAssocSwap x y z w = ·Assoc (x · y) z w ∙∙ cong (_· w) (·-commAssocr x y z) + ∙∙ sym (·Assoc (x · z) y w) + diff --git a/Cubical/Algebra/CommRing/RadicalIdeal.agda b/Cubical/Algebra/CommRing/RadicalIdeal.agda new file mode 100644 index 000000000..8c847e53c --- /dev/null +++ b/Cubical/Algebra/CommRing/RadicalIdeal.agda @@ -0,0 +1,122 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommRing.RadicalIdeal where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Sigma +open import Cubical.Data.Sum hiding (map) +open import Cubical.Data.FinData hiding (elim) +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ + ; +-comm to +ℕ-comm + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm + ; _choose_ to _ℕchoose_ ; snotz to ℕsnotz) +open import Cubical.Data.Nat.Order + +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Ideal +open import Cubical.Algebra.CommRing.FGIdeal +open import Cubical.Algebra.CommRing.BinomialThm +open import Cubical.Algebra.Ring.Properties +open import Cubical.Algebra.Ring.BigOps +open import Cubical.Algebra.RingSolver.ReflectionSolving + +private + variable + ℓ : Level + +module _ (R' : CommRing ℓ) where + private R = fst R' + open CommRingStr (snd R') + open RingTheory (CommRing→Ring R') + open Sum (CommRing→Ring R') + open CommRingTheory R' + open Exponentiation R' + open BinomialThm R' + open isCommIdeal + + + √ : ℙ R → ℙ R --\surd + √ I x = (∃[ n ∈ ℕ ] x ^ n ∈ I) , isPropPropTrunc + + ^∈√→∈√ : ∀ (I : ℙ R) (x : R) (n : ℕ) → x ^ n ∈ √ I → x ∈ √ I + ^∈√→∈√ I x n = + map (λ { (m , [xⁿ]ᵐ∈I) → (n ·ℕ m) , subst-∈ I (sym (^-rdist-·ℕ x n m)) [xⁿ]ᵐ∈I }) + + ∈→∈√ : ∀ (I : ℙ R) (x : R) → x ∈ I → x ∈ √ I + ∈→∈√ I _ x∈I = ∣ 1 , subst-∈ I (sym (·Rid _)) x∈I ∣ + + √OfIdealIsIdeal : ∀ (I : ℙ R) → isCommIdeal R' I → isCommIdeal R' (√ I) + +Closed (√OfIdealIsIdeal I ici) {x = x} {y = y} = map2 +ClosedΣ + where + +ClosedΣ : Σ[ n ∈ ℕ ] x ^ n ∈ I → Σ[ n ∈ ℕ ] y ^ n ∈ I → Σ[ n ∈ ℕ ] (x + y) ^ n ∈ I + +ClosedΣ (n , xⁿ∈I) (m , yᵐ∈I) = (n +ℕ m) + , subst-∈ I (sym (BinomialThm (n +ℕ m) _ _)) ∑Binomial∈I + where + binomialCoeff∈I : ∀ i → ((n +ℕ m) choose toℕ i) · x ^ toℕ i · y ^ (n +ℕ m ∸ toℕ i) ∈ I + binomialCoeff∈I i with ≤-+-split n m (toℕ i) (pred-≤-pred (toℕ>= normalise + just (varInfos , equation) ← returnTC (getVarsAndEquation hole′) + where + nothing + → typeError (strErr "Something went wrong when getting the variable names in " + ∷ termErr hole′ ∷ []) + adjustedCring ← returnTC (adjustDeBruijnIndex (length varInfos) cring) + just (lhs , rhs) ← returnTC (toAlgebraExpression adjustedCring (getArgs equation)) + where + nothing + → typeError( + strErr "Error while trying to buils ASTs for the equation " ∷ + termErr equation ∷ []) + let solution = constructSolution (length varInfos) varInfos adjustedCring lhs rhs + unify hole solution + +macro + solve : Term → Term → TC _ + solve = solve-macro + +fromℤ : (R : CommRing ℓ) → ℤ → fst R +fromℤ = scalar diff --git a/Cubical/Algebra/RingSolver/RingExpression.agda b/Cubical/Algebra/RingSolver/RingExpression.agda new file mode 100644 index 000000000..38f604ab7 --- /dev/null +++ b/Cubical/Algebra/RingSolver/RingExpression.agda @@ -0,0 +1,39 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.RingSolver.RingExpression where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.FinData +open import Cubical.Data.Nat using (ℕ) +open import Cubical.Data.Nat.Order using (zero-≤) +open import Cubical.Data.Vec.Base +open import Cubical.Algebra.RingSolver.AlmostRing +open import Cubical.Algebra.RingSolver.RawRing renaming (⟨_⟩ to ⟨_⟩ᵣ) + +private + variable + ℓ : Level + +infixl 6 _⊕_ +infixl 7 _⊗_ + +-- Expression in a ring on A with n variables +data Expr {ℓ} (A : Type ℓ) (n : ℕ) : Type ℓ where + K : A → Expr A n + ∣ : Fin n → Expr A n + _⊕_ : Expr A n → Expr A n → Expr A n + _⊗_ : Expr A n → Expr A n → Expr A n +-- _⊛_ : Expr A n → ℕ → Expr A n -- exponentiation + ⊝_ : Expr A n → Expr A n + +module Eval (R : RawRing ℓ) where + open import Cubical.Data.Vec + open RawRing R + + ⟦_⟧ : ∀ {n} → Expr ⟨ R ⟩ᵣ n → Vec ⟨ R ⟩ᵣ n → ⟨ R ⟩ᵣ + ⟦ K r ⟧ v = r + ⟦ ∣ k ⟧ v = lookup k v + ⟦ x ⊕ y ⟧ v = ⟦ x ⟧ v + ⟦ y ⟧ v + ⟦ x ⊗ y ⟧ v = ⟦ x ⟧ v · ⟦ y ⟧ v +-- ⟦ x ⊛ l ⟧ v = ⟦ x ⟧ v ^ l + ⟦ ⊝ x ⟧ v = - ⟦ x ⟧ v diff --git a/Cubical/Algebra/RingSolver/Solver.agda b/Cubical/Algebra/RingSolver/Solver.agda new file mode 100644 index 000000000..0b578fff5 --- /dev/null +++ b/Cubical/Algebra/RingSolver/Solver.agda @@ -0,0 +1,150 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.RingSolver.Solver where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.FinData +open import Cubical.Data.Nat using (ℕ) +open import Cubical.Data.Nat.Order using (zero-≤) +open import Cubical.Data.Vec.Base +open import Cubical.Algebra.RingSolver.AlmostRing +open import Cubical.Algebra.RingSolver.RawRing renaming (⟨_⟩ to ⟨_⟩ᵣ) +open import Cubical.Algebra.RingSolver.RingExpression +open import Cubical.Algebra.RingSolver.HornerForms +open import Cubical.Algebra.RingSolver.EvaluationHomomorphism + +private + variable + ℓ : Level + +module EqualityToNormalform (R : AlmostRing ℓ) where + νR = AlmostRing→RawRing R + open AlmostRing R + open Theory R + open Eval νR + open IteratedHornerOperations νR + open HomomorphismProperties R + + normalize : (n : ℕ) → Expr ⟨ R ⟩ n → IteratedHornerForms νR n + normalize n (K r) = Constant n νR r + normalize n (∣ k) = Variable n νR k + normalize n (x ⊕ y) = + (normalize n x) +ₕ (normalize n y) + normalize n (x ⊗ y) = + (normalize n x) ·ₕ (normalize n y) + normalize n (⊝ x) = -ₕ (normalize n x) + + isEqualToNormalform : + (n : ℕ) + (e : Expr ⟨ R ⟩ n) (xs : Vec ⟨ R ⟩ n) + → eval n (normalize n e) xs ≡ ⟦ e ⟧ xs + isEqualToNormalform ℕ.zero (K r) [] = refl + isEqualToNormalform (ℕ.suc n) (K r) (x ∷ xs) = + eval (ℕ.suc n) (Constant (ℕ.suc n) νR r) (x ∷ xs) ≡⟨ refl ⟩ + eval (ℕ.suc n) (0ₕ ·X+ Constant n νR r) (x ∷ xs) ≡⟨ refl ⟩ + eval (ℕ.suc n) 0ₕ (x ∷ xs) · x + eval n (Constant n νR r) xs + ≡⟨ cong (λ u → u · x + eval n (Constant n νR r) xs) (eval0H _ (x ∷ xs)) ⟩ + 0r · x + eval n (Constant n νR r) xs + ≡⟨ cong (λ u → u + eval n (Constant n νR r) xs) (0LeftAnnihilates _) ⟩ + 0r + eval n (Constant n νR r) xs ≡⟨ +Lid _ ⟩ + eval n (Constant n νR r) xs + ≡⟨ isEqualToNormalform n (K r) xs ⟩ + r ∎ + + isEqualToNormalform (ℕ.suc n) (∣ zero) (x ∷ xs) = + eval (ℕ.suc n) (1ₕ ·X+ 0ₕ) (x ∷ xs) ≡⟨ refl ⟩ + eval (ℕ.suc n) 1ₕ (x ∷ xs) · x + eval n 0ₕ xs ≡⟨ cong (λ u → u · x + eval n 0ₕ xs) + (eval1ₕ _ (x ∷ xs)) ⟩ + 1r · x + eval n 0ₕ xs ≡⟨ cong (λ u → 1r · x + u ) (eval0H _ xs) ⟩ + 1r · x + 0r ≡⟨ +Rid _ ⟩ + 1r · x ≡⟨ ·Lid _ ⟩ + x ∎ + isEqualToNormalform (ℕ.suc n) (∣ (suc k)) (x ∷ xs) = + eval (ℕ.suc n) (0ₕ ·X+ Variable n νR k) (x ∷ xs) ≡⟨ refl ⟩ + eval (ℕ.suc n) 0ₕ (x ∷ xs) · x + eval n (Variable n νR k) xs + ≡⟨ cong (λ u → u · x + eval n (Variable n νR k) xs) (eval0H _ (x ∷ xs)) ⟩ + 0r · x + eval n (Variable n νR k) xs + ≡⟨ cong (λ u → u + eval n (Variable n νR k) xs) (0LeftAnnihilates _) ⟩ + 0r + eval n (Variable n νR k) xs ≡⟨ +Lid _ ⟩ + eval n (Variable n νR k) xs + ≡⟨ isEqualToNormalform n (∣ k) xs ⟩ + ⟦ ∣ (suc k) ⟧ (x ∷ xs) ∎ + + isEqualToNormalform ℕ.zero (⊝ e) [] = + eval ℕ.zero (-ₕ (normalize ℕ.zero e)) [] ≡⟨ -evalDist ℕ.zero + (normalize ℕ.zero e) + [] ⟩ + - eval ℕ.zero (normalize ℕ.zero e) [] ≡⟨ cong -_ + (isEqualToNormalform + ℕ.zero e [] ) ⟩ + - ⟦ e ⟧ [] ∎ + isEqualToNormalform (ℕ.suc n) (⊝ e) (x ∷ xs) = + eval (ℕ.suc n) (-ₕ (normalize (ℕ.suc n) e)) (x ∷ xs) ≡⟨ -evalDist (ℕ.suc n) + (normalize + (ℕ.suc n) e) + (x ∷ xs) ⟩ + - eval (ℕ.suc n) (normalize (ℕ.suc n) e) (x ∷ xs) ≡⟨ cong -_ + (isEqualToNormalform + (ℕ.suc n) e (x ∷ xs) ) ⟩ + - ⟦ e ⟧ (x ∷ xs) ∎ + + isEqualToNormalform ℕ.zero (e ⊕ e₁) [] = + eval ℕ.zero (normalize ℕ.zero e +ₕ normalize ℕ.zero e₁) [] + ≡⟨ +Homeval ℕ.zero (normalize ℕ.zero e) _ [] ⟩ + eval ℕ.zero (normalize ℕ.zero e) [] + + eval ℕ.zero (normalize ℕ.zero e₁) [] + ≡⟨ cong (λ u → u + eval ℕ.zero (normalize ℕ.zero e₁) []) + (isEqualToNormalform ℕ.zero e []) ⟩ + ⟦ e ⟧ [] + + eval ℕ.zero (normalize ℕ.zero e₁) [] + ≡⟨ cong (λ u → ⟦ e ⟧ [] + u) (isEqualToNormalform ℕ.zero e₁ []) ⟩ + ⟦ e ⟧ [] + ⟦ e₁ ⟧ [] ∎ + isEqualToNormalform (ℕ.suc n) (e ⊕ e₁) (x ∷ xs) = + eval (ℕ.suc n) (normalize (ℕ.suc n) e + +ₕ normalize (ℕ.suc n) e₁) (x ∷ xs) + ≡⟨ +Homeval (ℕ.suc n) (normalize (ℕ.suc n) e) _ (x ∷ xs) ⟩ + eval (ℕ.suc n) (normalize (ℕ.suc n) e) (x ∷ xs) + + eval (ℕ.suc n) (normalize (ℕ.suc n) e₁) (x ∷ xs) + ≡⟨ cong (λ u → u + eval (ℕ.suc n) (normalize (ℕ.suc n) e₁) (x ∷ xs)) + (isEqualToNormalform (ℕ.suc n) e (x ∷ xs)) ⟩ + ⟦ e ⟧ (x ∷ xs) + + eval (ℕ.suc n) (normalize (ℕ.suc n) e₁) (x ∷ xs) + ≡⟨ cong (λ u → ⟦ e ⟧ (x ∷ xs) + u) + (isEqualToNormalform (ℕ.suc n) e₁ (x ∷ xs)) ⟩ + ⟦ e ⟧ (x ∷ xs) + ⟦ e₁ ⟧ (x ∷ xs) ∎ + + isEqualToNormalform ℕ.zero (e ⊗ e₁) [] = + eval ℕ.zero (normalize ℕ.zero e ·ₕ normalize ℕ.zero e₁) [] + ≡⟨ ·Homeval ℕ.zero (normalize ℕ.zero e) _ [] ⟩ + eval ℕ.zero (normalize ℕ.zero e) [] + · eval ℕ.zero (normalize ℕ.zero e₁) [] + ≡⟨ cong (λ u → u · eval ℕ.zero (normalize ℕ.zero e₁) []) + (isEqualToNormalform ℕ.zero e []) ⟩ + ⟦ e ⟧ [] + · eval ℕ.zero (normalize ℕ.zero e₁) [] + ≡⟨ cong (λ u → ⟦ e ⟧ [] · u) (isEqualToNormalform ℕ.zero e₁ []) ⟩ + ⟦ e ⟧ [] · ⟦ e₁ ⟧ [] ∎ + + isEqualToNormalform (ℕ.suc n) (e ⊗ e₁) (x ∷ xs) = + eval (ℕ.suc n) (normalize (ℕ.suc n) e + ·ₕ normalize (ℕ.suc n) e₁) (x ∷ xs) + ≡⟨ ·Homeval (ℕ.suc n) (normalize (ℕ.suc n) e) _ (x ∷ xs) ⟩ + eval (ℕ.suc n) (normalize (ℕ.suc n) e) (x ∷ xs) + · eval (ℕ.suc n) (normalize (ℕ.suc n) e₁) (x ∷ xs) + ≡⟨ cong (λ u → u · eval (ℕ.suc n) (normalize (ℕ.suc n) e₁) (x ∷ xs)) + (isEqualToNormalform (ℕ.suc n) e (x ∷ xs)) ⟩ + ⟦ e ⟧ (x ∷ xs) + · eval (ℕ.suc n) (normalize (ℕ.suc n) e₁) (x ∷ xs) + ≡⟨ cong (λ u → ⟦ e ⟧ (x ∷ xs) · u) + (isEqualToNormalform (ℕ.suc n) e₁ (x ∷ xs)) ⟩ + ⟦ e ⟧ (x ∷ xs) · ⟦ e₁ ⟧ (x ∷ xs) ∎ + + solve : + {n : ℕ} (e₁ e₂ : Expr ⟨ R ⟩ n) (xs : Vec ⟨ R ⟩ n) + (p : eval n (normalize n e₁) xs ≡ eval n (normalize n e₂) xs) + → ⟦ e₁ ⟧ xs ≡ ⟦ e₂ ⟧ xs + solve e₁ e₂ xs p = + ⟦ e₁ ⟧ xs ≡⟨ sym (isEqualToNormalform _ e₁ xs) ⟩ + eval _ (normalize _ e₁) xs ≡⟨ p ⟩ + eval _ (normalize _ e₂) xs ≡⟨ isEqualToNormalform _ e₂ xs ⟩ + ⟦ e₂ ⟧ xs ∎ diff --git a/Cubical/Algebra/Semigroup.agda b/Cubical/Algebra/Semigroup.agda new file mode 100644 index 000000000..53e0c0385 --- /dev/null +++ b/Cubical/Algebra/Semigroup.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Semigroup where + +open import Cubical.Algebra.Semigroup.Base public diff --git a/Cubical/Algebra/Semigroup/Base.agda b/Cubical/Algebra/Semigroup/Base.agda new file mode 100644 index 000000000..4388d929f --- /dev/null +++ b/Cubical/Algebra/Semigroup/Base.agda @@ -0,0 +1,100 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Semigroup.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Reflection.RecordEquiv +open import Cubical.Reflection.StrictEquiv + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open Iso + + +private + variable + ℓ : Level + +-- Semigroups as a record, inspired by the Agda standard library: +-- +-- https://github.com/agda/agda-stdlib/blob/master/src/Algebra/Bundles.agda#L48 +-- https://github.com/agda/agda-stdlib/blob/master/src/Algebra/Structures.agda#L50 +-- +-- Note that as we are using Path for all equations the IsMagma record +-- would only contain isSet A if we had it. +record IsSemigroup {A : Type ℓ} (_·_ : A → A → A) : Type ℓ where + no-eta-equality + constructor issemigroup + + field + is-set : isSet A + assoc : (x y z : A) → x · (y · z) ≡ (x · y) · z + +unquoteDecl IsSemigroupIsoΣ = declareRecordIsoΣ IsSemigroupIsoΣ (quote IsSemigroup) + +record SemigroupStr (A : Type ℓ) : Type ℓ where + + constructor semigroupstr + + field + _·_ : A → A → A + isSemigroup : IsSemigroup _·_ + + infixl 7 _·_ + + open IsSemigroup isSemigroup public + +Semigroup : ∀ ℓ → Type (ℓ-suc ℓ) +Semigroup ℓ = TypeWithStr ℓ SemigroupStr + +semigroup : (A : Type ℓ) (_·_ : A → A → A) (h : IsSemigroup _·_) → Semigroup ℓ +semigroup A _·_ h = A , semigroupstr _·_ h + +record IsSemigroupEquiv {A : Type ℓ} {B : Type ℓ} + (M : SemigroupStr A) (e : A ≃ B) (N : SemigroupStr B) + : Type ℓ + where + + -- Shorter qualified names + private + module M = SemigroupStr M + module N = SemigroupStr N + + field + isHom : (x y : A) → equivFun e (x M.· y) ≡ equivFun e x N.· equivFun e y + +open SemigroupStr +open IsSemigroup +open IsSemigroupEquiv + +SemigroupEquiv : (M N : Semigroup ℓ) → Type ℓ +SemigroupEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsSemigroupEquiv (M .snd) e (N .snd) + +isPropIsSemigroup : {A : Type ℓ} (_·_ : A → A → A) → isProp (IsSemigroup _·_) +isPropIsSemigroup _·_ = + isOfHLevelRetractFromIso 1 IsSemigroupIsoΣ + (isPropΣ + isPropIsSet + (λ isSetA → isPropΠ3 λ _ _ _ → isSetA _ _)) + +𝒮ᴰ-Semigroup : DUARel (𝒮-Univ ℓ) SemigroupStr ℓ +𝒮ᴰ-Semigroup = + 𝒮ᴰ-Record (𝒮-Univ _) IsSemigroupEquiv + (fields: + data[ _·_ ∣ autoDUARel _ _ ∣ isHom ] + prop[ isSemigroup ∣ (λ _ _ → isPropIsSemigroup _) ]) + +SemigroupPath : (M N : Semigroup ℓ) → SemigroupEquiv M N ≃ (M ≡ N) +SemigroupPath = ∫ 𝒮ᴰ-Semigroup .UARel.ua diff --git a/Cubical/Algebra/SymmetricGroup.agda b/Cubical/Algebra/SymmetricGroup.agda new file mode 100644 index 000000000..aa4395f7d --- /dev/null +++ b/Cubical/Algebra/SymmetricGroup.agda @@ -0,0 +1,26 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.SymmetricGroup where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.Data.Nat using (ℕ ; suc ; zero) +open import Cubical.Data.Fin using (Fin ; isSetFin) +open import Cubical.Data.Empty +open import Cubical.Relation.Nullary using (¬_) + +open import Cubical.Algebra.Group + +private + variable + ℓ : Level + +Symmetric-Group : (X : Type ℓ) → isSet X → Group ℓ +Symmetric-Group X isSetX = makeGroup (idEquiv X) compEquiv invEquiv (isOfHLevel≃ 2 isSetX isSetX) + compEquiv-assoc compEquivEquivId compEquivIdEquiv invEquiv-is-rinv invEquiv-is-linv + +-- Finite symmetrics groups + +Sym : ℕ → Group _ +Sym n = Symmetric-Group (Fin n) isSetFin diff --git a/Cubical/Algebra/ZariskiLattice/BasicOpens.agda b/Cubical/Algebra/ZariskiLattice/BasicOpens.agda new file mode 100644 index 000000000..41e9ecf5a --- /dev/null +++ b/Cubical/Algebra/ZariskiLattice/BasicOpens.agda @@ -0,0 +1,236 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.ZariskiLattice.BasicOpens where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Structure +open import Cubical.Functions.FunExtEquiv + +import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Bool +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ + ; +-comm to +ℕ-comm ; +-assoc to +ℕ-assoc + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm) +open import Cubical.Data.Sigma.Base +open import Cubical.Data.Sigma.Properties +open import Cubical.Data.FinData +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Algebra +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Localisation.Base +open import Cubical.Algebra.CommRing.Localisation.UniversalProperty +open import Cubical.Algebra.CommRing.Localisation.InvertingElements +open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebra.Properties +open import Cubical.Algebra.CommAlgebra.Localisation +open import Cubical.Algebra.RingSolver.ReflectionSolving + +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.HITs.PropositionalTruncation as PT + +open Iso +open BinaryRelation +open isEquivRel + +private + variable + ℓ ℓ' : Level + + +module Presheaf (A' : CommRing ℓ) where + open CommRingStr (snd A') renaming (_·_ to _·r_ ; ·-comm to ·r-comm ; ·Assoc to ·rAssoc + ; ·Lid to ·rLid ; ·Rid to ·rRid) + open Exponentiation A' + open CommRingTheory A' + open isMultClosedSubset + open CommAlgebraStr ⦃...⦄ + private + A = fst A' + + A[1/_] : A → CommAlgebra A' ℓ + A[1/ x ] = AlgLoc.S⁻¹RAsCommAlg A' ([_ⁿ|n≥0] A' x) (powersFormMultClosedSubset _ _) + + A[1/_]ˣ : (x : A) → ℙ (fst A[1/ x ]) + A[1/ x ]ˣ = (CommAlgebra→CommRing A[1/ x ]) ˣ + + + _≼_ : A → A → Type ℓ + x ≼ y = ∃[ n ∈ ℕ ] Σ[ z ∈ A ] x ^ n ≡ z ·r y -- rad(x) ⊆ rad(y) + +-- ≼ is a pre-order: + Refl≼ : isRefl _≼_ + Refl≼ x = PT.∣ 1 , 1r , ·r-comm _ _ ∣ + + Trans≼ : isTrans _≼_ + Trans≼ x y z = map2 Trans≼Σ + where + Trans≼Σ : Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a ·r y + → Σ[ n ∈ ℕ ] Σ[ a ∈ A ] y ^ n ≡ a ·r z + → Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a ·r z + Trans≼Σ (n , a , p) (m , b , q) = n ·ℕ m , (a ^ m ·r b) , path + where + path : x ^ (n ·ℕ m) ≡ a ^ m ·r b ·r z + path = x ^ (n ·ℕ m) ≡⟨ ^-rdist-·ℕ x n m ⟩ + (x ^ n) ^ m ≡⟨ cong (_^ m) p ⟩ + (a ·r y) ^ m ≡⟨ ^-ldist-· a y m ⟩ + a ^ m ·r y ^ m ≡⟨ cong (a ^ m ·r_) q ⟩ + a ^ m ·r (b ·r z) ≡⟨ ·rAssoc _ _ _ ⟩ + a ^ m ·r b ·r z ∎ + + + R : A → A → Type ℓ + R x y = x ≼ y × y ≼ x -- rad(x) ≡ rad(y) + + RequivRel : isEquivRel R + RequivRel .reflexive x = Refl≼ x , Refl≼ x + RequivRel .symmetric _ _ Rxy = (Rxy .snd) , (Rxy .fst) + RequivRel .transitive _ _ _ Rxy Ryz = Trans≼ _ _ _ (Rxy .fst) (Ryz .fst) + , Trans≼ _ _ _ (Ryz .snd) (Rxy .snd) + + RpropValued : isPropValued R + RpropValued x y = isProp× isPropPropTrunc isPropPropTrunc + + powerIs≽ : (x a : A) → x ∈ ([_ⁿ|n≥0] A' a) → a ≼ x + powerIs≽ x a = map powerIs≽Σ + where + powerIs≽Σ : Σ[ n ∈ ℕ ] (x ≡ a ^ n) → Σ[ n ∈ ℕ ] Σ[ z ∈ A ] (a ^ n ≡ z ·r x) + powerIs≽Σ (n , p) = n , 1r , sym p ∙ sym (·rLid _) + + module ≼ToLoc (x y : A) where + private + instance + _ = snd A[1/ x ] + + lemma : x ≼ y → y ⋆ 1a ∈ A[1/ x ]ˣ -- y/1 ∈ A[1/x]ˣ + lemma = PT.rec (A[1/ x ]ˣ (y ⋆ 1a) .snd) lemmaΣ + where + path1 : (y z : A) → 1r ·r (y ·r 1r ·r z) ·r 1r ≡ z ·r y + path1 = solve A' + path2 : (xn : A) → xn ≡ 1r ·r 1r ·r (1r ·r 1r ·r xn) + path2 = solve A' + + lemmaΣ : Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a ·r y → y ⋆ 1a ∈ A[1/ x ]ˣ + lemmaΣ (n , z , p) = [ z , (x ^ n) , PT.∣ n , refl ∣ ] -- xⁿ≡zy → y⁻¹ ≡ z/xⁿ + , eq/ _ _ ((1r , powersFormMultClosedSubset _ _ .containsOne) + , (path1 _ _ ∙∙ sym p ∙∙ path2 _)) + + module ≼PowerToLoc (x y : A) (x≼y : x ≼ y) where + private + [yⁿ|n≥0] = [_ⁿ|n≥0] A' y + instance + _ = snd A[1/ x ] + lemma : ∀ (s : A) → s ∈ [yⁿ|n≥0] → s ⋆ 1a ∈ A[1/ x ]ˣ + lemma _ s∈[yⁿ|n≥0] = ≼ToLoc.lemma _ _ (Trans≼ _ y _ x≼y (powerIs≽ _ _ s∈[yⁿ|n≥0])) + + + + 𝓞ᴰ : A / R → CommAlgebra A' ℓ + 𝓞ᴰ = rec→Gpd.fun isGroupoidCommAlgebra (λ a → A[1/ a ]) RCoh LocPathProp + where + RCoh : ∀ a b → R a b → A[1/ a ] ≡ A[1/ b ] + RCoh a b (a≼b , b≼a) = fst (isContrS₁⁻¹R≡S₂⁻¹R (≼PowerToLoc.lemma _ _ b≼a) + (≼PowerToLoc.lemma _ _ a≼b)) + where + open AlgLocTwoSubsets A' ([_ⁿ|n≥0] A' a) (powersFormMultClosedSubset _ _) + ([_ⁿ|n≥0] A' b) (powersFormMultClosedSubset _ _) + + LocPathProp : ∀ a b → isProp (A[1/ a ] ≡ A[1/ b ]) + LocPathProp a b = isPropS₁⁻¹R≡S₂⁻¹R + where + open AlgLocTwoSubsets A' ([_ⁿ|n≥0] A' a) (powersFormMultClosedSubset _ _) + ([_ⁿ|n≥0] A' b) (powersFormMultClosedSubset _ _) + + + -- The quotient A/R corresponds to the basic opens of the Zariski topology. + -- Multiplication lifts to the quotient and corresponds to intersection + -- of basic opens, i.e. we get a meet-semilattice with: + _∧/_ : A / R → A / R → A / R + _∧/_ = setQuotSymmBinOp (RequivRel .reflexive) (RequivRel .transitive) _·r_ ·r-comm ·r-lcoh + where + ·r-lcoh-≼ : (x y z : A) → x ≼ y → (x ·r z) ≼ (y ·r z) + ·r-lcoh-≼ x y z = map ·r-lcoh-≼Σ + where + path : (x z a y zn : A) → x ·r z ·r (a ·r y ·r zn) ≡ x ·r zn ·r a ·r (y ·r z) + path = solve A' + + ·r-lcoh-≼Σ : Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a ·r y + → Σ[ n ∈ ℕ ] Σ[ a ∈ A ] (x ·r z) ^ n ≡ a ·r (y ·r z) + ·r-lcoh-≼Σ (n , a , p) = suc n , (x ·r z ^ n ·r a) , (cong (x ·r z ·r_) (^-ldist-· _ _ _) + ∙∙ cong (λ v → x ·r z ·r (v ·r z ^ n)) p + ∙∙ path _ _ _ _ _) + + ·r-lcoh : (x y z : A) → R x y → R (x ·r z) (y ·r z) + ·r-lcoh x y z Rxy = ·r-lcoh-≼ x y z (Rxy .fst) , ·r-lcoh-≼ y x z (Rxy .snd) + + -- The induced partial order + _≼/_ : A / R → A / R → Type ℓ + x ≼/ y = x ≡ (x ∧/ y) + + -- coincides with our ≼ + ≼/CoincidesWith≼ : ∀ (x y : A) → [ x ] ≼/ [ y ] ≡ x ≼ y + ≼/CoincidesWith≼ x y = [ x ] ≼/ [ y ] -- ≡⟨ refl ⟩ [ x ] ≡ [ x ·r y ] + ≡⟨ isoToPath (isEquivRel→effectiveIso RpropValued RequivRel _ _) ⟩ + R x (x ·r y) + ≡⟨ hPropExt (RpropValued _ _) isPropPropTrunc ·To≼ ≼To· ⟩ + x ≼ y ∎ + where + x≼xy→x≼yΣ : Σ[ n ∈ ℕ ] Σ[ z ∈ A ] x ^ n ≡ z ·r (x ·r y) + → Σ[ n ∈ ℕ ] Σ[ z ∈ A ] x ^ n ≡ z ·r y + x≼xy→x≼yΣ (n , z , p) = n , (z ·r x) , p ∙ ·rAssoc _ _ _ + + ·To≼ : R x (x ·r y) → x ≼ y + ·To≼ (x≼xy , _) = PT.map x≼xy→x≼yΣ x≼xy + + x≼y→x≼xyΣ : Σ[ n ∈ ℕ ] Σ[ z ∈ A ] x ^ n ≡ z ·r y + → Σ[ n ∈ ℕ ] Σ[ z ∈ A ] x ^ n ≡ z ·r (x ·r y) + x≼y→x≼xyΣ (n , z , p) = suc n , z , cong (x ·r_) p ∙ ·-commAssocl _ _ _ + + ≼To· : x ≼ y → R x ( x ·r y) + ≼To· x≼y = PT.map x≼y→x≼xyΣ x≼y , PT.∣ 1 , y , ·rRid _ ∙ ·r-comm _ _ ∣ + + Refl≼/ : isRefl _≼/_ + Refl≼/ = SQ.elimProp (λ _ → squash/ _ _) λ _ → transport⁻ (≼/CoincidesWith≼ _ _) (Refl≼ _) + + Trans≼/ : isTrans _≼/_ + Trans≼/ = SQ.elimProp3 (λ _ _ _ → isPropΠ2 (λ _ _ → squash/ _ _)) + λ _ _ _ [a]≼/[b] [b]≼/[c] → transport⁻ (≼/CoincidesWith≼ _ _) + (Trans≼ _ _ _ (transport (≼/CoincidesWith≼ _ _) [a]≼/[b]) + (transport (≼/CoincidesWith≼ _ _) [b]≼/[c])) + + -- The restrictions: + ρᴰᴬ : (a b : A) → a ≼ b → isContr (CommAlgebraHom A[1/ b ] A[1/ a ]) + ρᴰᴬ _ b a≼b = A[1/b]HasUniversalProp _ (≼PowerToLoc.lemma _ _ a≼b) + where + open AlgLoc A' ([_ⁿ|n≥0] A' b) (powersFormMultClosedSubset _ _) + renaming (S⁻¹RHasAlgUniversalProp to A[1/b]HasUniversalProp) + + ρᴰᴬId : ∀ (a : A) (r : a ≼ a) → ρᴰᴬ a a r .fst ≡ idAlgHom + ρᴰᴬId a r = ρᴰᴬ a a r .snd _ + + ρᴰᴬComp : ∀ (a b c : A) (l : a ≼ b) (m : b ≼ c) + → ρᴰᴬ a c (Trans≼ _ _ _ l m) .fst ≡ ρᴰᴬ a b l .fst ∘a ρᴰᴬ b c m .fst + ρᴰᴬComp a _ c l m = ρᴰᴬ a c (Trans≼ _ _ _ l m) .snd _ + + + ρᴰ : (x y : A / R) → x ≼/ y → CommAlgebraHom (𝓞ᴰ y) (𝓞ᴰ x) + ρᴰ = elimContr2 λ _ _ → isOfHLevelΠ 0 + λ [a]≼/[b] → ρᴰᴬ _ _ (transport (≼/CoincidesWith≼ _ _) [a]≼/[b]) + + ρᴰId : ∀ (x : A / R) (r : x ≼/ x) → ρᴰ x x r ≡ idAlgHom + ρᴰId = SQ.elimProp (λ _ → isPropΠ (λ _ → isSetAlgebraHom _ _ _ _)) + λ a r → ρᴰᴬId a (transport (≼/CoincidesWith≼ _ _) r) + + ρᴰComp : ∀ (x y z : A / R) (l : x ≼/ y) (m : y ≼/ z) + → ρᴰ x z (Trans≼/ _ _ _ l m) ≡ ρᴰ x y l ∘a ρᴰ y z m + ρᴰComp = SQ.elimProp3 (λ _ _ _ → isPropΠ2 (λ _ _ → isSetAlgebraHom _ _ _ _)) + λ a b c _ _ → sym (ρᴰᴬ a c _ .snd _) ∙ ρᴰᴬComp a b c _ _ diff --git a/Cubical/AndersFile.agda b/Cubical/AndersFile.agda deleted file mode 100644 index 59285462e..000000000 --- a/Cubical/AndersFile.agda +++ /dev/null @@ -1,41 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.AndersFile where - - - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Data.Nat -open import Cubical.Data.HomotopyGroup - -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - - -FunLR : (n : ℕ) {f : A → B} → - ((typ ((Ω^ (suc n)) ((A → B) , f))) → - (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) -LR-id : (n : ℕ) {f : A → B} → FunLR n {f} refl ≡ (λ x → refl) -FunLR n {f} = {!!} -LR-id n {f} = {!!} - - -FunRL : (n : ℕ) {f : A → B} → - ((x : A) → (typ ((Ω^ (suc n)) (B , f x)))) → - ((typ ((Ω^ (suc n)) ((A → B) , f)))) -RL-id : (n : ℕ) {f : A → B} → - FunRL n (λ x → pt (((Ω^ (suc n)) (B , f x)))) ≡ pt ((Ω^ (suc n)) ((A → B) , f)) -FunRL n {f} = {!!} -RL-id n {f} = {!!} - - -wantToShow : (n : ℕ) (f : A → B) → - ((x : A) → (typ ((Ω^ n) (B , f x)))) ≡ (typ ((Ω^ n) ((A → B) , f))) -wantToShow n f = {!!} - - diff --git a/Cubical/Categories/Adjoint.agda b/Cubical/Categories/Adjoint.agda new file mode 100644 index 000000000..70ee18279 --- /dev/null +++ b/Cubical/Categories/Adjoint.agda @@ -0,0 +1,319 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Adjoint where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Sigma +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation +open import Cubical.Foundations.Isomorphism + +open Functor + +open Iso +open Precategory + +{- +============================================== + Overview +============================================== + +This module contains two definitions for adjoint +functors, and functions witnessing their +logical (and maybe eventually actual?) +equivalence. +-} + +private + variable + ℓC ℓC' ℓD ℓD' : Level + +{- +============================================== + Adjoint definitions +============================================== + +We provide two alternative definitions for +adjoint functors: the unit-counit +definition, followed by the natural bijection +definition. +-} + +module UnitCounit where + + -- Adjoint def 1: unit-counit + record _⊣_ {C : Precategory ℓC ℓC'} {D : Precategory ℓD ℓD'} (F : Functor C D) (G : Functor D C) + : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where + field + -- unit + η : 𝟙⟨ C ⟩ ⇒ (funcComp G F) + -- counit + ε : (funcComp F G) ⇒ 𝟙⟨ D ⟩ + -- triangle identities + Δ₁ : PathP (λ i → NatTrans (F-lUnit {F = F} i) (F-rUnit {F = F} i)) + (seqTransP F-assoc (F ∘ʳ η) (ε ∘ˡ F)) + (1[ F ]) + Δ₂ : PathP (λ i → NatTrans (F-rUnit {F = G} i) (F-lUnit {F = G} i)) + (seqTransP (sym F-assoc) (η ∘ˡ G) (G ∘ʳ ε)) + (1[ G ]) + + {- + Helper function for building unit-counit adjunctions between categories, + using that equality of natural transformations in a category is equality on objects + -} + + module _ {ℓC ℓC' ℓD ℓD'} + {C : Precategory ℓC ℓC'} {D : Precategory ℓD ℓD'} {F : Functor C D} {G : Functor D C} + ⦃ isCatC : isCategory C ⦄ ⦃ isCatD : isCategory D ⦄ + (η : 𝟙⟨ C ⟩ ⇒ (funcComp G F)) + (ε : (funcComp F G) ⇒ 𝟙⟨ D ⟩) + (Δ₁ : ∀ c → F ⟪ η ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ ε ⟦ F ⟅ c ⟆ ⟧ ≡ D .id (F ⟅ c ⟆)) + (Δ₂ : ∀ d → η ⟦ G ⟅ d ⟆ ⟧ ⋆⟨ C ⟩ G ⟪ ε ⟦ d ⟧ ⟫ ≡ C .id (G ⟅ d ⟆)) + where + + make⊣ : F ⊣ G + make⊣ ._⊣_.η = η + make⊣ ._⊣_.ε = ε + make⊣ ._⊣_.Δ₁ = + makeNatTransPathP F-lUnit F-rUnit + (funExt λ c → cong (D ._⋆_ (F ⟪ η ⟦ c ⟧ ⟫)) (transportRefl _) ∙ Δ₁ c) + make⊣ ._⊣_.Δ₂ = + makeNatTransPathP F-rUnit F-lUnit + (funExt λ d → cong (C ._⋆_ (η ⟦ G ⟅ d ⟆ ⟧)) (transportRefl _) ∙ Δ₂ d) + +module NaturalBijection where + -- Adjoint def 2: natural bijection + record _⊣_ {C : Precategory ℓC ℓC'} {D : Precategory ℓD ℓD'} (F : Functor C D) (G : Functor D C) : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where + field + adjIso : ∀ {c d} → Iso (D [ F ⟅ c ⟆ , d ]) (C [ c , G ⟅ d ⟆ ]) + + infix 40 _♭ + infix 40 _♯ + _♭ : ∀ {c d} → (D [ F ⟅ c ⟆ , d ]) → (C [ c , G ⟅ d ⟆ ]) + (_♭) {_} {_} = adjIso .fun + + _♯ : ∀ {c d} → (C [ c , G ⟅ d ⟆ ]) → (D [ F ⟅ c ⟆ , d ]) + (_♯) {_} {_} = adjIso .inv + + field + adjNatInD : ∀ {c : C .ob} {d d'} (f : D [ F ⟅ c ⟆ , d ]) (k : D [ d , d' ]) + → (f ⋆⟨ D ⟩ k) ♭ ≡ f ♭ ⋆⟨ C ⟩ G ⟪ k ⟫ + + adjNatInC : ∀ {c' c d} (g : C [ c , G ⟅ d ⟆ ]) (h : C [ c' , c ]) + → (h ⋆⟨ C ⟩ g) ♯ ≡ F ⟪ h ⟫ ⋆⟨ D ⟩ g ♯ + +{- +============================================== + Proofs of equivalence +============================================== + +This first unnamed module provides a function +adj'→adj which takes you from the second +definition to the first. + +The second unnamed module does the reverse. +-} + +module _ {C : Precategory ℓC ℓC'} {D : Precategory ℓD ℓD'} (F : Functor C D) (G : Functor D C) where + open UnitCounit + open NaturalBijection renaming (_⊣_ to _⊣²_) + module _ (adj : F ⊣² G) where + open _⊣²_ adj + open _⊣_ + + -- Naturality condition implies that a commutative square in C + -- appears iff the transpose in D is commutative as well + -- Used in adj'→adj + adjNat' : ∀ {c c' d d'} {f : D [ F ⟅ c ⟆ , d ]} {k : D [ d , d' ]} + → {h : C [ c , c' ]} {g : C [ c' , G ⟅ d' ⟆ ]} + -- commutativity of squares is iff + → ((f ⋆⟨ D ⟩ k ≡ F ⟪ h ⟫ ⋆⟨ D ⟩ g ♯) → (f ♭ ⋆⟨ C ⟩ G ⟪ k ⟫ ≡ h ⋆⟨ C ⟩ g)) + × ((f ♭ ⋆⟨ C ⟩ G ⟪ k ⟫ ≡ h ⋆⟨ C ⟩ g) → (f ⋆⟨ D ⟩ k ≡ F ⟪ h ⟫ ⋆⟨ D ⟩ g ♯)) + adjNat' {c} {c'} {d} {d'} {f} {k} {h} {g} = D→C , C→D + where + D→C : (f ⋆⟨ D ⟩ k ≡ F ⟪ h ⟫ ⋆⟨ D ⟩ g ♯) → (f ♭ ⋆⟨ C ⟩ G ⟪ k ⟫ ≡ h ⋆⟨ C ⟩ g) + D→C eq = f ♭ ⋆⟨ C ⟩ G ⟪ k ⟫ + ≡⟨ sym (adjNatInD _ _) ⟩ + ((f ⋆⟨ D ⟩ k) ♭) + ≡⟨ cong _♭ eq ⟩ + (F ⟪ h ⟫ ⋆⟨ D ⟩ g ♯) ♭ + ≡⟨ sym (cong _♭ (adjNatInC _ _)) ⟩ + (h ⋆⟨ C ⟩ g) ♯ ♭ + ≡⟨ adjIso .rightInv _ ⟩ + h ⋆⟨ C ⟩ g + ∎ + C→D : (f ♭ ⋆⟨ C ⟩ G ⟪ k ⟫ ≡ h ⋆⟨ C ⟩ g) → (f ⋆⟨ D ⟩ k ≡ F ⟪ h ⟫ ⋆⟨ D ⟩ g ♯) + C→D eq = f ⋆⟨ D ⟩ k + ≡⟨ sym (adjIso .leftInv _) ⟩ + (f ⋆⟨ D ⟩ k) ♭ ♯ + ≡⟨ cong _♯ (adjNatInD _ _) ⟩ + (f ♭ ⋆⟨ C ⟩ G ⟪ k ⟫) ♯ + ≡⟨ cong _♯ eq ⟩ + (h ⋆⟨ C ⟩ g) ♯ + ≡⟨ adjNatInC _ _ ⟩ + F ⟪ h ⟫ ⋆⟨ D ⟩ g ♯ + ∎ + + open NatTrans + + -- note : had to make this record syntax because termination checker was complaining + -- due to referencing η and ε from the definitions of Δs + adj'→adj : ⦃ isCatC : isCategory C ⦄ ⦃ isCatD : isCategory D ⦄ → F ⊣ G + adj'→adj = record + { η = η' + ; ε = ε' + ; Δ₁ = Δ₁' + ; Δ₂ = Δ₂' } + + where + -- ETA + + -- trivial commutative diagram between identities in D + commInD : ∀ {x y} (f : C [ x , y ]) → (D .id _) ⋆⟨ D ⟩ F ⟪ f ⟫ ≡ F ⟪ f ⟫ ⋆⟨ D ⟩ (D .id _) + commInD f = (D .⋆IdL _) ∙ sym (D .⋆IdR _) + + sharpen1 : ∀ {x y} (f : C [ x , y ]) → F ⟪ f ⟫ ⋆⟨ D ⟩ (D .id _) ≡ F ⟪ f ⟫ ⋆⟨ D ⟩ (D .id _) ♭ ♯ + sharpen1 f = cong (λ v → F ⟪ f ⟫ ⋆⟨ D ⟩ v) (sym (adjIso .leftInv _)) + + η' : 𝟙⟨ C ⟩ ⇒ G ∘F F + η' .N-ob x = (D .id _) ♭ + η' .N-hom f = sym (fst (adjNat') (commInD f ∙ sharpen1 f)) + + -- EPSILON + + -- trivial commutative diagram between identities in C + commInC : ∀ {x y} (g : D [ x , y ]) → (C .id _) ⋆⟨ C ⟩ G ⟪ g ⟫ ≡ G ⟪ g ⟫ ⋆⟨ C ⟩ (C .id _) + commInC g = (C .⋆IdL _) ∙ sym (C .⋆IdR _) + + sharpen2 : ∀ {x y} (g : D [ x , y ]) → (C .id _ ♯ ♭) ⋆⟨ C ⟩ G ⟪ g ⟫ ≡ (C .id _) ⋆⟨ C ⟩ G ⟪ g ⟫ + sharpen2 g = cong (λ v → v ⋆⟨ C ⟩ G ⟪ g ⟫) (adjIso .rightInv _) + + ε' : F ∘F G ⇒ 𝟙⟨ D ⟩ + ε' .N-ob x = (C .id _) ♯ + ε' .N-hom g = sym (snd adjNat' (sharpen2 g ∙ commInC g)) + + -- DELTA 1 + + expL : ∀ (c) + → (seqTransP F-assoc (F ∘ʳ η') (ε' ∘ˡ F) .N-ob c) + ≡ F ⟪ η' ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ ε' ⟦ F ⟅ c ⟆ ⟧ + expL c = seqTransP F-assoc (F ∘ʳ η') (ε' ∘ˡ F) .N-ob c + ≡⟨ refl ⟩ + seqP {C = D} {p = refl} (F ⟪ η' ⟦ c ⟧ ⟫) (ε' ⟦ F ⟅ c ⟆ ⟧) + ≡⟨ seqP≡seq {C = D} _ _ ⟩ + F ⟪ η' ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ ε' ⟦ F ⟅ c ⟆ ⟧ + ∎ + + body : ∀ (c) + → (idTrans F) ⟦ c ⟧ ≡ (seqTransP F-assoc (F ∘ʳ η') (ε' ∘ˡ F) .N-ob c) + body c = (idTrans F) ⟦ c ⟧ + ≡⟨ refl ⟩ + D .id _ + ≡⟨ sym (D .⋆IdL _) ⟩ + D .id _ ⋆⟨ D ⟩ D .id _ + ≡⟨ snd adjNat' (cong (λ v → (η' ⟦ c ⟧) ⋆⟨ C ⟩ v) (G .F-id)) ⟩ + F ⟪ η' ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ ε' ⟦ F ⟅ c ⟆ ⟧ + ≡⟨ sym (expL c) ⟩ + seqTransP F-assoc (F ∘ʳ η') (ε' ∘ˡ F) .N-ob c + ∎ + + Δ₁' : PathP (λ i → NatTrans (F-lUnit {F = F} i) (F-rUnit {F = F} i)) + (seqTransP F-assoc (F ∘ʳ η') (ε' ∘ˡ F)) + (1[ F ]) + Δ₁' = makeNatTransPathP F-lUnit F-rUnit (sym (funExt body)) + + -- DELTA 2 + + body2 : ∀ (d) + → seqP {C = C} {p = refl} ((η' ∘ˡ G) ⟦ d ⟧) ((G ∘ʳ ε') ⟦ d ⟧) ≡ C .id (G .F-ob d) + body2 d = seqP {C = C} {p = refl} ((η' ∘ˡ G) ⟦ d ⟧) ((G ∘ʳ ε') ⟦ d ⟧) + ≡⟨ seqP≡seq {C = C} _ _ ⟩ + ((η' ∘ˡ G) ⟦ d ⟧) ⋆⟨ C ⟩ ((G ∘ʳ ε') ⟦ d ⟧) + ≡⟨ refl ⟩ + (η' ⟦ G ⟅ d ⟆ ⟧) ⋆⟨ C ⟩ (G ⟪ ε' ⟦ d ⟧ ⟫) + ≡⟨ fst adjNat' (cong (λ v → v ⋆⟨ D ⟩ (ε' ⟦ d ⟧)) (sym (F .F-id))) ⟩ + C .id _ ⋆⟨ C ⟩ C .id _ + ≡⟨ C .⋆IdL _ ⟩ + C .id (G .F-ob d) + ∎ + + Δ₂' : PathP (λ i → NatTrans (F-rUnit {F = G} i) (F-lUnit {F = G} i)) + (seqTransP (sym F-assoc) (η' ∘ˡ G) (G ∘ʳ ε')) + (1[ G ]) + Δ₂' = makeNatTransPathP F-rUnit F-lUnit (funExt body2) + + + module _ (adj : F ⊣ G) where + open _⊣_ adj + open _⊣²_ + open NatTrans + + -- helper functions for working with this Adjoint definition + + δ₁ : ∀ {c} → (F ⟪ η ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ ε ⟦ F ⟅ c ⟆ ⟧) ≡ D .id _ + δ₁ {c} = (F ⟪ η ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ ε ⟦ F ⟅ c ⟆ ⟧) + ≡⟨ sym (seqP≡seq {C = D} _ _) ⟩ + seqP {C = D} {p = refl} (F ⟪ η ⟦ c ⟧ ⟫) (ε ⟦ F ⟅ c ⟆ ⟧) + ≡⟨ (λ j → (Δ₁ j) .N-ob c) ⟩ + D .id _ + ∎ + + δ₂ : ∀ {d} → (η ⟦ G ⟅ d ⟆ ⟧ ⋆⟨ C ⟩ G ⟪ ε ⟦ d ⟧ ⟫) ≡ C .id _ + δ₂ {d} = (η ⟦ G ⟅ d ⟆ ⟧ ⋆⟨ C ⟩ G ⟪ ε ⟦ d ⟧ ⟫) + ≡⟨ sym (seqP≡seq {C = C} _ _) ⟩ + seqP {C = C} {p = refl} (η ⟦ G ⟅ d ⟆ ⟧) (G ⟪ ε ⟦ d ⟧ ⟫) + ≡⟨ (λ j → (Δ₂ j) .N-ob d) ⟩ + C .id _ + ∎ + + + adj→adj' : F ⊣² G + -- ∀ {c d} → Iso (D [ F ⟅ c ⟆ , d ]) (C [ c , G ⟅ d ⟆ ]) + -- takes f to Gf precomposed with the unit + adj→adj' .adjIso {c = c} .fun f = η ⟦ c ⟧ ⋆⟨ C ⟩ G ⟪ f ⟫ + -- takes g to Fg postcomposed with the counit + adj→adj' .adjIso {d = d} .inv g = F ⟪ g ⟫ ⋆⟨ D ⟩ ε ⟦ d ⟧ + -- invertibility follows from the triangle identities + adj→adj' .adjIso {c = c} {d} .rightInv g + = η ⟦ c ⟧ ⋆⟨ C ⟩ G ⟪ F ⟪ g ⟫ ⋆⟨ D ⟩ ε ⟦ d ⟧ ⟫ -- step0 ∙ step1 ∙ step2 ∙ (C .⋆IdR _) + ≡⟨ cong (λ v → η ⟦ c ⟧ ⋆⟨ C ⟩ v) (G .F-seq _ _) ⟩ + η ⟦ c ⟧ ⋆⟨ C ⟩ (G ⟪ F ⟪ g ⟫ ⟫ ⋆⟨ C ⟩ G ⟪ ε ⟦ d ⟧ ⟫) + ≡⟨ sym (C .⋆Assoc _ _ _) ⟩ + η ⟦ c ⟧ ⋆⟨ C ⟩ G ⟪ F ⟪ g ⟫ ⟫ ⋆⟨ C ⟩ G ⟪ ε ⟦ d ⟧ ⟫ + -- apply naturality + ≡⟨ rPrecatWhisker {C = C} _ _ _ natu ⟩ + (g ⋆⟨ C ⟩ η ⟦ G ⟅ d ⟆ ⟧) ⋆⟨ C ⟩ G ⟪ ε ⟦ d ⟧ ⟫ + ≡⟨ C .⋆Assoc _ _ _ ⟩ + g ⋆⟨ C ⟩ (η ⟦ G ⟅ d ⟆ ⟧ ⋆⟨ C ⟩ G ⟪ ε ⟦ d ⟧ ⟫) + ≡⟨ lPrecatWhisker {C = C} _ _ _ δ₂ ⟩ + g ⋆⟨ C ⟩ C .id _ + ≡⟨ C .⋆IdR _ ⟩ + g + ∎ + where + natu : η ⟦ c ⟧ ⋆⟨ C ⟩ G ⟪ F ⟪ g ⟫ ⟫ ≡ g ⋆⟨ C ⟩ η ⟦ G ⟅ d ⟆ ⟧ + natu = sym (η .N-hom _) + adj→adj' .adjIso {c = c} {d} .leftInv f + = F ⟪ η ⟦ c ⟧ ⋆⟨ C ⟩ G ⟪ f ⟫ ⟫ ⋆⟨ D ⟩ ε ⟦ d ⟧ + ≡⟨ cong (λ v → v ⋆⟨ D ⟩ ε ⟦ d ⟧) (F .F-seq _ _) ⟩ + F ⟪ η ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ F ⟪ G ⟪ f ⟫ ⟫ ⋆⟨ D ⟩ ε ⟦ d ⟧ + ≡⟨ D .⋆Assoc _ _ _ ⟩ + F ⟪ η ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ (F ⟪ G ⟪ f ⟫ ⟫ ⋆⟨ D ⟩ ε ⟦ d ⟧) + -- apply naturality + ≡⟨ lPrecatWhisker {C = D} _ _ _ natu ⟩ + F ⟪ η ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ (ε ⟦ F ⟅ c ⟆ ⟧ ⋆⟨ D ⟩ f) + ≡⟨ sym (D .⋆Assoc _ _ _) ⟩ + F ⟪ η ⟦ c ⟧ ⟫ ⋆⟨ D ⟩ ε ⟦ F ⟅ c ⟆ ⟧ ⋆⟨ D ⟩ f + -- apply triangle identity + ≡⟨ rPrecatWhisker {C = D} _ _ _ δ₁ ⟩ + (D .id _) ⋆⟨ D ⟩ f + ≡⟨ D .⋆IdL _ ⟩ + f + ∎ + where + natu : F ⟪ G ⟪ f ⟫ ⟫ ⋆⟨ D ⟩ ε ⟦ d ⟧ ≡ ε ⟦ F ⟅ c ⟆ ⟧ ⋆⟨ D ⟩ f + natu = ε .N-hom _ + -- follows directly from functoriality + adj→adj' .adjNatInD {c = c} f k = cong (λ v → η ⟦ c ⟧ ⋆⟨ C ⟩ v) (G .F-seq _ _) ∙ (sym (C .⋆Assoc _ _ _)) + adj→adj' .adjNatInC {d = d} g h = cong (λ v → v ⋆⟨ D ⟩ ε ⟦ d ⟧) (F .F-seq _ _) ∙ D .⋆Assoc _ _ _ diff --git a/Cubical/Categories/Category.agda b/Cubical/Categories/Category.agda new file mode 100644 index 000000000..54f2d8ef9 --- /dev/null +++ b/Cubical/Categories/Category.agda @@ -0,0 +1,25 @@ +{- + Definition of various kinds of categories. + + This library follows the UniMath terminology, that is: + + Concept Ob C Hom C Univalence + + Precategory Type Type No + Category Type Set No + Univalent Category Type Set Yes + + This file also contains + - pathToIso : Turns a path between two objects into an isomorphism between them + - opposite categories + + +-} + +{-# OPTIONS --safe #-} + + +module Cubical.Categories.Category where + +open import Cubical.Categories.Category.Base public +open import Cubical.Categories.Category.Properties public diff --git a/Cubical/Categories/Category/Base.agda b/Cubical/Categories/Category/Base.agda new file mode 100644 index 000000000..c29fb204a --- /dev/null +++ b/Cubical/Categories/Category/Base.agda @@ -0,0 +1,101 @@ +{-# OPTIONS --safe #-} + + +module Cubical.Categories.Category.Base where + +open import Cubical.Core.Glue +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv + +private + variable + ℓ ℓ' : Level + +-- Precategories + +record Precategory ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where + -- no-eta-equality ; NOTE: need eta equality for `opop` + field + ob : Type ℓ + Hom[_,_] : ob → ob → Type ℓ' + id : ∀ x → Hom[ x , x ] + _⋆_ : ∀ {x y z} (f : Hom[ x , y ]) (g : Hom[ y , z ]) → Hom[ x , z ] + -- TODO: change these to implicit argument + ⋆IdL : ∀ {x y : ob} (f : Hom[ x , y ]) → (id x) ⋆ f ≡ f + ⋆IdR : ∀ {x y} (f : Hom[ x , y ]) → f ⋆ (id y) ≡ f + ⋆Assoc : ∀ {u v w x} (f : Hom[ u , v ]) (g : Hom[ v , w ]) (h : Hom[ w , x ]) → (f ⋆ g) ⋆ h ≡ f ⋆ (g ⋆ h) + + -- composition: alternative to diagramatic order + _∘_ : ∀ {x y z} (g : Hom[ y , z ]) (f : Hom[ x , y ]) → Hom[ x , z ] + g ∘ f = f ⋆ g + +open Precategory + + +-- Helpful syntax/notation + +_[_,_] : (C : Precategory ℓ ℓ') → (x y : C .ob) → Type ℓ' +_[_,_] = Hom[_,_] + +-- needed to define this in order to be able to make the subsequence syntax declaration +seq' : ∀ (C : Precategory ℓ ℓ') {x y z} (f : C [ x , y ]) (g : C [ y , z ]) → C [ x , z ] +seq' = _⋆_ + +infixl 15 seq' +syntax seq' C f g = f ⋆⟨ C ⟩ g + +-- composition +comp' : ∀ (C : Precategory ℓ ℓ') {x y z} (g : C [ y , z ]) (f : C [ x , y ]) → C [ x , z ] +comp' = _∘_ + +infixr 16 comp' +syntax comp' C g f = g ∘⟨ C ⟩ f + + +-- Categories + +record isCategory (C : Precategory ℓ ℓ') : Type (ℓ-max ℓ ℓ') where + field + isSetHom : ∀ {x y} → isSet (C [ x , y ]) + +-- Isomorphisms and paths in precategories + +record CatIso {C : Precategory ℓ ℓ'} (x y : C .Precategory.ob) : Type ℓ' where + constructor catiso + field + mor : C [ x , y ] + inv : C [ y , x ] + sec : inv ⋆⟨ C ⟩ mor ≡ C .id y + ret : mor ⋆⟨ C ⟩ inv ≡ C .id x + + +pathToIso : {C : Precategory ℓ ℓ'} (x y : C .ob) (p : x ≡ y) → CatIso {C = C} x y +pathToIso {C = C} x y p = J (λ z _ → CatIso x z) (catiso (C .id x) idx (C .⋆IdL idx) (C .⋆IdL idx)) p + where + idx = C .id x + +-- Univalent Categories + +record isUnivalent (C : Precategory ℓ ℓ') : Type (ℓ-max ℓ ℓ') where + field + univ : (x y : C .ob) → isEquiv (pathToIso {C = C} x y) + + -- package up the univalence equivalence + univEquiv : ∀ (x y : C .ob) → (x ≡ y) ≃ (CatIso x y) + univEquiv x y = (pathToIso {C = C} x y) , (univ x y) + +open isUnivalent public + +-- Opposite Categories + +_^op : Precategory ℓ ℓ' → Precategory ℓ ℓ' +(C ^op) .ob = C .ob +(C ^op) .Hom[_,_] x y = C .Hom[_,_] y x +(C ^op) .id = C .id +(C ^op) ._⋆_ f g = C ._⋆_ g f +(C ^op) .⋆IdL = C .⋆IdR +(C ^op) .⋆IdR = C .⋆IdL +(C ^op) .⋆Assoc f g h = sym (C .⋆Assoc _ _ _) + +open isCategory public diff --git a/Cubical/Categories/Category/Properties.agda b/Cubical/Categories/Category/Properties.agda new file mode 100644 index 000000000..35fc7b6a4 --- /dev/null +++ b/Cubical/Categories/Category/Properties.agda @@ -0,0 +1,93 @@ +{-# OPTIONS --safe #-} + + +module Cubical.Categories.Category.Properties where + +open import Cubical.Core.Glue +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Categories.Category.Base hiding (isSetHom) + +open Precategory + +private + variable + ℓ ℓ' : Level + +module _ {C : Precategory ℓ ℓ'} ⦃ isC : isCategory C ⦄ where + open isCategory isC + -- isSet where your allowed to compare paths where one side is only + -- equal up to path. Is there a generalization of this? + isSetHomP1 : ∀ {x y : C .ob} {n : C [ x , y ]} + → isOfHLevelDep 1 (λ m → m ≡ n) + isSetHomP1 {x = x} {y} {n} = isOfHLevel→isOfHLevelDep 1 (λ m → isSetHom m n) + + -- isSet where the arrows can be between non-definitionally equal obs + isSetHomP2l : ∀ {y : C .ob} + → isOfHLevelDep 2 (λ x → C [ x , y ]) + isSetHomP2l = isOfHLevel→isOfHLevelDep 2 (λ a → isSetHom {x = a}) + + isSetHomP2r : ∀ {x : C .ob} + → isOfHLevelDep 2 (λ y → C [ x , y ]) + isSetHomP2r = isOfHLevel→isOfHLevelDep 2 (λ a → isSetHom {y = a}) + + + +-- opposite of opposite is definitionally equal to itself +involutiveOp : ∀ {C : Precategory ℓ ℓ'} → (C ^op) ^op ≡ C +involutiveOp = refl + +module _ {C : Precategory ℓ ℓ'} where + -- Other useful operations on categories + + -- whisker the parallel morphisms g and g' with f + lPrecatWhisker : {x y z : C .ob} (f : C [ x , y ]) (g g' : C [ y , z ]) (p : g ≡ g') → f ⋆⟨ C ⟩ g ≡ f ⋆⟨ C ⟩ g' + lPrecatWhisker f _ _ p = cong (_⋆_ C f) p + + rPrecatWhisker : {x y z : C .ob} (f f' : C [ x , y ]) (g : C [ y , z ]) (p : f ≡ f') → f ⋆⟨ C ⟩ g ≡ f' ⋆⟨ C ⟩ g + rPrecatWhisker _ _ g p = cong (λ v → v ⋆⟨ C ⟩ g) p + + -- working with equal objects + idP : ∀ {x x'} {p : x ≡ x'} → C [ x , x' ] + idP {x = x} {x'} {p} = subst (λ v → C [ x , v ]) p (C .id x) + + -- heterogeneous seq + seqP : ∀ {x y y' z} {p : y ≡ y'} + → (f : C [ x , y ]) (g : C [ y' , z ]) + → C [ x , z ] + seqP {x = x} {_} {_} {z} {p} f g = f ⋆⟨ C ⟩ (subst (λ a → C [ a , z ]) (sym p) g) + + -- also heterogeneous seq, but substituting on the left + seqP' : ∀ {x y y' z} {p : y ≡ y'} + → (f : C [ x , y ]) (g : C [ y' , z ]) + → C [ x , z ] + seqP' {x = x} {_} {_} {z} {p} f g = subst (λ a → C [ x , a ]) p f ⋆⟨ C ⟩ g + + -- show that they're equal + seqP≡seqP' : ∀ {x y y' z} {p : y ≡ y'} + → (f : C [ x , y ]) (g : C [ y' , z ]) + → seqP {p = p} f g ≡ seqP' {p = p} f g + seqP≡seqP' {x = x} {z = z} {p = p} f g i = + (toPathP {A = λ i' → C [ x , p i' ]} {f} refl i) + ⋆⟨ C ⟩ + (toPathP {A = λ i' → C [ p (~ i') , z ]} {x = g} (sym refl) (~ i)) + + -- seqP is equal to normal seq when y ≡ y' + seqP≡seq : ∀ {x y z} + → (f : C [ x , y ]) (g : C [ y , z ]) + → seqP {p = refl} f g ≡ f ⋆⟨ C ⟩ g + seqP≡seq {y = y} {z} f g i = f ⋆⟨ C ⟩ toPathP {A = λ _ → C [ y , z ]} {x = g} refl (~ i) + + + -- whiskering with heterogenous seq -- (maybe should let z be heterogeneous too) + lPrecatWhiskerP : {x y z y' : C .ob} {p : y ≡ y'} + → (f : C [ x , y ]) (g : C [ y , z ]) (g' : C [ y' , z ]) + → (r : PathP (λ i → C [ p i , z ]) g g') + → f ⋆⟨ C ⟩ g ≡ seqP {p = p} f g' + lPrecatWhiskerP f g g' r = cong (λ v → f ⋆⟨ C ⟩ v) (sym (fromPathP (symP r))) + + rPrecatWhiskerP : {x y' y z : C .ob} {p : y' ≡ y} + → (f' : C [ x , y' ]) (f : C [ x , y ]) (g : C [ y , z ]) + → (r : PathP (λ i → C [ x , p i ]) f' f) + → f ⋆⟨ C ⟩ g ≡ seqP' {p = p} f' g + rPrecatWhiskerP f' f g r = cong (λ v → v ⋆⟨ C ⟩ g) (sym (fromPathP r)) diff --git a/Cubical/Categories/Commutativity.agda b/Cubical/Categories/Commutativity.agda new file mode 100644 index 000000000..b69c95a1b --- /dev/null +++ b/Cubical/Categories/Commutativity.agda @@ -0,0 +1,33 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Commutativity where + +open import Cubical.Foundations.Prelude +open import Cubical.Categories.Category + +private + variable + ℓ ℓ' : Level + +module _ {C : Precategory ℓ ℓ'} where + open Precategory C + + compSq : ∀ {x y z w u v} {f : C [ x , y ]} {g h} {k : C [ z , w ]} {l} {m} {n : C [ u , v ]} + -- square 1 + → f ⋆ g ≡ h ⋆ k + -- square 2 (sharing g) + → k ⋆ l ≡ m ⋆ n + → f ⋆ (g ⋆ l) ≡ (h ⋆ m) ⋆ n + compSq {f = f} {g} {h} {k} {l} {m} {n} p q + = f ⋆ (g ⋆ l) + ≡⟨ sym (⋆Assoc _ _ _) ⟩ + (f ⋆ g) ⋆ l + ≡⟨ cong (_⋆ l) p ⟩ + (h ⋆ k) ⋆ l + ≡⟨ ⋆Assoc _ _ _ ⟩ + h ⋆ (k ⋆ l) + ≡⟨ cong (h ⋆_) q ⟩ + h ⋆ (m ⋆ n) + ≡⟨ sym (⋆Assoc _ _ _) ⟩ + (h ⋆ m) ⋆ n + ∎ diff --git a/Cubical/Categories/Constructions/Elements.agda b/Cubical/Categories/Constructions/Elements.agda new file mode 100644 index 000000000..a8cad42f8 --- /dev/null +++ b/Cubical/Categories/Constructions/Elements.agda @@ -0,0 +1,137 @@ +{-# OPTIONS --safe #-} + +-- The Category of Elements + +open import Cubical.Categories.Category + +module Cubical.Categories.Constructions.Elements {ℓ ℓ'} {C : Precategory ℓ ℓ'} where + +open import Cubical.Categories.Instances.Sets +open import Cubical.Categories.Functor +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma + +import Cubical.Categories.Morphism as Morphism +import Cubical.Categories.Constructions.Slice as Slice + +-- some issues +-- * always need to specify objects during composition because can't infer isSet +open Precategory +open Functor + + +getIsSet : ∀ {ℓS} {C : Precategory ℓ ℓ'} (F : Functor C (SET ℓS)) → (c : C .ob) → isSet (fst (F ⟅ c ⟆)) +getIsSet F c = snd (F ⟅ c ⟆) + + +infix 50 ∫_ +∫_ : ∀ {ℓS} → Functor C (SET ℓS) → Precategory (ℓ-max ℓ ℓS) (ℓ-max ℓ' ℓS) +-- objects are (c , x) pairs where c ∈ C and x ∈ F c +(∫ F) .ob = Σ[ c ∈ C .ob ] fst (F ⟅ c ⟆) +-- morphisms are f : c → c' which take x to x' +(∫ F) .Hom[_,_] (c , x) (c' , x') = Σ[ f ∈ C [ c , c' ] ] x' ≡ (F ⟪ f ⟫) x +(∫ F) .id (c , x) = C .id c , sym (funExt⁻ (F .F-id) x ∙ refl) +(∫ F) ._⋆_ {c , x} {c₁ , x₁} {c₂ , x₂} (f , p) (g , q) + = (f ⋆⟨ C ⟩ g) , (x₂ + ≡⟨ q ⟩ + (F ⟪ g ⟫) x₁ -- basically expanding out function composition + ≡⟨ cong (F ⟪ g ⟫) p ⟩ + (F ⟪ g ⟫) ((F ⟪ f ⟫) x) + ≡⟨ funExt⁻ (sym (F .F-seq _ _)) _ ⟩ + (F ⟪ f ⋆⟨ C ⟩ g ⟫) x + ∎) +(∫ F) .⋆IdL o@{c , x} o1@{c' , x'} f'@(f , p) i + = (cIdL i) , isOfHLevel→isOfHLevelDep 1 (λ a → isS' x' ((F ⟪ a ⟫) x)) p' p cIdL i + where + isS = getIsSet F c + isS' = getIsSet F c' + cIdL = C .⋆IdL f + + -- proof from composition with id + p' : x' ≡ (F ⟪ C .id c ⋆⟨ C ⟩ f ⟫) x + p' = snd ((∫ F) ._⋆_ ((∫ F) .id o) f') +(∫ F) .⋆IdR o@{c , x} o1@{c' , x'} f'@(f , p) i + = (cIdR i) , isOfHLevel→isOfHLevelDep 1 (λ a → isS' x' ((F ⟪ a ⟫) x)) p' p cIdR i + where + cIdR = C .⋆IdR f + isS' = getIsSet F c' + + p' : x' ≡ (F ⟪ f ⋆⟨ C ⟩ C .id c' ⟫) x + p' = snd ((∫ F) ._⋆_ f' ((∫ F) .id o1)) +(∫ F) .⋆Assoc o@{c , x} o1@{c₁ , x₁} o2@{c₂ , x₂} o3@{c₃ , x₃} f'@(f , p) g'@(g , q) h'@(h , r) i + = (cAssoc i) , isOfHLevel→isOfHLevelDep 1 (λ a → isS₃ x₃ ((F ⟪ a ⟫) x)) p1 p2 cAssoc i + where + cAssoc = C .⋆Assoc f g h + isS₃ = getIsSet F c₃ + + p1 : x₃ ≡ (F ⟪ (f ⋆⟨ C ⟩ g) ⋆⟨ C ⟩ h ⟫) x + p1 = snd ((∫ F) ._⋆_ ((∫ F) ._⋆_ {o} {o1} {o2} f' g') h') + + p2 : x₃ ≡ (F ⟪ f ⋆⟨ C ⟩ (g ⋆⟨ C ⟩ h) ⟫) x + p2 = snd ((∫ F) ._⋆_ f' ((∫ F) ._⋆_ {o1} {o2} {o3} g' h')) + + +-- same thing but for presheaves +∫ᴾ_ : ∀ {ℓS} → Functor (C ^op) (SET ℓS) → Precategory (ℓ-max ℓ ℓS) (ℓ-max ℓ' ℓS) +-- objects are (c , x) pairs where c ∈ C and x ∈ F c +(∫ᴾ F) .ob = Σ[ c ∈ C .ob ] fst (F ⟅ c ⟆) +-- morphisms are f : c → c' which take x to x' +(∫ᴾ F) .Hom[_,_] (c , x) (c' , x') = Σ[ f ∈ C [ c , c' ] ] x ≡ (F ⟪ f ⟫) x' +(∫ᴾ F) .id (c , x) = C .id c , sym (funExt⁻ (F .F-id) x ∙ refl) +(∫ᴾ F) ._⋆_ {c , x} {c₁ , x₁} {c₂ , x₂} (f , p) (g , q) + = (f ⋆⟨ C ⟩ g) , (x + ≡⟨ p ⟩ + (F ⟪ f ⟫) x₁ -- basically expanding out function composition + ≡⟨ cong (F ⟪ f ⟫) q ⟩ + (F ⟪ f ⟫) ((F ⟪ g ⟫) x₂) + ≡⟨ funExt⁻ (sym (F .F-seq _ _)) _ ⟩ + (F ⟪ f ⋆⟨ C ⟩ g ⟫) x₂ + ∎) +(∫ᴾ F) .⋆IdL o@{c , x} o1@{c' , x'} f'@(f , p) i + = (cIdL i) , isOfHLevel→isOfHLevelDep 1 (λ a → isS x ((F ⟪ a ⟫) x')) p' p cIdL i + where + isS = getIsSet F c + isS' = getIsSet F c' + cIdL = C .⋆IdL f + + -- proof from composition with id + p' : x ≡ (F ⟪ C .id c ⋆⟨ C ⟩ f ⟫) x' + p' = snd ((∫ᴾ F) ._⋆_ ((∫ᴾ F) .id o) f') +(∫ᴾ F) .⋆IdR o@{c , x} o1@{c' , x'} f'@(f , p) i + = (cIdR i) , isOfHLevel→isOfHLevelDep 1 (λ a → isS x ((F ⟪ a ⟫) x')) p' p cIdR i + where + cIdR = C .⋆IdR f + isS = getIsSet F c + + p' : x ≡ (F ⟪ f ⋆⟨ C ⟩ C .id c' ⟫) x' + p' = snd ((∫ᴾ F) ._⋆_ f' ((∫ᴾ F) .id o1)) +(∫ᴾ F) .⋆Assoc o@{c , x} o1@{c₁ , x₁} o2@{c₂ , x₂} o3@{c₃ , x₃} f'@(f , p) g'@(g , q) h'@(h , r) i + = (cAssoc i) , isOfHLevel→isOfHLevelDep 1 (λ a → isS x ((F ⟪ a ⟫) x₃)) p1 p2 cAssoc i + where + cAssoc = C .⋆Assoc f g h + isS = getIsSet F c + + p1 : x ≡ (F ⟪ (f ⋆⟨ C ⟩ g) ⋆⟨ C ⟩ h ⟫) x₃ + p1 = snd ((∫ᴾ F) ._⋆_ ((∫ᴾ F) ._⋆_ {o} {o1} {o2} f' g') h') + + p2 : x ≡ (F ⟪ f ⋆⟨ C ⟩ (g ⋆⟨ C ⟩ h) ⟫) x₃ + p2 = snd ((∫ᴾ F) ._⋆_ f' ((∫ᴾ F) ._⋆_ {o1} {o2} {o3} g' h')) + +-- helpful results + +module _ {ℓS} {F : Functor (C ^op) (SET ℓS)} where + + -- morphisms are equal as long as the morphisms in C are equals + ∫ᴾhomEq : ∀ {o1 o1' o2 o2'} (f : (∫ᴾ F) [ o1 , o2 ]) (g : (∫ᴾ F) [ o1' , o2' ]) + → (p : o1 ≡ o1') (q : o2 ≡ o2') + → (eqInC : PathP (λ i → C [ fst (p i) , fst (q i) ]) (fst f) (fst g)) + → PathP (λ i → (∫ᴾ F) [ p i , q i ]) f g + ∫ᴾhomEq (f , eqf) (g , eqg) p q eqInC + = ΣPathP (eqInC + , isOfHLevel→isOfHLevelDep 1 {A = Σ[ (o1 , o2) ∈ (∫ᴾ F) .ob × (∫ᴾ F) .ob ] (C [ fst o1 , fst o2 ])} + {B = λ ((o1 , o2) , f) → snd o1 ≡ (F ⟪ f ⟫) (snd o2)} + (λ ((o1 , o2) , f) → snd (F ⟅ (fst o1) ⟆) (snd o1) ((F ⟪ f ⟫) (snd o2))) + eqf + eqg + λ i → ((p i , q i) , eqInC i)) diff --git a/Cubical/Categories/Constructions/Slice.agda b/Cubical/Categories/Constructions/Slice.agda new file mode 100644 index 000000000..2fb55752e --- /dev/null +++ b/Cubical/Categories/Constructions/Slice.agda @@ -0,0 +1,383 @@ +{-# OPTIONS --safe #-} + +open import Cubical.Categories.Category +open import Cubical.Categories.Morphism renaming (isIso to isIsoC) +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open Iso +open import Cubical.Foundations.HLevels +open Precategory +open import Cubical.Core.Glue +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport using (transpFill) + +module Cubical.Categories.Constructions.Slice {ℓ ℓ' : Level} (C : Precategory ℓ ℓ') (c : C .ob) {{isC : isCategory C}} where + +open import Cubical.Data.Sigma + + +-- just a helper to prevent redundency +TypeC : Type (ℓ-suc (ℓ-max ℓ ℓ')) +TypeC = Type (ℓ-max ℓ ℓ') + +-- Components of a slice category + +record SliceOb : TypeC where + constructor sliceob + field + {S-ob} : C .ob + S-arr : C [ S-ob , c ] + +open SliceOb public + +record SliceHom (a b : SliceOb) : Type ℓ' where + constructor slicehom + field + S-hom : C [ S-ob a , S-ob b ] + -- commutative diagram + S-comm : S-hom ⋆⟨ C ⟩ (S-arr b) ≡ S-arr a + +open SliceHom public + +-- Helpers for working with equality +-- can probably replace these by showing that SliceOb is isomorphic to Sigma and +-- that paths are isomorphic to Sigma? But sounds like that would need a lot of transp + +SliceOb-≡-intro : ∀ {a b} {f g} + → (p : a ≡ b) + → PathP (λ i → C [ p i , c ]) f g + → sliceob {a} f ≡ sliceob {b} g +SliceOb-≡-intro p q = λ i → sliceob {p i} (q i) + +module _ {xf yg : SliceOb} where + private + x = xf .S-ob + f = xf .S-arr + y = yg .S-ob + g = yg .S-arr + + -- a path between slice objects is the "same" as a pair of paths between C obs and C arrows + SOPathIsoPathΣ : Iso (xf ≡ yg) (Σ[ p ∈ x ≡ y ] PathP (λ i → C [ p i , c ]) f g) + SOPathIsoPathΣ .fun p = (λ i → (p i) .S-ob) , (λ i → (p i) .S-arr) + SOPathIsoPathΣ .inv (p , q) i = sliceob {p i} (q i) + SOPathIsoPathΣ .rightInv _ = refl + SOPathIsoPathΣ .leftInv _ = refl + + SOPath≃PathΣ = isoToEquiv SOPathIsoPathΣ + + SOPath≡PathΣ = ua (isoToEquiv SOPathIsoPathΣ) + +-- intro and elim for working with SliceHom equalities (is there a better way to do this?) +SliceHom-≡-intro : ∀ {a b} {f g} {c₁} {c₂} + → (p : f ≡ g) + → PathP (λ i → (p i) ⋆⟨ C ⟩ (S-arr b) ≡ S-arr a) c₁ c₂ + → slicehom f c₁ ≡ slicehom g c₂ +SliceHom-≡-intro p q = λ i → slicehom (p i) (q i) + +SliceHom-≡-elim : ∀ {a b} {f g} {c₁} {c₂} + → slicehom f c₁ ≡ slicehom g c₂ + → Σ[ p ∈ f ≡ g ] PathP (λ i → (p i) ⋆⟨ C ⟩ (S-arr b) ≡ S-arr a) c₁ c₂ +SliceHom-≡-elim r = (λ i → S-hom (r i)) , λ i → S-comm (r i) + + +SliceHom-≡-intro' : ∀ {a b} {f g : C [ a .S-ob , b .S-ob ]} {c₁} {c₂} + → (p : f ≡ g) + → slicehom f c₁ ≡ slicehom g c₂ +SliceHom-≡-intro' {a} {b} {f} {g} {c₁} {c₂} p i = slicehom (p i) (c₁≡c₂ i) + where + c₁≡c₂ : PathP (λ i → (p i) ⋆⟨ C ⟩ (b .S-arr) ≡ a .S-arr) c₁ c₂ + c₁≡c₂ = isOfHLevel→isOfHLevelDep 1 (λ _ → isC .isSetHom _ _) c₁ c₂ p + +-- SliceHom is isomorphic to the Sigma type with the same components +SliceHom-Σ-Iso : ∀ {a b} + → Iso (SliceHom a b) (Σ[ h ∈ C [ S-ob a , S-ob b ] ] h ⋆⟨ C ⟩ (S-arr b) ≡ S-arr a) +SliceHom-Σ-Iso .fun (slicehom h c) = h , c +SliceHom-Σ-Iso .inv (h , c) = slicehom h c +SliceHom-Σ-Iso .rightInv = λ x → refl +SliceHom-Σ-Iso .leftInv = λ x → refl + + +-- Precategory definition + +SliceCat : Precategory _ _ +SliceCat .ob = SliceOb +SliceCat .Hom[_,_] = SliceHom +SliceCat .id (sliceob {x} f) = slicehom (C .id x) (C .⋆IdL _) +SliceCat ._⋆_ {sliceob j} {sliceob k} {sliceob l} (slicehom f p) (slicehom g p') = + slicehom + (f ⋆⟨ C ⟩ g) + ( f ⋆⟨ C ⟩ g ⋆⟨ C ⟩ l + ≡⟨ C .⋆Assoc _ _ _ ⟩ + f ⋆⟨ C ⟩ (g ⋆⟨ C ⟩ l) + ≡⟨ cong (λ v → f ⋆⟨ C ⟩ v) p' ⟩ + f ⋆⟨ C ⟩ k + ≡⟨ p ⟩ + j + ∎) +SliceCat .⋆IdL (slicehom S-hom S-comm) = + SliceHom-≡-intro (⋆IdL C _) (toPathP (isC .isSetHom _ _ _ _)) +SliceCat .⋆IdR (slicehom S-hom S-comm) = + SliceHom-≡-intro (⋆IdR C _) (toPathP (isC .isSetHom _ _ _ _)) +SliceCat .⋆Assoc f g h = + SliceHom-≡-intro (⋆Assoc C _ _ _) (toPathP (isC .isSetHom _ _ _ _)) + + +-- SliceCat is a Category + +instance + isCatSlice : isCategory SliceCat + isCatSlice .isSetHom {a} {b} (slicehom f c₁) (slicehom g c₂) p q = cong isoP p'≡q' + where + -- paths between SliceHoms are equivalent to the projection paths + p' : Σ[ p ∈ f ≡ g ] PathP (λ i → (p i) ⋆⟨ C ⟩ (S-arr b) ≡ S-arr a) c₁ c₂ + p' = SliceHom-≡-elim p + q' : Σ[ p ∈ f ≡ g ] PathP (λ i → (p i) ⋆⟨ C ⟩ (S-arr b) ≡ S-arr a) c₁ c₂ + q' = SliceHom-≡-elim q + + -- we want all paths between (dependent) paths of this type to be equal + B = λ v → v ⋆⟨ C ⟩ (S-arr b) ≡ S-arr a + + -- need the groupoidness for dependent paths + isGroupoidDepHom : isOfHLevelDep 2 B + isGroupoidDepHom = isOfHLevel→isOfHLevelDep 2 (λ v x y → isSet→isGroupoid (isC .isSetHom) _ _ x y) + + -- we first prove that the projected paths are equal + p'≡q' : p' ≡ q' + p'≡q' = ΣPathP ((isC .isSetHom _ _ _ _) , toPathP (isGroupoidDepHom _ _ _ _ _)) + + -- and then we can use equivalence to lift these paths up + -- to actual SliceHom paths + isoP = λ g → cong (inv SliceHom-Σ-Iso) (fun (ΣPathIsoPathΣ) g) + +-- SliceCat is univalent if C is univalent + +module _ ⦃ isU : isUnivalent C ⦄ where + open CatIso + open Iso + + module _ { xf yg : SliceOb } where + private + x = xf .S-ob + y = yg .S-ob + + -- names for the equivalences/isos + + pathIsoEquiv : (x ≡ y) ≃ (CatIso x y) + pathIsoEquiv = univEquiv isU x y + + isoPathEquiv : (CatIso x y) ≃ (x ≡ y) + isoPathEquiv = invEquiv pathIsoEquiv + + pToIIso' : Iso (x ≡ y) (CatIso x y) + pToIIso' = equivToIso pathIsoEquiv + + -- the iso in SliceCat we're given induces an iso in C between x and y + module _ ( cIso@(catiso kc lc s r) : CatIso {C = SliceCat} xf yg ) where + extractIso' : CatIso {C = C} x y + extractIso' .mor = kc .S-hom + extractIso' .inv = lc .S-hom + extractIso' .sec i = (s i) .S-hom + extractIso' .ret i = (r i) .S-hom + + instance + preservesUnivalenceSlice : isUnivalent SliceCat + -- we prove the equivalence by going through Iso + preservesUnivalenceSlice .univ xf@(sliceob {x} f) yg@(sliceob {y} g) = isoToIsEquiv sIso + where + -- this is just here because the type checker can't seem to infer xf and yg + pToIIso : Iso (x ≡ y) (CatIso x y) + pToIIso = pToIIso' {xf = xf} {yg} + + -- the meat of the proof + sIso : Iso (xf ≡ yg) (CatIso xf yg) + sIso .fun p = pathToIso xf yg p -- we use the normal pathToIso via path induction to get an isomorphism + sIso .inv is@(catiso kc lc s r) = SliceOb-≡-intro x≡y (symP (sym (lc .S-comm) ◁ lf≡f)) + where + -- we get a path between xf and yg by combining paths between + -- x and y, and f and g + -- 1. x≡y follows from univalence of C + -- 2. f≡g is more tricky; by commutativity, we know that g ≡ l ⋆ f + -- so we want l to be id; we get this by showing: id ≡ pathToIso x y x≡y ≡ l + -- where the first step follows from path induction, and the second from univalence of C + + -- morphisms in C from kc and lc + k = kc .S-hom + l = lc .S-hom + + -- extract out the iso between x and y + extractIso : CatIso {C = C} x y + extractIso = extractIso' is + + -- and we can use univalence of C to get x ≡ y + x≡y : x ≡ y + x≡y = pToIIso .inv extractIso + + -- to show that f ≡ g, we show that l ≡ id + -- by using C's isomorphism + pToI≡id : PathP (λ i → C [ x≡y (~ i) , x ]) (pathToIso {C = C} x y x≡y .inv) (C .id x) + pToI≡id = J (λ y p → PathP (λ i → C [ p (~ i) , x ]) (pathToIso {C = C} x y p .inv) (C .id x)) + (λ j → JRefl pToIFam pToIBase j .inv) + x≡y + where + idx = C .id x + pToIFam = (λ z _ → CatIso {C = C} x z) + pToIBase = catiso (C .id x) idx (C .⋆IdL idx) (C .⋆IdL idx) + + l≡pToI : l ≡ pathToIso {C = C} x y x≡y .inv + l≡pToI i = pToIIso .rightInv extractIso (~ i) .inv + + l≡id : PathP (λ i → C [ x≡y (~ i) , x ]) l (C .id x) + l≡id = l≡pToI ◁ pToI≡id + + lf≡f : PathP (λ i → C [ x≡y (~ i) , c ]) (l ⋆⟨ C ⟩ f) f + lf≡f = (λ i → (l≡id i) ⋆⟨ C ⟩ f) ▷ C .⋆IdL _ + + sIso .rightInv is@(catiso kc lc s r) i = catiso (kc'≡kc i) (lc'≡lc i) (s'≡s i) (r'≡r i) + -- we prove rightInv using a combination of univalence and the fact that homs are an h-set + where + kc' = (sIso .fun) (sIso .inv is) .mor + lc' = (sIso .fun) (sIso .inv is) .inv + k' = kc' .S-hom + l' = lc' .S-hom + k = kc .S-hom + l = lc .S-hom + + extractIso : CatIso {C = C} x y + extractIso = extractIso' is + + -- we do the equality component wise + + -- mor + + k'≡k : k' ≡ k + k'≡k i = (pToIIso .rightInv extractIso) i .mor + + kcom'≡kcom : PathP (λ j → (k'≡k j) ⋆⟨ C ⟩ g ≡ f) (kc' .S-comm) (kc .S-comm) + kcom'≡kcom = isSetHomP1 _ _ λ i → (k'≡k i) ⋆⟨ C ⟩ g + + kc'≡kc : kc' ≡ kc + kc'≡kc i = slicehom (k'≡k i) (kcom'≡kcom i) + + -- inv + + l'≡l : l' ≡ l + l'≡l i = (pToIIso .rightInv extractIso) i .inv + + lcom'≡lcom : PathP (λ j → (l'≡l j) ⋆⟨ C ⟩ f ≡ g) (lc' .S-comm) (lc .S-comm) + lcom'≡lcom = isSetHomP1 _ _ λ i → (l'≡l i) ⋆⟨ C ⟩ f + + lc'≡lc : lc' ≡ lc + lc'≡lc i = slicehom (l'≡l i) (lcom'≡lcom i) + + -- sec + + s' = (sIso .fun) (sIso .inv is) .sec + s'≡s : PathP (λ i → lc'≡lc i ⋆⟨ SliceCat ⟩ kc'≡kc i ≡ SliceCat .id _) s' s + s'≡s = isSetHomP1 _ _ λ i → lc'≡lc i ⋆⟨ SliceCat ⟩ kc'≡kc i + + -- ret + + r' = (sIso .fun) (sIso .inv is) .ret + r'≡r : PathP (λ i → kc'≡kc i ⋆⟨ SliceCat ⟩ lc'≡lc i ≡ SliceCat .id _) r' r + r'≡r = isSetHomP1 _ _ λ i → kc'≡kc i ⋆⟨ SliceCat ⟩ lc'≡lc i + + sIso .leftInv p = p'≡p + -- to show that the round trip is equivalent to the identity + -- we show that this is true for each component (S-ob, S-arr) + -- and then combine + -- specifically, we show that p'Ob≡pOb and p'Mor≡pMor + -- and it follows that p'≡p + where + p' = (sIso .inv) (sIso .fun p) + + pOb : x ≡ y + pOb i = (p i) .S-ob + + p'Ob : x ≡ y + p'Ob i = (p' i) .S-ob + + + pMor : PathP (λ i → C [ pOb i , c ]) f g + pMor i = (p i) .S-arr + + p'Mor : PathP (λ i → C [ p'Ob i , c ]) f g + p'Mor i = (p' i) .S-arr + + -- we first show that it's equivalent to use sIso first then extract, or to extract first than use pToIIso + extractCom : extractIso' (sIso .fun p) ≡ pToIIso .fun pOb + extractCom = J (λ yg' p̃ → extractIso' (pathToIso xf yg' p̃) ≡ pToIIso' {xf = xf} {yg'} .fun (λ i → (p̃ i) .S-ob)) + (cong extractIso' (JRefl pToIFam' pToIBase') ∙ sym (JRefl pToIFam pToIBase)) + p + where + idx = C .id x + pToIFam = (λ z _ → CatIso {C = C} x z) + pToIBase = catiso (C .id x) idx (C .⋆IdL idx) (C .⋆IdL idx) + + idxf = SliceCat .id xf + pToIFam' = (λ z _ → CatIso {C = SliceCat} xf z) + pToIBase' = catiso (SliceCat .id xf) idxf (SliceCat .⋆IdL idxf) (SliceCat .⋆IdL idxf) + + -- why does this not follow definitionally? + -- from extractCom, we get that performing the roundtrip on pOb gives us back p'Ob + ppp : p'Ob ≡ (pToIIso .inv) (pToIIso .fun pOb) + ppp = cong (pToIIso .inv) extractCom + + -- apply univalence of C + -- this gives us the first component that we want + p'Ob≡pOb : p'Ob ≡ pOb + p'Ob≡pOb = ppp ∙ pToIIso .leftInv pOb + + -- isSetHom gives us the second component, path between morphisms + p'Mor≡pMor : PathP (λ j → PathP (λ i → C [ (p'Ob≡pOb j) i , c ]) f g) p'Mor pMor + p'Mor≡pMor = isSetHomP2l _ _ p'Mor pMor p'Ob≡pOb + + -- we can use the above paths to show that p' ≡ p + p'≡p : p' ≡ p + p'≡p i = comp (λ i' → SOPath≡PathΣ {xf = xf} {yg} (~ i')) + (λ j → λ { (i = i0) → left (~ j) ; (i = i1) → right (~ j) }) + (p'Σ≡pΣ i) + where + -- we break up p' and p into their constituent paths + -- first via transport and then via our component definitions from before + -- we show that p'ΣT ≡ p'Σ (and same for p) via univalence + -- and p'Σ≡pΣ follows from our work from above + p'ΣT : Σ[ p ∈ x ≡ y ] PathP (λ i → C [ p i , c ]) f g + p'ΣT = transport SOPath≡PathΣ p' + p'Σ : Σ[ p ∈ x ≡ y ] PathP (λ i → C [ p i , c ]) f g + p'Σ = (p'Ob , p'Mor) + + pΣT : Σ[ p ∈ x ≡ y ] PathP (λ i → C [ p i , c ]) f g + pΣT = transport SOPath≡PathΣ p + pΣ : Σ[ p ∈ x ≡ y ] PathP (λ i → C [ p i , c ]) f g + pΣ = (pOb , pMor)-- transport SOPathP≡PathPSO p + + -- using the computation rule to ua + p'ΣT≡p'Σ : p'ΣT ≡ p'Σ + p'ΣT≡p'Σ = uaβ SOPath≃PathΣ p' + + pΣT≡pΣ : pΣT ≡ pΣ + pΣT≡pΣ = uaβ SOPath≃PathΣ p + + p'Σ≡pΣ : p'Σ ≡ pΣ + p'Σ≡pΣ = ΣPathP (p'Ob≡pOb , p'Mor≡pMor) + + -- two sides of the square we're connecting + left : PathP (λ i → SOPath≡PathΣ {xf = xf} {yg} i) p' p'Σ + left = transport-filler SOPath≡PathΣ p' ▷ p'ΣT≡p'Σ + + right : PathP (λ i → SOPath≡PathΣ {xf = xf} {yg} i) p pΣ + right = transport-filler SOPath≡PathΣ p ▷ pΣT≡pΣ + +-- properties +-- TODO: move to own file + +open isIsoC renaming (inv to invC) + +-- make a slice isomorphism from just the hom +sliceIso : ∀ {a b} (f : C [ a .S-ob , b .S-ob ]) (c : (f ⋆⟨ C ⟩ b .S-arr) ≡ a .S-arr) + → isIsoC {C = C} f + → isIsoC {C = SliceCat} (slicehom f c) +sliceIso f c isof .invC = slicehom (isof .invC) (sym (invMoveL (isIso→areInv isof) c)) +sliceIso f c isof .sec = SliceHom-≡-intro' (isof .sec) +sliceIso f c isof .ret = SliceHom-≡-intro' (isof .ret) diff --git a/Cubical/Categories/Equivalence.agda b/Cubical/Categories/Equivalence.agda new file mode 100644 index 000000000..cea02fe38 --- /dev/null +++ b/Cubical/Categories/Equivalence.agda @@ -0,0 +1,7 @@ + +{-# OPTIONS --safe #-} + +module Cubical.Categories.Equivalence where + +open import Cubical.Categories.Equivalence.Base public +open import Cubical.Categories.Equivalence.Properties public diff --git a/Cubical/Categories/Equivalence/Base.agda b/Cubical/Categories/Equivalence/Base.agda new file mode 100644 index 000000000..b806f7fc4 --- /dev/null +++ b/Cubical/Categories/Equivalence/Base.agda @@ -0,0 +1,31 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Equivalence.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation + +open Precategory +open Functor + +private + variable + ℓC ℓC' ℓD ℓD' : Level + +-- Definition + +record isEquivalence {C : Precategory ℓC ℓC'} {D : Precategory ℓD ℓD'} + (func : Functor C D) : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where + field + invFunc : Functor D C + + η : 𝟙⟨ C ⟩ ≅ᶜ invFunc ∘F func + ε : func ∘F invFunc ≅ᶜ 𝟙⟨ D ⟩ + +record _≃ᶜ_ (C : Precategory ℓC ℓC') (D : Precategory ℓD ℓD') : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where + field + func : Functor C D + isEquiv : isEquivalence func + diff --git a/Cubical/Categories/Equivalence/Properties.agda b/Cubical/Categories/Equivalence/Properties.agda new file mode 100644 index 000000000..de92b0ce1 --- /dev/null +++ b/Cubical/Categories/Equivalence/Properties.agda @@ -0,0 +1,92 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Equivalence.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Sigma +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Morphism +open import Cubical.Categories.Equivalence.Base +open import Cubical.HITs.PropositionalTruncation.Base + +open Precategory +open Functor +open NatIso +open CatIso +open NatTrans +open isEquivalence + +private + variable + ℓC ℓC' ℓD ℓD' : Level + +-- Equivalence implies Full, Faithul, and Essentially Surjective + +module _ {C : Precategory ℓC ℓC'} {D : Precategory ℓD ℓD'} where + symEquiv : ∀ {F : Functor C D} → (e : isEquivalence F) → isEquivalence (e .invFunc) + symEquiv {F} record { invFunc = G ; η = η ; ε = ε } = record { invFunc = F ; η = symNatIso ε ; ε = symNatIso η } + + isEquiv→Faithful : ∀ {F : Functor C D} → isEquivalence F → isFaithful F + isEquiv→Faithful {F} record { invFunc = G + ; η = η + ; ε = _ } + c c' f g p = f + ≡⟨ sqRL η ⟩ + cIso .mor ⋆⟨ C ⟩ G ⟪ F ⟪ f ⟫ ⟫ ⋆⟨ C ⟩ c'Iso .inv + ≡⟨ cong (λ v → cIso .mor ⋆⟨ C ⟩ (G ⟪ v ⟫) ⋆⟨ C ⟩ c'Iso .inv) p ⟩ + cIso .mor ⋆⟨ C ⟩ G ⟪ F ⟪ g ⟫ ⟫ ⋆⟨ C ⟩ c'Iso .inv + ≡⟨ sym (sqRL η) ⟩ + g + ∎ + where + + -- isomorphism between c and GFc + cIso = isIso→CatIso (η .nIso c) + -- isomorphism between c' and GFc' + c'Iso = isIso→CatIso (η .nIso c') + +module _ {C : Precategory ℓC ℓC'} {D : Precategory ℓD ℓD'} where + isEquiv→Full : ∀ {F : Functor C D} → isEquivalence F → isFull F + isEquiv→Full {F} eq@record { invFunc = G + ; η = η + ; ε = _ } + c c' g = ∣ h , isEquiv→Faithful (symEquiv eq) _ _ _ _ GFh≡Gg ∣ -- apply faithfulness of G + where + -- isomorphism between c and GFc + cIso = isIso→CatIso (η .nIso c) + -- isomorphism between c' and GFc' + c'Iso = isIso→CatIso (η .nIso c') + + -- reverses + cIso⁻ = symCatIso cIso + c'Iso⁻ = symCatIso c'Iso + + h = cIso .mor ⋆⟨ C ⟩ G ⟪ g ⟫ ⋆⟨ C ⟩ c'Iso .inv + + -- we show that both `G ⟪ g ⟫` and `G ⟪ F ⟪ h ⟫ ⟫` + -- are equal to the same thing + -- namely : cIso .inv ⋆⟨ C ⟩ h ⋆⟨ C ⟩ c'Iso .mor + Gg≡ηhη : G ⟪ g ⟫ ≡ cIso .inv ⋆⟨ C ⟩ h ⋆⟨ C ⟩ c'Iso .mor + Gg≡ηhη = invMoveL cAreInv move-c' ∙ sym (C .⋆Assoc _ _ _) + where + cAreInv : areInv (cIso .mor) (cIso .inv) + cAreInv = CatIso→areInv cIso + + c'AreInv : areInv (c'Iso .mor) (c'Iso .inv) + c'AreInv = CatIso→areInv c'Iso + + move-c' : cIso .mor ⋆⟨ C ⟩ G ⟪ g ⟫ ≡ h ⋆⟨ C ⟩ c'Iso .mor + move-c' = invMoveR (symAreInv c'AreInv) refl + + GFh≡Gg : G ⟪ F ⟪ h ⟫ ⟫ ≡ G ⟪ g ⟫ + GFh≡Gg = G ⟪ F ⟪ h ⟫ ⟫ + ≡⟨ sqLR η ⟩ + cIso .inv ⋆⟨ C ⟩ h ⋆⟨ C ⟩ c'Iso .mor + ≡⟨ sym Gg≡ηhη ⟩ + G ⟪ g ⟫ + ∎ + + isEquiv→Surj : ∀ {F : Functor C D} → isEquivalence F → isEssentiallySurj F + isEquiv→Surj isE d = (isE .invFunc ⟅ d ⟆) , isIso→CatIso ((isE .ε .nIso) d) diff --git a/Cubical/Categories/Everything.agda b/Cubical/Categories/Everything.agda new file mode 100644 index 000000000..176b8ef2c --- /dev/null +++ b/Cubical/Categories/Everything.agda @@ -0,0 +1,20 @@ +{-# OPTIONS --safe --guardedness #-} +module Cubical.Categories.Everything where + +import Cubical.Categories.Adjoint +import Cubical.Categories.Category +import Cubical.Categories.Commutativity +import Cubical.Categories.Constructions.Elements +import Cubical.Categories.Constructions.Slice +import Cubical.Categories.Equivalence +import Cubical.Categories.Functor +import Cubical.Categories.Instances.Cospan +import Cubical.Categories.Instances.Functors +import Cubical.Categories.Instances.Sets +import Cubical.Categories.Limits +import Cubical.Categories.Morphism +import Cubical.Categories.NaturalTransformation +import Cubical.Categories.Presheaf +import Cubical.Categories.Sets +import Cubical.Categories.TypesOfCategories.TypeCategory +import Cubical.Categories.Yoneda diff --git a/Cubical/Categories/Functor.agda b/Cubical/Categories/Functor.agda new file mode 100644 index 000000000..f985a8b1b --- /dev/null +++ b/Cubical/Categories/Functor.agda @@ -0,0 +1,7 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Functor where + +open import Cubical.Categories.Functor.Base public +open import Cubical.Categories.Functor.Properties public +open import Cubical.Categories.Functor.Compose public diff --git a/Cubical/Categories/Functor/Base.agda b/Cubical/Categories/Functor/Base.agda new file mode 100644 index 000000000..59eb7a227 --- /dev/null +++ b/Cubical/Categories/Functor/Base.agda @@ -0,0 +1,89 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Functor.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Sigma +open import Cubical.Categories.Category + +private + variable + ℓC ℓC' ℓD ℓD' : Level + +record Functor (C : Precategory ℓC ℓC') (D : Precategory ℓD ℓD') : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where + no-eta-equality + open Precategory + + field + F-ob : C .ob → D .ob + F-hom : {x y : C .ob} → C [ x , y ] → D [(F-ob x) , (F-ob y)] + F-id : {x : C .ob} → F-hom (C .id x) ≡ D .id (F-ob x) + F-seq : {x y z : C .ob} (f : C [ x , y ]) (g : C [ y , z ]) → F-hom (f ⋆⟨ C ⟩ g) ≡ (F-hom f) ⋆⟨ D ⟩ (F-hom g) + + isFull = (x y : _) (F[f] : D [(F-ob x) , (F-ob y)]) → ∃ (C [ x , y ]) (λ f → F-hom f ≡ F[f]) + isFaithful = (x y : _) (f g : C [ x , y ]) → F-hom f ≡ F-hom g → f ≡ g + isEssentiallySurj = ∀ (d : D .ob) → Σ[ c ∈ C .ob ] CatIso {C = D} (F-ob c) d + +private + variable + ℓ ℓ' : Level + C D E : Precategory ℓ ℓ' + +open Precategory +open Functor + +-- Helpful notation + +-- action on objects +infix 30 _⟅_⟆ +_⟅_⟆ : (F : Functor C D) + → C .ob + → D .ob +_⟅_⟆ = F-ob + +-- action on morphisms +infix 30 _⟪_⟫ -- same infix level as on objects since these will never be used in the same context +_⟪_⟫ : (F : Functor C D) + → ∀ {x y} + → C [ x , y ] + → D [(F ⟅ x ⟆) , (F ⟅ y ⟆)] +_⟪_⟫ = F-hom + + +-- Functor constructions + +𝟙⟨_⟩ : ∀ (C : Precategory ℓ ℓ') → Functor C C +𝟙⟨ C ⟩ .F-ob x = x +𝟙⟨ C ⟩ .F-hom f = f +𝟙⟨ C ⟩ .F-id = refl +𝟙⟨ C ⟩ .F-seq _ _ = refl + +-- functor composition +funcComp : ∀ (G : Functor D E) (F : Functor C D) → Functor C E +(funcComp G F) .F-ob c = G ⟅ F ⟅ c ⟆ ⟆ +(funcComp G F) .F-hom f = G ⟪ F ⟪ f ⟫ ⟫ +(funcComp {D = D} {E = E} {C = C} G F) .F-id {c} + = (G ⟪ F ⟪ C .id c ⟫ ⟫) + ≡⟨ cong (G ⟪_⟫) (F .F-id) ⟩ + G .F-id + -- (G ⟪ D .id (F ⟅ c ⟆) ⟫) -- deleted this cause the extra refl composition was annoying + -- ≡⟨ G .F-id ⟩ + -- E .id (G ⟅ F ⟅ c ⟆ ⟆) + -- ∎ +(funcComp {D = D} {E = E} {C = C} G F) .F-seq {x} {y} {z} f g + = (G ⟪ F ⟪ f ⋆⟨ C ⟩ g ⟫ ⟫) + ≡⟨ cong (G ⟪_⟫) (F .F-seq _ _) ⟩ + G .F-seq _ _ + -- (G ⟪ (F ⟪ f ⟫) ⋆⟨ D ⟩ (F ⟪ g ⟫) ⟫) -- deleted for same reason as above + -- ≡⟨ G .F-seq _ _ ⟩ + -- (G ⟪ F ⟪ f ⟫ ⟫) ⋆⟨ E ⟩ (G ⟪ F ⟪ g ⟫ ⟫) + -- ∎ + +infixr 30 funcComp +syntax funcComp G F = G ∘F F + +_^opF : Functor C D → Functor (C ^op) (D ^op) +(F ^opF) .F-ob = F .F-ob +(F ^opF) .F-hom = F .F-hom +(F ^opF) .F-id = F .F-id +(F ^opF) .F-seq f g = F .F-seq g f diff --git a/Cubical/Categories/Functor/Compose.agda b/Cubical/Categories/Functor/Compose.agda new file mode 100644 index 000000000..e331c4543 --- /dev/null +++ b/Cubical/Categories/Functor/Compose.agda @@ -0,0 +1,47 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Functor.Compose where + +open import Cubical.Foundations.Prelude + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor.Base +open import Cubical.Categories.NaturalTransformation.Base + +open import Cubical.Categories.Instances.Functors + +module _ {ℓC ℓC' ℓD ℓD' ℓE ℓE'} + {C : Precategory ℓC ℓC'} {D : Precategory ℓD ℓD'} + (E : Precategory ℓE ℓE') ⦃ isCatE : isCategory E ⦄ + (F : Functor C D) + where + + open Functor + open NatTrans + + precomposeF : Functor (FUNCTOR D E) (FUNCTOR C E) + precomposeF .F-ob G = funcComp G F + precomposeF .F-hom α .N-ob c = α .N-ob (F .F-ob c) + precomposeF .F-hom α .N-hom f = α .N-hom (F .F-hom f) + precomposeF .F-id = refl + precomposeF .F-seq f g = refl + +module _ {ℓC ℓC' ℓD ℓD' ℓE ℓE'} + (C : Precategory ℓC ℓC') + {D : Precategory ℓD ℓD'} ⦃ isCatD : isCategory D ⦄ + {E : Precategory ℓE ℓE'} ⦃ isCatE : isCategory E ⦄ + (G : Functor D E) + where + + open Functor + open NatTrans + + postcomposeF : Functor (FUNCTOR C D) (FUNCTOR C E) + postcomposeF .F-ob F = funcComp G F + postcomposeF .F-hom α .N-ob c = G. F-hom (α .N-ob c) + postcomposeF .F-hom {F₀} {F₁} α .N-hom {x} {y} f = + sym (G .F-seq (F₀ ⟪ f ⟫) (α ⟦ y ⟧)) + ∙∙ cong (G ⟪_⟫) (α .N-hom f) + ∙∙ G .F-seq (α ⟦ x ⟧) (F₁ ⟪ f ⟫) + postcomposeF .F-id = makeNatTransPath (funExt λ _ → G .F-id) + postcomposeF .F-seq f g = makeNatTransPath (funExt λ _ → G .F-seq _ _) diff --git a/Cubical/Categories/Functor/Properties.agda b/Cubical/Categories/Functor/Properties.agda new file mode 100644 index 000000000..5073339a3 --- /dev/null +++ b/Cubical/Categories/Functor/Properties.agda @@ -0,0 +1,116 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Functor.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function renaming (_∘_ to _◍_) +open import Cubical.Foundations.GroupoidLaws using (lUnit; rUnit; assoc; cong-∙) +open import Cubical.Categories.Category +open import Cubical.Categories.Functor.Base + +private + variable + ℓ ℓ' ℓ'' : Level + B C D E : Precategory ℓ ℓ' + +open Precategory +open Functor + +{- +x ---p--- x' + ⇓ᵍ + g x' ---q--- y + ⇓ʰ + h y ---r--- z + +The path from `h (g x) ≡ z` obtained by + 1. first applying cong to p and composing with q; then applying cong again and composing with r + 2. first applying cong to q and composing with r; then applying a double cong to p and precomposing +are path equal. +-} +congAssoc : ∀ {X : Type ℓ} {Y : Type ℓ'} {Z : Type ℓ''} (g : X → Y) (h : Y → Z) {x x' : X} {y : Y} {z : Z} + → (p : x ≡ x') (q : g x' ≡ y) (r : h y ≡ z) + → cong (h ◍ g) p ∙ (cong h q ∙ r) ≡ cong h (cong g p ∙ q) ∙ r +congAssoc g h p q r + = cong (h ◍ g) p ∙ (cong h q ∙ r) + ≡⟨ assoc _ _ _ ⟩ + ((cong (h ◍ g) p) ∙ (cong h q)) ∙ r + ≡⟨ refl ⟩ + (cong h (cong g p) ∙ (cong h q)) ∙ r + ≡⟨ cong (_∙ r) (sym (cong-∙ h _ _)) ⟩ + cong h (cong g p ∙ q) ∙ r + ∎ + +-- composition is associative +F-assoc : {F : Functor B C} {G : Functor C D} {H : Functor D E} + → H ∘F (G ∘F F) ≡ (H ∘F G) ∘F F +F-assoc {F = F} {G} {H} i .F-ob x = H ⟅ G ⟅ F ⟅ x ⟆ ⟆ ⟆ +F-assoc {F = F} {G} {H} i .F-hom f = H ⟪ G ⟪ F ⟪ f ⟫ ⟫ ⟫ +F-assoc {F = F} {G} {H} i .F-id {x} = congAssoc (G ⟪_⟫) (H ⟪_⟫) (F .F-id {x}) (G .F-id {F ⟅ x ⟆}) (H .F-id) (~ i) +F-assoc {F = F} {G} {H} i .F-seq f g = congAssoc (G ⟪_⟫) (H ⟪_⟫) (F .F-seq f g) (G .F-seq _ _) (H .F-seq _ _) (~ i) + +-- Results about functors + +module _ {F : Functor C D} where + + -- the identity is the identity + F-lUnit : F ∘F 𝟙⟨ C ⟩ ≡ F + F-lUnit i .F-ob x = F ⟅ x ⟆ + F-lUnit i .F-hom f = F ⟪ f ⟫ + F-lUnit i .F-id {x} = lUnit (F .F-id) (~ i) + F-lUnit i .F-seq f g = lUnit (F .F-seq f g) (~ i) + + F-rUnit : 𝟙⟨ D ⟩ ∘F F ≡ F + F-rUnit i .F-ob x = F ⟅ x ⟆ + F-rUnit i .F-hom f = F ⟪ f ⟫ + F-rUnit i .F-id {x} = rUnit (F .F-id) (~ i) + F-rUnit i .F-seq f g = rUnit (F .F-seq f g) (~ i) + + -- functors preserve commutative diagrams (specificallysqures here) + preserveCommF : ∀ {x y z w} {f : C [ x , y ]} {g : C [ y , w ]} {h : C [ x , z ]} {k : C [ z , w ]} + → f ⋆⟨ C ⟩ g ≡ h ⋆⟨ C ⟩ k + → (F ⟪ f ⟫) ⋆⟨ D ⟩ (F ⟪ g ⟫) ≡ (F ⟪ h ⟫) ⋆⟨ D ⟩ (F ⟪ k ⟫) + preserveCommF {f = f} {g = g} {h = h} {k = k} eq + = (F ⟪ f ⟫) ⋆⟨ D ⟩ (F ⟪ g ⟫) + ≡⟨ sym (F .F-seq _ _) ⟩ + F ⟪ f ⋆⟨ C ⟩ g ⟫ + ≡⟨ cong (F ⟪_⟫) eq ⟩ + F ⟪ h ⋆⟨ C ⟩ k ⟫ + ≡⟨ F .F-seq _ _ ⟩ + (F ⟪ h ⟫) ⋆⟨ D ⟩ (F ⟪ k ⟫) + ∎ + + -- functors preserve isomorphisms + preserveIsosF : ∀ {x y} → CatIso {C = C} x y → CatIso {C = D} (F ⟅ x ⟆) (F ⟅ y ⟆) + preserveIsosF {x} {y} (catiso f f⁻¹ sec' ret') = + catiso + g g⁻¹ + -- sec + ( (g⁻¹ ⋆⟨ D ⟩ g) + ≡⟨ sym (F .F-seq f⁻¹ f) ⟩ + F ⟪ f⁻¹ ⋆⟨ C ⟩ f ⟫ + ≡⟨ cong (F .F-hom) sec' ⟩ + F ⟪ C .id y ⟫ + ≡⟨ F .F-id ⟩ + D .id y' + ∎ ) + -- ret + ( (g ⋆⟨ D ⟩ g⁻¹) + ≡⟨ sym (F .F-seq f f⁻¹) ⟩ + F ⟪ f ⋆⟨ C ⟩ f⁻¹ ⟫ + ≡⟨ cong (F .F-hom) ret' ⟩ + F ⟪ C .id x ⟫ + ≡⟨ F .F-id ⟩ + D .id x' + ∎ ) + + where + x' : D .ob + x' = F ⟅ x ⟆ + y' : D .ob + y' = F ⟅ y ⟆ + + g : D [ x' , y' ] + g = F ⟪ f ⟫ + g⁻¹ : D [ y' , x' ] + g⁻¹ = F ⟪ f⁻¹ ⟫ diff --git a/Cubical/Categories/Instances/Cospan.agda b/Cubical/Categories/Instances/Cospan.agda new file mode 100644 index 000000000..999bb3589 --- /dev/null +++ b/Cubical/Categories/Instances/Cospan.agda @@ -0,0 +1,60 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Instances.Cospan where + +open import Cubical.Foundations.Prelude +open import Cubical.Categories.Category +open import Cubical.Data.Unit +open import Cubical.Data.Empty + +open Precategory + +data 𝟛 : Type ℓ-zero where + ⓪ : 𝟛 + ① : 𝟛 + ② : 𝟛 + +CospanCat : Precategory ℓ-zero ℓ-zero +CospanCat .ob = 𝟛 + +CospanCat .Hom[_,_] ⓪ ① = Unit +CospanCat .Hom[_,_] ② ① = Unit +CospanCat .Hom[_,_] ⓪ ⓪ = Unit +CospanCat .Hom[_,_] ① ① = Unit +CospanCat .Hom[_,_] ② ② = Unit +CospanCat .Hom[_,_] _ _ = ⊥ + + +CospanCat ._⋆_ {x = ⓪} {⓪} {⓪} f g = tt +CospanCat ._⋆_ {x = ①} {①} {①} f g = tt +CospanCat ._⋆_ {x = ②} {②} {②} f g = tt +CospanCat ._⋆_ {x = ⓪} {①} {①} f g = tt +CospanCat ._⋆_ {x = ②} {①} {①} f g = tt +CospanCat ._⋆_ {x = ⓪} {⓪} {①} f g = tt +CospanCat ._⋆_ {x = ②} {②} {①} f g = tt + +CospanCat .id ⓪ = tt +CospanCat .id ① = tt +CospanCat .id ② = tt + +CospanCat .⋆IdL {⓪} {①} _ = refl +CospanCat .⋆IdL {②} {①} _ = refl +CospanCat .⋆IdL {⓪} {⓪} _ = refl +CospanCat .⋆IdL {①} {①} _ = refl +CospanCat .⋆IdL {②} {②} _ = refl + +CospanCat .⋆IdR {⓪} {①} _ = refl +CospanCat .⋆IdR {②} {①} _ = refl +CospanCat .⋆IdR {⓪} {⓪} _ = refl +CospanCat .⋆IdR {①} {①} _ = refl +CospanCat .⋆IdR {②} {②} _ = refl + +CospanCat .⋆Assoc {⓪} {⓪} {⓪} {⓪} _ _ _ = refl +CospanCat .⋆Assoc {⓪} {⓪} {⓪} {①} _ _ _ = refl +CospanCat .⋆Assoc {⓪} {⓪} {①} {①} _ _ _ = refl +CospanCat .⋆Assoc {⓪} {①} {①} {①} _ _ _ = refl +CospanCat .⋆Assoc {①} {①} {①} {①} _ _ _ = refl +CospanCat .⋆Assoc {②} {②} {②} {②} _ _ _ = refl +CospanCat .⋆Assoc {②} {②} {②} {①} _ _ _ = refl +CospanCat .⋆Assoc {②} {②} {①} {①} _ _ _ = refl +CospanCat .⋆Assoc {②} {①} {①} {①} _ _ _ = refl diff --git a/Cubical/Categories/Instances/Functors.agda b/Cubical/Categories/Instances/Functors.agda new file mode 100644 index 000000000..9d03f4436 --- /dev/null +++ b/Cubical/Categories/Instances/Functors.agda @@ -0,0 +1,59 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Instances.Functors where + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor.Base +open import Cubical.Categories.NaturalTransformation.Base +open import Cubical.Categories.NaturalTransformation.Properties +open import Cubical.Categories.Morphism renaming (isIso to isIsoC) +open import Cubical.Foundations.Prelude + +private + variable + ℓC ℓC' ℓD ℓD' : Level + +module _ (C : Precategory ℓC ℓC') (D : Precategory ℓD ℓD') ⦃ isCatD : isCategory D ⦄ where + open Precategory + open NatTrans + open Functor + + FUNCTOR : Precategory (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) (ℓ-max (ℓ-max ℓC ℓC') ℓD') + FUNCTOR .ob = Functor C D + FUNCTOR .Hom[_,_] = NatTrans + FUNCTOR .id = idTrans + FUNCTOR ._⋆_ = seqTrans + FUNCTOR .⋆IdL α = makeNatTransPath λ i x → D .⋆IdL (α .N-ob x) i + FUNCTOR .⋆IdR α = makeNatTransPath λ i x → D .⋆IdR (α .N-ob x) i + FUNCTOR .⋆Assoc α β γ = makeNatTransPath λ i x → D .⋆Assoc (α .N-ob x) (β .N-ob x) (γ .N-ob x) i + + isCatFUNCTOR : isCategory FUNCTOR + isCatFUNCTOR .isSetHom = isSetNat + + open isIsoC renaming (inv to invC) + -- component wise iso is an iso in Functor + FUNCTORIso : ∀ {F G : Functor C D} (α : F ⇒ G) + → (∀ (c : C .ob) → isIsoC {C = D} (α ⟦ c ⟧)) + → isIsoC {C = FUNCTOR} α + FUNCTORIso α is .invC .N-ob c = (is c) .invC + FUNCTORIso {F} {G} α is .invC .N-hom {c} {d} f + = invMoveL areInv-αc + ( α ⟦ c ⟧ ⋆⟨ D ⟩ (G ⟪ f ⟫ ⋆⟨ D ⟩ is d .invC) + ≡⟨ sym (D .⋆Assoc _ _ _) ⟩ + (α ⟦ c ⟧ ⋆⟨ D ⟩ G ⟪ f ⟫) ⋆⟨ D ⟩ is d .invC + ≡⟨ sym (invMoveR areInv-αd (α .N-hom f)) ⟩ + F ⟪ f ⟫ + ∎ ) + where + areInv-αc : areInv (α ⟦ c ⟧) ((is c) .invC) + areInv-αc = isIso→areInv (is c) + + areInv-αd : areInv (α ⟦ d ⟧) ((is d) .invC) + areInv-αd = isIso→areInv (is d) + FUNCTORIso α is .sec = makeNatTransPath (funExt (λ c → (is c) .sec)) + FUNCTORIso α is .ret = makeNatTransPath (funExt (λ c → (is c) .ret)) + +instance + ⦃isCatFUNCTOR⦄ : {C : Precategory ℓC ℓC'} {D : Precategory ℓD ℓD'} ⦃ isCatD : isCategory D ⦄ + → isCategory (FUNCTOR C D) + ⦃isCatFUNCTOR⦄ {C = C} {D} = isCatFUNCTOR C D diff --git a/Cubical/Categories/Instances/Sets.agda b/Cubical/Categories/Instances/Sets.agda new file mode 100644 index 000000000..5347b5311 --- /dev/null +++ b/Cubical/Categories/Instances/Sets.agda @@ -0,0 +1,106 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Instances.Sets where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation + +open Precategory + +module _ ℓ where + SET : Precategory (ℓ-suc ℓ) ℓ + SET .ob = Σ (Type ℓ) isSet + SET .Hom[_,_] (A , _) (B , _) = A → B + SET .id _ = λ x → x + SET ._⋆_ f g = λ x → g (f x) + SET .⋆IdL f = refl + SET .⋆IdR f = refl + SET .⋆Assoc f g h = refl + +module _ {ℓ} where + isSetExpIdeal : {A B : Type ℓ} → isSet B → isSet (A → B) + isSetExpIdeal B/set = isSetΠ λ _ → B/set + + isSetLift : {A : Type ℓ} → isSet A → isSet (Lift {ℓ} {ℓ-suc ℓ} A) + isSetLift = isOfHLevelLift 2 + + module _ {A B : SET ℓ .ob} where + -- monic/surjectiveness + open import Cubical.Categories.Morphism + isSurjSET : (f : SET ℓ [ A , B ]) → Type _ + isSurjSET f = ∀ (b : fst B) → Σ[ a ∈ fst A ] f a ≡ b + + -- isMonic→isSurjSET : {f : SET ℓ [ A , B ]} + -- → isEpic {C = SET ℓ} {x = A} {y = B} f + -- → isSurjSET f + -- isMonic→isSurjSET ism b = {!!} , {!!} + + instance + SET-category : isCategory (SET ℓ) + SET-category .isSetHom {_} {B , B/set} = isSetExpIdeal B/set + +private + variable + ℓ ℓ' : Level + +open Functor + +-- Hom functors +_[-,_] : (C : Precategory ℓ ℓ') → (c : C .ob) → ⦃ isCat : isCategory C ⦄ → Functor (C ^op) (SET ℓ') +(C [-, c ]) ⦃ isCat ⦄ .F-ob x = (C [ x , c ]) , isCat .isSetHom +(C [-, c ]) .F-hom f k = f ⋆⟨ C ⟩ k +(C [-, c ]) .F-id = funExt λ _ → C .⋆IdL _ +(C [-, c ]) .F-seq _ _ = funExt λ _ → C .⋆Assoc _ _ _ + +_[_,-] : (C : Precategory ℓ ℓ') → (c : C .ob) → ⦃ isCat : isCategory C ⦄ → Functor C (SET ℓ') +(C [ c ,-]) ⦃ isCat ⦄ .F-ob x = (C [ c , x ]) , isCat .isSetHom +(C [ c ,-]) .F-hom f k = k ⋆⟨ C ⟩ f +(C [ c ,-]) .F-id = funExt λ _ → C .⋆IdR _ +(C [ c ,-]) .F-seq _ _ = funExt λ _ → sym (C .⋆Assoc _ _ _) + +module _ {C : Precategory ℓ ℓ'} ⦃ _ : isCategory C ⦄ {F : Functor C (SET ℓ')} where + open NatTrans + + -- natural transformations by pre/post composition + preComp : {x y : C .ob} + → (f : C [ x , y ]) + → C [ x ,-] ⇒ F + → C [ y ,-] ⇒ F + preComp f α .N-ob c k = (α ⟦ c ⟧) (f ⋆⟨ C ⟩ k) + preComp f α .N-hom {x = c} {d} k + = (λ l → (α ⟦ d ⟧) (f ⋆⟨ C ⟩ (l ⋆⟨ C ⟩ k))) + ≡[ i ]⟨ (λ l → (α ⟦ d ⟧) (⋆Assoc C f l k (~ i))) ⟩ + (λ l → (α ⟦ d ⟧) (f ⋆⟨ C ⟩ l ⋆⟨ C ⟩ k)) + ≡[ i ]⟨ (λ l → (α .N-hom k) i (f ⋆⟨ C ⟩ l)) ⟩ + (λ l → (F ⟪ k ⟫) ((α ⟦ c ⟧) (f ⋆⟨ C ⟩ l))) + ∎ + +-- properties +-- TODO: move to own file +open CatIso renaming (inv to cInv) +open Iso + +Iso→CatIso : ∀ {A B : (SET ℓ) .ob} + → Iso (fst A) (fst B) + → CatIso {C = SET ℓ} A B +Iso→CatIso is .mor = is .fun +Iso→CatIso is .cInv = is .inv +Iso→CatIso is .sec = funExt λ b → is .rightInv b -- is .rightInv +Iso→CatIso is .ret = funExt λ b → is .leftInv b -- is .rightInv + + +-- TYPE category is has all types as objects +-- kind of useless +module _ ℓ where + TYPE : Precategory (ℓ-suc ℓ) ℓ + TYPE .ob = Type ℓ + TYPE .Hom[_,_] A B = A → B + TYPE .id A = λ x → x + TYPE ._⋆_ f g = λ x → g (f x) + TYPE .⋆IdL f = refl + TYPE .⋆IdR f = refl + TYPE .⋆Assoc f g h = refl diff --git a/Cubical/Categories/Limits.agda b/Cubical/Categories/Limits.agda new file mode 100644 index 000000000..6d2b58f3e --- /dev/null +++ b/Cubical/Categories/Limits.agda @@ -0,0 +1,7 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Limits where + +open import Cubical.Categories.Limits.Base public +open import Cubical.Categories.Limits.Terminal public +open import Cubical.Categories.Limits.Pullback public diff --git a/Cubical/Categories/Limits/Base.agda b/Cubical/Categories/Limits/Base.agda new file mode 100644 index 000000000..6eb20af4c --- /dev/null +++ b/Cubical/Categories/Limits/Base.agda @@ -0,0 +1,160 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Limits.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Unit +open import Cubical.Data.Sigma using (ΣPathP) +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Instances.Sets + +private + variable + ℓJ ℓJ' ℓC ℓC' : Level + ℓ ℓ' ℓ'' : Level + +module _ {J : Precategory ℓJ ℓJ'} + {C : Precategory ℓC ℓC'} where + open Precategory + open Functor + open NatTrans + + -- functor which sends all objects to c and all + -- morphisms to the identity + constF : (c : C .ob) → Functor J C + constF c .F-ob _ = c + constF c .F-hom _ = C .id c + constF c .F-id = refl + constF c .F-seq _ _ = sym (C .⋆IdL _) + + + -- a cone is a natural transformation from the constant functor at x to K + Cone : (K : Functor J C) → C .ob → Type _ + Cone K x = NatTrans (constF x) K + + + module _ {K : Functor J C} where + + -- precomposes a cone with a morphism + _◼_ : ∀ {d c : C .ob} + → (f : C [ d , c ]) + → Cone K c + → Cone K d + (f ◼ μ) .N-ob x = f ⋆⟨ C ⟩ μ ⟦ x ⟧ + (f ◼ μ) .N-hom {x = x} {y} k + = C .id _ ⋆⟨ C ⟩ (f ⋆⟨ C ⟩ μ ⟦ y ⟧) + ≡⟨ C .⋆IdL _ ⟩ + f ⋆⟨ C ⟩ μ ⟦ y ⟧ + ≡⟨ cong (λ v → f ⋆⟨ C ⟩ v) (sym (C .⋆IdL _)) ⟩ + f ⋆⟨ C ⟩ (C .id _ ⋆⟨ C ⟩ μ ⟦ y ⟧) + ≡⟨ cong (λ v → f ⋆⟨ C ⟩ v) (μ .N-hom k) ⟩ + f ⋆⟨ C ⟩ (μ ⟦ x ⟧ ⋆⟨ C ⟩ K ⟪ k ⟫) + ≡⟨ sym (C .⋆Assoc _ _ _) ⟩ + f ⋆⟨ C ⟩ μ ⟦ x ⟧ ⋆⟨ C ⟩ K ⟪ k ⟫ + ∎ + + -- μ factors ν if there's a morphism such that the natural transformation + -- from precomposing it with ν gives you back μ + _factors_ : {u v : C .ob} (μ : Cone K u) (ν : Cone K v) → Type _ + _factors_ {u} {v} μ ν = Σ[ mor ∈ C [ v , u ] ] ν ≡ (mor ◼ μ) + + -- μ uniquely factors ν if the factor from above is contractible + _uniquelyFactors_ : {u v : C .ob} (μ : Cone K u) (ν : Cone K v) → Type _ + _uniquelyFactors_ {u} {v} μ ν = isContr (μ factors ν) + + module _ (K : Functor J C) where + + -- a Limit is a cone such that all cones are uniquely factored by it + record isLimit (head : C .ob) : Type (ℓ-max (ℓ-max ℓJ ℓJ') (ℓ-max ℓC ℓC')) where + field + cone : Cone K head + -- TODO: change this to terminal object in category of Cones? + up : ∀ {v} (ν : Cone K v) → cone uniquelyFactors ν + + record Limit : Type (ℓ-max (ℓ-max ℓJ ℓJ') (ℓ-max ℓC ℓC')) where + field + head : C .ob + islim : isLimit head + +-- a Category is complete if it has all limits +complete' : {ℓJ ℓJ' : Level} (C : Precategory ℓC ℓC') → Type _ +complete' {ℓJ = ℓJ} {ℓJ'} C = (J : Precategory ℓJ ℓJ') (K : Functor J C) → Limit K + +complete : (C : Precategory ℓC ℓC') → Typeω +complete C = ∀ {ℓJ ℓJ'} → complete' {ℓJ = ℓJ} {ℓJ'} C + +open Limit +open NatTrans +open Precategory + + + +-- TODO: + +-- 1. every diagram has limits isomorphic to the limit of an equalizer of products + +-- 2. every equalizer can be made into a pullback + +-- 3. every product can be made into an equalizer + +-- 4. a category with all pullbacks and a terminal object has all limits + + +-- SET is complete + +-- notes: +-- didn't need to restrict to *finite* diagrams , why is that required in Set theoretic? +-- didn't use coinduction here because Agda didn't like me referencing 'cone' frome 'up' (termination check) + +isCompleteSET : ∀ {ℓJ ℓJ'} → complete' {ℓJ = ℓJ} {ℓJ'} (SET (ℓ-max ℓJ ℓJ')) +isCompleteSET J K = record + { head = head' + ; islim = record { cone = cone' ; up = up' } } + where + -- the limit is defined as the Set of all cones with head Unit + head' = Cone K (Unit* , isOfHLevelLift 2 isSetUnit) , isSetNat + + -- the legs are defined by taking a cone to its component at j + cone' : Cone K head' + cone' .N-ob j μ = (μ ⟦ j ⟧) tt* + -- Naturality follows from naturality of the Unit cone + cone' .N-hom {x = i} {j} f + = funExt λ μ → (μ ⟦ j ⟧) tt* + ≡[ i ]⟨ (μ .N-hom f i) tt* ⟩ + (K ⟪ f ⟫) ((μ ⟦ i ⟧) tt*) + ∎ + + -- Given another cone α, we want a unique function f from α → cone' which factors it + -- factorization property enforces that (cone' ⟦ j ⟧ ● f) ≡ α ⟦ j ⟧ + -- but cone' ⟦ j ⟧ simply takes the jth component the output Cone K Unit from f + -- so this enforces that for all x ∈ A, (f x) ⟦ j ⟧ ≡ α ⟦ j ⟧ x + -- this determines the *only* possible factoring morphism + up' : ∀ {A} (α : Cone K A) → cone' uniquelyFactors α + up' {A} α = (f , fact) , unique + where + f : fst A → Cone K (Unit* , isOfHLevelLift 2 isSetUnit) + f x = natTrans (λ j _ → α .N-ob j x) + (λ {m} {n} f → funExt λ μ i → α .N-hom f i x) + + fact : α ≡ (f ◼ cone') + fact = makeNatTransPath refl -- I LOVE DEFINITIONAL EQUALITY + + unique : (τ : cone' factors α) → (f , fact) ≡ τ + unique (f' , fact') = ΣPathP (f≡f' , fact≡fact') + where + f≡f' : f ≡ f' + f≡f' = funExt λ x → makeNatTransPath (funExt λ _ → sym eq2) + where + -- the factorization property enforces that f' must have the same behavior as f + eq1 : ∀ {x j} → ((cone' ⟦ j ⟧) (f' x)) ≡ (α ⟦ j ⟧) x + eq1 {x} {j} i = ((fact' (~ i)) ⟦ j ⟧) x + + eq2 : ∀ {x j} → (f' x) ⟦ j ⟧ ≡ λ _ → (α ⟦ j ⟧) x -- = (f x) ⟦ j ⟧ + eq2 {x} {j} = funExt λ _ → eq1 + + -- follows from Set having homsets + fact≡fact' : PathP (λ i → α ≡ ((f≡f' i) ◼ cone')) fact fact' + fact≡fact' = isOfHLevel→isOfHLevelDep 1 (λ β → isSetNat α β) fact fact' λ i → (f≡f' i) ◼ cone' diff --git a/Cubical/Categories/Limits/Pullback.agda b/Cubical/Categories/Limits/Pullback.agda new file mode 100644 index 000000000..38e0561fe --- /dev/null +++ b/Cubical/Categories/Limits/Pullback.agda @@ -0,0 +1,80 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Limits.Pullback where + +open import Cubical.Foundations.Prelude +open import Cubical.Categories.Limits.Base +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Data.Sigma +open import Cubical.Categories.Instances.Cospan public + + +private + variable + ℓ ℓ' : Level + +module _ {C : Precategory ℓ ℓ'} where + + open Precategory C + open Functor + + record Cospan : Type (ℓ-max ℓ ℓ') where + constructor cospan + field + l m r : ob + s₁ : C [ l , m ] + s₂ : C [ r , m ] + + record PullbackLegs (cspn : Cospan) (c : ob) : Type (ℓ-max ℓ ℓ') where + constructor pblegs + open Cospan cspn + field + p₁ : C [ c , l ] + p₂ : C [ c , r ] + + record PullbackCone (cspn : Cospan) (c : ob) : Type (ℓ-max ℓ ℓ') where + constructor cone + open Cospan cspn + field + pl : PullbackLegs cspn c + open PullbackLegs pl public + field + sq : p₁ ⋆⟨ C ⟩ s₁ ≡ p₂ ⋆⟨ C ⟩ s₂ + + + record isPullback (cspn : _) {c} (pb : PullbackLegs cspn c) : Type (ℓ-max ℓ ℓ') where + open Cospan cspn + open PullbackLegs + field + sq : pb .p₁ ⋆⟨ C ⟩ s₁ ≡ pb .p₂ ⋆⟨ C ⟩ s₂ + + open PullbackCone + field + up : ∀ {d} + → (pb' : PullbackCone cspn d) + → isContr (Σ[ f ∈ (C [ d , c ]) ] ((pb' .p₁ ≡ f ⋆⟨ C ⟩ pb .p₁) + × (pb' .p₂ ≡ f ⋆⟨ C ⟩ pb .p₂))) + + Cospan→Func : Cospan → Functor CospanCat C + Cospan→Func (cospan l m r f g) .F-ob ⓪ = l + Cospan→Func (cospan l m r f g) .F-ob ① = m + Cospan→Func (cospan l m r f g) .F-ob ② = r + Cospan→Func (cospan l m r f g) .F-hom {⓪} {①} k = f + Cospan→Func (cospan l m r f g) .F-hom {②} {①} k = g + Cospan→Func (cospan l m r f g) .F-hom {⓪} {⓪} k = id l + Cospan→Func (cospan l m r f g) .F-hom {①} {①} k = id m + Cospan→Func (cospan l m r f g) .F-hom {②} {②} k = id r + Cospan→Func (cospan l m r f g) .F-id {⓪} = refl + Cospan→Func (cospan l m r f g) .F-id {①} = refl + Cospan→Func (cospan l m r f g) .F-id {②} = refl + Cospan→Func (cospan l m r f g) .F-seq {⓪} {⓪} {⓪} φ ψ = sym (⋆IdL _) + Cospan→Func (cospan l m r f g) .F-seq {⓪} {⓪} {①} φ ψ = sym (⋆IdL _) + Cospan→Func (cospan l m r f g) .F-seq {⓪} {①} {①} φ ψ = sym (⋆IdR _) + Cospan→Func (cospan l m r f g) .F-seq {①} {①} {①} φ ψ = sym (⋆IdL _) + Cospan→Func (cospan l m r f g) .F-seq {②} {②} {②} φ ψ = sym (⋆IdL _) + Cospan→Func (cospan l m r f g) .F-seq {②} {②} {①} φ ψ = sym (⋆IdL _) + Cospan→Func (cospan l m r f g) .F-seq {②} {①} {①} φ ψ = sym (⋆IdR _) + + + -- TODO: show that this definition of Pullback is equivalent to the Cospan limit diff --git a/Cubical/Categories/Limits/Terminal.agda b/Cubical/Categories/Limits/Terminal.agda new file mode 100644 index 000000000..b86e3a137 --- /dev/null +++ b/Cubical/Categories/Limits/Terminal.agda @@ -0,0 +1,25 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Limits.Terminal where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Sigma +-- open import Cubical.Categories.Limits.Base +open import Cubical.Categories.Category +open import Cubical.Categories.Functor + +private + variable + ℓ ℓ' : Level + +module _ (C : Precategory ℓ ℓ') where + open Precategory C + + isInitial : (x : ob) → Type (ℓ-max ℓ ℓ') + isInitial x = ∀ (y : ob) → isContr (C [ x , y ]) + + isFinal : (x : ob) → Type (ℓ-max ℓ ℓ') + isFinal x = ∀ (y : ob) → isContr (C [ y , x ]) + + hasFinalOb : Type (ℓ-max ℓ ℓ') + hasFinalOb = Σ[ x ∈ ob ] isFinal x diff --git a/Cubical/Categories/Morphism.agda b/Cubical/Categories/Morphism.agda new file mode 100644 index 000000000..c075003e3 --- /dev/null +++ b/Cubical/Categories/Morphism.agda @@ -0,0 +1,113 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Morphism where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Sigma +open import Cubical.Categories.Category + +private + variable + ℓ ℓ' : Level + +module _ {C : Precategory ℓ ℓ'} where + open Precategory C + private + variable + x y z w : ob + + isMonic : (Hom[ x , y ]) → Type (ℓ-max ℓ ℓ') + isMonic {x} {y} f = ∀ {z : ob} {a a' : Hom[ z , x ]} + → (f ∘ a ≡ f ∘ a') → (a ≡ a') + + isEpic : (Hom[ x , y ]) → Type (ℓ-max ℓ ℓ') + isEpic {x} {y} g = ∀ {z : ob} {b b' : Hom[ y , z ]} + → (b ∘ g ≡ b' ∘ g) → (b ≡ b') + + -- A morphism is split monic if it has a *retraction* + isSplitMon : (Hom[ x , y ]) → Type ℓ' + isSplitMon {x} {y} f = ∃[ r ∈ Hom[ y , x ] ] r ∘ f ≡ id x + + -- A morphism is split epic if it has a *section* + isSplitEpi : (Hom[ x , y ]) → Type ℓ' + isSplitEpi {x} {y} g = ∃[ s ∈ Hom[ y , x ] ] g ∘ s ≡ id y + + record areInv (f : Hom[ x , y ]) (g : Hom[ y , x ]) : Type ℓ' where + field + sec : g ⋆ f ≡ id y + ret : f ⋆ g ≡ id x + + open areInv + + symAreInv : ∀ {f : Hom[ x , y ]} {g : Hom[ y , x ]} + → areInv f g + → areInv g f + symAreInv record { sec = sec ; ret = ret } = record { sec = ret ; ret = sec } + + -- equational reasoning with inverses + invMoveR : ∀ {f : Hom[ x , y ]} {g : Hom[ y , x ]} {h : Hom[ z , x ]} {k : Hom[ z , y ]} + → areInv f g + → h ⋆ f ≡ k + → h ≡ k ⋆ g + invMoveR {f = f} {g} {h} {k} ai p + = h + ≡⟨ sym (⋆IdR _) ⟩ + (h ⋆ id _) + ≡⟨ cong (h ⋆_) (sym (ai .ret)) ⟩ + (h ⋆ (f ⋆ g)) + ≡⟨ sym (⋆Assoc _ _ _) ⟩ + ((h ⋆ f) ⋆ g) + ≡⟨ cong (_⋆ g) p ⟩ + k ⋆ g + ∎ + + invMoveL : ∀ {f : Hom[ x , y ]} {g : Hom[ y , x ]} {h : Hom[ y , z ]} {k : Hom[ x , z ]} + → areInv f g + → f ⋆ h ≡ k + → h ≡ g ⋆ k + invMoveL {f = f} {g} {h} {k} ai p + = h + ≡⟨ sym (⋆IdL _) ⟩ + id _ ⋆ h + ≡⟨ cong (_⋆ h) (sym (ai .sec)) ⟩ + (g ⋆ f) ⋆ h + ≡⟨ ⋆Assoc _ _ _ ⟩ + g ⋆ (f ⋆ h) + ≡⟨ cong (g ⋆_) p ⟩ + (g ⋆ k) + ∎ + + record isIso (f : Hom[ x , y ]) : Type ℓ' where + field + inv : Hom[ y , x ] + sec : inv ⋆ f ≡ id y + ret : f ⋆ inv ≡ id x + + open isIso + + isIso→areInv : ∀ {f : Hom[ x , y ]} + → (isI : isIso f) + → areInv f (isI .inv) + isIso→areInv record { inv = inv ; sec = sec ; ret = ret } = record { sec = sec ; ret = ret } + + open CatIso + + -- isIso agrees with CatIso + isIso→CatIso : ∀ {f : C [ x , y ]} + → isIso f + → CatIso {C = C} x y + isIso→CatIso {f = f} record { inv = f⁻¹ ; sec = sec ; ret = ret } = catiso f f⁻¹ sec ret + + CatIso→isIso : (cIso : CatIso {C = C} x y) + → isIso (cIso .mor) + CatIso→isIso (catiso mor inv sec ret) = record { inv = inv ; sec = sec ; ret = ret } + + CatIso→areInv : (cIso : CatIso {C = C} x y) + → areInv (cIso .mor) (cIso .inv) + CatIso→areInv cIso = isIso→areInv (CatIso→isIso cIso) + + -- reverse of an iso is also an iso + symCatIso : ∀ {x y} + → CatIso {C = C} x y + → CatIso {C = C} y x + symCatIso (catiso mor inv sec ret) = catiso inv mor ret sec diff --git a/Cubical/Categories/NaturalTransformation.agda b/Cubical/Categories/NaturalTransformation.agda new file mode 100644 index 000000000..ccde33ef6 --- /dev/null +++ b/Cubical/Categories/NaturalTransformation.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.NaturalTransformation where + +open import Cubical.Categories.NaturalTransformation.Base public +open import Cubical.Categories.NaturalTransformation.Properties public diff --git a/Cubical/Categories/NaturalTransformation/Base.agda b/Cubical/Categories/NaturalTransformation/Base.agda new file mode 100644 index 000000000..93f538a63 --- /dev/null +++ b/Cubical/Categories/NaturalTransformation/Base.agda @@ -0,0 +1,222 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.NaturalTransformation.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism renaming (iso to iIso) +open import Cubical.Data.Sigma +open import Cubical.Categories.Category +open import Cubical.Categories.Functor.Base +open import Cubical.Categories.Functor.Properties +open import Cubical.Categories.Commutativity +open import Cubical.Categories.Morphism renaming (isIso to isIsoC) + +private + variable + ℓC ℓC' ℓD ℓD' : Level + +module _ {C : Precategory ℓC ℓC'} {D : Precategory ℓD ℓD'} where + -- syntax for sequencing in category D + infixl 15 _⋆ᴰ_ + private + _⋆ᴰ_ : ∀ {x y z} (f : D [ x , y ]) (g : D [ y , z ]) → D [ x , z ] + f ⋆ᴰ g = f ⋆⟨ D ⟩ g + + open Precategory + open Functor + + -- type aliases because it gets tedious typing it out all the time + N-ob-Type : (F G : Functor C D) → Type _ + N-ob-Type F G = (x : C .ob) → D [(F .F-ob x) , (G .F-ob x)] + + N-hom-Type : (F G : Functor C D) → N-ob-Type F G → Type _ + N-hom-Type F G ϕ = {x y : C .ob} (f : C [ x , y ]) → (F .F-hom f) ⋆ᴰ (ϕ y) ≡ (ϕ x) ⋆ᴰ (G .F-hom f) + + record NatTrans (F G : Functor C D) : Type (ℓ-max (ℓ-max ℓC ℓC') ℓD') where + constructor natTrans + field + -- components of the natural transformation + N-ob : N-ob-Type F G + -- naturality condition + N-hom : N-hom-Type F G N-ob + + record NatIso (F G : Functor C D): Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where + field + trans : NatTrans F G + open NatTrans trans + + field + nIso : ∀ (x : C .ob) → isIsoC {C = D} (N-ob x) + + open isIsoC + + -- the three other commuting squares + sqRL : ∀ {x y : C .ob} {f : C [ x , y ]} + → F ⟪ f ⟫ ≡ (N-ob x) ⋆ᴰ G ⟪ f ⟫ ⋆ᴰ (nIso y) .inv + sqRL {x} {y} {f} = invMoveR (isIso→areInv (nIso y)) (N-hom f) + + sqLL : ∀ {x y : C .ob} {f : C [ x , y ]} + → G ⟪ f ⟫ ⋆ᴰ (nIso y) .inv ≡ (nIso x) .inv ⋆ᴰ F ⟪ f ⟫ + sqLL {x} {y} {f} = invMoveL (isIso→areInv (nIso x)) (sym sqRL') + where + sqRL' : F ⟪ f ⟫ ≡ (N-ob x) ⋆ᴰ ( G ⟪ f ⟫ ⋆ᴰ (nIso y) .inv ) + sqRL' = sqRL ∙ (D .⋆Assoc _ _ _) + + sqLR : ∀ {x y : C .ob} {f : C [ x , y ]} + → G ⟪ f ⟫ ≡ (nIso x) .inv ⋆ᴰ F ⟪ f ⟫ ⋆ᴰ (N-ob y) + sqLR {x} {y} {f} = invMoveR (symAreInv (isIso→areInv (nIso y))) sqLL + + open NatTrans + open NatIso + + infix 10 NatTrans + syntax NatTrans F G = F ⇒ G + + infix 9 NatIso + syntax NatIso F G = F ≅ᶜ G -- c superscript to indicate that this is in the context of categories + + -- component of a natural transformation + infix 30 _⟦_⟧ + _⟦_⟧ : ∀ {F G : Functor C D} → (F ⇒ G) → (x : C .ob) → D [(F .F-ob x) , (G .F-ob x)] + _⟦_⟧ = N-ob + + idTrans : (F : Functor C D) → NatTrans F F + idTrans F .N-ob x = D .id (F .F-ob x) + idTrans F .N-hom f = + (F .F-hom f) ⋆ᴰ (idTrans F .N-ob _) + ≡⟨ D .⋆IdR _ ⟩ + F .F-hom f + ≡⟨ sym (D .⋆IdL _) ⟩ + (D .id (F .F-ob _)) ⋆ᴰ (F .F-hom f) + ∎ + + syntax idTrans F = 1[ F ] + + + -- vertical sequencing + seqTrans : {F G H : Functor C D} (α : NatTrans F G) (β : NatTrans G H) → NatTrans F H + seqTrans α β .N-ob x = (α .N-ob x) ⋆ᴰ (β .N-ob x) + seqTrans {F} {G} {H} α β .N-hom f = + (F .F-hom f) ⋆ᴰ ((α .N-ob _) ⋆ᴰ (β .N-ob _)) + ≡⟨ sym (D .⋆Assoc _ _ _) ⟩ + ((F .F-hom f) ⋆ᴰ (α .N-ob _)) ⋆ᴰ (β .N-ob _) + ≡[ i ]⟨ (α .N-hom f i) ⋆ᴰ (β .N-ob _) ⟩ + ((α .N-ob _) ⋆ᴰ (G .F-hom f)) ⋆ᴰ (β .N-ob _) + ≡⟨ D .⋆Assoc _ _ _ ⟩ + (α .N-ob _) ⋆ᴰ ((G .F-hom f) ⋆ᴰ (β .N-ob _)) + ≡[ i ]⟨ (α .N-ob _) ⋆ᴰ (β .N-hom f i) ⟩ + (α .N-ob _) ⋆ᴰ ((β .N-ob _) ⋆ᴰ (H .F-hom f)) + ≡⟨ sym (D .⋆Assoc _ _ _) ⟩ + ((α .N-ob _) ⋆ᴰ (β .N-ob _)) ⋆ᴰ (H .F-hom f) + ∎ + + compTrans : {F G H : Functor C D} (β : NatTrans G H) (α : NatTrans F G) → NatTrans F H + compTrans β α = seqTrans α β + + infixl 8 seqTrans + syntax seqTrans α β = α ●ᵛ β + + + -- vertically sequence natural transformations whose + -- common functor is not definitional equal + seqTransP : {F G G' H : Functor C D} (p : G ≡ G') + → (α : NatTrans F G) (β : NatTrans G' H) + → NatTrans F H + seqTransP {F} {G} {G'} {H} p α β .N-ob x + -- sequence morphisms with non-judgementally equal (co)domain + = seqP {C = D} {p = Gx≡G'x} (α ⟦ x ⟧) (β ⟦ x ⟧) + where + Gx≡G'x : ∀ {x} → G ⟅ x ⟆ ≡ G' ⟅ x ⟆ + Gx≡G'x {x} i = F-ob (p i) x + seqTransP {F} {G} {G'} {H} p α β .N-hom {x = x} {y} f + -- compose the two commuting squares + -- 1. α's commuting square + -- 2. β's commuting square, but extended to G since β is only G' ≡> H + = compSq {C = D} (α .N-hom f) βSq + where + -- functor equality implies equality of actions on objects and morphisms + Gx≡G'x : G ⟅ x ⟆ ≡ G' ⟅ x ⟆ + Gx≡G'x i = F-ob (p i) x + + Gy≡G'y : G ⟅ y ⟆ ≡ G' ⟅ y ⟆ + Gy≡G'y i = F-ob (p i) y + + Gf≡G'f : PathP (λ i → D [ Gx≡G'x i , Gy≡G'y i ]) (G ⟪ f ⟫) (G' ⟪ f ⟫) + Gf≡G'f i = p i ⟪ f ⟫ + + -- components of β extended out to Gx and Gy respectively + βx' = subst (λ a → D [ a , H ⟅ x ⟆ ]) (sym Gx≡G'x) (β ⟦ x ⟧) + βy' = subst (λ a → D [ a , H ⟅ y ⟆ ]) (sym Gy≡G'y) (β ⟦ y ⟧) + + -- extensions are equal to originals + βy'≡βy : PathP (λ i → D [ Gy≡G'y i , H ⟅ y ⟆ ]) βy' (β ⟦ y ⟧) + βy'≡βy = symP (toPathP {A = λ i → D [ Gy≡G'y (~ i) , H ⟅ y ⟆ ]} refl) + + βx≡βx' : PathP (λ i → D [ Gx≡G'x (~ i) , H ⟅ x ⟆ ]) (β ⟦ x ⟧) βx' + βx≡βx' = toPathP refl + + -- left wall of square + left : PathP (λ i → D [ Gx≡G'x i , H ⟅ y ⟆ ]) (G ⟪ f ⟫ ⋆⟨ D ⟩ βy') (G' ⟪ f ⟫ ⋆⟨ D ⟩ β ⟦ y ⟧) + left i = Gf≡G'f i ⋆⟨ D ⟩ βy'≡βy i + + -- right wall of square + right : PathP (λ i → D [ Gx≡G'x (~ i) , H ⟅ y ⟆ ]) (β ⟦ x ⟧ ⋆⟨ D ⟩ H ⟪ f ⟫) (βx' ⋆⟨ D ⟩ H ⟪ f ⟫) + right i = βx≡βx' i ⋆⟨ D ⟩ refl {x = H ⟪ f ⟫} i + + -- putting it all together + βSq : G ⟪ f ⟫ ⋆⟨ D ⟩ βy' ≡ βx' ⋆⟨ D ⟩ H ⟪ f ⟫ + βSq i = comp (λ k → D [ Gx≡G'x (~ k) , H ⟅ y ⟆ ]) + (λ j → λ { (i = i0) → left (~ j) + ; (i = i1) → right j }) + (β .N-hom f i) + + + module _ ⦃ isCatD : isCategory D ⦄ {F G : Functor C D} {α β : NatTrans F G} where + open Precategory + open Functor + open NatTrans + + makeNatTransPath : α .N-ob ≡ β .N-ob → α ≡ β + makeNatTransPath p i .N-ob = p i + makeNatTransPath p i .N-hom f = rem i + where + rem : PathP (λ i → (F .F-hom f) ⋆ᴰ (p i _) ≡ (p i _) ⋆ᴰ (G .F-hom f)) (α .N-hom f) (β .N-hom f) + rem = toPathP (isCatD .isSetHom _ _ _ _) + + module _ ⦃ isCatD : isCategory D ⦄ {F F' G G' : Functor C D} + {α : NatTrans F G} + {β : NatTrans F' G'} where + open Precategory + open Functor + open NatTrans + makeNatTransPathP : ∀ (p : F ≡ F') (q : G ≡ G') + → PathP (λ i → (x : C .ob) → D [ (p i) .F-ob x , (q i) .F-ob x ]) (α .N-ob) (β .N-ob) + → PathP (λ i → NatTrans (p i) (q i)) α β + makeNatTransPathP p q P i .N-ob = P i + makeNatTransPathP p q P i .N-hom f = rem i + where + rem : PathP (λ i → ((p i) .F-hom f) ⋆ᴰ (P i _) ≡ (P i _) ⋆ᴰ ((q i) .F-hom f)) (α .N-hom f) (β .N-hom f) + rem = toPathP (isCatD .isSetHom _ _ _ _) + + + +private + variable + ℓA ℓA' ℓB ℓB' : Level + +module _ {B : Precategory ℓB ℓB'} {C : Precategory ℓC ℓC'} {D : Precategory ℓD ℓD'} where + open NatTrans + -- whiskering + -- αF + _∘ˡ_ : ∀ {G H : Functor C D} (α : NatTrans G H) → (F : Functor B C) + → NatTrans (G ∘F F) (H ∘F F) + (_∘ˡ_ {G} {H} α F) .N-ob x = α ⟦ F ⟅ x ⟆ ⟧ + (_∘ˡ_ {G} {H} α F) .N-hom f = (α .N-hom) _ + + -- Kβ + _∘ʳ_ : ∀ (K : Functor C D) → {G H : Functor B C} (β : NatTrans G H) + → NatTrans (K ∘F G) (K ∘F H) + (_∘ʳ_ K {G} {H} β) .N-ob x = K ⟪ β ⟦ x ⟧ ⟫ + (_∘ʳ_ K {G} {H} β) .N-hom f = preserveCommF {C = C} {D = D} {K} (β .N-hom f) diff --git a/Cubical/Categories/NaturalTransformation/Properties.agda b/Cubical/Categories/NaturalTransformation/Properties.agda new file mode 100644 index 000000000..2ba0eacc4 --- /dev/null +++ b/Cubical/Categories/NaturalTransformation/Properties.agda @@ -0,0 +1,146 @@ + +{-# OPTIONS --safe #-} + +module Cubical.Categories.NaturalTransformation.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism renaming (iso to iIso) +open import Cubical.Data.Sigma +open import Cubical.Categories.Category +open import Cubical.Categories.Functor.Base +open import Cubical.Categories.Morphism renaming (isIso to isIsoC) +open import Cubical.Categories.NaturalTransformation.Base + +private + variable + ℓC ℓC' ℓD ℓD' : Level + +open isIsoC +open NatIso +open NatTrans +open Precategory +open Functor + +module _ {C : Precategory ℓC ℓC'} {D : Precategory ℓD ℓD'} where + private + _⋆ᴰ_ : ∀ {x y z} (f : D [ x , y ]) (g : D [ y , z ]) → D [ x , z ] + f ⋆ᴰ g = f ⋆⟨ D ⟩ g + + -- natural isomorphism is symmetric + symNatIso : ∀ {F G : Functor C D} + → F ≅ᶜ G + → G ≅ᶜ F + symNatIso η .trans .N-ob x = η .nIso x .inv + symNatIso η .trans .N-hom _ = sqLL η + symNatIso η .nIso x .inv = η .trans .N-ob x + symNatIso η .nIso x .sec = η .nIso x .ret + symNatIso η .nIso x .ret = η .nIso x .sec + + -- Properties + + -- path helpers + module NatTransP where + + module _ {F G : Functor C D} where + open Iso + + -- same as Sigma version + NatTransΣ = Σ[ ob ∈ ((x : C .ob) → D [(F .F-ob x) , (G .F-ob x)]) ] + ({x y : _ } (f : C [ x , y ]) → (F .F-hom f) ⋆ᴰ (ob y) ≡ (ob x) ⋆ᴰ (G .F-hom f)) + + NatTransIsoΣ : Iso (NatTrans F G) NatTransΣ + NatTransIsoΣ .fun (natTrans N-ob N-hom) = N-ob , N-hom + NatTransIsoΣ .inv (N-ob , N-hom) = (natTrans N-ob N-hom) + NatTransIsoΣ .rightInv _ = refl + NatTransIsoΣ .leftInv _ = refl + + NatTrans≡Σ = ua (isoToEquiv NatTransIsoΣ) + + -- introducing paths + NatTrans-≡-intro : ∀ {αo βo : N-ob-Type F G} + {αh : N-hom-Type F G αo} + {βh : N-hom-Type F G βo} + → (p : αo ≡ βo) + → PathP (λ i → ({x y : C .ob} (f : C [ x , y ]) → (F .F-hom f) ⋆ᴰ (p i y) ≡ (p i x) ⋆ᴰ (G .F-hom f))) αh βh + → natTrans {F = F} {G} αo αh ≡ natTrans βo βh + NatTrans-≡-intro p q i = natTrans (p i) (q i) + module _ {F G : Functor C D} {α β : NatTrans F G} where + open Iso + private + αOb = α .N-ob + βOb = β .N-ob + αHom = α .N-hom + βHom = β .N-hom + -- path between natural transformations is the same as a pair of paths (between ob and hom) + NTPathIsoPathΣ : Iso (α ≡ β) + (Σ[ p ∈ (αOb ≡ βOb) ] + (PathP (λ i → ({x y : _} (f : _) → F ⟪ f ⟫ ⋆ᴰ (p i y) ≡ (p i x) ⋆ᴰ G ⟪ f ⟫)) + αHom + βHom)) + NTPathIsoPathΣ .fun p = (λ i → p i .N-ob) , (λ i → p i .N-hom) + NTPathIsoPathΣ .inv (po , ph) i = record { N-ob = po i ; N-hom = ph i } + NTPathIsoPathΣ .rightInv pσ = refl + NTPathIsoPathΣ .leftInv p = refl + + NTPath≃PathΣ = isoToEquiv NTPathIsoPathΣ + + NTPath≡PathΣ = ua NTPath≃PathΣ + + module _ ⦃ isCatD : isCategory D ⦄ where + open NatTransP + + -- if the target category has hom Sets, then any natural transformation is a set + isSetNat : ∀ {F G : Functor C D} + → isSet (NatTrans F G) + isSetNat {F} {G} α β p1 p2 i = comp (λ i → NTPath≡PathΣ {F = F} {G} {α} {β} (~ i)) + (λ j → λ {(i = i0) → transport-filler NTPath≡PathΣ p1 (~ j) ; + (i = i1) → transport-filler NTPath≡PathΣ p2 (~ j)}) + (p1Σ≡p2Σ i) + where + αOb = α .N-ob + βOb = β .N-ob + αHom = α .N-hom + βHom = β .N-hom + + -- convert to sigmas so we can reason about constituent paths separately + p1Σ : Σ[ p ∈ (αOb ≡ βOb) ] + (PathP (λ i → ({x y : _} (f : _) → F ⟪ f ⟫ ⋆ᴰ (p i y) ≡ (p i x) ⋆ᴰ G ⟪ f ⟫)) + αHom + βHom) + p1Σ = transport NTPath≡PathΣ p1 + + p2Σ : Σ[ p ∈ (αOb ≡ βOb) ] + (PathP (λ i → ({x y : _} (f : _) → F ⟪ f ⟫ ⋆ᴰ (p i y) ≡ (p i x) ⋆ᴰ G ⟪ f ⟫)) + αHom + βHom) + p2Σ = transport NTPath≡PathΣ p2 + + -- type aliases + typeN-ob = (x : C .ob) → D [(F .F-ob x) , (G .F-ob x)] + typeN-hom : typeN-ob → Type _ + typeN-hom ϕ = {x y : C .ob} (f : C [ x , y ]) → (F .F-hom f) ⋆ᴰ (ϕ y) ≡ (ϕ x) ⋆ᴰ (G .F-hom f) + + -- the Ob function is a set + isSetN-ob : isSet ((x : C .ob) → D [(F .F-ob x) , (G .F-ob x)]) + isSetN-ob = isOfHLevelΠ 2 λ _ → isCatD .isSetHom + + -- the Hom function is a set + isSetN-hom : (ϕ : typeN-ob) → isSet (typeN-hom ϕ) + isSetN-hom γ = isProp→isSet (isPropImplicitΠ λ x → isPropImplicitΠ λ y → isPropΠ λ f → isCatD .isSetHom _ _) + + -- in fact it's a dependent Set, which we need because N-hom depends on N-ob + isSetN-homP : isOfHLevelDep 2 (λ γ → {x y : C .ob} (f : C [ x , y ]) → (F .F-hom f) ⋆ᴰ (γ y) ≡ (γ x) ⋆ᴰ (G .F-hom f)) + isSetN-homP = isOfHLevel→isOfHLevelDep 2 isSetN-hom + + -- components of the equality + p1Ob≡p2Ob : fst p1Σ ≡ fst p2Σ + p1Ob≡p2Ob = isSetN-ob _ _ (fst p1Σ) (fst p2Σ) + + p1Hom≡p2Hom : PathP (λ i → PathP (λ j → typeN-hom (p1Ob≡p2Ob i j)) αHom βHom) + (snd p1Σ) (snd p2Σ) + p1Hom≡p2Hom = isSetN-homP _ _ (snd p1Σ) (snd p2Σ) p1Ob≡p2Ob + + p1Σ≡p2Σ : p1Σ ≡ p2Σ + p1Σ≡p2Σ = ΣPathP (p1Ob≡p2Ob , p1Hom≡p2Hom) diff --git a/Cubical/Categories/Presheaf.agda b/Cubical/Categories/Presheaf.agda new file mode 100644 index 000000000..357f2f670 --- /dev/null +++ b/Cubical/Categories/Presheaf.agda @@ -0,0 +1,8 @@ +{-# OPTIONS --postfix-projections --safe #-} + +module Cubical.Categories.Presheaf where + +open import Cubical.Categories.Presheaf.Base public +open import Cubical.Categories.Presheaf.Properties public + +open import Cubical.Categories.Presheaf.KanExtension public diff --git a/Cubical/Categories/Presheaf/Base.agda b/Cubical/Categories/Presheaf/Base.agda new file mode 100644 index 000000000..074874757 --- /dev/null +++ b/Cubical/Categories/Presheaf/Base.agda @@ -0,0 +1,84 @@ +{-# OPTIONS --postfix-projections --safe #-} + +module Cubical.Categories.Presheaf.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.HITs.PropositionalTruncation + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Instances.Sets +open import Cubical.Categories.Instances.Functors + +module _ {ℓ ℓ'} where + + PreShv : Precategory ℓ ℓ' → (ℓS : Level) → Precategory (ℓ-max (ℓ-max ℓ ℓ') (ℓ-suc ℓS)) (ℓ-max (ℓ-max ℓ ℓ') ℓS) + PreShv C ℓS = FUNCTOR (C ^op) (SET ℓS) + +instance + isCatPreShv : ∀ {ℓ ℓ'} {C : Precategory ℓ ℓ'} {ℓS} + → isCategory (PreShv C ℓS) + isCatPreShv {C = C} {ℓS} = isCatFUNCTOR (C ^op) (SET ℓS) + +private + variable + ℓ ℓ' : Level + +module Yoneda (C : Precategory ℓ ℓ') ⦃ C-cat : isCategory C ⦄ where + open Functor + open NatTrans + open Precategory C + + yo : ob → Functor (C ^op) (SET ℓ') + yo x .F-ob y .fst = C [ y , x ] + yo x .F-ob y .snd = C-cat .isSetHom + yo x .F-hom f g = f ⋆⟨ C ⟩ g + yo x .F-id i f = ⋆IdL f i + yo x .F-seq f g i h = ⋆Assoc g f h i + + YO : Functor C (PreShv C ℓ') + YO .F-ob = yo + YO .F-hom f .N-ob z g = g ⋆⟨ C ⟩ f + YO .F-hom f .N-hom g i h = ⋆Assoc g h f i + YO .F-id = makeNatTransPath λ i _ → λ f → ⋆IdR f i + YO .F-seq f g = makeNatTransPath λ i _ → λ h → ⋆Assoc h f g (~ i) + + + module _ {x} (F : Functor (C ^op) (SET ℓ')) where + yo-yo-yo : NatTrans (yo x) F → F .F-ob x .fst + yo-yo-yo α = α .N-ob _ (id _) + + no-no-no : F .F-ob x .fst → NatTrans (yo x) F + no-no-no a .N-ob y f = F .F-hom f a + no-no-no a .N-hom f = funExt λ g i → F .F-seq g f i a + + yoIso : Iso (NatTrans (yo x) F) (F .F-ob x .fst) + yoIso .Iso.fun = yo-yo-yo + yoIso .Iso.inv = no-no-no + yoIso .Iso.rightInv b i = F .F-id i b + yoIso .Iso.leftInv a = makeNatTransPath (funExt λ _ → funExt rem) + where + rem : ∀ {z} (x₁ : C [ z , x ]) → F .F-hom x₁ (yo-yo-yo a) ≡ (a .N-ob z) x₁ + rem g = + F .F-hom g (yo-yo-yo a) + ≡[ i ]⟨ a .N-hom g (~ i) (id x) ⟩ + a .N-hom g i0 (id x) + ≡[ i ]⟨ a .N-ob _ (⋆IdR g i) ⟩ + (a .N-ob _) g + ∎ + + yoEquiv : NatTrans (yo x) F ≃ F .F-ob x .fst + yoEquiv = isoToEquiv yoIso + + + isFullYO : isFull YO + isFullYO x y F[f] = ∣ yo-yo-yo _ F[f] , yoIso {x} (yo y) .Iso.leftInv F[f] ∣ + + isFaithfulYO : isFaithful YO + isFaithfulYO x y f g p i = + hcomp + (λ j → λ{ (i = i0) → ⋆IdL f j; (i = i1) → ⋆IdL g j}) + (yo-yo-yo _ (p i)) diff --git a/Cubical/Categories/Presheaf/KanExtension.agda b/Cubical/Categories/Presheaf/KanExtension.agda new file mode 100644 index 000000000..da9be180a --- /dev/null +++ b/Cubical/Categories/Presheaf/KanExtension.agda @@ -0,0 +1,341 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} + +{- + Kan extension of a functor C → D to a functor PreShv C ℓ → PreShv D ℓ left or right adjoint to + precomposition. +-} + +module Cubical.Categories.Presheaf.KanExtension where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Functions.FunExtEquiv + +open import Cubical.HITs.SetQuotients + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Adjoint +open import Cubical.Categories.Presheaf.Base +open import Cubical.Categories.Instances.Functors +open import Cubical.Categories.Instances.Sets + +{- + Left Kan extension of a functor C → D to a functor PreShv C ℓ → PreShv D ℓ left adjoint to precomposition. +-} + +module Lan {ℓC ℓC' ℓD ℓD'} ℓS + {C : Precategory ℓC ℓC'} {D : Precategory ℓD ℓD'} + (F : Functor C D) + where + + open Functor + open NatTrans + + private + module C = Precategory C + module D = Precategory D + + {- + We want the category SET ℓ we're mapping into to be large enough that the coend will take presheaves + Cᵒᵖ → Set ℓ to presheaves Dᵒᵖ → Set ℓ, otherwise we get no adjunction with precomposition. + So we must have ℓC,ℓC',ℓD' ≤ ℓ; the parameter ℓS allows ℓ to be larger than their maximum. + -} + ℓ = ℓ-max (ℓ-max (ℓ-max ℓC ℓC') ℓD') ℓS + + module _ (G : Functor (C ^op) (SET ℓ)) where + + -- Definition of the coend + + module _ (d : D.ob) where + + Raw : Type ℓ + Raw = Σ[ c ∈ C.ob ] Σ[ g ∈ D.Hom[ d , F ⟅ c ⟆ ] ] G .F-ob c .fst + + data _≈_ : (u v : Raw) → Type ℓ where + shift : {c c' : C.ob} (g : D.Hom[ d , F ⟅ c ⟆ ]) (f : C.Hom[ c , c' ]) (a : (G ⟅ c' ⟆) .fst) + → (c' , (g D.⋆ F ⟪ f ⟫) , a) ≈ (c , g , (G ⟪ f ⟫) a) + + Quo = Raw / _≈_ + + pattern shift/ g f a i = eq/ _ _ (shift g f a) i + + -- Action of Quo on arrows in D + + mapR : {d d' : D.ob} (h : D.Hom[ d' , d ]) → Quo d → Quo d' + mapR h [ c , g , a ] = [ c , h D.⋆ g , a ] + mapR h (shift/ g f a i) = + hcomp + (λ j → λ + { (i = i0) → [ _ , D.⋆Assoc h g (F ⟪ f ⟫) j , a ] + ; (i = i1) → [ _ , h D.⋆ g , (G ⟪ f ⟫) a ] + }) + (shift/ (h D.⋆ g) f a i) + mapR h (squash/ t u p q i j) = + squash/ (mapR h t) (mapR h u) (cong (mapR h) p) (cong (mapR h) q) i j + + mapRId : (d : D.ob) → mapR (D.id d) ≡ (idfun _) + mapRId d = + funExt (elimProp (λ _ → squash/ _ _) (λ (c , g , a) i → [ c , D.⋆IdL g i , a ])) + + mapR∘ : {d d' d'' : D.ob} + (h' : D.Hom[ d'' , d' ]) (h : D.Hom[ d' , d ]) + → mapR (h' D.⋆ h) ≡ mapR h' ∘ mapR h + mapR∘ h' h = + funExt (elimProp (λ _ → squash/ _ _) (λ (c , g , a) i → [ c , D.⋆Assoc h' h g i , a ])) + + LanOb : Functor (C ^op) (SET ℓ) → Functor (D ^op) (SET _) + LanOb G .F-ob d .fst = Quo G d + LanOb G .F-ob d .snd = squash/ + LanOb G .F-hom = mapR G + LanOb G .F-id {d} = mapRId G d + LanOb G .F-seq h h' = mapR∘ G h' h + + -- Action of Quo on arrows in Cᵒᵖ → Set + + module _ {G G' : Functor (C ^op) (SET ℓ)} (α : NatTrans G G') where + + mapL : (d : D.ob) → Quo G d → Quo G' d + mapL d [ c , g , a ] = [ c , g , α .N-ob c a ] + mapL d (shift/ g f a i) = + hcomp + (λ j → λ + { (i = i0) → [ _ , (g D.⋆ F ⟪ f ⟫) , α .N-ob _ a ] + ; (i = i1) → [ _ , g , funExt⁻ (α .N-hom f) a (~ j) ] + }) + (shift/ g f ((α ⟦ _ ⟧) a) i) + mapL d (squash/ t u p q i j) = + squash/ (mapL d t) (mapL d u) (cong (mapL d) p) (cong (mapL d) q) i j + + mapLR : {d d' : D.ob} (h : D.Hom[ d' , d ]) + → mapL d' ∘ mapR G h ≡ mapR G' h ∘ mapL d + mapLR h = funExt (elimProp (λ _ → squash/ _ _) (λ _ → refl)) + + mapLId : (G : Functor (C ^op) (SET ℓ)) + (d : D.ob) → mapL (idTrans G) d ≡ idfun (Quo G d) + mapLId G d = funExt (elimProp (λ _ → squash/ _ _) (λ _ → refl)) + + mapL∘ : {G G' G'' : Functor (C ^op) (SET ℓ)} + (β : NatTrans G' G'') (α : NatTrans G G') + (d : D.ob) → mapL (seqTrans α β) d ≡ mapL β d ∘ mapL α d + mapL∘ β α d = funExt (elimProp (λ _ → squash/ _ _) (λ _ → refl)) + + LanHom : {G G' : Functor (C ^op) (SET ℓ)} + → NatTrans G G' → NatTrans (LanOb G) (LanOb G') + LanHom α .N-ob = mapL α + LanHom α .N-hom = mapLR α + + -- Definition of the left Kan extension functor + + Lan : Functor (FUNCTOR (C ^op) (SET ℓ)) (FUNCTOR (D ^op) (SET ℓ)) + Lan .F-ob = LanOb + Lan .F-hom = LanHom + Lan .F-id {G} = makeNatTransPath (funExt (mapLId G)) + Lan .F-seq α β = makeNatTransPath (funExt (mapL∘ β α)) + + -- Adjunction between the left Kan extension and precomposition + + private + F* = precomposeF (SET ℓ) (F ^opF) + + open UnitCounit + + η : 𝟙⟨ FUNCTOR (C ^op) (SET ℓ) ⟩ ⇒ funcComp F* Lan + η .N-ob G .N-ob c a = [ c , D.id _ , a ] + η .N-ob G .N-hom {c'} {c} f = + funExt λ a → + [ c , D.id _ , (G ⟪ f ⟫) a ] + ≡⟨ sym (shift/ (D.id _) f a) ⟩ + [ c' , ((D.id _) D.⋆ F ⟪ f ⟫) , a ] + ≡[ i ]⟨ [ c' , lem i , a ] ⟩ + [ c' , (F ⟪ f ⟫ D.⋆ (D.id _)) , a ] + ∎ + where + lem : (D.id _) D.⋆ F ⟪ f ⟫ ≡ F ⟪ f ⟫ D.⋆ (D.id _) + lem = D.⋆IdL (F ⟪ f ⟫) ∙ sym (D.⋆IdR (F ⟪ f ⟫)) + η .N-hom f = makeNatTransPath refl + + ε : funcComp Lan F* ⇒ 𝟙⟨ FUNCTOR (D ^op) (SET ℓ) ⟩ + ε .N-ob H .N-ob d = + elim + (λ _ → (H ⟅ d ⟆) .snd) + (λ (c , g , a) → (H ⟪ g ⟫) a) + (λ {_ _ (shift g f a) i → H .F-seq (F ⟪ f ⟫) g i a}) + ε .N-ob H .N-hom g' = + funExt (elimProp (λ _ → (H ⟅ _ ⟆) .snd _ _) (λ (c , g , a) → funExt⁻ (H .F-seq g g') a)) + ε .N-hom {H} {H'} α = + makeNatTransPath + (funExt₂ λ d → + elimProp (λ _ → (H' ⟅ _ ⟆) .snd _ _) + (λ (c , g , a) → sym (funExt⁻ (α .N-hom g) a))) + + Δ₁ : ∀ G → seqTrans (Lan ⟪ η ⟦ G ⟧ ⟫) (ε ⟦ Lan ⟅ G ⟆ ⟧) ≡ idTrans _ + Δ₁ G = + makeNatTransPath + (funExt₂ λ d → + elimProp (λ _ → squash/ _ _) + (λ (c , g , a) → + [ c , g D.⋆ D.id _ , a ] + ≡[ i ]⟨ [ c , (g D.⋆ F .F-id (~ i)) , a ] ⟩ + [ c , g D.⋆ (F ⟪ C.id _ ⟫) , a ] + ≡⟨ shift/ g (C.id _) a ⟩ + [ c , g , (G ⟪ C.id _ ⟫) a ] + ≡[ i ]⟨ [ c , g , G .F-id i a ] ⟩ + [ c , g , a ] + ∎)) + + Δ₂ : ∀ H → seqTrans (η ⟦ F* ⟅ H ⟆ ⟧) (F* ⟪ ε ⟦ H ⟧ ⟫) ≡ idTrans _ + Δ₂ H = makeNatTransPath (funExt λ c → H .F-id) + + adj : Lan ⊣ F* + adj = make⊣ η ε Δ₁ Δ₂ + +{- + Right Kan extension of a functor C → D to a functor PreShv C ℓ → PreShv D ℓ right adjoint to precomposition. +-} + +module Ran {ℓC ℓC' ℓD ℓD'} ℓS + {C : Precategory ℓC ℓC'} {D : Precategory ℓD ℓD'} + (F : Functor C D) + where + + open Functor + open NatTrans + + private + module C = Precategory C + module D = Precategory D + + {- + We want the category SET ℓ we're mapping into to be large enough that the coend will take presheaves + Cᵒᵖ → Set ℓ to presheaves Dᵒᵖ → Set ℓ, otherwise we get no adjunction with precomposition. + So we must have ℓC,ℓC',ℓD' ≤ ℓ; the parameter ℓS allows ℓ to be larger than their maximum. + -} + ℓ = ℓ-max (ℓ-max (ℓ-max ℓC ℓC') ℓD') ℓS + + module _ (G : Functor (C ^op) (SET ℓ)) where + + -- Definition of the end + + record End (d : D.ob) : Type ℓ where + field + fun : (c : C.ob) (g : D.Hom[ F ⟅ c ⟆ , d ]) → G .F-ob c .fst + coh : {c c' : C.ob} (f : C.Hom[ c , c' ]) (g : D.Hom[ F ⟅ c' ⟆ , d ]) + → fun c (F ⟪ f ⟫ ⋆⟨ D ⟩ g) ≡ (G ⟪ f ⟫) (fun c' g) + + open End + + end≡ : {d : D.ob} {x x' : End d} → (∀ c g → x .fun c g ≡ x' .fun c g) → x ≡ x' + end≡ h i .fun c g = h c g i + end≡ {_} {x} {x'} h i .coh f g = + isSet→isSet' (G .F-ob _ .snd) + (x .coh f g) + (x' .coh f g) + (h _ (F ⟪ f ⟫ ⋆⟨ D ⟩ g)) + (cong (G ⟪ f ⟫) (h _ g)) + i + + -- Action of End on arrows in D + + mapR : {d d' : D.ob} (h : D.Hom[ d' , d ]) → End d → End d' + mapR h x .fun c g = x .fun c (g ⋆⟨ D ⟩ h) + mapR h x .coh f g = cong (x .fun _) (D.⋆Assoc (F ⟪ f ⟫) g h) ∙ x .coh f (g ⋆⟨ D ⟩ h) + + mapRId : (d : D.ob) → mapR (D.id d) ≡ (idfun _) + mapRId h = funExt λ x → end≡ λ c g → cong (x .fun c) (D.⋆IdR g) + + mapR∘ : {d d' d'' : D.ob} + (h' : D.Hom[ d'' , d' ]) (h : D.Hom[ d' , d ]) + → mapR (h' D.⋆ h) ≡ mapR h' ∘ mapR h + mapR∘ h' h = funExt λ x → end≡ λ c g → cong (x .fun c) (sym (D.⋆Assoc g h' h)) + + open End + + RanOb : Functor (C ^op) (SET ℓ) → Functor (D ^op) (SET _) + RanOb G .F-ob d .fst = End G d + RanOb G .F-ob d .snd = + -- We use that End is equivalent to a Σ-type to prove its HLevel more easily + isOfHLevelRetract 2 + {B = + Σ[ z ∈ ((c : C.ob) (g : D.Hom[ F ⟅ c ⟆ , d ]) → G .F-ob c .fst) ] + ({c c' : C.ob} (f : C.Hom[ c , c' ]) (g : D.Hom[ F ⟅ c' ⟆ , d ]) + → z c (F ⟪ f ⟫ ⋆⟨ D ⟩ g) ≡ (G ⟪ f ⟫) (z c' g)) + } + (λ x → λ where .fst → x .fun; .snd → x .coh) + (λ σ → λ where .fun → σ .fst; .coh → σ .snd) + (λ _ → refl) + (isSetΣ + (isSetΠ2 λ _ _ → G .F-ob _ .snd) + (λ _ → isProp→isSet + (isPropImplicitΠ λ _ → isPropImplicitΠ λ _ → isPropΠ2 λ _ _ → G .F-ob _ .snd _ _))) + RanOb G .F-hom = mapR G + RanOb G .F-id {d} = mapRId G d + RanOb G .F-seq h h' = mapR∘ G h' h + + -- Action of End on arrows in Cᵒᵖ → Set + + module _ {G G' : Functor (C ^op) (SET ℓ)} (α : NatTrans G G') where + + mapL : (d : D.ob) → End G d → End G' d + mapL d x .fun c g = (α ⟦ c ⟧) (x .fun c g) + mapL d x .coh f g = + cong (α ⟦ _ ⟧) (x .coh f g) + ∙ funExt⁻ (α .N-hom f) (x .fun _ g) + + mapLR : {d d' : D.ob} (h : D.Hom[ d' , d ]) + → mapL d' ∘ mapR G h ≡ mapR G' h ∘ mapL d + mapLR h = funExt λ _ → end≡ _ λ _ _ → refl + + mapLId : (G : Functor (C ^op) (SET ℓ)) + (d : D.ob) → mapL (idTrans G) d ≡ idfun (End G d) + mapLId G d = funExt λ _ → end≡ _ λ _ _ → refl + + mapL∘ : {G G' G'' : Functor (C ^op) (SET ℓ)} + (β : NatTrans G' G'') (α : NatTrans G G') + (d : D.ob) → mapL (seqTrans α β) d ≡ mapL β d ∘ mapL α d + mapL∘ β α d = funExt λ _ → end≡ _ λ _ _ → refl + + RanHom : {G G' : Functor (C ^op) (SET ℓ)} + → NatTrans G G' → NatTrans (RanOb G) (RanOb G') + RanHom α .N-ob = mapL α + RanHom α .N-hom = mapLR α + + -- Definition of the right Kan extension functor + + Ran : Functor (FUNCTOR (C ^op) (SET ℓ)) (FUNCTOR (D ^op) (SET ℓ)) + Ran .F-ob = RanOb + Ran .F-hom = RanHom + Ran .F-id {G} = makeNatTransPath (funExt (mapLId G)) + Ran .F-seq α β = makeNatTransPath (funExt (mapL∘ β α)) + + -- Adjunction between precomposition and right Kan extension + + private + F* = precomposeF (SET ℓ) (F ^opF) + + open UnitCounit + + η : 𝟙⟨ FUNCTOR (D ^op) (SET ℓ) ⟩ ⇒ (funcComp Ran F*) + η .N-ob G .N-ob d a .fun c g = (G ⟪ g ⟫) a + η .N-ob G .N-ob d a .coh f g = funExt⁻ (G .F-seq g (F ⟪ f ⟫)) a + η .N-ob G .N-hom h = funExt λ a → end≡ _ λ c g → sym (funExt⁻ (G .F-seq h g) a) + η .N-hom {G} {G'} α = + makeNatTransPath (funExt₂ λ d a → end≡ _ λ c g → sym (funExt⁻ (α .N-hom g) a)) + + ε : funcComp F* Ran ⇒ 𝟙⟨ FUNCTOR (C ^op) (SET ℓ) ⟩ + ε .N-ob H .N-ob c x = x .fun c (D.id _) + ε .N-ob H .N-hom {c} {c'} g = + funExt λ x → + cong (x .fun c') (D.⋆IdL _ ∙ sym (D.⋆IdR _)) ∙ x .coh g (D.id _) + ε .N-hom {H} {H'} α = makeNatTransPath refl + + Δ₁ : ∀ G → seqTrans (F* ⟪ η ⟦ G ⟧ ⟫) (ε ⟦ F* ⟅ G ⟆ ⟧) ≡ idTrans _ + Δ₁ G = makeNatTransPath (funExt₂ λ c a → funExt⁻ (G .F-id) a) + + Δ₂ : ∀ H → seqTrans (η ⟦ Ran ⟅ H ⟆ ⟧) (Ran ⟪ ε ⟦ H ⟧ ⟫) ≡ idTrans _ + Δ₂ H = makeNatTransPath (funExt₂ λ c x → end≡ _ λ c' g → cong (x .fun c') (D.⋆IdL g)) + + adj : F* ⊣ Ran + adj = make⊣ η ε Δ₁ Δ₂ diff --git a/Cubical/Categories/Presheaf/Properties.agda b/Cubical/Categories/Presheaf/Properties.agda new file mode 100644 index 000000000..077979dd3 --- /dev/null +++ b/Cubical/Categories/Presheaf/Properties.agda @@ -0,0 +1,389 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Presheaf.Properties where + +open import Cubical.Categories.Category +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Instances.Sets +open import Cubical.Categories.Instances.Functors +open import Cubical.Categories.Functor +open import Cubical.Categories.Presheaf.Base +open import Cubical.Categories.Equivalence +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv using (fiber) +open import Cubical.Data.Sigma + +import Cubical.Categories.Morphism as Morphism +import Cubical.Categories.Constructions.Slice as Slice +import Cubical.Categories.Constructions.Elements as Elements +import Cubical.Functions.Fibration as Fibration + +private + variable + ℓ ℓ' : Level + e e' : Level + + +-- (PreShv C) / F ≃ᶜ PreShv (∫ᴾ F) +module _ {ℓS : Level} (C : Precategory ℓ ℓ') (F : Functor (C ^op) (SET ℓS)) where + open Precategory + open Functor + open _≃ᶜ_ + open isEquivalence + open NatTrans + open NatIso + open Slice (PreShv C ℓS) F ⦃ isCatPreShv {C = C} ⦄ + open Elements {C = C} + + open Fibration.ForSets + + -- specific case of fiber under natural transformation + fibersEqIfRepsEqNatTrans : ∀ {A} (ϕ : A ⇒ F) {c x x'} {px : x ≡ x'} {a' : fiber (ϕ ⟦ c ⟧) x} {b' : fiber (ϕ ⟦ c ⟧) x'} + → fst a' ≡ fst b' + → PathP (λ i → fiber (ϕ ⟦ c ⟧) (px i)) a' b' + fibersEqIfRepsEqNatTrans ϕ {c} {x} {x'} {px} {a , fiba} {b , fibb} p + = fibersEqIfRepsEq {isSetB = snd (F ⟅ c ⟆)} (ϕ ⟦ c ⟧) p + + -- ======================================== + -- K : Slice → PreShv + -- ======================================== + + -- action on (slice) objects + K-ob : (s : SliceCat .ob) → (PreShv (∫ᴾ F) ℓS .ob) + -- we take (c , x) to the fiber in A of ϕ over x + K-ob (sliceob {A} ϕ) .F-ob (c , x) + = (fiber (ϕ ⟦ c ⟧) x) + , isOfHLevelΣ 2 (snd (A ⟅ c ⟆)) λ _ → isSet→isGroupoid (snd (F ⟅ c ⟆)) _ _ + -- for morphisms, we just apply A ⟪ h ⟫ (plus equality proof) + K-ob (sliceob {A} ϕ) .F-hom {d , y} {c , x} (h , com) (b , eq) + = ((A ⟪ h ⟫) b) + , ((ϕ ⟦ c ⟧) ((A ⟪ h ⟫) b) + ≡[ i ]⟨ (ϕ .N-hom h) i b ⟩ + (F ⟪ h ⟫) ((ϕ ⟦ d ⟧) b) + ≡[ i ]⟨ (F ⟪ h ⟫) (eq i) ⟩ + (F ⟪ h ⟫) y + ≡⟨ sym com ⟩ + x + ∎) + -- functoriality follows from functoriality of A + K-ob (sliceob {A} ϕ) .F-id {x = (c , x)} + = funExt λ { (a , fibp) + → fibersEqIfRepsEqNatTrans ϕ (λ i → A .F-id i a) } + K-ob (sliceob {A} ϕ) .F-seq {x = (c , x)} {(d , y)} {(e , z)} (f' , eq1) (g' , eq2) + = funExt λ { ( a , fibp ) + → fibersEqIfRepsEqNatTrans ϕ (λ i → (A .F-seq f' g') i a) } + + + -- action on morphisms (in this case, natural transformation) + K-hom : {sA sB : SliceCat .ob} + → (ε : SliceCat [ sA , sB ]) + → (K-ob sA) ⇒ (K-ob sB) + K-hom {sA = s1@(sliceob {A} ϕ)} {s2@(sliceob {B} ψ)} (slicehom ε com) = natTrans η-ob (λ h → funExt (η-hom h)) + where + P = K-ob s1 + Q = K-ob s2 + + -- just apply the natural transformation (ε) we're given + -- this ensures that we stay in the fiber over x due to the commutativity given by slicenesss + η-ob : (el : (∫ᴾ F) .ob) → (fst (P ⟅ el ⟆) → fst (Q ⟅ el ⟆) ) + η-ob (c , x) (a , ϕa≡x) = ((ε ⟦ c ⟧) a) , εψ≡ϕ ∙ ϕa≡x + where + εψ≡ϕ : (ψ ⟦ c ⟧) ((ε ⟦ c ⟧) a) ≡ (ϕ ⟦ c ⟧) a + εψ≡ϕ i = ((com i) ⟦ c ⟧) a + + η-hom : ∀ {el1 el2} (h : (∫ᴾ F) [ el1 , el2 ]) (ae : fst (P ⟅ el2 ⟆)) → η-ob el1 ((P ⟪ h ⟫) ae) ≡ (Q ⟪ h ⟫) (η-ob el2 ae) + η-hom {el1 = (c , x)} {d , y} (h , eqh) (a , eqa) + = fibersEqIfRepsEqNatTrans ψ (λ i → ε .N-hom h i a) + + + K : Functor SliceCat (PreShv (∫ᴾ F) ℓS) + K .F-ob = K-ob + K .F-hom = K-hom + K .F-id = makeNatTransPath + (funExt λ cx@(c , x) + → funExt λ aeq@(a , eq) + → fibersEqIfRepsEq {isSetB = snd (F ⟅ c ⟆)} _ refl) + K .F-seq (slicehom α eqa) (slicehom β eqb) + = makeNatTransPath + (funExt λ cx@(c , x) + → funExt λ aeq@(a , eq) + → fibersEqIfRepsEq {isSetB = snd (F ⟅ c ⟆)} _ refl) + + + -- ======================================== + -- L : PreShv → Slice + -- ======================================== + + -- action on objects (presheaves) + L-ob : (P : PreShv (∫ᴾ F) ℓS .ob) + → SliceCat .ob + L-ob P = sliceob {S-ob = L-ob-ob} L-ob-hom + where + -- sends c to the disjoint union of all the images under P + LF-ob : (c : C .ob) → (SET _) .ob + LF-ob c = (Σ[ x ∈ fst (F ⟅ c ⟆) ] fst (P ⟅ c , x ⟆)) , isSetΣ (snd (F ⟅ c ⟆)) (λ x → snd (P ⟅ c , x ⟆)) + + -- defines a function piecewise over the fibers by applying P + LF-hom : ∀ {x y} + → (f : C [ y , x ]) + → (SET _) [ LF-ob x , LF-ob y ] + LF-hom {x = c} {d} f (x , a) = ((F ⟪ f ⟫) x) , (P ⟪ f , refl ⟫) a + + L-ob-ob : Functor (C ^op) (SET _) + L-ob-ob .F-ob = LF-ob + L-ob-ob .F-hom = LF-hom + L-ob-ob .F-id {x = c} + = funExt idFunExt + where + idFunExt : ∀ (un : fst (LF-ob c)) + → (LF-hom (C .id c) un) ≡ un + idFunExt (x , X) = ΣPathP (leftEq , rightEq) + where + leftEq : (F ⟪ C .id c ⟫) x ≡ x + leftEq i = F .F-id i x + + rightEq : PathP (λ i → fst (P ⟅ c , leftEq i ⟆)) + ((P ⟪ C .id c , refl ⟫) X) X + rightEq = left ▷ right + where + -- the id morphism in (∫ᴾ F) + ∫id = C .id c , sym (funExt⁻ (F .F-id) x ∙ refl) + + -- functoriality of P gives us close to what we want + right : (P ⟪ ∫id ⟫) X ≡ X + right i = P .F-id i X + + -- but need to do more work to show that (C .id c , refl) ≡ ∫id + left : PathP (λ i → fst (P ⟅ c , leftEq i ⟆)) + ((P ⟪ C .id c , refl ⟫) X) + ((P ⟪ ∫id ⟫) X) + left i = (P ⟪ ∫ᴾhomEq {F = F} (C .id c , refl) ∫id (λ i → (c , leftEq i)) refl refl i ⟫) X + L-ob-ob .F-seq {x = c} {d} {e} f g + = funExt seqFunEq + where + seqFunEq : ∀ (un : fst (LF-ob c)) + → (LF-hom (g ⋆⟨ C ⟩ f) un) ≡ (LF-hom g) (LF-hom f un) + seqFunEq un@(x , X) = ΣPathP (leftEq , rightEq) + where + -- the left component is comparing the action of F on x + -- equality follows from functoriality of F + -- leftEq : fst (LF-hom (g ⋆⟨ C ⟩ f) un) ≡ fst ((LF-hom g) (LF-hom f un)) + leftEq : (F ⟪ g ⋆⟨ C ⟩ f ⟫) x ≡ (F ⟪ g ⟫) ((F ⟪ f ⟫) x) + leftEq i = F .F-seq f g i x + + -- on the right, equality also follows from functoriality of P + -- but it's more complicated because of heterogeneity + -- since leftEq is not a definitional equality + rightEq : PathP (λ i → fst (P ⟅ e , leftEq i ⟆)) + ((P ⟪ g ⋆⟨ C ⟩ f , refl ⟫) X) + ((P ⟪ g , refl ⟫) ((P ⟪ f , refl ⟫) X)) + rightEq = left ▷ right + where + -- functoriality of P only gets us to this weird composition on the left + right : (P ⟪ (g , refl) ⋆⟨ (∫ᴾ F) ⟩ (f , refl) ⟫) X ≡ (P ⟪ g , refl ⟫) ((P ⟪ f , refl ⟫) X) + right i = P .F-seq (f , refl) (g , refl) i X + + -- so we need to show that this composition is actually equal to the one we want + left : PathP (λ i → fst (P ⟅ e , leftEq i ⟆)) + ((P ⟪ g ⋆⟨ C ⟩ f , refl ⟫) X) + ((P ⟪ (g , refl) ⋆⟨ (∫ᴾ F) ⟩ (f , refl) ⟫) X) + left i = (P ⟪ ∫ᴾhomEq {F = F} (g ⋆⟨ C ⟩ f , refl) ((g , refl) ⋆⟨ (∫ᴾ F) ⟩ (f , refl)) (λ i → (e , leftEq i)) refl refl i ⟫) X + L-ob-hom : L-ob-ob ⇒ F + L-ob-hom .N-ob c (x , _) = x + L-ob-hom .N-hom f = funExt λ (x , _) → refl + + -- action on morphisms (aka natural transformations between presheaves) + -- is essentially the identity (plus equality proofs for naturality and slice commutativity) + L-hom : ∀ {P Q} → PreShv (∫ᴾ F) ℓS [ P , Q ] → + SliceCat [ L-ob P , L-ob Q ] + L-hom {P} {Q} η = slicehom arr com + where + A = S-ob (L-ob P) + ϕ = S-arr (L-ob P) + B = S-ob (L-ob Q) + ψ = S-arr (L-ob Q) + arr : A ⇒ B + arr .N-ob c (x , X) = x , ((η ⟦ c , x ⟧) X) + arr .N-hom {c} {d} f = funExt natu + where + natuType : fst (A ⟅ c ⟆) → Type _ + natuType xX@(x , X) = ((F ⟪ f ⟫) x , (η ⟦ d , (F ⟪ f ⟫) x ⟧) ((P ⟪ f , refl ⟫) X)) ≡ ((F ⟪ f ⟫) x , (Q ⟪ f , refl ⟫) ((η ⟦ c , x ⟧) X)) + natu : ∀ (xX : fst (A ⟅ c ⟆)) → natuType xX + natu (x , X) = ΣPathP (refl , λ i → (η .N-hom (f , refl) i) X) + + com : arr ⋆⟨ PreShv C ℓS ⟩ ψ ≡ ϕ + com = makeNatTransPath (funExt comFunExt) + where + comFunExt : ∀ (c : C .ob) + → (arr ●ᵛ ψ) ⟦ c ⟧ ≡ ϕ ⟦ c ⟧ + comFunExt c = funExt λ x → refl + + L : Functor (PreShv (∫ᴾ F) ℓS) SliceCat + L .F-ob = L-ob + L .F-hom = L-hom + L .F-id {cx} = SliceHom-≡-intro' (makeNatTransPath (funExt λ c → refl)) + L .F-seq {cx} {dy} P Q = SliceHom-≡-intro' (makeNatTransPath (funExt λ c → refl)) + + -- ======================================== + -- η : 𝟙 ≅ LK + -- ======================================== + + module _ where + open Iso + open Morphism renaming (isIso to isIsoC) + -- the iso we need + -- a type is isomorphic to the disjoint union of all its fibers + typeSectionIso : ∀ {A B : Type ℓS} {isSetB : isSet B} → (ϕ : A → B) + → Iso A (Σ[ b ∈ B ] fiber ϕ b) + typeSectionIso ϕ .fun a = (ϕ a) , (a , refl) + typeSectionIso ϕ .inv (b , (a , eq)) = a + typeSectionIso {isSetB = isSetB} ϕ .rightInv (b , (a , eq)) + = ΣPathP (eq + , ΣPathP (refl + , isOfHLevel→isOfHLevelDep 1 (λ b' → isSetB _ _) refl eq eq)) + typeSectionIso ϕ .leftInv a = refl + + -- the natural transformation + -- just applies typeSectionIso + ηTrans : 𝟙⟨ SliceCat ⟩ ⇒ (L ∘F K) + ηTrans .N-ob sob@(sliceob {A} ϕ) = slicehom A⇒LK comm + where + LKA = S-ob (L ⟅ K ⟅ sob ⟆ ⟆) + ψ = S-arr (L ⟅ K ⟅ sob ⟆ ⟆) + + A⇒LK : A ⇒ LKA + A⇒LK .N-ob c = typeSectionIso {isSetB = snd (F ⟅ c ⟆)} (ϕ ⟦ c ⟧) .fun + A⇒LK .N-hom {c} {d} f = funExt homFunExt + where + homFunExt : (x : fst (A ⟅ c ⟆)) + → (((ϕ ⟦ d ⟧) ((A ⟪ f ⟫) x)) , ((A ⟪ f ⟫) x , refl)) ≡ ((F ⟪ f ⟫) ((ϕ ⟦ c ⟧) x) , (A ⟪ f ⟫) x , _) + homFunExt x = ΣPathP ((λ i → (ϕ .N-hom f i) x) , fibersEqIfRepsEqNatTrans ϕ refl) + + comm : (A⇒LK) ●ᵛ ψ ≡ ϕ + comm = makeNatTransPath (funExt λ x → refl) + ηTrans .N-hom {sliceob {A} α} {sliceob {B} β} (slicehom ϕ eq) + = SliceHom-≡-intro' (makeNatTransPath (funExt (λ c → funExt λ a → natFunExt c a))) + where + natFunExt : ∀ (c : C .ob) (a : fst (A ⟅ c ⟆)) + → ((β ⟦ c ⟧) ((ϕ ⟦ c ⟧) a) , (ϕ ⟦ c ⟧) a , _) ≡ ((α ⟦ c ⟧) a , (ϕ ⟦ c ⟧) a , _) + natFunExt c a = ΣPathP ((λ i → ((eq i) ⟦ c ⟧) a) , fibersEqIfRepsEqNatTrans β refl) + + -- isomorphism follows from typeSectionIso + ηIso : ∀ (sob : SliceCat .ob) + → isIsoC {C = SliceCat} (ηTrans ⟦ sob ⟧) + ηIso sob@(sliceob ϕ) = sliceIso _ _ (FUNCTORIso _ _ _ isIsoCf) + where + isIsoCf : ∀ (c : C .ob) + → isIsoC (ηTrans .N-ob sob .S-hom ⟦ c ⟧) + isIsoCf c = CatIso→isIso (Iso→CatIso (typeSectionIso {isSetB = snd (F ⟅ c ⟆)} (ϕ ⟦ c ⟧))) + + + -- ======================================== + -- ε : KL ≅ 𝟙 + -- ======================================== + + module _ where + open Iso + open Morphism renaming (isIso to isIsoC) + -- the iso we deserve + -- says that a type family at x is isomorphic to the fiber over x of that type family packaged up + typeFiberIso : ∀ {ℓ ℓ'} {A : Type ℓ} {isSetA : isSet A} {x} (B : A → Type ℓ') + → Iso (B x) (fiber {A = Σ[ a ∈ A ] B a} (λ (x , _) → x) x) + typeFiberIso {x = x} _ .fun b = (x , b) , refl + typeFiberIso _ .inv ((a , b) , eq) = subst _ eq b + typeFiberIso {isSetA = isSetA} {x = x} B .rightInv ((a , b) , eq) + = fibersEqIfRepsEq {isSetB = isSetA} (λ (x , _) → x) (ΣPathP (sym eq , symP (transport-filler (λ i → B (eq i)) b))) + typeFiberIso {x = x} _ .leftInv b = sym (transport-filler refl b) + + -- the natural isomorphism + -- applies typeFiberIso (inv) + εTrans : (K ∘F L) ⇒ 𝟙⟨ PreShv (∫ᴾ F) ℓS ⟩ + εTrans .N-ob P = natTrans γ-ob (λ f → funExt (λ a → γ-homFunExt f a)) + where + KLP = K ⟅ L ⟅ P ⟆ ⟆ + + γ-ob : (el : (∫ᴾ F) .ob) + → (fst (KLP ⟅ el ⟆) → fst (P ⟅ el ⟆) ) + γ-ob el@(c , _) = typeFiberIso {isSetA = snd (F ⟅ c ⟆)} (λ x → fst (P ⟅ c , x ⟆)) .inv + + -- naturality + -- the annoying part is all the substs + γ-homFunExt : ∀ {el2 el1} → (f' : (∫ᴾ F) [ el2 , el1 ]) + → (∀ (a : fst (KLP ⟅ el1 ⟆)) → γ-ob el2 ((KLP ⟪ f' ⟫) a) ≡ (P ⟪ f' ⟫) (γ-ob el1 a)) + γ-homFunExt {d , y} {c , x} f'@(f , comm) a@((x' , X') , eq) i + = comp (λ j → fst (P ⟅ d , eq' j ⟆)) (λ j → λ { (i = i0) → left j + ; (i = i1) → right j }) ((P ⟪ f , refl ⟫) X') + where + -- fiber equality proof that we get from an application of KLP + eq' = snd ((KLP ⟪ f' ⟫) a) + + -- top right of the commuting diagram + -- "remove" the subst from the inside + right : PathP (λ i → fst (P ⟅ d , eq' i ⟆)) ((P ⟪ f , refl ⟫) X') ((P ⟪ f , comm ⟫) (subst _ eq X')) + right i = (P ⟪ f , refl≡comm i ⟫) (X'≡subst i) + where + refl≡comm : PathP (λ i → (eq' i) ≡ (F ⟪ f ⟫) (eq i)) refl comm + refl≡comm = isOfHLevel→isOfHLevelDep 1 (λ (v , w) → snd (F ⟅ d ⟆) v ((F ⟪ f ⟫) w)) refl comm λ i → (eq' i , eq i) + + X'≡subst : PathP (λ i → fst (P ⟅ c , eq i ⟆)) X' (subst _ eq X') + X'≡subst = transport-filler (λ i → fst (P ⟅ c , eq i ⟆)) X' + + -- bottom left of the commuting diagram + -- "remove" the subst from the outside + left : PathP (λ i → fst (P ⟅ d , eq' i ⟆)) ((P ⟪ f , refl ⟫) X') (subst (λ v → fst (P ⟅ d , v ⟆)) eq' ((P ⟪ f , refl ⟫) X')) + left = transport-filler (λ i → fst (P ⟅ d , eq' i ⟆)) ((P ⟪ f , refl ⟫) X') + εTrans .N-hom {P} {Q} α = makeNatTransPath (funExt λ cx → funExt λ xX' → ε-homFunExt cx xX') + where + KLP = K ⟅ L ⟅ P ⟆ ⟆ + + -- naturality of the above construction applies a similar argument as in `γ-homFunExt` + ε-homFunExt : ∀ (cx@(c , x) : (∫ᴾ F) .ob) (xX'@((x' , X') , eq) : fst (KLP ⟅ cx ⟆)) + → subst (λ v → fst (Q ⟅ c , v ⟆)) (snd ((K ⟪ L ⟪ α ⟫ ⟫ ⟦ cx ⟧) xX')) ((α ⟦ c , x' ⟧) X') + ≡ (α ⟦ c , x ⟧) (subst _ eq X') + ε-homFunExt cx@(c , x) xX'@((x' , X') , eq) i + = comp (λ j → fst (Q ⟅ c , eq j ⟆)) (λ j → λ { (i = i0) → left j + ; (i = i1) → right j }) ((α ⟦ c , x' ⟧) X') + where + eq' : x' ≡ x + eq' = snd ((K ⟪ L ⟪ α ⟫ ⟫ ⟦ cx ⟧) xX') + + right : PathP (λ i → fst (Q ⟅ c , eq i ⟆)) ((α ⟦ c , x' ⟧) X') ((α ⟦ c , x ⟧) (subst _ eq X')) + right i = (α ⟦ c , eq i ⟧) (X'≡subst i) + where + -- this is exactly the same as the one from before, can refactor? + X'≡subst : PathP (λ i → fst (P ⟅ c , eq i ⟆)) X' (subst _ eq X') + X'≡subst = transport-filler _ _ + + -- extracted out type since need to use in in 'left' body as well + leftTy : (x' ≡ x) → Type _ + leftTy eq* = PathP (λ i → fst (Q ⟅ c , eq* i ⟆)) ((α ⟦ c , x' ⟧) X') (subst (λ v → fst (Q ⟅ c , v ⟆)) eq' ((α ⟦ c , x' ⟧) X')) + + left : leftTy eq + left = subst + (λ eq* → leftTy eq*) + eq'≡eq + (transport-filler _ _) + where + eq'≡eq : eq' ≡ eq + eq'≡eq = snd (F ⟅ c ⟆) _ _ eq' eq + + εIso : ∀ (P : PreShv (∫ᴾ F) ℓS .ob) + → isIsoC {C = PreShv (∫ᴾ F) ℓS} (εTrans ⟦ P ⟧) + εIso P = FUNCTORIso _ _ _ isIsoC' + where + isIsoC' : ∀ (cx : (∫ᴾ F) .ob) + → isIsoC {C = SET _} ((εTrans ⟦ P ⟧) ⟦ cx ⟧) + isIsoC' cx@(c , _) = CatIso→isIso (Iso→CatIso (invIso (typeFiberIso {isSetA = snd (F ⟅ c ⟆)} _))) + + + -- putting it all together + + preshvSlice≃preshvElem : SliceCat ≃ᶜ PreShv (∫ᴾ F) ℓS + preshvSlice≃preshvElem .func = K + preshvSlice≃preshvElem .isEquiv .invFunc = L + preshvSlice≃preshvElem .isEquiv .η .trans = ηTrans + preshvSlice≃preshvElem .isEquiv .η .nIso = ηIso + preshvSlice≃preshvElem .isEquiv .ε .trans = εTrans + preshvSlice≃preshvElem .isEquiv .ε .nIso = εIso diff --git a/Cubical/Categories/Sets.agda b/Cubical/Categories/Sets.agda new file mode 100644 index 000000000..caabf85a8 --- /dev/null +++ b/Cubical/Categories/Sets.agda @@ -0,0 +1,93 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Sets where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation + +open Precategory + +module _ ℓ where + SET : Precategory (ℓ-suc ℓ) ℓ + SET .ob = Σ (Type ℓ) isSet + SET .Hom[_,_] (A , _) (B , _) = A → B + SET .id _ = λ x → x + SET ._⋆_ f g = λ x → g (f x) + SET .⋆IdL f = refl + SET .⋆IdR f = refl + SET .⋆Assoc f g h = refl + +module _ {ℓ} where + isSetExpIdeal : {A B : Type ℓ} → isSet B → isSet (A → B) + isSetExpIdeal B/set = isSetΠ λ _ → B/set + + isSetLift : {A : Type ℓ} → isSet A → isSet (Lift {ℓ} {ℓ-suc ℓ} A) + isSetLift = isOfHLevelLift 2 + + module _ {A B : SET ℓ .ob} where + -- monic/surjectiveness + open import Cubical.Categories.Morphism + isSurjSET : (f : SET ℓ [ A , B ]) → Type _ + isSurjSET f = ∀ (b : fst B) → Σ[ a ∈ fst A ] f a ≡ b + + -- isMonic→isSurjSET : {f : SET ℓ [ A , B ]} + -- → isEpic {C = SET ℓ} {x = A} {y = B} f + -- → isSurjSET f + -- isMonic→isSurjSET ism b = {!!} , {!!} + + instance + SET-category : isCategory (SET ℓ) + SET-category .isSetHom {_} {B , B/set} = isSetExpIdeal B/set + +private + variable + ℓ ℓ' : Level + +open Functor + +-- Hom functors +_[-,_] : (C : Precategory ℓ ℓ') → (c : C .ob) → ⦃ isCat : isCategory C ⦄ → Functor (C ^op) (SET _) +(C [-, c ]) ⦃ isCat ⦄ .F-ob x = (C [ x , c ]) , isCat .isSetHom +(C [-, c ]) .F-hom f k = f ⋆⟨ C ⟩ k +(C [-, c ]) .F-id = funExt λ _ → C .⋆IdL _ +(C [-, c ]) .F-seq _ _ = funExt λ _ → C .⋆Assoc _ _ _ + +_[_,-] : (C : Precategory ℓ ℓ') → (c : C .ob) → ⦃ isCat : isCategory C ⦄ → Functor C (SET _) +(C [ c ,-]) ⦃ isCat ⦄ .F-ob x = (C [ c , x ]) , isCat .isSetHom +(C [ c ,-]) .F-hom f k = k ⋆⟨ C ⟩ f +(C [ c ,-]) .F-id = funExt λ _ → C .⋆IdR _ +(C [ c ,-]) .F-seq _ _ = funExt λ _ → sym (C .⋆Assoc _ _ _) + +module _ {C : Precategory ℓ ℓ'} ⦃ _ : isCategory C ⦄ {F : Functor C (SET ℓ')} where + open NatTrans + + -- natural transformations by pre/post composition + preComp : {x y : C .ob} + → (f : C [ x , y ]) + → C [ x ,-] ⇒ F + → C [ y ,-] ⇒ F + preComp f α .N-ob c k = (α ⟦ c ⟧) (f ⋆⟨ C ⟩ k) + preComp f α .N-hom {x = c} {d} k + = (λ l → (α ⟦ d ⟧) (f ⋆⟨ C ⟩ (l ⋆⟨ C ⟩ k))) + ≡[ i ]⟨ (λ l → (α ⟦ d ⟧) (⋆Assoc C f l k (~ i))) ⟩ + (λ l → (α ⟦ d ⟧) (f ⋆⟨ C ⟩ l ⋆⟨ C ⟩ k)) + ≡[ i ]⟨ (λ l → (α .N-hom k) i (f ⋆⟨ C ⟩ l)) ⟩ + (λ l → (F ⟪ k ⟫) ((α ⟦ c ⟧) (f ⋆⟨ C ⟩ l))) + ∎ + +-- properties +-- TODO: move to own file +open CatIso renaming (inv to cInv) +open Iso + +Iso→CatIso : ∀ {A B : (SET ℓ) .ob} + → Iso (fst A) (fst B) + → CatIso {C = SET ℓ} A B +Iso→CatIso is .mor = is .fun +Iso→CatIso is .cInv = is .inv +Iso→CatIso is .sec = funExt λ b → is .rightInv b -- is .rightInv +Iso→CatIso is .ret = funExt λ b → is .leftInv b -- is .rightInv diff --git a/Cubical/Categories/TypesOfCategories/TypeCategory.agda b/Cubical/Categories/TypesOfCategories/TypeCategory.agda new file mode 100644 index 000000000..4bd4ace5a --- /dev/null +++ b/Cubical/Categories/TypesOfCategories/TypeCategory.agda @@ -0,0 +1,187 @@ +{-# OPTIONS --postfix-projections --safe #-} + +module Cubical.Categories.TypesOfCategories.TypeCategory where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Data.Sigma +import Cubical.Functions.Fibration as Fibration + +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.Presheaf +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Instances.Sets + + +open Fibration.ForSets + +record isTypeCategory {ℓ ℓ' ℓ''} (C : Precategory ℓ ℓ') + : Type (ℓ-max ℓ (ℓ-max ℓ' (ℓ-suc ℓ''))) where + open Precategory C + open Cospan + open PullbackLegs + open isPullback + field + -- a Type of types over a context + Ty[_] : ob → Type ℓ'' + -- extend a context with a type + cext : ∀ (Γ : _) → (A : Ty[ Γ ]) → Σ[ ΓA ∈ ob ] (C [ ΓA , Γ ]) + + -- the new object from a context extension + _⍮_ : (Γ : _) → (A : Ty[ Γ ]) → ob + Γ ⍮ A = fst (cext Γ A) + + -- the projection from the extended context to the original + π : (Γ : _) → (A : Ty[ Γ ]) → C [ Γ ⍮ A , Γ ] + π Γ A = snd (cext Γ A) + + field + -- pullback over context extentions + reindex : ∀ {Γ' Γ} + → C [ Γ' , Γ ] + → (Ty[ Γ ] → Ty[ Γ' ]) + + q⟨_,_⟩ : ∀ {Γ' Γ} + → (f : C [ Γ' , Γ ]) + → (A : Ty[ Γ ]) + → C [ Γ' ⍮ (reindex f A) , Γ ⍮ A ] + + isPB : ∀ {Γ' Γ : ob} (f : C [ Γ' , Γ ]) (A : Ty[ Γ ]) + → isPullback {C = C} (cospan Γ' Γ (Γ ⍮ A) f (π Γ A)) + (pblegs (π Γ' (reindex f A)) q⟨ f , A ⟩) + +-- presheaves are type contexts +module _ {ℓ ℓ' ℓ'' : Level} (C : Precategory ℓ ℓ') where + open isTypeCategory + open Precategory + open Functor + open NatTrans + open isPullback + + private + -- types over Γ are types with a "projection" (aka surjection) to Γ + PSTy[_] : PreShv C ℓ'' .ob → Type _ + PSTy[ Γ ] = Σ[ ΓA ∈ PreShv C ℓ'' .ob ] + Σ[ π ∈ ΓA ⇒ Γ ] + (∀ (c : C .ob) + → isSurjSET {A = ΓA ⟅ c ⟆} {Γ ⟅ c ⟆} (π ⟦ c ⟧)) + + -- just directly use types from above as context extensions + PSCext : (Γ : _) → PSTy[ Γ ] → Σ[ ΓA ∈ PreShv C ℓ'' .ob ] ΓA ⇒ Γ + PSCext Γ (ΓA , π , _) = ΓA , π + + -- the pullback or reindexed set is the disjoint union of the fibers + -- from the projection + module _ {Δ Γ : PreShv C ℓ'' .ob} (γ : Δ ⇒ Γ) + (A'@(ΓA , π , isSurjπ) : PSTy[ Γ ]) where + ΔA : PreShv C ℓ'' .ob + ΔA .F-ob c = ΔATy , isSetΔA + where + ΔATy = (Σ[ x ∈ fst (Δ ⟅ c ⟆) ] fiber (π ⟦ c ⟧) ((γ ⟦ c ⟧) x)) + + isSetΔA : isSet ΔATy + isSetΔA = isOfHLevelΣ 2 (snd (Δ ⟅ c ⟆)) λ Γc → isOfHLevelΣ 2 (snd (ΓA ⟅ c ⟆)) λ ΓAc → isProp→isSet (snd (Γ ⟅ c ⟆) _ _) + -- for morphisms, we apply Δ ⟪ f ⟫ to the first component + -- and ΓA ⟪ f ⟫ to the second + -- the fiber rule + ΔA .F-hom {c} {d} f (δax , γax , eq) + = ((Δ ⟪ f ⟫) δax) + , (((ΓA ⟪ f ⟫) γax) + , ((π ⟦ d ⟧) ((ΓA ⟪ f ⟫) γax) + ≡[ i ]⟨ π .N-hom f i γax ⟩ + (Γ ⟪ f ⟫) ((π ⟦ c ⟧) γax) + ≡[ i ]⟨ (Γ ⟪ f ⟫) (eq i) ⟩ + (Γ ⟪ f ⟫) ((γ ⟦ c ⟧) δax) + ≡[ i ]⟨ γ .N-hom f (~ i) δax ⟩ + (γ ⟦ d ⟧) ((Δ ⟪ f ⟫) δax) + ∎)) + ΔA .F-id {x = c} + = funExt λ (δax , γax , eq) + → ΣPathP ((λ i → Δ .F-id i δax) + , fibersEqIfRepsEq {isSetB = snd (Γ ⟅ c ⟆)} _ + (λ i → ΓA .F-id i γax)) + ΔA .F-seq {a} {b} {c} f g + = funExt λ (δax , γax , eq) + → ΣPathP ((λ i → Δ .F-seq f g i δax) + , fibersEqIfRepsEq {isSetB = snd (Γ ⟅ c ⟆)} _ + λ i → ΓA .F-seq f g i γax) + π' : ΔA ⇒ Δ + π' .N-ob c (x , snd) = x + π' .N-hom f = refl + PSReindex : PSTy[ Δ ] + PSReindex = ΔA , (π' , isSurj) + where + + isSurj : ∀ (c : C .ob) → isSurjSET {A = ΔA ⟅ c ⟆} {B = Δ ⟅ c ⟆} (π' ⟦ c ⟧) + isSurj c δx = (δx , isSurjπ c ((γ ⟦ c ⟧) δx)) , refl + + PSq : ΔA ⇒ ΓA + PSq .N-ob c (δax , γax , eq) = γax + PSq .N-hom {c} {d} f = funExt λ (δax , γax , eq) → refl + + PSIsPB : isPullback {C = PreShv C ℓ''} + (cospan Δ Γ (fst (PSCext Γ A')) γ (snd (PSCext Γ A'))) + (pblegs (snd (PSCext Δ PSReindex)) (PSq)) + PSIsPB .sq = makeNatTransPath (funExt sqExt) + where + sqExt : ∀ (c : C .ob) → _ + sqExt c = funExt λ (δax , γax , eq) → sym eq + + PSIsPB .up {Θ} (cone (pblegs p₁ p₂) sq) + = ((α , eq) + , unique) + where + α : Θ ⇒ ΔA + α .N-ob c t = ((p₁ ⟦ c ⟧) t) + , (((p₂ ⟦ c ⟧) t) + , (λ i → (sq (~ i) ⟦ c ⟧) t)) + α .N-hom {d} {c} f = funExt αHomExt + where + αHomExt : ∀ (t : fst (Θ ⟅ d ⟆)) + → ((p₁ ⟦ c ⟧) ((Θ ⟪ f ⟫) t) , (p₂ ⟦ c ⟧) ((Θ ⟪ f ⟫) t), _) + ≡ ((Δ ⟪ f ⟫) ((p₁ ⟦ d ⟧) t) , (ΓA ⟪ f ⟫) ((p₂ ⟦ d ⟧) t) , _) + αHomExt t = ΣPathP ((λ i → p₁ .N-hom f i t) + , fibersEqIfRepsEq {isSetB = snd (Γ ⟅ c ⟆)} _ + (λ i → p₂ .N-hom f i t)) + + eq : _ + eq = makeNatTransPath (funExt (λ _ → funExt λ _ → refl)) + , makeNatTransPath (funExt (λ _ → funExt λ _ → refl)) + + unique : ∀ (βeq : Σ[ β ∈ Θ ⇒ ΔA ] _) + → (α , eq) ≡ βeq + unique (β , eqβ) = ΣPathP (α≡β , eq≡eqβ) + where + α≡β : α ≡ β + α≡β = makeNatTransPath (funExt λ c → funExt λ t → eqExt c t) + where + eqβ1 = eqβ .fst + eqβ2 = eqβ .snd + eqExt : ∀ (c : C .ob) + → (t : fst (Θ ⟅ c ⟆)) + → (α ⟦ c ⟧) t ≡ (β ⟦ c ⟧) t + eqExt c t = ΣPathP ((λ i → (eqβ1 i ⟦ c ⟧) t) + , fibersEqIfRepsEq {isSetB = snd (Γ ⟅ c ⟆)} _ + (λ i → (eqβ2 i ⟦ c ⟧) t)) + eq≡eqβ : PathP (λ i + → (p₁ ≡ (α≡β i) ●ᵛ π') + × (p₂ ≡ (α≡β i) ●ᵛ PSq)) eq eqβ + eq≡eqβ = ΣPathP ( isPropNatP1 (eq .fst) (eqβ .fst) α≡β + , isPropNatP2 (eq .snd) (eqβ .snd) α≡β) + where + isPropNatP1 : isOfHLevelDep 1 (λ γ → p₁ ≡ γ ●ᵛ π') + isPropNatP1 = isOfHLevel→isOfHLevelDep 1 (λ _ → isSetNat _ _) + + isPropNatP2 : isOfHLevelDep 1 (λ γ → p₂ ≡ γ ●ᵛ PSq) + isPropNatP2 = isOfHLevel→isOfHLevelDep 1 (λ _ → isSetNat _ _) + + -- putting everything together + isTypeCategoryPresheaf : isTypeCategory (PreShv C ℓ'') + isTypeCategoryPresheaf .Ty[_] Γ = PSTy[ Γ ] + isTypeCategoryPresheaf .cext = PSCext + isTypeCategoryPresheaf .reindex = PSReindex + isTypeCategoryPresheaf .q⟨_,_⟩ = PSq + isTypeCategoryPresheaf .isPB = PSIsPB diff --git a/Cubical/Categories/Yoneda.agda b/Cubical/Categories/Yoneda.agda new file mode 100644 index 000000000..a5178cc8d --- /dev/null +++ b/Cubical/Categories/Yoneda.agda @@ -0,0 +1,192 @@ +{-# OPTIONS --safe #-} + +module Cubical.Categories.Yoneda where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence using (ua) +open import Cubical.Foundations.Function renaming (_∘_ to _◍_) +open import Cubical.Data.Sigma +open import Cubical.Foundations.Equiv +open import Cubical.HITs.PropositionalTruncation + +open import Cubical.Categories.Category +open import Cubical.Categories.Instances.Sets +open import Cubical.Categories.Instances.Functors +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Functor +open import Cubical.Categories.Presheaf + +private + variable + ℓ ℓ' ℓ'' : Level + +-- THE YONEDA LEMMA + +open NatTrans +open NatTransP +open Functor +open Iso + + +module _ (A B : Type ℓ) (f : A → B) where + + isInj = ∀ (x y : A) → (f x ≡ f y) → x ≡ y + isSurj = ∀ (b : B) → Σ[ a ∈ A ] f a ≡ b + + bijectionToIso : isInj × isSurj + → isIso f + bijectionToIso (i , s) = (λ b → fst (s b)) , (λ b → snd (s b)) , λ a → i (fst (s (f a))) a (snd (s (f a))) + +module _ {C : Precategory ℓ ℓ'} ⦃ isCatC : isCategory C ⦄ where + open Precategory + yoneda : (F : Functor C (SET ℓ')) + → (c : C .ob) + → Iso ((FUNCTOR C (SET ℓ')) [ C [ c ,-] , F ]) (fst (F ⟅ c ⟆)) + yoneda F c = theIso + where + natType = (FUNCTOR C (SET ℓ')) [ C [ c ,-] , F ] + setType = fst (F ⟅ c ⟆) + + -- takes a natural transformation to what it does on id + ϕ : natType → setType + ϕ α = (α ⟦ _ ⟧) (C .id c) + + -- takes an element x of F c and sends it to the (only) natural transformation + -- which takes the identity to x + Ψ : setType → natType + Ψ x .N-ob c = λ f → (F ⟪ f ⟫) x + Ψ x .N-hom g + = funExt (λ f → (F ⟪ f ⋆⟨ C ⟩ g ⟫) x + ≡[ i ]⟨ (F .F-seq f g) i x ⟩ + (F ⟪ g ⟫) ((F ⟪ f ⟫) x) + ∎) + + theIso : Iso natType setType + theIso .fun = ϕ + theIso .inv = Ψ + theIso .rightInv x i = F .F-id i x + theIso .leftInv α@(natTrans αo αh) = NatTrans-≡-intro (sym αo≡βo) (symP αh≡βh) + where + β = Ψ (ϕ α) + βo = β .N-ob + βh = β .N-hom + + -- equivalence of action on objects follows + -- from simple equational reasoning using naturality + αo≡βo : αo ≡ βo + αo≡βo = funExt λ x → funExt λ f + → αo x f + ≡[ i ]⟨ αo x (C .⋆IdL f (~ i)) ⟩ -- expand into the bottom left of the naturality diagram + αo x (C .id c ⋆⟨ C ⟩ f) + ≡[ i ]⟨ αh f i (C .id c) ⟩ -- apply naturality + (F ⟪ f ⟫) ((αo _) (C .id c)) + ∎ + + -- type aliases for natural transformation + NOType = N-ob-Type (C [ c ,-]) F + NHType = N-hom-Type (C [ c ,-]) F + + -- equivalence of commutative squares follows from SET being a Category + αh≡βh : PathP (λ i → NHType (αo≡βo i)) αh βh -- αh βh + αh≡βh = isPropHomP αh βh αo≡βo + where + isProp-hom : ⦃ isCatSET : isCategory (SET ℓ') ⦄ → (ϕ : NOType) → isProp (NHType ϕ) + isProp-hom ⦃ isCatSET ⦄ γ = isPropImplicitΠ λ x + → isPropImplicitΠ λ y + → isPropΠ λ f + → isCatSET .isSetHom {x = (C [ c , x ]) , (isCatC .isSetHom)} {F ⟅ y ⟆} _ _ + + isPropHomP : isOfHLevelDep 1 (λ ηo → NHType ηo) + isPropHomP = isOfHLevel→isOfHLevelDep 1 λ a → isProp-hom a + + -- Naturality of the bijection + + -- in the functor + -- it's equivalent to apply ϕ to α then do β ⟦ c ⟧ + -- or apply ϕ the the composite nat trans α ◍ β + -- where ϕ takes a natural transformation to its representing element + yonedaIsNaturalInFunctor : ∀ {F G : Functor C (SET ℓ')} (c : C .ob) + → (β : F ⇒ G) + → (fun (yoneda G c) ◍ compTrans β) ≡ (β ⟦ c ⟧ ◍ fun (yoneda F c)) + yonedaIsNaturalInFunctor {F = F} {G} c β = funExt λ α → refl + + -- in the object + -- it's equivalent to apply ϕ and then F ⟪ f ⟫ + -- or to apply ϕ to the natural transformation obtained by precomposing with f + yonedaIsNaturalInOb : ∀ {F : Functor C (SET ℓ')} + → (c c' : C .ob) + → (f : C [ c , c' ]) + → yoneda F c' .fun ◍ preComp f ≡ F ⟪ f ⟫ ◍ yoneda F c .fun + yonedaIsNaturalInOb {F = F} c c' f = funExt (λ α + → (yoneda F c' .fun ◍ preComp f) α + ≡⟨ refl ⟩ + (α ⟦ c' ⟧) (f ⋆⟨ C ⟩ C .id c') + ≡[ i ]⟨ (α ⟦ c' ⟧) (C .⋆IdR f i) ⟩ + (α ⟦ c' ⟧) f + ≡[ i ]⟨ (α ⟦ c' ⟧) (C .⋆IdL f (~ i)) ⟩ + (α ⟦ c' ⟧) (C .id c ⋆⟨ C ⟩ f) + ≡[ i ]⟨ (α .N-hom f i) (C .id c) ⟩ + (F ⟪ f ⟫) ((α ⟦ c ⟧) (C .id c)) + ≡⟨ refl ⟩ + ((F ⟪ f ⟫) ◍ yoneda F c .fun) α + ∎) + +-- Yoneda embedding +-- TODO: probably want to rename/refactor +module _ {C : Precategory ℓ ℓ} ⦃ C-cat : isCategory C ⦄ where + open Functor + open NatTrans + open Precategory C + + yo : ob → Functor (C ^op) (SET ℓ) + yo x .F-ob y .fst = C [ y , x ] + yo x .F-ob y .snd = C-cat .isSetHom + yo x .F-hom f g = f ⋆⟨ C ⟩ g + yo x .F-id i f = ⋆IdL f i + yo x .F-seq f g i h = ⋆Assoc g f h i + + YO : Functor C (PreShv C ℓ) + YO .F-ob = yo + YO .F-hom f .N-ob z g = g ⋆⟨ C ⟩ f + YO .F-hom f .N-hom g i h = ⋆Assoc g h f i + YO .F-id = makeNatTransPath λ i _ → λ f → ⋆IdR f i + YO .F-seq f g = makeNatTransPath λ i _ → λ h → ⋆Assoc h f g (~ i) + + + module _ {x} (F : Functor (C ^op) (SET ℓ)) where + yo-yo-yo : NatTrans (yo x) F → F .F-ob x .fst + yo-yo-yo α = α .N-ob _ (id _) + + no-no-no : F .F-ob x .fst → NatTrans (yo x) F + no-no-no a .N-ob y f = F .F-hom f a + no-no-no a .N-hom f = funExt λ g i → F .F-seq g f i a + + yoIso : Iso (NatTrans (yo x) F) (F .F-ob x .fst) + yoIso .Iso.fun = yo-yo-yo + yoIso .Iso.inv = no-no-no + yoIso .Iso.rightInv b i = F .F-id i b + yoIso .Iso.leftInv a = makeNatTransPath (funExt λ _ → funExt rem) + where + rem : ∀ {z} (x₁ : C [ z , x ]) → F .F-hom x₁ (yo-yo-yo a) ≡ (a .N-ob z) x₁ + rem g = + F .F-hom g (yo-yo-yo a) + ≡[ i ]⟨ a .N-hom g (~ i) (id x) ⟩ + a .N-hom g i0 (id x) + ≡[ i ]⟨ a .N-ob _ (⋆IdR g i) ⟩ + (a .N-ob _) g + ∎ + + yoEquiv : NatTrans (yo x) F ≃ F .F-ob x .fst + yoEquiv = isoToEquiv yoIso + + + isFullYO : isFull YO + isFullYO x y F[f] = ∣ yo-yo-yo _ F[f] , yoIso {x} (yo y) .Iso.leftInv F[f] ∣ + + isFaithfulYO : isFaithful YO + isFaithfulYO x y f g p i = + hcomp + (λ j → λ{ (i = i0) → ⋆IdL f j; (i = i1) → ⋆IdL g j}) + (yo-yo-yo _ (p i)) diff --git a/Cubical/Codata/Conat.agda b/Cubical/Codata/Conat.agda index b8fb936b8..f35a316c9 100644 --- a/Cubical/Codata/Conat.agda +++ b/Cubical/Codata/Conat.agda @@ -1,4 +1,5 @@ -{-# OPTIONS --cubical #-} +{-# OPTIONS --guardedness #-} + module Cubical.Codata.Conat where open import Cubical.Codata.Conat.Base public diff --git a/Cubical/Codata/Conat/Base.agda b/Cubical/Codata/Conat/Base.agda index 10a1329f9..93974fb4e 100644 --- a/Cubical/Codata/Conat/Base.agda +++ b/Cubical/Codata/Conat/Base.agda @@ -18,7 +18,7 @@ The first approach is chosen to exploit guarded recursion and to avoid the use of Sized Types. -} -{-# OPTIONS --cubical --safe --guardedness #-} +{-# OPTIONS --safe --guardedness #-} module Cubical.Codata.Conat.Base where open import Cubical.Data.Unit diff --git a/Cubical/Codata/Conat/Properties.agda b/Cubical/Codata/Conat/Properties.agda index 031826651..b4469ccf5 100644 --- a/Cubical/Codata/Conat/Properties.agda +++ b/Cubical/Codata/Conat/Properties.agda @@ -20,12 +20,12 @@ The standard library also defines bisimulation on conaturals: https://github.com/agda/agda-stdlib/blob/master/src/Codata/Conat/Bisimilarity.agda -} -{-# OPTIONS --cubical --safe --guardedness #-} +{-# OPTIONS --safe --guardedness #-} module Cubical.Codata.Conat.Properties where open import Cubical.Data.Unit open import Cubical.Data.Sum -open import Cubical.Data.Empty +open import Cubical.Data.Empty as ⊥ open import Cubical.Core.Everything @@ -36,7 +36,6 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence open import Cubical.Relation.Nullary -open import Cubical.Relation.Nullary.DecidableEq open import Cubical.Codata.Conat.Base Unwrap-prev : Conat′ → Type₀ @@ -84,8 +83,26 @@ n+′∞≡∞′ (suc n) = λ i → suc (n+∞≡∞ n i) ∞+∞≡∞ : ∞ + ∞ ≡ ∞ force (∞+∞≡∞ i) = suc (∞+∞≡∞ i) ++-zeroˡ : ∀ n → 𝟘 + n ≡ n +force (+-zeroˡ n _) = force n + ++-zeroʳ : ∀ n → n + 𝟘 ≡ n ++′-zeroʳ : ∀ n → n +′ 𝟘 ≡ n + +force (+-zeroʳ n i) = +′-zeroʳ (force n) i ++′-zeroʳ zero _ = zero ++′-zeroʳ (suc n) i = suc (+-zeroʳ n i) + ++-assoc : ∀ m n p → (m + n) + p ≡ m + (n + p) ++′-assoc : ∀ m n p → (m +′ n) +′ p ≡ m +′ (n + p) + +force (+-assoc m n p i) = +′-assoc (force m) n p i ++′-assoc zero _ _ = refl ++′-assoc (suc m) n p i = suc (+-assoc m n p i) + + conat-absurd : ∀ {y : Conat} {ℓ} {Whatever : Type ℓ} → zero ≡ suc y → Whatever -conat-absurd eq = ⊥-elim (transport (cong diag eq) tt) +conat-absurd eq = ⊥.rec (transport (cong diag eq) tt) where diag : Conat′ → Type₀ diag zero = Unit @@ -99,11 +116,11 @@ module IsSet where ≡′-stable {zero} {zero} ¬¬p′ = refl ≡′-stable {suc x} {suc y} ¬¬p′ = cong′ suc (≡-stable λ ¬p → ¬¬p′ λ p → ¬p (cong pred′′ p)) - ≡′-stable {zero} {suc y} ¬¬p′ = ⊥-elim (¬¬p′ conat-absurd) - ≡′-stable {suc x} {zero} ¬¬p′ = ⊥-elim (¬¬p′ λ p → conat-absurd (sym p)) + ≡′-stable {zero} {suc y} ¬¬p′ = ⊥.rec (¬¬p′ conat-absurd) + ≡′-stable {suc x} {zero} ¬¬p′ = ⊥.rec (¬¬p′ λ p → conat-absurd (sym p)) isSetConat : isSet Conat - isSetConat _ _ = Stable≡→isSet (λ _ _ → ≡-stable) _ _ + isSetConat _ _ = Separated→isSet (λ _ _ → ≡-stable) _ _ isSetConat′ : isSet Conat′ isSetConat′ m n p′ q′ = cong (cong force) (isSetConat (conat m) (conat n) p q) diff --git a/Cubical/Codata/Everything.agda b/Cubical/Codata/Everything.agda index 2802bb300..3031b6a8a 100644 --- a/Cubical/Codata/Everything.agda +++ b/Cubical/Codata/Everything.agda @@ -1,4 +1,5 @@ -{-# OPTIONS --cubical #-} +{-# OPTIONS --guardedness #-} + module Cubical.Codata.Everything where open import Cubical.Codata.EverythingSafe public @@ -13,6 +14,19 @@ open import Cubical.Codata.Conat public open import Cubical.Codata.M public - -- Also uses {-# TERMINATING #-}. open import Cubical.Codata.M.Bisimilarity public + +{- +-- Alternative M type implemetation, based on +-- https://arxiv.org/pdf/1504.02949.pdf +-- "Non-wellfounded trees in Homotopy Type Theory" +-- Benedikt Ahrens, Paolo Capriotti, Régis Spadotti +-} + +open import Cubical.Codata.M.AsLimit.M +open import Cubical.Codata.M.AsLimit.Coalg +open import Cubical.Codata.M.AsLimit.helper +open import Cubical.Codata.M.AsLimit.Container +open import Cubical.Codata.M.AsLimit.itree +open import Cubical.Codata.M.AsLimit.stream diff --git a/Cubical/Codata/EverythingSafe.agda b/Cubical/Codata/EverythingSafe.agda index 22b314237..0f7a36468 100644 --- a/Cubical/Codata/EverythingSafe.agda +++ b/Cubical/Codata/EverythingSafe.agda @@ -1,2 +1,2 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Codata.EverythingSafe where diff --git a/Cubical/Codata/M.agda b/Cubical/Codata/M.agda index b85c44f04..4399a07c6 100644 --- a/Cubical/Codata/M.agda +++ b/Cubical/Codata/M.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe --guardedness #-} +{-# OPTIONS --safe --guardedness #-} module Cubical.Codata.M where open import Cubical.Foundations.Prelude @@ -32,15 +32,15 @@ module Helpers where open Helpers -IxCont : Type₀ → Type₁ -IxCont X = Σ (X → Type₀) \ S → ∀ x → S x → X → Type₀ +IxCont : ∀ {ℓ} -> Type ℓ → Type (ℓ-suc ℓ) +IxCont {ℓ} X = Σ (X → Type ℓ) λ S → ∀ x → S x → X → Type ℓ -⟦_⟧ : ∀ {X : Type₀} → IxCont X → (X → Type₀) → (X → Type₀) -⟦ (S , P) ⟧ X x = Σ (S x) \ s → ∀ y → P x s y → X y +⟦_⟧ : ∀ {ℓ} {X : Type ℓ} → IxCont X → (X → Type ℓ) → (X → Type ℓ) +⟦ (S , P) ⟧ X x = Σ (S x) λ s → ∀ y → P x s y → X y -record M {X : Type₀} (C : IxCont X) (x : X) : Type₀ where +record M {ℓ} {X : Type ℓ} (C : IxCont X) (x : X) : Type ℓ where -- Type₀ coinductive field head : C .fst x @@ -48,11 +48,14 @@ record M {X : Type₀} (C : IxCont X) (x : X) : Type₀ where open M public -module _ {X : Type₀} {C : IxCont X} where +module _ {ℓ} {X : Type ℓ} {C : IxCont X} where private F = ⟦ C ⟧ + inM : ∀ x → F (M C) x → M C x + inM x (head , tail) = record { head = head ; tails = tail } + out : ∀ x → M C x → F (M C) x out x a = (a .head) , (a .tails) diff --git a/Cubical/Codata/M/AsLimit/Coalg.agda b/Cubical/Codata/M/AsLimit/Coalg.agda new file mode 100644 index 000000000..4118df13c --- /dev/null +++ b/Cubical/Codata/M/AsLimit/Coalg.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.Coalg where + +open import Cubical.Codata.M.AsLimit.Coalg.Base public diff --git a/Cubical/Codata/M/AsLimit/Coalg/Base.agda b/Cubical/Codata/M/AsLimit/Coalg/Base.agda new file mode 100644 index 000000000..d4c7105c2 --- /dev/null +++ b/Cubical/Codata/M/AsLimit/Coalg/Base.agda @@ -0,0 +1,39 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.Coalg.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function using ( _∘_ ) +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Data.Unit +open import Cubical.Data.Nat +open import Cubical.Data.Prod +open import Cubical.Data.Sigma + +open import Cubical.Codata.M.AsLimit.Container +open import Cubical.Codata.M.AsLimit.helper + +------------------------------- +-- Definition of a Coalgebra -- +------------------------------- + +Coalg₀ : ∀ {ℓ} {S : Container ℓ} -> Type (ℓ-suc ℓ) +Coalg₀ {ℓ} {S = S} = Σ[ C ∈ Type ℓ ] (C → P₀ S C) + +-------------------------- +-- Definition of a Cone -- +-------------------------- + +Cone₀ : ∀ {ℓ} {S : Container ℓ} {C,γ : Coalg₀ {S = S}} -> Type ℓ +Cone₀ {S = S} {C , _} = (n : ℕ) → C → X (sequence S) n + +Cone₁ : ∀ {ℓ} {S : Container ℓ} {C,γ : Coalg₀ {S = S}} -> (f : Cone₀ {C,γ = C,γ}) -> Type ℓ +Cone₁ {S = S} {C , _} f = (n : ℕ) → π (sequence S) ∘ (f (suc n)) ≡ f n + +Cone : ∀ {ℓ} {S : Container ℓ} (C,γ : Coalg₀ {S = S}) -> Type ℓ +Cone {S = S} C,γ = Σ[ Cone ∈ (Cone₀ {C,γ = C,γ}) ] (Cone₁{C,γ = C,γ} Cone) diff --git a/Cubical/Codata/M/AsLimit/Container.agda b/Cubical/Codata/M/AsLimit/Container.agda new file mode 100644 index 000000000..f0c9cb6e8 --- /dev/null +++ b/Cubical/Codata/M/AsLimit/Container.agda @@ -0,0 +1,88 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.Container where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv using (_≃_) +open import Cubical.Foundations.Function using (_∘_) + +open import Cubical.Data.Unit +open import Cubical.Data.Prod +open import Cubical.Data.Nat as ℕ using (ℕ ; suc ; _+_ ) + +open import Cubical.Foundations.Transport + +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function + +open import Cubical.Data.Sum + +open import Cubical.Foundations.Structure + +open import Cubical.Codata.M.AsLimit.helper + +------------------------------------- +-- Container and Container Functor -- +------------------------------------- + +-- Σ[ A ∈ (Type ℓ) ] (A → Type ℓ) +Container : ∀ ℓ -> Type (ℓ-suc ℓ) +Container ℓ = TypeWithStr ℓ (λ x → x → Type ℓ) + +-- Polynomial functor (P₀ , P₁) defined over a container +-- https://ncatlab.org/nlab/show/polynomial+functor + +-- P₀ object part of functor +P₀ : ∀ {ℓ} (S : Container ℓ) -> Type ℓ -> Type ℓ +P₀ (A , B) X = Σ[ a ∈ A ] (B a -> X) + +-- P₁ morphism part of functor +P₁ : ∀ {ℓ} {S : Container ℓ} {X Y} (f : X -> Y) -> P₀ S X -> P₀ S Y +P₁ {S = S} f = λ { (a , g) -> a , f ∘ g } + +----------------------- +-- Chains and Limits -- +----------------------- + +record Chain ℓ : Type (ℓ-suc ℓ) where + constructor _,,_ + field + X : ℕ -> Type ℓ + π : {n : ℕ} -> X (suc n) -> X n + +open Chain public + +limit-of-chain : ∀ {ℓ} -> Chain ℓ → Type ℓ +limit-of-chain (x ,, pi) = Σ[ x ∈ ((n : ℕ) → x n) ] ((n : ℕ) → pi {n = n} (x (suc n)) ≡ x n) + +shift-chain : ∀ {ℓ} -> Chain ℓ -> Chain ℓ +shift-chain = λ X,π -> ((λ x → X X,π (suc x)) ,, λ {n} → π X,π {suc n}) + +------------------------------------------------------ +-- Limit type of a Container , and Shift of a Limit -- +------------------------------------------------------ + +Wₙ : ∀ {ℓ} -> Container ℓ -> ℕ -> Type ℓ +Wₙ S 0 = Lift Unit +Wₙ S (suc n) = P₀ S (Wₙ S n) + +πₙ : ∀ {ℓ} -> (S : Container ℓ) -> {n : ℕ} -> Wₙ S (suc n) -> Wₙ S n +πₙ {ℓ} S {0} _ = lift tt +πₙ S {suc n} = P₁ (πₙ S {n}) + +sequence : ∀ {ℓ} -> Container ℓ -> Chain ℓ +X (sequence S) n = Wₙ S n +π (sequence S) {n} = πₙ S {n} + +PX,Pπ : ∀ {ℓ} (S : Container ℓ) -> Chain ℓ +PX,Pπ {ℓ} S = + (λ n → P₀ S (X (sequence S) n)) ,, + (λ {n : ℕ} x → P₁ (λ z → z) (π (sequence S) {n = suc n} x )) + +----------------------------------- +-- M type is limit of a sequence -- +----------------------------------- + +M : ∀ {ℓ} -> Container ℓ → Type ℓ +M = limit-of-chain ∘ sequence diff --git a/Cubical/Codata/M/AsLimit/M.agda b/Cubical/Codata/M/AsLimit/M.agda new file mode 100644 index 000000000..22c3b9e65 --- /dev/null +++ b/Cubical/Codata/M/AsLimit/M.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.M where + +open import Cubical.Codata.M.AsLimit.M.Base public +open import Cubical.Codata.M.AsLimit.M.Properties public diff --git a/Cubical/Codata/M/AsLimit/M/Base.agda b/Cubical/Codata/M/AsLimit/M/Base.agda new file mode 100644 index 000000000..a07b2d811 --- /dev/null +++ b/Cubical/Codata/M/AsLimit/M/Base.agda @@ -0,0 +1,194 @@ +{-# OPTIONS --guardedness --safe #-} + +-- Construction of M-types from +-- https://arxiv.org/pdf/1504.02949.pdf +-- "Non-wellfounded trees in Homotopy Type Theory" +-- Benedikt Ahrens, Paolo Capriotti, Régis Spadotti + +module Cubical.Codata.M.AsLimit.M.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv using (_≃_) +open import Cubical.Foundations.Function using (_∘_) + +open import Cubical.Data.Unit +open import Cubical.Data.Prod +open import Cubical.Data.Nat as ℕ using (ℕ ; suc ; _+_ ) +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.Data.Nat.Algebra + +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Data.Sum + +open import Cubical.Codata.M.AsLimit.helper +open import Cubical.Codata.M.AsLimit.Container + +open import Cubical.Codata.M.AsLimit.Coalg.Base + +open Iso + +private + limit-collapse : ∀ {ℓ} {S : Container ℓ} (X : ℕ → Type ℓ) (l : (n : ℕ) → X n → X (suc n)) → (x₀ : X 0) → ∀ (n : ℕ) → X n + limit-collapse X l x₀ 0 = x₀ + limit-collapse {S = S} X l x₀ (suc n) = l n (limit-collapse {S = S} X l x₀ n) + +lemma11-Iso : + ∀ {ℓ} {S : Container ℓ} (X : ℕ → Type ℓ) (l : (n : ℕ) → X n → X (suc n)) + → Iso (Σ[ x ∈ ((n : ℕ) → X n)] ((n : ℕ) → x (suc n) ≡ l n (x n))) + (X 0) +fun (lemma11-Iso X l) (x , y) = x 0 +inv (lemma11-Iso {S = S} X l) x₀ = limit-collapse {S = S} X l x₀ , (λ n → refl {x = limit-collapse {S = S} X l x₀ (suc n)}) +rightInv (lemma11-Iso X l) _ = refl +leftInv (lemma11-Iso {ℓ = ℓ} {S = S} X l) (x , y) i = + let temp = χ-prop (x 0) (fst (inv (lemma11-Iso {S = S} X l) (fun (lemma11-Iso {S = S} X l) (x , y))) , refl , (λ n → refl {x = limit-collapse {S = S} X l (x 0) (suc n)})) (x , refl , y) + in temp i .fst , proj₂ (temp i .snd) + where + open AlgebraPropositionality + open NatSection + + X-fiber-over-ℕ : (x₀ : X 0) -> NatFiber NatAlgebraℕ ℓ + X-fiber-over-ℕ x₀ = record { Fiber = X ; fib-zero = x₀ ; fib-suc = λ {n : ℕ} xₙ → l n xₙ } + + X-section : (x₀ : X 0) → (z : Σ[ x ∈ ((n : ℕ) → X n)] (x 0 ≡ x₀) × (∀ n → (x (suc n)) ≡ l n (x n))) -> NatSection (X-fiber-over-ℕ x₀) + X-section = λ x₀ z → record { section = fst z ; sec-comm-zero = proj₁ (snd z) ; sec-comm-suc = proj₂ (snd z) } + + Z-is-Section : (x₀ : X 0) → + Iso (Σ[ x ∈ ((n : ℕ) → X n)] (x 0 ≡ x₀) × (∀ n → (x (suc n)) ≡ l n (x n))) + (NatSection (X-fiber-over-ℕ x₀)) + fun (Z-is-Section x₀) (x , (z , y)) = record { section = x ; sec-comm-zero = z ; sec-comm-suc = y } + inv (Z-is-Section x₀) x = NatSection.section x , (sec-comm-zero x , sec-comm-suc x) + rightInv (Z-is-Section x₀) _ = refl + leftInv (Z-is-Section x₀) (x , (z , y)) = refl + + -- S≡T + χ-prop' : (x₀ : X 0) → isProp (NatSection (X-fiber-over-ℕ x₀)) + χ-prop' x₀ a b = SectionProp.S≡T isNatInductiveℕ (X-section x₀ (inv (Z-is-Section x₀) a)) (X-section x₀ (inv (Z-is-Section x₀) b)) + + χ-prop : (x₀ : X 0) → isProp (Σ[ x ∈ ((n : ℕ) → X n) ] (x 0 ≡ x₀) × (∀ n → (x (suc n)) ≡ l n (x n))) + χ-prop x₀ = subst isProp (sym (isoToPath (Z-is-Section x₀))) (χ-prop' x₀) + +----------------------------------------------------- +-- Shifting the limit of a chain is an equivalence -- +----------------------------------------------------- + +-- Shift is equivalence (12) and (13) in the proof of Theorem 7 +-- https://arxiv.org/pdf/1504.02949.pdf +-- "Non-wellfounded trees in Homotopy Type Theory" +-- Benedikt Ahrens, Paolo Capriotti, Régis Spadotti + +-- TODO: This definition is inefficient, it should be updated to use some cubical features! +shift-iso : ∀ {ℓ} (S : Container ℓ) -> Iso (P₀ S (M S)) (M S) +shift-iso S@(A , B) = + P₀ S (M S) + Iso⟨ Σ-cong-iso-snd + (λ x → iso (λ f → (λ n z → f z .fst n) , λ n i a → f a .snd n i) + (λ (u , q) z → (λ n → u n z) , λ n i → q n i z) + (λ _ → refl) + (λ _ → refl)) ⟩ + (Σ[ a ∈ A ] (Σ[ u ∈ ((n : ℕ) → B a → X (sequence S) n) ] ((n : ℕ) → π (sequence S) ∘ (u (suc n)) ≡ u n))) + Iso⟨ invIso α-iso-step-5-Iso ⟩ + (Σ[ a ∈ (Σ[ a ∈ ((n : ℕ) → A) ] ((n : ℕ) → a (suc n) ≡ a n)) ] + Σ[ u ∈ ((n : ℕ) → B (a .fst n) → X (sequence S) n) ] + ((n : ℕ) → PathP (λ x → B (a .snd n x) → X (sequence S) n) + (π (sequence S) ∘ u (suc n)) + (u n))) + Iso⟨ α-iso-step-1-4-Iso-lem-12 ⟩ + M S ∎Iso + where + α-iso-step-5-Iso-helper0 : + ∀ (a : (ℕ -> A)) + → (p : (n : ℕ) → a (suc n) ≡ a n) + → (n : ℕ) + → a n ≡ a 0 + α-iso-step-5-Iso-helper0 a p 0 = refl + α-iso-step-5-Iso-helper0 a p (suc n) = p n ∙ α-iso-step-5-Iso-helper0 a p n + + α-iso-step-5-Iso-helper1-Iso : + ∀ (a : ℕ -> A) + → (p : (n : ℕ) → a (suc n) ≡ a n) + → (u : (n : ℕ) → B (a n) → Wₙ S n) + → (n : ℕ) + → (PathP (λ x → B (p n x) → Wₙ S n) (πₙ S ∘ u (suc n)) (u n)) ≡ + (πₙ S ∘ (subst (\k -> B k → Wₙ S (suc n)) (α-iso-step-5-Iso-helper0 a p (suc n))) (u (suc n)) + ≡ subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p n) (u n)) + α-iso-step-5-Iso-helper1-Iso a p u n = + PathP (λ x → B (p n x) → Wₙ S n) (πₙ S ∘ u (suc n)) (u n) + ≡⟨ PathP≡Path (λ x → B (p n x) → Wₙ S n) (πₙ S ∘ u (suc n)) (u n) ⟩ + subst (λ k → B k → Wₙ S n) (p n) (πₙ S ∘ u (suc n)) ≡ (u n) + ≡⟨ (λ i → transp (λ j → B (α-iso-step-5-Iso-helper0 a p n (i ∧ j)) → Wₙ S n) (~ i) (subst (λ k → B k → Wₙ S n) (p n) (πₙ S ∘ u (suc n))) + ≡ transp (λ j → B (α-iso-step-5-Iso-helper0 a p n (i ∧ j)) → Wₙ S n) (~ i) (u n)) ⟩ + subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p n) (subst (λ k → B k → Wₙ S n) (p n) (πₙ S ∘ u (suc n))) ≡ + subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p n) (u n) + ≡⟨ (λ i → sym (substComposite (λ k → B k → Wₙ S n) (p n) (α-iso-step-5-Iso-helper0 a p n) (πₙ S ∘ u (suc n))) i + ≡ subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p n) (u n)) ⟩ + subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p (suc n)) (πₙ S ∘ u (suc n)) ≡ + subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p n) (u n) + ≡⟨ (λ i → substCommSlice (λ k → B k → Wₙ S (suc n)) (λ k → B k → Wₙ S n) + (λ a x x₁ → (πₙ S) (x x₁)) + (α-iso-step-5-Iso-helper0 a p (suc n)) (u (suc n)) i + ≡ subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p n) (u n)) ⟩ + πₙ S ∘ subst (λ k → B k → Wₙ S (suc n)) (α-iso-step-5-Iso-helper0 a p (suc n)) (u (suc n)) + ≡ + subst (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 a p n) (u n) ∎ + + α-iso-step-5-Iso : + Iso + (Σ[ a ∈ (Σ[ a ∈ ((n : ℕ) → A) ] ((n : ℕ) → a (suc n) ≡ a n)) ] + Σ[ u ∈ ((n : ℕ) → B (a .fst n) → X (sequence S) n) ] + ((n : ℕ) → PathP (λ x → B (a .snd n x) → X (sequence S) n) + (π (sequence S) ∘ u (suc n)) + (u n))) + (Σ[ a ∈ A ] (Σ[ u ∈ ((n : ℕ) → B a → X (sequence S) n) ] ((n : ℕ) → π (sequence S) ∘ (u (suc n)) ≡ u n))) + α-iso-step-5-Iso = + Σ-cong-iso (lemma11-Iso {S = S} (λ _ → A) (λ _ x → x)) (λ a,p → + Σ-cong-iso (pathToIso (cong (λ k → (n : ℕ) → k n) (funExt λ n → cong (λ k → B k → Wₙ S n) (α-iso-step-5-Iso-helper0 (a,p .fst) (a,p .snd) n)))) λ u → + pathToIso (cong (λ k → (n : ℕ) → k n) (funExt λ n → α-iso-step-5-Iso-helper1-Iso (a,p .fst) (a,p .snd) u n))) + + α-iso-step-1-4-Iso-lem-12 : + Iso (Σ[ a ∈ (Σ[ a ∈ ((n : ℕ) → A)] ((n : ℕ) → a (suc n) ≡ a n)) ] + (Σ[ u ∈ ((n : ℕ) → B (a .fst n) → X (sequence S) n) ] + ((n : ℕ) → PathP (λ x → B (a .snd n x) → X (sequence S) n) + (π (sequence S) ∘ u (suc n)) + (u n)))) + (limit-of-chain (sequence S)) + fun α-iso-step-1-4-Iso-lem-12 (a , b) = (λ { 0 → lift tt ; (suc n) → (a .fst n) , (b .fst n)}) , λ { 0 → refl {x = lift tt} ; (suc m) i → a .snd m i , b .snd m i } + inv α-iso-step-1-4-Iso-lem-12 x = ((λ n → (x .fst) (suc n) .fst) , λ n i → (x .snd) (suc n) i .fst) , (λ n → (x .fst) (suc n) .snd) , λ n i → (x .snd) (suc n) i .snd + fst (rightInv α-iso-step-1-4-Iso-lem-12 (b , c) i) 0 = lift tt + fst (rightInv α-iso-step-1-4-Iso-lem-12 (b , c) i) (suc n) = refl i + snd (rightInv α-iso-step-1-4-Iso-lem-12 (b , c) i) 0 = refl + snd (rightInv α-iso-step-1-4-Iso-lem-12 (b , c) i) (suc n) = c (suc n) + leftInv α-iso-step-1-4-Iso-lem-12 (a , b) = refl + +shift : ∀ {ℓ} (S : Container ℓ) -> P₀ S (M S) ≡ M S +shift S = isoToPath (shift-iso S) -- lemma 13 & lemma 12 + +-- Transporting along shift + +in-fun : ∀ {ℓ} {S : Container ℓ} -> P₀ S (M S) -> M S +in-fun {S = S} = fun (shift-iso S) + +out-fun : ∀ {ℓ} {S : Container ℓ} -> M S -> P₀ S (M S) +out-fun {S = S} = inv (shift-iso S) + +-- Property of functions into M-types + +lift-to-M : ∀ {ℓ} {A : Type ℓ} {S : Container ℓ} + → (x : (n : ℕ) -> A → X (sequence S) n) + → ((n : ℕ) → (a : A) → π (sequence S) (x (suc n) a) ≡ x n a) + --------------- + → (A → M S) +lift-to-M x p a = (λ n → x n a) , λ n i → p n a i + +lift-direct-M : ∀ {ℓ} {S : Container ℓ} + → (x : (n : ℕ) → X (sequence S) n) + → ((n : ℕ) → π (sequence S) (x (suc n)) ≡ x n) + --------------- + → M S +lift-direct-M x p = x , p diff --git a/Cubical/Codata/M/AsLimit/M/Properties.agda b/Cubical/Codata/M/AsLimit/M/Properties.agda new file mode 100644 index 000000000..adabd33a1 --- /dev/null +++ b/Cubical/Codata/M/AsLimit/M/Properties.agda @@ -0,0 +1,61 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.M.Properties where + +open import Cubical.Data.Unit +open import Cubical.Data.Prod +open import Cubical.Data.Nat as ℕ using (ℕ ; suc ; _+_ ) +open import Cubical.Data.Sum +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv using (_≃_) +open import Cubical.Foundations.Function using (_∘_) +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv + +open import Cubical.Functions.Embedding + +open import Cubical.Codata.M.AsLimit.helper +open import Cubical.Codata.M.AsLimit.M.Base +open import Cubical.Codata.M.AsLimit.Container + +-- in-fun and out-fun are inverse + +open Iso + +in-inverse-out : ∀ {ℓ} {S : Container ℓ} → (in-fun {S = S} ∘ out-fun {S = S}) ≡ idfun (M S) +in-inverse-out {S = S} = subst (λ inv → in-fun {S = S} ∘ inv ≡ idfun (M S)) idpath def where + -- substituting refl makes type-checking work a lot faster, but introduces a transport + -- TODO (2020-05-23): revisit this at some point to see if it's still needed in future versions of agda + def : (in-fun {S = S} ∘ inv (shift-iso S)) ≡ idfun (M S) + def = funExt (rightInv (shift-iso S)) + idpath : inv (shift-iso S) ≡ out-fun {S = S} + idpath = refl + +out-inverse-in : ∀ {ℓ} {S : Container ℓ} → (out-fun {S = S} ∘ in-fun {S = S}) ≡ idfun (P₀ S (M S)) +out-inverse-in {S = S} = subst (λ fun → out-fun {S = S} ∘ fun ≡ idfun (P₀ S (M S))) idpath def where + def : (out-fun {S = S} ∘ fun (shift-iso S)) ≡ idfun (P₀ S (M S)) + def = funExt (leftInv (shift-iso S)) + idpath : fun (shift-iso S) ≡ in-fun {S = S} + idpath = refl + +in-out-id : ∀ {ℓ} {S : Container ℓ} -> ∀ {x y : M S} → (in-fun (out-fun x) ≡ in-fun (out-fun y)) ≡ (x ≡ y) +in-out-id {x = x} {y} i = (in-inverse-out i x) ≡ (in-inverse-out i y) + +-- constructor properties + +in-inj : ∀ {ℓ} {S : Container ℓ} {Z : Type ℓ} → ∀ {f g : Z → P₀ S (M S)} → (in-fun ∘ f ≡ in-fun ∘ g) ≡ (f ≡ g) +in-inj {ℓ} {S = S} {Z = Z} {f = f} {g = g} = iso→fun-Injection-Path {ℓ = ℓ} {A = P₀ S (M S)} {B = M S} {C = Z} (shift-iso S) {f = f} {g = g} + +out-inj : ∀ {ℓ} {S : Container ℓ} {Z : Type ℓ} → ∀ {f g : Z → M S} → (out-fun ∘ f ≡ out-fun ∘ g) ≡ (f ≡ g) +out-inj {ℓ} {S = S} {Z = Z} {f = f} {g = g} = iso→inv-Injection-Path {ℓ = ℓ} {A = P₀ S (M S)} {B = M S} {C = Z} (shift-iso S) {f = f} {g = g} + +in-inj-x : ∀ {ℓ} {S : Container ℓ} → ∀ {x y : P₀ S (M S)} → (in-fun x ≡ in-fun y) ≡ (x ≡ y) +in-inj-x {ℓ} {S = S} {x = x} {y} = iso→fun-Injection-Path-x (shift-iso S) {x} {y} + +out-inj-x : ∀ {ℓ} {S : Container ℓ} → ∀ {x y : M S} → (out-fun x ≡ out-fun y) ≡ (x ≡ y) +out-inj-x {ℓ} {S = S} {x = x} {y} = iso→inv-Injection-Path-x (shift-iso S) {x} {y} diff --git a/Cubical/Codata/M/AsLimit/helper.agda b/Cubical/Codata/M/AsLimit/helper.agda new file mode 100644 index 000000000..33be276fd --- /dev/null +++ b/Cubical/Codata/M/AsLimit/helper.agda @@ -0,0 +1,116 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.helper where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv using (_≃_) +open import Cubical.Foundations.Function using (_∘_) + +open import Cubical.Data.Unit +open import Cubical.Data.Prod +open import Cubical.Data.Nat as ℕ using (ℕ ; suc ; _+_ ) +open import Cubical.Data.Sigma hiding (_×_) + +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +open import Cubical.Functions.Embedding +open import Cubical.Functions.FunExtEquiv + +open Iso + +-- General + +iso→fun-Injection : + ∀ {ℓ} {A B C : Type ℓ} (isom : Iso A B) + → ∀ {f g : C -> A} + → (x : C) → (Iso.fun isom (f x) ≡ Iso.fun isom (g x)) ≡ (f x ≡ g x) +iso→fun-Injection {A = A} {B} {C} isom {f = f} {g} = + isEmbedding→Injection {A = A} {B} {C} (Iso.fun isom) (iso→isEmbedding {A = A} {B} isom) {f = f} {g = g} + +abstract + iso→Pi-fun-Injection : + ∀ {ℓ} {A B C : Type ℓ} (isom : Iso A B) + → ∀ {f g : C -> A} + → Iso (∀ x → (fun isom) (f x) ≡ (fun isom) (g x)) (∀ x → f x ≡ g x) + iso→Pi-fun-Injection {A = A} {B} {C} isom {f = f} {g} = + pathToIso (cong (λ k → ∀ x → k x) (funExt (iso→fun-Injection isom {f = f} {g = g}))) + +iso→fun-Injection-Iso : + ∀ {ℓ} {A B C : Type ℓ} (isom : Iso A B) + → ∀ {f g : C -> A} + → Iso (fun isom ∘ f ≡ fun isom ∘ g) (f ≡ g) +iso→fun-Injection-Iso {A = A} {B} {C} isom {f = f} {g} = + (fun isom) ∘ f ≡ (fun isom) ∘ g + Iso⟨ invIso funExtIso ⟩ + (∀ x → (fun isom) (f x) ≡ (fun isom) (g x)) + Iso⟨ iso→Pi-fun-Injection isom ⟩ + (∀ x → f x ≡ g x) + Iso⟨ funExtIso ⟩ + f ≡ g ∎Iso + +iso→fun-Injection-Path : + ∀ {ℓ} {A B C : Type ℓ} (isom : Iso A B) + → ∀ {f g : C -> A} + → (fun isom ∘ f ≡ fun isom ∘ g) ≡ (f ≡ g) +iso→fun-Injection-Path {A = A} {B} {C} isom {f = f} {g} = + isoToPath (iso→fun-Injection-Iso isom) + +iso→inv-Injection-Path : + ∀ {ℓ} {A B C : Type ℓ} (isom : Iso A B) → + ∀ {f g : C -> B} → + ----------------------- + ((inv isom) ∘ f ≡ (inv isom) ∘ g) ≡ (f ≡ g) +iso→inv-Injection-Path {A = A} {B} {C} isom {f = f} {g} = iso→fun-Injection-Path (invIso isom) + +iso→fun-Injection-Iso-x : + ∀ {ℓ} {A B : Type ℓ} + → (isom : Iso A B) + → ∀ {x y : A} + → Iso ((fun isom) x ≡ (fun isom) y) (x ≡ y) +iso→fun-Injection-Iso-x isom {x} {y} = + let tempx = λ {(lift tt) → x} + tempy = λ {(lift tt) → y} in + fun isom x ≡ fun isom y + Iso⟨ iso (λ x₁ t → x₁) + (λ x₁ → x₁ (lift tt)) + (λ x → refl) + (λ x → refl) ⟩ + (∀ (t : Lift Unit) -> (((fun isom) ∘ tempx) t ≡ ((fun isom) ∘ tempy) t)) + Iso⟨ iso→Pi-fun-Injection isom ⟩ + (∀ (t : Lift Unit) -> tempx t ≡ tempy t) + Iso⟨ iso (λ x₁ → x₁ (lift tt)) + (λ x₁ t → x₁) + (λ x → refl) + (λ x → refl) ⟩ + x ≡ y ∎Iso + +iso→inv-Injection-Iso-x : + ∀ {ℓ} {A B : Type ℓ} + → (isom : Iso A B) + → ∀ {x y : B} + → Iso ((inv isom) x ≡ (inv isom) y) (x ≡ y) +iso→inv-Injection-Iso-x {A = A} {B = B} isom = + iso→fun-Injection-Iso-x {A = B} {B = A} (invIso isom) + +iso→fun-Injection-Path-x : + ∀ {ℓ} {A B : Type ℓ} + → (isom : Iso A B) + → ∀ {x y : A} + → ((fun isom) x ≡ (fun isom) y) ≡ (x ≡ y) +iso→fun-Injection-Path-x isom {x} {y} = + isoToPath (iso→fun-Injection-Iso-x isom) + +iso→inv-Injection-Path-x : + ∀ {ℓ} {A B : Type ℓ} + → (isom : Iso A B) + → ∀ {x y : B} + → ((inv isom) x ≡ (inv isom) y) ≡ (x ≡ y) +iso→inv-Injection-Path-x isom = + isoToPath (iso→inv-Injection-Iso-x isom) + diff --git a/Cubical/Codata/M/AsLimit/itree.agda b/Cubical/Codata/M/AsLimit/itree.agda new file mode 100644 index 000000000..50d8e1075 --- /dev/null +++ b/Cubical/Codata/M/AsLimit/itree.agda @@ -0,0 +1,75 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.itree where + +open import Cubical.Data.Unit +open import Cubical.Data.Prod +open import Cubical.Data.Nat as ℕ using (ℕ ; suc ; _+_ ) +open import Cubical.Data.Sum +open import Cubical.Data.Empty +open import Cubical.Data.Bool + +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude + +open import Cubical.Codata.M.AsLimit.Container +open import Cubical.Codata.M.AsLimit.M +open import Cubical.Codata.M.AsLimit.Coalg + +-- Delay monad defined as an M-type +delay-S : (R : Type₀) -> Container ℓ-zero +delay-S R = (Unit ⊎ R) , λ { (inr _) -> ⊥ ; (inl tt) -> Unit } + +delay : (R : Type₀) -> Type₀ +delay R = M (delay-S R) + +delay-ret : {R : Type₀} -> R -> delay R +delay-ret r = in-fun (inr r , λ ()) + +delay-tau : {R : Type₀} -> delay R -> delay R +delay-tau S = in-fun (inl tt , λ x → S) + +-- Bottom element raised +data ⊥₁ : Type₁ where + +-- TREES +tree-S : (E : Type₀ -> Type₁) (R : Type₀) -> Container (ℓ-suc ℓ-zero) +tree-S E R = (R ⊎ (Σ[ A ∈ Type₀ ] (E A))) , (λ { (inl _) -> ⊥₁ ; (inr (A , e)) -> Lift A } ) + +tree : (E : Type₀ -> Type₁) (R : Type₀) -> Type₁ +tree E R = M (tree-S E R) + +tree-ret : ∀ {E} {R} -> R -> tree E R +tree-ret {E} {R} r = in-fun (inl r , λ ()) + +tree-vis : ∀ {E} {R} -> ∀ {A} -> E A -> (A -> tree E R) -> tree E R +tree-vis {A = A} e k = in-fun (inr (A , e) , λ { (lift x) -> k x } ) + +-- ITrees (and buildup examples) +-- https://arxiv.org/pdf/1906.00046.pdf +-- Interaction Trees: Representing Recursive and Impure Programs in Coq +-- Li-yao Xia, Yannick Zakowski, Paul He, Chung-Kil Hur, Gregory Malecha, Benjamin C. Pierce, Steve Zdancewic + +itree-S : ∀ (E : Type₀ -> Type₁) (R : Type₀) -> Container (ℓ-suc ℓ-zero) +itree-S E R = ((Unit ⊎ R) ⊎ (Σ[ A ∈ Type₀ ] (E A))) , (λ { (inl (inl _)) -> Lift Unit ; (inl (inr _)) -> ⊥₁ ; (inr (A , e)) -> Lift A } ) + +itree : ∀ (E : Type₀ -> Type₁) (R : Type₀) -> Type₁ +itree E R = M (itree-S E R) + +ret' : ∀ {E} {R} -> R -> P₀ (itree-S E R) (itree E R) +ret' {E} {R} r = inl (inr r) , λ () + +tau' : {E : Type₀ -> Type₁} -> {R : Type₀} -> itree E R -> P₀ (itree-S E R) (itree E R) +tau' t = inl (inl tt) , λ x → t + +vis' : ∀ {E} {R} -> ∀ {A : Type₀} -> E A -> (A -> itree E R) -> P₀ (itree-S E R) (itree E R) +vis' {A = A} e k = inr (A , e) , λ { (lift x) -> k x } + +ret : ∀ {E} {R} -> R -> itree E R +ret = in-fun ∘ ret' + +tau : {E : Type₀ -> Type₁} -> {R : Type₀} -> itree E R -> itree E R +tau = in-fun ∘ tau' + +vis : ∀ {E} {R} -> ∀ {A : Type₀} -> E A -> (A -> itree E R) -> itree E R +vis {A = A} e = in-fun ∘ (vis' {A = A} e) diff --git a/Cubical/Codata/M/AsLimit/stream.agda b/Cubical/Codata/M/AsLimit/stream.agda new file mode 100644 index 000000000..f6fc3c19e --- /dev/null +++ b/Cubical/Codata/M/AsLimit/stream.agda @@ -0,0 +1,30 @@ +{-# OPTIONS --guardedness --safe #-} + +module Cubical.Codata.M.AsLimit.stream where + +open import Cubical.Data.Unit + +open import Cubical.Foundations.Prelude + +open import Cubical.Codata.M.AsLimit.M +open import Cubical.Codata.M.AsLimit.helper +open import Cubical.Codata.M.AsLimit.Container + +-------------------------------------- +-- Stream definitions using M.AsLimit -- +-------------------------------------- + +stream-S : ∀ A -> Container ℓ-zero +stream-S A = (A , (λ _ → Unit)) + +stream : ∀ (A : Type₀) -> Type₀ +stream A = M (stream-S A) + +cons : ∀ {A} -> A -> stream A -> stream A +cons x xs = in-fun (x , λ { tt -> xs } ) + +hd : ∀ {A} -> stream A -> A +hd {A} S = out-fun S .fst + +tl : ∀ {A} -> stream A -> stream A +tl {A} S = out-fun S .snd tt diff --git a/Cubical/Codata/M/Bisimilarity.agda b/Cubical/Codata/M/Bisimilarity.agda index 15943a60e..b88086d6e 100644 --- a/Cubical/Codata/M/Bisimilarity.agda +++ b/Cubical/Codata/M/Bisimilarity.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --postfix-projections #-} +{-# OPTIONS --postfix-projections --guardedness #-} module Cubical.Codata.M.Bisimilarity where open import Cubical.Core.Everything diff --git a/Cubical/Codata/Stream.agda b/Cubical/Codata/Stream.agda index 0e22ce71a..45011fd31 100644 --- a/Cubical/Codata/Stream.agda +++ b/Cubical/Codata/Stream.agda @@ -1,4 +1,5 @@ -{-# OPTIONS --cubical #-} +{-# OPTIONS --guardedness #-} + module Cubical.Codata.Stream where open import Cubical.Codata.Stream.Base public diff --git a/Cubical/Codata/Stream/Base.agda b/Cubical/Codata/Stream/Base.agda index 21c4b2d92..d4a62600a 100644 --- a/Cubical/Codata/Stream/Base.agda +++ b/Cubical/Codata/Stream/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe --guardedness #-} +{-# OPTIONS --safe --guardedness #-} module Cubical.Codata.Stream.Base where open import Cubical.Core.Everything diff --git a/Cubical/Codata/Stream/Properties.agda b/Cubical/Codata/Stream/Properties.agda index 086435ac7..85927262c 100644 --- a/Cubical/Codata/Stream/Properties.agda +++ b/Cubical/Codata/Stream/Properties.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe --guardedness #-} +{-# OPTIONS --safe --guardedness #-} module Cubical.Codata.Stream.Properties where open import Cubical.Core.Everything @@ -101,11 +101,11 @@ module Equality≅Bisimulation where misibTransp p = J (λ _ p → cast p ≡ misib p) ((transportRefl refl≈) ∙ (sym misib-refl)) p module Stream≅Nat→ {A : Type₀} where - lookup : {A : Type₀} → Stream A → ℕ → A + lookup : Stream A → ℕ → A lookup xs zero = head xs lookup xs (suc n) = lookup (tail xs) n - tabulate : {A : Type₀} → (ℕ → A) → Stream A + tabulate : (ℕ → A) → Stream A head (tabulate f) = f zero tail (tabulate f) = tabulate (λ n → f (suc n)) diff --git a/Cubical/Core/Everything.agda b/Cubical/Core/Everything.agda index 4e658a94c..9ce877e58 100644 --- a/Cubical/Core/Everything.agda +++ b/Cubical/Core/Everything.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Core.Everything where -- Basic primitives (some are from Agda.Primitive) diff --git a/Cubical/Core/Glue.agda b/Cubical/Core/Glue.agda index 40c5edd5c..f214a683b 100644 --- a/Cubical/Core/Glue.agda +++ b/Cubical/Core/Glue.agda @@ -7,7 +7,7 @@ This file contains: - Glue types -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Core.Glue where open import Cubical.Core.Primitives @@ -64,22 +64,78 @@ unglue φ = prim^unglue {φ = φ} -- Cyril Cohen, Thierry Coquand, Simon Huber, Anders Mörtberg private - Glue-S : (A : Type ℓ) {φ : I} - → (Te : Partial φ (Σ[ T ∈ Type ℓ' ] T ≃ A)) - → Sub (Type ℓ') φ (λ { (φ = i1) → Te 1=1 .fst }) - Glue-S A Te = inS (Glue A Te) - - glue-S : - ∀ {A : Type ℓ} {φ : I} - → {T : Partial φ (Type ℓ')} {e : PartialP φ (λ o → T o ≃ A)} - → (t : PartialP φ T) - → Sub A φ (λ { (φ = i1) → e 1=1 .fst (t 1=1) }) - → Sub (primGlue A T e) φ (λ { (φ = i1) → t 1=1 }) - glue-S t s = inS (glue t (outS s)) - - unglue-S : - ∀ {A : Type ℓ} (φ : I) - {T : Partial φ (Type ℓ')} {e : PartialP φ (λ o → T o ≃ A)} - → (x : primGlue A T e) - → Sub A φ (λ { (φ = i1) → e 1=1 .fst x }) - unglue-S φ x = inS (prim^unglue {φ = φ} x) + module GluePrims (A : Type ℓ) {φ : I} (Te : Partial φ (Σ[ T ∈ Type ℓ' ] T ≃ A)) where + T : Partial φ (Type ℓ') + T φ1 = Te φ1 .fst + e : PartialP φ (λ φ → T φ ≃ A) + e φ1 = Te φ1 .snd + + -- Glue can be seen as a subtype of Type that, at φ, is definitionally equal to the left type + -- of the given equivalences. + Glue-S : Type ℓ' [ φ ↦ T ] + Glue-S = inS (Glue A Te) + + -- Which means partial elements of T are partial elements of Glue + coeT→G : + ∀ (t : PartialP φ T) + → Partial φ (Glue A Te) + coeT→G t (φ = i1) = t 1=1 + + -- ... and elements of Glue can be seen as partial elements of T + coeG→T : + ∀ (g : Glue A Te) + → PartialP φ T + coeG→T g (φ = i1) = g + + -- What about elements that are applied to the equivalences? + trans-e : + ∀ (t : PartialP φ T) + → Partial φ A + trans-e t ϕ1 = equivFun (e ϕ1) (t ϕ1) + + -- glue gives a partial element of Glue given an element of A. Note that it "undoes" + -- the application of the equivalences! + glue-S : + ∀ (t : PartialP φ T) + → A [ φ ↦ trans-e t ] + → Glue A Te [ φ ↦ coeT→G t ] + glue-S t s = inS (glue t (outS s)) + + -- typechecking glue enforces this, e.g. you can not simply write + -- glue-bad : (t : PartialP φ T) → A → Glue A Te + -- glue-bad t s = glue t s + + -- unglue does the inverse: + unglue-S : + ∀ (b : Glue A Te) + → A [ φ ↦ trans-e (coeG→T b) ] + unglue-S b = inS (unglue φ b) + + module GlueTransp (A : I → Type ℓ) (Te : (i : I) → Partial (i ∨ ~ i) (Σ[ T ∈ Type ℓ' ] T ≃ A i)) where + A0 A1 : Type ℓ + A0 = A i0 + A1 = A i1 + T : (i : I) → Partial (i ∨ ~ i) (Type ℓ') + T i φ = Te i φ .fst + e : (i : I) → PartialP (i ∨ ~ i) (λ φ → T i φ ≃ A i) + e i φ = Te i φ .snd + T0 T1 : Type ℓ' + T0 = T i0 1=1 + T1 = T i1 1=1 + e0 : T0 ≃ A0 + e0 = e i0 1=1 + e1 : T1 ≃ A1 + e1 = e i1 1=1 + + open import Cubical.Foundations.Prelude using (transport) + transportA : A0 → A1 + transportA = transport (λ i → A i) + + -- copied over from Foundations/Equiv for readability, can't directly import due to cyclic dependency + invEq : ∀ {X : Type ℓ'} {ℓ''} {Y : Type ℓ''} (w : X ≃ Y) → Y → X + invEq w y = w .snd .equiv-proof y .fst .fst + + -- transport in Glue reduces to transport in A + the application of the equivalences in forward and backward + -- direction. + transp-S : (t0 : T0) → T1 [ i1 ↦ (λ _ → invEq e1 (transportA (equivFun e0 t0))) ] + transp-S t0 = inS (transport (λ i → Glue (A i) (Te i)) t0) diff --git a/Cubical/Core/Id.agda b/Cubical/Core/Id.agda index 3fdb23e13..8229e5a75 100644 --- a/Cubical/Core/Id.agda +++ b/Cubical/Core/Id.agda @@ -1,5 +1,5 @@ {- This file exports the primitives of cubical Id types -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Core.Id where open import Cubical.Core.Primitives hiding ( _≡_ ) diff --git a/Cubical/Core/Primitives.agda b/Cubical/Core/Primitives.agda index a07d80f7b..da206a984 100644 --- a/Cubical/Core/Primitives.agda +++ b/Cubical/Core/Primitives.agda @@ -4,7 +4,7 @@ This file document and export the main primitives of Cubical Agda. It also defines some basic derived operations (composition and filling). -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Core.Primitives where open import Agda.Builtin.Cubical.Path public @@ -21,23 +21,23 @@ open import Agda.Primitive.Cubical public ; primHComp to hcomp ; primTransp to transp ; itIsOne to 1=1 ) + +-- These two are to make sure all the primitives are loaded and ready +-- to compute hcomp/transp for the universe. +import Agda.Builtin.Cubical.Glue +-- HCompU is already imported from Glue, and older Agda versions do +-- not have it. So we comment it out for now. +-- import Agda.Builtin.Cubical.HCompU + open import Agda.Primitive public using ( Level ) renaming ( lzero to ℓ-zero ; lsuc to ℓ-suc ; _⊔_ to ℓ-max + ; Set to Type ; Setω to Typeω ) open import Agda.Builtin.Sigma public -Type : (ℓ : Level) → Set (ℓ-suc ℓ) -Type ℓ = Set ℓ - -Type₀ : Type (ℓ-suc ℓ-zero) -Type₀ = Type ℓ-zero - -Type₁ : Type (ℓ-suc (ℓ-suc ℓ-zero)) -Type₁ = Type (ℓ-suc ℓ-zero) - -- This file document the Cubical Agda primitives. The primitives -- themselves are bound by the Agda files imported above. @@ -120,7 +120,7 @@ private -- * There are cubical subtypes as in CCHM. Note that these are not -- fibrant (hence in Typeω): -_[_↦_] : ∀ {ℓ} (A : Type ℓ) (φ : I) (u : Partial φ A) → Typeω +_[_↦_] : ∀ {ℓ} (A : Type ℓ) (φ : I) (u : Partial φ A) → _ A [ φ ↦ u ] = Sub A φ u infix 4 _[_↦_] @@ -146,9 +146,16 @@ infix 4 _[_↦_] -- * Generalized transport and homogeneous composition [CHM 18]. --- When calling "transp A φ a" Agda makes sure that "A" is constant on "φ". +-- When calling "transp A φ a" Agda makes sure that "A" is constant on "φ" (see below). -- transp : ∀ {ℓ} (A : I → Type ℓ) (φ : I) (a : A i0) → A i1 +-- "A" being constant on "φ" means that "A" should be a constant function whenever the +-- constraint "φ = i1" is satisfied. For example: +-- - If "φ" is "i0" then "A" can be anything, since this condition is vacuously true. +-- - If "φ" is "i1" then "A" must be a constant function. +-- - If "φ" is some in-scope variable "i" then "A" only needs to be a constant function +-- when substituting "i1" for "i". + -- When calling "hcomp A φ u a" Agda makes sure that "a" agrees with "u i0" on "φ". -- hcomp : ∀ {ℓ} {A : Type ℓ} {φ : I} (u : I → Partial φ A) (a : A) → A @@ -202,4 +209,3 @@ infix 2 Σ-syntax Σ-syntax = Σ syntax Σ-syntax A (λ x → B) = Σ[ x ∈ A ] B - diff --git a/Cubical/Data/BinNat.agda b/Cubical/Data/BinNat.agda index 45ad2a2b8..0832a1481 100644 --- a/Cubical/Data/BinNat.agda +++ b/Cubical/Data/BinNat.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.BinNat where open import Cubical.Data.BinNat.BinNat public diff --git a/Cubical/Data/BinNat/BinNat.agda b/Cubical/Data/BinNat/BinNat.agda index 3d0510577..1d2eca79b 100644 --- a/Cubical/Data/BinNat/BinNat.agda +++ b/Cubical/Data/BinNat/BinNat.agda @@ -31,7 +31,7 @@ definitions and examples are: principle where we transport a property that is infeasible to prove by computation for ℕ ([propDoubleℕ]): - 2^20 * 2^10 = 2^5 * (2^15 * 2^10) + 2^20 · 2^10 = 2^5 · (2^15 · 2^10) from the corresponding result on Binℕ which is proved instantly by refl ([propDoubleBinℕ]). @@ -58,7 +58,7 @@ but it has the virtue of being a bit simpler to prove equivalent to http://www.cs.bham.ac.uk/~mhe/agda-new/BinaryNaturals.html -} -{-# OPTIONS --cubical --no-exact-split --safe #-} +{-# OPTIONS --no-exact-split --safe #-} module Cubical.Data.BinNat.BinNat where open import Cubical.Foundations.Prelude @@ -318,13 +318,13 @@ s (NatImplℕ≡Binℕ i) = -- Doubling structures record Double {ℓ} (A : Type ℓ) : Type (ℓ-suc ℓ) where field - -- doubling function computing 2 * x + -- doubling function computing 2 · x double : A → A -- element to double elt : A open Double --- Compute: 2^n * x +-- Compute: 2^n · x doubles : ∀ {ℓ} {A : Type ℓ} (D : Double A) → ℕ → A → A doubles D n x = iter n (double D) x @@ -332,7 +332,7 @@ Doubleℕ : Double ℕ double Doubleℕ = doubleℕ elt Doubleℕ = n1024 where - -- 1024 = 2^8 * 2^2 = 2^10 + -- 1024 = 2^8 · 2^2 = 2^10 n1024 : ℕ n1024 = doublesℕ 8 4 @@ -375,7 +375,7 @@ elt (DoubleBinℕ≡Doubleℕ i) = transp (λ j → Binℕ≡ℕ (i ∨ ~ j)) i -- as a record so that Agda does not try to unfold it eagerly. record propDouble {ℓ} {A : Type ℓ} (D : Double A) : Type ℓ where field - -- 2^20 * e = 2^5 * (2^15 * e) + -- 2^20 · e = 2^5 · (2^15 · e) proof : doubles D 20 (elt D) ≡ doubles D 5 (doubles D 15 (elt D)) open propDouble @@ -404,8 +404,8 @@ propDoubleℕ = transport (λ i → propDouble (DoubleBinℕ≡Doubleℕ i)) pro data binnat : Type₀ where zero : binnat -- 0 - consOdd : binnat → binnat -- 2*n + 1 - consEven : binnat → binnat -- 2*{n+1} + consOdd : binnat → binnat -- 2·n + 1 + consEven : binnat → binnat -- 2·{n+1} binnat→ℕ : binnat → ℕ binnat→ℕ zero = 0 diff --git a/Cubical/Data/Bool.agda b/Cubical/Data/Bool.agda index 6918841ae..5a084c5f9 100644 --- a/Cubical/Data/Bool.agda +++ b/Cubical/Data/Bool.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.Bool where open import Cubical.Data.Bool.Base public diff --git a/Cubical/Data/Bool/Base.agda b/Cubical/Data/Bool/Base.agda index e4c5f3735..f2676e735 100644 --- a/Cubical/Data/Bool/Base.agda +++ b/Cubical/Data/Bool/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.Bool.Base where open import Cubical.Core.Everything @@ -6,15 +6,23 @@ open import Cubical.Core.Everything open import Cubical.Foundations.Prelude open import Cubical.Data.Empty +open import Cubical.Data.Sum.Base +open import Cubical.Data.Unit.Base -open import Cubical.Relation.Nullary +open import Cubical.Relation.Nullary.Base open import Cubical.Relation.Nullary.DecidableEq -- Obtain the booleans open import Agda.Builtin.Bool public +private + variable + ℓ : Level + A : Type ℓ + infixr 6 _and_ infixr 5 _or_ +infix 0 if_then_else_ not : Bool → Bool not true = false @@ -32,16 +40,46 @@ false and true = false true and false = false true and true = true -caseBool : ∀ {ℓ} → {A : Type ℓ} → (a0 aS : A) → Bool → A -caseBool att aff true = att -caseBool att aff false = aff +-- xor / mod-2 addition +_⊕_ : Bool → Bool → Bool +false ⊕ x = x +true ⊕ x = not x + +if_then_else_ : Bool → A → A → A +if true then x else y = x +if false then x else y = y _≟_ : Discrete Bool false ≟ false = yes refl -false ≟ true = no λ p → subst (caseBool ⊥ Bool) p true -true ≟ false = no λ p → subst (caseBool Bool ⊥) p true +false ≟ true = no λ p → subst (λ b → if b then ⊥ else Bool) p true +true ≟ false = no λ p → subst (λ b → if b then Bool else ⊥) p true true ≟ true = yes refl -Dec→Bool : ∀ {ℓ} {A : Type ℓ} → Dec A → Bool +Dec→Bool : Dec A → Bool Dec→Bool (yes p) = true Dec→Bool (no ¬p) = false + +-- Helpers for automatic proof +Bool→Type : Bool → Type₀ +Bool→Type true = Unit +Bool→Type false = ⊥ + +True : Dec A → Type₀ +True Q = Bool→Type (Dec→Bool Q) + +False : Dec A → Type₀ +False Q = Bool→Type (not (Dec→Bool Q)) + +toWitness : {Q : Dec A} → True Q → A +toWitness {Q = yes p} _ = p + +toWitnessFalse : {Q : Dec A} → False Q → ¬ A +toWitnessFalse {Q = no ¬p} _ = ¬p + +dichotomyBool : (x : Bool) → (x ≡ true) ⊎ (x ≡ false) +dichotomyBool true = inl refl +dichotomyBool false = inr refl + +-- TODO: this should be uncommented and implemented using instance arguments +-- _==_ : {dA : Discrete A} → A → A → Bool +-- _==_ {dA = dA} x y = Dec→Bool (dA x y) diff --git a/Cubical/Data/Bool/Properties.agda b/Cubical/Data/Bool/Properties.agda index bc95f5f31..be4cd01c9 100644 --- a/Cubical/Data/Bool/Properties.agda +++ b/Cubical/Data/Bool/Properties.agda @@ -1,11 +1,14 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.Bool.Properties where open import Cubical.Core.Everything +open import Cubical.Functions.Involution + open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport open import Cubical.Foundations.Univalence open import Cubical.Data.Bool.Base @@ -18,16 +21,25 @@ notnot : ∀ x → not (not x) ≡ x notnot true = refl notnot false = refl +notIso : Iso Bool Bool +Iso.fun notIso = not +Iso.inv notIso = not +Iso.rightInv notIso = notnot +Iso.leftInv notIso = notnot + notIsEquiv : isEquiv not -notIsEquiv = isoToIsEquiv (iso not not notnot notnot) +notIsEquiv = involIsEquiv {f = not} notnot notEquiv : Bool ≃ Bool -notEquiv = (not , notIsEquiv) +notEquiv = involEquiv {f = not} notnot notEq : Bool ≡ Bool -notEq = ua notEquiv +notEq = involPath {f = not} notnot private + variable + ℓ : Level + -- This computes to false as expected nfalse : Bool nfalse = transp (λ i → notEq i) i0 true @@ -36,14 +48,33 @@ private nfalsepath : nfalse ≡ false nfalsepath = refl +K-Bool + : (P : {b : Bool} → b ≡ b → Type ℓ) + → (∀{b} → P {b} refl) + → ∀{b} → (q : b ≡ b) → P q +K-Bool P Pr {false} = J (λ{ false q → P q ; true _ → Lift ⊥ }) Pr +K-Bool P Pr {true} = J (λ{ true q → P q ; false _ → Lift ⊥ }) Pr + isSetBool : isSet Bool -isSetBool = Discrete→isSet _≟_ +isSetBool a b = J (λ _ p → ∀ q → p ≡ q) (K-Bool (refl ≡_) refl) true≢false : ¬ true ≡ false -true≢false p = subst (caseBool Bool ⊥) p true +true≢false p = subst (λ b → if b then Bool else ⊥) p true false≢true : ¬ false ≡ true -false≢true p = subst (caseBool ⊥ Bool) p true +false≢true p = subst (λ b → if b then ⊥ else Bool) p true + +¬true→false : (x : Bool) → ¬ x ≡ true → x ≡ false +¬true→false false _ = refl +¬true→false true p = rec (p refl) + +¬false→true : (x : Bool) → ¬ x ≡ false → x ≡ true +¬false→true false p = rec (p refl) +¬false→true true _ = refl + +not≢const : ∀ x → ¬ not x ≡ x +not≢const false = true≢false +not≢const true = false≢true zeroˡ : ∀ x → true or x ≡ true zeroˡ false = refl @@ -85,3 +116,94 @@ or-assoc true y z = or-idem : ∀ x → x or x ≡ x or-idem false = refl or-idem true = refl + +⊕-identityʳ : ∀ x → x ⊕ false ≡ x +⊕-identityʳ false = refl +⊕-identityʳ true = refl + +⊕-comm : ∀ x y → x ⊕ y ≡ y ⊕ x +⊕-comm false false = refl +⊕-comm false true = refl +⊕-comm true false = refl +⊕-comm true true = refl + +⊕-assoc : ∀ x y z → x ⊕ (y ⊕ z) ≡ (x ⊕ y) ⊕ z +⊕-assoc false y z = refl +⊕-assoc true false z = refl +⊕-assoc true true z = notnot z + +not-⊕ˡ : ∀ x y → not (x ⊕ y) ≡ not x ⊕ y +not-⊕ˡ false y = refl +not-⊕ˡ true y = notnot y + +⊕-invol : ∀ x y → x ⊕ (x ⊕ y) ≡ y +⊕-invol false x = refl +⊕-invol true x = notnot x + +isEquiv-⊕ : ∀ x → isEquiv (x ⊕_) +isEquiv-⊕ x = involIsEquiv (⊕-invol x) + +⊕-Path : ∀ x → Bool ≡ Bool +⊕-Path x = involPath {f = x ⊕_} (⊕-invol x) + +⊕-Path-refl : ⊕-Path false ≡ refl +⊕-Path-refl = isInjectiveTransport refl + +¬transportNot : ∀(P : Bool ≡ Bool) b → ¬ (transport P (not b) ≡ transport P b) +¬transportNot P b eq = not≢const b sub + where + sub : not b ≡ b + sub = subst {A = Bool → Bool} (λ f → f (not b) ≡ f b) + (λ i c → transport⁻Transport P c i) (cong (transport⁻ P) eq) + +module BoolReflection where + data Table (A : Type₀) (P : Bool ≡ A) : Type₀ where + inspect : (b c : A) + → transport P false ≡ b + → transport P true ≡ c + → Table A P + + table : ∀ P → Table Bool P + table = J Table (inspect false true refl refl) + + reflLemma : (P : Bool ≡ Bool) + → transport P false ≡ false + → transport P true ≡ true + → transport P ≡ transport (⊕-Path false) + reflLemma P ff tt i false = ff i + reflLemma P ff tt i true = tt i + + notLemma : (P : Bool ≡ Bool) + → transport P false ≡ true + → transport P true ≡ false + → transport P ≡ transport (⊕-Path true) + notLemma P ft tf i false = ft i + notLemma P ft tf i true = tf i + + categorize : ∀ P → transport P ≡ transport (⊕-Path (transport P false)) + categorize P with table P + categorize P | inspect false true p q + = subst (λ b → transport P ≡ transport (⊕-Path b)) (sym p) (reflLemma P p q) + categorize P | inspect true false p q + = subst (λ b → transport P ≡ transport (⊕-Path b)) (sym p) (notLemma P p q) + categorize P | inspect false false p q + = rec (¬transportNot P false (q ∙ sym p)) + categorize P | inspect true true p q + = rec (¬transportNot P false (q ∙ sym p)) + + ⊕-complete : ∀ P → P ≡ ⊕-Path (transport P false) + ⊕-complete P = isInjectiveTransport (categorize P) + + ⊕-comp : ∀ p q → ⊕-Path p ∙ ⊕-Path q ≡ ⊕-Path (q ⊕ p) + ⊕-comp p q = isInjectiveTransport (λ i x → ⊕-assoc q p x i) + + open Iso + + reflectIso : Iso Bool (Bool ≡ Bool) + reflectIso .fun = ⊕-Path + reflectIso .inv P = transport P false + reflectIso .leftInv = ⊕-identityʳ + reflectIso .rightInv P = sym (⊕-complete P) + + reflectEquiv : Bool ≃ (Bool ≡ Bool) + reflectEquiv = isoToEquiv reflectIso diff --git a/Cubical/Data/Bool/SwitchStatement.agda b/Cubical/Data/Bool/SwitchStatement.agda new file mode 100644 index 000000000..f86bc476d --- /dev/null +++ b/Cubical/Data/Bool/SwitchStatement.agda @@ -0,0 +1,42 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Bool.SwitchStatement where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Bool.Base +open import Cubical.Data.Nat + +{- + Switch-case: + + _==_ : A → A → Bool + + _ : B + _ = switch (λ x → x == fixedValue) cases + case value1 ⇒ result1 break + case value2 ⇒ result2 break + ... + case valueN ⇒ resultN break + default⇒ defaultResult +-} + + +private + variable + ℓ ℓ′ : Level + + +infixr 6 default⇒_ +infixr 5 case_⇒_break_ +infixr 4 switch_cases_ + +switch_cases_ : {A : Type ℓ} {B : Type ℓ′} → (A → Bool) → ((A → Bool) → B) → B +switch caseIndicator cases caseData = caseData caseIndicator + +case_⇒_break_ : {A : Type ℓ} {B : Type ℓ′} → A → B → (otherCases : (A → Bool) → B) → (A → Bool) → B +case forValue ⇒ result break otherCases = λ caseIndicator → if (caseIndicator forValue) then result else (otherCases caseIndicator) + +default⇒_ : {A : Type ℓ} {B : Type ℓ′} → B → (A → Bool) → B +default⇒_ value caseIndicator = value diff --git a/Cubical/Data/DescendingList.agda b/Cubical/Data/DescendingList.agda index 0ece5c884..f98f3975e 100644 --- a/Cubical/Data/DescendingList.agda +++ b/Cubical/Data/DescendingList.agda @@ -1,5 +1,9 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.DescendingList where open import Cubical.Data.DescendingList.Base public open import Cubical.Data.DescendingList.Properties public +open import Cubical.Data.DescendingList.Examples public +open import Cubical.Data.DescendingList.Strict public + +open import Cubical.Data.DescendingList.Strict.Properties public diff --git a/Cubical/Data/DescendingList/Base.agda b/Cubical/Data/DescendingList/Base.agda index e2f0d46a5..4ef8e9220 100644 --- a/Cubical/Data/DescendingList/Base.agda +++ b/Cubical/Data/DescendingList/Base.agda @@ -2,7 +2,7 @@ -- Descending lists ------------------------------------------------------------------------ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} open import Cubical.Foundations.Everything diff --git a/Cubical/Data/DescendingList/Examples.agda b/Cubical/Data/DescendingList/Examples.agda index 6f697ebf0..91e364711 100644 --- a/Cubical/Data/DescendingList/Examples.agda +++ b/Cubical/Data/DescendingList/Examples.agda @@ -9,13 +9,13 @@ -- 2. "sorting" finite multisets by converting into sorted lists. ------------------------------------------------------------------------ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.DescendingList.Examples where open import Cubical.Foundations.Everything -open import Cubical.Data.Empty +open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Nat open import Cubical.Relation.Nullary @@ -31,11 +31,11 @@ open import Cubical.HITs.FiniteMultiset infix 4 _≤_ _≥_ -data _≤_ : ℕ → ℕ → Set where +data _≤_ : ℕ → ℕ → Type where z≤n : ∀ {n} → zero ≤ n s≤s : ∀ {m n} (m≤n : m ≤ n) → suc m ≤ suc n -_≥_ : ℕ → ℕ → Set +_≥_ : ℕ → ℕ → Type m ≥ n = n ≤ m ≤pred : {n m : ℕ} → suc n ≤ suc m → n ≤ m @@ -55,7 +55,7 @@ m ≥ n = n ≤ m ≰→≥ : {x y : ℕ} → ¬ (x ≥ y) → y ≥ x ≰→≥ {zero} {y} f = z≤n -≰→≥ {suc x} {zero} f = ⊥-elim (f z≤n) +≰→≥ {suc x} {zero} f = ⊥.rec (f z≤n) ≰→≥ {suc x} {suc y} f = s≤s (≰→≥ λ g → f (s≤s g)) ≥trans : {x y z : ℕ} → x ≥ y → y ≥ z → x ≥ z @@ -91,13 +91,14 @@ ValueOfl3 = refl l3=l2++l1 : l3 ≡ l2 ++ᴰᴸ l1 l3=l2++l1 = refl -LongerExample : l1 ++ᴰᴸ l2 ++ᴰᴸ l1 ++ᴰᴸ l1 ++ᴰᴸ l2 - ≡ l2 ++ᴰᴸ l1 ++ᴰᴸ l2 ++ᴰᴸ l1 ++ᴰᴸ l1 -LongerExample = refl +-- Commented as it was the slowest definition in the whole library :-) +-- LongerExample : l1 ++ᴰᴸ l2 ++ᴰᴸ l1 ++ᴰᴸ l1 ++ᴰᴸ l2 +-- ≡ l2 ++ᴰᴸ l1 ++ᴰᴸ l2 ++ᴰᴸ l1 ++ᴰᴸ l1 +-- LongerExample = refl ------------------------------------------------------------------------- --- A simple example of sorting finite multisets +-- ------------------------------------------------------------------------ +-- -- A simple example of sorting finite multisets m1 : FMSet ℕ m1 = 13 ∷ 9 ∷ 78 ∷ 31 ∷ 86 ∷ 3 ∷ 0 ∷ 99 ∷ [] diff --git a/Cubical/Data/DescendingList/Properties.agda b/Cubical/Data/DescendingList/Properties.agda index a7e5c0404..1992c02cc 100644 --- a/Cubical/Data/DescendingList/Properties.agda +++ b/Cubical/Data/DescendingList/Properties.agda @@ -10,18 +10,18 @@ -- properties by transporting those on finite multisets. ------------------------------------------------------------------------ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} open import Cubical.Foundations.Everything -open import Cubical.Data.Empty +open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Unit open import Cubical.Data.List using (List ; [] ; _∷_) open import Cubical.Relation.Nullary open import Cubical.Relation.Nullary.DecidableEq -open import Cubical.HITs.FiniteMultiset hiding ([_]) +open import Cubical.HITs.FiniteMultiset as FMSet hiding ([_]) module Cubical.Data.DescendingList.Properties @@ -175,7 +175,7 @@ module IsoToFMSet insert-swap x y [] | yes x≥y | no y≱x = cons≡ refl (cons≡ refl refl) insert-swap x y [] | no x≱y with ≥dec y x insert-swap x y [] | no x≱y | yes y≥x = cons≡ refl (cons≡ refl refl) - insert-swap x y [] | no x≱y | no y≱x = ⊥-elim (x≱y (≰→≥ y≱x)) + insert-swap x y [] | no x≱y | no y≱x = ⊥.rec (x≱y (≰→≥ y≱x)) insert-swap x y (cons z u z≥u) with ≥dec y z insert-swap x y (cons z u z≥u) | yes y≥z with ≥dec x y insert-swap x y (cons z u z≥u) | yes y≥z | yes x≥y with ≥dec x z @@ -183,28 +183,28 @@ module IsoToFMSet insert-swap x y (cons z u z≥u) | yes y≥z | yes x≥y | yes x≥z | yes y≥x = cons≡ (≤≥→≡ x≥y y≥x) (cons≡ (≤≥→≡ y≥x x≥y) refl) insert-swap x y (cons z u z≥u) | yes y≥z | yes x≥y | yes x≥z | no y≱x with ≥dec y z insert-swap x y (cons z u z≥u) | yes y≥z | yes x≥y | yes x≥z | no y≱x | yes y≥z' = cons≡ refl (cons≡ refl refl) - insert-swap x y (cons z u z≥u) | yes y≥z | yes x≥y | yes x≥z | no y≱x | no y≱z = ⊥-elim (y≱z y≥z) - insert-swap x y (cons z u z≥u) | yes y≥z | yes x≥y | no x≱z = ⊥-elim (x≱z (≥trans x≥y y≥z)) + insert-swap x y (cons z u z≥u) | yes y≥z | yes x≥y | yes x≥z | no y≱x | no y≱z = ⊥.rec (y≱z y≥z) + insert-swap x y (cons z u z≥u) | yes y≥z | yes x≥y | no x≱z = ⊥.rec (x≱z (≥trans x≥y y≥z)) insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y with ≥dec x z insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y | yes x≥z with ≥dec y x insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y | yes x≥z | yes y≥x = cons≡ refl (cons≡ refl refl) - insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y | yes x≥z | no y≱x = ⊥-elim (x≱y (≰→≥ y≱x)) + insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y | yes x≥z | no y≱x = ⊥.rec (x≱y (≰→≥ y≱x)) insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y | no x≱z with ≥dec y z insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y | no x≱z | yes y≥z' = cons≡ refl (cons≡ refl refl) - insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y | no x≱z | no y≱z = ⊥-elim (y≱z y≥z) + insert-swap x y (cons z u z≥u) | yes y≥z | no x≱y | no x≱z | no y≱z = ⊥.rec (y≱z y≥z) insert-swap x y (cons z u z≥u) | no y≱z with ≥dec x z insert-swap x y (cons z u z≥u) | no y≱z | yes x≥z with ≥dec y x - insert-swap x y (cons z u z≥u) | no y≱z | yes x≥z | yes y≥x = ⊥-elim (y≱z (≥trans y≥x x≥z)) + insert-swap x y (cons z u z≥u) | no y≱z | yes x≥z | yes y≥x = ⊥.rec (y≱z (≥trans y≥x x≥z)) insert-swap x y (cons z u z≥u) | no y≱z | yes x≥z | no y≱x with ≥dec y z - insert-swap x y (cons z u z≥u) | no y≱z | yes x≥z | no y≱x | yes y≥z = ⊥-elim (y≱z y≥z) + insert-swap x y (cons z u z≥u) | no y≱z | yes x≥z | no y≱x | yes y≥z = ⊥.rec (y≱z y≥z) insert-swap x y (cons z u z≥u) | no y≱z | yes x≥z | no y≱x | no y≱z' = cons≡ refl (cons≡ refl refl) insert-swap x y (cons z u z≥u) | no y≱z | no x≱z with ≥dec y z - insert-swap x y (cons z u z≥u) | no y≱z | no x≱z | yes y≥z = ⊥-elim (y≱z y≥z) + insert-swap x y (cons z u z≥u) | no y≱z | no x≱z | yes y≥z = ⊥.rec (y≱z y≥z) insert-swap x y (cons z u z≥u) | no y≱z | no x≱z | no y≱z' = cons≡ refl (insert-swap x y u) -- Insertion sort toDL : FMSet A → DL - toDL = FMSetRec.f isSetDL [] insert insert-swap + toDL = FMSet.Rec.f isSetDL [] insert insert-swap {- toDL [] = [] toDL (x ∷ u) = insert x (toDL u) @@ -217,7 +217,7 @@ module IsoToFMSet insert-cons x [] _ = cons≡ refl refl insert-cons x (cons y u _) _ with ≥dec x y insert-cons x (cons y u _) _ | yes x≥y = cons≡ refl refl - insert-cons x (cons y u _) (≥ᴴcons x≥y) | no x≱y = ⊥-elim (x≱y x≥y) + insert-cons x (cons y u _) (≥ᴴcons x≥y) | no x≱y = ⊥.rec (x≱y x≥y) toDL∘toFMSet≡id : (u : DL) → toDL (toFMSet u) ≡ u toDL∘toFMSet≡id [] = refl @@ -234,9 +234,9 @@ module IsoToFMSet insert-∷ x (cons y u _) | no x≱y = cong (λ z → y ∷ z) (insert-∷ x u) ∙ comm y x (toFMSet u) toFMSet∘toDL≡id : (u : FMSet A) → toFMSet (toDL u) ≡ u - toFMSet∘toDL≡id = FMSetElimProp.f (trunc _ _) - refl - (λ x {u} p → insert-∷ x (toDL u) ∙ cong (λ z → x ∷ z) p) + toFMSet∘toDL≡id = FMSet.ElimProp.f (trunc _ _) + refl + (λ x {u} p → insert-∷ x (toDL u) ∙ cong (λ z → x ∷ z) p) FMSet≡DL : FMSet A ≡ DL FMSet≡DL = isoToPath (iso toDL toFMSet toDL∘toFMSet≡id toFMSet∘toDL≡id) diff --git a/Cubical/Data/DescendingList/Strict.agda b/Cubical/Data/DescendingList/Strict.agda index 49f3ac2f3..4141c6d76 100644 --- a/Cubical/Data/DescendingList/Strict.agda +++ b/Cubical/Data/DescendingList/Strict.agda @@ -2,7 +2,7 @@ -- Strictly descending lists ------------------------------------------------------------------------ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} open import Cubical.Core.Everything diff --git a/Cubical/Data/DescendingList/Strict/Properties.agda b/Cubical/Data/DescendingList/Strict/Properties.agda index 14a11f249..ed23dc000 100644 --- a/Cubical/Data/DescendingList/Strict/Properties.agda +++ b/Cubical/Data/DescendingList/Strict/Properties.agda @@ -1,19 +1,22 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} -open import Cubical.Core.Everything -open import Cubical.Foundations.Everything -open import Agda.Primitive using (lzero) +open import Cubical.Foundations.Prelude module Cubical.Data.DescendingList.Strict.Properties (A : Type₀) (_>_ : A → A → Type₀) where +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + open import Cubical.Data.DescendingList.Strict A _>_ -open import Cubical.HITs.ListedFiniteSet renaming (_∈_ to _∈ʰ_) +open import Cubical.HITs.ListedFiniteSet as LFSet renaming (_∈_ to _∈ʰ_) -import Cubical.Data.Empty as Empty -open Empty using (⊥-elim) +import Cubical.Data.Empty as ⊥ open import Cubical.Relation.Nullary.DecidableEq open import Cubical.Relation.Nullary using (Dec; Discrete) renaming (¬_ to Type¬_) @@ -24,7 +27,7 @@ unsort (cons x xs x>xs) = x ∷ unsort xs module _ where open import Cubical.Relation.Nullary - data Tri (A B C : Set) : Set where + data Tri (A B C : Type) : Type where tri-< : A → ¬ B → ¬ C → Tri A B C tri-≡ : ¬ A → B → ¬ C → Tri A B C tri-> : ¬ A → ¬ B → C → Tri A B C @@ -37,41 +40,41 @@ module IsoToLFSet (>-irreflexive : ∀ {x} → Type¬ x > x) where - Tri' : A → A → Set + Tri' : A → A → Type Tri' x y = Tri (y > x) (x ≡ y) (x > y) - open import Cubical.Foundations.Logic - open import Cubical.HITs.PropositionalTruncation + open import Cubical.HITs.PropositionalTruncation as PropTrunc + open import Cubical.Functions.Logic -- Membership is defined via `LFSet`. -- This computes just as well as a direct inductive definition, -- and additionally lets us use the extra `comm` and `dup` paths to prove -- things about membership. - _∈ˡ_ : A → SDL → hProp lzero + _∈ˡ_ : A → SDL → hProp ℓ-zero a ∈ˡ l = a ∈ʰ unsort l - Memˡ : SDL → A → hProp lzero + Memˡ : SDL → A → hProp ℓ-zero Memˡ l a = a ∈ˡ l - Memʰ : LFSet A → A → hProp lzero + Memʰ : LFSet A → A → hProp ℓ-zero Memʰ l a = a ∈ʰ l >ᴴ-trans : ∀ x y zs → x > y → y >ᴴ zs → x >ᴴ zs >ᴴ-trans x y [] x>y y>zs = >ᴴ[] >ᴴ-trans x y (cons z zs _) x>y (>ᴴcons y>z) = >ᴴcons (>-trans x>y y>z) - ≡ₚ-sym : ∀ {A : Set} {x y : A} → [ x ≡ₚ y ] → [ y ≡ₚ x ] - ≡ₚ-sym p = recPropTrunc squash (λ p → ∣ sym p ∣) p + ≡ₚ-sym : ∀ {A : Type} {x y : A} → ⟨ x ≡ₚ y ⟩ → ⟨ y ≡ₚ x ⟩ + ≡ₚ-sym p = PropTrunc.rec squash (λ p → ∣ sym p ∣) p - >-all : ∀ x l → x >ᴴ l → ∀ a → [ a ∈ˡ l ] → x > a + >-all : ∀ x l → x >ᴴ l → ∀ a → ⟨ a ∈ˡ l ⟩ → x > a >-all x (cons y zs y>zs) (>ᴴcons x>y) a a∈l = ⊔-elim (a ≡ₚ y) (a ∈ˡ zs) (λ _ → (x > a) , >-isProp {x} {a}) (λ a≡ₚy → substₚ (λ q → x > q , >-isProp) (≡ₚ-sym a≡ₚy) x>y) (λ a∈zs → >-all x zs (>ᴴ-trans x y zs x>y y>zs) a a∈zs) a∈l - >-absent : ∀ x l → x >ᴴ l → [ ¬ (x ∈ˡ l) ] - >-absent x l x>l x∈l = ⊥-elim (>-irreflexive (>-all x l x>l x x∈l)) + >-absent : ∀ x l → x >ᴴ l → ⟨ ¬ (x ∈ˡ l) ⟩ + >-absent x l x>l x∈l = ⊥.rec (>-irreflexive (>-all x l x>l x x∈l)) >ᴴ-isProp : ∀ x xs → isProp (x >ᴴ xs) >ᴴ-isProp x _ >ᴴ[] >ᴴ[] = refl @@ -112,23 +115,23 @@ module IsoToLFSet abstract -- for some reason, making [exclude] non-abstract makes -- typechecking noticeably slower - exclude : A → (A → hProp lzero) → (A → hProp lzero) + exclude : A → (A → hProp ℓ-zero) → (A → hProp ℓ-zero) exclude x h a = ¬ a ≡ₚ x ⊓ h a >-excluded : ∀ x xs → x >ᴴ xs → exclude x (Memʰ (x ∷ unsort xs)) ≡ Memˡ xs >-excluded x xs x>xs = funExt (λ a → ⇔toPath (to a) (from a)) where - import Cubical.Data.Prod as D + import Cubical.Data.Sigma as D import Cubical.Data.Sum as D - from : ∀ a → [ a ∈ˡ xs ] → [ ¬ a ≡ₚ x ⊓ (a ≡ₚ x ⊔ a ∈ˡ xs) ] - from a a∈xs = (recPropTrunc (snd ⊥) a≢x) D., inr a∈xs where + from : ∀ a → ⟨ a ∈ˡ xs ⟩ → ⟨ ¬ a ≡ₚ x ⊓ (a ≡ₚ x ⊔ a ∈ˡ xs) ⟩ + from a a∈xs = (PropTrunc.rec (snd ⊥) a≢x) D., inr a∈xs where a≢x : Type¬ (a ≡ x) - a≢x = λ a≡x → (>-absent x xs x>xs (transport (λ i → [ a≡x i ∈ˡ xs ]) a∈xs )) + a≢x = λ a≡x → (>-absent x xs x>xs (transport (λ i → ⟨ a≡x i ∈ˡ xs ⟩) a∈xs )) - to : ∀ a → [ ¬ a ≡ₚ x ⊓ (a ≡ₚ x ⊔ a ∈ˡ xs) ] → [ a ∈ˡ xs ] - to a (a≢x D., x) = recPropTrunc (snd (a ∈ˡ xs)) (λ { - (D.inl a≡x) → ⊥-elim (a≢x a≡x); + to : ∀ a → ⟨ ¬ a ≡ₚ x ⊓ (a ≡ₚ x ⊔ a ∈ˡ xs) ⟩ → ⟨ a ∈ˡ xs ⟩ + to a (a≢x D., x) = PropTrunc.rec (snd (a ∈ˡ xs)) (λ { + (D.inl a≡x) → ⊥.rec (a≢x a≡x); (D.inr x) → x }) x cons-eq : ∀ x xs x>xs y ys y>ys → x ≡ y → xs ≡ ys → cons x xs x>xs ≡ cons y ys y>ys @@ -150,24 +153,24 @@ module IsoToLFSet Memˡ-inj : ∀ l₁ l₂ → Memˡ l₁ ≡ Memˡ l₂ → l₁ ≡ l₂ Memˡ-inj [] [] eq = refl - Memˡ-inj [] (cons y ys y>ys) eq = ⊥-elim (transport (λ i → [ eq (~ i) y ]) (inl ∣ refl ∣)) - Memˡ-inj (cons y ys y>ys) [] eq = ⊥-elim (transport (λ i → [ eq i y ]) (inl ∣ refl ∣)) + Memˡ-inj [] (cons y ys y>ys) eq = ⊥.rec (lower (transport (λ i → ⟨ eq (~ i) y ⟩) (inl ∣ refl ∣))) + Memˡ-inj (cons y ys y>ys) [] eq = ⊥.rec (lower (transport (λ i → ⟨ eq i y ⟩) (inl ∣ refl ∣))) Memˡ-inj (cons x xs x>xs) (cons y ys y>ys) e = ⊔-elim (x ≡ₚ y) (x ∈ʰ unsort ys) (λ _ → ((cons x xs x>xs) ≡ (cons y ys y>ys)) , SDL-isSet _ _) - (recPropTrunc (SDL-isSet _ _) with-x≡y) - (⊥-elim ∘ x∉ys) - (transport (λ i → [ e i x ]) (inl ∣ refl ∣)) where + (PropTrunc.rec (SDL-isSet _ _) with-x≡y) + (⊥.rec ∘ x∉ys) + (transport (λ i → ⟨ e i x ⟩) (inl ∣ refl ∣)) where xxs = cons x xs x>xs - x∉ys : [ ¬ x ∈ˡ ys ] - x∉ys x∈ys = ⊥-elim (>-irreflexive y>y) where + x∉ys : ⟨ ¬ x ∈ˡ ys ⟩ + x∉ys x∈ys = ⊥.rec (>-irreflexive y>y) where y>x : y > x y>x = (>-all y ys y>ys x x∈ys) - y∈xxs : [ y ∈ˡ (cons x xs x>xs) ] - y∈xxs = (transport (λ i → [ e (~ i) y ]) (inl ∣ refl ∣)) + y∈xxs : ⟨ y ∈ˡ (cons x xs x>xs) ⟩ + y∈xxs = (transport (λ i → ⟨ e (~ i) y ⟩) (inl ∣ refl ∣)) y>y : y > y y>y = >-all y xxs (>ᴴcons y>x) y y∈xxs @@ -205,11 +208,11 @@ module IsoToLFSet ) sort : LFSet A → SDL - sort = LFSetRec.f [] insert insert-swap insert-dup SDL-isSet + sort = LFSet.Rec.f [] insert insert-swap insert-dup SDL-isSet unsort∘sort : ∀ x → unsort (sort x) ≡ x unsort∘sort = - LFPropElim.f (λ x → unsort (sort x) ≡ x) + LFSet.PropElim.f refl (λ x {ys} ys-hyp → insert-correct x (sort ys) ∙ cong (λ q → x ∷ q) ys-hyp) (λ xs → trunc (unsort (sort xs)) xs) diff --git a/Cubical/Data/DiffInt.agda b/Cubical/Data/DiffInt.agda deleted file mode 100644 index eb47c48d0..000000000 --- a/Cubical/Data/DiffInt.agda +++ /dev/null @@ -1,5 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.Data.DiffInt where - -open import Cubical.Data.DiffInt.Base public -open import Cubical.Data.DiffInt.Properties public diff --git a/Cubical/Data/DiffInt/Base.agda b/Cubical/Data/DiffInt/Base.agda deleted file mode 100644 index f597c5ad7..000000000 --- a/Cubical/Data/DiffInt/Base.agda +++ /dev/null @@ -1,18 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.Data.DiffInt.Base where - -open import Cubical.Foundations.Prelude - -open import Cubical.HITs.SetQuotients.Base - -open import Cubical.Data.Prod -open import Cubical.Data.Nat - -rel : (ℕ ×Σ ℕ) → (ℕ ×Σ ℕ) → Type₀ -rel (a₀ , b₀) (a₁ , b₁) = x ≡ y - where - x = a₀ + b₁ - y = a₁ + b₀ - -ℤ = (ℕ ×Σ ℕ) / rel - diff --git a/Cubical/Data/Empty.agda b/Cubical/Data/Empty.agda index a72500f24..d3e44609c 100644 --- a/Cubical/Data/Empty.agda +++ b/Cubical/Data/Empty.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.Empty where open import Cubical.Data.Empty.Base public diff --git a/Cubical/Data/Empty/Base.agda b/Cubical/Data/Empty/Base.agda index cb7b69d4d..c6384d3b0 100644 --- a/Cubical/Data/Empty/Base.agda +++ b/Cubical/Data/Empty/Base.agda @@ -1,12 +1,20 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.Empty.Base where open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude + +private + variable + ℓ : Level data ⊥ : Type₀ where -⊥-elim : ∀ {ℓ} {A : Type ℓ} → ⊥ → A -⊥-elim () +⊥* : Type ℓ +⊥* = Lift ⊥ + +rec : {A : Type ℓ} → ⊥ → A +rec () -⊥-elimDep : ∀ {ℓ} {A : ⊥ → Type ℓ} → (x : ⊥) → A x -⊥-elimDep () +elim : {A : ⊥ → Type ℓ} → (x : ⊥) → A x +elim () diff --git a/Cubical/Data/Empty/Properties.agda b/Cubical/Data/Empty/Properties.agda index 3b6d6b323..979b8734e 100644 --- a/Cubical/Data/Empty/Properties.agda +++ b/Cubical/Data/Empty/Properties.agda @@ -1,15 +1,30 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.Empty.Properties where open import Cubical.Core.Everything open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism open import Cubical.Data.Empty.Base isProp⊥ : isProp ⊥ -isProp⊥ x = ⊥-elim x +isProp⊥ () + +isProp⊥* : ∀ {ℓ} → isProp {ℓ} ⊥* +isProp⊥* _ () isContr⊥→A : ∀ {ℓ} {A : Type ℓ} → isContr (⊥ → A) -fst isContr⊥→A = ⊥-elim +fst isContr⊥→A () snd isContr⊥→A f i () + +uninhabEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → (A → ⊥) → (B → ⊥) → A ≃ B +uninhabEquiv ¬a ¬b = isoToEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun a = rec (¬a a) + isom .inv b = rec (¬b b) + isom .rightInv b = rec (¬b b) + isom .leftInv a = rec (¬a a) diff --git a/Cubical/Data/Equality.agda b/Cubical/Data/Equality.agda index 81f478ff5..044216caa 100644 --- a/Cubical/Data/Equality.agda +++ b/Cubical/Data/Equality.agda @@ -8,7 +8,7 @@ and the inductively define equality types. TODO: reconsider naming scheme. -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.Equality where @@ -24,7 +24,7 @@ open import Cubical.Foundations.Isomorphism private variable ℓ : Level - A : Set ℓ + A : Type ℓ x y z : A ptoc : x ≡p y → x ≡c y diff --git a/Cubical/Data/Everything.agda b/Cubical/Data/Everything.agda index 536cc11c2..bcb81412e 100644 --- a/Cubical/Data/Everything.agda +++ b/Cubical/Data/Everything.agda @@ -1,23 +1,49 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe --guardedness #-} module Cubical.Data.Everything where -open import Cubical.Data.BinNat public -open import Cubical.Data.Bool public -open import Cubical.Data.Empty public -open import Cubical.Data.Equality public -open import Cubical.Data.Fin public -open import Cubical.Data.Nat public -open import Cubical.Data.Nat.Algebra public -open import Cubical.Data.NatMinusOne hiding (1+_; -1+_) public -open import Cubical.Data.NatMinusTwo public -open import Cubical.Data.Int public renaming (_+_ to _+Int_ ; +-assoc to +Int-assoc; +-comm to +Int-comm) -open import Cubical.Data.Sum public -open import Cubical.Data.Prod public -open import Cubical.Data.Unit public -open import Cubical.Data.Sigma public -open import Cubical.Data.DiffInt public -open import Cubical.Data.Group public hiding (_≃_) -open import Cubical.Data.HomotopyGroup public -open import Cubical.Data.List public -open import Cubical.Data.Graph public -open import Cubical.Data.InfNat public renaming (_*_ to _*ℕ+∞_; _+_ to _+ℕ+∞_; suc to ∞suc; zero to ∞zero) +import Cubical.Data.BinNat +import Cubical.Data.Bool +import Cubical.Data.Bool.SwitchStatement +import Cubical.Data.DescendingList +import Cubical.Data.Empty +import Cubical.Data.Equality +import Cubical.Data.Fin +import Cubical.Data.Fin.LehmerCode +import Cubical.Data.Fin.Recursive +import Cubical.Data.FinData +import Cubical.Data.FinInd +import Cubical.Data.FinSet +import Cubical.Data.FinSet.Binary.Large +import Cubical.Data.FinSet.Binary.Small +import Cubical.Data.Graph +import Cubical.Data.HomotopyGroup +import Cubical.Data.InfNat +import Cubical.Data.Int +import Cubical.Data.Int.MoreInts.BiInvInt +import Cubical.Data.Int.MoreInts.DeltaInt +import Cubical.Data.Int.MoreInts.DiffInt +import Cubical.Data.Int.MoreInts.QuoInt +import Cubical.Data.List +import Cubical.Data.Maybe +import Cubical.Data.Nat +import Cubical.Data.Nat.Algebra +import Cubical.Data.Nat.Coprime +import Cubical.Data.Nat.Divisibility +import Cubical.Data.Nat.GCD +import Cubical.Data.Nat.Literals +import Cubical.Data.Nat.Order +import Cubical.Data.Nat.Order.Recursive +import Cubical.Data.NatMinusOne +import Cubical.Data.NatPlusOne +import Cubical.Data.Prod +import Cubical.Data.Queue +import Cubical.Data.Queue.1List +import Cubical.Data.Queue.Truncated2List +import Cubical.Data.Queue.Untruncated2List +import Cubical.Data.Queue.Untruncated2ListInvariant +import Cubical.Data.Sigma +import Cubical.Data.SubFinSet +import Cubical.Data.Sum +import Cubical.Data.SumFin +import Cubical.Data.Unit +import Cubical.Data.Vec diff --git a/Cubical/Data/Fin.agda b/Cubical/Data/Fin.agda index 8914a32f3..458cb01e0 100644 --- a/Cubical/Data/Fin.agda +++ b/Cubical/Data/Fin.agda @@ -1,6 +1,7 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.Fin where -open import Cubical.Data.Fin.Base public +open import Cubical.Data.Fin.Base public open import Cubical.Data.Fin.Properties public +open import Cubical.Data.Fin.Literals public diff --git a/Cubical/Data/Fin/Base.agda b/Cubical/Data/Fin/Base.agda index cba6530c1..c3ad6c533 100644 --- a/Cubical/Data/Fin/Base.agda +++ b/Cubical/Data/Fin/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.Fin.Base where @@ -6,10 +6,12 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels -open import Cubical.Data.Empty -open import Cubical.Data.Nat +import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Nat using (ℕ; zero; suc) open import Cubical.Data.Nat.Order -open import Cubical.Data.Sum +open import Cubical.Data.Nat.Order.Recursive using () renaming (_≤_ to _≤′_) +open import Cubical.Data.Sigma +open import Cubical.Data.Sum using (_⊎_; _⊎?_; inl; inr) open import Cubical.Relation.Nullary @@ -42,7 +44,13 @@ toℕ = fst -- ... and injective. toℕ-injective : ∀{fj fk : Fin k} → toℕ fj ≡ toℕ fk → fj ≡ fk -toℕ-injective {fj = fj} {fk} = ΣProp≡ (λ _ → m≤n-isProp) +toℕ-injective {fj = fj} {fk} = Σ≡Prop (λ _ → m≤n-isProp) + +-- Conversion from ℕ with a recursive definition of ≤ + +fromℕ≤ : (m n : ℕ) → m ≤′ n → Fin (suc n) +fromℕ≤ zero _ _ = fzero +fromℕ≤ (suc m) (suc n) m≤n = fsuc (fromℕ≤ m n m≤n) -- A case analysis helper for induction. fsplit @@ -51,20 +59,39 @@ fsplit fsplit (0 , km = Empty.rec (Fin-inj′ n>m p) + +≤-·sk-cancel : ∀ {m} {k} {n} → m · suc k ≤ n · suc k → m ≤ n +≤-·sk-cancel {m} {k} {n} (d , p) = o , inj-·sm {m = k} goal where + r = d % suc k + o = d / suc k + resn·k : Residue (suc k) (n · suc k) + resn·k = ((r , n%sk_ asDiag {x , x'} {y , y'} (inr f) | yes p | yes _ = subst _ p (×Hom₂ x f ) _<$>_ asDiag {x , x'} {y , y'} f | yes p | no _ = subst _ p (×Hom₂ x (lower f) ) _<$>_ asDiag {x , x'} {y , y'} f | no _ | yes p' = subst _ p' (×Hom₁ (lower f) x') - diff --git a/Cubical/Data/Group.agda b/Cubical/Data/Group.agda deleted file mode 100644 index 97c1b5d00..000000000 --- a/Cubical/Data/Group.agda +++ /dev/null @@ -1,4 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.Data.Group where - -open import Cubical.Data.Group.Base public diff --git a/Cubical/Data/Group/Base.agda b/Cubical/Data/Group/Base.agda deleted file mode 100644 index efc370e83..000000000 --- a/Cubical/Data/Group/Base.agda +++ /dev/null @@ -1,162 +0,0 @@ -{-# OPTIONS --cubical --safe #-} - -module Cubical.Data.Group.Base where - -open import Cubical.Foundations.Prelude hiding ( comp ) - -import Cubical.Foundations.Isomorphism as I -import Cubical.Foundations.Equiv as E -import Cubical.Foundations.HAEquiv as HAE - -record isGroup {ℓ} (A : Type ℓ) : Type ℓ where - constructor group-struct - field - id : A - inv : A → A - comp : A → A → A - lUnit : ∀ a → comp id a ≡ a - rUnit : ∀ a → comp a id ≡ a - assoc : ∀ a b c → comp (comp a b) c ≡ comp a (comp b c) - lCancel : ∀ a → comp (inv a) a ≡ id - rCancel : ∀ a → comp a (inv a) ≡ id - -record Group ℓ : Type (ℓ-suc ℓ) where - constructor group - field - type : Type ℓ - setStruc : isSet type - groupStruc : isGroup type - -isMorph : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (f : (Group.type G → Group.type H)) → Type (ℓ-max ℓ ℓ') -isMorph (group G Gset (group-struct _ _ _⊙_ _ _ _ _ _)) - (group H Hset (group-struct _ _ _∘_ _ _ _ _ _)) f - = (g0 g1 : G) → f (g0 ⊙ g1) ≡ (f g0) ∘ (f g1) - -morph : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → Type (ℓ-max ℓ ℓ') -morph G H = Σ (Group.type G → Group.type H) (isMorph G H) - -record Iso {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ') where - constructor iso - field - fun : morph G H - inv : morph H G - rightInv : I.section (fun .fst) (inv .fst) - leftInv : I.retract (fun .fst) (inv .fst) - -record Iso' {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ') where - constructor iso' - field - isoSet : I.Iso (Group.type G) (Group.type H) - isoSetMorph : isMorph G H (I.Iso.fun isoSet) - -_≃_ : ∀ {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') → Type (ℓ-max ℓ ℓ') -A ≃ B = Σ (morph A B) \ f → (E.isEquiv (f .fst)) - -Iso'→Iso : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → Iso' G H → Iso G H -Iso'→Iso {G = group G Gset Ggroup} {H = group H Hset Hgroup} i = iso (fun , funMorph) (inv , invMorph) rightInv leftInv - where - G_ : Group _ - G_ = (group G Gset Ggroup) - - H_ : Group _ - H_ = (group H Hset Hgroup) - - open Iso' - - fun : G → H - fun = I.Iso.fun (isoSet i) - - inv : H → G - inv = I.Iso.inv (isoSet i) - - rightInv : I.section fun inv - rightInv = I.Iso.rightInv (isoSet i) - - leftInv : I.retract fun inv - leftInv = I.Iso.leftInv (isoSet i) - - e' : G E.≃ H - e' = I.isoToEquiv (I.iso fun inv rightInv leftInv) - - funMorph : isMorph G_ H_ fun - funMorph = isoSetMorph i - - _∘_ : H → H → H - _∘_ = isGroup.comp Hgroup - - _⊙_ : G → G → G - _⊙_ = isGroup.comp Ggroup - - invMorph : isMorph H_ G_ inv - invMorph h0 h1 = E.invEq (HAE.congEquiv e') - (fun (inv (h0 ∘ h1)) ≡⟨ rightInv (h0 ∘ h1) ⟩ - h0 ∘ h1 ≡⟨ cong (λ x → x ∘ h1) (sym (rightInv h0)) ⟩ - (fun (inv h0)) ∘ h1 ≡⟨ cong (λ x → fun (inv h0) ∘ x) (sym (rightInv h1)) ⟩ - (fun (inv h0)) ∘ (fun (inv h1)) ≡⟨ sym (funMorph (inv h0) (inv h1)) ⟩ - fun ((inv h0) ⊙ (inv h1)) ∎ ) - - -Equiv→Iso' : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → G ≃ H → Iso' G H -Equiv→Iso' {G = group G Gset Ggroup} - {H = group H Hset Hgroup} - e = iso' i' (e .fst .snd) - where - e' : G E.≃ H - e' = (e .fst .fst) , (e .snd) - - i' : I.Iso G H - i' = E.equivToIso e' - -compMorph : ∀ {ℓ} {F G H : Group ℓ} (I : morph F G) (J : morph G H) → morph F H -compMorph {ℓ} {group F Fset (group-struct _ _ _⊙_ _ _ _ _ _)} - {group G Gset (group-struct _ _ _∙_ _ _ _ _ _)} - {group H Hset (group-struct _ _ _∘_ _ _ _ _ _)} - (i , ic) (j , jc) = k , kc - where - k : F → H - k f = j (i f) - - kc : (g0 g1 : F) → k (g0 ⊙ g1) ≡ (k g0) ∘ (k g1) - kc g0 g1 = j (i (g0 ⊙ g1)) ≡⟨ cong j (ic _ _) ⟩ - j (i g0 ∙ i g1) ≡⟨ jc _ _ ⟩ - j (i g0) ∘ j (i g1) ∎ - -compIso : ∀ {ℓ} {F G H : Group ℓ} (I : Iso F G) (J : Iso G H) → Iso F H -compIso {ℓ} {F} {G} {H} - (iso F→G G→F fg gf) (iso G→H H→G gh hg ) = iso F→H H→F sec ret - where - F→H : morph F H - F→H = compMorph {ℓ} {F} {G} {H} F→G G→H - - H→F : morph H F - H→F = compMorph {ℓ} {H} {G} {F} H→G G→F - - open Group - - f→h : F .type → H .type - f→h = F→H .fst - - f→g : F .type → G .type - f→g = F→G .fst - - g→h : G .type → H .type - g→h = G→H .fst - - h→f : H .type → F .type - h→f = H→F .fst - - h→g : H .type → G .type - h→g = H→G .fst - - g→f : G .type → F .type - g→f = G→F .fst - - sec : I.section f→h h→f - sec h = f→h (h→f h) ≡⟨ cong g→h (fg (h→g h)) ⟩ - g→h (h→g h) ≡⟨ gh _ ⟩ - h ∎ - - ret : I.retract f→h h→f - ret f = h→f (f→h f) ≡⟨ cong g→f (hg (f→g f)) ⟩ - g→f (f→g f) ≡⟨ gf _ ⟩ - f ∎ diff --git a/Cubical/Data/HomotopyGroup.agda b/Cubical/Data/HomotopyGroup.agda index b457867c7..04ab4f13b 100644 --- a/Cubical/Data/HomotopyGroup.agda +++ b/Cubical/Data/HomotopyGroup.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.HomotopyGroup where open import Cubical.Data.HomotopyGroup.Base public diff --git a/Cubical/Data/HomotopyGroup/Base.agda b/Cubical/Data/HomotopyGroup/Base.agda index f7f088808..997b31dfe 100644 --- a/Cubical/Data/HomotopyGroup/Base.agda +++ b/Cubical/Data/HomotopyGroup/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.HomotopyGroup.Base where open import Cubical.Foundations.Prelude @@ -7,19 +7,14 @@ import Cubical.Foundations.GroupoidLaws as GL open import Cubical.Foundations.Pointed open import Cubical.Data.Nat -open import Cubical.Data.Group.Base +open import Cubical.Algebra.Group -open import Cubical.HITs.SetTruncation +open import Cubical.Homotopy.Loopspace -Ω : ∀ {ℓ} → Pointed ℓ → Pointed ℓ -Ω (A , a ) = ( (a ≡ a) , refl) - -Ω^_ : ∀ {ℓ} → ℕ → Pointed ℓ → Pointed ℓ -(Ω^ 0) p = p -(Ω^ (suc n)) p = Ω ((Ω^ n) p) +open import Cubical.HITs.SetTruncation as SetTrunc π^_ : ∀ {ℓ} → ℕ → Pointed ℓ → Group ℓ -π^_ {ℓ} n p = group ∥ A ∥₀ squash₀ g +π^_ {ℓ} n p = makeGroup e _⨀_ _⁻¹ isSetSetTrunc assoc rUnit lUnit rCancel lCancel where n' : ℕ n' = suc n @@ -27,34 +22,31 @@ open import Cubical.HITs.SetTruncation A : Type ℓ A = typ ((Ω^ n') p) - g : isGroup ∥ A ∥₀ - g = group-struct e _⁻¹ _⊙_ lUnit rUnit assoc lCancel rCancel - where - e : ∥ A ∥₀ - e = ∣ pt ((Ω^ n') p) ∣₀ + e : ∥ A ∥₂ + e = ∣ pt ((Ω^ n') p) ∣₂ - _⁻¹ : ∥ A ∥₀ → ∥ A ∥₀ - _⁻¹ = elimSetTrunc {B = λ _ → ∥ A ∥₀} (λ x → squash₀) λ a → ∣ sym a ∣₀ + _⁻¹ : ∥ A ∥₂ → ∥ A ∥₂ + _⁻¹ = SetTrunc.elim {B = λ _ → ∥ A ∥₂} (λ x → squash₂) λ a → ∣ sym a ∣₂ - _⊙_ : ∥ A ∥₀ → ∥ A ∥₀ → ∥ A ∥₀ - _⊙_ = elimSetTrunc2 (λ _ _ → squash₀) λ a₀ a₁ → ∣ a₀ ∙ a₁ ∣₀ + _⨀_ : ∥ A ∥₂ → ∥ A ∥₂ → ∥ A ∥₂ + _⨀_ = SetTrunc.elim2 (λ _ _ → squash₂) λ a₀ a₁ → ∣ a₀ ∙ a₁ ∣₂ - lUnit : (a : ∥ A ∥₀) → (e ⊙ a) ≡ a - lUnit = elimSetTrunc (λ _ → isProp→isSet (squash₀ _ _)) - (λ a → cong ∣_∣₀ (sym (GL.lUnit a) )) + lUnit : (a : ∥ A ∥₂) → (e ⨀ a) ≡ a + lUnit = SetTrunc.elim (λ _ → isProp→isSet (squash₂ _ _)) + (λ a → cong ∣_∣₂ (sym (GL.lUnit a) )) - rUnit : (a : ∥ A ∥₀) → a ⊙ e ≡ a - rUnit = elimSetTrunc (λ _ → isProp→isSet (squash₀ _ _)) - (λ a → cong ∣_∣₀ (sym (GL.rUnit a) )) + rUnit : (a : ∥ A ∥₂) → a ⨀ e ≡ a + rUnit = SetTrunc.elim (λ _ → isProp→isSet (squash₂ _ _)) + (λ a → cong ∣_∣₂ (sym (GL.rUnit a) )) - assoc : (a b c : ∥ A ∥₀) → ((a ⊙ b) ⊙ c) ≡ (a ⊙ (b ⊙ c)) - assoc = elimSetTrunc3 (λ _ _ _ → isProp→isSet (squash₀ _ _)) - (λ a b c → cong ∣_∣₀ (sym (GL.assoc _ _ _))) + assoc : (a b c : ∥ A ∥₂) → a ⨀ (b ⨀ c) ≡ (a ⨀ b) ⨀ c + assoc = SetTrunc.elim3 (λ _ _ _ → isProp→isSet (squash₂ _ _)) + (λ a b c → cong ∣_∣₂ (GL.assoc _ _ _)) - lCancel : (a : ∥ A ∥₀) → ((a ⁻¹) ⊙ a) ≡ e - lCancel = elimSetTrunc (λ _ → isProp→isSet (squash₀ _ _)) - λ a → cong ∣_∣₀ (GL.lCancel _) + lCancel : (a : ∥ A ∥₂) → ((a ⁻¹) ⨀ a) ≡ e + lCancel = SetTrunc.elim (λ _ → isProp→isSet (squash₂ _ _)) + λ a → cong ∣_∣₂ (GL.lCancel _) - rCancel : (a : ∥ A ∥₀) → (a ⊙ (a ⁻¹)) ≡ e - rCancel = elimSetTrunc (λ _ → isProp→isSet (squash₀ _ _)) - λ a → cong ∣_∣₀ (GL.rCancel _) + rCancel : (a : ∥ A ∥₂) → (a ⨀ (a ⁻¹)) ≡ e + rCancel = SetTrunc.elim (λ _ → isProp→isSet (squash₂ _ _)) + λ a → cong ∣_∣₂ (GL.rCancel _) diff --git a/Cubical/Data/InfNat.agda b/Cubical/Data/InfNat.agda index 61a885ab6..26cf4093d 100644 --- a/Cubical/Data/InfNat.agda +++ b/Cubical/Data/InfNat.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --no-exact-split --safe #-} +{-# OPTIONS --no-exact-split --safe #-} module Cubical.Data.InfNat where open import Cubical.Data.InfNat.Base public diff --git a/Cubical/Data/InfNat/Base.agda b/Cubical/Data/InfNat/Base.agda index 23f78c736..d36895464 100644 --- a/Cubical/Data/InfNat/Base.agda +++ b/Cubical/Data/InfNat/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --no-exact-split --safe #-} +{-# OPTIONS --no-exact-split --safe #-} module Cubical.Data.InfNat.Base where @@ -26,11 +26,11 @@ _+_ : ℕ+∞ → ℕ+∞ → ℕ+∞ fin n + ∞ = ∞ fin n + fin m = fin (n ℕ.+ m) -infixl 7 _*_ -_*_ : ℕ+∞ → ℕ+∞ → ℕ+∞ -fin m * fin n = fin (m ℕ.* n) -∞ * fin ℕ.zero = zero -fin ℕ.zero * ∞ = zero -∞ * ∞ = ∞ -∞ * fin (ℕ.suc _) = ∞ -fin (ℕ.suc _) * ∞ = ∞ +infixl 7 _·_ +_·_ : ℕ+∞ → ℕ+∞ → ℕ+∞ +fin m · fin n = fin (m ℕ.· n) +∞ · fin ℕ.zero = zero +fin ℕ.zero · ∞ = zero +∞ · ∞ = ∞ +∞ · fin (ℕ.suc _) = ∞ +fin (ℕ.suc _) · ∞ = ∞ diff --git a/Cubical/Data/InfNat/Properties.agda b/Cubical/Data/InfNat/Properties.agda index d2d0eb769..653de73ff 100644 --- a/Cubical/Data/InfNat/Properties.agda +++ b/Cubical/Data/InfNat/Properties.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --no-exact-split --safe #-} +{-# OPTIONS --no-exact-split --safe #-} module Cubical.Data.InfNat.Properties where diff --git a/Cubical/Data/Int.agda b/Cubical/Data/Int.agda index 782e4f6c8..57147d579 100644 --- a/Cubical/Data/Int.agda +++ b/Cubical/Data/Int.agda @@ -1,6 +1,7 @@ -{-# OPTIONS --cubical --safe #-} +-- This is the preferred version of the integers in the library. Other +-- versions can be found in the MoreInts directory. +{-# OPTIONS --safe #-} module Cubical.Data.Int where open import Cubical.Data.Int.Base public - open import Cubical.Data.Int.Properties public diff --git a/Cubical/Data/Int/Base.agda b/Cubical/Data/Int/Base.agda index b33008087..b7d9c48b5 100644 --- a/Cubical/Data/Int/Base.agda +++ b/Cubical/Data/Int/Base.agda @@ -1,34 +1,85 @@ -{-# OPTIONS --cubical --safe #-} +-- This is the preferred version of the integers in the library. Other +-- versions can be found in the MoreInts directory. +{-# OPTIONS --safe #-} module Cubical.Data.Int.Base where open import Cubical.Core.Everything +open import Cubical.Data.Bool +open import Cubical.Data.Nat hiding (_+_ ; _·_) renaming (isEven to isEvenℕ ; isOdd to isOddℕ) -open import Cubical.Data.Nat +infix 8 -_ +infixl 7 _·_ +infixl 6 _+_ _-_ -data Int : Type₀ where - pos : (n : ℕ) → Int - negsuc : (n : ℕ) → Int +data ℤ : Type₀ where + pos : (n : ℕ) → ℤ + negsuc : (n : ℕ) → ℤ -neg : (n : ℕ) → Int -neg zero = pos zero +neg : (n : ℕ) → ℤ +neg zero = pos zero neg (suc n) = negsuc n -sucInt : Int → Int -sucInt (pos n) = pos (suc n) -sucInt (negsuc zero) = pos zero -sucInt (negsuc (suc n)) = negsuc n +sucℤ : ℤ → ℤ +sucℤ (pos n) = pos (suc n) +sucℤ (negsuc zero) = pos zero +sucℤ (negsuc (suc n)) = negsuc n -predInt : Int → Int -predInt (pos zero) = negsuc zero -predInt (pos (suc n)) = pos n -predInt (negsuc n) = negsuc (suc n) +predℤ : ℤ → ℤ +predℤ (pos zero) = negsuc zero +predℤ (pos (suc n)) = pos n +predℤ (negsuc n) = negsuc (suc n) --- Natural number and negative integer literals for Int +isEven : ℤ → Bool +isEven (pos n) = isEvenℕ n +isEven (negsuc n) = isOddℕ n + +isOdd : ℤ → Bool +isOdd (pos n) = isOddℕ n +isOdd (negsuc n) = isEvenℕ n + +abs : ℤ → ℕ +abs (pos n) = n +abs (negsuc n) = suc n + +_ℕ-_ : ℕ → ℕ → ℤ +a ℕ- 0 = pos a +0 ℕ- suc b = negsuc b +suc a ℕ- suc b = a ℕ- b + +_+pos_ : ℤ → ℕ → ℤ +z +pos 0 = z +z +pos (suc n) = sucℤ (z +pos n) + +_+negsuc_ : ℤ → ℕ → ℤ +z +negsuc 0 = predℤ z +z +negsuc (suc n) = predℤ (z +negsuc n) + +_+_ : ℤ → ℤ → ℤ +m + pos n = m +pos n +m + negsuc n = m +negsuc n + +-_ : ℤ → ℤ +- pos zero = pos zero +- pos (suc n) = negsuc n +- negsuc n = pos (suc n) + +_-_ : ℤ → ℤ → ℤ +m - n = m + (- n) + +_·_ : ℤ → ℤ → ℤ +pos zero · m = pos zero +pos (suc n) · m = m + pos n · m +negsuc zero · m = - m +negsuc (suc n) · m = - m + negsuc n · m + +-- Natural number and negative integer literals for ℤ + +open import Cubical.Data.Nat.Literals public instance - fromNatInt : HasFromNat Int - fromNatInt = record { Constraint = λ _ → Unit ; fromNat = λ n → pos n } + fromNatℤ : HasFromNat ℤ + fromNatℤ = record { Constraint = λ _ → Unit ; fromNat = λ n → pos n } instance - fromNegInt : HasFromNeg Int - fromNegInt = record { Constraint = λ _ → Unit ; fromNeg = λ n → neg n } + fromNegℤ : HasFromNeg ℤ + fromNegℤ = record { Constraint = λ _ → Unit ; fromNeg = λ n → neg n } diff --git a/Cubical/Data/Int/MoreInts/BiInvInt.agda b/Cubical/Data/Int/MoreInts/BiInvInt.agda new file mode 100644 index 000000000..60adfd54d --- /dev/null +++ b/Cubical/Data/Int/MoreInts/BiInvInt.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.BiInvInt where + +open import Cubical.Data.Int.MoreInts.BiInvInt.Base public + +open import Cubical.Data.Int.MoreInts.BiInvInt.Properties public diff --git a/Cubical/HITs/Ints/BiInvInt/Base.agda b/Cubical/Data/Int/MoreInts/BiInvInt/Base.agda similarity index 76% rename from Cubical/HITs/Ints/BiInvInt/Base.agda rename to Cubical/Data/Int/MoreInts/BiInvInt/Base.agda index ceb230dba..83af6718b 100644 --- a/Cubical/HITs/Ints/BiInvInt/Base.agda +++ b/Cubical/Data/Int/MoreInts/BiInvInt/Base.agda @@ -13,8 +13,8 @@ This file contains: - versions of the point constructors of BiInvInt which satisfy the path constructors judgmentally -} -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Ints.BiInvInt.Base where +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.BiInvInt.Base where open import Cubical.Core.Everything @@ -26,26 +26,26 @@ open import Cubical.Data.Int open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.BiInvEquiv -open import Cubical.Foundations.HAEquiv open import Cubical.Foundations.Equiv open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Equiv.BiInvertible +open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Relation.Nullary -data BiInvInt : Type₀ where - zero : BiInvInt - suc : BiInvInt -> BiInvInt +data BiInvℤ : Type₀ where + zero : BiInvℤ + suc : BiInvℤ -> BiInvℤ -- suc is a bi-invertible equivalence: - predr : BiInvInt -> BiInvInt + predr : BiInvℤ -> BiInvℤ suc-predr : ∀ z -> suc (predr z) ≡ z - predl : BiInvInt -> BiInvInt + predl : BiInvℤ -> BiInvℤ predl-suc : ∀ z -> predl (suc z) ≡ z -suc-biinvequiv : BiInvEquiv BiInvInt BiInvInt +suc-biinvequiv : BiInvEquiv BiInvℤ BiInvℤ suc-biinvequiv = record { fun = suc ; invr = predr ; invr-rightInv = suc-predr ; invl = predl ; invl-leftInv = predl-suc } @@ -55,11 +55,11 @@ predr-suc = BiInvEquiv.invr-leftInv suc-biinvequiv -- since we want to use suc-adj and pred-adj (defined below) later on, we will need the -- definition of suc-predl taken from HAEquiv instead of from BiInvEquiv -suc-haequiv : HAEquiv BiInvInt BiInvInt +suc-haequiv : HAEquiv BiInvℤ BiInvℤ suc-haequiv = biInvEquiv→HAEquiv suc-biinvequiv suc-predl : ∀ z -> suc (predl z) ≡ z -suc-predl = isHAEquiv.ret (snd suc-haequiv) +suc-predl = isHAEquiv.rinv (snd suc-haequiv) -- predr and predl (as well as suc-predr and suc-predl - up to a path) are indistinguishable, -- so we can safely define 'pred' to just be predl @@ -72,7 +72,7 @@ suc-predl≡predr : ∀ z -> PathP (λ j → suc (predl≡predr z j) ≡ z) (suc suc-predl≡predr z i = cong snd (isContr→isProp (isContr-hasSection (biInvEquiv→Equiv-left suc-biinvequiv)) (predl , suc-predl) (predr , suc-predr)) i z -pred : BiInvInt -> BiInvInt +pred : BiInvℤ -> BiInvℤ pred = predl suc-pred = suc-predl @@ -86,38 +86,38 @@ pred-adj z = isHAEquiv.com-op (snd suc-haequiv) z --- Int ≡ BiInvInt +-- ℤ ≡ BiInvℤ -fwd : Int -> BiInvInt +fwd : ℤ -> BiInvℤ fwd (pos zero) = zero fwd (pos (suc n)) = suc (fwd (pos n)) fwd (negsuc zero) = pred zero fwd (negsuc (suc n)) = pred (fwd (negsuc n)) -bwd : BiInvInt -> Int +bwd : BiInvℤ -> ℤ bwd zero = pos zero -bwd (suc z) = sucInt (bwd z) -bwd (predr z) = predInt (bwd z) +bwd (suc z) = sucℤ (bwd z) +bwd (predr z) = predℤ (bwd z) bwd (suc-predr z i) = sucPred (bwd z) i -bwd (predl z) = predInt (bwd z) +bwd (predl z) = predℤ (bwd z) bwd (predl-suc z i) = predSuc (bwd z) i -bwd-fwd : ∀ (x : Int) -> bwd (fwd x) ≡ x +bwd-fwd : ∀ (x : ℤ) -> bwd (fwd x) ≡ x bwd-fwd (pos zero) = refl -bwd-fwd (pos (suc n)) = cong sucInt (bwd-fwd (pos n)) +bwd-fwd (pos (suc n)) = cong sucℤ (bwd-fwd (pos n)) bwd-fwd (negsuc zero) = refl -bwd-fwd (negsuc (suc n)) = cong predInt (bwd-fwd (negsuc n)) +bwd-fwd (negsuc (suc n)) = cong predℤ (bwd-fwd (negsuc n)) -fwd-sucInt : ∀ (x : Int) → fwd (sucInt x) ≡ suc (fwd x) -fwd-sucInt (pos n) = refl -fwd-sucInt (negsuc zero) = sym (suc-pred (fwd (pos zero))) -fwd-sucInt (negsuc (suc n)) = sym (suc-pred (fwd (negsuc n))) +fwd-sucℤ : ∀ (x : ℤ) → fwd (sucℤ x) ≡ suc (fwd x) +fwd-sucℤ (pos n) = refl +fwd-sucℤ (negsuc zero) = sym (suc-pred (fwd (pos zero))) +fwd-sucℤ (negsuc (suc n)) = sym (suc-pred (fwd (negsuc n))) -fwd-predInt : ∀ (x : Int) → fwd (predInt x) ≡ pred (fwd x) -fwd-predInt (pos zero) = refl -fwd-predInt (pos (suc n)) = sym (predl-suc (fwd (pos n))) -fwd-predInt (negsuc n) = refl +fwd-predℤ : ∀ (x : ℤ) → fwd (predℤ x) ≡ pred (fwd x) +fwd-predℤ (pos zero) = refl +fwd-predℤ (pos (suc n)) = sym (predl-suc (fwd (pos n))) +fwd-predℤ (negsuc n) = refl private sym-filler : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) @@ -130,8 +130,8 @@ private ; (i = i1) → y ; (j = i1) → p (i ∨ (~ k)) }) y -fwd-sucPred : ∀ (x : Int) - → Square {- (i = i0) -} (fwd-sucInt (predInt x) ∙ (λ i → suc (fwd-predInt x i))) +fwd-sucPred : ∀ (x : ℤ) + → Square {- (i = i0) -} (fwd-sucℤ (predℤ x) ∙ (λ i → suc (fwd-predℤ x i))) {- (i = i1) -} (λ _ → fwd x) {- (j = i0) -} (λ i → fwd (sucPred x i)) {- (j = i1) -} (suc-pred (fwd x)) @@ -140,7 +140,7 @@ fwd-sucPred : ∀ (x : Int) fwd-sucPred (pos zero) i j = hcomp (λ k → λ { (j = i0) → fwd (pos zero) ; (i = i0) → rUnit (sym (suc-pred (fwd (pos zero)))) k j - -- because fwd-sucInt (predInt (pos zero)) ≡ sym (suc-pred (fwd (pos zero))) + -- because fwd-sucℤ (predℤ (pos zero)) ≡ sym (suc-pred (fwd (pos zero))) ; (i = i1) → fwd (pos zero) ; (j = i1) → suc-pred (fwd (pos zero)) i }) @@ -149,7 +149,7 @@ fwd-sucPred (pos zero) i j fwd-sucPred (pos (suc n)) i j = hcomp (λ k → λ { (j = i0) → suc (fwd (pos n)) ; (i = i0) → lUnit (λ i → suc (sym (predl-suc (fwd (pos n))) i)) k j - -- because fwd-predInt (pos (suc n)) ≡ sym (predl-suc (fwd (pos n))) + -- because fwd-predℤ (pos (suc n)) ≡ sym (predl-suc (fwd (pos n))) ; (i = i1) → suc (fwd (pos n)) ; (j = i1) → suc-adj (fwd (pos n)) k i }) @@ -158,15 +158,15 @@ fwd-sucPred (pos (suc n)) i j fwd-sucPred (negsuc n) i j = hcomp (λ k → λ { (j = i0) → fwd (negsuc n) ; (i = i0) → rUnit (sym (suc-pred (fwd (negsuc n)))) k j - -- because fwd-sucInt (predInt (negsuc n)) ≡ sym (suc-pred (fwd (negsuc n))) + -- because fwd-sucℤ (predℤ (negsuc n)) ≡ sym (suc-pred (fwd (negsuc n))) ; (i = i1) → fwd (negsuc n) ; (j = i1) → suc-pred (fwd (negsuc n)) i }) (sym-filler (suc-pred (fwd (negsuc n))) i j) -fwd-predSuc : ∀ (x : Int) - → Square {- (i = i0) -} (fwd-predInt (sucInt x) ∙ (λ i → pred (fwd-sucInt x i))) +fwd-predSuc : ∀ (x : ℤ) + → Square {- (i = i0) -} (fwd-predℤ (sucℤ x) ∙ (λ i → pred (fwd-sucℤ x i))) {- (i = i1) -} (λ _ → fwd x) {- (j = i0) -} (λ i → fwd (predSuc x i)) {- (j = i1) -} (pred-suc (fwd x)) @@ -174,7 +174,7 @@ fwd-predSuc : ∀ (x : Int) fwd-predSuc (pos n) i j = hcomp (λ k → λ { (j = i0) → fwd (pos n) ; (i = i0) → rUnit (sym (pred-suc (fwd (pos n)))) k j - -- because fwd-predInt (sucInt (pos n)) ≡ sym (pred-suc (fwd (pos n))) + -- because fwd-predℤ (sucℤ (pos n)) ≡ sym (pred-suc (fwd (pos n))) ; (i = i1) → fwd (pos n) ; (j = i1) → pred-suc (fwd (pos n)) i }) @@ -183,7 +183,7 @@ fwd-predSuc (pos n) i j fwd-predSuc (negsuc zero) i j = hcomp (λ k → λ { (j = i0) → fwd (negsuc zero) ; (i = i0) → lUnit (λ i → pred (sym (suc-pred (fwd (pos zero))) i)) k j - -- because fwd-sucInt (negsuc zero) ≡ sym (suc-pred (fwd (pos zero))) + -- because fwd-sucℤ (negsuc zero) ≡ sym (suc-pred (fwd (pos zero))) ; (i = i1) → fwd (negsuc zero) ; (j = i1) → pred-adj (fwd (pos zero)) k i }) @@ -192,22 +192,22 @@ fwd-predSuc (negsuc zero) i j fwd-predSuc (negsuc (suc n)) i j = hcomp (λ k → λ { (j = i0) → fwd (negsuc (suc n)) ; (i = i0) → lUnit (λ i → pred (sym (suc-pred (fwd (negsuc n))) i)) k j - -- because fwd-sucInt (negsuc (suc n)) ≡ sym (suc-pred (fwd (negsuc n))) + -- because fwd-sucℤ (negsuc (suc n)) ≡ sym (suc-pred (fwd (negsuc n))) ; (i = i1) → fwd (negsuc (suc n)) ; (j = i1) → pred-adj (fwd (negsuc n)) k i }) (pred (sym-filler (suc-pred (fwd (negsuc n))) i j)) -fwd-bwd : ∀ (z : BiInvInt) -> fwd (bwd z) ≡ z +fwd-bwd : ∀ (z : BiInvℤ) -> fwd (bwd z) ≡ z fwd-bwd zero = refl -fwd-bwd (suc z) = fwd-sucInt (bwd z) ∙ (λ i → suc (fwd-bwd z i)) -fwd-bwd (predr z) = fwd-predInt (bwd z) ∙ (λ i → predl≡predr (fwd-bwd z i) i) -fwd-bwd (predl z) = fwd-predInt (bwd z) ∙ (λ i → pred (fwd-bwd z i)) +fwd-bwd (suc z) = fwd-sucℤ (bwd z) ∙ (λ i → suc (fwd-bwd z i)) +fwd-bwd (predr z) = fwd-predℤ (bwd z) ∙ (λ i → predl≡predr (fwd-bwd z i) i) +fwd-bwd (predl z) = fwd-predℤ (bwd z) ∙ (λ i → pred (fwd-bwd z i)) fwd-bwd (suc-predr z i) j = hcomp (λ k → λ { (j = i0) → fwd (sucPred (bwd z) i) - ; (i = i0) → (fwd-sucInt (predInt (bwd z)) - ∙ (λ i → suc (compPath-filler (fwd-predInt (bwd z)) + ; (i = i0) → (fwd-sucℤ (predℤ (bwd z)) + ∙ (λ i → suc (compPath-filler (fwd-predℤ (bwd z)) (λ i' → predl≡predr (fwd-bwd z i') i') k i))) j ; (i = i1) → fwd-bwd z (j ∧ k) @@ -215,8 +215,8 @@ fwd-bwd (suc-predr z i) j (fwd-sucPred (bwd z) i j) fwd-bwd (predl-suc z i) j = hcomp (λ k → λ { (j = i0) → fwd (predSuc (bwd z) i) - ; (i = i0) → (fwd-predInt (sucInt (bwd z)) - ∙ (λ i → pred (compPath-filler (fwd-sucInt (bwd z)) + ; (i = i0) → (fwd-predℤ (sucℤ (bwd z)) + ∙ (λ i → pred (compPath-filler (fwd-sucℤ (bwd z)) (λ i' → suc (fwd-bwd z i')) k i))) j ; (i = i1) → fwd-bwd z (j ∧ k) @@ -224,33 +224,33 @@ fwd-bwd (predl-suc z i) j (fwd-predSuc (bwd z) i j) -Int≡BiInvInt : Int ≡ BiInvInt -Int≡BiInvInt = isoToPath (iso fwd bwd fwd-bwd bwd-fwd) +ℤ≡BiInvℤ : ℤ ≡ BiInvℤ +ℤ≡BiInvℤ = isoToPath (iso fwd bwd fwd-bwd bwd-fwd) -discreteBiInvInt : Discrete BiInvInt -discreteBiInvInt = subst Discrete Int≡BiInvInt discreteInt +discreteBiInvℤ : Discrete BiInvℤ +discreteBiInvℤ = subst Discrete ℤ≡BiInvℤ discreteℤ -isSetBiInvInt : isSet BiInvInt -isSetBiInvInt = subst isSet Int≡BiInvInt isSetInt +isSetBiInvℤ : isSet BiInvℤ +isSetBiInvℤ = subst isSet ℤ≡BiInvℤ isSetℤ -- suc' is equal to suc (suc≡suc') but cancels pred *exactly* (see suc'-pred) -suc'ᵖ : (z : BiInvInt) -> Σ BiInvInt (suc z ≡_) +suc'ᵖ : (z : BiInvℤ) -> Σ BiInvℤ (suc z ≡_) suc' = fst ∘ suc'ᵖ suc≡suc' = snd ∘ suc'ᵖ suc'ᵖ zero = suc zero , refl suc'ᵖ (suc z) = suc (suc z) , refl suc'ᵖ (predr z) = z , suc-predr z -suc'ᵖ (suc-predr z i) = let filler : I → I → BiInvInt +suc'ᵖ (suc-predr z i) = let filler : I → I → BiInvℤ filler j i = hfill (λ j → λ { (i = i0) → suc (suc (predr z)) ; (i = i1) → suc≡suc' z j }) (inS (suc (suc-predr z i))) j in filler i1 i , λ j → filler j i suc'ᵖ (predl z) = z , suc-predl z -suc'ᵖ (predl-suc z i) = let filler : I → I → BiInvInt +suc'ᵖ (predl-suc z i) = let filler : I → I → BiInvℤ filler j i = hfill (λ j → λ { (i = i0) → suc-predl (suc z) j ; (i = i1) → suc≡suc' z j }) (inS (suc (predl-suc z i))) j @@ -263,39 +263,39 @@ private -- pred' is equal to pred (pred≡pred') but cancels suc *exactly* (see pred'-suc) -predr'ᵖ : (z : BiInvInt) -> Σ BiInvInt (predr z ≡_) +predr'ᵖ : (z : BiInvℤ) -> Σ BiInvℤ (predr z ≡_) predr' = fst ∘ predr'ᵖ predr≡predr' = snd ∘ predr'ᵖ predr'ᵖ zero = predr zero , refl predr'ᵖ (suc z) = z , predr-suc z predr'ᵖ (predr z) = predr (predr z) , refl -predr'ᵖ (suc-predr z i) = let filler : I → I → BiInvInt +predr'ᵖ (suc-predr z i) = let filler : I → I → BiInvℤ filler j i = hfill (λ j → λ { (i = i0) → predr-suc (predr z) j ; (i = i1) → predr≡predr' z j }) (inS (predr (suc-predr z i))) j in filler i1 i , λ j → filler j i predr'ᵖ (predl z) = predr (predl z) , refl -predr'ᵖ (predl-suc z i) = let filler : I → I → BiInvInt +predr'ᵖ (predl-suc z i) = let filler : I → I → BiInvℤ filler j i = hfill (λ j → λ { (i = i0) → predr (predl (suc z)) ; (i = i1) → predr≡predr' z j }) (inS (predr (predl-suc z i))) j in filler i1 i , λ j → filler j i -predl'ᵖ : (z : BiInvInt) -> Σ BiInvInt (predl z ≡_) +predl'ᵖ : (z : BiInvℤ) -> Σ BiInvℤ (predl z ≡_) predl' = fst ∘ predl'ᵖ predl≡predl' = snd ∘ predl'ᵖ predl'ᵖ zero = predl zero , refl predl'ᵖ (suc z) = z , predl-suc z predl'ᵖ (predr z) = predl (predr z) , refl -predl'ᵖ (suc-predr z i) = let filler : I → I → BiInvInt +predl'ᵖ (suc-predr z i) = let filler : I → I → BiInvℤ filler j i = hfill (λ j → λ { (i = i0) → predl-suc (predr z) j ; (i = i1) → predl≡predl' z j }) (inS (predl (suc-predr z i))) j in filler i1 i , λ j → filler j i predl'ᵖ (predl z) = predl (predl z) , refl -predl'ᵖ (predl-suc z i) = let filler : I → I → BiInvInt +predl'ᵖ (predl-suc z i) = let filler : I → I → BiInvℤ filler j i = hfill (λ j → λ { (i = i0) → predl (predl (suc z)) ; (i = i1) → predl≡predl' z j }) (inS (predl (predl-suc z i))) j diff --git a/Cubical/Data/Int/MoreInts/BiInvInt/Properties.agda b/Cubical/Data/Int/MoreInts/BiInvInt/Properties.agda new file mode 100644 index 000000000..c108cc8ab --- /dev/null +++ b/Cubical/Data/Int/MoreInts/BiInvInt/Properties.agda @@ -0,0 +1,243 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.BiInvInt.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Nat using (ℕ) +import Cubical.Data.Int as ℤ +open import Cubical.Data.Bool + +open import Cubical.Data.Int.MoreInts.BiInvInt.Base + +infixl 6 _+_ _-_ +infixl 7 _·_ + +-- To prove a property P about BiInvℤ, we need to show: +-- * P zero +-- * If P n, then P (suc n) +-- * If P n, then P (pred n) +BiInvℤ-ind-prop : + ∀ {ℓ} {P : BiInvℤ → Type ℓ} → (∀ n → isProp (P n)) → + P zero → (∀ n → P n → P (suc n)) → (∀ n → P n → P (pred n)) → + (n : BiInvℤ) → P n +BiInvℤ-ind-prop {P = P} P-isProp P-zero P-suc P-pred = φ + where + P-predr : ∀ n → P n → P (predr n) + P-predr n x = subst P (predl≡predr n) (P-pred n x) + + P-predl : ∀ n → P n → P (predl n) + P-predl = P-pred + + P-isProp' : {a b : BiInvℤ} (p : a ≡ b) (x : P a) (y : P b) → PathP (λ i → P (p i)) x y + P-isProp' _ _ _ = toPathP (P-isProp _ _ _) + + φ : (n : BiInvℤ) → P n + φ zero = P-zero + φ (suc n) = P-suc n (φ n) + φ (predr n) = P-predr n (φ n) + φ (suc-predr n i) = P-isProp' (suc-predr n) (P-suc (predr n) (P-predr n (φ n))) (φ n) i + φ (predl n) = P-predl n (φ n) + φ (predl-suc n i) = P-isProp' (predl-suc n) (P-predl (suc n) (P-suc n (φ n))) (φ n) i + +-- To define a function BiInvℤ → A, we need: +-- · a point z : A for zero +-- · an equivalence s : A ≃ A for suc/pred +BiInvℤ-rec : ∀ {ℓ} {A : Type ℓ} → A → A ≃ A → BiInvℤ → A +BiInvℤ-rec {A = A} z e = φ + where + e-Iso : Iso A A + e-Iso = equivToIso e + + φ : BiInvℤ → A + φ zero = z + φ (suc n) = Iso.fun e-Iso (φ n) + φ (predr n) = Iso.inv e-Iso (φ n) + φ (suc-predr n i) = Iso.rightInv e-Iso (φ n) i + φ (predl n) = Iso.inv e-Iso (φ n) + φ (predl-suc n i) = Iso.leftInv e-Iso (φ n) i + +sucIso : Iso BiInvℤ BiInvℤ +Iso.fun sucIso = suc +Iso.inv sucIso = pred +Iso.rightInv sucIso = suc-predl +Iso.leftInv sucIso = predl-suc + +sucEquiv : BiInvℤ ≃ BiInvℤ +sucEquiv = isoToEquiv sucIso + +-- addition + +-- the following equalities hold definitionally: +-- zero + n ≡ n +-- suc m + n ≡ suc (m + n) +-- pred m + n ≡ pred (m + n) +_+_ : BiInvℤ → BiInvℤ → BiInvℤ +_+_ = BiInvℤ-rec (idfun BiInvℤ) (postCompEquiv sucEquiv) + +-- properties of addition + ++-zero : ∀ n → n + zero ≡ n ++-zero = BiInvℤ-ind-prop (λ _ → isSetBiInvℤ _ _) refl (λ n p → cong suc p) (λ n p → cong pred p) + ++-suc : ∀ m n → m + suc n ≡ suc (m + n) ++-suc = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + (λ m → refl) + (λ m p n → cong suc (p n)) + (λ m p n → cong pred (p n) ∙ predl-suc (m + n) ∙ sym (suc-predl (m + n))) + ++-pred : ∀ m n → m + pred n ≡ pred (m + n) ++-pred = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + (λ m → refl) + (λ m p n → cong suc (p n) ∙ suc-predl (m + n) ∙ sym (predl-suc (m + n))) + (λ m p n → cong pred (p n)) + ++-comm : ∀ m n → m + n ≡ n + m ++-comm m n = +-comm' n m + where + +-comm' : ∀ n m → m + n ≡ n + m + +-comm' = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + +-zero + (λ n p m → +-suc m n ∙ cong suc (p m)) + (λ n p m → +-pred m n ∙ cong pred (p m)) + ++-assoc : ∀ m n o → m + (n + o) ≡ (m + n) + o ++-assoc = BiInvℤ-ind-prop (λ _ → isPropΠ2 λ _ _ → isSetBiInvℤ _ _) + (λ n o → refl) + (λ m p n o → cong suc (p n o)) + (λ m p n o → cong pred (p n o)) + +-- negation / subtraction + +-- the following equalities hold definitionally: +-- - zero ≡ zero +-- - (suc m) ≡ pred m +-- - (pred m) ≡ suc m +-_ : BiInvℤ → BiInvℤ +-_ = BiInvℤ-rec zero (invEquiv sucEquiv) + +_-_ : BiInvℤ → BiInvℤ → BiInvℤ +m - n = m + (- n) + ++-invˡ : ∀ n → (- n) + n ≡ zero ++-invˡ = BiInvℤ-ind-prop (λ _ → isSetBiInvℤ _ _) + refl + (λ n p → cong pred (+-suc (- n) n) ∙ predl-suc _ ∙ p) + (λ n p → cong suc (+-pred (- n) n) ∙ suc-predl _ ∙ p) + ++-invʳ : ∀ n → n + (- n) ≡ zero ++-invʳ = BiInvℤ-ind-prop (λ _ → isSetBiInvℤ _ _) + refl + (λ n p → cong suc (+-pred n (- n)) ∙ suc-predl _ ∙ p) + (λ n p → cong pred (+-suc n (- n)) ∙ predl-suc _ ∙ p) + +inv-hom : ∀ m n → - (m + n) ≡ (- m) + (- n) +inv-hom = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + (λ n → refl) + (λ m p n → cong pred (p n)) + (λ m p n → cong suc (p n)) + +-- natural injections from ℕ + +pos : ℕ → BiInvℤ +pos ℕ.zero = zero +pos (ℕ.suc n) = suc (pos n) + +neg : ℕ → BiInvℤ +neg ℕ.zero = zero +neg (ℕ.suc n) = pred (neg n) + +-- absolute value +-- (Note that there doesn't appear to be any way around using +-- bwd here! Any direct proof ends up doing the same work...) + +abs : BiInvℤ → ℕ +abs n = ℤ.abs (bwd n) + +Iso-n+ : (n : BiInvℤ) → Iso BiInvℤ BiInvℤ +Iso.fun (Iso-n+ n) = n +_ +Iso.inv (Iso-n+ n) = - n +_ +Iso.rightInv (Iso-n+ n) m = +-assoc n (- n) m ∙ cong (_+ m) (+-invʳ n) +Iso.leftInv (Iso-n+ n) m = +-assoc (- n) n m ∙ cong (_+ m) (+-invˡ n) + +isEquiv-n+ : ∀ n → isEquiv (n +_) +isEquiv-n+ n = isoToIsEquiv (Iso-n+ n) + +-- multiplication + +-- the following equalities hold definitionally: +-- zero · n ≡ zero +-- suc m · n ≡ n + m · n +-- pred m · n ≡ (- n) + m · n +_·_ : BiInvℤ → BiInvℤ → BiInvℤ +m · n = BiInvℤ-rec zero (n +_ , isEquiv-n+ n) m + +-- properties of multiplication + +·-zero : ∀ n → n · zero ≡ zero +·-zero = BiInvℤ-ind-prop (λ _ → isSetBiInvℤ _ _) refl (λ n p → p) (λ n p → p) + +·-suc : ∀ m n → m · suc n ≡ m + m · n +·-suc = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + (λ n → refl) + (λ m p n → cong suc + (cong (n +_) (p n) ∙ +-assoc n m (m · n) ∙ + cong (_+ m · n) (+-comm n m) ∙ sym (+-assoc m n (m · n)))) + (λ m p n → cong pred + (cong (- n +_) (p n) ∙ +-assoc (- n) m (m · n) ∙ + cong (_+ m · n) (+-comm (- n) m) ∙ sym (+-assoc m (- n) (m · n)))) + +·-pred : ∀ m n → m · pred n ≡ (- m) + m · n +·-pred = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + (λ n → refl) + (λ m p n → cong pred + (cong (n +_) (p n) ∙ +-assoc n (- m) (m · n) ∙ + cong (_+ m · n) (+-comm n (- m)) ∙ sym (+-assoc (- m) n (m · n)))) + (λ m p n → cong suc + (cong (- n +_) (p n) ∙ +-assoc (- n) (- m) (m · n) ∙ + cong (_+ m · n) (+-comm (- n) (- m)) ∙ sym (+-assoc (- m) (- n) (m · n)))) + +·-comm : ∀ m n → m · n ≡ n · m +·-comm = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + (λ n → sym (·-zero n)) + (λ m p n → cong (n +_) (p n) ∙ sym (·-suc n m)) + (λ m p n → cong (- n +_) (p n) ∙ sym (·-pred n m)) + +·-identityˡ : ∀ m → suc zero · m ≡ m +·-identityˡ = +-zero + +·-identityʳ : ∀ m → m · suc zero ≡ m +·-identityʳ m = ·-comm m (suc zero) ∙ ·-identityˡ m + +·-distribʳ : ∀ m n o → (m · o) + (n · o) ≡ (m + n) · o +·-distribʳ = BiInvℤ-ind-prop (λ _ → isPropΠ2 λ _ _ → isSetBiInvℤ _ _) + (λ n o → refl) + (λ m p n o → sym (+-assoc o (m · o) (n · o)) ∙ cong (o +_) (p n o)) + (λ m p n o → sym (+-assoc (- o) (m · o) (n · o)) ∙ cong (- o +_) (p n o)) + +·-distribˡ : ∀ o m n → (o · m) + (o · n) ≡ o · (m + n) +·-distribˡ o m n = + cong (_+ o · n) (·-comm o m) ∙ cong (m · o +_) (·-comm o n) ∙ + ·-distribʳ m n o ∙ ·-comm (m + n) o + +·-inv : ∀ m n → m · (- n) ≡ - (m · n) +·-inv = BiInvℤ-ind-prop (λ _ → isPropΠ λ _ → isSetBiInvℤ _ _) + (λ n → refl) + (λ m p n → cong (- n +_) (p n) ∙ sym (inv-hom n (m · n))) + (λ m p n → cong (- (- n) +_) (p n) ∙ sym (inv-hom (- n) (m · n))) + +inv-· : ∀ m n → (- m) · n ≡ - (m · n) +inv-· m n = ·-comm (- m) n ∙ ·-inv n m ∙ cong (-_) (·-comm n m) + +·-assoc : ∀ m n o → m · (n · o) ≡ (m · n) · o +·-assoc = BiInvℤ-ind-prop (λ _ → isPropΠ2 λ _ _ → isSetBiInvℤ _ _) + (λ n o → refl) + (λ m p n o → + cong (n · o +_) (p n o) ∙ ·-distribʳ n (m · n) o) + (λ m p n o → + cong (- (n · o) +_) (p n o) ∙ cong (_+ m · n · o) (sym (inv-· n o)) ∙ + ·-distribʳ (- n) (m · n) o) diff --git a/Cubical/Data/Int/MoreInts/DeltaInt.agda b/Cubical/Data/Int/MoreInts/DeltaInt.agda new file mode 100644 index 000000000..9c7ad0922 --- /dev/null +++ b/Cubical/Data/Int/MoreInts/DeltaInt.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.DeltaInt where + +open import Cubical.Data.Int.MoreInts.DeltaInt.Base public + +open import Cubical.Data.Int.MoreInts.DeltaInt.Properties public diff --git a/Cubical/HITs/Ints/DeltaInt/Base.agda b/Cubical/Data/Int/MoreInts/DeltaInt/Base.agda similarity index 64% rename from Cubical/HITs/Ints/DeltaInt/Base.agda rename to Cubical/Data/Int/MoreInts/DeltaInt/Base.agda index fc7f56517..96b80d35a 100644 --- a/Cubical/HITs/Ints/DeltaInt/Base.agda +++ b/Cubical/Data/Int/MoreInts/DeltaInt/Base.agda @@ -1,8 +1,9 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} {- -This file defines DiffInt (Cubical.Data.DiffInt) as HIT. +This file defines integers as equivalence classes of pairs of natural numbers +using a direct & untruncated HIT definition (cf. Data.Int.MoreInts.DiffInt) and some basic operations, and the zero value: @@ -13,8 +14,8 @@ zero : {a : ℕ} → DeltaInt and conversion function for ℕ and Int: fromℕ : ℕ → DeltaInt -fromInt : Int → DeltaInt -toInt : DeltaInt → Int +fromℤ : Int → DeltaInt +toℤ : DeltaInt → Int and a generalized version of cancel: @@ -22,11 +23,11 @@ cancelN : ∀ a b n → a ⊖ b ≡ (n + a) ⊖ n + b -} -module Cubical.HITs.Ints.DeltaInt.Base where +module Cubical.Data.Int.MoreInts.DeltaInt.Base where open import Cubical.Foundations.Prelude open import Cubical.Data.Nat hiding (zero) -open import Cubical.Data.Int hiding (abs; sgn; _+_) +open import Cubical.Data.Int hiding (abs; _+_) renaming (ℤ to Int) infixl 5 _⊖_ data DeltaInt : Type₀ where @@ -47,13 +48,13 @@ zero {a} = a ⊖ a fromℕ : ℕ → DeltaInt fromℕ n = n ⊖ 0 -fromInt : Int → DeltaInt -fromInt (pos n) = fromℕ n -fromInt (negsuc n) = 0 ⊖ suc n +fromℤ : Int → DeltaInt +fromℤ (pos n) = fromℕ n +fromℤ (negsuc n) = 0 ⊖ suc n -toInt : DeltaInt → Int -toInt (x ⊖ y) = x ℕ- y -toInt (cancel a b i) = a ℕ- b +toℤ : DeltaInt → Int +toℤ (x ⊖ y) = x ℕ- y +toℤ (cancel a b i) = a ℕ- b cancelN : ∀ a b n → a ⊖ b ≡ (n + a) ⊖ n + b cancelN a b 0 = refl diff --git a/Cubical/HITs/Ints/DeltaInt/Properties.agda b/Cubical/Data/Int/MoreInts/DeltaInt/Properties.agda similarity index 81% rename from Cubical/HITs/Ints/DeltaInt/Properties.agda rename to Cubical/Data/Int/MoreInts/DeltaInt/Properties.agda index 64ac2c543..f3f0119f5 100644 --- a/Cubical/HITs/Ints/DeltaInt/Properties.agda +++ b/Cubical/Data/Int/MoreInts/DeltaInt/Properties.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} {- @@ -28,15 +28,15 @@ isSetDeltaInt : isSet DeltaInt -} -module Cubical.HITs.Ints.DeltaInt.Properties where +module Cubical.Data.Int.MoreInts.DeltaInt.Properties where open import Cubical.Foundations.Everything open import Cubical.Data.Nat hiding (zero) -open import Cubical.Data.Int hiding (abs; sgn; _+_) -open import Cubical.HITs.Ints.DeltaInt.Base +open import Cubical.Data.Int hiding (abs; _+_) +open import Cubical.Data.Int.MoreInts.DeltaInt.Base open import Cubical.Relation.Nullary using (Discrete) -deltaIntSec : ∀ b → toInt (fromInt b) ≡ b +deltaIntSec : ∀ b → toℤ (fromℤ b) ≡ b deltaIntSec (pos n) = refl deltaIntSec (negsuc n) = refl @@ -55,7 +55,7 @@ cancelTriangle a b i j = hcomp (λ k → λ ; (j = i1) → cancel a b (i ∧ k) }) (a ⊖ b) -deltaIntRet : ∀ a → fromInt (toInt a) ≡ a +deltaIntRet : ∀ a → fromℤ (toℤ a) ≡ a deltaIntRet (x ⊖ 0) = refl deltaIntRet (0 ⊖ suc y) = refl deltaIntRet (suc x ⊖ suc y) = deltaIntRet (x ⊖ y) ∙ cancel x y @@ -63,14 +63,14 @@ deltaIntRet (cancel a 0 i) = cancelTriangle a 0 i deltaIntRet (cancel 0 (suc b) i) = cancelTriangle 0 (suc b) i deltaIntRet (cancel (suc a) (suc b) i) = deltaIntRet (cancel a b i) ∙ cancelDiamond a b i -DeltaInt≡Int : DeltaInt ≡ Int -DeltaInt≡Int = isoToPath (iso toInt fromInt deltaIntSec deltaIntRet) +DeltaInt≡ℤ : DeltaInt ≡ ℤ +DeltaInt≡ℤ = isoToPath (iso toℤ fromℤ deltaIntSec deltaIntRet) discreteDeltaInt : Discrete DeltaInt -discreteDeltaInt = subst Discrete (sym DeltaInt≡Int) discreteInt +discreteDeltaInt = subst Discrete (sym DeltaInt≡ℤ) discreteℤ isSetDeltaInt : isSet DeltaInt -isSetDeltaInt = subst isSet (sym DeltaInt≡Int) isSetInt +isSetDeltaInt = subst isSet (sym DeltaInt≡ℤ) isSetℤ succPred : ∀ n → succ (pred n) ≡ n succPred (x ⊖ y) i = cancel x y (~ i) diff --git a/Cubical/Data/Int/MoreInts/DiffInt.agda b/Cubical/Data/Int/MoreInts/DiffInt.agda new file mode 100644 index 000000000..66354a320 --- /dev/null +++ b/Cubical/Data/Int/MoreInts/DiffInt.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.DiffInt where + +open import Cubical.Data.Int.MoreInts.DiffInt.Base public +open import Cubical.Data.Int.MoreInts.DiffInt.Properties public diff --git a/Cubical/Data/Int/MoreInts/DiffInt/Base.agda b/Cubical/Data/Int/MoreInts/DiffInt/Base.agda new file mode 100644 index 000000000..28aed1e2f --- /dev/null +++ b/Cubical/Data/Int/MoreInts/DiffInt/Base.agda @@ -0,0 +1,84 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.DiffInt.Base where + +open import Cubical.Foundations.Prelude + +open import Cubical.HITs.SetQuotients +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat hiding (+-comm ; +-assoc) renaming (_+_ to _+ℕ_) +open import Cubical.Data.Int renaming (ℤ to Int) + +rel : (ℕ × ℕ) → (ℕ × ℕ) → Type₀ +rel (a₀ , b₀) (a₁ , b₁) = x ≡ y + where + x = a₀ +ℕ b₁ + y = a₁ +ℕ b₀ + +ℤ = (ℕ × ℕ) / rel + + +-- Proof of equivalence between Int and DiffInt + +private + -- Prove all the identities for ℕ×ℕ first and use that to build to ℤ + + Int→ℕ×ℕ : Int → (ℕ × ℕ) + Int→ℕ×ℕ (pos n) = ( n , zero ) + Int→ℕ×ℕ (negsuc n) = ( zero , suc n ) + + ℕ×ℕ→Int : (ℕ × ℕ) → Int + ℕ×ℕ→Int(n , m) = n ℕ- m + + relInt→ℕ×ℕ : ∀ m n → rel (Int→ℕ×ℕ (m ℕ- n)) (m , n) + relInt→ℕ×ℕ zero zero = refl + relInt→ℕ×ℕ zero (suc n) = refl + relInt→ℕ×ℕ (suc m) zero = refl + relInt→ℕ×ℕ (suc m) (suc n) = + fst (Int→ℕ×ℕ (m ℕ- n)) +ℕ suc n ≡⟨ +-suc (fst (Int→ℕ×ℕ (m ℕ- n))) n ⟩ + suc (fst (Int→ℕ×ℕ (m ℕ- n)) +ℕ n) ≡⟨ cong suc (relInt→ℕ×ℕ m n) ⟩ + suc (m +ℕ snd (Int→ℕ×ℕ (m ℕ- n))) ∎ + + Int→ℕ×ℕ→Int : ∀ n → ℕ×ℕ→Int (Int→ℕ×ℕ n) ≡ n + Int→ℕ×ℕ→Int (pos n) = refl + Int→ℕ×ℕ→Int (negsuc n) = refl + + ℕ×ℕ→Int→ℕ×ℕ : ∀ p → rel (Int→ℕ×ℕ (ℕ×ℕ→Int p)) p + ℕ×ℕ→Int→ℕ×ℕ (p₀ , p₁) = relInt→ℕ×ℕ p₀ p₁ + +-- Now build to ℤ using the above + +Int→ℤ : Int → ℤ +Int→ℤ n = [ Int→ℕ×ℕ n ] + +ℤ→Int : ℤ → Int +ℤ→Int [ z ] = ℕ×ℕ→Int z +ℤ→Int(eq/ a b r i) = lemℤeq a b r i + where lemℤeq : (a b : (ℕ × ℕ)) → rel a b → ℕ×ℕ→Int(a) ≡ ℕ×ℕ→Int(b) + lemℤeq (a₀ , a₁) (b₀ , b₁) r = a₀ ℕ- a₁ ≡⟨ pos- a₀ a₁ ⟩ + pos a₀ - pos a₁ ≡[ i ]⟨ ((pos a₀ - pos a₁) + -Cancel (pos b₁) (~ i)) ⟩ + (pos a₀ - pos a₁) + (pos b₁ - pos b₁) ≡⟨ +Assoc (pos a₀ + (- pos a₁)) (pos b₁) (- pos b₁) ⟩ + ((pos a₀ - pos a₁) + pos b₁) - pos b₁ ≡[ i ]⟨ +Assoc (pos a₀) (- pos a₁) (pos b₁) (~ i) + (- pos b₁) ⟩ + (pos a₀ + ((- pos a₁) + pos b₁)) - pos b₁ ≡[ i ]⟨ (pos a₀ + +Comm (- pos a₁) (pos b₁) i) - pos b₁ ⟩ + (pos a₀ + (pos b₁ - pos a₁)) - pos b₁ ≡[ i ]⟨ +Assoc (pos a₀) (pos b₁) (- pos a₁) i + (- pos b₁) ⟩ + ((pos a₀ + pos b₁) - pos a₁) - pos b₁ ≡[ i ]⟨ (pos+ a₀ b₁ (~ i) - pos a₁) - pos b₁ ⟩ + (pos (a₀ +ℕ b₁) - pos a₁) - pos b₁ ≡[ i ]⟨ (pos (r i) - pos a₁) - pos b₁ ⟩ + (pos (b₀ +ℕ a₁) - pos a₁) - pos b₁ ≡[ i ]⟨ (pos+ b₀ a₁ i - pos a₁) - pos b₁ ⟩ + ((pos b₀ + pos a₁) - pos a₁) - pos b₁ ≡[ i ]⟨ +Assoc (pos b₀) (pos a₁) (- pos a₁) (~ i) + (- pos b₁) ⟩ + (pos b₀ + (pos a₁ - pos a₁)) - pos b₁ ≡[ i ]⟨ (pos b₀ + (-Cancel (pos a₁) i)) - pos b₁ ⟩ + pos b₀ - pos b₁ ≡[ i ]⟨ pos- b₀ b₁ (~ i) ⟩ + b₀ ℕ- b₁ ∎ +ℤ→Int(squash/ x x₀ p q i j) = isSetℤ (ℤ→Int x) (ℤ→Int x₀) (cong ℤ→Int p) (cong ℤ→Int q) i j + +Int→ℤ→Int : ∀ z → ℤ→Int (Int→ℤ z) ≡ z +Int→ℤ→Int (pos n) = refl +Int→ℤ→Int (negsuc n) = refl + +ℤ→Int→ℤ : ∀ z → Int→ℤ (ℤ→Int z) ≡ z +ℤ→Int→ℤ = elimProp (λ z → squash/ (Int→ℤ (ℤ→Int z)) z) ℕ×ℕprf + where ℕ×ℕprf : (a : ℕ × ℕ) → Int→ℤ (ℤ→Int [ a ]) ≡ [ a ] + ℕ×ℕprf (a , b) = eq/ (Int→ℕ×ℕ (ℕ×ℕ→Int (a , b))) (a , b) (ℕ×ℕ→Int→ℕ×ℕ (a , b)) + +Int≡DiffInt : Int ≡ ℤ +Int≡DiffInt = isoToPath (iso Int→ℤ ℤ→Int ℤ→Int→ℤ Int→ℤ→Int) diff --git a/Cubical/Data/DiffInt/Properties.agda b/Cubical/Data/Int/MoreInts/DiffInt/Properties.agda similarity index 87% rename from Cubical/Data/DiffInt/Properties.agda rename to Cubical/Data/Int/MoreInts/DiffInt/Properties.agda index b50fc8431..b413c8301 100644 --- a/Cubical/Data/DiffInt/Properties.agda +++ b/Cubical/Data/Int/MoreInts/DiffInt/Properties.agda @@ -1,13 +1,12 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.Data.DiffInt.Properties where +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.DiffInt.Properties where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Univalence -open import Cubical.Data.DiffInt.Base +open import Cubical.Data.Int.MoreInts.DiffInt.Base open import Cubical.Data.Nat open import Cubical.Data.Sigma -open import Cubical.Data.Prod open import Cubical.Data.Bool open import Cubical.Relation.Binary.Base @@ -18,7 +17,7 @@ open import Cubical.HITs.SetQuotients open BinaryRelation relIsEquiv : isEquivRel rel -relIsEquiv = EquivRel {A = ℕ ×Σ ℕ} relIsRefl relIsSym relIsTrans +relIsEquiv = equivRel {A = ℕ × ℕ} relIsRefl relIsSym relIsTrans where relIsRefl : isRefl rel relIsRefl (a0 , a1) = refl @@ -33,7 +32,7 @@ relIsEquiv = EquivRel {A = ℕ ×Σ ℕ} relIsRefl relIsSym relIsTrans ((b1 + b0) + a0) + c1 ≡[ i ]⟨ +-comm (b1 + b0) a0 i + c1 ⟩ (a0 + (b1 + b0)) + c1 ≡[ i ]⟨ +-assoc a0 b1 b0 i + c1 ⟩ (a0 + b1) + b0 + c1 ≡⟨ sym (+-assoc (a0 + b1) b0 c1) ⟩ - (a0 + b1) + (b0 + c1) ≡⟨ cong (λ p → p . fst + p .snd) (transport (ua Σ≡) (p0 , p1))⟩ + (a0 + b1) + (b0 + c1) ≡⟨ cong (λ p → p . fst + p .snd) (transport ΣPath≡PathΣ (p0 , p1))⟩ (b0 + a1) + (c0 + b1) ≡⟨ sym (+-assoc b0 a1 (c0 + b1))⟩ b0 + (a1 + (c0 + b1)) ≡[ i ]⟨ b0 + (a1 + +-comm c0 b1 i) ⟩ b0 + (a1 + (b1 + c0)) ≡[ i ]⟨ b0 + +-comm a1 (b1 + c0) i ⟩ diff --git a/Cubical/Data/Int/MoreInts/QuoInt.agda b/Cubical/Data/Int/MoreInts/QuoInt.agda new file mode 100644 index 000000000..a40448242 --- /dev/null +++ b/Cubical/Data/Int/MoreInts/QuoInt.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.QuoInt where + +open import Cubical.Data.Int.MoreInts.QuoInt.Base public + +open import Cubical.Data.Int.MoreInts.QuoInt.Properties public diff --git a/Cubical/Data/Int/MoreInts/QuoInt/Base.agda b/Cubical/Data/Int/MoreInts/QuoInt/Base.agda new file mode 100644 index 000000000..7de09c85f --- /dev/null +++ b/Cubical/Data/Int/MoreInts/QuoInt/Base.agda @@ -0,0 +1,215 @@ +-- Define the integers as a HIT by identifying +0 and -0 +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.QuoInt.Base where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Relation.Nullary + +open import Cubical.Data.Int using () renaming (ℤ to Int ; discreteℤ to discreteInt ; isSetℤ to isSetInt) +open import Cubical.Data.Nat as ℕ using (ℕ; zero; suc) +open import Cubical.Data.Bool as Bool using (Bool; not; notnot) + +variable + l : Level + + +Sign : Type₀ +Sign = Bool + +pattern spos = Bool.false +pattern sneg = Bool.true + +_·S_ : Sign → Sign → Sign +_·S_ = Bool._⊕_ + + +data ℤ : Type₀ where + signed : (s : Sign) (n : ℕ) → ℤ + posneg : signed spos 0 ≡ signed sneg 0 + +pattern pos n = signed spos n +pattern neg n = signed sneg n + + +sign : ℤ → Sign +sign (signed _ zero) = spos +sign (signed s (suc _)) = s +sign (posneg i) = spos + +sign-pos : ∀ n → sign (pos n) ≡ spos +sign-pos zero = refl +sign-pos (suc n) = refl + +abs : ℤ → ℕ +abs (signed _ n) = n +abs (posneg i) = zero + +signed-inv : ∀ n → signed (sign n) (abs n) ≡ n +signed-inv (pos zero) = refl +signed-inv (neg zero) = posneg +signed-inv (signed s (suc n)) = refl +signed-inv (posneg i) j = posneg (i ∧ j) + +signed-zero : ∀ s₁ s₂ → signed s₁ zero ≡ signed s₂ zero +signed-zero spos spos = refl +signed-zero sneg sneg = refl +signed-zero spos sneg = posneg +signed-zero sneg spos = sym posneg + + +rec : ∀ {A : Type l} → (pos' neg' : ℕ → A) → pos' 0 ≡ neg' 0 → ℤ → A +rec pos' neg' eq (pos m) = pos' m +rec pos' neg' eq (neg m) = neg' m +rec pos' neg' eq (posneg i) = eq i + +elim : ∀ (P : ℤ → Type l) + → (pos' : ∀ n → P (pos n)) + → (neg' : ∀ n → P (neg n)) + → (λ i → P (posneg i)) [ pos' 0 ≡ neg' 0 ] + → ∀ z → P z +elim P pos' neg' eq (pos n) = pos' n +elim P pos' neg' eq (neg n) = neg' n +elim P pos' neg' eq (posneg i) = eq i + + +Int→ℤ : Int → ℤ +Int→ℤ (Int.pos n) = pos n +Int→ℤ (Int.negsuc n) = neg (suc n) + +ℤ→Int : ℤ → Int +ℤ→Int (pos n) = Int.pos n +ℤ→Int (neg zero) = Int.pos 0 +ℤ→Int (neg (suc n)) = Int.negsuc n +ℤ→Int (posneg _) = Int.pos 0 + +ℤ→Int→ℤ : ∀ (n : ℤ) → Int→ℤ (ℤ→Int n) ≡ n +ℤ→Int→ℤ (pos n) _ = pos n +ℤ→Int→ℤ (neg zero) i = posneg i +ℤ→Int→ℤ (neg (suc n)) _ = neg (suc n) +ℤ→Int→ℤ (posneg j) i = posneg (j ∧ i) + +Int→ℤ→Int : ∀ (n : Int) → ℤ→Int (Int→ℤ n) ≡ n +Int→ℤ→Int (Int.pos n) _ = Int.pos n +Int→ℤ→Int (Int.negsuc n) _ = Int.negsuc n + +Int≡ℤ : Int ≡ ℤ +Int≡ℤ = isoToPath (iso Int→ℤ ℤ→Int ℤ→Int→ℤ Int→ℤ→Int) + +discreteℤ : Discrete ℤ +discreteℤ = subst Discrete Int≡ℤ discreteInt + +isSetℤ : isSet ℤ +isSetℤ = subst isSet Int≡ℤ isSetInt + + +-_ : ℤ → ℤ +- signed s n = signed (not s) n +- posneg i = posneg (~ i) + +negate-invol : ∀ n → - - n ≡ n +negate-invol (signed s n) i = signed (notnot s i) n +negate-invol (posneg i) _ = posneg i + +negateEquiv : ℤ ≃ ℤ +negateEquiv = isoToEquiv (iso -_ -_ negate-invol negate-invol) + +negateEq : ℤ ≡ ℤ +negateEq = ua negateEquiv + + +infixl 6 _+_ +infixl 7 _·_ + +sucℤ : ℤ → ℤ +sucℤ (pos n) = pos (suc n) +sucℤ (neg zero) = pos 1 +sucℤ (neg (suc n)) = neg n +sucℤ (posneg _) = pos 1 + +predℤ : ℤ → ℤ +predℤ = subst (λ Z → (Z → Z)) negateEq sucℤ + -- definitionally equal to λ n → - (sucℤ (- n)) + -- strictly more useful than the direct pattern matching version, + -- see negateSuc and negatePred + +sucPredℤ : ∀ n → sucℤ (predℤ n) ≡ n +sucPredℤ (pos zero) = sym posneg +sucPredℤ (pos (suc _)) = refl +sucPredℤ (neg _) = refl +sucPredℤ (posneg i) j = posneg (i ∨ ~ j) + +predSucℤ : ∀ n → predℤ (sucℤ n) ≡ n +predSucℤ (pos _) = refl +predSucℤ (neg zero) = posneg +predSucℤ (neg (suc _)) = refl +predSucℤ (posneg i) j = posneg (i ∧ j) + +_+_ : ℤ → ℤ → ℤ +(signed _ zero) + n = n +(posneg _) + n = n +(pos (suc m)) + n = sucℤ (pos m + n) +(neg (suc m)) + n = predℤ (neg m + n) + + +sucPathℤ : ℤ ≡ ℤ +sucPathℤ = isoToPath (iso sucℤ predℤ sucPredℤ predSucℤ) + +-- We do the same trick as in Cubical.Data.Int to prove that addition +-- is an equivalence +addEqℤ : ℕ → ℤ ≡ ℤ +addEqℤ zero = refl +addEqℤ (suc n) = addEqℤ n ∙ sucPathℤ + +predPathℤ : ℤ ≡ ℤ +predPathℤ = isoToPath (iso predℤ sucℤ predSucℤ sucPredℤ) + +subEqℤ : ℕ → ℤ ≡ ℤ +subEqℤ zero = refl +subEqℤ (suc n) = subEqℤ n ∙ predPathℤ + +addℤ : ℤ → ℤ → ℤ +addℤ (pos m) n = transport (addEqℤ m) n +addℤ (neg m) n = transport (subEqℤ m) n +addℤ (posneg _) n = n + +isEquivAddℤ : (m : ℤ) → isEquiv (addℤ m) +isEquivAddℤ (pos n) = isEquivTransport (addEqℤ n) +isEquivAddℤ (neg n) = isEquivTransport (subEqℤ n) +isEquivAddℤ (posneg _) = isEquivTransport refl + +addℤ≡+ℤ : addℤ ≡ _+_ +addℤ≡+ℤ i (pos (suc m)) n = sucℤ (addℤ≡+ℤ i (pos m) n) +addℤ≡+ℤ i (neg (suc m)) n = predℤ (addℤ≡+ℤ i (neg m) n) +addℤ≡+ℤ i (pos zero) n = n +addℤ≡+ℤ i (neg zero) n = n +addℤ≡+ℤ _ (posneg _) n = n + +isEquiv+ℤ : (m : ℤ) → isEquiv (m +_) +isEquiv+ℤ = subst (λ _+_ → (m : ℤ) → isEquiv (m +_)) addℤ≡+ℤ isEquivAddℤ + + +_·_ : ℤ → ℤ → ℤ +m · n = signed (sign m ·S sign n) (abs m ℕ.· abs n) + +private + ·-abs : ∀ m n → abs (m · n) ≡ abs m ℕ.· abs n + ·-abs m n = refl + + +-- Natural number and negative integer literals for ℤ + +open import Cubical.Data.Nat.Literals public + +instance + fromNatℤ : HasFromNat ℤ + fromNatℤ = record { Constraint = λ _ → Unit ; fromNat = λ n → pos n } + +instance + fromNegℤ : HasFromNeg ℤ + fromNegℤ = record { Constraint = λ _ → Unit ; fromNeg = λ n → neg n } diff --git a/Cubical/Data/Int/MoreInts/QuoInt/Properties.agda b/Cubical/Data/Int/MoreInts/QuoInt/Properties.agda new file mode 100644 index 000000000..ac61e796b --- /dev/null +++ b/Cubical/Data/Int/MoreInts/QuoInt/Properties.agda @@ -0,0 +1,235 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Int.MoreInts.QuoInt.Properties where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Relation.Nullary + +open import Cubical.Data.Nat as ℕ using (ℕ; zero; suc) +open import Cubical.Data.Bool as Bool using (Bool; not; notnot) +open import Cubical.Data.Empty +open import Cubical.Data.Unit renaming (Unit to ⊤) + +open import Cubical.Data.Int.MoreInts.QuoInt.Base + +·S-comm : ∀ x y → x ·S y ≡ y ·S x +·S-comm = Bool.⊕-comm + +·S-assoc : ∀ x y z → x ·S (y ·S z) ≡ (x ·S y) ·S z +·S-assoc = Bool.⊕-assoc + +not-·Sˡ : ∀ x y → not (x ·S y) ≡ not x ·S y +not-·Sˡ = Bool.not-⊕ˡ + + +snotz : ∀ s n s' → ¬ (signed s (suc n) ≡ signed s' zero) +snotz s n s' eq = subst (λ { (signed s (suc n)) → ⊤ + ; _ → ⊥ }) eq tt + + ++-zeroʳ : ∀ s m → m + signed s zero ≡ m ++-zeroʳ s (signed s' zero) = signed-zero s s' ++-zeroʳ s (pos (suc n)) = cong sucℤ (+-zeroʳ s (pos n)) ++-zeroʳ s (neg (suc n)) = cong predℤ (+-zeroʳ s (neg n)) ++-zeroʳ spos (posneg i) j = posneg (i ∧ j) ++-zeroʳ sneg (posneg i) j = posneg (i ∨ ~ j) + ++-identityʳ : ∀ m → m + pos zero ≡ m ++-identityʳ = +-zeroʳ spos + +sucℤ-+ʳ : ∀ m n → sucℤ (m + n) ≡ m + sucℤ n +sucℤ-+ʳ (signed _ zero) _ = refl +sucℤ-+ʳ (posneg _) _ = refl +sucℤ-+ʳ (pos (suc m)) n = cong sucℤ (sucℤ-+ʳ (pos m) n) +sucℤ-+ʳ (neg (suc m)) n = + sucPredℤ (neg m + n) ∙∙ + sym (predSucℤ (neg m + n)) ∙∙ + cong predℤ (sucℤ-+ʳ (neg m) n) + +-- I wonder if we could somehow use negEq to get this for free from the above... +predℤ-+ʳ : ∀ m n → predℤ (m + n) ≡ m + predℤ n +predℤ-+ʳ (signed _ zero) _ = refl +predℤ-+ʳ (posneg _) _ = refl +predℤ-+ʳ (neg (suc m)) n = cong predℤ (predℤ-+ʳ (neg m) n) +predℤ-+ʳ (pos (suc m)) n = + predSucℤ (pos m + n) ∙∙ + sym (sucPredℤ (pos m + n)) ∙∙ + cong sucℤ (predℤ-+ʳ (pos m) n) + ++-comm : ∀ m n → m + n ≡ n + m ++-comm (signed s zero) n = sym (+-zeroʳ s n) ++-comm (pos (suc m)) n = cong sucℤ (+-comm (pos m) n) ∙ sucℤ-+ʳ n (pos m) ++-comm (neg (suc m)) n = cong predℤ (+-comm (neg m) n) ∙ predℤ-+ʳ n (neg m) ++-comm (posneg i) n = lemma n i + where + -- get some help from: + -- https://www.codewars.com/kata/reviews/5c878e3ef1abb10001e32eb1/groups/5cca3f9e840f4b0001d8ac98 + lemma : ∀ n → (λ i → (posneg i + n) ≡ (n + posneg i)) + [ sym (+-zeroʳ spos n) ≡ sym (+-zeroʳ sneg n) ] + lemma (pos zero) i j = posneg (i ∧ j) + lemma (pos (suc n)) i = cong sucℤ (lemma (pos n) i) + lemma (neg zero) i j = posneg (i ∨ ~ j) + lemma (neg (suc n)) i = cong predℤ (lemma (neg n) i) + lemma (posneg i) j k = posneg ((i ∧ ~ k) ∨ (i ∧ j) ∨ (j ∧ k)) + +sucℤ-+ˡ : ∀ m n → sucℤ (m + n) ≡ sucℤ m + n +sucℤ-+ˡ (pos _) n = refl +sucℤ-+ˡ (neg zero) n = refl +sucℤ-+ˡ (neg (suc m)) n = sucPredℤ _ +sucℤ-+ˡ (posneg _) n = refl + +-- I wonder if we could somehow use negEq to get this for free from the above... +predℤ-+ˡ : ∀ m n → predℤ (m + n) ≡ predℤ m + n +predℤ-+ˡ (pos zero) n = refl +predℤ-+ˡ (pos (suc m)) n = predSucℤ _ +predℤ-+ˡ (neg _) n = refl +predℤ-+ˡ (posneg _) n = refl + ++-assoc : ∀ m n o → m + (n + o) ≡ m + n + o ++-assoc (signed s zero) n o = refl ++-assoc (posneg i) n o = refl ++-assoc (pos (suc m)) n o = cong sucℤ (+-assoc (pos m) n o) ∙ sucℤ-+ˡ (pos m + n) o ++-assoc (neg (suc m)) n o = cong predℤ (+-assoc (neg m) n o) ∙ predℤ-+ˡ (neg m + n) o + + +sucℤ-inj : ∀ m n → sucℤ m ≡ sucℤ n → m ≡ n +sucℤ-inj m n p = sym (predSucℤ m) ∙ cong predℤ p ∙ predSucℤ n + +predℤ-inj : ∀ m n → predℤ m ≡ predℤ n → m ≡ n +predℤ-inj m n p = sym (sucPredℤ m) ∙ cong sucℤ p ∙ sucPredℤ n + ++-injˡ : ∀ o m n → o + m ≡ o + n → m ≡ n ++-injˡ (signed _ zero) _ _ p = p ++-injˡ (posneg _) _ _ p = p ++-injˡ (pos (suc o)) m n p = +-injˡ (pos o) m n (sucℤ-inj _ _ p) ++-injˡ (neg (suc o)) m n p = +-injˡ (neg o) m n (predℤ-inj _ _ p) + ++-injʳ : ∀ m n o → m + o ≡ n + o → m ≡ n ++-injʳ m n o p = +-injˡ o m n (+-comm o m ∙ p ∙ +-comm n o) + + +·-comm : ∀ m n → m · n ≡ n · m +·-comm m n i = signed (·S-comm (sign m) (sign n) i) (ℕ.·-comm (abs m) (abs n) i) + +·-identityˡ : ∀ n → pos 1 · n ≡ n +·-identityˡ n = cong (signed (sign n)) (ℕ.+-zero (abs n)) ∙ signed-inv n + +·-identityʳ : ∀ n → n · pos 1 ≡ n +·-identityʳ n = ·-comm n (pos 1) ∙ ·-identityˡ n + +·-zeroˡ : ∀ {s} n → signed s zero · n ≡ signed s zero +·-zeroˡ _ = signed-zero _ _ + +·-zeroʳ : ∀ {s} n → n · signed s zero ≡ signed s zero +·-zeroʳ n = cong (signed _) (sym (ℕ.0≡m·0 (abs n))) ∙ signed-zero _ _ + +·-signed-pos : ∀ {s} m n → signed s m · pos n ≡ signed s (m ℕ.· n) +·-signed-pos {s} zero n = ·-zeroˡ {s} (pos n) +·-signed-pos {s} (suc m) n i = signed (·S-comm s (sign-pos n i) i) (suc m ℕ.· n) + + +-- this proof is why we defined ℤ using `signed` instead of `pos` and `neg` +-- based on that in: https://github.com/danr/Agda-Numerics +·-assoc : ∀ m n o → m · (n · o) ≡ m · n · o + +·-assoc (signed s zero) n o = + ·-zeroˡ (n · o) +·-assoc m@(signed _ (suc _)) (signed s zero) o = + ·-zeroʳ {sign o} m ∙ signed-zero _ _ ∙ cong (_· o) (sym (·-zeroʳ {s} m)) +·-assoc m@(signed _ (suc _)) n@(signed _ (suc _)) (signed s zero) = + cong (m ·_) (·-zeroʳ {s} n) ∙ ·-zeroʳ {s} m ∙ sym (·-zeroʳ {s} (m · n)) + +·-assoc (signed sm (suc m)) (signed sn (suc n)) (signed so (suc o)) i = + signed (·S-assoc sm sn so i) (ℕ.·-assoc (suc m) (suc n) (suc o) i) + +·-assoc (posneg i) n o j = + isSet→isSet' isSetℤ (·-assoc (pos zero) n o) (·-assoc (neg zero) n o) + (λ i → posneg i · (n · o)) (λ i → posneg i · n · o) i j +·-assoc m@(signed _ (suc _)) (posneg i) o j = + isSet→isSet' isSetℤ (·-assoc m (pos zero) o) (·-assoc m (neg zero) o) + (λ i → m · (posneg i · o)) (λ i → m · posneg i · o) i j +·-assoc m@(signed _ (suc _)) n@(signed _ (suc _)) (posneg i) j = + isSet→isSet' isSetℤ (·-assoc m n (pos zero)) (·-assoc m n (neg zero)) + (λ i → m · (n · posneg i)) (λ i → m · n · posneg i) i j + + +negateSuc : ∀ n → - sucℤ n ≡ predℤ (- n) +negateSuc n i = - sucℤ (negate-invol n (~ i)) + +negatePred : ∀ n → - predℤ n ≡ sucℤ (- n) +negatePred n i = negate-invol (sucℤ (- n)) i + +negate-+ : ∀ m n → - (m + n) ≡ (- m) + (- n) +negate-+ (signed _ zero) n = refl +negate-+ (posneg _) n = refl +negate-+ (pos (suc m)) n = negateSuc (pos m + n) ∙ cong predℤ (negate-+ (pos m) n) +negate-+ (neg (suc m)) n = negatePred (neg m + n) ∙ cong sucℤ (negate-+ (neg m) n) + + +negate-·ˡ : ∀ m n → - (m · n) ≡ (- m) · n +negate-·ˡ (signed _ zero) n = signed-zero (not (sign n)) (sign n) +negate-·ˡ (signed ss (suc m)) n i = signed (not-·Sˡ ss (sign n) i) (suc m ℕ.· abs n) +negate-·ˡ (posneg i) n j = + isSet→isSet' isSetℤ (signed-zero (not (sign n)) _) (signed-zero _ _) + refl (λ i → posneg (~ i) · n) i j + + +signed-distrib : ∀ s m n → signed s (m ℕ.+ n) ≡ signed s m + signed s n +signed-distrib s zero n = refl +signed-distrib spos (suc m) n = cong sucℤ (signed-distrib spos m n) +signed-distrib sneg (suc m) n = cong predℤ (signed-distrib sneg m n) + +·-pos-suc : ∀ m n → pos (suc m) · n ≡ n + pos m · n +·-pos-suc m n = signed-distrib (sign n) (abs n) (m ℕ.· abs n) + ∙ (λ i → signed-inv n i + signed (sign-pos m (~ i) ·S sign n) (m ℕ.· abs n)) + + +-- the below is based on that in: https://github.com/danr/Agda-Numerics + +·-distribˡ-pos : ∀ o m n → (pos o · m) + (pos o · n) ≡ pos o · (m + n) +·-distribˡ-pos zero m n = signed-zero (sign n) (sign (m + n)) +·-distribˡ-pos (suc o) m n = + pos (suc o) · m + pos (suc o) · n ≡[ i ]⟨ ·-pos-suc o m i + ·-pos-suc o n i ⟩ + m + pos o · m + (n + pos o · n) ≡⟨ +-assoc (m + pos o · m) n (pos o · n) ⟩ + m + pos o · m + n + pos o · n ≡[ i ]⟨ +-assoc m (pos o · m) n (~ i) + pos o · n ⟩ + m + (pos o · m + n) + pos o · n ≡[ i ]⟨ m + +-comm (pos o · m) n i + pos o · n ⟩ + m + (n + pos o · m) + pos o · n ≡[ i ]⟨ +-assoc m n (pos o · m) i + pos o · n ⟩ + m + n + pos o · m + pos o · n ≡⟨ sym (+-assoc (m + n) (pos o · m) (pos o · n)) ⟩ + m + n + (pos o · m + pos o · n) ≡⟨ cong ((m + n) +_) (·-distribˡ-pos o m n) ⟩ + m + n + pos o · (m + n) ≡⟨ sym (·-pos-suc o (m + n)) ⟩ + pos (suc o) · (m + n) ∎ + +·-distribˡ-neg : ∀ o m n → (neg o · m) + (neg o · n) ≡ neg o · (m + n) +·-distribˡ-neg o m n = + neg o · m + neg o · n ≡[ i ]⟨ negate-·ˡ (pos o) m (~ i) + negate-·ˡ (pos o) n (~ i) ⟩ + - (pos o · m) + - (pos o · n) ≡⟨ sym (negate-+ (pos o · m) (pos o · n)) ⟩ + - (pos o · m + pos o · n) ≡⟨ cong -_ (·-distribˡ-pos o m n) ⟩ + - (pos o · (m + n)) ≡⟨ negate-·ˡ (pos o) (m + n) ⟩ + neg o · (m + n) ∎ + +·-distribˡ : ∀ o m n → (o · m) + (o · n) ≡ o · (m + n) +·-distribˡ (pos o) m n = ·-distribˡ-pos o m n +·-distribˡ (neg o) m n = ·-distribˡ-neg o m n +·-distribˡ (posneg i) m n j = + isSet→isSet' isSetℤ (·-distribˡ-pos zero m n) (·-distribˡ-neg zero m n) + (λ i → posneg i · n) (λ i → posneg i · (m + n)) i j + +·-distribʳ : ∀ m n o → (m · o) + (n · o) ≡ (m + n) · o +·-distribʳ m n o = (λ i → ·-comm m o i + ·-comm n o i) ∙ ·-distribˡ o m n ∙ ·-comm o (m + n) + + +sign-pos-suc-· : ∀ m n → sign (pos (suc m) · n) ≡ sign n +sign-pos-suc-· m (signed s zero) = sign-pos (suc m ℕ.· zero) +sign-pos-suc-· m (posneg i) = sign-pos (suc m ℕ.· zero) +sign-pos-suc-· m (signed s (suc n)) = refl + +·-injˡ : ∀ o m n → pos (suc o) · m ≡ pos (suc o) · n → m ≡ n +·-injˡ o m n p = sym (signed-inv m) ∙ (λ i → signed (sign-eq i) (abs-eq i)) ∙ signed-inv n + where sign-eq = sym (sign-pos-suc-· o m) ∙ cong sign p ∙ sign-pos-suc-· o n + abs-eq = ℕ.inj-sm· {o} (cong abs p) + +·-injʳ : ∀ m n o → m · pos (suc o) ≡ n · pos (suc o) → m ≡ n +·-injʳ m n o p = ·-injˡ o m n (·-comm (pos (suc o)) m ∙ p ∙ ·-comm n (pos (suc o))) diff --git a/Cubical/Data/Int/Properties.agda b/Cubical/Data/Int/Properties.agda index ea37539fd..14e2c120f 100644 --- a/Cubical/Data/Int/Properties.agda +++ b/Cubical/Data/Int/Properties.agda @@ -1,28 +1,4 @@ -{-# OPTIONS --cubical --safe #-} - -{- - -This file defines - -sucPred : ∀ (i : Int) → sucInt (predInt i) ≡ i -predSuc : ∀ (i : Int) → predInt (sucInt i) ≡ i - -discreteInt : discrete Int -isSetInt : isSet Int - -addition of Int is defined _+_ : Int → Int → Int - -as well as its commutativity and associativity -+-comm : ∀ (m n : Int) → m + n ≡ n + m -+-assoc : ∀ (m n o : Int) → m + (n + o) ≡ (m + n) + o - -An alternate definition of _+_ is defined via ua, -namely _+'_, which helps us to easily prove - -isEquivAddInt : (m : Int) → isEquiv (λ n → n + m) - --} - +{-# OPTIONS --safe #-} module Cubical.Data.Int.Properties where open import Cubical.Core.Everything @@ -35,7 +11,9 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence open import Cubical.Data.Empty -open import Cubical.Data.Nat hiding (_+_ ; +-assoc ; +-comm) +open import Cubical.Data.Nat + hiding (+-assoc ; +-comm ; ·-comm) + renaming (_·_ to _·ℕ_; _+_ to _+ℕ_ ; ·-assoc to ·ℕ-assoc) open import Cubical.Data.Bool open import Cubical.Data.Sum open import Cubical.Data.Int.Base @@ -43,140 +21,124 @@ open import Cubical.Data.Int.Base open import Cubical.Relation.Nullary open import Cubical.Relation.Nullary.DecidableEq -abs : Int → ℕ -abs (pos n) = n -abs (negsuc n) = suc n - -sgn : Int → Bool -sgn (pos n) = true -sgn (negsuc n) = false - -sucPred : ∀ i → sucInt (predInt i) ≡ i +sucPred : ∀ i → sucℤ (predℤ i) ≡ i sucPred (pos zero) = refl sucPred (pos (suc n)) = refl sucPred (negsuc n) = refl -predSuc : ∀ i → predInt (sucInt i) ≡ i +predSuc : ∀ i → predℤ (sucℤ i) ≡ i predSuc (pos n) = refl predSuc (negsuc zero) = refl predSuc (negsuc (suc n)) = refl --- TODO: define multiplication - injPos : ∀ {a b : ℕ} → pos a ≡ pos b → a ≡ b injPos {a} h = subst T h refl where - T : Int → Type₀ + T : ℤ → Type₀ T (pos x) = a ≡ x T (negsuc _) = ⊥ injNegsuc : ∀ {a b : ℕ} → negsuc a ≡ negsuc b → a ≡ b injNegsuc {a} h = subst T h refl where - T : Int → Type₀ + T : ℤ → Type₀ T (pos _) = ⊥ T (negsuc x) = a ≡ x posNotnegsuc : ∀ (a b : ℕ) → ¬ (pos a ≡ negsuc b) posNotnegsuc a b h = subst T h 0 where - T : Int → Type₀ + T : ℤ → Type₀ T (pos _) = ℕ T (negsuc _) = ⊥ negsucNotpos : ∀ (a b : ℕ) → ¬ (negsuc a ≡ pos b) negsucNotpos a b h = subst T h 0 where - T : Int → Type₀ + T : ℤ → Type₀ T (pos _) = ⊥ T (negsuc _) = ℕ -discreteInt : Discrete Int -discreteInt (pos n) (pos m) with discreteℕ n m +discreteℤ : Discrete ℤ +discreteℤ (pos n) (pos m) with discreteℕ n m ... | yes p = yes (cong pos p) ... | no p = no (λ x → p (injPos x)) -discreteInt (pos n) (negsuc m) = no (posNotnegsuc n m) -discreteInt (negsuc n) (pos m) = no (negsucNotpos n m) -discreteInt (negsuc n) (negsuc m) with discreteℕ n m +discreteℤ (pos n) (negsuc m) = no (posNotnegsuc n m) +discreteℤ (negsuc n) (pos m) = no (negsucNotpos n m) +discreteℤ (negsuc n) (negsuc m) with discreteℕ n m ... | yes p = yes (cong negsuc p) ... | no p = no (λ x → p (injNegsuc x)) -isSetInt : isSet Int -isSetInt = Discrete→isSet discreteInt - -_ℕ-_ : ℕ → ℕ → Int -a ℕ- 0 = pos a -0 ℕ- suc b = negsuc b -suc a ℕ- suc b = a ℕ- b - -_+pos_ : Int → ℕ → Int -z +pos 0 = z -z +pos (suc n) = sucInt (z +pos n) +isSetℤ : isSet ℤ +isSetℤ = Discrete→isSet discreteℤ -_+negsuc_ : Int → ℕ → Int -z +negsuc 0 = predInt z -z +negsuc (suc n) = predInt (z +negsuc n) +-pos : ∀ n → - (pos n) ≡ neg n +-pos zero = refl +-pos (suc n) = refl -_+_ : Int → Int → Int -m + pos n = m +pos n -m + negsuc n = m +negsuc n +-neg : ∀ n → - (neg n) ≡ pos n +-neg zero = refl +-neg (suc n) = refl -_-_ : Int → Int → Int -m - pos zero = m -m - pos (suc n) = m + negsuc n -m - negsuc n = m + pos (suc n) +-Involutive : ∀ z → - (- z) ≡ z +-Involutive (pos n) = cong -_ (-pos n) ∙ -neg n +-Involutive (negsuc n) = refl -sucInt+pos : ∀ n m → sucInt (m +pos n) ≡ (sucInt m) +pos n -sucInt+pos zero m = refl -sucInt+pos (suc n) m = cong sucInt (sucInt+pos n m) +sucℤ+pos : ∀ n m → sucℤ (m +pos n) ≡ (sucℤ m) +pos n +sucℤ+pos zero m = refl +sucℤ+pos (suc n) m = cong sucℤ (sucℤ+pos n m) -predInt+negsuc : ∀ n m → predInt (m +negsuc n) ≡ (predInt m) +negsuc n -predInt+negsuc zero m = refl -predInt+negsuc (suc n) m = cong predInt (predInt+negsuc n m) +predℤ+negsuc : ∀ n m → predℤ (m +negsuc n) ≡ (predℤ m) +negsuc n +predℤ+negsuc zero m = refl +predℤ+negsuc (suc n) m = cong predℤ (predℤ+negsuc n m) -sucInt+negsuc : ∀ n m → sucInt (m +negsuc n) ≡ (sucInt m) +negsuc n -sucInt+negsuc zero m = (sucPred _) ∙ (sym (predSuc _)) -sucInt+negsuc (suc n) m = _ ≡⟨ sucPred _ ⟩ +sucℤ+negsuc : ∀ n m → sucℤ (m +negsuc n) ≡ (sucℤ m) +negsuc n +sucℤ+negsuc zero m = (sucPred _) ∙ (sym (predSuc _)) +sucℤ+negsuc (suc n) m = _ ≡⟨ sucPred _ ⟩ m +negsuc n ≡[ i ]⟨ predSuc m (~ i) +negsuc n ⟩ - (predInt (sucInt m)) +negsuc n ≡⟨ sym (predInt+negsuc n (sucInt m)) ⟩ - predInt (sucInt m +negsuc n) ∎ + (predℤ (sucℤ m)) +negsuc n ≡⟨ sym (predℤ+negsuc n (sucℤ m)) ⟩ + predℤ (sucℤ m +negsuc n) ∎ -predInt+pos : ∀ n m → predInt (m +pos n) ≡ (predInt m) +pos n -predInt+pos zero m = refl -predInt+pos (suc n) m = _ ≡⟨ predSuc _ ⟩ +predℤ+pos : ∀ n m → predℤ (m +pos n) ≡ (predℤ m) +pos n +predℤ+pos zero m = refl +predℤ+pos (suc n) m = _ ≡⟨ predSuc _ ⟩ m +pos n ≡[ i ]⟨ sucPred m (~ i) + pos n ⟩ - (sucInt (predInt m)) +pos n ≡⟨ sym (sucInt+pos n (predInt m))⟩ - (predInt m) +pos (suc n) ∎ + (sucℤ (predℤ m)) +pos n ≡⟨ sym (sucℤ+pos n (predℤ m))⟩ + (predℤ m) +pos (suc n) ∎ -predInt+ : ∀ m n → predInt (m + n) ≡ (predInt m) + n -predInt+ m (pos n) = predInt+pos n m -predInt+ m (negsuc n) = predInt+negsuc n m +predℤ-pos : ∀ n → predℤ(- (pos n)) ≡ negsuc n +predℤ-pos zero = refl +predℤ-pos (suc n) = refl -+predInt : ∀ m n → predInt (m + n) ≡ m + (predInt n) -+predInt m (pos zero) = refl -+predInt m (pos (suc n)) = (predSuc (m + pos n)) ∙ (cong (_+_ m) (sym (predSuc (pos n)))) -+predInt m (negsuc n) = refl +predℤ+ : ∀ m n → predℤ (m + n) ≡ (predℤ m) + n +predℤ+ m (pos n) = predℤ+pos n m +predℤ+ m (negsuc n) = predℤ+negsuc n m -sucInt+ : ∀ m n → sucInt (m + n) ≡ (sucInt m) + n -sucInt+ m (pos n) = sucInt+pos n m -sucInt+ m (negsuc n) = sucInt+negsuc n m ++predℤ : ∀ m n → predℤ (m + n) ≡ m + (predℤ n) ++predℤ m (pos zero) = refl ++predℤ m (pos (suc n)) = (predSuc (m + pos n)) ∙ (cong (_+_ m) (sym (predSuc (pos n)))) ++predℤ m (negsuc n) = refl -+sucInt : ∀ m n → sucInt (m + n) ≡ m + (sucInt n) -+sucInt m (pos n) = refl -+sucInt m (negsuc zero) = sucPred _ -+sucInt m (negsuc (suc n)) = (sucPred (m +negsuc n)) ∙ (cong (_+_ m) (sym (sucPred (negsuc n)))) +sucℤ+ : ∀ m n → sucℤ (m + n) ≡ (sucℤ m) + n +sucℤ+ m (pos n) = sucℤ+pos n m +sucℤ+ m (negsuc n) = sucℤ+negsuc n m + ++sucℤ : ∀ m n → sucℤ (m + n) ≡ m + (sucℤ n) ++sucℤ m (pos n) = refl ++sucℤ m (negsuc zero) = sucPred _ ++sucℤ m (negsuc (suc n)) = (sucPred (m +negsuc n)) ∙ (cong (_+_ m) (sym (sucPred (negsuc n)))) pos0+ : ∀ z → z ≡ pos 0 + z pos0+ (pos zero) = refl -pos0+ (pos (suc n)) = cong sucInt (pos0+ (pos n)) +pos0+ (pos (suc n)) = cong sucℤ (pos0+ (pos n)) pos0+ (negsuc zero) = refl -pos0+ (negsuc (suc n)) = cong predInt (pos0+ (negsuc n)) +pos0+ (negsuc (suc n)) = cong predℤ (pos0+ (negsuc n)) -negsuc0+ : ∀ z → predInt z ≡ negsuc 0 + z +negsuc0+ : ∀ z → predℤ z ≡ negsuc 0 + z negsuc0+ (pos zero) = refl -negsuc0+ (pos (suc n)) = (sym (sucPred (pos n))) ∙ (cong sucInt (negsuc0+ _)) +negsuc0+ (pos (suc n)) = (sym (sucPred (pos n))) ∙ (cong sucℤ (negsuc0+ _)) negsuc0+ (negsuc zero) = refl -negsuc0+ (negsuc (suc n)) = cong predInt (negsuc0+ (negsuc n)) +negsuc0+ (negsuc (suc n)) = cong predℤ (negsuc0+ (negsuc n)) ind-comm : {A : Type₀} (_∙_ : A → A → A) (f : ℕ → A) (g : A → A) (p : ∀ {n} → f (suc n) ≡ g (f n)) @@ -213,67 +175,67 @@ ind-assoc _·_ f g p q base m n (suc o) = where IH = ind-assoc _·_ f g p q base m n o -+-comm : ∀ m n → m + n ≡ n + m -+-comm m (pos n) = ind-comm _+_ pos sucInt refl sucInt+ +sucInt pos0+ m n -+-comm m (negsuc n) = ind-comm _+_ negsuc predInt refl predInt+ +predInt negsuc0+ m n ++Comm : ∀ m n → m + n ≡ n + m ++Comm m (pos n) = ind-comm _+_ pos sucℤ refl sucℤ+ +sucℤ pos0+ m n ++Comm m (negsuc n) = ind-comm _+_ negsuc predℤ refl predℤ+ +predℤ negsuc0+ m n -+-assoc : ∀ m n o → m + (n + o) ≡ (m + n) + o -+-assoc m n (pos o) = ind-assoc _+_ pos sucInt +sucInt refl (λ _ _ → refl) m n o -+-assoc m n (negsuc o) = ind-assoc _+_ negsuc predInt +predInt refl +predInt m n o ++Assoc : ∀ m n o → m + (n + o) ≡ (m + n) + o ++Assoc m n (pos o) = ind-assoc _+_ pos sucℤ +sucℤ refl (λ _ _ → refl) m n o ++Assoc m n (negsuc o) = ind-assoc _+_ negsuc predℤ +predℤ refl +predℤ m n o --- Compose sucPathInt with itself n times. Transporting along this +-- Compose sucPathℤ with itself n times. Transporting along this -- will be addition, transporting with it backwards will be subtraction. -- Use this to define _+'_ for which is easier to prove -- isEquiv (λ n → n +' m) since _+'_ is defined by transport -sucPathInt : Int ≡ Int -sucPathInt = ua (sucInt , isoToIsEquiv (iso sucInt predInt sucPred predSuc)) +sucPathℤ : ℤ ≡ ℤ +sucPathℤ = ua (sucℤ , isoToIsEquiv (iso sucℤ predℤ sucPred predSuc)) -addEq : ℕ → Int ≡ Int +addEq : ℕ → ℤ ≡ ℤ addEq zero = refl -addEq (suc n) = (addEq n) ∙ sucPathInt +addEq (suc n) = (addEq n) ∙ sucPathℤ -predPathInt : Int ≡ Int -predPathInt = ua (predInt , isoToIsEquiv (iso predInt sucInt predSuc sucPred)) +predPathℤ : ℤ ≡ ℤ +predPathℤ = ua (predℤ , isoToIsEquiv (iso predℤ sucℤ predSuc sucPred)) -subEq : ℕ → Int ≡ Int +subEq : ℕ → ℤ ≡ ℤ subEq zero = refl -subEq (suc n) = (subEq n) ∙ predPathInt +subEq (suc n) = (subEq n) ∙ predPathℤ -_+'_ : Int → Int → Int +_+'_ : ℤ → ℤ → ℤ m +' pos n = transport (addEq n) m m +' negsuc n = transport (subEq (suc n)) m +'≡+ : _+'_ ≡ _+_ +'≡+ i m (pos zero) = m -+'≡+ i m (pos (suc n)) = sucInt (+'≡+ i m (pos n)) -+'≡+ i m (negsuc zero) = predInt m -+'≡+ i m (negsuc (suc n)) = predInt (+'≡+ i m (negsuc n)) -- - -- compPath (λ i → (+'≡+ i (predInt m) (negsuc n))) (sym (predInt+negsuc n m)) i ++'≡+ i m (pos (suc n)) = sucℤ (+'≡+ i m (pos n)) ++'≡+ i m (negsuc zero) = predℤ m ++'≡+ i m (negsuc (suc n)) = predℤ (+'≡+ i m (negsuc n)) -- + -- compPath (λ i → (+'≡+ i (predℤ m) (negsuc n))) (sym (predℤ+negsuc n m)) i -isEquivAddInt' : (m : Int) → isEquiv (λ n → n +' m) -isEquivAddInt' (pos n) = isEquivTransport (addEq n) -isEquivAddInt' (negsuc n) = isEquivTransport (subEq (suc n)) +isEquivAddℤ' : (m : ℤ) → isEquiv (λ n → n +' m) +isEquivAddℤ' (pos n) = isEquivTransport (addEq n) +isEquivAddℤ' (negsuc n) = isEquivTransport (subEq (suc n)) -isEquivAddInt : (m : Int) → isEquiv (λ n → n + m) -isEquivAddInt = subst (λ add → (m : Int) → isEquiv (λ n → add n m)) +'≡+ isEquivAddInt' +isEquivAddℤ : (m : ℤ) → isEquiv (λ n → n + m) +isEquivAddℤ = subst (λ add → (m : ℤ) → isEquiv (λ n → add n m)) +'≡+ isEquivAddℤ' --- below is an alternate proof of isEquivAddInt for comparison +-- below is an alternate proof of isEquivAddℤ for comparison -- We also have two useful lemma here. minusPlus : ∀ m n → (n - m) + m ≡ n minusPlus (pos zero) n = refl minusPlus (pos 1) = sucPred minusPlus (pos (suc (suc m))) n = - sucInt ((n +negsuc (suc m)) +pos (suc m)) ≡⟨ sucInt+pos (suc m) _ ⟩ - sucInt (n +negsuc (suc m)) +pos (suc m) ≡[ i ]⟨ sucPred (n +negsuc m) i +pos (suc m) ⟩ + sucℤ ((n +negsuc (suc m)) +pos (suc m)) ≡⟨ sucℤ+pos (suc m) _ ⟩ + sucℤ (n +negsuc (suc m)) +pos (suc m) ≡[ i ]⟨ sucPred (n +negsuc m) i +pos (suc m) ⟩ (n - pos (suc m)) +pos (suc m) ≡⟨ minusPlus (pos (suc m)) n ⟩ n ∎ minusPlus (negsuc zero) = predSuc minusPlus (negsuc (suc m)) n = - predInt (sucInt (sucInt (n +pos m)) +negsuc m) ≡⟨ predInt+negsuc m _ ⟩ - predInt (sucInt (sucInt (n +pos m))) +negsuc m ≡[ i ]⟨ predSuc (sucInt (n +pos m)) i +negsuc m ⟩ - sucInt (n +pos m) +negsuc m ≡⟨ minusPlus (negsuc m) n ⟩ + predℤ (sucℤ (sucℤ (n +pos m)) +negsuc m) ≡⟨ predℤ+negsuc m _ ⟩ + predℤ (sucℤ (sucℤ (n +pos m))) +negsuc m ≡[ i ]⟨ predSuc (sucℤ (n +pos m)) i +negsuc m ⟩ + sucℤ (n +pos m) +negsuc m ≡⟨ minusPlus (negsuc m) n ⟩ n ∎ plusMinus : ∀ m n → (n + m) - m ≡ n @@ -282,8 +244,215 @@ plusMinus (pos (suc m)) = minusPlus (negsuc m) plusMinus (negsuc m) = minusPlus (pos (suc m)) private - alternateProof : (m : Int) → isEquiv (λ n → n + m) + alternateProof : (m : ℤ) → isEquiv (λ n → n + m) alternateProof m = isoToIsEquiv (iso (λ n → n + m) (λ n → n - m) (minusPlus m) (plusMinus m)) + +-Cancel : ∀ z → z - z ≡ 0 +-Cancel z = cong (_- z) (pos0+ z) ∙ plusMinus z (pos zero) + +pos+ : ∀ m n → pos (m +ℕ n) ≡ pos m + pos n +pos+ zero zero = refl +pos+ zero (suc n) = + pos (zero +ℕ suc n) ≡⟨ +Comm (pos (suc n)) (pos zero) ⟩ + pos zero + pos (suc n) ∎ +pos+ (suc m) zero = + pos (suc (m +ℕ zero)) ≡⟨ cong pos (cong suc (+-zero m)) ⟩ + pos (suc m) + pos zero ∎ +pos+ (suc m) (suc n) = + pos (suc m +ℕ suc n) ≡⟨ cong pos (cong suc (+-suc m n)) ⟩ + sucℤ (pos (suc (m +ℕ n))) ≡⟨ cong sucℤ (cong sucℤ (pos+ m n)) ⟩ + sucℤ (sucℤ (pos m + pos n)) ≡⟨ sucℤ+ (pos m) (sucℤ (pos n)) ⟩ + pos (suc m) + pos (suc n) ∎ + +negsuc+ : ∀ m n → negsuc (m +ℕ n) ≡ negsuc m - pos n +negsuc+ zero zero = refl +negsuc+ zero (suc n) = + negsuc (zero +ℕ suc n) ≡⟨ negsuc0+ (negsuc n) ⟩ + negsuc zero + negsuc n ≡⟨ cong (negsuc zero +_) (-pos (suc n)) ⟩ + negsuc zero - pos (suc n) ∎ +negsuc+ (suc m) zero = + negsuc (suc m +ℕ zero) ≡⟨ cong negsuc (cong suc (+-zero m)) ⟩ + negsuc (suc m) - pos zero ∎ +negsuc+ (suc m) (suc n) = + negsuc (suc m +ℕ suc n) ≡⟨ cong negsuc (sym (+-suc m (suc n))) ⟩ + negsuc (m +ℕ suc (suc n)) ≡⟨ negsuc+ m (suc (suc n)) ⟩ + negsuc m - pos (suc (suc n)) ≡⟨ sym (+predℤ (negsuc m) (negsuc n)) ⟩ + predℤ (negsuc m + negsuc n ) ≡⟨ predℤ+ (negsuc m) (negsuc n) ⟩ + negsuc (suc m) - pos (suc n) ∎ + +neg+ : ∀ m n → neg (m +ℕ n) ≡ neg m + neg n +neg+ zero zero = refl +neg+ zero (suc n) = neg (zero +ℕ suc n) ≡⟨ +Comm (neg (suc n)) (pos zero) ⟩ + neg zero + neg (suc n) ∎ +neg+ (suc m) zero = neg (suc (m +ℕ zero)) ≡⟨ cong neg (cong suc (+-zero m)) ⟩ + neg (suc m) + neg zero ∎ +neg+ (suc m) (suc n) = neg (suc m +ℕ suc n) ≡⟨ negsuc+ m (suc n) ⟩ + neg (suc m) + neg (suc n) ∎ + +ℕ-AntiComm : ∀ m n → m ℕ- n ≡ - (n ℕ- m) +ℕ-AntiComm zero zero = refl +ℕ-AntiComm zero (suc n) = refl +ℕ-AntiComm (suc m) zero = refl +ℕ-AntiComm (suc m) (suc n) = suc m ℕ- suc n ≡⟨ ℕ-AntiComm m n ⟩ + - (suc n ℕ- suc m) ∎ + +pos- : ∀ m n → m ℕ- n ≡ pos m - pos n +pos- zero zero = refl +pos- zero (suc n) = zero ℕ- suc n ≡⟨ +Comm (negsuc n) (pos zero) ⟩ + pos zero - pos (suc n) ∎ +pos- (suc m) zero = refl +pos- (suc m) (suc n) = + suc m ℕ- suc n ≡⟨ pos- m n ⟩ + pos m - pos n ≡⟨ sym (sucPred (pos m - pos n)) ⟩ + sucℤ (predℤ (pos m - pos n)) ≡⟨ cong sucℤ (+predℤ (pos m) (- pos n)) ⟩ + sucℤ (pos m + predℤ (- (pos n))) ≡⟨ cong sucℤ (cong (pos m +_) (predℤ-pos n)) ⟩ + sucℤ (pos m + negsuc n) ≡⟨ sucℤ+negsuc n (pos m) ⟩ + pos (suc m) - pos (suc n) ∎ + +-AntiComm : ∀ m n → m - n ≡ - (n - m) +-AntiComm (pos n) (pos m) = pos n - pos m ≡⟨ sym (pos- n m) ⟩ + n ℕ- m ≡⟨ ℕ-AntiComm n m ⟩ + - (m ℕ- n) ≡⟨ cong -_ (pos- m n) ⟩ + - (pos m - pos n) ∎ +-AntiComm (pos n) (negsuc m) = + pos n - negsuc m ≡⟨ +Comm (pos n) (pos (suc m)) ⟩ + pos (suc m) + pos n ≡⟨ sym (pos+ (suc m) n) ⟩ + pos (suc m +ℕ n) ≡⟨ sym (-neg (suc m +ℕ n)) ⟩ + - neg (suc m +ℕ n) ≡⟨ cong -_ (neg+ (suc m) n) ⟩ + - (neg (suc m) + neg n) ≡⟨ cong -_ (cong (negsuc m +_) (sym (-pos n))) ⟩ + - (negsuc m - pos n) ∎ +-AntiComm (negsuc n) (pos m) = + negsuc n - pos m ≡⟨ sym (negsuc+ n m) ⟩ + negsuc (n +ℕ m) ≡⟨ cong -_ (pos+ (suc n) m) ⟩ + - (pos (suc n) + pos m) ≡⟨ cong -_ (+Comm (pos (suc n)) (pos m)) ⟩ + - (pos m - negsuc n) ∎ +-AntiComm (negsuc n) (negsuc m) = + negsuc n - negsuc m ≡⟨ +Comm (negsuc n) (pos (suc m)) ⟩ + pos (suc m) + negsuc n ≡⟨ sym (pos- (suc m) (suc n)) ⟩ + suc m ℕ- suc n ≡⟨ ℕ-AntiComm (suc m) (suc n) ⟩ + - (suc n ℕ- suc m) ≡⟨ cong -_ (pos- (suc n) (suc m)) ⟩ + - (pos (suc n) - pos (suc m)) ≡⟨ cong -_ (+Comm (pos (suc n)) (negsuc m)) ⟩ + - (negsuc m - negsuc n) ∎ + +-Dist+ : ∀ m n → - (m + n) ≡ (- m) + (- n) +-Dist+ (pos n) (pos m) = + - (pos n + pos m) ≡⟨ cong -_ (sym (pos+ n m)) ⟩ + - pos (n +ℕ m) ≡⟨ -pos (n +ℕ m) ⟩ + neg (n +ℕ m) ≡⟨ neg+ n m ⟩ + neg n + neg m ≡⟨ cong (neg n +_) (sym (-pos m)) ⟩ + neg n + (- pos m) ≡⟨ cong (_+ (- pos m)) (sym (-pos n)) ⟩ + (- pos n) + (- pos m) ∎ +-Dist+ (pos n) (negsuc m) = + - (pos n + negsuc m) ≡⟨ sym (-AntiComm (pos (suc m)) (pos n)) ⟩ + pos (suc m) - pos n ≡⟨ +Comm (pos (suc m)) (- pos n) ⟩ + (- pos n) + (- negsuc m) ∎ +-Dist+ (negsuc n) (pos m) = + - (negsuc n + pos m) ≡⟨ cong -_ (+Comm (negsuc n) (pos m)) ⟩ + - (pos m + negsuc n) ≡⟨ sym (-AntiComm (- negsuc n) (pos m)) ⟩ + (- negsuc n) + (- pos m) ∎ +-Dist+ (negsuc n) (negsuc m) = + - (negsuc n + negsuc m) ≡⟨ cong -_ (sym (neg+ (suc n) (suc m))) ⟩ + - neg (suc n +ℕ suc m) ≡⟨ pos+ (suc n) (suc m) ⟩ + (- negsuc n) + (- negsuc m) ∎ + + +-- multiplication + +pos·negsuc : (n m : ℕ) → pos n · negsuc m ≡ - (pos n · pos (suc m)) +pos·negsuc zero m = refl +pos·negsuc (suc n) m = + (λ i → negsuc m + (pos·negsuc n m i)) + ∙ sym (-Dist+ (pos (suc m)) (pos n · pos (suc m))) + +negsuc·pos : (n m : ℕ) → negsuc n · pos m ≡ - (pos (suc n) · pos m) +negsuc·pos zero m = refl +negsuc·pos (suc n) m = + cong ((- pos m) +_) (negsuc·pos n m) + ∙ sym (-Dist+ (pos m) (pos m + (pos n · pos m))) + +negsuc·negsuc : (n m : ℕ) → negsuc n · negsuc m ≡ pos (suc n) · pos (suc m) +negsuc·negsuc zero m = refl +negsuc·negsuc (suc n) m = cong (pos (suc m) +_) (negsuc·negsuc n m) + +·Comm : (x y : ℤ) → x · y ≡ y · x +·Comm (pos n) (pos m) = lem n m + where + lem : (n m : ℕ) → (pos n · pos m) ≡ (pos m · pos n) + lem zero zero = refl + lem zero (suc m) i = +Comm (lem zero m i) (pos zero) i + lem (suc n) zero i = +Comm (pos zero) (lem n zero i) i + lem (suc n) (suc m) = + (λ i → pos (suc m) + (lem n (suc m) i)) + ∙∙ +Assoc (pos (suc m)) (pos n) (pos m · pos n) + ∙∙ (λ i → sucℤ+ (pos m) (pos n) (~ i) + (pos m · pos n)) + ∙∙ (λ i → +Comm (pos m) (pos (suc n)) i + (pos m · pos n)) + ∙∙ sym (+Assoc (pos (suc n)) (pos m) (pos m · pos n)) + ∙∙ (λ i → pos (suc n) + (pos m + (lem n m (~ i)))) + ∙∙ λ i → pos (suc n) + (lem (suc n) m i) +·Comm (pos n) (negsuc m) = + pos·negsuc n m + ∙∙ cong -_ (·Comm (pos n) (pos (suc m))) + ∙∙ sym (negsuc·pos m n) +·Comm (negsuc n) (pos m) = + sym (pos·negsuc m n + ∙∙ cong -_ (·Comm (pos m) (pos (suc n))) + ∙∙ sym (negsuc·pos n m)) +·Comm (negsuc n) (negsuc m) = + negsuc·negsuc n m ∙∙ ·Comm (pos (suc n)) (pos (suc m)) ∙∙ sym (negsuc·negsuc m n) + +·Rid : (x : ℤ) → x · 1 ≡ x +·Rid x = ·Comm x 1 + +private + distrHelper : (x y z w : ℤ) → (x + y) + (z + w) ≡ ((x + z) + (y + w)) + distrHelper x y z w = + +Assoc (x + y) z w + ∙∙ cong (_+ w) (sym (+Assoc x y z) ∙∙ cong (x +_) (+Comm y z) ∙∙ +Assoc x z y) + ∙∙ sym (+Assoc (x + z) y w) + +·DistR+ : (x y z : ℤ) → x · (y + z) ≡ x · y + x · z +·DistR+ (pos zero) y z = refl +·DistR+ (pos (suc n)) y z = + cong ((y + z) +_) (·DistR+ (pos n) y z) + ∙ distrHelper y z (pos n · y) (pos n · z) +·DistR+ (negsuc zero) y z = -Dist+ y z +·DistR+ (negsuc (suc n)) y z = + cong₂ _+_ (-Dist+ y z) (·DistR+ (negsuc n) y z) + ∙ distrHelper (- y) (- z) (negsuc n · y) (negsuc n · z) + +·DistL+ : (x y z : ℤ) → (x + y) · z ≡ x · z + y · z +·DistL+ x y z = ·Comm (x + y) z ∙∙ ·DistR+ z x y ∙∙ cong₂ _+_ (·Comm z x) (·Comm z y) + +-DistL· : (b c : ℤ) → - (b · c) ≡ - b · c +-DistL· (pos zero) c = refl +-DistL· (pos (suc n)) (pos m) = sym (negsuc·pos n m) +-DistL· (pos (suc zero)) (negsuc m) = + -Dist+ (negsuc m) (pos zero · negsuc m) + ∙ cong (pos (suc m) +_) (-DistL· (pos zero) (negsuc m)) +-DistL· (pos (suc (suc n))) (negsuc m) = + -Dist+ (negsuc m) (pos (suc n) · negsuc m) + ∙ cong (pos (suc m) +_) (-DistL· (pos (suc n)) (negsuc m)) +-DistL· (negsuc zero) c = -Involutive c +-DistL· (negsuc (suc n)) c = + -Dist+ (- c) (negsuc n · c) + ∙∙ cong (_+ (- (negsuc n · c))) (-Involutive c) + ∙∙ cong (c +_) (-DistL· (negsuc n) c) + +-DistR· : (b c : ℤ) → - (b · c) ≡ b · - c +-DistR· b c = cong (-_) (·Comm b c) ∙∙ -DistL· c b ∙∙ ·Comm (- c) b + +·Assoc : (a b c : ℤ) → (a · (b · c)) ≡ ((a · b) · c) +·Assoc (pos zero) b c = refl +·Assoc (pos (suc n)) b c = + cong ((b · c) +_) (·Assoc (pos n) b c) + ∙∙ cong₂ _+_ (·Comm b c) (·Comm (pos n · b) c) + ∙∙ sym (·DistR+ c b (pos n · b)) + ∙ sym (·Comm _ c) +·Assoc (negsuc zero) = -DistL· +·Assoc (negsuc (suc n)) b c = + cong ((- (b · c)) +_) (·Assoc (negsuc n) b c) + ∙∙ cong (_+ ((negsuc n · b) · c)) (-DistL· b c) + ∙∙ sym (·DistL+ (- b) (negsuc n · b) c) diff --git a/Cubical/Data/List.agda b/Cubical/Data/List.agda index d19502daf..ec82b6a1b 100644 --- a/Cubical/Data/List.agda +++ b/Cubical/Data/List.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.List where open import Cubical.Data.List.Base public diff --git a/Cubical/Data/List/Base.agda b/Cubical/Data/List/Base.agda index 287224e99..022226595 100644 --- a/Cubical/Data/List/Base.agda +++ b/Cubical/Data/List/Base.agda @@ -1,13 +1,15 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.List.Base where -open import Agda.Builtin.List public +open import Agda.Builtin.List public open import Cubical.Core.Everything -open import Cubical.Data.Nat +open import Cubical.Data.Maybe.Base as Maybe +open import Cubical.Data.Nat.Base module _ {ℓ} {A : Type ℓ} where - infixr 5 _++_ _∷ʳ_ + infixr 5 _++_ + infixl 5 _∷ʳ_ [_] : A → List A [ a ] = a ∷ [] @@ -26,3 +28,25 @@ module _ {ℓ} {A : Type ℓ} where length : List A → ℕ length [] = 0 length (x ∷ l) = 1 + length l + + map : ∀ {ℓ'} {B : Type ℓ'} → (A → B) → List A → List B + map f [] = [] + map f (x ∷ xs) = f x ∷ map f xs + + map2 : ∀ {ℓ' ℓ''} {B : Type ℓ'} {C : Type ℓ''} + → (A → B → C) → List A → List B → List C + map2 f [] _ = [] + map2 f _ [] = [] + map2 f (x ∷ xs) (y ∷ ys) = f x y ∷ map2 f xs ys + + filterMap : ∀ {ℓ'} {B : Type ℓ'} → (A → Maybe B) → List A → List B + filterMap f [] = [] + filterMap f (x ∷ xs) = Maybe.rec (filterMap f xs) (_∷ filterMap f xs) (f x) + + foldr : ∀ {ℓ'} {B : Type ℓ'} → (A → B → B) → B → List A → B + foldr f b [] = b + foldr f b (x ∷ xs) = f x (foldr f b xs) + + foldl : ∀ {ℓ'} {B : Type ℓ'} → (B → A → B) → B → List A → B + foldl f b [] = b + foldl f b (x ∷ xs) = foldl f (f b x) xs diff --git a/Cubical/Data/List/Properties.agda b/Cubical/Data/List/Properties.agda index c0a02b0be..b7278a38c 100644 --- a/Cubical/Data/List/Properties.agda +++ b/Cubical/Data/List/Properties.agda @@ -1,13 +1,14 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.List.Properties where open import Agda.Builtin.List open import Cubical.Core.Everything +open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.HLevels open import Cubical.Foundations.Prelude -open import Cubical.Data.Empty +open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Nat -open import Cubical.Data.Prod +open import Cubical.Data.Sigma open import Cubical.Data.Unit open import Cubical.Relation.Nullary @@ -23,6 +24,10 @@ module _ {ℓ} {A : Type ℓ} where ++-assoc [] ys zs = refl ++-assoc (x ∷ xs) ys zs = cong (_∷_ x) (++-assoc xs ys zs) + rev-snoc : (xs : List A) (y : A) → rev (xs ++ [ y ]) ≡ y ∷ rev xs + rev-snoc [] y = refl + rev-snoc (x ∷ xs) y = cong (_++ [ x ]) (rev-snoc xs y) + rev-++ : (xs ys : List A) → rev (xs ++ ys) ≡ rev ys ++ rev xs rev-++ [] ys = sym (++-unit-r (rev ys)) rev-++ (x ∷ xs) ys = @@ -31,7 +36,30 @@ module _ {ℓ} {A : Type ℓ} where rev-rev : (xs : List A) → rev (rev xs) ≡ xs rev-rev [] = refl - rev-rev (x ∷ xs) = rev-++ (rev xs) [ x ] ∙ cong (_∷_ x) (rev-rev xs) + rev-rev (x ∷ xs) = rev-snoc (rev xs) x ∙ cong (_∷_ x) (rev-rev xs) + + rev-rev-snoc : (xs : List A) (y : A) → + Square (rev-rev (xs ++ [ y ])) (cong (_++ [ y ]) (rev-rev xs)) (cong rev (rev-snoc xs y)) refl + rev-rev-snoc [] y = sym (lUnit refl) + rev-rev-snoc (x ∷ xs) y i j = + hcomp + (λ k → λ + { (i = i1) → compPath-filler (rev-snoc (rev xs) x) (cong (x ∷_) (rev-rev xs)) k j ++ [ y ] + ; (j = i0) → rev (rev-snoc xs y i ++ [ x ]) + ; (j = i1) → x ∷ rev-rev-snoc xs y i k + }) + (rev-snoc (rev-snoc xs y i) x j) + + data SnocView : List A → Type ℓ where + nil : SnocView [] + snoc : (x : A) → (xs : List A) → (sx : SnocView xs) → SnocView (xs ∷ʳ x) + + snocView : (xs : List A) → SnocView xs + snocView xs = helper nil xs + where + helper : {l : List A} -> SnocView l -> (r : List A) -> SnocView (l ++ r) + helper {l} sl [] = subst SnocView (sym (++-unit-r l)) sl + helper {l} sl (x ∷ r) = subst SnocView (++-assoc l (x ∷ []) r) (helper (snoc x l sl) r) -- Path space of list type module ListPath {ℓ} {A : Type ℓ} where @@ -67,7 +95,7 @@ module ListPath {ℓ} {A : Type ℓ} where J (λ ys p → decode xs ys (encode xs ys p) ≡ p) (cong (decode xs xs) (encodeRefl xs) ∙ decodeRefl xs) - isOfHLevelCover : (n : ℕ) (p : isOfHLevel (suc (suc n)) A) + isOfHLevelCover : (n : HLevel) (p : isOfHLevel (suc (suc n)) A) (xs ys : List A) → isOfHLevel (suc n) (Cover xs ys) isOfHLevelCover n p [] [] = isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isPropUnit) @@ -76,9 +104,9 @@ module ListPath {ℓ} {A : Type ℓ} where isOfHLevelCover n p (x ∷ xs) [] = isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) isOfHLevelCover n p (x ∷ xs) (y ∷ ys) = - hLevelProd (suc n) (p x y) (isOfHLevelCover n p xs ys) + isOfHLevelΣ (suc n) (p x y) (\ _ → isOfHLevelCover n p xs ys) -isOfHLevelList : ∀ {ℓ} (n : ℕ) {A : Type ℓ} +isOfHLevelList : ∀ {ℓ} (n : HLevel) {A : Type ℓ} → isOfHLevel (suc (suc n)) A → isOfHLevel (suc (suc n)) (List A) isOfHLevelList n ofLevel xs ys = isOfHLevelRetract (suc n) @@ -127,12 +155,12 @@ cons≡rev-snoc : (x : A) → (xs : List A) → x ∷ rev xs ≡ rev (xs ∷ʳ x cons≡rev-snoc _ [] = refl cons≡rev-snoc x (y ∷ ys) = λ i → cons≡rev-snoc x ys i ++ y ∷ [] -nil≡nil-isContr : isContr (Path (List A) [] []) -nil≡nil-isContr = refl , ListPath.decodeEncode [] [] +isContr[]≡[] : isContr (Path (List A) [] []) +isContr[]≡[] = refl , ListPath.decodeEncode [] [] -list≡nil-isProp : {xs : List A} → isProp (xs ≡ []) -list≡nil-isProp {xs = []} = isOfHLevelSuc 0 nil≡nil-isContr -list≡nil-isProp {xs = x ∷ xs} = λ p _ → ⊥-elim (¬cons≡nil p) +isPropXs≡[] : {xs : List A} → isProp (xs ≡ []) +isPropXs≡[] {xs = []} = isOfHLevelSuc 0 isContr[]≡[] +isPropXs≡[] {xs = x ∷ xs} = λ p _ → ⊥.rec (¬cons≡nil p) discreteList : Discrete A → Discrete (List A) discreteList eqA [] [] = yes refl @@ -142,3 +170,7 @@ discreteList eqA (x ∷ xs) (y ∷ ys) with eqA x y | discreteList eqA xs ys ... | yes p | yes q = yes (λ i → p i ∷ q i) ... | yes _ | no ¬q = no (λ p → ¬q (cons-inj₂ p)) ... | no ¬p | _ = no (λ q → ¬p (cons-inj₁ q)) + +foldrCons : (xs : List A) → foldr _∷_ [] xs ≡ xs +foldrCons [] = refl +foldrCons (x ∷ xs) = cong (x ∷_) (foldrCons xs) diff --git a/Cubical/Data/Maybe.agda b/Cubical/Data/Maybe.agda index d9a7c7e67..bd40e1d2c 100644 --- a/Cubical/Data/Maybe.agda +++ b/Cubical/Data/Maybe.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.Maybe where open import Cubical.Data.Maybe.Base public diff --git a/Cubical/Data/Maybe/Base.agda b/Cubical/Data/Maybe/Base.agda index 39e5a155d..0f230b623 100644 --- a/Cubical/Data/Maybe/Base.agda +++ b/Cubical/Data/Maybe/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.Maybe.Base where open import Cubical.Core.Everything @@ -19,3 +19,7 @@ caseMaybe _ j (just _) = j map-Maybe : (A → B) → Maybe A → Maybe B map-Maybe _ nothing = nothing map-Maybe f (just x) = just (f x) + +rec : B → (A → B) → Maybe A → B +rec n j nothing = n +rec n j (just a) = j a diff --git a/Cubical/Data/Maybe/Properties.agda b/Cubical/Data/Maybe/Properties.agda index ed648cbba..23da4995c 100644 --- a/Cubical/Data/Maybe/Properties.agda +++ b/Cubical/Data/Maybe/Properties.agda @@ -1,13 +1,14 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.Maybe.Properties where open import Cubical.Core.Everything open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Embedding -open import Cubical.Data.Empty +open import Cubical.Functions.Embedding +open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Unit open import Cubical.Data.Nat open import Cubical.Relation.Nullary @@ -15,6 +16,10 @@ open import Cubical.Data.Sum open import Cubical.Data.Maybe.Base +map-Maybe-id : ∀ {ℓ} {A : Type ℓ} → ∀ m → map-Maybe (idfun A) m ≡ m +map-Maybe-id nothing = refl +map-Maybe-id (just _) = refl + -- Path space of Maybe type module MaybePath {ℓ} {A : Type ℓ} where Cover : Maybe A → Maybe A → Type ℓ @@ -59,7 +64,7 @@ module MaybePath {ℓ} {A : Type ℓ} where Cover≡Path c c' = isoToPath (iso (decode c c') (encode c c') (decodeEncode c c') (encodeDecode c c')) - isOfHLevelCover : (n : ℕ) + isOfHLevelCover : (n : HLevel) → isOfHLevel (suc (suc n)) A → ∀ c c' → isOfHLevel (suc n) (Cover c c') isOfHLevelCover n p nothing nothing = isOfHLevelLift (suc n) (isOfHLevelUnit (suc n)) @@ -67,7 +72,7 @@ module MaybePath {ℓ} {A : Type ℓ} where isOfHLevelCover n p (just a) nothing = isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) isOfHLevelCover n p (just a) (just a') = p a a' -isOfHLevelMaybe : ∀ {ℓ} (n : ℕ) {A : Type ℓ} +isOfHLevelMaybe : ∀ {ℓ} (n : HLevel) {A : Type ℓ} → isOfHLevel (suc (suc n)) A → isOfHLevel (suc (suc n)) (Maybe A) isOfHLevelMaybe n lA c c' = @@ -99,8 +104,14 @@ isEmbedding-just w z = MaybePath.Cover≃Path (just w) (just z) .snd ¬just≡nothing {A = A} {x = x} p = lower (subst (caseMaybe (Lift ⊥) (Maybe A)) p (just x)) isProp-x≡nothing : (x : Maybe A) → isProp (x ≡ nothing) -isProp-x≡nothing nothing x w = subst isProp (MaybePath.Cover≡Path nothing nothing) (isOfHLevelLift 1 isPropUnit) x w -isProp-x≡nothing (just _) p _ = ⊥-elim (¬just≡nothing p) +isProp-x≡nothing nothing x w = + subst isProp (MaybePath.Cover≡Path nothing nothing) (isOfHLevelLift 1 isPropUnit) x w +isProp-x≡nothing (just _) p _ = ⊥.rec (¬just≡nothing p) + +isProp-nothing≡x : (x : Maybe A) → isProp (nothing ≡ x) +isProp-nothing≡x nothing x w = + subst isProp (MaybePath.Cover≡Path nothing nothing) (isOfHLevelLift 1 isPropUnit) x w +isProp-nothing≡x (just _) p _ = ⊥.rec (¬nothing≡just p) isContr-nothing≡nothing : isContr (nothing {A = A} ≡ nothing) isContr-nothing≡nothing = inhProp→isContr refl (isProp-x≡nothing _) @@ -132,3 +143,16 @@ module SumUnit where Maybe≡SumUnit : Maybe A ≡ Unit ⊎ A Maybe≡SumUnit = isoToPath (iso SumUnit.Maybe→SumUnit SumUnit.SumUnit→Maybe SumUnit.SumUnit→Maybe→SumUnit SumUnit.Maybe→SumUnit→Maybe) + +congMaybeEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → A ≃ B → Maybe A ≃ Maybe B +congMaybeEquiv e = isoToEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun = map-Maybe (equivFun e) + isom .inv = map-Maybe (invEq e) + isom .rightInv nothing = refl + isom .rightInv (just b) = cong just (secEq e b) + isom .leftInv nothing = refl + isom .leftInv (just a) = cong just (retEq e a) diff --git a/Cubical/Data/Nat.agda b/Cubical/Data/Nat.agda index a6774b6a4..b00c8ec4f 100644 --- a/Cubical/Data/Nat.agda +++ b/Cubical/Data/Nat.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Data.Nat where open import Cubical.Data.Nat.Base public diff --git a/Cubical/Data/Nat/Algebra.agda b/Cubical/Data/Nat/Algebra.agda index b9903d0db..0ee04806b 100644 --- a/Cubical/Data/Nat/Algebra.agda +++ b/Cubical/Data/Nat/Algebra.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --no-exact-split --safe #-} +{-# OPTIONS --no-exact-split --safe #-} {- @@ -25,6 +25,8 @@ open import Cubical.Foundations.Path open import Cubical.Foundations.Isomorphism hiding (section) open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Reflection.StrictEquiv open import Cubical.Data.Nat.Base @@ -69,7 +71,7 @@ isNatInductive N ℓ = (S : NatFiber N ℓ) → NatSection S module AlgebraPropositionality {N : NatAlgebra ℓ'} where open NatAlgebra N isPropIsNatHInitial : isProp (isNatHInitial N ℓ) - isPropIsNatHInitial = isPropPi (λ _ → isPropIsContr) + isPropIsNatHInitial = isPropΠ (λ _ → isPropIsContr) -- under the assumption that some shape is nat-inductive, the type of sections over any fiber -- is propositional @@ -112,8 +114,11 @@ module AlgebraHInd→HInit {N : NatAlgebra ℓ'} (ind : isNatInductive N ℓ) (M section→morph : NatSection ConstFiberM → NatMorphism N M section→morph x = record { morph = section ; comm-zero = sec-comm-zero ; comm-suc = λ n i → sec-comm-suc i n } where open NatSection x + Morph≡Section : NatSection ConstFiberM ≡ NatMorphism N M - Morph≡Section = isoToPath (iso section→morph morph→section (λ _ → refl) (λ _ → refl)) + Morph≡Section = ua e + where + unquoteDecl e = declStrictEquiv e section→morph morph→section isContrMorph : isContr (NatMorphism N M) isContrMorph = subst isContr Morph≡Section (inhProp→isContr (ind ConstFiberM) (AlgebraPropositionality.SectionProp.S≡T ind)) @@ -181,7 +186,7 @@ module AlgebraHInit→Ind (N : NatAlgebra ℓ') ℓ (hinit : isNatHInitial N ( Q-zero : α (N .alg-zero) ≡ N .alg-zero Q-zero = ζ Q-suc : ∀ n → α (N .alg-suc n) ≡ N .alg-suc n - Q-suc n = σ n □ cong (N .alg-suc) (P n) + Q-suc n = σ n ∙ cong (N .alg-suc) (P n) -- but P and Q are the same up to homotopy P-zero : P (N .alg-zero) ≡ Q-zero @@ -194,7 +199,7 @@ module AlgebraHInit→Ind (N : NatAlgebra ℓ') ℓ (hinit : isNatHInitial N ( P-suc : ∀ n → P (N .alg-suc n) ≡ Q-suc n P-suc n i j = hcomp (λ k → λ where (i = i0) → lower (fst∘μ≡id j .comm-suc (~ k) n) - (i = i1) → compPath'-filler (σ n) (cong (N .alg-suc) (P n)) k j + (i = i1) → compPath-filler' (σ n) (cong (N .alg-suc) (P n)) k j (j = i0) → σ n (~ k) (j = i1) → N .alg-suc n ) (N .alg-suc (P n j)) @@ -212,7 +217,7 @@ module AlgebraHInit→Ind (N : NatAlgebra ℓ') ℓ (hinit : isNatHInitial N ( P (N .alg-suc n) ! α-h (N .alg-suc n) ≡[ i ]⟨ P-suc n i ! α-h _ ⟩ Q-suc n ! α-h (N .alg-suc n) - ≡⟨ substComposite-□ (F .Fiber) (σ n) (cong (N .alg-suc) (P n)) _ ⟩ + ≡⟨ substComposite (F .Fiber) (σ n) (cong (N .alg-suc) (P n)) _ ⟩ cong (N .alg-suc) (P n) ! (σ n ! α-h (N .alg-suc n)) ≡[ i ]⟨ cong (N .alg-suc) (P n) ! fromPathP (σ-h n) i ⟩ cong (N .alg-suc) (P n) ! (F .fib-suc (α-h n)) @@ -222,9 +227,7 @@ module AlgebraHInit→Ind (N : NatAlgebra ℓ') ℓ (hinit : isNatHInitial N ( isNatInductive≡isNatHInitial : {N : NatAlgebra ℓ'} (ℓ : Level) → isNatInductive N (ℓ-max ℓ' ℓ) ≡ isNatHInitial N (ℓ-max ℓ' ℓ) -isNatInductive≡isNatHInitial {ℓ'} {N} ℓ = - isoToPath (equivToIso (PropEquiv→Equiv isPropIsNatInductive isPropIsNatHInitial ind→init init→ind)) where - open import Cubical.Foundations.Equiv +isNatInductive≡isNatHInitial {_} {N} ℓ = hPropExt isPropIsNatInductive isPropIsNatHInitial ind→init init→ind where open AlgebraPropositionality open AlgebraHInit→Ind N ℓ renaming (Fsection to init→ind) open AlgebraHInd→HInit renaming (isContrMorph to ind→init) diff --git a/Cubical/Data/Nat/Base.agda b/Cubical/Data/Nat/Base.agda index cfc4404dc..798728362 100644 --- a/Cubical/Data/Nat/Base.agda +++ b/Cubical/Data/Nat/Base.agda @@ -1,40 +1,33 @@ -{-# OPTIONS --cubical --no-exact-split --safe #-} +{-# OPTIONS --no-exact-split --safe #-} module Cubical.Data.Nat.Base where open import Cubical.Core.Primitives open import Agda.Builtin.Nat public - using (zero; suc; _+_; _*_) - renaming (Nat to ℕ) - --- Allows for (constrained) natural number and negative integer --- literals for any type (e.g. ℕ, ℕ₋₁, ℕ₋₂, Int) -open import Agda.Builtin.FromNat public - renaming (Number to HasFromNat) -open import Agda.Builtin.FromNeg public - renaming (Negative to HasFromNeg) -open import Cubical.Data.Unit.Base public - --- Natural number literals for ℕ -instance - fromNatℕ : HasFromNat ℕ - fromNatℕ = record { Constraint = λ _ → Unit ; fromNat = λ n → n } + using (zero; suc; _+_) + renaming (Nat to ℕ; _-_ to _∸_; _*_ to _·_) + +open import Cubical.Data.Nat.Literals public +open import Cubical.Data.Bool.Base +open import Cubical.Data.Sum.Base hiding (elim) +open import Cubical.Data.Empty.Base hiding (elim) +open import Cubical.Data.Unit.Base predℕ : ℕ → ℕ -predℕ zero = 0 +predℕ zero = zero predℕ (suc n) = n caseNat : ∀ {ℓ} → {A : Type ℓ} → (a0 aS : A) → ℕ → A -caseNat a0 aS 0 = a0 +caseNat a0 aS zero = a0 caseNat a0 aS (suc n) = aS doubleℕ : ℕ → ℕ -doubleℕ 0 = 0 +doubleℕ zero = zero doubleℕ (suc x) = suc (suc (doubleℕ x)) --- doublesℕ n m = 2^n * m +-- doublesℕ n m = 2^n · m doublesℕ : ℕ → ℕ → ℕ -doublesℕ 0 m = m +doublesℕ zero m = m doublesℕ (suc n) m = doublesℕ n (doubleℕ m) -- iterate @@ -42,9 +35,27 @@ iter : ∀ {ℓ} {A : Type ℓ} → ℕ → (A → A) → A → A iter zero f z = z iter (suc n) f z = f (iter n f z) -ℕ-induction : ∀ {ℓ} {A : ℕ → Type ℓ} - → A 0 - → ((n : ℕ) → A n → A (suc n)) - → (n : ℕ) → A n -ℕ-induction a₀ _ zero = a₀ -ℕ-induction a₀ f (suc n) = f n ((ℕ-induction a₀ f n)) +elim : ∀ {ℓ} {A : ℕ → Type ℓ} + → A zero + → ((n : ℕ) → A n → A (suc n)) + → (n : ℕ) → A n +elim a₀ _ zero = a₀ +elim a₀ f (suc n) = f n (elim a₀ f n) + +isEven isOdd : ℕ → Bool +isEven zero = true +isEven (suc n) = isOdd n +isOdd zero = false +isOdd (suc n) = isEven n + +--Typed version +private + toType : Bool → Type + toType false = ⊥ + toType true = Unit + +isEvenT : ℕ → Type +isEvenT n = toType (isEven n) + +isOddT : ℕ → Type +isOddT n = isEvenT (suc n) diff --git a/Cubical/Data/Nat/Coprime.agda b/Cubical/Data/Nat/Coprime.agda new file mode 100644 index 000000000..d551ae28c --- /dev/null +++ b/Cubical/Data/Nat/Coprime.agda @@ -0,0 +1,95 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Nat.Coprime where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function + +open import Cubical.Data.Sigma +open import Cubical.Data.NatPlusOne + +open import Cubical.HITs.PropositionalTruncation as PropTrunc + +open import Cubical.Data.Nat.Base +open import Cubical.Data.Nat.Properties +open import Cubical.Data.Nat.Divisibility +open import Cubical.Data.Nat.GCD + +areCoprime : ℕ × ℕ → Type₀ +areCoprime (m , n) = isGCD m n 1 + +-- Any pair (m , n) can be converted to a coprime pair (m' , n') s.t. +-- m' ∣ m, n' ∣ n if and only if one of m or n is nonzero + +module ToCoprime ((m , n) : ℕ × ℕ₊₁) where + d = gcd m (ℕ₊₁→ℕ n) + d∣m = gcdIsGCD m (ℕ₊₁→ℕ n) .fst .fst + d∣n = gcdIsGCD m (ℕ₊₁→ℕ n) .fst .snd + gr = gcdIsGCD m (ℕ₊₁→ℕ n) .snd + + c₁ : ℕ + p₁ : c₁ · d ≡ m + c₁ = ∣-untrunc d∣m .fst; p₁ = ∣-untrunc d∣m .snd + + c₂ : ℕ₊₁ + p₂ : (ℕ₊₁→ℕ c₂) · d ≡ (ℕ₊₁→ℕ n) + c₂ = 1+ (∣s-untrunc d∣n .fst); p₂ = ∣s-untrunc d∣n .snd + + toCoprime : ℕ × ℕ₊₁ + toCoprime = (c₁ , c₂) + + private + lem : ∀ a {b c d e} → a · b ≡ c → c · d ≡ e → a · (b · d) ≡ e + lem a p q = ·-assoc a _ _ ∙ cong (_· _) p ∙ q + + gr' : ∀ d' → prediv d' c₁ → prediv d' (ℕ₊₁→ℕ c₂) → (d' · d) ∣ d + gr' d' (b₁ , q₁) (b₂ , q₂) = gr (d' · d) ((∣ b₁ , lem b₁ q₁ p₁ ∣) , + (∣ b₂ , lem b₂ q₂ p₂ ∣)) + + d-1 = m∣sn→z 0 (<=> m > 0 or n > 0) + d-cancelʳ : ∀ d' → (d' · d) ∣ d → d' ∣ 1 + d-cancelʳ d' div = ∣-cancelʳ d-1 (∣-trans (subst (λ x → (d' · x) ∣ x) q div) + (∣-refl (sym (·-identityˡ _)))) + + toCoprimeAreCoprime : areCoprime (c₁ , ℕ₊₁→ℕ c₂) + fst toCoprimeAreCoprime = ∣-oneˡ c₁ , ∣-oneˡ (ℕ₊₁→ℕ c₂) + snd toCoprimeAreCoprime d' (d'∣c₁ , d'∣c₂) = PropTrunc.rec isProp∣ (λ a → + PropTrunc.rec isProp∣ (λ b → + d-cancelʳ d' (gr' d' a b)) d'∣c₂) d'∣c₁ + + toCoprime∣ : (c₁ ∣ m) × (ℕ₊₁→ℕ c₂ ∣ ℕ₊₁→ℕ n) + toCoprime∣ = ∣ d , ·-comm d c₁ ∙ p₁ ∣ , ∣ d , ·-comm d (ℕ₊₁→ℕ c₂) ∙ p₂ ∣ + + toCoprime-idem : areCoprime (m , ℕ₊₁→ℕ n) → (c₁ , c₂) ≡ (m , n) + toCoprime-idem cp i = q₁ i , ℕ₊₁→ℕ-inj q₂ i + where q₁ = sym (·-identityʳ c₁) ∙ cong (c₁ ·_) (sym (isGCD→gcd≡ cp)) ∙ p₁ + q₂ = sym (·-identityʳ (ℕ₊₁→ℕ c₂)) ∙ cong (ℕ₊₁→ℕ c₂ ·_) (sym (isGCD→gcd≡ cp)) ∙ p₂ + +open ToCoprime using (toCoprime; toCoprimeAreCoprime; toCoprime∣; toCoprime-idem) public + + +toCoprime-cancelʳ : ∀ ((m , n) : ℕ × ℕ₊₁) k + → toCoprime (m · ℕ₊₁→ℕ k , n ·₊₁ k) ≡ toCoprime (m , n) +toCoprime-cancelʳ (m , n) (1+ k) i = + inj-·sm {c₁'} {d-1} {c₁} r₁ i , ℕ₊₁→ℕ-inj (inj-·sm {ℕ₊₁→ℕ c₂'} {d-1} {ℕ₊₁→ℕ c₂} r₂) i + where open ToCoprime (m , n) + open ToCoprime (m · suc k , n ·₊₁ (1+ k)) using () + renaming (c₁ to c₁'; p₁ to p₁'; c₂ to c₂'; p₂ to p₂') + + q₁ : c₁' · d · suc k ≡ m · suc k + q₁ = sym (·-assoc c₁' (ToCoprime.d (m , n)) (suc k)) + ∙ cong (c₁' ·_) (sym (gcd-factorʳ m (ℕ₊₁→ℕ n) (suc k))) + ∙ p₁' + q₂ : ℕ₊₁→ℕ c₂' · (ToCoprime.d (m , n)) · suc k ≡ ℕ₊₁→ℕ n · suc k + q₂ = sym (·-assoc (ℕ₊₁→ℕ c₂') (ToCoprime.d (m , n)) (suc k)) + ∙ cong (ℕ₊₁→ℕ c₂' ·_) (sym (gcd-factorʳ m (ℕ₊₁→ℕ n) (suc k))) + ∙ p₂' + + r₁ : c₁' · suc d-1 ≡ c₁ · suc d-1 + r₁ = subst (λ z → c₁' · z ≡ c₁ · z) q (inj-·sm q₁ ∙ sym p₁) + r₂ : ℕ₊₁→ℕ c₂' · suc d-1 ≡ ℕ₊₁→ℕ c₂ · suc d-1 + r₂ = subst (λ z → ℕ₊₁→ℕ c₂' · z ≡ ℕ₊₁→ℕ c₂ · z) q (inj-·sm q₂ ∙ sym p₂) diff --git a/Cubical/Data/Nat/Divisibility.agda b/Cubical/Data/Nat/Divisibility.agda new file mode 100644 index 000000000..27354cf55 --- /dev/null +++ b/Cubical/Data/Nat/Divisibility.agda @@ -0,0 +1,104 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Nat.Divisibility where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ + +open import Cubical.HITs.PropositionalTruncation as PropTrunc + +open import Cubical.Data.Nat.Base +open import Cubical.Data.Nat.Properties +open import Cubical.Data.Nat.Order + +private + variable + l m n : ℕ + +_∣_ : ℕ → ℕ → Type₀ +m ∣ n = ∃[ c ∈ ℕ ] c · m ≡ n + +isProp∣ : isProp (m ∣ n) +isProp∣ = squash + +prediv : ℕ → ℕ → Type₀ +prediv m n = Σ[ c ∈ ℕ ] c · m ≡ n + +-- an alternate definition of m ∣ n that doesn't use truncation + +_∣'_ : ℕ → ℕ → Type₀ +zero ∣' n = zero ≡ n +suc m ∣' n = Σ[ c ∈ ℕ ] c · suc m ≡ n + +isProp∣' : isProp (m ∣' n) +isProp∣' {zero} {n} = isSetℕ _ _ +isProp∣' {suc m} {n} (c₁ , p₁) (c₂ , p₂) = + Σ≡Prop (λ _ → isSetℕ _ _) (inj-·sm {c₁} {m} {c₂} (p₁ ∙ sym p₂)) + +∣≃∣' : (m ∣ n) ≃ (m ∣' n) +∣≃∣' {zero} = propBiimpl→Equiv isProp∣ isProp∣' + (PropTrunc.rec (isSetℕ _ _) λ { (c , p) → 0≡m·0 c ∙ p }) + (λ p → ∣ zero , p ∣) +∣≃∣' {suc m} = propTruncIdempotent≃ isProp∣' + +∣-untrunc : m ∣ n → Σ[ c ∈ ℕ ] c · m ≡ n +∣-untrunc {zero} p = zero , equivFun ∣≃∣' p +∣-untrunc {suc m} = equivFun ∣≃∣' + + +-- basic properties of ∣ + +∣-refl : m ≡ n → m ∣ n +∣-refl p = ∣ 1 , +-zero _ ∙ p ∣ + +∣-trans : l ∣ m → m ∣ n → l ∣ n +∣-trans = PropTrunc.map2 λ { + (c₁ , p) (c₂ , q) → (c₂ · c₁ , sym (·-assoc c₂ c₁ _) ∙ cong (c₂ ·_) p ∙ q) } + +∣-left : ∀ k → m ∣ (m · k) +∣-left k = ∣ k , ·-comm k _ ∣ + +∣-right : ∀ k → m ∣ (k · m) +∣-right k = ∣ k , refl ∣ + +∣-cancelʳ : ∀ k → (m · suc k) ∣ (n · suc k) → m ∣ n +∣-cancelʳ k = PropTrunc.map λ { + (c , p) → (c , inj-·sm (sym (·-assoc c _ (suc k)) ∙ p)) } + +∣-multʳ : ∀ k → m ∣ n → (m · k) ∣ (n · k) +∣-multʳ k = PropTrunc.map λ { + (c , p) → (c , ·-assoc c _ k ∙ cong (_· k) p) } + +∣-zeroˡ : zero ∣ m → zero ≡ m +∣-zeroˡ = equivFun ∣≃∣' + +∣-zeroʳ : ∀ m → m ∣ zero +∣-zeroʳ m = ∣ zero , refl ∣ + +∣-oneˡ : ∀ m → 1 ∣ m +∣-oneˡ m = ∣ m , ·-identityʳ m ∣ + +-- if n > 0, then the constant c (s.t. c · m ≡ n) is also > 0 +∣s-untrunc : m ∣ suc n → Σ[ c ∈ ℕ ] (suc c) · m ≡ suc n +∣s-untrunc ∣p∣ with ∣-untrunc ∣p∣ +... | (zero , p) = ⊥.rec (znots p) +... | (suc c , p) = (c , p) + +m∣sn→m≤sn : m ∣ suc n → m ≤ suc n +m∣sn→m≤sn {m} {n} = f ∘ ∣s-untrunc + where f : Σ[ c ∈ ℕ ] (suc c) · m ≡ suc n → Σ[ c ∈ ℕ ] c + m ≡ suc n + f (c , p) = (c · m) , (+-comm (c · m) m ∙ p) + +m∣sn→z trySigma <|> tryParam <|> tryPi <|> tryUnit <|> useGeneric + where + tryUniv : R.TC Unit + tryUniv = R.unify (R.con (quote UARelDesc.univ) [ varg R.unknown ]) hole + + tryBinary : R.Name → R.TC Unit + tryBinary name = + newMeta R.unknown >>= λ hole₁ → + newMeta R.unknown >>= λ hole₂ → + R.unify (R.con name (hole₁ v∷ hole₂ v∷ [])) hole >> + autoUARelDesc n hole₁ >> + autoDUARelDesc n hole₂ + + tryParam : R.TC Unit + tryParam = + newMeta R.unknown >>= λ paramTy → + newMeta R.unknown >>= λ hole₁ → + R.unify (R.con (quote UARelDesc.param) (paramTy v∷ hole₁ v∷ [])) hole >> + autoUARelDesc n hole₁ + + trySigma = tryBinary (quote UARelDesc.sigma) + tryPi = tryBinary (quote UARelDesc.pi) + + tryUnit : R.TC Unit + tryUnit = R.unify (R.con (quote UARelDesc.unit) []) hole + + useGeneric : R.TC Unit + useGeneric = R.unify (R.con (quote UARelDesc.generic) []) hole + + autoUARelReindex : ℕ → R.Term → R.TC Unit + autoUARelReindex zero hole = R.typeError [ R.strErr "Out of fuel" ] + autoUARelReindex (suc n) hole = + tryId <|> tryFst <|> trySnd <|> tryApp + where + tryId : R.TC Unit + tryId = R.unify (R.con (quote UARelReindex.id) []) hole + + tryUnary : R.Name → R.TC Unit + tryUnary name = + newMeta R.unknown >>= λ hole₁ → + R.unify (R.con name [ varg hole₁ ]) hole >> + autoUARelReindex n hole₁ + + tryFst = tryUnary (quote UARelReindex.∘fst) + trySnd = tryUnary (quote UARelReindex.∘snd) + + tryApp : R.TC Unit + tryApp = + newMeta R.unknown >>= λ hole₁ → + newMeta R.unknown >>= λ param → + R.unify (R.con (quote UARelReindex.∘app) (hole₁ v∷ param v∷ [])) hole >> + autoUARelReindex n hole₁ + + autoSubstRelDesc : ℕ → R.Term → R.TC Unit + autoSubstRelDesc zero hole = R.typeError [ R.strErr "Out of fuel" ] + autoSubstRelDesc (suc n) hole = + tryConstant <|> tryEl <|> tryEl <|> trySigma <|> tryPi <|> useGeneric + where + tryConstant : R.TC Unit + tryConstant = + R.unify (R.con (quote SubstRelDesc.constant) []) hole + + tryEl : R.TC Unit + tryEl = + newMeta R.unknown >>= λ hole₁ → + R.unify (R.con (quote SubstRelDesc.el) [ varg hole₁ ]) hole >> + autoUARelReindex n hole₁ + + tryBinary : R.Name → R.TC Unit + tryBinary name = + newMeta R.unknown >>= λ hole₁ → + newMeta R.unknown >>= λ hole₂ → + R.unify (R.con name (hole₁ v∷ hole₂ v∷ [])) hole >> + autoSubstRelDesc n hole₁ >> + autoSubstRelDesc n hole₂ + + trySigma = tryBinary (quote SubstRelDesc.sigma) + tryPi = tryBinary (quote SubstRelDesc.pi) + + useGeneric : R.TC Unit + useGeneric = R.unify (R.con (quote SubstRelDesc.generic) []) hole + + autoDUARelDesc : ℕ → R.Term → R.TC Unit + autoDUARelDesc zero hole = R.typeError [ R.strErr "Out of fuel" ] + autoDUARelDesc (suc n) hole = + tryConstant <|> tryEl <|> trySigma <|> tryPiˢ <|> tryPi <|> useGeneric + where + tryConstant : R.TC Unit + tryConstant = + newMeta R.unknown >>= λ hole₁ → + R.unify (R.con (quote DUARelDesc.constant) [ varg hole₁ ]) hole >> + autoUARelDesc n hole₁ + + tryEl : R.TC Unit + tryEl = + newMeta R.unknown >>= λ hole₁ → + R.unify (R.con (quote DUARelDesc.el) [ varg hole₁ ]) hole >> + autoUARelReindex n hole₁ + + tryBinary : R.Name → R.TC Unit + tryBinary name = + newMeta R.unknown >>= λ hole₁ → + newMeta R.unknown >>= λ hole₂ → + R.unify (R.con name (hole₁ v∷ hole₂ v∷ [])) hole >> + autoDUARelDesc n hole₁ >> + autoDUARelDesc n hole₂ + + trySigma = tryBinary (quote DUARelDesc.sigma) + tryPi = tryBinary (quote DUARelDesc.pi) + + tryPiˢ : R.TC Unit + tryPiˢ = + newMeta R.unknown >>= λ hole₁ → + newMeta R.unknown >>= λ hole₂ → + R.unify (R.con (quote DUARelDesc.piˢ) (hole₁ v∷ hole₂ v∷ [])) hole >> + autoSubstRelDesc n hole₁ >> + autoDUARelDesc n hole₂ + + useGeneric : R.TC Unit + useGeneric = R.unify (R.con (quote DUARelDesc.generic) []) hole + +module DisplayedAutoMacro where + autoUARel : ∀ {ℓA} (A : Type ℓA) → ℕ → R.Term → R.TC Unit + autoUARel A n hole = + R.quoteTC A >>= λ `A` → + newMeta R.unknown >>= λ desc → + makeAuxiliaryDef "autoUA" + (R.def (quote UARel) (`A` v∷ R.unknown v∷ [])) + (R.def (quote getUARel) [ varg desc ]) + >>= λ uaTerm → + R.unify hole uaTerm >> + autoUARelDesc n desc + + autoDUARel : ∀ {ℓA ℓ≅A ℓB} {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) (B : A → Type ℓB) + → ℕ → R.Term → R.TC Unit + autoDUARel 𝒮-A B n hole = + R.quoteTC 𝒮-A >>= λ `𝒮-A` → + R.quoteTC B >>= λ `B` → + newMeta R.unknown >>= λ desc → + makeAuxiliaryDef "autoDUA" + (R.def (quote DUARel) (`𝒮-A` v∷ `B` v∷ R.unknown v∷ [])) + (R.def (quote getDUARel) [ varg desc ]) + >>= λ duaTerm → + R.unify hole duaTerm >> + autoDUARelDesc n desc + +macro + autoUARel : ∀ {ℓA} (A : Type ℓA) → R.Term → R.TC Unit + autoUARel A = DisplayedAutoMacro.autoUARel A FUEL + + autoDUARel : ∀ {ℓA ℓ≅A ℓB} {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) (B : A → Type ℓB) + → R.Term → R.TC Unit + autoDUARel 𝒮-A B = DisplayedAutoMacro.autoDUARel 𝒮-A B FUEL + +private + module Example (A : Type) (a₀ : A) where + + example0 : DUARel (autoUARel Type) (λ X → X → A × X) ℓ-zero + example0 = autoDUARel _ _ + + example0' : {X Y : Type} (e : X ≃ Y) + (f : X → A × X) (g : Y → A × Y) + → (∀ x → (f x .fst ≡ g (e .fst x) .fst) × (e .fst (f x .snd) ≡ g (e .fst x) .snd)) + → PathP (λ i → ua e i → A × ua e i) f g + example0' e f g = example0 .DUARel.uaᴰ f e g .fst + + -- An example where a DUARel is parameterized over a pair of types + + example1 : DUARel (autoUARel (Type × Type)) (λ (X , Z) → X → Z) ℓ-zero + example1 = autoDUARel _ _ + + example1' : {X Y : Type} (e : X ≃ Y) {Z W : Type} (h : Z ≃ W) + (f : X → Z) (g : Y → W) + → (∀ x → h .fst (f x) ≡ g (e .fst x)) + → PathP (λ i → ua e i → ua h i) f g + example1' e h f g = example1 .DUARel.uaᴰ f (e , h) g .fst + + -- An example where a DUARel is parameterized over a family of types + + example2 : DUARel (autoUARel (A → Type)) (λ B → B a₀) ℓ-zero + example2 = autoDUARel _ _ + + example2' : {B C : A → Type} (e : (a : A) → B a ≃ C a) + (b : B a₀) (c : C a₀) + → e a₀ .fst b ≡ c + → PathP (λ i → ua (e a₀) i) b c + example2' e b c = example2 .DUARel.uaᴰ b e c .fst diff --git a/Cubical/Displayed/Base.agda b/Cubical/Displayed/Base.agda new file mode 100644 index 000000000..9491845a7 --- /dev/null +++ b/Cubical/Displayed/Base.agda @@ -0,0 +1,87 @@ +{- + + Definition of univalent and displayed univalent relations + +-} +{-# OPTIONS --safe #-} +module Cubical.Displayed.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Transport + +open import Cubical.Data.Sigma + +open import Cubical.Relation.Binary + +private + variable + ℓ ℓA ℓA' ℓ≅A ℓB ℓB' ℓ≅B ℓC ℓ≅C : Level + +record UARel (A : Type ℓA) (ℓ≅A : Level) : Type (ℓ-max ℓA (ℓ-suc ℓ≅A)) where + no-eta-equality + constructor uarel + field + _≅_ : A → A → Type ℓ≅A + ua : (a a' : A) → (a ≅ a') ≃ (a ≡ a') + + uaIso : (a a' : A) → Iso (a ≅ a') (a ≡ a') + uaIso a a' = equivToIso (ua a a') + + ≅→≡ : {a a' : A} (p : a ≅ a') → a ≡ a' + ≅→≡ {a} {a'} = Iso.fun (uaIso a a') + ≡→≅ : {a a' : A} (p : a ≡ a') → a ≅ a' + ≡→≅ {a} {a'} = Iso.inv (uaIso a a') + + ρ : (a : A) → a ≅ a + ρ a = ≡→≅ refl + +open BinaryRelation + +-- another constructor for UARel using contractibility of relational singletons +make-𝒮 : {A : Type ℓA} {ℓ≅A : Level} {_≅_ : A → A → Type ℓ≅A} + (ρ : isRefl _≅_) (contrTotal : contrRelSingl _≅_) → UARel A ℓ≅A +UARel._≅_ (make-𝒮 {_≅_ = _≅_} _ _) = _≅_ +UARel.ua (make-𝒮 {_≅_ = _≅_} ρ c) = contrRelSingl→isUnivalent _≅_ ρ c + +record DUARel {A : Type ℓA} {ℓ≅A : Level} (𝒮-A : UARel A ℓ≅A) + (B : A → Type ℓB) (ℓ≅B : Level) : Type (ℓ-max (ℓ-max ℓA ℓB) (ℓ-max ℓ≅A (ℓ-suc ℓ≅B))) where + no-eta-equality + constructor duarel + open UARel 𝒮-A + + field + _≅ᴰ⟨_⟩_ : {a a' : A} → B a → a ≅ a' → B a' → Type ℓ≅B + uaᴰ : {a a' : A} (b : B a) (p : a ≅ a') (b' : B a') → (b ≅ᴰ⟨ p ⟩ b') ≃ PathP (λ i → B (≅→≡ p i)) b b' + + fiberRel : (a : A) → Rel (B a) (B a) ℓ≅B + fiberRel a = _≅ᴰ⟨ ρ a ⟩_ + + uaᴰρ : {a : A} (b b' : B a) → b ≅ᴰ⟨ ρ a ⟩ b' ≃ (b ≡ b') + uaᴰρ {a} b b' = + compEquiv + (uaᴰ b (ρ _) b') + (substEquiv (λ q → PathP (λ i → B (q i)) b b') (secEq (ua a a) refl)) + + ρᴰ : {a : A} → (b : B a) → b ≅ᴰ⟨ ρ a ⟩ b + ρᴰ {a} b = invEq (uaᴰρ b b) refl + + +-- total UARel induced by a DUARel + +module _ {A : Type ℓA} {ℓ≅A : Level} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} {ℓ≅B : Level} + (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + where + + open UARel 𝒮-A + open DUARel 𝒮ᴰ-B + + ∫ : UARel (Σ A B) (ℓ-max ℓ≅A ℓ≅B) + UARel._≅_ ∫ (a , b) (a' , b') = Σ[ p ∈ a ≅ a' ] (b ≅ᴰ⟨ p ⟩ b') + UARel.ua ∫ (a , b) (a' , b') = + compEquiv + (Σ-cong-equiv (ua a a') (λ p → uaᴰ b p b')) + ΣPath≃PathΣ + diff --git a/Cubical/Displayed/Constant.agda b/Cubical/Displayed/Constant.agda new file mode 100644 index 000000000..0689104c2 --- /dev/null +++ b/Cubical/Displayed/Constant.agda @@ -0,0 +1,39 @@ +{- + + Functions building DUARels on constant families + +-} +{-# OPTIONS --safe #-} +module Cubical.Displayed.Constant where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Subst + +private + variable + ℓ ℓA ℓA' ℓP ℓ≅A ℓ≅A' ℓB ℓB' ℓ≅B ℓ≅B' ℓC ℓ≅C : Level + +-- constant DUARel + +module _ {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) + {B : Type ℓB} (𝒮-B : UARel B ℓ≅B) where + + open UARel 𝒮-B + open DUARel + + 𝒮ᴰ-const : DUARel 𝒮-A (λ _ → B) ℓ≅B + 𝒮ᴰ-const ._≅ᴰ⟨_⟩_ b _ b' = b ≅ b' + 𝒮ᴰ-const .uaᴰ b p b' = ua b b' + +-- SubstRel for an arbitrary constant family + +module _ {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) (B : Type ℓB) where + + open SubstRel + + 𝒮ˢ-const : SubstRel 𝒮-A (λ _ → B) + 𝒮ˢ-const .SubstRel.act _ = idEquiv B + 𝒮ˢ-const .SubstRel.uaˢ p b = transportRefl b diff --git a/Cubical/Displayed/Everything.agda b/Cubical/Displayed/Everything.agda new file mode 100644 index 000000000..d4e083945 --- /dev/null +++ b/Cubical/Displayed/Everything.agda @@ -0,0 +1,16 @@ +{-# OPTIONS --safe --guardedness #-} +module Cubical.Displayed.Everything where + +import Cubical.Displayed.Auto +import Cubical.Displayed.Base +import Cubical.Displayed.Constant +import Cubical.Displayed.Function +import Cubical.Displayed.Generic +import Cubical.Displayed.Morphism +import Cubical.Displayed.Prop +import Cubical.Displayed.Properties +import Cubical.Displayed.Record +import Cubical.Displayed.Sigma +import Cubical.Displayed.Subst +import Cubical.Displayed.Unit +import Cubical.Displayed.Universe diff --git a/Cubical/Displayed/Function.agda b/Cubical/Displayed/Function.agda new file mode 100644 index 000000000..9132432b3 --- /dev/null +++ b/Cubical/Displayed/Function.agda @@ -0,0 +1,163 @@ +{- + + Functions building UARels and DUARels on function types + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Displayed.Function where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path + +open import Cubical.Functions.FunExtEquiv +open import Cubical.Functions.Implicit + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Constant +open import Cubical.Displayed.Morphism +open import Cubical.Displayed.Subst +open import Cubical.Displayed.Sigma + +private + variable + ℓA ℓ≅A ℓB ℓ≅B ℓC ℓ≅C : Level + +-- UARel on dependent function type +-- from UARel on domain and DUARel on codomain + +module _ {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) {B : A → Type ℓB} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) where + + open UARel 𝒮-A + open DUARel 𝒮ᴰ-B + + 𝒮-Π : UARel ((a : A) → B a) (ℓ-max ℓA ℓ≅B) + UARel._≅_ 𝒮-Π f f' = ∀ a → f a ≅ᴰ⟨ ρ a ⟩ f' a + UARel.ua 𝒮-Π f f' = + compEquiv + (equivΠCod λ a → uaᴰρ (f a) (f' a)) + funExtEquiv + +-- Parameterize UARel by type + +_→𝒮_ : (A : Type ℓA) {B : Type ℓB} (𝒮-B : UARel B ℓ≅B) → UARel (A → B) (ℓ-max ℓA ℓ≅B) +(A →𝒮 𝒮-B) .UARel._≅_ f f' = ∀ a → 𝒮-B .UARel._≅_ (f a) (f' a) +(A →𝒮 𝒮-B) .UARel.ua f f' = + compEquiv + (equivΠCod λ a → 𝒮-B .UARel.ua (f a) (f' a)) + funExtEquiv + +𝒮-app : {A : Type ℓA} {B : Type ℓB} {𝒮-B : UARel B ℓ≅B} + → A → UARelHom (A →𝒮 𝒮-B) 𝒮-B +𝒮-app a .UARelHom.fun f = f a +𝒮-app a .UARelHom.rel h = h a +𝒮-app a .UARelHom.ua h = refl + +-- DUARel on dependent function type +-- from DUARels on domain and codomain + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + {C : (a : A) → B a → Type ℓC} (𝒮ᴰ-C : DUARel (∫ 𝒮ᴰ-B) (uncurry C) ℓ≅C) + where + + open UARel 𝒮-A + private + module B = DUARel 𝒮ᴰ-B + module C = DUARel 𝒮ᴰ-C + + 𝒮ᴰ-Π : DUARel 𝒮-A (λ a → (b : B a) → C a b) (ℓ-max (ℓ-max ℓB ℓ≅B) ℓ≅C) + DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-Π f p f' = + ∀ {b b'} (q : b B.≅ᴰ⟨ p ⟩ b') → f b C.≅ᴰ⟨ p , q ⟩ f' b' + DUARel.uaᴰ 𝒮ᴰ-Π f p f' = + compEquiv + (equivImplicitΠCod λ {b} → + (equivImplicitΠCod λ {b'} → + (equivΠ (B.uaᴰ b p b') (λ q → C.uaᴰ (f b) (p , q) (f' b'))))) + funExtDepEquiv + +_→𝒮ᴰ_ : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + {C : A → Type ℓC} (𝒮ᴰ-C : DUARel 𝒮-A C ℓ≅C) + → DUARel 𝒮-A (λ a → B a → C a) (ℓ-max (ℓ-max ℓB ℓ≅B) ℓ≅C) +𝒮ᴰ-B →𝒮ᴰ 𝒮ᴰ-C = + 𝒮ᴰ-Π 𝒮ᴰ-B (𝒮ᴰ-Lift _ 𝒮ᴰ-C 𝒮ᴰ-B) + +-- DUARel on dependent function type +-- from a SubstRel on the domain and DUARel on the codomain + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} (𝒮ˢ-B : SubstRel 𝒮-A B) + {C : (a : A) → B a → Type ℓC} (𝒮ᴰ-C : DUARel (∫ˢ 𝒮ˢ-B) (uncurry C) ℓ≅C) + where + + open UARel 𝒮-A + open SubstRel 𝒮ˢ-B + open DUARel 𝒮ᴰ-C + + 𝒮ᴰ-Πˢ : DUARel 𝒮-A (λ a → (b : B a) → C a b) (ℓ-max ℓB ℓ≅C) + DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-Πˢ f p f' = + (b : B _) → f b ≅ᴰ⟨ p , refl ⟩ f' (act p .fst b) + DUARel.uaᴰ 𝒮ᴰ-Πˢ f p f' = + compEquiv + (compEquiv + (equivΠCod λ b → Jequiv (λ b' q → f b ≅ᴰ⟨ p , q ⟩ f' b')) + (invEquiv implicit≃Explicit)) + (DUARel.uaᴰ (𝒮ᴰ-Π (Subst→DUA 𝒮ˢ-B) 𝒮ᴰ-C) f p f') + +-- SubstRel on a dependent function type +-- from a SubstRel on the domain and SubstRel on the codomain + +equivΠ' : ∀ {ℓA ℓA' ℓB ℓB'} {A : Type ℓA} {A' : Type ℓA'} + {B : A → Type ℓB} {B' : A' → Type ℓB'} + (eA : A ≃ A') + (eB : {a : A} {a' : A'} → eA .fst a ≡ a' → B a ≃ B' a') + → ((a : A) → B a) ≃ ((a' : A') → B' a') +equivΠ' {B' = B'} eA eB = isoToEquiv isom + where + open Iso + + isom : Iso _ _ + isom .fun f a' = + eB (secEq eA a') .fst (f (invEq eA a')) + isom .inv f' a = + invEq (eB refl) (f' (eA .fst a)) + isom .rightInv f' = + funExt λ a' → + J (λ a'' p → eB p .fst (invEq (eB refl) (f' (p i0))) ≡ f' a'') + (secEq (eB refl) (f' (eA .fst (invEq eA a')))) + (secEq eA a') + isom .leftInv f = + funExt λ a → + subst + (λ p → invEq (eB refl) (eB p .fst (f (invEq eA (eA .fst a)))) ≡ f a) + (sym (commPathIsEq (eA .snd) a)) + (J (λ a'' p → invEq (eB refl) (eB (cong (eA .fst) p) .fst (f (invEq eA (eA .fst a)))) ≡ f a'') + (retEq (eB refl) (f (invEq eA (eA .fst a)))) + (retEq eA a)) + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} (𝒮ˢ-B : SubstRel 𝒮-A B) + {C : Σ A B → Type ℓC} (𝒮ˢ-C : SubstRel (∫ˢ 𝒮ˢ-B) C) + where + + open UARel 𝒮-A + open SubstRel + private + module B = SubstRel 𝒮ˢ-B + module C = SubstRel 𝒮ˢ-C + + 𝒮ˢ-Π : SubstRel 𝒮-A (λ a → (b : B a) → C (a , b)) + 𝒮ˢ-Π .act p = equivΠ' (B.act p) (λ q → C.act (p , q)) + 𝒮ˢ-Π .uaˢ p f = + fromPathP + (DUARel.uaᴰ (𝒮ᴰ-Π (Subst→DUA 𝒮ˢ-B) (Subst→DUA 𝒮ˢ-C)) f p (equivFun (𝒮ˢ-Π .act p) f) .fst + (λ {b} → + J (λ b' q → + equivFun (C.act (p , q)) (f b) + ≡ equivFun (equivΠ' (𝒮ˢ-B .act p) (λ q → C.act (p , q))) f b') + (λ i → + C.act (p , λ j → commSqIsEq (𝒮ˢ-B .act p .snd) b (~ i) j) .fst + (f (retEq (𝒮ˢ-B .act p) b (~ i)))))) diff --git a/Cubical/Displayed/Generic.agda b/Cubical/Displayed/Generic.agda new file mode 100644 index 000000000..d28d705bc --- /dev/null +++ b/Cubical/Displayed/Generic.agda @@ -0,0 +1,36 @@ +{- + + Generic UARel, DUARel, and SubstRel: the path relation is always univalent + +-} +{-# OPTIONS --safe #-} +module Cubical.Displayed.Generic where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Transport + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Subst + +private + variable + ℓ ℓA ℓA' ℓP ℓ≅A ℓ≅A' ℓB ℓB' ℓ≅B ℓ≅B' ℓC ℓ≅C : Level + +-- UARel for an arbitrary type + +𝒮-generic : (A : Type ℓA) → UARel A ℓA +UARel._≅_ (𝒮-generic A) = _≡_ +UARel.ua (𝒮-generic A) a a' = idEquiv (a ≡ a') + +-- DUARel for an arbitrary family + +𝒮ᴰ-generic : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) (B : A → Type ℓB) → DUARel 𝒮-A B ℓB +𝒮ᴰ-generic 𝒮-A B .DUARel._≅ᴰ⟨_⟩_ b p b' = PathP (λ i → B (UARel.≅→≡ 𝒮-A p i)) b b' +𝒮ᴰ-generic 𝒮-A B .DUARel.uaᴰ b p b' = idEquiv _ + +-- SubstRel for an arbitrary family + +𝒮ˢ-generic : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) (B : A → Type ℓB) → SubstRel 𝒮-A B +𝒮ˢ-generic 𝒮-A B .SubstRel.act p = substEquiv B (UARel.≅→≡ 𝒮-A p) +𝒮ˢ-generic 𝒮-A B .SubstRel.uaˢ p b = refl diff --git a/Cubical/Displayed/Morphism.agda b/Cubical/Displayed/Morphism.agda new file mode 100644 index 000000000..94b463a36 --- /dev/null +++ b/Cubical/Displayed/Morphism.agda @@ -0,0 +1,67 @@ +{- + A morphism of UARels is a function between the structures with an action on the relations that + commutes with the equivalence to PathP. + + We can reindex a DUARel or SubstRel along one of these. +-} +{-# OPTIONS --safe #-} +module Cubical.Displayed.Morphism where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Subst + +private + variable + ℓ ℓA ℓA' ℓ≅A ℓB ℓB' ℓ≅B ℓC ℓ≅C : Level + +record UARelHom {A : Type ℓA} {B : Type ℓB} (𝒮-A : UARel A ℓ≅A) (𝒮-B : UARel B ℓ≅B) + : Type (ℓ-max (ℓ-max ℓA ℓ≅A) (ℓ-max ℓB ℓ≅B)) where + no-eta-equality + constructor uarelhom + field + fun : A → B + rel : ∀ {a a'} → UARel._≅_ 𝒮-A a a' → UARel._≅_ 𝒮-B (fun a) (fun a') + ua : ∀ {a a'} (p : UARel._≅_ 𝒮-A a a') + → cong fun (UARel.≅→≡ 𝒮-A p) ≡ UARel.≅→≡ 𝒮-B (rel p) + +open UARelHom + +𝒮-id : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) → UARelHom 𝒮-A 𝒮-A +𝒮-id 𝒮-A .fun = idfun _ +𝒮-id 𝒮-A .rel = idfun _ +𝒮-id 𝒮-A .ua _ = refl + +𝒮-∘ : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} {B : Type ℓB} {𝒮-B : UARel B ℓ≅B} + {C : Type ℓC} {𝒮-C : UARel C ℓ≅C} + → UARelHom 𝒮-B 𝒮-C + → UARelHom 𝒮-A 𝒮-B + → UARelHom 𝒮-A 𝒮-C +𝒮-∘ g f .fun = g .fun ∘ f .fun +𝒮-∘ g f .rel = g .rel ∘ f .rel +𝒮-∘ {𝒮-A = 𝒮-A} {𝒮-B = 𝒮-B} {𝒮-C = 𝒮-C} g f .ua p = + cong (cong (g .fun)) (f .ua p) ∙ g .ua (f .rel p) + +𝒮ᴰ-reindex : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} {B : Type ℓB} {𝒮-B : UARel B ℓ≅B} {C : B → Type ℓC} + (f : UARelHom 𝒮-A 𝒮-B) + → DUARel 𝒮-B C ℓ≅C + → DUARel 𝒮-A (C ∘ fun f) ℓ≅C +𝒮ᴰ-reindex f 𝒮ᴰ-C .DUARel._≅ᴰ⟨_⟩_ c p c' = 𝒮ᴰ-C .DUARel._≅ᴰ⟨_⟩_ c (f .rel p) c' +𝒮ᴰ-reindex {C = C} f 𝒮ᴰ-C .DUARel.uaᴰ c p c' = + compEquiv + (𝒮ᴰ-C .DUARel.uaᴰ c (f .rel p) c') + (substEquiv (λ q → PathP (λ i → C (q i)) c c') (sym (f .ua p))) + +𝒮ˢ-reindex : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} {B : Type ℓB} {𝒮-B : UARel B ℓ≅B} {C : B → Type ℓC} + (f : UARelHom 𝒮-A 𝒮-B) + → SubstRel 𝒮-B C + → SubstRel 𝒮-A (C ∘ fun f) +𝒮ˢ-reindex f 𝒮ˢ-C .SubstRel.act p = 𝒮ˢ-C .SubstRel.act (f .rel p) +𝒮ˢ-reindex {C = C} f 𝒮ˢ-C .SubstRel.uaˢ p c = + cong (λ q → subst C q c) (f .ua p) + ∙ 𝒮ˢ-C .SubstRel.uaˢ (f .rel p) c diff --git a/Cubical/Displayed/Prop.agda b/Cubical/Displayed/Prop.agda new file mode 100644 index 000000000..d57a0d5e2 --- /dev/null +++ b/Cubical/Displayed/Prop.agda @@ -0,0 +1,52 @@ +{- + + Functions building UARels and DUARels on propositions / propositional families + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Displayed.Prop where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Sigma +open import Cubical.Data.Unit + +open import Cubical.Displayed.Base + +private + variable + ℓA ℓ≅A ℓB ℓ≅B ℓP : Level + +𝒮-prop : (P : hProp ℓP) → UARel (P .fst) ℓ-zero +𝒮-prop P .UARel._≅_ _ _ = Unit +𝒮-prop P .UARel.ua _ _ = + invEquiv (isContr→≃Unit (isOfHLevelPath' 0 (P .snd) _ _)) + +𝒮ᴰ-prop : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) (P : A → hProp ℓP) + → DUARel 𝒮-A (λ a → P a .fst) ℓ-zero +𝒮ᴰ-prop 𝒮-A P .DUARel._≅ᴰ⟨_⟩_ _ _ _ = Unit +𝒮ᴰ-prop 𝒮-A P .DUARel.uaᴰ _ _ _ = + invEquiv (isContr→≃Unit (isOfHLevelPathP' 0 (P _ .snd) _ _)) + +𝒮-subtype : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) {P : A → Type ℓP} + → (∀ a → isProp (P a)) + → UARel (Σ A P) ℓ≅A +𝒮-subtype 𝒮-A propP .UARel._≅_ (a , _) (a' , _) = 𝒮-A .UARel._≅_ a a' +𝒮-subtype 𝒮-A propP .UARel.ua (a , _) (a' , _) = + compEquiv (𝒮-A .UARel.ua a a') (Σ≡PropEquiv propP) + +𝒮ᴰ-subtype : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + {P : (a : A) → B a → Type ℓP} + → (∀ a b → isProp (P a b)) + → DUARel 𝒮-A (λ a → Σ[ b ∈ B a ] (P a b)) ℓ≅B +𝒮ᴰ-subtype 𝒮ᴰ-B propP .DUARel._≅ᴰ⟨_⟩_ (b , _) p (b' , _) = 𝒮ᴰ-B .DUARel._≅ᴰ⟨_⟩_ b p b' +𝒮ᴰ-subtype 𝒮ᴰ-B propP .DUARel.uaᴰ (b , _) p (b' , _) = + compEquiv + (𝒮ᴰ-B .DUARel.uaᴰ b p b') + (compEquiv + (invEquiv (Σ-contractSnd λ _ → isOfHLevelPathP' 0 (propP _ b') _ _)) + ΣPath≃PathΣ) diff --git a/Cubical/Displayed/Properties.agda b/Cubical/Displayed/Properties.agda new file mode 100644 index 000000000..037a20bed --- /dev/null +++ b/Cubical/Displayed/Properties.agda @@ -0,0 +1,143 @@ +{-# OPTIONS --safe #-} +module Cubical.Displayed.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence using (pathToEquiv; univalence; ua-ungluePath-Equiv) + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma + +open import Cubical.Relation.Binary +open BinaryRelation + +open import Cubical.Displayed.Base + +private + variable + ℓ ℓA ℓA' ℓP ℓ≅A ℓ≅A' ℓB ℓB' ℓ≅B ℓ≅B' ℓC ℓ≅C : Level + +-- induction principles + +module _ {A : Type ℓA} {ℓ≅A : Level} (𝒮-A : UARel A ℓ≅A) where + open UARel 𝒮-A + 𝒮-J : {a : A} + (P : (a' : A) → (p : a ≡ a') → Type ℓ) + (d : P a refl) + {a' : A} + (p : a ≅ a') + → P a' (≅→≡ p) + 𝒮-J {a} P d {a'} p + = J (λ y q → P y q) + d + (≅→≡ p) + + 𝒮-J-2 : {a : A} + (P : (a' : A) → (p : a ≅ a') → Type ℓ) + (d : P a (ρ a)) + {a' : A} + (p : a ≅ a') + → P a' p + 𝒮-J-2 {a = a} P d {a'} p + = subst (λ r → P a' r) (Iso.leftInv (uaIso a a') p) g + where + g : P a' (≡→≅ (≅→≡ p)) + g = J (λ y q → P y (≡→≅ q)) d (≅→≡ p) + + +-- constructors + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} + (_≅ᴰ⟨_⟩_ : {a a' : A} → B a → UARel._≅_ 𝒮-A a a' → B a' → Type ℓ≅B) + where + + open UARel 𝒮-A + + -- constructor that reduces ua to the case where p = ρ a by induction on p + private + 𝒮ᴰ-make-aux : (uni : {a : A} (b b' : B a) → b ≅ᴰ⟨ ρ a ⟩ b' ≃ (b ≡ b')) + → ({a a' : A} (b : B a) (p : a ≅ a') (b' : B a') → (b ≅ᴰ⟨ p ⟩ b') ≃ PathP (λ i → B (≅→≡ p i)) b b') + 𝒮ᴰ-make-aux uni {a} {a'} b p + = 𝒮-J-2 𝒮-A + (λ y q → (b' : B y) → (b ≅ᴰ⟨ q ⟩ b') ≃ PathP (λ i → B (≅→≡ q i)) b b') + (λ b' → uni' b') + p + where + g : (b' : B a) → (b ≡ b') ≡ PathP (λ i → B (≅→≡ (ρ a) i)) b b' + g b' = subst (λ r → (b ≡ b') ≡ PathP (λ i → B (r i)) b b') + (sym (Iso.rightInv (uaIso a a) refl)) + refl + uni' : (b' : B a) → b ≅ᴰ⟨ ρ a ⟩ b' ≃ PathP (λ i → B (≅→≡ (ρ a) i)) b b' + uni' b' = compEquiv (uni b b') (pathToEquiv (g b')) + + 𝒮ᴰ-make-1 : (uni : {a : A} (b b' : B a) → b ≅ᴰ⟨ ρ a ⟩ b' ≃ (b ≡ b')) + → DUARel 𝒮-A B ℓ≅B + DUARel._≅ᴰ⟨_⟩_ (𝒮ᴰ-make-1 uni) = _≅ᴰ⟨_⟩_ + DUARel.uaᴰ (𝒮ᴰ-make-1 uni) = 𝒮ᴰ-make-aux uni + + -- constructor that reduces univalence further to contractibility of relational singletons + + 𝒮ᴰ-make-2 : (ρᴰ : {a : A} → isRefl _≅ᴰ⟨ ρ a ⟩_) + (contrTotal : (a : A) → contrRelSingl _≅ᴰ⟨ UARel.ρ 𝒮-A a ⟩_) + → DUARel 𝒮-A B ℓ≅B + DUARel._≅ᴰ⟨_⟩_ (𝒮ᴰ-make-2 ρᴰ contrTotal) = _≅ᴰ⟨_⟩_ + DUARel.uaᴰ (𝒮ᴰ-make-2 ρᴰ contrTotal) + = 𝒮ᴰ-make-aux (contrRelSingl→isUnivalent _ ρᴰ (contrTotal _)) + +-- relational isomorphisms + +𝒮-iso→iso : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) + {B : Type ℓB} (𝒮-B : UARel B ℓ≅B) + (F : RelIso (UARel._≅_ 𝒮-A) (UARel._≅_ 𝒮-B)) + → Iso A B +𝒮-iso→iso 𝒮-A 𝒮-B F + = RelIso→Iso (UARel._≅_ 𝒮-A) + (UARel._≅_ 𝒮-B) + (UARel.≅→≡ 𝒮-A) + (UARel.≅→≡ 𝒮-B) + F + +-- fiberwise relational isomorphisms + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {A' : Type ℓA'} {𝒮-A' : UARel A' ℓ≅A'} + (F : Iso A A') + {B : A → Type ℓB} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + {B' : A' → Type ℓB'} (𝒮ᴰ-B' : DUARel 𝒮-A' B' ℓ≅B') where + + open UARel 𝒮-A + open DUARel 𝒮ᴰ-B renaming (_≅ᴰ⟨_⟩_ to _≅B⟨_⟩_ + ; uaᴰ to uaB + ; fiberRel to fiberRelB + ; uaᴰρ to uaᴰρB) + open DUARel 𝒮ᴰ-B' renaming (_≅ᴰ⟨_⟩_ to _≅B'⟨_⟩_ + ; uaᴰ to uaB' + ; fiberRel to fiberRelB' + ; uaᴰρ to uaᴰρB') + + private + f = Iso.fun F + + -- the following can of course be done slightly more generally + -- for fiberwise binary relations + + F*fiberRelB' : (a : A) → Rel (B' (f a)) (B' (f a)) ℓ≅B' + F*fiberRelB' a = fiberRelB' (f a) + + module _ (G : (a : A) → RelIso (fiberRelB a) (F*fiberRelB' a)) where + private + fiberIsoOver : (a : A) → Iso (B a) (B' (f a)) + fiberIsoOver a + = RelIso→Iso (fiberRelB a) + (F*fiberRelB' a) + (equivFun (uaᴰρB _ _)) + (equivFun (uaᴰρB' _ _)) + (G a) + + -- DUARelFiberIsoOver→TotalIso produces an isomorphism of total spaces + -- from a relational isomorphism between B a and (F * B) a + 𝒮ᴰ-fiberIsoOver→totalIso : Iso (Σ A B) (Σ A' B') + 𝒮ᴰ-fiberIsoOver→totalIso = Σ-cong-iso F fiberIsoOver diff --git a/Cubical/Displayed/Record.agda b/Cubical/Displayed/Record.agda new file mode 100644 index 000000000..bfec660d9 --- /dev/null +++ b/Cubical/Displayed/Record.agda @@ -0,0 +1,231 @@ +{- + +Generate univalent reflexive graph characterizations for record types from +characterizations of the field types using reflection. + +See end of file for an example. + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Displayed.Record where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Data.Sigma +open import Cubical.Data.List as List +open import Cubical.Data.Unit +open import Cubical.Data.Bool +open import Cubical.Data.Maybe as Maybe + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Properties +open import Cubical.Displayed.Prop +open import Cubical.Displayed.Sigma +open import Cubical.Displayed.Unit +open import Cubical.Displayed.Universe +open import Cubical.Displayed.Auto + +import Agda.Builtin.Reflection as R +open import Cubical.Reflection.Base +import Cubical.Reflection.RecordEquiv as RE + +{- + `DUAFields` + A collection of DURG characterizations for fields of a record is described by an element of this inductive + family. If you just want to see how to use it, have a look at the end of the file first. + + An element of `DUAFields 𝒮-A R _≅R⟨_⟩_ πS 𝒮ᴰ-S πS≅` describes a mapping + - from a structure `R : A → Type _` and notion of structured equivalence `_≅R⟨_⟩_`, + which are meant to be defined as parameterized record types, + - to a DURG `𝒮ᴰ-S`, + the underlying structure of which will be an iterated Σ-type, + - via projections `πS` and `πS≅`. + + `𝒮-A`, `R`, and `_≅R⟨_⟩_` are parameters, while `πS`, `𝒮ᴰ-S`, and `πS≅` are indices; + the user builds up the Σ-type representation of the record using the DUAFields constructors. + + A DUAFields representation is _total_ when the projections `πS` and `πS≅` are equivalences, in which case we + obtain a DURG on `R` with `_≅R⟨_⟩_` as the notion of structured equivalence---see `𝒮ᴰ-Fields` below. + + When `R`, and `_≅R⟨_⟩_` are defined by record types, we can use reflection to automatically generate proofs + `πS` and `πS≅` are equivalences---see `𝒮ᴰ-Record` below. + +-} +data DUAFields {ℓA ℓ≅A ℓR ℓ≅R} {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) + (R : A → Type ℓR) (_≅R⟨_⟩_ : {a a' : A} → R a → UARel._≅_ 𝒮-A a a' → R a' → Type ℓ≅R) + : ∀ {ℓS ℓ≅S} {S : A → Type ℓS} + (πS : ∀ {a} → R a → S a) (𝒮ᴰ-S : DUARel 𝒮-A S ℓ≅S) + (πS≅ : ∀ {a} {r : R a} {e} {r' : R a} → r ≅R⟨ e ⟩ r' → DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-S (πS r) e (πS r')) + → Typeω + where + + -- `fields:` + -- Base case, no fields yet recorded in `𝒮ᴰ-S`. + fields: : DUAFields 𝒮-A R _≅R⟨_⟩_ (λ _ → tt) (𝒮ᴰ-Unit 𝒮-A) (λ _ → tt) + + -- `… data[ πF ∣ 𝒮ᴰ-F ∣ πF≅ ]` + -- Add a new field with a DURG. `πF` should be the name of the field in the structure record `R` and `πF≅` + -- the name of the corresponding field in the equivalence record `_≅R⟨_⟩_`, while `𝒮ᴰ-F` is a DURG for the + -- field's type over `𝒮-A`. Data fields that depend on previous fields of the record are not currently + -- supported. + _data[_∣_∣_] : ∀ {ℓS ℓ≅S} {S : A → Type ℓS} + {πS : ∀ {a} → R a → S a} {𝒮ᴰ-S : DUARel 𝒮-A S ℓ≅S} + {πS≅ : ∀ {a} {r : R a} {e} {r' : R a} → r ≅R⟨ e ⟩ r' → DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-S (πS r) e (πS r')} + → DUAFields 𝒮-A R _≅R⟨_⟩_ πS 𝒮ᴰ-S πS≅ + → ∀ {ℓF ℓ≅F} {F : A → Type ℓF} + (πF : ∀ {a} → (r : R a) → F a) + (𝒮ᴰ-F : DUARel 𝒮-A F ℓ≅F) + (πF≅ : ∀ {a} {r : R a} {e} {r' : R a} (p : r ≅R⟨ e ⟩ r') → DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-F (πF r) e (πF r')) + → DUAFields 𝒮-A R _≅R⟨_⟩_ (λ r → πS r , πF r) (𝒮ᴰ-S ×𝒮ᴰ 𝒮ᴰ-F) (λ p → πS≅ p , πF≅ p) + + -- `… prop[ πF ∣ propF ]` + -- Add a new propositional field. `πF` should be the name of the field in the structure record `R`, while + -- propF is a proof that this field is a proposition. + _prop[_∣_] : ∀ {ℓS ℓ≅S} {S : A → Type ℓS} + {πS : ∀ {a} → R a → S a} {𝒮ᴰ-S : DUARel 𝒮-A S ℓ≅S} + {πS≅ : ∀ {a} {r : R a} {e} {r' : R a} → r ≅R⟨ e ⟩ r' → DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-S (πS r) e (πS r')} + → DUAFields 𝒮-A R _≅R⟨_⟩_ πS 𝒮ᴰ-S πS≅ + → ∀ {ℓF} {F : (a : A) → S a → Type ℓF} + (πF : ∀ {a} → (r : R a) → F a (πS r)) + (propF : ∀ a s → isProp (F a s)) + → DUAFields 𝒮-A R _≅R⟨_⟩_ (λ r → πS r , πF r) (𝒮ᴰ-subtype 𝒮ᴰ-S propF) (λ p → πS≅ p) + +module _ {ℓA ℓ≅A} {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {ℓR ℓ≅R} {R : A → Type ℓR} (_≅R⟨_⟩_ : {a a' : A} → R a → UARel._≅_ 𝒮-A a a' → R a' → Type ℓ≅R) + {ℓS ℓ≅S} {S : A → Type ℓS} + {πS : ∀ {a} → R a → S a} {𝒮ᴰ-S : DUARel 𝒮-A S ℓ≅S} + {πS≅ : ∀ {a} {r : R a} {e} {r' : R a} → r ≅R⟨ e ⟩ r' → DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-S (πS r) e (πS r')} + (fs : DUAFields 𝒮-A R _≅R⟨_⟩_ πS 𝒮ᴰ-S πS≅) + where + + open UARel 𝒮-A + open DUARel 𝒮ᴰ-S + + 𝒮ᴰ-Fields : + (e : ∀ a → Iso (R a) (S a)) + (e≅ : ∀ a a' (r : R a) p (r' : R a') → Iso (r ≅R⟨ p ⟩ r') ((e a .Iso.fun r ≅ᴰ⟨ p ⟩ e a' .Iso.fun r'))) + → DUARel 𝒮-A R ℓ≅R + DUARel._≅ᴰ⟨_⟩_ (𝒮ᴰ-Fields e e≅) r p r' = r ≅R⟨ p ⟩ r' + DUARel.uaᴰ (𝒮ᴰ-Fields e e≅) r p r' = + isoToEquiv + (compIso + (e≅ _ _ r p r') + (compIso + (equivToIso (uaᴰ (e _ .Iso.fun r) p (e _ .Iso.fun r'))) + (invIso (congPathIso λ i → isoToEquiv (e _))))) + +module DisplayedRecordMacro where + + -- Extract a name from a term + findName : R.Term → R.TC R.Name + findName t = + Maybe.rec + (R.typeError (R.strErr "Not a name: " ∷ R.termErr t ∷ [])) + (λ s → s) + (go t) + where + go : R.Term → Maybe (R.TC R.Name) + go (R.meta x _) = just (R.blockOnMeta x) + go (R.def name _) = just (R.returnTC name) + go (R.lam _ (R.abs _ t)) = go t + go t = nothing + + -- ℓA ℓ≅A ℓR ℓ≅R A 𝒮-A R _≅R⟨_⟩_ + pattern family∷ hole = _ h∷ _ h∷ _ h∷ _ h∷ _ h∷ _ h∷ _ h∷ _ h∷ hole + + -- ℓS ℓ≅S S πS 𝒮ᴰ-S πS≅ + pattern indices∷ hole = _ h∷ _ h∷ _ h∷ _ h∷ _ h∷ _ h∷ hole + + {- + Takes a reflected DUAFields term as input and collects lists of structure field names and equivalence + field names. (These are returned in reverse order. + -} + parseFields : R.Term → R.TC (List R.Name × List R.Name) + parseFields (R.con (quote fields:) _) = R.returnTC ([] , []) + parseFields (R.con (quote _data[_∣_∣_]) (family∷ (indices∷ (fs v∷ ℓF h∷ ℓ≅F h∷ F h∷ πF v∷ 𝒮ᴰ-F v∷ πF≅ v∷ _)))) = + parseFields fs >>= λ (fs , f≅s) → + findName πF >>= λ f → + findName πF≅ >>= λ f≅ → + R.returnTC (f ∷ fs , f≅ ∷ f≅s) + parseFields (R.con (quote _prop[_∣_]) (family∷ (indices∷ (fs v∷ ℓF h∷ F h∷ πF v∷ _)))) = + parseFields fs >>= λ (fs , f≅s) → + findName πF >>= λ f → + R.returnTC (f ∷ fs , f≅s) + parseFields (R.meta x _) = R.blockOnMeta x + parseFields t = R.typeError (R.strErr "Malformed specification: " ∷ R.termErr t ∷ []) + + {- + Given a list of record field names (in reverse order), generates a ΣFormat (in the sense of + Cubical.Reflection.RecordEquiv) associating the record fields with the fields of a left-associated + iterated Σ-type + -} + List→LeftAssoc : List R.Name → RE.ΣFormat + List→LeftAssoc [] = RE.unit + List→LeftAssoc (x ∷ xs) = List→LeftAssoc xs RE., RE.leaf x + + module _ {ℓA ℓ≅A} {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) + {ℓR ℓ≅R} {R : A → Type ℓR} (≅R : {a a' : A} → R a → UARel._≅_ 𝒮-A a a' → R a' → Type ℓ≅R) + {ℓS ℓ≅S} {S : A → Type ℓS} + {πS : ∀ {a} → R a → S a} {𝒮ᴰ-S : DUARel 𝒮-A S ℓ≅S} + {πS≅ : ∀ {a} {r : R a} {e} {r' : R a} → ≅R r e r' → DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-S (πS r) e (πS r')} + where + + {- + "𝒮ᴰ-Record ... : DUARel 𝒮-A R ℓ≅R" + Requires that `R` and `_≅R⟨_⟩_` are defined by records and `πS` and `πS≅` are equivalences. + The proofs of equivalence are generated using Cubical.Reflection.RecordEquiv and then + `𝒮ᴰ-Fields` is applied. + -} + 𝒮ᴰ-Record : DUAFields 𝒮-A R ≅R πS 𝒮ᴰ-S πS≅ → R.Term → R.TC Unit + 𝒮ᴰ-Record fs hole = + R.quoteTC (DUARel 𝒮-A R ℓ≅R) >>= R.checkType hole >>= λ hole → + R.quoteωTC fs >>= λ `fs` → + parseFields `fs` >>= λ (fields , ≅fields) → + R.freshName "fieldsIso" >>= λ fieldsIso → + R.freshName "≅fieldsIso" >>= λ ≅fieldsIso → + R.quoteTC R >>= R.normalise >>= λ `R` → + R.quoteTC {A = {a a' : A} → R a → UARel._≅_ 𝒮-A a a' → R a' → Type ℓ≅R} ≅R >>= R.normalise >>= λ `≅R` → + findName `R` >>= RE.declareRecordIsoΣ' fieldsIso (List→LeftAssoc fields) >> + findName `≅R` >>= RE.declareRecordIsoΣ' ≅fieldsIso (List→LeftAssoc ≅fields) >> + R.unify hole + (R.def (quote 𝒮ᴰ-Fields) + (`≅R` v∷ `fs` v∷ + vlam "_" (R.def fieldsIso []) v∷ + vlam "a" (vlam "a'" (vlam "r" (vlam "p" (vlam "r'" (R.def ≅fieldsIso []))))) v∷ + [])) + +macro + 𝒮ᴰ-Record = DisplayedRecordMacro.𝒮ᴰ-Record + +-- Example + +private + module Example where + + record Example (A : Type) : Type where + no-eta-equality -- works with or without eta equality + field + dog : A → A → A + cat : A → A → A + mouse : Unit + + open Example + + record ExampleEquiv {A B : Type} (x : Example A) (e : A ≃ B) (y : Example B) : Type where + no-eta-equality -- works with or without eta equality + field + dogEq : ∀ a a' → e .fst (x .dog a a') ≡ y .dog (e .fst a) (e .fst a') + catEq : ∀ a a' → e .fst (x .cat a a') ≡ y .cat (e .fst a) (e .fst a') + + open ExampleEquiv + + example : DUARel (𝒮-Univ ℓ-zero) Example ℓ-zero + example = + 𝒮ᴰ-Record (𝒮-Univ ℓ-zero) ExampleEquiv + (fields: + data[ dog ∣ autoDUARel _ _ ∣ dogEq ] + data[ cat ∣ autoDUARel _ _ ∣ catEq ] + prop[ mouse ∣ (λ _ _ → isPropUnit) ]) diff --git a/Cubical/Displayed/Sigma.agda b/Cubical/Displayed/Sigma.agda new file mode 100644 index 000000000..8eb37612b --- /dev/null +++ b/Cubical/Displayed/Sigma.agda @@ -0,0 +1,101 @@ +{- + + Functions building UARels and DUARels on Σ-types + +-} +{-# OPTIONS --safe #-} +module Cubical.Displayed.Sigma where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Sigma + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Subst +open import Cubical.Displayed.Morphism +open import Cubical.Displayed.Constant + +private + variable + ℓ ℓA ℓA' ℓP ℓ≅A ℓ≅A' ℓB ℓB' ℓ≅B ℓ≅B' ℓC ℓ≅C : Level + +-- UARel on a Σ-type + +∫ˢ : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} {B : A → Type ℓB} (𝒮ˢ-B : SubstRel 𝒮-A B) + → UARel (Σ A B) (ℓ-max ℓ≅A ℓB) +∫ˢ 𝒮ˢ-B = ∫ (Subst→DUA 𝒮ˢ-B) + +_×𝒮_ : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) {B : Type ℓB} (𝒮-B : UARel B ℓ≅B) + → UARel (A × B) (ℓ-max ℓ≅A ℓ≅B) +𝒮-A ×𝒮 𝒮-B = ∫ (𝒮ᴰ-const 𝒮-A 𝒮-B) + +-- Projection UARel morphisms + +𝒮-fst : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} {B : A → Type ℓB} {𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B} + → UARelHom (∫ 𝒮ᴰ-B) 𝒮-A +𝒮-fst .UARelHom.fun = fst +𝒮-fst .UARelHom.rel = fst +𝒮-fst .UARelHom.ua p = refl + +𝒮-snd : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} {B : Type ℓB} {𝒮-B : UARel B ℓ≅B} + → UARelHom (𝒮-A ×𝒮 𝒮-B) 𝒮-B +𝒮-snd .UARelHom.fun = snd +𝒮-snd .UARelHom.rel = snd +𝒮-snd .UARelHom.ua p = refl + +-- Lift a DUARel to live over a Σ-type + +𝒮ᴰ-Lift : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) {B : A → Type ℓB} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + {C : A → Type ℓC} (𝒮ᴰ-C : DUARel 𝒮-A C ℓ≅C) + → DUARel (∫ 𝒮ᴰ-C) (λ (a , _) → B a) ℓ≅B +𝒮ᴰ-Lift _ 𝒮ᴰ-B _ = 𝒮ᴰ-reindex 𝒮-fst 𝒮ᴰ-B + +-- DUARel on a Σ-type + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} {ℓ≅B : Level} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + {C : Σ A B → Type ℓC} {ℓ≅C : Level} (𝒮ᴰ-C : DUARel (∫ 𝒮ᴰ-B) C ℓ≅C) + where + + open UARel 𝒮-A + private + module B = DUARel 𝒮ᴰ-B + module C = DUARel 𝒮ᴰ-C + + 𝒮ᴰ-Σ : DUARel 𝒮-A (λ a → Σ[ b ∈ B a ] C (a , b)) (ℓ-max ℓ≅B ℓ≅C) + DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-Σ (b , c) p (b' , c') = + Σ[ q ∈ b B.≅ᴰ⟨ p ⟩ b' ] (c C.≅ᴰ⟨ p , q ⟩ c') + DUARel.uaᴰ 𝒮ᴰ-Σ (b , c) p (b' , c') = + compEquiv + (Σ-cong-equiv (B.uaᴰ b p b') (λ q → C.uaᴰ c (p , q) c')) + ΣPath≃PathΣ + +-- DUARel on a non-dependent Σ-type + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} {ℓ≅B : Level} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) + {C : A → Type ℓC} {ℓ≅C : Level} (𝒮ᴰ-C : DUARel 𝒮-A C ℓ≅C) + where + + _×𝒮ᴰ_ : DUARel 𝒮-A (λ a → B a × C a) (ℓ-max ℓ≅B ℓ≅C) + _×𝒮ᴰ_ = 𝒮ᴰ-Σ 𝒮ᴰ-B (𝒮ᴰ-Lift 𝒮-A 𝒮ᴰ-C 𝒮ᴰ-B) + +-- SubstRel on a Σ-type + +module _ {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} + {B : A → Type ℓB} (𝒮ˢ-B : SubstRel 𝒮-A B) + {C : Σ A B → Type ℓC} (𝒮ˢ-C : SubstRel (∫ˢ 𝒮ˢ-B) C) + where + + open UARel 𝒮-A + open SubstRel + private + module B = SubstRel 𝒮ˢ-B + module C = SubstRel 𝒮ˢ-C + + 𝒮ˢ-Σ : SubstRel 𝒮-A (λ a → Σ[ b ∈ B a ] C (a , b)) + 𝒮ˢ-Σ .act p = Σ-cong-equiv (B.act p) (λ b → C.act (p , refl)) + 𝒮ˢ-Σ .uaˢ p _ = + fromPathP + (DUARel.uaᴰ (𝒮ᴰ-Σ (Subst→DUA 𝒮ˢ-B) (Subst→DUA 𝒮ˢ-C)) _ p _ .fst (refl , refl)) diff --git a/Cubical/Displayed/Subst.agda b/Cubical/Displayed/Subst.agda new file mode 100644 index 000000000..76f7b08a3 --- /dev/null +++ b/Cubical/Displayed/Subst.agda @@ -0,0 +1,60 @@ +{- + Given a type A with a UARel and a family B over A, + a SubstRel on B is a family of functions a ≅ a' → B a ≃ B a' path-equal to transport in that family. + + Any SubstRel gives rise to a DUARel in which b and b' are related over p when the transport of b along p is + equial to b'. +-} + +{-# OPTIONS --safe #-} +module Cubical.Displayed.Subst where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence + +open import Cubical.Displayed.Base + +private + variable + ℓA ℓ≅A ℓB : Level + +record SubstRel {A : Type ℓA} {ℓ≅A : Level} (𝒮-A : UARel A ℓ≅A) (B : A → Type ℓB) + : Type (ℓ-max (ℓ-max ℓA ℓB) ℓ≅A) + where + + no-eta-equality + constructor substrel + open UARel 𝒮-A + + field + act : {a a' : A} → a ≅ a' → B a ≃ B a' + uaˢ : {a a' : A} (p : a ≅ a') (b : B a) → subst B (≅→≡ p) b ≡ equivFun (act p) b + + uaˢ⁻ : {a a' : A} (p : a ≅ a') (b : B a') → subst B (sym (≅→≡ p)) b ≡ invEq (act p) b + uaˢ⁻ p b = + subst B (sym (≅→≡ p)) b + ≡⟨ cong (subst B (sym (≅→≡ p))) (sym (secEq (act p) b)) ⟩ + subst B (sym (≅→≡ p)) (equivFun (act p) (invEq (act p) b)) + ≡⟨ cong (subst B (sym (≅→≡ p))) (sym (uaˢ p (invEq (act p) b))) ⟩ + subst B (sym (≅→≡ p)) (subst B (≅→≡ p) (invEq (act p) b)) + ≡⟨ pathToIso (cong B (≅→≡ p)) .Iso.leftInv (invEq (act p) b) ⟩ + invEq (act p) b + ∎ + +Subst→DUA : {A : Type ℓA} {ℓ≅A : Level} {𝒮-A : UARel A ℓ≅A} {B : A → Type ℓB} + → SubstRel 𝒮-A B → DUARel 𝒮-A B ℓB +DUARel._≅ᴰ⟨_⟩_ (Subst→DUA 𝒮ˢ-B) b p b' = + equivFun (SubstRel.act 𝒮ˢ-B p) b ≡ b' +DUARel.uaᴰ (Subst→DUA {𝒮-A = 𝒮-A} {B = B} 𝒮ˢ-B) b p b' = + equivFun (SubstRel.act 𝒮ˢ-B p) b ≡ b' + ≃⟨ invEquiv (compPathlEquiv (sym (SubstRel.uaˢ 𝒮ˢ-B p b))) ⟩ + subst B (≅→≡ p) b ≡ b' + ≃⟨ invEquiv (PathP≃Path (λ i → B (≅→≡ p i)) b b') ⟩ + PathP (λ i → B (UARel.≅→≡ 𝒮-A p i)) b b' + ■ + where + open UARel 𝒮-A diff --git a/Cubical/Displayed/Unit.agda b/Cubical/Displayed/Unit.agda new file mode 100644 index 000000000..6d24f8539 --- /dev/null +++ b/Cubical/Displayed/Unit.agda @@ -0,0 +1,27 @@ +{- + + DUARel for the constant unit family + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Displayed.Unit where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Unit + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Constant + +private + variable + ℓA ℓ≅A : Level + +𝒮-Unit : UARel Unit ℓ-zero +𝒮-Unit .UARel._≅_ _ _ = Unit +𝒮-Unit .UARel.ua _ _ = invEquiv (isContr→≃Unit (isProp→isContrPath isPropUnit _ _)) + +𝒮ᴰ-Unit : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) → DUARel 𝒮-A (λ _ → Unit) ℓ-zero +𝒮ᴰ-Unit 𝒮-A = 𝒮ᴰ-const 𝒮-A 𝒮-Unit diff --git a/Cubical/Displayed/Universe.agda b/Cubical/Displayed/Universe.agda new file mode 100644 index 000000000..5a2da6f0e --- /dev/null +++ b/Cubical/Displayed/Universe.agda @@ -0,0 +1,32 @@ +{- + + - UARel given by a universe and equivalences + - SubstRel and DUARel for the element family over the universe + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Displayed.Universe where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Subst + +private + variable + ℓA ℓ≅A ℓB ℓ≅B ℓP : Level + +𝒮-Univ : ∀ ℓ → UARel (Type ℓ) ℓ +𝒮-Univ ℓ .UARel._≅_ = _≃_ +𝒮-Univ ℓ .UARel.ua _ _ = isoToEquiv (invIso univalenceIso) + +𝒮ˢ-El : ∀ ℓ → SubstRel (𝒮-Univ ℓ) (λ X → X) +𝒮ˢ-El ℓ .SubstRel.act e = e +𝒮ˢ-El ℓ .SubstRel.uaˢ e a = uaβ e a + +𝒮ᴰ-El : ∀ ℓ → DUARel (𝒮-Univ ℓ) (λ X → X) ℓ +𝒮ᴰ-El ℓ .DUARel._≅ᴰ⟨_⟩_ a e a' = e .fst a ≡ a' +𝒮ᴰ-El ℓ .DUARel.uaᴰ a e a' = invEquiv (ua-ungluePath-Equiv e) diff --git a/Cubical/Experiments/Brunerie.agda b/Cubical/Experiments/Brunerie.agda index b2b130892..f32633564 100644 --- a/Cubical/Experiments/Brunerie.agda +++ b/Cubical/Experiments/Brunerie.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical #-} +{-# OPTIONS --safe #-} module Cubical.Experiments.Brunerie where open import Cubical.Foundations.Everything @@ -11,9 +11,10 @@ open import Cubical.HITs.S2 open import Cubical.HITs.S3 open import Cubical.HITs.Join open import Cubical.HITs.Hopf -open import Cubical.HITs.SetTruncation -open import Cubical.HITs.GroupoidTruncation -open import Cubical.HITs.2GroupoidTruncation +open import Cubical.HITs.SetTruncation as SetTrunc +open import Cubical.HITs.GroupoidTruncation as GroupoidTrunc +open import Cubical.HITs.2GroupoidTruncation as 2GroupoidTrunc +open import Cubical.Homotopy.Loopspace -- This code is adapted from examples/brunerie3.ctt on the pi4s3_nobug branch of cubicaltt @@ -23,9 +24,9 @@ S¹∙ = (S¹ , base) S²∙ = (S² , base) S³∙ = (S³ , base) -∥_∥₁∙ ∥_∥₂∙ : Pointed₀ → Pointed₀ -∥ A , a ∥₁∙ = ∥ A ∥₁ , ∣ a ∣₁ -∥ A , a ∥₂∙ = ∥ A ∥₂ , ∣ a ∣₂ +∥_∥₃∙ ∥_∥₄∙ : Pointed₀ → Pointed₀ +∥ A , a ∥₃∙ = ∥ A ∥₃ , ∣ a ∣₃ +∥ A , a ∥₄∙ = ∥ A ∥₄ , ∣ a ∣₄ join∙ : Pointed₀ → Type₀ → Pointed₀ join∙ (A , a) B = join A B , inl a @@ -126,46 +127,46 @@ fibContrΩ³Hopf p i j k = h : Ω³ S²∙ .fst → Ω³ (join∙ S¹∙ S¹) .fst h p i j k = tee (p i j k) (fibContrΩ³Hopf p i j k) -multTwoAux : (x : S²) → Path (Path ∥ S² ∥₂ ∣ x ∣₂ ∣ x ∣₂) refl refl -multTwoAux base i j = ∣ surf i j ∣₂ +multTwoAux : (x : S²) → Path (Path ∥ S² ∥₄ ∣ x ∣₄ ∣ x ∣₄) refl refl +multTwoAux base i j = ∣ surf i j ∣₄ multTwoAux (surf k l) i j = hcomp (λ m → λ - { (i = i0) → ∣ surf k l ∣₂ - ; (i = i1) → ∣ surf k l ∣₂ - ; (j = i0) → ∣ surf k l ∣₂ - ; (j = i1) → ∣ surf k l ∣₂ - ; (k = i0) → ∣ surf i j ∣₂ - ; (k = i1) → ∣ surf i j ∣₂ - ; (l = i0) → ∣ surf i j ∣₂ - ; (l = i1) → squash₂ _ _ _ _ _ _ (λ k i j → step₁ k i j) refl m k i j + { (i = i0) → ∣ surf k l ∣₄ + ; (i = i1) → ∣ surf k l ∣₄ + ; (j = i0) → ∣ surf k l ∣₄ + ; (j = i1) → ∣ surf k l ∣₄ + ; (k = i0) → ∣ surf i j ∣₄ + ; (k = i1) → ∣ surf i j ∣₄ + ; (l = i0) → ∣ surf i j ∣₄ + ; (l = i1) → squash₄ _ _ _ _ _ _ (λ k i j → step₁ k i j) refl m k i j }) (step₁ k i j) where - step₁ : I → I → I → ∥ S² ∥₂ + step₁ : I → I → I → ∥ S² ∥₄ step₁ k i j = - hcomp {A = ∥ S² ∥₂} + hcomp {A = ∥ S² ∥₄} (λ m → λ - { (i = i0) → ∣ surf k (l ∧ m) ∣₂ - ; (i = i1) → ∣ surf k (l ∧ m) ∣₂ - ; (j = i0) → ∣ surf k (l ∧ m) ∣₂ - ; (j = i1) → ∣ surf k (l ∧ m) ∣₂ - ; (k = i0) → ∣ surf i j ∣₂ - ; (k = i1) → ∣ surf i j ∣₂ - ; (l = i0) → ∣ surf i j ∣₂ + { (i = i0) → ∣ surf k (l ∧ m) ∣₄ + ; (i = i1) → ∣ surf k (l ∧ m) ∣₄ + ; (j = i0) → ∣ surf k (l ∧ m) ∣₄ + ; (j = i1) → ∣ surf k (l ∧ m) ∣₄ + ; (k = i0) → ∣ surf i j ∣₄ + ; (k = i1) → ∣ surf i j ∣₄ + ; (l = i0) → ∣ surf i j ∣₄ }) - ∣ surf i j ∣₂ + ∣ surf i j ∣₄ -multTwoTildeAux : (t : ∥ S² ∥₂) → Path (Path ∥ S² ∥₂ t t) refl refl -multTwoTildeAux ∣ x ∣₂ = multTwoAux x -multTwoTildeAux (squash₂ _ _ _ _ _ _ t u k l m n) i j = - squash₂ _ _ _ _ _ _ +multTwoTildeAux : (t : ∥ S² ∥₄) → Path (Path ∥ S² ∥₄ t t) refl refl +multTwoTildeAux ∣ x ∣₄ = multTwoAux x +multTwoTildeAux (squash₄ _ _ _ _ _ _ t u k l m n) i j = + squash₄ _ _ _ _ _ _ (λ k l m → multTwoTildeAux (t k l m) i j) (λ k l m → multTwoTildeAux (u k l m) i j) k l m n -multTwoEquivAux : Path (Path (∥ S² ∥₂ ≃ ∥ S² ∥₂) (idEquiv _) (idEquiv _)) refl refl +multTwoEquivAux : Path (Path (∥ S² ∥₄ ≃ ∥ S² ∥₄) (idEquiv _) (idEquiv _)) refl refl multTwoEquivAux i j = ( f i j , hcomp @@ -182,41 +183,41 @@ multTwoEquivAux i j = (transp (λ k → isEquiv (f i (j ∧ k))) (i ∨ ~ i ∨ ~ j) (idIsEquiv _)) ) where - f : I → I → ∥ S² ∥₂ → ∥ S² ∥₂ + f : I → I → ∥ S² ∥₄ → ∥ S² ∥₄ f i j t = multTwoTildeAux t i j tHopf³ : S³ → Type₀ -tHopf³ base = ∥ S² ∥₂ +tHopf³ base = ∥ S² ∥₄ tHopf³ (surf i j k) = - Glue ∥ S² ∥₂ - (λ { (i = i0) → (∥ S² ∥₂ , idEquiv _) - ; (i = i1) → (∥ S² ∥₂ , idEquiv _) - ; (j = i0) → (∥ S² ∥₂ , idEquiv _) - ; (j = i1) → (∥ S² ∥₂ , idEquiv _) - ; (k = i0) → (∥ S² ∥₂ , multTwoEquivAux i j) - ; (k = i1) → (∥ S² ∥₂ , idEquiv _) + Glue ∥ S² ∥₄ + (λ { (i = i0) → (∥ S² ∥₄ , idEquiv _) + ; (i = i1) → (∥ S² ∥₄ , idEquiv _) + ; (j = i0) → (∥ S² ∥₄ , idEquiv _) + ; (j = i1) → (∥ S² ∥₄ , idEquiv _) + ; (k = i0) → (∥ S² ∥₄ , multTwoEquivAux i j) + ; (k = i1) → (∥ S² ∥₄ , idEquiv _) }) -π₃S³ : Ω³ S³∙ .fst → Ω² ∥ S²∙ ∥₂∙ .fst -π₃S³ p i j = transp (λ k → tHopf³ (p j k i)) i0 ∣ base ∣₂ +π₃S³ : Ω³ S³∙ .fst → Ω² ∥ S²∙ ∥₄∙ .fst +π₃S³ p i j = transp (λ k → tHopf³ (p j k i)) i0 ∣ base ∣₄ codeS² : S² → hGroupoid _ -codeS² s = ∥ HopfS² s ∥₁ , squash₁ +codeS² s = ∥ HopfS² s ∥₃ , squash₃ -codeTruncS² : ∥ S² ∥₂ → hGroupoid _ -codeTruncS² = rec2GroupoidTrunc (isOfHLevelHLevel 3) codeS² +codeTruncS² : ∥ S² ∥₄ → hGroupoid _ +codeTruncS² = 2GroupoidTrunc.rec (isOfHLevelTypeOfHLevel 3) codeS² -encodeTruncS² : Ω ∥ S²∙ ∥₂∙ .fst → ∥ S¹ ∥₁ -encodeTruncS² p = transp (λ i → codeTruncS² (p i) .fst) i0 ∣ base ∣₁ +encodeTruncS² : Ω ∥ S²∙ ∥₄∙ .fst → ∥ S¹ ∥₃ +encodeTruncS² p = transp (λ i → codeTruncS² (p i) .fst) i0 ∣ base ∣₃ codeS¹ : S¹ → hSet _ -codeS¹ s = ∥ helix s ∥₀ , squash₀ +codeS¹ s = ∥ helix s ∥₂ , squash₂ -codeTruncS¹ : ∥ S¹ ∥₁ → hSet _ -codeTruncS¹ = recGroupoidTrunc (isOfHLevelHLevel 2) codeS¹ +codeTruncS¹ : ∥ S¹ ∥₃ → hSet _ +codeTruncS¹ = GroupoidTrunc.rec (isOfHLevelTypeOfHLevel 2) codeS¹ -encodeTruncS¹ : Ω ∥ S¹∙ ∥₁∙ .fst → ∥ Int ∥₀ -encodeTruncS¹ p = transp (λ i → codeTruncS¹ (p i) .fst) i0 ∣ pos zero ∣₀ +encodeTruncS¹ : Ω ∥ S¹∙ ∥₃∙ .fst → ∥ ℤ ∥₂ +encodeTruncS¹ p = transp (λ i → codeTruncS¹ (p i) .fst) i0 ∣ pos zero ∣₂ -- THE BIG GAME @@ -233,25 +234,25 @@ f5 = h f6 : Ω³ (join∙ S¹∙ S¹) .fst → Ω³ S³∙ .fst f6 = mapΩ³refl joinS¹S¹→S³ -f7 : Ω³ S³∙ .fst → Ω² ∥ S²∙ ∥₂∙ .fst +f7 : Ω³ S³∙ .fst → Ω² ∥ S²∙ ∥₄∙ .fst f7 = π₃S³ -g8 : Ω² ∥ S²∙ ∥₂∙ .fst → Ω ∥ S¹∙ ∥₁∙ .fst +g8 : Ω² ∥ S²∙ ∥₄∙ .fst → Ω ∥ S¹∙ ∥₃∙ .fst g8 = mapΩrefl encodeTruncS² -g9 : Ω ∥ S¹∙ ∥₁∙ .fst → ∥ Int ∥₀ +g9 : Ω ∥ S¹∙ ∥₃∙ .fst → ∥ ℤ ∥₂ g9 = encodeTruncS¹ -g10 : ∥ Int ∥₀ → Int -g10 = elimSetTrunc (λ _ → isSetInt) (idfun Int) +g10 : ∥ ℤ ∥₂ → ℤ +g10 = SetTrunc.rec isSetℤ (idfun ℤ) -- don't run me -brunerie : Int +brunerie : ℤ brunerie = g10 (g9 (g8 (f7 (f6 (f5 (f4 (f3 (λ i j k → surf i j k)))))))) -- simpler tests -test63 : ℕ → Int +test63 : ℕ → ℤ test63 n = g10 (g9 (g8 (f7 (63n n)))) where 63n : ℕ → Ω³ S³∙ .fst @@ -293,5 +294,5 @@ sorghum i j k = }) (surf k i)) -goo : Ω³ S²∙ .fst → Int +goo : Ω³ S²∙ .fst → ℤ goo x = g10 (g9 (g8 (f7 (f6 (f5 x))))) diff --git a/Cubical/Experiments/CohomologyGroups.agda b/Cubical/Experiments/CohomologyGroups.agda new file mode 100644 index 000000000..1f2403aac --- /dev/null +++ b/Cubical/Experiments/CohomologyGroups.agda @@ -0,0 +1,140 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.CohomologyGroups where + +open import Cubical.Experiments.ZCohomologyOld.Base +open import Cubical.Experiments.ZCohomologyOld.Properties +open import Cubical.Experiments.ZCohomologyOld.MayerVietorisUnreduced +open import Cubical.Experiments.ZCohomologyOld.Groups.Unit +open import Cubical.Experiments.ZCohomologyOld.KcompPrelims +open import Cubical.Experiments.ZCohomologyOld.Groups.Sn + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Isomorphism + +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) + hiding (map) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; ∥_∥ to ∥_∥₁ ; ∣_∣ to ∣_∣₁) + hiding (map) + +open import Cubical.Data.Bool +open import Cubical.Data.Sigma +open import Cubical.Data.Int + +open import Cubical.Algebra.Group + +open GroupIso +open GroupHom +open BijectionIso + +-- --------------------------H¹(S¹) ----------------------------------- +{- +In order to apply Mayer-Vietoris, we need the following lemma. +Given the following diagram + a ↦ (a , 0) ψ ϕ + A --> A × A -------> B ---> C +If ψ is an isomorphism and ϕ is surjective with ker ϕ ≡ {ψ (a , a) ∣ a ∈ A}, then C ≅ B +-} + + +diagonalIso : ∀ {ℓ ℓ' ℓ''} {A : Group {ℓ}} (B : Group {ℓ'}) {C : Group {ℓ''}} + (ψ : GroupIso (dirProd A A) B) (ϕ : GroupHom B C) + → isSurjective _ _ ϕ + → ((x : ⟨ B ⟩) → isInKer B C ϕ x + → ∃[ y ∈ ⟨ A ⟩ ] x ≡ (fun (map ψ)) (y , y)) + → ((x : ⟨ B ⟩) → (∃[ y ∈ ⟨ A ⟩ ] x ≡ (fun (map ψ)) (y , y)) + → isInKer B C ϕ x) + → GroupIso A C +diagonalIso {A = A} B {C = C} ψ ϕ issurj ker→diag diag→ker = BijectionIsoToGroupIso bijIso + where + open GroupStr + module A = GroupStr (snd A) + module B = GroupStr (snd B) + module C = GroupStr (snd C) + module A×A = GroupStr (snd (dirProd A A)) + module ψ = GroupIso ψ + module ϕ = GroupHom ϕ + ψ⁻ = inv ψ + + fstProj : GroupHom A (dirProd A A) + fun fstProj a = a , GroupStr.0g (snd A) + isHom fstProj g0 g1 i = (g0 A.+ g1) , GroupStr.lid (snd A) (GroupStr.0g (snd A)) (~ i) + + bijIso : BijectionIso A C + map' bijIso = compGroupHom fstProj (compGroupHom (map ψ) ϕ) + inj bijIso a inker = pRec (isSetCarrier A _ _) + (λ {(a' , id) → (cong fst (sym (leftInv ψ (a , GroupStr.0g (snd A))) ∙∙ cong ψ⁻ id ∙∙ leftInv ψ (a' , a'))) + ∙ cong snd (sym (leftInv ψ (a' , a')) ∙∙ cong ψ⁻ (sym id) ∙∙ leftInv ψ (a , GroupStr.0g (snd A)))}) + (ker→diag _ inker) + surj bijIso c = + pRec isPropPropTrunc + (λ { (b , id) → ∣ (fst (ψ⁻ b) A.+ (A.- snd (ψ⁻ b))) + , ((sym (GroupStr.rid (snd C) _) + ∙∙ cong ((fun ϕ) ((fun (map ψ)) (fst (ψ⁻ b) A.+ (A.- snd (ψ⁻ b)) , GroupStr.0g (snd A))) C.+_) + (sym (diag→ker (fun (map ψ) ((snd (ψ⁻ b)) , (snd (ψ⁻ b)))) + ∣ (snd (ψ⁻ b)) , refl ∣₁)) + ∙∙ sym ((isHom ϕ) _ _)) + ∙∙ cong (fun ϕ) (sym ((isHom (map ψ)) _ _) + ∙∙ cong (fun (map ψ)) (ΣPathP (sym (GroupStr.assoc (snd A) _ _ _) + ∙∙ cong (fst (ψ⁻ b) A.+_) (GroupStr.invl (snd A) _) + ∙∙ GroupStr.rid (snd A) _ + , (GroupStr.lid (snd A) _))) + ∙∙ rightInv ψ b) + ∙∙ id) ∣₁ }) + (issurj c) + +H¹-S¹≅ℤ : GroupIso intGroup (coHomGr 1 (S₊ 1)) +H¹-S¹≅ℤ = + diagonalIso (coHomGr 0 (S₊ 0)) + (invGroupIso H⁰-S⁰≅ℤ×ℤ) + (K.d 0) + (λ x → K.Ker-i⊂Im-d 0 x + (ΣPathP (isOfHLevelSuc 0 (isContrHⁿ-Unit 0) _ _ + , isOfHLevelSuc 0 (isContrHⁿ-Unit 0) _ _))) + ((sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 isPropPropTrunc) + (λ x inker + → pRec isPropPropTrunc + (λ {((f , g) , id') → helper x f g id' inker}) + ((K.Ker-d⊂Im-Δ 0 ∣ x ∣₂ inker))))) + ((sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ F surj + → pRec (isSetSetTrunc _ _) + (λ { (x , id) → K.Im-Δ⊂Ker-d 0 ∣ F ∣₂ + ∣ (∣ (λ _ → x) ∣₂ , ∣ (λ _ → 0) ∣₂) , + (cong ∣_∣₂ (funExt (surjHelper x))) ∙ sym id ∣₁ }) + surj) ) + □ invGroupIso (coHomPushout≅coHomSn 0 1) + where + module K = MV Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) + + surjHelper : (x : Int) (x₁ : S₊ 0) → x -[ 0 ]ₖ 0 ≡ S0→Int (x , x) x₁ + surjHelper x true = Iso.leftInv (Iso-Kn-ΩKn+1 0) x + surjHelper x false = Iso.leftInv (Iso-Kn-ΩKn+1 0) x + + helper : (F : S₊ 0 → Int) (f g : ∥ (Unit → Int) ∥₂) + (id : GroupHom.fun (K.Δ 0) (f , g) ≡ ∣ F ∣₂) + → isInKer (coHomGr 0 (S₊ 0)) + (coHomGr 1 (Pushout (λ _ → tt) (λ _ → tt))) + (K.d 0) + ∣ F ∣₂ + → ∃[ x ∈ Int ] ∣ F ∣₂ ≡ inv H⁰-S⁰≅ℤ×ℤ (x , x) + helper F = + sElim2 (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 isPropPropTrunc) + λ f g id inker + → pRec isPropPropTrunc + (λ ((a , b) , id2) + → sElim2 {C = λ f g → GroupHom.fun (K.Δ 0) (f , g) ≡ ∣ F ∣₂ → _ } + (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 isPropPropTrunc) + (λ f g id → ∣ (helper2 f g .fst) , (sym id ∙ sym (helper2 f g .snd)) ∣₁) + a b id2) + (MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₂ inker) + where + helper2 : (f g : Unit → Int) + → Σ[ x ∈ Int ] (inv H⁰-S⁰≅ℤ×ℤ (x , x)) + ≡ GroupHom.fun (K.Δ 0) (∣ f ∣₂ , ∣ g ∣₂) + helper2 f g = (f _ -[ 0 ]ₖ g _) , cong ∣_∣₂ (funExt λ {true → refl ; false → refl}) diff --git a/Cubical/Experiments/EscardoSIP.agda b/Cubical/Experiments/EscardoSIP.agda index 31ffce97f..56f3acfa6 100644 --- a/Cubical/Experiments/EscardoSIP.agda +++ b/Cubical/Experiments/EscardoSIP.agda @@ -5,22 +5,18 @@ https://www.cs.bham.ac.uk/~mhe/HoTT-UF-in-Agda-Lecture-Notes/HoTT-UF-Agda.html#s All the needed preliminary results from the lecture notes are stated and proven in this file. It would be interesting to compare the proves with the one in Cubical.Foundations.SIP -} - - - -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Experiments.EscardoSIP where open import Cubical.Core.Everything open import Cubical.Foundations.Everything -open import Cubical.Foundations.HAEquiv +open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Data.Sigma.Properties - private - variable - ℓ ℓ' ℓ'' : Level - + variable + ℓ ℓ' ℓ'' : Level + S : Type ℓ → Type ℓ' -- We prove several useful equalities and equivalences between Σ-types all the proofs are taken from -- Martin Hötzel-Escardó's lecture notes. @@ -29,7 +25,7 @@ private Σ-≡-≃ : {X : Type ℓ} {A : X → Type ℓ'} → (σ τ : Σ X A) → ((σ ≡ τ) ≃ (Σ[ p ∈ (σ .fst) ≡ (τ .fst) ] (subst A p (σ .snd) ≡ (τ .snd)))) -Σ-≡-≃ {A = A} σ τ = pathToEquiv (pathSigma≡sigmaPath σ τ) +Σ-≡-≃ {A = A} σ τ = invEquiv (ΣPathTransport≃PathΣ σ τ) @@ -59,10 +55,10 @@ NatΣ τ (x , a) = (x , τ x a) g x = equivFun (invEquiv (φ x)) η : (x : X) → (a : A x) → (g x) ((f x) a) ≡ a - η x = retEq (invEquiv (φ x)) + η x = secEq (invEquiv (φ x)) ε : (x : X) → (b : B x) → f x (g x b) ≡ b - ε x = secEq (invEquiv (φ x)) + ε x = retEq (invEquiv (φ x)) NatΣ-η : (w : Σ X A) → NatΣ g (NatΣ f w) ≡ w NatΣ-η (x , a) = (x , g x (f x a)) ≡⟨ Σ-to-PathP (η x a) ⟩ @@ -82,9 +78,9 @@ NatΣ τ (x , a) = (x , τ x a) g : Y → X g = isHAEquiv.g isHAEquivf ε : (x : X) → (g (f x)) ≡ x - ε = isHAEquiv.sec isHAEquivf + ε = isHAEquiv.linv isHAEquivf η : (y : Y) → f (g y) ≡ y - η = isHAEquiv.ret isHAEquivf + η = isHAEquiv.rinv isHAEquivf τ : (x : X) → cong f (ε x) ≡ η (f x) τ = isHAEquiv.com isHAEquivf @@ -95,23 +91,23 @@ NatΣ τ (x , a) = (x , τ x a) ψ (y , a) = (g y , subst A (sym (η y)) a) φψ : (z : (Σ Y A)) → φ (ψ z) ≡ z - φψ (y , a) = sigmaPath→pathSigma (φ (ψ (y , a))) (y , a) - (η y , transportTransport⁻ (λ i → A (η y i)) a) + φψ (y , a) = + ΣPathTransport→PathΣ _ _ (η y , transportTransport⁻ (λ i → A (η y i)) a) -- last term proves transp (λ i → A (η y i)) i0 (transp (λ i → A (η y (~ i))) i0 a) ≡ a ψφ : (z : (Σ X (A ∘ f))) → ψ (φ z) ≡ z - ψφ (x , a) = sigmaPath→pathSigma (ψ (φ (x , a))) (x , a) (ε x , q) + ψφ (x , a) = ΣPathTransport→PathΣ _ _ (ε x , q) where b : A (f (g (f x))) b = (transp (λ i → A (η (f x) (~ i))) i0 a) q : transp (λ i → A (f (ε x i))) i0 (transp (λ i → A (η (f x) (~ i))) i0 a) ≡ a - q = transp (λ i → A (f (ε x i))) i0 b ≡⟨ S ⟩ + q = transp (λ i → A (f (ε x i))) i0 b ≡⟨ i ⟩ transp (λ i → A (η (f x) i)) i0 b ≡⟨ transportTransport⁻ (λ i → A (η (f x) i)) a ⟩ a ∎ where - S : (transp (λ i → A (f (ε x i))) i0 b) ≡ (transp (λ i → A (η (f x) i)) i0 b) - S = subst (λ p → (transp (λ i → A (f (ε x i))) i0 b) ≡ (transp (λ i → A (p i)) i0 b)) + i : (transp (λ i → A (f (ε x i))) i0 b) ≡ (transp (λ i → A (η (f x) i)) i0 b) + i = subst (λ p → (transp (λ i → A (f (ε x i))) i0 b) ≡ (transp (λ i → A (p i)) i0 b)) (τ x) refl @@ -122,28 +118,6 @@ NatΣ τ (x , a) = (x , τ x a) isoToEquiv (Σ-change-of-variable-Iso f (equiv→HAEquiv (f , isEquivf) .snd)) - - -Σ-assoc-Iso : (X : Type ℓ) (A : X → Type ℓ') (P : Σ X A → Type ℓ'') - → (Iso (Σ (Σ X A) P) (Σ[ x ∈ X ] (Σ[ a ∈ A x ] P (x , a)))) -Σ-assoc-Iso X A P = iso f g ε η - where - f : (Σ (Σ X A) P) → (Σ[ x ∈ X ] (Σ[ a ∈ A x ] P (x , a))) - f ((x , a) , p) = (x , (a , p)) - g : (Σ[ x ∈ X ] (Σ[ a ∈ A x ] P (x , a))) → (Σ (Σ X A) P) - g (x , (a , p)) = ((x , a) , p) - ε : section f g - ε n = refl - η : retract f g - η m = refl - -Σ-assoc-≃ : (X : Type ℓ) (A : X → Type ℓ') (P : Σ X A → Type ℓ'') - → (Σ (Σ X A) P) ≃ (Σ[ x ∈ X ] (Σ[ a ∈ A x ] P (x , a))) -Σ-assoc-≃ X A P = isoToEquiv (Σ-assoc-Iso X A P) - - - - -- A structure is a type-family S : Type ℓ → Type ℓ', i.e. for X : Type ℓ and s : S X, the pair (X , s) -- means that X is equipped with a S-structure, which is witnessed by s. -- An S-structure should have a notion of S-homomorphism, or rather S-isomorphism. @@ -155,84 +129,69 @@ NatΣ τ (x , a) = (x , τ x a) -- as S-structures. This we call a standard notion of structure. -SNS : (S : Type ℓ → Type ℓ') - → ((A B : Σ[ X ∈ (Type ℓ) ] (S X)) → ((A .fst) ≃ (B .fst)) → Type ℓ'') - → Type (ℓ-max (ℓ-max(ℓ-suc ℓ) ℓ') ℓ'') +SNS : (S : Type ℓ → Type ℓ') (ι : StrEquiv S ℓ'') → Type (ℓ-max (ℓ-max (ℓ-suc ℓ) ℓ') ℓ'') SNS {ℓ = ℓ} S ι = ∀ {X : (Type ℓ)} (s t : S X) → ((s ≡ t) ≃ ι (X , s) (X , t) (idEquiv X)) -- Escardo's ρ can actually be defined from this: -ρ : {S : Type ℓ → Type ℓ'} - → {ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → ((A .fst) ≃ (B .fst)) → Type ℓ''} - → (SNS S ι) - → (A : Σ[ X ∈ (Type ℓ) ] (S X)) → (ι A A (idEquiv (A .fst))) -ρ θ A = equivFun (θ (A .snd) (A .snd)) refl +ρ : {ι : StrEquiv S ℓ''} (θ : SNS S ι) (A : TypeWithStr ℓ S) → (ι A A (idEquiv (typ A))) +ρ θ A = equivFun (θ (str A) (str A)) refl -- We introduce the notation a bit differently: -_≃[_]_ : {S : Type ℓ → Type ℓ'} → (Σ[ X ∈ (Type ℓ) ] (S X)) - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → ((A .fst) ≃ (B .fst)) → Type ℓ'') - → (Σ[ X ∈ (Type ℓ) ] (S X)) - → (Type (ℓ-max ℓ ℓ'')) -A ≃[ ι ] B = Σ[ f ∈ ((A .fst) ≃ (B. fst)) ] (ι A B f) +_≃[_]_ : (A : TypeWithStr ℓ S) (ι : StrEquiv S ℓ'') (B : TypeWithStr ℓ S) → (Type (ℓ-max ℓ ℓ'')) +A ≃[ ι ] B = Σ[ f ∈ ((typ A) ≃ (typ B)) ] (ι A B f) -Id→homEq : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → ((A .fst) ≃ (B .fst)) → Type ℓ'') - → (ρ : (A : Σ[ X ∈ (Type ℓ) ] (S X)) → ι A A (idEquiv (A .fst))) - → (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A ≡ B → (A ≃[ ι ] B) -Id→homEq S ι ρ A B p = J (λ y x → A ≃[ ι ] y) (idEquiv (A .fst) , ρ A) p +Id→homEq : (S : Type ℓ → Type ℓ') (ι : StrEquiv S ℓ'') + → (ρ : (A : TypeWithStr ℓ S) → ι A A (idEquiv (typ A))) + → (A B : TypeWithStr ℓ S) → A ≡ B → (A ≃[ ι ] B) +Id→homEq S ι ρ A B p = J (λ y x → A ≃[ ι ] y) (idEquiv (typ A) , ρ A) p -- Use a PathP version of Escardó's homomorphism-lemma -hom-lemma-dep : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → ((A .fst) ≃ (B .fst)) → Type ℓ'') - → (θ : SNS S ι) - → (A B : Σ[ X ∈ (Type ℓ) ] (S X)) - → (p : (A .fst) ≡ (B. fst)) - → (PathP (λ i → S (p i)) (A .snd) (B .snd)) ≃ (ι A B (pathToEquiv p)) -hom-lemma-dep S ι θ A B p = (J P (λ s → γ s) p) (B .snd) +hom-lemma-dep : (S : Type ℓ → Type ℓ') (ι : StrEquiv S ℓ'') (θ : SNS S ι) + → (A B : TypeWithStr ℓ S) + → (p : (typ A) ≡ (typ B)) + → (PathP (λ i → S (p i)) (str A) (str B)) ≃ (ι A B (pathToEquiv p)) +hom-lemma-dep S ι θ A B p = (J P (λ s → γ s) p) (str B) where - P = (λ y x → (s : S y) → PathP (λ i → S (x i)) (A .snd) s ≃ ι A (y , s) (pathToEquiv x)) + P = (λ y x → (s : S y) → PathP (λ i → S (x i)) (str A) s ≃ ι A (y , s) (pathToEquiv x)) - γ : (s : S (A .fst)) → ((A .snd) ≡ s) ≃ ι A ((A .fst) , s) (pathToEquiv refl) - γ s = subst (λ f → ((A .snd) ≡ s) ≃ ι A ((A .fst) , s) f) (sym pathToEquivRefl) (θ (A. snd) s) + γ : (s : S (typ A)) → ((str A) ≡ s) ≃ ι A ((typ A) , s) (pathToEquiv refl) + γ s = subst (λ f → ((str A) ≡ s) ≃ ι A ((typ A) , s) f) (sym pathToEquivRefl) (θ (str A) s) -- Define the inverse of Id→homEq directly. ua-lemma : (A B : Type ℓ) (e : A ≃ B) → (pathToEquiv (ua e)) ≡ e -ua-lemma A B e = EquivJ (λ b a f → (pathToEquiv (ua f)) ≡ f) - (λ x → subst (λ r → pathToEquiv r ≡ idEquiv x) (sym uaIdEquiv) pathToEquivRefl) - B A e +ua-lemma A B e = EquivJ (λ A f → (pathToEquiv (ua f)) ≡ f) + (subst (λ r → pathToEquiv r ≡ idEquiv B) (sym uaIdEquiv) pathToEquivRefl) + e -homEq→Id : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → ((A .fst) ≃ (B .fst)) → Type ℓ'') - → (θ : SNS S ι) - → (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → (A ≃[ ι ] B) → A ≡ B +homEq→Id : (S : Type ℓ → Type ℓ') (ι : StrEquiv S ℓ'') (θ : SNS S ι) + → (A B : TypeWithStr ℓ S) → (A ≃[ ι ] B) → A ≡ B homEq→Id S ι θ A B (f , φ) = ΣPathP (p , q) where p = ua f ψ : ι A B (pathToEquiv p) - ψ = subst (λ g → ι A B g) (sym (ua-lemma (A .fst) (B. fst) f)) φ + ψ = subst (λ g → ι A B g) (sym (ua-lemma (typ A) (typ B) f)) φ - q : PathP (λ i → S (p i)) (A .snd) (B .snd) + q : PathP (λ i → S (p i)) (str A) (str B) q = equivFun (invEquiv (hom-lemma-dep S ι θ A B p)) ψ -- Proof of the SIP: -SIP : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → ((A .fst) ≃ (B .fst)) → Type ℓ'') - → (θ : SNS S ι) - → (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → ((A ≡ B) ≃ (A ≃[ ι ] B)) +SIP : (S : Type ℓ → Type ℓ') (ι : StrEquiv S ℓ'') (θ : SNS S ι) + → (A B : TypeWithStr ℓ S) → ((A ≡ B) ≃ (A ≃[ ι ] B)) SIP S ι θ A B = - (A ≡ B) ≃⟨ i ⟩ - (Σ[ p ∈ (A .fst) ≡ (B. fst) ] PathP (λ i → S (p i)) (A .snd) (B .snd)) ≃⟨ ii ⟩ - (Σ[ p ∈ (A .fst) ≡ (B. fst) ] (ι A B (pathToEquiv p))) ≃⟨ iii ⟩ + (A ≡ B) ≃⟨ i ⟩ + (Σ[ p ∈ (typ A) ≡ (typ B) ] PathP (λ i → S (p i)) (str A) (str B)) ≃⟨ ii ⟩ + (Σ[ p ∈ (typ A) ≡ (typ B) ] (ι A B (pathToEquiv p))) ≃⟨ iii ⟩ (A ≃[ ι ] B) ■ where - i = invEquiv Σ≡ + i = invEquiv ΣPath≃PathΣ ii = Σ-cong-≃ (hom-lemma-dep S ι θ A B) iii = Σ-change-of-variable-≃ pathToEquiv (equivIsEquiv univalence) diff --git a/Cubical/Experiments/Everything.agda b/Cubical/Experiments/Everything.agda index f31eafb9f..570538491 100644 --- a/Cubical/Experiments/Everything.agda +++ b/Cubical/Experiments/Everything.agda @@ -1,11 +1,16 @@ -- Export only the experiments that are expected to compile (without -- any holes) -{-# OPTIONS --cubical #-} module Cubical.Experiments.Everything where open import Cubical.Experiments.Brunerie public open import Cubical.Experiments.EscardoSIP public open import Cubical.Experiments.Generic public +open import Cubical.Experiments.NatMinusTwo +open import Cubical.Experiments.IsoInt +open import Cubical.Experiments.HAEquivInt open import Cubical.Experiments.Problem open import Cubical.Experiments.FunExtFromUA public - +open import Cubical.Experiments.HoTT-UF +open import Cubical.Experiments.ZCohomologyOld.Base +open import Cubical.Experiments.ZCohomologyOld.KcompPrelims +open import Cubical.Experiments.ZCohomologyOld.Properties diff --git a/Cubical/Experiments/FunExtFromUA.agda b/Cubical/Experiments/FunExtFromUA.agda index a1a8c460e..2cf2ae5cf 100644 --- a/Cubical/Experiments/FunExtFromUA.agda +++ b/Cubical/Experiments/FunExtFromUA.agda @@ -1,6 +1,6 @@ {- Voevodsky's proof that univalence implies funext -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Experiments.FunExtFromUA where @@ -15,27 +15,22 @@ funext : ∀ ℓ ℓ' → Type (ℓ-suc(ℓ-max ℓ ℓ')) funext ℓ ℓ' = {X : Type ℓ} {Y : Type ℓ'} {f g : X → Y} → f ∼ g → f ≡ g -elimEquivFun' : ∀ {ℓ} (P : (A B : Type ℓ) → (A → B) → Type ℓ) - → (r : (B : Type ℓ) → P B B (λ x → x)) - → (A B : Type ℓ) → (e : A ≃ B) → P A B (e .fst) -elimEquivFun' P r A B = elimEquivFun B (λ A → P A B) (r B) A - pre-comp-is-equiv : (X Y : Type ℓ) (f : X → Y) (Z : Type ℓ) → isEquiv f → isEquiv (λ (g : Y → Z) → g ∘ f) -pre-comp-is-equiv {ℓ} X Y f Z e = elimEquivFun' P r X Y (f , e) +pre-comp-is-equiv {ℓ} X Y f Z e = elimEquivFun P r (f , e) where - P : (X Y : Type ℓ) → (X → Y) → Type ℓ - P X Y f = isEquiv (λ (g : Y → Z) → g ∘ f) - r : (B : Type ℓ) → P B B (λ x → x) - r B = idIsEquiv (B → Z) + P : (X : Type ℓ) → (X → Y) → Type ℓ + P X f = isEquiv (λ (g : Y → Z) → g ∘ f) + r : P Y (λ x → x) + r = idIsEquiv (Y → Z) leftCancellable : {X : Type ℓ} {Y : Type ℓ'} → (X → Y) → Type (ℓ-max ℓ ℓ') leftCancellable f = ∀ {x x'} → f x ≡ f x' → x ≡ x' equivLC : {X : Type ℓ} {Y : Type ℓ'} (f : X → Y) → isEquiv f → leftCancellable f -equivLC f e {x} {x'} p i = hcomp (λ j → \ {(i = i0) → secEq (f , e) x j ; - (i = i1) → secEq (f , e) x' j}) +equivLC f e {x} {x'} p i = hcomp (λ j → \ {(i = i0) → retEq (f , e) x j ; + (i = i1) → retEq (f , e) x' j}) (invEq (f , e) (p i)) univalence-gives-funext : funext ℓ' ℓ diff --git a/Cubical/Experiments/Generic.agda b/Cubical/Experiments/Generic.agda index 776fe379a..97a247578 100644 --- a/Cubical/Experiments/Generic.agda +++ b/Cubical/Experiments/Generic.agda @@ -1,6 +1,6 @@ -- Two fun examples of generic programming using univalence -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Experiments.Generic where open import Agda.Builtin.String @@ -86,35 +86,35 @@ data Company (A : Type₀) : Type₀ where -- A small example -anders : Employee Int +anders : Employee ℤ anders = E (P "Anders" "Pittsburgh") (S (pos 2500)) -andrea : Employee Int +andrea : Employee ℤ andrea = E (P "Andrea" "Copenhagen") (S (pos 2000)) -andreas : Employee Int +andreas : Employee ℤ andreas = E (P "Andreas" "Gothenburg") (S (pos 3000)) -- For now we have a small company -genCom : Company Int +genCom : Company ℤ genCom = C ( D "Research" andreas (PU anders ∷ PU andrea ∷ []) ∷ []) -- Increase the salary for everyone by 1 -incSalary : Company Int → Company Int -incSalary c = transport (λ i → Company (sucPathInt i)) c +incSalary : Company ℤ → Company ℤ +incSalary c = transport (λ i → Company (sucPathℤ i)) c -genCom1 : Company Int +genCom1 : Company ℤ genCom1 = incSalary genCom -- Increase the salary more -incSalaryℕ : ℕ → Company Int → Company Int +incSalaryℕ : ℕ → Company ℤ → Company ℤ incSalaryℕ n c = transport (λ i → Company (addEq n i)) c -genCom2 : Company Int +genCom2 : Company ℤ genCom2 = incSalaryℕ 2 genCom -genCom10 : Company Int +genCom10 : Company ℤ genCom10 = incSalaryℕ 10 genCom diff --git a/Cubical/Experiments/HAEquivInt.agda b/Cubical/Experiments/HAEquivInt.agda new file mode 100644 index 000000000..819a03aac --- /dev/null +++ b/Cubical/Experiments/HAEquivInt.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.HAEquivInt where + +open import Cubical.Experiments.HAEquivInt.Base public + diff --git a/Cubical/HITs/Ints/HAEquivInt/Base.agda b/Cubical/Experiments/HAEquivInt/Base.agda similarity index 67% rename from Cubical/HITs/Ints/HAEquivInt/Base.agda rename to Cubical/Experiments/HAEquivInt/Base.agda index 560ee5e87..7c7c96db3 100644 --- a/Cubical/HITs/Ints/HAEquivInt/Base.agda +++ b/Cubical/Experiments/HAEquivInt/Base.agda @@ -1,8 +1,8 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Ints.HAEquivInt.Base where +{-# OPTIONS --safe #-} +module Cubical.Experiments.HAEquivInt.Base where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.HAEquiv +open import Cubical.Foundations.Equiv.HalfAdjoint data HAEquivInt : Type₀ where @@ -17,7 +17,7 @@ data HAEquivInt : Type₀ where suc-haequiv : HAEquiv HAEquivInt HAEquivInt -suc-haequiv = suc , record { g = pred ; sec = pred-suc ; ret = suc-pred ; com = coh } +suc-haequiv = suc , record { g = pred ; linv = pred-suc ; rinv = suc-pred ; com = coh } -- OPEN: prove HAEquivInt ≃ Int! See Experiments/HInt.agda diff --git a/Cubical/Experiments/HInt.agda b/Cubical/Experiments/HInt.agda index 8eb72ee27..7532ab79f 100644 --- a/Cubical/Experiments/HInt.agda +++ b/Cubical/Experiments/HInt.agda @@ -17,7 +17,6 @@ https://github.com/RedPRL/redtt/blob/master/library/cool/biinv-int.red It might be interesting to port that example one day. -} -{-# OPTIONS --cubical #-} module Cubical.Experiments.HInt where open import Cubical.Foundations.Prelude diff --git a/Cubical/Foundations/HoTT-UF.agda b/Cubical/Experiments/HoTT-UF.agda similarity index 96% rename from Cubical/Foundations/HoTT-UF.agda rename to Cubical/Experiments/HoTT-UF.agda index 51eadf7e6..d5c7d169a 100644 --- a/Cubical/Foundations/HoTT-UF.agda +++ b/Cubical/Experiments/HoTT-UF.agda @@ -11,9 +11,9 @@ For the moment, this requires the development version of Agda. -} -{-# OPTIONS --cubical --exact-split --safe #-} +{-# OPTIONS --exact-split --safe #-} -module Cubical.Foundations.HoTT-UF where +module Cubical.Experiments.HoTT-UF where open import Cubical.Core.Primitives hiding ( _≡_ ) open import Cubical.Core.Id public diff --git a/Cubical/Experiments/IsoInt.agda b/Cubical/Experiments/IsoInt.agda new file mode 100644 index 000000000..b3dbb7872 --- /dev/null +++ b/Cubical/Experiments/IsoInt.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.IsoInt where + +open import Cubical.Experiments.IsoInt.Base public + diff --git a/Cubical/HITs/Ints/IsoInt/Base.agda b/Cubical/Experiments/IsoInt/Base.agda similarity index 83% rename from Cubical/HITs/Ints/IsoInt/Base.agda rename to Cubical/Experiments/IsoInt/Base.agda index 950bc3637..487f3d61f 100644 --- a/Cubical/HITs/Ints/IsoInt/Base.agda +++ b/Cubical/Experiments/IsoInt/Base.agda @@ -3,11 +3,11 @@ The naive, but incorrect, way to define the integers as a HIT. This file mainly contains a proof that IsoInt ≢ Int, and ends with a - demonstration of how the same proof strategy fails for BiInvInt. + demonstration of how the same proof strategy fails for BiInvℤ. -} -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Ints.IsoInt.Base where +{-# OPTIONS --safe #-} +module Cubical.Experiments.IsoInt.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function @@ -72,25 +72,25 @@ module NonTrivial where ¬isSet-IsoInt : ¬ (isSet IsoInt) ¬isSet-IsoInt pf = NonTrivial.p₁≢p₂ (pf _ _ NonTrivial.p₁ NonTrivial.p₂) -¬Int≡IsoInt : ¬ (Int ≡ IsoInt) -¬Int≡IsoInt p = ¬isSet-IsoInt (subst isSet p isSetInt) +¬Int≡IsoInt : ¬ (ℤ ≡ IsoInt) +¬Int≡IsoInt p = ¬isSet-IsoInt (subst isSet p isSetℤ) private - -- Note: this same proof strategy fails for BiInvInt! + -- Note: this same proof strategy fails for BiInvℤ! - open import Cubical.HITs.Ints.BiInvInt hiding (zero; suc; pred; suc-pred; pred-suc) - import Cubical.HITs.Ints.BiInvInt as BiI + open import Cubical.Data.Int.MoreInts.BiInvInt hiding (zero; suc; pred; suc-pred; pred-suc) + import Cubical.Data.Int.MoreInts.BiInvInt as BiI - p₁ p₂ : Path BiInvInt (BiI.suc (BiI.pred (BiI.suc BiI.zero))) (BiI.suc BiI.zero) + p₁ p₂ : Path BiInvℤ (BiI.suc (BiI.pred (BiI.suc BiI.zero))) (BiI.suc BiI.zero) p₁ i = BiI.suc-pred (BiI.suc BiI.zero) i p₂ i = BiI.suc (BiI.pred-suc BiI.zero i) open import Cubical.HITs.S1 - toS¹ : BiInvInt → S¹ + toS¹ : BiInvℤ → S¹ toS¹ BiI.zero = base toS¹ (BiI.suc x) = toS¹ x toS¹ (BiI.predr x) = toS¹ x diff --git a/Cubical/Experiments/List.agda b/Cubical/Experiments/List.agda new file mode 100644 index 000000000..9421961e2 --- /dev/null +++ b/Cubical/Experiments/List.agda @@ -0,0 +1,285 @@ +{- + +An experiment of transporting rev-++-distr from lists to lists where +the arguments to cons have been flipped inspired by section 2 of +https://arxiv.org/abs/2010.00774 + +Note that Agda doesn't care about the order of constructors so we +can't do exactly the same example. + +-} +{-# OPTIONS --safe #-} +module Cubical.Experiments.List where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Sigma + +infixr 5 _∷_ +infixl 5 _∷'_ +infixr 5 _++_ + +-- Normal lists +data List (A : Type) : Type where + [] : List A + _∷_ : (x : A) (xs : List A) → List A + +-- Lists where the arguments to cons have been flipped +data List' (A : Type) : Type where + [] : List' A + _∷'_ : (xs : List' A) (x : A) → List' A + +variable + A : Type + + +-- Some operations and properties for List + +_++_ : List A → List A → List A +[] ++ ys = ys +(x ∷ xs) ++ ys = x ∷ xs ++ ys + +rev : List A → List A +rev [] = [] +rev (x ∷ xs) = rev xs ++ (x ∷ []) + +++-unit-r : (xs : List A) → xs ++ [] ≡ xs +++-unit-r [] = refl +++-unit-r (x ∷ xs) = cong (_∷_ x) (++-unit-r xs) + +++-assoc : (xs ys zs : List A) → (xs ++ ys) ++ zs ≡ xs ++ ys ++ zs +++-assoc [] ys zs = refl +++-assoc (x ∷ xs) ys zs = cong (_∷_ x) (++-assoc xs ys zs) + +rev-++-distr : (xs ys : List A) → rev (xs ++ ys) ≡ rev ys ++ rev xs +rev-++-distr [] ys = sym (++-unit-r (rev ys)) +rev-++-distr (x ∷ xs) ys = cong (_++ _) (rev-++-distr xs ys) ∙ ++-assoc (rev ys) (rev xs) (x ∷ []) + + + +-- We now want to transport this to List'. For this we first establish +-- an isomorphism of the types. + +toList' : List A → List' A +toList' [] = [] +toList' (x ∷ xs) = toList' xs ∷' x + +fromList' : List' A → List A +fromList' [] = [] +fromList' (xs ∷' x) = x ∷ fromList' xs + +toFrom : (xs : List' A) → toList' (fromList' xs) ≡ xs +toFrom [] = refl +toFrom (xs ∷' x) i = toFrom xs i ∷' x + +fromTo : (xs : List A) → fromList' (toList' xs) ≡ xs +fromTo [] = refl +fromTo (x ∷ xs) i = x ∷ fromTo xs i + +ListIso : Iso (List A) (List' A) +ListIso = iso toList' fromList' toFrom fromTo + +ListEquiv : List A ≃ List' A +ListEquiv = isoToEquiv ListIso + +-- We then use univalence to turn this into a path +ListPath : (A : Type) → List A ≡ List' A +ListPath A = isoToPath (ListIso {A = A}) + + +-- We can now use this path to transport the operations and properties +-- from List to List' +module transport where + + -- First make a suitable Σ-type packaging what we need for the + -- transport (note that _++_ and rev here are part of the Σ-type). + -- It should be possible to automatically generate this given a module/file. + T : Type → Type + T X = Σ[ _++_ ∈ (X → X → X) ] Σ[ rev ∈ (X → X) ] ((xs ys : X) → rev (xs ++ ys) ≡ rev ys ++ rev xs) + + -- We can now transport the instance of T for List to List' + T-List' : T (List' A) + T-List' {A = A} = transport (λ i → T (ListPath A i)) (_++_ , rev , rev-++-distr) + + -- Getting the operations and property for List' is then just a matter of projecting them out + _++'_ : List' A → List' A → List' A + _++'_ = T-List' .fst + + rev' : List' A → List' A + rev' = T-List' .snd .fst + + rev-++-distr' : (xs ys : List' A) → rev' (xs ++' ys) ≡ rev' ys ++' rev' xs + rev-++-distr' = T-List' .snd .snd + + +-- To connect this with the Cubical Agda paper consider the following +-- (painfully) manual transport. This is really what the above code +-- unfolds to. +module manualtransport where + + _++'_ : List' A → List' A → List' A + _++'_ {A = A} = transport (λ i → ListPath A i → ListPath A i → ListPath A i) _++_ + + rev' : List' A → List' A + rev' {A = A} = transport (λ i → ListPath A i → ListPath A i) rev + + rev-++-distr' : (xs ys : List' A) → rev' (xs ++' ys) ≡ rev' ys ++' rev' xs + rev-++-distr' {A = A} = transport (λ i → (xs ys : ListPath A i) + → revP i (appP i xs ys) ≡ appP i (revP i ys) (revP i xs)) + rev-++-distr + where + appP : PathP (λ i → ListPath A i → ListPath A i → ListPath A i) _++_ _++'_ + appP i = transp (λ j → ListPath A (i ∧ j) → ListPath A (i ∧ j) → ListPath A (i ∧ j)) (~ i) _++_ + + revP : PathP (λ i → ListPath A i → ListPath A i) rev rev' + revP i = transp (λ j → ListPath A (i ∧ j) → ListPath A (i ∧ j)) (~ i) rev + + + +-- The above operations for List' are derived by going back and +-- forth. With the SIP we can do better and transport properties for +-- user defined operations (assuming that the operations are +-- well-defined wrt to the forward direction of the equivalence). +open import Cubical.Foundations.SIP +open import Cubical.Structures.Axioms +open import Cubical.Structures.Product +open import Cubical.Structures.Pointed +open import Cubical.Structures.Function + + +-- For illustrative purposes we first apply the SIP manually. This +-- requires quite a bit of boilerplate code which is automated in the +-- next module. +module manualSIP (A : Type) where + + -- First define the raw structure without axioms. This is just the + -- signature of _++_ and rev. + RawStruct : Type → Type + RawStruct X = (X → X → X) × (X → X) + + -- Some boilerplate code which can be automated + e1 : StrEquiv (λ x → x → x → x) ℓ-zero + e1 = FunctionEquivStr+ pointedEquivAction + (FunctionEquivStr+ pointedEquivAction PointedEquivStr) + + e2 : StrEquiv (λ x → x → x) ℓ-zero + e2 = FunctionEquivStr+ pointedEquivAction PointedEquivStr + + RawEquivStr : StrEquiv RawStruct _ + RawEquivStr = ProductEquivStr e1 e2 + + rawUnivalentStr : UnivalentStr _ RawEquivStr + rawUnivalentStr = productUnivalentStr e1 he1 e2 he2 + where + he2 : UnivalentStr (λ z → z → z) e2 + he2 = functionUnivalentStr+ pointedEquivAction pointedTransportStr + PointedEquivStr pointedUnivalentStr + + he1 : UnivalentStr (λ z → z → z → z) e1 + he1 = functionUnivalentStr+ pointedEquivAction pointedTransportStr e2 he2 + + -- Now the property that we want to transport + P : (X : Type) → RawStruct X → Type + P X (_++_ , rev) = ((xs ys : X) → rev (xs ++ ys) ≡ rev ys ++ rev xs) + + -- Package things up for List + List-Struct : Σ[ X ∈ Type ] (Σ[ s ∈ RawStruct X ] (P X s)) + List-Struct = List A , (_++_ , rev) , rev-++-distr + + + -- We now give direct definitions of ++' and rev' for List' + _++'_ : List' A → List' A → List' A + [] ++' ys = ys + (xs ∷' x) ++' ys = (xs ++' ys) ∷' x + + rev' : List' A → List' A + rev' [] = [] + rev' (xs ∷' x) = rev' xs ++' ([] ∷' x) + + -- We then package this up into a raw structure on List' + List'-RawStruct : Σ[ X ∈ Type ] (RawStruct X) + List'-RawStruct = List' A , (_++'_ , rev') + + -- Finally we prove that toList' commutes with _++_ and rev. Note + -- that this could be a lot more complex, see for example the Matrix + -- example (Cubical.Algebra.Matrix). + toList'-++ : (xs ys : List A) → toList' (xs ++ ys) ≡ toList' xs ++' toList' ys + toList'-++ [] ys = refl + toList'-++ (x ∷ xs) ys i = toList'-++ xs ys i ∷' x + + toList'-rev : (xs : List A) → toList' (rev xs) ≡ rev' (toList' xs) + toList'-rev [] = refl + toList'-rev (x ∷ xs) = toList'-++ (rev xs) (x ∷ []) ∙ cong (_++' ([] ∷' x)) (toList'-rev xs) + + -- We can now get the property for ++' and rev' via the SIP + rev-++-distr' : (xs ys : List' A) → rev' (xs ++' ys) ≡ rev' ys ++' rev' xs + rev-++-distr' = transferAxioms rawUnivalentStr List-Struct List'-RawStruct + (ListEquiv , toList'-++ , toList'-rev) + + -- Note that rev-++-distr' is really talking about the direct + -- definitions of ++' and rev', not the transported operations as in + -- the previous attempt. + + + +-- We now automate parts of the above construction +open import Cubical.Structures.Auto + +module SIP-auto (A : Type) where + + -- First define the raw structure without axioms. This is just the + -- signature of _++_ and rev. + RawStruct : Type → Type + RawStruct X = (X → X → X) × (X → X) + + -- Some automated SIP magic + RawEquivStr : _ + RawEquivStr = AutoEquivStr RawStruct + + rawUnivalentStr : UnivalentStr _ RawEquivStr + rawUnivalentStr = autoUnivalentStr RawStruct + + -- Now the property that we want to transport + P : (X : Type) → RawStruct X → Type + P X (_++_ , rev) = ((xs ys : X) → rev (xs ++ ys) ≡ rev ys ++ rev xs) + + -- Package things up for List + List-Struct : Σ[ X ∈ Type ] (Σ[ s ∈ RawStruct X ] (P X s)) + List-Struct = List A , (_++_ , rev) , rev-++-distr + + + -- We now give direct definitions of ++' and rev' for List' + _++'_ : List' A → List' A → List' A + [] ++' ys = ys + (xs ∷' x) ++' ys = (xs ++' ys) ∷' x + + rev' : List' A → List' A + rev' [] = [] + rev' (xs ∷' x) = rev' xs ++' ([] ∷' x) + + -- We then package this up into a raw structure on List' + List'-RawStruct : Σ[ X ∈ Type ] (RawStruct X) + List'-RawStruct = List' A , (_++'_ , rev') + + -- Finally we prove that toList' commutes with _++_ and rev. Note + -- that this could be a lot more complex, see for example the Matrix + -- example (Cubical.Algebra.Matrix). + toList'-++ : (xs ys : List A) → toList' (xs ++ ys) ≡ toList' xs ++' toList' ys + toList'-++ [] ys = refl + toList'-++ (x ∷ xs) ys i = toList'-++ xs ys i ∷' x + + toList'-rev : (xs : List A) → toList' (rev xs) ≡ rev' (toList' xs) + toList'-rev [] = refl + toList'-rev (x ∷ xs) = toList'-++ (rev xs) (x ∷ []) ∙ cong (_++' ([] ∷' x)) (toList'-rev xs) + + -- We can now get the property for ++' and rev' via the SIP + rev-++-distr' : (xs ys : List' A) → rev' (xs ++' ys) ≡ rev' ys ++' rev' xs + rev-++-distr' = transferAxioms rawUnivalentStr List-Struct List'-RawStruct + (ListEquiv , toList'-++ , toList'-rev) + + -- Note that rev-++-distr' is really talking about the direct + -- definitions of ++' and rev', not the transported operations as in + -- the previous attempt. diff --git a/Cubical/Experiments/LocalisationDefs.agda b/Cubical/Experiments/LocalisationDefs.agda new file mode 100644 index 000000000..f0e2da1c0 --- /dev/null +++ b/Cubical/Experiments/LocalisationDefs.agda @@ -0,0 +1,146 @@ +-- This file contains several ways to define localisation +-- and proves them all to be equivalent + +{-# OPTIONS --safe #-} +module Cubical.Experiments.LocalisationDefs where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Transport +open import Cubical.Functions.FunExtEquiv + +import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Bool +open import Cubical.Data.Vec +open import Cubical.Data.Sigma.Base +open import Cubical.Data.Sigma.Properties +open import Cubical.Data.FinData +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary + +open import Cubical.Algebra.Group +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing + +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.HITs.PropositionalTruncation as PT + +open Iso + +private + variable + ℓ ℓ' : Level + A : Type ℓ + +record isMultClosedSubset (R' : CommRing {ℓ}) (S' : ℙ (R' .fst)) : Type ℓ where + constructor + multclosedsubset + field + containsOne : (R' .snd .CommRingStr.1r) ∈ S' + multClosed : ∀ {s t} → s ∈ S' → t ∈ S' → (R' .snd .CommRingStr._·_ s t) ∈ S' + +module _ (R' : CommRing {ℓ}) (S' : ℙ (R' .fst)) (SMultClosedSubset : isMultClosedSubset R' S') where + open isMultClosedSubset + private R = R' .fst + open CommRingStr (R' .snd) + open Theory (CommRing→Ring R') + + S = Σ[ s ∈ R ] (s ∈ S') + + -- HIT definition + data S⁻¹R : Type ℓ where + _/ₗ_ : R → S → S⁻¹R + zd : (r₁ r₂ : R) (s s₁ s₂ : S) + → fst s · r₁ · fst s₂ ≡ fst s · r₂ · fst s₁ + → r₁ /ₗ s₁ ≡ r₂ /ₗ s₂ + trunc : isSet S⁻¹R + + infixr 5 _/ₗ_ + + + module Elim {ℓ'} {B : S⁻¹R → Type ℓ'} + (_/*_ : (r : R) (s : S) → B (r /ₗ s)) + (zd* : (r₁ r₂ : R) (s s₁ s₂ : S) + → (p : fst s · r₁ · fst s₂ ≡ fst s · r₂ · fst s₁) + → PathP (λ i → B (zd r₁ r₂ s s₁ s₂ p i)) (r₁ /* s₁) (r₂ /* s₂)) + (trunc* : (q : S⁻¹R) → isSet (B q)) where + + + f : (q : S⁻¹R) → B q + f (r /ₗ s) = r /* s + f (zd r₁ r₂ s s₁ s₂ p i) = zd* r₁ r₂ s s₁ s₂ p i + f (trunc q₁ q₂ x y i j) = isOfHLevel→isOfHLevelDep 2 trunc* (f q₁) (f q₂) (cong f x) (cong f y) + (trunc q₁ q₂ x y) i j + + + module ElimProp {ℓ'} {B : S⁻¹R → Type ℓ'} (Bprop : {q : S⁻¹R} → isProp (B q)) + (_/*_ : (r : R) → (s : S) → B (r /ₗ s)) where + + + f : (q : S⁻¹R) → B q + f = Elim.f _/*_ (λ r₁ r₂ s s₁ s₂ p + → toPathP (Bprop (transp (λ i → B (zd r₁ r₂ s s₁ s₂ p i)) i0 (r₁ /* s₁)) + (r₂ /* s₂))) + λ q → isProp→isSet Bprop + + + module Rec {ℓ'} {B : Type ℓ'} (BType : isSet B) + (_/*_ : R → S → B) + (zd* : (r₁ r₂ : R) (s s₁ s₂ : S) + → (p : fst s · r₁ · fst s₂ ≡ fst s · r₂ · fst s₁) + → r₁ /* s₁ ≡ r₂ /* s₂) + where + + f : S⁻¹R → B + f = Elim.f _/*_ zd* (λ _ → BType) + + + -- approach using set quotients + _≈_ : R × S → R × S → Type ℓ + (r₁ , s₁) ≈ (r₂ , s₂) = ∃[ s ∈ S ] (fst s · r₁ · fst s₂ ≡ fst s · r₂ · fst s₁) + + S⁻¹R/ = (R × S) / _≈_ + + -- proving equivalence of the two types + φ : S⁻¹R/ → S⁻¹R + φ = SQ.rec trunc (λ (r , s) → r /ₗ s) β + where + α : ((r₁ , s₁) (r₂ , s₂) : R × S) → Σ[ s ∈ S ] (fst s · r₁ · fst s₂ ≡ fst s · r₂ · fst s₁) + → r₁ /ₗ s₁ ≡ r₂ /ₗ s₂ + α _ _ (s , p) = zd _ _ s _ _ p + + β : ((r₁ , s₁) (r₂ , s₂) : R × S) → ∃[ s ∈ S ] (fst s · r₁ · fst s₂ ≡ fst s · r₂ · fst s₁) + → r₁ /ₗ s₁ ≡ r₂ /ₗ s₂ + β _ _ = PT.rec (trunc _ _) (α _ _) + + ψ : S⁻¹R → S⁻¹R/ + ψ (r /ₗ s) = [ r , s ] + ψ (zd r₁ r₂ s s₁ s₂ p i) = eq/ (r₁ , s₁) (r₂ , s₂) ∣ s , p ∣ i + ψ (trunc x y p q i j) = squash/ (ψ x) (ψ y) (cong ψ p) (cong ψ q) i j + + η : section φ ψ + η = ElimProp.f (trunc _ _) λ _ _ → refl + + ε : retract φ ψ + ε = elimProp (λ _ → squash/ _ _) λ _ → refl + + S⁻¹R/≃S⁻¹R : S⁻¹R/ ≃ S⁻¹R + S⁻¹R/≃S⁻¹R = isoToEquiv (iso φ ψ η ε) + + + -- Set quotients but with Σ, this is the type used in Algebra.Localisation.Base + -- as this is easiest to use + _≈'_ : R × S → R × S → Type ℓ + (r₁ , s₁) ≈' (r₂ , s₂) = Σ[ s ∈ S ] (fst s · r₁ · fst s₂ ≡ fst s · r₂ · fst s₁) + + S⁻¹R/' = (R × S) / _≈'_ + + S⁻¹R/'≃S⁻¹R/ : S⁻¹R/' ≃ S⁻¹R/ + S⁻¹R/'≃S⁻¹R/ = SQ.truncRelEquiv diff --git a/Cubical/Experiments/NatMinusTwo.agda b/Cubical/Experiments/NatMinusTwo.agda new file mode 100644 index 000000000..e0a676e45 --- /dev/null +++ b/Cubical/Experiments/NatMinusTwo.agda @@ -0,0 +1,21 @@ +{- + This type ℕ₋₂ was originally used as the index to n-truncation in order to + be consistent with the notation in the HoTT book. However, ℕ was already + being used as an analogous index in Foundations.HLevels, and it became + clear that having two different indexing schemes for truncation levels was + very inconvenient. In the end, having slightly nicer notation was not worth + the hassle of having to use this type everywhere where truncation levels + were needed. So for this library, use the type `HLevel = ℕ` instead. + + See the discussions below for more context: + - https://github.com/agda/cubical/issues/266 + - https://github.com/agda/cubical/pull/238 +-} +{-# OPTIONS --safe #-} +module Cubical.Experiments.NatMinusTwo where + +open import Cubical.Experiments.NatMinusTwo.Base public + +open import Cubical.Experiments.NatMinusTwo.Properties public + +open import Cubical.Experiments.NatMinusTwo.ToNatMinusOne using (1+_; ℕ₋₁→ℕ₋₂; -1+Path) public diff --git a/Cubical/Experiments/NatMinusTwo/Base.agda b/Cubical/Experiments/NatMinusTwo/Base.agda new file mode 100644 index 000000000..e450c7adb --- /dev/null +++ b/Cubical/Experiments/NatMinusTwo/Base.agda @@ -0,0 +1,41 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Experiments.NatMinusTwo.Base where + +open import Cubical.Core.Primitives +open import Cubical.Data.Nat +open import Cubical.Data.Empty + +record ℕ₋₂ : Type₀ where + constructor -2+_ + field + n : ℕ + +pattern neg2 = -2+ zero +pattern neg1 = -2+ (suc zero) +pattern ℕ→ℕ₋₂ n = -2+ (suc (suc n)) +pattern -1+_ n = -2+ (suc n) + +2+_ : ℕ₋₂ → ℕ +2+ (-2+ n) = n + +pred₋₂ : ℕ₋₂ → ℕ₋₂ +pred₋₂ neg2 = neg2 +pred₋₂ neg1 = neg2 +pred₋₂ (ℕ→ℕ₋₂ zero) = neg1 +pred₋₂ (ℕ→ℕ₋₂ (suc n)) = (ℕ→ℕ₋₂ n) + +suc₋₂ : ℕ₋₂ → ℕ₋₂ +suc₋₂ (-2+ n) = -2+ (suc n) + +-- Natural number and negative integer literals for ℕ₋₂ + +open import Cubical.Data.Nat.Literals public + +instance + fromNatℕ₋₂ : HasFromNat ℕ₋₂ + fromNatℕ₋₂ = record { Constraint = λ _ → Unit ; fromNat = ℕ→ℕ₋₂ } + +instance + fromNegℕ₋₂ : HasFromNeg ℕ₋₂ + fromNegℕ₋₂ = record { Constraint = λ { (suc (suc (suc _))) → ⊥ ; _ → Unit } + ; fromNeg = λ { zero → 0 ; (suc zero) → neg1 ; (suc (suc zero)) → neg2 } } diff --git a/Cubical/Experiments/NatMinusTwo/Properties.agda b/Cubical/Experiments/NatMinusTwo/Properties.agda new file mode 100644 index 000000000..d72234843 --- /dev/null +++ b/Cubical/Experiments/NatMinusTwo/Properties.agda @@ -0,0 +1,11 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Experiments.NatMinusTwo.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Data.Nat +open import Cubical.Data.NatMinusOne using (ℕ₋₁) +open import Cubical.Experiments.NatMinusTwo.Base + +-2+Path : ℕ ≡ ℕ₋₂ +-2+Path = isoToPath (iso -2+_ 2+_ (λ _ → refl) (λ _ → refl)) diff --git a/Cubical/Experiments/NatMinusTwo/ToNatMinusOne.agda b/Cubical/Experiments/NatMinusTwo/ToNatMinusOne.agda new file mode 100644 index 000000000..49392bb13 --- /dev/null +++ b/Cubical/Experiments/NatMinusTwo/ToNatMinusOne.agda @@ -0,0 +1,24 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Experiments.NatMinusTwo.ToNatMinusOne where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.NatMinusOne as ℕ₋₁ using (ℕ₋₁) +open import Cubical.Experiments.NatMinusTwo.Base as ℕ₋₂ using (ℕ₋₂; -2+_) + +-- Conversions to/from ℕ₋₁ + +-1+_ : ℕ₋₁ → ℕ₋₂ +-1+ (ℕ₋₁.-1+ n) = -2+ n + +1+_ : ℕ₋₂ → ℕ₋₁ +1+ (-2+ n) = ℕ₋₁.-1+ n + +ℕ₋₁→ℕ₋₂ : ℕ₋₁ → ℕ₋₂ +ℕ₋₁→ℕ₋₂ (ℕ₋₁.-1+ n) = ℕ₋₂.-1+ n + +-- Properties + +-1+Path : ℕ₋₁ ≡ ℕ₋₂ +-1+Path = isoToPath (iso -1+_ 1+_ (λ _ → refl) (λ _ → refl)) diff --git a/Cubical/Experiments/Problem.agda b/Cubical/Experiments/Problem.agda index 2a5b70b51..1ee86657f 100644 --- a/Cubical/Experiments/Problem.agda +++ b/Cubical/Experiments/Problem.agda @@ -1,5 +1,5 @@ -- An example of something where normalization is surprisingly slow -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Experiments.Problem where open import Cubical.Foundations.Prelude @@ -71,78 +71,78 @@ problem i = transp (λ j → helix (innerpath i j)) i0 (pos 0) -- Lots of tests: (thanks Evan!) -winding2 : Path (Path S² base base) refl refl → Int +winding2 : Path (Path S² base base) refl refl → ℤ winding2 p = winding (λ j → transp (λ i → HopfS² (p i j)) i0 base) -test0 : Int +test0 : ℤ test0 = winding2 (λ i j → surf i j) -test1 : Int +test1 : ℤ test1 = winding2 (λ i j → surf j i) -test2 : Int +test2 : ℤ test2 = winding2 (λ i j → hcomp (λ _ → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) (surf i j)) -test3 : Int +test3 : ℤ test3 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → surf j k ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) base) -test4 : Int +test4 : ℤ test4 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → surf j k ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) base) -test5 : Int +test5 : ℤ test5 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → surf k i ; (j = i1) → base}) base) -test6 : Int +test6 : ℤ test6 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → base ; (j = i1) → surf k i}) base) -test7 : Int +test7 : ℤ test7 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → surf j k ; (j = i0) → base ; (j = i1) → base}) (surf i j)) -test8 : Int +test8 : ℤ test8 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → surf k i ; (j = i1) → base}) (surf i j)) -test9 : Int +test9 : ℤ test9 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → base ; (j = i1) → surf k i}) (surf i j)) -test10 : Int +test10 : ℤ test10 = winding2 (λ i j → hcomp (λ k → λ { (i = i0) → surf j k ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) (surf i j)) -- Tests using HopfS²' -winding2' : Path (Path S² base base) refl refl → Int +winding2' : Path (Path S² base base) refl refl → ℤ winding2' p = winding (λ j → transp (λ i → HopfS²' (p i j)) i0 base) -test0' : Int +test0' : ℤ test0' = winding2' (λ i j → surf i j) -test1' : Int +test1' : ℤ test1' = winding2' (λ i j → surf j i) -test2' : Int +test2' : ℤ test2' = winding2' (λ i j → hcomp (λ _ → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) (surf i j)) -test3' : Int +test3' : ℤ test3' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → surf j k ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) base) -test4' : Int +test4' : ℤ test4' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → surf j k ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) base) -test5' : Int +test5' : ℤ test5' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → surf k i ; (j = i1) → base}) base) -test6' : Int +test6' : ℤ test6' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → base ; (j = i1) → surf k i}) base) -test7' : Int +test7' : ℤ test7' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → surf j k ; (j = i0) → base ; (j = i1) → base}) (surf i j)) -test8' : Int +test8' : ℤ test8' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → surf k i ; (j = i1) → base}) (surf i j)) -test9' : Int +test9' : ℤ test9' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → base ; (i = i1) → base ; (j = i0) → base ; (j = i1) → surf k i}) (surf i j)) -test10' : Int +test10' : ℤ test10' = winding2' (λ i j → hcomp (λ k → λ { (i = i0) → surf j k ; (i = i1) → base ; (j = i0) → base ; (j = i1) → base}) (surf i j)) diff --git a/Cubical/Experiments/ZCohomology/Benchmarks.agda b/Cubical/Experiments/ZCohomology/Benchmarks.agda new file mode 100644 index 000000000..f1e403259 --- /dev/null +++ b/Cubical/Experiments/ZCohomology/Benchmarks.agda @@ -0,0 +1,385 @@ +{- + +Please do not move this file. Changes should only be made if +necessary. + +This file contains benchmarks for the paper: + +Synthetic Cohomology Theory in Cubical Agda + + +Command to run the benchmarks and get timings: + +agda -v profile.definitions:10 Benchmarks.agda + +This assumes that there is no Benchmarks.agdai file. If there is one, +then it should be removed before the above command is run. + +-} + +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Experiments.ZCohomology.Benchmarks where + +open import Cubical.Foundations.Everything +open import Cubical.Data.Nat +open import Cubical.Data.Bool +open import Cubical.Data.Int +open import Cubical.HITs.Sn +open import Cubical.Algebra.Group hiding (ℤ ; Bool) +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure hiding (_+ₕ_) renaming (_+'ₕ_ to _+ₕ_) +-- _+'ₕ_ is just (λ x y → (x +ₕ 0ₕ) +ₕ (x +ₕ 0ₕ)) +-- For technical reason, this gives nicer reductions and computes better +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Groups.Torus +open import Cubical.ZCohomology.Groups.KleinBottle +open import Cubical.ZCohomology.Groups.WedgeOfSpheres +open import Cubical.ZCohomology.Groups.RP2 +open import Cubical.ZCohomology.Groups.CP2 +open import Cubical.Data.Sigma + +open import Cubical.HITs.KleinBottle +open import Cubical.HITs.RPn.Base +open import Cubical.HITs.SetTruncation +open import Cubical.HITs.Pushout +open import Cubical.HITs.Hopf +open import Cubical.HITs.Truncation +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 + + +open IsGroupHom +open Iso + +-- S¹ (everything fast) +module S1-tests where + + ϕ : coHom 1 (S₊ 1) → ℤ + ϕ = fun (fst (Hⁿ-Sⁿ≅ℤ 0)) + + ϕ⁻¹ : ℤ → coHom 1 (S₊ 1) + ϕ⁻¹ = inv (fst (Hⁿ-Sⁿ≅ℤ 0)) + + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- <10ms + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ 1) ≡ 1 -- <10ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- <10ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 12ms + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 13ms + test₅ = refl + + test₆ : ϕ (ϕ⁻¹ -3 +ₕ ϕ⁻¹ 4) ≡ 1 -- 37ms + test₆ = refl + + test₇ : ϕ (ϕ⁻¹ -5 +ₕ ϕ⁻¹ -2) ≡ -7 -- 38ms + test₇ = refl + +-- S² +module S2-tests where + + ϕ : coHom 2 (S₊ 2) → ℤ + ϕ = fun (fst (Hⁿ-Sⁿ≅ℤ 1)) + + ϕ⁻¹ : ℤ → coHom 2 (S₊ 2) + ϕ⁻¹ = inv (fst (Hⁿ-Sⁿ≅ℤ 1)) + + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- 13ms + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ 1) ≡ 1 -- 17ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 1,152ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 1,235ms + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 1,208ms + test₅ = refl + + test₆ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 1) ≡ 2 -- 1,153ms + test₆ = refl + + test₇ : ϕ (ϕ⁻¹ 2 +ₕ ϕ⁻¹ 4) ≡ 6 -- 1,365ms + test₇ = refl + +-- S³ +module S3-tests where + + ϕ : coHom 3 (S₊ 3) → ℤ + ϕ = fun (fst (Hⁿ-Sⁿ≅ℤ 2)) + + ϕ⁻¹ : ℤ → coHom 3 (S₊ 3) + ϕ⁻¹ = inv (fst (Hⁿ-Sⁿ≅ℤ 2)) + + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- 228ms + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ 1) ≡ 1 -- 231ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 3) ≡ 3 -- 325ms + test₃ = refl + +{- + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- nope + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- nope + test₅ = refl +-} + +-- S⁴ +module S4-tests where + + ϕ : coHom 4 (S₊ 4) → ℤ + ϕ = fun (fst (Hⁿ-Sⁿ≅ℤ 3)) + + ϕ⁻¹ : ℤ → coHom 4 (S₊ 4) + ϕ⁻¹ = inv (fst (Hⁿ-Sⁿ≅ℤ 3)) + +{- _+ₕ_ Fails already here... + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- nope + test₁ = refl +-} + +-- ϕ can handle 0ₕ fast + test₂ : ϕ (0ₕ _) ≡ 0 -- < 10ms + test₂ = refl + +{- It fails to map the generator to 1, however. + test₂ : ϕ (∣ ∣_∣ ∣₂) ≡ 1 -- nope + test₂ = refl +-} + +module S1∨S1∨S2-tests₁ where -- everything fast + + ϕ : coHom 1 S²⋁S¹⋁S¹ → ℤ × ℤ + ϕ = fun (fst H¹-S²⋁S¹⋁S¹) + + ϕ⁻¹ : ℤ × ℤ → coHom 1 S²⋁S¹⋁S¹ + ϕ⁻¹ = inv (fst H¹-S²⋁S¹⋁S¹) + + test₁ : ϕ (ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- 11ms + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ (3 , 1)) ≡ (3 , 1) -- 23ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ (0 , 0) +ₕ ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- 19ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ (0 , 1) +ₕ ϕ⁻¹ (1 , 0)) ≡ (1 , 1) -- 26ms + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ (3 , 2) +ₕ ϕ⁻¹ (-1 , 5)) ≡ (2 , 7) -- 62ms + test₅ = refl + + +module S1∨S1∨S2-tests₂ where + + ϕ : coHom 2 S²⋁S¹⋁S¹ → ℤ + ϕ = fun (fst H²-S²⋁S¹⋁S¹) + + ϕ⁻¹ : ℤ → coHom 2 S²⋁S¹⋁S¹ + ϕ⁻¹ = inv (fst H²-S²⋁S¹⋁S¹) + + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- 106ms + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ 3) ≡ 3 -- 125ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 9,689ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 9,235ms + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 9,748ms + test₅ = refl + + test₆ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 1) ≡ 2 -- 9,136ms + test₆ = refl + + test₇ : ϕ (ϕ⁻¹ 2 +ₕ ϕ⁻¹ 4) ≡ 6 -- 9,557ms + test₇ = refl + + +module Torus-test₁ where -- fast + + ϕ : coHom 1 (S₊ 1 × S₊ 1) → ℤ × ℤ + ϕ = fun (fst H¹-T²≅ℤ×ℤ) + + ϕ⁻¹ : ℤ × ℤ → coHom 1 (S₊ 1 × S₊ 1) + ϕ⁻¹ = inv (fst H¹-T²≅ℤ×ℤ) + + test₁ : ϕ (ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- 11ms + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ (3 , 1)) ≡ (3 , 1) -- 17ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ (0 , 0) +ₕ ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- 19ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ (0 , 1) +ₕ ϕ⁻¹ (1 , 0)) ≡ (1 , 1) -- 26ms + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ (-3 , 2) +ₕ ϕ⁻¹ (-1 , 5)) ≡ (-4 , 7) -- 61ms + test₅ = refl + + +module Torus-test₂ where + + ϕ : coHom 2 (S₊ 1 × S₊ 1) → ℤ + ϕ = fun (fst H²-T²≅ℤ) + + ϕ⁻¹ : ℤ → coHom 2 (S₊ 1 × S₊ 1) + ϕ⁻¹ = inv (fst H²-T²≅ℤ) + + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- 136sm + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ 3) ≡ 3 -- 154ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 12,790ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 12,366ms + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 12,257ms + test₅ = refl + + test₆ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 1) ≡ 2 -- 13,092ms + test₆ = refl + + test₇ : ϕ (ϕ⁻¹ 2 +ₕ ϕ⁻¹ 4) ≡ 6 -- 12,528ms + test₇ = refl + + +module Klein-test₁ where -- fast + + ϕ : coHom 1 KleinBottle → ℤ + ϕ = fun (fst H¹-𝕂²≅ℤ) + + ϕ⁻¹ : ℤ → coHom 1 KleinBottle + ϕ⁻¹ = inv (fst H¹-𝕂²≅ℤ) + + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- <10ms + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ 3) ≡ 3 -- 13ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 10ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 14ms + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 14ms + test₅ = refl + + test₆ : ϕ (ϕ⁻¹ -3 +ₕ ϕ⁻¹ 4) ≡ 1 -- 38ms + test₆ = refl + + test₇ : ϕ (ϕ⁻¹ -5 +ₕ ϕ⁻¹ -2) ≡ -7 -- 38ms + test₇ = refl + + -- The example in the paper: + test : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 2) ≡ 3 -- 22ms + test = refl + + +module Klein-test₂ where + ϕ : coHom 2 KleinBottle → Bool + ϕ = fun (fst H²-𝕂²≅Bool) + + ϕ⁻¹ : Bool → coHom 2 KleinBottle + ϕ⁻¹ = inv (fst H²-𝕂²≅Bool) + +{- + test₀ : ϕ (0ₕ _) ≡ true -- fails already here... + test₀ = refl +-} + + +module RP2-test₂ where + ϕ : coHom 2 RP² → Bool + ϕ = fun (fst H²-RP²≅Bool) + + ϕ⁻¹ : Bool → coHom 2 RP² + ϕ⁻¹ = inv (fst H²-RP²≅Bool) + + test₀ : ϕ (0ₕ _) ≡ true -- 1,328ms (unlike for Klein, this works) + test₀ = refl + +{- + test₁ : ϕ (ϕ⁻¹ true) ≡ true -- nope + test₁ = refl +-} + + +module CP2-test₂ where + ϕ : coHom 2 CP² → ℤ + ϕ = fun (fst H²CP²≅ℤ) + + ϕ⁻¹ : ℤ → coHom 2 CP² + ϕ⁻¹ = inv (fst H²CP²≅ℤ) + + -- For explicitly constructed elmenet g : H²CP², ϕ works well + test₀ : ϕ (0ₕ _) ≡ 0 -- <10ms + test₀ = refl + + generator : coHom 2 CP² + generator = ∣ (λ { (inl x) → ∣ x ∣ ; (inr x) → 0ₖ _ ; (push a i) → p a i}) ∣₂ + where + ind : (B : TotalHopf → Type) → ((x : _) → isOfHLevel 3 (B x)) → B (north , base) → (x : _) → B x + ind = + transport (λ i → (B : isoToPath IsoS³TotalHopf i → Type) + → ((x : _) → isOfHLevel 3 (B x)) + → B (transp (λ j → isoToPath IsoS³TotalHopf (i ∨ ~ j)) i (north , base)) → (x : _) → B x) + λ B hLev ind → sphereElim _ (λ _ → hLev _) ind + + p : (a : TotalHopf) → ∣ fst a ∣ ≡ 0ₖ 2 + p = ind _ (λ _ → isOfHLevelTrunc 4 _ _) refl + + test₁ : ϕ generator ≡ 1 -- 24ms + test₁ = refl + + + -- For _+ₕ_ too + test₂ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 1,343ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 1) ≡ 2 -- 1,302ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ 2 +ₕ ϕ⁻¹ 2) ≡ 4 -- 1,410ms + test₄ = refl + + +module CP2-test₄ where + ϕ : coHom 4 CP² → ℤ + ϕ = fun (fst H⁴CP²≅ℤ) + + ϕ⁻¹ : ℤ → coHom 4 CP² + ϕ⁻¹ = inv (fst H⁴CP²≅ℤ) + +{- + test₀ : ϕ (0ₕ _) ≡ true -- fails already here... + test₀ = refl +-} diff --git a/Cubical/Experiments/ZCohomologyExperiments.agda b/Cubical/Experiments/ZCohomologyExperiments.agda new file mode 100644 index 000000000..62ab820d1 --- /dev/null +++ b/Cubical/Experiments/ZCohomologyExperiments.agda @@ -0,0 +1,79 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.ZCohomologyExperiments where +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.Groups.WedgeOfSpheres +open import Cubical.Foundations.Prelude + +open import Cubical.HITs.Truncation +open import Cubical.HITs.SetTruncation +open import Cubical.HITs.Pushout +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.Data.Int +open import Cubical.Foundations.Equiv +open import Cubical.Data.Sigma +open import Cubical.Foundations.Isomorphism +open import Cubical.Algebra.Group hiding (Int) + +private + + ⋁-to : coHom 2 S²⋁S¹⋁S¹ → Int + ⋁-to = Iso.fun (GroupIso.isom H²-S²⋁S¹⋁S¹) + ⋁-from : Int → coHom 2 S²⋁S¹⋁S¹ + ⋁-from = Iso.inv (GroupIso.isom H²-S²⋁S¹⋁S¹) + + g : S²⋁S¹⋁S¹ → ∥ S₊ 2 ∥ 4 + g (inl x) = ∣ x ∣ + g (inr x) = 0ₖ 2 + g (push a i) = 0ₖ 2 + + G = ∣ g ∣₂ + +{- +This computes: +⋁-to (⋁-from 1 +ₕ ⋁-from 1) ≡ 2 +⋁-to = refl + +We have ⋁-from 1 ≡ G and G is much simpler + +But this does not compute: +test₀ : ⋁-to (G +ₕ G) ≡ 2 +test₀ = refl + +With the strange addition it does: +test₁ : ⋁-to (G +'ₕ G) ≡ 2 +test₁ = refl +-} + +-- Similar stuff happens with Sⁿ. +private + S²-to : coHom 2 (S₊ 2) → Int + S²-to = Iso.fun (GroupIso.isom (Hⁿ-Sⁿ≅ℤ 1)) + + S²-from : Int → coHom 2 (S₊ 2) + S²-from = Iso.inv (GroupIso.isom (Hⁿ-Sⁿ≅ℤ 1)) + + one : coHom 2 (S₊ 2) + one = ∣ ∣_∣ ∣₂ + +{- +Does not compute +test₂ : S²-to (S²-from 1 +ₕ S²-from 1) ≡ 2 +test₂ = refl + +But this does +test₂ : S²-to (S²-from 1 +'ₕ S²-from 1) ≡ 2 +test₂ = refl + +This doesn't +test₃ : S²-to (one +ₕ one) ≡ 2 +test₃ = refl + +But this does +test₃ : S²-to (one +'ₕ one) ≡ 2 +test₃ = refl +-} diff --git a/Cubical/Experiments/ZCohomologyOld/Base.agda b/Cubical/Experiments/ZCohomologyOld/Base.agda new file mode 100644 index 000000000..83854ddbf --- /dev/null +++ b/Cubical/Experiments/ZCohomologyOld/Base.agda @@ -0,0 +1,46 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.ZCohomologyOld.Base where + +open import Cubical.Data.Int.Base hiding (_+_) +open import Cubical.Data.Nat.Base +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Pointed.Base + +open import Cubical.HITs.Nullification.Base +open import Cubical.HITs.SetTruncation.Base +open import Cubical.HITs.Sn.Base +open import Cubical.HITs.S1.Base +open import Cubical.HITs.Susp.Base +open import Cubical.HITs.Truncation.Base + +private + variable + ℓ : Level + A : Type ℓ + +--- Cohomology --- + +{- EM-spaces Kₙ from Brunerie 2016 -} +coHomK : (n : ℕ) → Type₀ +coHomK zero = ℤ +coHomK (suc n) = ∥ S₊ (suc n) ∥ (2 + suc n) + +{- Cohomology -} +coHom : (n : ℕ) → Type ℓ → Type ℓ +coHom n A = ∥ (A → coHomK n) ∥₂ + +--- Reduced cohomology --- + +coHom-pt : (n : ℕ) → coHomK n +coHom-pt 0 = 0 +coHom-pt 1 = ∣ base ∣ +coHom-pt (suc (suc n)) = ∣ north ∣ + +{- Pointed version of Kₙ -} +coHomK-ptd : (n : ℕ) → Pointed (ℓ-zero) +coHomK-ptd n = coHomK n , coHom-pt n + +{- Reduced cohomology -} +coHomRed : (n : ℕ) → (A : Pointed ℓ) → Type ℓ +coHomRed n A = ∥ A →∙ coHomK-ptd n ∥₂ diff --git a/Cubical/Experiments/ZCohomologyOld/KcompPrelims.agda b/Cubical/Experiments/ZCohomologyOld/KcompPrelims.agda new file mode 100644 index 000000000..df2464b9f --- /dev/null +++ b/Cubical/Experiments/ZCohomologyOld/KcompPrelims.agda @@ -0,0 +1,212 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.ZCohomologyOld.KcompPrelims where + +open import Cubical.ZCohomology.Base +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Hopf +-- open import Cubical.Homotopy.Freudenthal hiding (encode) +open import Cubical.HITs.Sn +open import Cubical.HITs.S1 +open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; map to trMap) + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Equiv.HalfAdjoint + +open import Cubical.Data.Int renaming (_+_ to +Int) hiding (_·_) +open import Cubical.Data.Nat hiding (_·_) +open import Cubical.Data.Unit + +open import Cubical.HITs.Susp +open import Cubical.HITs.Nullification +open import Cubical.Data.Prod.Base +open import Cubical.Homotopy.Loopspace +open import Cubical.Data.Bool +open import Cubical.Data.Sum.Base +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.Foundations.Function +open import Cubical.Foundations.Pointed +open import Cubical.HITs.S3 + +private + variable + ℓ : Level + A : Type ℓ + +{- We want to prove that Kn≃ΩKn+1. For this we use the map ϕ-} + +-- Proof of Kₙ ≃ ∥ ΩSⁿ⁺¹ ∥ₙ for $n ≥ 2$ +-- Entirely based on Cavallos proof of Freudenthal in Cubical.Homotopy.Freudenthal +module miniFreudenthal (n : HLevel) where + σ : S₊ (2 + n) → typ (Ω (S₊∙ (3 + n))) + σ a = merid a ∙ merid north ⁻¹ + + S2+n = S₊ (2 + n) + 4n+2 = (2 + n) + (2 + n) + + module WC-S (p : north ≡ north) where + P : (a b : S₊ (2 + n)) → Type₀ + P a b = σ b ≡ p → hLevelTrunc 4n+2 (fiber (λ x → merid x ∙ merid a ⁻¹) p) + + hLevelP : (a b : S₊ (2 + n)) → isOfHLevel 4n+2 (P a b) + hLevelP _ _ = isOfHLevelΠ 4n+2 λ _ → isOfHLevelTrunc 4n+2 + + leftFun : (a : S₊ (2 + n)) → P a north + leftFun a r = ∣ a , (rCancel' (merid a) ∙ rCancel' (merid north) ⁻¹) ∙ r ∣ + + rightFun : (b : S₊ (2 + n)) → P north b + rightFun b r = ∣ b , r ∣ + + funsAgree : leftFun north ≡ rightFun north + funsAgree = + funExt λ r → (λ i → ∣ north , rCancel' (rCancel' (merid north)) i ∙ r ∣) + ∙ λ i → ∣ north , lUnit r (~ i) ∣ + + totalFun : (a b : S2+n) → P a b + totalFun = wedgeconFun (suc n) (suc n) hLevelP rightFun leftFun funsAgree + + leftId : (λ x → totalFun x north) ≡ leftFun + leftId x i = wedgeconRight (suc n) (suc n) hLevelP rightFun leftFun funsAgree i x + + fwd : (p : north ≡ north) (a : S2+n) + → hLevelTrunc 4n+2 (fiber σ p) + → hLevelTrunc 4n+2 (fiber (λ x → merid x ∙ merid a ⁻¹) p) + fwd p a = trRec (isOfHLevelTrunc 4n+2) (uncurry (WC-S.totalFun p a)) + + fwdnorth : (p : north ≡ north) → fwd p north ≡ idfun _ + fwdnorth p = funExt (trElim (λ _ → isOfHLevelPath 4n+2 (isOfHLevelTrunc 4n+2) _ _) + λ p → refl) + + isEquivFwd : (p : north ≡ north) (a : S2+n) → isEquiv (fwd p a) + isEquivFwd p = + suspToPropElim (ptSn (suc n)) + (λ _ → isPropIsEquiv _) + helper + where + helper : isEquiv (fwd p north) + helper = subst isEquiv (sym (fwdnorth p)) (idIsEquiv _) + + interpolate : (a : S2+n) + → PathP (λ i → S2+n → north ≡ merid a i) (λ x → merid x ∙ merid a ⁻¹) merid + interpolate a i x j = compPath-filler (merid x) (merid a ⁻¹) (~ i) j + + Code : (y : Susp S2+n) → north ≡ y → Type₀ + Code north p = hLevelTrunc 4n+2 (fiber σ p) + Code south q = hLevelTrunc 4n+2 (fiber merid q) + Code (merid a i) p = + Glue + (hLevelTrunc 4n+2 (fiber (interpolate a i) p)) + (λ + { (i = i0) → _ , (fwd p a , isEquivFwd p a) + ; (i = i1) → _ , idEquiv _ + }) + + encodeS : (y : S₊ (3 + n)) (p : north ≡ y) → Code y p + encodeS y = J Code ∣ north , rCancel' (merid north) ∣ + + encodeMerid : (a : S2+n) → encodeS south (merid a) ≡ ∣ a , refl ∣ + encodeMerid a = + cong (transport (λ i → gluePath i)) + (funExt⁻ (funExt⁻ (WC-S.leftId refl) a) _ ∙ λ i → ∣ a , lem (rCancel' (merid a)) (rCancel' (merid north)) i ∣) + ∙ transport (PathP≡Path gluePath _ _) + (λ i → ∣ a , (λ j k → rCancel-filler' (merid a) i j k) ∣) + where + gluePath : I → Type _ + gluePath i = hLevelTrunc 4n+2 (fiber (interpolate a i) (λ j → merid a (i ∧ j))) + + lem : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : z ≡ y) → (p ∙ q ⁻¹) ∙ q ≡ p + lem p q = assoc p (q ⁻¹) q ⁻¹ ∙ cong (p ∙_) (lCancel q) ∙ rUnit p ⁻¹ + + contractCodeSouth : (p : north ≡ south) (c : Code south p) → encodeS south p ≡ c + contractCodeSouth p = + trElim + (λ _ → isOfHLevelPath 4n+2 (isOfHLevelTrunc 4n+2) _ _) + (uncurry λ a → + J (λ p r → encodeS south p ≡ ∣ a , r ∣) (encodeMerid a)) + + isConnectedMerid : isConnectedFun 4n+2 (merid {A = S2+n}) + isConnectedMerid p = encodeS south p , contractCodeSouth p + + isConnectedσ : isConnectedFun 4n+2 σ + isConnectedσ = + transport (λ i → isConnectedFun 4n+2 (interpolate north (~ i))) isConnectedMerid + +isConnectedσ-Sn : (n : ℕ) → isConnectedFun (4 + n) (miniFreudenthal.σ n) +isConnectedσ-Sn n = + isConnectedFunSubtr _ n _ + (subst (λ x → isConnectedFun x (miniFreudenthal.σ n)) + helper + (miniFreudenthal.isConnectedσ n)) + where + helper : 2 + (n + (2 + n)) ≡ n + (4 + n) + helper = cong suc (sym (+-suc n _)) ∙ sym (+-suc n _) + +stabSpheres-n≥2 : (n : ℕ) → Iso (hLevelTrunc (4 + n) (S₊ (2 + n))) + (hLevelTrunc (4 + n) (typ (Ω (S₊∙ (3 + n))))) +stabSpheres-n≥2 n = connectedTruncIso (4 + n) (miniFreudenthal.σ n) (isConnectedσ-Sn n) + +-- + +ϕ : (pt a : A) → typ (Ω (Susp A , north)) +ϕ pt a = (merid a) ∙ sym (merid pt) + +private + Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) + Kn→ΩKn+1 zero x i = ∣ intLoop x i ∣ + Kn→ΩKn+1 (suc zero) = trRec (isOfHLevelTrunc 4 ∣ north ∣ ∣ north ∣) + λ a i → ∣ ϕ base a i ∣ + Kn→ΩKn+1 (suc (suc n)) = trRec (isOfHLevelTrunc (2 + (3 + n)) ∣ north ∣ ∣ north ∣) + λ a i → ∣ ϕ north a i ∣ + + d-map : typ (Ω ((Susp S¹) , north)) → S¹ + d-map p = subst HopfSuspS¹ p base + + d-mapId : (r : S¹) → d-map (ϕ base r) ≡ r + d-mapId r = substComposite HopfSuspS¹ (merid r) (sym (merid base)) base ∙ + rotLemma r + where + rotLemma : (r : S¹) → r · base ≡ r + rotLemma base = refl + rotLemma (loop i) = refl + +sphereConnectedSpecCase : isConnected 4 (Susp (Susp S¹)) +sphereConnectedSpecCase = sphereConnected 3 + +d-mapComp : Iso (fiber d-map base) (Path (S₊ 3) north north) +d-mapComp = compIso (IsoΣPathTransportPathΣ {B = HopfSuspS¹} _ _) + (congIso (invIso IsoS³TotalHopf)) + +is1Connected-dmap : isConnectedFun 3 d-map +is1Connected-dmap = toPropElim (λ _ → isPropIsOfHLevel 0) + (isConnectedRetractFromIso 3 d-mapComp + (isOfHLevelRetractFromIso 0 (invIso (PathIdTruncIso 3)) + contrHelper)) + where + contrHelper : isContr (Path (∥ Susp (Susp S¹) ∥ 4) ∣ north ∣ ∣ north ∣) + fst contrHelper = refl + snd contrHelper = isOfHLevelPlus {n = 0} 2 (sphereConnected 3) ∣ north ∣ ∣ north ∣ refl + +d-Iso : Iso (∥ Path (S₊ 2) north north ∥ 3) (coHomK 1) +d-Iso = connectedTruncIso _ d-map is1Connected-dmap + +d-mapId2 : Iso.fun d-Iso ∘ trMap (ϕ base) ≡ idfun (coHomK 1) +d-mapId2 = funExt (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ a i → ∣ d-mapId a i ∣) + +Iso∥ϕ₁∥ : Iso (coHomK 1) (∥ Path (S₊ 2) north north ∥ 3) +Iso∥ϕ₁∥ = composesToId→Iso d-Iso (trMap (ϕ base)) d-mapId2 + +Iso-Kn-ΩKn+1 : (n : HLevel) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) +Iso-Kn-ΩKn+1 zero = invIso (compIso (congIso (truncIdempotentIso _ isGroupoidS¹)) ΩS¹Isoℤ) +Iso-Kn-ΩKn+1 (suc zero) = compIso Iso∥ϕ₁∥ (invIso (PathIdTruncIso 3)) +Iso-Kn-ΩKn+1 (suc (suc n)) = compIso (stabSpheres-n≥2 n) + (invIso (PathIdTruncIso (4 + n))) + where + helper : n + (4 + n) ≡ 2 + (n + (2 + n)) + helper = +-suc n (3 + n) ∙ (λ i → suc (+-suc n (2 + n) i)) diff --git a/Cubical/Experiments/ZCohomologyOld/MayerVietorisUnreduced.agda b/Cubical/Experiments/ZCohomologyOld/MayerVietorisUnreduced.agda new file mode 100644 index 000000000..8ad1449b2 --- /dev/null +++ b/Cubical/Experiments/ZCohomologyOld/MayerVietorisUnreduced.agda @@ -0,0 +1,384 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.ZCohomologyOld.MayerVietorisUnreduced where + +open import Cubical.ZCohomology.Base +open import Cubical.Experiments.ZCohomologyOld.Properties +open import Cubical.Experiments.ZCohomologyOld.KcompPrelims hiding (ϕ) + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.Sigma +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.S1 +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₁ ; ∣_∣ to ∣_∣₁) +open import Cubical.Data.Nat +open import Cubical.Data.Prod hiding (_×_) +open import Cubical.Algebra.Group +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) + +open GroupHom + +-- "Very" short exact sequences +-- i.e. an exact sequence A → B → C → D where A and D are trivial +record vSES {ℓ ℓ' ℓ'' ℓ'''} (A : Group {ℓ}) (B : Group {ℓ'}) (leftGr : Group {ℓ''}) (rightGr : Group {ℓ'''}) + : Type (ℓ-suc (ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')))) where + constructor vses + + field + isTrivialLeft : isProp ⟨ leftGr ⟩ + isTrivialRight : isProp ⟨ rightGr ⟩ + + left : GroupHom leftGr A + right : GroupHom B rightGr + ϕ : GroupHom A B + + Ker-ϕ⊂Im-left : (x : ⟨ A ⟩) + → isInKer ϕ x + → isInIm left x + Ker-right⊂Im-ϕ : (x : ⟨ B ⟩) + → isInKer right x + → isInIm ϕ x + +open BijectionIso +open vSES + +vSES→GroupIso : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group {ℓ}} {B : Group {ℓ'}} (leftGr : Group {ℓ''}) (rightGr : Group {ℓ'''}) + → vSES A B leftGr rightGr + → GroupIso A B +vSES→GroupIso {A = A} lGr rGr isvses = BijectionIsoToGroupIso theIso + where + theIso : BijectionIso _ _ + fun theIso = vSES.ϕ isvses + inj theIso a inker = pRec (isSetCarrier A _ _) + (λ (a , p) → sym p + ∙∙ cong (fun (left isvses)) (isTrivialLeft isvses a _) + ∙∙ morph1g→1g lGr A (left isvses)) + (Ker-ϕ⊂Im-left isvses a inker) + surj theIso a = Ker-right⊂Im-ϕ isvses a (isTrivialRight isvses _ _) + +vSES→GroupEquiv : {ℓ ℓ₁ ℓ₂ ℓ₃ : Level} {A : Group {ℓ}} {B : Group {ℓ₁}} (leftGr : Group {ℓ₂}) (rightGr : Group {ℓ₃}) + → vSES A B leftGr rightGr + → GroupEquiv A B +vSES→GroupEquiv lGr rGr isvses = GrIsoToGrEquiv (vSES→GroupIso lGr rGr isvses) + +module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) where + -- Proof from Brunerie 2016. + -- We first define the three morphisms involved: i, Δ and d. + + private + i* : (n : ℕ) → coHom n (Pushout f g) → coHom n A × coHom n B + i* _ = sRec (isSet× isSetSetTrunc isSetSetTrunc) λ δ → ∣ (λ x → δ (inl x)) ∣₂ , ∣ (λ x → δ (inr x)) ∣₂ + + iIsHom : (n : ℕ) → isGroupHom (coHomGr n (Pushout f g)) (×coHomGr n A B) (i* n) + iIsHom _ = sElim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) λ _ _ → refl + + i : (n : ℕ) → GroupHom (coHomGr n (Pushout f g)) (×coHomGr n A B) + GroupHom.fun (i n) = i* n + GroupHom.isHom (i n) = iIsHom n + + + private + distrLem : (n : ℕ) (x y z w : coHomK n) → (x +[ n ]ₖ y) -[ n ]ₖ (z +[ n ]ₖ w) ≡ (x -[ n ]ₖ z) +[ n ]ₖ (y -[ n ]ₖ w) + distrLem n x y z w = + cong (ΩKn+1→Kn n) (cong₂ (λ q p → q ∙ sym p) (+ₖ→∙ n x y) (+ₖ→∙ n z w) + ∙∙ cong ((Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) ∙_) (symDistr (Kn→ΩKn+1 n z) (Kn→ΩKn+1 n w)) + ∙∙ ((sym (assoc (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) _)) + ∙∙ cong (Kn→ΩKn+1 n x ∙_) (assoc (Kn→ΩKn+1 n y) (sym (Kn→ΩKn+1 n w)) (sym (Kn→ΩKn+1 n z))) + ∙∙ (cong (Kn→ΩKn+1 n x ∙_) (isCommΩK (suc n) _ _) + ∙∙ assoc _ _ _ + ∙∙ cong₂ _∙_ (sym (Iso.rightInv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ sym (Kn→ΩKn+1 n z)))) + (sym (Iso.rightInv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n y ∙ sym (Kn→ΩKn+1 n w))))))) + + + Δ' : (n : ℕ) → ⟨ ×coHomGr n A B ⟩ → ⟨ coHomGr n C ⟩ + Δ' n (α , β) = coHomFun n f α -[ n ]ₕ coHomFun n g β + + Δ'-isMorph : (n : ℕ) → isGroupHom (×coHomGr n A B) (coHomGr n C) (Δ' n) + Δ'-isMorph n = + prodElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _ ) + λ f' x1 g' x2 i → ∣ (λ x → distrLem n (f' (f x)) (g' (f x)) (x1 (g x)) (x2 (g x)) i) ∣₂ + + Δ : (n : ℕ) → GroupHom (×coHomGr n A B) (coHomGr n C) + GroupHom.fun (Δ n) = Δ' n + GroupHom.isHom (Δ n) = Δ'-isMorph n + + d-pre : (n : ℕ) → (C → coHomK n) → Pushout f g → coHomK (suc n) + d-pre n γ (inl x) = 0ₖ (suc n) + d-pre n γ (inr x) = 0ₖ (suc n) + d-pre zero γ (push a i) = Kn→ΩKn+1 zero (γ a) i + d-pre (suc n) γ (push a i) = Kn→ΩKn+1 (suc n) (γ a) i + + dHomHelperPath : (n : ℕ) (h l : C → coHomK n) (a : C) → I → I → coHomK (suc n) + dHomHelperPath zero h l a i j = + hcomp (λ k → λ { (i = i0) → lUnitₖ 1 (0ₖ 1) (~ j) + ; (i = i1) → lUnitₖ 1 (0ₖ 1) (~ j) + ; (j = i0) → +ₖ→∙ 0 (h a) (l a) (~ k) i + ; (j = i1) → cong₂Funct (λ x y → x +[ 1 ]ₖ y) + (Kn→ΩKn+1 0 (h a)) (Kn→ΩKn+1 0 (l a)) (~ k) i}) + (bottom i j) + where + bottom : I → I → coHomK 1 + bottom i j = hcomp (λ k → λ { (i = i0) → lUnitₖ 1 (0ₖ 1) (~ j) + ; (i = i1) → lUnitₖ 1 (Kn→ΩKn+1 0 (l a) k) (~ j) }) + (anotherbottom i j) + + where + anotherbottom : I → I → coHomK 1 + anotherbottom i j = hcomp (λ k → λ { (i = i0) → rUnitlUnit0 1 k (~ j) + ; (i = i1) → rUnitlUnit0 1 k (~ j) + ; (j = i0) → Kn→ΩKn+1 0 (h a) i + ; (j = i1) → Kn→ΩKn+1 0 (h a) i +[ 1 ]ₖ 0ₖ 1 }) + (rUnitₖ 1 (Kn→ΩKn+1 0 (h a) i) (~ j)) + dHomHelperPath (suc n) h l a i j = + hcomp (λ k → λ { (i = i0) → lUnitₖ (2 + n) (0ₖ (2 + n)) (~ j) + ; (i = i1) → lUnitₖ (2 + n) (0ₖ (2 + n)) (~ j) + ; (j = i0) → +ₖ→∙ (suc n) (h a) (l a) (~ k) i + ; (j = i1) → cong₂Funct (λ x y → x +[ 2 + n ]ₖ y) + (Kn→ΩKn+1 (suc n) (h a)) (Kn→ΩKn+1 (suc n) (l a)) (~ k) i}) + (bottom i j) + where + bottom : I → I → coHomK (2 + n) + bottom i j = hcomp (λ k → λ { (i = i0) → lUnitₖ (2 + n) (0ₖ (2 + n)) (~ j) + ; (i = i1) → lUnitₖ (2 + n) (Kn→ΩKn+1 (suc n) (l a) k) (~ j) }) + (anotherbottom i j) + + where + anotherbottom : I → I → coHomK (2 + n) + anotherbottom i j = hcomp (λ k → λ { (i = i0) → rUnitlUnit0 (2 + n) k (~ j) + ; (i = i1) → rUnitlUnit0 (2 + n) k (~ j) + ; (j = i0) → Kn→ΩKn+1 (suc n) (h a) i + ; (j = i1) → Kn→ΩKn+1 (suc n) (h a) i +[ 2 + n ]ₖ (0ₖ (2 + n)) }) + (rUnitₖ (2 + n) (Kn→ΩKn+1 (suc n) (h a) i) (~ j)) + + dHomHelper : (n : ℕ) (h l : C → coHomK n) (x : Pushout f g) + → d-pre n (λ x → h x +[ n ]ₖ l x) x ≡ d-pre n h x +[ suc n ]ₖ d-pre n l x + dHomHelper n h l (inl x) = sym (lUnitₖ (suc n) (0ₖ (suc n))) + dHomHelper n h l (inr x) = sym (lUnitₖ (suc n) (0ₖ (suc n))) + dHomHelper zero h l (push a i) j = dHomHelperPath zero h l a i j + dHomHelper (suc n) h l (push a i) j = dHomHelperPath (suc n) h l a i j + + dIsHom : (n : ℕ) → isGroupHom (coHomGr n C) (coHomGr (suc n) (Pushout f g)) (sRec isSetSetTrunc λ a → ∣ d-pre n a ∣₂) + dIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f g i → ∣ funExt (λ x → dHomHelper zero f g x) i ∣₂ + dIsHom (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f g i → ∣ funExt (λ x → dHomHelper (suc n) f g x) i ∣₂ + + d : (n : ℕ) → GroupHom (coHomGr n C) (coHomGr (suc n) (Pushout f g)) + GroupHom.fun (d n) = sRec isSetSetTrunc λ a → ∣ d-pre n a ∣₂ + GroupHom.isHom (d n) = dIsHom n + + -- The long exact sequence + Im-d⊂Ker-i : (n : ℕ) (x : ⟨ (coHomGr (suc n) (Pushout f g)) ⟩) + → isInIm (d n) x + → isInKer (i (suc n)) x + Im-d⊂Ker-i n = sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ a → pRec (isOfHLevelPath' 1 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + (sigmaElim (λ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ δ b i → sRec (isSet× isSetSetTrunc isSetSetTrunc) + (λ δ → ∣ (λ x → δ (inl x)) ∣₂ , ∣ (λ x → δ (inr x)) ∣₂ ) (b (~ i))) + + + Ker-i⊂Im-d : (n : ℕ) (x : ⟨ coHomGr (suc n) (Pushout f g) ⟩) + → isInKer (i (suc n)) x + → isInIm (d n) x + Ker-i⊂Im-d zero = + sElim (λ _ → isSetΠ λ _ → isProp→isSet isPropPropTrunc) + λ a p → pRec {A = (λ x → a (inl x)) ≡ λ _ → 0ₖ 1} + (isProp→ isPropPropTrunc) + (λ p1 → pRec isPropPropTrunc λ p2 → ∣ ∣ (λ c → ΩKn+1→Kn 0 (sym (cong (λ F → F (f c)) p1) + ∙∙ cong a (push c) + ∙∙ cong (λ F → F (g c)) p2)) ∣₂ + , cong ∣_∣₂ (funExt (λ δ → helper a p1 p2 δ)) ∣₁) + (Iso.fun PathIdTrunc₀Iso (cong fst p)) + (Iso.fun PathIdTrunc₀Iso (cong snd p)) + + where + helper : (F : (Pushout f g) → hLevelTrunc 3 (S₊ 1)) + (p1 : Path (_ → hLevelTrunc 3 (S₊ 1)) (λ a₁ → F (inl a₁)) (λ _ → ∣ base ∣)) + (p2 : Path (_ → hLevelTrunc 3 (S₊ 1)) (λ a₁ → F (inr a₁)) (λ _ → ∣ base ∣)) + → (δ : Pushout f g) + → d-pre 0 (λ c → ΩKn+1→Kn 0 ((λ i₁ → p1 (~ i₁) (f c)) + ∙∙ cong F (push c) + ∙∙ cong (λ F → F (g c)) p2)) δ + ≡ F δ + helper F p1 p2 (inl x) = sym (cong (λ f → f x) p1) + helper F p1 p2 (inr x) = sym (cong (λ f → f x) p2) + helper F p1 p2 (push a i) j = + hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) + ; (i = i1) → p2 (~ j) (g a) + ; (j = i0) → Iso.rightInv (Iso-Kn-ΩKn+1 0) ((λ i₁ → p1 (~ i₁) (f a)) + ∙∙ cong F (push a) + ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i + ; (j = i1) → F (push a i)}) + (doubleCompPath-filler (sym (cong (λ F → F (f a)) p1)) (cong F (push a)) (cong (λ F → F (g a)) p2) (~ j) i) + Ker-i⊂Im-d (suc n) = + sElim (λ _ → isSetΠ λ _ → isProp→isSet isPropPropTrunc) + λ a p → pRec {A = (λ x → a (inl x)) ≡ λ _ → 0ₖ (2 + n)} (isProp→ isPropPropTrunc) + (λ p1 → pRec isPropPropTrunc λ p2 → ∣ ∣ (λ c → ΩKn+1→Kn (suc n) (sym (cong (λ F → F (f c)) p1) + ∙∙ cong a (push c) + ∙∙ cong (λ F → F (g c)) p2)) ∣₂ + , cong ∣_∣₂ (funExt (λ δ → helper a p1 p2 δ)) ∣₁) + (Iso.fun PathIdTrunc₀Iso (cong fst p)) + (Iso.fun PathIdTrunc₀Iso (cong snd p)) + + where + helper : (F : (Pushout f g) → hLevelTrunc (4 + n) (S₊ (2 + n))) + (p1 : Path (_ → hLevelTrunc (4 + n) (S₊ (2 + n))) (λ a₁ → F (inl a₁)) (λ _ → ∣ north ∣)) + (p2 : Path (_ → hLevelTrunc (4 + n) (S₊ (2 + n))) (λ a₁ → F (inr a₁)) (λ _ → ∣ north ∣)) + → (δ : (Pushout f g)) + → d-pre (suc n) (λ c → ΩKn+1→Kn (suc n) ((λ i₁ → p1 (~ i₁) (f c)) + ∙∙ cong F (push c) + ∙∙ cong (λ F → F (g c)) p2)) δ + ≡ F δ + helper F p1 p2 (inl x) = sym (cong (λ f → f x) p1) + helper F p1 p2 (inr x) = sym (cong (λ f → f x) p2) + helper F p1 p2 (push a i) j = + hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) + ; (i = i1) → p2 (~ j) (g a) + ; (j = i0) → Iso.rightInv (Iso-Kn-ΩKn+1 (suc n)) ((λ i₁ → p1 (~ i₁) (f a)) + ∙∙ cong F (push a) + ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i + ; (j = i1) → F (push a i)}) + (doubleCompPath-filler (sym (cong (λ F → F (f a)) p1)) (cong F (push a)) (cong (λ F → F (g a)) p2) (~ j) i) + + open GroupHom + + Im-i⊂Ker-Δ : (n : ℕ) (x : ⟨ ×coHomGr n A B ⟩) + → isInIm (i n) x + → isInKer (Δ n) x + Im-i⊂Ker-Δ n (Fa , Fb) = + sElim {B = λ Fa → (Fb : _) → isInIm (i n) (Fa , Fb) + → isInKer (Δ n) (Fa , Fb)} + (λ _ → isSetΠ2 λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (λ Fa → sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ Fb → pRec (isSetSetTrunc _ _) + (sigmaElim (λ x → isProp→isSet (isSetSetTrunc _ _)) + λ Fd p → helper n Fa Fb Fd p)) + Fa + Fb + where + helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) (Fd : (Pushout f g) → coHomK n) + → (fun (i n) ∣ Fd ∣₂ ≡ (∣ Fa ∣₂ , ∣ Fb ∣₂)) + → (fun (Δ n)) (∣ Fa ∣₂ , ∣ Fb ∣₂) ≡ 0ₕ n + helper zero Fa Fb Fd p = cong (fun (Δ zero)) (sym p) + ∙∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₂ -[ 0 ]ₕ ∣ (λ x → Fd (push x (~ i))) ∣₂ ) + ∙∙ cancelₕ 0 ∣ (λ x → Fd (inl (f x))) ∣₂ + helper (suc n) Fa Fb Fd p = cong (fun (Δ (suc n))) (sym p) + ∙∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₂ -[ (suc n) ]ₕ ∣ (λ x → Fd (push x (~ i))) ∣₂) + ∙∙ cancelₕ (suc n) ∣ (λ x → Fd (inl (f x))) ∣₂ + + Ker-Δ⊂Im-i : (n : ℕ) (a : ⟨ ×coHomGr n A B ⟩) + → isInKer (Δ n) a + → isInIm (i n) a + Ker-Δ⊂Im-i n = prodElim (λ _ → isSetΠ (λ _ → isProp→isSet isPropPropTrunc)) + (λ Fa Fb p → pRec isPropPropTrunc + (λ q → ∣ ∣ helpFun Fa Fb q ∣₂ , refl ∣₁) + (helper Fa Fb p)) + where + helper : (Fa : A → coHomK n) (Fb : B → coHomK n) + → fun (Δ n) (∣ Fa ∣₂ , ∣ Fb ∣₂) ≡ 0ₕ n + → ∥ Path (_ → _) (λ c → Fa (f c)) (λ c → Fb (g c)) ∥₁ + helper Fa Fb p = Iso.fun PathIdTrunc₀Iso + (sym (-+cancelₕ n ∣ (λ c → Fa (f c)) ∣₂ ∣ (λ c → Fb (g c)) ∣₂) + ∙∙ cong (λ x → x +[ n ]ₕ ∣ (λ c → Fb (g c)) ∣₂) p + ∙∙ lUnitₕ n _) + + helpFun : (Fa : A → coHomK n) (Fb : B → coHomK n) + → ((λ c → Fa (f c)) ≡ (λ c → Fb (g c))) + → Pushout f g → coHomK n + helpFun Fa Fb p (inl x) = Fa x + helpFun Fa Fb p (inr x) = Fb x + helpFun Fa Fb p (push a i) = p i a + + + private + distrHelper : (n : ℕ) (p q : _) + → ΩKn+1→Kn n p +[ n ]ₖ (-[ n ]ₖ ΩKn+1→Kn n q) ≡ ΩKn+1→Kn n (p ∙ sym q) + distrHelper n p q i = + ΩKn+1→Kn n (Iso.rightInv (Iso-Kn-ΩKn+1 n) p i + ∙ Iso.rightInv (Iso-Kn-ΩKn+1 n) (sym (Iso.rightInv (Iso-Kn-ΩKn+1 n) q i)) i) + + Ker-d⊂Im-Δ : (n : ℕ) (a : coHom n C) + → isInKer (d n) a + → isInIm (Δ n) a + Ker-d⊂Im-Δ zero = + sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 isPropPropTrunc) + λ Fc p → pRec isPropPropTrunc (λ p → ∣ (∣ (λ a → ΩKn+1→Kn 0 (cong (λ f → f (inl a)) p)) ∣₂ , + ∣ (λ b → ΩKn+1→Kn 0 (cong (λ f → f (inr b)) p)) ∣₂) , + Iso.inv (PathIdTrunc₀Iso) ∣ funExt (λ c → helper2 Fc p c) ∣₁ ∣₁) + (Iso.fun (PathIdTrunc₀Iso) p) + + where + + helper2 : (Fc : C → coHomK 0) + (p : d-pre 0 Fc ≡ (λ _ → ∣ base ∣)) (c : C) + → ΩKn+1→Kn 0 (λ i₁ → p i₁ (inl (f c))) -[ 0 ]ₖ (ΩKn+1→Kn 0 (λ i₁ → p i₁ (inr (g c)))) ≡ Fc c + helper2 Fc p c = cong₂ (λ x y → ΩKn+1→Kn 0 (x ∙ sym y)) (Iso.rightInv (Iso-Kn-ΩKn+1 0) (λ i₁ → p i₁ (inl (f c)))) + (Iso.rightInv (Iso-Kn-ΩKn+1 0) (λ i₁ → p i₁ (inr (g c)))) + ∙∙ cong (ΩKn+1→Kn 0) (sym ((PathP→compPathR (cong (λ f → cong f (push c)) p)) + ∙ (λ i → (λ i₁ → p i₁ (inl (f c))) + ∙ (lUnit (sym (λ i₁ → p i₁ (inr (g c)))) (~ i))))) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 zero) (Fc c) + Ker-d⊂Im-Δ (suc n) = + sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 isPropPropTrunc) + λ Fc p → pRec isPropPropTrunc (λ p → ∣ (∣ (λ a → ΩKn+1→Kn (suc n) (cong (λ f → f (inl a)) p)) ∣₂ , + ∣ (λ b → ΩKn+1→Kn (suc n) (cong (λ f → f (inr b)) p)) ∣₂) , + Iso.inv (PathIdTrunc₀Iso) ∣ funExt (λ c → helper2 Fc p c) ∣₁ ∣₁) + (Iso.fun (PathIdTrunc₀Iso) p) + + where + + helper2 : (Fc : C → coHomK (suc n)) + (p : d-pre (suc n) Fc ≡ (λ _ → ∣ north ∣)) (c : C) + → ΩKn+1→Kn (suc n) (λ i₁ → p i₁ (inl (f c))) -[ (suc n) ]ₖ (ΩKn+1→Kn (suc n) (λ i₁ → p i₁ (inr (g c)))) ≡ Fc c + helper2 Fc p c = cong₂ (λ x y → ΩKn+1→Kn (suc n) (x ∙ sym y)) (Iso.rightInv (Iso-Kn-ΩKn+1 (suc n)) (λ i₁ → p i₁ (inl (f c)))) + (Iso.rightInv (Iso-Kn-ΩKn+1 (suc n)) (λ i₁ → p i₁ (inr (g c)))) + ∙∙ cong (ΩKn+1→Kn (suc n)) (sym ((PathP→compPathR (cong (λ f → cong f (push c)) p)) + ∙ (λ i → (λ i₁ → p i₁ (inl (f c))) + ∙ (lUnit (sym (λ i₁ → p i₁ (inr (g c)))) (~ i))))) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 (suc n)) (Fc c) + + Im-Δ⊂Ker-d : (n : ℕ) (a : coHom n C) + → isInIm (Δ n) a + → isInKer (d n) a + Im-Δ⊂Ker-d n = + sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ Fc → pRec (isOfHLevelPath' 1 isSetSetTrunc _ _) + (sigmaProdElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ Fa Fb p → pRec (isOfHLevelPath' 1 isSetSetTrunc _ _) + (λ q → ((λ i → fun (d n) ∣ (q (~ i)) ∣₂) ∙ dΔ-Id n Fa Fb)) + (Iso.fun (PathIdTrunc₀Iso) p)) + + where + d-preLeftId : (n : ℕ) (Fa : A → coHomK n)(d : (Pushout f g)) + → d-pre n (Fa ∘ f) d ≡ 0ₖ (suc n) + d-preLeftId zero Fa (inl x) = Kn→ΩKn+1 0 (Fa x) + d-preLeftId (suc n) Fa (inl x) = Kn→ΩKn+1 (suc n) (Fa x) + d-preLeftId zero Fa (inr x) = refl + d-preLeftId (suc n) Fa (inr x) = refl + d-preLeftId zero Fa (push a i) j = Kn→ΩKn+1 zero (Fa (f a)) (j ∨ i) + d-preLeftId (suc n) Fa (push a i) j = Kn→ΩKn+1 (suc n) (Fa (f a)) (j ∨ i) + + d-preRightId : (n : ℕ) (Fb : B → coHomK n) (d : (Pushout f g)) + → d-pre n (Fb ∘ g) d ≡ 0ₖ (suc n) + d-preRightId n Fb (inl x) = refl + d-preRightId zero Fb (inr x) = sym (Kn→ΩKn+1 0 (Fb x)) + d-preRightId (suc n) Fb (inr x) = sym (Kn→ΩKn+1 (suc n) (Fb x)) + d-preRightId zero Fb (push a i) j = Kn→ΩKn+1 zero (Fb (g a)) (~ j ∧ i) + d-preRightId (suc n) Fb (push a i) j = Kn→ΩKn+1 (suc n) (Fb (g a)) (~ j ∧ i) + + dΔ-Id : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) + → fun (d n) (fun (Δ n) (∣ Fa ∣₂ , ∣ Fb ∣₂)) ≡ 0ₕ (suc n) + dΔ-Id zero Fa Fb = -distrLemma 0 1 (d zero) ∣ Fa ∘ f ∣₂ ∣ Fb ∘ g ∣₂ + ∙∙ (λ i → ∣ (λ x → d-preLeftId zero Fa x i) ∣₂ -[ 1 ]ₕ ∣ (λ x → d-preRightId zero Fb x i) ∣₂) + ∙∙ cancelₕ 1 (0ₕ 1) + dΔ-Id (suc n) Fa Fb = -distrLemma (suc n) (2 + n) (d (suc n)) ∣ Fa ∘ f ∣₂ ∣ Fb ∘ g ∣₂ + ∙∙ (λ i → ∣ (λ x → d-preLeftId (suc n) Fa x i) ∣₂ -[ (2 + n) ]ₕ ∣ (λ x → d-preRightId (suc n) Fb x i) ∣₂) + ∙∙ cancelₕ (2 + n) (0ₕ (2 + n)) diff --git a/Cubical/Experiments/ZCohomologyOld/Properties.agda b/Cubical/Experiments/ZCohomologyOld/Properties.agda new file mode 100644 index 000000000..afcbe3246 --- /dev/null +++ b/Cubical/Experiments/ZCohomologyOld/Properties.agda @@ -0,0 +1,634 @@ +{-# OPTIONS --safe #-} +module Cubical.Experiments.ZCohomologyOld.Properties where + +open import Cubical.ZCohomology.Base +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed hiding (id) +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Data.Empty +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; isSetSetTrunc to §) +open import Cubical.Data.Int renaming (_+_ to _ℤ+_) hiding (-_) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) hiding (map2) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Freudenthal +open import Cubical.Algebra.Group renaming (Unit to trivialGroup ; ℤ to ℤGroup) +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Data.NatMinusOne + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup + +open import Cubical.Experiments.ZCohomologyOld.KcompPrelims + +open Iso renaming (inv to inv') + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + A' : Pointed ℓ + +infixr 34 _+ₖ_ +infixr 34 _+ₕ_ + +is2ConnectedKn : (n : ℕ) → isConnected 2 (coHomK (suc n)) +is2ConnectedKn zero = ∣ ∣ base ∣ ∣ + , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isOfHLevelTrunc 2)) _ _) + (toPropElim (λ _ → isOfHLevelTrunc 2 _ _) refl)) +is2ConnectedKn (suc n) = ∣ ∣ north ∣ ∣ + , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isOfHLevelTrunc 2 _ _)) + (suspToPropElim (ptSn (suc n)) (λ _ → isOfHLevelTrunc 2 _ _) refl)) + +isConnectedKn : (n : ℕ) → isConnected (2 + n) (coHomK (suc n)) +isConnectedKn n = isOfHLevelRetractFromIso 0 (invIso (truncOfTruncIso (2 + n) 1)) (sphereConnected (suc n)) + +-- Induction principles for cohomology groups +-- If we want to show a proposition about some x : Hⁿ(A), it suffices to show it under the +-- assumption that x = ∣f∣₂ and that f is pointed + +coHomPointedElim : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → Type ℓ'} + → ((x : coHom (suc n) A) → isProp (B x)) + → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) + → (x : coHom (suc n) A) → B x +coHomPointedElim {ℓ' = ℓ'} {A = A} n a isprop indp = + sElim (λ _ → isOfHLevelSuc 1 (isprop _)) + λ f → helper n isprop indp f (f a) refl + where + helper : (n : ℕ) {B : coHom (suc n) A → Type ℓ'} + → ((x : coHom (suc n) A) → isProp (B x)) + → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) + → (f : A → coHomK (suc n)) + → (x : coHomK (suc n)) + → f a ≡ x → B ∣ f ∣₂ + -- pattern matching a bit extra to avoid isOfHLevelPlus' + helper zero isprop ind f = + trElim (λ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ λ _ → isprop _)) + (toPropElim (λ _ → isPropΠ λ _ → isprop _) (ind f)) + helper (suc zero) isprop ind f = + trElim (λ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ λ _ → isprop _)) + (suspToPropElim base (λ _ → isPropΠ λ _ → isprop _) (ind f)) + helper (suc (suc zero)) isprop ind f = + trElim (λ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ λ _ → isprop _)) + (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) + helper (suc (suc (suc n))) isprop ind f = + trElim (λ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ λ _ → isprop _)) + (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) + +coHomPointedElim2 : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} + → ((x y : coHom (suc n) A) → isProp (B x y)) + → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) + → (x y : coHom (suc n) A) → B x y +coHomPointedElim2 {ℓ' = ℓ'} {A = A} n a isprop indp = sElim2 (λ _ _ → isOfHLevelSuc 1 (isprop _ _)) + λ f g → helper n a isprop indp f g (f a) (g a) refl refl + where + helper : (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} + → ((x y : coHom (suc n) A) → isProp (B x y)) + → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) + → (f g : A → coHomK (suc n)) + → (x y : coHomK (suc n)) + → f a ≡ x → g a ≡ y + → B ∣ f ∣₂ ∣ g ∣₂ + helper zero a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ2 λ _ _ → isprop _ _)) + (toPropElim2 (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + helper (suc zero) a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ2 λ _ _ → isprop _ _)) + (suspToPropElim2 base (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + helper (suc (suc zero)) a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ2 λ _ _ → isprop _ _)) + (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + helper (suc (suc (suc n))) a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ2 λ _ _ → isprop _ _)) + (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + + +{- Equivalence between cohomology of A and reduced cohomology of (A + 1) -} +coHomRed+1Equiv : (n : ℕ) → + (A : Type ℓ) → + (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) +coHomRed+1Equiv zero A i = ∥ helpLemma {C = (_ , pos 0)} i ∥₂ + module coHomRed+1 where + helpLemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →∙ C))) + helpLemma {C = C} = isoToPath (iso map1 + map2 + (λ b → linvPf b) + (λ _ → refl)) + where + map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →∙ C)) + map1 f = map1' , refl + module helpmap where + map1' : A ⊎ Unit → fst C + map1' (inl x) = f x + map1' (inr x) = pt C + + map2 : ((((A ⊎ Unit) , inr (tt)) →∙ C)) → (A → typ C) + map2 (g , pf) x = g (inl x) + + linvPf : (b :((((A ⊎ Unit) , inr (tt)) →∙ C))) → map1 (map2 b) ≡ b + linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) + where + helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x + helper (inl x) = refl + helper (inr tt) = sym snd +coHomRed+1Equiv (suc zero) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK 1 , ∣ base ∣)} i ∥₂ +coHomRed+1Equiv (suc (suc n)) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK (2 + n) , ∣ north ∣)} i ∥₂ + +----------- + +Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) +Kn→ΩKn+1 n = Iso.fun (Iso-Kn-ΩKn+1 n) + +ΩKn+1→Kn : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n +ΩKn+1→Kn n = Iso.inv (Iso-Kn-ΩKn+1 n) + +Kn≃ΩKn+1 : {n : ℕ} → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) +Kn≃ΩKn+1 {n = n} = isoToEquiv (Iso-Kn-ΩKn+1 n) + +---------- Algebra/Group stuff -------- + +0ₖ : (n : ℕ) → coHomK n +0ₖ = coHom-pt + +_+ₖ_ : {n : ℕ} → coHomK n → coHomK n → coHomK n +_+ₖ_ {n = n} x y = ΩKn+1→Kn n (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) + +-ₖ_ : {n : ℕ} → coHomK n → coHomK n +-ₖ_ {n = n} x = ΩKn+1→Kn n (sym (Kn→ΩKn+1 n x)) + +-- subtraction as a binary operator +_-ₖ_ : {n : ℕ} → coHomK n → coHomK n → coHomK n +_-ₖ_ {n = n} x y = ΩKn+1→Kn n (Kn→ΩKn+1 n x ∙ sym (Kn→ΩKn+1 n y)) + ++ₖ-syntax : (n : ℕ) → coHomK n → coHomK n → coHomK n ++ₖ-syntax n = _+ₖ_ {n = n} + +-ₖ-syntax : (n : ℕ) → coHomK n → coHomK n +-ₖ-syntax n = -ₖ_ {n = n} + +-'ₖ-syntax : (n : ℕ) → coHomK n → coHomK n → coHomK n +-'ₖ-syntax n = _-ₖ_ {n = n} + +syntax +ₖ-syntax n x y = x +[ n ]ₖ y +syntax -ₖ-syntax n x = -[ n ]ₖ x +syntax -'ₖ-syntax n x y = x -[ n ]ₖ y + +Kn→ΩKn+10ₖ : (n : ℕ) → Kn→ΩKn+1 n (0ₖ n) ≡ refl +Kn→ΩKn+10ₖ zero = sym (rUnit refl) +Kn→ΩKn+10ₖ (suc zero) i j = ∣ (rCancel (merid base) i j) ∣ +Kn→ΩKn+10ₖ (suc (suc n)) i j = ∣ (rCancel (merid north) i j) ∣ + +ΩKn+1→Kn-refl : (n : ℕ) → ΩKn+1→Kn n refl ≡ 0ₖ n +ΩKn+1→Kn-refl zero = refl +ΩKn+1→Kn-refl (suc zero) = refl +ΩKn+1→Kn-refl (suc (suc zero)) = refl +ΩKn+1→Kn-refl (suc (suc (suc zero))) = refl +ΩKn+1→Kn-refl (suc (suc (suc (suc zero)))) = refl +ΩKn+1→Kn-refl (suc (suc (suc (suc (suc n))))) = refl + +-0ₖ : {n : ℕ} → -[ n ]ₖ (0ₖ n) ≡ (0ₖ n) +-0ₖ {n = n} = (λ i → ΩKn+1→Kn n (sym (Kn→ΩKn+10ₖ n i))) + ∙∙ (λ i → ΩKn+1→Kn n (Kn→ΩKn+10ₖ n (~ i))) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) (0ₖ n) + ++ₖ→∙ : (n : ℕ) (a b : coHomK n) → Kn→ΩKn+1 n (a +[ n ]ₖ b) ≡ Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b ++ₖ→∙ n a b = Iso.rightInv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b) + +lUnitₖ : (n : ℕ) (x : coHomK n) → (0ₖ n) +[ n ]ₖ x ≡ x +lUnitₖ 0 x = Iso.leftInv (Iso-Kn-ΩKn+1 zero) x +lUnitₖ (suc zero) = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ x → Iso.leftInv (Iso-Kn-ΩKn+1 1) ∣ x ∣ +lUnitₖ (suc (suc n)) x = + (λ i → ΩKn+1→Kn (2 + n) (Kn→ΩKn+10ₖ (2 + n) i ∙ Kn→ΩKn+1 (2 + n) x)) ∙∙ + (cong (ΩKn+1→Kn (2 + n)) (sym (lUnit (Kn→ΩKn+1 (2 + n) x)))) ∙∙ + Iso.leftInv (Iso-Kn-ΩKn+1 (2 + n)) x +rUnitₖ : (n : ℕ) (x : coHomK n) → x +[ n ]ₖ (0ₖ n) ≡ x +rUnitₖ 0 x = Iso.leftInv (Iso-Kn-ΩKn+1 zero) x +rUnitₖ (suc zero) = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ x → Iso.leftInv (Iso-Kn-ΩKn+1 1) ∣ x ∣ +rUnitₖ (suc (suc n)) x = + (λ i → ΩKn+1→Kn (2 + n) (Kn→ΩKn+1 (2 + n) x ∙ Kn→ΩKn+10ₖ (2 + n) i)) + ∙∙ (cong (ΩKn+1→Kn (2 + n)) (sym (rUnit (Kn→ΩKn+1 (2 + n) x)))) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 (2 + n)) x + +rCancelₖ : (n : ℕ) (x : coHomK n) → x +[ n ]ₖ (-[ n ]ₖ x) ≡ (0ₖ n) +rCancelₖ zero x = (λ i → ΩKn+1→Kn 0 (Kn→ΩKn+1 zero x ∙ Iso.rightInv (Iso-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i)) ∙ + cong (ΩKn+1→Kn 0) (rCancel (Kn→ΩKn+1 zero x)) +rCancelₖ (suc n) x = (λ i → ΩKn+1→Kn (suc n) (Kn→ΩKn+1 (1 + n) x ∙ Iso.rightInv (Iso-Kn-ΩKn+1 (1 + n)) (sym (Kn→ΩKn+1 (1 + n) x)) i)) ∙ + cong (ΩKn+1→Kn (suc n)) (rCancel (Kn→ΩKn+1 (1 + n) x)) ∙ + (λ i → ΩKn+1→Kn (suc n) (Kn→ΩKn+10ₖ (suc n) (~ i))) ∙ + Iso.leftInv (Iso-Kn-ΩKn+1 (suc n)) (0ₖ (suc n)) + +lCancelₖ : (n : ℕ) (x : coHomK n) → (-[ n ]ₖ x) +[ n ]ₖ x ≡ (0ₖ n) +lCancelₖ 0 x = (λ i → ΩKn+1→Kn 0 (Iso.rightInv (Iso-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i ∙ Kn→ΩKn+1 zero x)) ∙ + cong (ΩKn+1→Kn 0) (lCancel (Kn→ΩKn+1 zero x)) +lCancelₖ (suc n) x = (λ i → ΩKn+1→Kn (suc n) (Iso.rightInv (Iso-Kn-ΩKn+1 (1 + n)) (sym (Kn→ΩKn+1 (1 + n) x)) i ∙ Kn→ΩKn+1 (1 + n) x)) ∙ + cong (ΩKn+1→Kn (suc n)) (lCancel (Kn→ΩKn+1 (1 + n) x)) ∙ + (λ i → (ΩKn+1→Kn (suc n)) (Kn→ΩKn+10ₖ (suc n) (~ i))) ∙ + Iso.leftInv (Iso-Kn-ΩKn+1 (suc n)) (0ₖ (suc n)) + +assocₖ : (n : ℕ) (x y z : coHomK n) → ((x +[ n ]ₖ y) +[ n ]ₖ z) ≡ (x +[ n ]ₖ (y +[ n ]ₖ z)) +assocₖ n x y z = ((λ i → ΩKn+1→Kn n (Kn→ΩKn+1 n (ΩKn+1→Kn n (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y)) ∙ Kn→ΩKn+1 n z)) ∙∙ + (λ i → ΩKn+1→Kn n (Iso.rightInv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i ∙ Kn→ΩKn+1 n z)) ∙∙ + (λ i → ΩKn+1→Kn n (assoc (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) (Kn→ΩKn+1 n z) (~ i)))) ∙ + (λ i → ΩKn+1→Kn n ((Kn→ΩKn+1 n x) ∙ Iso.rightInv (Iso-Kn-ΩKn+1 n) ((Kn→ΩKn+1 n y ∙ Kn→ΩKn+1 n z)) (~ i))) + +cancelₖ : (n : ℕ) (x : coHomK n) → x -[ n ]ₖ x ≡ (0ₖ n) +cancelₖ zero x = cong (ΩKn+1→Kn 0) (rCancel (Kn→ΩKn+1 zero x)) +cancelₖ (suc zero) x = cong (ΩKn+1→Kn 1) (rCancel (Kn→ΩKn+1 1 x)) +cancelₖ (suc (suc zero)) x = cong (ΩKn+1→Kn 2) (rCancel (Kn→ΩKn+1 2 x)) +cancelₖ (suc (suc (suc zero))) x = cong (ΩKn+1→Kn 3) (rCancel (Kn→ΩKn+1 3 x)) +cancelₖ (suc (suc (suc (suc zero)))) x = cong (ΩKn+1→Kn 4) (rCancel (Kn→ΩKn+1 4 x)) +cancelₖ (suc (suc (suc (suc (suc n))))) x = cong (ΩKn+1→Kn (5 + n)) (rCancel (Kn→ΩKn+1 (5 + n) x)) + +-rUnitₖ : (n : ℕ) (x : coHomK n) → x -[ n ]ₖ 0ₖ n ≡ x +-rUnitₖ zero x = rUnitₖ zero x +-rUnitₖ (suc n) x = cong (λ y → ΩKn+1→Kn (suc n) (Kn→ΩKn+1 (suc n) x ∙ sym y)) (Kn→ΩKn+10ₖ (suc n)) + ∙∙ cong (ΩKn+1→Kn (suc n)) (sym (rUnit (Kn→ΩKn+1 (suc n) x))) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 (suc n)) x + +abstract + isCommΩK1 : (n : ℕ) → isComm∙ ((Ω^ n) (coHomK-ptd 1)) + isCommΩK1 zero = isCommA→isCommTrunc 2 comm-ΩS¹ isGroupoidS¹ + isCommΩK1 (suc n) = Eckmann-Hilton n + + open Iso renaming (inv to inv') + isCommΩK : (n : ℕ) → isComm∙ (coHomK-ptd n) + isCommΩK zero p q = isSetℤ _ _ (p ∙ q) (q ∙ p) + isCommΩK (suc zero) = isCommA→isCommTrunc 2 comm-ΩS¹ isGroupoidS¹ + isCommΩK (suc (suc n)) = subst isComm∙ (λ i → coHomK (2 + n) , ΩKn+1→Kn-refl (2 + n) i) (ptdIso→comm {A = (_ , _)} (invIso (Iso-Kn-ΩKn+1 (2 + n))) (Eckmann-Hilton 0)) + +commₖ : (n : ℕ) (x y : coHomK n) → (x +[ n ]ₖ y) ≡ (y +[ n ]ₖ x) +commₖ 0 x y i = ΩKn+1→Kn 0 (isCommΩK1 0 (Kn→ΩKn+1 0 x) (Kn→ΩKn+1 0 y) i) +commₖ 1 x y i = ΩKn+1→Kn 1 (ptdIso→comm {A = ((∣ north ∣ ≡ ∣ north ∣) , snd ((Ω^ 1) (coHomK 3 , ∣ north ∣)))} + {B = coHomK 2} + (invIso (Iso-Kn-ΩKn+1 2)) (Eckmann-Hilton 0) (Kn→ΩKn+1 1 x) (Kn→ΩKn+1 1 y) i) +commₖ 2 x y i = ΩKn+1→Kn 2 (ptdIso→comm {A = (∣ north ∣ ≡ ∣ north ∣) , snd ((Ω^ 1) (coHomK 4 , ∣ north ∣))} + {B = coHomK 3} + (invIso (Iso-Kn-ΩKn+1 3)) (Eckmann-Hilton 0) (Kn→ΩKn+1 2 x) (Kn→ΩKn+1 2 y) i) +commₖ 3 x y i = ΩKn+1→Kn 3 (ptdIso→comm {A = (∣ north ∣ ≡ ∣ north ∣) , snd ((Ω^ 1) (coHomK 5 , ∣ north ∣))} + {B = coHomK 4} + (invIso (Iso-Kn-ΩKn+1 4)) (Eckmann-Hilton 0) (Kn→ΩKn+1 3 x) (Kn→ΩKn+1 3 y) i) +commₖ (suc (suc (suc (suc n)))) x y i = + ΩKn+1→Kn (4 + n) (ptdIso→comm {A = (∣ north ∣ ≡ ∣ north ∣) , snd ((Ω^ 1) (coHomK (6 + n) , ∣ north ∣))} + {B = coHomK (5 + n)} + (invIso (Iso-Kn-ΩKn+1 (5 + n))) (Eckmann-Hilton 0) (Kn→ΩKn+1 (4 + n) x) (Kn→ΩKn+1 (4 + n) y) i) + + +rUnitₖ' : (n : ℕ) (x : coHomK n) → x +[ n ]ₖ (0ₖ n) ≡ x +rUnitₖ' n x = commₖ n x (0ₖ n) ∙ lUnitₖ n x + +-distrₖ : (n : ℕ) (x y : coHomK n) → -[ n ]ₖ (x +[ n ]ₖ y) ≡ (-[ n ]ₖ x) +[ n ]ₖ (-[ n ]ₖ y) +-distrₖ n x y = ((λ i → ΩKn+1→Kn n (sym (Kn→ΩKn+1 n (ΩKn+1→Kn n (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y))))) ∙∙ + (λ i → ΩKn+1→Kn n (sym (Iso.rightInv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i))) ∙∙ + (λ i → ΩKn+1→Kn n (symDistr (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) i))) ∙∙ + (λ i → ΩKn+1→Kn n (Iso.rightInv (Iso-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n y)) (~ i) ∙ (Iso.rightInv (Iso-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n x)) (~ i)))) ∙∙ + commₖ n (-[ n ]ₖ y) (-[ n ]ₖ x) + +private + rCancelLem : (n : ℕ) (x : coHomK n) → ΩKn+1→Kn n ((Kn→ΩKn+1 n x) ∙ refl) ≡ ΩKn+1→Kn n (Kn→ΩKn+1 n x) + rCancelLem zero x = refl + rCancelLem (suc n) x = cong (ΩKn+1→Kn (suc n)) (sym (rUnit (Kn→ΩKn+1 (suc n) x))) + + lCancelLem : (n : ℕ) (x : coHomK n) → ΩKn+1→Kn n (refl ∙ (Kn→ΩKn+1 n x)) ≡ ΩKn+1→Kn n (Kn→ΩKn+1 n x) + lCancelLem zero x = refl + lCancelLem (suc n) x = cong (ΩKn+1→Kn (suc n)) (sym (lUnit (Kn→ΩKn+1 (suc n) x))) + + +-cancelRₖ : (n : ℕ) (x y : coHomK n) → (y +[ n ]ₖ x) -[ n ]ₖ x ≡ y +-cancelRₖ n x y = (cong (ΩKn+1→Kn n) ((cong (_∙ sym (Kn→ΩKn+1 n x)) (+ₖ→∙ n y x)) + ∙∙ sym (assoc _ _ _) + ∙∙ cong (Kn→ΩKn+1 n y ∙_) (rCancel _))) + ∙∙ rCancelLem n y + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) y + +-cancelLₖ : (n : ℕ) (x y : coHomK n) → (x +[ n ]ₖ y) -[ n ]ₖ x ≡ y +-cancelLₖ n x y = cong (λ z → z -[ n ]ₖ x) (commₖ n x y) ∙ -cancelRₖ n x y + +-+cancelₖ : (n : ℕ) (x y : coHomK n) → (x -[ n ]ₖ y) +[ n ]ₖ y ≡ x +-+cancelₖ n x y = (cong (ΩKn+1→Kn n) ((cong (_∙ (Kn→ΩKn+1 n y)) (Iso.rightInv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ sym (Kn→ΩKn+1 n y)))) + ∙∙ sym (assoc _ _ _) + ∙∙ cong (Kn→ΩKn+1 n x ∙_) (lCancel _))) + ∙∙ rCancelLem n x + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) x + +---- Group structure of cohomology groups --- + +_+ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A +_+ₕ_ {n = n} = sRec2 § λ a b → ∣ (λ x → a x +[ n ]ₖ b x) ∣₂ + +-ₕ_ : {n : ℕ} → coHom n A → coHom n A +-ₕ_ {n = n} = sRec § λ a → ∣ (λ x → -[ n ]ₖ a x) ∣₂ + +_-ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A +_-ₕ_ {n = n} = sRec2 § λ a b → ∣ (λ x → a x -[ n ]ₖ b x) ∣₂ + ++ₕ-syntax : (n : ℕ) → coHom n A → coHom n A → coHom n A ++ₕ-syntax n = _+ₕ_ {n = n} + +-ₕ-syntax : (n : ℕ) → coHom n A → coHom n A +-ₕ-syntax n = -ₕ_ {n = n} + +-ₕ'-syntax : (n : ℕ) → coHom n A → coHom n A → coHom n A +-ₕ'-syntax n = _-ₕ_ {n = n} + +syntax +ₕ-syntax n x y = x +[ n ]ₕ y +syntax -ₕ-syntax n x = -[ n ]ₕ x +syntax -ₕ'-syntax n x y = x -[ n ]ₕ y + +0ₕ : (n : ℕ) → coHom n A +0ₕ n = ∣ (λ _ → (0ₖ n)) ∣₂ + +rUnitₕ : (n : ℕ) (x : coHom n A) → x +[ n ]ₕ (0ₕ n) ≡ x +rUnitₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rUnitₖ n (a x)) i ∣₂ + +lUnitₕ : (n : ℕ) (x : coHom n A) → (0ₕ n) +[ n ]ₕ x ≡ x +lUnitₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lUnitₖ n (a x)) i ∣₂ + +rCancelₕ : (n : ℕ) (x : coHom n A) → x +[ n ]ₕ (-[ n ]ₕ x) ≡ 0ₕ n +rCancelₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rCancelₖ n (a x)) i ∣₂ + +lCancelₕ : (n : ℕ) (x : coHom n A) → (-[ n ]ₕ x) +[ n ]ₕ x ≡ 0ₕ n +lCancelₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lCancelₖ n (a x)) i ∣₂ + +assocₕ : (n : ℕ) (x y z : coHom n A) → ((x +[ n ]ₕ y) +[ n ]ₕ z) ≡ (x +[ n ]ₕ (y +[ n ]ₕ z)) +assocₕ n = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b c i → ∣ funExt (λ x → assocₖ n (a x) (b x) (c x)) i ∣₂ + +commₕ : (n : ℕ) (x y : coHom n A) → (x +[ n ]ₕ y) ≡ (y +[ n ]ₕ x) +commₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ funExt (λ x → commₖ n (a x) (b x)) i ∣₂ + +cancelₕ : (n : ℕ) (x : coHom n A) → x -[ n ]ₕ x ≡ 0ₕ n +cancelₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → cancelₖ n (a x)) i ∣₂ + +-ₖ-ₖ : (n : ℕ) (x : coHomK n) → (-[ n ]ₖ (-[ n ]ₖ x)) ≡ x +-ₖ-ₖ n x = cong ((ΩKn+1→Kn n) ∘ sym) (Iso.rightInv (Iso-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n x))) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) x + +-- Proof that rUnitₖ and lUnitₖ agree on 0ₖ. Needed for Mayer-Vietoris. +private + rUnitlUnitGen : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {b : B} (e : Iso A (b ≡ b)) + (0A : A) + (0fun : fun e 0A ≡ refl) + → Path (inv' e (fun e 0A ∙ fun e 0A) ≡ 0A) + (cong (inv' e) (cong (_∙ fun e 0A) 0fun) ∙∙ cong (inv' e) (sym (lUnit (fun e 0A))) ∙∙ Iso.leftInv e 0A) + (cong (inv' e) (cong (fun e 0A ∙_) 0fun) ∙∙ cong (inv' e) (sym (rUnit (fun e 0A))) ∙∙ Iso.leftInv e 0A) + rUnitlUnitGen e 0A 0fun = + (λ i → cong (inv' e) (cong (_∙ fun e 0A) 0fun) ∙∙ rUnit (cong (inv' e) (sym (lUnit (fun e 0A)))) i ∙∙ Iso.leftInv e 0A) + ∙ ((λ i → (λ j → inv' e (0fun (~ i ∧ j) ∙ 0fun (j ∧ i))) + ∙∙ ((λ j → inv' e (0fun (~ i ∨ j) ∙ 0fun i)) + ∙∙ cong (inv' e) (sym (lUnit (0fun i))) + ∙∙ λ j → inv' e (0fun (i ∧ (~ j)))) + ∙∙ Iso.leftInv e 0A) + ∙∙ (λ i → (λ j → inv' e (fun e 0A ∙ 0fun j)) + ∙∙ (λ j → inv' e (0fun (j ∧ ~ i) ∙ refl)) + ∙∙ cong (inv' e) (sym (rUnit (0fun (~ i)))) + ∙∙ (λ j → inv' e (0fun (~ i ∧ ~ j))) + ∙∙ Iso.leftInv e 0A) + ∙∙ λ i → cong (inv' e) (cong (fun e 0A ∙_) 0fun) + ∙∙ rUnit (cong (inv' e) (sym (rUnit (fun e 0A)))) (~ i) + ∙∙ Iso.leftInv e 0A) + +rUnitlUnit0 : (n : ℕ) → rUnitₖ n (0ₖ n) ≡ lUnitₖ n (0ₖ n) +rUnitlUnit0 0 = refl +rUnitlUnit0 (suc zero) = refl +rUnitlUnit0 (suc (suc n)) = + sym (rUnitlUnitGen (Iso-Kn-ΩKn+1 (2 + n)) + (0ₖ (2 + n)) + (Kn→ΩKn+10ₖ (2 + n))) + +-cancelLₕ : (n : ℕ) (x y : coHom n A) → (x +[ n ]ₕ y) -[ n ]ₕ x ≡ y +-cancelLₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelLₖ n (a x) (b x) i) ∣₂ + +-cancelRₕ : (n : ℕ) (x y : coHom n A) → (y +[ n ]ₕ x) -[ n ]ₕ x ≡ y +-cancelRₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelRₖ n (a x) (b x) i) ∣₂ + +-+cancelₕ : (n : ℕ) (x y : coHom n A) → (x -[ n ]ₕ y) +[ n ]ₕ y ≡ x +-+cancelₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -+cancelₖ n (a x) (b x) i) ∣₂ + + +-- Group structure of reduced cohomology groups (in progress - might need K to compute properly first) --- + ++ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A → coHomRed n A ++ₕ∙ zero = sRec2 § λ { (a , pa) (b , pb) → ∣ (λ x → a x +[ zero ]ₖ b x) , (λ i → (pa i +[ zero ]ₖ pb i)) ∣₂ } ++ₕ∙ (suc zero) = sRec2 § λ { (a , pa) (b , pb) → ∣ (λ x → a x +[ 1 ]ₖ b x) , (λ i → pa i +[ 1 ]ₖ pb i) ∙ lUnitₖ 1 (0ₖ 1) ∣₂ } ++ₕ∙ (suc (suc n)) = sRec2 § λ { (a , pa) (b , pb) → ∣ (λ x → a x +[ (2 + n) ]ₖ b x) , (λ i → pa i +[ (2 + n) ]ₖ pb i) ∙ lUnitₖ (2 + n) (0ₖ (2 + n)) ∣₂ } + +open IsSemigroup +open IsMonoid +open GroupStr +open IsGroupHom + +coHomGr : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → Group ℓ +coHomGr n A = coHom n A , coHomGrnA + where + coHomGrnA : GroupStr (coHom n A) + 1g coHomGrnA = 0ₕ n + GroupStr._·_ coHomGrnA = λ x y → x +[ n ]ₕ y + inv coHomGrnA = λ x → -[ n ]ₕ x + isGroup coHomGrnA = helper + where + abstract + helper : IsGroup {G = coHom n A} (0ₕ n) (λ x y → x +[ n ]ₕ y) (λ x → -[ n ]ₕ x) + helper = makeIsGroup § (λ x y z → sym (assocₕ n x y z)) (rUnitₕ n) (lUnitₕ n) (rCancelₕ n) (lCancelₕ n) + +×coHomGr : (n : ℕ) (A : Type ℓ) (B : Type ℓ') → Group _ +×coHomGr n A B = DirProd (coHomGr n A) (coHomGr n B) + +coHomFun : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) (f : A → B) → coHom n B → coHom n A +coHomFun n f = sRec § λ β → ∣ β ∘ f ∣₂ + +-distrLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : ℕ) (f : GroupHom (coHomGr n A) (coHomGr m B)) + (x y : coHom n A) + → fst f (x -[ n ]ₕ y) ≡ fst f x -[ m ]ₕ fst f y +-distrLemma n m f' x y = sym (-cancelRₕ m (f y) (f (x -[ n ]ₕ y))) + ∙∙ cong (λ x → x -[ m ]ₕ f y) (sym (f' .snd .pres· (x -[ n ]ₕ y) y)) + ∙∙ cong (λ x → x -[ m ]ₕ f y) ( cong f (-+cancelₕ n _ _)) + where + f = fst f' + +--- the loopspace of Kₙ is commutative regardless of base + +addIso : (n : ℕ) (x : coHomK n) → Iso (coHomK n) (coHomK n) +fun (addIso n x) y = y +[ n ]ₖ x +inv' (addIso n x) y = y -[ n ]ₖ x +rightInv (addIso n x) y = -+cancelₖ n y x +leftInv (addIso n x) y = -cancelRₖ n x y + +isCommΩK-based : (n : ℕ) (x : coHomK n) → isComm∙ (coHomK n , x) +isCommΩK-based zero x p q = isSetℤ _ _ (p ∙ q) (q ∙ p) +isCommΩK-based (suc zero) x = + subst isComm∙ (λ i → coHomK 1 , lUnitₖ 1 x i) + (ptdIso→comm {A = (_ , 0ₖ 1)} (addIso 1 x) + (isCommΩK 1)) +isCommΩK-based (suc (suc n)) x = + subst isComm∙ (λ i → coHomK (suc (suc n)) , lUnitₖ (suc (suc n)) x i) + (ptdIso→comm {A = (_ , 0ₖ (suc (suc n)))} (addIso (suc (suc n)) x) + (isCommΩK (suc (suc n)))) + +addLemma : (a b : ℤ) → a +[ 0 ]ₖ b ≡ (a ℤ+ b) +addLemma a b = (cong (ΩKn+1→Kn 0) (sym (congFunct ∣_∣ (intLoop a) (intLoop b)))) + ∙∙ (λ i → ΩKn+1→Kn 0 (cong ∣_∣ (intLoop-hom a b i))) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 0) (a ℤ+ b) + +--- +-- hidden versions of cohom stuff using the "lock" hack. The locked versions can be used when proving things. +-- Swapping "key" for "tt*" will then give computing functions. + +Unit' : Type₀ +Unit' = lockUnit {ℓ-zero} + +lock : ∀ {ℓ} {A : Type ℓ} → Unit' → A → A +lock unlock = λ x → x + +module lockedCohom (key : Unit') where + +K : (n : ℕ) → coHomK n → coHomK n → coHomK n + +K n = lock key (_+ₖ_ {n = n}) + + -K : (n : ℕ) → coHomK n → coHomK n + -K n = lock key (-ₖ_ {n = n}) + + -Kbin : (n : ℕ) → coHomK n → coHomK n → coHomK n + -Kbin n = lock key (_-ₖ_ {n = n}) + + rUnitK : (n : ℕ) (x : coHomK n) → +K n x (0ₖ n) ≡ x + rUnitK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) x (0ₖ n) ≡ x + pm unlock = rUnitₖ n x + + lUnitK : (n : ℕ) (x : coHomK n) → +K n (0ₖ n) x ≡ x + lUnitK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (0ₖ n) x ≡ x + pm unlock = lUnitₖ n x + + rCancelK : (n : ℕ) (x : coHomK n) → +K n x (-K n x) ≡ 0ₖ n + rCancelK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) x (lock t (-ₖ_ {n = n}) x) ≡ 0ₖ n + pm unlock = rCancelₖ n x + + lCancelK : (n : ℕ) (x : coHomK n) → +K n (-K n x) x ≡ 0ₖ n + lCancelK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (lock t (-ₖ_ {n = n}) x) x ≡ 0ₖ n + pm unlock = lCancelₖ n x + + -cancelRK : (n : ℕ) (x y : coHomK n) → -Kbin n (+K n y x) x ≡ y + -cancelRK n x y = pm key + where + pm : (t : Unit') → lock t (_-ₖ_ {n = n}) (lock t (_+ₖ_ {n = n}) y x) x ≡ y + pm unlock = -cancelRₖ n x y + + -cancelLK : (n : ℕ) (x y : coHomK n) → -Kbin n (+K n x y) x ≡ y + -cancelLK n x y = pm key + where + pm : (t : Unit') → lock t (_-ₖ_ {n = n}) (lock t (_+ₖ_ {n = n}) x y) x ≡ y + pm unlock = -cancelLₖ n x y + + -+cancelK : (n : ℕ) (x y : coHomK n) → +K n (-Kbin n x y) y ≡ x + -+cancelK n x y = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (lock t (_-ₖ_ {n = n}) x y) y ≡ x + pm unlock = -+cancelₖ n x y + + cancelK : (n : ℕ) (x : coHomK n) → -Kbin n x x ≡ 0ₖ n + cancelK n x = pm key + where + pm : (t : Unit') → (lock t (_-ₖ_ {n = n}) x x) ≡ 0ₖ n + pm unlock = cancelₖ n x + + assocK : (n : ℕ) (x y z : coHomK n) → +K n (+K n x y) z ≡ +K n x (+K n y z) + assocK n x y z = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (lock t (_+ₖ_ {n = n}) x y) z + ≡ lock t (_+ₖ_ {n = n}) x (lock t (_+ₖ_ {n = n}) y z) + pm unlock = assocₖ n x y z + + commK : (n : ℕ) (x y : coHomK n) → +K n x y ≡ +K n y x + commK n x y = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) x y ≡ lock t (_+ₖ_ {n = n}) y x + pm unlock = commₖ n x y + + -- cohom + + +H : (n : ℕ) (x y : coHom n A) → coHom n A + +H n = sRec2 § λ a b → ∣ (λ x → +K n (a x) (b x)) ∣₂ + + -H : (n : ℕ) (x : coHom n A) → coHom n A + -H n = sRec § λ a → ∣ (λ x → -K n (a x)) ∣₂ + + -Hbin : (n : ℕ) → coHom n A → coHom n A → coHom n A + -Hbin n = sRec2 § λ a b → ∣ (λ x → -Kbin n (a x) (b x)) ∣₂ + + rUnitH : (n : ℕ) (x : coHom n A) → +H n x (0ₕ n) ≡ x + rUnitH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rUnitK n (a x)) i ∣₂ + + lUnitH : (n : ℕ) (x : coHom n A) → +H n (0ₕ n) x ≡ x + lUnitH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lUnitK n (a x)) i ∣₂ + + rCancelH : (n : ℕ) (x : coHom n A) → +H n x (-H n x) ≡ 0ₕ n + rCancelH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rCancelK n (a x)) i ∣₂ + + lCancelH : (n : ℕ) (x : coHom n A) → +H n (-H n x) x ≡ 0ₕ n + lCancelH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lCancelK n (a x)) i ∣₂ + + assocH : (n : ℕ) (x y z : coHom n A) → (+H n (+H n x y) z) ≡ (+H n x (+H n y z)) + assocH n = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b c i → ∣ funExt (λ x → assocK n (a x) (b x) (c x)) i ∣₂ + + commH : (n : ℕ) (x y : coHom n A) → (+H n x y) ≡ (+H n y x) + commH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ funExt (λ x → commK n (a x) (b x)) i ∣₂ + + -cancelRH : (n : ℕ) (x y : coHom n A) → -Hbin n (+H n y x) x ≡ y + -cancelRH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelRK n (a x) (b x) i) ∣₂ + + -cancelLH : (n : ℕ) (x y : coHom n A) → -Hbin n (+H n x y) x ≡ y + -cancelLH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelLK n (a x) (b x) i) ∣₂ + + -+cancelH : (n : ℕ) (x y : coHom n A) → +H n (-Hbin n x y) y ≡ x + -+cancelH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -+cancelK n (a x) (b x) i) ∣₂ + + ++K→∙ : (key : Unit') (n : ℕ) (a b : coHomK n) → Kn→ΩKn+1 n (lockedCohom.+K key n a b) ≡ Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b ++K→∙ unlock = +ₖ→∙ + ++H≡+ₕ : (key : Unit') (n : ℕ) → lockedCohom.+H key {A = A} n ≡ _+ₕ_ {n = n} ++H≡+ₕ unlock _ = refl + +rUnitlUnit0K : (key : Unit') (n : ℕ) → lockedCohom.rUnitK key n (0ₖ n) ≡ lockedCohom.lUnitK key n (0ₖ n) +rUnitlUnit0K unlock = rUnitlUnit0 diff --git a/Cubical/Foundations/CartesianKanOps.agda b/Cubical/Foundations/CartesianKanOps.agda index ac06580c6..6aa33b870 100644 --- a/Cubical/Foundations/CartesianKanOps.agda +++ b/Cubical/Foundations/CartesianKanOps.agda @@ -1,11 +1,11 @@ -- This file derives some of the Cartesian Kan operations using transp -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.CartesianKanOps where open import Cubical.Foundations.Prelude coe0→1 : ∀ {ℓ} (A : I → Type ℓ) → A i0 → A i1 -coe0→1 A a = transp A i0 a +coe0→1 A a = transp (\ i → A i) i0 a -- "coe filler" coe0→i : ∀ {ℓ} (A : I → Type ℓ) (i : I) → A i0 → A i @@ -48,7 +48,7 @@ coei1→0 A a = refl -- unlike in cartesian cubes, we don't get coei→i = id definitionally coei→j : ∀ {ℓ} (A : I → Type ℓ) (i j : I) → A i → A j coei→j A i j a = - fill A + fill (\ i → A i) (λ j → λ { (i = i0) → coe0→i A j a ; (i = i1) → coe1→i A j a }) @@ -118,9 +118,9 @@ filli→j : ∀ {ℓ} (A : ∀ i → Type ℓ) --------------------------- (j : I) → A j filli→j A {φ = φ} u i ui j = - fill A + fill (\ i → A i) (λ j → λ { (φ = i1) → u j 1=1 - ; (i = i0) → fill A u ui j + ; (i = i0) → fill (\ i → A i) (\ i → u i) ui j ; (i = i1) → fill1→i A u ui j }) (inS (filli→0 A u i ui)) diff --git a/Cubical/Foundations/Embedding.agda b/Cubical/Foundations/Embedding.agda deleted file mode 100644 index abaf1b787..000000000 --- a/Cubical/Foundations/Embedding.agda +++ /dev/null @@ -1,130 +0,0 @@ -{-# OPTIONS --cubical --safe #-} - -module Cubical.Foundations.Embedding where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Function -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Transport -open import Cubical.Foundations.Isomorphism - -private - variable - ℓ : Level - A B : Type ℓ - f : A → B - w x : A - y z : B - --- Embeddings are generalizations of injections. The usual --- definition of injection as: --- --- f x ≡ f y → x ≡ y --- --- is not well-behaved with higher h-levels, while embeddings --- are. -isEmbedding : (A → B) → Type _ -isEmbedding f = ∀ w x → isEquiv {A = w ≡ x} (cong f) - -isEmbeddingIsProp : isProp (isEmbedding f) -isEmbeddingIsProp {f = f} - = isPropPi λ w → isPropPi λ x → isPropIsEquiv (cong f) - --- If A and B are h-sets, then injective functions between --- them are embeddings. --- --- Note: It doesn't appear to be possible to omit either of --- the `isSet` hypotheses. -injEmbedding - : {f : A → B} - → isSet A → isSet B - → (∀{w x} → f w ≡ f x → w ≡ x) - → isEmbedding f -injEmbedding {f = f} iSA iSB inj w x - = isoToIsEquiv (iso (cong f) inj sect retr) - where - sect : section (cong f) inj - sect p = iSB (f w) (f x) _ p - - retr : retract (cong f) inj - retr p = iSA w x _ p - -private - lemma₀ : (p : y ≡ z) → fiber f y ≡ fiber f z - lemma₀ {f = f} p = λ i → fiber f (p i) - - lemma₁ : isEmbedding f → ∀ x → isContr (fiber f (f x)) - lemma₁ {f = f} iE x = value , path - where - value : fiber f (f x) - value = (x , refl) - - path : ∀(fi : fiber f (f x)) → value ≡ fi - path (w , p) i - = case equiv-proof (iE w x) p of λ - { ((q , sq) , _) - → hfill (λ j → λ { (i = i0) → (x , refl) - ; (i = i1) → (w , sq j) - }) - (inS (q (~ i) , λ j → f (q (~ i ∨ j)))) - i1 - } - --- If `f` is an embedding, we'd expect the fibers of `f` to be --- propositions, like an injective function. -hasPropFibers : (A → B) → Type _ -hasPropFibers f = ∀ y → isProp (fiber f y) - -hasPropFibersIsProp : isProp (hasPropFibers f) -hasPropFibersIsProp = isPropPi (λ _ → isPropIsProp) - -isEmbedding→hasPropFibers : isEmbedding f → hasPropFibers f -isEmbedding→hasPropFibers iE y (x , p) - = subst (λ f → isProp f) (lemma₀ p) (isContr→isProp (lemma₁ iE x)) (x , p) - -private - fibCong→PathP - : {f : A → B} - → (p : f w ≡ f x) - → (fi : fiber (cong f) p) - → PathP (λ i → fiber f (p i)) (w , refl) (x , refl) - fibCong→PathP p (q , r) i = q i , λ j → r j i - - PathP→fibCong - : {f : A → B} - → (p : f w ≡ f x) - → (pp : PathP (λ i → fiber f (p i)) (w , refl) (x , refl)) - → fiber (cong f) p - PathP→fibCong p pp = (λ i → fst (pp i)) , (λ j i → snd (pp i) j) - -PathP≡fibCong - : {f : A → B} - → (p : f w ≡ f x) - → PathP (λ i → fiber f (p i)) (w , refl) (x , refl) ≡ fiber (cong f) p -PathP≡fibCong p - = isoToPath (iso (PathP→fibCong p) (fibCong→PathP p) (λ _ → refl) (λ _ → refl)) - -hasPropFibers→isEmbedding : hasPropFibers f → isEmbedding f -hasPropFibers→isEmbedding {f = f} iP w x .equiv-proof p - = subst isContr (PathP≡fibCong p) (isProp→isContrPathP (λ i → iP (p i)) fw fx) - where - fw : fiber f (f w) - fw = (w , refl) - - fx : fiber f (f x) - fx = (x , refl) - -isEmbedding≡hasPropFibers : isEmbedding f ≡ hasPropFibers f -isEmbedding≡hasPropFibers - = isoToPath - (iso isEmbedding→hasPropFibers - hasPropFibers→isEmbedding - (λ _ → hasPropFibersIsProp _ _) - (λ _ → isEmbeddingIsProp _ _)) - -isEquiv→hasPropFibers : isEquiv f → hasPropFibers f -isEquiv→hasPropFibers e b = isContr→isProp (equiv-proof e b) - -isEquiv→isEmbedding : isEquiv f → isEmbedding f -isEquiv→isEmbedding e = hasPropFibers→isEmbedding (isEquiv→hasPropFibers e) diff --git a/Cubical/Foundations/Equiv.agda b/Cubical/Foundations/Equiv.agda index cf8fc610a..9af465309 100644 --- a/Cubical/Foundations/Equiv.agda +++ b/Cubical/Foundations/Equiv.agda @@ -14,7 +14,7 @@ There are more statements about equivalences in Equiv/Properties.agda: - if f is an equivalence then postcomposition with f is an equivalence -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Equiv where open import Cubical.Foundations.Function @@ -22,34 +22,24 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws -open import Cubical.Core.Glue public - using ( isEquiv ; equiv-proof ; _≃_ ; equivFun ; equivProof ) +open import Cubical.Foundations.Equiv.Base public +open import Cubical.Data.Sigma.Base private variable ℓ ℓ' ℓ'' : Level - A B C : Type ℓ + A B C D : Type ℓ -fiber : ∀ {A : Type ℓ} {B : Type ℓ'} (f : A → B) (y : B) → Type (ℓ-max ℓ ℓ') -fiber {A = A} f y = Σ[ x ∈ A ] f x ≡ y - -equivIsEquiv : ∀ {A : Type ℓ} {B : Type ℓ'} (e : A ≃ B) → isEquiv (equivFun e) +equivIsEquiv : (e : A ≃ B) → isEquiv (equivFun e) equivIsEquiv e = snd e -equivCtr : ∀ {A : Type ℓ} {B : Type ℓ'} (e : A ≃ B) (y : B) → fiber (equivFun e) y +equivCtr : (e : A ≃ B) (y : B) → fiber (equivFun e) y equivCtr e y = e .snd .equiv-proof y .fst -equivCtrPath : ∀ {A : Type ℓ} {B : Type ℓ'} (e : A ≃ B) (y : B) → +equivCtrPath : (e : A ≃ B) (y : B) → (v : fiber (equivFun e) y) → Path _ (equivCtr e y) v equivCtrPath e y = e .snd .equiv-proof y .snd --- The identity equivalence -idIsEquiv : ∀ (A : Type ℓ) → isEquiv (idfun A) -equiv-proof (idIsEquiv A) y = - ((y , refl) , λ z i → z .snd (~ i) , λ j → z .snd (~ i ∨ j)) - -idEquiv : ∀ (A : Type ℓ) → A ≃ A -idEquiv A = (idfun A , idIsEquiv A) -- Proof using isPropIsContr. This is slow and the direct proof below is better @@ -72,88 +62,212 @@ equiv-proof (isPropIsEquiv f p q i) y = ; (j = i1) → w }) (p2 w (i ∨ j)) -equivEq : (e f : A ≃ B) → (h : e .fst ≡ f .fst) → e ≡ f -equivEq e f h = λ i → (h i) , isProp→PathP (λ i → isPropIsEquiv (h i)) (e .snd) (f .snd) i +equivEq : {e f : A ≃ B} → (h : e .fst ≡ f .fst) → e ≡ f +equivEq {e = e} {f = f} h = λ i → (h i) , isProp→PathP (λ i → isPropIsEquiv (h i)) (e .snd) (f .snd) i -module _ (w : A ≃ B) where - invEq : B → A - invEq y = fst (fst (snd w .equiv-proof y)) +module _ {f : A → B} (equivF : isEquiv f) where + funIsEq : A → B + funIsEq = f - secEq : section invEq (w .fst) - secEq x = λ i → fst (snd (snd w .equiv-proof (fst w x)) (x , (λ j → fst w x)) i) + invIsEq : B → A + invIsEq y = equivF .equiv-proof y .fst .fst - retEq : retract invEq (w .fst) - retEq y = λ i → snd (fst (snd w .equiv-proof y)) i + secIsEq : section f invIsEq + secIsEq y = equivF .equiv-proof y .fst .snd -equivToIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → A ≃ B → Iso A B -equivToIso {A = A} {B = B} e = iso (e .fst) (invEq e ) (retEq e) (secEq e) + retIsEq : retract f invIsEq + retIsEq a i = equivF .equiv-proof (f a) .snd (a , refl) i .fst -invEquiv : A ≃ B → B ≃ A -invEquiv f = isoToEquiv (iso (invEq f) (fst f) (secEq f) (retEq f)) + commSqIsEq : ∀ a → Square (secIsEq (f a)) refl (cong f (retIsEq a)) refl + commSqIsEq a i = equivF .equiv-proof (f a) .snd (a , refl) i .snd -invEquivIdEquiv : (A : Type ℓ) → invEquiv (idEquiv A) ≡ idEquiv A -invEquivIdEquiv _ = equivEq _ _ refl + commPathIsEq : ∀ a → secIsEq (f a) ≡ cong f (retIsEq a) + commPathIsEq a i j = + hcomp + (λ k → λ + { (i = i0) → secIsEq (f a) j + ; (i = i1) → f (retIsEq a (j ∨ ~ k)) + ; (j = i0) → f (retIsEq a (i ∧ ~ k)) + ; (j = i1) → f a + }) + (commSqIsEq a i j) -compEquiv : A ≃ B → B ≃ C → A ≃ C -compEquiv f g = isoToEquiv - (iso (λ x → g .fst (f .fst x)) - (λ x → invEq f (invEq g x)) - (λ y → (cong (g .fst) (retEq f (invEq g y))) ∙ (retEq g y)) - (λ y → (cong (invEq f) (secEq g (f .fst y))) ∙ (secEq f y))) +module _ (w : A ≃ B) where + invEq : B → A + invEq = invIsEq (snd w) -compEquivIdEquiv : {A B : Type ℓ} (e : A ≃ B) → compEquiv (idEquiv A) e ≡ e -compEquivIdEquiv e = equivEq _ _ refl + retEq : retract (w .fst) invEq + retEq = retIsEq (snd w) -LiftEquiv : {A : Type ℓ} → A ≃ Lift {i = ℓ} {j = ℓ'} A -LiftEquiv = isoToEquiv (iso lift lower (λ _ → refl) (λ _ → refl)) + secEq : section (w .fst) invEq + secEq = secIsEq (snd w) --- module _ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} where --- invEquivInvol : (f : A ≃ B) → invEquiv (invEquiv f) ≡ f --- invEquivInvol f i .fst = fst f --- invEquivInvol f i .snd = propIsEquiv (fst f) (snd (invEquiv (invEquiv f))) (snd f) i +open Iso -Contr→Equiv : isContr A → isContr B → A ≃ B -Contr→Equiv Actr Bctr = isoToEquiv (iso (λ _ → fst Bctr) (λ _ → fst Actr) (snd Bctr) (snd Actr)) +equivToIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → A ≃ B → Iso A B +fun (equivToIso e) = e .fst +inv (equivToIso e) = invEq e +rightInv (equivToIso e) = secEq e +leftInv (equivToIso e) = retEq e -PropEquiv→Equiv : (Aprop : isProp A) (Bprop : isProp B) (f : A → B) (g : B → A) → (A ≃ B) -PropEquiv→Equiv Aprop Bprop f g = isoToEquiv (iso f g (λ b → Bprop (f (g b)) b) λ a → Aprop (g (f a)) a) +-- TODO: there should be a direct proof of this that doesn't use equivToIso +invEquiv : A ≃ B → B ≃ A +invEquiv e = isoToEquiv (invIso (equivToIso e)) + +invEquivIdEquiv : (A : Type ℓ) → invEquiv (idEquiv A) ≡ idEquiv A +invEquivIdEquiv _ = equivEq refl -homotopyNatural : {f g : A → B} (H : ∀ a → f a ≡ g a) {x y : A} (p : x ≡ y) → - H x ∙ cong g p ≡ cong f p ∙ H y -homotopyNatural H p = homotopyNatural' H p ∙ □≡∙ _ _ +compEquiv : A ≃ B → B ≃ C → A ≃ C +compEquiv f g .fst = g .fst ∘ f .fst +compEquiv {A = A} {C = C} f g .snd .equiv-proof c = contr + where + contractG = g .snd .equiv-proof c .snd + + secFiller : (a : A) (p : g .fst (f .fst a) ≡ c) → _ {- square in A -} + secFiller a p = compPath-filler (cong (invEq f ∘ fst) (contractG (_ , p))) (retEq f a) + + contr : isContr (fiber (g .fst ∘ f .fst) c) + contr .fst .fst = invEq f (invEq g c) + contr .fst .snd = cong (g .fst) (secEq f (invEq g c)) ∙ secEq g c + contr .snd (a , p) i .fst = secFiller a p i1 i + contr .snd (a , p) i .snd j = + hcomp + (λ k → λ + { (i = i1) → fSquare k + ; (j = i0) → g .fst (f .fst (secFiller a p k i)) + ; (j = i1) → contractG (_ , p) i .snd k + }) + (g .fst (secEq f (contractG (_ , p) i .fst) j)) + where + fSquare : I → C + fSquare k = + hcomp + (λ l → λ + { (j = i0) → g .fst (f .fst (retEq f a k)) + ; (j = i1) → p (k ∧ l) + ; (k = i0) → g .fst (secEq f (f .fst a) j) + ; (k = i1) → p (j ∧ l) + }) + (g .fst (f .snd .equiv-proof (f .fst a) .snd (a , refl) k .snd j)) + +compEquivIdEquiv : (e : A ≃ B) → compEquiv (idEquiv A) e ≡ e +compEquivIdEquiv e = equivEq refl + +compEquivEquivId : (e : A ≃ B) → compEquiv e (idEquiv B) ≡ e +compEquivEquivId e = equivEq refl + +invEquiv-is-rinv : (e : A ≃ B) → compEquiv e (invEquiv e) ≡ idEquiv A +invEquiv-is-rinv e = equivEq (funExt (retEq e)) + +invEquiv-is-linv : (e : A ≃ B) → compEquiv (invEquiv e) e ≡ idEquiv B +invEquiv-is-linv e = equivEq (funExt (secEq e)) + +compEquiv-assoc : (f : A ≃ B) (g : B ≃ C) (h : C ≃ D) + → compEquiv f (compEquiv g h) ≡ compEquiv (compEquiv f g) h +compEquiv-assoc f g h = equivEq refl + +LiftEquiv : A ≃ Lift {i = ℓ} {j = ℓ'} A +LiftEquiv .fst a .lower = a +LiftEquiv .snd .equiv-proof = strictContrFibers lower + +Lift≃Lift : (e : A ≃ B) → Lift {j = ℓ'} A ≃ Lift {j = ℓ''} B +Lift≃Lift e .fst a .lower = e .fst (a .lower) +Lift≃Lift e .snd .equiv-proof b .fst .fst .lower = invEq e (b .lower) +Lift≃Lift e .snd .equiv-proof b .fst .snd i .lower = + e .snd .equiv-proof (b .lower) .fst .snd i +Lift≃Lift e .snd .equiv-proof b .snd (a , p) i .fst .lower = + e .snd .equiv-proof (b .lower) .snd (a .lower , cong lower p) i .fst +Lift≃Lift e .snd .equiv-proof b .snd (a , p) i .snd j .lower = + e .snd .equiv-proof (b .lower) .snd (a .lower , cong lower p) i .snd j + +isContr→Equiv : isContr A → isContr B → A ≃ B +isContr→Equiv Actr Bctr = isoToEquiv (isContr→Iso Actr Bctr) + +propBiimpl→Equiv : (Aprop : isProp A) (Bprop : isProp B) (f : A → B) (g : B → A) → A ≃ B +propBiimpl→Equiv Aprop Bprop f g = f , hf where - homotopyNatural' : {f g : A → B} (H : ∀ a → f a ≡ g a) {x y : A} (p : x ≡ y) → - H x ∙ cong g p ≡ cong f p □ H y - homotopyNatural' {f = f} {g = g} H {x} {y} p i j = - hcomp (λ k → λ { (i = i0) → compPath-filler (H x) (cong g p) k j - ; (i = i1) → compPath'-filler (cong f p) (H y) k j - ; (j = i0) → cong f p (i ∧ (~ k)) - ; (j = i1) → cong g p (i ∨ k) }) - (H (p i) j) - -Hfa≡fHa : ∀ {A : Type ℓ} (f : A → A) → (H : ∀ a → f a ≡ a) → ∀ a → H (f a) ≡ cong f (H a) -Hfa≡fHa {A = A} f H a = - H (f a) ≡⟨ rUnit (H (f a)) ⟩ - H (f a) ∙ refl ≡⟨ cong (_∙_ (H (f a))) (sym (rCancel (H a))) ⟩ - H (f a) ∙ H a ∙ sym (H a) ≡⟨ assoc _ _ _ ⟩ - (H (f a) ∙ H a) ∙ sym (H a) ≡⟨ cong (λ x → x ∙ (sym (H a))) (homotopyNatural H (H a)) ⟩ - (cong f (H a) ∙ H a) ∙ sym (H a) ≡⟨ sym (assoc _ _ _) ⟩ - cong f (H a) ∙ H a ∙ sym (H a) ≡⟨ cong (_∙_ (cong f (H a))) (rCancel _) ⟩ - cong f (H a) ∙ refl ≡⟨ sym (rUnit _) ⟩ - cong f (H a) ∎ - -equivPi - : ∀{F : A → Set ℓ} {G : A → Set ℓ'} - → ((x : A) → F x ≃ G x) → (((x : A) → F x) ≃ ((x : A) → G x)) -equivPi k .fst f x = k x .fst (f x) -equivPi k .snd .equiv-proof f - .fst .fst x = equivCtr (k x) (f x) .fst -equivPi k .snd .equiv-proof f - .fst .snd i x = equivCtr (k x) (f x) .snd i -equivPi k .snd .equiv-proof f - .snd (g , p) i .fst x = equivCtrPath (k x) (f x) (g x , λ j → p j x) i .fst -equivPi k .snd .equiv-proof f - .snd (g , p) i .snd j x = equivCtrPath (k x) (f x) (g x , λ k → p k x) i .snd j + hf : isEquiv f + hf .equiv-proof y .fst = (g y , Bprop (f (g y)) y) + hf .equiv-proof y .snd h i .fst = Aprop (g y) (h .fst) i + hf .equiv-proof y .snd h i .snd = isProp→isSet' Bprop (Bprop (f (g y)) y) (h .snd) + (cong f (Aprop (g y) (h .fst))) refl i + +isEquivPropBiimpl→Equiv : isProp A → isProp B + → ((A → B) × (B → A)) ≃ (A ≃ B) +isEquivPropBiimpl→Equiv {A = A} {B = B} Aprop Bprop = isoToEquiv isom where + isom : Iso (Σ (A → B) (λ _ → B → A)) (A ≃ B) + isom .fun (f , g) = propBiimpl→Equiv Aprop Bprop f g + isom .inv e = equivFun e , invEq e + isom .rightInv e = equivEq refl + isom .leftInv _ = refl + +equivΠCod : ∀ {F : A → Type ℓ} {G : A → Type ℓ'} + → ((x : A) → F x ≃ G x) → ((x : A) → F x) ≃ ((x : A) → G x) +equivΠCod k .fst f x = k x .fst (f x) +equivΠCod k .snd .equiv-proof f .fst .fst x = equivCtr (k x) (f x) .fst +equivΠCod k .snd .equiv-proof f .fst .snd i x = equivCtr (k x) (f x) .snd i +equivΠCod k .snd .equiv-proof f .snd (g , p) i .fst x = + equivCtrPath (k x) (f x) (g x , λ j → p j x) i .fst +equivΠCod k .snd .equiv-proof f .snd (g , p) i .snd j x = + equivCtrPath (k x) (f x) (g x , λ k → p k x) i .snd j + +equivImplicitΠCod : ∀ {F : A → Type ℓ} {G : A → Type ℓ'} + → ({x : A} → F x ≃ G x) → ({x : A} → F x) ≃ ({x : A} → G x) +equivImplicitΠCod k .fst f {x} = k {x} .fst (f {x}) +equivImplicitΠCod k .snd .equiv-proof f .fst .fst {x} = equivCtr (k {x}) (f {x}) .fst +equivImplicitΠCod k .snd .equiv-proof f .fst .snd i {x} = equivCtr (k {x}) (f {x}) .snd i +equivImplicitΠCod k .snd .equiv-proof f .snd (g , p) i .fst {x} = + equivCtrPath (k {x}) (f {x}) (g {x} , λ j → p j {x}) i .fst +equivImplicitΠCod k .snd .equiv-proof f .snd (g , p) i .snd j {x} = + equivCtrPath (k {x}) (f {x}) (g {x} , λ k → p k {x}) i .snd j + +equiv→Iso : (A ≃ B) → (C ≃ D) → Iso (A → C) (B → D) +equiv→Iso h k .Iso.fun f b = equivFun k (f (invEq h b)) +equiv→Iso h k .Iso.inv g a = invEq k (g (equivFun h a)) +equiv→Iso h k .Iso.rightInv g = funExt λ b → secEq k _ ∙ cong g (secEq h b) +equiv→Iso h k .Iso.leftInv f = funExt λ a → retEq k _ ∙ cong f (retEq h a) + +equiv→ : (A ≃ B) → (C ≃ D) → (A → C) ≃ (B → D) +equiv→ h k = isoToEquiv (equiv→Iso h k) + +equivΠ : ∀ {ℓA ℓA' ℓB ℓB'} {A : Type ℓA} {A' : Type ℓA'} + {B : A → Type ℓB} {B' : A' → Type ℓB'} + (eA : A ≃ A') + (eB : (a : A) → B a ≃ B' (eA .fst a)) + → ((a : A) → B a) ≃ ((a' : A') → B' a') +equivΠ {B' = B'} eA eB = isoToEquiv isom + where + open Iso + + isom : Iso _ _ + isom .fun f a' = + subst B' (secEq eA a') (eB _ .fst (f (invEq eA a'))) + isom .inv f' a = + invEq (eB _) (f' (eA .fst a)) + isom .rightInv f' = + funExt λ a' → + cong (subst B' (secEq eA a')) (secEq (eB _) _) + ∙ fromPathP (cong f' (secEq eA a')) + isom .leftInv f = + funExt λ a → + invEq (eB a) (subst B' (secEq eA _) (eB _ .fst (f (invEq eA (eA .fst a))))) + ≡⟨ cong (λ t → invEq (eB a) (subst B' t (eB _ .fst (f (invEq eA (eA .fst a)))))) + (commPathIsEq (snd eA) a) ⟩ + invEq (eB a) (subst B' (cong (eA .fst) (retEq eA a)) (eB _ .fst (f (invEq eA (eA .fst a))))) + ≡⟨ cong (invEq (eB a)) (fromPathP (λ i → eB _ .fst (f (retEq eA a i)))) ⟩ + invEq (eB a) (eB a .fst (f a)) + ≡⟨ retEq (eB _) (f a) ⟩ + f a + ∎ + +equivCompIso : (A ≃ B) → (C ≃ D) → Iso (A ≃ C) (B ≃ D) +equivCompIso h k .Iso.fun f = compEquiv (compEquiv (invEquiv h) f) k +equivCompIso h k .Iso.inv g = compEquiv (compEquiv h g) (invEquiv k) +equivCompIso h k .Iso.rightInv g = equivEq (equiv→Iso h k .Iso.rightInv (equivFun g)) +equivCompIso h k .Iso.leftInv f = equivEq (equiv→Iso h k .Iso.leftInv (equivFun f)) + +equivComp : (A ≃ B) → (C ≃ D) → (A ≃ C) ≃ (B ≃ D) +equivComp h k = isoToEquiv (equivCompIso h k) -- Some helpful notation: _≃⟨_⟩_ : (X : Type ℓ) → (X ≃ B) → (B ≃ C) → (X ≃ C) @@ -164,3 +278,12 @@ _■ = idEquiv infixr 0 _≃⟨_⟩_ infix 1 _■ + +composesToId→Equiv : (f : A → B) (g : B → A) → f ∘ g ≡ idfun B → isEquiv f → isEquiv g +composesToId→Equiv f g id iseqf = + isoToIsEquiv + (iso g f + (λ b → (λ i → equiv-proof iseqf (f b) .snd (g (f b) , cong (λ h → h (f b)) id) (~ i) .fst) + ∙∙ cong (λ x → equiv-proof iseqf (f b) .fst .fst) id + ∙∙ λ i → equiv-proof iseqf (f b) .snd (b , refl) i .fst) + λ a i → id i a) diff --git a/Cubical/Foundations/Equiv/Base.agda b/Cubical/Foundations/Equiv/Base.agda new file mode 100644 index 000000000..ff35a2028 --- /dev/null +++ b/Cubical/Foundations/Equiv/Base.agda @@ -0,0 +1,27 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv.Base where + +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude + +open import Cubical.Core.Glue public + using ( isEquiv ; equiv-proof ; _≃_ ; equivFun ; equivProof ) + +fiber : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) (y : B) → Type (ℓ-max ℓ ℓ') +fiber {A = A} f y = Σ[ x ∈ A ] f x ≡ y + +-- Helper function for constructing equivalences from pairs (f,g) that cancel each other up to definitional +-- equality. For such (f,g), the result type simplifies to isContr (fiber f b). +strictContrFibers : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {f : A → B} (g : B → A) (b : B) + → Σ[ t ∈ fiber f (f (g b)) ] + ((t' : fiber f b) → Path (fiber f (f (g b))) t (g (f (t' .fst)) , cong (f ∘ g) (t' .snd))) +strictContrFibers {f = f} g b .fst = (g b , refl) +strictContrFibers {f = f} g b .snd (a , p) i = (g (p (~ i)) , λ j → f (g (p (~ i ∨ j)))) + +-- The identity equivalence +idIsEquiv : ∀ {ℓ} (A : Type ℓ) → isEquiv (idfun A) +idIsEquiv A .equiv-proof = strictContrFibers (idfun A) + +idEquiv : ∀ {ℓ} (A : Type ℓ) → A ≃ A +idEquiv A .fst = idfun A +idEquiv A .snd = idIsEquiv A diff --git a/Cubical/Foundations/BiInvEquiv.agda b/Cubical/Foundations/Equiv/BiInvertible.agda similarity index 87% rename from Cubical/Foundations/BiInvEquiv.agda rename to Cubical/Foundations/Equiv/BiInvertible.agda index a7839f9c3..383852814 100644 --- a/Cubical/Foundations/BiInvEquiv.agda +++ b/Cubical/Foundations/Equiv/BiInvertible.agda @@ -8,14 +8,14 @@ Some theory about Bi-Invertible Equivalences - Iso to BiInvEquiv -} -{-# OPTIONS --cubical --safe #-} -module Cubical.Foundations.BiInvEquiv where +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv.BiInvertible 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.HAEquiv +open import Cubical.Foundations.Equiv.HalfAdjoint record BiInvEquiv {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') : Type (ℓ-max ℓ ℓ') where @@ -33,16 +33,16 @@ record BiInvEquiv {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') : Type (ℓ-max ℓ invl b ∎ invr-leftInv : retract fun invr - invr-leftInv a = invr≡invl (fun a) □ (invl-leftInv a) + invr-leftInv a = invr≡invl (fun a) ∙ (invl-leftInv a) invr≡invl-leftInv : ∀ a → PathP (λ j → invr≡invl (fun a) j ≡ a) (invr-leftInv a) (invl-leftInv a) - invr≡invl-leftInv a j i = compPath'-filler (invr≡invl (fun a)) (invl-leftInv a) (~ j) i + invr≡invl-leftInv a j i = compPath-filler' (invr≡invl (fun a)) (invl-leftInv a) (~ j) i invl-rightInv : section fun invl - invl-rightInv a = sym (cong fun (invr≡invl a)) □ (invr-rightInv a) + invl-rightInv a = sym (cong fun (invr≡invl a)) ∙ (invr-rightInv a) invr≡invl-rightInv : ∀ a → PathP (λ j → fun (invr≡invl a j) ≡ a) (invr-rightInv a) (invl-rightInv a) - invr≡invl-rightInv a j i = compPath'-filler (sym (cong fun (invr≡invl a))) (invr-rightInv a) j i + invr≡invl-rightInv a j i = compPath-filler' (sym (cong fun (invr≡invl a))) (invr-rightInv a) j i module _ {ℓ} {A B : Type ℓ} (e : BiInvEquiv A B) where diff --git a/Cubical/Foundations/Equiv/Fiberwise.agda b/Cubical/Foundations/Equiv/Fiberwise.agda index dc4b8dcf1..b0e0db5b3 100644 --- a/Cubical/Foundations/Equiv/Fiberwise.agda +++ b/Cubical/Foundations/Equiv/Fiberwise.agda @@ -1,9 +1,13 @@ -{-# OPTIONS --cubical #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Equiv.Fiberwise where open import Cubical.Core.Everything -open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma module _ {a p q} {A : Type a} (P : A → Type p) (Q : A → Type q) (f : ∀ x → P x → Q x) @@ -26,23 +30,28 @@ module _ {a p q} {A : Type a} (P : A → Type p) (Q : A → Type q) g-h {xv} ((a , p) , eq) = J (λ _ eq₁ → g (h ((a , p) , eq₁)) ≡ ((a , p) , eq₁)) (cong g (JRefl (λ xv₁ eq₁ → fiber (f (xv₁ .fst)) (xv₁ .snd)) (p , refl))) eq - -- half of Thm 4.7.7 (fiberwise equivalences) + -- Thm 4.7.7 (fiberwise equivalences) fiberEquiv : ([tf] : isEquiv total) → ∀ x → isEquiv (f x) fiberEquiv [tf] x .equiv-proof y = isContrRetract (fibers-total .Iso.inv) (fibers-total .Iso.fun) (fibers-total .Iso.rightInv) ([tf] .equiv-proof (x , y)) + totalEquiv : (fx-equiv : ∀ x → isEquiv (f x)) + → isEquiv total + totalEquiv fx-equiv .equiv-proof (x , v) = isContrRetract (fibers-total .Iso.fun) (fibers-total .Iso.inv) (fibers-total .Iso.leftInv) + (fx-equiv x .equiv-proof v) + module _ {ℓ : Level} {U : Type ℓ} {ℓr} (_~_ : U → U → Type ℓr) (idTo~ : ∀ {A B} → A ≡ B → A ~ B) - (c : ∀ A → isContr (Σ U \ X → A ~ X)) + (c : ∀ A → ∃![ X ∈ U ] (A ~ X)) where isContrToUniv : ∀ {A B} → isEquiv (idTo~ {A} {B}) isContrToUniv {A} {B} = fiberEquiv (λ z → A ≡ z) (λ z → A ~ z) (\ B → idTo~ {A} {B}) (λ { .equiv-proof y - → isContrΣ ((_ , refl) , (\ z → contrSingl (z .snd))) + → isContrΣ (isContrSingl _) \ a → isContr→isContrPath (c A) _ _ }) B diff --git a/Cubical/Foundations/Equiv/HalfAdjoint.agda b/Cubical/Foundations/Equiv/HalfAdjoint.agda new file mode 100644 index 000000000..e589d5d06 --- /dev/null +++ b/Cubical/Foundations/Equiv/HalfAdjoint.agda @@ -0,0 +1,166 @@ +{- + +Half adjoint equivalences ([HAEquiv]) + +- Iso to HAEquiv ([iso→HAEquiv]) +- Equiv to HAEquiv ([equiv→HAEquiv]) +- Cong is an equivalence ([congEquiv]) + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv.HalfAdjoint where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + +record isHAEquiv {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) : Type (ℓ-max ℓ ℓ') where + field + g : B → A + linv : ∀ a → g (f a) ≡ a + rinv : ∀ b → f (g b) ≡ b + com : ∀ a → cong f (linv a) ≡ rinv (f a) + + -- from redtt's ha-equiv/symm + com-op : ∀ b → cong g (rinv b) ≡ linv (g b) + com-op b j i = hcomp (λ k → λ { (i = i0) → linv (g b) (j ∧ (~ k)) + ; (j = i0) → g (rinv b i) + ; (j = i1) → linv (g b) (i ∨ (~ k)) + ; (i = i1) → g b }) + (cap1 j i) + + where cap0 : Square {- (j = i0) -} (λ i → f (g (rinv b i))) + {- (j = i1) -} (λ i → rinv b i) + {- (i = i0) -} (λ j → f (linv (g b) j)) + {- (i = i1) -} (λ j → rinv b j) + + cap0 j i = hcomp (λ k → λ { (i = i0) → com (g b) (~ k) j + ; (j = i0) → f (g (rinv b i)) + ; (j = i1) → rinv b i + ; (i = i1) → rinv b j }) + (rinv (rinv b i) j) + + filler : I → I → A + filler j i = hfill (λ k → λ { (i = i0) → g (rinv b k) + ; (i = i1) → g b }) + (inS (linv (g b) i)) j + + cap1 : Square {- (j = i0) -} (λ i → g (rinv b i)) + {- (j = i1) -} (λ i → g b) + {- (i = i0) -} (λ j → linv (g b) j) + {- (i = i1) -} (λ j → g b) + + cap1 j i = hcomp (λ k → λ { (i = i0) → linv (linv (g b) j) k + ; (j = i0) → linv (g (rinv b i)) k + ; (j = i1) → filler i k + ; (i = i1) → filler j k }) + (g (cap0 j i)) + + isHAEquiv→Iso : Iso A B + Iso.fun isHAEquiv→Iso = f + Iso.inv isHAEquiv→Iso = g + Iso.rightInv isHAEquiv→Iso = rinv + Iso.leftInv isHAEquiv→Iso = linv + + isHAEquiv→isEquiv : isEquiv f + isHAEquiv→isEquiv .equiv-proof y = (g y , rinv y) , isCenter where + isCenter : ∀ xp → (g y , rinv y) ≡ xp + isCenter (x , p) i = gy≡x i , ry≡p i where + gy≡x : g y ≡ x + gy≡x = sym (cong g p) ∙∙ refl ∙∙ linv x + + lem0 : Square (cong f (linv x)) p (cong f (linv x)) p + lem0 i j = invSides-filler p (sym (cong f (linv x))) (~ i) j + + ry≡p : Square (rinv y) p (cong f gy≡x) refl + ry≡p i j = hcomp (λ k → λ { (i = i0) → cong rinv p k j + ; (i = i1) → lem0 k j + ; (j = i0) → f (doubleCompPath-filler (sym (cong g p)) refl (linv x) k i) + ; (j = i1) → p k }) + (com x (~ i) j) + +open isHAEquiv using (isHAEquiv→Iso; isHAEquiv→isEquiv) public + +HAEquiv : (A : Type ℓ) (B : Type ℓ') → Type (ℓ-max ℓ ℓ') +HAEquiv A B = Σ[ f ∈ (A → B) ] isHAEquiv f + +-- vogt's lemma (https://ncatlab.org/nlab/show/homotopy+equivalence#vogts_lemma) +iso→HAEquiv : Iso A B → HAEquiv A B +iso→HAEquiv e = f , isHAEquivf + where + f = Iso.fun e + g = Iso.inv e + ε = Iso.rightInv e + η = Iso.leftInv e + + Hfa≡fHa : (f : A → A) → (H : ∀ a → f a ≡ a) → ∀ a → H (f a) ≡ cong f (H a) + Hfa≡fHa f H = J (λ f p → ∀ a → funExt⁻ (sym p) (f a) ≡ cong f (funExt⁻ (sym p) a)) + (λ a → refl) + (sym (funExt H)) + + isHAEquivf : isHAEquiv f + isHAEquiv.g isHAEquivf = g + isHAEquiv.linv isHAEquivf = η + isHAEquiv.rinv isHAEquivf b i = + hcomp (λ j → λ { (i = i0) → ε (f (g b)) j + ; (i = i1) → ε b j }) + (f (η (g b) i)) + isHAEquiv.com isHAEquivf a i j = + hcomp (λ k → λ { (i = i0) → ε (f (η a j)) k + ; (j = i0) → ε (f (g (f a))) k + ; (j = i1) → ε (f a) k}) + (f (Hfa≡fHa (λ x → g (f x)) η a (~ i) j)) + +equiv→HAEquiv : A ≃ B → HAEquiv A B +equiv→HAEquiv e = e .fst , λ where + .isHAEquiv.g → invIsEq (snd e) + .isHAEquiv.linv → retIsEq (snd e) + .isHAEquiv.rinv → secIsEq (snd e) + .isHAEquiv.com a → sym (commPathIsEq (snd e) a) + +congIso : {x y : A} (e : Iso A B) → Iso (x ≡ y) (Iso.fun e x ≡ Iso.fun e y) +congIso {x = x} {y} e = goal + where + open isHAEquiv (iso→HAEquiv e .snd) + open Iso + + goal : Iso (x ≡ y) (Iso.fun e x ≡ Iso.fun e y) + fun goal = cong (iso→HAEquiv e .fst) + inv goal p = sym (linv x) ∙∙ cong g p ∙∙ linv y + rightInv goal p i j = + hcomp (λ k → λ { (i = i0) → iso→HAEquiv e .fst + (doubleCompPath-filler (sym (linv x)) (cong g p) (linv y) k j) + ; (i = i1) → rinv (p j) k + ; (j = i0) → com x i k + ; (j = i1) → com y i k }) + (iso→HAEquiv e .fst (g (p j))) + leftInv goal p i j = + hcomp (λ k → λ { (i = i1) → p j + ; (j = i0) → Iso.leftInv e x (i ∨ k) + ; (j = i1) → Iso.leftInv e y (i ∨ k) }) + (Iso.leftInv e (p j) i) + +invCongFunct : {x : A} (e : Iso A B) (p : Iso.fun e x ≡ Iso.fun e x) (q : Iso.fun e x ≡ Iso.fun e x) + → Iso.inv (congIso e) (p ∙ q) ≡ Iso.inv (congIso e) p ∙ Iso.inv (congIso e) q +invCongFunct {x = x} e p q = helper (Iso.inv e) _ _ _ + where + helper : {x : A} {y : B} (f : A → B) (r : f x ≡ y) (p q : x ≡ x) + → (sym r ∙∙ cong f (p ∙ q) ∙∙ r) ≡ (sym r ∙∙ cong f p ∙∙ r) ∙ (sym r ∙∙ cong f q ∙∙ r) + helper {x = x} f = + J (λ y r → (p q : x ≡ x) + → (sym r ∙∙ cong f (p ∙ q) ∙∙ r) ≡ (sym r ∙∙ cong f p ∙∙ r) ∙ (sym r ∙∙ cong f q ∙∙ r)) + λ p q → (λ i → rUnit (congFunct f p q i) (~ i)) + ∙ λ i → rUnit (cong f p) i ∙ rUnit (cong f q) i + +invCongRefl : {x : A} (e : Iso A B) → Iso.inv (congIso {x = x} {y = x} e) refl ≡ refl +invCongRefl {x = x} e = (λ i → (λ j → Iso.leftInv e x (i ∨ ~ j)) ∙∙ refl ∙∙ (λ j → Iso.leftInv e x (i ∨ j))) ∙ sym (rUnit refl) diff --git a/Cubical/Foundations/PathSplitEquiv.agda b/Cubical/Foundations/Equiv/PathSplit.agda similarity index 73% rename from Cubical/Foundations/PathSplitEquiv.agda rename to Cubical/Foundations/Equiv/PathSplit.agda index c64fab187..8e1f4ca60 100644 --- a/Cubical/Foundations/PathSplitEquiv.agda +++ b/Cubical/Foundations/Equiv/PathSplit.agda @@ -15,8 +15,8 @@ The module starts with a couple of general facts about equivalences: (those are not in 'Equiv.agda' because they need Univalence.agda (which imports Equiv.agda)) -} -{-# OPTIONS --cubical --safe #-} -module Cubical.Foundations.PathSplitEquiv where +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv.PathSplit where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function @@ -26,6 +26,8 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv.Properties +open import Cubical.Data.Sigma + record isPathSplitEquiv {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) : Type (ℓ-max ℓ ℓ') where field sec : hasSection f @@ -37,35 +39,36 @@ PathSplitEquiv A B = Σ[ f ∈ (A → B) ] isPathSplitEquiv f open isPathSplitEquiv idIsPathSplitEquiv : ∀ {ℓ} {A : Type ℓ} → isPathSplitEquiv (λ (x : A) → x) -sec idIsPathSplitEquiv = (λ x → x) , (λ x → refl) -secCong idIsPathSplitEquiv = λ x y → (λ p → p) , λ p _ → p +sec idIsPathSplitEquiv = (λ x → x) , (λ _ → refl) +secCong idIsPathSplitEquiv = λ _ _ → (λ p → p) , λ p _ → p module _ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} where + + open Iso + + toIso : (f : A → B) → isPathSplitEquiv f → Iso A B + fun (toIso f _) = f + inv (toIso _ p) = p .sec .fst + rightInv (toIso _ p) = p .sec .snd + leftInv (toIso f p) x = p .secCong (p .sec .fst (f x)) x .fst (p .sec .snd (f x)) + toIsEquiv : (f : A → B) → isPathSplitEquiv f → isEquiv f - toIsEquiv f record { sec = sec ; secCong = secCong } = - (isoToEquiv (iso f (fst sec) (snd sec) (λ x → (secCong (fst sec (f x)) x).fst (snd sec (f x))))) .snd + toIsEquiv f p = isoToIsEquiv (toIso f p) sectionOfEquiv' : (f : A → B) → isEquiv f → B → A - sectionOfEquiv' f record { equiv-proof = all-fibers-contractible } x = - all-fibers-contractible x .fst .fst + sectionOfEquiv' f isEqv x = isEqv .equiv-proof x .fst .fst isSec : (f : A → B) → (pf : isEquiv f) → section f (sectionOfEquiv' f pf) - isSec f record { equiv-proof = all-fibers-contractible } x = - all-fibers-contractible x .fst .snd + isSec f isEqv x = isEqv .equiv-proof x .fst .snd sectionOfEquiv : (f : A → B) → isEquiv f → hasSection f sectionOfEquiv f e = sectionOfEquiv' f e , isSec f e module _ {ℓ} {A B : Type ℓ} where - abstract - fromIsEquiv : (f : A → B) → isEquiv f → isPathSplitEquiv f - sec (fromIsEquiv f pf) = sectionOfEquiv' f pf , isSec f pf - secCong (fromIsEquiv f pf) x y = sectionOfEquiv cong-f eq-cong - where - cong-f : x ≡ y → f x ≡ f y - cong-f = λ (p : x ≡ y) → cong f p - eq-cong : isEquiv cong-f - eq-cong = isEquivCong (f , pf) + + fromIsEquiv : (f : A → B) → isEquiv f → isPathSplitEquiv f + sec (fromIsEquiv f pf) = sectionOfEquiv' f pf , isSec f pf + secCong (fromIsEquiv f pf) x y = sectionOfEquiv (cong f) (isEquivCong (f , pf)) pathSplitToEquiv : PathSplitEquiv A B → A ≃ B fst (pathSplitToEquiv (f , _)) = f @@ -75,28 +78,22 @@ module _ {ℓ} {A B : Type ℓ} where fst (equivToPathSplit (f , _)) = f snd (equivToPathSplit (_ , e)) = fromIsEquiv _ e - equivHasUniqueSection : (f : A → B) - → isEquiv f → isContr (Σ (B → A) (section f)) + equivHasUniqueSection : (f : A → B) → isEquiv f → ∃![ g ∈ (B → A) ] section f g equivHasUniqueSection f eq = helper' where - idB = λ (x : B) → x - abstract - helper : isContr (fiber (λ (φ : B → A) → f ∘ φ) idB) - helper = (equiv-proof (snd (postCompEquiv (f , eq)))) idB - helper' : isContr (Σ[ φ ∈ (B → A) ] ((x : B) → f (φ x) ≡ x)) - fst helper' = (φ , λ x i → η i x) - where φ = fst (fst helper) - η : f ∘ φ ≡ idB - η = snd (fst helper) - (snd helper') y i = (fst (η i) , λ b j → snd (η i) j b) - where η = (snd helper) (fst y , λ i b → snd y b i) + helper : isContr (fiber (λ (φ : B → A) → f ∘ φ) (idfun B)) + helper = (equiv-proof (snd (postCompEquiv (f , eq)))) (idfun B) + + helper' : ∃![ φ ∈ (B → A) ] ((x : B) → f (φ x) ≡ x) + fst helper' = (helper .fst .fst , λ x i → helper .fst .snd i x) + snd helper' y i = (fst (η i) , λ b j → snd (η i) j b) + where η = helper .snd (fst y , λ i b → snd y b i) {- PathSplitEquiv is a proposition and the type of path split equivs is equivalent to the type of equivalences -} -isPropIsPathSplitEquiv : ∀ {ℓ} {A B : Type ℓ} (f : A → B) - → isProp (isPathSplitEquiv f) +isPropIsPathSplitEquiv : ∀ {ℓ} {A B : Type ℓ} (f : A → B) → isProp (isPathSplitEquiv f) isPropIsPathSplitEquiv {_} {A} {B} f record { sec = sec-φ ; secCong = secCong-φ } record { sec = sec-ψ ; secCong = secCong-ψ } i diff --git a/Cubical/Foundations/Equiv/Properties.agda b/Cubical/Foundations/Equiv/Properties.agda index e59190f11..33a776a2b 100644 --- a/Cubical/Foundations/Equiv/Properties.agda +++ b/Cubical/Foundations/Equiv/Properties.agda @@ -6,9 +6,10 @@ A couple of general facts about equivalences: - if f is an equivalence then pre- and postcomposition with f are equivalences ([preCompEquiv], [postCompEquiv]) - if f is an equivalence then (Σ[ g ] section f g) and (Σ[ g ] retract f g) are contractible ([isContr-section], [isContr-retract]) +- isHAEquiv is a proposition [isPropIsHAEquiv] (these are not in 'Equiv.agda' because they need Univalence.agda (which imports Equiv.agda)) -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Equiv.Properties where open import Cubical.Core.Everything @@ -18,64 +19,171 @@ open import Cubical.Data.Sigma open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.Univalence open import Cubical.Foundations.Isomorphism - -open import Cubical.Foundations.FunExtEquiv - - -isEquivCong : ∀ {ℓ} {A B : Type ℓ} {x y : A} (e : A ≃ B) → isEquiv (λ (p : x ≡ y) → (cong (fst e) p)) -isEquivCong e = EquivJ (λ (B' A' : Type _) (e' : A' ≃ B') → - (x' y' : A') → isEquiv (λ (p : x' ≡ y') → cong (fst e') p)) - (λ _ x' y' → idIsEquiv (x' ≡ y')) _ _ e _ _ - -congEquiv : ∀ {ℓ} {A B : Type ℓ} {x y : A} (e : A ≃ B) → (x ≡ y) ≃ (e .fst x ≡ e .fst y) -congEquiv e = ((λ (p : _ ≡ _) → cong (fst e) p) , isEquivCong e) - -isEquivPreComp : ∀ {ℓ ℓ′} {A B : Type ℓ} {C : Type ℓ′} (e : A ≃ B) - → isEquiv (λ (φ : B → C) → φ ∘ e .fst) -isEquivPreComp {A = A} {C = C} e = EquivJ - (λ (B A : Type _) (e' : A ≃ B) → isEquiv (λ (φ : B → C) → φ ∘ e' .fst)) - (λ A → idIsEquiv (A → C)) _ _ e - -isEquivPostComp : ∀ {ℓ ℓ′} {A B : Type ℓ} {C : Type ℓ′} (e : A ≃ B) - → isEquiv (λ (φ : C → A) → e .fst ∘ φ) -isEquivPostComp {A = A} {C = C} e = EquivJ - (λ (B A : Type _) (e' : A ≃ B) → isEquiv (λ (φ : C → A) → e' .fst ∘ φ)) - (λ A → idIsEquiv (C → A)) _ _ e - -preCompEquiv : ∀ {ℓ ℓ′} {A B : Type ℓ} {C : Type ℓ′} (e : A ≃ B) - → (B → C) ≃ (A → C) -preCompEquiv e = (λ φ x → φ (fst e x)) , isEquivPreComp e - -postCompEquiv : ∀ {ℓ ℓ′} {A B : Type ℓ} {C : Type ℓ′} (e : A ≃ B) - → (C → A) ≃ (C → B) -postCompEquiv e = (λ φ x → fst e (φ x)) , isEquivPostComp e - - -hasSection : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (A → B) → Type _ -hasSection {A = A} {B} f = Σ[ g ∈ (B → A) ] section f g - -hasRetract : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (A → B) → Type _ -hasRetract {A = A} {B} f = Σ[ g ∈ (B → A) ] retract f g - -isContr-hasSection : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) → isContr (hasSection (fst e)) -fst (isContr-hasSection e) = invEq e , retEq e -snd (isContr-hasSection e) (f , ε) i = (λ b → fst (p b i)) , (λ b → snd (p b i)) - where p : ∀ b → (invEq e b , retEq e b) ≡ (f b , ε b) - p b = snd (equiv-proof (snd e) b) (f b , ε b) - --- there is a (much slower) alternate proof that also works for retract - -isContr-hasSection' : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) → isContr (hasSection (fst e)) -isContr-hasSection' {_} {A} {B} e = transport (λ i → isContr (Σ[ g ∈ (B → A) ] eq g i)) - (equiv-proof (isEquivPostComp e) (idfun _)) - where eq : ∀ (g : B → A) → ((fst e) ∘ g ≡ idfun _) ≡ (section (fst e) g) - eq g = sym (funExtPath {f = (fst e) ∘ g} {g = idfun _}) - -isContr-hasRetract : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) → isContr (hasRetract (fst e)) -isContr-hasRetract {_} {A} {B} e = transport (λ i → isContr (Σ[ g ∈ (B → A) ] eq g i)) - (equiv-proof (isEquivPreComp e) (idfun _)) - where eq : ∀ (g : B → A) → (g ∘ (fst e) ≡ idfun _) ≡ (retract (fst e) g) - eq g = sym (funExtPath {f = g ∘ (fst e)} {g = idfun _}) - +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels + +open import Cubical.Functions.FunExtEquiv + +private + variable + ℓ ℓ′ : Level + A B C : Type ℓ + +isEquivInvEquiv : isEquiv (λ (e : A ≃ B) → invEquiv e) +isEquivInvEquiv = isoToIsEquiv goal where + open Iso + goal : Iso (A ≃ B) (B ≃ A) + goal .fun = invEquiv + goal .inv = invEquiv + goal .rightInv g = equivEq refl + goal .leftInv f = equivEq refl + +invEquivEquiv : (A ≃ B) ≃ (B ≃ A) +invEquivEquiv = _ , isEquivInvEquiv + +isEquivCong : {x y : A} (e : A ≃ B) → isEquiv (λ (p : x ≡ y) → cong (equivFun e) p) +isEquivCong e = isoToIsEquiv (congIso (equivToIso e)) + +congEquiv : {x y : A} (e : A ≃ B) → (x ≡ y) ≃ (equivFun e x ≡ equivFun e y) +congEquiv e = isoToEquiv (congIso (equivToIso e)) + +equivAdjointEquiv : (e : A ≃ B) → ∀ {a b} → (a ≡ invEq e b) ≃ (equivFun e a ≡ b) +equivAdjointEquiv e = compEquiv (congEquiv e) (compPathrEquiv (secEq e _)) + +invEq≡→equivFun≡ : (e : A ≃ B) → ∀ {a b} → invEq e b ≡ a → equivFun e a ≡ b +invEq≡→equivFun≡ e = equivFun (equivAdjointEquiv e) ∘ sym + +isEquivPreComp : (e : A ≃ B) → isEquiv (λ (φ : B → C) → φ ∘ equivFun e) +isEquivPreComp e = snd (equiv→ (invEquiv e) (idEquiv _)) + +preCompEquiv : (e : A ≃ B) → (B → C) ≃ (A → C) +preCompEquiv e = (λ φ → φ ∘ fst e) , isEquivPreComp e + +isEquivPostComp : (e : A ≃ B) → isEquiv (λ (φ : C → A) → e .fst ∘ φ) +isEquivPostComp e = snd (equivΠCod (λ _ → e)) + +postCompEquiv : (e : A ≃ B) → (C → A) ≃ (C → B) +postCompEquiv e = _ , isEquivPostComp e + +-- see also: equivΠCod for a dependent version of postCompEquiv + +hasSection : (A → B) → Type _ +hasSection {A = A} {B = B} f = Σ[ g ∈ (B → A) ] section f g + +hasRetract : (A → B) → Type _ +hasRetract {A = A} {B = B} f = Σ[ g ∈ (B → A) ] retract f g + +isEquiv→isContrHasSection : {f : A → B} → isEquiv f → isContr (hasSection f) +fst (isEquiv→isContrHasSection isEq) = invIsEq isEq , secIsEq isEq +snd (isEquiv→isContrHasSection isEq) (f , ε) i = (λ b → fst (p b i)) , (λ b → snd (p b i)) + where p : ∀ b → (invIsEq isEq b , secIsEq isEq b) ≡ (f b , ε b) + p b = isEq .equiv-proof b .snd (f b , ε b) + +isEquiv→hasSection : {f : A → B} → isEquiv f → hasSection f +isEquiv→hasSection = fst ∘ isEquiv→isContrHasSection + +isContr-hasSection : (e : A ≃ B) → isContr (hasSection (fst e)) +isContr-hasSection e = isEquiv→isContrHasSection (snd e) + +isEquiv→isContrHasRetract : {f : A → B} → isEquiv f → isContr (hasRetract f) +fst (isEquiv→isContrHasRetract isEq) = invIsEq isEq , retIsEq isEq +snd (isEquiv→isContrHasRetract {f = f} isEq) (g , η) = + λ i → (λ b → p b i) , (λ a → q a i) + where p : ∀ b → invIsEq isEq b ≡ g b + p b = sym (η (invIsEq isEq b)) ∙' cong g (secIsEq isEq b) + -- one square from the definition of invIsEq + ieSq : ∀ a → Square (cong g (secIsEq isEq (f a))) + refl + (cong (g ∘ f) (retIsEq isEq a)) + refl + ieSq a k j = g (commSqIsEq isEq a k j) + -- one square from η + ηSq : ∀ a → Square (η (invIsEq isEq (f a))) + (η a) + (cong (g ∘ f) (retIsEq isEq a)) + (retIsEq isEq a) + ηSq a i j = η (retIsEq isEq a i) j + -- and one last square from the definition of p + pSq : ∀ b → Square (η (invIsEq isEq b)) + refl + (cong g (secIsEq isEq b)) + (p b) + pSq b i j = compPath'-filler (sym (η (invIsEq isEq b))) (cong g (secIsEq isEq b)) j i + q : ∀ a → Square (retIsEq isEq a) (η a) (p (f a)) refl + q a i j = hcomp (λ k → λ { (i = i0) → ηSq a j k + ; (i = i1) → η a (j ∧ k) + ; (j = i0) → pSq (f a) i k + ; (j = i1) → η a k + }) + (ieSq a j i) + +isEquiv→hasRetract : {f : A → B} → isEquiv f → hasRetract f +isEquiv→hasRetract = fst ∘ isEquiv→isContrHasRetract + +isContr-hasRetract : (e : A ≃ B) → isContr (hasRetract (fst e)) +isContr-hasRetract e = isEquiv→isContrHasRetract (snd e) + +cong≃ : (F : Type ℓ → Type ℓ′) → (A ≃ B) → F A ≃ F B +cong≃ F e = pathToEquiv (cong F (ua e)) + +cong≃-char : (F : Type ℓ → Type ℓ′) {A B : Type ℓ} (e : A ≃ B) → ua (cong≃ F e) ≡ cong F (ua e) +cong≃-char F e = ua-pathToEquiv (cong F (ua e)) + +cong≃-idEquiv : (F : Type ℓ → Type ℓ′) (A : Type ℓ) → cong≃ F (idEquiv A) ≡ idEquiv (F A) +cong≃-idEquiv F A = cong≃ F (idEquiv A) ≡⟨ cong (λ p → pathToEquiv (cong F p)) uaIdEquiv ⟩ + pathToEquiv refl ≡⟨ pathToEquivRefl ⟩ + idEquiv (F A) ∎ + +isPropIsHAEquiv : {f : A → B} → isProp (isHAEquiv f) +isPropIsHAEquiv {f = f} ishaef = goal ishaef where + equivF : isEquiv f + equivF = isHAEquiv→isEquiv ishaef + + rCoh1 : (sec : hasSection f) → Type _ + rCoh1 (g , ε) = Σ[ η ∈ retract f g ] ∀ x → cong f (η x) ≡ ε (f x) + + rCoh2 : (sec : hasSection f) → Type _ + rCoh2 (g , ε) = Σ[ η ∈ retract f g ] ∀ x → Square (ε (f x)) refl (cong f (η x)) refl + + rCoh3 : (sec : hasSection f) → Type _ + rCoh3 (g , ε) = ∀ x → Σ[ ηx ∈ g (f x) ≡ x ] Square (ε (f x)) refl (cong f ηx) refl + + rCoh4 : (sec : hasSection f) → Type _ + rCoh4 (g , ε) = ∀ x → Path (fiber f (f x)) (g (f x) , ε (f x)) (x , refl) + + characterization : isHAEquiv f ≃ Σ _ rCoh4 + characterization = + isHAEquiv f + -- first convert between Σ and record + ≃⟨ isoToEquiv (iso (λ e → (e .g , e .rinv) , (e .linv , e .com)) + (λ e → record { g = e .fst .fst ; rinv = e .fst .snd + ; linv = e .snd .fst ; com = e .snd .snd }) + (λ _ → refl) λ _ → refl) ⟩ + Σ _ rCoh1 + -- secondly, convert the path into a dependent path for later convenience + ≃⟨ Σ-cong-equiv-snd (λ s → Σ-cong-equiv-snd + λ η → equivΠCod + λ x → compEquiv (flipSquareEquiv {a₀₀ = f x}) (invEquiv slideSquareEquiv)) ⟩ + Σ _ rCoh2 + ≃⟨ Σ-cong-equiv-snd (λ s → invEquiv Σ-Π-≃) ⟩ + Σ _ rCoh3 + ≃⟨ Σ-cong-equiv-snd (λ s → equivΠCod λ x → ΣPath≃PathΣ) ⟩ + Σ _ rCoh4 + ■ + where open isHAEquiv + + goal : isProp (isHAEquiv f) + goal = subst isProp (sym (ua characterization)) + (isPropΣ (isContr→isProp (isEquiv→isContrHasSection equivF)) + λ s → isPropΠ λ x → isProp→isSet (isContr→isProp (equivF .equiv-proof (f x))) _ _) + +-- composition on the right induces an equivalence of path types +compr≡Equiv : {A : Type ℓ} {a b c : A} (p q : a ≡ b) (r : b ≡ c) → (p ≡ q) ≃ (p ∙ r ≡ q ∙ r) +compr≡Equiv p q r = congEquiv ((λ s → s ∙ r) , compPathr-isEquiv r) + +-- composition on the left induces an equivalence of path types +compl≡Equiv : {A : Type ℓ} {a b c : A} (p : a ≡ b) (q r : b ≡ c) → (q ≡ r) ≃ (p ∙ q ≡ p ∙ r) +compl≡Equiv p q r = congEquiv ((λ s → p ∙ s) , (compPathl-isEquiv p)) diff --git a/Cubical/Foundations/Everything.agda b/Cubical/Foundations/Everything.agda index f3713adb7..8810a1338 100644 --- a/Cubical/Foundations/Everything.agda +++ b/Cubical/Foundations/Everything.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Everything where -- Basic cubical prelude @@ -14,14 +14,14 @@ open import Cubical.Foundations.Id ; ap to apId ; equivFun to equivFunId ; equivCtr to equivCtrId - ; fiber to fiberId + ; fiber to fiberId ; funExt to funExtId ; isContr to isContrId ; isProp to isPropId ; isSet to isSetId ; isEquiv to isEquivId ; equivIsEquiv to equivIsEquivId - ; refl to reflId + ; refl to reflId ; ∥_∥ to propTruncId ; ∣_∣ to incId ; isPropIsContr to isPropIsContrId @@ -29,24 +29,23 @@ open import Cubical.Foundations.Id ) open import Cubical.Foundations.GroupoidLaws public -open import Cubical.Foundations.CartesianKanOps public open import Cubical.Foundations.Function public -open import Cubical.Foundations.Embedding public open import Cubical.Foundations.Equiv public open import Cubical.Foundations.Equiv.Properties public -open import Cubical.Foundations.PathSplitEquiv public -open import Cubical.Foundations.BiInvEquiv public -open import Cubical.Foundations.FunExtEquiv public +open import Cubical.Foundations.Equiv.Fiberwise +open import Cubical.Foundations.Equiv.PathSplit public +open import Cubical.Foundations.Equiv.BiInvertible public +open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.HLevels public open import Cubical.Foundations.Path public open import Cubical.Foundations.Pointed public +open import Cubical.Foundations.RelationalStructure public +open import Cubical.Foundations.Structure public open import Cubical.Foundations.Transport public open import Cubical.Foundations.Univalence public -open import Cubical.Foundations.UnivalenceId public +open import Cubical.Foundations.Univalence.Universe open import Cubical.Foundations.GroupoidLaws public open import Cubical.Foundations.Isomorphism public -open import Cubical.Foundations.Surjection public -open import Cubical.Foundations.TotalFiber public -open import Cubical.Foundations.Logic +open import Cubical.Foundations.CartesianKanOps +open import Cubical.Foundations.Powerset open import Cubical.Foundations.SIP -open import Cubical.Foundations.HoTT-UF diff --git a/Cubical/Foundations/FunExtEquiv.agda b/Cubical/Foundations/FunExtEquiv.agda deleted file mode 100644 index 218e7f2c2..000000000 --- a/Cubical/Foundations/FunExtEquiv.agda +++ /dev/null @@ -1,28 +0,0 @@ -{-# OPTIONS --cubical --safe #-} - -module Cubical.Foundations.FunExtEquiv where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Univalence - --- Function extensionality is an equivalence. -module _ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {f g : (x : A) → B x} where - private - fib : (p : f ≡ g) → fiber (funExt {B = B}) p - fib p = (funExt⁻ p , refl) - - funExt-fiber-isContr - : (p : f ≡ g) - → (fi : fiber (funExt {B = B}) p) - → fib p ≡ fi - funExt-fiber-isContr p (h , eq) i = (funExt⁻ (eq (~ i)) , λ j → eq (~ i ∨ j)) - - funExt-isEquiv : isEquiv (funExt {B = B}) - equiv-proof funExt-isEquiv p = (fib p , funExt-fiber-isContr p) - - funExtEquiv : (∀ x → f x ≡ g x) ≃ (f ≡ g) - funExtEquiv = (funExt {B = B} , funExt-isEquiv) - - funExtPath : (∀ x → f x ≡ g x) ≡ (f ≡ g) - funExtPath = ua funExtEquiv diff --git a/Cubical/Foundations/Function.agda b/Cubical/Foundations/Function.agda index 1f2cd2153..45a6a963d 100644 --- a/Cubical/Foundations/Function.agda +++ b/Cubical/Foundations/Function.agda @@ -1,49 +1,64 @@ {- Definitions for functions -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Function where open import Cubical.Foundations.Prelude +private + variable + ℓ ℓ' ℓ'' : Level + A : Type ℓ + B : A → Type ℓ + C : (a : A) → B a → Type ℓ + D : (a : A) (b : B a) → C a b → Type ℓ + -- The identity function -idfun : ∀ {ℓ} → (A : Type ℓ) → A → A +idfun : (A : Type ℓ) → A → A idfun _ x = x infixr 9 _∘_ -_∘_ : ∀ {ℓ ℓ′ ℓ″} {A : Type ℓ} {B : A → Type ℓ′} {C : (a : A) → B a → Type ℓ″} - (g : {a : A} → (b : B a) → C a b) → (f : (a : A) → B a) → (a : A) → C a (f a) +_∘_ : (g : {a : A} → (b : B a) → C a b) → (f : (a : A) → B a) → (a : A) → C a (f a) g ∘ f = λ x → g (f x) +{-# INLINE _∘_ #-} -∘-assoc : ∀ {ℓ ℓ′ ℓ″ ℓ‴} {A : Type ℓ} {B : A → Type ℓ′} {C : (a : A) → B a → Type ℓ″} {D : (a : A) (b : B a) → C a b → Type ℓ‴} - (h : {a : A} {b : B a} → (c : C a b) → D a b c) (g : {a : A} → (b : B a) → C a b) (f : (a : A) → B a) - → (h ∘ g) ∘ f ≡ h ∘ (g ∘ f) +∘-assoc : (h : {a : A} {b : B a} → (c : C a b) → D a b c) + (g : {a : A} → (b : B a) → C a b) + (f : (a : A) → B a) + → (h ∘ g) ∘ f ≡ h ∘ (g ∘ f) ∘-assoc h g f i x = h (g (f x)) -∘-idˡ : ∀ {ℓ ℓ′} {A : Type ℓ} {B : A → Type ℓ′} (f : (a : A) → B a) → f ∘ idfun A ≡ f +∘-idˡ : (f : (a : A) → B a) → f ∘ idfun A ≡ f ∘-idˡ f i x = f x -∘-idʳ : ∀ {ℓ ℓ′} {A : Type ℓ} {B : A → Type ℓ′} (f : (a : A) → B a) → (λ {a} → idfun (B a)) ∘ f ≡ f +∘-idʳ : (f : (a : A) → B a) → (λ {a} → idfun (B a)) ∘ f ≡ f ∘-idʳ f i x = f x +flip : {B : Type ℓ'} {C : A → B → Type ℓ''} → + ((x : A) (y : B) → C x y) → ((y : B) (x : A) → C x y) +flip f y x = f x y +{-# INLINE flip #-} -const : ∀ {ℓ ℓ′} {A : Type ℓ} {B : Type ℓ′} → A → B → A +const : {B : Type ℓ} → A → B → A const x = λ _ → x - +{-# INLINE const #-} case_of_ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} → (x : A) → (∀ x → B x) → B x case x of f = f x +{-# INLINE case_of_ #-} case_return_of_ : ∀ {ℓ ℓ'} {A : Type ℓ} (x : A) (B : A → Type ℓ') → (∀ x → B x) → B x case x return P of f = f x +{-# INLINE case_return_of_ #-} -uncurry - : ∀{ℓ ℓ′ ℓ″} {A : Type ℓ} {B : A → Type ℓ′} {C : (a : A) → B a → Type ℓ″} - → ((x : A) → (y : B x) → C x y) - → (p : Σ A B) → C (fst p) (snd p) +uncurry : ((x : A) → (y : B x) → C x y) → (p : Σ A B) → C (fst p) (snd p) uncurry f (x , y) = f x y +curry : ((p : Σ A B) → C (fst p) (snd p)) → (x : A) → (y : B x) → C x y +curry f x y = f (x , y) + module _ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} where -- Notions of 'coherently constant' functions for low dimensions. -- These are the properties of functions necessary to e.g. eliminate @@ -102,3 +117,21 @@ module _ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} where ; (j = i1) → link y x (i ∧ ~ k) }) (downleft x y i j) + +homotopyNatural : {B : Type ℓ'} {f g : A → B} (H : ∀ a → f a ≡ g a) {x y : A} (p : x ≡ y) → + H x ∙ cong g p ≡ cong f p ∙ H y +homotopyNatural {f = f} {g = g} H {x} {y} p i j = + hcomp (λ k → λ { (i = i0) → compPath-filler (H x) (cong g p) k j + ; (i = i1) → compPath-filler' (cong f p) (H y) k j + ; (j = i0) → cong f p (i ∧ ~ k) + ; (j = i1) → cong g p (i ∨ k) }) + (H (p i) j) + +homotopySymInv : {f : A → A} (H : ∀ a → f a ≡ a) (a : A) + → Path (f a ≡ f a) (λ i → H (H a (~ i)) i) refl +homotopySymInv {f = f} H a j i = + hcomp (λ k → λ { (i = i0) → f a + ; (i = i1) → H a (j ∧ ~ k) + ; (j = i0) → H (H a (~ i)) i + ; (j = i1) → H a (i ∧ ~ k)}) + (H (H a (~ i ∨ j)) i) diff --git a/Cubical/Foundations/GroupoidLaws.agda b/Cubical/Foundations/GroupoidLaws.agda index 660132dc3..8440e8e10 100644 --- a/Cubical/Foundations/GroupoidLaws.agda +++ b/Cubical/Foundations/GroupoidLaws.agda @@ -4,7 +4,7 @@ This file proves the higher groupoid structure of types for homogeneous and heterogeneous paths -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.GroupoidLaws where open import Cubical.Foundations.Prelude @@ -13,7 +13,7 @@ private variable ℓ : Level A : Type ℓ - x y z w : A + x y z w v : A _⁻¹ : (x ≡ y) → (y ≡ x) x≡y ⁻¹ = sym x≡y @@ -64,47 +64,27 @@ rCancel-filler {x = x} p k j i = rCancel : (p : x ≡ y) → p ∙ p ⁻¹ ≡ refl rCancel {x = x} p j i = rCancel-filler p i1 j i -lCancel : (p : x ≡ y) → p ⁻¹ ∙ p ≡ refl -lCancel p = rCancel (p ⁻¹) - --- The filler of the three-out-of-four identification: 3outof4-filler α p β = --- PathP (λ i → PathP (λ j → PathP (λ k → A) (α i i0) (α i i1)) --- (λ j → α i j) (λ j → β i j)) (λ j i → α i0 i) (3outof4 α p β) - -3outof4-filler : (α : I → I → A) → (p : α i1 i0 ≡ α i1 i1) → - (β : PathP (λ j → Path A (α j i0) (α j i1)) (λ i → α i0 i) p) → I → I → I → A -3outof4-filler α p β k j i = - hfill (λ k → λ { (i = i0) → α k i0 - ; (i = i1) → α k i1 - ; (j = i0) → α k i - ; (j = i1) → β k i - }) (inS (α i0 i)) k - -3outof4 : (α : I → I → A) → (p : α i1 i0 ≡ α i1 i1) → - (β : PathP (λ j → Path A (α j i0) (α j i1)) (λ i → α i0 i) p) → (λ i → α i1 i) ≡ p -3outof4 α p β j i = 3outof4-filler α p β i1 j i - --- The filler of the pre-associative square: preassoc p q r = --- PathP (λ i → PathP (λ j → PathP (λ k → A) x (compPath-filler q r i j)) --- (refl i) (λ j → compPath-filler (p ∙ q) r i j)) (λ j i → compPath-filler p q j i) (preassoc p q r) - -preassoc-filler : {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) → I → I → I → A -preassoc-filler {x = x} p q r k j i = - hfill (λ k → λ { (i = i0) → x - ; (i = i1) → compPath-filler q r k j - ; (j = i0) → p i - -- ; (j = i1) → compPath-filler (p ∙ q) r k i - }) (inS (compPath-filler p q j i)) k +rCancel-filler' : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → (i j k : I) → A +rCancel-filler' {x = x} {y} p i j k = + hfill + (λ i → λ + { (j = i1) → p (~ i ∧ k) + ; (k = i0) → x + ; (k = i1) → p (~ i) + }) + (inS (p k)) + (~ i) -preassoc : (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) → - PathP (λ j → Path A x ((q ∙ r) j)) p ((p ∙ q) ∙ r) -preassoc {x = x} p q r j i = preassoc-filler p q r i1 j i +rCancel' : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → p ∙ p ⁻¹ ≡ refl +rCancel' p j k = rCancel-filler' p i0 j k --- deducing associativity for compPath +lCancel : (p : x ≡ y) → p ⁻¹ ∙ p ≡ refl +lCancel p = rCancel (p ⁻¹) assoc : (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) → p ∙ q ∙ r ≡ (p ∙ q) ∙ r -assoc p q r = 3outof4 (compPath-filler p (q ∙ r)) ((p ∙ q) ∙ r) (preassoc p q r) +assoc p q r k = (compPath-filler p q k) ∙ compPath-filler' q r (~ k) + -- heterogeneous groupoid laws @@ -114,7 +94,7 @@ symInvoP p = refl rUnitP : {A : I → Type ℓ} → {x : A i0} → {y : A i1} → (p : PathP A x y) → PathP (λ j → PathP (λ i → rUnit (λ i → A i) j i) x y) p (compPathP p refl) -rUnitP p j i = compPathP-filler p refl i j +rUnitP p j i = compPathP-filler p refl j i lUnitP : {A : I → Type ℓ} → {x : A i0} → {y : A i1} → (p : PathP A x y) → PathP (λ j → PathP (λ i → lUnit (λ i → A i) j i) x y) p (compPathP refl p) @@ -138,88 +118,68 @@ lCancelP : {A : I → Type ℓ} → {x : A i0} → {y : A i1} → (p : PathP A x PathP (λ j → PathP (λ i → lCancel (λ i → A i) j i) y y) (compPathP (symP p) p) refl lCancelP p = rCancelP (symP p) -3outof4P : {A : I → I → Type ℓ} {P : (A i0 i1) ≡ (A i1 i1)} - {B : PathP (λ j → Path (Type ℓ) (A i0 j) (A i1 j)) (λ i → A i i0) P} - (α : ∀ (i j : I) → A j i) - (p : PathP (λ i → P i) (α i1 i0) (α i1 i1)) → - (β : PathP (λ j → PathP (λ i → B j i) (α j i0) (α j i1)) (λ i → α i0 i) p) → - PathP (λ j → PathP (λ i → 3outof4 (λ j i → A i j) P B j i) (α i1 i0) (α i1 i1)) (λ i → α i1 i) p -3outof4P {A = A} {P} {B} α p β j i = - comp (λ k → 3outof4-filler (λ j i → A i j) P B k j i) - (λ k → λ { (i = i0) → α k i0 - ; (i = i1) → α k i1 - ; (j = i0) → α k i - ; (j = i1) → β k i - }) (α i0 i) - -preassocP : {A : I → Type ℓ} {x : A i0} {y : A i1} {B_i1 : Type ℓ} {B : (A i1) ≡ B_i1} {z : B i1} - {C_i1 : Type ℓ} {C : (B i1) ≡ C_i1} {w : C i1} (p : PathP A x y) (q : PathP (λ i → B i) y z) (r : PathP (λ i → C i) z w) → - PathP (λ j → PathP (λ i → preassoc (λ i → A i) B C j i) x ((compPathP q r) j)) p (compPathP (compPathP p q) r) -preassocP {A = A} {x = x} {B = B} {C = C} p q r j i = - comp (λ k → preassoc-filler (λ i → A i) B C k j i) - (λ k → λ { (i = i0) → x - ; (i = i1) → compPathP-filler q r j k - ; (j = i0) → p i - -- ; (j = i1) → compPathP-filler (compPathP p q) r i k - }) (compPathP-filler p q i j) + assocP : {A : I → Type ℓ} {x : A i0} {y : A i1} {B_i1 : Type ℓ} {B : (A i1) ≡ B_i1} {z : B i1} {C_i1 : Type ℓ} {C : (B i1) ≡ C_i1} {w : C i1} (p : PathP A x y) (q : PathP (λ i → B i) y z) (r : PathP (λ i → C i) z w) → PathP (λ j → PathP (λ i → assoc (λ i → A i) B C j i) x w) (compPathP p (compPathP q r)) (compPathP (compPathP p q) r) -assocP p q r = - 3outof4P (λ i j → compPathP-filler p (compPathP q r) j i) (compPathP (compPathP p q) r) (preassocP p q r) +assocP {A = A} {B = B} {C = C} p q r k i = + comp (\ j' → hfill (λ j → λ { + (i = i0) → A i0 + ; (i = i1) → compPath-filler' (λ i₁ → B i₁) (λ i₁ → C i₁) (~ k) j }) + (inS (compPath-filler (λ i₁ → A i₁) (λ i₁ → B i₁) k i)) j') + (λ j → λ + { (i = i0) → p i0 + ; (i = i1) → + comp (\ j' → hfill ((λ l → λ + { (j = i0) → B k + ; (j = i1) → C l + ; (k = i1) → C (j ∧ l) + })) (inS (B ( j ∨ k)) ) j') + (λ l → λ + { (j = i0) → q k + ; (j = i1) → r l + ; (k = i1) → r (j ∧ l) + }) + (q (j ∨ k)) + }) + (compPathP-filler p q k i) -- Loic's code below - --- simultaneaous composition on both sides of a path - -doubleCompPath-filler : {ℓ : Level} {A : Type ℓ} {w x y z : A} → w ≡ x → x ≡ y → y ≡ z → - I → I → A -doubleCompPath-filler p q r i = - hfill (λ t → λ { (i = i0) → p (~ t) - ; (i = i1) → r t }) - (inS (q i)) - -doubleCompPath : {ℓ : Level} {A : Type ℓ} {w x y z : A} → w ≡ x → x ≡ y → y ≡ z → w ≡ z -doubleCompPath p q r i = doubleCompPath-filler p q r i i1 - -_∙∙_∙∙_ : {ℓ : Level} {A : Type ℓ} {w x y z : A} → w ≡ x → x ≡ y → y ≡ z → w ≡ z -p ∙∙ q ∙∙ r = doubleCompPath p q r - -- some exchange law for doubleCompPath and refl -rhombus-filler : {ℓ : Level} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → I → I → A -rhombus-filler p q i j = - hcomp (λ t → λ { (i = i0) → p (~ t ∨ j) - ; (i = i1) → q (t ∧ j) - ; (j = i0) → p (~ t ∨ i) - ; (j = i1) → q (t ∧ i) }) - (p i1) +invSides-filler : {x y z : A} (p : x ≡ y) (q : x ≡ z) → Square p (sym q) q (sym p) +invSides-filler {x = x} p q i j = + hcomp (λ k → λ { (i = i0) → p (k ∧ j) + ; (i = i1) → q (~ j ∧ k) + ; (j = i0) → q (i ∧ k) + ; (j = i1) → p (~ i ∧ k)}) + x leftright : {ℓ : Level} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → (refl ∙∙ p ∙∙ q) ≡ (p ∙∙ q ∙∙ refl) leftright p q i j = hcomp (λ t → λ { (j = i0) → p (i ∧ (~ t)) ; (j = i1) → q (t ∨ i) }) - (rhombus-filler p q i j) + (invSides-filler q (sym p) (~ i) j) -- equating doubleCompPath and a succession of two compPath split-leftright : {ℓ : Level} {A : Type ℓ} {w x y z : A} (p : w ≡ x) (q : x ≡ y) (r : y ≡ z) → (p ∙∙ q ∙∙ r) ≡ (refl ∙∙ (p ∙∙ q ∙∙ refl) ∙∙ r) -split-leftright p q r i j = - hcomp (λ t → λ { (j = i0) → p (~ i ∧ ~ t) - ; (j = i1) → r t }) +split-leftright p q r j i = + hcomp (λ t → λ { (i = i0) → p (~ j ∧ ~ t) + ; (i = i1) → r t }) (doubleCompPath-filler p q refl j i) split-leftright' : {ℓ : Level} {A : Type ℓ} {w x y z : A} (p : w ≡ x) (q : x ≡ y) (r : y ≡ z) → (p ∙∙ q ∙∙ r) ≡ (p ∙∙ (refl ∙∙ q ∙∙ r) ∙∙ refl) -split-leftright' p q r i j = - hcomp (λ t → λ { (j = i0) → p (~ t) - ; (j = i1) → r (i ∨ t) }) +split-leftright' p q r j i = + hcomp (λ t → λ { (i = i0) → p (~ t) + ; (i = i1) → r (j ∨ t) }) (doubleCompPath-filler refl q r j i) doubleCompPath-elim : {ℓ : Level} {A : Type ℓ} {w x y z : A} (p : w ≡ x) (q : x ≡ y) @@ -230,22 +190,18 @@ doubleCompPath-elim' : {ℓ : Level} {A : Type ℓ} {w x y z : A} (p : w ≡ x) (r : y ≡ z) → (p ∙∙ q ∙∙ r) ≡ p ∙ (q ∙ r) doubleCompPath-elim' p q r = (split-leftright' p q r) ∙ (sym (leftright p (q ∙ r))) - -cong-∙ : ∀ {B : Type ℓ} (f : A → B) (p : x ≡ y) (q : y ≡ z) - → cong f (p ∙ q) ≡ (cong f p) ∙ (cong f q) -cong-∙ f p q j i = hcomp (λ k → λ { (j = i0) → f (compPath-filler p q k i) - ; (i = i0) → f (p i0) - ; (i = i1) → f (q k) }) - (f (p i)) - cong-∙∙ : ∀ {B : Type ℓ} (f : A → B) (p : w ≡ x) (q : x ≡ y) (r : y ≡ z) → cong f (p ∙∙ q ∙∙ r) ≡ (cong f p) ∙∙ (cong f q) ∙∙ (cong f r) -cong-∙∙ f p q r j i = hcomp (λ k → λ { (j = i0) → f (doubleCompPath-filler p q r i k) +cong-∙∙ f p q r j i = hcomp (λ k → λ { (j = i0) → f (doubleCompPath-filler p q r k i) ; (i = i0) → f (p (~ k)) ; (i = i1) → f (r k) }) (f (q i)) -hcomp-unique : ∀ {ℓ} {A : Set ℓ} {φ} → (u : I → Partial φ A) → (u0 : A [ φ ↦ u i0 ]) → +cong-∙ : ∀ {B : Type ℓ} (f : A → B) (p : x ≡ y) (q : y ≡ z) + → cong f (p ∙ q) ≡ (cong f p) ∙ (cong f q) +cong-∙ f p q = cong-∙∙ f refl p q + +hcomp-unique : ∀ {ℓ} {A : Type ℓ} {φ} → (u : I → Partial φ A) → (u0 : A [ φ ↦ u i0 ]) → (h2 : ∀ i → A [ (φ ∨ ~ i) ↦ (\ { (φ = i1) → u i 1=1; (i = i0) → outS u0}) ]) → (hcomp u (outS u0) ≡ outS (h2 i1)) [ φ ↦ (\ { (φ = i1) → (\ i → u i1 1=1)}) ] hcomp-unique {φ = φ} u u0 h2 = inS (\ i → hcomp (\ k → \ { (φ = i1) → u k 1=1 @@ -253,7 +209,7 @@ hcomp-unique {φ = φ} u u0 h2 = inS (\ i → hcomp (\ k → \ { (φ = i1) → u (outS u0)) -lid-unique : ∀ {ℓ} {A : Set ℓ} {φ} → (u : I → Partial φ A) → (u0 : A [ φ ↦ u i0 ]) → +lid-unique : ∀ {ℓ} {A : Type ℓ} {φ} → (u : I → Partial φ A) → (u0 : A [ φ ↦ u i0 ]) → (h1 h2 : ∀ i → A [ (φ ∨ ~ i) ↦ (\ { (φ = i1) → u i 1=1; (i = i0) → outS u0}) ]) → (outS (h1 i1) ≡ outS (h2 i1)) [ φ ↦ (\ { (φ = i1) → (\ i → u i1 1=1)}) ] lid-unique {φ = φ} u u0 h1 h2 = inS (\ i → hcomp (\ k → \ { (φ = i1) → u k 1=1 @@ -262,21 +218,257 @@ lid-unique {φ = φ} u u0 h1 h2 = inS (\ i → hcomp (\ k → \ { (φ = i1) → (outS u0)) -transp-hcomp : ∀ {ℓ} (φ : I) {A' : Set ℓ} - (A : (i : I) → Set ℓ [ φ ↦ (λ _ → A') ]) (let B = \ (i : I) → outS (A i)) +transp-hcomp : ∀ {ℓ} (φ : I) {A' : Type ℓ} + (A : (i : I) → Type ℓ [ φ ↦ (λ _ → A') ]) (let B = \ (i : I) → outS (A i)) → ∀ {ψ} (u : I → Partial ψ (B i0)) → (u0 : B i0 [ ψ ↦ u i0 ]) → - (transp B φ (hcomp u (outS u0)) ≡ hcomp (\ i o → transp B φ (u i o)) (transp B φ (outS u0))) - [ ψ ↦ (\ { (ψ = i1) → (\ i → transp B φ (u i1 1=1))}) ] + (transp (\ i → B i) φ (hcomp u (outS u0)) ≡ hcomp (\ i o → transp (\ i → B i) φ (u i o)) (transp (\ i → B i) φ (outS u0))) + [ ψ ↦ (\ { (ψ = i1) → (\ i → transp (\ i → B i) φ (u i1 1=1))}) ] transp-hcomp φ A u u0 = inS (sym (outS (hcomp-unique - ((\ i o → transp B φ (u i o))) (inS (transp B φ (outS u0))) - \ i → inS (transp B φ (hfill u u0 i))))) + ((\ i o → transp (\ i → B i) φ (u i o))) (inS (transp (\ i → B i) φ (outS u0))) + \ i → inS (transp (\ i → B i) φ (hfill u u0 i))))) where B = \ (i : I) → outS (A i) -hcomp-cong : ∀ {ℓ} {A : Set ℓ} {φ} → (u : I → Partial φ A) → (u0 : A [ φ ↦ u i0 ]) → - (u' : I → Partial φ A) → (u0' : A [ φ ↦ u' i0 ]) → - +hcomp-cong : ∀ {ℓ} {A : Type ℓ} {φ} → (u : I → Partial φ A) → (u0 : A [ φ ↦ u i0 ]) → + (u' : I → Partial φ A) → (u0' : A [ φ ↦ u' i0 ]) → (ueq : ∀ i → PartialP φ (\ o → u i o ≡ u' i o)) → (outS u0 ≡ outS u0') [ φ ↦ (\ { (φ = i1) → ueq i0 1=1}) ] → (hcomp u (outS u0) ≡ hcomp u' (outS u0')) [ φ ↦ (\ { (φ = i1) → ueq i1 1=1 }) ] hcomp-cong u u0 u' u0' ueq 0eq = inS (\ j → hcomp (\ i o → ueq i o j) (outS 0eq j)) + + +congFunct-filler : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y z : A} (f : A → B) (p : x ≡ y) (q : y ≡ z) + → I → I → I → B +congFunct-filler {x = x} f p q i j z = + hfill (λ k → λ { (i = i0) → f x + ; (i = i1) → f (q k) + ; (j = i0) → f (compPath-filler p q k i)}) + (inS (f (p i))) + z + +congFunct : ∀ {ℓ} {B : Type ℓ} (f : A → B) (p : x ≡ y) (q : y ≡ z) → cong f (p ∙ q) ≡ cong f p ∙ cong f q +congFunct f p q j i = congFunct-filler f p q i j i1 + + +-- congFunct for dependent types +congFunct-dep : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {x y z : A} (f : (a : A) → B a) (p : x ≡ y) (q : y ≡ z) + → PathP (λ i → PathP (λ j → B (compPath-filler p q i j)) (f x) (f (q i))) (cong f p) (cong f (p ∙ q)) +congFunct-dep {B = B} {x = x} f p q i j = f (compPath-filler p q i j) + +cong₂Funct : ∀ {ℓ ℓ'} {A : Type ℓ} {x y : A} {B : Type ℓ'} (f : A → A → B) → + (p : x ≡ y) → + {u v : A} (q : u ≡ v) → + cong₂ f p q ≡ cong (λ x → f x u) p ∙ cong (f y) q +cong₂Funct {x = x} {y = y} f p {u = u} {v = v} q j i = + hcomp (λ k → λ { (i = i0) → f x u + ; (i = i1) → f y (q k) + ; (j = i0) → f (p i) (q (i ∧ k))}) + (f (p i) u) + +symDistr-filler : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → I → I → I → A +symDistr-filler {A = A} {z = z} p q i j k = + hfill (λ k → λ { (i = i0) → q (k ∨ j) + ; (i = i1) → p (~ k ∧ j) }) + (inS (invSides-filler q (sym p) i j)) + k + +symDistr : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → sym (p ∙ q) ≡ sym q ∙ sym p +symDistr p q i j = symDistr-filler p q j i i1 + +-- we can not write hcomp-isEquiv : {ϕ : I} → (p : I → Partial ϕ A) → isEquiv (λ (a : A [ ϕ ↦ p i0 ]) → hcomp p a) +-- due to size issues. But what we can write (compare to hfill) is: +hcomp-equivFillerSub : {ϕ : I} → (p : I → Partial ϕ A) → (a : A [ ϕ ↦ p i0 ]) + → (i : I) + → A [ ϕ ∨ i ∨ ~ i ↦ (λ { (i = i0) → outS a + ; (i = i1) → hcomp (λ i → p (~ i)) (hcomp p (outS a)) + ; (ϕ = i1) → p i0 1=1 }) ] +hcomp-equivFillerSub {ϕ = ϕ} p a i = + inS (hcomp (λ k → λ { (i = i1) → hfill (λ j → p (~ j)) (inS (hcomp p (outS a))) k + ; (i = i0) → outS a + ; (ϕ = i1) → p (~ k ∧ i) 1=1 }) + (hfill p a i)) + +hcomp-equivFiller : {ϕ : I} → (p : I → Partial ϕ A) → (a : A [ ϕ ↦ p i0 ]) + → (i : I) → A +hcomp-equivFiller p a i = outS (hcomp-equivFillerSub p a i) + + +pentagonIdentity : (p : x ≡ y) → (q : y ≡ z) → (r : z ≡ w) → (s : w ≡ v) + → + (assoc p q (r ∙ s) ∙ assoc (p ∙ q) r s) + ≡ + cong (p ∙_) (assoc q r s) ∙∙ assoc p (q ∙ r) s ∙∙ cong (_∙ s) (assoc p q r) + +pentagonIdentity {x = x} {y} p q r s = + (λ i → + (λ j → cong (p ∙_) (assoc q r s) (i ∧ j)) + ∙∙ (λ j → lemma₀₀ i j ∙ lemma₀₁ i j) + ∙∙ (λ j → lemma₁₀ i j ∙ lemma₁₁ i j) + ) + where + + + lemma₀₀ : ( i j : I) → _ ≡ _ + lemma₀₀ i j i₁ = + hcomp + (λ k → λ { (j = i0) → p i₁ + ; (i₁ = i0) → x + ; (i₁ = i1) → hcomp + (λ k₁ → λ { (i = i0) → (q (j ∧ k)) + ; (k = i0) → y + ; (j = i0) → y + ; (j = i1)(k = i1) → r (k₁ ∧ i)}) + (q (j ∧ k)) + }) (p i₁) + + lemma₀₁ : ( i j : I) → hcomp + (λ k → λ {(i = i0) → q j + ; (j = i0) → y + ; (j = i1) → r (k ∧ i) + }) + (q j) ≡ _ + lemma₀₁ i j i₁ = (hcomp + (λ k → λ { (j = i1) → hcomp + (λ k₁ → λ { (i₁ = i0) → r i + ; (k = i0) → r i + ; (i = i1) → s (k₁ ∧ k ∧ i₁) + ; (i₁ = i1)(k = i1) → s k₁ }) + (r ((i₁ ∧ k) ∨ i)) + ; (i₁ = i0) → compPath-filler q r i j + ; (i₁ = i1) → hcomp + (λ k₁ → λ { (k = i0) → r i + ; (k = i1) → s k₁ + ; (i = i1) → s (k ∧ k₁)}) + (r (i ∨ k))}) + (hfill + (λ k → λ { (j = i1) → r k + ; (i₁ = i1) → r k + ; (i₁ = i0)(j = i0) → y }) + (inS (q (i₁ ∨ j))) i)) + + lemma₁₁ : ( i j : I) → (r (i ∨ j)) ≡ _ + lemma₁₁ i j i₁ = + hcomp + (λ k → λ { (i = i1) → s (i₁ ∧ k) + ; (j = i1) → s (i₁ ∧ k) + ; (i₁ = i0) → r (i ∨ j) + ; (i₁ = i1) → s k + }) (r (i ∨ j ∨ i₁)) + + + lemma₁₀-back : I → I → I → _ + lemma₁₀-back i j i₁ = + hcomp + (λ k → λ { + (i₁ = i0) → x + ; (i₁ = i1) → hcomp + (λ k₁ → λ { (k = i0) → q (j ∨ ~ i) + ; (k = i1) → r (k₁ ∧ j) + ; (j = i0) → q (k ∨ ~ i) + ; (j = i1) → r (k₁ ∧ k) + ; (i = i0) → r (k ∧ j ∧ k₁) + }) + (q (k ∨ j ∨ ~ i)) + ; (i = i0)(j = i0) → (p ∙ q) i₁ + }) + (hcomp + (λ k → λ { (i₁ = i0) → x + ; (i₁ = i1) → q ((j ∨ ~ i ) ∧ k) + ; (j = i0)(i = i1) → p i₁ + }) + (p i₁)) + + + lemma₁₀-front : I → I → I → _ + lemma₁₀-front i j i₁ = + (((λ _ → x) ∙∙ compPath-filler p q j ∙∙ + (λ i₁ → + hcomp + (λ k → λ { (i₁ = i0) → q j + ; (i₁ = i1) → r (k ∧ (j ∨ i)) + ; (j = i0)(i = i0) → q i₁ + ; (j = i1) → r (i₁ ∧ k) + }) + (q (j ∨ i₁)) + )) i₁) + + compPath-filler-in-filler : + (p : _ ≡ y) → (q : _ ≡ _ ) + → _≡_ {A = Square (p ∙ q) (p ∙ q) (λ _ → x) (λ _ → z)} + (λ i j → hcomp + (λ i₂ → + λ { (j = i0) → x + ; (j = i1) → q (i₂ ∨ ~ i) + ; (i = i0) → (p ∙ q) j + }) + (compPath-filler p q (~ i) j)) + (λ _ → p ∙ q) + compPath-filler-in-filler p q z i j = + hcomp + (λ k → λ { + (j = i0) → p i0 + ; (j = i1) → q (k ∨ ~ i ∧ ~ z) + ; (i = i0) → hcomp + (λ i₂ → λ { + (j = i0) → p i0 + ;(j = i1) → q ((k ∨ ~ z) ∧ i₂) + ;(z = i1) (k = i0) → p j + }) + (p j) + ; (i = i1) → compPath-filler p (λ i₁ → q (k ∧ i₁)) k j + ; (z = i0) → hfill + ((λ i₂ → λ { (j = i0) → p i0 + ; (j = i1) → q (i₂ ∨ ~ i) + ; (i = i0) → (p ∙ q) j + })) + (inS ((compPath-filler p q (~ i) j))) k + ; (z = i1) → compPath-filler p q k j + }) + (compPath-filler p q (~ i ∧ ~ z) j) + + + cube-comp₋₀₋ : + (c : I → I → I → A) + → {a' : Square _ _ _ _} + → (λ i i₁ → c i i0 i₁) ≡ a' + → (I → I → I → A) + cube-comp₋₀₋ c p i j k = + hcomp + (λ l → λ { + (i = i0) → c i0 j k + ;(i = i1) → c i1 j k + ;(j = i0) → p l i k + ;(j = i1) → c i i1 k + ;(k = i0) → c i j i0 + ;(k = i1) → c i j i1 + }) + (c i j k) + + cube-comp₀₋₋ : + (c : I → I → I → A) + → {a' : Square _ _ _ _} + → (λ i i₁ → c i0 i i₁) ≡ a' + → (I → I → I → A) + cube-comp₀₋₋ c p i j k = + hcomp + (λ l → λ { + (i = i0) → p l j k + ;(i = i1) → c i1 j k + ;(j = i0) → c i i0 k + ;(j = i1) → c i i1 k + ;(k = i0) → c i j i0 + ;(k = i1) → c i j i1 + }) + (c i j k) + + + + lemma₁₀-back' : _ + lemma₁₀-back' k j i₁ = + (cube-comp₋₀₋ (lemma₁₀-back) + (compPath-filler-in-filler p q)) k j i₁ + + + lemma₁₀ : ( i j : I) → _ ≡ _ + lemma₁₀ i j i₁ = + (cube-comp₀₋₋ lemma₁₀-front (sym lemma₁₀-back')) i j i₁ diff --git a/Cubical/Foundations/HAEquiv.agda b/Cubical/Foundations/HAEquiv.agda deleted file mode 100644 index 89a56de68..000000000 --- a/Cubical/Foundations/HAEquiv.agda +++ /dev/null @@ -1,147 +0,0 @@ -{- - -Half adjoint equivalences ([HAEquiv]) - -- Iso to HAEquiv ([iso→HAEquiv]) -- Equiv to HAEquiv ([equiv→HAEquiv]) -- Cong is an equivalence ([congEquiv]) - --} -{-# OPTIONS --cubical --safe #-} -module Cubical.Foundations.HAEquiv where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.GroupoidLaws - -open import Cubical.Data.Nat.Base - -record isHAEquiv {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) : Type (ℓ-max ℓ ℓ') where - field - g : B → A - sec : ∀ a → g (f a) ≡ a - ret : ∀ b → f (g b) ≡ b - com : ∀ a → cong f (sec a) ≡ ret (f a) - - -- from redtt's ha-equiv/symm - com-op : ∀ b → cong g (ret b) ≡ sec (g b) - com-op b j i = hcomp (λ k → λ { (i = i0) → sec (g b) (j ∧ (~ k)) - ; (j = i0) → g (ret b i) - ; (j = i1) → sec (g b) (i ∨ (~ k)) - ; (i = i1) → g b }) - (cap1 j i) - - where cap0 : Square {- (j = i0) -} (λ i → f (g (ret b i))) - {- (j = i1) -} (λ i → ret b i) - {- (i = i0) -} (λ j → f (sec (g b) j)) - {- (i = i1) -} (λ j → ret b j) - - cap0 j i = hcomp (λ k → λ { (i = i0) → com (g b) (~ k) j - ; (j = i0) → f (g (ret b i)) - ; (j = i1) → ret b i - ; (i = i1) → ret b j }) - (ret (ret b i) j) - - filler : I → I → A - filler j i = hfill (λ k → λ { (i = i0) → g (ret b k) - ; (i = i1) → g b }) - (inS (sec (g b) i)) j - - cap1 : Square {- (j = i0) -} (λ i → g (ret b i)) - {- (j = i1) -} (λ i → g b) - {- (i = i0) -} (λ j → sec (g b) j) - {- (i = i1) -} (λ j → g b) - - cap1 j i = hcomp (λ k → λ { (i = i0) → sec (sec (g b) j) k - ; (j = i0) → sec (g (ret b i)) k - ; (j = i1) → filler i k - ; (i = i1) → filler j k }) - (g (cap0 j i)) - - -HAEquiv : ∀ {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') → Type (ℓ-max ℓ ℓ') -HAEquiv A B = Σ (A → B) λ f → isHAEquiv f - -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - -iso→HAEquiv : Iso A B → HAEquiv A B -iso→HAEquiv {A = A} {B = B} (iso f g ε η) = f , (record { g = g ; sec = η ; ret = ret ; com = com }) - where - sides : ∀ b i j → Partial (~ i ∨ i) B - sides b i j = λ { (i = i0) → ε (f (g b)) j - ; (i = i1) → ε b j } - - bot : ∀ b i → B - bot b i = cong f (η (g b)) i - - ret : (b : B) → f (g b) ≡ b - ret b i = hcomp (sides b i) (bot b i) - - com : (a : A) → cong f (η a) ≡ ret (f a) - com a i j = hcomp (λ k → λ { (i = i0) → ε (f (η a j)) k - ; (i = i1) → hfill (sides (f a) j) (inS (bot (f a) j)) k - ; (j = i0) → ε (f (g (f a))) k - ; (j = i1) → ε (f a) k}) - (cong (cong f) (sym (Hfa≡fHa (λ x → g (f x)) η a)) i j) - -equiv→HAEquiv : A ≃ B → HAEquiv A B -equiv→HAEquiv e = iso→HAEquiv (equivToIso e) - -congEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A} (e : A ≃ B) → (x ≡ y) ≃ (e .fst x ≡ e .fst y) -congEquiv {A = A} {B} {x} {y} e = isoToEquiv (iso intro elim intro-elim elim-intro) - where - e' : HAEquiv A B - e' = equiv→HAEquiv e - - f : A → B - f = e' .fst - - g : B → A - g = isHAEquiv.g (e' .snd) - - sec : ∀ a → g (f a) ≡ a - sec = isHAEquiv.sec (e' .snd) - - ret : ∀ b → f (g b) ≡ b - ret = isHAEquiv.ret (e' .snd) - - com : ∀ a → cong f (sec a) ≡ ret (f a) - com = isHAEquiv.com (e' .snd) - - intro : x ≡ y → f x ≡ f y - intro = cong f - - elim-sides : ∀ p i j → Partial (~ i ∨ i) A - elim-sides p i j = λ { (i = i0) → sec x j - ; (i = i1) → sec y j } - - elim-bot : ∀ p i → A - elim-bot p i = cong g p i - - elim : f x ≡ f y → x ≡ y - elim p i = hcomp (elim-sides p i) (elim-bot p i) - - intro-elim : ∀ p → intro (elim p) ≡ p - intro-elim p i j = - hcomp (λ k → λ { (i = i0) → f (hfill (elim-sides p j) - (inS (elim-bot p j)) k) - ; (i = i1) → ret (p j) k - ; (j = i0) → com x i k - ; (j = i1) → com y i k }) - (f (g (p j))) - - elim-intro : ∀ p → elim (intro p) ≡ p - elim-intro p i j = - hcomp (λ k → λ { (i = i0) → hfill (λ l → λ { (j = i0) → secEq e x l - ; (j = i1) → secEq e y l }) - (inS (cong (λ z → g (f z)) p j)) k - ; (i = i1) → p j - ; (j = i0) → secEq e x (i ∨ k) - ; (j = i1) → secEq e y (i ∨ k) }) - (secEq e (p j) i) diff --git a/Cubical/Foundations/HLevels.agda b/Cubical/Foundations/HLevels.agda index 0ce434b90..9a128fbfd 100644 --- a/Cubical/Foundations/HLevels.agda +++ b/Cubical/Foundations/HLevels.agda @@ -7,63 +7,93 @@ Basic theory about h-levels/n-types: - Hedberg's theorem can be found in Cubical/Relation/Nullary/DecidableEq -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.HLevels where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function -open import Cubical.Foundations.FunExtEquiv +open import Cubical.Foundations.Structure +open import Cubical.Functions.FunExtEquiv open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Path open import Cubical.Foundations.Transport -open import Cubical.Foundations.HAEquiv using (congEquiv) -open import Cubical.Foundations.Univalence using (ua; univalence) +open import Cubical.Foundations.Univalence using (ua ; univalenceIso) -open import Cubical.Data.Sigma using (ΣPathP; sigmaPath→pathSigma; pathSigma≡sigmaPath; _Σ≡T_) +open import Cubical.Data.Sigma open import Cubical.Data.Nat using (ℕ; zero; suc; _+_; +-zero; +-comm) +HLevel : Type₀ +HLevel = ℕ + private variable - ℓ ℓ' : Level + ℓ ℓ' ℓ'' ℓ''' ℓ'''' ℓ''''' : Level A : Type ℓ B : A → Type ℓ - x y : A - n : ℕ + C : (x : A) → B x → Type ℓ + D : (x : A) (y : B x) → C x y → Type ℓ + E : (x : A) (y : B x) → (z : C x y) → D x y z → Type ℓ + w x y z : A + n : HLevel -isOfHLevel : ℕ → Type ℓ → Type ℓ +isOfHLevel : HLevel → Type ℓ → Type ℓ isOfHLevel 0 A = isContr A isOfHLevel 1 A = isProp A isOfHLevel (suc (suc n)) A = (x y : A) → isOfHLevel (suc n) (x ≡ y) -HLevel : ∀ ℓ → ℕ → Type (ℓ-suc ℓ) -HLevel ℓ n = Σ[ A ∈ Type ℓ ] (isOfHLevel n A) +isOfHLevelFun : (n : HLevel) {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ-max ℓ ℓ') +isOfHLevelFun n f = ∀ b → isOfHLevel n (fiber f b) + +isOfHLevelΩ→isOfHLevel : + ∀ {ℓ} {A : Type ℓ} (n : ℕ) + → ((x : A) → isOfHLevel (suc n) (x ≡ x)) → isOfHLevel (2 + n) A +isOfHLevelΩ→isOfHLevel zero hΩ x y = + J (λ y p → (q : x ≡ y) → p ≡ q) (hΩ x refl) +isOfHLevelΩ→isOfHLevel (suc n) hΩ x y = + J (λ y p → (q : x ≡ y) → isOfHLevel (suc n) (p ≡ q)) (hΩ x refl) + +TypeOfHLevel : ∀ ℓ → HLevel → Type (ℓ-suc ℓ) +TypeOfHLevel ℓ n = TypeWithStr ℓ (isOfHLevel n) hProp hSet hGroupoid h2Groupoid : ∀ ℓ → Type (ℓ-suc ℓ) -hProp ℓ = HLevel ℓ 1 -hSet ℓ = HLevel ℓ 2 -hGroupoid ℓ = HLevel ℓ 3 -h2Groupoid ℓ = HLevel ℓ 4 +hProp ℓ = TypeOfHLevel ℓ 1 +hSet ℓ = TypeOfHLevel ℓ 2 +hGroupoid ℓ = TypeOfHLevel ℓ 3 +h2Groupoid ℓ = TypeOfHLevel ℓ 4 -- lower h-levels imply higher h-levels -isOfHLevelSuc : (n : ℕ) → isOfHLevel n A → isOfHLevel (suc n) A +isOfHLevelSuc : (n : HLevel) → isOfHLevel n A → isOfHLevel (suc n) A isOfHLevelSuc 0 = isContr→isProp isOfHLevelSuc 1 = isProp→isSet isOfHLevelSuc (suc (suc n)) h a b = isOfHLevelSuc (suc n) (h a b) -isOfHLevelPlus : (m : ℕ) → isOfHLevel n A → isOfHLevel (m + n) A +isSet→isGroupoid : isSet A → isGroupoid A +isSet→isGroupoid = isOfHLevelSuc 2 + +isGroupoid→is2Groupoid : isGroupoid A → is2Groupoid A +isGroupoid→is2Groupoid = isOfHLevelSuc 3 + +isOfHLevelPlus : (m : HLevel) → isOfHLevel n A → isOfHLevel (m + n) A isOfHLevelPlus zero hA = hA isOfHLevelPlus (suc m) hA = isOfHLevelSuc _ (isOfHLevelPlus m hA) -isContr→isOfHLevel : (n : ℕ) → isContr A → isOfHLevel n A -isContr→isOfHLevel {A = A} n cA = subst (λ m → isOfHLevel m A) (+-zero n) (isOfHLevelPlus n cA) +isContr→isOfHLevel : (n : HLevel) → isContr A → isOfHLevel n A +isContr→isOfHLevel zero cA = cA +isContr→isOfHLevel (suc n) cA = isOfHLevelSuc _ (isContr→isOfHLevel n cA) + +isProp→isOfHLevelSuc : (n : HLevel) → isProp A → isOfHLevel (suc n) A +isProp→isOfHLevelSuc zero pA = pA +isProp→isOfHLevelSuc (suc n) pA = isOfHLevelSuc _ (isProp→isOfHLevelSuc n pA) -isProp→isOfHLevelSuc : (n : ℕ) → isProp A → isOfHLevel (suc n) A -isProp→isOfHLevelSuc {A = A} n pA = subst (λ m → isOfHLevel m A) (+-comm n 1) (isOfHLevelPlus n pA) +isOfHLevelPlus' : (m : HLevel) → isOfHLevel m A → isOfHLevel (m + n) A +isOfHLevelPlus' {n = n} 0 = isContr→isOfHLevel n +isOfHLevelPlus' {n = n} 1 = isProp→isOfHLevelSuc n +isOfHLevelPlus' {n = n} (suc (suc m)) hA a₀ a₁ = isOfHLevelPlus' (suc m) (hA a₀ a₁) --- hlevel of path and dependent path types +-- hlevel of path types isProp→isContrPath : isProp A → (x y : A) → isContr (x ≡ y) isProp→isContrPath h x y = h x y , isProp→isSet h x y _ @@ -71,33 +101,21 @@ isProp→isContrPath h x y = h x y , isProp→isSet h x y _ isContr→isContrPath : isContr A → (x y : A) → isContr (x ≡ y) isContr→isContrPath cA = isProp→isContrPath (isContr→isProp cA) -isOfHLevelPath' : (n : ℕ) → isOfHLevel (suc n) A → (x y : A) → isOfHLevel n (x ≡ y) +isOfHLevelPath' : (n : HLevel) → isOfHLevel (suc n) A → (x y : A) → isOfHLevel n (x ≡ y) isOfHLevelPath' 0 = isProp→isContrPath isOfHLevelPath' (suc n) h x y = h x y -isOfHLevelPath : (n : ℕ) → isOfHLevel n A → (x y : A) → isOfHLevel n (x ≡ y) +isOfHLevelPath'⁻ : (n : HLevel) → ((x y : A) → isOfHLevel n (x ≡ y)) → isOfHLevel (suc n) A +isOfHLevelPath'⁻ zero h x y = h x y .fst +isOfHLevelPath'⁻ (suc n) h = h + +isOfHLevelPath : (n : HLevel) → isOfHLevel n A → (x y : A) → isOfHLevel n (x ≡ y) isOfHLevelPath 0 h x y = isContr→isContrPath h x y isOfHLevelPath (suc n) h x y = isOfHLevelSuc n (isOfHLevelPath' n h x y) -isOfHLevelPathP' : {A : I → Type ℓ} (n : ℕ) - → (∀ i → isOfHLevel (suc n) (A i)) - → (x : A i0) (y : A i1) → isOfHLevel n (PathP A x y) -isOfHLevelPathP' {A = A} n h x y = transport⁻ (λ i → isOfHLevel n (PathP≡Path A x y i)) - (isOfHLevelPath' n (h i1) _ _) - -isOfHLevelPathP : {A : I → Type ℓ} (n : ℕ) - → (∀ i → isOfHLevel n (A i)) - → (x : A i0) (y : A i1) → isOfHLevel n (PathP A x y) -isOfHLevelPathP {A = A} n h x y = transport⁻ (λ i → isOfHLevel n (PathP≡Path A x y i)) - (isOfHLevelPath n (h i1) _ _) - -isProp→isContrPathP : {A : I → Type ℓ} → (∀ i → isProp (A i)) - → (x : A i0) (y : A i1) → isContr (PathP A x y) -isProp→isContrPathP h x y = isProp→PathP h x y , isOfHLevelPathP 1 h x y _ - -- h-level of isOfHLevel -isPropIsOfHLevel : (n : ℕ) → isProp (isOfHLevel n A) +isPropIsOfHLevel : (n : HLevel) → isProp (isOfHLevel n A) isPropIsOfHLevel 0 = isPropIsContr isPropIsOfHLevel 1 = isPropIsProp isPropIsOfHLevel (suc (suc n)) f g i a b = @@ -106,32 +124,14 @@ isPropIsOfHLevel (suc (suc n)) f g i a b = isPropIsSet : isProp (isSet A) isPropIsSet = isPropIsOfHLevel 2 --- Fillers for cubes from h-level +isPropIsGroupoid : isProp (isGroupoid A) +isPropIsGroupoid = isPropIsOfHLevel 3 -isSet→isSet' : isSet A → isSet' A -isSet→isSet' {A = A} Aset a₀₋ a₁₋ a₋₀ a₋₁ = - transport⁻ (PathP≡Path (λ i → a₋₀ i ≡ a₋₁ i) a₀₋ a₁₋) (Aset _ _ _ _) +isPropIs2Groupoid : isProp (is2Groupoid A) +isPropIs2Groupoid = isPropIsOfHLevel 4 -isSet'→isSet : isSet' A → isSet A -isSet'→isSet {A = A} Aset' x y p q = Aset' p q refl refl - -isGroupoid→isGroupoid' : isGroupoid A → isGroupoid' A -isGroupoid→isGroupoid' {A = A} Agpd a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ = - transport⁻ (PathP≡Path (λ i → Square (a₋₀₋ i) (a₋₁₋ i) (a₋₋₀ i) (a₋₋₁ i)) a₀₋₋ a₁₋₋) - (isGroupoid→isPropSquare _ _ _ _ _ _) - where - isGroupoid→isPropSquare : - {a₀₀ a₀₁ : A} (a₀₋ : a₀₀ ≡ a₀₁) - {a₁₀ a₁₁ : A} (a₁₋ : a₁₀ ≡ a₁₁) - (a₋₀ : a₀₀ ≡ a₁₀) (a₋₁ : a₀₁ ≡ a₁₁) - → isProp (Square a₀₋ a₁₋ a₋₀ a₋₁) - isGroupoid→isPropSquare a₀₋ a₁₋ a₋₀ a₋₁ = - transport⁻ - (cong isProp (PathP≡Path (λ i → a₋₀ i ≡ a₋₁ i) a₀₋ a₁₋)) - (Agpd _ _ _ _) - -isGroupoid'→isGroupoid : isGroupoid' A → isGroupoid A -isGroupoid'→isGroupoid Agpd' x y p q r s = Agpd' r s refl refl refl refl +TypeOfHLevel≡ : (n : HLevel) {X Y : TypeOfHLevel ℓ n} → ⟨ X ⟩ ≡ ⟨ Y ⟩ → X ≡ Y +TypeOfHLevel≡ n = Σ≡Prop (λ _ → isPropIsOfHLevel n) -- hlevels are preserved by retracts (and consequently equivalences) @@ -140,7 +140,8 @@ isContrRetract → (f : A → B) (g : B → A) → (h : retract f g) → (v : isContr B) → isContr A -isContrRetract f g h (b , p) = (g b , λ x → (cong g (p (f x))) ∙ (h x)) +fst (isContrRetract f g h (b , p)) = g b +snd (isContrRetract f g h (b , p)) x = (cong g (p (f x))) ∙ (h x) isPropRetract : {B : Type ℓ} @@ -154,40 +155,152 @@ isPropRetract f g h p x y i = ; (i = i1) → h y j}) (g (p (f x) (f y) i)) +isSetRetract + : {B : Type ℓ} + (f : A → B) (g : B → A) + (h : (x : A) → g (f x) ≡ x) + → isSet B → isSet A +isSetRetract f g h set x y p q i j = + hcomp (λ k → λ { (i = i0) → h (p j) k + ; (i = i1) → h (q j) k + ; (j = i0) → h x k + ; (j = i1) → h y k}) + (g (set (f x) (f y) + (cong f p) (cong f q) i j)) + +isGroupoidRetract + : {B : Type ℓ} + (f : A → B) (g : B → A) + (h : (x : A) → g (f x) ≡ x) + → isGroupoid B → isGroupoid A +isGroupoidRetract f g h grp x y p q P Q i j k = + hcomp ((λ l → λ { (i = i0) → h (P j k) l + ; (i = i1) → h (Q j k) l + ; (j = i0) → h (p k) l + ; (j = i1) → h (q k) l + ; (k = i0) → h x l + ; (k = i1) → h y l})) + (g (grp (f x) (f y) (cong f p) (cong f q) + (cong (cong f) P) (cong (cong f) Q) i j k)) + +is2GroupoidRetract + : {B : Type ℓ} + (f : A → B) (g : B → A) + (h : (x : A) → g (f x) ≡ x) + → is2Groupoid B → is2Groupoid A +is2GroupoidRetract f g h grp x y p q P Q R S i j k l = + hcomp (λ r → λ { (i = i0) → h (R j k l) r + ; (i = i1) → h (S j k l) r + ; (j = i0) → h (P k l) r + ; (j = i1) → h (Q k l) r + ; (k = i0) → h (p l) r + ; (k = i1) → h (q l) r + ; (l = i0) → h x r + ; (l = i1) → h y r}) + (g (grp (f x) (f y) (cong f p) (cong f q) + (cong (cong f) P) (cong (cong f) Q) + (cong (cong (cong f)) R) (cong (cong (cong f)) S) i j k l)) + isOfHLevelRetract - : (n : ℕ) {B : Type ℓ} + : (n : HLevel) {B : Type ℓ} (f : A → B) (g : B → A) (h : (x : A) → g (f x) ≡ x) → isOfHLevel n B → isOfHLevel n A isOfHLevelRetract 0 = isContrRetract isOfHLevelRetract 1 = isPropRetract -isOfHLevelRetract (suc (suc n)) f g h ofLevel x y = - isOfHLevelRetract (suc n) - (cong f) - (λ q i → - hcomp - (λ j → λ - { (i = i0) → h x j - ; (i = i1) → h y j}) - (g (q i))) - (λ p k i → - hcomp - (λ j → λ - { (i = i0) → h x (j ∨ k) - ; (i = i1) → h y (j ∨ k) - ; (k = i1) → p i}) - (h (p i) k)) - (ofLevel (f x) (f y)) - -isOfHLevelRespectEquiv : {A : Type ℓ} {B : Type ℓ'} → (n : ℕ) → A ≃ B → isOfHLevel n A → isOfHLevel n B -isOfHLevelRespectEquiv n eq = isOfHLevelRetract n (invEq eq) (eq .fst) (retEq eq) +isOfHLevelRetract 2 = isSetRetract +isOfHLevelRetract 3 = isGroupoidRetract +isOfHLevelRetract 4 = is2GroupoidRetract +isOfHLevelRetract (suc (suc (suc (suc (suc n))))) f g h ofLevel x y p q P Q R S = + isOfHLevelRetract (suc n) (cong (cong (cong (cong f)))) + (λ s i j k l → + hcomp (λ r → λ { (i = i0) → h (R j k l) r + ; (i = i1) → h (S j k l) r + ; (j = i0) → h (P k l) r + ; (j = i1) → h (Q k l) r + ; (k = i0) → h (p l) r + ; (k = i1) → h (q l) r + ; (l = i0) → h x r + ; (l = i1) → h y r}) + (g (s i j k l))) + (λ s i j k l m → + hcomp (λ n → λ { (i = i1) → s j k l m + ; (j = i0) → h (R k l m) (i ∨ n) + ; (j = i1) → h (S k l m) (i ∨ n) + ; (k = i0) → h (P l m) (i ∨ n) + ; (k = i1) → h (Q l m) (i ∨ n) + ; (l = i0) → h (p m) (i ∨ n) + ; (l = i1) → h (q m) (i ∨ n) + ; (m = i0) → h x (i ∨ n) + ; (m = i1) → h y (i ∨ n) }) + (h (s j k l m) i)) + (ofLevel (f x) (f y) + (cong f p) (cong f q) + (cong (cong f) P) (cong (cong f) Q) + (cong (cong (cong f)) R) (cong (cong (cong f)) S)) + +isOfHLevelRetractFromIso : {A : Type ℓ} {B : Type ℓ'} (n : HLevel) → Iso A B → isOfHLevel n B → isOfHLevel n A +isOfHLevelRetractFromIso n e hlev = isOfHLevelRetract n (Iso.fun e) (Iso.inv e) (Iso.leftInv e) hlev + +isOfHLevelRespectEquiv : {A : Type ℓ} {B : Type ℓ'} → (n : HLevel) → A ≃ B → isOfHLevel n A → isOfHLevel n B +isOfHLevelRespectEquiv n eq = isOfHLevelRetract n (invEq eq) (eq .fst) (secEq eq) + +isContrRetractOfConstFun : {A : Type ℓ} {B : Type ℓ'} (b₀ : B) + → Σ[ f ∈ (B → A) ] ((x : A) → (f ∘ (λ _ → b₀)) x ≡ x) + → isContr A +fst (isContrRetractOfConstFun b₀ ret) = ret .fst b₀ +snd (isContrRetractOfConstFun b₀ ret) y = ret .snd y + +-- h-level of dependent path types + +isOfHLevelPathP' : {A : I → Type ℓ} (n : HLevel) + → isOfHLevel (suc n) (A i1) + → (x : A i0) (y : A i1) → isOfHLevel n (PathP A x y) +isOfHLevelPathP' {A = A} n h x y = + isOfHLevelRetractFromIso n (PathPIsoPath _ x y) (isOfHLevelPath' n h _ _) + +isOfHLevelPathP : {A : I → Type ℓ} (n : HLevel) + → isOfHLevel n (A i1) + → (x : A i0) (y : A i1) → isOfHLevel n (PathP A x y) +isOfHLevelPathP {A = A} n h x y = + isOfHLevelRetractFromIso n (PathPIsoPath _ x y) (isOfHLevelPath n h _ _) + +-- Fillers for cubes from h-level + +isSet→isSet' : isSet A → isSet' A +isSet→isSet' Aset _ _ _ _ = toPathP (Aset _ _ _ _) + +isSet'→isSet : isSet' A → isSet A +isSet'→isSet {A = A} Aset' x y p q = Aset' p q refl refl + +isSet→SquareP : + {A : I → I → Type ℓ} + (isSet : (i j : I) → isSet (A i j)) + {a₀₀ : A i0 i0} {a₀₁ : A i0 i1} (a₀₋ : PathP (λ j → A i0 j) a₀₀ a₀₁) + {a₁₀ : A i1 i0} {a₁₁ : A i1 i1} (a₁₋ : PathP (λ j → A i1 j) a₁₀ a₁₁) + (a₋₀ : PathP (λ i → A i i0) a₀₀ a₁₀) (a₋₁ : PathP (λ i → A i i1) a₀₁ a₁₁) + → SquareP A a₀₋ a₁₋ a₋₀ a₋₁ +isSet→SquareP isset a₀₋ a₁₋ a₋₀ a₋₁ = + PathPIsoPath _ _ _ .Iso.inv (isOfHLevelPathP' 1 (isset _ _) _ _ _ _ ) + +isGroupoid→isGroupoid' : isGroupoid A → isGroupoid' A +isGroupoid→isGroupoid' {A = A} Agpd a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ = + PathPIsoPath (λ i → Square (a₋₀₋ i) (a₋₁₋ i) (a₋₋₀ i) (a₋₋₁ i)) a₀₋₋ a₁₋₋ .Iso.inv + (isGroupoid→isPropSquare _ _ _ _ _ _) + where + isGroupoid→isPropSquare : + {a₀₀ a₀₁ : A} (a₀₋ : a₀₀ ≡ a₀₁) + {a₁₀ a₁₁ : A} (a₁₋ : a₁₀ ≡ a₁₁) + (a₋₀ : a₀₀ ≡ a₁₀) (a₋₁ : a₀₁ ≡ a₁₁) + → isProp (Square a₀₋ a₁₋ a₋₀ a₋₁) + isGroupoid→isPropSquare a₀₋ a₁₋ a₋₀ a₋₁ = + isOfHLevelRetractFromIso 1 (PathPIsoPath (λ i → a₋₀ i ≡ a₋₁ i) a₀₋ a₁₋) (Agpd _ _ _ _) +isGroupoid'→isGroupoid : isGroupoid' A → isGroupoid A +isGroupoid'→isGroupoid Agpd' x y p q r s = Agpd' r s refl refl refl refl -- h-level of Σ-types -isContrΣ - : isContr A - → ((x : A) → isContr (B x)) - → isContr (Σ[ x ∈ A ] B x) +isContrΣ : isContr A → ((x : A) → isContr (B x)) → isContr (Σ A B) isContrΣ {A = A} {B = B} (a , p) q = let h : (x : A) (y : B x) → (q x) .fst ≡ y h x y = (q x) .snd y @@ -195,96 +308,237 @@ isContrΣ {A = A} {B = B} (a , p) q = , ( λ x i → p (x .fst) i , h (p (x .fst) i) (transp (λ j → B (p (x .fst) (i ∨ ~ j))) i (x .snd)) i)) -ΣProp≡ - : ((x : A) → isProp (B x)) → {u v : Σ[ a ∈ A ] B a} - → (p : u .fst ≡ v .fst) → u ≡ v -ΣProp≡ pB {u} {v} p i = (p i) , isProp→PathP (λ i → pB (p i)) (u .snd) (v .snd) i +isContrΣ′ : (ca : isContr A) → isContr (B (fst ca)) → isContr (Σ A B) +isContrΣ′ ca cb = isContrΣ ca (λ x → subst _ (snd ca x) cb) + +section-Σ≡Prop + : (pB : (x : A) → isProp (B x)) {u v : Σ A B} + → section (Σ≡Prop pB {u} {v}) (cong fst) +section-Σ≡Prop {A = A} pB {u} {v} p j i = + (p i .fst) , isProp→PathP (λ i → isOfHLevelPath 1 (pB (fst (p i))) + (Σ≡Prop pB {u} {v} (cong fst p) i .snd) + (p i .snd) ) + refl refl i j -ΣProp≡-equiv - : (pB : (x : A) → isProp (B x)) {u v : Σ[ a ∈ A ] B a} - → isEquiv (ΣProp≡ pB {u} {v}) -ΣProp≡-equiv {A = A} pB {u} {v} = isoToIsEquiv (iso (ΣProp≡ pB) (cong fst) sq (λ _ → refl)) - where sq : (p : u ≡ v) → ΣProp≡ pB (cong fst p) ≡ p - sq p j i = (p i .fst) , isProp→PathP (λ i → isOfHLevelPath 1 (pB (fst (p i))) - (ΣProp≡ pB {u} {v} (cong fst p) i .snd) - (p i .snd) ) - refl refl i j +isEquiv-Σ≡Prop + : (pB : (x : A) → isProp (B x)) {u v : Σ A B} + → isEquiv (Σ≡Prop pB {u} {v}) +isEquiv-Σ≡Prop {A = A} pB {u} {v} = isoToIsEquiv (iso (Σ≡Prop pB) (cong fst) (section-Σ≡Prop pB) (λ _ → refl)) -isPropΣ : isProp A → ((x : A) → isProp (B x)) → isProp (Σ[ x ∈ A ] B x) -isPropΣ pA pB t u = ΣProp≡ pB (pA (t .fst) (u .fst)) +isPropΣ : isProp A → ((x : A) → isProp (B x)) → isProp (Σ A B) +isPropΣ pA pB t u = Σ≡Prop pB (pA (t .fst) (u .fst)) isOfHLevelΣ : ∀ n → isOfHLevel n A → ((x : A) → isOfHLevel n (B x)) - → isOfHLevel n (Σ A B) + → isOfHLevel n (Σ A B) isOfHLevelΣ 0 = isContrΣ isOfHLevelΣ 1 = isPropΣ isOfHLevelΣ {B = B} (suc (suc n)) h1 h2 x y = - let h3 : isOfHLevel (suc n) (x Σ≡T y) - h3 = isOfHLevelΣ (suc n) (h1 (fst x) (fst y)) λ p → h2 (p i1) - (subst B p (snd x)) (snd y) - in transport (λ i → isOfHLevel (suc n) (pathSigma≡sigmaPath x y (~ i))) h3 + isOfHLevelRetractFromIso (suc n) + (invIso (IsoΣPathTransportPathΣ _ _)) + (isOfHLevelΣ (suc n) (h1 (fst x) (fst y)) λ x → h2 _ _ _) + +isSetΣ : isSet A → ((x : A) → isSet (B x)) → isSet (Σ A B) +isSetΣ = isOfHLevelΣ 2 + +isGroupoidΣ : isGroupoid A → ((x : A) → isGroupoid (B x)) → isGroupoid (Σ A B) +isGroupoidΣ = isOfHLevelΣ 3 + +is2GroupoidΣ : is2Groupoid A → ((x : A) → is2Groupoid (B x)) → is2Groupoid (Σ A B) +is2GroupoidΣ = isOfHLevelΣ 4 + +-- h-level of × + +isProp× : {A : Type ℓ} {B : Type ℓ'} → isProp A → isProp B → isProp (A × B) +isProp× pA pB = isPropΣ pA (λ _ → pB) + +isProp×2 : {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + → isProp A → isProp B → isProp C → isProp (A × B × C) +isProp×2 pA pB pC = isProp× pA (isProp× pB pC) + +isProp×3 : {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + → isProp A → isProp B → isProp C → isProp D → isProp (A × B × C × D) +isProp×3 pA pB pC pD = isProp×2 pA pB (isProp× pC pD) + +isProp×4 : {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} {E : Type ℓ''''} + → isProp A → isProp B → isProp C → isProp D → isProp E → isProp (A × B × C × D × E) +isProp×4 pA pB pC pD pE = isProp×3 pA pB pC (isProp× pD pE) + +isProp×5 : {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} {E : Type ℓ''''} {F : Type ℓ'''''} + → isProp A → isProp B → isProp C → isProp D → isProp E → isProp F + → isProp (A × B × C × D × E × F) +isProp×5 pA pB pC pD pE pF = isProp×4 pA pB pC pD (isProp× pE pF) + + +isOfHLevel× : ∀ {A : Type ℓ} {B : Type ℓ'} n → isOfHLevel n A → isOfHLevel n B + → isOfHLevel n (A × B) +isOfHLevel× n hA hB = isOfHLevelΣ n hA (λ _ → hB) + +isSet× : ∀ {A : Type ℓ} {B : Type ℓ'} → isSet A → isSet B → isSet (A × B) +isSet× = isOfHLevel× 2 + +isGroupoid× : ∀ {A : Type ℓ} {B : Type ℓ'} → isGroupoid A → isGroupoid B + → isGroupoid (A × B) +isGroupoid× = isOfHLevel× 3 + +is2Groupoid× : ∀ {A : Type ℓ} {B : Type ℓ'} → is2Groupoid A → is2Groupoid B + → is2Groupoid (A × B) +is2Groupoid× = isOfHLevel× 4 -- h-level of Π-types -isOfHLevelPi - : ∀ n - → ((x : A) → isOfHLevel n (B x)) - → isOfHLevel n ((x : A) → B x) -isOfHLevelPi 0 h = (λ x → fst (h x)) , λ f i y → snd (h y) (f y) i -isOfHLevelPi 1 h f g i x = (h x) (f x) (g x) i -isOfHLevelPi (suc (suc n)) h f g = - subst (isOfHLevel (suc n)) funExtPath (isOfHLevelPi (suc n) λ x → h x (f x) (g x)) - -isPropPi : (h : (x : A) → isProp (B x)) → isProp ((x : A) → B x) -isPropPi = isOfHLevelPi 1 - -isSetPi : ((x : A) → isSet (B x)) → isSet ((x : A) → B x) -isSetPi = isOfHLevelPi 2 - -isOfHLevelPi⁻ : ∀ {A : Type ℓ} {B : Type ℓ'} n - → isOfHLevel n (A → B) - → (A → isOfHLevel n B) -isOfHLevelPi⁻ 0 h x = fst h x , λ y → funExt⁻ (snd h (const y)) x -isOfHLevelPi⁻ 1 h x y z = funExt⁻ (h (const y) (const z)) x -isOfHLevelPi⁻ (suc (suc n)) h x y z = - isOfHLevelPi⁻ (suc n) (subst (isOfHLevel (suc n)) (sym funExtPath) (h (const y) (const z))) x +isOfHLevelΠ : ∀ n → ((x : A) → isOfHLevel n (B x)) + → isOfHLevel n ((x : A) → B x) +isOfHLevelΠ 0 h = (λ x → fst (h x)) , λ f i y → snd (h y) (f y) i +isOfHLevelΠ 1 h f g i x = (h x) (f x) (g x) i +isOfHLevelΠ 2 h f g F G i j z = h z (f z) (g z) (funExt⁻ F z) (funExt⁻ G z) i j +isOfHLevelΠ 3 h f g p q P Q i j k z = + h z (f z) (g z) + (funExt⁻ p z) (funExt⁻ q z) + (cong (λ f → funExt⁻ f z) P) (cong (λ f → funExt⁻ f z) Q) i j k +isOfHLevelΠ 4 h f g p q P Q R S i j k l z = + h z (f z) (g z) + (funExt⁻ p z) (funExt⁻ q z) + (cong (λ f → funExt⁻ f z) P) (cong (λ f → funExt⁻ f z) Q) + (cong (cong (λ f → funExt⁻ f z)) R) (cong (cong (λ f → funExt⁻ f z)) S) i j k l +isOfHLevelΠ (suc (suc (suc (suc (suc n))))) h f g p q P Q R S = + isOfHLevelRetract (suc n) + (cong (cong (cong funExt⁻))) (cong (cong (cong funExt))) (λ _ → refl) + (isOfHLevelΠ (suc (suc (suc (suc n)))) (λ x → h x (f x) (g x)) + (funExt⁻ p) (funExt⁻ q) + (cong funExt⁻ P) (cong funExt⁻ Q) + (cong (cong funExt⁻) R) (cong (cong funExt⁻) S)) + +isPropΠ : (h : (x : A) → isProp (B x)) → isProp ((x : A) → B x) +isPropΠ = isOfHLevelΠ 1 + +isPropΠ2 : (h : (x : A) (y : B x) → isProp (C x y)) + → isProp ((x : A) (y : B x) → C x y) +isPropΠ2 h = isPropΠ λ x → isPropΠ λ y → h x y + +isPropΠ3 : (h : (x : A) (y : B x) (z : C x y) → isProp (D x y z)) + → isProp ((x : A) (y : B x) (z : C x y) → D x y z) +isPropΠ3 h = isPropΠ λ x → isPropΠ λ y → isPropΠ λ z → h x y z + +isPropΠ4 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) → isProp (E x y z w)) + → isProp ((x : A) (y : B x) (z : C x y) (w : D x y z) → E x y z w) +isPropΠ4 h = isPropΠ λ _ → isPropΠ3 λ _ → h _ _ + +isPropImplicitΠ : (h : (x : A) → isProp (B x)) → isProp ({x : A} → B x) +isPropImplicitΠ h f g i {x} = h x (f {x}) (g {x}) i + +isPropImplicitΠ2 : (h : (x : A) (y : B x) → isProp (C x y)) → isProp ({x : A} {y : B x} → C x y) +isPropImplicitΠ2 h = isPropImplicitΠ (λ x → isPropImplicitΠ (λ y → h x y)) + +isProp→ : {A : Type ℓ} {B : Type ℓ'} → isProp B → isProp (A → B) +isProp→ pB = isPropΠ λ _ → pB + +isSetΠ : ((x : A) → isSet (B x)) → isSet ((x : A) → B x) +isSetΠ = isOfHLevelΠ 2 + +isSetΠ2 : (h : (x : A) (y : B x) → isSet (C x y)) + → isSet ((x : A) (y : B x) → C x y) +isSetΠ2 h = isSetΠ λ x → isSetΠ λ y → h x y + +isSetΠ3 : (h : (x : A) (y : B x) (z : C x y) → isSet (D x y z)) + → isSet ((x : A) (y : B x) (z : C x y) → D x y z) +isSetΠ3 h = isSetΠ λ x → isSetΠ λ y → isSetΠ λ z → h x y z + +isGroupoidΠ : ((x : A) → isGroupoid (B x)) → isGroupoid ((x : A) → B x) +isGroupoidΠ = isOfHLevelΠ 3 + +isGroupoidΠ2 : (h : (x : A) (y : B x) → isGroupoid (C x y)) → isGroupoid ((x : A) (y : B x) → C x y) +isGroupoidΠ2 h = isGroupoidΠ λ _ → isGroupoidΠ λ _ → h _ _ + +isGroupoidΠ3 : (h : (x : A) (y : B x) (z : C x y) → isGroupoid (D x y z)) + → isGroupoid ((x : A) (y : B x) (z : C x y) → D x y z) +isGroupoidΠ3 h = isGroupoidΠ λ _ → isGroupoidΠ2 λ _ → h _ _ + +isGroupoidΠ4 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) → isGroupoid (E x y z w)) + → isGroupoid ((x : A) (y : B x) (z : C x y) (w : D x y z) → E x y z w) +isGroupoidΠ4 h = isGroupoidΠ λ _ → isGroupoidΠ3 λ _ → h _ _ + +is2GroupoidΠ : ((x : A) → is2Groupoid (B x)) → is2Groupoid ((x : A) → B x) +is2GroupoidΠ = isOfHLevelΠ 4 + +isOfHLevelΠ⁻ : ∀ {A : Type ℓ} {B : Type ℓ'} n + → isOfHLevel n (A → B) → (A → isOfHLevel n B) +isOfHLevelΠ⁻ 0 h x = fst h x , λ y → funExt⁻ (snd h (const y)) x +isOfHLevelΠ⁻ 1 h x y z = funExt⁻ (h (const y) (const z)) x +isOfHLevelΠ⁻ (suc (suc n)) h x y z = + isOfHLevelΠ⁻ (suc n) (isOfHLevelRetractFromIso (suc n) funExtIso (h _ _)) x -- h-level of A ≃ B and A ≡ B -isOfHLevel≃ : ∀ n → {A B : Type ℓ} (hA : isOfHLevel n A) (hB : isOfHLevel n B) → isOfHLevel n (A ≃ B) -isOfHLevel≃ zero {A = A} {B = B} hA hB = A≃B , contr +isOfHLevel≃ + : ∀ n {A : Type ℓ} {B : Type ℓ'} + → (hA : isOfHLevel n A) (hB : isOfHLevel n B) → isOfHLevel n (A ≃ B) +isOfHLevel≃ zero {A = A} {B = B} hA hB = isContr→Equiv hA hB , contr where - A≃B : A ≃ B - A≃B = isoToEquiv (iso (λ _ → fst hB) (λ _ → fst hA) (snd hB ) (snd hA)) + contr : (y : A ≃ B) → isContr→Equiv hA hB ≡ y + contr y = Σ≡Prop isPropIsEquiv (funExt (λ a → snd hB (fst y a))) - contr : (y : A ≃ B) → A≃B ≡ y - contr y = ΣProp≡ isPropIsEquiv (funExt (λ a → snd hB (fst y a))) - -isOfHLevel≃ (suc n) hA hB = - isOfHLevelΣ (suc n) (isOfHLevelPi (suc n) (λ _ → hB)) - (λ a → subst (λ n → isOfHLevel n (isEquiv a)) (+-comm n 1) (isOfHLevelPlus n (isPropIsEquiv a))) +isOfHLevel≃ (suc n) {A = A} {B = B} hA hB = + isOfHLevelΣ (suc n) (isOfHLevelΠ _ λ _ → hB) + (λ f → isProp→isOfHLevelSuc n (isPropIsEquiv f)) isOfHLevel≡ : ∀ n → {A B : Type ℓ} (hA : isOfHLevel n A) (hB : isOfHLevel n B) → isOfHLevel n (A ≡ B) -isOfHLevel≡ n hA hB = isOfHLevelRespectEquiv n (invEquiv univalence) (isOfHLevel≃ n hA hB) +isOfHLevel≡ n hA hB = isOfHLevelRetractFromIso n univalenceIso (isOfHLevel≃ n hA hB) --- h-level of HLevel +isOfHLevel⁺≃ₗ + : ∀ n {A : Type ℓ} {B : Type ℓ'} + → isOfHLevel (suc n) A → isOfHLevel (suc n) (A ≃ B) +isOfHLevel⁺≃ₗ zero pA e = isOfHLevel≃ 1 pA (isOfHLevelRespectEquiv 1 e pA) e +isOfHLevel⁺≃ₗ (suc n) hA e = isOfHLevel≃ m hA (isOfHLevelRespectEquiv m e hA) e + where + m = suc (suc n) + +isOfHLevel⁺≃ᵣ + : ∀ n {A : Type ℓ} {B : Type ℓ'} + → isOfHLevel (suc n) B → isOfHLevel (suc n) (A ≃ B) +isOfHLevel⁺≃ᵣ zero pB e + = isOfHLevel≃ 1 (isPropRetract (e .fst) (invEq e) (retEq e) pB) pB e +isOfHLevel⁺≃ᵣ (suc n) hB e + = isOfHLevel≃ m (isOfHLevelRetract m (e .fst) (invEq e) (retEq e) hB) hB e + where + m = suc (suc n) + +isOfHLevel⁺≡ₗ + : ∀ n → {A B : Type ℓ} + → isOfHLevel (suc n) A → isOfHLevel (suc n) (A ≡ B) +isOfHLevel⁺≡ₗ zero pA P = isOfHLevel≡ 1 pA (subst isProp P pA) P +isOfHLevel⁺≡ₗ (suc n) hA P + = isOfHLevel≡ m hA (subst (isOfHLevel m) P hA) P + where + m = suc (suc n) + +isOfHLevel⁺≡ᵣ + : ∀ n → {A B : Type ℓ} + → isOfHLevel (suc n) B → isOfHLevel (suc n) (A ≡ B) +isOfHLevel⁺≡ᵣ zero pB P = isOfHLevel≡ 1 (subst⁻ isProp P pB) pB P +isOfHLevel⁺≡ᵣ (suc n) hB P + = isOfHLevel≡ m (subst⁻ (isOfHLevel m) P hB) hB P + where + m = suc (suc n) + +-- h-level of TypeOfHLevel -isPropHContr : isProp (HLevel ℓ 0) -isPropHContr x y = ΣProp≡ (λ _ → isPropIsContr) ((isOfHLevel≡ 0 (x .snd) (y .snd) .fst)) +isPropHContr : isProp (TypeOfHLevel ℓ 0) +isPropHContr x y = Σ≡Prop (λ _ → isPropIsContr) (isOfHLevel≡ 0 (x .snd) (y .snd) .fst) -isOfHLevelHLevel : ∀ n → isOfHLevel (suc n) (HLevel ℓ n) -isOfHLevelHLevel 0 = isPropHContr -isOfHLevelHLevel (suc n) x y = subst (isOfHLevel (suc n)) eq (isOfHLevel≡ (suc n) (snd x) (snd y)) - where eq : ∀ {A B : Type ℓ} {hA : isOfHLevel (suc n) A} {hB : isOfHLevel (suc n) B} - → (A ≡ B) ≡ ((A , hA) ≡ (B , hB)) - eq = ua (_ , ΣProp≡-equiv (λ _ → isPropIsOfHLevel (suc n))) +isOfHLevelTypeOfHLevel : ∀ n → isOfHLevel (suc n) (TypeOfHLevel ℓ n) +isOfHLevelTypeOfHLevel zero = isPropHContr +isOfHLevelTypeOfHLevel (suc n) (X , a) (Y , b) = + isOfHLevelRetract (suc n) (cong fst) (Σ≡Prop λ _ → isPropIsOfHLevel (suc n)) + (section-Σ≡Prop λ _ → isPropIsOfHLevel (suc n)) + (isOfHLevel≡ (suc n) a b) isSetHProp : isSet (hProp ℓ) -isSetHProp = isOfHLevelHLevel 1 +isSetHProp = isOfHLevelTypeOfHLevel 1 -- h-level of lifted type -isOfHLevelLift : ∀ {ℓ ℓ'} (n : ℕ) {A : Type ℓ} → isOfHLevel n A → isOfHLevel n (Lift {j = ℓ'} A) +isOfHLevelLift : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} → isOfHLevel n A → isOfHLevel n (Lift {j = ℓ'} A) isOfHLevelLift n = isOfHLevelRetract n lower lift λ _ → refl ---------------------------- @@ -294,6 +548,9 @@ isOfHLevelLift n = isOfHLevelRetract n lower lift λ _ → refl inhProp→isContr : A → isProp A → isContr A inhProp→isContr x h = x , h x +extend : isContr A → (∀ φ → (u : Partial φ A) → Sub A φ u) +extend (x , p) φ u = inS (hcomp (λ { j (φ = i1) → p (u 1=1) j }) x) + isContrPartial→isContr : ∀ {ℓ} {A : Type ℓ} → (extend : ∀ φ → Partial φ A → A) → (∀ u → u ≡ (extend i1 λ { _ → u})) @@ -309,12 +566,12 @@ isContrPartial→isContr {A = A} extend law -- Dependent h-level over a type -isOfHLevelDep : ℕ → {A : Type ℓ} (B : A → Type ℓ') → Type (ℓ-max ℓ ℓ') +isOfHLevelDep : HLevel → {A : Type ℓ} (B : A → Type ℓ') → Type (ℓ-max ℓ ℓ') isOfHLevelDep 0 {A = A} B = {a : A} → Σ[ b ∈ B a ] ({a' : A} (b' : B a') (p : a ≡ a') → PathP (λ i → B (p i)) b b') isOfHLevelDep 1 {A = A} B = {a0 a1 : A} (b0 : B a0) (b1 : B a1) (p : a0 ≡ a1) → PathP (λ i → B (p i)) b0 b1 isOfHLevelDep (suc (suc n)) {A = A} B = {a0 a1 : A} (b0 : B a0) (b1 : B a1) → isOfHLevelDep (suc n) {A = a0 ≡ a1} (λ p → PathP (λ i → B (p i)) b0 b1) -isOfHLevel→isOfHLevelDep : (n : ℕ) +isOfHLevel→isOfHLevelDep : (n : HLevel) → {A : Type ℓ} {B : A → Type ℓ'} (h : (a : A) → isOfHLevel n (B a)) → isOfHLevelDep n {A = A} B isOfHLevel→isOfHLevelDep 0 h {a} = (h a .fst , λ b' p → isProp→PathP (λ i → isContr→isProp (h (p i))) (h a .fst) b') @@ -324,6 +581,59 @@ isOfHLevel→isOfHLevelDep (suc (suc n)) {A = A} {B} h {a0} {a1} b0 b1 = where helper : (a1 : A) (p : a0 ≡ a1) (b1 : B a1) → isOfHLevel (suc n) (PathP (λ i → B (p i)) b0 b1) - helper a1 p b1 = J - (λ a1 p → ∀ b1 → isOfHLevel (suc n) (PathP (λ i → B (p i)) b0 b1)) + helper a1 p b1 = J (λ a1 p → ∀ b1 → isOfHLevel (suc n) (PathP (λ i → B (p i)) b0 b1)) (λ _ → h _ _ _) p b1 + +isContrDep→isPropDep : isOfHLevelDep 0 B → isOfHLevelDep 1 B +isContrDep→isPropDep {B = B} Bctr {a0 = a0} b0 b1 p i + = comp (λ k → B (p (i ∧ k))) (λ k → λ where + (i = i0) → Bctr .snd b0 refl k + (i = i1) → Bctr .snd b1 p k) + (c0 .fst) + where + c0 = Bctr {a0} + +isPropDep→isSetDep : isOfHLevelDep 1 B → isOfHLevelDep 2 B +isPropDep→isSetDep {B = B} Bprp b0 b1 b2 b3 p i j + = comp (λ k → B (p (i ∧ k) (j ∧ k))) (λ k → λ where + (j = i0) → Bprp b0 b0 refl k + (i = i0) → Bprp b0 (b2 j) (λ k → p i0 (j ∧ k)) k + (i = i1) → Bprp b0 (b3 j) (λ k → p k (j ∧ k)) k + (j = i1) → Bprp b0 b1 (λ k → p (i ∧ k) (j ∧ k)) k) + b0 + +isOfHLevelDepSuc : (n : HLevel) → isOfHLevelDep n B → isOfHLevelDep (suc n) B +isOfHLevelDepSuc 0 = isContrDep→isPropDep +isOfHLevelDepSuc 1 = isPropDep→isSetDep +isOfHLevelDepSuc (suc (suc n)) Blvl b0 b1 = isOfHLevelDepSuc (suc n) (Blvl b0 b1) + +isPropDep→isSetDep' + : isOfHLevelDep 1 B + → {p : w ≡ x} {q : y ≡ z} {r : w ≡ y} {s : x ≡ z} + → {tw : B w} {tx : B x} {ty : B y} {tz : B z} + → (sq : Square p q r s) + → (tp : PathP (λ i → B (p i)) tw tx) + → (tq : PathP (λ i → B (q i)) ty tz) + → (tr : PathP (λ i → B (r i)) tw ty) + → (ts : PathP (λ i → B (s i)) tx tz) + → SquareP (λ i j → B (sq i j)) tp tq tr ts +isPropDep→isSetDep' {B = B} Bprp {p} {q} {r} {s} {tw} sq tp tq tr ts i j + = comp (λ k → B (sq (i ∧ k) (j ∧ k))) (λ k → λ where + (i = i0) → Bprp tw (tp j) (λ k → p (k ∧ j)) k + (i = i1) → Bprp tw (tq j) (λ k → sq (i ∧ k) (j ∧ k)) k + (j = i0) → Bprp tw (tr i) (λ k → r (k ∧ i)) k + (j = i1) → Bprp tw (ts i) (λ k → sq (k ∧ i) (j ∧ k)) k) + tw + +isOfHLevelΣ' : ∀ n → isOfHLevel n A → isOfHLevelDep n B → isOfHLevel n (Σ A B) +isOfHLevelΣ' 0 Actr Bctr .fst = (Actr .fst , Bctr .fst) +isOfHLevelΣ' 0 Actr Bctr .snd (x , y) i + = Actr .snd x i , Bctr .snd y (Actr .snd x) i +isOfHLevelΣ' 1 Alvl Blvl (w , y) (x , z) i .fst = Alvl w x i +isOfHLevelΣ' 1 Alvl Blvl (w , y) (x , z) i .snd = Blvl y z (Alvl w x) i +isOfHLevelΣ' {A = A} {B = B} (suc (suc n)) Alvl Blvl (w , y) (x , z) + = isOfHLevelRetract (suc n) + (λ p → (λ i → p i .fst) , λ i → p i .snd) + ΣPathP + (λ x → refl) + (isOfHLevelΣ' (suc n) (Alvl w x) (Blvl y z)) diff --git a/Cubical/Foundations/Id.agda b/Cubical/Foundations/Id.agda index 81bea8bbd..8b747ae7c 100644 --- a/Cubical/Foundations/Id.agda +++ b/Cubical/Foundations/Id.agda @@ -17,7 +17,7 @@ This file contains: - Propositional truncation and its elimination principle -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Id where open import Cubical.Foundations.Prelude public @@ -40,17 +40,18 @@ open import Cubical.Foundations.Equiv renaming ( fiber to fiberPath ; isEquiv to isEquivPath ; _≃_ to EquivPath - ; equivFun to equivFunPath ) - hiding ( isPropIsEquiv - ; equivCtr + ; equivFun to equivFunPath + ; isPropIsEquiv to isPropIsEquivPath ) + hiding ( equivCtr ; equivIsEquiv ) open import Cubical.Foundations.Univalence renaming ( EquivContr to EquivContrPath ) +open import Cubical.Foundations.Isomorphism open import Cubical.HITs.PropositionalTruncation public renaming ( squash to squashPath - ; recPropTrunc to recPropTruncPath - ; elimPropTrunc to elimPropTruncPath ) + ; rec to recPropTruncPath + ; elim to elimPropTruncPath ) open import Cubical.Core.Id public private @@ -205,14 +206,14 @@ isPropToIsPropPath H x y i = idToPath (H x y) i -- isContrPath and isContr: helper1 : ∀ {A B : Type ℓ} (f : A → B) (g : B → A) - (h : ∀ y → Path _ (f (g y)) y) → isContrPath A → isContr B + (h : (y : B) → Path B (f (g y)) y) → isContrPath A → isContr B helper1 f g h (x , p) = (f x , λ y → pathToId (λ i → hcomp (λ j → λ { (i = i0) → f x ; (i = i1) → h y j }) (f (p (g y) i)))) helper2 : ∀ {A B : Type ℓ} (f : A → B) (g : B → A) - (h : ∀ y → Path _ (g (f y)) y) → isContr B → isContrPath A + (h : (y : A) → Path A (g (f y)) y) → isContr B → isContrPath A helper2 {A = A} f g h (x , p) = (g x , λ y → idToPath (rem y)) where rem : ∀ (y : A) → g x ≡ y @@ -259,17 +260,16 @@ equivToEquiv (f , p) i = -- We can finally prove univalence with Id everywhere from the one for Path EquivContr : ∀ (A : Type ℓ) → isContr (Σ[ T ∈ Type ℓ ] (T ≃ A)) -EquivContr A = helper1 f1 f2 f12 (EquivContrPath A) +EquivContr {ℓ = ℓ} A = helper1 f1 f2 f12 (EquivContrPath A) where - f1 : ∀ {ℓ} {A : Type ℓ} → Σ[ T ∈ Type ℓ ] (EquivPath T A) → Σ[ T ∈ Type ℓ ] (T ≃ A) + f1 : {A : Type ℓ} → Σ[ T ∈ Type ℓ ] (EquivPath T A) → Σ[ T ∈ Type ℓ ] (T ≃ A) f1 (x , p) = x , equivPathToEquiv p - f2 : ∀ {ℓ} {A : Type ℓ} → Σ[ T ∈ Type ℓ ] (T ≃ A) → Σ[ T ∈ Type ℓ ] (EquivPath T A) + f2 : {A : Type ℓ} → Σ[ T ∈ Type ℓ ] (T ≃ A) → Σ[ T ∈ Type ℓ ] (EquivPath T A) f2 (x , p) = x , equivToEquivPath p - f12 : ∀ {ℓ} {A : Type ℓ} → (y : Σ[ T ∈ Type ℓ ] (T ≃ A)) → Path _ (f1 (f2 y)) y - f12 (x , p) = λ i → x , equivToEquiv p i - + f12 : (y : Σ[ T ∈ Type ℓ ] (T ≃ A)) → Path (Σ[ T ∈ Type ℓ ] (T ≃ A)) (f1 (f2 y)) y + f12 (x , p) i = x , equivToEquiv {A = x} {B = A} p i -- Propositional truncation @@ -282,3 +282,31 @@ EquivContr A = helper1 f1 f2 f12 (EquivContrPath A) ∥∥-induction : ∀ {A : Type ℓ} {P : ∥ A ∥ → Type ℓ} → ((a : ∥ A ∥) → isProp (P a)) → ((x : A) → P ∣ x ∣) → (a : ∥ A ∥) → P a ∥∥-induction Pprop f x = elimPropTruncPath (λ a → isPropToIsPropPath (Pprop a)) f x + + +-- Univalence + +path≡Id : ∀ {ℓ} {A B : Type ℓ} → Path _ (Path _ A B) (Id A B) +path≡Id = isoToPath (iso pathToId idToPath idToPathToId pathToIdToPath ) + +equivPathToEquivPath : ∀ {ℓ} {A : Type ℓ} {B : Type ℓ} → (p : EquivPath A B) → + Path _ (equivToEquivPath (equivPathToEquiv p)) p +equivPathToEquivPath (f , p) i = + ( f , isPropIsEquivPath f (equivToEquivPath (equivPathToEquiv (f , p)) .pr₂) p i ) + +equivPath≡Equiv : ∀ {ℓ} {A B : Type ℓ} → Path _ (EquivPath A B) (A ≃ B) +equivPath≡Equiv {ℓ} = isoToPath (iso (equivPathToEquiv {ℓ}) equivToEquivPath equivToEquiv equivPathToEquivPath) + +univalenceId : ∀ {ℓ} {A B : Type ℓ} → (A ≡ B) ≃ (A ≃ B) +univalenceId {ℓ} {A = A} {B = B} = equivPathToEquiv rem + where + rem0 : Path _ (Lift (EquivPath A B)) (Lift (A ≃ B)) + rem0 = congPath Lift equivPath≡Equiv + + rem1 : Path _ (Id A B) (Lift (A ≃ B)) + rem1 i = hcomp (λ j → λ { (i = i0) → path≡Id {A = A} {B = B} j + ; (i = i1) → rem0 j }) + (univalencePath {A = A} {B = B} i) + + rem : EquivPath (Id A B) (A ≃ B) + rem = compEquiv (eqweqmap rem1) (invEquiv LiftEquiv) diff --git a/Cubical/Foundations/Isomorphism.agda b/Cubical/Foundations/Isomorphism.agda index 9654c3acf..1746086b1 100644 --- a/Cubical/Foundations/Isomorphism.agda +++ b/Cubical/Foundations/Isomorphism.agda @@ -7,17 +7,21 @@ Theory about isomorphisms - Any isomorphism is an equivalence ([isoToEquiv]) -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Isomorphism where open import Cubical.Core.Everything open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv.Base + open import Cubical.Foundations.Function private variable - ℓ : Level + ℓ ℓ' : Level + A B C : Type ℓ -- Section and retract module _ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} where @@ -29,15 +33,25 @@ module _ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} where retract f g = ∀ a → g (f a) ≡ a record Iso {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') : Type (ℓ-max ℓ ℓ') where + no-eta-equality constructor iso field fun : A → B inv : B → A rightInv : section fun inv - leftInv : retract fun inv + leftInv : retract fun inv + +isIso : (A → B) → Type _ +isIso {A = A} {B = B} f = Σ[ g ∈ (B → A) ] Σ[ _ ∈ section f g ] retract f g + +isoFunInjective : (f : Iso A B) → (x y : A) → Iso.fun f x ≡ Iso.fun f y → x ≡ y +isoFunInjective f x y h = sym (Iso.leftInv f x) ∙∙ cong (Iso.inv f) h ∙∙ Iso.leftInv f y + +isoInvInjective : (f : Iso A B) → (x y : B) → Iso.inv f x ≡ Iso.inv f y → x ≡ y +isoInvInjective f x y h = sym (Iso.rightInv f x) ∙∙ cong (Iso.fun f) h ∙∙ Iso.rightInv f y -- Any iso is an equivalence -module _ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (i : Iso A B) where +module _ (i : Iso A B) where open Iso i renaming ( fun to f ; inv to g ; rightInv to s @@ -87,19 +101,80 @@ module _ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (i : Iso A B) where isoToIsEquiv .equiv-proof y .snd z = lemIso y (g y) (fst z) (s y) (snd z) -isoToEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso A B → A ≃ B -isoToEquiv i = _ , isoToIsEquiv i +isoToEquiv : Iso A B → A ≃ B +isoToEquiv i .fst = _ +isoToEquiv i .snd = isoToIsEquiv i -isoToPath : ∀ {ℓ} {A B : Type ℓ} → (Iso A B) → A ≡ B +isoToPath : Iso A B → A ≡ B isoToPath {A = A} {B = B} f i = - Glue B (λ { (i = i0) → (A , (Iso.fun f , isoToIsEquiv f)) - ; (i = i1) → (B , (λ x → x) , - record { equiv-proof = λ y → (y , refl) - , λ z i → z .snd (~ i) - , λ j → z .snd (~ i ∨ j)})}) - -compIso : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} - → Iso A B → Iso B C → Iso A C -compIso (iso fun inv rightInv leftInv) (iso fun₁ inv₁ rightInv₁ leftInv₁) = iso (fun₁ ∘ fun) (inv ∘ inv₁) - (λ b → cong fun₁ (rightInv (inv₁ b)) ∙ (rightInv₁ b)) - (λ a → cong inv (leftInv₁ (fun a) ) ∙ leftInv a ) + Glue B (λ { (i = i0) → (A , isoToEquiv f) + ; (i = i1) → (B , idEquiv B) }) + +open Iso + +invIso : Iso A B → Iso B A +fun (invIso f) = inv f +inv (invIso f) = fun f +rightInv (invIso f) = leftInv f +leftInv (invIso f) = rightInv f + +compIso : Iso A B → Iso B C → Iso A C +fun (compIso i j) = fun j ∘ fun i +inv (compIso i j) = inv i ∘ inv j +rightInv (compIso i j) b = cong (fun j) (rightInv i (inv j b)) ∙ rightInv j b +leftInv (compIso i j) a = cong (inv i) (leftInv j (fun i a)) ∙ leftInv i a + +composesToId→Iso : (G : Iso A B) (g : B → A) → G .fun ∘ g ≡ idfun B → Iso B A +fun (composesToId→Iso _ g _) = g +inv (composesToId→Iso j _ _) = fun j +rightInv (composesToId→Iso i g path) b = + sym (leftInv i (g (fun i b))) ∙∙ cong (λ g → inv i (g (fun i b))) path ∙∙ leftInv i b +leftInv (composesToId→Iso _ _ path) b i = path i b + +idIso : Iso A A +fun idIso = idfun _ +inv idIso = idfun _ +rightInv idIso _ = refl +leftInv idIso _ = refl + +LiftIso : Iso A (Lift {i = ℓ} {j = ℓ'} A) +fun LiftIso = lift +inv LiftIso = lower +rightInv LiftIso _ = refl +leftInv LiftIso _ = refl + +isContr→Iso : isContr A → isContr B → Iso A B +fun (isContr→Iso _ Bctr) _ = Bctr .fst +inv (isContr→Iso Actr _) _ = Actr .fst +rightInv (isContr→Iso _ Bctr) = Bctr .snd +leftInv (isContr→Iso Actr _) = Actr .snd + +isProp→Iso : (Aprop : isProp A) (Bprop : isProp B) (f : A → B) (g : B → A) → Iso A B +fun (isProp→Iso _ _ f _) = f +inv (isProp→Iso _ _ _ g) = g +rightInv (isProp→Iso _ Bprop f g) b = Bprop (f (g b)) b +leftInv (isProp→Iso Aprop _ f g) a = Aprop (g (f a)) a + +domIso : ∀ {ℓ} {C : Type ℓ} → Iso A B → Iso (A → C) (B → C) +fun (domIso e) f b = f (inv e b) +inv (domIso e) f a = f (fun e a) +rightInv (domIso e) f i x = f (rightInv e x i) +leftInv (domIso e) f i x = f (leftInv e x i) + +-- Helpful notation +_Iso⟨_⟩_ : ∀ {ℓ ℓ' ℓ''} {B : Type ℓ'} {C : Type ℓ''} (X : Type ℓ) → Iso X B → Iso B C → Iso X C +_ Iso⟨ f ⟩ g = compIso f g + +_∎Iso : ∀ {ℓ} (A : Type ℓ) → Iso A A +A ∎Iso = idIso {A = A} + +infixr 0 _Iso⟨_⟩_ +infix 1 _∎Iso + +codomainIso : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + → Iso B C + → Iso (A → B) (A → C) +Iso.fun (codomainIso is) f a = Iso.fun is (f a) +Iso.inv (codomainIso is) f a = Iso.inv is (f a) +Iso.rightInv (codomainIso is) f = funExt λ a → Iso.rightInv is (f a) +Iso.leftInv (codomainIso is) f = funExt λ a → Iso.leftInv is (f a) diff --git a/Cubical/Foundations/Path.agda b/Cubical/Foundations/Path.agda index cb9aaa067..29b7934f9 100644 --- a/Cubical/Foundations/Path.agda +++ b/Cubical/Foundations/Path.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Path where open import Cubical.Foundations.Prelude @@ -6,6 +6,9 @@ open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence + +open import Cubical.Reflection.StrictEquiv private variable @@ -16,64 +19,67 @@ private cong′ : ∀ {B : Type ℓ'} (f : A → B) {x y : A} (p : x ≡ y) → Path B (f x) (f y) cong′ f = cong f +{-# INLINE cong′ #-} PathP≡Path : ∀ (P : I → Type ℓ) (p : P i0) (q : P i1) → PathP P p q ≡ Path (P i1) (transport (λ i → P i) p) q -PathP≡Path P p q i = PathP (λ j → P (i ∨ j)) (transp (λ j → P (i ∧ j)) (~ i) p) q - -PathP≃Path : ∀ (P : I → Type ℓ) (p : P i0) (q : P i1) → - PathP P p q ≃ Path (P i1) (transport (λ i → P i) p) q -PathP≃Path P p q = transportEquiv (PathP≡Path P p q) - --- Alternative more unfolded proof -toPathP-isEquiv : ∀ (A : I → Set ℓ) {x y} → isEquiv (toPathP {A = A} {x} {y}) -toPathP-isEquiv A {x} {y} = isoToIsEquiv (iso toPathP fromPathP to-from from-to) - where - to-from : ∀ (p : PathP A x y) → toPathP (fromPathP p) ≡ p - to-from p h i = outS (hcomp-unique (λ { j (i = i0) → x ; j (i = i1) → fromPathP p j }) - (inS (transp (λ j → A (i ∧ j)) (~ i) x)) - \ h → inS (sq1 h i)) - h - where - sq1 : (\ h → A [ x ≡ transp (\ j → A (h ∨ j)) h (p h) ]) [ (\ i → transp (λ j → A (i ∧ j)) (~ i) x) ≡ p ] - sq1 = \ h i → comp (\ z → (hcomp (\ w → - \ { (z = i1) → A (i ∧ (w ∨ h)) - ; (z = i0) → A (i ∧ h) - ; (i = i0) → A i0 - ; (i = i1) → A (h ∨ (w ∧ z)) - ; (h = i0) → A (i ∧ (w ∧ z)) - ; (h = i1) → A i}) - ((A (i ∧ h))))) - (\ z → \ { (i = i0) → x - ; (i = i1) → transp (\ j → A (h ∨ (z ∧ j))) (h ∨ ~ z) (p h) - ; (h = i0) → transp (λ j → A ((i ∧ z) ∧ j)) (~ (i ∧ z)) x - ; (h = i1) → p i }) - (p (i ∧ h)) - from-to : ∀ (q : transp A i0 x ≡ y) → fromPathP (toPathP {A = A} q) ≡ q - from-to q = (\ h i → outS (transp-hcomp i {A' = A i1} (\ j → inS (A (i ∨ j))) - ((λ { j (i = i0) → x ; j (i = i1) → q j })) - (inS ((transp (λ j → A (i ∧ j)) (~ i) x)))) - h) - ∙ (\ h i → outS (hcomp-unique {A = A i1} ((λ { j (i = i0) → transp A i0 x ; j (i = i1) → q j })) - (inS ((transp (λ j → A (i ∨ j)) i (transp (λ j → A (i ∧ j)) (~ i) x)))) - \ h → inS (sq2 h i)) - h) - ∙ sym (lUnit q) - where - sq2 : (\ h → transp A i0 x ≡ q h) [ (\ i → transp (\ j → A (i ∨ j)) i (transp (\ j → A (i ∧ j)) (~ i) x)) ≡ refl ∙ q ] - sq2 = \ h i → comp (\ z → hcomp (\ w → \ { (i = i1) → A i1 - ; (i = i0) → A (h ∨ (w ∧ z)) - ; (h = i0) → A (i ∨ (w ∧ z)) - ; (h = i1) → A i1 - ; (z = i0) → A (i ∨ h) - ; (z = i1) → A ((i ∨ h) ∨ w) }) - (A (i ∨ h))) - (\ z → \ { (i = i0) → transp (λ j → A ((z ∨ h) ∧ j)) (~ z ∧ ~ h) x - ; (i = i1) → q (z ∧ h) - ; (h = i1) → compPath-filler refl q z i - ; (h = i0) → transp (\ j → A (i ∨ (z ∧ j))) (i ∨ ~ z) (transp (\ j → A (i ∧ j)) (~ i) x) - }) - (transp (\ j → A ((i ∨ h) ∧ j)) (~ (i ∨ h)) x) +PathP≡Path P p q i = PathP (λ j → P (i ∨ j)) (transport-filler (λ j → P j) p i) q + +PathP≡Path⁻ : ∀ (P : I → Type ℓ) (p : P i0) (q : P i1) → + PathP P p q ≡ Path (P i0) p (transport⁻ (λ i → P i) q) +PathP≡Path⁻ P p q i = PathP (λ j → P (~ i ∧ j)) p (transport⁻-filler (λ j → P j) q i) + +PathPIsoPath : ∀ (A : I → Type ℓ) (x : A i0) (y : A i1) → Iso (PathP A x y) (transport (λ i → A i) x ≡ y) +PathPIsoPath A x y .Iso.fun = fromPathP +PathPIsoPath A x y .Iso.inv = toPathP +PathPIsoPath A x y .Iso.rightInv q k i = + hcomp + (λ j → λ + { (i = i0) → slide (j ∨ ~ k) + ; (i = i1) → q j + ; (k = i0) → transp (λ l → A (i ∨ l)) i (fromPathPFiller j) + ; (k = i1) → ∧∨Square i j + }) + (transp (λ l → A (i ∨ ~ k ∨ l)) (i ∨ ~ k) + (transp (λ l → (A (i ∨ (~ k ∧ l)))) (k ∨ i) + (transp (λ l → A (i ∧ l)) (~ i) + x))) + where + fromPathPFiller : _ + fromPathPFiller = + hfill + (λ j → λ + { (i = i0) → x + ; (i = i1) → q j }) + (inS (transp (λ j → A (i ∧ j)) (~ i) x)) + + slide : I → _ + slide i = transp (λ l → A (i ∨ l)) i (transp (λ l → A (i ∧ l)) (~ i) x) + + ∧∨Square : I → I → _ + ∧∨Square i j = + hcomp + (λ l → λ + { (i = i0) → slide j + ; (i = i1) → q (j ∧ l) + ; (j = i0) → slide i + ; (j = i1) → q (i ∧ l) + }) + (slide (i ∨ j)) +PathPIsoPath A x y .Iso.leftInv q k i = + outS + (hcomp-unique + (λ j → λ + { (i = i0) → x + ; (i = i1) → transp (λ l → A (j ∨ l)) j (q j) + }) + (inS (transp (λ l → A (i ∧ l)) (~ i) x)) + (λ j → inS (transp (λ l → A (i ∧ (j ∨ l))) (~ i ∨ j) (q (i ∧ j))))) + k + +PathP≃Path : (A : I → Type ℓ) (x : A i0) (y : A i1) → + PathP A x y ≃ (transport (λ i → A i) x ≡ y) +PathP≃Path A x y = isoToEquiv (PathPIsoPath A x y) PathP≡compPath : ∀ {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) (r : x ≡ z) → (PathP (λ i → x ≡ q i) p r) ≡ (p ∙ q ≡ r) @@ -82,12 +88,12 @@ PathP≡compPath p q r k = PathP (λ i → p i0 ≡ q (i ∨ k)) (λ j → compP PathP≡doubleCompPathˡ : ∀ {A : Type ℓ} {w x y z : A} (p : w ≡ y) (q : w ≡ x) (r : y ≡ z) (s : x ≡ z) → (PathP (λ i → p i ≡ s i) q r) ≡ (p ⁻¹ ∙∙ q ∙∙ s ≡ r) PathP≡doubleCompPathˡ p q r s k = PathP (λ i → p (i ∨ k) ≡ s (i ∨ k)) - (λ j → doubleCompPath-filler (p ⁻¹) q s j k) r + (λ j → doubleCompPath-filler (p ⁻¹) q s k j) r PathP≡doubleCompPathʳ : ∀ {A : Type ℓ} {w x y z : A} (p : w ≡ y) (q : w ≡ x) (r : y ≡ z) (s : x ≡ z) → (PathP (λ i → p i ≡ s i) q r) ≡ (q ≡ p ∙∙ r ∙∙ s ⁻¹) PathP≡doubleCompPathʳ p q r s k = PathP (λ i → p (i ∧ (~ k)) ≡ s (i ∧ (~ k))) - q (λ j → doubleCompPath-filler p r (s ⁻¹) j k) + q (λ j → doubleCompPath-filler p r (s ⁻¹) k j) compPathl-cancel : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : x ≡ z) → p ∙ (sym p ∙ q) ≡ q compPathl-cancel p q = p ∙ (sym p ∙ q) ≡⟨ assoc p (sym p) q ⟩ @@ -96,20 +102,209 @@ compPathl-cancel p q = p ∙ (sym p ∙ q) ≡⟨ assoc p (sym p) q ⟩ q ∎ compPathr-cancel : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : z ≡ y) (q : x ≡ y) → (q ∙ sym p) ∙ p ≡ q -compPathr-cancel p q = (q ∙ sym p) ∙ p ≡⟨ sym (assoc q (sym p) p) ⟩ - q ∙ (sym p ∙ p) ≡⟨ cong (q ∙_) (lCancel p) ⟩ - q ∙ refl ≡⟨ sym (rUnit q) ⟩ - q ∎ +compPathr-cancel {x = x} p q i j = + hcomp-equivFiller (doubleComp-faces (λ _ → x) (sym p) j) (inS (q j)) (~ i) -compPathl-isEquiv : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) → isEquiv (λ (q : y ≡ z) → p ∙ q) +compPathl-isEquiv : {x y z : A} (p : x ≡ y) → isEquiv (λ (q : y ≡ z) → p ∙ q) compPathl-isEquiv p = isoToIsEquiv (iso (p ∙_) (sym p ∙_) (compPathl-cancel p) (compPathl-cancel (sym p))) -compPathr-isEquiv : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : y ≡ z) → isEquiv (λ (q : x ≡ y) → q ∙ p) +compPathlEquiv : {x y z : A} (p : x ≡ y) → (y ≡ z) ≃ (x ≡ z) +compPathlEquiv p = (p ∙_) , compPathl-isEquiv p + +compPathr-isEquiv : {x y z : A} (p : y ≡ z) → isEquiv (λ (q : x ≡ y) → q ∙ p) compPathr-isEquiv p = isoToIsEquiv (iso (_∙ p) (_∙ sym p) (compPathr-cancel p) (compPathr-cancel (sym p))) --- Variation of isProp→isSet for PathP -isProp→isSet-PathP : ∀ {ℓ} {B : I → Type ℓ} → ((i : I) → isProp (B i)) +compPathrEquiv : {x y z : A} (p : y ≡ z) → (x ≡ y) ≃ (x ≡ z) +compPathrEquiv p = (_∙ p) , compPathr-isEquiv p + +-- Variations of isProp→isSet for PathP +isProp→SquareP : ∀ {B : I → I → Type ℓ} → ((i j : I) → isProp (B i j)) + → {a : B i0 i0} {b : B i0 i1} {c : B i1 i0} {d : B i1 i1} + → (r : PathP (λ j → B j i0) a c) (s : PathP (λ j → B j i1) b d) + → (t : PathP (λ j → B i0 j) a b) (u : PathP (λ j → B i1 j) c d) + → SquareP B t u r s +isProp→SquareP {B = B} isPropB {a = a} r s t u i j = + hcomp (λ { k (i = i0) → isPropB i0 j (base i0 j) (t j) k + ; k (i = i1) → isPropB i1 j (base i1 j) (u j) k + ; k (j = i0) → isPropB i i0 (base i i0) (r i) k + ; k (j = i1) → isPropB i i1 (base i i1) (s i) k + }) (base i j) where + base : (i j : I) → B i j + base i j = transport (λ k → B (i ∧ k) (j ∧ k)) a + +isProp→isPropPathP : ∀ {ℓ} {B : I → Type ℓ} → ((i : I) → isProp (B i)) → (b0 : B i0) (b1 : B i1) → isProp (PathP (λ i → B i) b0 b1) -isProp→isSet-PathP {B = B} hB b0 b1 = - transport (λ i → isProp (PathP≡Path B b0 b1 (~ i))) (isProp→isSet (hB i1) _ _) +isProp→isPropPathP {B = B} hB b0 b1 = isProp→SquareP (λ _ → hB) refl refl + +isProp→isContrPathP : {A : I → Type ℓ} → (∀ i → isProp (A i)) + → (x : A i0) (y : A i1) + → isContr (PathP A x y) +isProp→isContrPathP h x y = isProp→PathP h x y , isProp→isPropPathP h x y _ + +-- Flipping a square along its diagonal + +flipSquare : {a₀₀ a₀₁ : A} {a₀₋ : a₀₀ ≡ a₀₁} + {a₁₀ a₁₁ : A} {a₁₋ : a₁₀ ≡ a₁₁} + {a₋₀ : a₀₀ ≡ a₁₀} {a₋₁ : a₀₁ ≡ a₁₁} + → Square a₀₋ a₁₋ a₋₀ a₋₁ → Square a₋₀ a₋₁ a₀₋ a₁₋ +flipSquare sq i j = sq j i + +module _ {a₀₀ a₀₁ : A} {a₀₋ : a₀₀ ≡ a₀₁} {a₁₀ a₁₁ : A} {a₁₋ : a₁₀ ≡ a₁₁} + {a₋₀ : a₀₀ ≡ a₁₀} {a₋₁ : a₀₁ ≡ a₁₁} + where + + flipSquareEquiv : Square a₀₋ a₁₋ a₋₀ a₋₁ ≃ Square a₋₀ a₋₁ a₀₋ a₁₋ + unquoteDef flipSquareEquiv = defStrictEquiv flipSquareEquiv flipSquare flipSquare + + flipSquarePath : Square a₀₋ a₁₋ a₋₀ a₋₁ ≡ Square a₋₀ a₋₁ a₀₋ a₁₋ + flipSquarePath = ua flipSquareEquiv + +module _ {a₀₀ a₁₁ : A} {a₋ : a₀₀ ≡ a₁₁} + {a₁₀ : A} {a₁₋ : a₁₀ ≡ a₁₁} {a₋₀ : a₀₀ ≡ a₁₀} where + + slideSquareFaces : (i j k : I) → Partial (i ∨ ~ i ∨ j ∨ ~ j) A + slideSquareFaces i j k (i = i0) = a₋ (j ∧ ~ k) + slideSquareFaces i j k (i = i1) = a₁₋ j + slideSquareFaces i j k (j = i0) = a₋₀ i + slideSquareFaces i j k (j = i1) = a₋ (i ∨ ~ k) + + slideSquare : Square a₋ a₁₋ a₋₀ refl → Square refl a₁₋ a₋₀ a₋ + slideSquare sq i j = hcomp (slideSquareFaces i j) (sq i j) + + slideSquareEquiv : (Square a₋ a₁₋ a₋₀ refl) ≃ (Square refl a₁₋ a₋₀ a₋) + slideSquareEquiv = isoToEquiv (iso slideSquare slideSquareInv fillerTo fillerFrom) where + slideSquareInv : Square refl a₁₋ a₋₀ a₋ → Square a₋ a₁₋ a₋₀ refl + slideSquareInv sq i j = hcomp (λ k → slideSquareFaces i j (~ k)) (sq i j) + fillerTo : ∀ p → slideSquare (slideSquareInv p) ≡ p + fillerTo p k i j = hcomp-equivFiller (λ k → slideSquareFaces i j (~ k)) (inS (p i j)) (~ k) + fillerFrom : ∀ p → slideSquareInv (slideSquare p) ≡ p + fillerFrom p k i j = hcomp-equivFiller (slideSquareFaces i j) (inS (p i j)) (~ k) + +-- The type of fillers of a square is equivalent to the double composition identites +Square≃doubleComp : {a₀₀ a₀₁ a₁₀ a₁₁ : A} + (a₀₋ : a₀₀ ≡ a₀₁) (a₁₋ : a₁₀ ≡ a₁₁) + (a₋₀ : a₀₀ ≡ a₁₀) (a₋₁ : a₀₁ ≡ a₁₁) + → Square a₀₋ a₁₋ a₋₀ a₋₁ ≃ (a₋₀ ⁻¹ ∙∙ a₀₋ ∙∙ a₋₁ ≡ a₁₋) +Square≃doubleComp a₀₋ a₁₋ a₋₀ a₋₁ = transportEquiv (PathP≡doubleCompPathˡ a₋₀ a₀₋ a₁₋ a₋₁) + +-- Flipping a square in Ω²A is the same as inverting it +sym≡flipSquare : {x : A} (P : Square (refl {x = x}) refl refl refl) + → sym P ≡ flipSquare P +sym≡flipSquare {x = x} = helper x x refl refl + where + helper : ∀ {ℓ} {A : Type ℓ} (x y : A) (p q : x ≡ y) (P : Square p q refl refl) + → PathP (λ j → PathP (λ i → Path A (p (i ∧ j)) (q (i ∨ ~ j))) + (λ k → q (~ j ∧ k)) + λ k → p (j ∨ k)) + (sym P) + (flipSquare P) + helper {A = A} x y = + J (λ y p → (q : x ≡ y) → (P : Square p q refl refl) → + PathP (λ j → PathP (λ i → Path A (p (i ∧ j)) (q (i ∨ ~ j))) + (λ k → q (~ j ∧ k)) + (λ k → p (j ∨ k))) + (sym P) + (flipSquare P)) + λ q → J (λ q P → PathP (λ j → PathP (λ i → Path A x (q (i ∨ ~ j))) + (λ k → q (~ j ∧ k)) refl) + (λ i → P (~ i)) λ i j → P j i) refl + +-- Inverting both interval arguments of a square in Ω²A is the same as doing nothing +sym-cong-sym≡id : {x : A} (P : Square (refl {x = x}) refl refl refl) + → P ≡ λ i j → P (~ i) (~ j) +sym-cong-sym≡id P = + transport (λ i → doubleCompPath-filler (sym (rUnit refl)) P (lUnit refl) (~ i) + ≡ doubleCompPath-filler (sym (rUnit refl)) + (λ i j → P (~ i) (~ j)) (lUnit refl) (~ i)) (helper _ _ refl refl P) + where + helper : ∀ {ℓ} {A : Type ℓ} (x y : A) → (p q : x ≡ y) (P : Square p q refl refl) + → PathP (λ i → PathP (λ j → p i ≡ q (~ i)) + ((λ j → p (i ∨ j)) ∙ (λ j → q (~ i ∨ ~ j))) + ((λ j → p (i ∧ ~ j)) ∙ (λ j → q (~ i ∧ j)))) + (sym (rUnit p) ∙∙ P ∙∙ lUnit _) + (sym (lUnit (sym q)) ∙∙ (λ i j → P (~ i) (~ j)) ∙∙ rUnit (sym p)) + helper x y = + J (λ y p → (q : x ≡ y) (P : Square p q refl refl) + → PathP (λ i → PathP (λ j → p i ≡ q (~ i)) ((λ j → p (i ∨ j)) ∙ (λ j → q (~ i ∨ ~ j))) + ((λ j → p (i ∧ ~ j)) ∙ (λ j → q (~ i ∧ j)))) + (sym (rUnit p) ∙∙ P ∙∙ lUnit _) + (sym (lUnit (sym q)) ∙∙ (λ i j → P (~ i) (~ j)) ∙∙ rUnit (sym p))) + λ q → + J (λ q P → PathP (λ i → (λ j → x) ∙ (λ j → q (~ i ∨ ~ j)) ≡ + (λ j → x) ∙ (λ j → q (~ i ∧ j))) + ((λ i → rUnit refl (~ i)) ∙∙ P ∙∙ lUnit q) + ((λ i → lUnit (λ i₁ → q (~ i₁)) (~ i)) ∙∙ (λ i j → P (~ i) (~ j)) ∙∙ rUnit refl)) + refl + +-- Applying cong sym is the same as flipping a square in Ω²A +flipSquare≡cong-sym : ∀ {ℓ} {A : Type ℓ} {x : A} (P : Square (refl {x = x}) refl refl refl) + → flipSquare P ≡ λ i j → P i (~ j) +flipSquare≡cong-sym P = sym (sym≡flipSquare P) ∙ sym (sym-cong-sym≡id (cong sym P)) + +-- Applying cong sym is the same as inverting a square in Ω²A +sym≡cong-sym : ∀ {ℓ} {A : Type ℓ} {x : A} (P : Square (refl {x = x}) refl refl refl) + → sym P ≡ cong sym P +sym≡cong-sym P = sym≡flipSquare _ ∙ flipSquare≡cong-sym P + +-- sym induces an equivalence on identity types of paths +symIso : {a b : A} (p q : a ≡ b) → Iso (p ≡ q) (q ≡ p) +symIso p q = iso sym sym (λ _ → refl) λ _ → refl + +-- J is an equivalence +Jequiv : {x : A} (P : ∀ y → x ≡ y → Type ℓ') → P x refl ≃ (∀ {y} (p : x ≡ y) → P y p) +Jequiv P = isoToEquiv isom + where + isom : Iso _ _ + Iso.fun isom = J P + Iso.inv isom f = f refl + Iso.rightInv isom f = + implicitFunExt λ {_} → + funExt λ t → + J (λ _ t → J P (f refl) t ≡ f t) (JRefl P (f refl)) t + Iso.leftInv isom = JRefl P + +-- Action of PathP on equivalences (without relying on univalence) + +congPathIso : ∀ {ℓ ℓ'} {A : I → Type ℓ} {B : I → Type ℓ'} + (e : ∀ i → A i ≃ B i) {a₀ : A i0} {a₁ : A i1} + → Iso (PathP A a₀ a₁) (PathP B (e i0 .fst a₀) (e i1 .fst a₁)) +congPathIso {A = A} {B} e {a₀} {a₁} .Iso.fun p i = e i .fst (p i) +congPathIso {A = A} {B} e {a₀} {a₁} .Iso.inv q i = + hcomp + (λ j → λ + { (i = i0) → retEq (e i0) a₀ j + ; (i = i1) → retEq (e i1) a₁ j + }) + (invEq (e i) (q i)) +congPathIso {A = A} {B} e {a₀} {a₁} .Iso.rightInv q k i = + hcomp + (λ j → λ + { (i = i0) → commSqIsEq (e i0 .snd) a₀ j k + ; (i = i1) → commSqIsEq (e i1 .snd) a₁ j k + ; (k = i0) → + e i .fst + (hfill + (λ j → λ + { (i = i0) → retEq (e i0) a₀ j + ; (i = i1) → retEq (e i1) a₁ j + }) + (inS (invEq (e i) (q i))) + j) + ; (k = i1) → q i + }) + (secEq (e i) (q i) k) + where b = commSqIsEq +congPathIso {A = A} {B} e {a₀} {a₁} .Iso.leftInv p k i = + hcomp + (λ j → λ + { (i = i0) → retEq (e i0) a₀ (j ∨ k) + ; (i = i1) → retEq (e i1) a₁ (j ∨ k) + ; (k = i1) → p i + }) + (retEq (e i) (p i) k) + +congPathEquiv : ∀ {ℓ ℓ'} {A : I → Type ℓ} {B : I → Type ℓ'} + (e : ∀ i → A i ≃ B i) {a₀ : A i0} {a₁ : A i1} + → PathP A a₀ a₁ ≃ PathP B (e i0 .fst a₀) (e i1 .fst a₁) +congPathEquiv e = isoToEquiv (congPathIso e) diff --git a/Cubical/Foundations/Pointed.agda b/Cubical/Foundations/Pointed.agda index 3d79f2765..f81c75846 100644 --- a/Cubical/Foundations/Pointed.agda +++ b/Cubical/Foundations/Pointed.agda @@ -1,7 +1,9 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Pointed where open import Cubical.Foundations.Pointed.Base public open import Cubical.Foundations.Pointed.Properties public +open import Cubical.Foundations.Pointed.FunExt public +open import Cubical.Foundations.Pointed.Homotopy public open import Cubical.Foundations.Pointed.Homogeneous diff --git a/Cubical/Foundations/Pointed/Base.agda b/Cubical/Foundations/Pointed/Base.agda index d114a4a55..5ed40cc0c 100644 --- a/Cubical/Foundations/Pointed/Base.agda +++ b/Cubical/Foundations/Pointed/Base.agda @@ -1,22 +1,77 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Pointed.Base where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Structure using (typ) public +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence Pointed : (ℓ : Level) → Type (ℓ-suc ℓ) -Pointed ℓ = Σ[ A ∈ Type ℓ ] A +Pointed ℓ = TypeWithStr ℓ (λ x → x) + +pt : ∀ {ℓ} (A∙ : Pointed ℓ) → typ A∙ +pt = str Pointed₀ = Pointed ℓ-zero -typ : ∀ {ℓ} (A∙ : Pointed ℓ) → Type ℓ -typ = fst +{- Pointed functions -} +_→∙_ : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → Type (ℓ-max ℓ ℓ') +(A , a) →∙ (B , b) = Σ[ f ∈ (A → B) ] f a ≡ b -pt : ∀ {ℓ} (A∙ : Pointed ℓ) → typ A∙ -pt = snd +_→∙_∙ : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +(A →∙ B ∙) .fst = A →∙ B +(A →∙ B ∙) .snd .fst x = pt B +(A →∙ B ∙) .snd .snd = refl -{- Pointed functions -} -_→*_ : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → Type (ℓ-max ℓ ℓ') -_→*_ A B = Σ (typ A → typ B) λ f → f (pt A) ≡ pt B +idfun∙ : ∀ {ℓ} (A : Pointed ℓ) → A →∙ A +idfun∙ A .fst x = x +idfun∙ A .snd = refl + +ua∙ : ∀ {ℓ} {A B : Pointed ℓ} (e : fst A ≃ fst B) + → fst e (snd A) ≡ snd B → A ≡ B +fst (ua∙ e p i) = ua e i +snd (ua∙ {A = A} e p i) = ua-gluePath e p i + +{- HIT allowing for pattern matching on pointed types -} +data Pointer {ℓ} (A : Pointed ℓ) : Type ℓ where + pt₀ : Pointer A + ⌊_⌋ : typ A → Pointer A + id : ⌊ pt A ⌋ ≡ pt₀ + +IsoPointedPointer : ∀ {ℓ} {A : Pointed ℓ} → Iso (typ A) (Pointer A) +Iso.fun IsoPointedPointer = ⌊_⌋ +Iso.inv (IsoPointedPointer {A = A}) pt₀ = pt A +Iso.inv IsoPointedPointer ⌊ x ⌋ = x +Iso.inv (IsoPointedPointer {A = A}) (id i) = pt A +Iso.rightInv IsoPointedPointer pt₀ = id +Iso.rightInv IsoPointedPointer ⌊ x ⌋ = refl +Iso.rightInv IsoPointedPointer (id i) j = id (i ∧ j) +Iso.leftInv IsoPointedPointer x = refl + +Pointed≡Pointer : ∀ {ℓ} {A : Pointed ℓ} → typ A ≡ Pointer A +Pointed≡Pointer = isoToPath IsoPointedPointer + +Pointer∙ : ∀ {ℓ} (A : Pointed ℓ) → Pointed ℓ +Pointer∙ A .fst = Pointer A +Pointer∙ A .snd = pt₀ + +Pointed≡∙Pointer : ∀ {ℓ} {A : Pointed ℓ} → A ≡ (Pointer A , pt₀) +Pointed≡∙Pointer {A = A} i = (Pointed≡Pointer {A = A} i) , helper i + where + helper : PathP (λ i → Pointed≡Pointer {A = A} i) (pt A) pt₀ + helper = ua-gluePath (isoToEquiv (IsoPointedPointer {A = A})) id + +pointerFun : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → Pointer A → Pointer B +pointerFun f pt₀ = pt₀ +pointerFun f ⌊ x ⌋ = ⌊ fst f x ⌋ +pointerFun f (id i) = (cong ⌊_⌋ (snd f) ∙ id) i -_→*_* : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') -A →* B * = (A →* B) , (λ x → pt B) , refl +pointerFun∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) + → Pointer∙ A →∙ Pointer∙ B +pointerFun∙ f .fst = pointerFun f +pointerFun∙ f .snd = refl diff --git a/Cubical/Foundations/Pointed/FunExt.agda b/Cubical/Foundations/Pointed/FunExt.agda new file mode 100644 index 000000000..69530642c --- /dev/null +++ b/Cubical/Foundations/Pointed/FunExt.agda @@ -0,0 +1,48 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Pointed.FunExt where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv + +open import Cubical.Foundations.Pointed.Base +open import Cubical.Foundations.Pointed.Properties +open import Cubical.Foundations.Pointed.Homotopy + +private + variable + ℓ ℓ' : Level + +module _ {A : Pointed ℓ} {B : typ A → Type ℓ'} {ptB : B (pt A)} where + + -- pointed function extensionality + funExt∙P : {f g : Π∙ A B ptB} → f ∙∼P g → f ≡ g + funExt∙P (h , h∙) i .fst x = h x i + funExt∙P (h , h∙) i .snd = h∙ i + + -- inverse of pointed function extensionality + funExt∙P⁻ : {f g : Π∙ A B ptB} → f ≡ g → f ∙∼P g + funExt∙P⁻ p .fst a i = p i .fst a + funExt∙P⁻ p .snd i = p i .snd + + -- function extensionality is an isomorphism, PathP version + funExt∙PIso : (f g : Π∙ A B ptB) → Iso (f ∙∼P g) (f ≡ g) + Iso.fun (funExt∙PIso f g) = funExt∙P {f = f} {g = g} + Iso.inv (funExt∙PIso f g) = funExt∙P⁻ {f = f} {g = g} + Iso.rightInv (funExt∙PIso f g) p i j = p j + Iso.leftInv (funExt∙PIso f g) h _ = h + + -- transformed to equivalence + funExt∙P≃ : (f g : Π∙ A B ptB) → (f ∙∼P g) ≃ (f ≡ g) + funExt∙P≃ f g = isoToEquiv (funExt∙PIso f g) + + -- funExt∙≃ using the other kind of pointed homotopy + funExt∙≃ : (f g : Π∙ A B ptB) → (f ∙∼ g) ≃ (f ≡ g) + funExt∙≃ f g = compEquiv (∙∼≃∙∼P f g) (funExt∙P≃ f g) + + -- standard pointed function extensionality and its inverse + funExt∙ : {f g : Π∙ A B ptB} → f ∙∼ g → f ≡ g + funExt∙ {f = f} {g = g} = equivFun (funExt∙≃ f g) + + funExt∙⁻ : {f g : Π∙ A B ptB} → f ≡ g → f ∙∼ g + funExt∙⁻ {f = f} {g = g} = equivFun (invEquiv (funExt∙≃ f g)) diff --git a/Cubical/Foundations/Pointed/Homogeneous.agda b/Cubical/Foundations/Pointed/Homogeneous.agda index 71507124d..5315e2685 100644 --- a/Cubical/Foundations/Pointed/Homogeneous.agda +++ b/Cubical/Foundations/Pointed/Homogeneous.agda @@ -6,7 +6,7 @@ Portions of this file adapted from Nicolai Kraus' code here: https://bitbucket.org/nicolaikraus/agda/src/e30d70c72c6af8e62b72eefabcc57623dd921f04/trunc-inverse.lagda -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Pointed.Homogeneous where open import Cubical.Foundations.Prelude @@ -14,30 +14,71 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence open import Cubical.Foundations.Path -open import Cubical.Data.Prod -open import Cubical.Data.Empty +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ open import Cubical.Relation.Nullary +open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Pointed.Base open import Cubical.Foundations.Pointed.Properties +open import Cubical.Structures.Pointed isHomogeneous : ∀ {ℓ} → Pointed ℓ → Type (ℓ-suc ℓ) isHomogeneous {ℓ} (A , x) = ∀ y → Path (Pointed ℓ) (A , x) (A , y) +-- Pointed functions into a homogeneous type are equal as soon as they are equal +-- as unpointed functions +→∙Homogeneous≡ : ∀ {ℓ ℓ'} {A∙ : Pointed ℓ} {B∙ : Pointed ℓ'} {f∙ g∙ : A∙ →∙ B∙} + (h : isHomogeneous B∙) → f∙ .fst ≡ g∙ .fst → f∙ ≡ g∙ +→∙Homogeneous≡ {A∙ = A∙@(_ , a₀)} {B∙@(B , _)} {f∙@(_ , f₀)} {g∙@(_ , g₀)} h p = + subst (λ Q∙ → PathP (λ i → A∙ →∙ Q∙ i) f∙ g∙) (sym (flipSquare fix)) badPath + where + badPath : PathP (λ i → A∙ →∙ (B , (sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) f∙ g∙ + badPath i .fst = p i + badPath i .snd j = doubleCompPath-filler (sym f₀) (funExt⁻ p a₀) g₀ j i + + fix : PathP (λ i → B∙ ≡ (B , (sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) refl refl + fix i = + hcomp + (λ j → λ + { (i = i0) → lCancel (h (pt B∙)) j + ; (i = i1) → lCancel (h (pt B∙)) j + }) + (sym (h (pt B∙)) ∙ h ((sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) isHomogeneousPi : ∀ {ℓ ℓ'} {A : Type ℓ} {B∙ : A → Pointed ℓ'} - → (∀ a → isHomogeneous (B∙ a)) → isHomogeneous (Π∙ A B∙) -isHomogeneousPi h f i = (∀ a → typ (h a (f a) i)) , (λ a → pt (h a (f a) i)) + → (∀ a → isHomogeneous (B∙ a)) → isHomogeneous (Πᵘ∙ A B∙) +isHomogeneousPi h f i .fst = ∀ a → typ (h a (f a) i) +isHomogeneousPi h f i .snd a = pt (h a (f a) i) + +isHomogeneous→∙ : ∀ {ℓ ℓ'} {A∙ : Pointed ℓ} {B∙ : Pointed ℓ'} + → isHomogeneous B∙ → isHomogeneous (A∙ →∙ B∙ ∙) +isHomogeneous→∙ {A∙ = A∙} {B∙} h f∙ = + ΣPathP + ( (λ i → Π∙ A∙ (λ a → T a i) (t₀ i)) + , PathPIsoPath _ _ _ .Iso.inv + (→∙Homogeneous≡ h + (PathPIsoPath (λ i → (a : typ A∙) → T a i) (λ _ → pt B∙) _ .Iso.fun + (λ i a → pt (h (f∙ .fst a) i)))) + ) + where + T : ∀ a → typ B∙ ≡ typ B∙ + T a i = typ (h (f∙ .fst a) i) + + t₀ : PathP (λ i → T (pt A∙) i) (pt B∙) (pt B∙) + t₀ = cong pt (h (f∙ .fst (pt A∙))) ▷ f∙ .snd isHomogeneousProd : ∀ {ℓ ℓ'} {A∙ : Pointed ℓ} {B∙ : Pointed ℓ'} → isHomogeneous A∙ → isHomogeneous B∙ → isHomogeneous (A∙ ×∙ B∙) -isHomogeneousProd hA hB (a , b) i = (typ (hA a i)) × (typ (hB b i)) , (pt (hA a i) , pt (hB b i)) +isHomogeneousProd hA hB (a , b) i .fst = typ (hA a i) × typ (hB b i) +isHomogeneousProd hA hB (a , b) i .snd .fst = pt (hA a i) +isHomogeneousProd hA hB (a , b) i .snd .snd = pt (hB b i) isHomogeneousPath : ∀ {ℓ} (A : Type ℓ) {x y : A} (p : x ≡ y) → isHomogeneous ((x ≡ y) , p) -isHomogeneousPath A {x} {y} p q i - = ua eqv i , ua-gluePath eqv (compPathr-cancel p q) i +isHomogeneousPath A {x} {y} p q + = pointed-sip ((x ≡ y) , p) ((x ≡ y) , q) (eqv , compPathr-cancel p q) where eqv : (x ≡ y) ≃ (x ≡ y) - eqv = ((q ∙ sym p) ∙_) , compPathl-isEquiv (q ∙ sym p) + eqv = compPathlEquiv (q ∙ sym p) module HomogeneousDiscrete {ℓ} {A∙ : Pointed ℓ} (dA : Discrete (typ A∙)) (y : typ A∙) where @@ -52,7 +93,7 @@ module HomogeneousDiscrete {ℓ} {A∙ : Pointed ℓ} (dA : Discrete (typ A∙)) switch-ptA∙ : switch (pt A∙) ≡ y switch-ptA∙ with dA (pt A∙) (pt A∙) ... | yes _ = refl - ... | no ¬p = ⊥-elim (¬p refl) + ... | no ¬p = ⊥.rec (¬p refl) switch-idp : ∀ x → switch (switch x) ≡ x switch-idp x with dA x (pt A∙) @@ -60,23 +101,23 @@ module HomogeneousDiscrete {ℓ} {A∙ : Pointed ℓ} (dA : Discrete (typ A∙)) switch-idp x | yes p | yes q = q ∙ sym p switch-idp x | yes p | no _ with dA y y switch-idp x | yes p | no _ | yes _ = sym p - switch-idp x | yes p | no _ | no ¬p = ⊥-elim (¬p refl) + switch-idp x | yes p | no _ | no ¬p = ⊥.rec (¬p refl) switch-idp x | no ¬p with dA x y switch-idp x | no ¬p | yes p with dA y (pt A∙) - switch-idp x | no ¬p | yes p | yes q = ⊥-elim (¬p (p ∙ q)) + switch-idp x | no ¬p | yes p | yes q = ⊥.rec (¬p (p ∙ q)) switch-idp x | no ¬p | yes p | no _ with dA (pt A∙) (pt A∙) switch-idp x | no ¬p | yes p | no _ | yes _ = sym p - switch-idp x | no ¬p | yes p | no _ | no ¬q = ⊥-elim (¬q refl) + switch-idp x | no ¬p | yes p | no _ | no ¬q = ⊥.rec (¬q refl) switch-idp x | no ¬p | no ¬q with dA x (pt A∙) - switch-idp x | no ¬p | no ¬q | yes p = ⊥-elim (¬p p) + switch-idp x | no ¬p | no ¬q | yes p = ⊥.rec (¬p p) switch-idp x | no ¬p | no ¬q | no _ with dA x y - switch-idp x | no ¬p | no ¬q | no _ | yes q = ⊥-elim (¬q q) + switch-idp x | no ¬p | no ¬q | no _ | yes q = ⊥.rec (¬q q) switch-idp x | no ¬p | no ¬q | no _ | no _ = refl switch-eqv : typ A∙ ≃ typ A∙ switch-eqv = isoToEquiv (iso switch switch switch-idp switch-idp) isHomogeneousDiscrete : ∀ {ℓ} {A∙ : Pointed ℓ} (dA : Discrete (typ A∙)) → isHomogeneous A∙ -isHomogeneousDiscrete {ℓ} {A∙} dA y i - = ua switch-eqv i , ua-gluePath switch-eqv switch-ptA∙ i +isHomogeneousDiscrete {ℓ} {A∙} dA y + = pointed-sip (typ A∙ , pt A∙) (typ A∙ , y) (switch-eqv , switch-ptA∙) where open HomogeneousDiscrete {ℓ} {A∙} dA y diff --git a/Cubical/Foundations/Pointed/Homotopy.agda b/Cubical/Foundations/Pointed/Homotopy.agda new file mode 100644 index 000000000..06612f6b2 --- /dev/null +++ b/Cubical/Foundations/Pointed/Homotopy.agda @@ -0,0 +1,119 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Pointed.Homotopy where + +{- + This module defines two kinds of pointed homotopies, + ∙∼ and ∙∼P, and proves their equivalence +-} + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Fiberwise +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Path +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Pointed.Base +open import Cubical.Foundations.Pointed.Properties +open import Cubical.Homotopy.Base +open import Cubical.Data.Sigma + +private + variable + ℓ ℓ' : Level + +module _ {A : Pointed ℓ} {B : typ A → Type ℓ'} {ptB : B (pt A)} where + + ⋆ = pt A + + -- pointed homotopy as pointed Π. This is just a Σ-type, see ∙∼Σ + _∙∼_ : (f g : Π∙ A B ptB) → Type (ℓ-max ℓ ℓ') + (f₁ , f₂) ∙∼ (g₁ , g₂) = Π∙ A (λ x → f₁ x ≡ g₁ x) (f₂ ∙ g₂ ⁻¹) + + -- pointed homotopy with PathP. Also a Σ-type, see ∙∼PΣ + _∙∼P_ : (f g : Π∙ A B ptB) → Type (ℓ-max ℓ ℓ') + (f₁ , f₂) ∙∼P (g₁ , g₂) = Σ[ h ∈ f₁ ∼ g₁ ] PathP (λ i → h ⋆ i ≡ ptB) f₂ g₂ + + -- Proof that f ∙∼ g ≃ f ∙∼P g + -- using equivalence of the total map of φ + private + module _ (f g : Π∙ A B ptB) (H : f .fst ∼ g .fst) where + -- convenient notation + f₁ = fst f + f₂ = snd f + g₁ = fst g + g₂ = snd g + + -- P is the predicate on a homotopy H to be pointed of the ∙∼ kind + P : Type ℓ' + P = H ⋆ ≡ f₂ ∙ g₂ ⁻¹ + + -- Q is the predicate on a homotopy H to be pointed of the ∙∼P kind + Q : Type ℓ' + Q = PathP (λ i → H ⋆ i ≡ ptB) f₂ g₂ + + -- simplify the notation even more to see that P≡Q + -- is just a jingle of paths + p = H ⋆ + r = f₂ + s = g₂ + P≡Q : P ≡ Q + P≡Q = p ≡ r ∙ s ⁻¹ + ≡⟨ isoToPath (symIso p (r ∙ s ⁻¹)) ⟩ + r ∙ s ⁻¹ ≡ p + ≡⟨ cong (r ∙ s ⁻¹ ≡_) (rUnit p ∙∙ cong (p ∙_) (sym (rCancel s)) ∙∙ assoc p s (s ⁻¹)) ⟩ + r ∙ s ⁻¹ ≡ (p ∙ s) ∙ s ⁻¹ + ≡⟨ sym (ua (compr≡Equiv r (p ∙ s) (s ⁻¹))) ⟩ + r ≡ p ∙ s + ≡⟨ ua (compl≡Equiv (p ⁻¹) r (p ∙ s)) ⟩ + p ⁻¹ ∙ r ≡ p ⁻¹ ∙ (p ∙ s) + ≡⟨ cong (p ⁻¹ ∙ r ≡_ ) (assoc (p ⁻¹) p s ∙∙ (cong (_∙ s) (lCancel p)) ∙∙ sym (lUnit s)) ⟩ + p ⁻¹ ∙ r ≡ s + ≡⟨ cong (λ z → p ⁻¹ ∙ z ≡ s) (rUnit r) ⟩ + p ⁻¹ ∙ (r ∙ refl) ≡ s + ≡⟨ cong (_≡ s) (sym (doubleCompPath-elim' (p ⁻¹) r refl)) ⟩ + p ⁻¹ ∙∙ r ∙∙ refl ≡ s + ≡⟨ sym (ua (Square≃doubleComp r s p refl)) ⟩ + PathP (λ i → p i ≡ ptB) r s ∎ + + -- φ is a fiberwise transformation (H : f ∼ g) → P H → Q H + -- φ is even a fiberwise equivalence by P≡Q + φ : P → Q + φ = transport P≡Q + + -- The total map corresponding to φ + totφ : (f g : Π∙ A B ptB) → f ∙∼ g → f ∙∼P g + totφ f g p .fst = p .fst + totφ f g p .snd = φ f g (p .fst) (p .snd) + + -- transformation of the homotopies using totφ + ∙∼→∙∼P : (f g : Π∙ A B ptB) → (f ∙∼ g) → (f ∙∼P g) + ∙∼→∙∼P f g = totφ f g + + -- Proof that ∙∼ and ∙∼P are equivalent using the fiberwise equivalence φ + ∙∼≃∙∼P : (f g : Π∙ A B ptB) → (f ∙∼ g) ≃ (f ∙∼P g) + ∙∼≃∙∼P f g = Σ-cong-equiv-snd (λ H → transportEquiv (P≡Q f g H)) + + -- inverse of ∙∼→∙∼P extracted from the equivalence + ∙∼P→∙∼ : {f g : Π∙ A B ptB} → f ∙∼P g → f ∙∼ g + ∙∼P→∙∼ {f = f} {g = g} = invEq (∙∼≃∙∼P f g) + + -- ∙∼≃∙∼P transformed to a path + ∙∼≡∙∼P : (f g : Π∙ A B ptB) → (f ∙∼ g) ≡ (f ∙∼P g) + ∙∼≡∙∼P f g = ua (∙∼≃∙∼P f g) + + -- Verifies that the pointed homotopies actually correspond + -- to their Σ-type versions + _∙∼Σ_ : (f g : Π∙ A B ptB) → Type (ℓ-max ℓ ℓ') + f ∙∼Σ g = Σ[ H ∈ f .fst ∼ g .fst ] (P f g H) + + _∙∼PΣ_ : (f g : Π∙ A B ptB) → Type (ℓ-max ℓ ℓ') + f ∙∼PΣ g = Σ[ H ∈ f .fst ∼ g .fst ] (Q f g H) + + ∙∼≡∙∼Σ : (f g : Π∙ A B ptB) → f ∙∼ g ≡ f ∙∼Σ g + ∙∼≡∙∼Σ f g = refl + + ∙∼P≡∙∼PΣ : (f g : Π∙ A B ptB) → f ∙∼P g ≡ f ∙∼PΣ g + ∙∼P≡∙∼PΣ f g = refl diff --git a/Cubical/Foundations/Pointed/Properties.agda b/Cubical/Foundations/Pointed/Properties.agda index b093c0e1e..d56e5b595 100644 --- a/Cubical/Foundations/Pointed/Properties.agda +++ b/Cubical/Foundations/Pointed/Properties.agda @@ -1,15 +1,83 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Pointed.Properties where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed.Base -open import Cubical.Data.Prod +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Isomorphism -Π∙ : ∀ {ℓ ℓ'} (A : Type ℓ) (B∙ : A → Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') -Π∙ A B∙ = (∀ a → typ (B∙ a)) , (λ a → pt (B∙ a)) +open import Cubical.Data.Sigma -Σ∙ : ∀ {ℓ ℓ'} (A∙ : Pointed ℓ) (B∙ : typ A∙ → Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') -Σ∙ A∙ B∙ = (Σ[ a ∈ typ A∙ ] typ (B∙ a)) , (pt A∙ , pt (B∙ (pt A∙))) +private + variable + ℓ ℓ' ℓA ℓB ℓC ℓD : Level -_×∙_ : ∀ {ℓ ℓ'} (A∙ : Pointed ℓ) (B∙ : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') -A∙ ×∙ B∙ = ((typ A∙) × (typ B∙)) , (pt A∙ , pt B∙) +-- the default pointed Π-type: A is pointed, and B has a base point in the chosen fiber +Π∙ : (A : Pointed ℓ) (B : typ A → Type ℓ') (ptB : B (pt A)) → Type (ℓ-max ℓ ℓ') +Π∙ A B ptB = Σ[ f ∈ ((a : typ A) → B a) ] f (pt A) ≡ ptB + +-- the unpointed Π-type becomes a pointed type if the fibers are all pointed +Πᵘ∙ : (A : Type ℓ) (B : A → Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +Πᵘ∙ A B .fst = ∀ a → typ (B a) +Πᵘ∙ A B .snd a = pt (B a) + +-- if the base and all fibers are pointed, we have the pointed pointed Π-type +Πᵖ∙ : (A : Pointed ℓ) (B : typ A → Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +Πᵖ∙ A B .fst = Π∙ A (typ ∘ B) (pt (B (pt A))) +Πᵖ∙ A B .snd .fst a = pt (B a) +Πᵖ∙ A B .snd .snd = refl + +-- the default pointed Σ-type is just the Σ-type, but as a pointed type +Σ∙ : (A : Pointed ℓ) (B : typ A → Type ℓ') (ptB : B (pt A)) → Pointed (ℓ-max ℓ ℓ') +Σ∙ A B ptB .fst = Σ[ a ∈ typ A ] B a +Σ∙ A B ptB .snd .fst = pt A +Σ∙ A B ptB .snd .snd = ptB + +-- version if B is a family of pointed types +Σᵖ∙ : (A : Pointed ℓ) (B : typ A → Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +Σᵖ∙ A B = Σ∙ A (typ ∘ B) (pt (B (pt A))) + +_×∙_ : (A∙ : Pointed ℓ) (B∙ : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +(A∙ ×∙ B∙) .fst = (typ A∙) × (typ B∙) +(A∙ ×∙ B∙) .snd .fst = pt A∙ +(A∙ ×∙ B∙) .snd .snd = pt B∙ + +-- composition of pointed maps +_∘∙_ : {A : Pointed ℓA} {B : Pointed ℓB} {C : Pointed ℓC} + (g : B →∙ C) (f : A →∙ B) → (A →∙ C) +((g , g∙) ∘∙ (f , f∙)) .fst x = g (f x) +((g , g∙) ∘∙ (f , f∙)) .snd = (cong g f∙) ∙ g∙ + +-- pointed identity +id∙ : (A : Pointed ℓA) → (A →∙ A) +id∙ A .fst x = x +id∙ A .snd = refl + +-- constant pointed map +const∙ : (A : Pointed ℓA) (B : Pointed ℓB) → (A →∙ B) +const∙ _ B .fst _ = B .snd +const∙ _ B .snd = refl + +-- left identity law for pointed maps +∘∙-idˡ : {A : Pointed ℓA} {B : Pointed ℓB} (f : A →∙ B) → f ∘∙ id∙ A ≡ f +∘∙-idˡ f = ΣPathP ( refl , (lUnit (f .snd)) ⁻¹ ) + +-- right identity law for pointed maps +∘∙-idʳ : {A : Pointed ℓA} {B : Pointed ℓB} (f : A →∙ B) → id∙ B ∘∙ f ≡ f +∘∙-idʳ f = ΣPathP ( refl , (rUnit (f .snd)) ⁻¹ ) + +-- associativity for composition of pointed maps +∘∙-assoc : {A : Pointed ℓA} {B : Pointed ℓB} {C : Pointed ℓC} {D : Pointed ℓD} + (h : C →∙ D) (g : B →∙ C) (f : A →∙ B) + → (h ∘∙ g) ∘∙ f ≡ h ∘∙ (g ∘∙ f) +∘∙-assoc (h , h∙) (g , g∙) (f , f∙) = ΣPathP (refl , q) + where + q : (cong (h ∘ g) f∙) ∙ (cong h g∙ ∙ h∙) ≡ cong h (cong g f∙ ∙ g∙) ∙ h∙ + q = ( (cong (h ∘ g) f∙) ∙ (cong h g∙ ∙ h∙) + ≡⟨ refl ⟩ + (cong h (cong g f∙)) ∙ (cong h g∙ ∙ h∙) + ≡⟨ assoc (cong h (cong g f∙)) (cong h g∙) h∙ ⟩ + (cong h (cong g f∙) ∙ cong h g∙) ∙ h∙ + ≡⟨ cong (λ p → p ∙ h∙) ((cong-∙ h (cong g f∙) g∙) ⁻¹) ⟩ + (cong h (cong g f∙ ∙ g∙) ∙ h∙) ∎ ) diff --git a/Cubical/Foundations/Powerset.agda b/Cubical/Foundations/Powerset.agda new file mode 100644 index 000000000..043a0dfed --- /dev/null +++ b/Cubical/Foundations/Powerset.agda @@ -0,0 +1,65 @@ +{- + +This file introduces the "powerset" of a type in the style of +Escardó's lecture notes: + +https://www.cs.bham.ac.uk/~mhe/HoTT-UF-in-Agda-Lecture-Notes/HoTT-UF-Agda.html#propositionalextensionality + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Powerset where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence using (hPropExt) + +open import Cubical.Data.Sigma + +private + variable + ℓ : Level + X : Type ℓ + +ℙ : Type ℓ → Type (ℓ-suc ℓ) +ℙ X = X → hProp _ + +isSetℙ : isSet (ℙ X) +isSetℙ = isSetΠ λ x → isSetHProp + +infix 5 _∈_ + +_∈_ : {X : Type ℓ} → X → ℙ X → Type ℓ +x ∈ A = ⟨ A x ⟩ + +_⊆_ : {X : Type ℓ} → ℙ X → ℙ X → Type ℓ +A ⊆ B = ∀ x → x ∈ A → x ∈ B + +∈-isProp : (A : ℙ X) (x : X) → isProp (x ∈ A) +∈-isProp A = snd ∘ A + +⊆-isProp : (A B : ℙ X) → isProp (A ⊆ B) +⊆-isProp A B = isPropΠ2 (λ x _ → ∈-isProp B x) + +⊆-refl : (A : ℙ X) → A ⊆ A +⊆-refl A x = idfun (x ∈ A) + +subst-∈ : (A : ℙ X) {x y : X} → x ≡ y → x ∈ A → y ∈ A +subst-∈ A = subst (_∈ A) + +⊆-refl-consequence : (A B : ℙ X) → A ≡ B → (A ⊆ B) × (B ⊆ A) +⊆-refl-consequence A B p = subst (A ⊆_) p (⊆-refl A) + , subst (B ⊆_) (sym p) (⊆-refl B) + +⊆-extensionality : (A B : ℙ X) → (A ⊆ B) × (B ⊆ A) → A ≡ B +⊆-extensionality A B (φ , ψ) = + funExt (λ x → TypeOfHLevel≡ 1 (hPropExt (A x .snd) (B x .snd) (φ x) (ψ x))) + +⊆-extensionalityEquiv : (A B : ℙ X) → (A ⊆ B) × (B ⊆ A) ≃ (A ≡ B) +⊆-extensionalityEquiv A B = isoToEquiv (iso (⊆-extensionality A B) + (⊆-refl-consequence A B) + (λ _ → isSetℙ A B _ _) + (λ _ → isPropΣ (⊆-isProp A B) (λ _ → ⊆-isProp B A) _ _)) diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 060ba61a9..8d5731cf4 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -20,7 +20,7 @@ This file proves a variety of basic results about paths: - Export universe lifting -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Prelude where open import Cubical.Core.Primitives public @@ -28,6 +28,7 @@ open import Cubical.Core.Primitives public infixr 30 _∙_ infix 3 _∎ infixr 2 _≡⟨_⟩_ +infixr 2.5 _≡⟨_⟩≡⟨_⟩_ -- Basic theory about paths. These proofs should typically be -- inlined. This module also makes equational reasoning work with @@ -38,13 +39,15 @@ private ℓ ℓ' : Level A : Type ℓ B : A → Type ℓ - x y z : A + x y z w : A refl : x ≡ x refl {x = x} = λ _ → x +{-# INLINE refl #-} sym : x ≡ y → y ≡ x sym p i = p (~ i) +{-# INLINE sym #-} symP : {A : I → Type ℓ} → {x : A i0} → {y : A i1} → (p : PathP A x y) → PathP (λ i → A (~ i)) y x @@ -53,6 +56,7 @@ symP p j = p (~ j) cong : ∀ (f : (a : A) → B a) (p : x ≡ y) → PathP (λ i → B (p i)) (f x) (f y) cong f p i = f (p i) +{-# INLINE cong #-} cong₂ : ∀ {C : (a : A) → (b : B a) → Type ℓ} → (f : (a : A) → (b : B a) → C a b) → @@ -60,65 +64,154 @@ cong₂ : ∀ {C : (a : A) → (b : B a) → Type ℓ} → {u : B x} {v : B y} (q : PathP (λ i → B (p i)) u v) → PathP (λ i → C (p i) (q i)) (f x u) (f y v) cong₂ f p q i = f (p i) (q i) +{-# INLINE cong₂ #-} --- The filler of homogeneous path composition: --- compPath-filler p q = PathP (λ i → x ≡ q i) p (p ∙ q) +{- The most natural notion of homogenous path composition + in a cubical setting is double composition: -compPath-filler : ∀ {x y z : A} → x ≡ y → y ≡ z → I → I → A -compPath-filler {x = x} p q j i = - hfill (λ j → λ { (i = i0) → x - ; (i = i1) → q j }) (inS (p i)) j + x ∙ ∙ ∙ > w + ^ ^ + p⁻¹ | | r ^ + | | j | + y — — — > z ∙ — > + q i + + `p ∙∙ q ∙∙ r` gives the line at the top, + `doubleCompPath-filler p q r` gives the whole square +-} + +doubleComp-faces : {x y z w : A } (p : x ≡ y) (r : z ≡ w) + → (i : I) (j : I) → Partial (i ∨ ~ i) A +doubleComp-faces p r i j (i = i0) = p (~ j) +doubleComp-faces p r i j (i = i1) = r j + +_∙∙_∙∙_ : w ≡ x → x ≡ y → y ≡ z → w ≡ z +(p ∙∙ q ∙∙ r) i = + hcomp (doubleComp-faces p r i) (q i) + +doubleCompPath-filler : (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → PathP (λ j → p (~ j) ≡ r j) q (p ∙∙ q ∙∙ r) +doubleCompPath-filler p q r j i = + hfill (doubleComp-faces p r i) (inS (q i)) j + +-- any two definitions of double composition are equal +compPath-unique : ∀ (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → (α β : Σ[ s ∈ x ≡ w ] PathP (λ j → p (~ j) ≡ r j) q s) + → α ≡ β +compPath-unique p q r (α , α-filler) (β , β-filler) t + = (λ i → cb i1 i) , (λ j i → cb j i) + where cb : I → I → _ + cb j i = hfill (λ j → λ { (t = i0) → α-filler j i + ; (t = i1) → β-filler j i + ; (i = i0) → p (~ j) + ; (i = i1) → r j }) + (inS (q i)) j + +{- For single homogenous path composition, we take `p = refl`: + + x ∙ ∙ ∙ > z + ‖ ^ + ‖ | r ^ + ‖ | j | + x — — — > y ∙ — > + q i + + `q ∙ r` gives the line at the top, + `compPath-filler q r` gives the whole square +-} _∙_ : x ≡ y → y ≡ z → x ≡ z -(p ∙ q) j = compPath-filler p q i1 j +p ∙ q = refl ∙∙ p ∙∙ q + +compPath-filler : (p : x ≡ y) (q : y ≡ z) → PathP (λ j → x ≡ q j) p (p ∙ q) +compPath-filler p q = doubleCompPath-filler refl p q + +-- We could have also defined single composition by taking `r = refl`: + +_∙'_ : x ≡ y → y ≡ z → x ≡ z +p ∙' q = p ∙∙ q ∙∙ refl + +compPath'-filler : (p : x ≡ y) (q : y ≡ z) → PathP (λ j → p (~ j) ≡ z) q (p ∙' q) +compPath'-filler p q = doubleCompPath-filler p q refl + +-- It's easy to show that `p ∙ q` also has such a filler: +compPath-filler' : (p : x ≡ y) (q : y ≡ z) → PathP (λ j → p (~ j) ≡ z) q (p ∙ q) +compPath-filler' {z = z} p q j i = + hcomp (λ k → λ { (i = i0) → p (~ j) + ; (i = i1) → q k + ; (j = i0) → q (i ∧ k) }) + (p (i ∨ ~ j)) +-- Note: We can omit a (j = i1) case here since when (j = i1), the whole expression is +-- definitionally equal to `p ∙ q`. (Notice that `p ∙ q` is also an hcomp.) Nevertheless, +-- we could have given `compPath-filler p q k i` as the (j = i1) case. + +-- From this, we can show that these two notions of composition are the same +compPath≡compPath' : (p : x ≡ y) (q : y ≡ z) → p ∙ q ≡ p ∙' q +compPath≡compPath' p q j = + compPath-unique p q refl (p ∙ q , compPath-filler' p q) + (p ∙' q , compPath'-filler p q) j .fst + +-- Heterogeneous path composition and its filler: + +compPathP : ∀ {A : I → Type ℓ} {x : A i0} {y : A i1} {B_i1 : Type ℓ} {B : (A i1) ≡ B_i1} {z : B i1} + → PathP A x y → PathP (λ i → B i) y z → PathP (λ j → ((λ i → A i) ∙ B) j) x z +compPathP {A = A} {x = x} {B = B} p q i = + comp (λ j → compPath-filler (λ i → A i) B j i) + (λ j → λ { (i = i0) → x ; + (i = i1) → q j }) + (p i) + + +compPathP' : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {x y z : A} {x' : B x} {y' : B y} {z' : B z} {p : x ≡ y} {q : y ≡ z} + (P : PathP (λ i → B (p i)) x' y') (Q : PathP (λ i → B (q i)) y' z') + → PathP (λ i → B ((p ∙ q) i)) x' z' +compPathP' {B = B} {x' = x'} {p = p} {q = q} P Q i = + comp (λ j → B (compPath-filler p q j i)) + (λ j → λ { (i = i0) → x' ; + (i = i1) → Q j }) + (P i) --- The filler of heterogeneous path composition: --- compPathP-filler p q = PathP (λ i → PathP (λ j → (compPath-filler (λ i → A i) B i j)) x (q i)) p (compPathP p q) -compPathP-filler : {A : I → Type ℓ} → {x : A i0} → {y : A i1} → {B_i1 : Type ℓ} {B : A i1 ≡ B_i1} → {z : B i1} → - (p : PathP A x y) → (q : PathP (λ i → B i) y z) → ∀ (i j : I) → compPath-filler (λ i → A i) B j i -compPathP-filler {A = A} {x = x} {B = B} p q i = +compPathP-filler : ∀ {A : I → Type ℓ} {x : A i0} {y : A i1} {B_i1 : Type ℓ} {B : A i1 ≡ B_i1} {z : B i1} + → (p : PathP A x y) (q : PathP (λ i → B i) y z) + → PathP (λ j → PathP (λ k → (compPath-filler (λ i → A i) B j k)) x (q j)) p (compPathP p q) +compPathP-filler {A = A} {x = x} {B = B} p q j i = fill (λ j → compPath-filler (λ i → A i) B j i) (λ j → λ { (i = i0) → x ; - (i = i1) → q j }) (inS (p i)) + (i = i1) → q j }) + (inS (p i)) j -compPathP : {A : I → Type ℓ} → {x : A i0} → {y : A i1} → {B_i1 : Type ℓ} {B : (A i1) ≡ B_i1} → {z : B i1} → - (p : PathP A x y) → (q : PathP (λ i → B i) y z) → PathP (λ j → ((λ i → A i) ∙ B) j) x z -compPathP p q j = compPathP-filler p q j i1 + + +compPathP'-filler : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {x y z : A} {x' : B x} {y' : B y} {z' : B z} {p : x ≡ y} {q : y ≡ z} + (P : PathP (λ i → B (p i)) x' y') (Q : PathP (λ i → B (q i)) y' z') + → PathP (λ j → PathP (λ i → B (compPath-filler p q j i)) x' (Q j)) P (compPathP' {x = x} {y = y} {x' = x'} {y' = y'} P Q) +compPathP'-filler {B = B} {x' = x'} {p = p} {q = q} P Q j i = + fill (λ j → B (compPath-filler p q j i)) + (λ j → λ { (i = i0) → x' ; + (i = i1) → Q j }) + (inS (P i)) + j _≡⟨_⟩_ : (x : A) → x ≡ y → y ≡ z → x ≡ z _ ≡⟨ x≡y ⟩ y≡z = x≡y ∙ y≡z - ≡⟨⟩-syntax : (x : A) → x ≡ y → y ≡ z → x ≡ z ≡⟨⟩-syntax = _≡⟨_⟩_ infixr 2 ≡⟨⟩-syntax syntax ≡⟨⟩-syntax x (λ i → B) y = x ≡[ i ]⟨ B ⟩ y -_∎ : (x : A) → x ≡ x -_ ∎ = refl +≡⟨⟩⟨⟩-syntax : (x y : A) → x ≡ y → y ≡ z → z ≡ w → x ≡ w +≡⟨⟩⟨⟩-syntax x y p q r = p ∙∙ q ∙∙ r +infixr 3 ≡⟨⟩⟨⟩-syntax +syntax ≡⟨⟩⟨⟩-syntax x y B C = x ≡⟨ B ⟩≡ y ≡⟨ C ⟩≡ --- another definition of composition, useful for some proofs -compPath'-filler : ∀ {x y z : A} → x ≡ y → y ≡ z → I → I → A -compPath'-filler {z = z} p q j i = - hfill (λ j → λ { (i = i0) → p (~ j) - ; (i = i1) → z }) (inS (q i)) j +_≡⟨_⟩≡⟨_⟩_ : (x : A) → x ≡ y → y ≡ z → z ≡ w → x ≡ w +_ ≡⟨ x≡y ⟩≡⟨ y≡z ⟩ z≡w = x≡y ∙∙ y≡z ∙∙ z≡w -_□_ : x ≡ y → y ≡ z → x ≡ z -(p □ q) j = compPath'-filler p q i1 j +_∎ : (x : A) → x ≡ x +_ ∎ = refl -□≡∙ : (p : x ≡ y) (q : y ≡ z) → p □ q ≡ p ∙ q -□≡∙ {x = x} {y = y} {z = z} p q i j = hcomp (λ k → \ { (i = i0) → compPath'-filler p q k j - ; (i = i1) → compPath-filler p q k j - ; (j = i0) → p ( ~ i ∧ ~ k) - ; (j = i1) → q (k ∨ ~ i) }) (helper i j) - where - helper : PathP (λ i → p (~ i) ≡ q (~ i)) q p - helper i j = hcomp (λ k → \ { (i = i0) → q (k ∧ j) - ; (i = i1) → p (~ k ∨ j) - ; (j = i0) → p (~ i ∨ ~ k) - ; (j = i1) → q (~ i ∧ k) }) - y -- Transport, subst and functional extensionality @@ -131,18 +224,45 @@ transport p a = transp (λ i → p i) i0 a transportRefl : (x : A) → transport refl x ≡ x transportRefl {A = A} x i = transp (λ _ → A) i x +transport-filler : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) (x : A) + → PathP (λ i → p i) x (transport p x) +transport-filler p x i = transp (λ j → p (i ∧ j)) (~ i) x + -- We want B to be explicit in subst subst : (B : A → Type ℓ') (p : x ≡ y) → B x → B y subst B p pa = transport (λ i → B (p i)) pa +subst2 : ∀ {ℓ' ℓ''} {B : Type ℓ'} {z w : B} (C : A → B → Type ℓ'') + (p : x ≡ y) (q : z ≡ w) → C x z → C y w +subst2 B p q b = transport (λ i → B (p i) (q i)) b + substRefl : (px : B x) → subst B refl px ≡ px substRefl px = transportRefl px -funExt : {f g : (x : A) → B x} → ((x : A) → f x ≡ g x) → f ≡ g +subst-filler : (B : A → Type ℓ') (p : x ≡ y) (b : B x) + → PathP (λ i → B (p i)) b (subst B p b) +subst-filler B p = transport-filler (cong B p) + +funExt : {B : A → I → Type ℓ'} + {f : (x : A) → B x i0} {g : (x : A) → B x i1} + → ((x : A) → PathP (B x) (f x) (g x)) + → PathP (λ i → (x : A) → B x i) f g funExt p i x = p x i --- the inverse to funExt (see Foundations.FunExtEquiv) -funExt⁻ : ∀ {f g : (x : A) → B x} → f ≡ g → (x : A) → f x ≡ g x +implicitFunExt : {B : A → I → Type ℓ'} + {f : {x : A} → B x i0} {g : {x : A} → B x i1} + → ({x : A} → PathP (B x) (f {x}) (g {x})) + → PathP (λ i → {x : A} → B x i) f g +implicitFunExt p i {x} = p {x} i + +-- the inverse to funExt (see Functions.FunExtEquiv), converting paths +-- between functions to homotopies; `funExt⁻` is called `happly` and +-- defined by path induction in the HoTT book (see function 2.9.2 in +-- section 2.9) +funExt⁻ : {B : A → I → Type ℓ'} + {f : (x : A) → B x i0} {g : (x : A) → B x i1} + → PathP (λ i → (x : A) → B x i) f g + → ((x : A) → PathP (B x) (f x) (g x)) funExt⁻ eq x i = eq i x -- J for paths and its computation rule @@ -154,26 +274,36 @@ module _ (P : ∀ y → x ≡ y → Type ℓ') (d : P x refl) where JRefl : J refl ≡ d JRefl = transportRefl d --- Contractibility of singletons - -singl : (a : A) → Type _ -singl {A = A} a = Σ[ x ∈ A ] (a ≡ x) - -contrSingl : (p : x ≡ y) → Path (singl x) (x , refl) (y , p) -contrSingl p i = (p i , λ j → p (i ∧ j)) - + J-∙ : (p : x ≡ y) (q : y ≡ z) + → J (p ∙ q) ≡ transport (λ i → P (q i) (λ j → compPath-filler p q i j)) (J p) + J-∙ p q k = + transp + (λ i → P (q (i ∨ ~ k)) + (λ j → compPath-filler p q (i ∨ ~ k) j)) (~ k) + (J (λ j → compPath-filler p q (~ k) j)) -- Converting to and from a PathP module _ {A : I → Type ℓ} {x : A i0} {y : A i1} where - toPathP : transp A i0 x ≡ y → PathP A x y + toPathP : transp (\ i → A i) i0 x ≡ y → PathP A x y toPathP p i = hcomp (λ j → λ { (i = i0) → x ; (i = i1) → p j }) (transp (λ j → A (i ∧ j)) (~ i) x) - fromPathP : PathP A x y → transp A i0 x ≡ y + fromPathP : PathP A x y → transp (\ i → A i) i0 x ≡ y fromPathP p i = transp (λ j → A (i ∨ j)) i (p i) +-- Whiskering a dependent path by a path + +_◁_ : ∀ {ℓ} {A : I → Type ℓ} {a₀ a₀' : A i0} {a₁ : A i1} + → a₀ ≡ a₀' → PathP A a₀' a₁ → PathP A a₀ a₁ +(p ◁ q) i = + hcomp (λ j → λ {(i = i0) → p (~ j); (i = i1) → q i1}) (q i) + +_▷_ : ∀ {ℓ} {A : I → Type ℓ} {a₀ : A i0} {a₁ a₁' : A i1} + → PathP A a₀ a₁ → a₁ ≡ a₁' → PathP A a₀ a₁' +(p ▷ q) i = + hcomp (λ j → λ {(i = i0) → p i0; (i = i1) → q j}) (p i) -- Direct definitions of lower h-levels @@ -186,6 +316,36 @@ isProp A = (x y : A) → x ≡ y isSet : Type ℓ → Type ℓ isSet A = (x y : A) → isProp (x ≡ y) +isGroupoid : Type ℓ → Type ℓ +isGroupoid A = ∀ a b → isSet (Path A a b) + +is2Groupoid : Type ℓ → Type ℓ +is2Groupoid A = ∀ a b → isGroupoid (Path A a b) + +-- Contractibility of singletons + +singlP : (A : I → Type ℓ) (a : A i0) → Type _ +singlP A a = Σ[ x ∈ A i1 ] PathP A a x + +singl : (a : A) → Type _ +singl {A = A} a = singlP (λ _ → A) a + +isContrSingl : (a : A) → isContr (singl a) +isContrSingl a = (a , refl) , λ p i → p .snd i , λ j → p .snd (i ∧ j) + +isContrSinglP : (A : I → Type ℓ) (a : A i0) → isContr (singlP A a) +isContrSinglP A a .fst = _ , transport-filler (λ i → A i) a +isContrSinglP A a .snd (x , p) i = + _ , λ j → fill (\ i → A i) (λ j → λ {(i = i0) → transport-filler (λ i → A i) a j; (i = i1) → p j}) (inS a) j + +SquareP : + (A : I → I → Type ℓ) + {a₀₀ : A i0 i0} {a₀₁ : A i0 i1} (a₀₋ : PathP (λ j → A i0 j) a₀₀ a₀₁) + {a₁₀ : A i1 i0} {a₁₁ : A i1 i1} (a₁₋ : PathP (λ j → A i1 j) a₁₀ a₁₁) + (a₋₀ : PathP (λ i → A i i0) a₀₀ a₁₀) (a₋₁ : PathP (λ i → A i i1) a₀₁ a₁₁) + → Type ℓ +SquareP A a₀₋ a₁₋ a₋₀ a₋₁ = PathP (λ i → PathP (λ j → A i j) (a₋₀ i) (a₋₁ i)) a₀₋ a₁₋ + Square : {a₀₀ a₀₁ : A} (a₀₋ : a₀₀ ≡ a₀₁) {a₁₀ a₁₁ : A} (a₁₋ : a₁₀ ≡ a₁₁) @@ -200,9 +360,6 @@ isSet' A = (a₋₀ : a₀₀ ≡ a₁₀) (a₋₁ : a₀₁ ≡ a₁₁) → Square a₀₋ a₁₋ a₋₀ a₋₁ -isGroupoid : Type ℓ → Type ℓ -isGroupoid A = ∀ a b → isSet (Path A a b) - Cube : {a₀₀₀ a₀₀₁ : A} {a₀₀₋ : a₀₀₀ ≡ a₀₀₁} {a₀₁₀ a₀₁₁ : A} {a₀₁₋ : a₀₁₀ ≡ a₀₁₁} @@ -222,7 +379,7 @@ Cube : Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ = PathP (λ i → Square (a₋₀₋ i) (a₋₁₋ i) (a₋₋₀ i) (a₋₋₁ i)) a₀₋₋ a₁₋₋ -isGroupoid' : Set ℓ → Set ℓ +isGroupoid' : Type ℓ → Type ℓ isGroupoid' A = {a₀₀₀ a₀₀₁ : A} {a₀₀₋ : a₀₀₀ ≡ a₀₀₁} {a₀₁₀ a₀₁₁ : A} {a₀₁₋ : a₀₁₀ ≡ a₀₁₁} @@ -240,9 +397,6 @@ isGroupoid' A = (a₋₋₁ : Square a₀₋₁ a₁₋₁ a₋₀₁ a₋₁₁) → Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ -is2Groupoid : Type ℓ → Type ℓ -is2Groupoid A = ∀ a b → isGroupoid (Path A a b) - -- Essential consequences of isProp and isContr isProp→PathP : ∀ {B : I → Type ℓ} → ((i : I) → isProp (B i)) → (b0 : B i0) (b1 : B i1) @@ -250,13 +404,11 @@ isProp→PathP : ∀ {B : I → Type ℓ} → ((i : I) → isProp (B i)) isProp→PathP hB b0 b1 = toPathP (hB _ _ _) isPropIsContr : isProp (isContr A) -isPropIsContr z0 z1 j = - ( z0 .snd (z1 .fst) j - , λ x i → hcomp (λ k → λ { (i = i0) → z0 .snd (z1 .fst) j - ; (i = i1) → z0 .snd x (j ∨ k) - ; (j = i0) → z0 .snd x (i ∧ k) - ; (j = i1) → z1 .snd x i }) - (z0 .snd (z1 .snd x i) j)) +isPropIsContr (c0 , h0) (c1 , h1) j = + h0 c1 j , λ y i → hcomp (λ k → λ { (i = i0) → h0 (h0 c1 j) k; + (i = i1) → h0 y k; + (j = i0) → h0 (h0 y i) k; + (j = i1) → h0 (h1 y i) k}) c0 isContr→isProp : isContr A → isProp A isContr→isProp (x , p) a b i = @@ -270,9 +422,22 @@ isProp→isSet h a b p q j i = ; (j = i0) → h a (p i) k ; (j = i1) → h a (q i) k }) a +isProp→isSet' : isProp A → isSet' A +isProp→isSet' h {a} p q r s i j = + hcomp (λ k → λ { (i = i0) → h a (p j) k + ; (i = i1) → h a (q j) k + ; (j = i0) → h a (r i) k + ; (j = i1) → h a (s i) k}) a + isPropIsProp : isProp (isProp A) isPropIsProp f g i a b = isProp→isSet f a b (f a b) (g a b) i +isPropSingl : {a : A} → isProp (singl a) +isPropSingl = isContr→isProp (isContrSingl _) + +isPropSinglP : {A : I → Type ℓ} {a : A i0} → isProp (singlP A a) +isPropSinglP = isContr→isProp (isContrSinglP _ _) + -- Universe lifting record Lift {i j} (A : Type i) : Type (ℓ-max i j) where @@ -281,3 +446,74 @@ record Lift {i j} (A : Type i) : Type (ℓ-max i j) where lower : A open Lift public + +liftExt : ∀ {A : Type ℓ} {a b : Lift {ℓ} {ℓ'} A} → (lower a ≡ lower b) → a ≡ b +liftExt x i = lift (x i) + +doubleCompPath-filler∙ : {a b c d : A} (p : a ≡ b) (q : b ≡ c) (r : c ≡ d) + → PathP (λ i → p i ≡ r (~ i)) (p ∙ q ∙ r) q +doubleCompPath-filler∙ {A = A} {b = b} p q r j i = + hcomp (λ k → λ { (i = i0) → p j + ; (i = i1) → side j k + ; (j = i1) → q (i ∧ k)}) + (p (j ∨ i)) + where + side : I → I → A + side i j = + hcomp (λ k → λ { (i = i1) → q j + ; (j = i0) → b + ; (j = i1) → r (~ i ∧ k)}) + (q j) + +PathP→compPathL : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} + → PathP (λ i → p i ≡ q i) r s + → sym p ∙ r ∙ q ≡ s +PathP→compPathL {p = p} {q = q} {r = r} {s = s} P j i = + hcomp (λ k → λ { (i = i0) → p (j ∨ k) + ; (i = i1) → q (j ∨ k) + ; (j = i0) → doubleCompPath-filler∙ (sym p) r q (~ k) i + ; (j = i1) → s i }) + (P j i) + +PathP→compPathR : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} + → PathP (λ i → p i ≡ q i) r s + → r ≡ p ∙ s ∙ sym q +PathP→compPathR {p = p} {q = q} {r = r} {s = s} P j i = + hcomp (λ k → λ { (i = i0) → p (j ∧ (~ k)) + ; (i = i1) → q (j ∧ (~ k)) + ; (j = i0) → r i + ; (j = i1) → doubleCompPath-filler∙ p s (sym q) (~ k) i}) + (P j i) + + +-- otherdir + +compPathL→PathP : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} + → sym p ∙ r ∙ q ≡ s + → PathP (λ i → p i ≡ q i) r s +compPathL→PathP {p = p} {q = q} {r = r} {s = s} P j i = + hcomp (λ k → λ { (i = i0) → p (~ k ∨ j) + ; (i = i1) → q (~ k ∨ j) + ; (j = i0) → doubleCompPath-filler∙ (sym p) r q k i + ; (j = i1) → s i}) + (P j i) + +compPathR→PathP : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} + → r ≡ p ∙ s ∙ sym q + → PathP (λ i → p i ≡ q i) r s +compPathR→PathP {p = p} {q = q} {r = r} {s = s} P j i = + hcomp (λ k → λ { (i = i0) → p (k ∧ j) + ; (i = i1) → q (k ∧ j) + ; (j = i0) → r i + ; (j = i1) → doubleCompPath-filler∙ p s (sym q) k i}) + (P j i) + +compPathR→PathP∙∙ : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} + → r ≡ p ∙∙ s ∙∙ sym q + → PathP (λ i → p i ≡ q i) r s +compPathR→PathP∙∙ {p = p} {q = q} {r = r} {s = s} P j i = + hcomp (λ k → λ { (i = i0) → p (k ∧ j) + ; (i = i1) → q (k ∧ j) + ; (j = i0) → r i + ; (j = i1) → doubleCompPath-filler p s (sym q) (~ k) i}) + (P j i) diff --git a/Cubical/Foundations/RelationalStructure.agda b/Cubical/Foundations/RelationalStructure.agda new file mode 100644 index 000000000..75ff04e9c --- /dev/null +++ b/Cubical/Foundations/RelationalStructure.agda @@ -0,0 +1,269 @@ +{- + +Definition of what it means to be a notion of relational structure + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.RelationalStructure where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Univalence +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation as Trunc +open import Cubical.HITs.SetQuotients +open import Cubical.Relation.Binary.Base +open import Cubical.Relation.ZigZag.Base + +open BinaryRelation +open isEquivRel +open isQuasiEquivRel + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + +-- A notion of structured relation for a structure S assigns a relation on S X and S Y to every relation on X +-- and Y. We require the output to be proposition-valued when the input is proposition-valued. +StrRel : (S : Type ℓ → Type ℓ') (ℓ'' : Level) → Type (ℓ-max (ℓ-suc (ℓ-max ℓ ℓ'')) ℓ') +StrRel {ℓ = ℓ} S ℓ'' = ∀ {A B} (R : Rel A B ℓ) → Rel (S A) (S B) ℓ'' + +-- Given a type A and relation R, a quotient structure is a structure on the set quotient A/R such that +-- the graph of [_] : A → A/R is a structured relation +InducedQuotientStr : (S : Type ℓ → Type ℓ') (ρ : StrRel S ℓ'') + (A : TypeWithStr ℓ S) (R : Rel (typ A) (typ A) ℓ) + → Type (ℓ-max ℓ' ℓ'') +InducedQuotientStr S ρ A R = + Σ[ s ∈ S (typ A / R) ] ρ (graphRel [_]) (A .snd) s + +-- A structured equivalence relation R on a structured type A should induce a structure on A/R +InducesQuotientStr : (S : Type ℓ → Type ℓ') (ρ : StrRel S ℓ'') → Type _ +InducesQuotientStr {ℓ = ℓ} S ρ = + (A : TypeWithStr ℓ S) (R : EquivPropRel (typ A) ℓ) + → ρ (R .fst .fst) (A .snd) (A .snd) + → ∃![ s ∈ S (typ A / R .fst .fst) ] ρ (graphRel [_]) (A .snd) s + +-- The identity should be a structured relation +isReflexiveStrRel : {S : Type ℓ → Type ℓ'} (ρ : StrRel S ℓ'') → Type _ +isReflexiveStrRel {ℓ = ℓ} {S = S} ρ = + {X : Type ℓ} → (s : S X) → ρ (idPropRel X .fst) s s + +-- The inverse of a structured relation should be structured +isSymmetricStrRel : {S : Type ℓ → Type ℓ'} (ρ : StrRel S ℓ'') → Type _ +isSymmetricStrRel {ℓ = ℓ} {S = S} ρ = + {X Y : Type ℓ} (R : PropRel X Y ℓ) + {sx : S X} {sy : S Y} + → ρ (R .fst) sx sy + → ρ (invPropRel R .fst) sy sx + +-- The composite of structured relations should be structured +isTransitiveStrRel : {S : Type ℓ → Type ℓ'} (ρ : StrRel S ℓ'') → Type _ +isTransitiveStrRel {ℓ = ℓ} {S = S} ρ = + {X Y Z : Type ℓ} + (R : PropRel X Y ℓ) (R' : PropRel Y Z ℓ) + {sx : S X} {sy : S Y} {sz : S Z} + → ρ (R .fst) sx sy + → ρ (R' .fst) sy sz + → ρ (compPropRel R R' .fst) sx sz + +-- The type of structures on a set should be a set +preservesSetsStr : (S : Type ℓ → Type ℓ') → Type (ℓ-max (ℓ-suc ℓ) ℓ') +preservesSetsStr S = ∀ {X} → isSet X → isSet (S X) + +-- The type of structures on a prop-valued relation should be a prop +preservesPropsStrRel : {S : Type ℓ → Type ℓ'} (ρ : StrRel S ℓ'') → Type _ +preservesPropsStrRel {ℓ = ℓ} {S = S} ρ = + {X Y : Type ℓ} {R : Rel X Y ℓ} + → (∀ x y → isProp (R x y)) + → (sx : S X) (sy : S Y) + → isProp (ρ R sx sy) + +record SuitableStrRel (S : Type ℓ → Type ℓ') (ρ : StrRel S ℓ'') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) ℓ') ℓ'') + where + field + quo : InducesQuotientStr S ρ + symmetric : isSymmetricStrRel ρ + transitive : isTransitiveStrRel ρ + set : preservesSetsStr S + prop : preservesPropsStrRel ρ + +open SuitableStrRel public + +quotientPropRel : ∀ {ℓ} {A : Type ℓ} (R : Rel A A ℓ) → PropRel A (A / R) ℓ +quotientPropRel R .fst a t = [ a ] ≡ t +quotientPropRel R .snd _ _ = squash/ _ _ + +-- We can also ask for a notion of structured relations to agree with some notion of structured equivalences. + +StrRelMatchesEquiv : {S : Type ℓ → Type ℓ'} + → StrRel S ℓ'' → StrEquiv S ℓ''' → Type _ +StrRelMatchesEquiv {S = S} ρ ι = + (A B : TypeWithStr _ S) (e : typ A ≃ typ B) → + ρ (graphRel (e .fst)) (A .snd) (B .snd) ≃ ι A B e + +-- Additional conditions for a "positive" notion of structured relation + +isDetransitiveStrRel : {S : Type ℓ → Type ℓ'} (ρ : StrRel S ℓ'') → Type _ +isDetransitiveStrRel {ℓ = ℓ} {S = S} ρ = + {X Y Z : Type ℓ} + (R : PropRel X Y ℓ) (R' : PropRel Y Z ℓ) + {sx : S X} {sz : S Z} + → ρ (compPropRel R R' .fst) sx sz + → ∃[ sy ∈ S Y ] ρ (R .fst) sx sy × ρ (R' .fst) sy sz + +record StrRelAction {S : Type ℓ → Type ℓ'} (ρ : StrRel S ℓ'') + : Type (ℓ-max (ℓ-suc ℓ) (ℓ-max ℓ' ℓ'')) + where + field + actStr : ∀ {X Y} → (X → Y) → S X → S Y + actStrId : ∀ {X} (s : S X) → actStr (idfun X) s ≡ s + actRel : ∀ {X₀ Y₀ X₁ Y₁} {f : X₀ → X₁} {g : Y₀ → Y₁} + {R₀ : X₀ → Y₀ → Type ℓ} {R₁ : X₁ → Y₁ → Type ℓ} + → (∀ x y → R₀ x y → R₁ (f x) (g y)) + → ∀ sx sy → ρ R₀ sx sy → ρ R₁ (actStr f sx) (actStr g sy) + +open StrRelAction public + +strRelQuotientComparison : {S : Type ℓ → Type ℓ'} {ρ : StrRel S ℓ''} + (θ : SuitableStrRel S ρ) (α : StrRelAction ρ) + {X : Type ℓ} (R : EquivPropRel X ℓ) + → (S X / ρ (R .fst .fst)) → S (X / R .fst .fst) +strRelQuotientComparison θ α R [ s ] = α .actStr [_] s +strRelQuotientComparison {ρ = ρ} θ α R (eq/ s t r i) = + (sym leftEq ∙ rightEq) i + where + ρs : ρ (R .fst .fst) s s + ρs = + subst (λ R' → ρ R' s s) + (funExt₂ λ x x' → + hPropExt squash (R .fst .snd x x') + (Trunc.rec (R .fst .snd x x') + (λ {(_ , r , r') → R .snd .transitive _ _ _ r (R .snd .symmetric _ _ r')})) + (λ r → ∣ x' , r , R .snd .reflexive x' ∣)) + (θ .transitive (R .fst) (invPropRel (R .fst)) r (θ .symmetric (R .fst) r)) + + leftEq : θ .quo (_ , s) R ρs .fst .fst ≡ α .actStr [_] s + leftEq = + cong fst + (θ .quo (_ , s) R ρs .snd + ( α .actStr [_] s + , subst + (λ s' → ρ (graphRel [_]) s' (α .actStr [_] s)) + (α .actStrId s) + (α .actRel eq/ s s ρs) + )) + + rightEq : θ .quo (_ , s) R ρs .fst .fst ≡ α .actStr [_] t + rightEq = + cong fst + (θ .quo (_ , s) R ρs .snd + ( α .actStr [_] t + , subst + (λ s' → ρ (graphRel [_]) s' (α .actStr [_] t)) + (α .actStrId s) + (α .actRel eq/ s t r) + )) +strRelQuotientComparison θ α R (squash/ _ _ p q i j) = + θ .set squash/ _ _ + (cong (strRelQuotientComparison θ α R) p) + (cong (strRelQuotientComparison θ α R) q) + i j + +record PositiveStrRel {S : Type ℓ → Type ℓ'} {ρ : StrRel S ℓ''} (θ : SuitableStrRel S ρ) + : Type (ℓ-max (ℓ-suc ℓ) (ℓ-max ℓ' ℓ'')) + where + field + act : StrRelAction ρ + reflexive : isReflexiveStrRel ρ + detransitive : isDetransitiveStrRel ρ + quo : {X : Type ℓ} (R : EquivPropRel X ℓ) → isEquiv (strRelQuotientComparison θ act R) + +open PositiveStrRel public + +posRelReflexive : {S : Type ℓ → Type ℓ'} {ρ : StrRel S ℓ''} {θ : SuitableStrRel S ρ} + → PositiveStrRel θ + → {X : Type ℓ} (R : EquivPropRel X ℓ) + → (s : S X) → ρ (R .fst .fst) s s +posRelReflexive {ρ = ρ} σ R s = + subst + (uncurry (ρ (R .fst .fst))) + (ΣPathP (σ .act .actStrId s , σ .act .actStrId s)) + (σ .act .actRel + (λ x y → + Trunc.rec (R .fst .snd _ _) + (λ p → subst (R .fst .fst x) p (R .snd .reflexive x))) + s s + (σ .reflexive s)) + +-- Given a suitable notion of structured relation, if we have a structured quasi equivalence relation R +-- between structured types A and B, we get induced structures on the quotients A/(R ∙ R⁻¹) and B/(R⁻¹ ∙ R), +-- and the induced equivalence e : A/(R ∙ R⁻¹) ≃ B/(R⁻¹ ∙ R) is structured with respect to those quotient +-- structures. + +record QERDescends (S : Type ℓ → Type ℓ') (ρ : StrRel S ℓ'') + (A B : TypeWithStr ℓ S) (R : QuasiEquivRel (typ A) (typ B) ℓ) : Type (ℓ-max ℓ' ℓ'') + where + private + module E = QER→Equiv R + + field + quoᴸ : InducedQuotientStr S ρ A E.Rᴸ + quoᴿ : InducedQuotientStr S ρ B E.Rᴿ + rel : ρ (graphRel (E.Thm .fst)) (quoᴸ .fst) (quoᴿ .fst) + +open QERDescends + +structuredQER→structuredEquiv : {S : Type ℓ → Type ℓ'} {ρ : StrRel S ℓ''} + (θ : SuitableStrRel S ρ) + (A B : TypeWithStr ℓ S) (R : QuasiEquivRel (typ A) (typ B) ℓ) + → ρ (R .fst .fst) (A .snd) (B .snd) + → QERDescends S ρ A B R +structuredQER→structuredEquiv {ρ = ρ} θ (X , s) (Y , t) R r .quoᴸ = + θ .quo (X , s) (QER→EquivRel R) + (θ .transitive (R .fst) (invPropRel (R .fst)) r (θ .symmetric (R .fst) r)) + .fst + +structuredQER→structuredEquiv {ρ = ρ} θ (X , s) (Y , t) R r .quoᴿ = + θ .quo (Y , t) (QER→EquivRel (invQER R)) + (θ .transitive (invPropRel (R .fst)) (R .fst) (θ .symmetric (R .fst) r) r) + .fst + +structuredQER→structuredEquiv {ρ = ρ} θ (X , s) (Y , t) R r .rel = + subst (λ R' → ρ R' (quol .fst) (quor .fst)) correction + (θ .transitive (compPropRel (invPropRel (quotientPropRel E.Rᴸ)) (R .fst)) (quotientPropRel E.Rᴿ) + (θ .transitive (invPropRel (quotientPropRel E.Rᴸ)) (R .fst) + (θ .symmetric (quotientPropRel E.Rᴸ) (quol .snd)) + r) + (quor .snd)) + where + module E = QER→Equiv R + quol = structuredQER→structuredEquiv {ρ = ρ} θ (X , s) (Y , t) R r .quoᴸ + quor = structuredQER→structuredEquiv {ρ = ρ} θ (X , s) (Y , t) R r .quoᴿ + [R] = compPropRel (compPropRel (invPropRel (quotientPropRel E.Rᴸ)) (R .fst)) (quotientPropRel E.Rᴿ) + + correction : [R] .fst ≡ graphRel (E.Thm .fst) + correction = + funExt₂ λ qx qy → + (hPropExt squash (squash/ _ _) + (Trunc.rec (squash/ _ _) + (λ {(y , qr , py) → + Trunc.rec + (squash/ _ _) + (λ {(x , px , r) → + (cong (E.Thm .fst) (sym px) + ∙ E.relToFwd≡ r) + ∙ py}) + qr})) + (elimProp {B = λ qx → E.Thm .fst qx ≡ qy → [R] .fst qx qy} + (λ _ → isPropΠ λ _ → squash) + (λ x → + elimProp {B = λ qy → E.Thm .fst [ x ] ≡ qy → [R] .fst [ x ] qy} + (λ _ → isPropΠ λ _ → squash) + (λ y p → ∣ _ , ∣ _ , refl , E.fwd≡ToRel p ∣ , refl ∣) + qy) + qx)) + diff --git a/Cubical/Foundations/SIP.agda b/Cubical/Foundations/SIP.agda index 314a31686..b68515407 100644 --- a/Cubical/Foundations/SIP.agda +++ b/Cubical/Foundations/SIP.agda @@ -6,509 +6,119 @@ structure identity principle: https://www.cs.bham.ac.uk/~mhe/HoTT-UF-in-Agda-Lecture-Notes/HoTT-UF-Agda.html#sns -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.SIP where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Univalence renaming (ua-pathToEquiv to ua-pathToEquiv') open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function open import Cubical.Foundations.Path open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.FunExtEquiv open import Cubical.Foundations.Equiv -open import Cubical.Foundations.HAEquiv -open import Cubical.Foundations.HLevels open import Cubical.Data.Sigma -open import Cubical.Data.Prod.Base hiding (_×_) renaming (_×Σ_ to _×_) -private - variable - ℓ ℓ' ℓ'' ℓ''' ℓ₁ ℓ₂ ℓ₃ ℓ₄ ℓ₅ : Level - --- For technical reasons we reprove ua-pathToEquiv using the --- particular proof constructed by iso→HAEquiv. The reason is that we --- want to later be able to extract --- --- eq : ua-au (ua e) ≡ cong ua (au-ua e) --- -uaHAEquiv : (A B : Type ℓ) → HAEquiv (A ≃ B) (A ≡ B) -uaHAEquiv A B = iso→HAEquiv (iso ua pathToEquiv ua-pathToEquiv' pathToEquiv-ua) -open isHAEquiv - --- We now extract the particular proof constructed from iso→HAEquiv --- for reasons explained above. -ua-pathToEquiv : {A B : Type ℓ} (e : A ≡ B) → ua (pathToEquiv e) ≡ e -ua-pathToEquiv e = uaHAEquiv _ _ .snd .ret e +open import Cubical.Foundations.Structure public +private + variable + ℓ ℓ₁ ℓ₂ ℓ₃ ℓ₄ ℓ₅ : Level + S : Type ℓ₁ → Type ℓ₂ --- A structure is a type-family S : Type ℓ → Type ℓ', i.e. for X : Type ℓ and s : S X, the pair (X , s) --- means that X is equipped with a S-structure, which is witnessed by s. --- An S-structure should have a notion of S-homomorphism, or rather S-isomorphism. --- This will be implemented by a function ι --- that gives us for any two types with S-structure (X , s) and (Y , t) a family: --- ι (X , s) (Y , t) : (X ≃ Y) → Type ℓ'' -- Note that for any equivalence (f , e) : X ≃ Y the type ι (X , s) (Y , t) (f , e) need not to be -- a proposition. Indeed this type should correspond to the ways s and t can be identified -- as S-structures. This we call a standard notion of structure or SNS. -- We will use a different definition, but the two definitions are interchangeable. -SNS : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - → Type (ℓ-max (ℓ-max(ℓ-suc ℓ) ℓ') ℓ'') -SNS {ℓ = ℓ} S ι = ∀ {X : (Type ℓ)} (s t : S X) → ((s ≡ t) ≃ ι (X , s) (X , t) (idEquiv X)) - - --- We introduce the notation for structure preserving equivalences a bit differently, --- but this definition doesn't actually change from Escardó's notes. -_≃[_]_ : {S : Type ℓ → Type ℓ'} - → (A : Σ[ X ∈ (Type ℓ) ] (S X)) - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - → (B : Σ[ X ∈ (Type ℓ) ] (S X)) - → Type (ℓ-max ℓ ℓ'') -A ≃[ ι ] B = Σ[ f ∈ ((A .fst) ≃ (B. fst)) ] (ι A B f) - --- Before we can formulate our version of an SNS we introduce a bit of --- notation and prove a few basic results. First, we define the --- "cong-≃": -_⋆_ : (S : Type ℓ → Type ℓ') → {X Y : Type ℓ} → (X ≃ Y) → S X ≃ S Y -S ⋆ f = pathToEquiv (cong S (ua f)) - - --- Next, we prove a couple of helpful results about this ⋆ operation: -⋆-idEquiv : (S : Type ℓ → Type ℓ') (X : Type ℓ) → S ⋆ (idEquiv X) ≡ idEquiv (S X) -⋆-idEquiv S X = S ⋆ (idEquiv X) ≡⟨ cong (λ p → pathToEquiv (cong S p)) uaIdEquiv ⟩ - pathToEquiv refl ≡⟨ pathToEquivRefl ⟩ - idEquiv (S X) ∎ - -⋆-char : (S : Type ℓ → Type ℓ') {X Y : Type ℓ} (f : X ≃ Y) → ua (S ⋆ f) ≡ cong S (ua f) -⋆-char S f = ua-pathToEquiv (cong S (ua f)) - -PathP-⋆-lemma : (S : Type ℓ → Type ℓ') (A B : Σ[ X ∈ (Type ℓ) ] (S X)) (f : A .fst ≃ B .fst) - → (PathP (λ i → ua (S ⋆ f) i) (A .snd) (B .snd)) ≡ (PathP (λ i → S ((ua f) i)) (A .snd) (B .snd)) -PathP-⋆-lemma S A B f i = PathP (λ j → (⋆-char S f) i j) (A .snd) (B .snd) +SNS : (S : Type ℓ₁ → Type ℓ₂) (ι : StrEquiv S ℓ₃) → Type (ℓ-max (ℓ-max (ℓ-suc ℓ₁) ℓ₂) ℓ₃) +SNS {ℓ₁} S ι = ∀ {X : Type ℓ₁} (s t : S X) → ι (X , s) (X , t) (idEquiv X) ≃ (s ≡ t) +-- We introduce the notation for structure preserving equivalences a +-- bit differently, but this definition doesn't actually change from +-- Escardó's notes. +_≃[_]_ : (A : TypeWithStr ℓ₁ S) (ι : StrEquiv S ℓ₂) (B : TypeWithStr ℓ₁ S) → Type (ℓ-max ℓ₁ ℓ₂) +A ≃[ ι ] B = Σ[ e ∈ typ A ≃ typ B ] (ι A B e) --- Our new definition of standard notion of structure SNS' using the ⋆ notation. --- This is easier to work with than SNS wrt Glue-types -SNS' : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - → Type (ℓ-max (ℓ-max(ℓ-suc ℓ) ℓ') ℓ'') -SNS' S ι = (A B : Σ[ X ∈ (Type _) ] (S X)) → (f : A .fst ≃ B .fst) - → (equivFun (S ⋆ f) (A .snd) ≡ (B .snd)) ≃ (ι A B f) - --- We can unfold SNS' as follows: -SNS'' : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - → Type (ℓ-max (ℓ-max(ℓ-suc ℓ) ℓ') ℓ'') -SNS'' S ι = (A B : Σ[ X ∈ (Type _) ] (S X)) → (f : A .fst ≃ B .fst) - → (transport (λ i → S (ua f i)) (A .snd) ≡ (B .snd)) ≃ (ι A B f) - -SNS'≡SNS'' : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - → SNS' S ι ≡ SNS'' S ι -SNS'≡SNS'' S ι = refl - +-- The following PathP version of SNS is a bit easier to work with +-- for the proof of the SIP +UnivalentStr : (S : Type ℓ₁ → Type ℓ₂) (ι : StrEquiv S ℓ₃) → Type (ℓ-max (ℓ-max (ℓ-suc ℓ₁) ℓ₂) ℓ₃) +UnivalentStr {ℓ₁} S ι = + {A B : TypeWithStr ℓ₁ S} (e : typ A ≃ typ B) + → ι A B e ≃ PathP (λ i → S (ua e i)) (str A) (str B) -- A quick sanity-check that our definition is interchangeable with --- Escardó's. The direction SNS→SNS' corresponds more or less to an --- EquivJ formulation of Escardó's homomorphism-lemma. -SNS'→SNS : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - → SNS' S ι → SNS S ι -SNS'→SNS S ι θ {X = X} s t = subst (λ x → (equivFun x s ≡ t) ≃ ι (X , s) (X , t) (idEquiv X)) (⋆-idEquiv S X) θ-id +-- Escardó's. The direction SNS→UnivalentStr corresponds more or less +-- to a dependent EquivJ formulation of Escardó's homomorphism-lemma. +UnivalentStr→SNS : (S : Type ℓ₁ → Type ℓ₂) (ι : StrEquiv S ℓ₃) + → UnivalentStr S ι → SNS S ι +UnivalentStr→SNS S ι θ {X = X} s t = + ι (X , s) (X , t) (idEquiv X) + ≃⟨ θ (idEquiv X) ⟩ + PathP (λ i → S (ua (idEquiv X) i)) s t + ≃⟨ transportEquiv (λ j → PathP (λ i → S (uaIdEquiv {A = X} j i)) s t) ⟩ + s ≡ t + ■ + + +SNS→UnivalentStr : (ι : StrEquiv S ℓ₃) → SNS S ι → UnivalentStr S ι +SNS→UnivalentStr {S = S} ι θ {A = A} {B = B} e = EquivJ P C e (str A) (str B) where - θ-id = θ (X , s) (X , t) (idEquiv X) - -SNS→SNS' : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - → SNS S ι → SNS' S ι -SNS→SNS' S ι θ A B f = EquivJ P C (B .fst) (A .fst) f (B .snd) (A .snd) + Y = typ B + + P : (X : Type _) → X ≃ Y → Type _ + P X e' = (s : S X) (t : S Y) → ι (X , s) (Y , t) e' ≃ PathP (λ i → S (ua e' i)) s t + + C : (s t : S Y) → ι (Y , s) (Y , t) (idEquiv Y) ≃ PathP (λ i → S (ua (idEquiv Y) i)) s t + C s t = + ι (Y , s) (Y , t) (idEquiv Y) + ≃⟨ θ s t ⟩ + s ≡ t + ≃⟨ transportEquiv (λ j → PathP (λ i → S (uaIdEquiv {A = Y} (~ j) i)) s t) ⟩ + PathP (λ i → S (ua (idEquiv Y) i)) s t + ■ + +TransportStr : {S : Type ℓ → Type ℓ₁} (α : EquivAction S) → Type (ℓ-max (ℓ-suc ℓ) ℓ₁) +TransportStr {ℓ} {S = S} α = + {X Y : Type ℓ} (e : X ≃ Y) (s : S X) → equivFun (α e) s ≡ subst S (ua e) s + +TransportStr→UnivalentStr : {S : Type ℓ → Type ℓ₁} (α : EquivAction S) + → TransportStr α → UnivalentStr S (EquivAction→StrEquiv α) +TransportStr→UnivalentStr {S = S} α τ {X , s} {Y , t} e = + equivFun (α e) s ≡ t + ≃⟨ pathToEquiv (cong (_≡ t) (τ e s)) ⟩ + subst S (ua e) s ≡ t + ≃⟨ invEquiv (PathP≃Path _ _ _) ⟩ + PathP (λ i → S (ua e i)) s t + ■ + +UnivalentStr→TransportStr : {S : Type ℓ → Type ℓ₁} (α : EquivAction S) + → UnivalentStr S (EquivAction→StrEquiv α) → TransportStr α +UnivalentStr→TransportStr {S = S} α θ e s = + invEq (θ e) (transport-filler (cong S (ua e)) s) + +invTransportStr : {S : Type ℓ → Type ℓ₂} (α : EquivAction S) (τ : TransportStr α) + {X Y : Type ℓ} (e : X ≃ Y) (t : S Y) → invEq (α e) t ≡ subst⁻ S (ua e) t +invTransportStr {S = S} α τ e t = + sym (transport⁻Transport (cong S (ua e)) (invEq (α e) t)) + ∙∙ sym (cong (subst⁻ S (ua e)) (τ e (invEq (α e) t))) + ∙∙ cong (subst⁻ S (ua e)) (secEq (α e) t) + + +--- We can now define an invertible function +--- +--- sip : A ≃[ ι ] B → A ≡ B + +module _ {S : Type ℓ₁ → Type ℓ₂} {ι : StrEquiv S ℓ₃} + (θ : UnivalentStr S ι) (A B : TypeWithStr ℓ₁ S) where - P : (X Y : Type _) → Y ≃ X → Type _ - P X Y g = (s : S X) (t : S Y) → (equivFun (S ⋆ g) t ≡ s) ≃ ι (Y , t) (X , s) g - - C : (X : Type _) → (s t : S X) → (equivFun (S ⋆ (idEquiv X)) t ≡ s) ≃ ι (X , t) (X , s) (idEquiv X) - C X s t = subst (λ u → (u ≡ s) ≃ (ι (X , t) (X , s) (idEquiv X))) - (sym ( cong (λ f → (equivFun f) t) (⋆-idEquiv S X))) (θ t s) - - - --- The following transport-free version of SNS'' is a bit easier to --- work with for the proof of the SIP -SNS''' : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - → Type (ℓ-max (ℓ-max(ℓ-suc ℓ) ℓ') ℓ'') -SNS''' S ι = (A B : Σ[ X ∈ (Type _) ] (S X)) → (e : A .fst ≃ B .fst) - → (PathP (λ i → S (ua e i)) (A .snd) (B .snd)) ≃ ι A B e - --- We can easily go between SNS'' (which is def. equal to SNS') and SNS''' --- We should be able to find are more direct version of PathP≃Path for the family (λ i → S (ua f i)) --- using glue and unglue terms. -SNS''→SNS''' : {S : Type ℓ → Type ℓ'} - → {ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ''} - → SNS'' S ι → SNS''' S ι -SNS''→SNS''' {S = S} h A B f = PathP (λ i → S (ua f i)) (A .snd) (B .snd) - ≃⟨ PathP≃Path (λ i → S (ua f i)) (A .snd) (B .snd) ⟩ - h A B f - -SNS'''→SNS'' : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - → SNS''' S ι → SNS'' S ι -SNS'''→SNS'' S ι h A B f = transport (λ i → S (ua f i)) (A .snd) ≡ (B .snd) - ≃⟨ invEquiv (PathP≃Path (λ i → S (ua f i)) (A .snd) (B .snd)) ⟩ - h A B f - - --- We can now directly define a function --- sip : A ≃[ ι ] B → A ≡ B --- together with is inverse. --- Here, these functions use SNS''' and are expressed using a Σ-type instead as it is a bit --- easier to work with -sip : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - → (θ : SNS''' S ι) - → (A B : Σ[ X ∈ (Type ℓ) ] (S X)) - → A ≃[ ι ] B - → Σ (A .fst ≡ B .fst) (λ p → PathP (λ i → S (p i)) (A .snd) (B .snd)) -sip S ι θ A B (e , p) = ua e , invEq (θ A B e) p - --- The inverse to sip using the following little lemma -lem : (S : Type ℓ → Type ℓ') - (A B : Σ[ X ∈ (Type ℓ) ] (S X)) - (e : A .fst ≡ B .fst) - → PathP (λ i → S (ua (pathToEquiv e) i)) (A .snd) (B .snd) ≡ - PathP (λ i → S (e i)) (A .snd) (B .snd) -lem S A B e i = PathP (λ j → S (ua-pathToEquiv e i j)) (A .snd) (B .snd) - -sip⁻ : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - → (θ : SNS''' S ι) - → (A B : Σ[ X ∈ (Type ℓ) ] (S X)) - → Σ (A .fst ≡ B .fst) (λ p → PathP (λ i → S (p i)) (A .snd) (B .snd)) - → A ≃[ ι ] B -sip⁻ S ι θ A B (e , r) = pathToEquiv e , θ A B (pathToEquiv e) .fst q - where - q : PathP (λ i → S (ua (pathToEquiv e) i)) (A .snd) (B .snd) - q = transport (λ i → lem S A B e (~ i)) r - - --- we can rather directly show that sip and sip⁻ are mutually inverse: -sip-sip⁻ : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - → (θ : SNS''' S ι) - → (A B : Σ[ X ∈ (Type ℓ) ] (S X)) - → (r : Σ (A .fst ≡ B .fst) (λ p → PathP (λ i → S (p i)) (A .snd) (B .snd))) - → sip S ι θ A B (sip⁻ S ι θ A B r) ≡ r -sip-sip⁻ S ι θ A B (p , q) = - sip S ι θ A B (sip⁻ S ι θ A B (p , q)) - ≡⟨ refl ⟩ - ua (pathToEquiv p) , invEq (θ A B (pathToEquiv p)) (θ A B (pathToEquiv p) .fst (transport (λ i → lem S A B p (~ i)) q)) - ≡⟨ (λ i → ua (pathToEquiv p) , secEq (θ A B (pathToEquiv p)) (transport (λ i → lem S A B p (~ i)) q) i) ⟩ - ua (pathToEquiv p) , transport (λ i → lem S A B p (~ i)) q - ≡⟨ (λ i → ua-pathToEquiv p i , - transp (λ k → PathP (λ j → S (ua-pathToEquiv p (i ∧ k) j)) (A .snd) (B .snd)) - (~ i) - (transport (λ i → lem S A B p (~ i)) q)) ⟩ - p , transport (λ i → lem S A B p i) (transport (λ i → lem S A B p (~ i)) q) - ≡⟨ (λ i → p , transportTransport⁻ (lem S A B p) q i) ⟩ - p , q ∎ - - --- The trickier direction: -sip⁻-sip : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - → (θ : SNS''' S ι) - → (A B : Σ[ X ∈ (Type ℓ) ] (S X)) - → (r : A ≃[ ι ] B) - → sip⁻ S ι θ A B (sip S ι θ A B r) ≡ r -sip⁻-sip S ι θ A B (e , p) = - sip⁻ S ι θ A B (sip S ι θ A B (e , p)) - ≡⟨ refl ⟩ - pathToEquiv (ua e) , θ A B (pathToEquiv (ua e)) .fst (f⁺ p') - ≡⟨ (λ i → pathToEquiv-ua e i , θ A B (pathToEquiv-ua e i) .fst (pth' i)) ⟩ - e , θ A B e .fst (f⁻ (f⁺ p')) - ≡⟨ (λ i → e , θ A B e .fst (transportTransport⁻ (lem S A B (ua e)) p' i)) ⟩ - e , θ A B e .fst (invEq (θ A B e) p) - ≡⟨ (λ i → e , (retEq (θ A B e) p i)) ⟩ - e , p ∎ - where - p' : PathP (λ i → S (ua e i)) (A .snd) (B .snd) - p' = invEq (θ A B e) p - - f⁺ : PathP (λ i → S (ua e i)) (A .snd) (B .snd) - → PathP (λ i → S (ua (pathToEquiv (ua e)) i)) (A .snd) (B .snd) - f⁺ = transport (λ i → PathP (λ j → S (ua-pathToEquiv (ua e) (~ i) j)) (A .snd) (B .snd)) - - f⁻ : PathP (λ i → S (ua (pathToEquiv (ua e)) i)) (A .snd) (B .snd) - → PathP (λ i → S (ua e i)) (A .snd) (B .snd) - f⁻ = transport (λ i → PathP (λ j → S (ua-pathToEquiv (ua e) i j)) (A .snd) (B .snd)) - - -- We can prove the following as in sip∘pis-id, but the type is not - -- what we want as it should be "cong ua (pathToEquiv-ua e)" - pth : PathP (λ j → PathP (λ k → S (ua-pathToEquiv (ua e) j k)) (A .snd) (B .snd)) - (f⁺ p') (f⁻ (f⁺ p')) - pth i = transp (λ k → PathP (λ j → S (ua-pathToEquiv (ua e) (i ∧ k) j)) (A .snd) (B .snd)) - (~ i) - (f⁺ p') - - -- So we build an equality that we want to cast the types with - casteq : PathP (λ j → PathP (λ k → S (ua-pathToEquiv (ua e) j k)) (A .snd) (B .snd)) - (f⁺ p') (f⁻ (f⁺ p')) - ≡ PathP (λ j → PathP (λ k → S (cong ua (pathToEquiv-ua e) j k)) (A .snd) (B .snd)) - (f⁺ p') (f⁻ (f⁺ p')) - casteq i = PathP (λ j → PathP (λ k → S (eq i j k)) (A .snd) (B .snd)) (f⁺ p') (f⁻ (f⁺ p')) - where - -- This is where we need the half-adjoint equivalence property - eq : ua-pathToEquiv (ua e) ≡ cong ua (pathToEquiv-ua e) - eq = sym (uaHAEquiv (A .fst) (B .fst) .snd .com e) - - -- We then get a term of the type we need - pth' : PathP (λ j → PathP (λ k → S (cong ua (pathToEquiv-ua e) j k)) (A .snd) (B .snd)) - (f⁺ p') (f⁻ (f⁺ p')) - pth' = transport (λ i → casteq i) pth - - --- Finally package everything up to get the cubical SIP -SIP : (S : Type ℓ → Type ℓ') - → (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - → (θ : SNS''' S ι) - → (A B : Σ[ X ∈ (Type ℓ) ] (S X)) - → A ≃[ ι ] B ≃ (A ≡ B) -SIP S ι θ A B = (A ≃[ ι ] B ) ≃⟨ eq ⟩ Σ≡ - where - eq : A ≃[ ι ] B ≃ Σ (A .fst ≡ B .fst) (λ p → PathP (λ i → S (p i)) (A .snd) (B .snd)) - eq = isoToEquiv (iso (sip S ι θ A B) (sip⁻ S ι θ A B) - (sip-sip⁻ S ι θ A B) (sip⁻-sip S ι θ A B)) - - --- Now, we want to add axioms (i.e. propositions) to our Structure S that don't affect the ι. --- For this and the remainder of this file we will work with SNS' --- We use a lemma due to Zesen Qian, which can now be found in Foundations.Prelude: --- https://github.com/riaqn/cubical/blob/hgroup/Cubical/Data/Group/Properties.agda#L83 -add-to-structure : (S : Type ℓ → Type ℓ') - (axioms : (X : Type ℓ) → (S X) → Type ℓ''') - → Type ℓ → Type (ℓ-max ℓ' ℓ''') -add-to-structure S axioms X = Σ[ s ∈ S X ] (axioms X s) - -add-to-iso : (S : Type ℓ → Type ℓ') - (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - (axioms : (X : Type ℓ) → (S X) → Type ℓ''') - → (A B : Σ[ X ∈ (Type ℓ) ] (add-to-structure S axioms X)) → A .fst ≃ B .fst - → Type ℓ'' -add-to-iso S ι axioms (X , (s , a)) (Y , (t , b)) f = ι (X , s) (Y , t) f - -add-⋆-lemma : (S : Type ℓ → Type ℓ') - (axioms : (X : Type ℓ) → (S X) → Type ℓ''') - (axioms-are-Props : (X : Type ℓ) (s : S X) → isProp (axioms X s)) - {X Y : Type ℓ} {s : S X} {t : S Y} {a : axioms X s} {b : axioms Y t} - (f : X ≃ Y) - → (equivFun (add-to-structure S axioms ⋆ f) (s , a) ≡ (t , b)) ≃ (equivFun (S ⋆ f) s ≡ t) -add-⋆-lemma S axioms axioms-are-Props {Y = Y} {s = s} {t = t} {a = a} {b = b} f = isoToEquiv (iso φ ψ η ε) - where - φ : equivFun ((add-to-structure S axioms) ⋆ f) (s , a) ≡ (t , b) - → equivFun (S ⋆ f) s ≡ t - φ r i = r i .fst - - ψ : equivFun (S ⋆ f) s ≡ t - → equivFun ((add-to-structure S axioms) ⋆ f) (s , a) ≡ (t , b) - ψ p i = p i , isProp→PathP (λ j → axioms-are-Props Y (p j)) (equivFun (add-to-structure S axioms ⋆ f) (s , a) .snd) b i - - η : section φ ψ - η p = refl - - ε : retract φ ψ - ε r i j = r j .fst - , isProp→isSet-PathP (λ k → axioms-are-Props Y (r k .fst)) _ _ - (isProp→PathP (λ j → axioms-are-Props Y (r j .fst)) - (equivFun (add-to-structure S axioms ⋆ f) (s , a) .snd) b) - (λ k → r k .snd) i j - - -add-axioms-SNS' : (S : Type ℓ → Type ℓ') - (ι : (A B : Σ[ X ∈ (Type ℓ) ] (S X)) → A .fst ≃ B .fst → Type ℓ'') - (axioms : (X : Type ℓ) → (S X) → Type ℓ''') - (axioms-are-Props : (X : Type ℓ) (s : S X) → isProp (axioms X s)) - (θ : SNS' S ι) - → SNS' (add-to-structure S axioms) (add-to-iso S ι axioms) -add-axioms-SNS' S ι axioms axioms-are-Props θ (X , (s , a)) (Y , (t , b)) f = - equivFun (add-to-structure S axioms ⋆ f) (s , a) ≡ (t , b) ≃⟨ add-⋆-lemma S axioms axioms-are-Props f ⟩ - equivFun (S ⋆ f) s ≡ t ≃⟨ θ (X , s) (Y , t) f ⟩ - add-to-iso S ι axioms (X , (s , a)) (Y , (t , b)) f ■ - - --- Now, we want to join two structures. Together with the adding of --- axioms this will allow us to prove that a lot of mathematical --- structures are a standard notion of structure -technical-×-lemma : {A : Type ℓ₁} {B : Type ℓ₂} {C : Type ℓ₃} {D : Type ℓ₄} - → A ≃ C → B ≃ D → A × B ≃ C × D -technical-×-lemma {A = A} {B = B} {C = C} {D = D} f g = isoToEquiv (iso φ ψ η ε) - where - φ : (A × B) → (C × D) - φ (a , b) = equivFun f a , equivFun g b - - ψ : (C × D) → (A × B) - ψ (c , d) = equivFun (invEquiv f) c , equivFun (invEquiv g) d - - η : section φ ψ - η (c , d) i = retEq f c i , retEq g d i - - ε : retract φ ψ - ε (a , b) i = secEq f a i , secEq g b i - - -join-structure : (S₁ : Type ℓ₁ → Type ℓ₂) (S₂ : Type ℓ₁ → Type ℓ₄) - → Type ℓ₁ → Type (ℓ-max ℓ₂ ℓ₄) -join-structure S₁ S₂ X = S₁ X × S₂ X - -join-iso : {S₁ : Type ℓ₁ → Type ℓ₂} - (ι₁ : (A B : Σ[ X ∈ (Type ℓ₁) ] (S₁ X)) → A .fst ≃ B .fst → Type ℓ₃) - {S₂ : Type ℓ₁ → Type ℓ₄} - (ι₂ : (A B : Σ[ X ∈ (Type ℓ₁) ] (S₂ X)) → A .fst ≃ B .fst → Type ℓ₅) - (A B : Σ[ X ∈ (Type ℓ₁) ] (join-structure S₁ S₂ X)) - → A .fst ≃ B .fst - → Type (ℓ-max ℓ₃ ℓ₅) -join-iso ι₁ ι₂ (X , s₁ , s₂) (Y , t₁ , t₂) f = (ι₁ (X , s₁) (Y , t₁) f) × (ι₂ (X , s₂) (Y , t₂) f) - -join-⋆-lemma : (S₁ : Type ℓ₁ → Type ℓ₂) (S₂ : Type ℓ₁ → Type ℓ₄) - {X Y : Type ℓ₁} {s₁ : S₁ X} {s₂ : S₂ X} {t₁ : S₁ Y} {t₂ : S₂ Y} (f : X ≃ Y) - → (equivFun (join-structure S₁ S₂ ⋆ f) (s₁ , s₂) ≡ (t₁ , t₂)) ≃ - (equivFun (S₁ ⋆ f) s₁ ≡ t₁) × (equivFun (S₂ ⋆ f) s₂ ≡ t₂) -join-⋆-lemma S₁ S₂ {Y = Y} {s₁ = s₁} {s₂ = s₂} {t₁ = t₁} {t₂ = t₂} f = isoToEquiv (iso φ ψ η ε) - where - φ : equivFun (join-structure S₁ S₂ ⋆ f) (s₁ , s₂) ≡ (t₁ , t₂) - → (equivFun (S₁ ⋆ f) s₁ ≡ t₁) × (equivFun (S₂ ⋆ f) s₂ ≡ t₂) - φ p = (λ i → p i .fst) , (λ i → p i .snd) - - ψ : (equivFun (S₁ ⋆ f) s₁ ≡ t₁) × (equivFun (S₂ ⋆ f) s₂ ≡ t₂) - → equivFun (join-structure S₁ S₂ ⋆ f) (s₁ , s₂) ≡ (t₁ , t₂) - ψ (p , q) i = (p i) , (q i) - - η : section φ ψ - η (p , q) = refl - - ε : retract φ ψ - ε p = refl - -join-SNS' : (S₁ : Type ℓ₁ → Type ℓ₂) - (ι₁ : (A B : Σ[ X ∈ (Type ℓ₁) ] (S₁ X)) → A .fst ≃ B .fst → Type ℓ₃) - (θ₁ : SNS' S₁ ι₁) - (S₂ : Type ℓ₁ → Type ℓ₄) - (ι₂ : (A B : Σ[ X ∈ (Type ℓ₁) ] (S₂ X)) → A .fst ≃ B .fst → Type ℓ₅) - (θ₂ : SNS' S₂ ι₂) - → SNS' (join-structure S₁ S₂) (join-iso ι₁ ι₂) -join-SNS' S₁ ι₁ θ₁ S₂ ι₂ θ₂ (X , s₁ , s₂) (Y , t₁ , t₂) f = - equivFun (join-structure S₁ S₂ ⋆ f) (s₁ , s₂) ≡ (t₁ , t₂) - ≃⟨ join-⋆-lemma S₁ S₂ f ⟩ - (equivFun (S₁ ⋆ f) s₁ ≡ t₁) × (equivFun (S₂ ⋆ f) s₂ ≡ t₂) - ≃⟨ technical-×-lemma (θ₁ (X , s₁) (Y , t₁) f) (θ₂ (X , s₂) (Y , t₂) f) ⟩ - join-iso ι₁ ι₂ (X , s₁ , s₂) (Y , t₁ , t₂) f ■ - - --- Now, we want to do Monoids as an example. We begin by showing SNS' for simple structures, i.e. pointed types and ∞-magmas. --- Pointed types with SNS' -pointed-structure : Type ℓ → Type ℓ -pointed-structure X = X - -Pointed-Type : Type (ℓ-suc ℓ) -Pointed-Type {ℓ = ℓ} = Σ (Type ℓ) pointed-structure - -pointed-iso : (A B : Pointed-Type) → A .fst ≃ B .fst → Type ℓ -pointed-iso A B f = equivFun f (A .snd) ≡ B .snd - -pointed-is-SNS' : SNS' {ℓ = ℓ} pointed-structure pointed-iso -pointed-is-SNS' A B f = transportEquiv (λ i → transportRefl (equivFun f (A .snd)) i ≡ B .snd) - - --- ∞-Magmas with SNS' --- We need function extensionality for binary functions. --- TODO: upstream -funExtBin : {A : Type ℓ} {B : A → Type ℓ'} {C : (x : A) → B x → Type ℓ''} {f g : (x : A) → (y : B x) → C x y} - → ((x : A) (y : B x) → f x y ≡ g x y) → f ≡ g -funExtBin p i x y = p x y i -module _ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : A → Type ℓ'} {C : (x : A) → B x → Type ℓ''} {f g : (x : A) → (y : B x) → C x y} where - private - appl₂ : f ≡ g → ∀ x y → f x y ≡ g x y - appl₂ eq x y i = eq i x y - - fib : (p : f ≡ g) → fiber funExtBin p - fib p = (appl₂ p , refl) - - funExtBin-fiber-isContr - : (p : f ≡ g) - → (fi : fiber funExtBin p) - → fib p ≡ fi - funExtBin-fiber-isContr p (h , eq) i = (appl₂ (eq (~ i)) , λ j → eq (~ i ∨ j)) - - funExtBin-isEquiv : isEquiv funExtBin - equiv-proof funExtBin-isEquiv p = (fib p , funExtBin-fiber-isContr p) - - funExtBinEquiv : (∀ x y → f x y ≡ g x y) ≃ (f ≡ g) - funExtBinEquiv = (funExtBin , funExtBin-isEquiv) - --- ∞-Magmas -∞-magma-structure : Type ℓ → Type ℓ -∞-magma-structure X = X → X → X - -∞-Magma : Type (ℓ-suc ℓ) -∞-Magma {ℓ = ℓ} = Σ (Type ℓ) ∞-magma-structure - -∞-magma-iso : (A B : ∞-Magma) → A .fst ≃ B .fst → Type ℓ -∞-magma-iso (X , _·_) (Y , _∗_) f = (x x' : X) → equivFun f (x · x') ≡ (equivFun f x) ∗ (equivFun f x') - -∞-magma-is-SNS' : SNS' {ℓ = ℓ} ∞-magma-structure ∞-magma-iso -∞-magma-is-SNS' (X , _·_) (Y , _∗_) f = SNS→SNS' ∞-magma-structure ∞-magma-iso C (X , _·_) (Y , _∗_) f - where - C : {X : Type ℓ} (_·_ _∗_ : X → X → X) → (_·_ ≡ _∗_) ≃ ((x x' : X) → (x · x') ≡ (x ∗ x')) - C _·_ _∗_ = invEquiv funExtBinEquiv - - - --- Now we're getting serious: Monoids -raw-monoid-structure : Type ℓ → Type ℓ -raw-monoid-structure X = X × (X → X → X) - - -raw-monoid-iso : (M N : Σ (Type ℓ) raw-monoid-structure) → (M .fst) ≃ (N .fst) → Type ℓ -raw-monoid-iso (M , e , _·_) (N , d , _∗_) f = (equivFun f e ≡ d) - × ((x y : M) → equivFun f (x · y) ≡ (equivFun f x) ∗ (equivFun f y)) - --- If we ignore the axioms we get something like a "raw" monoid, which essentially is the join of a pointed type and an ∞-magma -raw-monoid-is-SNS' : SNS' {ℓ = ℓ} raw-monoid-structure raw-monoid-iso -raw-monoid-is-SNS' = join-SNS' pointed-structure pointed-iso pointed-is-SNS' - ∞-magma-structure ∞-magma-iso ∞-magma-is-SNS' - --- Now define monoids -monoid-axioms : (X : Type ℓ) → raw-monoid-structure X → Type ℓ -monoid-axioms X (e , _·_ ) = isSet X - × ((x y z : X) → (x · (y · z)) ≡ ((x · y) · z)) - × ((x : X) → (x · e) ≡ x) - × ((x : X) → (e · x) ≡ x) - -monoid-structure : Type ℓ → Type ℓ -monoid-structure = add-to-structure (raw-monoid-structure) monoid-axioms - --- TODO: it might be nicer to formulate the SIP lemmas so that they're --- easier to use for things that are not "completely packaged" -Monoids : Type (ℓ-suc ℓ) -Monoids = Σ (Type _) monoid-structure - -monoid-iso : (M N : Monoids) → M .fst ≃ N .fst → Type ℓ -monoid-iso = add-to-iso raw-monoid-structure raw-monoid-iso monoid-axioms --- We have to show that the monoid axioms are indeed Propositions -monoid-axioms-are-Props : (X : Type ℓ) (s : raw-monoid-structure X) → isProp (monoid-axioms X s) -monoid-axioms-are-Props X (e , _·_) s = β s - where - α = s .fst - β = isOfHLevelΣ 1 isPropIsSet - λ _ → isOfHLevelΣ 1 (isOfHLevelPi 1 (λ x → isOfHLevelPi 1 λ y → isOfHLevelPi 1 λ z → α (x · (y · z)) ((x · y) · z))) - λ _ → isOfHLevelΣ 1 (isOfHLevelPi 1 λ x → α (x · e) x) - λ _ → isOfHLevelPi 1 λ x → α (e · x) x + sip : A ≃[ ι ] B → A ≡ B + sip (e , p) i = ua e i , θ e .fst p i -monoid-is-SNS' : SNS' {ℓ = ℓ} monoid-structure monoid-iso -monoid-is-SNS' = add-axioms-SNS' raw-monoid-structure raw-monoid-iso - monoid-axioms monoid-axioms-are-Props raw-monoid-is-SNS' + SIP : A ≃[ ι ] B ≃ (A ≡ B) + SIP = + sip , isoToIsEquiv (compIso (Σ-cong-iso (invIso univalenceIso) (equivToIso ∘ θ)) ΣPathIsoPathΣ) -MonoidPath : (M N : Monoids {ℓ}) → (M ≃[ monoid-iso ] N) ≃ (M ≡ N) -MonoidPath M N = SIP monoid-structure monoid-iso (SNS''→SNS''' monoid-is-SNS') M N + sip⁻ : A ≡ B → A ≃[ ι ] B + sip⁻ = invEq SIP diff --git a/Cubical/Foundations/Structure.agda b/Cubical/Foundations/Structure.agda new file mode 100644 index 000000000..5f0b31c08 --- /dev/null +++ b/Cubical/Foundations/Structure.agda @@ -0,0 +1,44 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Structure where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude + +private + variable + ℓ ℓ' ℓ'' : Level + S : Type ℓ → Type ℓ' + +-- A structure is a type-family S : Type ℓ → Type ℓ', i.e. for X : Type ℓ and s : S X, +-- the pair (X , s) : TypeWithStr ℓ S means that X is equipped with a S-structure, witnessed by s. + +TypeWithStr : (ℓ : Level) (S : Type ℓ → Type ℓ') → Type (ℓ-max (ℓ-suc ℓ) ℓ') +TypeWithStr ℓ S = Σ[ X ∈ Type ℓ ] S X + +typ : TypeWithStr ℓ S → Type ℓ +typ = fst + +str : (A : TypeWithStr ℓ S) → S (typ A) +str = snd + +-- Alternative notation for typ +⟨_⟩ : TypeWithStr ℓ S → Type ℓ +⟨_⟩ = typ + +-- An S-structure should have a notion of S-homomorphism, or rather S-isomorphism. +-- This will be implemented by a function ι : StrEquiv S ℓ' +-- that gives us for any two types with S-structure (X , s) and (Y , t) a family: +-- ι (X , s) (Y , t) : (X ≃ Y) → Type ℓ'' +StrEquiv : (S : Type ℓ → Type ℓ'') (ℓ' : Level) → Type (ℓ-max (ℓ-suc (ℓ-max ℓ ℓ')) ℓ'') +StrEquiv {ℓ} S ℓ' = (A B : TypeWithStr ℓ S) → typ A ≃ typ B → Type ℓ' + +-- An S-structure may instead be equipped with an action on equivalences, which will +-- induce a notion of S-homomorphism + +EquivAction : (S : Type ℓ → Type ℓ'') → Type (ℓ-max (ℓ-suc ℓ) ℓ'') +EquivAction {ℓ} S = {X Y : Type ℓ} → X ≃ Y → S X ≃ S Y + +EquivAction→StrEquiv : {S : Type ℓ → Type ℓ''} + → EquivAction S → StrEquiv S ℓ'' +EquivAction→StrEquiv α (X , s) (Y , t) e = equivFun (α e) s ≡ t + diff --git a/Cubical/Foundations/Surjection.agda b/Cubical/Foundations/Surjection.agda deleted file mode 100644 index 6972630d5..000000000 --- a/Cubical/Foundations/Surjection.agda +++ /dev/null @@ -1,53 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.Foundations.Surjection where - -open import Cubical.Core.Everything -open import Cubical.Data.Prod -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Embedding -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv -open import Cubical.HITs.PropositionalTruncation - -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - f : A → B - -isSurjection : (A → B) → Type _ -isSurjection f = ∀ b → ∥ fiber f b ∥ - -section→isSurjection : {g : B → A} → section f g → isSurjection f -section→isSurjection {g = g} s b = ∣ g b , s b ∣ - -isSurjectionIsProp : isProp (isSurjection f) -isSurjectionIsProp = isPropPi λ _ → squash - -isEquiv→isSurjection : isEquiv f → isSurjection f -isEquiv→isSurjection e b = ∣ fst (equiv-proof e b) ∣ - -isEquiv→isEmbedding×isSurjection : isEquiv f → isEmbedding f × isSurjection f -isEquiv→isEmbedding×isSurjection e = isEquiv→isEmbedding e , isEquiv→isSurjection e - -isEmbedding×isSurjection→isEquiv : isEmbedding f × isSurjection f → isEquiv f -equiv-proof (isEmbedding×isSurjection→isEquiv {f = f} (emb , sur)) b = - inhProp→isContr (recPropTrunc fib' (λ x → x) fib) fib' - where - hpf : hasPropFibers f - hpf = isEmbedding→hasPropFibers emb - - fib : ∥ fiber f b ∥ - fib = sur b - - fib' : isProp (fiber f b) - fib' = hpf b - -isEquiv≃isEmbedding×isSurjection : isEquiv f ≃ isEmbedding f × isSurjection f -isEquiv≃isEmbedding×isSurjection = isoToEquiv (iso - isEquiv→isEmbedding×isSurjection - isEmbedding×isSurjection→isEquiv - (λ _ → hLevelProd 1 isEmbeddingIsProp isSurjectionIsProp _ _) - (λ _ → isPropIsEquiv _ _ _)) diff --git a/Cubical/Foundations/TotalFiber.agda b/Cubical/Foundations/TotalFiber.agda deleted file mode 100644 index ef5d768eb..000000000 --- a/Cubical/Foundations/TotalFiber.agda +++ /dev/null @@ -1,65 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.Foundations.TotalFiber where - -open import Cubical.Core.Everything -open import Cubical.Data.Prod -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Surjection -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv -open import Cubical.HITs.PropositionalTruncation - -private - variable - ℓ ℓ' : Level - -module _ {A : Type ℓ} {B : Type ℓ'} (f : A → B) where - private - ℓ'' : Level - ℓ'' = ℓ-max ℓ ℓ' - - LiftA : Type ℓ'' - LiftA = Lift {j = ℓ'} A - - Total-fiber : Type ℓ'' - Total-fiber = Σ B \ b → fiber f b - - A→Total-fiber : A → Total-fiber - A→Total-fiber a = f a , a , refl - - Total-fiber→A : Total-fiber → A - Total-fiber→A (b , a , p) = a - - Total-fiber→A→Total-fiber : ∀ t → A→Total-fiber (Total-fiber→A t) ≡ t - Total-fiber→A→Total-fiber (b , a , p) i = p i , a , λ j → p (i ∧ j) - - A→Total-fiber→A : ∀ a → Total-fiber→A (A→Total-fiber a) ≡ a - A→Total-fiber→A a = refl - - LiftA→Total-fiber : LiftA → Total-fiber - LiftA→Total-fiber x = A→Total-fiber (lower x) - - Total-fiber→LiftA : Total-fiber → LiftA - Total-fiber→LiftA x = lift (Total-fiber→A x) - - Total-fiber→LiftA→Total-fiber : ∀ t → LiftA→Total-fiber (Total-fiber→LiftA t) ≡ t - Total-fiber→LiftA→Total-fiber (b , a , p) i = p i , a , λ j → p (i ∧ j) - - LiftA→Total-fiber→LiftA : ∀ a → Total-fiber→LiftA (LiftA→Total-fiber a) ≡ a - LiftA→Total-fiber→LiftA a = refl - - A≡Total : Lift A ≡ Total-fiber - A≡Total = isoToPath (iso - LiftA→Total-fiber - Total-fiber→LiftA - Total-fiber→LiftA→Total-fiber - LiftA→Total-fiber→LiftA) - - -- HoTT Lemma 4.8.2 - A≃Total : Lift A ≃ Total-fiber - A≃Total = isoToEquiv (iso - LiftA→Total-fiber - Total-fiber→LiftA - Total-fiber→LiftA→Total-fiber - LiftA→Total-fiber→LiftA) diff --git a/Cubical/Foundations/Transport.agda b/Cubical/Foundations/Transport.agda index 0ec2ea036..3fdf2f1dc 100644 --- a/Cubical/Foundations/Transport.agda +++ b/Cubical/Foundations/Transport.agda @@ -4,13 +4,15 @@ - transport is an equivalence ([transportEquiv]) -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Transport where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Function using (_∘_) -- Direct definition of transport filler, note that we have to -- explicitly tell Agda that the type is constant (like in CHM) @@ -25,40 +27,90 @@ transpFill φ A u0 i = transp (λ j → outS (A (i ∧ j))) (~ i ∨ φ) u0 transport⁻ : ∀ {ℓ} {A B : Type ℓ} → A ≡ B → B → A transport⁻ p = transport (λ i → p (~ i)) +subst⁻ : ∀ {ℓ ℓ'} {A : Type ℓ} {x y : A} (B : A → Type ℓ') (p : x ≡ y) → B y → B x +subst⁻ B p pa = transport⁻ (λ i → B (p i)) pa + +transport-fillerExt : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) + → PathP (λ i → A → p i) (λ x → x) (transport p) +transport-fillerExt p i x = transport-filler p x i + +transport⁻-fillerExt : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) + → PathP (λ i → p i → A) (λ x → x) (transport⁻ p) +transport⁻-fillerExt p i x = transp (λ j → p (i ∧ ~ j)) (~ i) x + +transport-fillerExt⁻ : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) + → PathP (λ i → p i → B) (transport p) (λ x → x) +transport-fillerExt⁻ p = symP (transport⁻-fillerExt (sym p)) + +transport⁻-fillerExt⁻ : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) + → PathP (λ i → B → p i) (transport⁻ p) (λ x → x) +transport⁻-fillerExt⁻ p = symP (transport-fillerExt (sym p)) + + +transport⁻-filler : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) (x : B) + → PathP (λ i → p (~ i)) x (transport⁻ p x) +transport⁻-filler p x = transport-filler (λ i → p (~ i)) x + transport⁻Transport : ∀ {ℓ} {A B : Type ℓ} → (p : A ≡ B) → (a : A) → transport⁻ p (transport p a) ≡ a -transport⁻Transport p a j = - transp (λ i → p (~ i ∧ ~ j)) j (transp (λ i → p (i ∧ ~ j)) j a) +transport⁻Transport p a j = transport⁻-fillerExt p (~ j) (transport-fillerExt p (~ j) a) transportTransport⁻ : ∀ {ℓ} {A B : Type ℓ} → (p : A ≡ B) → (b : B) → transport p (transport⁻ p b) ≡ b -transportTransport⁻ p b j = - transp (λ i → p (i ∨ j)) j (transp (λ i → p (~ i ∨ j)) j b) +transportTransport⁻ p b j = transport-fillerExt⁻ p j (transport⁻-fillerExt⁻ p j b) -- Transport is an equivalence isEquivTransport : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) → isEquiv (transport p) isEquivTransport {A = A} {B = B} p = - transport (λ i → isEquiv (λ x → transp (λ j → p (i ∧ j)) (~ i) x)) (idIsEquiv A) + transport (λ i → isEquiv (transport-fillerExt p i)) (idIsEquiv A) transportEquiv : ∀ {ℓ} {A B : Type ℓ} → A ≡ B → A ≃ B transportEquiv p = (transport p , isEquivTransport p) +substEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {a a' : A} (P : A → Type ℓ') (p : a ≡ a') → P a ≃ P a' +substEquiv P p = (subst P p , isEquivTransport (λ i → P (p i))) + +liftEquiv : ∀ {ℓ ℓ'} {A B : Type ℓ} (P : Type ℓ → Type ℓ') (e : A ≃ B) → P A ≃ P B +liftEquiv P e = substEquiv P (ua e) + +transpEquiv : ∀ {ℓ} {A B : Type ℓ} (p : A ≡ B) → ∀ i → p i ≃ B +transpEquiv P i .fst = transp (λ j → P (i ∨ j)) i +transpEquiv P i .snd + = transp (λ k → isEquiv (transp (λ j → P (i ∨ (j ∧ k))) (i ∨ ~ k))) + i (idIsEquiv (P i)) + +uaTransportη : ∀ {ℓ} {A B : Type ℓ} (P : A ≡ B) → ua (transportEquiv P) ≡ P +uaTransportη P i j + = Glue (P i1) λ where + (j = i0) → P i0 , transportEquiv P + (i = i1) → P j , transpEquiv P j + (j = i1) → P i1 , idEquiv (P i1) + pathToIso : ∀ {ℓ} {A B : Type ℓ} → A ≡ B → Iso A B -pathToIso x = iso (transport x) (transport⁻ x ) ( transportTransport⁻ x) (transport⁻Transport x) +Iso.fun (pathToIso x) = transport x +Iso.inv (pathToIso x) = transport⁻ x +Iso.rightInv (pathToIso x) = transportTransport⁻ x +Iso.leftInv (pathToIso x) = transport⁻Transport x isInjectiveTransport : ∀ {ℓ : Level} {A B : Type ℓ} {p q : A ≡ B} → transport p ≡ transport q → p ≡ q isInjectiveTransport {p = p} {q} α i = hcomp (λ j → λ - { (i = i0) → secEq univalence p j - ; (i = i1) → secEq univalence q j + { (i = i0) → retEq univalence p j + ; (i = i1) → retEq univalence q j }) (invEq univalence ((λ a → α i a) , t i)) where t : PathP (λ i → isEquiv (λ a → α i a)) (pathToEquiv p .snd) (pathToEquiv q .snd) t = isProp→PathP (λ i → isPropIsEquiv (λ a → α i a)) _ _ +transportUaInv : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) → transport (ua (invEquiv e)) ≡ transport (sym (ua e)) +transportUaInv e = cong transport (uaInvEquiv e) +-- notice that transport (ua e) would reduce, thus an alternative definition using EquivJ can give +-- refl for the case of idEquiv: +-- transportUaInv e = EquivJ (λ _ e → transport (ua (invEquiv e)) ≡ transport (sym (ua e))) refl e + isSet-subst : ∀ {ℓ ℓ′} {A : Type ℓ} {B : A → Type ℓ′} → (isSet-A : isSet A) → ∀ {a : A} @@ -66,14 +118,11 @@ isSet-subst : ∀ {ℓ ℓ′} {A : Type ℓ} {B : A → Type ℓ′} isSet-subst {B = B} isSet-A p x = subst (λ p′ → subst B p′ x ≡ x) (isSet-A _ _ refl p) (substRefl {B = B} x) -- substituting along a composite path is equivalent to substituting twice -substComposite-□ : ∀ {ℓ ℓ′} {A : Type ℓ} → (B : A → Type ℓ′) - → {x y z : A} (p : x ≡ y) (q : y ≡ z) (u : B x) - → subst B (p □ q) u ≡ subst B q (subst B p u) -substComposite-□ B p q Bx = sym (substRefl {B = B} _) ∙ helper where - compSq : I → I → _ - compSq = compPath'-filler p q - helper : subst B refl (subst B (p □ q) Bx) ≡ subst B q (subst B p Bx) - helper i = subst B (λ k → compSq (~ i ∧ ~ k) (~ i ∨ k)) (subst B (λ k → compSq (~ i ∨ ~ k) (~ i ∧ k)) Bx) +substComposite : ∀ {ℓ ℓ′} {A : Type ℓ} → (B : A → Type ℓ′) + → {x y z : A} (p : x ≡ y) (q : y ≡ z) (u : B x) + → subst B (p ∙ q) u ≡ subst B q (subst B p u) +substComposite B p q Bx i = + transport (cong B (compPath-filler' p q (~ i))) (transport-fillerExt (cong B p) i Bx) -- substitution commutes with morphisms in slices substCommSlice : ∀ {ℓ ℓ′} {A : Type ℓ} @@ -81,11 +130,20 @@ substCommSlice : ∀ {ℓ ℓ′} {A : Type ℓ} → (F : ∀ i → B i → C i) → {x y : A} (p : x ≡ y) (u : B x) → subst C p (F x u) ≡ F y (subst B p u) -substCommSlice B C F p Bx i = comp pathC (λ k → λ where - (i = i0) → toPathP {A = pathC} (λ _ → subst C p (F _ Bx)) k - (i = i1) → F (p k) (toPathP {A = pathB} (λ _ → subst B p Bx) k) - ) (F _ Bx) where - pathC : I → Type _ - pathC i = cong C p i - pathB : I → Type _ - pathB i = cong B p i +substCommSlice B C F p Bx i = + transport-fillerExt⁻ (cong C p) i (F _ (transport-fillerExt (cong B p) i Bx)) + +-- transporting over (λ i → B (p i) → C (p i)) divides the transport into +-- transports over (λ i → C (p i)) and (λ i → B (p (~ i))) +funTypeTransp : ∀ {ℓ ℓ'} {A : Type ℓ} (B C : A → Type ℓ') {x y : A} (p : x ≡ y) (f : B x → C x) + → PathP (λ i → B (p i) → C (p i)) f (subst C p ∘ f ∘ subst B (sym p)) +funTypeTransp B C {x = x} p f i b = + transp (λ j → C (p (j ∧ i))) (~ i) (f (transp (λ j → B (p (i ∧ ~ j))) (~ i) b)) + +-- transports between loop spaces preserve path composition +overPathFunct : ∀ {ℓ} {A : Type ℓ} {x y : A} (p q : x ≡ x) (P : x ≡ y) + → transport (λ i → P i ≡ P i) (p ∙ q) + ≡ transport (λ i → P i ≡ P i) p ∙ transport (λ i → P i ≡ P i) q +overPathFunct p q = + J (λ y P → transport (λ i → P i ≡ P i) (p ∙ q) ≡ transport (λ i → P i ≡ P i) p ∙ transport (λ i → P i ≡ P i) q) + (transportRefl (p ∙ q) ∙ cong₂ _∙_ (sym (transportRefl p)) (sym (transportRefl q))) diff --git a/Cubical/Foundations/Univalence.agda b/Cubical/Foundations/Univalence.agda index 91ad029c6..b113d83bd 100644 --- a/Cubical/Foundations/Univalence.agda +++ b/Cubical/Foundations/Univalence.agda @@ -12,7 +12,7 @@ various consequences of univalence - Isomorphism induction ([elimIso]) -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Univalence where open import Cubical.Foundations.Prelude @@ -21,9 +21,13 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.Sigma.Base + open import Cubical.Core.Glue public using ( Glue ; glue ; unglue ; lineToEquiv ) +open import Cubical.Reflection.StrictEquiv + private variable ℓ ℓ' : Level @@ -36,26 +40,74 @@ ua {A = A} {B = B} e i = Glue B (λ { (i = i0) → (A , e) uaIdEquiv : {A : Type ℓ} → ua (idEquiv A) ≡ refl uaIdEquiv {A = A} i j = Glue A {φ = i ∨ ~ j ∨ j} (λ _ → A , idEquiv A) --- Give detailed type to unglue, mainly for documentation purposes -ua-unglue : ∀ {A B : Type ℓ} → (e : A ≃ B) → (i : I) (x : ua e i) - → B [ _ ↦ (λ { (i = i0) → e .fst x ; (i = i1) → x }) ] -ua-unglue e i x = inS (unglue (i ∨ ~ i) x) +-- Propositional extensionality +hPropExt : {A B : Type ℓ} → isProp A → isProp B → (A → B) → (B → A) → A ≡ B +hPropExt Aprop Bprop f g = ua (propBiimpl→Equiv Aprop Bprop f g) + +-- the unglue and glue primitives specialized to the case of ua + +ua-unglue : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : ua e i) + → B {- [ _ ↦ (λ { (i = i0) → e .fst x ; (i = i1) → x }) ] -} +ua-unglue e i x = unglue (i ∨ ~ i) x + +ua-glue : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : Partial (~ i) A) + (y : B [ _ ↦ (λ { (i = i0) → e .fst (x 1=1) }) ]) + → ua e i {- [ _ ↦ (λ { (i = i0) → x 1=1 ; (i = i1) → outS y }) ] -} +ua-glue e i x y = glue {φ = i ∨ ~ i} (λ { (i = i0) → x 1=1 ; (i = i1) → outS y }) (outS y) + +module _ {A B : Type ℓ} (e : A ≃ B) {x : A} {y : B} where + -- sometimes more useful are versions of these functions with the (i : I) factored in + + ua-ungluePath : PathP (λ i → ua e i) x y → e .fst x ≡ y + ua-ungluePath p i = ua-unglue e i (p i) + + ua-gluePath : e .fst x ≡ y → PathP (λ i → ua e i) x y + ua-gluePath p i = ua-glue e i (λ { (i = i0) → x }) (inS (p i)) + + -- ua-ungluePath and ua-gluePath are definitional inverses + ua-ungluePath-Equiv : (PathP (λ i → ua e i) x y) ≃ (e .fst x ≡ y) + unquoteDef ua-ungluePath-Equiv = + defStrictEquiv ua-ungluePath-Equiv ua-ungluePath ua-gluePath + +-- ua-unglue and ua-glue are also definitional inverses, in a way +-- strengthening the types of ua-unglue and ua-glue gives a nicer formulation of this, see below + +ua-unglue-glue : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : Partial (~ i) A) (y : B [ _ ↦ _ ]) + → ua-unglue e i (ua-glue e i x y) ≡ outS y +ua-unglue-glue _ _ _ _ = refl + +ua-glue-unglue : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : ua e i) + → ua-glue e i (λ { (i = i0) → x }) (inS (ua-unglue e i x)) ≡ x +ua-glue-unglue _ _ _ = refl -ua-ungluePath : ∀ {A B : Type ℓ} (e : A ≃ B) {x : A} {y : B} - → PathP (λ i → ua e i) x y - → e .fst x ≡ y -ua-ungluePath e p i = unglue (i ∨ ~ i) (p i) +-- mainly for documentation purposes, ua-unglue and ua-glue wrapped in cubical subtypes --- Give detailed type to glue -ua-glue : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : A) (y : B) - → B [ _ ↦ (λ { (i = i0) → e .fst x ; (i = i1) → y }) ] - → (ua e i) [ _ ↦ (λ { (i = i0) → x ; (i = i1) → y }) ] -ua-glue e i x y s = inS (glue (λ { (i = i0) → x ; (i = i1) → y }) (outS s)) +ua-unglueS : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : A) (y : B) + → ua e i [ _ ↦ (λ { (i = i0) → x ; (i = i1) → y }) ] + → B [ _ ↦ (λ { (i = i0) → e .fst x ; (i = i1) → y }) ] +ua-unglueS e i x y s = inS (ua-unglue e i (outS s)) + +ua-glueS : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : A) (y : B) + → B [ _ ↦ (λ { (i = i0) → e .fst x ; (i = i1) → y }) ] + → ua e i [ _ ↦ (λ { (i = i0) → x ; (i = i1) → y }) ] +ua-glueS e i x y s = inS (ua-glue e i (λ { (i = i0) → x }) (inS (outS s))) + +ua-unglueS-glueS : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : A) (y : B) + (s : B [ _ ↦ (λ { (i = i0) → e .fst x ; (i = i1) → y }) ]) + → outS (ua-unglueS e i x y (ua-glueS e i x y s)) ≡ outS s +ua-unglueS-glueS _ _ _ _ _ = refl + +ua-glueS-unglueS : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : A) (y : B) + (s : ua e i [ _ ↦ (λ { (i = i0) → x ; (i = i1) → y }) ]) + → outS (ua-glueS e i x y (ua-unglueS e i x y s)) ≡ outS s +ua-glueS-unglueS _ _ _ _ _ = refl + + +-- a version of ua-glue with a single endpoint, identical to `ua-gluePath e {x} refl i` +ua-gluePt : ∀ {A B : Type ℓ} (e : A ≃ B) (i : I) (x : A) + → ua e i {- [ _ ↦ (λ { (i = i0) → x ; (i = i1) → e .fst x }) ] -} +ua-gluePt e i x = ua-glue e i (λ { (i = i0) → x }) (inS (e .fst x)) -ua-gluePath : ∀ {A B : Type ℓ} (e : A ≃ B) {x : A} {y : B} - → e .fst x ≡ y - → PathP (λ i → ua e i) x y -ua-gluePath e {x} {y} p i = glue (λ { (i = i0) → x ; (i = i1) → y }) (p i) -- Proof of univalence using that unglue is an equivalence: @@ -95,7 +147,7 @@ unglueEquiv A φ f = ( unglue φ , unglueIsEquiv A φ f ) -- unglue is an equivalence. The standard formulation can be found in -- Cubical/Basics/Univalence. -- -EquivContr : ∀ (A : Type ℓ) → isContr (Σ[ T ∈ Type ℓ ] T ≃ A) +EquivContr : ∀ (A : Type ℓ) → ∃![ T ∈ Type ℓ ] (T ≃ A) EquivContr {ℓ = ℓ} A = ( (A , idEquiv A) , idEquiv≡ ) @@ -114,37 +166,36 @@ contrSinglEquiv {A = A} {B = B} e = isContr→isProp (EquivContr B) (B , idEquiv B) (A , e) -- Equivalence induction -EquivJ : (P : (A B : Type ℓ) → (e : B ≃ A) → Type ℓ') - → (r : (A : Type ℓ) → P A A (idEquiv A)) - → (A B : Type ℓ) → (e : B ≃ A) → P A B e -EquivJ P r A B e = subst (λ x → P A (x .fst) (x .snd)) (contrSinglEquiv e) (r A) - --- Eliminate equivalences by just looking at the underlying function -elimEquivFun : (B : Type ℓ) (P : (A : Type ℓ) → (A → B) → Type ℓ') - → (r : P B (λ x → x)) - → (A : Type ℓ) → (e : A ≃ B) → P A (e .fst) -elimEquivFun B P r a e = subst (λ x → P (x .fst) (x .snd .fst)) (contrSinglEquiv e) r +EquivJ : {A B : Type ℓ} (P : (A : Type ℓ) → (e : A ≃ B) → Type ℓ') + → (r : P B (idEquiv B)) → (e : A ≃ B) → P A e +EquivJ P r e = subst (λ x → P (x .fst) (x .snd)) (contrSinglEquiv e) r -- Assuming that we have an inverse to ua we can easily prove univalence module Univalence (au : ∀ {ℓ} {A B : Type ℓ} → A ≡ B → A ≃ B) (aurefl : ∀ {ℓ} {A B : Type ℓ} → au refl ≡ idEquiv A) where ua-au : {A B : Type ℓ} (p : A ≡ B) → ua (au p) ≡ p - ua-au {B = B} p = J (λ b p → ua (au p) ≡ p) (cong ua (aurefl {B = B}) ∙ uaIdEquiv) p + ua-au {B = B} = J (λ _ p → ua (au p) ≡ p) + (cong ua (aurefl {B = B}) ∙ uaIdEquiv) au-ua : {A B : Type ℓ} (e : A ≃ B) → au (ua e) ≡ e - au-ua {B = B} e = EquivJ (λ b a f → au (ua f) ≡ f) - (λ x → subst (λ r → au r ≡ idEquiv x) (sym uaIdEquiv) (aurefl {B = B})) - _ _ e + au-ua {B = B} = EquivJ (λ _ f → au (ua f) ≡ f) + (subst (λ r → au r ≡ idEquiv _) (sym uaIdEquiv) (aurefl {B = B})) + + isoThm : ∀ {ℓ} {A B : Type ℓ} → Iso (A ≡ B) (A ≃ B) + isoThm .Iso.fun = au + isoThm .Iso.inv = ua + isoThm .Iso.rightInv = au-ua + isoThm .Iso.leftInv = ua-au thm : ∀ {ℓ} {A B : Type ℓ} → isEquiv au - thm {A = A} {B = B} = isoToIsEquiv {B = A ≃ B} (iso au ua au-ua ua-au) + thm {A = A} {B = B} = isoToIsEquiv {B = A ≃ B} isoThm pathToEquiv : {A B : Type ℓ} → A ≡ B → A ≃ B pathToEquiv p = lineToEquiv (λ i → p i) pathToEquivRefl : {A : Type ℓ} → pathToEquiv refl ≡ idEquiv A -pathToEquivRefl {A = A} = equivEq _ _ (λ i x → transp (λ _ → A) i x) +pathToEquivRefl {A = A} = equivEq (λ i x → transp (λ _ → A) i x) pathToEquiv-ua : {A B : Type ℓ} (e : A ≃ B) → pathToEquiv (ua e) ≡ e pathToEquiv-ua = Univalence.au-ua pathToEquiv pathToEquivRefl @@ -153,6 +204,9 @@ ua-pathToEquiv : {A B : Type ℓ} (p : A ≡ B) → ua (pathToEquiv p) ≡ p ua-pathToEquiv = Univalence.ua-au pathToEquiv pathToEquivRefl -- Univalence +univalenceIso : {A B : Type ℓ} → Iso (A ≡ B) (A ≃ B) +univalenceIso = Univalence.isoThm pathToEquiv pathToEquivRefl + univalence : {A B : Type ℓ} → (A ≡ B) ≃ (A ≃ B) univalence = ( pathToEquiv , Univalence.thm pathToEquiv pathToEquivRefl ) @@ -175,17 +229,58 @@ univalencePath = ua (compEquiv univalence LiftEquiv) -- The computation rule for ua. Because of "ghcomp" it is now very -- simple compared to cubicaltt: -- https://github.com/mortberg/cubicaltt/blob/master/examples/univalence.ctt#L202 -uaβ : {A B : Type ℓ} (e : A ≃ B) (x : A) → transport (ua e) x ≡ e .fst x -uaβ e x = transportRefl (e .fst x) +uaβ : {A B : Type ℓ} (e : A ≃ B) (x : A) → transport (ua e) x ≡ equivFun e x +uaβ e x = transportRefl (equivFun e x) uaη : ∀ {A B : Type ℓ} → (P : A ≡ B) → ua (pathToEquiv P) ≡ P uaη = J (λ _ q → ua (pathToEquiv q) ≡ q) (cong ua pathToEquivRefl ∙ uaIdEquiv) --- Alternative version of EquivJ that only requires a predicate on --- functions -elimEquiv : {B : Type ℓ} (P : {A : Type ℓ} → (A → B) → Type ℓ') → - (d : P (idfun B)) → {A : Type ℓ} → (e : A ≃ B) → P (e .fst) -elimEquiv P d e = subst (λ x → P (x .snd .fst)) (contrSinglEquiv e) d +-- Lemmas for constructing and destructing dependent paths in a function type where the domain is ua. +ua→ : ∀ {ℓ ℓ'} {A₀ A₁ : Type ℓ} {e : A₀ ≃ A₁} {B : (i : I) → Type ℓ'} + {f₀ : A₀ → B i0} {f₁ : A₁ → B i1} + → ((a : A₀) → PathP B (f₀ a) (f₁ (e .fst a))) + → PathP (λ i → ua e i → B i) f₀ f₁ +ua→ {e = e} {f₀ = f₀} {f₁} h i a = + hcomp + (λ j → λ + { (i = i0) → f₀ a + ; (i = i1) → f₁ (lem a j) + }) + (h (transp (λ j → ua e (~ j ∧ i)) (~ i) a) i) + where + lem : ∀ a₁ → e .fst (transport (sym (ua e)) a₁) ≡ a₁ + lem a₁ = secEq e _ ∙ transportRefl _ + +ua→⁻ : ∀ {ℓ ℓ'} {A₀ A₁ : Type ℓ} {e : A₀ ≃ A₁} {B : (i : I) → Type ℓ'} + {f₀ : A₀ → B i0} {f₁ : A₁ → B i1} + → PathP (λ i → ua e i → B i) f₀ f₁ + → ((a : A₀) → PathP B (f₀ a) (f₁ (e .fst a))) +ua→⁻ {e = e} {f₀ = f₀} {f₁} p a i = + hcomp + (λ k → λ + { (i = i0) → f₀ a + ; (i = i1) → f₁ (uaβ e a k) + }) + (p i (transp (λ j → ua e (j ∧ i)) (~ i) a)) + +-- Useful lemma for unfolding a transported function over ua +-- If we would have regularity this would be refl +transportUAop₁ : ∀ {A B : Type ℓ} → (e : A ≃ B) (f : A → A) (x : B) + → transport (λ i → ua e i → ua e i) f x ≡ equivFun e (f (invEq e x)) +transportUAop₁ e f x i = transportRefl (equivFun e (f (invEq e (transportRefl x i)))) i + +-- Binary version +transportUAop₂ : ∀ {ℓ} {A B : Type ℓ} → (e : A ≃ B) (f : A → A → A) (x y : B) + → transport (λ i → ua e i → ua e i → ua e i) f x y ≡ + equivFun e (f (invEq e x) (invEq e y)) +transportUAop₂ e f x y i = + transportRefl (equivFun e (f (invEq e (transportRefl x i)) + (invEq e (transportRefl y i)))) i + +-- Alternative version of EquivJ that only requires a predicate on functions +elimEquivFun : {A B : Type ℓ} (P : (A : Type ℓ) → (A → B) → Type ℓ') + → (r : P B (idfun B)) → (e : A ≃ B) → P A (e .fst) +elimEquivFun P r e = subst (λ x → P (x .fst) (x .snd .fst)) (contrSinglEquiv e) r -- Isomorphism induction elimIso : {B : Type ℓ} → (Q : {A : Type ℓ} → (A → B) → (B → A) → Type ℓ') → @@ -193,24 +288,22 @@ elimIso : {B : Type ℓ} → (Q : {A : Type ℓ} → (A → B) → (B → A) → (f : A → B) → (g : B → A) → section f g → retract f g → Q f g elimIso {ℓ} {ℓ'} {B} Q h {A} f g sfg rfg = rem1 f g sfg rfg where - P : {A : Type ℓ} → (f : A → B) → Type (ℓ-max ℓ' ℓ) - P {A} f = (g : B → A) → section f g → retract f g → Q f g + P : (A : Type ℓ) → (f : A → B) → Type (ℓ-max ℓ' ℓ) + P A f = (g : B → A) → section f g → retract f g → Q f g - rem : P (idfun B) + rem : P B (idfun B) rem g sfg rfg = subst (Q (idfun B)) (λ i b → (sfg b) (~ i)) h - rem1 : {A : Type ℓ} → (f : A → B) → P f - rem1 f g sfg rfg = elimEquiv P rem (f , isoToIsEquiv (iso f g sfg rfg)) g sfg rfg + rem1 : {A : Type ℓ} → (f : A → B) → P A f + rem1 f g sfg rfg = elimEquivFun P rem (f , isoToIsEquiv (iso f g sfg rfg)) g sfg rfg uaInvEquiv : ∀ {A B : Type ℓ} → (e : A ≃ B) → ua (invEquiv e) ≡ sym (ua e) -uaInvEquiv e = EquivJ (λ _ _ e → ua (invEquiv e) ≡ sym (ua e)) rem _ _ e - where - rem : (A : Type ℓ) → ua (invEquiv (idEquiv A)) ≡ sym (ua (idEquiv A)) - rem A = cong ua (invEquivIdEquiv A) +uaInvEquiv {B = B} = EquivJ (λ _ e → ua (invEquiv e) ≡ sym (ua e)) + (cong ua (invEquivIdEquiv B)) uaCompEquiv : ∀ {A B C : Type ℓ} → (e : A ≃ B) (f : B ≃ C) → ua (compEquiv e f) ≡ ua e ∙ ua f -uaCompEquiv {C = C} = EquivJ (λ A B e → (f : A ≃ C) → ua (compEquiv e f) ≡ ua e ∙ ua f) rem _ _ - where - rem : (A : Type _) (f : A ≃ C) → ua (compEquiv (idEquiv A) f) ≡ ua (idEquiv A) ∙ ua f - rem _ f = cong ua (compEquivIdEquiv f) ∙ sym (cong (λ x → x ∙ ua f) uaIdEquiv ∙ sym (lUnit (ua f))) +uaCompEquiv {B = B} {C} = EquivJ (λ _ e → (f : B ≃ C) → ua (compEquiv e f) ≡ ua e ∙ ua f) + (λ f → cong ua (compEquivIdEquiv f) + ∙ sym (cong (λ x → x ∙ ua f) uaIdEquiv + ∙ sym (lUnit (ua f)))) diff --git a/Cubical/Foundations/Univalence/Universe.agda b/Cubical/Foundations/Univalence/Universe.agda index 9f261cd53..e41e2db5a 100644 --- a/Cubical/Foundations/Univalence/Universe.agda +++ b/Cubical/Foundations/Univalence/Universe.agda @@ -1,25 +1,24 @@ -{-# OPTIONS --cubical --safe --postfix-projections #-} +{-# OPTIONS --safe --postfix-projections #-} open import Cubical.Core.Everything -open import Cubical.Foundations.Embedding +open import Cubical.Functions.Embedding open import Cubical.Foundations.Equiv open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Prelude open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport -- A helper module for deriving univalence for a higher inductive-recursive -- universe. -- -- U is the type of codes -- El is the decoding --- uaf is a higher constructor that requires paths between codes to exist +-- un is a higher constructor that requires paths between codes to exist -- for equivalences of decodings --- comp is intended to be the computational behavior of El on uaf, although +-- comp is intended to be the computational behavior of El on un, although -- it seems that being a path is sufficient. --- ret is a higher constructor that fills out the equivalence structure --- for uaf and the computational behavior of El. -- -- Given a universe defined as above, it's possible to show that the path -- space of the code type is equivalent to the path space of the actual @@ -33,20 +32,77 @@ open import Cubical.Foundations.Univalence module Cubical.Foundations.Univalence.Universe {ℓ ℓ'} (U : Type ℓ) (El : U → Type ℓ') - (uaf : ∀{s t} → El s ≃ El t → s ≡ t) - (comp : ∀{s t} (e : El s ≃ El t) → cong El (uaf e) ≡ ua e) - (ret : ∀{s t : U} → (p : s ≡ t) → uaf (lineToEquiv (λ i → El (p i))) ≡ p) + (un : ∀ s t → El s ≃ El t → s ≡ t) + (comp : ∀{s t} (e : El s ≃ El t) → cong El (un s t e) ≡ ua e) where +private + variable + A : Type ℓ' -minivalence : ∀{s t} → (s ≡ t) ≃ (El s ≡ El t) -minivalence {s} {t} = isoToEquiv mini - where - open Iso - mini : Iso (s ≡ t) (El s ≡ El t) - mini .fun = cong El - mini .inv = uaf ∘ pathToEquiv - mini .rightInv p = comp (pathToEquiv p) ∙ uaη p - mini .leftInv = ret +module UU-Lemmas where + reg : transport (λ _ → A) ≡ idfun A + reg {A} i z = transp (λ _ → A) i z + + nu : ∀ x y → x ≡ y → El x ≃ El y + nu x y p = transportEquiv (cong El p) + + cong-un-te + : ∀ x y (p : El x ≡ El y) + → cong El (un x y (transportEquiv p)) ≡ p + cong-un-te x y p + = comp (transportEquiv p) ∙ uaTransportη p + + nu-un : ∀ x y (e : El x ≃ El y) → nu x y (un x y e) ≡ e + nu-un x y e + = equivEq {e = nu x y (un x y e)} {f = e} λ i z + → (cong (λ p → transport p z) (comp e) ∙ uaβ e z) i + + El-un-equiv : ∀ x i → El (un x x (idEquiv _) i) ≃ El x + El-un-equiv x i = λ where + .fst → transp (λ j → p j) (i ∨ ~ i) + .snd → transp (λ j → isEquiv (transp (λ k → p (j ∧ k)) (~ j ∨ i ∨ ~ i))) + (i ∨ ~ i) (idIsEquiv T) + where + T = El (un x x (idEquiv _) i) + p : T ≡ El x + p j = (comp (idEquiv _) ∙ uaIdEquiv {A = El x}) j i + + un-refl : ∀ x → un x x (idEquiv (El x)) ≡ refl + un-refl x i j + = hcomp (λ k → λ where + (i = i0) → un x x (idEquiv (El x)) j + (i = i1) → un x x (idEquiv (El x)) (j ∨ k) + (j = i0) → un x x (idEquiv (El x)) (~ i ∨ k) + (j = i1) → x) + (un (un x x (idEquiv (El x)) (~ i)) x (El-un-equiv x (~ i)) j) + + nu-refl : ∀ x → nu x x refl ≡ idEquiv (El x) + nu-refl x = equivEq {e = nu x x refl} {f = idEquiv (El x)} reg + + un-nu : ∀ x y (p : x ≡ y) → un x y (nu x y p) ≡ p + un-nu x y p + = J (λ z q → un x z (nu x z q) ≡ q) (cong (un x x) (nu-refl x) ∙ un-refl x) p + +open UU-Lemmas +open Iso + +equivIso : ∀ s t → Iso (s ≡ t) (El s ≃ El t) +equivIso s t .fun = nu s t +equivIso s t .inv = un s t +equivIso s t .rightInv = nu-un s t +equivIso s t .leftInv = un-nu s t + +pathIso : ∀ s t → Iso (s ≡ t) (El s ≡ El t) +pathIso s t .fun = cong El +pathIso s t .inv = un s t ∘ transportEquiv +pathIso s t .rightInv = cong-un-te s t +pathIso s t .leftInv = un-nu s t + +minivalence : ∀{s t} → (s ≡ t) ≃ (El s ≃ El t) +minivalence {s} {t} = isoToEquiv (equivIso s t) + +path-reflection : ∀{s t} → (s ≡ t) ≃ (El s ≡ El t) +path-reflection {s} {t} = isoToEquiv (pathIso s t) isEmbeddingEl : isEmbedding El -isEmbeddingEl s t = snd minivalence +isEmbeddingEl s t = snd path-reflection diff --git a/Cubical/Foundations/UnivalenceId.agda b/Cubical/Foundations/UnivalenceId.agda deleted file mode 100644 index 9523d0ece..000000000 --- a/Cubical/Foundations/UnivalenceId.agda +++ /dev/null @@ -1,39 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.Foundations.UnivalenceId where - -open import Cubical.Foundations.Prelude public - hiding ( _≡_ ; _≡⟨_⟩_ ; _∎ ) -open import Cubical.Foundations.Id -open import Cubical.Foundations.Equiv - renaming ( isEquiv to isEquivPath - ; _≃_ to EquivPath - ; equivFun to equivFunPath - ; isPropIsEquiv to isPropIsEquivPath ) -open import Cubical.Foundations.Univalence - renaming ( EquivContr to EquivContrPath ) -open import Cubical.Foundations.Isomorphism - -path≡Id : ∀ {ℓ} {A B : Type ℓ} → Path _ (Path _ A B) (Id A B) -path≡Id = isoToPath (iso pathToId idToPath idToPathToId pathToIdToPath ) - -equivPathToEquivPath : ∀ {ℓ} {A : Type ℓ} {B : Type ℓ} → (p : EquivPath A B) → - Path _ (equivToEquivPath (equivPathToEquiv p)) p -equivPathToEquivPath (f , p) i = - ( f , isPropIsEquivPath f (equivToEquivPath (equivPathToEquiv (f , p)) .snd) p i ) - -equivPath≡Equiv : ∀ {ℓ} {A B : Type ℓ} → Path _ (EquivPath A B) (A ≃ B) -equivPath≡Equiv {ℓ} = isoToPath (iso (equivPathToEquiv {ℓ}) equivToEquivPath equivToEquiv equivPathToEquivPath) - -univalenceId : ∀ {ℓ} {A B : Type ℓ} → (A ≡ B) ≃ (A ≃ B) -univalenceId {ℓ} {A = A} {B = B} = equivPathToEquiv rem - where - rem0 : Path _ (Lift (EquivPath A B)) (Lift (A ≃ B)) - rem0 = congPath Lift equivPath≡Equiv - - rem1 : Path _ (Id A B) (Lift (A ≃ B)) - rem1 i = hcomp (λ j → λ { (i = i0) → path≡Id {A = A} {B = B} j - ; (i = i1) → rem0 j }) - (univalencePath {A = A} {B = B} i) - - rem : EquivPath (Id A B) (A ≃ B) - rem = compEquiv (eqweqmap rem1) (invEquiv LiftEquiv) diff --git a/Cubical/Functions/Bundle.agda b/Cubical/Functions/Bundle.agda new file mode 100644 index 000000000..98638451f --- /dev/null +++ b/Cubical/Functions/Bundle.agda @@ -0,0 +1,50 @@ +{-# OPTIONS --safe #-} +module Cubical.Functions.Bundle where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Functions.Fibration + +open import Cubical.Foundations.Structure +open import Cubical.Structures.TypeEqvTo + +open import Cubical.Data.Sigma.Properties +open import Cubical.HITs.PropositionalTruncation + +module _ {ℓb ℓf} {B : Type ℓb} {F : Type ℓf} {ℓ} where + +{- + A fiber bundle with base space B and fibers F is a map `p⁻¹ : B → TypeEqvTo F` + taking points in the base space to their respective fibers. + + e.g. a double cover is a map `B → TypeEqvTo Bool` +-} + + Total : (p⁻¹ : B → TypeEqvTo ℓ F) → Type (ℓ-max ℓb ℓ) + Total p⁻¹ = Σ[ x ∈ B ] p⁻¹ x .fst + + pr : (p⁻¹ : B → TypeEqvTo ℓ F) → Total p⁻¹ → B + pr p⁻¹ = fst + + inc : (p⁻¹ : B → TypeEqvTo ℓ F) (x : B) → p⁻¹ x .fst → Total p⁻¹ + inc p⁻¹ x = (x ,_) + + fibPrEquiv : (p⁻¹ : B → TypeEqvTo ℓ F) (x : B) → fiber (pr p⁻¹) x ≃ p⁻¹ x .fst + fibPrEquiv p⁻¹ x = fiberEquiv (λ x → typ (p⁻¹ x)) x + +module _ {ℓb ℓf} (B : Type ℓb) (ℓ : Level) (F : Type ℓf) where + private + ℓ' = ℓ-max ℓb ℓ + +{- + Equivalently, a fiber bundle with base space B and fibers F is a type E and + a map `p : E → B` with fibers merely equivalent to F. +-} + + bundleEquiv : (B → TypeEqvTo ℓ' F) ≃ (Σ[ E ∈ Type ℓ' ] Σ[ p ∈ (E → B) ] ∀ x → ∥ fiber p x ≃ F ∥) + bundleEquiv = compEquiv (compEquiv Σ-Π-≃ (pathToEquiv p)) Σ-assoc-≃ + where e = fibrationEquiv B ℓ + p : (Σ[ p⁻¹ ∈ (B → Type ℓ') ] ∀ x → ∥ p⁻¹ x ≃ F ∥) + ≡ (Σ[ p ∈ (Σ[ E ∈ Type ℓ' ] (E → B)) ] ∀ x → ∥ fiber (snd p) x ≃ F ∥ ) + p i = Σ[ q ∈ ua e (~ i) ] ∀ x → ∥ ua-unglue e (~ i) q x ≃ F ∥ diff --git a/Cubical/Functions/Embedding.agda b/Cubical/Functions/Embedding.agda new file mode 100644 index 000000000..3f83f2e55 --- /dev/null +++ b/Cubical/Functions/Embedding.agda @@ -0,0 +1,427 @@ +{-# OPTIONS --safe #-} +module Cubical.Functions.Embedding where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence using (ua; univalence) +open import Cubical.Functions.Fibration + +open import Cubical.Data.Sigma +open import Cubical.Functions.Fibration +open import Cubical.Functions.FunExtEquiv +open import Cubical.Relation.Nullary using (Discrete; yes; no) +open import Cubical.Structures.Axioms + +open import Cubical.Reflection.StrictEquiv + +open import Cubical.Data.Nat using (ℕ; zero; suc) +open import Cubical.Data.Sigma + +private + variable + ℓ ℓ₁ ℓ₂ : Level + A B : Type ℓ + f h : A → B + w x : A + y z : B + +-- Embeddings are generalizations of injections. The usual +-- definition of injection as: +-- +-- f x ≡ f y → x ≡ y +-- +-- is not well-behaved with higher h-levels, while embeddings +-- are. +isEmbedding : (A → B) → Type _ +isEmbedding f = ∀ w x → isEquiv {A = w ≡ x} (cong f) + +isPropIsEmbedding : isProp (isEmbedding f) +isPropIsEmbedding {f = f} = isPropΠ2 λ _ _ → isPropIsEquiv (cong f) + +-- If A and B are h-sets, then injective functions between +-- them are embeddings. +-- +-- Note: It doesn't appear to be possible to omit either of +-- the `isSet` hypotheses. +injEmbedding + : {f : A → B} + → isSet A → isSet B + → (∀{w x} → f w ≡ f x → w ≡ x) + → isEmbedding f +injEmbedding {f = f} iSA iSB inj w x + = isoToIsEquiv (iso (cong f) inj sect retr) + where + sect : section (cong f) inj + sect p = iSB (f w) (f x) _ p + + retr : retract (cong f) inj + retr p = iSA w x _ p + +-- If `f` is an embedding, we'd expect the fibers of `f` to be +-- propositions, like an injective function. +hasPropFibers : (A → B) → Type _ +hasPropFibers f = ∀ y → isProp (fiber f y) + +-- This can be relaxed to having all prop fibers over the image, see [hasPropFibersOfImage→isEmbedding] +hasPropFibersOfImage : (A → B) → Type _ +hasPropFibersOfImage f = ∀ x → isProp (fiber f (f x)) + +-- some notation +_↪_ : Type ℓ₁ → Type ℓ₂ → Type (ℓ-max ℓ₁ ℓ₂) +A ↪ B = Σ[ f ∈ (A → B) ] isEmbedding f + +hasPropFibersIsProp : isProp (hasPropFibers f) +hasPropFibersIsProp = isPropΠ (λ _ → isPropIsProp) + +private + lemma₀ : (p : y ≡ z) → fiber f y ≡ fiber f z + lemma₀ {f = f} p = λ i → fiber f (p i) + + lemma₁ : isEmbedding f → ∀ x → isContr (fiber f (f x)) + lemma₁ {f = f} iE x = value , path + where + value : fiber f (f x) + value = (x , refl) + + path : ∀(fi : fiber f (f x)) → value ≡ fi + path (w , p) i + = case equiv-proof (iE w x) p of λ + { ((q , sq) , _) + → hfill (λ j → λ { (i = i0) → (x , refl) + ; (i = i1) → (w , sq j) + }) + (inS (q (~ i) , λ j → f (q (~ i ∨ j)))) + i1 + } + +isEmbedding→hasPropFibers : isEmbedding f → hasPropFibers f +isEmbedding→hasPropFibers iE y (x , p) + = subst (λ f → isProp f) (lemma₀ p) (isContr→isProp (lemma₁ iE x)) (x , p) + +private + fibCong→PathP + : {f : A → B} + → (p : f w ≡ f x) + → (fi : fiber (cong f) p) + → PathP (λ i → fiber f (p i)) (w , refl) (x , refl) + fibCong→PathP p (q , r) i = q i , λ j → r j i + + PathP→fibCong + : {f : A → B} + → (p : f w ≡ f x) + → (pp : PathP (λ i → fiber f (p i)) (w , refl) (x , refl)) + → fiber (cong f) p + PathP→fibCong p pp = (λ i → fst (pp i)) , (λ j i → snd (pp i) j) + +PathP≡fibCong + : {f : A → B} + → (p : f w ≡ f x) + → PathP (λ i → fiber f (p i)) (w , refl) (x , refl) ≡ fiber (cong f) p +PathP≡fibCong p + = isoToPath (iso (PathP→fibCong p) (fibCong→PathP p) (λ _ → refl) (λ _ → refl)) + +hasPropFibers→isEmbedding : hasPropFibers f → isEmbedding f +hasPropFibers→isEmbedding {f = f} iP w x .equiv-proof p + = subst isContr (PathP≡fibCong p) (isProp→isContrPathP (λ i → iP (p i)) fw fx) + where + fw : fiber f (f w) + fw = (w , refl) + + fx : fiber f (f x) + fx = (x , refl) + +hasPropFibersOfImage→hasPropFibers : hasPropFibersOfImage f → hasPropFibers f +hasPropFibersOfImage→hasPropFibers {f = f} fibImg y a b = + subst (λ y → isProp (fiber f y)) (snd a) (fibImg (fst a)) a b + +hasPropFibersOfImage→isEmbedding : hasPropFibersOfImage f → isEmbedding f +hasPropFibersOfImage→isEmbedding = hasPropFibers→isEmbedding ∘ hasPropFibersOfImage→hasPropFibers + +isEmbedding≡hasPropFibers : isEmbedding f ≡ hasPropFibers f +isEmbedding≡hasPropFibers + = isoToPath + (iso isEmbedding→hasPropFibers + hasPropFibers→isEmbedding + (λ _ → hasPropFibersIsProp _ _) + (λ _ → isPropIsEmbedding _ _)) + +isEquiv→hasPropFibers : isEquiv f → hasPropFibers f +isEquiv→hasPropFibers e b = isContr→isProp (equiv-proof e b) + +isEquiv→isEmbedding : isEquiv f → isEmbedding f +isEquiv→isEmbedding e = λ _ _ → congEquiv (_ , e) .snd + +Equiv→Embedding : A ≃ B → A ↪ B +Equiv→Embedding (f , isEquivF) = (f , isEquiv→isEmbedding isEquivF) + +iso→isEmbedding : ∀ {ℓ} {A B : Type ℓ} + → (isom : Iso A B) + ------------------------------- + → isEmbedding (Iso.fun isom) +iso→isEmbedding {A = A} {B} isom = (isEquiv→isEmbedding (equivIsEquiv (isoToEquiv isom))) + +isEmbedding→Injection : + ∀ {ℓ} {A B C : Type ℓ} + → (a : A → B) + → (e : isEmbedding a) + ---------------------- + → ∀ {f g : C → A} → + ∀ x → (a (f x) ≡ a (g x)) ≡ (f x ≡ g x) +isEmbedding→Injection a e {f = f} {g} x = sym (ua (cong a , e (f x) (g x))) + +-- if `f` has a retract, then `cong f` has, as well. If `B` is a set, then `cong f` +-- further has a section, making `f` an embedding. +module _ {f : A → B} (retf : hasRetract f) where + open Σ retf renaming (fst to g ; snd to ϕ) + + congRetract : f w ≡ f x → w ≡ x + congRetract {w = w} {x = x} p = sym (ϕ w) ∙∙ cong g p ∙∙ ϕ x + + isRetractCongRetract : retract (cong {x = w} {y = x} f) congRetract + isRetractCongRetract p = transport (PathP≡doubleCompPathˡ _ _ _ _) (λ i j → ϕ (p j) i) + + hasRetract→hasRetractCong : hasRetract (cong {x = w} {y = x} f) + hasRetract→hasRetractCong = congRetract , isRetractCongRetract + + retractableIntoSet→isEmbedding : isSet B → isEmbedding f + retractableIntoSet→isEmbedding setB w x = + isoToIsEquiv (iso (cong f) congRetract (λ _ → setB _ _ _ _) (hasRetract→hasRetractCong .snd)) + +Embedding-into-Discrete→Discrete : A ↪ B → Discrete B → Discrete A +Embedding-into-Discrete→Discrete (f , isEmbeddingF) _≟_ x y with f x ≟ f y +... | yes p = yes (invIsEq (isEmbeddingF x y) p) +... | no ¬p = no (¬p ∘ cong f) + +Embedding-into-isProp→isProp : A ↪ B → isProp B → isProp A +Embedding-into-isProp→isProp (f , isEmbeddingF) isProp-B x y + = invIsEq (isEmbeddingF x y) (isProp-B (f x) (f y)) + +Embedding-into-isSet→isSet : A ↪ B → isSet B → isSet A +Embedding-into-isSet→isSet (f , isEmbeddingF) isSet-B x y p q = + p ≡⟨ sym (retIsEq isEquiv-cong-f p) ⟩ + cong-f⁻¹ (cong f p) ≡⟨ cong cong-f⁻¹ cong-f-p≡cong-f-q ⟩ + cong-f⁻¹ (cong f q) ≡⟨ retIsEq isEquiv-cong-f q ⟩ + q ∎ + where + cong-f-p≡cong-f-q = isSet-B (f x) (f y) (cong f p) (cong f q) + isEquiv-cong-f = isEmbeddingF x y + cong-f⁻¹ = invIsEq isEquiv-cong-f + +Embedding-into-hLevel→hLevel + : ∀ n → A ↪ B → isOfHLevel (suc n) B → isOfHLevel (suc n) A +Embedding-into-hLevel→hLevel zero = Embedding-into-isProp→isProp +Embedding-into-hLevel→hLevel (suc n) (f , isEmbeddingF) Blvl x y + = isOfHLevelRespectEquiv (suc n) (invEquiv equiv) subLvl + where + equiv : (x ≡ y) ≃ (f x ≡ f y) + equiv .fst = cong f + equiv .snd = isEmbeddingF x y + subLvl : isOfHLevel (suc n) (f x ≡ f y) + subLvl = Blvl (f x) (f y) + +-- We now show that the powerset is the subtype classifier +-- i.e. ℙ X ≃ Σ[A ∈ Type ℓ] (A ↪ X) +Embedding→Subset : {X : Type ℓ} → Σ[ A ∈ Type ℓ ] (A ↪ X) → ℙ X +Embedding→Subset (_ , f , isEmbeddingF) x = fiber f x , isEmbedding→hasPropFibers isEmbeddingF x + +Subset→Embedding : {X : Type ℓ} → ℙ X → Σ[ A ∈ Type ℓ ] (A ↪ X) +Subset→Embedding {X = X} A = D , fst , Ψ + where + D = Σ[ x ∈ X ] x ∈ A + + Ψ : isEmbedding fst + Ψ w x = isEmbeddingFstΣProp (∈-isProp A) + +Subset→Embedding→Subset : {X : Type ℓ} → section (Embedding→Subset {ℓ} {X}) (Subset→Embedding {ℓ} {X}) +Subset→Embedding→Subset _ = funExt λ x → Σ≡Prop (λ _ → isPropIsProp) (ua (FiberIso.fiberEquiv _ x)) + +Embedding→Subset→Embedding : {X : Type ℓ} → retract (Embedding→Subset {ℓ} {X}) (Subset→Embedding {ℓ} {X}) +Embedding→Subset→Embedding {ℓ = ℓ} {X = X} (A , f , ψ) = + cong (equivFun Σ-assoc-≃) (Σ≡Prop (λ _ → isPropIsEmbedding) (retEq (fibrationEquiv X ℓ) (A , f))) + +Subset≃Embedding : {X : Type ℓ} → ℙ X ≃ (Σ[ A ∈ Type ℓ ] (A ↪ X)) +Subset≃Embedding = isoToEquiv (iso Subset→Embedding Embedding→Subset + Embedding→Subset→Embedding Subset→Embedding→Subset) + +Subset≡Embedding : {X : Type ℓ} → ℙ X ≡ (Σ[ A ∈ Type ℓ ] (A ↪ X)) +Subset≡Embedding = ua Subset≃Embedding + +isEmbedding-∘ : isEmbedding f → isEmbedding h → isEmbedding (f ∘ h) +isEmbedding-∘ {f = f} {h = h} Embf Embh w x + = compEquiv (cong h , Embh w x) (cong f , Embf (h w) (h x)) .snd + +isEmbedding→embedsFibersIntoSingl + : isEmbedding f + → ∀ z → fiber f z ↪ singl z +isEmbedding→embedsFibersIntoSingl {f = f} isE z = e , isEmbE where + e : fiber f z → singl z + e x = f (fst x) , sym (snd x) + + isEmbE : isEmbedding e + isEmbE u v = goal where + -- "adjust" ΣeqCf by trivial equivalences that hold judgementally, which should save compositions + Dom′ : ∀ u v → Type _ + Dom′ u v = Σ[ p ∈ fst u ≡ fst v ] PathP (λ i → f (p i) ≡ z) (snd u) (snd v) + Cod′ : ∀ u v → Type _ + Cod′ u v = Σ[ p ∈ f (fst u) ≡ f (fst v) ] PathP (λ i → p i ≡ z) (snd u) (snd v) + ΣeqCf : Dom′ u v ≃ Cod′ u v + ΣeqCf = Σ-cong-equiv-fst (_ , isE _ _) + + dom→ : u ≡ v → Dom′ u v + dom→ p = cong fst p , cong snd p + dom← : Dom′ u v → u ≡ v + dom← p i = p .fst i , p .snd i + + cod→ : e u ≡ e v → Cod′ u v + cod→ p = cong fst p , cong (sym ∘ snd) p + cod← : Cod′ u v → e u ≡ e v + cod← p i = p .fst i , sym (p .snd i) + + goal : isEquiv (cong e) + goal .equiv-proof x .fst .fst = + dom← (equivCtr ΣeqCf (cod→ x) .fst) + goal .equiv-proof x .fst .snd j = + cod← (equivCtr ΣeqCf (cod→ x) .snd j) + goal .equiv-proof x .snd (g , p) i .fst = + dom← (equivCtrPath ΣeqCf (cod→ x) (dom→ g , cong cod→ p) i .fst) + goal .equiv-proof x .snd (g , p) i .snd j = + cod← (equivCtrPath ΣeqCf (cod→ x) (dom→ g , cong cod→ p) i .snd j) + +isEmbedding→hasPropFibers′ : isEmbedding f → hasPropFibers f +isEmbedding→hasPropFibers′ {f = f} iE z = + Embedding-into-isProp→isProp (isEmbedding→embedsFibersIntoSingl iE z) isPropSingl + +universeEmbedding : + ∀ {ℓ ℓ₁ : Level} + → (F : Type ℓ → Type ℓ₁) + → (∀ X → F X ≃ X) + → isEmbedding F +universeEmbedding F liftingEquiv = hasPropFibersOfImage→isEmbedding propFibersF where + lemma : ∀ A B → (F A ≡ F B) ≃ (B ≡ A) + lemma A B = (F A ≡ F B) ≃⟨ univalence ⟩ + (F A ≃ F B) ≃⟨ equivComp (liftingEquiv A) (liftingEquiv B) ⟩ + (A ≃ B) ≃⟨ invEquivEquiv ⟩ + (B ≃ A) ≃⟨ invEquiv univalence ⟩ + (B ≡ A) ■ + fiberSingl : ∀ X → fiber F (F X) ≃ singl X + fiberSingl X = Σ-cong-equiv-snd (λ _ → lemma _ _) + propFibersF : hasPropFibersOfImage F + propFibersF X = Embedding-into-isProp→isProp (Equiv→Embedding (fiberSingl X)) isPropSingl + +liftEmbedding : (ℓ ℓ₁ : Level) + → isEmbedding (Lift {i = ℓ} {j = ℓ₁}) +liftEmbedding ℓ ℓ₁ = universeEmbedding (Lift {j = ℓ₁}) (λ _ → invEquiv LiftEquiv) + +module FibrationIdentityPrinciple {B : Type ℓ} {ℓ₁} where + -- note that fibrationEquiv (for good reason) uses ℓ₁ = ℓ-max ℓ ℓ₁, so we have to work + -- some universe magic to achieve good universe polymorphism + + -- First, prove it for the case that's dealt with in fibrationEquiv + Fibration′ = Fibration B (ℓ-max ℓ ℓ₁) + + module Lifted (f g : Fibration′) where + f≃g′ : Type (ℓ-max ℓ ℓ₁) + f≃g′ = ∀ b → fiber (f .snd) b ≃ fiber (g .snd) b + + Fibration′IP : f≃g′ ≃ (f ≡ g) + Fibration′IP = + f≃g′ + ≃⟨ equivΠCod (λ _ → invEquiv univalence) ⟩ + (∀ b → fiber (f .snd) b ≡ fiber (g .snd) b) + ≃⟨ funExtEquiv ⟩ + fiber (f .snd) ≡ fiber (g .snd) + ≃⟨ invEquiv (congEquiv (fibrationEquiv B ℓ₁)) ⟩ + f ≡ g + ■ + + -- Then embed into the above case by lifting the type + L : Type _ → Type _ -- local synonym fixing the levels of Lift + L = Lift {i = ℓ₁} {j = ℓ} + + liftFibration : Fibration B ℓ₁ → Fibration′ + liftFibration (A , f) = L A , f ∘ lower + + hasPropFibersLiftFibration : hasPropFibers liftFibration + hasPropFibersLiftFibration (A , f) = + Embedding-into-isProp→isProp (Equiv→Embedding fiberChar) + (isPropΣ (isEmbedding→hasPropFibers (liftEmbedding _ _) A) + λ _ → isEquiv→hasPropFibers (snd (invEquiv (preCompEquiv LiftEquiv))) _) + where + fiberChar : fiber liftFibration (A , f) + ≃ (Σ[ (E , eq) ∈ fiber L A ] fiber (_∘ lower) (transport⁻ (λ i → eq i → B) f)) + fiberChar = + fiber liftFibration (A , f) + ≃⟨ Σ-cong-equiv-snd (λ _ → invEquiv ΣPath≃PathΣ) ⟩ + (Σ[ (E , g) ∈ Fibration B ℓ₁ ] Σ[ eq ∈ (L E ≡ A) ] PathP (λ i → eq i → B) (g ∘ lower) f) + ≃⟨ boringSwap ⟩ + (Σ[ (E , eq) ∈ fiber L A ] Σ[ g ∈ (E → B) ] PathP (λ i → eq i → B) (g ∘ lower) f) + ≃⟨ Σ-cong-equiv-snd (λ _ → Σ-cong-equiv-snd λ _ → transportEquiv (PathP≡Path⁻ _ _ _)) ⟩ + (Σ[ (E , eq) ∈ fiber L A ] fiber (_∘ lower) (transport⁻ (λ i → eq i → B) f)) + ■ where + unquoteDecl boringSwap = + declStrictEquiv boringSwap + (λ ((E , g) , (eq , p)) → ((E , eq) , (g , p))) + (λ ((E , g) , (eq , p)) → ((E , eq) , (g , p))) + + isEmbeddingLiftFibration : isEmbedding liftFibration + isEmbeddingLiftFibration = hasPropFibers→isEmbedding hasPropFibersLiftFibration + + -- and finish off + module _ (f g : Fibration B ℓ₁) where + open Lifted (liftFibration f) (liftFibration g) + f≃g : Type (ℓ-max ℓ ℓ₁) + f≃g = ∀ b → fiber (f .snd) b ≃ fiber (g .snd) b + + FibrationIP : f≃g ≃ (f ≡ g) + FibrationIP = + f≃g ≃⟨ equivΠCod (λ b → equivComp (Σ-cong-equiv-fst LiftEquiv) + (Σ-cong-equiv-fst LiftEquiv)) ⟩ + f≃g′ ≃⟨ Fibration′IP ⟩ + (liftFibration f ≡ liftFibration g) ≃⟨ invEquiv (_ , isEmbeddingLiftFibration _ _) ⟩ + (f ≡ g) ■ +open FibrationIdentityPrinciple renaming (f≃g to _≃Fib_) using (FibrationIP) public + +Embedding : (B : Type ℓ₁) → (ℓ : Level) → Type (ℓ-max ℓ₁ (ℓ-suc ℓ)) +Embedding B ℓ = Σ[ A ∈ Type ℓ ] A ↪ B + +module EmbeddingIdentityPrinciple {B : Type ℓ} {ℓ₁} (f g : Embedding B ℓ₁) where + module _ where + open Σ f renaming (fst to F) public + open Σ g renaming (fst to G) public + open Σ (f .snd) renaming (fst to ffun; snd to isEmbF) public + open Σ (g .snd) renaming (fst to gfun; snd to isEmbG) public + f≃g : Type _ + f≃g = (∀ b → fiber ffun b → fiber gfun b) × + (∀ b → fiber gfun b → fiber ffun b) + toFibr : Embedding B ℓ₁ → Fibration B ℓ₁ + toFibr (A , (f , _)) = (A , f) + + isEmbeddingToFibr : isEmbedding toFibr + isEmbeddingToFibr w x = fullEquiv .snd where + -- carefully managed such that (cong toFibr) is the equivalence + fullEquiv : (w ≡ x) ≃ (toFibr w ≡ toFibr x) + fullEquiv = compEquiv (congEquiv (invEquiv Σ-assoc-≃)) (invEquiv (Σ≡PropEquiv (λ _ → isPropIsEmbedding))) + + EmbeddingIP : f≃g ≃ (f ≡ g) + EmbeddingIP = + f≃g + ≃⟨ strictIsoToEquiv (invIso toProdIso) ⟩ + (∀ b → (fiber ffun b → fiber gfun b) × (fiber gfun b → fiber ffun b)) + ≃⟨ equivΠCod (λ _ → isEquivPropBiimpl→Equiv (isEmbedding→hasPropFibers isEmbF _) + (isEmbedding→hasPropFibers isEmbG _)) ⟩ + (∀ b → (fiber (f .snd .fst) b) ≃ (fiber (g .snd .fst) b)) + ≃⟨ FibrationIP (toFibr f) (toFibr g) ⟩ + (toFibr f ≡ toFibr g) + ≃⟨ invEquiv (_ , isEmbeddingToFibr _ _) ⟩ + f ≡ g + ■ +open EmbeddingIdentityPrinciple renaming (f≃g to _≃Emb_) using (EmbeddingIP) public diff --git a/Cubical/Functions/Everything.agda b/Cubical/Functions/Everything.agda new file mode 100644 index 000000000..08579b121 --- /dev/null +++ b/Cubical/Functions/Everything.agda @@ -0,0 +1,13 @@ +{-# OPTIONS --safe --guardedness #-} +module Cubical.Functions.Everything where + +import Cubical.Functions.Bundle +import Cubical.Functions.Embedding +import Cubical.Functions.Fibration +import Cubical.Functions.Fixpoint +import Cubical.Functions.FunExtEquiv +import Cubical.Functions.Implicit +import Cubical.Functions.Involution +import Cubical.Functions.Logic +import Cubical.Functions.Morphism +import Cubical.Functions.Surjection diff --git a/Cubical/Functions/Fibration.agda b/Cubical/Functions/Fibration.agda new file mode 100644 index 000000000..704da5cb9 --- /dev/null +++ b/Cubical/Functions/Fibration.agda @@ -0,0 +1,105 @@ +{-# OPTIONS --safe #-} +module Cubical.Functions.Fibration where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels using (isOfHLevel→isOfHLevelDep) +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Data.Sigma + +private + variable + ℓ ℓb : Level + B : Type ℓb + +module FiberIso {ℓ} (p⁻¹ : B → Type ℓ) (x : B) where + + p : Σ B p⁻¹ → B + p = fst + + fwd : fiber p x → p⁻¹ x + fwd ((x' , y) , q) = subst (λ z → p⁻¹ z) q y + + bwd : p⁻¹ x → fiber p x + bwd y = (x , y) , refl + + fwd-bwd : ∀ x → fwd (bwd x) ≡ x + fwd-bwd y = transportRefl y + + bwd-fwd : ∀ x → bwd (fwd x) ≡ x + bwd-fwd ((x' , y) , q) i = h (r i) + where h : Σ[ s ∈ singl x ] p⁻¹ (s .fst) → fiber p x + h ((x , p) , y) = (x , y) , sym p + r : Path (Σ[ s ∈ singl x ] p⁻¹ (s .fst)) + ((x , refl ) , subst p⁻¹ q y) + ((x' , sym q) , y ) + r = ΣPathP (isContrSingl x .snd (x' , sym q) + , toPathP (transport⁻Transport (λ i → p⁻¹ (q i)) y)) + + -- HoTT Lemma 4.8.1 + fiberEquiv : fiber p x ≃ p⁻¹ x + fiberEquiv = isoToEquiv (iso fwd bwd fwd-bwd bwd-fwd) + +open FiberIso using (fiberEquiv) public + +module _ {ℓ} {E : Type ℓ} (p : E → B) where + + -- HoTT Lemma 4.8.2 + totalEquiv : E ≃ Σ B (fiber p) + totalEquiv = isoToEquiv isom + where isom : Iso E (Σ B (fiber p)) + Iso.fun isom x = p x , x , refl + Iso.inv isom (b , x , q) = x + Iso.leftInv isom x i = x + Iso.rightInv isom (b , x , q) i = q i , x , λ j → q (i ∧ j) + +module _ (B : Type ℓb) (ℓ : Level) where + private + ℓ' = ℓ-max ℓb ℓ + + -- HoTT Theorem 4.8.3 + fibrationEquiv : (Σ[ E ∈ Type ℓ' ] (E → B)) ≃ (B → Type ℓ') + fibrationEquiv = isoToEquiv isom + where isom : Iso (Σ[ E ∈ Type ℓ' ] (E → B)) (B → Type ℓ') + Iso.fun isom (E , p) = fiber p + Iso.inv isom p⁻¹ = Σ B p⁻¹ , fst + Iso.rightInv isom p⁻¹ i x = ua (fiberEquiv p⁻¹ x) i + Iso.leftInv isom (E , p) i = ua e (~ i) , fst ∘ ua-unglue e (~ i) + where e = totalEquiv p + + +module ForSets {E : Type ℓ} {isSetB : isSet B} (f : E → B) where + module _ {x x'} {px : x ≡ x'} {a' : fiber f x} {b' : fiber f x'} where + -- fibers are equal when their representatives are equal + fibersEqIfRepsEq : fst a' ≡ fst b' + → PathP (λ i → fiber f (px i)) a' b' + fibersEqIfRepsEq p = ΣPathP (p , + (isOfHLevel→isOfHLevelDep 1 + (λ (v , w) → isSetB (f v) w) + (snd a') (snd b') + (λ i → (p i , px i)))) +-- The path type in a fiber of f is equivalent to a fiber of (cong f) +open import Cubical.Foundations.Function + +fiberPath : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {f : A → B} {b : B} (h h' : fiber f b) → + (Σ[ p ∈ (fst h ≡ fst h') ] (PathP (λ i → f (p i) ≡ b) (snd h) (snd h'))) + ≡ fiber (cong f) (h .snd ∙∙ refl ∙∙ sym (h' .snd)) +fiberPath h h' = cong (Σ (h .fst ≡ h' .fst)) (funExt λ p → flipSquarePath ∙ PathP≡doubleCompPathʳ _ _ _ _) + +fiber≡ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {f : A → B} {b : B} (h h' : fiber f b) + → (h ≡ h') ≡ fiber (cong f) (h .snd ∙∙ refl ∙∙ sym (h' .snd)) +fiber≡ {f = f} {b = b} h h' = + ΣPath≡PathΣ ⁻¹ ∙ + fiberPath h h' + +FibrationStr : (B : Type ℓb) → Type ℓ → Type (ℓ-max ℓ ℓb) +FibrationStr B A = A → B + +Fibration : (B : Type ℓb) → (ℓ : Level) → Type (ℓ-max ℓb (ℓ-suc ℓ)) +Fibration {ℓb = ℓb} B ℓ = Σ[ A ∈ Type ℓ ] FibrationStr B A diff --git a/Cubical/Functions/Fixpoint.agda b/Cubical/Functions/Fixpoint.agda new file mode 100644 index 000000000..24eb44d95 --- /dev/null +++ b/Cubical/Functions/Fixpoint.agda @@ -0,0 +1,43 @@ +{- + Definition of function fixpoint and Kraus' lemma +-} +{-# OPTIONS --safe #-} +module Cubical.Functions.Fixpoint where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws + +private + variable + ℓ : Level + A : Type ℓ + +Fixpoint : (A → A) → Type _ +Fixpoint {A = A} f = Σ A (λ x → f x ≡ x) + +fixpoint : {f : A → A} → Fixpoint f → A +fixpoint = fst + +fixpointPath : {f : A → A} → (p : Fixpoint f) → f (fixpoint p) ≡ fixpoint p +fixpointPath = snd + +-- Kraus' lemma +-- a version not using cubical features can be found at +-- https://www.cs.bham.ac.uk/~mhe/GeneralizedHedberg/html/GeneralizedHedberg.html#21576 +2-Constant→isPropFixpoint : (f : A → A) → 2-Constant f → isProp (Fixpoint f) +2-Constant→isPropFixpoint f fconst (x , p) (y , q) i = s i , t i where + noose : ∀ x y → f x ≡ f y + noose x y = sym (fconst x x) ∙ fconst x y + -- the main idea is that for any path p, cong f p does not depend on p + -- but only on its endpoints and the structure of 2-Constant f + KrausInsight : ∀ {x y} → (p : x ≡ y) → noose x y ≡ cong f p + KrausInsight {x} = J (λ y p → noose x y ≡ cong f p) (lCancel (fconst x x)) + -- Need to solve for a path s : x ≡ y, such that: + -- transport (λ i → cong f s i ≡ s i) p ≡ q + s : x ≡ y + s = sym p ∙∙ noose x y ∙∙ q + t' : PathP (λ i → noose x y i ≡ s i) p q + t' i j = doubleCompPath-filler (sym p) (noose x y) q j i + t : PathP (λ i → cong f s i ≡ s i) p q + t = subst (λ kraus → PathP (λ i → kraus i ≡ s i) p q) (KrausInsight s) t' diff --git a/Cubical/Functions/FunExtEquiv.agda b/Cubical/Functions/FunExtEquiv.agda new file mode 100644 index 000000000..21a5efdab --- /dev/null +++ b/Cubical/Functions/FunExtEquiv.agda @@ -0,0 +1,193 @@ +{-# OPTIONS --safe #-} +module Cubical.Functions.FunExtEquiv where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.CartesianKanOps +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Vec.Base +open import Cubical.Data.Vec.NAry +open import Cubical.Data.Nat + +open import Cubical.Reflection.StrictEquiv + +private + variable + ℓ ℓ₁ ℓ₂ ℓ₃ : Level + +-- Function extensionality is an equivalence +module _ {A : Type ℓ} {B : A → I → Type ℓ₁} + {f : (x : A) → B x i0} {g : (x : A) → B x i1} where + + funExtEquiv : (∀ x → PathP (B x) (f x) (g x)) ≃ PathP (λ i → ∀ x → B x i) f g + unquoteDef funExtEquiv = defStrictEquiv funExtEquiv funExt funExt⁻ + + funExtPath : (∀ x → PathP (B x) (f x) (g x)) ≡ PathP (λ i → ∀ x → B x i) f g + funExtPath = ua funExtEquiv + + funExtIso : Iso (∀ x → PathP (B x) (f x) (g x)) (PathP (λ i → ∀ x → B x i) f g) + funExtIso = iso funExt funExt⁻ (λ x → refl {x = x}) (λ x → refl {x = x}) + +-- Function extensionality for binary functions +funExt₂ : {A : Type ℓ} {B : A → Type ℓ₁} {C : (x : A) → B x → I → Type ℓ₂} + {f : (x : A) → (y : B x) → C x y i0} + {g : (x : A) → (y : B x) → C x y i1} + → ((x : A) (y : B x) → PathP (C x y) (f x y) (g x y)) + → PathP (λ i → ∀ x y → C x y i) f g +funExt₂ p i x y = p x y i + +-- Function extensionality for binary functions is an equivalence +module _ {A : Type ℓ} {B : A → Type ℓ₁} {C : (x : A) → B x → I → Type ℓ₂} + {f : (x : A) → (y : B x) → C x y i0} + {g : (x : A) → (y : B x) → C x y i1} where + private + appl₂ : PathP (λ i → ∀ x y → C x y i) f g → ∀ x y → PathP (C x y) (f x y) (g x y) + appl₂ eq x y i = eq i x y + + funExt₂Equiv : (∀ x y → PathP (C x y) (f x y) (g x y)) ≃ (PathP (λ i → ∀ x y → C x y i) f g) + unquoteDef funExt₂Equiv = defStrictEquiv funExt₂Equiv funExt₂ appl₂ + + funExt₂Path : (∀ x y → PathP (C x y) (f x y) (g x y)) ≡ (PathP (λ i → ∀ x y → C x y i) f g) + funExt₂Path = ua funExt₂Equiv + + +-- Function extensionality for ternary functions +funExt₃ : {A : Type ℓ} {B : A → Type ℓ₁} {C : (x : A) → B x → Type ℓ₂} + {D : (x : A) → (y : B x) → C x y → I → Type ℓ₃} + {f : (x : A) → (y : B x) → (z : C x y) → D x y z i0} + {g : (x : A) → (y : B x) → (z : C x y) → D x y z i1} + → ((x : A) (y : B x) (z : C x y) → PathP (D x y z) (f x y z) (g x y z)) + → PathP (λ i → ∀ x y z → D x y z i) f g +funExt₃ p i x y z = p x y z i + +-- Function extensionality for ternary functions is an equivalence +module _ {A : Type ℓ} {B : A → Type ℓ₁} {C : (x : A) → B x → Type ℓ₂} + {D : (x : A) → (y : B x) → C x y → I → Type ℓ₃} + {f : (x : A) → (y : B x) → (z : C x y) → D x y z i0} + {g : (x : A) → (y : B x) → (z : C x y) → D x y z i1} where + private + appl₃ : PathP (λ i → ∀ x y z → D x y z i) f g → ∀ x y z → PathP (D x y z) (f x y z) (g x y z) + appl₃ eq x y z i = eq i x y z + + funExt₃Equiv : (∀ x y z → PathP (D x y z) (f x y z) (g x y z)) ≃ (PathP (λ i → ∀ x y z → D x y z i) f g) + unquoteDef funExt₃Equiv = defStrictEquiv funExt₃Equiv funExt₃ appl₃ + + funExt₃Path : (∀ x y z → PathP (D x y z) (f x y z) (g x y z)) ≡ (PathP (λ i → ∀ x y z → D x y z i) f g) + funExt₃Path = ua funExt₃Equiv + + +-- n-ary non-dependent funext +nAryFunExt : {X : Type ℓ} {Y : I → Type ℓ₁} (n : ℕ) (fX : nAryOp n X (Y i0)) (fY : nAryOp n X (Y i1)) + → ((xs : Vec X n) → PathP Y (fX $ⁿ xs) (fY $ⁿ map (λ x → x) xs)) + → PathP (λ i → nAryOp n X (Y i)) fX fY +nAryFunExt zero fX fY p = p [] +nAryFunExt (suc n) fX fY p i x = nAryFunExt n (fX x) (fY x) (λ xs → p (x ∷ xs)) i + +-- n-ary funext⁻ +nAryFunExt⁻ : (n : ℕ) {X : Type ℓ} {Y : I → Type ℓ₁} (fX : nAryOp n X (Y i0)) (fY : nAryOp n X (Y i1)) + → PathP (λ i → nAryOp n X (Y i)) fX fY + → ((xs : Vec X n) → PathP Y (fX $ⁿ xs) (fY $ⁿ map (λ x → x) xs)) +nAryFunExt⁻ zero fX fY p [] = p +nAryFunExt⁻ (suc n) fX fY p (x ∷ xs) = nAryFunExt⁻ n (fX x) (fY x) (λ i → p i x) xs + +nAryFunExtEquiv : (n : ℕ) {X : Type ℓ} {Y : I → Type ℓ₁} (fX : nAryOp n X (Y i0)) (fY : nAryOp n X (Y i1)) + → ((xs : Vec X n) → PathP Y (fX $ⁿ xs) (fY $ⁿ map (λ x → x) xs)) ≃ PathP (λ i → nAryOp n X (Y i)) fX fY +nAryFunExtEquiv n {X} {Y} fX fY = isoToEquiv (iso (nAryFunExt n fX fY) (nAryFunExt⁻ n fX fY) + (linv n fX fY) (rinv n fX fY)) + where + linv : (n : ℕ) (fX : nAryOp n X (Y i0)) (fY : nAryOp n X (Y i1)) + (p : PathP (λ i → nAryOp n X (Y i)) fX fY) + → nAryFunExt n fX fY (nAryFunExt⁻ n fX fY p) ≡ p + linv zero fX fY p = refl + linv (suc n) fX fY p i j x = linv n (fX x) (fY x) (λ k → p k x) i j + + rinv : (n : ℕ) (fX : nAryOp n X (Y i0)) (fY : nAryOp n X (Y i1)) + (p : (xs : Vec X n) → PathP Y (fX $ⁿ xs) (fY $ⁿ map (λ x → x) xs)) + → nAryFunExt⁻ n fX fY (nAryFunExt n fX fY p) ≡ p + rinv zero fX fY p i [] = p [] + rinv (suc n) fX fY p i (x ∷ xs) = rinv n (fX x) (fY x) (λ ys i → p (x ∷ ys) i) i xs + +-- Funext when the domain also depends on the interval + +funExtDep : {A : I → Type ℓ} {B : (i : I) → A i → Type ℓ₁} + {f : (x : A i0) → B i0 x} {g : (x : A i1) → B i1 x} + → ({x₀ : A i0} {x₁ : A i1} (p : PathP A x₀ x₁) → PathP (λ i → B i (p i)) (f x₀) (g x₁)) + → PathP (λ i → (x : A i) → B i x) f g +funExtDep {A = A} {B} {f} {g} h i x = + comp + (λ k → B i (coei→i A i x k)) + (λ k → λ + { (i = i0) → f (coei→i A i0 x k) + ; (i = i1) → g (coei→i A i1 x k) + }) + (h (λ j → coei→j A i j x) i) + +funExtDep⁻ : {A : I → Type ℓ} {B : (i : I) → A i → Type ℓ₁} + {f : (x : A i0) → B i0 x} {g : (x : A i1) → B i1 x} + → PathP (λ i → (x : A i) → B i x) f g + → ({x₀ : A i0} {x₁ : A i1} (p : PathP A x₀ x₁) → PathP (λ i → B i (p i)) (f x₀) (g x₁)) +funExtDep⁻ q p i = q i (p i) + +funExtDepEquiv : {A : I → Type ℓ} {B : (i : I) → A i → Type ℓ₁} + {f : (x : A i0) → B i0 x} {g : (x : A i1) → B i1 x} + → ({x₀ : A i0} {x₁ : A i1} (p : PathP A x₀ x₁) → PathP (λ i → B i (p i)) (f x₀) (g x₁)) + ≃ PathP (λ i → (x : A i) → B i x) f g +funExtDepEquiv {A = A} {B} {f} {g} = isoToEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun = funExtDep + isom .inv = funExtDep⁻ + isom .rightInv q m i x = + comp + (λ k → B i (coei→i A i x (k ∨ m))) + (λ k → λ + { (i = i0) → f (coei→i A i0 x (k ∨ m)) + ; (i = i1) → g (coei→i A i1 x (k ∨ m)) + ; (m = i1) → q i x + }) + (q i (coei→i A i x m)) + isom .leftInv h m p i = + comp + (λ k → B i (lemi→i m k)) + (λ k → λ + { (i = i0) → f (lemi→i m k) + ; (i = i1) → g (lemi→i m k) + ; (m = i1) → h p i + }) + (h (λ j → lemi→j j m) i) + where + lemi→j : ∀ j → coei→j A i j (p i) ≡ p j + lemi→j j = + coei→j (λ k → coei→j A i k (p i) ≡ p k) i j (coei→i A i (p i)) + + lemi→i : PathP (λ m → lemi→j i m ≡ p i) (coei→i A i (p i)) refl + lemi→i = + sym (coei→i (λ k → coei→j A i k (p i) ≡ p k) i (coei→i A i (p i))) + ◁ λ m k → lemi→j i (m ∨ k) + +heteroHomotopy≃Homotopy : {A : I → Type ℓ} {B : (i : I) → Type ℓ₁} + {f : A i0 → B i0} {g : A i1 → B i1} + → ({x₀ : A i0} {x₁ : A i1} → PathP A x₀ x₁ → PathP B (f x₀) (g x₁)) + ≃ ((x₀ : A i0) → PathP B (f x₀) (g (transport (λ i → A i) x₀))) +heteroHomotopy≃Homotopy {A = A} {B} {f} {g} = isoToEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun h x₀ = h (isContrSinglP A x₀ .fst .snd) + isom .inv k {x₀} {x₁} p = + subst (λ fib → PathP B (f x₀) (g (fib .fst))) (isContrSinglP A x₀ .snd (x₁ , p)) (k x₀) + isom .rightInv k = funExt λ x₀ → + cong (λ α → subst (λ fib → PathP B (f x₀) (g (fib .fst))) α (k x₀)) + (isProp→isSet isPropSinglP (isContrSinglP A x₀ .fst) _ + (isContrSinglP A x₀ .snd (isContrSinglP A x₀ .fst)) + refl) + ∙ transportRefl (k x₀) + isom .leftInv h j {x₀} {x₁} p = + transp + (λ i → PathP B (f x₀) (g (isContrSinglP A x₀ .snd (x₁ , p) (i ∨ j) .fst))) + j + (h (isContrSinglP A x₀ .snd (x₁ , p) j .snd)) diff --git a/Cubical/Functions/Implicit.agda b/Cubical/Functions/Implicit.agda new file mode 100644 index 000000000..a98a2d3a6 --- /dev/null +++ b/Cubical/Functions/Implicit.agda @@ -0,0 +1,17 @@ +{-# OPTIONS --safe #-} +module Cubical.Functions.Implicit where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +implicit≃Explicit : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} + → ({a : A} → B a) ≃ ((a : A) → B a) +implicit≃Explicit = isoToEquiv isom + where + isom : Iso _ _ + Iso.fun isom f a = f + Iso.inv isom f = f _ + Iso.rightInv isom f = funExt λ _ → refl + Iso.leftInv isom f = implicitFunExt refl + diff --git a/Cubical/Functions/Involution.agda b/Cubical/Functions/Involution.agda new file mode 100644 index 000000000..9f9e3bfa9 --- /dev/null +++ b/Cubical/Functions/Involution.agda @@ -0,0 +1,42 @@ +{-# OPTIONS --safe #-} + +module Cubical.Functions.Involution where + +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence + +isInvolution : ∀{ℓ} {A : Type ℓ} → (A → A) → Type _ +isInvolution f = ∀ x → f (f x) ≡ x + +module _ {ℓ} {A : Type ℓ} {f : A → A} (invol : isInvolution f) where + open Iso + + involIso : Iso A A + involIso .fun = f + involIso .inv = f + involIso .rightInv = invol + involIso .leftInv = invol + + involIsEquiv : isEquiv f + involIsEquiv = isoToIsEquiv involIso + + involEquiv : A ≃ A + involEquiv = f , involIsEquiv + + involPath : A ≡ A + involPath = ua involEquiv + + involEquivComp : compEquiv involEquiv involEquiv ≡ idEquiv A + involEquivComp + = equivEq (λ i x → invol x i) + + involPathComp : involPath ∙ involPath ≡ refl + involPathComp + = sym (uaCompEquiv involEquiv involEquiv) ∙∙ cong ua involEquivComp ∙∙ uaIdEquiv + + involPath² : Square involPath refl refl involPath + involPath² + = subst (λ s → Square involPath s refl involPath) + involPathComp (compPath-filler involPath involPath) diff --git a/Cubical/Foundations/Logic.agda b/Cubical/Functions/Logic.agda similarity index 53% rename from Cubical/Foundations/Logic.agda rename to Cubical/Functions/Logic.agda index 2d966ea3c..3e5439c20 100644 --- a/Cubical/Foundations/Logic.agda +++ b/Cubical/Functions/Logic.agda @@ -1,22 +1,45 @@ -{-# OPTIONS --cubical --safe #-} - -module Cubical.Foundations.Logic where - -import Cubical.Data.Empty as D -import Cubical.Data.Prod as D -import Cubical.Data.Sum as D -import Cubical.Data.Unit as D +-- Various functions for manipulating hProps. +-- +-- This file used to be part of Foundations, but it turned out to be +-- not very useful so moved here. Feel free to upstream content. +-- +-- Note that it is often a bad idea to use hProp instead of having the +-- isProp proof separate. The reason is that Agda can rarely infer +-- isProp proofs making it easier to just give them explicitly instead +-- of having them bundled up with the type. +-- +{-# OPTIONS --safe #-} +module Cubical.Functions.Logic where open import Cubical.Foundations.Prelude - -open import Cubical.HITs.PropositionalTruncation - -open import Cubical.Foundations.HLevels using (hProp; ΣProp≡; isPropPi) public +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Structure open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence using (hPropExt) + +import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sum as ⊎ using (_⊎_) +open import Cubical.Data.Unit +open import Cubical.Data.Sigma + +open import Cubical.HITs.PropositionalTruncation as PropTrunc open import Cubical.Relation.Nullary hiding (¬_) + +-------------------------------------------------------------------------------- +-- The type hProp of mere propositions +-- the definition hProp is given in Foundations.HLevels +-- hProp ℓ = Σ (Type ℓ) isProp + +private + variable + ℓ ℓ' ℓ'' : Level + P Q R : hProp ℓ + A B C : Type ℓ + infix 10 ¬_ infixr 8 _⊔_ infixr 8 _⊔′_ @@ -36,71 +59,59 @@ infix 2 ∀[]-syntax infix 2 ⇒∶_⇐∶_ infix 2 ⇐∶_⇒∶_ --------------------------------------------------------------------------------- --- The type hProp of mere propositions --- the definition hProp is given in Foundations.HLevels --- hProp {ℓ} = Σ (Type ℓ) isProp - -private - variable - ℓ ℓ' ℓ'' : Level - P Q R : hProp ℓ - A B C : Type ℓ - -[_] : hProp ℓ → Type ℓ -[_] = fst - ∥_∥ₚ : Type ℓ → hProp ℓ -∥ A ∥ₚ = ∥ A ∥ , propTruncIsProp +∥ A ∥ₚ = ∥ A ∥ , isPropPropTrunc _≡ₚ_ : (x y : A) → hProp _ x ≡ₚ y = ∥ x ≡ y ∥ₚ -hProp≡ : [ P ] ≡ [ Q ] → P ≡ Q -hProp≡ p = ΣProp≡ (\ _ → isPropIsProp) p +hProp≡ : ⟨ P ⟩ ≡ ⟨ Q ⟩ → P ≡ Q +hProp≡ = TypeOfHLevel≡ 1 + +isProp⟨⟩ : (A : hProp ℓ) → isProp ⟨ A ⟩ +isProp⟨⟩ = snd -------------------------------------------------------------------------------- -- Logical implication of mere propositions _⇒_ : (A : hProp ℓ) → (B : hProp ℓ') → hProp _ -A ⇒ B = ([ A ] → [ B ]) , isPropPi λ _ → B .snd +A ⇒ B = (⟨ A ⟩ → ⟨ B ⟩) , isPropΠ λ _ → isProp⟨⟩ B -⇔toPath : [ P ⇒ Q ] → [ Q ⇒ P ] → P ≡ Q -⇔toPath {P = P} {Q = Q} P⇒Q Q⇒P = hProp≡ (isoToPath - (iso P⇒Q Q⇒P (λ b → Q .snd (P⇒Q (Q⇒P b)) b) λ a → P .snd (Q⇒P (P⇒Q a)) a)) +⇔toPath : ⟨ P ⇒ Q ⟩ → ⟨ Q ⇒ P ⟩ → P ≡ Q +⇔toPath {P = P} {Q = Q} P⇒Q Q⇒P = hProp≡ (hPropExt (isProp⟨⟩ P) (isProp⟨⟩ Q) P⇒Q Q⇒P) -pathTo⇒ : P ≡ Q → [ P ⇒ Q ] +pathTo⇒ : P ≡ Q → ⟨ P ⇒ Q ⟩ pathTo⇒ p x = subst fst p x -pathTo⇐ : P ≡ Q → [ Q ⇒ P ] +pathTo⇐ : P ≡ Q → ⟨ Q ⇒ P ⟩ pathTo⇐ p x = subst fst (sym p) x -substₚ : {x y : A} (B : A → hProp ℓ) → [ x ≡ₚ y ⇒ B x ⇒ B y ] -substₚ {x = x} {y = y} B = elimPropTrunc (λ _ → isPropPi λ _ → B y .snd) (subst (fst ∘ B)) +substₚ : {x y : A} (B : A → hProp ℓ) → ⟨ x ≡ₚ y ⇒ B x ⇒ B y ⟩ +substₚ {x = x} {y = y} B = PropTrunc.elim (λ _ → isPropΠ λ _ → isProp⟨⟩ (B y)) (subst (fst ∘ B)) -------------------------------------------------------------------------------- -- Mixfix notations for ⇔-toPath -- see ⊔-identityˡ and ⊔-identityʳ for the difference -⇒∶_⇐∶_ : [ P ⇒ Q ] → [ Q ⇒ P ] → P ≡ Q +⇒∶_⇐∶_ : ⟨ P ⇒ Q ⟩ → ⟨ Q ⇒ P ⟩ → P ≡ Q ⇒∶_⇐∶_ = ⇔toPath -⇐∶_⇒∶_ : [ Q ⇒ P ] → [ P ⇒ Q ] → P ≡ Q +⇐∶_⇒∶_ : ⟨ Q ⇒ P ⟩ → ⟨ P ⇒ Q ⟩ → P ≡ Q ⇐∶ g ⇒∶ f = ⇔toPath f g -------------------------------------------------------------------------------- -- False and True ⊥ : hProp _ -⊥ = D.⊥ , λ () +⊥ = ⊥.⊥ , λ () -⊤ : hProp _ -⊤ = D.Unit , (λ _ _ _ → D.tt) +⊤ : ∀ {ℓ} → hProp ℓ +⊤ = Unit* , (λ _ _ _ → tt*) -------------------------------------------------------------------------------- -- Pseudo-complement of mere propositions ¬_ : hProp ℓ → hProp _ -¬ A = ([ A ] → D.⊥) , isPropPi λ _ → D.isProp⊥ +¬ A = (⟨ A ⟩ → ⊥.⊥) , isPropΠ λ _ → ⊥.isProp⊥ _≢ₚ_ : (x y : A) → hProp _ x ≢ₚ y = ¬ x ≡ₚ y @@ -109,32 +120,32 @@ x ≢ₚ y = ¬ x ≡ₚ y -- Disjunction of mere propositions _⊔′_ : Type ℓ → Type ℓ' → Type _ -A ⊔′ B = ∥ A D.⊎ B ∥ +A ⊔′ B = ∥ A ⊎ B ∥ _⊔_ : hProp ℓ → hProp ℓ' → hProp _ -P ⊔ Q = ∥ [ P ] D.⊎ [ Q ] ∥ₚ +P ⊔ Q = ∥ ⟨ P ⟩ ⊎ ⟨ Q ⟩ ∥ₚ inl : A → A ⊔′ B -inl x = ∣ D.inl x ∣ +inl x = ∣ ⊎.inl x ∣ inr : B → A ⊔′ B -inr x = ∣ D.inr x ∣ +inr x = ∣ ⊎.inr x ∣ -⊔-elim : (P : hProp ℓ) (Q : hProp ℓ') (R : [ P ⊔ Q ] → hProp ℓ'') - → (∀ x → [ R (inl x) ]) → (∀ y → [ R (inr y) ]) → (∀ z → [ R z ]) -⊔-elim _ _ R P⇒R Q⇒R = elimPropTrunc (snd ∘ R) (D.elim-⊎ P⇒R Q⇒R) +⊔-elim : (P : hProp ℓ) (Q : hProp ℓ') (R : ⟨ P ⊔ Q ⟩ → hProp ℓ'') + → (∀ x → ⟨ R (inl x) ⟩) → (∀ y → ⟨ R (inr y) ⟩) → (∀ z → ⟨ R z ⟩) +⊔-elim _ _ R P⇒R Q⇒R = PropTrunc.elim (snd ∘ R) (⊎.elim P⇒R Q⇒R) -------------------------------------------------------------------------------- -- Conjunction of mere propositions _⊓′_ : Type ℓ → Type ℓ' → Type _ -A ⊓′ B = A D.× B +A ⊓′ B = A × B _⊓_ : hProp ℓ → hProp ℓ' → hProp _ -A ⊓ B = [ A ] ⊓′ [ B ] , D.hLevelProd 1 (A .snd) (B .snd) +A ⊓ B = ⟨ A ⟩ ⊓′ ⟨ B ⟩ , isOfHLevelΣ 1 (isProp⟨⟩ A) (\ _ → isProp⟨⟩ B) -⊓-intro : (P : hProp ℓ) (Q : [ P ] → hProp ℓ') (R : [ P ] → hProp ℓ'') - → (∀ a → [ Q a ]) → (∀ a → [ R a ]) → (∀ (a : [ P ]) → [ Q a ⊓ R a ] ) -⊓-intro _ _ _ = D.intro-× +⊓-intro : (P : hProp ℓ) (Q : ⟨ P ⟩ → hProp ℓ') (R : ⟨ P ⟩ → hProp ℓ'') + → (∀ a → ⟨ Q a ⟩) → (∀ a → ⟨ R a ⟩) → (∀ (a : ⟨ P ⟩) → ⟨ Q a ⊓ R a ⟩ ) +⊓-intro _ _ _ = \ f g a → f a , g a -------------------------------------------------------------------------------- -- Logical bi-implication of mere propositions @@ -142,43 +153,47 @@ A ⊓ B = [ A ] ⊓′ [ B ] , D.hLevelProd 1 (A .snd) (B .snd) _⇔_ : hProp ℓ → hProp ℓ' → hProp _ A ⇔ B = (A ⇒ B) ⊓ (B ⇒ A) +⇔-id : (P : hProp ℓ) → ⟨ P ⇔ P ⟩ +⇔-id P = (idfun ⟨ P ⟩) , (idfun ⟨ P ⟩) + -------------------------------------------------------------------------------- -- Universal Quantifier ∀[∶]-syntax : (A → hProp ℓ) → hProp _ -∀[∶]-syntax {A = A} P = (∀ x → [ P x ]) , isPropPi (snd ∘ P) +∀[∶]-syntax {A = A} P = (∀ x → ⟨ P x ⟩) , isPropΠ (isProp⟨⟩ ∘ P) ∀[]-syntax : (A → hProp ℓ) → hProp _ -∀[]-syntax {A = A} P = (∀ x → [ P x ]) , isPropPi (snd ∘ P) +∀[]-syntax {A = A} P = (∀ x → ⟨ P x ⟩) , isPropΠ (isProp⟨⟩ ∘ P) syntax ∀[∶]-syntax {A = A} (λ a → P) = ∀[ a ∶ A ] P syntax ∀[]-syntax (λ a → P) = ∀[ a ] P + -------------------------------------------------------------------------------- -- Existential Quantifier - ∃[]-syntax : (A → hProp ℓ) → hProp _ -∃[]-syntax {A = A} P = ∥ Σ A (fst ∘ P) ∥ₚ +∃[]-syntax {A = A} P = ∥ Σ A (⟨_⟩ ∘ P) ∥ₚ ∃[∶]-syntax : (A → hProp ℓ) → hProp _ -∃[∶]-syntax {A = A} P = ∥ Σ A (fst ∘ P) ∥ₚ +∃[∶]-syntax {A = A} P = ∥ Σ A (⟨_⟩ ∘ P) ∥ₚ + +syntax ∃[∶]-syntax {A = A} (λ x → P) = ∃[ x ∶ A ] P +syntax ∃[]-syntax (λ x → P) = ∃[ x ] P -syntax ∃[]-syntax {A = A} (λ x → P) = ∃[ x ∶ A ] P -syntax ∃[∶]-syntax (λ x → P) = ∃[ x ] P -------------------------------------------------------------------------------- -- Decidable mere proposition Decₚ : (P : hProp ℓ) → hProp ℓ -Decₚ P = Dec [ P ] , isPropDec (snd P) +Decₚ P = Dec ⟨ P ⟩ , isPropDec (isProp⟨⟩ P) -------------------------------------------------------------------------------- -- Negation commutes with truncation -∥¬A∥≡¬∥A∥ : (A : Type ℓ) → ∥ (A → D.⊥) ∥ₚ ≡ (¬ ∥ A ∥ₚ) +∥¬A∥≡¬∥A∥ : (A : Type ℓ) → ∥ (A → ⊥.⊥) ∥ₚ ≡ (¬ ∥ A ∥ₚ) ∥¬A∥≡¬∥A∥ _ = - ⇒∶ (λ ¬A A → elimPropTrunc (λ _ → D.isProp⊥) - (elimPropTrunc (λ _ → isPropPi λ _ → D.isProp⊥) (λ ¬p p → ¬p p) ¬A) A) + ⇒∶ (λ ¬A A → PropTrunc.elim (λ _ → ⊥.isProp⊥) + (PropTrunc.elim (λ _ → isPropΠ λ _ → ⊥.isProp⊥) (λ ¬p p → ¬p p) ¬A) A) ⇐∶ λ ¬p → ∣ (λ p → ¬p ∣ p ∣) ∣ -------------------------------------------------------------------------------- @@ -194,11 +209,11 @@ Decₚ P = Dec [ P ] , isPropDec (snd P) ⇐∶ assoc2 where assoc2 : (A ⊔′ B) ⊔′ C → A ⊔′ (B ⊔′ C) - assoc2 ∣ D.inr a ∣ = ∣ D.inr ∣ D.inr a ∣ ∣ - assoc2 ∣ D.inl ∣ D.inr b ∣ ∣ = ∣ D.inr ∣ D.inl b ∣ ∣ - assoc2 ∣ D.inl ∣ D.inl c ∣ ∣ = ∣ D.inl c ∣ - assoc2 ∣ D.inl (squash x y i) ∣ = propTruncIsProp (assoc2 ∣ D.inl x ∣) (assoc2 ∣ D.inl y ∣) i - assoc2 (squash x y i) = propTruncIsProp (assoc2 x) (assoc2 y) i + assoc2 ∣ ⊎.inr a ∣ = ∣ ⊎.inr ∣ ⊎.inr a ∣ ∣ + assoc2 ∣ ⊎.inl ∣ ⊎.inr b ∣ ∣ = ∣ ⊎.inr ∣ ⊎.inl b ∣ ∣ + assoc2 ∣ ⊎.inl ∣ ⊎.inl c ∣ ∣ = ∣ ⊎.inl c ∣ + assoc2 ∣ ⊎.inl (squash x y i) ∣ = isPropPropTrunc (assoc2 ∣ ⊎.inl x ∣) (assoc2 ∣ ⊎.inl y ∣) i + assoc2 (squash x y i) = isPropPropTrunc (assoc2 x) (assoc2 y) i ⊔-idem : (P : hProp ℓ) → P ⊔ P ≡ P ⊔-idem P = @@ -224,20 +239,20 @@ Decₚ P = Dec [ P ] , isPropDec (snd P) ⊓-assoc : (P : hProp ℓ) (Q : hProp ℓ') (R : hProp ℓ'') → P ⊓ Q ⊓ R ≡ (P ⊓ Q) ⊓ R ⊓-assoc _ _ _ = - ⇒∶ (λ {(x D., (y D., z)) → (x D., y) D., z}) - ⇐∶ (λ {((x D., y) D., z) → x D., (y D., z) }) + ⇒∶ (λ {(x , (y , z)) → (x , y) , z}) + ⇐∶ (λ {((x , y) , z) → x , (y , z) }) ⊓-comm : (P : hProp ℓ) (Q : hProp ℓ') → P ⊓ Q ≡ Q ⊓ P -⊓-comm _ _ = ⇔toPath D.swap D.swap +⊓-comm _ _ = ⇔toPath (\ p → p .snd , p .fst) (\ p → p .snd , p .fst) ⊓-idem : (P : hProp ℓ) → P ⊓ P ≡ P -⊓-idem _ = ⇔toPath D.proj₁ (λ x → x D., x) +⊓-idem _ = ⇔toPath fst (λ x → x , x) -⊓-identityˡ : (P : hProp ℓ) → ⊤ ⊓ P ≡ P -⊓-identityˡ _ = ⇔toPath D.proj₂ λ x → D.tt D., x +⊓-identityˡ : (P : hProp ℓ) → ⊤ {ℓ} ⊓ P ≡ P +⊓-identityˡ _ = ⇔toPath snd λ x → tt* , x -⊓-identityʳ : (P : hProp ℓ) → P ⊓ ⊤ ≡ P -⊓-identityʳ _ = ⇔toPath D.proj₁ λ x → x D., D.tt +⊓-identityʳ : (P : hProp ℓ) → P ⊓ ⊤ {ℓ} ≡ P +⊓-identityʳ _ = ⇔toPath fst λ x → x , tt* -------------------------------------------------------------------------------- -- Distributive laws @@ -245,31 +260,31 @@ Decₚ P = Dec [ P ] , isPropDec (snd P) ⇒-⊓-distrib : (P : hProp ℓ) (Q : hProp ℓ')(R : hProp ℓ'') → P ⇒ (Q ⊓ R) ≡ (P ⇒ Q) ⊓ (P ⇒ R) ⇒-⊓-distrib _ _ _ = - ⇒∶ (λ f → (D.proj₁ ∘ f) D., (D.proj₂ ∘ f)) - ⇐∶ (λ { (f D., g) x → f x D., g x}) + ⇒∶ (λ f → (fst ∘ f) , (snd ∘ f)) + ⇐∶ (λ { (f , g) x → f x , g x}) ⊓-⊔-distribˡ : (P : hProp ℓ) (Q : hProp ℓ')(R : hProp ℓ'') → P ⊓ (Q ⊔ R) ≡ (P ⊓ Q) ⊔ (P ⊓ R) ⊓-⊔-distribˡ P Q R = - ⇒∶ (λ { (x D., a) → ⊔-elim Q R (λ _ → (P ⊓ Q) ⊔ (P ⊓ R)) - (λ y → ∣ D.inl (x D., y) ∣ ) - (λ z → ∣ D.inr (x D., z) ∣ ) a }) + ⇒∶ (λ { (x , a) → ⊔-elim Q R (λ _ → (P ⊓ Q) ⊔ (P ⊓ R)) + (λ y → ∣ ⊎.inl (x , y) ∣ ) + (λ z → ∣ ⊎.inr (x , z) ∣ ) a }) ⇐∶ ⊔-elim (P ⊓ Q) (P ⊓ R) (λ _ → P ⊓ Q ⊔ R) - (λ y → D.proj₁ y D., inl (D.proj₂ y)) - (λ z → D.proj₁ z D., inr (D.proj₂ z)) + (λ y → fst y , inl (snd y)) + (λ z → fst z , inr (snd z)) ⊔-⊓-distribˡ : (P : hProp ℓ) (Q : hProp ℓ')(R : hProp ℓ'') → P ⊔ (Q ⊓ R) ≡ (P ⊔ Q) ⊓ (P ⊔ R) ⊔-⊓-distribˡ P Q R = ⇒∶ ⊔-elim P (Q ⊓ R) (λ _ → (P ⊔ Q) ⊓ (P ⊔ R) ) - (D.intro-× inl inl) (D.map-× inr inr) + (\ x → inl x , inl x) (\ p → inr (p .fst) , inr (p .snd)) - ⇐∶ (λ { (x D., y) → ⊔-elim P R (λ _ → P ⊔ Q ⊓ R) inl - (λ z → ⊔-elim P Q (λ _ → P ⊔ Q ⊓ R) inl (λ y → inr (y D., z)) x) y }) + ⇐∶ (λ { (x , y) → ⊔-elim P R (λ _ → P ⊔ Q ⊓ R) inl + (λ z → ⊔-elim P Q (λ _ → P ⊔ Q ⊓ R) inl (λ y → inr (y , z)) x) y }) ⊓-∀-distrib : (P : A → hProp ℓ) (Q : A → hProp ℓ') → (∀[ a ∶ A ] P a) ⊓ (∀[ a ∶ A ] Q a) ≡ (∀[ a ∶ A ] (P a ⊓ Q a)) ⊓-∀-distrib P Q = - ⇒∶ (λ {(p D., q) a → p a D., q a}) - ⇐∶ λ pq → (D.proj₁ ∘ pq ) D., (D.proj₂ ∘ pq) + ⇒∶ (λ {(p , q) a → p a , q a}) + ⇐∶ λ pq → (fst ∘ pq ) , (snd ∘ pq) diff --git a/Cubical/Functions/Morphism.agda b/Cubical/Functions/Morphism.agda new file mode 100644 index 000000000..c4886cfdf --- /dev/null +++ b/Cubical/Functions/Morphism.agda @@ -0,0 +1,84 @@ +{- + General lemmas about morphisms (defined as loosely as possible) +-} +{-# OPTIONS --safe #-} +module Cubical.Functions.Morphism where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function + +open Iso +module ax {ℓ : Level} (A : Type ℓ) (_+A_ : A → A → A) (a₀ : A) where + rUnit = (a : A) → a +A a₀ ≡ a + lUnit = (a : A) → a₀ +A a ≡ a + + rCancel : (-A_ : A → A) → Type ℓ + rCancel -A_ = (a : A) → a +A (-A a) ≡ a₀ + + lCancel : (-A_ : A → A) → Type ℓ + lCancel -A_ = (a : A) → (-A a) +A a ≡ a₀ + + assoc = (x y z : A) → x +A (y +A z) ≡ ((x +A y) +A z) + + comm = (x y : A) → x +A y ≡ y +A x + +module morphLemmas {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} + (_+A_ : A → A → A) (_+B_ : B → B → B) + (f : A → B) (f-hom : (x y : A) → f (x +A y) ≡ f x +B f y) + where + + 0↦0 : (a₀ : A) (b₀ : B) (-A_ : A → A) (-B_ : B → B) + → ax.rUnit A _+A_ a₀ + → ax.rUnit B _+B_ b₀ + → ax.rCancel B _+B_ b₀ -B_ + → ax.assoc B _+B_ b₀ + → f a₀ ≡ b₀ + 0↦0 a₀ b₀ -A_ -B_ rUnitA rUnitB rCancelB assocB = + sym (rUnitB (f a₀)) + ∙∙ cong (f a₀ +B_) (sym (rCancelB (f a₀))) + ∙∙ assocB (f a₀) (f a₀) (-B f a₀) + ∙∙ sym (cong (_+B (-B f a₀)) (cong f (sym (rUnitA a₀)) ∙ f-hom a₀ a₀)) + ∙∙ rCancelB (f a₀) + + distrMinus : (a₀ : A) (b₀ : B) (-A_ : A → A) (-B_ : B → B) + → ax.lUnit B _+B_ b₀ + → ax.rUnit B _+B_ b₀ + → ax.lCancel A _+A_ a₀ -A_ + → ax.rCancel B _+B_ b₀ -B_ + → ax.assoc B _+B_ b₀ + → (0↦0 : f a₀ ≡ b₀) + → (x : A) → f (-A x) ≡ -B (f x) + distrMinus a₀ b₀ -A_ -B_ lUnitB rUnitB lCancelA rCancelB assocB 0↦0 x = + sym (rUnitB _) + ∙∙ cong (f (-A x) +B_) (sym (rCancelB (f x))) + ∙∙ assocB _ _ _ + ∙∙ cong (_+B (-B (f x))) (sym (f-hom (-A x) x) ∙∙ cong f (lCancelA x) ∙∙ 0↦0) + ∙∙ lUnitB _ + + distrMinus' : (a₀ : A) (b₀ : B) (-A_ : A → A) (-B_ : B → B) + → ax.lUnit B _+B_ b₀ + → ax.rUnit B _+B_ b₀ + → ax.rUnit A _+A_ a₀ + → ax.lCancel A _+A_ a₀ -A_ + → ax.rCancel B _+B_ b₀ -B_ + → ax.assoc A _+A_ a₀ + → ax.assoc B _+B_ b₀ + → f a₀ ≡ b₀ -- not really needed, but it can be useful to specify the proof yourself + → (x y : A) + → f (x +A (-A y)) ≡ (f x +B (-B (f y))) + distrMinus' a₀ b₀ -A_ -B_ lUnitB rUnitB rUnitA lCancelA rCancelB assocA assocB 0↦0 x y = + sym (rUnitB _) + ∙∙ cong (f (x +A (-A y)) +B_) (sym (rCancelB (f y))) ∙ assocB _ _ _ + ∙∙ cong (_+B (-B f y)) (sym (f-hom (x +A (-A y)) y) + ∙ cong f (sym (assocA x (-A y) y) + ∙∙ cong (x +A_) (lCancelA y) + ∙∙ rUnitA x)) + + isMorphInv : (g : B → A) → section f g → retract f g + → (x y : B) + → (g (x +B y) ≡ (g x +A g y)) + isMorphInv g sect retr x y = + cong g (cong₂ _+B_ (sym (sect x)) (sym (sect y))) + ∙∙ cong g (sym (f-hom (g x) (g y))) + ∙∙ retr (g x +A g y) diff --git a/Cubical/Functions/Surjection.agda b/Cubical/Functions/Surjection.agda new file mode 100644 index 000000000..1c833049a --- /dev/null +++ b/Cubical/Functions/Surjection.agda @@ -0,0 +1,80 @@ +{-# OPTIONS --safe #-} +module Cubical.Functions.Surjection where + +open import Cubical.Core.Everything +open import Cubical.Data.Sigma +open import Cubical.Data.Unit +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Functions.Embedding +open import Cubical.HITs.PropositionalTruncation as PropTrunc + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + f : A → B + +isSurjection : (A → B) → Type _ +isSurjection f = ∀ b → ∥ fiber f b ∥ + +_↠_ : Type ℓ → Type ℓ' → Type (ℓ-max ℓ ℓ') +A ↠ B = Σ[ f ∈ (A → B) ] isSurjection f + +section→isSurjection : {g : B → A} → section f g → isSurjection f +section→isSurjection {g = g} s b = ∣ g b , s b ∣ + +isPropIsSurjection : isProp (isSurjection f) +isPropIsSurjection = isPropΠ λ _ → squash + +isEquiv→isSurjection : isEquiv f → isSurjection f +isEquiv→isSurjection e b = ∣ fst (equiv-proof e b) ∣ + +isEquiv→isEmbedding×isSurjection : isEquiv f → isEmbedding f × isSurjection f +isEquiv→isEmbedding×isSurjection e = isEquiv→isEmbedding e , isEquiv→isSurjection e + +isEmbedding×isSurjection→isEquiv : isEmbedding f × isSurjection f → isEquiv f +equiv-proof (isEmbedding×isSurjection→isEquiv {f = f} (emb , sur)) b = + inhProp→isContr (PropTrunc.rec fib' (λ x → x) fib) fib' + where + hpf : hasPropFibers f + hpf = isEmbedding→hasPropFibers emb + + fib : ∥ fiber f b ∥ + fib = sur b + + fib' : isProp (fiber f b) + fib' = hpf b + +isEquiv≃isEmbedding×isSurjection : isEquiv f ≃ isEmbedding f × isSurjection f +isEquiv≃isEmbedding×isSurjection = isoToEquiv (iso + isEquiv→isEmbedding×isSurjection + isEmbedding×isSurjection→isEquiv + (λ _ → isOfHLevelΣ 1 isPropIsEmbedding (\ _ → isPropIsSurjection) _ _) + (λ _ → isPropIsEquiv _ _ _)) + +-- obs: for epi⇒surjective to go through we require a stronger +-- hypothesis that one would expect: +-- f must cancel functions from a higher universe. +rightCancellable : (f : A → B) → Type _ +rightCancellable {ℓ} {A} {ℓ'} {B} f = ∀ {C : Type (ℓ-suc (ℓ-max ℓ ℓ'))} + → ∀ (g g' : B → C) → (∀ x → g (f x) ≡ g' (f x)) → ∀ y → g y ≡ g' y + +-- This statement is in Mac Lane & Moerdijk (page 143, corollary 5). +epi⇒surjective : (f : A → B) → rightCancellable f → isSurjection f +epi⇒surjective f rc y = transport (fact₂ y) tt* + where hasPreimage : (A → B) → B → _ + hasPreimage f y = ∥ fiber f y ∥ + + fact₁ : ∀ x → Unit* ≡ hasPreimage f (f x) + fact₁ x = hPropExt isPropUnit* + isPropPropTrunc + (λ _ → ∣ (x , refl) ∣) + (λ _ → tt*) + + fact₂ : ∀ y → Unit* ≡ hasPreimage f y + fact₂ = rc _ _ fact₁ diff --git a/Cubical/HITs/2GroupoidTruncation.agda b/Cubical/HITs/2GroupoidTruncation.agda index 5b690a474..8178890b8 100644 --- a/Cubical/HITs/2GroupoidTruncation.agda +++ b/Cubical/HITs/2GroupoidTruncation.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.2GroupoidTruncation where open import Cubical.HITs.2GroupoidTruncation.Base public diff --git a/Cubical/HITs/2GroupoidTruncation/Base.agda b/Cubical/HITs/2GroupoidTruncation/Base.agda index e8f0909bc..0a83ce4ca 100644 --- a/Cubical/HITs/2GroupoidTruncation/Base.agda +++ b/Cubical/HITs/2GroupoidTruncation/Base.agda @@ -5,13 +5,13 @@ This file contains: - Definition of 2-groupoid truncations -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.2GroupoidTruncation.Base where open import Cubical.Foundations.Prelude -- 2-groupoid truncation as a higher inductive type: -data ∥_∥₂ {ℓ} (A : Type ℓ) : Type ℓ where - ∣_∣₂ : A → ∥ A ∥₂ - squash₂ : ∀ (x y : ∥ A ∥₂) (p q : x ≡ y) (r s : p ≡ q) (t u : r ≡ s) → t ≡ u +data ∥_∥₄ {ℓ} (A : Type ℓ) : Type ℓ where + ∣_∣₄ : A → ∥ A ∥₄ + squash₄ : ∀ (x y : ∥ A ∥₄) (p q : x ≡ y) (r s : p ≡ q) (t u : r ≡ s) → t ≡ u diff --git a/Cubical/HITs/2GroupoidTruncation/Properties.agda b/Cubical/HITs/2GroupoidTruncation/Properties.agda index 88805968b..30dd85a57 100644 --- a/Cubical/HITs/2GroupoidTruncation/Properties.agda +++ b/Cubical/HITs/2GroupoidTruncation/Properties.agda @@ -5,77 +5,64 @@ This file contains: - Properties of 2-groupoid truncations -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.2GroupoidTruncation.Properties where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + open import Cubical.HITs.2GroupoidTruncation.Base -rec2GroupoidTrunc : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (gB : is2Groupoid B) → (A → B) → (∥ A ∥₂ → B) -rec2GroupoidTrunc gB f ∣ x ∣₂ = f x -rec2GroupoidTrunc gB f (squash₂ _ _ _ _ _ _ t u i j k l) = - gB _ _ _ _ _ _ - (λ m n o → rec2GroupoidTrunc gB f (t m n o)) - (λ m n o → rec2GroupoidTrunc gB f (u m n o)) +private + variable + ℓ : Level + A : Type ℓ + +rec : ∀ {B : Type ℓ} → is2Groupoid B → (A → B) → ∥ A ∥₄ → B +rec gB f ∣ x ∣₄ = f x +rec gB f (squash₄ _ _ _ _ _ _ t u i j k l) = + gB _ _ _ _ _ _ (λ m n o → rec gB f (t m n o)) (λ m n o → rec gB f (u m n o)) + i j k l + +elim : {B : ∥ A ∥₄ → Type ℓ} + (bG : (x : ∥ A ∥₄) → is2Groupoid (B x)) + (f : (x : A) → B ∣ x ∣₄) (x : ∥ A ∥₄) → B x +elim bG f ∣ x ∣₄ = f x +elim bG f (squash₄ x y p q r s u v i j k l) = + isOfHLevel→isOfHLevelDep 4 bG _ _ _ _ _ _ + (λ j k l → elim bG f (u j k l)) (λ j k l → elim bG f (v j k l)) + (squash₄ x y p q r s u v) i j k l -g2TruncFib : ∀ {ℓ ℓ'} {A : Type ℓ} (P : A → Type ℓ') - {a b : A} (sPb : is2Groupoid (P b)) - {p q : a ≡ b} {r s : p ≡ q} {u v : r ≡ s} (w : u ≡ v) {a1 : P a} {b1 : P b} - {p1 : PathP (λ i → P (p i)) a1 b1} - {q1 : PathP (λ i → P (q i)) a1 b1} - {r1 : PathP (λ i → PathP (λ j → P (r i j)) a1 b1) p1 q1} - {s1 : PathP (λ i → PathP (λ j → P (s i j)) a1 b1) p1 q1} - (u1 : PathP (λ i → PathP (λ j → PathP (λ k → P (u i j k)) a1 b1) p1 q1) r1 s1) - (v1 : PathP (λ i → PathP (λ j → PathP (λ k → P (v i j k)) a1 b1) p1 q1) r1 s1) - → PathP (λ i → PathP (λ j → PathP (λ k → PathP (λ l → P (w i j k l)) a1 b1) p1 q1) r1 s1) u1 v1 -g2TruncFib {A} P {a} {b} sPb {p} {q} {r} {s} {u} {v} w - {a1} {b1} {p1} {q1} {r1} {s1} u1 v1 i j k l - = hcomp (λ m → λ { (i = i0) → u1 j k l - ; (i = i1) → v1 j k l - ; (j = i0) → r1 k l - ; (j = i1) → s1 k l - ; (k = i0) → p1 l - ; (k = i1) → q1 l - ; (l = i0) → a1 - ; (l = i1) → sPb b1 b1 refl refl refl refl L refl m i j k - }) - (Lb i j k l) +elim2 : {B : ∥ A ∥₄ → ∥ A ∥₄ → Type ℓ} + (gB : ((x y : ∥ A ∥₄) → is2Groupoid (B x y))) + (g : (a b : A) → B ∣ a ∣₄ ∣ b ∣₄) + (x y : ∥ A ∥₄) → B x y +elim2 gB g = elim (λ _ → is2GroupoidΠ (λ _ → gB _ _)) + (λ a → elim (λ _ → gB _ _) (g a)) + +elim3 : {B : (x y z : ∥ A ∥₄) → Type ℓ} + (gB : ((x y z : ∥ A ∥₄) → is2Groupoid (B x y z))) + (g : (a b c : A) → B ∣ a ∣₄ ∣ b ∣₄ ∣ c ∣₄) + (x y z : ∥ A ∥₄) → B x y z +elim3 gB g = elim2 (λ _ _ → is2GroupoidΠ (λ _ → gB _ _ _)) + (λ a b → elim (λ _ → gB _ _ _) (g a b)) + +is2Groupoid2GroupoidTrunc : is2Groupoid ∥ A ∥₄ +is2Groupoid2GroupoidTrunc a b p q r s = squash₄ a b p q r s + +2GroupoidTruncIdempotent≃ : is2Groupoid A → ∥ A ∥₄ ≃ A +2GroupoidTruncIdempotent≃ {A = A} hA = isoToEquiv f where - L : Path (Path (b1 ≡ b1) refl refl) refl refl - L i j k = comp (λ l → P (w i j k l)) - (λ l → λ { (i = i0) → u1 j k l - ; (i = i1) → v1 j k l - ; (j = i0) → r1 k l - ; (j = i1) → s1 k l - ; (k = i0) → p1 l - ; (k = i1) → q1 l - }) - a1 - Lb : PathP (λ i → PathP (λ j → PathP (λ k → PathP (λ l → P (w i j k l)) a1 (L i j k)) p1 q1) r1 s1) u1 v1 - Lb i j k l = fill (λ l → P (w i j k l)) - (λ l → λ { (i = i0) → u1 j k l - ; (i = i1) → v1 j k l - ; (j = i0) → r1 k l - ; (j = i1) → s1 k l - ; (k = i0) → p1 l - ; (k = i1) → q1 l - }) - (inS a1) l + f : Iso ∥ A ∥₄ A + Iso.fun f = rec hA (idfun A) + Iso.inv f x = ∣ x ∣₄ + Iso.rightInv f _ = refl + Iso.leftInv f = elim (λ _ → isOfHLevelSuc 4 is2Groupoid2GroupoidTrunc _ _) (λ _ → refl) -g2TruncElim : ∀ {ℓ ℓ'} (A : Type ℓ) (B : ∥ A ∥₂ → Type ℓ') - (bG : (x : ∥ A ∥₂) → is2Groupoid (B x)) - (f : (x : A) → B ∣ x ∣₂) (x : ∥ A ∥₂) → B x -g2TruncElim A B bG f ∣ x ∣₂ = f x -g2TruncElim A B bG f (squash₂ x y p q r s u v i j k l) = - g2TruncFib {A = ∥ A ∥₂} B {x} {y} (bG y) {p} {q} {r} {s} {u} {v} - (squash₂ x y p q r s u v) - {g2TruncElim A B bG f x} - {g2TruncElim A B bG f y} - {λ i → g2TruncElim A B bG f (p i)} - {λ i → g2TruncElim A B bG f (q i)} - {λ i j → g2TruncElim A B bG f (r i j)} - {λ i j → g2TruncElim A B bG f (s i j)} - (λ i j k → g2TruncElim A B bG f (u i j k)) - (λ i j k → g2TruncElim A B bG f (v i j k)) - i j k l +2GroupoidTruncIdempotent : is2Groupoid A → ∥ A ∥₄ ≡ A +2GroupoidTruncIdempotent hA = ua (2GroupoidTruncIdempotent≃ hA) diff --git a/Cubical/HITs/AssocList.agda b/Cubical/HITs/AssocList.agda new file mode 100644 index 000000000..746a04c1f --- /dev/null +++ b/Cubical/HITs/AssocList.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.AssocList where + + +open import Cubical.HITs.AssocList.Base public +open import Cubical.HITs.AssocList.Properties public diff --git a/Cubical/HITs/AssocList/Base.agda b/Cubical/HITs/AssocList/Base.agda new file mode 100644 index 000000000..26f82b040 --- /dev/null +++ b/Cubical/HITs/AssocList/Base.agda @@ -0,0 +1,69 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.AssocList.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Nat using (ℕ; _+_) + +private + variable + ℓ : Level + A : Type ℓ + +infixr 5 ⟨_,_⟩∷_ + +data AssocList (A : Type ℓ) : Type ℓ where + ⟨⟩ : AssocList A + ⟨_,_⟩∷_ : (a : A) (n : ℕ) (xs : AssocList A) → AssocList A + per : ∀ a b xs → ⟨ a , 1 ⟩∷ ⟨ b , 1 ⟩∷ xs + ≡ ⟨ b , 1 ⟩∷ ⟨ a , 1 ⟩∷ xs + agg : ∀ a m n xs → ⟨ a , m ⟩∷ ⟨ a , n ⟩∷ xs + ≡ ⟨ a , m + n ⟩∷ xs + del : ∀ a xs → ⟨ a , 0 ⟩∷ xs ≡ xs + trunc : isSet (AssocList A) + +pattern ⟨_⟩ a = ⟨ a , 1 ⟩∷ ⟨⟩ + + +-- Elimination and recursion principle for association lists +module Elim {ℓ'} {B : AssocList A → Type ℓ'} + (⟨⟩* : B ⟨⟩) (⟨_,_⟩∷*_ : (x : A) (n : ℕ) {xs : AssocList A} → B xs → B (⟨ x , n ⟩∷ xs)) + (per* : (x y : A) {xs : AssocList A} (b : B xs) + → PathP (λ i → B (per x y xs i)) (⟨ x , 1 ⟩∷* ⟨ y , 1 ⟩∷* b) (⟨ y , 1 ⟩∷* ⟨ x , 1 ⟩∷* b)) + (agg* : (x : A) (m n : ℕ) {xs : AssocList A} (b : B xs) + → PathP (λ i → B (agg x m n xs i)) (⟨ x , m ⟩∷* ⟨ x , n ⟩∷* b) (⟨ x , m + n ⟩∷* b)) + (del* : (x : A) {xs : AssocList A} (b : B xs) + → PathP (λ i → B (del x xs i)) (⟨ x , 0 ⟩∷* b) b) + (trunc* : (xs : AssocList A) → isSet (B xs)) where + + f : (xs : AssocList A) → B xs + f ⟨⟩ = ⟨⟩* + f (⟨ a , n ⟩∷ xs) = ⟨ a , n ⟩∷* f xs + f (per a b xs i) = per* a b (f xs) i + f (agg a m n xs i) = agg* a m n (f xs) i + f (del a xs i) = del* a (f xs) i + f (trunc xs ys p q i j) = isOfHLevel→isOfHLevelDep 2 trunc* (f xs) (f ys) (cong f p) (cong f q) (trunc xs ys p q) i j + + + +module ElimProp {ℓ'} {B : AssocList A → Type ℓ'} (BProp : {xs : AssocList A} → isProp (B xs)) + (⟨⟩* : B ⟨⟩) (⟨_,_⟩∷*_ : (x : A) (n : ℕ) {xs : AssocList A} → B xs → B (⟨ x , n ⟩∷ xs)) where + + f : (xs : AssocList A) → B xs + f = Elim.f ⟨⟩* ⟨_,_⟩∷*_ + (λ x y {xs} b → toPathP (BProp (transp (λ i → B (per x y xs i)) i0 (⟨ x , 1 ⟩∷* ⟨ y , 1 ⟩∷* b)) (⟨ y , 1 ⟩∷* ⟨ x , 1 ⟩∷* b))) + (λ x m n {xs} b → toPathP (BProp (transp (λ i → B (agg x m n xs i)) i0 (⟨ x , m ⟩∷* ⟨ x , n ⟩∷* b)) (⟨ x , m + n ⟩∷* b))) + (λ x {xs} b → toPathP (BProp (transp (λ i → B (del x xs i)) i0 (⟨ x , 0 ⟩∷* b)) b)) + (λ xs → isProp→isSet BProp) + + + +module Rec {ℓ'} {B : Type ℓ'} (BType : isSet B) + (⟨⟩* : B) (⟨_,_⟩∷*_ : (x : A) (n : ℕ) → B → B) + (per* : (x y : A) (b : B) → (⟨ x , 1 ⟩∷* ⟨ y , 1 ⟩∷* b) ≡ (⟨ y , 1 ⟩∷* ⟨ x , 1 ⟩∷* b)) + (agg* : (x : A) (m n : ℕ) (b : B) → (⟨ x , m ⟩∷* ⟨ x , n ⟩∷* b) ≡ (⟨ x , m + n ⟩∷* b)) + (del* : (x : A) (b : B) → (⟨ x , 0 ⟩∷* b) ≡ b) where + + f : AssocList A → B + f = Elim.f ⟨⟩* (λ x n b → ⟨ x , n ⟩∷* b) (λ x y b → per* x y b) (λ x m n b → agg* x m n b) (λ x b → del* x b) (λ _ → BType) diff --git a/Cubical/HITs/AssocList/Properties.agda b/Cubical/HITs/AssocList/Properties.agda new file mode 100644 index 000000000..213396161 --- /dev/null +++ b/Cubical/HITs/AssocList/Properties.agda @@ -0,0 +1,162 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.AssocList.Properties where + +open import Cubical.HITs.AssocList.Base as AL +open import Cubical.Foundations.Everything +open import Cubical.Foundations.SIP +open import Cubical.HITs.FiniteMultiset as FMS +open import Cubical.Data.Nat using (ℕ; zero; suc; _+_; +-assoc; isSetℕ) +open import Cubical.Structures.MultiSet +open import Cubical.Relation.Nullary +open import Cubical.Relation.Nullary.DecidableEq + +private + variable + ℓ : Level + A : Type ℓ + + + +multiPer : (a b : A) (m n : ℕ) (xs : AssocList A) + → ⟨ a , m ⟩∷ ⟨ b , n ⟩∷ xs ≡ ⟨ b , n ⟩∷ ⟨ a , m ⟩∷ xs +multiPer a b zero n xs = del a (⟨ b , n ⟩∷ xs) ∙ cong (λ ys → ⟨ b , n ⟩∷ ys) (sym (del a xs)) +multiPer a b (suc m) zero xs = cong (λ ys → ⟨ a , suc m ⟩∷ ys) (del b xs) ∙ sym (del b (⟨ a , suc m ⟩∷ xs)) +multiPer a b (suc m) (suc n) xs = + + ⟨ a , suc m ⟩∷ ⟨ b , suc n ⟩∷ xs ≡⟨ sym (agg a 1 m (⟨ b , suc n ⟩∷ xs)) ⟩ + ⟨ a , 1 ⟩∷ ⟨ a , m ⟩∷ ⟨ b , suc n ⟩∷ xs ≡⟨ cong (λ ys → ⟨ a , 1 ⟩∷ ys) (multiPer a b m (suc n) xs) ⟩ + ⟨ a , 1 ⟩∷ ⟨ b , suc n ⟩∷ ⟨ a , m ⟩∷ xs ≡⟨ cong (λ ys → ⟨ a , 1 ⟩∷ ys) (sym (agg b 1 n (⟨ a , m ⟩∷ xs))) ⟩ + ⟨ a , 1 ⟩∷ ⟨ b , 1 ⟩∷ ⟨ b , n ⟩∷ ⟨ a , m ⟩∷ xs ≡⟨ per a b (⟨ b , n ⟩∷ ⟨ a , m ⟩∷ xs) ⟩ + ⟨ b , 1 ⟩∷ ⟨ a , 1 ⟩∷ ⟨ b , n ⟩∷ ⟨ a , m ⟩∷ xs ≡⟨ cong (λ ys → ⟨ b , 1 ⟩∷ ⟨ a , 1 ⟩∷ ys) (multiPer b a n m xs) ⟩ + ⟨ b , 1 ⟩∷ ⟨ a , 1 ⟩∷ ⟨ a , m ⟩∷ ⟨ b , n ⟩∷ xs ≡⟨ cong (λ ys → ⟨ b , 1 ⟩∷ ys) (agg a 1 m (⟨ b , n ⟩∷ xs)) ⟩ + ⟨ b , 1 ⟩∷ ⟨ a , suc m ⟩∷ ⟨ b , n ⟩∷ xs ≡⟨ cong (λ ys → ⟨ b , 1 ⟩∷ ys) (multiPer a b (suc m) n xs) ⟩ + ⟨ b , 1 ⟩∷ ⟨ b , n ⟩∷ ⟨ a , suc m ⟩∷ xs ≡⟨ agg b 1 n (⟨ a , suc m ⟩∷ xs) ⟩ + ⟨ b , suc n ⟩∷ ⟨ a , suc m ⟩∷ xs ∎ + + + + + +-- Show that association lists and finite multisets are equivalent +multi-∷ : A → ℕ → FMSet A → FMSet A +multi-∷ x zero xs = xs +multi-∷ x (suc n) xs = x ∷ multi-∷ x n xs + +multi-∷-agg : (x : A) (m n : ℕ) (b : FMSet A) → multi-∷ x m (multi-∷ x n b) ≡ multi-∷ x (m + n) b +multi-∷-agg x zero n b = refl +multi-∷-agg x (suc m) n b i = x ∷ (multi-∷-agg x m n b i) + +AL→FMS : AssocList A → FMSet A +AL→FMS = AL.Rec.f FMS.trunc [] multi-∷ comm multi-∷-agg λ _ _ → refl + + +FMS→AL : FMSet A → AssocList A +FMS→AL = FMS.Rec.f AL.trunc ⟨⟩ (λ x xs → ⟨ x , 1 ⟩∷ xs) per + + + +AL→FMS∘FMS→AL≡id : section {A = AssocList A} AL→FMS FMS→AL +AL→FMS∘FMS→AL≡id = FMS.ElimProp.f (FMS.trunc _ _) refl (λ x p → cong (λ ys → x ∷ ys) p) + + +-- need a little lemma for other direction +multi-∷-id : (x : A) (n : ℕ) (u : FMSet A) + → FMS→AL (multi-∷ x n u) ≡ ⟨ x , n ⟩∷ FMS→AL u +multi-∷-id x zero u = sym (del x (FMS→AL u)) +multi-∷-id x (suc n) u = FMS→AL (multi-∷ x (suc n) u) ≡⟨ cong (λ ys → ⟨ x , 1 ⟩∷ ys) (multi-∷-id x n u) ⟩ + ⟨ x , 1 ⟩∷ ⟨ x , n ⟩∷ (FMS→AL u) ≡⟨ agg x 1 n (FMS→AL u) ⟩ + ⟨ x , (suc n) ⟩∷ (FMS→AL u) ∎ + +FMS→AL∘AL→FMS≡id : retract {A = AssocList A} AL→FMS FMS→AL +FMS→AL∘AL→FMS≡id = AL.ElimProp.f (AL.trunc _ _) refl (λ x n {xs} p → (multi-∷-id x n (AL→FMS xs)) ∙ cong (λ ys → ⟨ x , n ⟩∷ ys) p) + + + + +AssocList≃FMSet : AssocList A ≃ FMSet A +AssocList≃FMSet = isoToEquiv (iso AL→FMS FMS→AL AL→FMS∘FMS→AL≡id FMS→AL∘AL→FMS≡id) + +FMSet≃AssocList : FMSet A ≃ AssocList A +FMSet≃AssocList = isoToEquiv (iso FMS→AL AL→FMS FMS→AL∘AL→FMS≡id AL→FMS∘FMS→AL≡id) + + +AssocList≡FMSet : AssocList A ≡ FMSet A +AssocList≡FMSet = ua AssocList≃FMSet + + + + +-- We want to define a multiset structure on AssocList A, we use the recursor to define the count-function + + +module _(discA : Discrete A) where + setA = Discrete→isSet discA + + + + ALcount-⟨,⟩∷* : A → A → ℕ → ℕ → ℕ + ALcount-⟨,⟩∷* a x n xs with discA a x + ... | yes a≡x = n + xs + ... | no a≢x = xs + + + ALcount-per* : (a x y : A) (xs : ℕ) + → ALcount-⟨,⟩∷* a x 1 (ALcount-⟨,⟩∷* a y 1 xs) + ≡ ALcount-⟨,⟩∷* a y 1 (ALcount-⟨,⟩∷* a x 1 xs) + ALcount-per* a x y xs with discA a x | discA a y + ALcount-per* a x y xs | yes a≡x | yes a≡y = refl + ALcount-per* a x y xs | yes a≡x | no a≢y = refl + ALcount-per* a x y xs | no a≢x | yes a≡y = refl + ALcount-per* a x y xs | no a≢x | no a≢y = refl + + + ALcount-agg* : (a x : A) (m n xs : ℕ) + → ALcount-⟨,⟩∷* a x m (ALcount-⟨,⟩∷* a x n xs) + ≡ ALcount-⟨,⟩∷* a x (m + n) xs + ALcount-agg* a x m n xs with discA a x + ... | yes _ = +-assoc m n xs + ... | no _ = refl + + ALcount-del* : (a x : A) (xs : ℕ) → ALcount-⟨,⟩∷* a x 0 xs ≡ xs + ALcount-del* a x xs with discA a x + ... | yes _ = refl + ... | no _ = refl + + + ALcount : A → AssocList A → ℕ + ALcount a = AL.Rec.f isSetℕ 0 (ALcount-⟨,⟩∷* a) (ALcount-per* a) (ALcount-agg* a) (ALcount-del* a) + + + + AL-with-str : MultiSet A setA + AL-with-str = (AssocList A , ⟨⟩ , ⟨_, 1 ⟩∷_ , ALcount) + + +-- We want to show that Al-with-str ≅ FMS-with-str as multiset-structures + FMS→AL-EquivStr : MultiSetEquivStr A setA (FMS-with-str discA) (AL-with-str) FMSet≃AssocList + FMS→AL-EquivStr = refl , (λ a xs → refl) , φ + where + φ : ∀ a xs → FMScount discA a xs ≡ ALcount a (FMS→AL xs) + φ a = FMS.ElimProp.f (isSetℕ _ _) refl ψ + where + ψ : (x : A) {xs : FMSet A} + → FMScount discA a xs ≡ ALcount a (FMS→AL xs) + → FMScount discA a (x ∷ xs) ≡ ALcount a (FMS→AL (x ∷ xs)) + ψ x {xs} p = subst B α θ + where + B = λ ys → FMScount discA a (x ∷ xs) ≡ ALcount a ys + + α : ⟨ x , 1 ⟩∷ FMS→AL xs ≡ FMS→AL (x ∷ xs) + α = sym (multi-∷-id x 1 xs) + + θ : FMScount discA a (x ∷ xs) ≡ ALcount a (⟨ x , 1 ⟩∷ (FMS→AL xs)) + θ with discA a x + ... | yes _ = cong suc p + ... | no ¬p = p + + + + FMS-with-str≡AL-with-str : FMS-with-str discA ≡ AL-with-str + FMS-with-str≡AL-with-str = sip (multiSetUnivalentStr A setA) + (FMS-with-str discA) AL-with-str + (FMSet≃AssocList , FMS→AL-EquivStr) diff --git a/Cubical/HITs/Colimit.agda b/Cubical/HITs/Colimit.agda index cf0c75bb8..18c956c57 100644 --- a/Cubical/HITs/Colimit.agda +++ b/Cubical/HITs/Colimit.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Colimit where open import Cubical.HITs.Colimit.Base public diff --git a/Cubical/HITs/Colimit/Base.agda b/Cubical/HITs/Colimit/Base.agda index c66229d1b..40210d185 100644 --- a/Cubical/HITs/Colimit/Base.agda +++ b/Cubical/HITs/Colimit/Base.agda @@ -3,7 +3,7 @@ Homotopy colimits of graphs -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Colimit.Base where open import Cubical.Foundations.Prelude @@ -14,7 +14,6 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Data.Graph - -- Cones under a diagram record Cocone ℓ {ℓd ℓv ℓe} {I : Graph ℓv ℓe} (F : Diag ℓd I) (X : Type ℓ) @@ -58,11 +57,11 @@ module _ {ℓ ℓd ℓv ℓe} {I : Graph ℓv ℓe} {F : Diag ℓd I} {X : Type isUniversalAt : ∀ ℓq → Cocone ℓ F X → Type (ℓ-max ℓ (ℓ-suc (ℓ-max ℓq (ℓ-max (ℓ-max ℓv ℓe) (ℓ-suc ℓd))))) isUniversalAt ℓq C = ∀ (Y : Type ℓq) → isEquiv {A = (X → Y)} {B = Cocone ℓq F Y} (postcomp C) -- (unfolding isEquiv, this ^ is equivalent to what one might expect:) - -- ∀ (Y : Type ℓ) (D : Cocone ℓ F Y) → isContr (Σ[ h ∈ (X → Y) ] (h *) C ≡ D) + -- ∀ (Y : Type ℓ) (D : Cocone ℓ F Y) → ∃![ h ∈ (X → Y) ] (h *) C ≡ D -- (≡ isContr (CoconeMor (X , C) (Y , D))) isPropIsUniversalAt : ∀ ℓq (C : Cocone ℓ F X) → isProp (isUniversalAt ℓq C) - isPropIsUniversalAt ℓq C = isPropPi (λ Y → isPropIsEquiv (postcomp C)) + isPropIsUniversalAt ℓq C = isPropΠ (λ Y → isPropIsEquiv (postcomp C)) isUniversal : Cocone ℓ F X → Typeω isUniversal C = ∀ ℓq → isUniversalAt ℓq C @@ -82,7 +81,7 @@ module _ {ℓ ℓ' ℓd ℓv ℓe} {I : Graph ℓv ℓe} {F : Diag ℓd I} {X : postcomp⁻¹ cl = invEq (_ , univ cl _ Y) postcomp⁻¹-inv : (cl : isColimit F X) (D : Cocone ℓ' F Y) → (postcomp (cone cl) (postcomp⁻¹ cl D)) ≡ D - postcomp⁻¹-inv cl D = retEq (_ , univ cl _ Y) D + postcomp⁻¹-inv cl D = secEq (_ , univ cl _ Y) D postcomp⁻¹-mor : (cl : isColimit F X) (D : Cocone ℓ' F Y) → CoconeMor (X , cone cl) (Y , D) postcomp⁻¹-mor cl D = (postcomp⁻¹ cl D) , (postcomp⁻¹-inv cl D) @@ -117,18 +116,17 @@ module _ {ℓd ℓv ℓe} {I : Graph ℓv ℓe} {F : Diag ℓd I} where leg colimCone = colim-leg com colimCone = colim-com - colim-rec : ∀ {ℓ} {X : Type ℓ} → Cocone ℓ F X → (colim F → X) - colim-rec C (colim-leg j A) = leg C j A - colim-rec C (colim-com f i A) = com C f i A + rec : ∀ {ℓ} {X : Type ℓ} → Cocone ℓ F X → (colim F → X) + rec C (colim-leg j A) = leg C j A + rec C (colim-com f i A) = com C f i A colimIsColimit : isColimit F (colim F) cone colimIsColimit = colimCone univ colimIsColimit ℓq Y = isoToIsEquiv record { fun = postcomp colimCone - ; inv = colim-rec + ; inv = rec ; rightInv = λ C → refl ; leftInv = λ h → funExt (eq h) } - where eq : ∀ h (x : colim _) → colim-rec (postcomp colimCone h) x ≡ h x + where eq : ∀ h (x : colim _) → rec (postcomp colimCone h) x ≡ h x eq h (colim-leg j A) = refl eq h (colim-com f i A) = refl - diff --git a/Cubical/HITs/Colimit/Examples.agda b/Cubical/HITs/Colimit/Examples.agda index 95963f5d9..e2e926d2f 100644 --- a/Cubical/HITs/Colimit/Examples.agda +++ b/Cubical/HITs/Colimit/Examples.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Colimit.Examples where open import Cubical.Foundations.Prelude diff --git a/Cubical/HITs/Cost.agda b/Cubical/HITs/Cost.agda new file mode 100644 index 000000000..6e9e6dc96 --- /dev/null +++ b/Cubical/HITs/Cost.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Cost where + +open import Cubical.HITs.Cost.Base diff --git a/Cubical/HITs/Cost/Base.agda b/Cubical/HITs/Cost/Base.agda new file mode 100644 index 000000000..483d1bc7f --- /dev/null +++ b/Cubical/HITs/Cost/Base.agda @@ -0,0 +1,124 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Cost.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.HITs.PropositionalTruncation +open import Cubical.Data.Nat +open import Cubical.Data.Sigma + +private + variable + ℓ : Level + A B C : Type ℓ + +Cost : (A : Type ℓ) → Type ℓ +Cost A = A × ∥ ℕ ∥ + +-- To compare two elements of Cost A we only need to look at the first parts +Cost≡ : (x y : Cost A) → x .fst ≡ y .fst → x ≡ y +Cost≡ _ _ = Σ≡Prop λ _ → squash + +-- To make it easier to program with Cost A we prove that it forms a +-- monad which counts the number of calls to >>=. We could also turn +-- it into a proper state monad which tracks the cost, but for the +-- programs we write later this simple version suffices +_>>=_ : Cost A → (A → Cost B) → Cost B +(x , m) >>= g with g x +... | (y , n) = (y , map suc (map2 _+_ m n)) + +return : A → Cost A +return x = (x , ∣ 0 ∣) + +-- The monad laws all hold by Cost≡ + +>>=-return : (f : Cost A) → f >>= return ≡ f +>>=-return f = Cost≡ _ _ refl + +return->>= : (a : A) (f : A → Cost B) → return a >>= f ≡ f a +return->>= a f = Cost≡ _ _ refl + +>>=-assoc : (f : Cost A) (g : A → Cost B) (h : B → Cost C) + → (f >>= g) >>= h ≡ f >>= (λ x → g x >>= h) +>>=-assoc f g h = Cost≡ _ _ refl + + +-- An inefficient version of addition which recurses on both arguments +addBad : ℕ → ℕ → Cost ℕ +addBad 0 0 = return 0 +addBad 0 (suc y) = do + x ← addBad 0 y + return (suc x) +addBad (suc x) y = do + z ← addBad x y + return (suc z) + +-- More efficient addition which only recurse on the first argument +add : ℕ → ℕ → Cost ℕ +add 0 y = return y +add (suc x) y = do + z ← add x y + return (suc z) + +private + -- addBad x y will do x + y calls + _ : addBad 3 5 ≡ (8 , ∣ 8 ∣) + _ = refl + + -- add x y will only do x recursive calls + _ : add 3 5 ≡ (8 , ∣ 3 ∣) + _ = refl + +-- Despite addBad and add having different cost we can still prove +-- that they are equal functions +add≡addBad : add ≡ addBad +add≡addBad i x y = Cost≡ (add x y) (addBad x y) (rem x y) i + where + rem : (x y : ℕ) → add x y .fst ≡ addBad x y .fst + rem 0 0 = refl + rem 0 (suc y) = cong suc (rem 0 y) + rem (suc x) y = cong suc (rem x y) + + +-- Another example: computing Fibonacci numbers + +fib : ℕ → Cost ℕ +fib 0 = return 0 +fib 1 = return 1 +fib (suc (suc x)) = do + y ← fib (suc x) + z ← fib x + return (y + z) + +fibTail : ℕ → Cost ℕ +fibTail 0 = return 0 +fibTail (suc x) = fibAux x 1 0 + where + fibAux : ℕ → ℕ → ℕ → Cost ℕ + fibAux 0 res _ = return res + fibAux (suc x) res prev = do + r ← fibAux x (res + prev) res + return r + + +private + _ : fib 10 ≡ (55 , ∣ 176 ∣) + _ = refl + + _ : fibTail 10 ≡ (55 , ∣ 9 ∣) + _ = refl + + _ : fib 20 ≡ (6765 , ∣ 21890 ∣) + _ = refl + + _ : fibTail 20 ≡ (6765 , ∣ 19 ∣) + _ = refl + +-- Exercise left to the reader: prove that fib and fibTail are equal functions +-- fib≡fibTail : fib ≡ fibTail +-- fib≡fibTail i x = Cost≡ (fib x) (fibTail x) (rem x) i +-- where +-- rem : (x : ℕ) → fib x .fst ≡ fibTail x .fst +-- rem zero = refl +-- rem (suc zero) = refl +-- rem (suc (suc x)) = {!!} diff --git a/Cubical/HITs/Cylinder.agda b/Cubical/HITs/Cylinder.agda index eb84245c9..36c27f010 100644 --- a/Cubical/HITs/Cylinder.agda +++ b/Cubical/HITs/Cylinder.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Cylinder where open import Cubical.HITs.Cylinder.Base public diff --git a/Cubical/HITs/Cylinder/Base.agda b/Cubical/HITs/Cylinder/Base.agda index 0d1dfc6f0..6331d0384 100644 --- a/Cubical/HITs/Cylinder/Base.agda +++ b/Cubical/HITs/Cylinder/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Cylinder.Base where @@ -6,12 +6,12 @@ open import Cubical.Core.Everything open import Cubical.Foundations.Everything -open import Cubical.HITs.PropositionalTruncation +open import Cubical.Data.Sigma +open import Cubical.Data.Unit +open import Cubical.Data.Sum using (_⊎_; inl; inr) -import Cubical.Data.Everything as Data -open Data hiding (inl; inr) - -open import Cubical.HITs.Interval +open import Cubical.HITs.PropositionalTruncation using (∥_∥; ∣_∣; squash) +open import Cubical.HITs.Interval using (Interval; zero; one; seg) -- Cylinder A is a cylinder object in the category of cubical types. -- @@ -35,28 +35,28 @@ module _ {ℓ} {A : Type ℓ} where -- -- include is the first part of the factorization. include : A ⊎ A → Cylinder A - include (Data.inl x) = inl x - include (Data.inr x) = inr x + include (inl x) = inl x + include (inr x) = inr x -- The above inclusion is surjective - includeSurjective : ∀ c → ∥ Σ[ s ∈ A ⊎ A ] include s ≡ c ∥ - includeSurjective (inl x) = ∣ Data.inl x , refl ∣ - includeSurjective (inr x) = ∣ Data.inr x , refl ∣ + includeSurjective : ∀ c → ∃[ s ∈ A ⊎ A ] include s ≡ c + includeSurjective (inl x) = ∣ inl x , refl ∣ + includeSurjective (inr x) = ∣ inr x , refl ∣ includeSurjective (cross x i) = squash - ∣ Data.inl x , (λ j → cross x (i ∧ j)) ∣ - ∣ Data.inr x , (λ j → cross x (i ∨ ~ j)) ∣ + ∣ inl x , (λ j → cross x (i ∧ j)) ∣ + ∣ inr x , (λ j → cross x (i ∨ ~ j)) ∣ i - elimCyl + elim : ∀{ℓ'} {B : Cylinder A → Type ℓ'} → (f : (x : A) → B (inl x)) → (g : (x : A) → B (inr x)) → (p : ∀ x → PathP (λ i → B (cross x i)) (f x) (g x)) → (c : Cylinder A) → B c - elimCyl f _ _ (inl x) = f x - elimCyl _ g _ (inr x) = g x - elimCyl _ _ p (cross x i) = p x i + elim f _ _ (inl x) = f x + elim _ g _ (inr x) = g x + elim _ _ p (cross x i) = p x i private out : Cylinder A → A @@ -117,22 +117,22 @@ module Functorial where B : Type ℓb C : Type ℓc - mapCylinder : (A → B) → Cylinder A → Cylinder B - mapCylinder f (inl x) = inl (f x) - mapCylinder f (inr x) = inr (f x) - mapCylinder f (cross x i) = cross (f x) i + map : (A → B) → Cylinder A → Cylinder B + map f (inl x) = inl (f x) + map f (inr x) = inr (f x) + map f (cross x i) = cross (f x) i - mapCylinderId : mapCylinder (λ(x : A) → x) ≡ (λ x → x) - mapCylinderId i (inl x) = inl x - mapCylinderId i (inr x) = inr x - mapCylinderId i (cross x j) = cross x j + mapId : map (λ(x : A) → x) ≡ (λ x → x) + mapId i (inl x) = inl x + mapId i (inr x) = inr x + mapId i (cross x j) = cross x j - mapCylinder∘ + map∘ : (f : A → B) → (g : B → C) - → mapCylinder (λ x → g (f x)) ≡ (λ x → mapCylinder g (mapCylinder f x)) - mapCylinder∘ f g i (inl x) = inl (g (f x)) - mapCylinder∘ f g i (inr x) = inr (g (f x)) - mapCylinder∘ f g i (cross x j) = cross (g (f x)) j + → map (λ x → g (f x)) ≡ (λ x → map g (map f x)) + map∘ f g i (inl x) = inl (g (f x)) + map∘ f g i (inr x) = inr (g (f x)) + map∘ f g i (cross x j) = cross (g (f x)) j -- There is an adjunction between the cylinder and coyclinder -- functors. @@ -263,4 +263,3 @@ module Push {ℓ} {A : Type ℓ} where Cylinder→Pushout Cylinder→Pushout→Cylinder Pushout→Cylinder→Pushout) - diff --git a/Cubical/HITs/Delooping/Two/Base.agda b/Cubical/HITs/Delooping/Two/Base.agda new file mode 100644 index 000000000..655cadf9c --- /dev/null +++ b/Cubical/HITs/Delooping/Two/Base.agda @@ -0,0 +1,55 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.Delooping.Two.Base where + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude + +-- Explicit construction of the delooping of the two-element group +data Bℤ₂ : Type₀ where + base : Bℤ₂ + loop : base ≡ base + loop² : Square loop refl refl loop + trunc : isGroupoid Bℤ₂ + +private + variable + ℓ : Level + A : Type ℓ + +module Elim where + rec : (x : A) + → (p : x ≡ x) + → (sq : Square p refl refl p) + → isGroupoid A + → Bℤ₂ → A + rec x p sq Agpd = go + where + go : _ → _ + go base = x + go (loop i) = p i + go (loop² i j) = sq i j + go (trunc x y p q r s i j k) + = Agpd + (go x) (go y) + (cong go p) (cong go q) + (cong (cong go) r) (cong (cong go) s) + i j k + + elim : (P : Bℤ₂ → Type ℓ) + → (x : P base) + → (p : PathP (λ i → P (loop i)) x x) + → (sq : SquareP (λ i j → P (loop² i j)) p (λ i → x) (λ i → x) p) + → isOfHLevelDep 3 P + → (z : Bℤ₂) → P z + elim P x p sq Pgpd = go + where + go : (z : Bℤ₂) → P z + go base = x + go (loop i) = p i + go (loop² i j) = sq i j + go (trunc x y p q r s i j k) + = Pgpd (go x) (go y) + (cong go p) (cong go q) + (cong (cong go) r) (cong (cong go) s) + (trunc x y p q r s) i j k diff --git a/Cubical/HITs/Delooping/Two/Properties.agda b/Cubical/HITs/Delooping/Two/Properties.agda new file mode 100644 index 000000000..9a5b7bb96 --- /dev/null +++ b/Cubical/HITs/Delooping/Two/Properties.agda @@ -0,0 +1,198 @@ +{-# OPTIONS --safe #-} + +module Cubical.HITs.Delooping.Two.Properties where + +open import Cubical.Functions.Involution + +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Bool +open import Cubical.Data.Empty +open import Cubical.Data.Unit + +open import Cubical.HITs.Delooping.Two.Base +open import Cubical.HITs.PropositionalTruncation + +private + variable + ℓ : Level + +module Embed where + isPropDepIsSet : isOfHLevelDep {ℓ' = ℓ} 1 isSet + isPropDepIsSet = isOfHLevel→isOfHLevelDep 1 (λ _ → isPropIsSet) + + notSet : PathP (λ i → isSet (notEq i)) isSetBool isSetBool + notSet = isPropDepIsSet isSetBool isSetBool notEq + + notNot² : Square notEq refl refl notEq + notNot² = involPath² notnot + + notNotSet + : SquareP (λ i j → isSet (notNot² i j)) notSet refl refl notSet + notNotSet = isPropDep→isSetDep' + isPropDepIsSet + (involPath² notnot) + notSet refl refl notSet + + Code : Bℤ₂ → hSet ℓ-zero + Code = Elim.rec + (Bool , isSetBool) + (λ i → notEq i , notSet i) + (λ i j → λ where + .fst → notNot² i j + .snd → notNotSet i j) + (isOfHLevelTypeOfHLevel 2) + + El : Bℤ₂ → Type₀ + El b = Code b .fst + +module BINARY where + open import Cubical.Data.FinSet.Binary.Large + + sem : Bℤ₂ → Binary _ + sem = Elim.rec Base Loop Loop² isGroupoidBinary + + loop? : Bool → base ≡ base + loop? false = refl + loop? true = loop + + Loop²-coh : (a b c : Bool) → Type₀ + Loop²-coh false false false = Unit + Loop²-coh false true true = Unit + Loop²-coh true false true = Unit + Loop²-coh true true false = Unit + Loop²-coh _ _ _ = ⊥ + + rf : Bool ≡ Bool → Bool + rf P = transport P false + + Loop²-coh-lemma₀ + : ∀(p q r : Bool) + → r ⊕ p ≡ q + → Loop²-coh p q r + Loop²-coh-lemma₀ false false false sq = _ + Loop²-coh-lemma₀ false true true sq = _ + Loop²-coh-lemma₀ true false true sq = _ + Loop²-coh-lemma₀ true true false sq = _ + Loop²-coh-lemma₀ false true false = false≢true + Loop²-coh-lemma₀ false false true = true≢false + Loop²-coh-lemma₀ true false false = true≢false + Loop²-coh-lemma₀ true true true = false≢true + + Loop²-coh-lemma + : ∀(P Q R : Bool ≡ Bool) + → Square P Q refl R + → Loop²-coh (rf P) (rf Q) (rf R) + Loop²-coh-lemma P Q R sq = Loop²-coh-lemma₀ p q r eqn + where + p = rf P + q = rf Q + r = rf R + + open BoolReflection + + cmp : P ∙ R ≡ Q + cmp i j + = hcomp (λ k → λ where + (i = i0) → compPath-filler P R k j + (i = i1) → Q j + (j = i0) → Bool + (j = i1) → R (i ∨ k)) + (sq i j) + + rcmp : ⊕-Path (r ⊕ p) ≡ ⊕-Path q + rcmp = ⊕-Path (r ⊕ p) + ≡[ i ]⟨ ⊕-comp p r (~ i) ⟩ + ⊕-Path p ∙ ⊕-Path r + ≡[ i ]⟨ ⊕-complete P (~ i) ∙ ⊕-complete R (~ i) ⟩ + P ∙ R + ≡⟨ cmp ⟩ + Q + ≡⟨ ⊕-complete Q ⟩ + ⊕-Path q ∎ + open Iso + eqn : r ⊕ p ≡ q + eqn = transport (λ i → + reflectIso .leftInv (r ⊕ p) i ≡ reflectIso .leftInv q i) + (cong (reflectIso .inv) rcmp) + + loop²? + : ∀ p q r → Loop²-coh p q r + → Square (loop? p) (loop? q) refl (loop? r) + loop²? false false false _ = refl + loop²? false true true _ = λ i j → loop (i ∧ j) + loop²? true false true _ = loop² + loop²? true true false _ = refl + + module _ (B : Type₀) where + based : (P : Bool ≃ B) → Bℤ₂ + based _ = base + + pull₀ : (P Q : Bool ≃ B) → Bool ≡ Bool + pull₀ P Q i + = hcomp (λ k → λ where + (i = i0) → ua P (~ k) + (i = i1) → ua Q (~ k)) + B + + pull₁ : (P Q : Bool ≃ B) → Square (ua P) (ua Q) (pull₀ P Q) refl + pull₁ P Q i j + = hcomp (λ k → λ where + (i = i0) → ua P (~ k ∨ j) + (i = i1) → ua Q (~ k ∨ j) + (j = i1) → B) + B + + pull₂ + : (P Q R : Bool ≃ B) + → Square (pull₀ P Q) (pull₀ P R) refl (pull₀ Q R) + pull₂ P Q R i j + = hcomp (λ k → λ where + (j = i0) → ua P (~ k) + (i = i0) (j = i1) → ua Q (~ k) + (i = i1) (j = i1) → ua R (~ k)) + B + + pull₃ + : (P Q R : Bool ≃ B) + → Cube (pull₁ P Q) (pull₁ P R) + (λ _ → ua P) (pull₁ Q R) + (pull₂ P Q R) (λ _ _ → B) + pull₃ P Q R i j k + = hcomp (λ τ → λ where + (j = i0) → ua P (~ τ ∨ k) + (i = i0) (j = i1) → ua Q (~ τ ∨ k) + (i = i1) (j = i1) → ua R (~ τ ∨ k) + (k = i1) → B) + B + + looped : (P Q : Bool ≃ B) → based P ≡ based Q + looped P Q = loop? b + where + b : Bool + b = rf (pull₀ P Q) + + looped² + : (P Q R : Bool ≃ B) + → Square (looped P Q) (looped P R) refl (looped Q R) + looped² P Q R = loop²? pq pr qr pqr + where + pq = rf (pull₀ P Q) + pr = rf (pull₀ P R) + qr = rf (pull₀ Q R) + + pqr : Loop²-coh pq pr qr + pqr = Loop²-coh-lemma (pull₀ P Q) (pull₀ P R) (pull₀ Q R) (pull₂ P Q R) + + syn : Binary _ → Bℤ₂ + syn (B , tP) = rec→Gpd trunc (based B) 3k tP + where + open 3-Constant + 3k : 3-Constant (based B) + 3k .link = looped B + 3k .coh₁ = looped² B diff --git a/Cubical/HITs/DunceCap.agda b/Cubical/HITs/DunceCap.agda index b9003d468..42c01490a 100644 --- a/Cubical/HITs/DunceCap.agda +++ b/Cubical/HITs/DunceCap.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.DunceCap where diff --git a/Cubical/HITs/DunceCap/Base.agda b/Cubical/HITs/DunceCap/Base.agda index 4c3c53d49..57f87839b 100644 --- a/Cubical/HITs/DunceCap/Base.agda +++ b/Cubical/HITs/DunceCap/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.DunceCap.Base where diff --git a/Cubical/HITs/DunceCap/Properties.agda b/Cubical/HITs/DunceCap/Properties.agda index fbdcd20fa..a655e4d41 100644 --- a/Cubical/HITs/DunceCap/Properties.agda +++ b/Cubical/HITs/DunceCap/Properties.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.DunceCap.Properties where @@ -61,4 +61,4 @@ snd isContr-Dunce = sym ∘ contrDunce Dunce≡DunceCone : Dunce ≡ DunceCone -Dunce≡DunceCone = ua (Contr→Equiv isContr-Dunce isContr-DunceCone) +Dunce≡DunceCone = ua (isContr→Equiv isContr-Dunce isContr-DunceCone) diff --git a/Cubical/HITs/EilenbergMacLane1.agda b/Cubical/HITs/EilenbergMacLane1.agda new file mode 100644 index 000000000..28ea97a27 --- /dev/null +++ b/Cubical/HITs/EilenbergMacLane1.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --cubical --no-import-sorts --safe #-} +module Cubical.HITs.EilenbergMacLane1 where + +open import Cubical.HITs.EilenbergMacLane1.Base public +open import Cubical.HITs.EilenbergMacLane1.Properties public diff --git a/Cubical/HITs/EilenbergMacLane1/Base.agda b/Cubical/HITs/EilenbergMacLane1/Base.agda new file mode 100644 index 000000000..b28a48d27 --- /dev/null +++ b/Cubical/HITs/EilenbergMacLane1/Base.agda @@ -0,0 +1,43 @@ +{- + +This file contains: + +- The first Eilenberg–Mac Lane type as a HIT + +-} +{-# OPTIONS --cubical --no-import-sorts --safe #-} +module Cubical.HITs.EilenbergMacLane1.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Algebra.Group.Base + +private + variable ℓ : Level + +module _ (Group@(G , str) : Group ℓ) where + open GroupStr str + + data EM₁ : Type ℓ where + embase : EM₁ + emloop : G → embase ≡ embase + emcomp : (g h : G) + → PathP (λ j → embase ≡ emloop h j) (emloop g) (emloop (g · h)) + emsquash : ∀ (x y : EM₁) (p q : x ≡ y) (r s : p ≡ q) → r ≡ s + + {- The emcomp constructor fills the square: +^ + emloop (g · h) + [a]— — — >[c] + ‖ ^ + ‖ | emloop h ^ + ‖ | j | + [a]— — — >[b] ∙ — > + emloop g i + + We use this to give another constructor-like construction: + -} + + emloop-comp : (g h : G) → emloop (g · h) ≡ emloop g ∙ emloop h + emloop-comp g h i = compPath-unique refl (emloop g) (emloop h) + (emloop (g · h) , emcomp g h) + (emloop g ∙ emloop h , compPath-filler (emloop g) (emloop h)) i .fst diff --git a/Cubical/HITs/EilenbergMacLane1/Properties.agda b/Cubical/HITs/EilenbergMacLane1/Properties.agda new file mode 100644 index 000000000..7af601e5d --- /dev/null +++ b/Cubical/HITs/EilenbergMacLane1/Properties.agda @@ -0,0 +1,100 @@ +{- + +Eilenberg–Mac Lane type K(G, 1) + +-} + +{-# OPTIONS --cubical --no-import-sorts --safe #-} +module Cubical.HITs.EilenbergMacLane1.Properties where + +open import Cubical.HITs.EilenbergMacLane1.Base + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Sigma + +open import Cubical.Algebra.Group.Base + +open import Cubical.HITs.PropositionalTruncation as PropTrunc using (∥_∥; ∣_∣; squash) +open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂; ∣_∣₂; squash₂) + +private + variable + ℓG ℓ : Level + +module _ ((G , str) : Group ℓG) where + + open GroupStr str + + elimSet : {B : EM₁ (G , str) → Type ℓ} + → ((x : EM₁ (G , str)) → isSet (B x)) + → (b : B embase) + → ((g : G) → PathP (λ i → B (emloop g i)) b b) + → (x : EM₁ (G , str)) + → B x + elimSet Bset b bloop embase = b + elimSet Bset b bloop (emloop g i) = bloop g i + elimSet Bset b bloop (emcomp g h i j) = + isSet→SquareP + (λ i j → Bset (emcomp g h i j)) + (λ j → bloop g j) (λ j → bloop (g · h) j) + (λ i → b) (λ i → bloop h i) + i j + elimSet Bset b bloop (emsquash x y p q r s i j k) = + isOfHLevel→isOfHLevelDep 3 (λ x → isSet→isGroupoid (Bset x)) + _ _ _ _ (λ j k → g (r j k)) (λ j k → g (s j k)) (emsquash x y p q r s) i j k + where + g = elimSet Bset b bloop + + elimProp : {B : EM₁ (G , str) → Type ℓ} + → ((x : EM₁ (G , str)) → isProp (B x)) + → B embase + → (x : EM₁ (G , str)) + → B x + elimProp Bprop b x = + elimSet + (λ x → isProp→isSet (Bprop x)) + b + (λ g → isProp→PathP (λ i → Bprop ((emloop g) i)) b b) + x + + elimProp2 : {C : EM₁ (G , str) → EM₁ (G , str) → Type ℓ} + → ((x y : EM₁ (G , str)) → isProp (C x y)) + → C embase embase + → (x y : EM₁ (G , str)) + → C x y + elimProp2 Cprop c = + elimProp + (λ x → isPropΠ (λ y → Cprop x y)) + (elimProp (λ y → Cprop embase y) c) + + elim : {B : EM₁ (G , str) → Type ℓ} + → ((x : EM₁ (G , str)) → isGroupoid (B x)) + → (b : B embase) + → (bloop : (g : G) → PathP (λ i → B (emloop g i)) b b) + → ((g h : G) → SquareP (λ i j → B (emcomp g h i j)) + (bloop g) (bloop (g · h)) (λ j → b) (bloop h)) + → (x : EM₁ (G , str)) + → B x + elim Bgpd b bloop bcomp embase = b + elim Bgpd b bloop bcomp (emloop g i) = bloop g i + elim Bgpd b bloop bcomp (emcomp g h i j) = bcomp g h i j + elim Bgpd b bloop bcomp (emsquash x y p q r s i j k) = + isOfHLevel→isOfHLevelDep 3 Bgpd + _ _ _ _ (λ j k → g (r j k)) (λ j k → g (s j k)) (emsquash x y p q r s) i j k + where + g = elim Bgpd b bloop bcomp + + rec : {B : Type ℓ} + → isGroupoid B + → (b : B) + → (bloop : G → b ≡ b) + → ((g h : G) → Square (bloop g) (bloop (g · h)) refl (bloop h)) + → (x : EM₁ (G , str)) + → B + rec Bgpd = elim (λ _ → Bgpd) diff --git a/Cubical/HITs/Everything.agda b/Cubical/HITs/Everything.agda index 2f4ea4cab..26ceba567 100644 --- a/Cubical/HITs/Everything.agda +++ b/Cubical/HITs/Everything.agda @@ -1,37 +1,46 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe --guardedness #-} module Cubical.HITs.Everything where -open import Cubical.HITs.Cylinder public -open import Cubical.HITs.Hopf public -open import Cubical.HITs.Interval public -open import Cubical.HITs.Ints.BiInvInt public hiding ( pred ; suc-pred ; pred-suc ) -open import Cubical.HITs.Ints.DeltaInt public hiding ( pred ; succ ; zero ) -open import Cubical.HITs.Ints.HAEquivInt public hiding ( suc-haequiv ) -open import Cubical.HITs.Ints.IsoInt public -open import Cubical.HITs.Ints.QuoInt public -open import Cubical.HITs.Join public -open import Cubical.HITs.ListedFiniteSet public -open import Cubical.HITs.Pushout public -open import Cubical.HITs.Modulo public -open import Cubical.HITs.S1 public -open import Cubical.HITs.S2 public -open import Cubical.HITs.S3 public -open import Cubical.HITs.Rational public -open import Cubical.HITs.Susp public -open import Cubical.HITs.SmashProduct public renaming (comm to Smash-comm) -open import Cubical.HITs.Torus public -open import Cubical.HITs.PropositionalTruncation public -open import Cubical.HITs.SetTruncation public -open import Cubical.HITs.GroupoidTruncation public -open import Cubical.HITs.2GroupoidTruncation public -open import Cubical.HITs.SetQuotients public -open import Cubical.HITs.FiniteMultiset public hiding ( _++_ ; [_] ; assoc-++ ) -open import Cubical.HITs.Sn public -open import Cubical.HITs.Truncation public -open import Cubical.HITs.Colimit -open import Cubical.HITs.MappingCones -open import Cubical.HITs.InfNat public -open import Cubical.HITs.KleinBottle -open import Cubical.HITs.DunceCap -open import Cubical.HITs.Localization public -open import Cubical.HITs.Nullification public +import Cubical.HITs.2GroupoidTruncation +import Cubical.HITs.AssocList +import Cubical.HITs.Colimit +import Cubical.HITs.Cost +import Cubical.HITs.Cylinder +import Cubical.HITs.Delooping.Two.Base +import Cubical.HITs.Delooping.Two.Properties +import Cubical.HITs.DunceCap +import Cubical.HITs.EilenbergMacLane1 +import Cubical.HITs.FiniteMultiset +import Cubical.HITs.FiniteMultiset.CountExtensionality +import Cubical.HITs.GroupoidQuotients +import Cubical.HITs.GroupoidTruncation +import Cubical.HITs.Hopf +import Cubical.HITs.InfNat +import Cubical.HITs.Interval +import Cubical.HITs.Join +import Cubical.HITs.KleinBottle +import Cubical.HITs.ListedFiniteSet +import Cubical.HITs.Localization +import Cubical.HITs.MappingCones +import Cubical.HITs.Modulo +import Cubical.HITs.Nullification +import Cubical.HITs.PropositionalTruncation +import Cubical.HITs.PropositionalTruncation.Monad +import Cubical.HITs.Pushout +import Cubical.HITs.Pushout.Flattening +import Cubical.HITs.RPn +import Cubical.HITs.Rationals.HITQ +import Cubical.HITs.Rationals.QuoQ +import Cubical.HITs.Rationals.SigmaQ +import Cubical.HITs.S1 +import Cubical.HITs.S2 +import Cubical.HITs.S3 +import Cubical.HITs.SetQuotients +import Cubical.HITs.SetTruncation +import Cubical.HITs.SmashProduct +import Cubical.HITs.Sn +import Cubical.HITs.Susp +import Cubical.HITs.Torus +import Cubical.HITs.Truncation +import Cubical.HITs.TypeQuotients +import Cubical.HITs.Wedge diff --git a/Cubical/HITs/FiniteMultiset.agda b/Cubical/HITs/FiniteMultiset.agda index 13c3e24b7..cc015b38f 100644 --- a/Cubical/HITs/FiniteMultiset.agda +++ b/Cubical/HITs/FiniteMultiset.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.FiniteMultiset where diff --git a/Cubical/HITs/FiniteMultiset/Base.agda b/Cubical/HITs/FiniteMultiset/Base.agda index 6f951de7a..8ca45984c 100644 --- a/Cubical/HITs/FiniteMultiset/Base.agda +++ b/Cubical/HITs/FiniteMultiset/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.FiniteMultiset.Base where open import Cubical.Foundations.Prelude @@ -7,11 +7,12 @@ open import Cubical.Foundations.HLevels private variable - A : Type₀ + ℓ : Level + A : Type ℓ infixr 5 _∷_ -data FMSet (A : Type₀) : Type₀ where +data FMSet (A : Type ℓ) : Type ℓ where [] : FMSet A _∷_ : (x : A) → (xs : FMSet A) → FMSet A comm : ∀ x y xs → x ∷ y ∷ xs ≡ y ∷ x ∷ xs @@ -19,7 +20,7 @@ data FMSet (A : Type₀) : Type₀ where pattern [_] x = x ∷ [] -module FMSetElim {ℓ} {B : FMSet A → Type ℓ} +module Elim {ℓ'} {B : FMSet A → Type ℓ'} ([]* : B []) (_∷*_ : (x : A) {xs : FMSet A} → B xs → B (x ∷ xs)) (comm* : (x y : A) {xs : FMSet A} (b : B xs) → PathP (λ i → B (comm x y xs i)) (x ∷* (y ∷* b)) (y ∷* (x ∷* b))) @@ -32,18 +33,18 @@ module FMSetElim {ℓ} {B : FMSet A → Type ℓ} f (trunc xs zs p q i j) = isOfHLevel→isOfHLevelDep 2 trunc* (f xs) (f zs) (cong f p) (cong f q) (trunc xs zs p q) i j -module FMSetElimProp {ℓ} {B : FMSet A → Type ℓ} (BProp : {xs : FMSet A} → isProp (B xs)) +module ElimProp {ℓ'} {B : FMSet A → Type ℓ'} (BProp : {xs : FMSet A} → isProp (B xs)) ([]* : B []) (_∷*_ : (x : A) {xs : FMSet A} → B xs → B (x ∷ xs)) where f : (xs : FMSet A) → B xs - f = FMSetElim.f []* _∷*_ + f = Elim.f []* _∷*_ (λ x y {xs} b → toPathP (BProp (transp (λ i → B (comm x y xs i)) i0 (x ∷* (y ∷* b))) (y ∷* (x ∷* b)))) (λ xs → isProp→isSet BProp) -module FMSetRec {ℓ} {B : Type ℓ} (BType : isSet B) +module Rec {ℓ'} {B : Type ℓ'} (BType : isSet B) ([]* : B) (_∷*_ : A → B → B) (comm* : (x y : A) (b : B) → x ∷* (y ∷* b) ≡ y ∷* (x ∷* b)) where f : FMSet A → B - f = FMSetElim.f []* (λ x b → x ∷* b) (λ x y b → comm* x y b) (λ _ → BType) + f = Elim.f []* (λ x b → x ∷* b) (λ x y b → comm* x y b) (λ _ → BType) diff --git a/Cubical/HITs/FiniteMultiset/CountExtensionality.agda b/Cubical/HITs/FiniteMultiset/CountExtensionality.agda new file mode 100644 index 000000000..79bea106e --- /dev/null +++ b/Cubical/HITs/FiniteMultiset/CountExtensionality.agda @@ -0,0 +1,191 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.FiniteMultiset.CountExtensionality where + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sum +open import Cubical.Relation.Nullary +open import Cubical.HITs.FiniteMultiset.Base +open import Cubical.HITs.FiniteMultiset.Properties as FMS +open import Cubical.Structures.MultiSet +open import Cubical.Relation.Nullary.DecidableEq + + +private + variable + ℓ : Level + + +-- We define a partial order on FMSet A and use it to proof +-- a strong induction principle for finite multi-sets. +-- Finally, we use this stronger elimination principle to show +-- that any two FMSets can be identified, if they have the same count for every a : A +module _{A : Type ℓ} (discA : Discrete A) where + _≼_ : FMSet A → FMSet A → Type ℓ + xs ≼ ys = ∀ a → FMScount discA a xs ≤ FMScount discA a ys + + ≼-refl : ∀ xs → xs ≼ xs + ≼-refl xs a = ≤-refl + + ≼-trans : ∀ xs ys zs → xs ≼ ys → ys ≼ zs → xs ≼ zs + ≼-trans xs ys zs xs≼ys ys≼zs a = ≤-trans (xs≼ys a) (ys≼zs a) + + + ≼[]→≡[] : ∀ xs → xs ≼ [] → xs ≡ [] + ≼[]→≡[] xs xs≼[] = FMScount-0-lemma discA xs λ a → ≤0→≡0 (xs≼[] a) + + + + + ≼-remove1 : ∀ a xs → remove1 discA a xs ≼ xs + ≼-remove1 a xs b with discA a b + ... | yes a≡b = subst (λ n → n ≤ FMScount discA b xs) (sym path) (≤-predℕ) + where + path : FMScount discA b (remove1 discA a xs) ≡ predℕ (FMScount discA b xs) + path = cong (λ c → FMScount discA b (remove1 discA c xs)) a≡b ∙ remove1-predℕ-lemma discA b xs + ... | no a≢b = subst (λ n → n ≤ FMScount discA b xs) + (sym (FMScount-remove1-≢-lemma discA xs λ b≡a → a≢b (sym b≡a))) ≤-refl + + + ≼-remove1-lemma : ∀ x xs ys → ys ≼ (x ∷ xs) → (remove1 discA x ys) ≼ xs + ≼-remove1-lemma x xs ys ys≼x∷xs a with discA a x + ... | yes a≡x = ≤-trans (≤-trans (0 , p₁) (predℕ-≤-predℕ (ys≼x∷xs a))) + (0 , cong predℕ (FMScount-≡-lemma discA xs a≡x)) + where + p₁ : FMScount discA a (remove1 discA x ys) ≡ predℕ (FMScount discA a ys) + p₁ = subst (λ b → FMScount discA a (remove1 discA b ys) ≡ predℕ (FMScount discA a ys)) a≡x (remove1-predℕ-lemma discA a ys) + ... | no a≢x = ≤-trans (≤-trans (0 , FMScount-remove1-≢-lemma discA ys a≢x) (ys≼x∷xs a)) + (0 , FMScount-≢-lemma discA xs a≢x) + + + ≼-Dichotomy : ∀ x xs ys → ys ≼ (x ∷ xs) → (ys ≼ xs) ⊎ (ys ≡ x ∷ (remove1 discA x ys)) + ≼-Dichotomy x xs ys ys≼x∷xs with (FMScount discA x ys) ≟ suc (FMScount discA x xs) + ... | lt suc = ⊥.rec (¬msuc (ys≼x∷xs x)) (0 , FMScount-≡-lemma-refl discA xs) + + + + -- proving a strong elimination principle for finite multisets + module ≼-ElimProp {ℓ'} {B : FMSet A → Type ℓ'} + (BisProp : ∀ {xs} → isProp (B xs)) (b₀ : B []) + (B-≼-hyp : ∀ x xs → (∀ ys → ys ≼ xs → B ys) → B (x ∷ xs)) where + + C : FMSet A → Type (ℓ-max ℓ ℓ') + C xs = ∀ ys → ys ≼ xs → B ys + + g : ∀ xs → C xs + g = ElimProp.f (isPropΠ2 (λ _ _ → BisProp)) c₀ θ + where + c₀ : C [] + c₀ ys ys≼[] = subst B (sym (≼[]→≡[] ys ys≼[])) b₀ + + θ : ∀ x {xs} → C xs → C (x ∷ xs) + θ x {xs} hyp ys ys≼x∷xs with ≼-Dichotomy x xs ys ys≼x∷xs + ... | inl ys≼xs = hyp ys ys≼xs + ... | inr ys≡x∷zs = subst B (sym ys≡x∷zs) (B-≼-hyp x zs χ) + where + zs = remove1 discA x ys + χ : ∀ vs → vs ≼ zs → B vs + χ vs vs≼zs = hyp vs (≼-trans vs zs xs vs≼zs (≼-remove1-lemma x xs ys ys≼x∷xs)) + + f : ∀ xs → B xs + f = C→B g + where + C→B : (∀ xs → C xs) → (∀ xs → B xs) + C→B C-hyp xs = C-hyp xs xs (≼-refl xs) + + + ≼-ElimPropBin : ∀ {ℓ'} {B : FMSet A → FMSet A → Type ℓ'} + → (∀ {xs} {ys} → isProp (B xs ys)) + → (B [] []) + → (∀ x xs ys → (∀ vs ws → vs ≼ xs → ws ≼ ys → B vs ws) → B (x ∷ xs) ys) + → (∀ x xs ys → (∀ vs ws → vs ≼ xs → ws ≼ ys → B vs ws) → B xs (x ∷ ys)) + ------------------------------------------------------------------------------- + → (∀ xs ys → B xs ys) + ≼-ElimPropBin {B = B} BisProp b₀₀ left-hyp right-hyp = ≼-ElimProp.f (isPropΠ (λ _ → BisProp)) θ χ + where + θ : ∀ ys → B [] ys + θ = ≼-ElimProp.f BisProp b₀₀ h₁ + where + h₁ : ∀ x ys → (∀ ws → ws ≼ ys → B [] ws) → B [] (x ∷ ys) + h₁ x ys mini-h = right-hyp x [] ys h₂ + where + h₂ : ∀ vs ws → vs ≼ [] → ws ≼ ys → B vs ws + h₂ vs ws vs≼[] ws≼ys = subst (λ zs → B zs ws) (sym (≼[]→≡[] vs vs≼[])) (mini-h ws ws≼ys) + + χ : ∀ x xs → (∀ zs → zs ≼ xs → (∀ ys → B zs ys)) → ∀ ys → B (x ∷ xs) ys + χ x xs h ys = left-hyp x xs ys λ vs ws vs≼xs _ → h vs vs≼xs ws + + + + ≼-ElimPropBinSym : ∀ {ℓ'} {B : FMSet A → FMSet A → Type ℓ'} + → (∀ {xs} {ys} → isProp (B xs ys)) + → (∀ {xs} {ys} → B xs ys → B ys xs) + → (B [] []) + → (∀ x xs ys → (∀ vs ws → vs ≼ xs → ws ≼ ys → B vs ws) → B (x ∷ xs) ys) + ---------------------------------------------------------------------------- + → (∀ xs ys → B xs ys) + ≼-ElimPropBinSym {B = B} BisProp BisSym b₀₀ left-hyp = ≼-ElimPropBin BisProp b₀₀ left-hyp right-hyp + where + right-hyp : ∀ x xs ys → (∀ vs ws → vs ≼ xs → ws ≼ ys → B vs ws) → B xs (x ∷ ys) + right-hyp x xs ys h₁ = BisSym (left-hyp x ys xs (λ vs ws vs≼ys ws≼xs → BisSym (h₁ ws vs ws≼xs vs≼ys))) + + + -- The main result + module FMScountExt where + B : FMSet A → FMSet A → Type ℓ + B xs ys = (∀ a → FMScount discA a xs ≡ FMScount discA a ys) → xs ≡ ys + + BisProp : ∀ {xs} {ys} → isProp (B xs ys) + BisProp = isPropΠ λ _ → trunc _ _ + + BisSym : ∀ {xs} {ys} → B xs ys → B ys xs + BisSym {xs} {ys} b h = sym (b (λ a → sym (h a))) + + b₀₀ : B [] [] + b₀₀ _ = refl + + left-hyp : ∀ x xs ys → (∀ vs ws → vs ≼ xs → ws ≼ ys → B vs ws) → B (x ∷ xs) ys + left-hyp x xs ys hyp h₁ = (λ i → x ∷ (hyp-path i)) ∙ sym path + where + eq₁ : FMScount discA x ys ≡ suc (FMScount discA x xs) + eq₁ = sym (h₁ x) ∙ FMScount-≡-lemma-refl discA xs + + path : ys ≡ x ∷ (remove1 discA x ys) + path = remove1-suc-lemma discA x (FMScount discA x xs) ys eq₁ + + hyp-path : xs ≡ remove1 discA x ys + hyp-path = hyp xs (remove1 discA x ys) (≼-refl xs) (≼-remove1 x ys) θ + where + θ : ∀ a → FMScount discA a xs ≡ FMScount discA a (remove1 discA x ys) + θ a with discA a x + ... | yes a≡x = subst (λ b → FMScount discA b xs ≡ FMScount discA b (remove1 discA x ys)) (sym a≡x) eq₂ + where + eq₂ : FMScount discA x xs ≡ FMScount discA x (remove1 discA x ys) + eq₂ = FMScount discA x xs ≡⟨ cong predℕ (sym (FMScount-≡-lemma-refl discA xs)) ⟩ + predℕ (FMScount discA x (x ∷ xs)) ≡⟨ cong predℕ (h₁ x) ⟩ + predℕ (FMScount discA x ys) ≡⟨ sym (remove1-predℕ-lemma discA x ys) ⟩ + FMScount discA x (remove1 discA x ys) ∎ + ... | no a≢x = + FMScount discA a xs ≡⟨ sym (FMScount-≢-lemma discA xs a≢x) ⟩ + FMScount discA a (x ∷ xs) ≡⟨ h₁ a ⟩ + FMScount discA a ys ≡⟨ cong (FMScount discA a) path ⟩ + FMScount discA a (x ∷ (remove1 discA x ys)) ≡⟨ FMScount-≢-lemma discA (remove1 discA x ys) a≢x ⟩ + FMScount discA a (remove1 discA x ys) ∎ + + + Thm : ∀ xs ys → (∀ a → FMScount discA a xs ≡ FMScount discA a ys) → xs ≡ ys + Thm = ≼-ElimPropBinSym BisProp BisSym b₀₀ left-hyp diff --git a/Cubical/HITs/FiniteMultiset/Properties.agda b/Cubical/HITs/FiniteMultiset/Properties.agda index 8d206e841..1a4e24028 100644 --- a/Cubical/HITs/FiniteMultiset/Properties.agda +++ b/Cubical/HITs/FiniteMultiset/Properties.agda @@ -1,14 +1,19 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.FiniteMultiset.Properties where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels - +open import Cubical.Data.Nat +open import Cubical.Data.Empty as ⊥ +open import Cubical.Relation.Nullary open import Cubical.HITs.FiniteMultiset.Base +open import Cubical.Structures.MultiSet +open import Cubical.Relation.Nullary.DecidableEq private variable - A : Type₀ + ℓ : Level + A : Type ℓ infixr 30 _++_ @@ -23,22 +28,20 @@ unitl-++ : ∀ (xs : FMSet A) → [] ++ xs ≡ xs unitl-++ xs = refl unitr-++ : ∀ (xs : FMSet A) → xs ++ [] ≡ xs -unitr-++ = FMSetElimProp.f (trunc _ _) - refl - (λ x p → cong (_∷_ x) p) +unitr-++ = ElimProp.f (trunc _ _) refl (λ x p → cong (_∷_ x) p) assoc-++ : ∀ (xs ys zs : FMSet A) → xs ++ (ys ++ zs) ≡ (xs ++ ys) ++ zs -assoc-++ = FMSetElimProp.f (isPropPi (λ _ → isPropPi (λ _ → trunc _ _))) - (λ ys zs → refl) - (λ x p ys zs → cong (_∷_ x) (p ys zs)) +assoc-++ = ElimProp.f (isPropΠ2 (λ _ _ → trunc _ _)) + (λ ys zs → refl) + (λ x p ys zs → cong (_∷_ x) (p ys zs)) cons-++ : ∀ (x : A) (xs : FMSet A) → x ∷ xs ≡ xs ++ [ x ] -cons-++ x = FMSetElimProp.f (trunc _ _) +cons-++ x = ElimProp.f (trunc _ _) refl (λ y {xs} p → comm x y xs ∙ cong (_∷_ y) p) comm-++ : ∀ (xs ys : FMSet A) → xs ++ ys ≡ ys ++ xs -comm-++ = FMSetElimProp.f (isPropPi (λ _ → trunc _ _)) +comm-++ = ElimProp.f (isPropΠ (λ _ → trunc _ _)) (λ ys → sym (unitr-++ ys)) (λ x {xs} p ys → cong (x ∷_) (p ys) ∙ cong (_++ xs) (cons-++ x ys) @@ -51,7 +54,7 @@ module FMSetUniversal {ℓ} {M : Type ℓ} (MSet : isSet M) (f : A → M) where f-extend : FMSet A → M - f-extend = FMSetRec.f MSet e (λ x m → f x ⊗ m) + f-extend = Rec.f MSet e (λ x m → f x ⊗ m) (λ x y m → comm-⊗ (f x) (f y ⊗ m) ∙ sym (assoc-⊗ (f y) m (f x)) ∙ cong (f y ⊗_) (comm-⊗ m (f x))) f-extend-nil : f-extend [] ≡ e @@ -64,7 +67,7 @@ module FMSetUniversal {ℓ} {M : Type ℓ} (MSet : isSet M) f-extend-sing x = comm-⊗ (f x) e ∙ unit-⊗ (f x) f-extend-++ : ∀ xs ys → f-extend (xs ++ ys) ≡ f-extend xs ⊗ f-extend ys - f-extend-++ = FMSetElimProp.f (isPropPi λ _ → MSet _ _) + f-extend-++ = ElimProp.f (isPropΠ λ _ → MSet _ _) (λ ys → sym (unit-⊗ (f-extend ys))) (λ x {xs} p ys → cong (f x ⊗_) (p ys) ∙ assoc-⊗ (f x) (f-extend xs) (f-extend ys)) @@ -72,6 +75,145 @@ module FMSetUniversal {ℓ} {M : Type ℓ} (MSet : isSet M) (h-++ : ∀ xs ys → h (xs ++ ys) ≡ h xs ⊗ h ys) where f-extend-unique : h ≡ f-extend - f-extend-unique = funExt (FMSetElimProp.f (MSet _ _) + f-extend-unique = funExt (ElimProp.f (MSet _ _) h-nil (λ x {xs} p → (h-++ [ x ] xs) ∙ cong (_⊗ h xs) (h-sing x) ∙ cong (f x ⊗_) p)) + + + +-- We want to construct a multiset-structure on FMSet A, the empty set and insertion are given by the constructors, +-- for the count part we use the recursor + + +-- If A has decidable equality we can define the count-function: +module _(discA : Discrete A) where + FMScount-∷* : A → A → ℕ → ℕ + FMScount-∷* a x n with discA a x + ... | yes a≡x = suc n + ... | no a≢x = n + + FMScount-comm* : (a x y : A) (n : ℕ) + → FMScount-∷* a x (FMScount-∷* a y n) + ≡ FMScount-∷* a y (FMScount-∷* a x n) + FMScount-comm* a x y n with discA a x | discA a y + ... | yes a≡x | yes a≡y = refl + ... | yes a≡x | no a≢y = refl + ... | no a≢x | yes a≡y = refl + ... | no a≢x | no a≢y = refl + + FMScount : A → FMSet A → ℕ + FMScount a = Rec.f isSetℕ 0 (FMScount-∷* a) (FMScount-comm* a) + + + FMS-with-str : MultiSet A (Discrete→isSet discA) + FMS-with-str = (FMSet A , [] , _∷_ , FMScount) + + + + -- We prove some useful properties of the FMScount function + FMScount-≡-lemma : ∀ {a} {x} xs → a ≡ x → FMScount a (x ∷ xs) ≡ suc (FMScount a xs) + FMScount-≡-lemma {a} {x} xs a≡x with discA a x + ... | yes _ = refl + ... | no a≢x = ⊥.rec (a≢x a≡x) + + + FMScount-≡-lemma-refl : ∀ {x} xs → FMScount x (x ∷ xs) ≡ suc (FMScount x xs) + FMScount-≡-lemma-refl {x} xs = FMScount-≡-lemma xs refl + + + FMScount-≢-lemma : ∀ {a} {x} xs → ¬ a ≡ x → FMScount a (x ∷ xs) ≡ FMScount a xs + FMScount-≢-lemma {a} {x} xs a≢x with discA a x + ... | yes a≡x = ⊥.rec (a≢x a≡x) + ... | no _ = refl + + + FMScount-0-lemma : ∀ xs → (∀ a → FMScount a xs ≡ 0) → xs ≡ [] + FMScount-0-lemma = ElimProp.f (isPropΠ λ _ → trunc _ _) (λ _ → refl) θ + where + θ : ∀ x {xs} → ((∀ a → FMScount a xs ≡ 0) → xs ≡ []) + → ((∀ a → FMScount a (x ∷ xs) ≡ 0) → (x ∷ xs) ≡ []) + θ x {xs} _ p = ⊥.rec (snotz (sym (FMScount-≡-lemma-refl xs) ∙ p x)) + + + -- we define a function that removes an element a from a finite multiset once + -- by simultaneously defining two lemmas about it + remove1 : A → FMSet A → FMSet A + remove1 a [] = [] + remove1 a (x ∷ xs) with (discA a x) + ... | yes _ = xs + ... | no _ = x ∷ remove1 a xs + remove1 a (comm x y xs i) with discA a x with discA a y + ... | yes a≡x | yes a≡y = ((sym a≡y ∙ a≡x) i) ∷ xs + ... | yes a≡x | no _ = y ∷ (eq i) + where + eq : xs ≡ remove1 a (x ∷ xs) + eq with discA a x + eq | yes _ = refl + eq | no a≢x = ⊥.rec (a≢x a≡x) + remove1 a (comm x y xs i) | no _ | yes a≡y = x ∷ (eq i) + where + eq : remove1 a (y ∷ xs) ≡ xs + eq with discA a y + eq | yes _ = refl + eq | no a≢y = ⊥.rec (a≢y a≡y) + remove1 a (comm x y xs i) | no a≢x | no a≢y = ((λ i → x ∷ (p i)) ∙∙ comm x y (remove1 a xs) ∙∙ (λ i → y ∷ (sym q i))) i + where + p : remove1 a (y ∷ xs) ≡ y ∷ (remove1 a xs) + p with discA a y + p | yes a≡y = ⊥.rec (a≢y a≡y) + p | no _ = refl + q : remove1 a (x ∷ xs) ≡ x ∷ (remove1 a xs) + q with discA a x + q | yes a≡x = ⊥.rec (a≢x a≡x) + q | no _ = refl + remove1 a (trunc xs ys p q i j) = trunc (remove1 a xs) (remove1 a ys) (cong (remove1 a) p) (cong (remove1 a) q) i j + + + remove1-≡-lemma : ∀ {a} {x} xs → a ≡ x → xs ≡ remove1 a (x ∷ xs) + remove1-≡-lemma {a} {x} xs a≡x with discA a x + ... | yes _ = refl + ... | no a≢x = ⊥.rec (a≢x a≡x) + + remove1-≢-lemma : ∀ {a} {x} xs → ¬ a ≡ x → remove1 a (x ∷ xs) ≡ x ∷ remove1 a xs + remove1-≢-lemma {a} {x} xs a≢x with discA a x + ... | yes a≡x = ⊥.rec (a≢x a≡x) + ... | no _ = refl + + + remove1-predℕ-lemma : ∀ a xs → FMScount a (remove1 a xs) ≡ predℕ (FMScount a xs) + remove1-predℕ-lemma a = ElimProp.f (isSetℕ _ _) refl θ + where + θ : ∀ x {xs} → FMScount a (remove1 a xs) ≡ predℕ (FMScount a xs) + → FMScount a (remove1 a (x ∷ xs)) ≡ predℕ (FMScount a (x ∷ xs)) + θ x {xs} p with discA a x + ... | yes _ = refl + ... | no a≢x = FMScount-≢-lemma (remove1 a xs) a≢x ∙ p + + + remove1-zero-lemma : ∀ a xs → FMScount a xs ≡ zero → xs ≡ remove1 a xs + remove1-zero-lemma a = ElimProp.f (isPropΠ λ _ → trunc _ _) (λ _ → refl) θ + where + θ : ∀ x {xs} → (FMScount a xs ≡ zero → xs ≡ remove1 a xs) + → FMScount a (x ∷ xs) ≡ zero → x ∷ xs ≡ remove1 a (x ∷ xs) + θ x {xs} hyp p with discA a x + ... | yes _ = ⊥.rec (snotz p) + ... | no _ = cong (x ∷_) (hyp p) + + + remove1-suc-lemma : ∀ a n xs → FMScount a xs ≡ suc n → xs ≡ a ∷ (remove1 a xs) + remove1-suc-lemma a n = ElimProp.f (isPropΠ λ _ → trunc _ _) (λ p → ⊥.rec (znots p)) θ + where + θ : ∀ x {xs} → (FMScount a xs ≡ suc n → xs ≡ a ∷ (remove1 a xs)) + → FMScount a (x ∷ xs) ≡ suc n → x ∷ xs ≡ a ∷ (remove1 a (x ∷ xs)) + θ x {xs} hyp p with discA a x + ... | yes a≡x = (λ i → (sym a≡x i) ∷ xs) + ... | no a≢x = cong (x ∷_) (hyp p) ∙ comm x a (remove1 a xs) + + + FMScount-remove1-≢-lemma : ∀ {a} {x} xs → ¬ a ≡ x → FMScount a (remove1 x xs) ≡ FMScount a xs + FMScount-remove1-≢-lemma {a} {x} xs a≢x with discreteℕ (FMScount x xs) zero + ... | yes p = cong (FMScount a) (sym (remove1-zero-lemma x xs p)) + ... | no ¬p = sym (FMScount-≢-lemma (remove1 x xs) a≢x) ∙ cong (FMScount a) eq₁ + where + eq₁ : (x ∷ remove1 x xs) ≡ xs + eq₁ = sym (remove1-suc-lemma x (predℕ (FMScount x xs)) xs (suc-predℕ (FMScount x xs) ¬p)) diff --git a/Cubical/HITs/GroupoidQuotients.agda b/Cubical/HITs/GroupoidQuotients.agda new file mode 100644 index 000000000..ef2758b56 --- /dev/null +++ b/Cubical/HITs/GroupoidQuotients.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --cubical --no-import-sorts --safe #-} +module Cubical.HITs.GroupoidQuotients where + +open import Cubical.HITs.GroupoidQuotients.Base public +open import Cubical.HITs.GroupoidQuotients.Properties public diff --git a/Cubical/HITs/GroupoidQuotients/Base.agda b/Cubical/HITs/GroupoidQuotients/Base.agda new file mode 100644 index 000000000..0e141fea7 --- /dev/null +++ b/Cubical/HITs/GroupoidQuotients/Base.agda @@ -0,0 +1,43 @@ +{- + +This file contains: + +- Definition of groupoid quotients + +-} +{-# OPTIONS --cubical --no-import-sorts --safe #-} +module Cubical.HITs.GroupoidQuotients.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Relation.Binary.Base + +-- Groupoid quotients as a higher inductive type: +-- For the definition, only transitivity is needed +data _//_ {ℓ ℓ'} (A : Type ℓ) {R : A → A → Type ℓ'} + (Rt : BinaryRelation.isTrans R) : Type (ℓ-max ℓ ℓ') where + [_] : (a : A) → A // Rt + eq// : {a b : A} → (r : R a b) → [ a ] ≡ [ b ] + comp// : {a b c : A} → (r : R a b) → (s : R b c) + → PathP (λ j → [ a ] ≡ eq// s j) (eq// r) (eq// (Rt a b c r s)) + squash// : ∀ (x y : A // Rt) (p q : x ≡ y) (r s : p ≡ q) → r ≡ s + +{- The comp// constructor fills the square: + + eq// (Rt a b c r s) + [a]— — — >[c] + ‖ ^ + ‖ | eq// s ^ + ‖ | j | + [a]— — — >[b] ∙ — > + eq// r i + + We use this to give another constructor-like construction: +-} + +comp'// : {ℓ ℓ' : Level} (A : Type ℓ) {R : A → A → Type ℓ'} + (Rt : BinaryRelation.isTrans R) + {a b c : A} → (r : R a b) → (s : R b c) + → eq// {Rt = Rt} (Rt a b c r s) ≡ eq// r ∙ eq// s +comp'// A Rt r s i = compPath-unique refl (eq// r) (eq// s) + (eq// {Rt = Rt} (Rt _ _ _ r s) , comp// r s) + (eq// r ∙ eq// s , compPath-filler (eq// r) (eq// s)) i .fst diff --git a/Cubical/HITs/GroupoidQuotients/Properties.agda b/Cubical/HITs/GroupoidQuotients/Properties.agda new file mode 100644 index 000000000..aff8f3313 --- /dev/null +++ b/Cubical/HITs/GroupoidQuotients/Properties.agda @@ -0,0 +1,106 @@ +{- + +Groupoid quotients: + +-} + +{-# OPTIONS --cubical --no-import-sorts --safe #-} +module Cubical.HITs.GroupoidQuotients.Properties where + +open import Cubical.HITs.GroupoidQuotients.Base + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels + +open import Cubical.Functions.Surjection + +open import Cubical.Data.Sigma + +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary.Base + +open import Cubical.HITs.PropositionalTruncation as PropTrunc using (∥_∥; ∣_∣; squash) +open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂; ∣_∣₂; squash₂) + +-- Type quotients + +private + variable + ℓA ℓR ℓ : Level + A : Type ℓA + R : A → A → Type ℓR + +elimSet : (Rt : BinaryRelation.isTrans R) + → {B : A // Rt → Type ℓ} + → ((x : A // Rt) → isSet (B x)) + → (f : (a : A) → B [ a ]) + → ({a b : A} (r : R a b) → PathP (λ i → B (eq// r i)) (f a) (f b)) + → (x : A // Rt) + → B x +elimSet Rt Bset f feq [ a ] = f a +elimSet Rt Bset f feq (eq// r i) = feq r i +elimSet Rt Bset f feq (comp// {a} {b} {c} r s i j) = + isSet→SquareP (λ i j → Bset (comp// r s i j)) + (λ j → feq r j) (λ j → feq (Rt a b c r s) j) + (λ i → f a) (λ i → feq s i) i j +elimSet Rt Bset f feq (squash// x y p q r s i j k) = + isOfHLevel→isOfHLevelDep 3 (λ x → isSet→isGroupoid (Bset x)) + _ _ _ _ (λ j k → g (r j k)) (λ j k → g (s j k)) (squash// x y p q r s) i j k + where + g = elimSet Rt Bset f feq + +elimProp : (Rt : BinaryRelation.isTrans R) + → {B : A // Rt → Type ℓ} + → ((x : A // Rt) → isProp (B x)) + → ((a : A) → B [ a ]) + → (x : A // Rt) + → B x +elimProp Rt Brop f x = + elimSet Rt (λ x → isProp→isSet (Brop x)) f (λ r → isProp→PathP (λ i → Brop (eq// r i)) (f _) (f _)) x + +elimProp2 : (Rt : BinaryRelation.isTrans R) + → {C : A // Rt → A // Rt → Type ℓ} + → ((x y : A // Rt) → isProp (C x y)) + → ((a b : A) → C [ a ] [ b ]) + → (x y : A // Rt) + → C x y +elimProp2 Rt Cprop f = elimProp Rt (λ x → isPropΠ (λ y → Cprop x y)) + (λ x → elimProp Rt (λ y → Cprop [ x ] y) (f x)) + +isSurjective[] : (Rt : BinaryRelation.isTrans R) + → isSurjection (λ a → [ a ]) +isSurjective[] Rt = elimProp Rt (λ x → squash) (λ a → ∣ a , refl ∣) + +elim : (Rt : BinaryRelation.isTrans R) + → {B : A // Rt → Type ℓ} + → ((x : A // Rt) → isGroupoid (B x)) + → (f : (a : A) → B [ a ]) + → (feq : {a b : A} (r : R a b) → PathP (λ i → B (eq// r i)) (f a) (f b)) + → ({a b c : A} (r : R a b) (s : R b c) + → SquareP (λ i j → B (comp// r s i j)) + (feq r) (feq (Rt a b c r s)) (λ j → f a) (feq s)) + → (x : A // Rt) + → B x +elim Rt Bgpd f feq fcomp [ a ] = f a +elim Rt Bgpd f feq fcomp (eq// r i) = feq r i +elim Rt Bgpd f feq fcomp (comp// r s i j) = fcomp r s i j +elim Rt Bgpd f feq fcomp (squash// x y p q r s i j k) = + isOfHLevel→isOfHLevelDep 3 Bgpd + _ _ _ _ (λ j k → g (r j k)) (λ j k → g (s j k)) (squash// x y p q r s) i j k + where + g = elim Rt Bgpd f feq fcomp + +rec : (Rt : BinaryRelation.isTrans R) + → {B : Type ℓ} + → isGroupoid B + → (f : A → B) + → (feq : {a b : A} (r : R a b) → f a ≡ f b) + → ({a b c : A} (r : R a b) (s : R b c) + → Square (feq r) (feq (Rt a b c r s)) refl (feq s)) + → (x : A // Rt) + → B +rec Rt Bgpd = elim Rt (λ _ → Bgpd) diff --git a/Cubical/HITs/GroupoidTruncation.agda b/Cubical/HITs/GroupoidTruncation.agda index ae7442b70..62e8c4db0 100644 --- a/Cubical/HITs/GroupoidTruncation.agda +++ b/Cubical/HITs/GroupoidTruncation.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.GroupoidTruncation where open import Cubical.HITs.GroupoidTruncation.Base public diff --git a/Cubical/HITs/GroupoidTruncation/Base.agda b/Cubical/HITs/GroupoidTruncation/Base.agda index 5d1b2049e..80ec0f201 100644 --- a/Cubical/HITs/GroupoidTruncation/Base.agda +++ b/Cubical/HITs/GroupoidTruncation/Base.agda @@ -5,13 +5,13 @@ This file contains: - Definition of groupoid truncations -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.GroupoidTruncation.Base where open import Cubical.Core.Primitives -- groupoid truncation as a higher inductive type: -data ∥_∥₁ {ℓ} (A : Type ℓ) : Type ℓ where - ∣_∣₁ : A → ∥ A ∥₁ - squash₁ : ∀ (x y : ∥ A ∥₁) (p q : x ≡ y) (r s : p ≡ q) → r ≡ s +data ∥_∥₃ {ℓ} (A : Type ℓ) : Type ℓ where + ∣_∣₃ : A → ∥ A ∥₃ + squash₃ : ∀ (x y : ∥ A ∥₃) (p q : x ≡ y) (r s : p ≡ q) → r ≡ s diff --git a/Cubical/HITs/GroupoidTruncation/Properties.agda b/Cubical/HITs/GroupoidTruncation/Properties.agda index e792c409e..c46f41a38 100644 --- a/Cubical/HITs/GroupoidTruncation/Properties.agda +++ b/Cubical/HITs/GroupoidTruncation/Properties.agda @@ -5,60 +5,63 @@ This file contains: - Properties of groupoid truncations -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.GroupoidTruncation.Properties where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + open import Cubical.HITs.GroupoidTruncation.Base -recGroupoidTrunc : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (gB : isGroupoid B) → (A → B) → (∥ A ∥₁ → B) -recGroupoidTrunc gB f ∣ x ∣₁ = f x -recGroupoidTrunc gB f (squash₁ _ _ _ _ r s i j k) = - gB _ _ _ _ - (λ m n → recGroupoidTrunc gB f (r m n)) - (λ m n → recGroupoidTrunc gB f (s m n)) +private + variable + ℓ : Level + A : Type ℓ + +rec : ∀ {B : Type ℓ} → isGroupoid B → (A → B) → ∥ A ∥₃ → B +rec gB f ∣ x ∣₃ = f x +rec gB f (squash₃ x y p q r s i j k) = + gB _ _ _ _ (λ m n → rec gB f (r m n)) (λ m n → rec gB f (s m n)) i j k + +elim : {B : ∥ A ∥₃ → Type ℓ} + (bG : (x : ∥ A ∥₃) → isGroupoid (B x)) + (f : (x : A) → B ∣ x ∣₃) (x : ∥ A ∥₃) → B x +elim bG f (∣ x ∣₃) = f x +elim bG f (squash₃ x y p q r s i j k) = + isOfHLevel→isOfHLevelDep 3 bG _ _ _ _ + (λ j k → elim bG f (r j k)) (λ j k → elim bG f (s j k)) + (squash₃ x y p q r s) i j k -groupoidTruncFib : ∀ {ℓ ℓ'} {A : Type ℓ} (P : A → Type ℓ') - {a b : A} (sPb : isGroupoid (P b)) - {p q : a ≡ b} {r s : p ≡ q} (u : r ≡ s) {a1 : P a} {b1 : P b} - {p1 : PathP (λ i → P (p i)) a1 b1} - {q1 : PathP (λ i → P (q i)) a1 b1} - (r1 : PathP (λ i → PathP (λ j → P (r i j)) a1 b1) p1 q1) - (s1 : PathP (λ i → PathP (λ j → P (s i j)) a1 b1) p1 q1) - → PathP (λ i → PathP (λ j → PathP (λ k → P (u i j k)) a1 b1) p1 q1) r1 s1 -groupoidTruncFib P {a} {b} sPb u {a1} {b1} {p1} {q1} r1 s1 i j k = - hcomp (λ l → λ { (i = i0) → r1 j k - ; (i = i1) → s1 j k - ; (j = i0) → p1 k - ; (j = i1) → q1 k - ; (k = i0) → a1 - ; (k = i1) → sPb b1 b1 refl refl (λ i j → Lb i j i1) refl l i j - }) - (Lb i j k) +elim2 : {B : ∥ A ∥₃ → ∥ A ∥₃ → Type ℓ} + (gB : ((x y : ∥ A ∥₃) → isGroupoid (B x y))) + (g : (a b : A) → B ∣ a ∣₃ ∣ b ∣₃) + (x y : ∥ A ∥₃) → B x y +elim2 gB g = elim (λ _ → isGroupoidΠ (λ _ → gB _ _)) + (λ a → elim (λ _ → gB _ _) (g a)) + +elim3 : {B : (x y z : ∥ A ∥₃) → Type ℓ} + (gB : ((x y z : ∥ A ∥₃) → isGroupoid (B x y z))) + (g : (a b c : A) → B ∣ a ∣₃ ∣ b ∣₃ ∣ c ∣₃) + (x y z : ∥ A ∥₃) → B x y z +elim3 gB g = elim2 (λ _ _ → isGroupoidΠ (λ _ → gB _ _ _)) + (λ a b → elim (λ _ → gB _ _ _) (g a b)) + +isGroupoidGroupoidTrunc : isGroupoid ∥ A ∥₃ +isGroupoidGroupoidTrunc a b p q r s = squash₃ a b p q r s + +groupoidTruncIdempotent≃ : isGroupoid A → ∥ A ∥₃ ≃ A +groupoidTruncIdempotent≃ {A = A} hA = isoToEquiv f where - L : (i j : I) → P b - L i j = comp (λ k → P (u i j k)) - (λ k → λ { (i = i0) → r1 j k - ; (i = i1) → s1 j k - ; (j = i0) → p1 k - ; (j = i1) → q1 k }) - a1 - Lb : PathP (λ i → PathP (λ j → PathP (λ k → P (u i j k)) a1 (L i j)) p1 q1) r1 s1 - Lb i j k = fill (λ k → P (u i j k)) - (λ k → λ { (i = i0) → r1 j k - ; (i = i1) → s1 j k - ; (j = i0) → p1 k - ; (j = i1) → q1 k }) - (inS a1) k + f : Iso ∥ A ∥₃ A + Iso.fun f = rec hA (idfun A) + Iso.inv f x = ∣ x ∣₃ + Iso.rightInv f _ = refl + Iso.leftInv f = elim (λ _ → isGroupoid→is2Groupoid isGroupoidGroupoidTrunc _ _) (λ _ → refl) -groupoidTruncElim : ∀ {ℓ ℓ'} (A : Type ℓ) (B : ∥ A ∥₁ → Type ℓ') - (bG : (x : ∥ A ∥₁) → isGroupoid (B x)) - (f : (x : A) → B ∣ x ∣₁) (x : ∥ A ∥₁) → B x -groupoidTruncElim A B bG f (∣ x ∣₁) = f x -groupoidTruncElim A B bG f (squash₁ x y p q r s i j k) = - groupoidTruncFib B (bG y) - (squash₁ x y p q r s) - (λ i j → groupoidTruncElim A B bG f (r i j)) - (λ i j → groupoidTruncElim A B bG f (s i j)) - i j k +groupoidTruncIdempotent : isGroupoid A → ∥ A ∥₃ ≡ A +groupoidTruncIdempotent hA = ua (groupoidTruncIdempotent≃ hA) diff --git a/Cubical/HITs/Hopf.agda b/Cubical/HITs/Hopf.agda index 51008e845..0b72621fb 100644 --- a/Cubical/HITs/Hopf.agda +++ b/Cubical/HITs/Hopf.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Hopf where open import Cubical.Foundations.Prelude @@ -7,19 +7,21 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence -open import Cubical.Data.Int -open import Cubical.Data.Prod +open import Cubical.Data.Int hiding (_·_) +open import Cubical.Data.Sigma +open import Cubical.Foundations.Function open import Cubical.HITs.S1 open import Cubical.HITs.S2 open import Cubical.HITs.S3 +open import Cubical.HITs.Sn open import Cubical.HITs.Susp open import Cubical.HITs.Join open import Cubical.HITs.Interval renaming ( zero to I0 ; one to I1 ) Border : (x : S¹) → (j : I) → Partial (j ∨ ~ j) (Σ Type₀ (λ T → T ≃ S¹)) -Border x j (j = i0) = S¹ , rot x , rotIsEquiv x +Border x j (j = i0) = S¹ , (x ·_) , rotIsEquiv x Border x j (j = i1) = S¹ , idEquiv S¹ -- Hopf fibration using SuspS¹ @@ -54,7 +56,7 @@ TotalHopf = Σ SuspS¹ HopfSuspS¹ filler-1 : I → (j : I) → (y : S¹) → Glue S¹ (Border y j) → join S¹ S¹ filler-1 i j y x = hfill (λ t → λ { (j = i0) → inl (rotInv-1 x y t) ; (j = i1) → inr x }) - (inS (push ((unglue (j ∨ ~ j) x) * inv y) (unglue (j ∨ ~ j) x) j)) i + (inS (push ((unglue (j ∨ ~ j) x) · invLooper y) (unglue (j ∨ ~ j) x) j)) i TotalHopf→JoinS¹S¹ : TotalHopf → join S¹ S¹ TotalHopf→JoinS¹S¹ (north , x) = inl x @@ -66,7 +68,7 @@ JoinS¹S¹→TotalHopf : join S¹ S¹ → TotalHopf JoinS¹S¹→TotalHopf (inl x) = (north , x) JoinS¹S¹→TotalHopf (inr x) = (south , x) JoinS¹S¹→TotalHopf (push y x j) = - (merid (inv y * x) j + (merid (invLooper y · x) j , glue (λ { (j = i0) → y ; (j = i1) → x }) (rotInv-2 x y j)) -- Now for the homotopies, we will need to fill squares indexed by x y : S¹ with value in S¹ @@ -79,8 +81,8 @@ JoinS¹S¹→TotalHopf (push y x j) = -- this should be generalized to a constant fibration over a connected space with -- discrete fiber -fibInt : S¹ → S¹ → Type₀ -fibInt _ _ = Int +fibℤ : S¹ → S¹ → Type₀ +fibℤ _ _ = ℤ S¹→HSet : (A : Type₀) (p : isSet A) (F : S¹ → A) (x : S¹) → F base ≡ F x S¹→HSet A p F base = refl {x = F base} @@ -93,17 +95,17 @@ S¹→HSet A p F (loop i) = f' i f' : PathP (λ i → F base ≡ F (loop i)) (refl {x = F base}) (refl {x = F base}) f' = transport (λ i → PathP (λ j → F base ≡ F (loop j)) refl (L i)) f -constant-loop : (F : S¹ → S¹ → Int) → (x y : S¹) → F base base ≡ F x y +constant-loop : (F : S¹ → S¹ → ℤ) → (x y : S¹) → F base base ≡ F x y constant-loop F x y = L0 ∙ L1 where - p : isSet (S¹ → Int) - p = isOfHLevelPi 2 (λ _ → isSetInt) + p : isSet (S¹ → ℤ) + p = isSetΠ (λ _ → isSetℤ) L : F base ≡ F x - L = S¹→HSet (S¹ → Int) p F x + L = S¹→HSet (S¹ → ℤ) p F x L0 : F base base ≡ F x base L0 i = L i base L1 : F x base ≡ F x y - L1 = S¹→HSet Int isSetInt (F x) y + L1 = S¹→HSet ℤ isSetℤ (F x) y discretefib : (F : S¹ → S¹ → Type₀) → Type₀ discretefib F = (a : (x y : S¹) → F x y) → @@ -111,8 +113,8 @@ discretefib F = (a : (x y : S¹) → F x y) → (a base base ≡ b base base) → (x y : S¹) → a x y ≡ b x y -discretefib-fibInt : discretefib fibInt -discretefib-fibInt a b h x y i = +discretefib-fibℤ : discretefib fibℤ +discretefib-fibℤ a b h x y i = hcomp (λ t → λ { (i = i0) → constant-loop a x y t ; (i = i1) → constant-loop b x y t }) (h i) @@ -121,13 +123,13 @@ discretefib-fibInt a b h x y i = assocFiller-3-aux : I → I → I → I → S¹ assocFiller-3-aux x y j i = - hfill (λ t → λ { (i = i0) → rotInv-1 (loop y) (loop (~ y) * loop x) t + hfill (λ t → λ { (i = i0) → rotInv-1 (loop y) (loop (~ y) · loop x) t ; (i = i1) → rotInv-3 (loop y) (loop x) t ; (x = i0) (y = i0) → base ; (x = i0) (y = i1) → base ; (x = i1) (y = i0) → base ; (x = i1) (y = i1) → base }) - (inS ((rotInv-2 (loop x) (loop y) i) * (inv (loop (~ y) * loop x)))) j + (inS ((rotInv-2 (loop x) (loop y) i) · (invLooper (loop (~ y) · loop x)))) j -- assocFiller-3-endpoint is used only in the type of the next function, to specify the -- second endpoint. @@ -140,8 +142,8 @@ assocFiller-3-endpoint base (loop y) i = assocFiller-3-aux i0 y i1 i assocFiller-3-endpoint (loop x) (loop y) i = assocFiller-3-aux x y i1 i assocFiller-3 : (x : S¹) → (y : S¹) → - PathP (λ j → rotInv-1 y (inv y * x) j ≡ rotInv-3 y x j) - (λ i → ((rotInv-2 x y i) * (inv (inv y * x)))) + PathP (λ j → rotInv-1 y (invLooper y · x) j ≡ rotInv-3 y x j) + (λ i → ((rotInv-2 x y i) · (invLooper (invLooper y · x)))) (assocFiller-3-endpoint x y) assocFiller-3 base base j i = base assocFiller-3 (loop x) base j i = assocFiller-3-aux x i0 j i @@ -151,12 +153,12 @@ assocFiller-3 (loop x) (loop y) j i = assocFiller-3-aux x y j i assoc-3 : (_ y : S¹) → basedΩS¹ y assoc-3 x y i = assocFiller-3 x y i1 i -fibInt≡fibAssoc-3 : fibInt ≡ (λ _ y → basedΩS¹ y) -fibInt≡fibAssoc-3 i = λ x y → basedΩS¹≡Int y (~ i) +fibℤ≡fibAssoc-3 : fibℤ ≡ (λ _ y → basedΩS¹ y) +fibℤ≡fibAssoc-3 i = λ x y → basedΩS¹≡ℤ y (~ i) discretefib-fibAssoc-3 : discretefib (λ _ y → basedΩS¹ y) discretefib-fibAssoc-3 = - transp (λ i → discretefib (fibInt≡fibAssoc-3 i)) i0 discretefib-fibInt + transp (λ i → discretefib (fibℤ≡fibAssoc-3 i)) i0 discretefib-fibℤ assocConst-3 : (x y : S¹) → assoc-3 x y ≡ refl assocConst-3 x y = discretefib-fibAssoc-3 assoc-3 (λ _ _ → refl) refl x y @@ -170,13 +172,13 @@ assocSquare-3 i j x y = hcomp (λ t → λ { (i = i0) → assocFiller-3 x y j i0 filler-3 : I → I → S¹ → S¹ → join S¹ S¹ filler-3 i j y x = - hcomp (λ t → λ { (i = i0) → filler-1 t j (inv y * x) + hcomp (λ t → λ { (i = i0) → filler-1 t j (invLooper y · x) (glue (λ { (j = i0) → y ; (j = i1) → x }) (rotInv-2 x y j)) ; (i = i1) → push (rotInv-3 y x t) x j ; (j = i0) → inl (assocSquare-3 i t x y) ; (j = i1) → inr x }) - (push ((rotInv-2 x y (i ∨ j)) * (inv (inv y * x))) (rotInv-2 x y (i ∨ j)) j) + (push ((rotInv-2 x y (i ∨ j)) · (invLooper (invLooper y · x))) (rotInv-2 x y (i ∨ j)) j) JoinS¹S¹→TotalHopf→JoinS¹S¹ : ∀ x → TotalHopf→JoinS¹S¹ (JoinS¹S¹→TotalHopf x) ≡ x JoinS¹S¹→TotalHopf→JoinS¹S¹ (inl x) i = inl x @@ -201,41 +203,41 @@ PseudoHopf-π2 (_ , x) = x assocFiller-4-aux : I → I → I → I → S¹ assocFiller-4-aux x y j i = - hfill (λ t → λ { (i = i0) → ((inv (loop y * loop x * loop (~ y))) * (loop y * loop x)) - * (rotInv-1 (loop x) (loop y) t) - ; (i = i1) → (rotInv-4 (loop y) (loop y * loop x) (~ t)) * loop x + hfill (λ t → λ { (i = i0) → ((invLooper (loop y · loop x · loop (~ y))) · (loop y · loop x)) + · (rotInv-1 (loop x) (loop y) t) + ; (i = i1) → (rotInv-4 (loop y) (loop y · loop x) (~ t)) · loop x ; (x = i0) (y = i0) → base ; (x = i0) (y = i1) → base ; (x = i1) (y = i0) → base ; (x = i1) (y = i1) → base }) - (inS (rotInv-2 (loop y * loop x) (loop y * loop x * loop (~ y)) i)) j + (inS (rotInv-2 (loop y · loop x) (loop y · loop x · loop (~ y)) i)) j -- See assocFiller-3-endpoint -- TODO : use cubical extension types when available to remove assocFiller-4-endpoint -assocFiller-4-endpoint : (x y : S¹) → basedΩS¹ (((inv (y * x * inv y)) * (y * x)) * x) +assocFiller-4-endpoint : (x y : S¹) → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x) assocFiller-4-endpoint base base i = base assocFiller-4-endpoint (loop x) base i = assocFiller-4-aux x i0 i1 i assocFiller-4-endpoint base (loop y) i = assocFiller-4-aux i0 y i1 i assocFiller-4-endpoint (loop x) (loop y) i = assocFiller-4-aux x y i1 i assocFiller-4 : (x y : S¹) → - PathP (λ j → ((inv (y * x * inv y)) * (y * x)) * (rotInv-1 x y j) ≡ (rotInv-4 y (y * x) (~ j)) * x) - (λ i → (rotInv-2 (y * x) (y * x * inv y) i)) + PathP (λ j → ((invLooper (y · x · invLooper y)) · (y · x)) · (rotInv-1 x y j) ≡ (rotInv-4 y (y · x) (~ j)) · x) + (λ i → (rotInv-2 (y · x) (y · x · invLooper y) i)) (assocFiller-4-endpoint x y) assocFiller-4 base base j i = base assocFiller-4 (loop x) base j i = assocFiller-4-aux x i0 j i assocFiller-4 base (loop y) j i = assocFiller-4-aux i0 y j i assocFiller-4 (loop x) (loop y) j i = assocFiller-4-aux x y j i -assoc-4 : (x y : S¹) → basedΩS¹ (((inv (y * x * inv y)) * (y * x)) * x) +assoc-4 : (x y : S¹) → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x) assoc-4 x y i = assocFiller-4 x y i1 i -fibInt≡fibAssoc-4 : fibInt ≡ (λ x y → basedΩS¹ (((inv (y * x * inv y)) * (y * x)) * x)) -fibInt≡fibAssoc-4 i = λ x y → basedΩS¹≡Int (((inv (y * x * inv y)) * (y * x)) * x) (~ i) +fibℤ≡fibAssoc-4 : fibℤ ≡ (λ x y → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x)) +fibℤ≡fibAssoc-4 i = λ x y → basedΩS¹≡ℤ (((invLooper (y · x · invLooper y)) · (y · x)) · x) (~ i) -discretefib-fibAssoc-4 : discretefib (λ x y → basedΩS¹ (((inv (y * x * inv y)) * (y * x)) * x)) +discretefib-fibAssoc-4 : discretefib (λ x y → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x)) discretefib-fibAssoc-4 = - transp (λ i → discretefib (fibInt≡fibAssoc-4 i)) i0 discretefib-fibInt + transp (λ i → discretefib (fibℤ≡fibAssoc-4 i)) i0 discretefib-fibℤ assocConst-4 : (x y : S¹) → assoc-4 x y ≡ refl assocConst-4 x y = discretefib-fibAssoc-4 assoc-4 (λ _ _ → refl) refl x y @@ -251,18 +253,18 @@ assocSquare-4 i j x y = filler-4-0 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf filler-4-0 i j y x = let x' = unglue (j ∨ ~ j) x in - hfill (λ t → λ { (j = i0) → ((inv (y * x * inv y) * (y * x) , I0) - , inv (y * x * inv y) * (y * x) * (rotInv-1 x y t)) - ; (j = i1) → ((inv (x * inv y) * x , I1) , x) }) - (inS ((inv (x' * inv y) * x' , seg j) , rotInv-2 x' (x' * inv y) j)) i + hfill (λ t → λ { (j = i0) → ((invLooper (y · x · invLooper y) · (y · x) , I0) + , invLooper (y · x · invLooper y) · (y · x) · (rotInv-1 x y t)) + ; (j = i1) → ((invLooper (x · invLooper y) · x , I1) , x) }) + (inS ((invLooper (x' · invLooper y) · x' , seg j) , rotInv-2 x' (x' · invLooper y) j)) i filler-4-1 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf filler-4-1 i j y x = let x' = unglue (j ∨ ~ j) x in - hfill (λ t → λ { (j = i0) → ((inv (y * x * inv y) * (y * x) , I0) - , (rotInv-4 y (y * x) (~ t)) * x) - ; (j = i1) → ((inv (x * inv y) * x , I1) , x) }) - (inS ((inv (x' * inv y) * x' , seg j) , unglue (j ∨ ~ j) x)) i + hfill (λ t → λ { (j = i0) → ((invLooper (y · x · invLooper y) · (y · x) , I0) + , (rotInv-4 y (y · x) (~ t)) · x) + ; (j = i1) → ((invLooper (x · invLooper y) · x , I1) , x) }) + (inS ((invLooper (x' · invLooper y) · x' , seg j) , unglue (j ∨ ~ j) x)) i filler-4-2 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → TotalHopf filler-4-2 i j y x = @@ -273,25 +275,25 @@ filler-4-2 i j y x = (PseudoHopf-π2 (filler-4-0 t j y x))) ; (j = i0) → (north , rotInv-1 x y t) ; (j = i1) → (south , x) }) - (merid (inv (x' * inv y) * x') j - , glue (λ { (j = i0) → y * x * inv y ; (j = i1) → x }) (rotInv-2 x' (x' * inv y) j)) + (merid (invLooper (x' · invLooper y) · x') j + , glue (λ { (j = i0) → y · x · invLooper y ; (j = i1) → x }) (rotInv-2 x' (x' · invLooper y) j)) filler-4-3 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf filler-4-3 i j y x = let x' = unglue (j ∨ ~ j) x in hcomp (λ t → λ { (i = i0) → filler-4-0 t j y x ; (i = i1) → filler-4-1 t j y x - ; (j = i0) → ((inv (y * x * inv y) * (y * x) , I0) , assocSquare-4 i t x y) - ; (j = i1) → ((inv (x * inv y) * x , I1) , x) }) - ((inv (x' * inv y) * x' , seg j) , rotInv-2 x' (x' * inv y) (i ∨ j)) + ; (j = i0) → ((invLooper (y · x · invLooper y) · (y · x) , I0) , assocSquare-4 i t x y) + ; (j = i1) → ((invLooper (x · invLooper y) · x , I1) , x) }) + ((invLooper (x' · invLooper y) · x' , seg j) , rotInv-2 x' (x' · invLooper y) (i ∨ j)) filler-4-4 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf filler-4-4 i j y x = let x' = unglue (j ∨ ~ j) x in hcomp (λ t → λ { (i = i0) → filler-4-1 t j y x ; (i = i1) → ((y , seg j) , unglue (j ∨ ~ j) x) - ; (j = i0) → ((rotInv-4 y (y * x) i , I0) - , (rotInv-4 y (y * x) (i ∨ ~ t)) * x) + ; (j = i0) → ((rotInv-4 y (y · x) i , I0) + , (rotInv-4 y (y · x) (i ∨ ~ t)) · x) ; (j = i1) → ((rotInv-4 y x i , I1) , x) }) ((rotInv-4 y x' i , seg j) , x') @@ -320,3 +322,19 @@ JoinS¹S¹≡TotalHopf = isoToPath (iso JoinS¹S¹→TotalHopf S³≡TotalHopf : S³ ≡ TotalHopf S³≡TotalHopf = S³≡joinS¹S¹ ∙ JoinS¹S¹≡TotalHopf + +open Iso +IsoS³TotalHopf : Iso (S₊ 3) TotalHopf +fun IsoS³TotalHopf x = JoinS¹S¹→TotalHopf (S³→joinS¹S¹ (inv IsoS³S3 x)) +inv IsoS³TotalHopf x = fun IsoS³S3 (joinS¹S¹→S³ (TotalHopf→JoinS¹S¹ x)) +rightInv IsoS³TotalHopf x = + cong (JoinS¹S¹→TotalHopf ∘ S³→joinS¹S¹) + (leftInv IsoS³S3 (joinS¹S¹→S³ (TotalHopf→JoinS¹S¹ x))) + ∙∙ cong JoinS¹S¹→TotalHopf + (joinS¹S¹→S³→joinS¹S¹ (TotalHopf→JoinS¹S¹ x)) + ∙∙ TotalHopf→JoinS¹S¹→TotalHopf x +leftInv IsoS³TotalHopf x = + cong (fun IsoS³S3 ∘ joinS¹S¹→S³) + (JoinS¹S¹→TotalHopf→JoinS¹S¹ (S³→joinS¹S¹ (inv IsoS³S3 x))) + ∙∙ cong (fun IsoS³S3) (S³→joinS¹S¹→S³ (inv IsoS³S3 x)) + ∙∙ Iso.rightInv IsoS³S3 x diff --git a/Cubical/HITs/InfNat.agda b/Cubical/HITs/InfNat.agda index 3af94c17e..d45d38cb4 100644 --- a/Cubical/HITs/InfNat.agda +++ b/Cubical/HITs/InfNat.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --no-exact-split --safe #-} +{-# OPTIONS --no-exact-split --safe #-} module Cubical.HITs.InfNat where diff --git a/Cubical/HITs/InfNat/Base.agda b/Cubical/HITs/InfNat/Base.agda index c20be7b7b..992ab56fc 100644 --- a/Cubical/HITs/InfNat/Base.agda +++ b/Cubical/HITs/InfNat/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --no-exact-split --safe #-} +{-# OPTIONS --no-exact-split --safe #-} module Cubical.HITs.InfNat.Base where diff --git a/Cubical/HITs/InfNat/Properties.agda b/Cubical/HITs/InfNat/Properties.agda index 985d7c7a8..f974ba7a9 100644 --- a/Cubical/HITs/InfNat/Properties.agda +++ b/Cubical/HITs/InfNat/Properties.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --no-exact-split --safe #-} +{-# OPTIONS --no-exact-split --safe #-} module Cubical.HITs.InfNat.Properties where diff --git a/Cubical/HITs/Interval.agda b/Cubical/HITs/Interval.agda index 21fb18178..ca584ce13 100644 --- a/Cubical/HITs/Interval.agda +++ b/Cubical/HITs/Interval.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Interval where open import Cubical.HITs.Interval.Base public diff --git a/Cubical/HITs/Interval/Base.agda b/Cubical/HITs/Interval/Base.agda index a454a7fb9..54865c63b 100644 --- a/Cubical/HITs/Interval/Base.agda +++ b/Cubical/HITs/Interval/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Interval.Base where open import Cubical.Core.Everything @@ -26,14 +26,14 @@ funExtInterval A B f g p = λ i x → hmtpy x (seg i) hmtpy x one = g x hmtpy x (seg i) = p x i -intervalElim : (A : Interval → Type₀) (x : A zero) (y : A one) +elim : (A : Interval → Type₀) (x : A zero) (y : A one) (p : PathP (λ i → A (seg i)) x y) → (x : Interval) → A x -intervalElim A x y p zero = x -intervalElim A x y p one = y -intervalElim A x y p (seg i) = p i +elim A x y p zero = x +elim A x y p one = y +elim A x y p (seg i) = p i -- Note that this is not definitional (it is not proved by refl) -intervalEta : ∀ {A : Type₀} (f : Interval → A) → intervalElim _ (f zero) (f one) (λ i → f (seg i)) ≡ f +intervalEta : ∀ {A : Type₀} (f : Interval → A) → elim _ (f zero) (f one) (λ i → f (seg i)) ≡ f intervalEta f i zero = f zero intervalEta f i one = f one intervalEta f i (seg j) = f (seg j) diff --git a/Cubical/HITs/Ints/BiInvInt.agda b/Cubical/HITs/Ints/BiInvInt.agda deleted file mode 100644 index a0d3529fe..000000000 --- a/Cubical/HITs/Ints/BiInvInt.agda +++ /dev/null @@ -1,6 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Ints.BiInvInt where - -open import Cubical.HITs.Ints.BiInvInt.Base public - -open import Cubical.HITs.Ints.BiInvInt.Properties public diff --git a/Cubical/HITs/Ints/BiInvInt/Properties.agda b/Cubical/HITs/Ints/BiInvInt/Properties.agda deleted file mode 100644 index 9a5451a0b..000000000 --- a/Cubical/HITs/Ints/BiInvInt/Properties.agda +++ /dev/null @@ -1,148 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Ints.BiInvInt.Properties where - -open import Cubical.Core.Everything - -open import Cubical.Foundations.Prelude - -open import Cubical.Data.Nat hiding (_+_; +-comm) -open import Cubical.Data.Int -open import Cubical.Data.Bool - -open import Cubical.HITs.Ints.BiInvInt.Base - --- addition - -_+ᴮ_ : BiInvInt → BiInvInt → BiInvInt -m +ᴮ zero = m -m +ᴮ suc n = suc (m +ᴮ n) -m +ᴮ predr n = predr (m +ᴮ n) -m +ᴮ predl n = predl (m +ᴮ n) -m +ᴮ suc-predr n i = suc-predr (m +ᴮ n) i -m +ᴮ predl-suc n i = predl-suc (m +ᴮ n) i - --- properties of addition - -+ᴮ-assoc : ∀ l m n → (l +ᴮ m) +ᴮ n ≡ l +ᴮ (m +ᴮ n) -+ᴮ-assoc l m zero i = l +ᴮ m -+ᴮ-assoc l m (suc n) i = suc (+ᴮ-assoc l m n i) -+ᴮ-assoc l m (predr n) i = predr (+ᴮ-assoc l m n i) -+ᴮ-assoc l m (predl n) i = predl (+ᴮ-assoc l m n i) -+ᴮ-assoc l m (suc-predr n i) j = suc-predr (+ᴮ-assoc l m n j) i -+ᴮ-assoc l m (predl-suc n i) j = predl-suc (+ᴮ-assoc l m n j) i - -+ᴮ-unitʳ : ∀ n → n +ᴮ zero ≡ n -+ᴮ-unitʳ n i = n - -+ᴮ-unitˡ : ∀ n → zero +ᴮ n ≡ n -+ᴮ-unitˡ zero i = zero -+ᴮ-unitˡ (suc n) i = suc (+ᴮ-unitˡ n i) -+ᴮ-unitˡ (predr n) i = predr (+ᴮ-unitˡ n i) -+ᴮ-unitˡ (predl n) i = predl (+ᴮ-unitˡ n i) -+ᴮ-unitˡ (suc-predr n i) j = suc-predr (+ᴮ-unitˡ n j) i -+ᴮ-unitˡ (predl-suc n i) j = predl-suc (+ᴮ-unitˡ n j) i - --- TODO: a direct proof of commutatitivty --- (for now, we use Data.Int) - -fwd-+≡+ᴮ : ∀ m n → fwd (m + n) ≡ (fwd m) +ᴮ (fwd n) -fwd-+≡+ᴮ m (pos zero) = refl -fwd-+≡+ᴮ m (pos (suc n)) = fwd-sucInt (m +pos n) ∙ cong suc (fwd-+≡+ᴮ m (pos n)) -fwd-+≡+ᴮ m (negsuc zero) = fwd-predInt m -fwd-+≡+ᴮ m (negsuc (suc n)) = fwd-predInt (m +negsuc n) ∙ cong pred (fwd-+≡+ᴮ m (negsuc n)) - -+ᴮ≡+ : ∀ m n → m +ᴮ n ≡ fwd ((bwd m) + (bwd n)) -+ᴮ≡+ m n = sym (fwd-+≡+ᴮ (bwd m) (bwd n) ∙ (λ i → (fwd-bwd m i) +ᴮ (fwd-bwd n i))) - -+ᴮ-comm : ∀ m n → m +ᴮ n ≡ n +ᴮ m -+ᴮ-comm m n = +ᴮ≡+ m n ∙ cong fwd (+-comm (bwd m) (bwd n)) ∙ sym (+ᴮ≡+ n m) - --- some of the lemmas needed for a direct proof +ᴮ-comm are corollaries of +ᴮ-comm - -suc-+ᴮ : ∀ m n → (suc m) +ᴮ n ≡ suc (m +ᴮ n) -suc-+ᴮ m n = +ᴮ-comm (suc m) n ∙ (λ i → suc (+ᴮ-comm n m i)) --- suc-+ᴮ m zero i = suc m --- suc-+ᴮ m (suc n) i = suc (suc-+ᴮ m n i) --- suc-+ᴮ m (predr n) = cong predr (suc-+ᴮ m n) ∙ predr-suc (m +ᴮ n) ∙ sym (suc-predr (m +ᴮ n)) --- suc-+ᴮ m (predl n) = cong predl (suc-+ᴮ m n) ∙ predl-suc (m +ᴮ n) ∙ sym (suc-predl (m +ᴮ n)) --- suc-+ᴮ m (suc-predr n i) j = {!!} --- suc-+ᴮ m (predl-suc n i) j = {!!} - -predr-+ᴮ : ∀ m n → (predr m) +ᴮ n ≡ predr (m +ᴮ n) -predr-+ᴮ m n = +ᴮ-comm (predr m) n ∙ (λ i → predr (+ᴮ-comm n m i)) - -predl-+ᴮ : ∀ m n → (predl m) +ᴮ n ≡ predl (m +ᴮ n) -predl-+ᴮ m n = +ᴮ-comm (predl m) n ∙ (λ i → predl (+ᴮ-comm n m i)) - --- +ᴮ-comm : ∀ m n → n +ᴮ m ≡ m +ᴮ n --- +ᴮ-comm m zero = +ᴮ-unitˡ m --- +ᴮ-comm m (suc n) = suc-+ᴮ n m ∙ cong suc (+ᴮ-comm m n) --- +ᴮ-comm m (predr n) = predr-+ᴮ n m ∙ cong predr (+ᴮ-comm m n) --- +ᴮ-comm m (predl n) = predl-+ᴮ n m ∙ cong predl (+ᴮ-comm m n) --- +ᴮ-comm m (suc-predr n i) j = {!!} --- +ᴮ-comm m (predl-suc n i) j = {!!} - - --- negation / subtraction - --ᴮ_ : BiInvInt → BiInvInt --ᴮ zero = zero --ᴮ suc n = predl (-ᴮ n) --ᴮ predr n = suc (-ᴮ n) --ᴮ predl n = suc (-ᴮ n) --ᴮ suc-predr n i = predl-suc (-ᴮ n) i --ᴮ predl-suc n i = suc-predl (-ᴮ n) i - -_-ᴮ_ : BiInvInt → BiInvInt → BiInvInt -m -ᴮ n = m +ᴮ (-ᴮ n) - --- TODO: properties of negation - --- +ᴮ-invˡ : ∀ n → (-ᴮ n) +ᴮ n ≡ zero --- +ᴮ-invˡ zero = refl --- +ᴮ-invˡ (suc n) = (λ i → suc (predl-+ᴮ (-ᴮ n) n i)) ∙ (λ i → suc-pred (+ᴮ-invˡ n i) i) --- +ᴮ-invˡ (predr n) = (λ i → predr (suc-+ᴮ (-ᴮ n) n i)) ∙ (λ i → predr-suc (+ᴮ-invˡ n i) i) --- +ᴮ-invˡ (predl n) = (λ i → predl (suc-+ᴮ (-ᴮ n) n i)) ∙ (λ i → predl-suc (+ᴮ-invˡ n i) i) --- +ᴮ-invˡ (suc-predr n i) j = {!!} --- +ᴮ-invˡ (predl-suc n i) j = {!!} - --- +ᴮ-invʳ : ∀ n → n +ᴮ (-ᴮ n) ≡ zero --- +ᴮ-invʳ n = {!!} - - --- natural injections from ℕ - -posᴮ : ℕ → BiInvInt -posᴮ zero = zero -posᴮ (suc n) = suc (posᴮ n) - -negᴮ : ℕ → BiInvInt -negᴮ zero = zero -negᴮ (suc n) = pred (negᴮ n) - --- absolute value and sign --- (Note that there doesn't appear to be any way around using --- bwd here! Any direct proof ends up doing the same work...) - -absᴮ : BiInvInt → ℕ -absᴮ n = abs (bwd n) - -sgnᴮ : BiInvInt → Bool -sgnᴮ n = sgn (bwd n) - - --- TODO: a direct definition of multiplication using +ᴮ-invˡ/ʳ --- (for now we use abs and sgn, as in agda's stdlib) - -_*ᴮ_ : BiInvInt → BiInvInt → BiInvInt -m *ᴮ n = caseBool posᴮ negᴮ (sgnᴮ m and sgnᴮ n) (absᴮ m * absᴮ n) --- m *ᴮ zero = zero --- m *ᴮ suc n = (m *ᴮ n) +ᴮ m --- m *ᴮ predr n = (m *ᴮ n) -ᴮ m --- m *ᴮ predl n = (m *ᴮ n) -ᴮ m --- m *ᴮ suc-predr n i = ( +ᴮ-assoc (m *ᴮ n) (-ᴮ m) m --- ∙ cong ((m *ᴮ n) +ᴮ_) (+ᴮ-invˡ m) --- ∙ +ᴮ-unitʳ (m *ᴮ n)) i --- m *ᴮ predl-suc n i = ( +ᴮ-assoc (m *ᴮ n) m (-ᴮ m) --- ∙ cong ((m *ᴮ n) +ᴮ_) (+ᴮ-invʳ m) --- ∙ +ᴮ-unitʳ (m *ᴮ n)) i diff --git a/Cubical/HITs/Ints/DeltaInt.agda b/Cubical/HITs/Ints/DeltaInt.agda deleted file mode 100644 index bec6e34e8..000000000 --- a/Cubical/HITs/Ints/DeltaInt.agda +++ /dev/null @@ -1,6 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Ints.DeltaInt where - -open import Cubical.HITs.Ints.DeltaInt.Base public - -open import Cubical.HITs.Ints.DeltaInt.Properties public diff --git a/Cubical/HITs/Ints/HAEquivInt.agda b/Cubical/HITs/Ints/HAEquivInt.agda deleted file mode 100644 index 956fba76b..000000000 --- a/Cubical/HITs/Ints/HAEquivInt.agda +++ /dev/null @@ -1,5 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Ints.HAEquivInt where - -open import Cubical.HITs.Ints.HAEquivInt.Base public - diff --git a/Cubical/HITs/Ints/IsoInt.agda b/Cubical/HITs/Ints/IsoInt.agda deleted file mode 100644 index 97cd54f36..000000000 --- a/Cubical/HITs/Ints/IsoInt.agda +++ /dev/null @@ -1,5 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Ints.IsoInt where - -open import Cubical.HITs.Ints.IsoInt.Base public - diff --git a/Cubical/HITs/Ints/QuoInt.agda b/Cubical/HITs/Ints/QuoInt.agda deleted file mode 100644 index 1478c4e40..000000000 --- a/Cubical/HITs/Ints/QuoInt.agda +++ /dev/null @@ -1,6 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Ints.QuoInt where - -open import Cubical.HITs.Ints.QuoInt.Base public - --- open import Cubical.HITs.Ints.QuoInt.Properties public diff --git a/Cubical/HITs/Ints/QuoInt/Base.agda b/Cubical/HITs/Ints/QuoInt/Base.agda deleted file mode 100644 index 4bd36f57e..000000000 --- a/Cubical/HITs/Ints/QuoInt/Base.agda +++ /dev/null @@ -1,180 +0,0 @@ --- Define the integers as a HIT by identifying +0 and -0 -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Ints.QuoInt.Base where - -open import Cubical.Core.Everything - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Transport -open import Cubical.Foundations.Isomorphism - -open import Cubical.Data.Int hiding (neg; abs; sgn) -open import Cubical.Data.Nat - -variable - l : Level - -data ℤ : Type₀ where - pos : (n : ℕ) → ℤ - neg : (n : ℕ) → ℤ - posneg : pos 0 ≡ neg 0 - -recℤ : ∀ {A : Type l} → (pos' neg' : ℕ → A) → pos' 0 ≡ neg' 0 → ℤ → A -recℤ pos' neg' eq (pos m) = pos' m -recℤ pos' neg' eq (neg m) = neg' m -recℤ pos' neg' eq (posneg i) = eq i - -indℤ : ∀ (P : ℤ → Type l) - → (pos' : ∀ n → P (pos n)) - → (neg' : ∀ n → P (neg n)) - → (λ i → P (posneg i)) [ pos' 0 ≡ neg' 0 ] - → ∀ z → P z -indℤ P pos' neg' eq (pos n) = pos' n -indℤ P pos' neg' eq (neg n) = neg' n -indℤ P pos' neg' eq (posneg i) = eq i - -Int→ℤ : Int → ℤ -Int→ℤ (pos n) = pos n -Int→ℤ (negsuc n) = neg (suc n) - -ℤ→Int : ℤ → Int -ℤ→Int (pos n) = pos n -ℤ→Int (neg zero) = pos 0 -ℤ→Int (neg (suc n)) = negsuc n -ℤ→Int (posneg _) = pos 0 - -ℤ→Int→ℤ : ∀ (n : ℤ) → Int→ℤ (ℤ→Int n) ≡ n -ℤ→Int→ℤ (pos n) _ = pos n -ℤ→Int→ℤ (neg zero) i = posneg i -ℤ→Int→ℤ (neg (suc n)) _ = neg (suc n) -ℤ→Int→ℤ (posneg j) i = posneg (j ∧ i) - -Int→ℤ→Int : ∀ (n : Int) → ℤ→Int (Int→ℤ n) ≡ n -Int→ℤ→Int (pos n) _ = pos n -Int→ℤ→Int (negsuc n) _ = negsuc n - -Int≡ℤ : Int ≡ ℤ -Int≡ℤ = isoToPath (iso Int→ℤ ℤ→Int ℤ→Int→ℤ Int→ℤ→Int) - -isSetℤ : isSet ℤ -isSetℤ = subst isSet Int≡ℤ isSetInt - -sucℤ : ℤ → ℤ -sucℤ (pos n) = pos (suc n) -sucℤ (neg zero) = pos 1 -sucℤ (neg (suc n)) = neg n -sucℤ (posneg _) = pos 1 - -predℤ : ℤ → ℤ -predℤ (pos zero) = neg 1 -predℤ (pos (suc n)) = pos n -predℤ (neg n) = neg (suc n) -predℤ (posneg _) = neg 1 - -sucPredℤ : ∀ n → sucℤ (predℤ n) ≡ n -sucPredℤ (pos zero) = sym posneg -sucPredℤ (pos (suc _)) = refl -sucPredℤ (neg _) = refl -sucPredℤ (posneg i) j = posneg (i ∨ ~ j) - -predSucℤ : ∀ n → predℤ (sucℤ n) ≡ n -predSucℤ (pos _) = refl -predSucℤ (neg zero) = posneg -predSucℤ (neg (suc _)) = refl -predSucℤ (posneg i) j = posneg (i ∧ j) - -_+ℤ_ : ℤ → ℤ → ℤ -m +ℤ (pos (suc n)) = sucℤ (m +ℤ pos n) -m +ℤ (neg (suc n)) = predℤ (m +ℤ neg n) -m +ℤ _ = m - -sucPathℤ : ℤ ≡ ℤ -sucPathℤ = isoToPath (iso sucℤ predℤ sucPredℤ predSucℤ) - --- We do the same trick as in Cubical.Data.Int to prove that addition --- is an equivalence -addEqℤ : ℕ → ℤ ≡ ℤ -addEqℤ zero = refl -addEqℤ (suc n) = addEqℤ n ∙ sucPathℤ - -predPathℤ : ℤ ≡ ℤ -predPathℤ = isoToPath (iso predℤ sucℤ predSucℤ sucPredℤ) - -subEqℤ : ℕ → ℤ ≡ ℤ -subEqℤ zero = refl -subEqℤ (suc n) = subEqℤ n ∙ predPathℤ - -addℤ : ℤ → ℤ → ℤ -addℤ m (pos n) = transport (addEqℤ n) m -addℤ m (neg n) = transport (subEqℤ n) m -addℤ m (posneg _) = m - -isEquivAddℤ : (m : ℤ) → isEquiv (λ n → addℤ n m) -isEquivAddℤ (pos n) = isEquivTransport (addEqℤ n) -isEquivAddℤ (neg n) = isEquivTransport (subEqℤ n) -isEquivAddℤ (posneg _) = isEquivTransport refl - -addℤ≡+ℤ : addℤ ≡ _+ℤ_ -addℤ≡+ℤ i m (pos (suc n)) = sucℤ (addℤ≡+ℤ i m (pos n)) -addℤ≡+ℤ i m (neg (suc n)) = predℤ (addℤ≡+ℤ i m (neg n)) -addℤ≡+ℤ i m (pos zero) = m -addℤ≡+ℤ i m (neg zero) = m -addℤ≡+ℤ _ m (posneg _) = m - -isEquiv+ℤ : (m : ℤ) → isEquiv (λ n → n +ℤ m) -isEquiv+ℤ = subst (λ _+_ → (m : ℤ) → isEquiv (λ n → n + m)) addℤ≡+ℤ isEquivAddℤ - -data Sign : Type₀ where - pos neg : Sign - -sign : ℤ → Sign -sign (pos n) = pos -sign (neg 0) = pos -sign (neg (suc n)) = neg -sign (posneg i) = pos - -abs : ℤ → ℕ -abs (pos n) = n -abs (neg n) = n -abs (posneg i) = 0 - -signed : Sign → ℕ → ℤ -signed Sign.pos n = pos n -signed Sign.neg n = neg n - -signed-inv : ∀ z → signed (sign z) (abs z) ≡ z -signed-inv (pos n) = refl -signed-inv (neg zero) = posneg -signed-inv (neg (suc n)) = refl -signed-inv (posneg i) = \ j → posneg (i ∧ j) -{- - - The square for signed-inv (posneg i) - - - posneg i - ---------------------> - ^ ^ - | | - pos 0 | | posneg j - | | - | | - | | - ----------------------> - = pos 0 - = signed Sign.pos 0 - signed (sign (posneg i)) - (abs (posneg i)) --} - - --- * Multiplication - -_*S_ : Sign → Sign → Sign -pos *S neg = neg -neg *S pos = neg -_ *S _ = pos - -_*ℤ_ : ℤ → ℤ → ℤ -m *ℤ n = signed (sign m *S sign n) (abs m * abs n) diff --git a/Cubical/HITs/Ints/QuoInt/Properties.agda b/Cubical/HITs/Ints/QuoInt/Properties.agda deleted file mode 100644 index 96ddcf8fa..000000000 --- a/Cubical/HITs/Ints/QuoInt/Properties.agda +++ /dev/null @@ -1,68 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Ints.QuoInt.Properties where - -open import Cubical.Core.Everything -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.GroupoidLaws - -open import Cubical.HITs.Ints.QuoInt.Base - renaming (_+ℤ_ to _+_) -open import Cubical.Data.Nat - hiding (+-comm; +-zero; _+_) - -+-pos-0 : ∀ a → pos 0 + a ≡ a -+-pos-0 (pos zero) = refl -+-pos-0 (pos (suc n)) = cong sucℤ (+-pos-0 (pos n)) -+-pos-0 (neg zero) = posneg -+-pos-0 (neg (suc n)) = cong predℤ (+-pos-0 (neg n)) -+-pos-0 (posneg i) j = posneg (i ∧ j) - -+-neg-0 : ∀ a → neg 0 + a ≡ a -+-neg-0 (neg zero) = refl -+-neg-0 (neg (suc n)) = cong predℤ (+-neg-0 (neg n)) -+-neg-0 (pos zero) = sym posneg -+-neg-0 (pos (suc n)) = cong sucℤ (+-neg-0 (pos n)) -+-neg-0 (posneg i) j = posneg (i ∨ ~ j) - -+-pos-suc : ∀ a b → sucℤ (pos b + a) ≡ (pos (suc b) + a) -+-pos-suc (pos zero) b = refl -+-pos-suc (pos (suc n)) b = cong sucℤ (+-pos-suc (pos n) b) -+-pos-suc (neg zero) b = refl -+-pos-suc (posneg i) b = refl -+-pos-suc (neg (suc n)) b = - sucPredℤ (pos b + neg n) ∙∙ - sym (predSucℤ (pos b + neg n)) ∙∙ - cong predℤ (+-pos-suc (neg n) b) - -- the following hcomp does not termination-check: - -- hcomp (λ j → λ - -- { (i = i0) → sucPredℤ (pos b + neg n) (~ j) - -- ; (i = i1) → {!predℤ ((+-pos-suc (neg n) b) j)!} - -- }) (predSucℤ (pos b + neg n) (~ i)) - -+-neg-pred : ∀ a b → predℤ (neg b + a) ≡ (neg (suc b) + a) -+-neg-pred (pos zero) b = refl -+-neg-pred (neg zero) b = refl -+-neg-pred (neg (suc n)) b = cong predℤ (+-neg-pred (neg n) b) -+-neg-pred (posneg i) b = refl -+-neg-pred (pos (suc n)) b = - predSucℤ (neg b + pos n) ∙∙ - sym (sucPredℤ (neg b + pos n)) ∙∙ - cong sucℤ (+-neg-pred (pos n) b) - -+-comm : ∀ a b → a + b ≡ b + a -+-comm a (pos zero) = sym (+-pos-0 a) -+-comm a (neg zero) = sym (+-neg-0 a) -+-comm a (pos (suc b)) = cong sucℤ (+-comm a (pos b)) ∙ +-pos-suc a b -+-comm a (neg (suc b)) = cong predℤ (+-comm a (neg b)) ∙ +-neg-pred a b -+-comm a (posneg i) = lemma a i - where - -- get some help from: - -- https://www.codewars.com/kata/reviews/5c878e3ef1abb10001e32eb1/groups/5cca3f9e840f4b0001d8ac98 - lemma : ∀ a → (λ i → (a + posneg i) ≡ (posneg i + a)) - [ sym (+-pos-0 a) ≡ sym (+-neg-0 a) ] - lemma (pos zero) i j = posneg (i ∧ j) - lemma (pos (suc n)) i = cong sucℤ (lemma (pos n) i) - lemma (neg zero) i j = posneg (i ∨ ~ j) - lemma (neg (suc n)) i = cong predℤ (lemma (neg n) i) - lemma (posneg i) j k = posneg ((i ∧ ~ k) ∨ (i ∧ j) ∨ (j ∧ k)) - diff --git a/Cubical/HITs/Join.agda b/Cubical/HITs/Join.agda index 5b5ac1937..9fce34009 100644 --- a/Cubical/HITs/Join.agda +++ b/Cubical/HITs/Join.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Join where open import Cubical.HITs.Join.Base public diff --git a/Cubical/HITs/Join/Base.agda b/Cubical/HITs/Join/Base.agda index 5040c23d9..d94f9551a 100644 --- a/Cubical/HITs/Join/Base.agda +++ b/Cubical/HITs/Join/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Join.Base where open import Cubical.Foundations.Prelude @@ -76,5 +76,12 @@ joinS¹S¹→S³→joinS¹S¹ (push base (loop k) i) l = push base (loop k) (i joinS¹S¹→S³→joinS¹S¹ (push (loop k) base i) l = facek01 (~ i) k (~ l) joinS¹S¹→S³→joinS¹S¹ (push (loop j) (loop k) i) l = border-contraction i j k (~ l) +S³IsojoinS¹S¹ : Iso S³ (join S¹ S¹) +Iso.fun S³IsojoinS¹S¹ = S³→joinS¹S¹ +Iso.inv S³IsojoinS¹S¹ = joinS¹S¹→S³ +Iso.rightInv S³IsojoinS¹S¹ = joinS¹S¹→S³→joinS¹S¹ +Iso.leftInv S³IsojoinS¹S¹ = S³→joinS¹S¹→S³ + + S³≡joinS¹S¹ : S³ ≡ join S¹ S¹ -S³≡joinS¹S¹ = isoToPath (iso S³→joinS¹S¹ joinS¹S¹→S³ joinS¹S¹→S³→joinS¹S¹ S³→joinS¹S¹→S³) +S³≡joinS¹S¹ = isoToPath S³IsojoinS¹S¹ diff --git a/Cubical/HITs/Join/Properties.agda b/Cubical/HITs/Join/Properties.agda index b21b97afe..fe0efa926 100644 --- a/Cubical/HITs/Join/Properties.agda +++ b/Cubical/HITs/Join/Properties.agda @@ -10,7 +10,7 @@ This file contains: -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Join.Properties where @@ -19,7 +19,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws -open import Cubical.Data.Prod +open import Cubical.Data.Sigma renaming (fst to proj₁; snd to proj₂) open import Cubical.HITs.Join.Base open import Cubical.HITs.Pushout @@ -39,7 +39,7 @@ joinPushout-iso-join A B = iso joinPushout→join join→joinPushout join→join joinPushout→join : joinPushout A B → join A B joinPushout→join (inl x) = inl x joinPushout→join (inr x) = inr x - joinPushout→join (push y i) = push (proj₁ y) (proj₂ y) i + joinPushout→join (push x i) = push (proj₁ x) (proj₂ x) i join→joinPushout : join A B → joinPushout A B join→joinPushout (inl x) = inl x @@ -179,12 +179,12 @@ join-assoc A B C = (joinPushout≡join (join A B) C) ⁻¹ H1 : (x : 3x3-span.A□2 span) → proj₁ (A□2→A×join x) ≡ A□0→A (3x3-span.f□1 span x) H1 (inl (a , b)) = refl H1 (inr (a , c)) = refl - H1 (push (a , (b , c)) i) j = A□0→A (doubleCompPath-filler refl (λ i → push (a , c) i) refl i j) + H1 (push (a , (b , c)) i) j = A□0→A (doubleCompPath-filler refl (λ i → push (a , c) i) refl j i) H2 : (x : 3x3-span.A□2 span) → proj₂ (A□2→A×join x) ≡ fst (joinPushout≃join _ _) (3x3-span.f□3 span x) H2 (inl (a , b)) = refl H2 (inr (a , c)) = refl - H2 (push (a , (b , c)) i) j = fst (joinPushout≃join _ _) (doubleCompPath-filler refl (λ i → push (b , c) i) refl i j) + H2 (push (a , (b , c)) i) j = fst (joinPushout≃join _ _) (doubleCompPath-filler refl (λ i → push (b , c) i) refl j i) -- the second span appearing in 3x3 lemma sp3 : 3-span @@ -261,12 +261,12 @@ join-assoc A B C = (joinPushout≡join (join A B) C) ⁻¹ H3 : (x : 3x3-span.A2□ span) → proj₂ (A2□→join×C x) ≡ A4□→C (3x3-span.f3□ span x) H3 (inl (a , c)) = refl H3 (inr (b , c)) = refl - H3 (push (a , (b , c)) i) j = A4□→C (doubleCompPath-filler refl (λ i → push (a , c) i) refl i j) + H3 (push (a , (b , c)) i) j = A4□→C (doubleCompPath-filler refl (λ i → push (a , c) i) refl j i) H4 : (x : 3x3-span.A2□ span) → proj₁ (A2□→join×C x) ≡ fst (joinPushout≃join _ _) (3x3-span.f1□ span x) H4 (inl (a , c)) = refl H4 (inr (b , c)) = refl - H4 (push (a , (b , c)) i) j = fst (joinPushout≃join _ _) (doubleCompPath-filler refl (λ i → push (a , b) i) refl i j) + H4 (push (a , (b , c)) i) j = fst (joinPushout≃join _ _) (doubleCompPath-filler refl (λ i → push (a , b) i) refl j i) {- Direct proof of an associativity-related property. Combined with @@ -274,17 +274,17 @@ join-assoc A B C = (joinPushout≡join (join A B) C) ⁻¹ -} joinSwitch : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → join (join A B) C ≃ join (join C B) A -joinSwitch = isoToEquiv (iso map map invol invol) +joinSwitch = isoToEquiv (iso switch switch invol invol) where - map : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + switch : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → join (join A B) C → join (join C B) A - map (inl (inl a)) = inr a - map (inl (inr b)) = inl (inr b) - map (inl (push a b i)) = push (inr b) a (~ i) - map (inr c) = inl (inl c) - map (push (inl a) c j) = push (inl c) a (~ j) - map (push (inr b) c j) = inl (push c b (~ j)) - map (push (push a b i) c j) = + switch (inl (inl a)) = inr a + switch (inl (inr b)) = inl (inr b) + switch (inl (push a b i)) = push (inr b) a (~ i) + switch (inr c) = inl (inl c) + switch (push (inl a) c j) = push (inl c) a (~ j) + switch (push (inr b) c j) = inl (push c b (~ j)) + switch (push (push a b i) c j) = hcomp (λ k → λ { (i = i0) → push (inl c) a (~ j ∨ ~ k) @@ -295,7 +295,7 @@ joinSwitch = isoToEquiv (iso map map invol invol) (push (push c b (~ j)) a (~ i)) invol : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} - (u : join (join A B) C) → map (map u) ≡ u + (u : join (join A B) C) → switch (switch u) ≡ u invol (inl (inl a)) = refl invol (inl (inr b)) = refl invol (inl (push a b i)) = refl diff --git a/Cubical/HITs/KleinBottle.agda b/Cubical/HITs/KleinBottle.agda index 5636709cf..bc1b0a0bd 100644 --- a/Cubical/HITs/KleinBottle.agda +++ b/Cubical/HITs/KleinBottle.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.KleinBottle where open import Cubical.HITs.KleinBottle.Base public diff --git a/Cubical/HITs/KleinBottle/Base.agda b/Cubical/HITs/KleinBottle/Base.agda index b4bcc0a90..3e9645c0b 100644 --- a/Cubical/HITs/KleinBottle/Base.agda +++ b/Cubical/HITs/KleinBottle/Base.agda @@ -3,12 +3,12 @@ Definition of the Klein bottle as a HIT -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.KleinBottle.Base where open import Cubical.Core.Everything -data KleinBottle : Set where +data KleinBottle : Type where point : KleinBottle line1 : point ≡ point line2 : point ≡ point diff --git a/Cubical/HITs/KleinBottle/Properties.agda b/Cubical/HITs/KleinBottle/Properties.agda index a878f2165..0bb163bfa 100644 --- a/Cubical/HITs/KleinBottle/Properties.agda +++ b/Cubical/HITs/KleinBottle/Properties.agda @@ -3,7 +3,7 @@ Definition of the Klein bottle as a HIT -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.KleinBottle.Properties where open import Cubical.Core.Everything @@ -15,12 +15,10 @@ open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.HLevels open import Cubical.Foundations.Transport open import Cubical.Foundations.Univalence -open import Cubical.Data.Nat -open import Cubical.Data.Int public renaming (_+_ to _+Int_ ; +-assoc to +Int-assoc; +-comm to +Int-comm) -open import Cubical.Data.Prod +open import Cubical.Data.Int open import Cubical.Data.Sigma open import Cubical.HITs.S1 -open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation as PropTrunc open import Cubical.HITs.KleinBottle.Base @@ -28,16 +26,16 @@ loop1 : S¹ → KleinBottle loop1 base = point loop1 (loop i) = line1 i -invS¹Loop : S¹ → Set +invS¹Loop : S¹ → Type invS¹Loop base = S¹ invS¹Loop (loop i) = invS¹Path i -loop1Inv : (s : S¹) → loop1 (inv s) ≡ loop1 s +loop1Inv : (s : S¹) → loop1 (invLooper s) ≡ loop1 s loop1Inv base = line2 loop1Inv (loop i) = square i -twist : (s : S¹) → PathP (λ i → invS¹Path i) s (inv s) -twist s i = glue (λ {(i = i0) → s; (i = i1) → inv s}) (inv s) +twist : (s : S¹) → PathP (λ i → invS¹Path i) s (invLooper s) +twist s i = glue (λ {(i = i0) → s; (i = i1) → invLooper s}) (invLooper s) twistBaseLoop : (s : S¹) → invS¹Loop s twistBaseLoop base = base @@ -77,7 +75,7 @@ kleinBottle≃Σ = isoToEquiv (iso fro to froTo toFro) froLoop1 (loop i) = refl froLoop1Inv : - PathP (λ k → (s : S¹) → froLoop1 (inv s) k ≡ froLoop1 s k) + PathP (λ k → (s : S¹) → froLoop1 (invLooper s) k ≡ froLoop1 s k) (λ s l → fro (loop1Inv s l)) (λ s l → loop (~ l) , twist s (~ l)) froLoop1Inv k base l = loop (~ l) , twist base (~ l) @@ -109,42 +107,40 @@ isGroupoidKleinBottle = transport (λ i → isGroupoid (ua kleinBottle≃Σ (~ i))) (isOfHLevelΣ 3 isGroupoidS¹ (λ s → - recPropTrunc + PropTrunc.rec (isPropIsOfHLevel 3) (λ p → subst (λ s → isGroupoid (invS¹Loop s)) p isGroupoidS¹) (isConnectedS¹ s))) -- Transport across the following is too slow :( -ΩKlein≡Int² : Path KleinBottle point point ≡ Int × Int -ΩKlein≡Int² = +ΩKlein≡ℤ² : Path KleinBottle point point ≡ ℤ × ℤ +ΩKlein≡ℤ² = Path KleinBottle point point ≡⟨ (λ i → basePath i ≡ basePath i) ⟩ Path (Σ S¹ invS¹Loop) (base , base) (base , base) - ≡⟨ sym (ua Σ≡) ⟩ + ≡⟨ sym ΣPath≡PathΣ ⟩ Σ ΩS¹ (λ p → PathP (λ j → invS¹Loop (p j)) base base) ≡⟨ (λ i → Σ ΩS¹ (λ p → PathP (λ j → invS¹Loop (p (j ∨ i))) (twistBaseLoop (p i)) base)) ⟩ - Σ ΩS¹ (λ _ → ΩS¹) - ≡⟨ sym A×B≡A×ΣB ⟩ ΩS¹ × ΩS¹ - ≡⟨ (λ i → ΩS¹≡Int i × ΩS¹≡Int i) ⟩ - Int × Int ∎ + ≡⟨ (λ i → ΩS¹≡ℤ i × ΩS¹≡ℤ i) ⟩ + ℤ × ℤ ∎ where basePath : PathP (λ i → ua kleinBottle≃Σ i) point (base , base) basePath i = glue (λ {(i = i0) → point; (i = i1) → base , base}) (base , base) -- We can at least define the winding function directly and get results on small examples -windingKlein : Path KleinBottle point point → Int × Int +windingKlein : Path KleinBottle point point → ℤ × ℤ windingKlein p = (z₀ , z₁) where step₀ : Path (Σ S¹ invS¹Loop) (base , base) (base , base) step₀ = (λ i → kleinBottle≃Σ .fst (p i)) - z₀ : Int + z₀ : ℤ z₀ = winding (λ i → kleinBottle≃Σ .fst (p i) .fst) - z₁ : Int + z₁ : ℤ z₁ = winding (transport (λ i → PathP (λ j → invS¹Loop (step₀ (j ∨ i) .fst)) (twistBaseLoop (step₀ i .fst)) base) diff --git a/Cubical/HITs/ListedFiniteSet.agda b/Cubical/HITs/ListedFiniteSet.agda index f2761e5cc..4a1d5e8a5 100644 --- a/Cubical/HITs/ListedFiniteSet.agda +++ b/Cubical/HITs/ListedFiniteSet.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.ListedFiniteSet where open import Cubical.HITs.ListedFiniteSet.Base public diff --git a/Cubical/HITs/ListedFiniteSet/Base.agda b/Cubical/HITs/ListedFiniteSet/Base.agda index ab100b8f4..d1f6e7381 100644 --- a/Cubical/HITs/ListedFiniteSet/Base.agda +++ b/Cubical/HITs/ListedFiniteSet/Base.agda @@ -1,33 +1,36 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.ListedFiniteSet.Base where -open import Cubical.Core.Everything -open import Cubical.Foundations.Logic -open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sum as ⊎ using (_⊎_; inl; inr) + +open import Cubical.Functions.Logic private variable - A : Type₀ + ℓ : Level + A B : Type ℓ infixr 20 _∷_ -infix 30 _∈_ - +-- infix 30 _∈_ +infixr 5 _++_ -data LFSet (A : Type₀) : Type₀ where +data LFSet (A : Type ℓ) : Type ℓ where [] : LFSet A _∷_ : (x : A) → (xs : LFSet A) → LFSet A dup : ∀ x xs → x ∷ x ∷ xs ≡ x ∷ xs comm : ∀ x y xs → x ∷ y ∷ xs ≡ y ∷ x ∷ xs trunc : isSet (LFSet A) - -- Membership. -- -- Doing some proofs with equational reasoning adds an extra "_∙ refl" -- at the end. -- We might want to avoid it, or come up with a more clever equational reasoning. -_∈_ : A → LFSet A → hProp _ -z ∈ [] = ⊥ +_∈_ : {A : Type ℓ} → A → LFSet A → hProp ℓ +z ∈ [] = Lift ⊥.⊥ , isOfHLevelLift 1 isProp⊥ z ∈ (y ∷ xs) = (z ≡ₚ y) ⊔ (z ∈ xs) z ∈ dup x xs i = proof i where @@ -42,11 +45,10 @@ z ∈ comm x y xs i = proof i (z ≡ₚ x ⊔ z ≡ₚ y) ⊔ z ∈ xs ≡⟨ cong (_⊔ (z ∈ xs)) (⊔-comm (z ≡ₚ x) (z ≡ₚ y)) ⟩ (z ≡ₚ y ⊔ z ≡ₚ x) ⊔ z ∈ xs ≡⟨ sym (⊔-assoc (z ≡ₚ y) (z ≡ₚ x) (z ∈ xs)) ⟩ z ≡ₚ y ⊔ (z ≡ₚ x ⊔ z ∈ xs) ∎ - x ∈ trunc xs ys p q i j = isSetHProp (x ∈ xs) (x ∈ ys) (cong (x ∈_) p) (cong (x ∈_) q) i j -module LFSetElim {ℓ} - (B : LFSet A → Type ℓ) +module Elim {ℓ} + {B : LFSet A → Type ℓ} ([]* : B []) (_∷*_ : (x : A) {xs : LFSet A} → B xs → B (x ∷ xs)) (comm* : (x y : A) {xs : LFSet A} (b : B xs) @@ -66,7 +68,7 @@ module LFSetElim {ℓ} (λ i → f (p i)) (λ i → f (q i)) (trunc x y p q) i j -module LFSetRec {ℓ} {B : Type ℓ} +module Rec {ℓ} {B : Type ℓ} ([]* : B) (_∷*_ : (x : A) → B → B) (comm* : (x y : A) (xs : B) → (x ∷* (y ∷* xs)) ≡ (y ∷* (x ∷* xs))) @@ -74,19 +76,40 @@ module LFSetRec {ℓ} {B : Type ℓ} (trunc* : isSet B) where f : LFSet A → B - f = LFSetElim.f _ - []* (λ x xs → x ∷* xs) - (λ x y b → comm* x y b) (λ x b → dup* x b) - λ _ → trunc* + f = + Elim.f + []* (λ x xs → x ∷* xs) + (λ x y b → comm* x y b) (λ x b → dup* x b) + λ _ → trunc* -module LFPropElim {ℓ} - (B : LFSet A → Type ℓ) +module PropElim {ℓ} + {B : LFSet A → Type ℓ} ([]* : B []) (_∷*_ : (x : A) {xs : LFSet A} → B xs → B (x ∷ xs)) (trunc* : (xs : LFSet A) → isProp (B xs)) where f : ∀ x → B x - f = LFSetElim.f _ - []* _∷*_ - (λ _ _ _ → isOfHLevel→isOfHLevelDep 1 trunc* _ _ _) - (λ _ _ → isOfHLevel→isOfHLevelDep 1 trunc* _ _ _) - λ xs → isProp→isSet (trunc* xs) + f = + Elim.f + []* _∷*_ + (λ _ _ _ → isOfHLevel→isOfHLevelDep 1 trunc* _ _ _) + (λ _ _ → isOfHLevel→isOfHLevelDep 1 trunc* _ _ _) + λ xs → isProp→isSet (trunc* xs) + +_++_ : ∀ (xs ys : LFSet A) → LFSet A +[] ++ ys = ys +(x ∷ xs) ++ ys = x ∷ (xs ++ ys) +dup x xs i ++ ys = dup x (xs ++ ys) i +comm x y xs i ++ ys = comm x y (xs ++ ys) i +trunc xs zs p q i j ++ ys = + trunc (xs ++ ys) (zs ++ ys) (cong (_++ ys) p) (cong (_++ ys) q) i j + +map : (A → B) → LFSet A → LFSet B +map f [] = [] +map f (x ∷ xs) = f x ∷ map f xs +map f (dup x xs i) = dup (f x) (map f xs) i +map f (comm x y xs i) = comm (f x) (f y) (map f xs) i +map f (trunc xs ys p q i j) = + trunc (map f xs) (map f ys) (cong (map f) p) (cong (map f) q) i j + +disj-union : LFSet A → LFSet B → LFSet (A ⊎ B) +disj-union xs ys = map ⊎.inl xs ++ map ⊎.inr ys diff --git a/Cubical/HITs/ListedFiniteSet/Properties.agda b/Cubical/HITs/ListedFiniteSet/Properties.agda index 313d81160..7bee9f313 100644 --- a/Cubical/HITs/ListedFiniteSet/Properties.agda +++ b/Cubical/HITs/ListedFiniteSet/Properties.agda @@ -1,41 +1,106 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.ListedFiniteSet.Properties where -open import Cubical.Core.Everything -open import Cubical.Foundations.Logic -open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Data.Prod using (_×_; _,_) open import Cubical.HITs.ListedFiniteSet.Base private variable - A : Type₀ - -_++_ : ∀ (xs ys : LFSet A) → LFSet A -[] ++ ys = ys -(x ∷ xs) ++ ys = x ∷ (xs ++ ys) ---------------------------------------------- -dup x xs i ++ ys = proof i - where - proof : - -- Need: (x ∷ x ∷ xs) ++ ys ≡ (x ∷ xs) ++ ys - -- which reduces to: - x ∷ x ∷ (xs ++ ys) ≡ x ∷ (xs ++ ys) - proof = dup x (xs ++ ys) -comm x y xs i ++ ys = comm x y (xs ++ ys) i -trunc xs zs p q i j ++ ys - = trunc (xs ++ ys) (zs ++ ys) (cong (_++ ys) p) (cong (_++ ys) q) i j - + ℓ : Level + A B : Type ℓ assoc-++ : ∀ (xs : LFSet A) ys zs → xs ++ (ys ++ zs) ≡ (xs ++ ys) ++ zs assoc-++ [] ys zs = refl -assoc-++ (x ∷ xs) ys zs - = cong (x ∷_) (assoc-++ xs ys zs) ------------------------------------- -assoc-++ (dup x xs i) ys zs j - = dup x (assoc-++ xs ys zs j) i -assoc-++ (comm x y xs i) ys zs j - = comm x y (assoc-++ xs ys zs j) i -assoc-++ (trunc xs xs' p q i k) ys zs j - = trunc (assoc-++ xs ys zs j) (assoc-++ xs' ys zs j) - (cong (\ xs -> assoc-++ xs ys zs j) p) (cong (\ xs -> assoc-++ xs ys zs j) q) i k +assoc-++ (x ∷ xs) ys zs = cong (x ∷_) (assoc-++ xs ys zs) +assoc-++ (dup x xs i) ys zs j = dup x (assoc-++ xs ys zs j) i +assoc-++ (comm x y xs i) ys zs j = comm x y (assoc-++ xs ys zs j) i +assoc-++ (trunc xs xs' p q i k) ys zs j = + trunc + (assoc-++ xs ys zs j) (assoc-++ xs' ys zs j) + (cong (λ xs → assoc-++ xs ys zs j) p) (cong (λ xs → assoc-++ xs ys zs j) q) + i k + +comm-++-[] : ∀ (xs : LFSet A) → xs ++ [] ≡ [] ++ xs +comm-++-[] xs = + PropElim.f + refl + (λ x {xs} ind → + (x ∷ xs) ++ [] ≡⟨ refl ⟩ + x ∷ (xs ++ []) ≡⟨ cong (x ∷_) ind ⟩ + x ∷ xs ≡⟨ refl ⟩ + [] ++ (x ∷ xs) ∎ + ) + (λ _ → trunc _ _) + xs + +comm-++-∷ + : ∀ (z : A) xs ys + → xs ++ (z ∷ ys) ≡ (z ∷ xs) ++ ys +comm-++-∷ z xs ys = + PropElim.f + refl + (λ x {xs} ind → + x ∷ (xs ++ (z ∷ ys)) ≡⟨ cong (x ∷_) ind ⟩ + x ∷ z ∷ (xs ++ ys) ≡⟨ comm x z (xs ++ ys) ⟩ + z ∷ x ∷ (xs ++ ys) ∎ + ) + (λ _ → trunc _ _) + xs + +comm-++ : (xs ys : LFSet A) → xs ++ ys ≡ ys ++ xs +comm-++ xs ys = + PropElim.f + (comm-++-[] xs) + (λ y {ys} ind → + xs ++ (y ∷ ys) ≡⟨ comm-++-∷ y xs ys ⟩ + y ∷ (xs ++ ys) ≡⟨ cong (y ∷_) ind ⟩ + y ∷ (ys ++ xs) ≡⟨ refl ⟩ + (y ∷ ys) ++ xs ∎ + ) + (λ _ → trunc _ _) + ys + +idem-++ : (xs : LFSet A) → xs ++ xs ≡ xs +idem-++ = + PropElim.f + refl + (λ x {xs} ind → + (x ∷ xs) ++ (x ∷ xs) ≡⟨ refl ⟩ + x ∷ (xs ++ (x ∷ xs)) ≡⟨ refl ⟩ + x ∷ (xs ++ ((x ∷ []) ++ xs)) ≡⟨ cong (x ∷_) (assoc-++ xs (x ∷ []) xs) ⟩ + x ∷ ((xs ++ (x ∷ [])) ++ xs) + ≡⟨ cong (λ h → x ∷ (h ++ xs)) (comm-++ xs (x ∷ [])) ⟩ + x ∷ x ∷ (xs ++ xs) ≡⟨ cong (λ ys → x ∷ x ∷ ys) ind ⟩ + x ∷ x ∷ xs ≡⟨ dup x xs ⟩ + x ∷ xs ∎ + ) + (λ xs → trunc (xs ++ xs) xs) + +cart-product : LFSet A → LFSet B → LFSet (A × B) +cart-product [] ys = [] +cart-product (x ∷ xs) ys = map (x ,_) ys ++ cart-product xs ys +cart-product (dup x xs i) ys = + ( map (x ,_) ys ++ (map (x ,_) ys ++ cart-product xs ys) + ≡⟨ assoc-++ (map (x ,_) ys) (map (x ,_) ys) (cart-product xs ys) ⟩ + (map (x ,_) ys ++ map (x ,_) ys) ++ cart-product xs ys + ≡⟨ cong (_++ cart-product xs ys) (idem-++ (map (x ,_) ys)) ⟩ + map (x ,_) ys ++ cart-product xs ys + ∎ + ) i +cart-product (comm x y xs i) ys = + ( map (x ,_) ys ++ (map (y ,_) ys ++ cart-product xs ys) + ≡⟨ assoc-++ (map (x ,_) ys) (map (y ,_) ys) (cart-product xs ys) ⟩ + (map (x ,_) ys ++ map (y ,_) ys) ++ cart-product xs ys + ≡⟨ cong (_++ cart-product xs ys) (comm-++ (map (x ,_) ys) (map (y ,_) ys)) ⟩ + (map (y ,_) ys ++ map (x ,_) ys) ++ cart-product xs ys + ≡⟨ sym (assoc-++ (map (y ,_) ys) (map (x ,_) ys) (cart-product xs ys)) ⟩ + map (y ,_) ys ++ (map (x ,_) ys ++ cart-product xs ys) + ∎ + ) i +cart-product (trunc xs xs′ p q i j) ys = + trunc + (cart-product xs ys) (cart-product xs′ ys) + (λ k → cart-product (p k) ys) (λ k → cart-product (q k) ys) + i j diff --git a/Cubical/HITs/Localization.agda b/Cubical/HITs/Localization.agda index c385e78ed..e8bd1f017 100644 --- a/Cubical/HITs/Localization.agda +++ b/Cubical/HITs/Localization.agda @@ -1,6 +1,6 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Localization where open import Cubical.HITs.Localization.Base public -open import Cubical.HITs.Localization.Properties public renaming (rec to Localize-rec) +open import Cubical.HITs.Localization.Properties public diff --git a/Cubical/HITs/Localization/Base.agda b/Cubical/HITs/Localization/Base.agda index dbf08134f..d6c378d72 100644 --- a/Cubical/HITs/Localization/Base.agda +++ b/Cubical/HITs/Localization/Base.agda @@ -1,9 +1,9 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Localization.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function -open import Cubical.Foundations.PathSplitEquiv +open import Cubical.Foundations.Equiv.PathSplit open isPathSplitEquiv module _ {ℓα ℓs ℓt} {A : Type ℓα} {S : A → Type ℓs} {T : A → Type ℓt} where diff --git a/Cubical/HITs/Localization/Properties.agda b/Cubical/HITs/Localization/Properties.agda index c9be10f93..0a36c9429 100644 --- a/Cubical/HITs/Localization/Properties.agda +++ b/Cubical/HITs/Localization/Properties.agda @@ -1,11 +1,11 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Localization.Properties where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.PathSplitEquiv +open import Cubical.Foundations.Equiv.PathSplit open isPathSplitEquiv open import Cubical.HITs.Localization.Base diff --git a/Cubical/HITs/MappingCones.agda b/Cubical/HITs/MappingCones.agda index b8b35a743..29543c1de 100644 --- a/Cubical/HITs/MappingCones.agda +++ b/Cubical/HITs/MappingCones.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.MappingCones where diff --git a/Cubical/HITs/MappingCones/Base.agda b/Cubical/HITs/MappingCones/Base.agda index d22ccd02f..0c533ea3e 100644 --- a/Cubical/HITs/MappingCones/Base.agda +++ b/Cubical/HITs/MappingCones/Base.agda @@ -3,7 +3,7 @@ Mapping cones or the homotopy cofiber/cokernel -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.MappingCones.Base where diff --git a/Cubical/HITs/MappingCones/Properties.agda b/Cubical/HITs/MappingCones/Properties.agda index 76d9dcc0b..3c344b2fd 100644 --- a/Cubical/HITs/MappingCones/Properties.agda +++ b/Cubical/HITs/MappingCones/Properties.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.MappingCones.Properties where diff --git a/Cubical/HITs/Modulo.agda b/Cubical/HITs/Modulo.agda index f23d60661..50dfa2e9b 100644 --- a/Cubical/HITs/Modulo.agda +++ b/Cubical/HITs/Modulo.agda @@ -1,6 +1,8 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Modulo where open import Cubical.HITs.Modulo.Base public open import Cubical.HITs.Modulo.Properties public + +open import Cubical.HITs.Modulo.FinEquiv public diff --git a/Cubical/HITs/Modulo/Base.agda b/Cubical/HITs/Modulo/Base.agda index c7005a407..ef6172280 100644 --- a/Cubical/HITs/Modulo/Base.agda +++ b/Cubical/HITs/Modulo/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Modulo.Base where @@ -7,9 +7,8 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Prelude -open import Cubical.Data.Empty -open import Cubical.Data.Fin -open import Cubical.Data.Nat +open import Cubical.Data.Empty using (⊥) +open import Cubical.Data.Nat using (ℕ; zero; suc; _+_) open import Cubical.Data.Unit renaming (Unit to ⊤) open import Cubical.Relation.Nullary diff --git a/Cubical/HITs/Modulo/FinEquiv.agda b/Cubical/HITs/Modulo/FinEquiv.agda index 5c83075e1..e7ee0bb7d 100644 --- a/Cubical/HITs/Modulo/FinEquiv.agda +++ b/Cubical/HITs/Modulo/FinEquiv.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Modulo.FinEquiv where diff --git a/Cubical/HITs/Modulo/Properties.agda b/Cubical/HITs/Modulo/Properties.agda index 7086dbf1e..70c5c4588 100644 --- a/Cubical/HITs/Modulo/Properties.agda +++ b/Cubical/HITs/Modulo/Properties.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Modulo.Properties where diff --git a/Cubical/HITs/Nullification.agda b/Cubical/HITs/Nullification.agda index 9b0875139..5a3c2be8b 100644 --- a/Cubical/HITs/Nullification.agda +++ b/Cubical/HITs/Nullification.agda @@ -1,6 +1,6 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Nullification where open import Cubical.HITs.Nullification.Base public -open import Cubical.HITs.Nullification.Properties public renaming (rec to Null-rec; ind to Null-ind) +open import Cubical.HITs.Nullification.Properties public diff --git a/Cubical/HITs/Nullification/Base.agda b/Cubical/HITs/Nullification/Base.agda index a035eb088..822dd48a5 100644 --- a/Cubical/HITs/Nullification/Base.agda +++ b/Cubical/HITs/Nullification/Base.agda @@ -1,9 +1,9 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Nullification.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function -open import Cubical.Foundations.PathSplitEquiv +open import Cubical.Foundations.Equiv.PathSplit open isPathSplitEquiv isNull : ∀ {ℓ ℓ'} (S : Type ℓ) (A : Type ℓ') → Type (ℓ-max ℓ ℓ') diff --git a/Cubical/HITs/Nullification/Properties.agda b/Cubical/HITs/Nullification/Properties.agda index e1c647d30..b52a35a20 100644 --- a/Cubical/HITs/Nullification/Properties.agda +++ b/Cubical/HITs/Nullification/Properties.agda @@ -1,15 +1,26 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Nullification.Properties where -open import Cubical.Foundations.Everything -open isPathSplitEquiv -open import Cubical.Modalities.Everything -open Modality -open import Cubical.HITs.Localization +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.PathSplit +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Univalence + +open import Cubical.Modalities.Modality + +open import Cubical.Functions.FunExtEquiv +open import Cubical.HITs.Localization renaming (rec to Localize-rec) open import Cubical.Data.Unit open import Cubical.HITs.Nullification.Base +open Modality +open isPathSplitEquiv + rec : ∀ {ℓ ℓ' ℓ''} {S : Type ℓ} {A : Type ℓ'} {B : Type ℓ''} → (nB : isNull S B) → (A → B) → Null S A → B rec nB g ∣ x ∣ = g x @@ -34,52 +45,52 @@ module _ {ℓ ℓ' ℓ''} {S : Type ℓ} {A : Type ℓ'} {B : Null S A → Type → (∀ (bx : B x) (by : B y) → hasSection (cong₂ (λ x (b : B x) (_ : S) → b) p)) secCongDep' nB p = secCongDep (λ _ → const) p (λ a → secCong (nB a)) - ind : (nB : (x : Null S A) → isNull S (B x)) → ((a : A) → B ∣ a ∣) → (x : Null S A) → B x - ind nB g ∣ x ∣ = g x - ind nB g (hub f) + elim : (nB : (x : Null S A) → isNull S (B x)) → ((a : A) → B ∣ a ∣) → (x : Null S A) → B x + elim nB g ∣ x ∣ = g x + elim nB g (hub f) = fst (sec (nB (hub f))) - (λ s → transport (λ i → B (spoke f s (~ i))) (ind nB g (f s))) + (λ s → transport (λ i → B (spoke f s (~ i))) (elim nB g (f s))) - ind nB g (spoke f s i) + elim nB g (spoke f s i) = toPathP⁻ (λ i → B (spoke f s i)) (funExt⁻ ( snd (sec (nB (hub f))) - (λ s → transport (λ i → B (spoke f s (~ i))) (ind nB g (f s))) ) s) i - - ind nB g (≡hub {x} {y} p i) - = hcomp (λ k → λ { (i = i0) → transportRefl (ind nB g x) k - ; (i = i1) → transportRefl (ind nB g y) k }) - (fst (secCongDep' nB (≡hub p) (transport refl (ind nB g x)) (transport refl (ind nB g y))) - (λ i s → transport (λ j → B (≡spoke p s (~ j) i)) (ind nB g (p s i))) i) - - ind nB g (≡spoke {x} {y} p s j i) - = hcomp (λ k → λ { (i = i0) → toPathP⁻-sq (ind nB g x) k j - ; (i = i1) → toPathP⁻-sq (ind nB g y) k j - ; (j = i1) → ind nB g (p s i) }) + (λ s → transport (λ i → B (spoke f s (~ i))) (elim nB g (f s))) ) s) i + + elim nB g (≡hub {x} {y} p i) + = hcomp (λ k → λ { (i = i0) → transportRefl (elim nB g x) k + ; (i = i1) → transportRefl (elim nB g y) k }) + (fst (secCongDep' nB (≡hub p) (transport refl (elim nB g x)) (transport refl (elim nB g y))) + (λ i s → transport (λ j → B (≡spoke p s (~ j) i)) (elim nB g (p s i))) i) + + elim nB g (≡spoke {x} {y} p s j i) + = hcomp (λ k → λ { (i = i0) → toPathP⁻-sq (elim nB g x) k j + ; (i = i1) → toPathP⁻-sq (elim nB g y) k j + ; (j = i1) → elim nB g (p s i) }) (q₂ j i) - where q₁ : Path (PathP (λ i → B (≡hub p i)) (transport refl (ind nB g x)) (transport refl (ind nB g y))) - (fst (secCongDep' nB (≡hub p) (transport refl (ind nB g x)) (transport refl (ind nB g y))) - (λ i s → transport (λ j → B (≡spoke p s (~ j) i)) (ind nB g (p s i)))) - (λ i → transport (λ j → B (≡spoke p s (~ j) i)) (ind nB g (p s i))) - q₁ j i = snd (secCongDep' nB (≡hub p) (transport refl (ind nB g x)) (transport refl (ind nB g y))) - (λ i s → transport (λ j → B (≡spoke p s (~ j) i)) (ind nB g (p s i))) j i s - - q₂ : PathP (λ j → PathP (λ i → B (≡spoke p s j i)) (toPathP⁻ (λ _ → B x) (λ _ → transport refl (ind nB g x)) j) - (toPathP⁻ (λ _ → B y) (λ _ → transport refl (ind nB g y)) j)) - (fst (secCongDep' nB (≡hub p) (transport refl (ind nB g x)) (transport refl (ind nB g y))) - (λ i s → transport (λ j → B (≡spoke p s (~ j) i)) (ind nB g (p s i)))) - (λ i → ind nB g (p s i)) + where q₁ : Path (PathP (λ i → B (≡hub p i)) (transport refl (elim nB g x)) (transport refl (elim nB g y))) + (fst (secCongDep' nB (≡hub p) (transport refl (elim nB g x)) (transport refl (elim nB g y))) + (λ i s → transport (λ j → B (≡spoke p s (~ j) i)) (elim nB g (p s i)))) + (λ i → transport (λ j → B (≡spoke p s (~ j) i)) (elim nB g (p s i))) + q₁ j i = snd (secCongDep' nB (≡hub p) (transport refl (elim nB g x)) (transport refl (elim nB g y))) + (λ i s → transport (λ j → B (≡spoke p s (~ j) i)) (elim nB g (p s i))) j i s + + q₂ : PathP (λ j → PathP (λ i → B (≡spoke p s j i)) (toPathP⁻ (λ _ → B x) (λ _ → transport refl (elim nB g x)) j) + (toPathP⁻ (λ _ → B y) (λ _ → transport refl (elim nB g y)) j)) + (fst (secCongDep' nB (≡hub p) (transport refl (elim nB g x)) (transport refl (elim nB g y))) + (λ i s → transport (λ j → B (≡spoke p s (~ j) i)) (elim nB g (p s i)))) + (λ i → elim nB g (p s i)) q₂ j i = toPathP⁻ (λ j → B (≡spoke p s j i)) (λ j → q₁ j i) j -- nullification is a modality NullModality : ∀ {ℓ} (S : Type ℓ) → Modality ℓ isModal (NullModality S) = isNull S -isModalIsProp (NullModality S) = isPropIsPathSplitEquiv _ +isPropIsModal (NullModality S) = isPropIsPathSplitEquiv _ ◯ (NullModality S) = Null S ◯-isModal (NullModality S) = isNull-Null η (NullModality S) = ∣_∣ -◯-elim (NullModality S) = ind +◯-elim (NullModality S) = elim ◯-elim-β (NullModality S) = λ _ _ _ → refl ◯-=-isModal (NullModality S) x y = fromIsEquiv _ e where e : isEquiv (λ (p : x ≡ y) → const {B = S} p) diff --git a/Cubical/HITs/PropositionalTruncation.agda b/Cubical/HITs/PropositionalTruncation.agda index 9fcac4a51..74160c678 100644 --- a/Cubical/HITs/PropositionalTruncation.agda +++ b/Cubical/HITs/PropositionalTruncation.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.PropositionalTruncation where open import Cubical.HITs.PropositionalTruncation.Base public diff --git a/Cubical/HITs/PropositionalTruncation/Base.agda b/Cubical/HITs/PropositionalTruncation/Base.agda index e5802ab1c..a453e7a16 100644 --- a/Cubical/HITs/PropositionalTruncation/Base.agda +++ b/Cubical/HITs/PropositionalTruncation/Base.agda @@ -5,7 +5,7 @@ This file contains: - Definition of propositional truncation -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.PropositionalTruncation.Base where open import Cubical.Core.Primitives diff --git a/Cubical/HITs/PropositionalTruncation/MagicTrick.agda b/Cubical/HITs/PropositionalTruncation/MagicTrick.agda index 00be68b22..e663da1d4 100644 --- a/Cubical/HITs/PropositionalTruncation/MagicTrick.agda +++ b/Cubical/HITs/PropositionalTruncation/MagicTrick.agda @@ -11,7 +11,7 @@ Also see the follow-up post by Jason Gross: https://homotopytypetheory.org/2014/02/24/composition-is-not-what-you-think-it-is-why-nearly-invertible-isnt/ -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.PropositionalTruncation.MagicTrick where @@ -30,8 +30,7 @@ module Recover {ℓ} (A∙ : Pointed ℓ) (h : isHomogeneous A∙) where a = pt A∙ toEquivPtd : ∥ A ∥ → Σ[ B∙ ∈ Pointed ℓ ] (A , a) ≡ B∙ - toEquivPtd = recPropTrunc (isContr→isProp (_ , λ p → contrSingl (snd p))) - (λ x → (A , x) , h x) + toEquivPtd = rec isPropSingl (λ x → (A , x) , h x) private B∙ : ∥ A ∥ → Pointed ℓ B∙ tx = fst (toEquivPtd tx) diff --git a/Cubical/HITs/PropositionalTruncation/Monad.agda b/Cubical/HITs/PropositionalTruncation/Monad.agda new file mode 100644 index 000000000..7ecb829a5 --- /dev/null +++ b/Cubical/HITs/PropositionalTruncation/Monad.agda @@ -0,0 +1,33 @@ +{-# OPTIONS --safe #-} +{- +Implements the monadic interface of propositional truncation, for reasoning in do-syntax. +-} + +module Cubical.HITs.PropositionalTruncation.Monad where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Functions.Logic +open import Cubical.HITs.PropositionalTruncation + +private + variable + ℓ : Level + P Q : Type ℓ + +infix 1 proof_by_ +proof_by_ : (P : hProp ℓ) → ∥ ⟨ P ⟩ ∥ → ⟨ P ⟩ +proof P by p = rec (isProp⟨⟩ P) (λ p → p) p + +return : P → ∥ P ∥ +return p = ∣ p ∣ + +exact_ : ∥ P ∥ → ∥ P ∥ +exact p = p + +_>>=_ : ∥ P ∥ → (P → ∥ Q ∥) → ∥ Q ∥ +p >>= f = rec isPropPropTrunc f p + +_>>_ : ∥ P ∥ → ∥ Q ∥ → ∥ Q ∥ +_ >> q = q diff --git a/Cubical/HITs/PropositionalTruncation/Properties.agda b/Cubical/HITs/PropositionalTruncation/Properties.agda index 259007a3c..a1655f385 100644 --- a/Cubical/HITs/PropositionalTruncation/Properties.agda +++ b/Cubical/HITs/PropositionalTruncation/Properties.agda @@ -5,7 +5,7 @@ This file contains: - Eliminator for propositional truncation -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.PropositionalTruncation.Properties where open import Cubical.Core.Everything @@ -15,45 +15,83 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Sigma +open import Cubical.Data.Sum hiding (rec ; elim ; map) open import Cubical.HITs.PropositionalTruncation.Base private variable - ℓ : Level - A B : Type ℓ - -recPropTrunc : ∀ {P : Type ℓ} → isProp P → (A → P) → ∥ A ∥ → P -recPropTrunc Pprop f ∣ x ∣ = f x -recPropTrunc Pprop f (squash x y i) = - Pprop (recPropTrunc Pprop f x) (recPropTrunc Pprop f y) i - -propTruncIsProp : isProp ∥ A ∥ -propTruncIsProp x y = squash x y - -elimPropTrunc : ∀ {P : ∥ A ∥ → Type ℓ} → ((a : ∥ A ∥) → isProp (P a)) → - ((x : A) → P ∣ x ∣) → (a : ∥ A ∥) → P a -elimPropTrunc Pprop f ∣ x ∣ = f x -elimPropTrunc {A = A} {P = P} Pprop f (squash x y i) = - PpropOver (squash x y) (elimPropTrunc Pprop f x) (elimPropTrunc Pprop f y) i - where - PpropOver : {a b : ∥ A ∥} → (sq : a ≡ b) → ∀ x y → PathP (λ i → P (sq i)) x y - PpropOver {a} = J (λ b (sq : a ≡ b) → ∀ x y → PathP (λ i → P (sq i)) x y) (Pprop a) - + ℓ ℓ′ : Level + A B C : Type ℓ + A′ : Type ℓ′ + +∥∥-isPropDep : (P : A → Type ℓ) → isOfHLevelDep 1 (λ x → ∥ P x ∥) +∥∥-isPropDep P = isOfHLevel→isOfHLevelDep 1 (λ _ → squash) + +rec : {P : Type ℓ} → isProp P → (A → P) → ∥ A ∥ → P +rec Pprop f ∣ x ∣ = f x +rec Pprop f (squash x y i) = Pprop (rec Pprop f x) (rec Pprop f y) i + +rec2 : {P : Type ℓ} → isProp P → (A → B → P) → ∥ A ∥ → ∥ B ∥ → P +rec2 Pprop f ∣ x ∣ ∣ y ∣ = f x y +rec2 Pprop f ∣ x ∣ (squash y z i) = Pprop (rec2 Pprop f ∣ x ∣ y) (rec2 Pprop f ∣ x ∣ z) i +rec2 Pprop f (squash x y i) z = Pprop (rec2 Pprop f x z) (rec2 Pprop f y z) i + +-- Old version +-- rec2 : ∀ {P : Type ℓ} → isProp P → (A → A → P) → ∥ A ∥ → ∥ A ∥ → P +-- rec2 Pprop f = rec (isProp→ Pprop) (λ a → rec Pprop (f a)) + +elim : {P : ∥ A ∥ → Type ℓ} → ((a : ∥ A ∥) → isProp (P a)) + → ((x : A) → P ∣ x ∣) → (a : ∥ A ∥) → P a +elim Pprop f ∣ x ∣ = f x +elim Pprop f (squash x y i) = + isOfHLevel→isOfHLevelDep 1 Pprop + (elim Pprop f x) (elim Pprop f y) (squash x y) i + +elim2 : {P : ∥ A ∥ → ∥ B ∥ → Type ℓ} + (Pprop : (x : ∥ A ∥) (y : ∥ B ∥) → isProp (P x y)) + (f : (a : A) (b : B) → P ∣ a ∣ ∣ b ∣) + (x : ∥ A ∥) (y : ∥ B ∥) → P x y +elim2 Pprop f = + elim (λ _ → isPropΠ (λ _ → Pprop _ _)) + (λ a → elim (λ _ → Pprop _ _) (f a)) + +elim3 : {P : ∥ A ∥ → ∥ B ∥ → ∥ C ∥ → Type ℓ} + (Pprop : ((x : ∥ A ∥) (y : ∥ B ∥) (z : ∥ C ∥) → isProp (P x y z))) + (g : (a : A) (b : B) (c : C) → P (∣ a ∣) ∣ b ∣ ∣ c ∣) + (x : ∥ A ∥) (y : ∥ B ∥) (z : ∥ C ∥) → P x y z +elim3 Pprop g = elim2 (λ _ _ → isPropΠ (λ _ → Pprop _ _ _)) + (λ a b → elim (λ _ → Pprop _ _ _) (g a b)) + +isPropPropTrunc : isProp ∥ A ∥ +isPropPropTrunc x y = squash x y + +propTruncIdempotent≃ : isProp A → ∥ A ∥ ≃ A +propTruncIdempotent≃ {A = A} hA = isoToEquiv f + where + f : Iso ∥ A ∥ A + Iso.fun f = rec hA (idfun A) + Iso.inv f x = ∣ x ∣ + Iso.rightInv f _ = refl + Iso.leftInv f = elim (λ _ → isProp→isSet isPropPropTrunc _ _) (λ _ → refl) -propId : isProp A → ∥ A ∥ ≡ A -propId {A = A} hA = isoToPath (iso (recPropTrunc hA (idfun A)) (λ x → ∣ x ∣) (λ _ → refl) rinv) - where - rinv : ∀ (x : ∥ A ∥) → ∣ recPropTrunc hA (idfun A) x ∣ ≡ x - rinv x = elimPropTrunc {P = λ x → ∣ recPropTrunc hA (idfun A) x ∣ ≡ x} - (λ _ → isProp→isSet propTruncIsProp _ _) - (λ _ → refl) x +propTruncIdempotent : isProp A → ∥ A ∥ ≡ A +propTruncIdempotent hA = ua (propTruncIdempotent≃ hA) -- We could also define the eliminator using the recursor -elimPropTrunc' : ∀ {P : ∥ A ∥ → Type ℓ} → ((a : ∥ A ∥) → isProp (P a)) → - ((x : A) → P ∣ x ∣) → (a : ∥ A ∥) → P a -elimPropTrunc' {P = P} Pprop f a = - recPropTrunc (Pprop a) (λ x → transp (λ i → P (squash ∣ x ∣ a i)) i0 (f x)) a +elim' : {P : ∥ A ∥ → Type ℓ} → ((a : ∥ A ∥) → isProp (P a)) → + ((x : A) → P ∣ x ∣) → (a : ∥ A ∥) → P a +elim' {P = P} Pprop f a = + rec (Pprop a) (λ x → transp (λ i → P (squash ∣ x ∣ a i)) i0 (f x)) a + +map : (A → B) → (∥ A ∥ → ∥ B ∥) +map f = rec squash (∣_∣ ∘ f) + +map2 : (A → B → C) → (∥ A ∥ → ∥ B ∥ → ∥ C ∥) +map2 f = rec (isPropΠ λ _ → squash) (map ∘ f) -- The propositional truncation can be eliminated into non-propositional -- types as long as the function used in the eliminator is 'coherently @@ -64,13 +102,12 @@ module SetElim (Bset : isSet B) where Bset' : isSet' B Bset' = isSet→isSet' Bset - recPropTrunc→Set : (f : A → B) (kf : 2-Constant f) → ∥ A ∥ → B - helper - : (f : A → B) (kf : 2-Constant f) → (t u : ∥ A ∥) - → recPropTrunc→Set f kf t ≡ recPropTrunc→Set f kf u + rec→Set : (f : A → B) (kf : 2-Constant f) → ∥ A ∥ → B + helper : (f : A → B) (kf : 2-Constant f) → (t u : ∥ A ∥) + → rec→Set f kf t ≡ rec→Set f kf u - recPropTrunc→Set f kf ∣ x ∣ = f x - recPropTrunc→Set f kf (squash t u i) = helper f kf t u i + rec→Set f kf ∣ x ∣ = f x + rec→Set f kf (squash t u i) = helper f kf t u i helper f kf ∣ x ∣ ∣ y ∣ = kf x y helper f kf (squash t u i) v @@ -78,33 +115,32 @@ module SetElim (Bset : isSet B) where helper f kf t (squash u v i) = Bset' (helper f kf t u) (helper f kf t v) refl (helper f kf u v) i - kcomp : ∀(f : ∥ A ∥ → B) → 2-Constant (f ∘ ∣_∣) + kcomp : (f : ∥ A ∥ → B) → 2-Constant (f ∘ ∣_∣) kcomp f x y = cong f (squash ∣ x ∣ ∣ y ∣) Fset : isSet (A → B) - Fset = isSetPi (const Bset) + Fset = isSetΠ (const Bset) Kset : (f : A → B) → isSet (2-Constant f) - Kset f = isSetPi (λ _ → isSetPi (λ _ → isProp→isSet (Bset _ _))) + Kset f = isSetΠ (λ _ → isSetΠ (λ _ → isProp→isSet (Bset _ _))) setRecLemma : (f : ∥ A ∥ → B) - → recPropTrunc→Set (f ∘ ∣_∣) (kcomp f) ≡ f + → rec→Set (f ∘ ∣_∣) (kcomp f) ≡ f setRecLemma f i t - = elimPropTrunc {P = λ t → recPropTrunc→Set (f ∘ ∣_∣) (kcomp f) t ≡ f t} + = elim {P = λ t → rec→Set (f ∘ ∣_∣) (kcomp f) t ≡ f t} (λ t → Bset _ _) (λ x → refl) t i mkKmap : (∥ A ∥ → B) → Σ (A → B) 2-Constant mkKmap f = f ∘ ∣_∣ , kcomp f fib : (g : Σ (A → B) 2-Constant) → fiber mkKmap g - fib (g , kg) = recPropTrunc→Set g kg , refl + fib (g , kg) = rec→Set g kg , refl - eqv : (g : Σ (A → B) 2-Constant) - → ∀ fi → fib g ≡ fi + eqv : (g : Σ (A → B) 2-Constant) → ∀ fi → fib g ≡ fi eqv g (f , p) = - ΣProp≡ (λ f → isOfHLevelΣ 2 Fset Kset _ _) - (cong (uncurry recPropTrunc→Set) (sym p) ∙ setRecLemma f) + Σ≡Prop (λ f → isOfHLevelΣ 2 Fset Kset _ _) + (cong (uncurry rec→Set) (sym p) ∙ setRecLemma f) trunc→Set≃ : (∥ A ∥ → B) ≃ (Σ (A → B) 2-Constant) trunc→Set≃ .fst = mkKmap @@ -127,7 +163,7 @@ module SetElim (Bset : isSet B) where e-isEquiv a₀ = isoToIsEquiv (iso e (eval a₀) (e-eval a₀) λ _ → refl) preEquiv₁ : ∥ A ∥ → B ≃ Σ (A → B) 2-Constant - preEquiv₁ t = e , recPropTrunc (isPropIsEquiv e) e-isEquiv t + preEquiv₁ t = e , rec (isPropIsEquiv e) e-isEquiv t preEquiv₂ : (∥ A ∥ → Σ (A → B) 2-Constant) ≃ Σ (A → B) 2-Constant preEquiv₂ = isoToEquiv (iso to const (λ _ → refl) retr) @@ -147,18 +183,18 @@ module SetElim (Bset : isSet B) where i trunc→Set≃₂ : (∥ A ∥ → B) ≃ Σ (A → B) 2-Constant - trunc→Set≃₂ = compEquiv (equivPi preEquiv₁) preEquiv₂ + trunc→Set≃₂ = compEquiv (equivΠCod preEquiv₁) preEquiv₂ -open SetElim public using (recPropTrunc→Set; trunc→Set≃) +open SetElim public using (rec→Set; trunc→Set≃) -elimPropTrunc→Set - : {P : ∥ A ∥ → Set ℓ} +elim→Set + : {P : ∥ A ∥ → Type ℓ} → (∀ t → isSet (P t)) → (f : (x : A) → P ∣ x ∣) → (kf : ∀ x y → PathP (λ i → P (squash ∣ x ∣ ∣ y ∣ i)) (f x) (f y)) → (t : ∥ A ∥) → P t -elimPropTrunc→Set {A = A} {P = P} Pset f kf t - = recPropTrunc→Set (Pset t) g gk t +elim→Set {A = A} {P = P} Pset f kf t + = rec→Set (Pset t) g gk t where g : A → P t g x = transp (λ i → P (squash ∣ x ∣ t i)) i0 (f x) @@ -167,7 +203,7 @@ elimPropTrunc→Set {A = A} {P = P} Pset f kf t gk x y i = transp (λ j → P (squash (squash ∣ x ∣ ∣ y ∣ i) t j)) i0 (kf x y i) RecHProp : (P : A → hProp ℓ) (kP : ∀ x y → P x ≡ P y) → ∥ A ∥ → hProp ℓ -RecHProp P kP = recPropTrunc→Set isSetHProp P kP +RecHProp P kP = rec→Set isSetHProp P kP module GpdElim (Bgpd : isGroupoid B) where Bgpd' : isGroupoid' B @@ -176,8 +212,8 @@ module GpdElim (Bgpd : isGroupoid B) where module _ (f : A → B) (3kf : 3-Constant f) where open 3-Constant 3kf - recPropTrunc→Gpd : ∥ A ∥ → B - pathHelper : (t u : ∥ A ∥) → recPropTrunc→Gpd t ≡ recPropTrunc→Gpd u + rec→Gpd : ∥ A ∥ → B + pathHelper : (t u : ∥ A ∥) → rec→Gpd t ≡ rec→Gpd u triHelper₁ : (t u v : ∥ A ∥) → Square (pathHelper t u) (pathHelper t v) refl (pathHelper u v) @@ -185,8 +221,8 @@ module GpdElim (Bgpd : isGroupoid B) where : (t u v : ∥ A ∥) → Square (pathHelper t v) (pathHelper u v) (pathHelper t u) refl - recPropTrunc→Gpd ∣ x ∣ = f x - recPropTrunc→Gpd (squash t u i) = pathHelper t u i + rec→Gpd ∣ x ∣ = f x + rec→Gpd (squash t u i) = pathHelper t u i pathHelper ∣ x ∣ ∣ y ∣ = link x y pathHelper (squash t u j) v = triHelper₂ t u v j @@ -307,12 +343,112 @@ module GpdElim (Bgpd : isGroupoid B) where e-isEquiv a₀ = isoToIsEquiv (iso e (eval a₀) (e-eval a₀) λ _ → refl) preEquiv₂ : ∥ A ∥ → B ≃ Σ (A → B) 3-Constant - preEquiv₂ t = e , recPropTrunc (isPropIsEquiv e) e-isEquiv t + preEquiv₂ t = e , rec (isPropIsEquiv e) e-isEquiv t trunc→Gpd≃ : (∥ A ∥ → B) ≃ Σ (A → B) 3-Constant - trunc→Gpd≃ = compEquiv (equivPi preEquiv₂) preEquiv₁ + trunc→Gpd≃ = compEquiv (equivΠCod preEquiv₂) preEquiv₁ + +open GpdElim using (rec→Gpd; trunc→Gpd≃) public -open GpdElim using (recPropTrunc→Gpd; trunc→Gpd≃) public +squashᵗ + : ∀(x y z : A) + → Square (squash ∣ x ∣ ∣ y ∣) (squash ∣ x ∣ ∣ z ∣) refl (squash ∣ y ∣ ∣ z ∣) +squashᵗ x y z i = squash ∣ x ∣ (squash ∣ y ∣ ∣ z ∣ i) + +elim→Gpd + : (P : ∥ A ∥ → Type ℓ) + → (∀ t → isGroupoid (P t)) + → (f : (x : A) → P ∣ x ∣) + → (kf : ∀ x y → PathP (λ i → P (squash ∣ x ∣ ∣ y ∣ i)) (f x) (f y)) + → (3kf : ∀ x y z + → SquareP (λ i j → P (squashᵗ x y z i j)) (kf x y) (kf x z) refl (kf y z)) + → (t : ∥ A ∥) → P t +elim→Gpd {A = A} P Pgpd f kf 3kf t = rec→Gpd (Pgpd t) g 3kg t + where + g : A → P t + g x = transp (λ i → P (squash ∣ x ∣ t i)) i0 (f x) -RecHSet : (P : A → HLevel ℓ 2) → 3-Constant P → ∥ A ∥ → HLevel ℓ 2 -RecHSet P 3kP = recPropTrunc→Gpd (isOfHLevelHLevel 2) P 3kP + open 3-Constant + + 3kg : 3-Constant g + 3kg .link x y i + = transp (λ j → P (squash (squash ∣ x ∣ ∣ y ∣ i) t j)) i0 (kf x y i) + 3kg .coh₁ x y z i j + = transp (λ k → P (squash (squashᵗ x y z i j) t k)) i0 (3kf x y z i j) + +RecHSet : (P : A → TypeOfHLevel ℓ 2) → 3-Constant P → ∥ A ∥ → TypeOfHLevel ℓ 2 +RecHSet P 3kP = rec→Gpd (isOfHLevelTypeOfHLevel 2) P 3kP + +∥∥-IdempotentL-⊎-≃ : ∥ ∥ A ∥ ⊎ A′ ∥ ≃ ∥ A ⊎ A′ ∥ +∥∥-IdempotentL-⊎-≃ = isoToEquiv ∥∥-IdempotentL-⊎-Iso + where ∥∥-IdempotentL-⊎-Iso : Iso (∥ ∥ A ∥ ⊎ A′ ∥) (∥ A ⊎ A′ ∥) + Iso.fun ∥∥-IdempotentL-⊎-Iso x = rec squash lem x + where lem : ∥ A ∥ ⊎ A′ → ∥ A ⊎ A′ ∥ + lem (inl x) = map (λ a → inl a) x + lem (inr x) = ∣ inr x ∣ + Iso.inv ∥∥-IdempotentL-⊎-Iso x = map lem x + where lem : A ⊎ A′ → ∥ A ∥ ⊎ A′ + lem (inl x) = inl ∣ x ∣ + lem (inr x) = inr x + Iso.rightInv ∥∥-IdempotentL-⊎-Iso x = squash (Iso.fun ∥∥-IdempotentL-⊎-Iso (Iso.inv ∥∥-IdempotentL-⊎-Iso x)) x + Iso.leftInv ∥∥-IdempotentL-⊎-Iso x = squash (Iso.inv ∥∥-IdempotentL-⊎-Iso (Iso.fun ∥∥-IdempotentL-⊎-Iso x)) x + +∥∥-IdempotentL-⊎ : ∥ ∥ A ∥ ⊎ A′ ∥ ≡ ∥ A ⊎ A′ ∥ +∥∥-IdempotentL-⊎ = ua ∥∥-IdempotentL-⊎-≃ + +∥∥-IdempotentR-⊎-≃ : ∥ A ⊎ ∥ A′ ∥ ∥ ≃ ∥ A ⊎ A′ ∥ +∥∥-IdempotentR-⊎-≃ = isoToEquiv ∥∥-IdempotentR-⊎-Iso + where ∥∥-IdempotentR-⊎-Iso : Iso (∥ A ⊎ ∥ A′ ∥ ∥) (∥ A ⊎ A′ ∥) + Iso.fun ∥∥-IdempotentR-⊎-Iso x = rec squash lem x + where lem : A ⊎ ∥ A′ ∥ → ∥ A ⊎ A′ ∥ + lem (inl x) = ∣ inl x ∣ + lem (inr x) = map (λ a → inr a) x + Iso.inv ∥∥-IdempotentR-⊎-Iso x = map lem x + where lem : A ⊎ A′ → A ⊎ ∥ A′ ∥ + lem (inl x) = inl x + lem (inr x) = inr ∣ x ∣ + Iso.rightInv ∥∥-IdempotentR-⊎-Iso x = squash (Iso.fun ∥∥-IdempotentR-⊎-Iso (Iso.inv ∥∥-IdempotentR-⊎-Iso x)) x + Iso.leftInv ∥∥-IdempotentR-⊎-Iso x = squash (Iso.inv ∥∥-IdempotentR-⊎-Iso (Iso.fun ∥∥-IdempotentR-⊎-Iso x)) x + +∥∥-IdempotentR-⊎ : ∥ A ⊎ ∥ A′ ∥ ∥ ≡ ∥ A ⊎ A′ ∥ +∥∥-IdempotentR-⊎ = ua ∥∥-IdempotentR-⊎-≃ + +∥∥-Idempotent-⊎ : {A : Type ℓ} {A′ : Type ℓ′} → ∥ ∥ A ∥ ⊎ ∥ A′ ∥ ∥ ≡ ∥ A ⊎ A′ ∥ +∥∥-Idempotent-⊎ {A = A} {A′} = ∥ ∥ A ∥ ⊎ ∥ A′ ∥ ∥ ≡⟨ ∥∥-IdempotentR-⊎ ⟩ + ∥ ∥ A ∥ ⊎ A′ ∥ ≡⟨ ∥∥-IdempotentL-⊎ ⟩ + ∥ A ⊎ A′ ∥ ∎ + +∥∥-IdempotentL-×-≃ : ∥ ∥ A ∥ × A′ ∥ ≃ ∥ A × A′ ∥ +∥∥-IdempotentL-×-≃ = isoToEquiv ∥∥-IdempotentL-×-Iso + where ∥∥-IdempotentL-×-Iso : Iso (∥ ∥ A ∥ × A′ ∥) (∥ A × A′ ∥) + Iso.fun ∥∥-IdempotentL-×-Iso x = rec squash lem x + where lem : ∥ A ∥ × A′ → ∥ A × A′ ∥ + lem (a , a′) = map2 (λ a a′ → a , a′) a ∣ a′ ∣ + Iso.inv ∥∥-IdempotentL-×-Iso x = map lem x + where lem : A × A′ → ∥ A ∥ × A′ + lem (a , a′) = ∣ a ∣ , a′ + Iso.rightInv ∥∥-IdempotentL-×-Iso x = squash (Iso.fun ∥∥-IdempotentL-×-Iso (Iso.inv ∥∥-IdempotentL-×-Iso x)) x + Iso.leftInv ∥∥-IdempotentL-×-Iso x = squash (Iso.inv ∥∥-IdempotentL-×-Iso (Iso.fun ∥∥-IdempotentL-×-Iso x)) x + +∥∥-IdempotentL-× : ∥ ∥ A ∥ × A′ ∥ ≡ ∥ A × A′ ∥ +∥∥-IdempotentL-× = ua ∥∥-IdempotentL-×-≃ + +∥∥-IdempotentR-×-≃ : ∥ A × ∥ A′ ∥ ∥ ≃ ∥ A × A′ ∥ +∥∥-IdempotentR-×-≃ = isoToEquiv ∥∥-IdempotentR-×-Iso + where ∥∥-IdempotentR-×-Iso : Iso (∥ A × ∥ A′ ∥ ∥) (∥ A × A′ ∥) + Iso.fun ∥∥-IdempotentR-×-Iso x = rec squash lem x + where lem : A × ∥ A′ ∥ → ∥ A × A′ ∥ + lem (a , a′) = map2 (λ a a′ → a , a′) ∣ a ∣ a′ + Iso.inv ∥∥-IdempotentR-×-Iso x = map lem x + where lem : A × A′ → A × ∥ A′ ∥ + lem (a , a′) = a , ∣ a′ ∣ + Iso.rightInv ∥∥-IdempotentR-×-Iso x = squash (Iso.fun ∥∥-IdempotentR-×-Iso (Iso.inv ∥∥-IdempotentR-×-Iso x)) x + Iso.leftInv ∥∥-IdempotentR-×-Iso x = squash (Iso.inv ∥∥-IdempotentR-×-Iso (Iso.fun ∥∥-IdempotentR-×-Iso x)) x + +∥∥-IdempotentR-× : ∥ A × ∥ A′ ∥ ∥ ≡ ∥ A × A′ ∥ +∥∥-IdempotentR-× = ua ∥∥-IdempotentR-×-≃ + +∥∥-Idempotent-× : {A : Type ℓ} {A′ : Type ℓ′} → ∥ ∥ A ∥ × ∥ A′ ∥ ∥ ≡ ∥ A × A′ ∥ +∥∥-Idempotent-× {A = A} {A′} = ∥ ∥ A ∥ × ∥ A′ ∥ ∥ ≡⟨ ∥∥-IdempotentR-× ⟩ + ∥ ∥ A ∥ × A′ ∥ ≡⟨ ∥∥-IdempotentL-× ⟩ + ∥ A × A′ ∥ ∎ diff --git a/Cubical/HITs/Pushout.agda b/Cubical/HITs/Pushout.agda index 240af15e8..e1b04e650 100644 --- a/Cubical/HITs/Pushout.agda +++ b/Cubical/HITs/Pushout.agda @@ -1,5 +1,6 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Pushout where open import Cubical.HITs.Pushout.Base public open import Cubical.HITs.Pushout.Properties public +open import Cubical.HITs.Pushout.KrausVonRaumer public diff --git a/Cubical/HITs/Pushout/Base.agda b/Cubical/HITs/Pushout/Base.agda index 58520d1f4..2be7530cd 100644 --- a/Cubical/HITs/Pushout/Base.agda +++ b/Cubical/HITs/Pushout/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Pushout.Base where open import Cubical.Foundations.Prelude @@ -7,7 +7,7 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Data.Unit -open import Cubical.HITs.Susp +open import Cubical.HITs.Susp.Base data Pushout {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (f : A → B) (g : A → C) : Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) where @@ -15,6 +15,12 @@ data Pushout {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} inr : C → Pushout f g push : (a : A) → inl (f a) ≡ inr (g a) +-- cofiber (equivalent to Cone in Cubical.HITs.MappingCones.Base) +cofib : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type _ +cofib f = Pushout (λ _ → tt) f + +cfcod : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → B → cofib f +cfcod f = inr -- Suspension defined as a pushout @@ -43,8 +49,15 @@ PushoutSusp→Susp→PushoutSusp (inl _) = refl PushoutSusp→Susp→PushoutSusp (inr _) = refl PushoutSusp→Susp→PushoutSusp (push _ _) = refl +PushoutSuspIsoSusp : ∀ {ℓ} {A : Type ℓ} → Iso (PushoutSusp A) (Susp A) +Iso.fun PushoutSuspIsoSusp = PushoutSusp→Susp +Iso.inv PushoutSuspIsoSusp = Susp→PushoutSusp +Iso.rightInv PushoutSuspIsoSusp = Susp→PushoutSusp→Susp +Iso.leftInv PushoutSuspIsoSusp = PushoutSusp→Susp→PushoutSusp + + PushoutSusp≃Susp : ∀ {ℓ} {A : Type ℓ} → PushoutSusp A ≃ Susp A -PushoutSusp≃Susp = isoToEquiv (iso PushoutSusp→Susp Susp→PushoutSusp Susp→PushoutSusp→Susp PushoutSusp→Susp→PushoutSusp) +PushoutSusp≃Susp = isoToEquiv PushoutSuspIsoSusp PushoutSusp≡Susp : ∀ {ℓ} {A : Type ℓ} → PushoutSusp A ≡ Susp A -PushoutSusp≡Susp = isoToPath (iso PushoutSusp→Susp Susp→PushoutSusp Susp→PushoutSusp→Susp PushoutSusp→Susp→PushoutSusp) +PushoutSusp≡Susp = isoToPath PushoutSuspIsoSusp diff --git a/Cubical/HITs/Pushout/Flattening.agda b/Cubical/HITs/Pushout/Flattening.agda new file mode 100644 index 000000000..f9528296f --- /dev/null +++ b/Cubical/HITs/Pushout/Flattening.agda @@ -0,0 +1,91 @@ +{- + + The flattening lemma for pushouts (Lemma 8.5.3 in the HoTT book) proved in a cubical style. + + The proof in the HoTT book (the core lying in Lemma 6.12.2, the flattening lemma for coequalizers) + consists mostly of long strings of equalities about transport. This proof follows almost + entirely from definitional equalities involving glue/unglue. + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.Pushout.Flattening where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Sigma + +open import Cubical.HITs.Pushout.Base + +module FlatteningLemma {ℓa ℓb ℓc} {A : Type ℓa} {B : Type ℓb} {C : Type ℓc} (f : A → B) (g : A → C) + {ℓ} (F : B → Type ℓ) (G : C → Type ℓ) (e : ∀ a → F (f a) ≃ G (g a)) where + + E : Pushout f g → Type ℓ + E (inl x) = F x + E (inr x) = G x + E (push a i) = ua (e a) i + + Σf : Σ[ a ∈ A ] F (f a) → Σ[ b ∈ B ] F b + Σf (a , x) = (f a , x) + + Σg : Σ[ a ∈ A ] F (f a) → Σ[ c ∈ C ] G c + Σg (a , x) = (g a , (e a) .fst x) + + module FlattenIso where + + fwd : Pushout Σf Σg → Σ (Pushout f g) E + fwd (inl (b , x)) = (inl b , x) + fwd (inr (c , x)) = (inr c , x) + fwd (push (a , x) i) = (push a i , ua-gluePt (e a) i x) + + bwd : Σ (Pushout f g) E → Pushout Σf Σg + bwd (inl b , x) = inl (b , x) + bwd (inr c , x) = inr (c , x) + bwd (push a i , x) = hcomp (λ j → λ { (i = i0) → push (a , x) (~ j) + ; (i = i1) → inr (g a , x) }) + (inr (g a , ua-unglue (e a) i x)) + + bwd-fwd : ∀ x → bwd (fwd x) ≡ x + bwd-fwd (inl (b , x)) = refl + bwd-fwd (inr (c , x)) = refl + bwd-fwd (push (a , x) i) j = + hcomp (λ k → λ { (i = i0) → push (a , ua-gluePt (e a) i0 x) (~ k) + ; (i = i1) → inr (g a , ua-gluePt (e a) i1 x) + ; (j = i1) → push (a , x) (i ∨ ~ k) }) + (inr (g a , ua-unglue (e a) i (ua-gluePt (e a) i x))) + -- Note: the (j = i1) case typechecks because of the definitional equalities: + -- ua-gluePt e i0 x ≡ x , ua-gluePt e i1 x ≡ e .fst x, + -- ua-unglue-glue : ua-unglue e i (ua-gluePt e i x) ≡ e .fst x + + -- essentially: ua-glue e (i ∨ ~ k) ∘ ua-unglue e i + sq : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) + → SquareP (λ i k → ua e i → ua e (i ∨ ~ k)) + {- i = i0 -} (λ k x → ua-gluePt e (~ k) x) + {- i = i1 -} (λ k x → x) + {- k = i0 -} (λ i x → ua-unglue e i x) + {- k = i1 -} (λ i x → x) + sq e i k x = ua-glue e (i ∨ ~ k) (λ { ((i ∨ ~ k) = i0) → x }) + (inS (ua-unglue e i x)) + -- Note: this typechecks because of the definitional equalities: + -- ua-unglue e i0 x ≡ e .fst x, ua-glue e i1 _ (inS y) ≡ y, ua-unglue e i1 x ≡ x, + -- ua-glue-unglue : ua-glue e i (λ { (i = i0) → x }) (inS (ua-unglue e i x)) ≡ x + + fwd-bwd : ∀ x → fwd (bwd x) ≡ x + fwd-bwd (inl b , x) = refl + fwd-bwd (inr c , x) = refl + fwd-bwd (push a i , x) j = + -- `fwd` (or any function) takes hcomps to comps on a constant family, so we must use a comp here + comp (λ _ → Σ (Pushout f g) E) + (λ k → λ { (i = i0) → push a (~ k) , ua-gluePt (e a) (~ k) x + ; (i = i1) → inr (g a) , x + ; (j = i1) → push a (i ∨ ~ k) , sq (e a) i k x }) + (inr (g a) , ua-unglue (e a) i x) + + isom : Iso (Σ (Pushout f g) E) (Pushout Σf Σg) + isom = iso bwd fwd bwd-fwd fwd-bwd + + flatten : Σ (Pushout f g) E ≃ Pushout Σf Σg + flatten = isoToEquiv FlattenIso.isom diff --git a/Cubical/HITs/Pushout/KrausVonRaumer.agda b/Cubical/HITs/Pushout/KrausVonRaumer.agda new file mode 100644 index 000000000..c6881e69a --- /dev/null +++ b/Cubical/HITs/Pushout/KrausVonRaumer.agda @@ -0,0 +1,137 @@ +{- + +An induction principle for paths in a pushout, described in +Kraus and von Raumer, "Path Spaces of Higher Inductive Types in Homotopy Type Theory" +https://arxiv.org/abs/1901.06022 + +-} + +{-# OPTIONS --safe #-} +module Cubical.HITs.Pushout.KrausVonRaumer where + +open import Cubical.Foundations.Everything +open import Cubical.Functions.Embedding +open import Cubical.Data.Sigma +open import Cubical.HITs.Pushout.Base + +private + interpolate : ∀ {ℓ} {A : Type ℓ} {x y z : A} (q : y ≡ z) + → PathP (λ i → x ≡ q i → x ≡ z) (_∙ q) (idfun _) + interpolate q i p j = + hcomp + (λ k → λ + { (j = i0) → p i0 + ; (j = i1) → q (i ∨ k) + ; (i = i1) → p j + }) + (p j) + + interpolateCompPath : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) {z : A} (q : y ≡ z) + → (λ i → interpolate q i (λ j → compPath-filler p q i j)) ≡ refl + interpolateCompPath p = + J (λ z q → (λ i → interpolate q i (λ j → compPath-filler p q i j)) ≡ refl) + (homotopySymInv (λ p i j → compPath-filler p refl (~ i) j) p) + +module ElimL {ℓ ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + {f : A → B} {g : A → C} {b₀ : B} + (P : ∀ b → Path (Pushout f g) (inl b₀) (inl b) → Type ℓ''') + (Q : ∀ c → Path (Pushout f g) (inl b₀) (inr c) → Type ℓ''') + (r : P b₀ refl) + (e : (a : A) (q : inl b₀ ≡ inl (f a)) → P (f a) q ≃ Q (g a) (q ∙ push a)) + where + + Codes : (d : Pushout f g) (q : inl b₀ ≡ d) → Type ℓ''' + Codes (inl b) q = P b q + Codes (inr c) q = Q c q + Codes (push a i) q = + Glue + (Q (g a) (interpolate (push a) i q)) + (λ + { (i = i0) → _ , e a q + ; (i = i1) → _ , idEquiv (Q (g a) q) + }) + + elimL : ∀ b q → P b q + elimL _ = J Codes r + + elimR : ∀ c q → Q c q + elimR _ = J Codes r + + refl-β : elimL b₀ refl ≡ r + refl-β = transportRefl _ + + push-β : (a : A) (q : inl b₀ ≡ inl (f a)) + → elimR (g a) (q ∙ push a) ≡ e a q .fst (elimL (f a) q) + push-β a q = + J-∙ Codes r q (push a) + ∙ fromPathP + (subst + (λ α → PathP (λ i → Q (g a) (α i)) (e a q .fst (elimL (f a) q)) (e a q .fst (elimL (f a) q))) + (interpolateCompPath q (push a) ⁻¹) + refl) + +module ElimR {ℓ ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + {f : A → B} {g : A → C} {c₀ : C} + (P : ∀ b → Path (Pushout f g) (inr c₀) (inl b) → Type ℓ''') + (Q : ∀ c → Path (Pushout f g) (inr c₀) (inr c) → Type ℓ''') + (r : Q c₀ refl) + (e : (a : A) (q : inr c₀ ≡ inl (f a)) → P (f a) q ≃ Q (g a) (q ∙ push a)) + where + + Codes : (d : Pushout f g) (q : inr c₀ ≡ d) → Type ℓ''' + Codes (inl b) q = P b q + Codes (inr c) q = Q c q + Codes (push a i) q = + Glue + (Q (g a) (interpolate (push a) i q)) + (λ + { (i = i0) → _ , e a q + ; (i = i1) → _ , idEquiv (Q (g a) q) + }) + + elimL : ∀ b q → P b q + elimL _ = J Codes r + + elimR : ∀ c q → Q c q + elimR _ = J Codes r + + refl-β : elimR c₀ refl ≡ r + refl-β = transportRefl _ + + push-β : (a : A) (q : inr c₀ ≡ inl (f a)) + → elimR (g a) (q ∙ push a) ≡ e a q .fst (elimL (f a) q) + push-β a q = + J-∙ Codes r q (push a) + ∙ fromPathP + (subst + (λ α → PathP (λ i → Q (g a) (α i)) (e a q .fst (elimL (f a) q)) (e a q .fst (elimL (f a) q))) + (interpolateCompPath q (push a) ⁻¹) + refl) + +-- Example application: pushouts preserve embeddings + +isEmbeddingInr : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + {f : A → B} (g : A → C) + → isEmbedding f → isEmbedding (inr {f = f} {g = g}) +isEmbeddingInr {f = f} g fEmb c₀ c₁ = + isoToIsEquiv (iso _ (fst ∘ bwd c₁) (snd ∘ bwd c₁) bwdCong) + where + Q : ∀ c → inr c₀ ≡ inr c → Type _ + Q _ q = fiber (cong inr) q + + P : ∀ b → inr c₀ ≡ inl b → Type _ + P b p = Σ[ u ∈ fiber f b ] Q _ (p ∙ cong inl (u .snd ⁻¹) ∙ push (u .fst)) + + module Bwd = ElimR P Q + (refl , refl) + (λ a p → + subst + (P (f a) p ≃_) + (cong (λ w → fiber (cong inr) (p ∙ w)) (lUnit (push a) ⁻¹)) + (Σ-contractFst (inhProp→isContr (a , refl) (isEmbedding→hasPropFibers fEmb (f a))))) + + bwd : ∀ c → (t : inr c₀ ≡ inr c) → fiber (cong inr) t + bwd = Bwd.elimR + + bwdCong : ∀ {c} → (r : c₀ ≡ c) → bwd c (cong inr r) .fst ≡ r + bwdCong = J (λ c r → bwd c (cong inr r) .fst ≡ r) (cong fst Bwd.refl-β) diff --git a/Cubical/HITs/Pushout/Properties.agda b/Cubical/HITs/Pushout/Properties.agda index 5df36f5aa..8429160d9 100644 --- a/Cubical/HITs/Pushout/Properties.agda +++ b/Cubical/HITs/Pushout/Properties.agda @@ -10,19 +10,20 @@ This file contains: -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Pushout.Properties where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.HAEquiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Univalence open import Cubical.Foundations.Transport -open import Cubical.Data.Prod.Base +open import Cubical.Data.Sigma open import Cubical.Data.Unit open import Cubical.HITs.Pushout.Base @@ -37,6 +38,9 @@ record 3-span : Type₁ where f1 : A2 → A0 f3 : A2 → A4 +3span : {A0 A2 A4 : Type₀} → (A2 → A0) → (A2 → A4) → 3-span +3span f1 f3 = record { f1 = f1 ; f3 = f3 } + spanPushout : (s : 3-span) → Type₀ spanPushout s = Pushout (3-span.f1 s) (3-span.f3 s) @@ -222,10 +226,10 @@ record 3x3-span : Type₁ where forward-r (push a i) = push (inr a) i forward-filler : A22 → I → I → I → A○□ - forward-filler a i j = hfill (λ t → λ { (i = i0) → inl (doubleCompPath-filler (λ k → inl (H11 a (~ k))) (push (f12 a)) (λ k → inr (H13 a k)) j (~ t)) - ; (i = i1) → inr (doubleCompPath-filler (λ k → inl (H31 a (~ k))) (push (f32 a)) (λ k → inr (H33 a k)) j (~ t)) - ; (j = i0) → forward-l (doubleCompPath-filler (λ k → inl (H11 a k)) (push (f21 a)) (λ k → inr (H31 a (~ k))) i t) - ; (j = i1) → forward-r (doubleCompPath-filler (λ k → inl (H13 a k)) (push (f23 a)) (λ k → inr (H33 a (~ k))) i t) }) + forward-filler a i j = hfill (λ t → λ { (i = i0) → inl (doubleCompPath-filler (λ k → inl (H11 a (~ k))) (push (f12 a)) (λ k → inr (H13 a k)) (~ t) j) + ; (i = i1) → inr (doubleCompPath-filler (λ k → inl (H31 a (~ k))) (push (f32 a)) (λ k → inr (H33 a k)) (~ t) j) + ; (j = i0) → forward-l (doubleCompPath-filler (λ k → inl (H11 a k)) (push (f21 a)) (λ k → inr (H31 a (~ k))) t i) + ; (j = i1) → forward-r (doubleCompPath-filler (λ k → inl (H13 a k)) (push (f23 a)) (λ k → inr (H33 a (~ k))) t i) }) (inS (push (push a j) i)) A□○→A○□ : A□○ → A○□ @@ -247,10 +251,10 @@ record 3x3-span : Type₁ where backward-r (push a i) = push (inr a) i backward-filler : A22 → I → I → I → A□○ - backward-filler a i j = hfill (λ t → λ { (i = i0) → inl (doubleCompPath-filler (λ k → inl (H11 a k)) (push (f21 a)) (λ k → inr (H31 a (~ k))) j (~ t)) - ; (i = i1) → inr (doubleCompPath-filler (λ k → inl (H13 a k)) (push (f23 a)) (λ k → inr (H33 a (~ k))) j (~ t)) - ; (j = i0) → backward-l (doubleCompPath-filler (λ k → inl (H11 a (~ k))) (push (f12 a)) (λ k → inr (H13 a k)) i t) - ; (j = i1) → backward-r (doubleCompPath-filler (λ k → inl (H31 a (~ k))) (push (f32 a)) (λ k → inr (H33 a k)) i t) }) + backward-filler a i j = hfill (λ t → λ { (i = i0) → inl (doubleCompPath-filler (λ k → inl (H11 a k)) (push (f21 a)) (λ k → inr (H31 a (~ k))) (~ t) j) + ; (i = i1) → inr (doubleCompPath-filler (λ k → inl (H13 a k)) (push (f23 a)) (λ k → inr (H33 a (~ k))) (~ t) j) + ; (j = i0) → backward-l (doubleCompPath-filler (λ k → inl (H11 a (~ k))) (push (f12 a)) (λ k → inr (H13 a k)) t i) + ; (j = i1) → backward-r (doubleCompPath-filler (λ k → inl (H31 a (~ k))) (push (f32 a)) (λ k → inr (H33 a k)) t i) }) (inS (push (push a j) i)) A○□→A□○ : A○□ → A□○ @@ -277,10 +281,10 @@ record 3x3-span : Type₁ where A○□→A□○→A○□ (push (inl x) i) k = push (inl x) i A○□→A□○→A○□ (push (inr x) i) k = push (inr x) i A○□→A□○→A○□ (push (push a i) j) k = - hcomp (λ t → λ { (i = i0) → forward-l (doubleCompPath-filler (λ k → inl (H11 a k)) (push (f21 a)) (λ k → inr (H31 a (~ k))) j (~ t)) - ; (i = i1) → forward-r (doubleCompPath-filler (λ k → inl (H13 a k)) (push (f23 a)) (λ k → inr (H33 a (~ k))) j (~ t)) - ; (j = i0) → homotopy1-l (doubleCompPath-filler (λ k → inl (H11 a (~ k))) (push (f12 a)) (λ k → inr (H13 a k)) i t) k - ; (j = i1) → homotopy1-r (doubleCompPath-filler (λ k → inl (H31 a (~ k))) (push (f32 a)) (λ k → inr (H33 a k)) i t) k + hcomp (λ t → λ { (i = i0) → forward-l (doubleCompPath-filler (λ k → inl (H11 a k)) (push (f21 a)) (λ k → inr (H31 a (~ k))) (~ t) j) + ; (i = i1) → forward-r (doubleCompPath-filler (λ k → inl (H13 a k)) (push (f23 a)) (λ k → inr (H33 a (~ k))) (~ t) j) + ; (j = i0) → homotopy1-l (doubleCompPath-filler (λ k → inl (H11 a (~ k))) (push (f12 a)) (λ k → inr (H13 a k)) t i) k + ; (j = i1) → homotopy1-r (doubleCompPath-filler (λ k → inl (H31 a (~ k))) (push (f32 a)) (λ k → inr (H33 a k)) t i) k ; (k = i0) → A□○→A○□ (backward-filler a i j t) ; (k = i1) → forward-filler a j i (~ t) }) (forward-filler a j i i1) @@ -302,10 +306,23 @@ record 3x3-span : Type₁ where A□○→A○□→A□○ (push (inl x) i) k = push (inl x) i A□○→A○□→A□○ (push (inr x) i) k = push (inr x) i A□○→A○□→A□○ (push (push a i) j) k = - hcomp (λ t → λ { (i = i0) → backward-l (doubleCompPath-filler (λ k → inl (H11 a (~ k))) (push (f12 a)) (λ k → inr (H13 a k)) j (~ t)) - ; (i = i1) → backward-r (doubleCompPath-filler (λ k → inl (H31 a (~ k))) (push (f32 a)) (λ k → inr (H33 a k)) j (~ t)) - ; (j = i0) → homotopy2-l (doubleCompPath-filler (λ k → inl (H11 a k)) (push (f21 a)) (λ k → inr (H31 a (~ k))) i t) k - ; (j = i1) → homotopy2-r (doubleCompPath-filler (λ k → inl (H13 a k)) (push (f23 a)) (λ k → inr (H33 a (~ k))) i t) k + hcomp (λ t → λ { (i = i0) → backward-l (doubleCompPath-filler (λ k → inl (H11 a (~ k))) (push (f12 a)) (λ k → inr (H13 a k)) (~ t) j) + ; (i = i1) → backward-r (doubleCompPath-filler (λ k → inl (H31 a (~ k))) (push (f32 a)) (λ k → inr (H33 a k)) (~ t) j) + ; (j = i0) → homotopy2-l (doubleCompPath-filler (λ k → inl (H11 a k)) (push (f21 a)) (λ k → inr (H31 a (~ k))) t i) k + ; (j = i1) → homotopy2-r (doubleCompPath-filler (λ k → inl (H13 a k)) (push (f23 a)) (λ k → inr (H33 a (~ k))) t i) k ; (k = i0) → A○□→A□○ (forward-filler a i j t) ; (k = i1) → backward-filler a j i (~ t) }) (backward-filler a j i i1) + + + +PushoutToProp : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {f : A → B} {g : A → C} + {D : Pushout f g → Type ℓ'''} + → ((x : Pushout f g) → isProp (D x)) + → ((a : B) → D (inl a)) + → ((c : C) → D (inr c)) + → (x : Pushout f g) → D x +PushoutToProp isset baseB baseC (inl x) = baseB x +PushoutToProp isset baseB baseC (inr x) = baseC x +PushoutToProp {f = f} {g = g} isset baseB baseC (push a i) = + isOfHLevel→isOfHLevelDep 1 isset (baseB (f a)) (baseC (g a)) (push a) i diff --git a/Cubical/HITs/RPn.agda b/Cubical/HITs/RPn.agda new file mode 100644 index 000000000..019b3ee2a --- /dev/null +++ b/Cubical/HITs/RPn.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.RPn where + +open import Cubical.HITs.RPn.Base public diff --git a/Cubical/HITs/RPn/Base.agda b/Cubical/HITs/RPn/Base.agda new file mode 100644 index 000000000..f3df232e6 --- /dev/null +++ b/Cubical/HITs/RPn/Base.agda @@ -0,0 +1,320 @@ +{- + + A defintion of the real projective spaces following: + + [BR17] U. Buchholtz, E. Rijke, The real projective spaces in homotopy type theory. + (2017) https://arxiv.org/abs/1704.05770 + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.RPn.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Functions.Bundle + +open import Cubical.Foundations.SIP +open import Cubical.Structures.Pointed +open import Cubical.Structures.TypeEqvTo + +open import Cubical.Data.Bool +open import Cubical.Data.Nat hiding (elim) +open import Cubical.Data.NatMinusOne +open import Cubical.Data.Sigma +open import Cubical.Data.Unit +open import Cubical.Data.Empty as ⊥ hiding (elim) +open import Cubical.Data.Sum as ⊎ hiding (elim) + +open import Cubical.HITs.PropositionalTruncation as PropTrunc hiding (elim) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Join +open import Cubical.HITs.Pushout +open import Cubical.HITs.Pushout.Flattening + +private + variable + ℓ ℓ' ℓ'' : Level + +-- PR² as a HIT +data RP² : Type₀ where + point : RP² + line : point ≡ point + square : line ≡ sym line + +-- Definition II.1 in [BR17], see also Cubical.Functions.Bundle + +2-EltType₀ = TypeEqvTo ℓ-zero Bool -- Σ[ X ∈ Type₀ ] ∥ X ≃ Bool ∥ +2-EltPointed₀ = PointedEqvTo ℓ-zero Bool -- Σ[ X ∈ Type₀ ] X × ∥ X ≃ Bool ∥ + +Bool* : 2-EltType₀ +Bool* = Bool , ∣ idEquiv _ ∣ + + +-- Our first goal is to 'lift' `_⊕_ : Bool → Bool ≃ Bool` to a function `_⊕_ : A → A ≃ Bool` +-- for any 2-element type (A, ∣e∣). + +-- `isContrBoolPointedEquiv` and `isContr-2-EltPointedEquiv` are contained in the proof +-- of Lemma II.2 in [BR17], though we prove `isContr-BoolPointedEquiv` more directly +-- with ⊕ -- [BR17] proves it for just the x = false case and uses notEquiv to get +-- the x = true case. + +-- (λ y → x ⊕ y) is the unqiue pointed isomorphism (Bool , false) ≃ (Bool , x) +isContrBoolPointedEquiv : ∀ x → isContr ((Bool , false) ≃[ PointedEquivStr ] (Bool , x)) +fst (isContrBoolPointedEquiv x) = ((λ y → x ⊕ y) , isEquiv-⊕ x) , ⊕-comm x false +snd (isContrBoolPointedEquiv x) (e , p) + = Σ≡Prop (λ e → isSetBool (equivFun e false) x) + (Σ≡Prop isPropIsEquiv (funExt λ { false → ⊕-comm x false ∙ sym p + ; true → ⊕-comm x true ∙ sym q })) + where q : e .fst true ≡ not x + q with dichotomyBool (invEq e (not x)) + ... | inl q = invEq≡→equivFun≡ e q + ... | inr q = ⊥.rec (not≢const x (sym (invEq≡→equivFun≡ e q) ∙ p)) + +-- Since isContr is a mere proposition, we can eliminate a witness ∣e∣ : ∥ X ≃ Bool ∥ to get +-- that there is therefore a unique pointed isomorphism (Bool , false) ≃ (X , x) for any +-- 2-element pointed type (X , x, ∣e∣). +isContr-2-EltPointedEquiv : (X∙ : 2-EltPointed₀) + → isContr ((Bool , false , ∣ idEquiv Bool ∣) ≃[ PointedEqvToEquivStr Bool ] X∙) +isContr-2-EltPointedEquiv (X , x , ∣e∣) + = PropTrunc.rec isPropIsContr + (λ e → J (λ X∙ _ → isContr ((Bool , false) ≃[ PointedEquivStr ] X∙)) + (isContrBoolPointedEquiv (e .fst x)) + (sym (pointed-sip _ _ (e , refl)))) + ∣e∣ + +-- This unique isomorphism must be _⊕_ 'lifted' to X. This idea is alluded to at the end of the +-- proof of Theorem III.4 in [BR17], where the authors reference needing ⊕-comm. +module ⊕* (X : 2-EltType₀) where + + _⊕*_ : typ X → typ X → Bool + y ⊕* z = invEquiv (fst (fst (isContr-2-EltPointedEquiv (fst X , y , snd X)))) .fst z + + -- we've already shown that this map is an equivalence on the right + + isEquivʳ : (y : typ X) → isEquiv (y ⊕*_) + isEquivʳ y = invEquiv (fst (fst (isContr-2-EltPointedEquiv (fst X , y , snd X)))) .snd + + Equivʳ : typ X → typ X ≃ Bool + Equivʳ y = (y ⊕*_) , isEquivʳ y + + -- any mere proposition that holds for (Bool, _⊕_) holds for (typ X, _⊕*_) + -- this amounts to just carefully unfolding the PropTrunc.elim and J in isContr-2-EltPointedEquiv + elim : ∀ {ℓ'} (P : (A : Type₀) (_⊕'_ : A → A → Bool) → Type ℓ') (propP : ∀ A _⊕'_ → isProp (P A _⊕'_)) + → P Bool _⊕_ → P (typ X) _⊕*_ + elim {ℓ'} P propP r = PropTrunc.elim {P = λ ∣e∣ → P (typ X) (R₁ ∣e∣)} (λ _ → propP _ _) + (λ e → EquivJ (λ A e → P A (R₂ A e)) r e) + (snd X) + where R₁ : ∥ fst X ≃ Bool ∥ → typ X → typ X → Bool + R₁ ∣e∣ y = invEq (fst (fst (isContr-2-EltPointedEquiv (fst X , y , ∣e∣)))) + R₂ : (B : Type₀) → B ≃ Bool → B → B → Bool + R₂ A e y = invEq (fst (fst (J (λ A∙ _ → isContr ((Bool , false) ≃[ PointedEquivStr ] A∙)) + (isContrBoolPointedEquiv (e .fst y)) + (sym (pointed-sip (A , y) (Bool , e .fst y) (e , refl)))))) + + -- as a consequence, we get that ⊕* is commutative, and is therefore also an equivalence on the left + + comm : (y z : typ X) → y ⊕* z ≡ z ⊕* y + comm = elim (λ A _⊕'_ → (x y : A) → x ⊕' y ≡ y ⊕' x) + (λ _ _ → isPropΠ2 (λ _ _ → isSetBool _ _)) + ⊕-comm + + isEquivˡ : (y : typ X) → isEquiv (_⊕* y) + isEquivˡ y = subst isEquiv (funExt (λ z → comm y z)) (isEquivʳ y) + + Equivˡ : typ X → typ X ≃ Bool + Equivˡ y = (_⊕* y) , isEquivˡ y + +-- Lemma II.2 in [BR17], though we do not use it here +-- Note: Lemma II.3 is `pointed-sip`, used in `PointedEqvTo-sip` +isContr-2-EltPointed : isContr (2-EltPointed₀) +fst isContr-2-EltPointed = (Bool , false , ∣ idEquiv Bool ∣) +snd isContr-2-EltPointed A∙ = PointedEqvTo-sip Bool _ A∙ (fst (isContr-2-EltPointedEquiv A∙)) + + +-------------------------------------------------------------------------------- + +-- Now we mutually define RP n and its double cover (Definition III.1 in [BR17]), +-- and show that the total space of this double cover is S n (Theorem III.4). + +RP : ℕ₋₁ → Type₀ +cov⁻¹ : (n : ℕ₋₁) → RP n → 2-EltType₀ -- (see Cubical.Functions.Bundle) + +RP neg1 = ⊥ +RP (ℕ→ℕ₋₁ n) = Pushout (pr (cov⁻¹ (-1+ n))) (λ _ → tt) +{- + tt + Total (cov⁻¹ (n-1)) — — — > Unit + | ∙ + pr | ∙ inr + | ∙ + V V + RP (n-1) ∙ ∙ ∙ ∙ ∙ ∙ > RP n := Pushout pr (const tt) + inl +-} + +cov⁻¹ neg1 x = Bool* +cov⁻¹ (ℕ→ℕ₋₁ n) (inl x) = cov⁻¹ (-1+ n) x +cov⁻¹ (ℕ→ℕ₋₁ n) (inr _) = Bool* +cov⁻¹ (ℕ→ℕ₋₁ n) (push (x , y) i) = ua ((λ z → y ⊕* z) , ⊕*.isEquivʳ (cov⁻¹ (-1+ n) x) y) i , ∣p∣ i + where open ⊕* (cov⁻¹ (-1+ n) x) + ∣p∣ = isProp→PathP (λ i → squash {A = ua (⊕*.Equivʳ (cov⁻¹ (-1+ n) x) y) i ≃ Bool}) + (str (cov⁻¹ (-1+ n) x)) (∣ idEquiv _ ∣) +{- + tt + Total (cov⁻¹ (n-1)) — — — > Unit + | | + pr | // ua α // | Bool + | | + V V + RP (n-1) - - - - - - > Type + cov⁻¹ (n-1) + + where α : ∀ (x : Total (cov⁻¹ (n-1))) → cov⁻¹ (n-1) (pr x) ≃ Bool + α (x , y) = (λ z → y ⊕* z) , ⊕*.isEquivʳ y +-} + +TotalCov≃Sn : ∀ n → Total (cov⁻¹ n) ≃ S n +TotalCov≃Sn neg1 = isoToEquiv (iso (λ { () }) (λ { () }) (λ { () }) (λ { () })) +TotalCov≃Sn (ℕ→ℕ₋₁ n) = + Total (cov⁻¹ (ℕ→ℕ₋₁ n)) ≃⟨ i ⟩ + Pushout Σf Σg ≃⟨ ii ⟩ + join (Total (cov⁻¹ (-1+ n))) Bool ≃⟨ iii ⟩ + S (ℕ→ℕ₋₁ n) ■ + where +{- + (i) First we want to show that `Total (cov⁻¹ (ℕ→ℕ₋₁ n))` is equivalent to a pushout. + We do this using the flattening lemma, which states: + + Given f,g,F,G,e such that the following square commutes: + + g + A — — — — > C Define: E : Pushout f g → Type + | | E (inl b) = F b + f | ua e | G E (inr c) = G c + V V E (push a i) = ua (e a) i + B — — — — > Type + F + + Then, the total space `Σ (Pushout f g) E` is the following pushout: + + Σg := (g , e a) + Σ[ a ∈ A ] F (f a) — — — — — — — — > Σ[ c ∈ C ] G c + | ∙ + Σf := (f , id) | ∙ + V V + Σ[ b ∈ B ] F b ∙ ∙ ∙ ∙ ∙ ∙ ∙ ∙ ∙ ∙ > Σ (Pushout f g) E + + In our case, setting `f = pr (cov⁻¹ (n-1))`, `g = λ _ → tt`, `F = cov⁻¹ (n-1)`, `G = λ _ → Bool`, + and `e = λ (x , y) → ⊕*.Equivʳ (cov⁻¹ (n-1) x) y` makes E equal (up to funExt) to `cov⁻¹ n`. + + Thus the flattening lemma gives us that `Total (cov⁻¹ n) ≃ Pushout Σf Σg`. +-} + open FlatteningLemma {- f = -} (λ x → pr (cov⁻¹ (-1+ n)) x) {- g = -} (λ _ → tt) + {- F = -} (λ x → typ (cov⁻¹ (-1+ n) x)) {- G = -} (λ _ → Bool) + {- e = -} (λ { (x , y) → ⊕*.Equivʳ (cov⁻¹ (-1+ n) x) y }) + hiding (Σf ; Σg) + + cov⁻¹≃E : ∀ x → typ (cov⁻¹ (ℕ→ℕ₋₁ n) x) ≃ E x + cov⁻¹≃E (inl x) = idEquiv _ + cov⁻¹≃E (inr x) = idEquiv _ + cov⁻¹≃E (push a i) = idEquiv _ + + -- for easier reference, we copy these definitons here + Σf : Σ[ x ∈ Total (cov⁻¹ (-1+ n)) ] typ (cov⁻¹ (-1+ n) (fst x)) → Total (cov⁻¹ (-1+ n)) + Σg : Σ[ x ∈ Total (cov⁻¹ (-1+ n)) ] typ (cov⁻¹ (-1+ n) (fst x)) → Unit × Bool + Σf ((x , y) , z) = (x , z) -- ≡ (f a , x) + Σg ((x , y) , z) = (tt , y ⊕* z) -- ≡ (g a , (e a) .fst x) + where open ⊕* (cov⁻¹ (-1+ n) x) + + i : Total (cov⁻¹ (ℕ→ℕ₋₁ n)) ≃ Pushout Σf Σg + i = (Σ[ x ∈ RP (ℕ→ℕ₋₁ n) ] typ (cov⁻¹ (ℕ→ℕ₋₁ n) x)) ≃⟨ Σ-cong-equiv-snd cov⁻¹≃E ⟩ + (Σ[ x ∈ RP (ℕ→ℕ₋₁ n) ] E x) ≃⟨ flatten ⟩ + Pushout Σf Σg ■ +{- + (ii) Next we want to show that `Pushout Σf Σg` is equivalent to `join (Total (cov⁻¹ (n-1))) Bool`. + Since both are pushouts, this can be done by defining a diagram equivalence: + + Σf Σg + Total (cov⁻¹ (n-1)) < — — Σ[ x ∈ Total (cov⁻¹ (n-1)) ] cov⁻¹ (n-1) (pr x) — — > Unit × Bool + | ∙ | + id |≃ u ∙≃ snd |≃ + V V V + Total (cov⁻¹ (n-1)) < — — — — — — — Total (cov⁻¹ (n-1)) × Bool — — — — — — — — — > Bool + proj₁ proj₂ + + where the equivalence u above must therefore satisfy: `u .fst x ≡ (Σf x , snd (Σg x))` + Unfolding this, we get: `u .fst ((x , y) , z) ≡ ((x , z) , (y ⊕* z))` + + It suffices to show that the map y ↦ y ⊕* z is an equivalence, since we can then express u as + the following composition of equivalences: + ((x , y) , z) ↦ (x , (y , z)) ↦ (x , (z , y)) ↦ (x , (z , y ⊕* z)) ↦ ((x , z) , y ⊕* z) + + This was proved above by ⊕*.isEquivˡ. +-} + u : ∀ {n} → (Σ[ x ∈ Total (cov⁻¹ n) ] typ (cov⁻¹ n (fst x))) ≃ (Total (cov⁻¹ n) × Bool) + u {n} = Σ[ x ∈ Total (cov⁻¹ n) ] typ (cov⁻¹ n (fst x)) ≃⟨ Σ-assoc-≃ ⟩ + Σ[ x ∈ RP n ] (typ (cov⁻¹ n x)) × (typ (cov⁻¹ n x)) ≃⟨ Σ-cong-equiv-snd (λ x → Σ-swap-≃) ⟩ + Σ[ x ∈ RP n ] (typ (cov⁻¹ n x)) × (typ (cov⁻¹ n x)) ≃⟨ Σ-cong-equiv-snd + (λ x → Σ-cong-equiv-snd + (λ y → ⊕*.Equivˡ (cov⁻¹ n x) y)) ⟩ + Σ[ x ∈ RP n ] (typ (cov⁻¹ n x)) × Bool ≃⟨ invEquiv Σ-assoc-≃ ⟩ + Total (cov⁻¹ n) × Bool ■ + + H : ∀ x → u .fst x ≡ (Σf x , snd (Σg x)) + H x = refl + + nat : 3-span-equiv (3span Σf Σg) (3span {A2 = Total (cov⁻¹ (-1+ n)) × Bool} fst snd) + nat = record { e0 = idEquiv _ ; e2 = u ; e4 = ΣUnit _ + ; H1 = λ x → cong fst (H x) + ; H3 = λ x → cong snd (H x) } + + ii : Pushout Σf Σg ≃ join (Total (cov⁻¹ (-1+ n))) Bool + ii = compEquiv (pathToEquiv (spanEquivToPushoutPath nat)) (joinPushout≃join _ _) + +{- + (iii) Finally, it's trivial to show that `join (Total (cov⁻¹ (n-1))) Bool` is equivalent to + `Susp (Total (cov⁻¹ (n-1)))`. Induction then gives us that `Susp (Total (cov⁻¹ (n-1)))` + is equivalent to `S n`, which completes the proof. +-} + + iii : join (Total (cov⁻¹ (-1+ n))) Bool ≃ S (ℕ→ℕ₋₁ n) + iii = join (Total (cov⁻¹ (-1+ n))) Bool ≃⟨ invEquiv Susp≃joinBool ⟩ + Susp (Total (cov⁻¹ (-1+ n))) ≃⟨ congSuspEquiv (TotalCov≃Sn (-1+ n)) ⟩ + S (ℕ→ℕ₋₁ n) ■ + +-- the usual covering map S n → RP n, with fibers exactly cov⁻¹ + +cov : (n : ℕ₋₁) → S n → RP n +cov n = pr (cov⁻¹ n) ∘ invEq (TotalCov≃Sn n) + +fibcov≡cov⁻¹ : ∀ n (x : RP n) → fiber (cov n) x ≡ cov⁻¹ n x .fst +fibcov≡cov⁻¹ n x = + fiber (cov n) x ≡[ i ]⟨ fiber {A = ua e i} (pr (cov⁻¹ n) ∘ ua-unglue e i) x ⟩ + fiber (pr (cov⁻¹ n)) x ≡⟨ ua (fibPrEquiv (cov⁻¹ n) x) ⟩ + cov⁻¹ n x .fst ∎ + where e = invEquiv (TotalCov≃Sn n) + + +-------------------------------------------------------------------------------- + +-- Finally, we state the trivial equivalences for RP 0 and RP 1 (Example III.3 in [BR17]) + +RP0≃Unit : RP 0 ≃ Unit +RP0≃Unit = isoToEquiv (iso (λ _ → tt) (λ _ → inr tt) (λ _ → refl) (λ { (inr tt) → refl })) + +RP1≡S1 : RP 1 ≡ S 1 +RP1≡S1 = Pushout {A = Total (cov⁻¹ 0)} {B = RP 0} (pr (cov⁻¹ 0)) (λ _ → tt) ≡⟨ i ⟩ + Pushout {A = Total (cov⁻¹ 0)} {B = Unit} (λ _ → tt) (λ _ → tt) ≡⟨ ii ⟩ + Pushout {A = S 0} {B = Unit} (λ _ → tt) (λ _ → tt) ≡⟨ PushoutSusp≡Susp ⟩ + S 1 ∎ + where i = λ i → Pushout {A = Total (cov⁻¹ 0)} + {B = ua RP0≃Unit i} + (λ x → ua-gluePt RP0≃Unit i (pr (cov⁻¹ 0) x)) + (λ _ → tt) + ii = λ j → Pushout {A = ua (TotalCov≃Sn 0) j} (λ _ → tt) (λ _ → tt) diff --git a/Cubical/HITs/Rational.agda b/Cubical/HITs/Rational.agda deleted file mode 100644 index 4a5e17dcc..000000000 --- a/Cubical/HITs/Rational.agda +++ /dev/null @@ -1,6 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Rational where - -open import Cubical.HITs.Rational.Base public - --- open import Cubical.HITs.Rational.Properties public diff --git a/Cubical/HITs/Rational/Base.agda b/Cubical/HITs/Rational/Base.agda deleted file mode 100644 index a2275e2d2..000000000 --- a/Cubical/HITs/Rational/Base.agda +++ /dev/null @@ -1,18 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Rational.Base where - -open import Cubical.Relation.Nullary -open import Cubical.Core.Everything -open import Cubical.Foundations.Prelude -open import Cubical.HITs.Ints.QuoInt -open import Cubical.Data.Nat -open import Cubical.Data.Empty -open import Cubical.Data.Unit - -data ℚ : Type₀ where - con : (u : ℤ) (a : ℤ) → ¬ (a ≡ pos 0) → ℚ - path : ∀ u a v b {p q} → (u *ℤ b) ≡ (v *ℤ a) → con u a p ≡ con v b q - trunc : isSet ℚ - -int : ℤ → ℚ -int z = con z (pos 1) \ p → snotz (cong abs p) diff --git a/Cubical/HITs/Rationals/HITQ.agda b/Cubical/HITs/Rationals/HITQ.agda new file mode 100644 index 000000000..ce7734b93 --- /dev/null +++ b/Cubical/HITs/Rationals/HITQ.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.HITQ where + +open import Cubical.HITs.Rationals.HITQ.Base public + +-- open import Cubical.HITs.Rationals.HITQ.Properties public diff --git a/Cubical/HITs/Rationals/HITQ/Base.agda b/Cubical/HITs/Rationals/HITQ/Base.agda new file mode 100644 index 000000000..f93a022ae --- /dev/null +++ b/Cubical/HITs/Rationals/HITQ/Base.agda @@ -0,0 +1,36 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.HITQ.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Relation.Nullary + +open import Cubical.Data.Int + +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.NatPlusOne +open import Cubical.Data.Sigma + + +-- ℚ as a higher inductive type + +data ℚ : Type₀ where + con : (u : ℤ) (a : ℤ) → ¬ (a ≡ pos 0) → ℚ + path : ∀ u a v b {p q} → (u · b) ≡ (v · a) → con u a p ≡ con v b q + trunc : isSet ℚ + +[_] : ℤ × ℕ₊₁ → ℚ +[ a , 1+ b ] = con a (pos (suc b)) (ℕ.snotz ∘ cong abs) + + +-- Natural number and negative integer literals for ℚ + +open import Cubical.Data.Nat.Literals public + +instance + fromNatℚ : HasFromNat ℚ + fromNatℚ = record { Constraint = λ _ → Unit ; fromNat = λ n → [ pos n , 1 ] } + +instance + fromNegℚ : HasFromNeg ℚ + fromNegℚ = record { Constraint = λ _ → Unit ; fromNeg = λ n → [ neg n , 1 ] } diff --git a/Cubical/HITs/Rationals/QuoQ.agda b/Cubical/HITs/Rationals/QuoQ.agda new file mode 100644 index 000000000..3cd943266 --- /dev/null +++ b/Cubical/HITs/Rationals/QuoQ.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.QuoQ where + +open import Cubical.HITs.Rationals.QuoQ.Base public + +open import Cubical.HITs.Rationals.QuoQ.Properties public diff --git a/Cubical/HITs/Rationals/QuoQ/Base.agda b/Cubical/HITs/Rationals/QuoQ/Base.agda new file mode 100644 index 000000000..5af5d6d57 --- /dev/null +++ b/Cubical/HITs/Rationals/QuoQ/Base.agda @@ -0,0 +1,76 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.QuoQ.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Nat as ℕ using (discreteℕ) +open import Cubical.Data.NatPlusOne +open import Cubical.Data.Sigma +open import Cubical.Data.Int.MoreInts.QuoInt + +open import Cubical.HITs.SetQuotients as SetQuotient + using ([_]; eq/; discreteSetQuotients) renaming (_/_ to _//_) public + +open import Cubical.Relation.Nullary +open import Cubical.Relation.Binary.Base +open BinaryRelation + +ℕ₊₁→ℤ : ℕ₊₁ → ℤ +ℕ₊₁→ℤ n = pos (ℕ₊₁→ℕ n) + +private + ℕ₊₁→ℤ-hom : ∀ m n → ℕ₊₁→ℤ (m ·₊₁ n) ≡ ℕ₊₁→ℤ m · ℕ₊₁→ℤ n + ℕ₊₁→ℤ-hom _ _ = refl + + +-- ℚ as a set quotient of ℤ × ℕ₊₁ (as in the HoTT book) + +_∼_ : ℤ × ℕ₊₁ → ℤ × ℕ₊₁ → Type₀ +(a , b) ∼ (c , d) = a · ℕ₊₁→ℤ d ≡ c · ℕ₊₁→ℤ b + +ℚ : Type₀ +ℚ = (ℤ × ℕ₊₁) // _∼_ + + +isSetℚ : isSet ℚ +isSetℚ = SetQuotient.squash/ + +[_/_] : ℤ → ℕ₊₁ → ℚ +[ a / b ] = [ a , b ] + + +isEquivRel∼ : isEquivRel _∼_ +isEquivRel.reflexive isEquivRel∼ (a , b) = refl +isEquivRel.symmetric isEquivRel∼ (a , b) (c , d) = sym +isEquivRel.transitive isEquivRel∼ (a , b) (c , d) (e , f) p q = ·-injʳ _ _ _ r + where r = (a · ℕ₊₁→ℤ f) · ℕ₊₁→ℤ d ≡[ i ]⟨ ·-comm a (ℕ₊₁→ℤ f) i · ℕ₊₁→ℤ d ⟩ + (ℕ₊₁→ℤ f · a) · ℕ₊₁→ℤ d ≡⟨ sym (·-assoc (ℕ₊₁→ℤ f) a (ℕ₊₁→ℤ d)) ⟩ + ℕ₊₁→ℤ f · (a · ℕ₊₁→ℤ d) ≡[ i ]⟨ ℕ₊₁→ℤ f · p i ⟩ + ℕ₊₁→ℤ f · (c · ℕ₊₁→ℤ b) ≡⟨ ·-assoc (ℕ₊₁→ℤ f) c (ℕ₊₁→ℤ b) ⟩ + (ℕ₊₁→ℤ f · c) · ℕ₊₁→ℤ b ≡[ i ]⟨ ·-comm (ℕ₊₁→ℤ f) c i · ℕ₊₁→ℤ b ⟩ + (c · ℕ₊₁→ℤ f) · ℕ₊₁→ℤ b ≡[ i ]⟨ q i · ℕ₊₁→ℤ b ⟩ + (e · ℕ₊₁→ℤ d) · ℕ₊₁→ℤ b ≡⟨ sym (·-assoc e (ℕ₊₁→ℤ d) (ℕ₊₁→ℤ b)) ⟩ + e · (ℕ₊₁→ℤ d · ℕ₊₁→ℤ b) ≡[ i ]⟨ e · ·-comm (ℕ₊₁→ℤ d) (ℕ₊₁→ℤ b) i ⟩ + e · (ℕ₊₁→ℤ b · ℕ₊₁→ℤ d) ≡⟨ ·-assoc e (ℕ₊₁→ℤ b) (ℕ₊₁→ℤ d) ⟩ + (e · ℕ₊₁→ℤ b) · ℕ₊₁→ℤ d ∎ + +eq/⁻¹ : ∀ x y → Path ℚ [ x ] [ y ] → x ∼ y +eq/⁻¹ = SetQuotient.effective (λ _ _ → isSetℤ _ _) isEquivRel∼ + +discreteℚ : Discrete ℚ +discreteℚ = discreteSetQuotients (discreteΣ discreteℤ (λ _ → subst Discrete 1+Path discreteℕ)) + (λ _ _ → isSetℤ _ _) isEquivRel∼ (λ _ _ → discreteℤ _ _) + + +-- Natural number and negative integer literals for ℚ + +open import Cubical.Data.Nat.Literals public + +instance + fromNatℚ : HasFromNat ℚ + fromNatℚ = record { Constraint = λ _ → Unit ; fromNat = λ n → [ pos n / 1 ] } + +instance + fromNegℚ : HasFromNeg ℚ + fromNegℚ = record { Constraint = λ _ → Unit ; fromNeg = λ n → [ neg n / 1 ] } diff --git a/Cubical/HITs/Rationals/QuoQ/Properties.agda b/Cubical/HITs/Rationals/QuoQ/Properties.agda new file mode 100644 index 000000000..cc8932b22 --- /dev/null +++ b/Cubical/HITs/Rationals/QuoQ/Properties.agda @@ -0,0 +1,245 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.QuoQ.Properties where + +open import Cubical.Foundations.Everything hiding (_⁻¹) + +open import Cubical.Data.Int.MoreInts.QuoInt as ℤ using (ℤ; Sign; signed; pos; neg; posneg; sign) +open import Cubical.HITs.SetQuotients as SetQuotient using () renaming (_/_ to _//_) + +open import Cubical.Data.Nat as ℕ using (ℕ; zero; suc) +open import Cubical.Data.NatPlusOne +open import Cubical.Data.Sigma + +open import Cubical.Data.Sum +open import Cubical.Relation.Nullary + +open import Cubical.HITs.Rationals.QuoQ.Base + +ℚ-cancelˡ : ∀ {a b} (c : ℕ₊₁) → [ ℕ₊₁→ℤ c ℤ.· a / c ·₊₁ b ] ≡ [ a / b ] +ℚ-cancelˡ {a} {b} c = eq/ _ _ + (cong (ℤ._· ℕ₊₁→ℤ b) (ℤ.·-comm (ℕ₊₁→ℤ c) a) ∙ sym (ℤ.·-assoc a (ℕ₊₁→ℤ c) (ℕ₊₁→ℤ b))) + +ℚ-cancelʳ : ∀ {a b} (c : ℕ₊₁) → [ a ℤ.· ℕ₊₁→ℤ c / b ·₊₁ c ] ≡ [ a / b ] +ℚ-cancelʳ {a} {b} c = eq/ _ _ + (sym (ℤ.·-assoc a (ℕ₊₁→ℤ c) (ℕ₊₁→ℤ b)) ∙ cong (a ℤ.·_) (ℤ.·-comm (ℕ₊₁→ℤ c) (ℕ₊₁→ℤ b))) + +-- useful functions for defining operations on ℚ + +onCommonDenom : + (g : ℤ × ℕ₊₁ → ℤ × ℕ₊₁ → ℤ) + (g-eql : ∀ ((a , b) (c , d) (e , f) : ℤ × ℕ₊₁) (p : a ℤ.· ℕ₊₁→ℤ d ≡ c ℤ.· ℕ₊₁→ℤ b) + → ℕ₊₁→ℤ d ℤ.· (g (a , b) (e , f)) ≡ ℕ₊₁→ℤ b ℤ.· (g (c , d) (e , f))) + (g-eqr : ∀ ((a , b) (c , d) (e , f) : ℤ × ℕ₊₁) (p : c ℤ.· ℕ₊₁→ℤ f ≡ e ℤ.· ℕ₊₁→ℤ d) + → (g (a , b) (c , d)) ℤ.· ℕ₊₁→ℤ f ≡ (g (a , b) (e , f)) ℤ.· ℕ₊₁→ℤ d) + → ℚ → ℚ → ℚ +onCommonDenom g g-eql g-eqr = SetQuotient.rec2 isSetℚ + (λ { (a , b) (c , d) → [ g (a , b) (c , d) / b ·₊₁ d ] }) + (λ { (a , b) (c , d) (e , f) p → eql (a , b) (c , d) (e , f) p }) + (λ { (a , b) (c , d) (e , f) p → eqr (a , b) (c , d) (e , f) p }) + where eql : ∀ ((a , b) (c , d) (e , f) : ℤ × ℕ₊₁) (p : a ℤ.· ℕ₊₁→ℤ d ≡ c ℤ.· ℕ₊₁→ℤ b) + → [ g (a , b) (e , f) / b ·₊₁ f ] ≡ [ g (c , d) (e , f) / d ·₊₁ f ] + eql (a , b) (c , d) (e , f) p = + [ g (a , b) (e , f) / b ·₊₁ f ] + ≡⟨ sym (ℚ-cancelˡ d) ⟩ + [ ℕ₊₁→ℤ d ℤ.· (g (a , b) (e , f)) / d ·₊₁ (b ·₊₁ f) ] + ≡[ i ]⟨ [ ℕ₊₁→ℤ d ℤ.· (g (a , b) (e , f)) / ·₊₁-assoc d b f i ] ⟩ + [ ℕ₊₁→ℤ d ℤ.· (g (a , b) (e , f)) / (d ·₊₁ b) ·₊₁ f ] + ≡[ i ]⟨ [ g-eql (a , b) (c , d) (e , f) p i / ·₊₁-comm d b i ·₊₁ f ] ⟩ + [ ℕ₊₁→ℤ b ℤ.· (g (c , d) (e , f)) / (b ·₊₁ d) ·₊₁ f ] + ≡[ i ]⟨ [ ℕ₊₁→ℤ b ℤ.· (g (c , d) (e , f)) / ·₊₁-assoc b d f (~ i) ] ⟩ + [ ℕ₊₁→ℤ b ℤ.· (g (c , d) (e , f)) / b ·₊₁ (d ·₊₁ f) ] + ≡⟨ ℚ-cancelˡ b ⟩ + [ g (c , d) (e , f) / d ·₊₁ f ] ∎ + eqr : ∀ ((a , b) (c , d) (e , f) : ℤ × ℕ₊₁) (p : c ℤ.· ℕ₊₁→ℤ f ≡ e ℤ.· ℕ₊₁→ℤ d) + → [ g (a , b) (c , d) / b ·₊₁ d ] ≡ [ g (a , b) (e , f) / b ·₊₁ f ] + eqr (a , b) (c , d) (e , f) p = + [ g (a , b) (c , d) / b ·₊₁ d ] + ≡⟨ sym (ℚ-cancelʳ f) ⟩ + [ (g (a , b) (c , d)) ℤ.· ℕ₊₁→ℤ f / (b ·₊₁ d) ·₊₁ f ] + ≡[ i ]⟨ [ (g (a , b) (c , d)) ℤ.· ℕ₊₁→ℤ f / ·₊₁-assoc b d f (~ i) ] ⟩ + [ (g (a , b) (c , d)) ℤ.· ℕ₊₁→ℤ f / b ·₊₁ (d ·₊₁ f) ] + ≡[ i ]⟨ [ g-eqr (a , b) (c , d) (e , f) p i / b ·₊₁ ·₊₁-comm d f i ] ⟩ + [ (g (a , b) (e , f)) ℤ.· ℕ₊₁→ℤ d / b ·₊₁ (f ·₊₁ d) ] + ≡[ i ]⟨ [ (g (a , b) (e , f)) ℤ.· ℕ₊₁→ℤ d / ·₊₁-assoc b f d i ] ⟩ + [ (g (a , b) (e , f)) ℤ.· ℕ₊₁→ℤ d / (b ·₊₁ f) ·₊₁ d ] + ≡⟨ ℚ-cancelʳ d ⟩ + [ g (a , b) (e , f) / b ·₊₁ f ] ∎ + +onCommonDenomSym : + (g : ℤ × ℕ₊₁ → ℤ × ℕ₊₁ → ℤ) + (g-sym : ∀ x y → g x y ≡ g y x) + (g-eql : ∀ ((a , b) (c , d) (e , f) : ℤ × ℕ₊₁) (p : a ℤ.· ℕ₊₁→ℤ d ≡ c ℤ.· ℕ₊₁→ℤ b) + → ℕ₊₁→ℤ d ℤ.· (g (a , b) (e , f)) ≡ ℕ₊₁→ℤ b ℤ.· (g (c , d) (e , f))) + → ℚ → ℚ → ℚ +onCommonDenomSym g g-sym g-eql = onCommonDenom g g-eql q-eqr + where q-eqr : ∀ ((a , b) (c , d) (e , f) : ℤ × ℕ₊₁) (p : c ℤ.· ℕ₊₁→ℤ f ≡ e ℤ.· ℕ₊₁→ℤ d) + → (g (a , b) (c , d)) ℤ.· ℕ₊₁→ℤ f ≡ (g (a , b) (e , f)) ℤ.· ℕ₊₁→ℤ d + q-eqr (a , b) (c , d) (e , f) p = + (g (a , b) (c , d)) ℤ.· ℕ₊₁→ℤ f ≡[ i ]⟨ ℤ.·-comm (g-sym (a , b) (c , d) i) (ℕ₊₁→ℤ f) i ⟩ + ℕ₊₁→ℤ f ℤ.· (g (c , d) (a , b)) ≡⟨ g-eql (c , d) (e , f) (a , b) p ⟩ + ℕ₊₁→ℤ d ℤ.· (g (e , f) (a , b)) ≡[ i ]⟨ ℤ.·-comm (ℕ₊₁→ℤ d) (g-sym (e , f) (a , b) i) i ⟩ + (g (a , b) (e , f)) ℤ.· ℕ₊₁→ℤ d ∎ + +onCommonDenomSym-comm : ∀ {g} g-sym {g-eql} (x y : ℚ) + → onCommonDenomSym g g-sym g-eql x y ≡ + onCommonDenomSym g g-sym g-eql y x +onCommonDenomSym-comm g-sym = SetQuotient.elimProp2 (λ _ _ → isSetℚ _ _) + (λ { (a , b) (c , d) i → [ g-sym (a , b) (c , d) i / ·₊₁-comm b d i ] }) + + +-- basic arithmetic operations on ℚ + +infixl 6 _+_ +infixl 7 _·_ + +private + lem₁ : ∀ a b c d e (p : a ℤ.· b ≡ c ℤ.· d) → b ℤ.· (a ℤ.· e) ≡ d ℤ.· (c ℤ.· e) + lem₁ a b c d e p = ℤ.·-assoc b a e + ∙ cong (ℤ._· e) (ℤ.·-comm b a ∙ p ∙ ℤ.·-comm c d) + ∙ sym (ℤ.·-assoc d c e) + + lem₂ : ∀ a b c → a ℤ.· (b ℤ.· c) ≡ c ℤ.· (b ℤ.· a) + lem₂ a b c = cong (a ℤ.·_) (ℤ.·-comm b c) ∙ ℤ.·-assoc a c b + ∙ cong (ℤ._· b) (ℤ.·-comm a c) ∙ sym (ℤ.·-assoc c a b) + ∙ cong (c ℤ.·_) (ℤ.·-comm a b) + +_+_ : ℚ → ℚ → ℚ +_+_ = onCommonDenomSym + (λ { (a , b) (c , d) → a ℤ.· (ℕ₊₁→ℤ d) ℤ.+ c ℤ.· (ℕ₊₁→ℤ b) }) + (λ { (a , b) (c , d) → ℤ.+-comm (a ℤ.· (ℕ₊₁→ℤ d)) (c ℤ.· (ℕ₊₁→ℤ b)) }) + (λ { (a , b) (c , d) (e , f) p → + ℕ₊₁→ℤ d ℤ.· (a ℤ.· ℕ₊₁→ℤ f ℤ.+ e ℤ.· ℕ₊₁→ℤ b) + ≡⟨ sym (ℤ.·-distribˡ (ℕ₊₁→ℤ d) (a ℤ.· ℕ₊₁→ℤ f) (e ℤ.· ℕ₊₁→ℤ b)) ⟩ + ℕ₊₁→ℤ d ℤ.· (a ℤ.· ℕ₊₁→ℤ f) ℤ.+ ℕ₊₁→ℤ d ℤ.· (e ℤ.· ℕ₊₁→ℤ b) + ≡[ i ]⟨ lem₁ a (ℕ₊₁→ℤ d) c (ℕ₊₁→ℤ b) (ℕ₊₁→ℤ f) p i ℤ.+ lem₂ (ℕ₊₁→ℤ d) e (ℕ₊₁→ℤ b) i ⟩ + ℕ₊₁→ℤ b ℤ.· (c ℤ.· ℕ₊₁→ℤ f) ℤ.+ ℕ₊₁→ℤ b ℤ.· (e ℤ.· ℕ₊₁→ℤ d) + ≡⟨ ℤ.·-distribˡ (ℕ₊₁→ℤ b) (c ℤ.· ℕ₊₁→ℤ f) (e ℤ.· ℕ₊₁→ℤ d) ⟩ + ℕ₊₁→ℤ b ℤ.· (c ℤ.· ℕ₊₁→ℤ f ℤ.+ e ℤ.· ℕ₊₁→ℤ d) ∎ }) + ++-comm : ∀ x y → x + y ≡ y + x ++-comm = onCommonDenomSym-comm + (λ { (a , b) (c , d) → ℤ.+-comm (a ℤ.· (ℕ₊₁→ℤ d)) (c ℤ.· (ℕ₊₁→ℤ b)) }) + ++-identityˡ : ∀ x → 0 + x ≡ x ++-identityˡ = SetQuotient.elimProp (λ _ → isSetℚ _ _) + (λ { (a , b) i → [ ℤ.·-identityʳ a i / ·₊₁-identityˡ b i ] }) + ++-identityʳ : ∀ x → x + 0 ≡ x ++-identityʳ x = +-comm x _ ∙ +-identityˡ x + ++-assoc : ∀ x y z → x + (y + z) ≡ (x + y) + z ++-assoc = SetQuotient.elimProp3 (λ _ _ _ → isSetℚ _ _) + (λ { (a , b) (c , d) (e , f) i → [ eq a (ℕ₊₁→ℤ b) c (ℕ₊₁→ℤ d) e (ℕ₊₁→ℤ f) i / ·₊₁-assoc b d f i ] }) + where eq₁ : ∀ a b c → (a ℤ.· b) ℤ.· c ≡ a ℤ.· (c ℤ.· b) + eq₁ a b c = sym (ℤ.·-assoc a b c) ∙ cong (a ℤ.·_) (ℤ.·-comm b c) + eq₂ : ∀ a b c → (a ℤ.· b) ℤ.· c ≡ (a ℤ.· c) ℤ.· b + eq₂ a b c = eq₁ a b c ∙ ℤ.·-assoc a c b + + eq : ∀ a b c d e f → Path ℤ _ _ + eq a b c d e f = + a ℤ.· (d ℤ.· f) ℤ.+ (c ℤ.· f ℤ.+ e ℤ.· d) ℤ.· b + ≡[ i ]⟨ a ℤ.· (d ℤ.· f) ℤ.+ ℤ.·-distribʳ (c ℤ.· f) (e ℤ.· d) b (~ i) ⟩ + a ℤ.· (d ℤ.· f) ℤ.+ ((c ℤ.· f) ℤ.· b ℤ.+ (e ℤ.· d) ℤ.· b) + ≡[ i ]⟨ ℤ.+-assoc (ℤ.·-assoc a d f i) (eq₂ c f b i) (eq₁ e d b i) i ⟩ + ((a ℤ.· d) ℤ.· f ℤ.+ (c ℤ.· b) ℤ.· f) ℤ.+ e ℤ.· (b ℤ.· d) + ≡[ i ]⟨ ℤ.·-distribʳ (a ℤ.· d) (c ℤ.· b) f i ℤ.+ e ℤ.· (b ℤ.· d) ⟩ + (a ℤ.· d ℤ.+ c ℤ.· b) ℤ.· f ℤ.+ e ℤ.· (b ℤ.· d) ∎ + + +_·_ : ℚ → ℚ → ℚ +_·_ = onCommonDenomSym + (λ { (a , _) (c , _) → a ℤ.· c }) + (λ { (a , _) (c , _) → ℤ.·-comm a c }) + (λ { (a , b) (c , d) (e , _) p → lem₁ a (ℕ₊₁→ℤ d) c (ℕ₊₁→ℤ b) e p }) + +·-comm : ∀ x y → x · y ≡ y · x +·-comm = onCommonDenomSym-comm (λ { (a , _) (c , _) → ℤ.·-comm a c }) + +·-identityˡ : ∀ x → 1 · x ≡ x +·-identityˡ = SetQuotient.elimProp (λ _ → isSetℚ _ _) + (λ { (a , b) i → [ ℤ.·-identityˡ a i / ·₊₁-identityˡ b i ] }) + +·-identityʳ : ∀ x → x · 1 ≡ x +·-identityʳ = SetQuotient.elimProp (λ _ → isSetℚ _ _) + (λ { (a , b) i → [ ℤ.·-identityʳ a i / ·₊₁-identityʳ b i ] }) + +·-zeroˡ : ∀ x → 0 · x ≡ 0 +·-zeroˡ = SetQuotient.elimProp (λ _ → isSetℚ _ _) + (λ { (a , b) → (λ i → [ p a b i / 1 ·₊₁ b ]) ∙ ℚ-cancelʳ b }) + where p : ∀ a b → 0 ℤ.· a ≡ 0 ℤ.· ℕ₊₁→ℤ b + p a b = ℤ.·-zeroˡ {ℤ.spos} a ∙ sym (ℤ.·-zeroˡ {ℤ.spos} (ℕ₊₁→ℤ b)) + +·-zeroʳ : ∀ x → x · 0 ≡ 0 +·-zeroʳ = SetQuotient.elimProp (λ _ → isSetℚ _ _) + (λ { (a , b) → (λ i → [ p a b i / b ·₊₁ 1 ]) ∙ ℚ-cancelˡ b }) + where p : ∀ a b → a ℤ.· 0 ≡ ℕ₊₁→ℤ b ℤ.· 0 + p a b = ℤ.·-zeroʳ {ℤ.spos} a ∙ sym (ℤ.·-zeroʳ {ℤ.spos} (ℕ₊₁→ℤ b)) + +·-assoc : ∀ x y z → x · (y · z) ≡ (x · y) · z +·-assoc = SetQuotient.elimProp3 (λ _ _ _ → isSetℚ _ _) + (λ { (a , b) (c , d) (e , f) i → [ ℤ.·-assoc a c e i / ·₊₁-assoc b d f i ] }) + +·-distribˡ : ∀ x y z → (x · y) + (x · z) ≡ x · (y + z) +·-distribˡ = SetQuotient.elimProp3 (λ _ _ _ → isSetℚ _ _) + (λ { (a , b) (c , d) (e , f) → eq a b c d e f }) + where lem : ∀ {ℓ} {A : Type ℓ} (_·_ : A → A → A) + (·-comm : ∀ x y → x · y ≡ y · x) + (·-assoc : ∀ x y z → x · (y · z) ≡ (x · y) · z) + a c b d + → (a · c) · (b · d) ≡ (a · (c · d)) · b + lem _·_ ·-comm ·-assoc a c b d = + (a · c) · (b · d) ≡[ i ]⟨ (a · c) · ·-comm b d i ⟩ + (a · c) · (d · b) ≡⟨ ·-assoc (a · c) d b ⟩ + ((a · c) · d) · b ≡[ i ]⟨ ·-assoc a c d (~ i) · b ⟩ + (a · (c · d)) · b ∎ + + lemℤ = lem ℤ._·_ ℤ.·-comm ℤ.·-assoc + lemℕ₊₁ = lem _·₊₁_ ·₊₁-comm ·₊₁-assoc + + eq : ∀ a b c d e f → + [ (a ℤ.· c) ℤ.· ℕ₊₁→ℤ (b ·₊₁ f) ℤ.+ (a ℤ.· e) ℤ.· ℕ₊₁→ℤ (b ·₊₁ d) + / (b ·₊₁ d) ·₊₁ (b ·₊₁ f) ] + ≡ [ a ℤ.· (c ℤ.· ℕ₊₁→ℤ f ℤ.+ e ℤ.· ℕ₊₁→ℤ d) + / b ·₊₁ (d ·₊₁ f) ] + eq a b c d e f = + (λ i → [ lemℤ a c (ℕ₊₁→ℤ b) (ℕ₊₁→ℤ f) i ℤ.+ lemℤ a e (ℕ₊₁→ℤ b) (ℕ₊₁→ℤ d) i + / lemℕ₊₁ b d b f i ]) ∙ + (λ i → [ ℤ.·-distribʳ (a ℤ.· (c ℤ.· ℕ₊₁→ℤ f)) (a ℤ.· (e ℤ.· ℕ₊₁→ℤ d)) (ℕ₊₁→ℤ b) i + / (b ·₊₁ (d ·₊₁ f)) ·₊₁ b ]) ∙ + ℚ-cancelʳ {a ℤ.· (c ℤ.· ℕ₊₁→ℤ f) ℤ.+ a ℤ.· (e ℤ.· ℕ₊₁→ℤ d)} {b ·₊₁ (d ·₊₁ f)} b ∙ + (λ i → [ ℤ.·-distribˡ a (c ℤ.· ℕ₊₁→ℤ f) (e ℤ.· ℕ₊₁→ℤ d) i + / b ·₊₁ (d ·₊₁ f) ]) + +·-distribʳ : ∀ x y z → (x · z) + (y · z) ≡ (x + y) · z +·-distribʳ x y z = (λ i → ·-comm x z i + ·-comm y z i) ∙ ·-distribˡ z x y ∙ ·-comm z (x + y) + + +-_ : ℚ → ℚ +- x = -1 · x + +negate-invol : ∀ x → - - x ≡ x +negate-invol x = ·-assoc -1 -1 x ∙ ·-identityˡ x + +negateEquiv : ℚ ≃ ℚ +negateEquiv = isoToEquiv (iso -_ -_ negate-invol negate-invol) + +negateEq : ℚ ≡ ℚ +negateEq = ua negateEquiv + ++-inverseˡ : ∀ x → (- x) + x ≡ 0 ++-inverseˡ x = (λ i → (-1 · x) + ·-identityˡ x (~ i)) ∙ ·-distribʳ -1 1 x ∙ ·-zeroˡ x + +_-_ : ℚ → ℚ → ℚ +x - y = x + (- y) + ++-inverseʳ : ∀ x → x - x ≡ 0 ++-inverseʳ x = +-comm x (- x) ∙ +-inverseˡ x + ++-injˡ : ∀ x y z → x + y ≡ x + z → y ≡ z ++-injˡ x y z p = sym (q y) ∙ cong ((- x) +_) p ∙ q z + where q : ∀ y → (- x) + (x + y) ≡ y + q y = +-assoc (- x) x y ∙ cong (_+ y) (+-inverseˡ x) ∙ +-identityˡ y + ++-injʳ : ∀ x y z → x + y ≡ z + y → x ≡ z ++-injʳ x y z p = +-injˡ y x z (+-comm y x ∙ p ∙ +-comm z y) + diff --git a/Cubical/HITs/Rationals/SigmaQ.agda b/Cubical/HITs/Rationals/SigmaQ.agda new file mode 100644 index 000000000..46a9efbd4 --- /dev/null +++ b/Cubical/HITs/Rationals/SigmaQ.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.SigmaQ where + +open import Cubical.HITs.Rationals.SigmaQ.Base public + +open import Cubical.HITs.Rationals.SigmaQ.Properties public diff --git a/Cubical/HITs/Rationals/SigmaQ/Base.agda b/Cubical/HITs/Rationals/SigmaQ/Base.agda new file mode 100644 index 000000000..8dda85f8b --- /dev/null +++ b/Cubical/HITs/Rationals/SigmaQ/Base.agda @@ -0,0 +1,54 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.SigmaQ.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Int.MoreInts.QuoInt + +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.NatPlusOne +open import Cubical.Data.Sigma + +open import Cubical.Data.Nat.GCD +open import Cubical.Data.Nat.Coprime + + +-- ℚ as the set of coprime pairs in ℤ × ℕ₊₁ + +ℚ : Type₀ +ℚ = Σ[ (a , b) ∈ ℤ × ℕ₊₁ ] areCoprime (abs a , ℕ₊₁→ℕ b) + +isSetℚ : isSet ℚ +isSetℚ = isSetΣ (isSet× isSetℤ (subst isSet 1+Path isSetℕ)) (λ _ → isProp→isSet isPropIsGCD) + + +signedPair : Sign → ℕ × ℕ₊₁ → ℤ × ℕ₊₁ +signedPair s (a , b) = (signed s a , b) + +[_] : ℤ × ℕ₊₁ → ℚ +[ signed s a , b ] = signedPair s (toCoprime (a , b)) , toCoprimeAreCoprime (a , b) +[ posneg i , b ] = (posneg i , 1) , toCoprimeAreCoprime (0 , b) + +[]-cancelʳ : ∀ ((a , b) : ℤ × ℕ₊₁) k → [ a · pos (ℕ₊₁→ℕ k) , b ·₊₁ k ] ≡ [ a , b ] +[]-cancelʳ (signed s zero , b) k = + Σ≡Prop (λ _ → isPropIsGCD) (λ i → signed-zero spos s i , 1) +[]-cancelʳ (signed s (suc a) , b) k = + Σ≡Prop (λ _ → isPropIsGCD) (λ i → signedPair (·S-comm s spos i) + (toCoprime-cancelʳ (suc a , b) k i)) +[]-cancelʳ (posneg i , b) k j = + isSet→isSet' isSetℚ ([]-cancelʳ (pos zero , b) k) ([]-cancelʳ (neg zero , b) k) + (λ i → [ posneg i · pos (ℕ₊₁→ℕ k) , b ·₊₁ k ]) (λ i → [ posneg i , b ]) i j + + +-- Natural number and negative integer literals for ℚ + +open import Cubical.Data.Nat.Literals public + +instance + fromNatℚ : HasFromNat ℚ + fromNatℚ = record { Constraint = λ _ → Unit ; fromNat = λ n → (pos n , 1) , oneGCD n } + +instance + fromNegℚ : HasFromNeg ℚ + fromNegℚ = record { Constraint = λ _ → Unit ; fromNeg = λ n → (neg n , 1) , oneGCD n } diff --git a/Cubical/HITs/Rationals/SigmaQ/Properties.agda b/Cubical/HITs/Rationals/SigmaQ/Properties.agda new file mode 100644 index 000000000..8dcedfe4c --- /dev/null +++ b/Cubical/HITs/Rationals/SigmaQ/Properties.agda @@ -0,0 +1,78 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Rationals.SigmaQ.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Int.MoreInts.QuoInt +import Cubical.HITs.SetQuotients as SetQuotient + +open import Cubical.Data.Nat as ℕ hiding (_·_) +open import Cubical.Data.NatPlusOne +open import Cubical.Data.Sigma + +open import Cubical.Data.Nat.GCD +open import Cubical.Data.Nat.Coprime + +open import Cubical.HITs.Rationals.QuoQ as Quo using (ℕ₊₁→ℤ) +import Cubical.HITs.Rationals.SigmaQ.Base as Sigma + + +reduce : Quo.ℚ → Sigma.ℚ +reduce = SetQuotient.rec Sigma.isSetℚ Sigma.[_] + (λ { (a , b) (c , d) p → sym (Sigma.[]-cancelʳ (a , b) d) + ∙ (λ i → Sigma.[ p i , ·₊₁-comm b d i ]) + ∙ Sigma.[]-cancelʳ (c , d) b }) + +private + toCoprime-eq₁ : ∀ s ((a , b) : ℕ × ℕ₊₁) + → signed s (toCoprime (a , b) .fst) · ℕ₊₁→ℤ b + ≡ signed s a · ℕ₊₁→ℤ (toCoprime (a , b) .snd) + toCoprime-eq₁ s (a , b) = + signed s c₁ · ℕ₊₁→ℤ b ≡⟨ ·-signed-pos c₁ (ℕ₊₁→ℕ b) ⟩ + signed s (c₁ ℕ.· ℕ₊₁→ℕ b) ≡[ i ]⟨ signed s (c₁ ℕ.· p₂ (~ i)) ⟩ + signed s (c₁ ℕ.· (ℕ₊₁→ℕ c₂ ℕ.· d)) ≡[ i ]⟨ signed s (c₁ ℕ.· ℕ.·-comm (ℕ₊₁→ℕ c₂) d i) ⟩ + signed s (c₁ ℕ.· (d ℕ.· ℕ₊₁→ℕ c₂)) ≡[ i ]⟨ signed s (ℕ.·-assoc c₁ d (ℕ₊₁→ℕ c₂) i) ⟩ + signed s ((c₁ ℕ.· d) ℕ.· ℕ₊₁→ℕ c₂) ≡[ i ]⟨ signed s (p₁ i ℕ.· ℕ₊₁→ℕ c₂) ⟩ + signed s (a ℕ.· ℕ₊₁→ℕ c₂) ≡⟨ sym (·-signed-pos a (ℕ₊₁→ℕ c₂)) ⟩ + signed s a · ℕ₊₁→ℤ c₂ ∎ + where open ToCoprime (a , b) + +[]-reduce : ∀ x → Quo.[ reduce x .fst ] ≡ x +[]-reduce = SetQuotient.elimProp (λ _ → Quo.isSetℚ _ _) + (λ { (signed s a , b) → Quo.eq/ _ _ (toCoprime-eq₁ s (a , b)) + ; (posneg i , b) j → isSet→isSet' Quo.isSetℚ + (Quo.eq/ (fst (reduce Quo.[ pos 0 , b ])) (pos 0 , b) (toCoprime-eq₁ spos (0 , b))) + (Quo.eq/ (fst (reduce Quo.[ neg 0 , b ])) (neg 0 , b) (toCoprime-eq₁ sneg (0 , b))) + (λ i → Quo.[ reduce (Quo.[ posneg i , b ]) .fst ]) (λ i → Quo.[ posneg i , b ]) i j }) + +private + toCoprime-eq₂ : ∀ s ((a , b) : ℕ × ℕ₊₁) (cp : areCoprime (a , ℕ₊₁→ℕ b)) + → Sigma.signedPair s (toCoprime (a , b)) ≡ Sigma.signedPair s (a , b) + toCoprime-eq₂ s (a , b) cp i = Sigma.signedPair s (toCoprime-idem (a , b) cp i) + +reduce-[] : ∀ x → reduce Quo.[ x .fst ] ≡ x + -- equivalent to: Sigma.[ s .fst ] ≡ x +reduce-[] ((signed s a , b) , cp) = + Σ≡Prop (λ _ → isPropIsGCD) (toCoprime-eq₂ s (a , b) cp) +reduce-[] ((posneg i , b) , cp) j = + isSet→isSet' Sigma.isSetℚ + (Σ≡Prop (λ _ → isPropIsGCD) (toCoprime-eq₂ spos (0 , b) cp)) + (Σ≡Prop (λ _ → isPropIsGCD) (toCoprime-eq₂ sneg (0 , b) cp)) + (λ i → Sigma.[ posneg i , b ]) (λ i → (posneg i , b) , cp) i j + + +Quoℚ-iso-Sigmaℚ : Iso Quo.ℚ Sigma.ℚ +Iso.fun Quoℚ-iso-Sigmaℚ = reduce +Iso.inv Quoℚ-iso-Sigmaℚ = Quo.[_] ∘ fst +Iso.rightInv Quoℚ-iso-Sigmaℚ = reduce-[] +Iso.leftInv Quoℚ-iso-Sigmaℚ = []-reduce + +Quoℚ≃Sigmaℚ : Quo.ℚ ≃ Sigma.ℚ +Quoℚ≃Sigmaℚ = isoToEquiv Quoℚ-iso-Sigmaℚ + +Quoℚ≡Sigmaℚ : Quo.ℚ ≡ Sigma.ℚ +Quoℚ≡Sigmaℚ = isoToPath Quoℚ-iso-Sigmaℚ diff --git a/Cubical/HITs/S1.agda b/Cubical/HITs/S1.agda index c6879da01..9490001b2 100644 --- a/Cubical/HITs/S1.agda +++ b/Cubical/HITs/S1.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.S1 where open import Cubical.HITs.S1.Base public diff --git a/Cubical/HITs/S1/Base.agda b/Cubical/HITs/S1/Base.agda index 37484354d..cfea54dc7 100644 --- a/Cubical/HITs/S1/Base.agda +++ b/Cubical/HITs/S1/Base.agda @@ -3,18 +3,22 @@ Definition of the circle as a HIT with a proof that Ω(S¹) ≡ ℤ -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.S1.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Path open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport open import Cubical.Foundations.Univalence open import Cubical.Data.Nat - hiding (_+_ ; _*_ ; +-assoc ; +-comm) -open import Cubical.Data.Int + hiding (_+_ ; _·_ ; +-assoc ; +-comm) +open import Cubical.Data.Int hiding (_·_) data S¹ : Type₀ where base : S¹ @@ -26,12 +30,13 @@ module _ where transpS¹ φ u0 = refl compS1 : ∀ (φ : I) (u : ∀ i → Partial φ S¹) (u0 : S¹ [ φ ↦ u i0 ]) → - comp (λ _ → S¹) u (outS u0) ≡ hcomp u (outS u0) + comp (λ _ → S¹) (\ i → u i) (outS u0) ≡ hcomp u (outS u0) compS1 φ u u0 = refl +-- ΩS¹ ≡ ℤ helix : S¹ → Type₀ -helix base = Int -helix (loop i) = sucPathInt i +helix base = ℤ +helix (loop i) = sucPathℤ i ΩS¹ : Type₀ ΩS¹ = base ≡ base @@ -39,29 +44,28 @@ helix (loop i) = sucPathInt i encode : ∀ x → base ≡ x → helix x encode x p = subst helix p (pos zero) -winding : ΩS¹ → Int +winding : ΩS¹ → ℤ winding = encode base -intLoop : Int → ΩS¹ +intLoop : ℤ → ΩS¹ intLoop (pos zero) = refl intLoop (pos (suc n)) = intLoop (pos n) ∙ loop intLoop (negsuc zero) = sym loop intLoop (negsuc (suc n)) = intLoop (negsuc n) ∙ sym loop -decodeSquare : (n : Int) → PathP (λ i → base ≡ loop i) (intLoop (predInt n)) (intLoop n) +decodeSquare : (n : ℤ) → PathP (λ i → base ≡ loop i) (intLoop (predℤ n)) (intLoop n) decodeSquare (pos zero) i j = loop (i ∨ ~ j) decodeSquare (pos (suc n)) i j = hfill (λ k → λ { (j = i0) → base ; (j = i1) → loop k } ) (inS (intLoop (pos n) j)) i -decodeSquare (negsuc n) i j = hcomp (λ k → λ { (i = i1) → intLoop (negsuc n) j - ; (j = i0) → base - ; (j = i1) → loop (i ∨ ~ k) }) - (intLoop (negsuc n) j) +decodeSquare (negsuc n) i j = hfill (λ k → λ { (j = i0) → base + ; (j = i1) → loop (~ k) }) + (inS (intLoop (negsuc n) j)) (~ i) decode : (x : S¹) → helix x → base ≡ x decode base = intLoop decode (loop i) y j = - let n : Int + let n : ℤ n = unglue (i ∨ ~ i) y in hcomp (λ k → λ { (i = i0) → intLoop (predSuc y k) j ; (i = i1) → intLoop y j @@ -78,55 +82,69 @@ isSetΩS¹ p q r s j i = ; (i = i1) → decodeEncode base q k ; (j = i0) → decodeEncode base (r i) k ; (j = i1) → decodeEncode base (s i) k }) - (decode base (isSetInt (winding p) (winding q) (cong winding r) (cong winding s) j i)) + (decode base (isSetℤ (winding p) (winding q) (cong winding r) (cong winding s) j i)) -- This proof does not rely on rewriting hcomp with empty systems in --- Int as ghcomp has been implemented! -windingIntLoop : (n : Int) → winding (intLoop n) ≡ n -windingIntLoop (pos zero) = refl -windingIntLoop (pos (suc n)) = cong sucInt (windingIntLoop (pos n)) -windingIntLoop (negsuc zero) = refl -windingIntLoop (negsuc (suc n)) = cong predInt (windingIntLoop (negsuc n)) - -ΩS¹≡Int : ΩS¹ ≡ Int -ΩS¹≡Int = isoToPath (iso winding intLoop windingIntLoop (decodeEncode base)) +-- ℤ as ghcomp has been implemented! +windingℤLoop : (n : ℤ) → winding (intLoop n) ≡ n +windingℤLoop (pos zero) = refl +windingℤLoop (pos (suc n)) = cong sucℤ (windingℤLoop (pos n)) +windingℤLoop (negsuc zero) = refl +windingℤLoop (negsuc (suc n)) = cong predℤ (windingℤLoop (negsuc n)) + +ΩS¹Isoℤ : Iso ΩS¹ ℤ +Iso.fun ΩS¹Isoℤ = winding +Iso.inv ΩS¹Isoℤ = intLoop +Iso.rightInv ΩS¹Isoℤ = windingℤLoop +Iso.leftInv ΩS¹Isoℤ = decodeEncode base + +ΩS¹≡ℤ : ΩS¹ ≡ ℤ +ΩS¹≡ℤ = isoToPath ΩS¹Isoℤ -- intLoop and winding are group homomorphisms private - intLoop-sucInt : (z : Int) → intLoop (sucInt z) ≡ intLoop z ∙ loop - intLoop-sucInt (pos n) = refl - intLoop-sucInt (negsuc zero) = sym (lCancel loop) - intLoop-sucInt (negsuc (suc n)) = + intLoop-sucℤ : (z : ℤ) → intLoop (sucℤ z) ≡ intLoop z ∙ loop + intLoop-sucℤ (pos n) = refl + intLoop-sucℤ (negsuc zero) = sym (lCancel loop) + intLoop-sucℤ (negsuc (suc n)) = rUnit (intLoop (negsuc n)) ∙ (λ i → intLoop (negsuc n) ∙ lCancel loop (~ i)) ∙ assoc (intLoop (negsuc n)) (sym loop) loop - intLoop-predInt : (z : Int) → intLoop (predInt z) ≡ intLoop z ∙ sym loop - intLoop-predInt (pos zero) = lUnit (sym loop) - intLoop-predInt (pos (suc n)) = + intLoop-predℤ : (z : ℤ) → intLoop (predℤ z) ≡ intLoop z ∙ sym loop + intLoop-predℤ (pos zero) = lUnit (sym loop) + intLoop-predℤ (pos (suc n)) = rUnit (intLoop (pos n)) ∙ (λ i → intLoop (pos n) ∙ (rCancel loop (~ i))) ∙ assoc (intLoop (pos n)) loop (sym loop) - intLoop-predInt (negsuc n) = refl + intLoop-predℤ (negsuc n) = refl -intLoop-hom : (a b : Int) → (intLoop a) ∙ (intLoop b) ≡ intLoop (a + b) +intLoop-hom : (a b : ℤ) → (intLoop a) ∙ (intLoop b) ≡ intLoop (a + b) intLoop-hom a (pos zero) = sym (rUnit (intLoop a)) intLoop-hom a (pos (suc n)) = assoc (intLoop a) (intLoop (pos n)) loop ∙ (λ i → (intLoop-hom a (pos n) i) ∙ loop) - ∙ sym (intLoop-sucInt (a + pos n)) -intLoop-hom a (negsuc zero) = sym (intLoop-predInt a) + ∙ sym (intLoop-sucℤ (a + pos n)) +intLoop-hom a (negsuc zero) = sym (intLoop-predℤ a) intLoop-hom a (negsuc (suc n)) = assoc (intLoop a) (intLoop (negsuc n)) (sym loop) ∙ (λ i → (intLoop-hom a (negsuc n) i) ∙ (sym loop)) - ∙ sym (intLoop-predInt (a + negsuc n)) + ∙ sym (intLoop-predℤ (a + negsuc n)) winding-hom : (a b : ΩS¹) → winding (a ∙ b) ≡ (winding a) + (winding b) winding-hom a b i = hcomp (λ t → λ { (i = i0) → winding (decodeEncode base a t ∙ decodeEncode base b t) - ; (i = i1) → windingIntLoop (winding a + winding b) t }) + ; (i = i1) → windingℤLoop (winding a + winding b) t }) (winding (intLoop-hom (winding a) (winding b) i)) +-- Commutativity +comm-ΩS¹ : (p q : ΩS¹) → p ∙ q ≡ q ∙ p +comm-ΩS¹ p q = sym (cong₂ (_∙_) (decodeEncode base p) (decodeEncode base q)) + ∙ intLoop-hom (winding p) (winding q) + ∙ cong intLoop (+Comm (winding p) (winding q)) + ∙ sym (intLoop-hom (winding q) (winding p)) + ∙ (cong₂ (_∙_) (decodeEncode base q) (decodeEncode base p)) + -- Based homotopy group basedΩS¹ : (x : S¹) → Type₀ @@ -148,42 +166,43 @@ private ; (j = i1) → loop (i ∧ (~ t)) }) (inS (x j)) l -ΩS¹→basedΩS¹ : (i : I) → ΩS¹ → basedΩS¹ (loop i) -ΩS¹→basedΩS¹ i x j = ΩS¹→basedΩS¹-filler i1 i x j + ΩS¹→basedΩS¹ : (i : I) → ΩS¹ → basedΩS¹ (loop i) + ΩS¹→basedΩS¹ i x j = ΩS¹→basedΩS¹-filler i1 i x j -basedΩS¹→ΩS¹ : (i : I) → basedΩS¹ (loop i) → ΩS¹ -basedΩS¹→ΩS¹ i x j = basedΩS¹→ΩS¹-filler i1 i x j + basedΩS¹→ΩS¹ : (i : I) → basedΩS¹ (loop i) → ΩS¹ + basedΩS¹→ΩS¹ i x j = basedΩS¹→ΩS¹-filler i1 i x j -basedΩS¹→ΩS¹→basedΩS¹ : (i : I) → (x : basedΩS¹ (loop i)) - → ΩS¹→basedΩS¹ i (basedΩS¹→ΩS¹ i x) ≡ x -basedΩS¹→ΩS¹→basedΩS¹ i x j k = - hcomp (λ t → λ { (j = i1) → basedΩS¹→ΩS¹-filler (~ t) i x k - ; (j = i0) → ΩS¹→basedΩS¹ i (basedΩS¹→ΩS¹ i x) k - ; (k = i0) → loop (i ∧ (t ∨ (~ j))) - ; (k = i1) → loop (i ∧ (t ∨ (~ j))) }) - (ΩS¹→basedΩS¹-filler (~ j) i (basedΩS¹→ΩS¹ i x) k) -ΩS¹→basedΩS¹→ΩS¹ : (i : I) → (x : ΩS¹) - → basedΩS¹→ΩS¹ i (ΩS¹→basedΩS¹ i x) ≡ x -ΩS¹→basedΩS¹→ΩS¹ i x j k = - hcomp (λ t → λ { (j = i1) → ΩS¹→basedΩS¹-filler (~ t) i x k - ; (j = i0) → basedΩS¹→ΩS¹ i (ΩS¹→basedΩS¹ i x) k - ; (k = i0) → loop (i ∧ ((~ t) ∧ j)) - ; (k = i1) → loop (i ∧ ((~ t) ∧ j)) }) - (basedΩS¹→ΩS¹-filler (~ j) i (ΩS¹→basedΩS¹ i x) k) --- from the existence of our quasi-inverse, we deduce that the basechange is an equivalence --- for all loop i + basedΩS¹→ΩS¹→basedΩS¹ : (i : I) → (x : basedΩS¹ (loop i)) + → ΩS¹→basedΩS¹ i (basedΩS¹→ΩS¹ i x) ≡ x + basedΩS¹→ΩS¹→basedΩS¹ i x j k = + hcomp (λ t → λ { (j = i1) → basedΩS¹→ΩS¹-filler (~ t) i x k + ; (j = i0) → ΩS¹→basedΩS¹ i (basedΩS¹→ΩS¹ i x) k + ; (k = i0) → loop (i ∧ (t ∨ (~ j))) + ; (k = i1) → loop (i ∧ (t ∨ (~ j))) }) + (ΩS¹→basedΩS¹-filler (~ j) i (basedΩS¹→ΩS¹ i x) k) -basedΩS¹→ΩS¹-isequiv : (i : I) → isEquiv (basedΩS¹→ΩS¹ i) -basedΩS¹→ΩS¹-isequiv i = isoToIsEquiv (iso (basedΩS¹→ΩS¹ i) (ΩS¹→basedΩS¹ i) - (ΩS¹→basedΩS¹→ΩS¹ i) (basedΩS¹→ΩS¹→basedΩS¹ i)) + ΩS¹→basedΩS¹→ΩS¹ : (i : I) → (x : ΩS¹) + → basedΩS¹→ΩS¹ i (ΩS¹→basedΩS¹ i x) ≡ x + ΩS¹→basedΩS¹→ΩS¹ i x j k = + hcomp (λ t → λ { (j = i1) → ΩS¹→basedΩS¹-filler (~ t) i x k + ; (j = i0) → basedΩS¹→ΩS¹ i (ΩS¹→basedΩS¹ i x) k + ; (k = i0) → loop (i ∧ ((~ t) ∧ j)) + ; (k = i1) → loop (i ∧ ((~ t) ∧ j)) }) + (basedΩS¹→ΩS¹-filler (~ j) i (ΩS¹→basedΩS¹ i x) k) + -- from the existence of our quasi-inverse, we deduce that the basechange is an equivalence + -- for all loop i --- now extend the basechange so that both ends match --- (and therefore we get a basechange for any x : S¹) + basedΩS¹→ΩS¹-isequiv : (i : I) → isEquiv (basedΩS¹→ΩS¹ i) + basedΩS¹→ΩS¹-isequiv i = isoToIsEquiv (iso (basedΩS¹→ΩS¹ i) (ΩS¹→basedΩS¹ i) + (ΩS¹→basedΩS¹→ΩS¹ i) (basedΩS¹→ΩS¹→basedΩS¹ i)) + + + -- now extend the basechange so that both ends match + -- (and therefore we get a basechange for any x : S¹) -private loop-conjugation : basedΩS¹→ΩS¹ i1 ≡ λ x → x loop-conjugation i x = let p = (doubleCompPath-elim loop x (sym loop)) @@ -194,10 +213,10 @@ private ∙ (λ t → intLoop (winding-hom (intLoop (pos (suc zero)) ∙ x) (intLoop (negsuc zero)) t)) ∙ (λ t → intLoop ((winding-hom (intLoop (pos (suc zero))) x t) - + (windingIntLoop (negsuc zero) t))) - ∙ (λ t → intLoop (((windingIntLoop (pos (suc zero)) t) + (winding x)) + (negsuc zero))) - ∙ (λ t → intLoop ((+-comm (pos (suc zero)) (winding x) t) + (negsuc zero))) - ∙ (λ t → intLoop (+-assoc (winding x) (pos (suc zero)) (negsuc zero) (~ t))) + + (windingℤLoop (negsuc zero) t))) + ∙ (λ t → intLoop (((windingℤLoop (pos (suc zero)) t) + (winding x)) + (negsuc zero))) + ∙ (λ t → intLoop ((+Comm (pos (suc zero)) (winding x) t) + (negsuc zero))) + ∙ (λ t → intLoop (+Assoc (winding x) (pos (suc zero)) (negsuc zero) (~ t))) ∙ (decodeEncode base x)) i refl-conjugation : basedΩS¹→ΩS¹ i0 ≡ λ x → x @@ -235,12 +254,102 @@ private (basechange-isequiv-aux i0) t }) (basechange-isequiv-aux i) - basedΩS¹≡ΩS¹ : (x : S¹) → basedΩS¹ x ≡ ΩS¹ - basedΩS¹≡ΩS¹ x = ua (basechange x , basechange-isequiv x) - -basedΩS¹≡Int : (x : S¹) → basedΩS¹ x ≡ Int -basedΩS¹≡Int x = (basedΩS¹≡ΩS¹ x) ∙ ΩS¹≡Int - + basedΩS¹≡ΩS¹' : (x : S¹) → basedΩS¹ x ≡ ΩS¹ + basedΩS¹≡ΩS¹' x = ua (basechange x , basechange-isequiv x) + + basedΩS¹≡ℤ' : (x : S¹) → basedΩS¹ x ≡ ℤ + basedΩS¹≡ℤ' x = (basedΩS¹≡ΩS¹' x) ∙ ΩS¹≡ℤ + + +---- Alternative proof of the same thing ----- + +toPropElim : ∀ {ℓ} {B : S¹ → Type ℓ} + → ((s : S¹) → isProp (B s)) + → B base + → (s : S¹) → B s +toPropElim isprop b base = b +toPropElim isprop b (loop i) = + isOfHLevel→isOfHLevelDep 1 isprop b b loop i + +toPropElim2 : ∀ {ℓ} {B : S¹ → S¹ → Type ℓ} + → ((s t : S¹) → isProp (B s t)) + → B base base + → (s t : S¹) → B s t +toPropElim2 isprop b base base = b +toPropElim2 isprop b base (loop i) = + isProp→PathP (λ i → isprop base (loop i)) b b i +toPropElim2 isprop b (loop i) base = + isProp→PathP (λ i → isprop (loop i) base) b b i +toPropElim2 {B = B} isprop b (loop i) (loop j) = + isSet→SquareP (λ _ _ → isOfHLevelSuc 1 (isprop _ _)) + (isProp→PathP (λ i₁ → isprop base (loop i₁)) b b) + (isProp→PathP (λ i₁ → isprop base (loop i₁)) b b) + (isProp→PathP (λ i₁ → isprop (loop i₁) base) b b) + (isProp→PathP (λ i₁ → isprop (loop i₁) base) b b) i j + +toSetElim2 : ∀ {ℓ} {B : S¹ → S¹ → Type ℓ} + → ((s t : S¹) → isSet (B s t)) + → (x : B base base) + → PathP (λ i → B base (loop i)) x x + → PathP (λ i → B (loop i) base) x x + → (s t : S¹) → B s t +toSetElim2 isset b right left base base = b +toSetElim2 isset b right left base (loop i) = right i +toSetElim2 isset b right left (loop i) base = left i +toSetElim2 isset b right left (loop i) (loop j) = + isSet→SquareP (λ _ _ → isset _ _) right right left left i j + +isSetΩx : (x : S¹) → isSet (x ≡ x) +isSetΩx = toPropElim (λ _ → isPropIsSet) isSetΩS¹ + +basechange2 : (x : S¹) → ΩS¹ → (x ≡ x) +basechange2 base p = p +basechange2 (loop i) p = + hcomp (λ k → λ { (i = i0) → lUnit (rUnit p (~ k)) (~ k) + ; (i = i1) → (cong ((sym loop) ∙_) (comm-ΩS¹ p loop) + ∙ assoc (sym loop) loop p + ∙ cong (_∙ p) (lCancel loop) + ∙ sym (lUnit _)) k }) + ((λ j → loop (~ j ∧ i)) ∙ p ∙ λ j → loop (j ∧ i)) +basechange2⁻ : (x : S¹) → (x ≡ x) → ΩS¹ +basechange2⁻ base p = p +basechange2⁻ (loop i) p = + hcomp (λ k → λ { (i = i0) → lUnit (rUnit p (~ k)) (~ k) + ; (i = i1) → (cong (loop ∙_) (comm-ΩS¹ p (sym loop)) + ∙ assoc loop (sym loop) p + ∙ cong (_∙ p) (rCancel loop) + ∙ sym (lUnit _)) k }) + ((λ j → loop (i ∧ j)) ∙ p ∙ λ j → loop (i ∧ (~ j))) +basechange2-sect : (x : S¹) → section (basechange2 x) (basechange2⁻ x) +basechange2-sect = + toPropElim (λ _ → isOfHLevelΠ 1 λ _ → isSetΩx _ _ _ ) + λ _ → refl + +basechange2-retr : (x : S¹) → retract (basechange2 x) (basechange2⁻ x) +basechange2-retr = + toPropElim (λ s → isOfHLevelΠ 1 λ x → isSetΩx _ _ _) + λ _ → refl + +Iso-basedΩS¹-ΩS¹ : (x : S¹) → Iso (basedΩS¹ x) ΩS¹ +Iso.fun (Iso-basedΩS¹-ΩS¹ x) = basechange2⁻ x +Iso.inv (Iso-basedΩS¹-ΩS¹ x) = basechange2 x +Iso.rightInv (Iso-basedΩS¹-ΩS¹ x) = basechange2-retr x +Iso.leftInv (Iso-basedΩS¹-ΩS¹ x) = basechange2-sect x + +Iso-basedΩS¹-ℤ : (x : S¹) → Iso (basedΩS¹ x) ℤ +Iso-basedΩS¹-ℤ x = compIso (Iso-basedΩS¹-ΩS¹ x) ΩS¹Isoℤ + +basedΩS¹≡ℤ : (x : S¹) → basedΩS¹ x ≡ ℤ +basedΩS¹≡ℤ x = isoToPath (Iso-basedΩS¹-ℤ x) + +-- baschange2⁻ is a morphism + +basechange2⁻-morph : (x : S¹) (p q : x ≡ x) → + basechange2⁻ x (p ∙ q) ≡ basechange2⁻ x p ∙ basechange2⁻ x q +basechange2⁻-morph = + toPropElim {B = λ x → (p q : x ≡ x) → basechange2⁻ x (p ∙ q) ≡ basechange2⁻ x p ∙ basechange2⁻ x q} + (λ _ → isPropΠ2 λ _ _ → isSetΩS¹ _ _) + λ _ _ → refl -- Some tests module _ where @@ -253,21 +362,29 @@ module _ where -- the inverse when S¹ is seen as a group -inv : S¹ → S¹ -inv base = base -inv (loop i) = loop (~ i) +invLooper : S¹ → S¹ +invLooper base = base +invLooper (loop i) = loop (~ i) -invInvolutive : section inv inv +invInvolutive : section invLooper invLooper invInvolutive base = refl invInvolutive (loop i) = refl invS¹Equiv : S¹ ≃ S¹ -invS¹Equiv = isoToEquiv (iso inv inv invInvolutive invInvolutive) +invS¹Equiv = isoToEquiv theIso + where + theIso : Iso S¹ S¹ + Iso.fun theIso = invLooper + Iso.inv theIso = invLooper + Iso.rightInv theIso = invInvolutive + Iso.leftInv theIso = invInvolutive invS¹Path : S¹ ≡ S¹ invS¹Path = ua invS¹Equiv -- rot, used in the Hopf fibration +-- we will denote rotation by _·_ +-- it is called μ in the HoTT-book in "8.5.2 The Hopf construction" rotLoop : (a : S¹) → a ≡ a rotLoop base = loop @@ -277,14 +394,11 @@ rotLoop (loop i) j = ; (j = i0) → loop (i ∨ ~ k) ; (j = i1) → loop (i ∧ k)}) base -rot : S¹ → S¹ → S¹ -rot base x = x -rot (loop i) x = rotLoop x i - -_*_ : S¹ → S¹ → S¹ -a * b = rot a b +_·_ : S¹ → S¹ → S¹ +base · x = x +(loop i) · x = rotLoop x i -infixl 30 _*_ +infixl 30 _·_ -- rot i j = filler-rot i j i1 @@ -299,23 +413,16 @@ isPropFamS¹ : ∀ {ℓ} (P : S¹ → Type ℓ) (pP : (x : S¹) → isProp (P x) isPropFamS¹ P pP b0 i = pP (loop i) (transp (λ j → P (loop (i ∧ j))) (~ i) b0) (transp (λ j → P (loop (i ∨ ~ j))) i b0) i -rotIsEquiv : (a : S¹) → isEquiv (rot a) +rotIsEquiv : (a : S¹) → isEquiv (a ·_) rotIsEquiv base = idIsEquiv S¹ -rotIsEquiv (loop i) = isPropFamS¹ (λ x → isEquiv (rot x)) - (λ x → isPropIsEquiv (rot x)) +rotIsEquiv (loop i) = isPropFamS¹ (λ x → isEquiv (x ·_)) + (λ x → isPropIsEquiv (x ·_)) (idIsEquiv _) i -- more direct definition of the rot (loop i) equivalence rotLoopInv : (a : S¹) → PathP (λ i → rotLoop (rotLoop a (~ i)) i ≡ a) refl refl -rotLoopInv a i j = - hcomp - (λ k → λ { - (i = i0) → a; - (i = i1) → rotLoop a (j ∧ ~ k); - (j = i0) → rotLoop (rotLoop a (~ i)) i; - (j = i1) → rotLoop a (i ∧ ~ k)}) - (rotLoop (rotLoop a (~ i ∨ j)) i) +rotLoopInv a i j = homotopySymInv rotLoop a j i rotLoopEquiv : (i : I) → S¹ ≃ S¹ rotLoopEquiv i = @@ -329,15 +436,15 @@ rotLoopEquiv i = private rotInv-aux-1 : I → I → I → I → S¹ rotInv-aux-1 j k i = - hfill (λ l → λ { (k = i0) → (loop (i ∧ ~ l)) * loop j + hfill (λ l → λ { (k = i0) → (loop (i ∧ ~ l)) · loop j ; (k = i1) → loop j - ; (i = i0) → (loop k * loop j) * loop (~ k) - ; (i = i1) → loop (~ k ∧ ~ l) * loop j }) - (inS ((loop (k ∨ i) * loop j) * loop (~ k))) + ; (i = i0) → (loop k · loop j) · loop (~ k) + ; (i = i1) → loop (~ k ∧ ~ l) · loop j }) + (inS ((loop (k ∨ i) · loop j) · loop (~ k))) rotInv-aux-2 : I → I → I → S¹ rotInv-aux-2 i j k = - hcomp (λ l → λ { (k = i0) → inv (filler-rot (~ i) (~ j) l) + hcomp (λ l → λ { (k = i0) → invLooper (filler-rot (~ i) (~ j) l) ; (k = i1) → loop (j ∧ l) ; (i = i0) → filler-rot k j l ; (i = i1) → loop (j ∧ l) @@ -349,37 +456,37 @@ private rotInv-aux-3 j k i = hfill (λ l → λ { (k = i0) → rotInv-aux-2 i j l ; (k = i1) → loop j - ; (i = i0) → loop (k ∨ l) * loop j - ; (i = i1) → loop k * (inv (loop (~ j) * loop k)) }) - (inS (loop k * (inv (loop (~ j) * loop (k ∨ ~ i))))) + ; (i = i0) → loop (k ∨ l) · loop j + ; (i = i1) → loop k · (invLooper (loop (~ j) · loop k)) }) + (inS (loop k · (invLooper (loop (~ j) · loop (k ∨ ~ i))))) rotInv-aux-4 : I → I → I → I → S¹ rotInv-aux-4 j k i = hfill (λ l → λ { (k = i0) → rotInv-aux-2 i j l ; (k = i1) → loop j - ; (i = i0) → loop j * loop (k ∨ l) - ; (i = i1) → (inv (loop (~ j) * loop k)) * loop k }) - (inS ((inv (loop (~ j) * loop (k ∨ ~ i))) * loop k)) + ; (i = i0) → loop j · loop (k ∨ l) + ; (i = i1) → (invLooper (loop (~ j) · loop k)) · loop k }) + (inS ((invLooper (loop (~ j) · loop (k ∨ ~ i))) · loop k)) -rotInv-1 : (a b : S¹) → b * a * inv b ≡ a +rotInv-1 : (a b : S¹) → b · a · invLooper b ≡ a rotInv-1 base base i = base rotInv-1 base (loop k) i = rotInv-aux-1 i0 k i i1 rotInv-1 (loop j) base i = loop j rotInv-1 (loop j) (loop k) i = rotInv-aux-1 j k i i1 -rotInv-2 : (a b : S¹) → inv b * a * b ≡ a +rotInv-2 : (a b : S¹) → invLooper b · a · b ≡ a rotInv-2 base base i = base rotInv-2 base (loop k) i = rotInv-aux-1 i0 (~ k) i i1 rotInv-2 (loop j) base i = loop j rotInv-2 (loop j) (loop k) i = rotInv-aux-1 j (~ k) i i1 -rotInv-3 : (a b : S¹) → b * (inv (inv a * b)) ≡ a +rotInv-3 : (a b : S¹) → b · (invLooper (invLooper a · b)) ≡ a rotInv-3 base base i = base rotInv-3 base (loop k) i = rotInv-aux-3 i0 k (~ i) i1 rotInv-3 (loop j) base i = loop j rotInv-3 (loop j) (loop k) i = rotInv-aux-3 j k (~ i) i1 -rotInv-4 : (a b : S¹) → inv (b * inv a) * b ≡ a +rotInv-4 : (a b : S¹) → invLooper (b · invLooper a) · b ≡ a rotInv-4 base base i = base rotInv-4 base (loop k) i = rotInv-aux-4 i0 k (~ i) i1 rotInv-4 (loop j) base i = loop j diff --git a/Cubical/HITs/S1/Properties.agda b/Cubical/HITs/S1/Properties.agda index 9316dc20b..7ba091704 100644 --- a/Cubical/HITs/S1/Properties.agda +++ b/Cubical/HITs/S1/Properties.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.S1.Properties where open import Cubical.Foundations.Prelude @@ -9,7 +9,7 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence open import Cubical.HITs.S1.Base -open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation as PropTrunc isConnectedS¹ : (s : S¹) → ∥ base ≡ s ∥ isConnectedS¹ base = ∣ refl ∣ @@ -18,11 +18,17 @@ isConnectedS¹ (loop i) = isGroupoidS¹ : isGroupoid S¹ isGroupoidS¹ s t = - recPropTrunc isPropIsSet + PropTrunc.rec isPropIsSet (λ p → subst (λ s → isSet (s ≡ t)) p - (recPropTrunc isPropIsSet + (PropTrunc.rec isPropIsSet (λ q → subst (λ t → isSet (base ≡ t)) q isSetΩS¹) (isConnectedS¹ t))) (isConnectedS¹ s) +IsoFunSpaceS¹ : ∀ {ℓ} {A : Type ℓ} → Iso (S¹ → A) (Σ[ x ∈ A ] x ≡ x) +Iso.fun IsoFunSpaceS¹ f = (f base) , (cong f loop) +Iso.inv IsoFunSpaceS¹ (x , p) base = x +Iso.inv IsoFunSpaceS¹ (x , p) (loop i) = p i +Iso.rightInv IsoFunSpaceS¹ (x , p) = refl +Iso.leftInv IsoFunSpaceS¹ f = funExt λ {base → refl ; (loop i) → refl} diff --git a/Cubical/HITs/S2.agda b/Cubical/HITs/S2.agda index 425338926..1c783bef7 100644 --- a/Cubical/HITs/S2.agda +++ b/Cubical/HITs/S2.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.S2 where open import Cubical.HITs.S2.Base public diff --git a/Cubical/HITs/S2/Base.agda b/Cubical/HITs/S2/Base.agda index ab9736d32..0e4a267aa 100644 --- a/Cubical/HITs/S2/Base.agda +++ b/Cubical/HITs/S2/Base.agda @@ -1,8 +1,17 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.S2.Base where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels data S² : Type₀ where base : S² surf : PathP (λ i → base ≡ base) refl refl + +S²ToSetRec : ∀ {ℓ} {A : S² → Type ℓ} + → ((x : S²) → isSet (A x)) + → A base + → (x : S²) → A x +S²ToSetRec set b base = b +S²ToSetRec set b (surf i j) = + isOfHLevel→isOfHLevelDep 2 set b b {a0 = refl} {a1 = refl} refl refl surf i j diff --git a/Cubical/HITs/S3.agda b/Cubical/HITs/S3.agda index ab7195ee1..e5c16e1cc 100644 --- a/Cubical/HITs/S3.agda +++ b/Cubical/HITs/S3.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.S3 where open import Cubical.HITs.S3.Base public diff --git a/Cubical/HITs/S3/Base.agda b/Cubical/HITs/S3/Base.agda index 031cbdd61..dd765533b 100644 --- a/Cubical/HITs/S3/Base.agda +++ b/Cubical/HITs/S3/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.S3.Base where open import Cubical.Foundations.Prelude @@ -6,4 +6,3 @@ open import Cubical.Foundations.Prelude data S³ : Type₀ where base : S³ surf : PathP (λ j → PathP (λ i → base ≡ base) refl refl) refl refl - diff --git a/Cubical/HITs/SetQuotients.agda b/Cubical/HITs/SetQuotients.agda index 87f14f6b1..7339c4344 100644 --- a/Cubical/HITs/SetQuotients.agda +++ b/Cubical/HITs/SetQuotients.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.SetQuotients where open import Cubical.HITs.SetQuotients.Base public diff --git a/Cubical/HITs/SetQuotients/Base.agda b/Cubical/HITs/SetQuotients/Base.agda index fe55ed52c..2c47ec7d2 100644 --- a/Cubical/HITs/SetQuotients/Base.agda +++ b/Cubical/HITs/SetQuotients/Base.agda @@ -5,7 +5,7 @@ This file contains: - Definition of set quotients -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.SetQuotients.Base where open import Cubical.Core.Primitives diff --git a/Cubical/HITs/SetQuotients/Properties.agda b/Cubical/HITs/SetQuotients/Properties.agda index 6fd17f0be..c8a5c1ca0 100644 --- a/Cubical/HITs/SetQuotients/Properties.agda +++ b/Cubical/HITs/SetQuotients/Properties.agda @@ -4,7 +4,7 @@ Set quotients: -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.SetQuotients.Properties where open import Cubical.HITs.SetQuotients.Base @@ -12,123 +12,246 @@ open import Cubical.HITs.SetQuotients.Base open import Cubical.Core.Everything open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels -open import Cubical.Foundations.HAEquiv +open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.Univalence -open import Cubical.Data.Nat +open import Cubical.Functions.FunExtEquiv + open import Cubical.Data.Sigma open import Cubical.Relation.Nullary open import Cubical.Relation.Binary.Base -open import Cubical.HITs.PropositionalTruncation -open import Cubical.HITs.SetTruncation +open import Cubical.HITs.TypeQuotients as TypeQuot using (_/ₜ_ ; [_] ; eq/) +open import Cubical.HITs.PropositionalTruncation as PropTrunc using (∥_∥ ; ∣_∣ ; squash) +open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂ ; ∣_∣₂ ; squash₂ + ; isSetSetTrunc) --- Type quotients private variable ℓ ℓ' ℓ'' : Level A : Type ℓ - R : A → A → Type ℓ' - B : A / R → Type ℓ'' - -elimEq/ : (Bprop : (x : A / R ) → isProp (B x)) - {x y : A / R} - (eq : x ≡ y) - (bx : B x) - (by : B y) → - PathP (λ i → B (eq i)) bx by -elimEq/ {B = B} Bprop {x = x} = - J (λ y eq → ∀ bx by → PathP (λ i → B (eq i)) bx by) (λ bx by → Bprop x bx by) - - -elimSetQuotientsProp : ((x : A / R ) → isProp (B x)) → - (f : (a : A) → B ( [ a ])) → - (x : A / R) → B x -elimSetQuotientsProp Bprop f [ x ] = f x -elimSetQuotientsProp Bprop f (squash/ x y p q i j) = + R : A → A → Type ℓ + B : A / R → Type ℓ + C : A / R → A / R → Type ℓ + D : A / R → A / R → A / R → Type ℓ + +elimProp : ((x : A / R ) → isProp (B x)) + → ((a : A) → B ( [ a ])) + → (x : A / R) + → B x +elimProp Bprop f [ x ] = f x +elimProp Bprop f (squash/ x y p q i j) = isOfHLevel→isOfHLevelDep 2 (λ x → isProp→isSet (Bprop x)) - (g x) (g y) (cong g p) (cong g q) (squash/ x y p q) i j + (g x) (g y) (cong g p) (cong g q) (squash/ x y p q) i j where - g = elimSetQuotientsProp Bprop f -elimSetQuotientsProp Bprop f (eq/ a b r i) = elimEq/ Bprop (eq/ a b r) (f a) (f b) i + g = elimProp Bprop f +elimProp Bprop f (eq/ a b r i) = isProp→PathP (λ i → Bprop ((eq/ a b r) i)) (f a) (f b) i + +elimProp2 : ((x y : A / R ) → isProp (C x y)) + → ((a b : A) → C [ a ] [ b ]) + → (x y : A / R) + → C x y +elimProp2 Cprop f = elimProp (λ x → isPropΠ (λ y → Cprop x y)) + (λ x → elimProp (λ y → Cprop [ x ] y) (f x)) + +elimProp3 : ((x y z : A / R ) → isProp (D x y z)) + → ((a b c : A) → D [ a ] [ b ] [ c ]) + → (x y z : A / R) + → D x y z +elimProp3 Dprop f = elimProp (λ x → isPropΠ2 (λ y z → Dprop x y z)) + (λ x → elimProp2 (λ y z → Dprop [ x ] y z) (f x)) + +-- sometimes more convenient: +elimContr : (∀ (a : A) → isContr (B [ a ])) + → (x : A / R) → B x +elimContr Bcontr = elimProp (elimProp (λ _ → isPropIsProp) λ _ → isContr→isProp (Bcontr _)) + λ _ → Bcontr _ .fst + +elimContr2 : (∀ (a b : A) → isContr (C [ a ] [ b ])) + → (x y : A / R) → C x y +elimContr2 Ccontr = elimContr λ _ → isOfHLevelΠ 0 + (elimContr λ _ → inhProp→isContr (Ccontr _ _) isPropIsContr) -- lemma 6.10.2 in hott book --- TODO: defined truncated Sigma as ∃ -[]surjective : (x : A / R) → ∥ Σ[ a ∈ A ] [ a ] ≡ x ∥ -[]surjective = elimSetQuotientsProp (λ x → squash) (λ a → ∣ a , refl ∣) - -elimSetQuotients : {B : A / R → Type ℓ} → - (Bset : (x : A / R) → isSet (B x)) → - (f : (a : A) → (B [ a ])) → - (feq : (a b : A) (r : R a b) → - PathP (λ i → B (eq/ a b r i)) (f a) (f b)) → - (x : A / R) → B x -elimSetQuotients Bset f feq [ a ] = f a -elimSetQuotients Bset f feq (eq/ a b r i) = feq a b r i -elimSetQuotients Bset f feq (squash/ x y p q i j) = +[]surjective : (x : A / R) → ∃[ a ∈ A ] [ a ] ≡ x +[]surjective = elimProp (λ x → squash) (λ a → ∣ a , refl ∣) + +elim : {B : A / R → Type ℓ} + → ((x : A / R) → isSet (B x)) + → (f : (a : A) → (B [ a ])) + → ((a b : A) (r : R a b) → PathP (λ i → B (eq/ a b r i)) (f a) (f b)) + → (x : A / R) + → B x +elim Bset f feq [ a ] = f a +elim Bset f feq (eq/ a b r i) = feq a b r i +elim Bset f feq (squash/ x y p q i j) = isOfHLevel→isOfHLevelDep 2 Bset (g x) (g y) (cong g p) (cong g q) (squash/ x y p q) i j where - g = elimSetQuotients Bset f feq + g = elim Bset f feq + +rec : {B : Type ℓ} + (Bset : isSet B) + (f : A → B) + (feq : (a b : A) (r : R a b) → f a ≡ f b) + → A / R → B +rec Bset f feq [ a ] = f a +rec Bset f feq (eq/ a b r i) = feq a b r i +rec Bset f feq (squash/ x y p q i j) = Bset (g x) (g y) (cong g p) (cong g q) i j + where + g = rec Bset f feq +rec2 : {B : Type ℓ} (Bset : isSet B) + (f : A → A → B) (feql : (a b c : A) (r : R a b) → f a c ≡ f b c) + (feqr : (a b c : A) (r : R b c) → f a b ≡ f a c) + → A / R → A / R → B +rec2 Bset f feql feqr = rec (isSetΠ (λ _ → Bset)) + (λ a → rec Bset (f a) (feqr a)) + (λ a b r → funExt (elimProp (λ _ → Bset _ _) + (λ c → feql a b c r))) -setQuotUniversal : {B : Type ℓ} (Bset : isSet B) → - (A / R → B) ≃ (Σ[ f ∈ (A → B) ] ((a b : A) → R a b → f a ≡ f b)) -setQuotUniversal Bset = isoToEquiv (iso intro elim elimRightInv elimLeftInv) +-- the recursor for maps into groupoids: +-- i.e. for any type A with a binary relation R and groupoid B, +-- we can construct a map A / R → B from a map A → B satisfying the conditions +-- (i) ∀ (a b : A) → R a b → f a ≡ f b +-- (ii) ∀ (a b : A) → isProp (f a ≡ f b) + +-- We start by proving that we can recover the set-quotient +-- by set-truncating the (non-truncated type quotient) +typeQuotSetTruncIso : Iso (A / R) ∥ A /ₜ R ∥₂ +Iso.fun typeQuotSetTruncIso = rec isSetSetTrunc (λ a → ∣ [ a ] ∣₂) + λ a b r → cong ∣_∣₂ (eq/ a b r) +Iso.inv typeQuotSetTruncIso = SetTrunc.rec squash/ (TypeQuot.rec [_] eq/) +Iso.rightInv typeQuotSetTruncIso = SetTrunc.elim (λ _ → isProp→isSet (squash₂ _ _)) + (TypeQuot.elimProp (λ _ → squash₂ _ _) λ _ → refl) +Iso.leftInv typeQuotSetTruncIso = elimProp (λ _ → squash/ _ _) λ _ → refl + +module rec→Gpd {A : Type ℓ} {R : A → A → Type ℓ'} {B : Type ℓ''} (Bgpd : isGroupoid B) + (f : A → B) + (feq : ∀ (a b : A) → R a b → f a ≡ f b) + (fprop : ∀ (a b : A) → isProp (f a ≡ f b)) where + + fun : A / R → B + fun = f₁ ∘ f₂ where - intro = λ g → (λ a → g [ a ]) , λ a b r i → g (eq/ a b r i) - elim = λ h → elimSetQuotients (λ x → Bset) (fst h) (snd h) + f₁ : ∥ A /ₜ R ∥₂ → B + f₁ = SetTrunc.rec→Gpd.fun Bgpd f/ congF/Const + where + f/ : A /ₜ R → B + f/ = TypeQuot.rec f feq + congF/Const : (a b : A /ₜ R) (p q : a ≡ b) → cong f/ p ≡ cong f/ q + congF/Const = TypeQuot.elimProp2 (λ _ _ → isPropΠ2 λ _ _ → Bgpd _ _ _ _) + λ a b p q → fprop a b (cong f/ p) (cong f/ q) + + f₂ : A / R → ∥ A /ₜ R ∥₂ + f₂ = Iso.fun typeQuotSetTruncIso - elimRightInv : ∀ h → intro (elim h) ≡ h - elimRightInv h = refl - elimLeftInv : ∀ g → elim (intro g) ≡ g - elimLeftInv = λ g → funExt (λ x → elimPropTrunc {P = λ sur → elim (intro g) x ≡ g x} - (λ sur → Bset (elim (intro g) x) (g x)) - (λ sur → cong (elim (intro g)) (sym (snd sur)) ∙ (cong g (snd sur))) ([]surjective x) - ) +setQuotUniversalIso : {B : Type ℓ} (Bset : isSet B) + → Iso (A / R → B) (Σ[ f ∈ (A → B) ] ((a b : A) → R a b → f a ≡ f b)) +Iso.fun (setQuotUniversalIso Bset) g = (λ a → g [ a ]) , λ a b r i → g (eq/ a b r i) +Iso.inv (setQuotUniversalIso Bset) h = elim (λ x → Bset) (fst h) (snd h) +Iso.rightInv (setQuotUniversalIso Bset) h = refl +Iso.leftInv (setQuotUniversalIso Bset) g = + funExt (λ x → PropTrunc.elim (λ sur → Bset (out (intro g) x) (g x)) + (λ sur → cong (out (intro g)) (sym (snd sur)) ∙ (cong g (snd sur))) ([]surjective x)) + where + intro = Iso.fun (setQuotUniversalIso Bset) + out = Iso.inv (setQuotUniversalIso Bset) + +setQuotUniversal : {B : Type ℓ} (Bset : isSet B) → + (A / R → B) ≃ (Σ[ f ∈ (A → B) ] ((a b : A) → R a b → f a ≡ f b)) +setQuotUniversal Bset = isoToEquiv (setQuotUniversalIso Bset) open BinaryRelation -effective : (Rprop : isPropValued R) (Requiv : isEquivRel R) (a b : A) → [ a ] ≡ [ b ] → R a b -effective {A = A} {R = R} Rprop (EquivRel R/refl R/sym R/trans) a b p = transport aa≡ab (R/refl _) +setQuotUnaryOp : (-_ : A → A) + → (∀ a a' → R a a' → R (- a) (- a')) + → (A / R → A / R) +setQuotUnaryOp -_ h = Iso.inv (setQuotUniversalIso squash/) ((λ a → [ - a ]) , λ a b x → eq/ _ _ (h _ _ x)) + +-- characterisation of binary functions/operations on set-quotients +setQuotUniversal2Iso : {B : Type ℓ} (Bset : isSet B) → isRefl R + → Iso (A / R → A / R → B) + (Σ[ _∗_ ∈ (A → A → B) ] ((a a' b b' : A) → R a a' → R b b' → a ∗ b ≡ a' ∗ b')) +Iso.fun (setQuotUniversal2Iso {A = A} {R = R} {B = B} Bset isReflR) _∗/_ = _∗_ , h + where + _∗_ = λ a b → [ a ] ∗/ [ b ] + h : (a a' b b' : A) → R a a' → R b b' → a ∗ b ≡ a' ∗ b' + h a a' b b' ra rb = cong (_∗/ [ b ]) (eq/ _ _ ra) ∙ cong ([ a' ] ∗/_) (eq/ _ _ rb) +Iso.inv (setQuotUniversal2Iso {A = A} {R = R} {B = B} Bset isReflR) (_∗_ , h) = + rec2 Bset _∗_ hleft hright + where + hleft : ∀ a b c → R a b → (a ∗ c) ≡ (b ∗ c) + hleft _ _ c r = h _ _ _ _ r (isReflR c) + hright : ∀ a b c → R b c → (a ∗ b) ≡ (a ∗ c) + hright a _ _ r = h _ _ _ _ (isReflR a) r +Iso.rightInv (setQuotUniversal2Iso {A = A} {R = R} {B = B} Bset isReflR) (_∗_ , h) = + Σ≡Prop (λ _ → isPropΠ4 λ _ _ _ _ → isPropΠ2 λ _ _ → Bset _ _) refl +Iso.leftInv (setQuotUniversal2Iso {A = A} {R = R} {B = B} Bset isReflR) _∗/_ = + funExt₂ (elimProp2 (λ _ _ → Bset _ _) λ _ _ → refl) + +setQuotUniversal2 : {B : Type ℓ} (Bset : isSet B) → isRefl R + → (A / R → A / R → B) + ≃ (Σ[ _∗_ ∈ (A → A → B) ] ((a a' b b' : A) → R a a' → R b b' → a ∗ b ≡ a' ∗ b')) +setQuotUniversal2 Bset isReflR = isoToEquiv (setQuotUniversal2Iso Bset isReflR) + +-- corollary for binary operations +-- TODO: prove truncated inverse for effective relations +setQuotBinOp : isRefl R + → (_∗_ : A → A → A) + → (∀ a a' b b' → R a a' → R b b' → R (a ∗ b) (a' ∗ b')) + → (A / R → A / R → A / R) +setQuotBinOp isReflR _∗_ h = Iso.inv (setQuotUniversal2Iso squash/ isReflR) + ((λ a b → [ a ∗ b ]) , λ _ _ _ _ ra rb → eq/ _ _ (h _ _ _ _ ra rb)) + +setQuotSymmBinOp : isRefl R → isTrans R + → (_∗_ : A → A → A) + → (∀ a b → a ∗ b ≡ b ∗ a) + → (∀ a a' b → R a a' → R (a ∗ b) (a' ∗ b)) + → (A / R → A / R → A / R) +setQuotSymmBinOp {A = A} {R = R} isReflR isTransR _∗_ ∗-symm h = setQuotBinOp isReflR _∗_ h' where - helper : A / R → hProp _ - helper = elimSetQuotients (λ _ → isSetHProp) (λ c → (R a c , Rprop a c)) - (λ c d cd → ΣProp≡ (λ _ → isPropIsProp) - (ua (PropEquiv→Equiv (Rprop a c) (Rprop a d) - (λ ac → R/trans _ _ _ ac cd) (λ ad → R/trans _ _ _ ad (R/sym _ _ cd))))) + h' : ∀ a a' b b' → R a a' → R b b' → R (a ∗ b) (a' ∗ b') + h' a a' b b' ra rb = isTransR _ _ _ (h a a' b ra) + (transport (λ i → R (∗-symm b a' i) (∗-symm b' a' i)) (h b b' a' rb)) - aa≡ab : R a a ≡ R a b - aa≡ab i = fst (helper (p i)) -isEquivRel→isEffective : isPropValued R → isEquivRel R → isEffective R -isEquivRel→isEffective {R = R} Rprop Req a b = isoToEquiv (iso intro elim intro-elim elim-intro) +effective : (Rprop : isPropValued R) (Requiv : isEquivRel R) (a b : A) → [ a ] ≡ [ b ] → R a b +effective {A = A} {R = R} Rprop (equivRel R/refl R/sym R/trans) a b p = transport aa≡ab (R/refl _) where - intro : [ a ] ≡ [ b ] → R a b - intro = effective Rprop Req a b + helper : A / R → hProp _ + helper = + rec isSetHProp (λ c → (R a c , Rprop a c)) + (λ c d cd → Σ≡Prop (λ _ → isPropIsProp) + (hPropExt (Rprop a c) (Rprop a d) + (λ ac → R/trans _ _ _ ac cd) (λ ad → R/trans _ _ _ ad (R/sym _ _ cd)))) - elim : R a b → [ a ] ≡ [ b ] - elim = eq/ a b + aa≡ab : R a a ≡ R a b + aa≡ab i = helper (p i) .fst - intro-elim : ∀ x → intro (elim x) ≡ x - intro-elim ab = Rprop a b _ _ +isEquivRel→effectiveIso : isPropValued R → isEquivRel R → (a b : A) → Iso ([ a ] ≡ [ b ]) (R a b) +Iso.fun (isEquivRel→effectiveIso {R = R} Rprop Req a b) = effective Rprop Req a b +Iso.inv (isEquivRel→effectiveIso {R = R} Rprop Req a b) = eq/ a b +Iso.rightInv (isEquivRel→effectiveIso {R = R} Rprop Req a b) _ = Rprop a b _ _ +Iso.leftInv (isEquivRel→effectiveIso {R = R} Rprop Req a b) _ = squash/ _ _ _ _ - elim-intro : ∀ x → elim (intro x) ≡ x - elim-intro eq = squash/ _ _ _ _ +isEquivRel→isEffective : isPropValued R → isEquivRel R → isEffective R +isEquivRel→isEffective Rprop Req a b = isoToIsEquiv (invIso (isEquivRel→effectiveIso Rprop Req a b)) discreteSetQuotients : Discrete A → isPropValued R → isEquivRel R → (∀ a₀ a₁ → Dec (R a₀ a₁)) → Discrete (A / R) discreteSetQuotients {A = A} {R = R} Adis Rprop Req Rdec = - elimSetQuotients ((λ a₀ → isSetPi (λ a₁ → isProp→isSet (isPropDec (squash/ a₀ a₁))))) - discreteSetQuotients' discreteSetQuotients'-eq + elim (λ a₀ → isSetΠ (λ a₁ → isProp→isSet (isPropDec (squash/ a₀ a₁)))) + discreteSetQuotients' discreteSetQuotients'-eq where discreteSetQuotients' : (a : A) (y : A / R) → Dec ([ a ] ≡ y) - discreteSetQuotients' a₀ = elimSetQuotients ((λ a₁ → isProp→isSet (isPropDec (squash/ [ a₀ ] a₁)))) dis dis-eq + discreteSetQuotients' a₀ = elim (λ a₁ → isProp→isSet (isPropDec (squash/ [ a₀ ] a₁))) dis dis-eq where dis : (a₁ : A) → Dec ([ a₀ ] ≡ [ a₁ ]) dis a₁ with Rdec a₀ a₁ @@ -147,3 +270,29 @@ discreteSetQuotients {A = A} {R = R} Adis Rprop Req Rdec = J (λ b ab → ∀ k → PathP (λ i → (y : A / R) → Dec (ab i ≡ y)) (discreteSetQuotients' a) k) (λ k → funExt (λ x → isPropDec (squash/ _ _) _ _)) (eq/ a b ab) (discreteSetQuotients' b) + + +-- Quotienting by the truncated relation is equivalent to quotienting by untruncated relation +truncRelIso : Iso (A / R) (A / (λ a b → ∥ R a b ∥)) +Iso.fun truncRelIso = rec squash/ [_] λ _ _ r → eq/ _ _ ∣ r ∣ +Iso.inv truncRelIso = rec squash/ [_] λ _ _ → PropTrunc.rec (squash/ _ _) λ r → eq/ _ _ r +Iso.rightInv truncRelIso = elimProp (λ _ → squash/ _ _) λ _ → refl +Iso.leftInv truncRelIso = elimProp (λ _ → squash/ _ _) λ _ → refl + +truncRelEquiv : A / R ≃ A / (λ a b → ∥ R a b ∥) +truncRelEquiv = isoToEquiv truncRelIso + +-- Using this we can obtain a useful characterization of +-- path-types for equivalence relations (not prop-valued) +-- and their quotients + +isEquivRel→TruncIso : isEquivRel R → (a b : A) → Iso ([ a ] ≡ [ b ]) ∥ R a b ∥ +isEquivRel→TruncIso {A = A} {R = R} Req a b = compIso (isProp→Iso (squash/ _ _) (squash/ _ _) + (cong (Iso.fun truncRelIso)) (cong (Iso.inv truncRelIso))) + (isEquivRel→effectiveIso (λ _ _ → PropTrunc.isPropPropTrunc) ∥R∥eq a b) + where + open isEquivRel + ∥R∥eq : isEquivRel λ a b → ∥ R a b ∥ + reflexive ∥R∥eq a = ∣ reflexive Req a ∣ + symmetric ∥R∥eq a b = PropTrunc.map (symmetric Req a b) + transitive ∥R∥eq a b c = PropTrunc.map2 (transitive Req a b c) diff --git a/Cubical/HITs/SetTruncation.agda b/Cubical/HITs/SetTruncation.agda index 2d1e1e1b0..402435485 100644 --- a/Cubical/HITs/SetTruncation.agda +++ b/Cubical/HITs/SetTruncation.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.SetTruncation where open import Cubical.HITs.SetTruncation.Base public diff --git a/Cubical/HITs/SetTruncation/Base.agda b/Cubical/HITs/SetTruncation/Base.agda index dff751554..44b0ecdac 100644 --- a/Cubical/HITs/SetTruncation/Base.agda +++ b/Cubical/HITs/SetTruncation/Base.agda @@ -5,13 +5,13 @@ This file contains: - Definition of set truncations -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.SetTruncation.Base where open import Cubical.Core.Primitives -- set truncation as a higher inductive type: -data ∥_∥₀ {ℓ} (A : Type ℓ) : Type ℓ where - ∣_∣₀ : A → ∥ A ∥₀ - squash₀ : ∀ (x y : ∥ A ∥₀) (p q : x ≡ y) → p ≡ q +data ∥_∥₂ {ℓ} (A : Type ℓ) : Type ℓ where + ∣_∣₂ : A → ∥ A ∥₂ + squash₂ : ∀ (x y : ∥ A ∥₂) (p q : x ≡ y) → p ≡ q diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index 0217f3ea2..e48526721 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -5,73 +5,315 @@ This file contains: - Properties of set truncations -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.SetTruncation.Properties where open import Cubical.HITs.SetTruncation.Base open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; elim to pElim) hiding (elim2 ; elim3 ; rec2 ; map) private variable - ℓ : Level - A : Type ℓ + ℓ ℓ' ℓ'' : Level + A B C D : Type ℓ + +rec : isSet B → (A → B) → ∥ A ∥₂ → B +rec Bset f ∣ x ∣₂ = f x +rec Bset f (squash₂ x y p q i j) = + Bset _ _ (cong (rec Bset f) p) (cong (rec Bset f) q) i j + +rec2 : isSet C → (A → B → C) → ∥ A ∥₂ → ∥ B ∥₂ → C +rec2 Cset f ∣ x ∣₂ ∣ y ∣₂ = f x y +rec2 Cset f ∣ x ∣₂ (squash₂ y z p q i j) = + Cset _ _ (cong (rec2 Cset f ∣ x ∣₂) p) (cong (rec2 Cset f ∣ x ∣₂) q) i j +rec2 Cset f (squash₂ x y p q i j) z = + Cset _ _ (cong (λ a → rec2 Cset f a z) p) (cong (λ a → rec2 Cset f a z) q) i j + +-- Old version: +-- rec2 Cset f = rec (isSetΠ λ _ → Cset) λ x → rec Cset (f x) -- lemma 6.9.1 in HoTT book -elimSetTrunc : {B : ∥ A ∥₀ → Type ℓ} → - (Bset : (x : ∥ A ∥₀) → isSet (B x)) → - (g : (a : A) → B (∣ a ∣₀)) → - (x : ∥ A ∥₀) → B x -elimSetTrunc Bset g ∣ a ∣₀ = g a -elimSetTrunc {A = A} {B = B} Bset g (squash₀ x y p q i j) = - isOfHLevel→isOfHLevelDep 2 Bset (elimSetTrunc Bset g x) (elimSetTrunc Bset g y) - (cong (elimSetTrunc Bset g) p) (cong (elimSetTrunc Bset g) q) (squash₀ x y p q) i j - -setTruncUniversal : {B : Type ℓ} → (isSet B) → (∥ A ∥₀ → B) ≃ (A → B) -setTruncUniversal Bset = isoToEquiv (iso intro elim leftInv rightInv) +elim : {B : ∥ A ∥₂ → Type ℓ} + (Bset : (x : ∥ A ∥₂) → isSet (B x)) + (f : (a : A) → B (∣ a ∣₂)) + (x : ∥ A ∥₂) → B x +elim Bset f ∣ a ∣₂ = f a +elim Bset f (squash₂ x y p q i j) = + isOfHLevel→isOfHLevelDep 2 Bset _ _ + (cong (elim Bset f) p) (cong (elim Bset f) q) (squash₂ x y p q) i j + +elim2 : {C : ∥ A ∥₂ → ∥ B ∥₂ → Type ℓ} + (Cset : ((x : ∥ A ∥₂) (y : ∥ B ∥₂) → isSet (C x y))) + (f : (a : A) (b : B) → C ∣ a ∣₂ ∣ b ∣₂) + (x : ∥ A ∥₂) (y : ∥ B ∥₂) → C x y +elim2 Cset f ∣ x ∣₂ ∣ y ∣₂ = f x y +elim2 Cset f ∣ x ∣₂ (squash₂ y z p q i j) = + isOfHLevel→isOfHLevelDep 2 (λ a → Cset ∣ x ∣₂ a) _ _ + (cong (elim2 Cset f ∣ x ∣₂) p) (cong (elim2 Cset f ∣ x ∣₂) q) (squash₂ y z p q) i j +elim2 Cset f (squash₂ x y p q i j) z = + isOfHLevel→isOfHLevelDep 2 (λ a → Cset a z) _ _ + (cong (λ a → elim2 Cset f a z) p) (cong (λ a → elim2 Cset f a z) q) (squash₂ x y p q) i j + +-- Old version: +-- elim2 Cset f = elim (λ _ → isSetΠ (λ _ → Cset _ _)) +-- (λ a → elim (λ _ → Cset _ _) (f a)) + +-- TODO: generalize +elim3 : {B : (x y z : ∥ A ∥₂) → Type ℓ} + (Bset : ((x y z : ∥ A ∥₂) → isSet (B x y z))) + (g : (a b c : A) → B ∣ a ∣₂ ∣ b ∣₂ ∣ c ∣₂) + (x y z : ∥ A ∥₂) → B x y z +elim3 Bset g = elim2 (λ _ _ → isSetΠ (λ _ → Bset _ _ _)) + (λ a b → elim (λ _ → Bset _ _ _) (g a b)) + + +-- the recursor for maps into groupoids following the "HIT proof" in: +-- https://arxiv.org/abs/1507.01150 +-- i.e. for any type A and groupoid B we can construct a map ∥ A ∥₂ → B +-- from a map A → B satisfying the condition +-- ∀ (a b : A) (p q : a ≡ b) → cong f p ≡ cong f q +-- TODO: prove that this is an equivalence +module rec→Gpd {A : Type ℓ} {B : Type ℓ'} (Bgpd : isGroupoid B) (f : A → B) + (congFConst : ∀ (a b : A) (p q : a ≡ b) → cong f p ≡ cong f q) where + + data H : Type ℓ where + η : A → H + ε : ∀ (a b : A) → ∥ a ≡ b ∥ → η a ≡ η b -- prop. trunc. of a≡b + δ : ∀ (a b : A) (p : a ≡ b) → ε a b ∣ p ∣ ≡ cong η p + gtrunc : isGroupoid H + + -- write elimination principle for H + module Helim {P : H → Type ℓ''} (Pgpd : ∀ h → isGroupoid (P h)) + (η* : (a : A) → P (η a)) + (ε* : ∀ (a b : A) (∣p∣ : ∥ a ≡ b ∥) + → PathP (λ i → P (ε a b ∣p∣ i)) (η* a) (η* b)) + (δ* : ∀ (a b : A) (p : a ≡ b) + → PathP (λ i → PathP (λ j → P (δ a b p i j)) (η* a) (η* b)) + (ε* a b ∣ p ∣) (cong η* p)) where + + fun : (h : H) → P h + fun (η a) = η* a + fun (ε a b ∣p∣ i) = ε* a b ∣p∣ i + fun (δ a b p i j) = δ* a b p i j + fun (gtrunc x y p q α β i j k) = isOfHLevel→isOfHLevelDep 3 Pgpd + (fun x) (fun y) + (cong fun p) (cong fun q) + (cong (cong fun) α) (cong (cong fun) β) + (gtrunc x y p q α β) i j k + + module Hrec {C : Type ℓ''} (Cgpd : isGroupoid C) + (η* : A → C) + (ε* : ∀ (a b : A) → ∥ a ≡ b ∥ → η* a ≡ η* b) + (δ* : ∀ (a b : A) (p : a ≡ b) → ε* a b ∣ p ∣ ≡ cong η* p) where + + fun : H → C + fun (η a) = η* a + fun (ε a b ∣p∣ i) = ε* a b ∣p∣ i + fun (δ a b p i j) = δ* a b p i j + fun (gtrunc x y p q α β i j k) = Cgpd (fun x) (fun y) (cong fun p) (cong fun q) + (cong (cong fun) α) (cong (cong fun) β) i j k + + module HelimProp {P : H → Type ℓ''} (Pprop : ∀ h → isProp (P h)) + (η* : (a : A) → P (η a)) where + + fun : ∀ h → P h + fun = Helim.fun (λ _ → isSet→isGroupoid (isProp→isSet (Pprop _))) η* + (λ a b ∣p∣ → isOfHLevel→isOfHLevelDep 1 Pprop _ _ (ε a b ∣p∣)) + λ a b p → isOfHLevel→isOfHLevelDep 1 + {B = λ p → PathP (λ i → P (p i)) (η* a) (η* b)} + (λ _ → isOfHLevelPathP 1 (Pprop _) _ _) _ _ (δ a b p) + + -- The main trick: eliminating into hsets is easy + -- i.e. H has the universal property of set truncation... + module HelimSet {P : H → Type ℓ''} (Pset : ∀ h → isSet (P h)) + (η* : ∀ a → P (η a)) where + + fun : (h : H) → P h + fun = Helim.fun (λ _ → isSet→isGroupoid (Pset _)) η* ε* + λ a b p → isOfHLevel→isOfHLevelDep 1 + {B = λ p → PathP (λ i → P (p i)) (η* a) (η* b)} + (λ _ → isOfHLevelPathP' 1 (Pset _) _ _) _ _ (δ a b p) + where + ε* : (a b : A) (∣p∣ : ∥ a ≡ b ∥) → PathP (λ i → P (ε a b ∣p∣ i)) (η* a) (η* b) + ε* a b = pElim (λ _ → isOfHLevelPathP' 1 (Pset _) (η* a) (η* b)) + λ p → subst (λ x → PathP (λ i → P (x i)) (η* a) (η* b)) + (sym (δ a b p)) (cong η* p) + + + -- Now we need to prove that H is a set. + -- We start with a little lemma: + localHedbergLemma : {X : Type ℓ''} (P : X → Type ℓ'') + → (∀ x → isProp (P x)) + → ((x : X) → P x → (y : X) → P y → x ≡ y) + -------------------------------------------------- + → (x : X) → P x → (y : X) → isProp (x ≡ y) + localHedbergLemma {X = X} P Pprop P→≡ x px y = isPropRetract + (λ p → subst P p px) (λ py → sym (P→≡ x px x px) ∙ P→≡ x px y py) + isRetract (Pprop y) + where + isRetract : (p : x ≡ y) → (sym (P→≡ x px x px)) ∙ P→≡ x px y (subst P p px) ≡ p + isRetract = J (λ y' p' → (sym (P→≡ x px x px)) ∙ P→≡ x px y' (subst P p' px) ≡ p') + (subst (λ px' → sym (P→≡ x px x px) ∙ P→≡ x px x px' ≡ refl) + (sym (substRefl {B = P} px)) (lCancel (P→≡ x px x px))) + + Hset : isSet H + Hset = HelimProp.fun (λ _ → isPropΠ λ _ → isPropIsProp) baseCaseLeft + where + baseCaseLeft : (a₀ : A) (y : H) → isProp (η a₀ ≡ y) + baseCaseLeft a₀ = localHedbergLemma (λ x → Q x .fst) (λ x → Q x .snd) Q→≡ _ ∣ refl ∣ + where + Q : H → hProp ℓ + Q = HelimSet.fun (λ _ → isSetHProp) λ b → ∥ a₀ ≡ b ∥ , isPropPropTrunc + -- Q (η b) = ∥ a ≡ b ∥ + + Q→≡ : (x : H) → Q x .fst → (y : H) → Q y .fst → x ≡ y + Q→≡ = HelimSet.fun (λ _ → isSetΠ3 λ _ _ _ → gtrunc _ _) + λ a p → HelimSet.fun (λ _ → isSetΠ λ _ → gtrunc _ _) + λ b q → sym (ε a₀ a p) ∙ ε a₀ b q + + -- our desired function will split through H, + -- i.e. we get a function ∥ A ∥₂ → H → B + fun : ∥ A ∥₂ → B + fun = f₁ ∘ f₂ + where + f₁ : H → B + f₁ = Hrec.fun Bgpd f εᶠ λ _ _ _ → refl + where + εᶠ : (a b : A) → ∥ a ≡ b ∥ → f a ≡ f b + εᶠ a b = rec→Set (Bgpd _ _) (cong f) λ p q → congFConst a b p q + -- this is the inductive step, + -- we use that maps ∥ A ∥ → B for an hset B + -- correspond to 2-Constant maps A → B (which cong f is by assumption) + f₂ : ∥ A ∥₂ → H + f₂ = rec Hset η + + +map : (A → B) → ∥ A ∥₂ → ∥ B ∥₂ +map f = rec squash₂ λ x → ∣ f x ∣₂ + +setTruncUniversal : isSet B → (∥ A ∥₂ → B) ≃ (A → B) +setTruncUniversal {B = B} Bset = + isoToEquiv (iso (λ h x → h ∣ x ∣₂) (rec Bset) (λ _ → refl) rinv) where - intro = (λ h a → h ∣ a ∣₀) - elim = elimSetTrunc (λ x → Bset) - - leftInv : ∀ g → intro (elim g) ≡ g - leftInv g = refl - - rightInv : ∀ h → elim (intro h) ≡ h - rightInv h i x = elimSetTrunc (λ x → isProp→isSet (Bset (elim (intro h) x) (h x))) - (λ a → refl) x i - -elimSetTrunc2 : {B : ∥ A ∥₀ → ∥ A ∥₀ → Type ℓ} - (Bset : ((x y : ∥ A ∥₀) → isSet (B x y))) - (g : (a b : A) → B ∣ a ∣₀ ∣ b ∣₀) - (x y : ∥ A ∥₀) → B x y -elimSetTrunc2 Bset g = elimSetTrunc (λ _ → isOfHLevelPi 2 (λ _ → Bset _ _)) (λ a → - elimSetTrunc (λ _ → Bset _ _) (λ b → g a b)) - -elimSetTrunc3 : {B : (x y z : ∥ A ∥₀) → Type ℓ} - (Bset : ((x y z : ∥ A ∥₀) → isSet (B x y z))) - (g : (a b c : A) → B ∣ a ∣₀ ∣ b ∣₀ ∣ c ∣₀) - (x y z : ∥ A ∥₀) → B x y z -elimSetTrunc3 Bset g = elimSetTrunc2 (λ _ _ → isOfHLevelPi 2 λ _ → Bset _ _ _) (λ a b → - elimSetTrunc (λ _ → Bset _ _ _) (λ c → g a b c)) - -setTruncIsSet : isSet ∥ A ∥₀ -setTruncIsSet a b p q = squash₀ a b p q - -setId : isSet A → ∥ A ∥₀ ≡ A -setId {A = A} isset = isoToPath (iso (elimSetTrunc {A = A} - (λ _ → isset) - (λ x → x)) - (λ x → ∣ x ∣₀) - (λ b → refl) - λ b → idLemma b) + rinv : (f : ∥ A ∥₂ → B) → rec Bset (λ x → f ∣ x ∣₂) ≡ f + rinv f i x = + elim (λ x → isProp→isSet (Bset (rec Bset (λ x → f ∣ x ∣₂) x) (f x))) + (λ _ → refl) x i + +isSetSetTrunc : isSet ∥ A ∥₂ +isSetSetTrunc a b p q = squash₂ a b p q + +setTruncIdempotentIso : isSet A → Iso ∥ A ∥₂ A +Iso.fun (setTruncIdempotentIso hA) = rec hA (idfun _) +Iso.inv (setTruncIdempotentIso hA) x = ∣ x ∣₂ +Iso.rightInv (setTruncIdempotentIso hA) _ = refl +Iso.leftInv (setTruncIdempotentIso hA) = elim (λ _ → isSet→isGroupoid isSetSetTrunc _ _) (λ _ → refl) + +setTruncIdempotent≃ : isSet A → ∥ A ∥₂ ≃ A +setTruncIdempotent≃ {A = A} hA = isoToEquiv (setTruncIdempotentIso hA) + +setTruncIdempotent : isSet A → ∥ A ∥₂ ≡ A +setTruncIdempotent hA = ua (setTruncIdempotent≃ hA) + +isContr→isContrSetTrunc : isContr A → isContr (∥ A ∥₂) +isContr→isContrSetTrunc contr = ∣ fst contr ∣₂ + , elim (λ _ → isOfHLevelPath 2 (isSetSetTrunc) _ _) + λ a → cong ∣_∣₂ (snd contr a) + + +setTruncIso : Iso A B → Iso ∥ A ∥₂ ∥ B ∥₂ +Iso.fun (setTruncIso is) = rec isSetSetTrunc (λ x → ∣ Iso.fun is x ∣₂) +Iso.inv (setTruncIso is) = rec isSetSetTrunc (λ x → ∣ Iso.inv is x ∣₂) +Iso.rightInv (setTruncIso is) = + elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ a → cong ∣_∣₂ (Iso.rightInv is a) +Iso.leftInv (setTruncIso is) = + elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ a → cong ∣_∣₂ (Iso.leftInv is a) + +setSigmaIso : {B : A → Type ℓ} → Iso ∥ Σ A B ∥₂ ∥ Σ A (λ x → ∥ B x ∥₂) ∥₂ +setSigmaIso {A = A} {B = B} = iso fun funinv sect retr where - idLemma : ∀ (b : ∥ A ∥₀) → ∣ elimSetTrunc (λ x → isset) (λ x → x) b ∣₀ ≡ b - idLemma b = elimSetTrunc {B = (λ x → ∣ elimSetTrunc (λ _ → isset) (λ x → x) x ∣₀ ≡ x)} - (λ x → (isOfHLevelSuc 2 (setTruncIsSet {A = A})) - ∣ elimSetTrunc (λ _ → isset) (λ x₁ → x₁) x ∣₀ - x) - (λ _ → refl) - b + {- writing it out explicitly to avoid yellow highlighting -} + fun : ∥ Σ A B ∥₂ → ∥ Σ A (λ x → ∥ B x ∥₂) ∥₂ + fun = rec isSetSetTrunc λ {(a , p) → ∣ a , ∣ p ∣₂ ∣₂} + funinv : ∥ Σ A (λ x → ∥ B x ∥₂) ∥₂ → ∥ Σ A B ∥₂ + funinv = rec isSetSetTrunc (λ {(a , p) → rec isSetSetTrunc (λ p → ∣ a , p ∣₂) p}) + sect : section fun funinv + sect = elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ { (a , p) → elim {B = λ p → fun (funinv ∣ a , p ∣₂) ≡ ∣ a , p ∣₂} + (λ p → isOfHLevelPath 2 isSetSetTrunc _ _) (λ _ → refl) p } + retr : retract fun funinv + retr = elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ { _ → refl } + +sigmaElim : {B : ∥ A ∥₂ → Type ℓ} {C : Σ ∥ A ∥₂ B → Type ℓ'} + (Bset : (x : Σ ∥ A ∥₂ B) → isSet (C x)) + (g : (a : A) (b : B ∣ a ∣₂) → C (∣ a ∣₂ , b)) + (x : Σ ∥ A ∥₂ B) → C x +sigmaElim {B = B} {C = C} set g (x , y) = + elim {B = λ x → (y : B x) → C (x , y)} (λ _ → isSetΠ λ _ → set _) g x y + +sigmaProdElim : {C : ∥ A ∥₂ × ∥ B ∥₂ → Type ℓ} {D : Σ (∥ A ∥₂ × ∥ B ∥₂) C → Type ℓ'} + (Bset : (x : Σ (∥ A ∥₂ × ∥ B ∥₂) C) → isSet (D x)) + (g : (a : A) (b : B) (c : C (∣ a ∣₂ , ∣ b ∣₂)) → D ((∣ a ∣₂ , ∣ b ∣₂) , c)) + (x : Σ (∥ A ∥₂ × ∥ B ∥₂) C) → D x +sigmaProdElim {B = B} {C = C} {D = D} set g ((x , y) , c) = + elim {B = λ x → (y : ∥ B ∥₂) (c : C (x , y)) → D ((x , y) , c)} + (λ _ → isSetΠ λ _ → isSetΠ λ _ → set _) + (λ x → elim (λ _ → isSetΠ λ _ → set _) (g x)) + x y c + +prodElim : {C : ∥ A ∥₂ × ∥ B ∥₂ → Type ℓ} + → ((x : ∥ A ∥₂ × ∥ B ∥₂) → isSet (C x)) + → ((a : A) (b : B) → C (∣ a ∣₂ , ∣ b ∣₂)) + → (x : ∥ A ∥₂ × ∥ B ∥₂) → C x +prodElim setC f (a , b) = elim2 (λ x y → setC (x , y)) f a b + +prodRec : {C : Type ℓ} → isSet C → (A → B → C) → ∥ A ∥₂ × ∥ B ∥₂ → C +prodRec setC f (a , b) = rec2 setC f a b + +prodElim2 : {E : (∥ A ∥₂ × ∥ B ∥₂) → (∥ C ∥₂ × ∥ D ∥₂) → Type ℓ} + → ((x : ∥ A ∥₂ × ∥ B ∥₂) (y : ∥ C ∥₂ × ∥ D ∥₂) → isSet (E x y)) + → ((a : A) (b : B) (c : C) (d : D) → E (∣ a ∣₂ , ∣ b ∣₂) (∣ c ∣₂ , ∣ d ∣₂)) + → ((x : ∥ A ∥₂ × ∥ B ∥₂) (y : ∥ C ∥₂ × ∥ D ∥₂) → (E x y)) +prodElim2 isset f = prodElim (λ _ → isSetΠ λ _ → isset _ _) + λ a b → prodElim (λ _ → isset _ _) + λ c d → f a b c d + +setTruncOfProdIso : Iso ∥ A × B ∥₂ (∥ A ∥₂ × ∥ B ∥₂) +Iso.fun setTruncOfProdIso = rec (isSet× isSetSetTrunc isSetSetTrunc) λ { (a , b) → ∣ a ∣₂ , ∣ b ∣₂ } +Iso.inv setTruncOfProdIso = prodRec isSetSetTrunc λ a b → ∣ a , b ∣₂ +Iso.rightInv setTruncOfProdIso = + prodElim (λ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) λ _ _ → refl +Iso.leftInv setTruncOfProdIso = + elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ {(a , b) → refl} + +IsoSetTruncateSndΣ : {A : Type ℓ} {B : A → Type ℓ'} → Iso ∥ Σ A B ∥₂ ∥ Σ A (λ x → ∥ B x ∥₂) ∥₂ +Iso.fun IsoSetTruncateSndΣ = map λ a → (fst a) , ∣ snd a ∣₂ +Iso.inv IsoSetTruncateSndΣ = rec isSetSetTrunc (uncurry λ x → map λ b → x , b) +Iso.rightInv IsoSetTruncateSndΣ = + elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry λ a → elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ _ → refl) +Iso.leftInv IsoSetTruncateSndΣ = + elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ _ → refl + +PathIdTrunc₀Iso : {a b : A} → Iso (∣ a ∣₂ ≡ ∣ b ∣₂) ∥ a ≡ b ∥ +Iso.fun (PathIdTrunc₀Iso {b = b}) p = + transport (λ i → rec {B = TypeOfHLevel _ 1} (isOfHLevelTypeOfHLevel 1) + (λ a → ∥ a ≡ b ∥ , squash) (p (~ i)) .fst) + ∣ refl ∣ +Iso.inv PathIdTrunc₀Iso = pRec (squash₂ _ _) (cong ∣_∣₂) +Iso.rightInv PathIdTrunc₀Iso _ = squash _ _ +Iso.leftInv PathIdTrunc₀Iso _ = squash₂ _ _ _ _ diff --git a/Cubical/HITs/SmashProduct.agda b/Cubical/HITs/SmashProduct.agda index 59842f918..00c7b1ceb 100644 --- a/Cubical/HITs/SmashProduct.agda +++ b/Cubical/HITs/SmashProduct.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.SmashProduct where open import Cubical.HITs.SmashProduct.Base public diff --git a/Cubical/HITs/SmashProduct/Base.agda b/Cubical/HITs/SmashProduct/Base.agda index d24450144..59cd41698 100644 --- a/Cubical/HITs/SmashProduct/Base.agda +++ b/Cubical/HITs/SmashProduct/Base.agda @@ -1,32 +1,44 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.SmashProduct.Base where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.HITs.Pushout.Base +open import Cubical.Data.Unit +open import Cubical.Data.Prod +open import Cubical.HITs.Wedge +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws --- This should be upstreamed to Basics when we develop some theory --- about pointed types -ptType : Type₁ -ptType = Σ[ A ∈ Type₀ ] A - -pt : ∀ (A : ptType) → A .fst -pt A = A .snd - -data Smash (A B : ptType) : Type₀ where +data Smash {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') : Type (ℓ-max ℓ ℓ') where basel : Smash A B baser : Smash A B - proj : (x : A .fst) → (y : B .fst) → Smash A B - gluel : (a : A .fst) → proj a (pt B) ≡ basel - gluer : (b : B .fst) → proj (pt A) b ≡ baser + proj : (x : typ A) → (y : typ B) → Smash A B + gluel : (a : typ A) → proj a (pt B) ≡ basel + gluer : (b : typ B) → proj (pt A) b ≡ baser + +private + variable + ℓ ℓ' : Level + A B C D : Pointed ℓ + +Smash-map : (f : A →∙ C) (g : B →∙ D) → Smash A B → Smash C D +Smash-map f g basel = basel +Smash-map f g baser = baser +Smash-map (f , fpt) (g , gpt) (proj x y) = proj (f x) (g y) +Smash-map (f , fpt) (g , gpt) (gluel a i) = ((λ j → proj (f a) (gpt j)) ∙ gluel (f a)) i +Smash-map (f , fpt) (g , gpt) (gluer b i) = ((λ j → proj (fpt j) (g b)) ∙ gluer (g b)) i -- Commutativity -comm : {A B : ptType} → Smash A B → Smash B A +comm : Smash A B → Smash B A comm basel = baser comm baser = basel comm (proj x y) = proj y x comm (gluel a i) = gluer a i comm (gluer b i) = gluel b i -commK : ∀ {A B : ptType} → (x : Smash A B) → comm (comm x) ≡ x +commK : (x : Smash A B) → comm (comm x) ≡ x commK basel = refl commK baser = refl commK (proj x y) = refl @@ -35,18 +47,167 @@ commK (gluer b x) = refl -- WIP below -SmashPt : (A B : ptType) → ptType +SmashPt : (A : Pointed ℓ) (B : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') SmashPt A B = (Smash A B , basel) --- -- A (B C) = C (B A) = C (A B) = (A B) C --- rearrange : ∀ {A B C : ptType} → Smash A (SmashPt B C) → Smash C (SmashPt B A) --- rearrange basel = baser --- rearrange baser = basel --- rearrange {B = B} {C = C} (proj x basel) = proj (C .snd) baser --- rearrange {C = C} (proj x baser) = proj (C .snd) basel -- ? --- rearrange (proj x (proj y z)) = proj z (proj y x) --- rearrange {C = C} (proj x (gluel a i)) = proj (C .snd) {!!} --- rearrange (proj x (gluer b i)) = {!!} --- rearrange (gluel a i) = {!!} --- rearrange (gluer b i) = {!gluel ? i!} +SmashPtProj : (A : Pointed ℓ) (B : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +SmashPtProj A B = Smash A B , (proj (snd A) (snd B)) + +--- Alternative definition + +i∧ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → A ⋁ B → (typ A) × (typ B) +i∧ {A = A , ptA} {B = B , ptB} (inl x) = x , ptB +i∧ {A = A , ptA} {B = B , ptB} (inr x) = ptA , x +i∧ {A = A , ptA} {B = B , ptB} (push tt i) = ptA , ptB + +_⋀_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') +A ⋀ B = Pushout {A = (A ⋁ B)} (λ _ → tt) i∧ + +_⋀∙_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ') +A ⋀∙ B = Pushout {A = (A ⋁ B)} (λ _ → tt) i∧ , (inl tt) + + +_⋀→_ : (f : A →∙ C) (g : B →∙ D) → A ⋀ B → C ⋀ D +(f ⋀→ g) (inl tt) = inl tt +((f , fpt) ⋀→ (g , gpt)) (inr (x , x₁)) = inr (f x , g x₁) +_⋀→_ {B = B} {D = D} (f , fpt) (b , gpt) (push (inl x) i) = (push (inl (f x)) ∙ (λ i → inr (f x , gpt (~ i)))) i +_⋀→_ (f , fpt) (g , gpt) (push (inr x) i) = (push (inr (g x)) ∙ (λ i → inr (fpt (~ i) , g x))) i +_⋀→_ {A = A} {C = C} {B = B} {D = D} (f , fpt) (g , gpt) (push (push tt j) i) = + hcomp (λ k → λ { (i = i0) → inl tt + ; (i = i1) → inr (fpt (~ k) , gpt (~ k)) + ; (j = i0) → compPath-filler (push (inl (fpt (~ k)))) + ((λ i → inr (fpt (~ k) , gpt (~ i)))) k i + ; (j = i1) → compPath-filler (push (inr (gpt (~ k)))) + ((λ i → inr (fpt (~ i) , gpt (~ k)))) k i}) + (push (push tt j) i) + +⋀→Smash : A ⋀ B → Smash A B +⋀→Smash (inl x) = basel +⋀→Smash (inr (x , x₁)) = proj x x₁ +⋀→Smash (push (inl x) i) = gluel x (~ i) +⋀→Smash {A = A} {B = B} (push (inr x) i) = (sym (gluel (snd A)) ∙∙ gluer (snd B) ∙∙ sym (gluer x)) i +⋀→Smash {A = A} {B = B} (push (push a j) i) = + hcomp (λ k → λ { (i = i0) → gluel (snd A) (k ∨ ~ j) + ; (i = i1) → gluer (snd B) (~ k ∧ j) + ; (j = i0) → gluel (snd A) (~ i)}) + (invSides-filler (gluel (snd A)) (gluer (snd B)) j (~ i)) + +Smash→⋀ : Smash A B → A ⋀ B +Smash→⋀ basel = inl tt +Smash→⋀ baser = inl tt +Smash→⋀ (proj x y) = inr (x , y) +Smash→⋀ (gluel a i) = push (inl a) (~ i) +Smash→⋀ (gluer b i) = push (inr b) (~ i) + +{- associativity maps for smash produts. Proof pretty much direcly translated from https://github.com/ecavallo/redtt/blob/master/library/pointed/smash.red -} +private + pivotl : (b b' : typ B) + → Path (Smash A B) (proj (snd A) b) (proj (snd A) b') + pivotl b b' i = (gluer b ∙ sym (gluer b')) i + + pivotr : (a a' : typ A) + → Path (Smash A B) (proj a (snd B)) (proj a' (snd B)) + pivotr a a' i = (gluel a ∙ sym (gluel a')) i + + pivotlrId : {A : Pointed ℓ} {B : Pointed ℓ'} → _ + pivotlrId {A = A} {B = B} = rCancel (gluer (snd B)) ∙ sym (rCancel (gluel (snd A))) + + rearrange-proj : (c : fst C) + → (Smash A B) → Smash (SmashPtProj C B) A + rearrange-proj c basel = baser + rearrange-proj c baser = basel + rearrange-proj c (proj x y) = proj (proj c y) x + rearrange-proj {C = C} c (gluel a i) = + hcomp (λ k → λ { (i = i0) → proj (pivotr (snd C) c k) a + ; (i = i1) → baser}) + (gluer a i) + rearrange-proj c (gluer b i) = gluel (proj c b) i + + rearrange-gluel : (s : Smash A B) + → Path (Smash (SmashPtProj C B) A) basel (rearrange-proj (snd C) s) + rearrange-gluel {A = A} {B = B} {C = C} basel = sym (gluel (proj (snd C) (snd B))) ∙ + gluer (snd A) + rearrange-gluel baser = refl + rearrange-gluel {A = A} {B = B} {C = C} (proj a b) i = + hcomp (λ k → λ { (i = i0) → (sym (gluel (proj (snd C) (snd B))) ∙ + gluer (snd A)) (~ k) + ; (i = i1) → proj (pivotl (snd B) b k) a}) + (gluer a (~ i)) + rearrange-gluel {A = A} {B = B} {C = C} (gluel a i) j = + hcomp (λ k → λ { (i = i1) → ((λ i₁ → gluel (proj (snd C) (snd B)) (~ i₁)) ∙ + gluer (snd A)) (~ k ∨ j) + ; (j = i0) → ((λ i₁ → gluel (proj (snd C) (snd B)) (~ i₁)) ∙ + gluer (snd A)) (~ k) + ; (j = i1) → top-filler i k}) + (gluer a (i ∨ ~ j)) + where + top-filler : I → I → Smash (SmashPtProj C B) A + top-filler i j = + hcomp (λ k → λ { (i = i0) → side-filler k j + ; (i = i1) → gluer a (j ∨ k) + ; (j = i0) → gluer a (i ∧ k)}) + (gluer a (i ∧ j)) + where + side-filler : I → I → Smash (SmashPtProj C B) A + side-filler i j = + hcomp (λ k → λ { (i = i0) → proj (proj (snd C) (snd B)) a + ; (i = i1) → proj ((rCancel (gluel (snd C)) ∙ sym (rCancel (gluer (snd B)))) k j) a + ; (j = i0) → proj (proj (snd C) (snd B)) a + ; (j = i1) → (proj ((gluel (snd C) ∙ sym (gluel (snd C))) i) a)}) + (proj ((gluel (snd C) ∙ sym (gluel (snd C))) (j ∧ i)) a) + rearrange-gluel {A = A} {B = B} {C = C} (gluer b i) j = + hcomp (λ k → λ {(i = i1) → ((sym (gluel (proj (snd C) (snd B)))) ∙ gluer (snd A)) (~ k) + ; (j = i0) → ((sym (gluel (proj (snd C) (snd B)))) ∙ gluer (snd A)) (~ k) + ; (j = i1) → top-filler1 i k}) + (gluer (snd A) (i ∨ (~ j))) + where + top-filler1 : I → I → Smash (SmashPtProj C B) A + top-filler1 i j = + hcomp (λ k → λ { (i = i0) → congFunct (λ x → proj x (snd A)) (gluer (snd B)) (sym (gluer b)) (~ k) j + ; (i = i1) → (sym (gluel (proj (snd C) (snd B))) ∙ gluer (snd A)) (~ j) + ; (j = i0) → gluer (snd A) i + ; (j = i1) → gluel (proj (snd C) b) i}) + (top-filler2 i j) + where + top-filler2 : I → I → Smash (SmashPtProj C B) A + top-filler2 i j = + hcomp (λ k → λ { (j = i0) → gluer (snd A) (i ∧ k) + ; (j = i1) → gluel (gluer b (~ k)) i}) + (hcomp (λ k → λ { (j = i0) → gluel (gluer (snd B) i0) (~ k ∧ (~ i)) + ; (j = i1) → gluel (baser) (~ k ∨ i) + ; (i = i0) → gluel (gluer (snd B) j) (~ k) + ; (i = i1) → gluel (proj (snd C) (snd B)) j }) + (gluel (proj (snd C) (snd B)) (j ∨ (~ i)))) + + rearrange : Smash (SmashPtProj A B) C → Smash (SmashPtProj C B) A + rearrange basel = basel + rearrange baser = baser + rearrange (proj x y) = rearrange-proj y x + rearrange (gluel a i) = rearrange-gluel a (~ i) + rearrange {A = A} {B = B} {C = C} (gluer b i) = ((λ j → proj (pivotr b (snd C) j) (snd A)) ∙ + gluer (snd A)) i + + ⋀∙→SmashPtProj : (A ⋀∙ B) →∙ SmashPtProj A B + ⋀∙→SmashPtProj {A = A} {B = B} = fun , refl + where + fun : (A ⋀ B) → Smash A B + fun (inl x) = proj (snd A) (snd B) + fun (inr (x , x₁)) = proj x x₁ + fun (push (inl x) i) = pivotr (snd A) x i + fun (push (inr x) i) = pivotl (snd B) x i + fun (push (push a j) i) = pivotlrId (~ j) i + + SmashPtProj→⋀∙ : (SmashPtProj A B) →∙ (A ⋀∙ B) + SmashPtProj→⋀∙ {A = A} {B = B} = Smash→⋀ , sym (push (inr (snd B))) + +SmashAssociate : Smash (SmashPtProj A B) C → Smash A (SmashPtProj B C) +SmashAssociate = comm ∘ Smash-map (comm , refl) (idfun∙ _) ∘ rearrange + +SmashAssociate⁻ : Smash A (SmashPtProj B C) → Smash (SmashPtProj A B) C +SmashAssociate⁻ = rearrange ∘ comm ∘ Smash-map (idfun∙ _) (comm , refl) + +⋀-associate : (A ⋀∙ B) ⋀ C → A ⋀ (B ⋀∙ C) +⋀-associate = (idfun∙ _ ⋀→ SmashPtProj→⋀∙) ∘ Smash→⋀ ∘ SmashAssociate ∘ ⋀→Smash ∘ (⋀∙→SmashPtProj ⋀→ idfun∙ _) +⋀-associate⁻ : A ⋀ (B ⋀∙ C) → (A ⋀∙ B) ⋀ C +⋀-associate⁻ = (SmashPtProj→⋀∙ ⋀→ idfun∙ _) ∘ Smash→⋀ ∘ SmashAssociate⁻ ∘ ⋀→Smash ∘ (idfun∙ _ ⋀→ ⋀∙→SmashPtProj) diff --git a/Cubical/HITs/Sn.agda b/Cubical/HITs/Sn.agda index a85d6f439..069cfd5e0 100644 --- a/Cubical/HITs/Sn.agda +++ b/Cubical/HITs/Sn.agda @@ -1,4 +1,5 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Sn where open import Cubical.HITs.Sn.Base public +open import Cubical.HITs.Sn.Properties public diff --git a/Cubical/HITs/Sn/Base.agda b/Cubical/HITs/Sn/Base.agda index 230229976..0fb8635f5 100644 --- a/Cubical/HITs/Sn/Base.agda +++ b/Cubical/HITs/Sn/Base.agda @@ -1,15 +1,29 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Sn.Base where open import Cubical.HITs.Susp +open import Cubical.Foundations.Pointed open import Cubical.Data.Nat open import Cubical.Data.NatMinusOne open import Cubical.Data.Empty open import Cubical.Foundations.Prelude +open import Cubical.Data.Bool +open import Cubical.HITs.S1 S : ℕ₋₁ → Type₀ S neg1 = ⊥ -S (suc n) = Susp (S n) +S (ℕ→ℕ₋₁ n) = Susp (S (-1+ n)) S₊ : ℕ → Type₀ -S₊ n = S (ℕ→ℕ₋₁ n) +S₊ 0 = Bool +S₊ 1 = S¹ +S₊ (suc (suc n)) = Susp (S₊ (suc n)) + +ptSn : (n : ℕ) → S₊ n +ptSn zero = true +ptSn (suc zero) = base +ptSn (suc (suc n)) = north + +-- Pointed version +S₊∙ : (n : ℕ) → Pointed₀ +S₊∙ n = (S₊ n) , ptSn n diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda new file mode 100644 index 000000000..4dc17ee2e --- /dev/null +++ b/Cubical/HITs/Sn/Properties.agda @@ -0,0 +1,310 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Sn.Properties where + +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Path +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.HITs.S1 +open import Cubical.Data.Nat hiding (elim) +open import Cubical.Data.Sigma +open import Cubical.HITs.Sn.Base +open import Cubical.HITs.Susp +open import Cubical.HITs.Truncation +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected + +private + variable + ℓ : Level + +-- Elimination principles for spheres +sphereElim : (n : ℕ) {A : (S₊ (suc n)) → Type ℓ} → ((x : S₊ (suc n)) → isOfHLevel (suc n) (A x)) + → A (ptSn (suc n)) + → (x : S₊ (suc n)) → A x +sphereElim zero hlev pt = toPropElim hlev pt +sphereElim (suc n) hlev pt north = pt +sphereElim (suc n) {A = A} hlev pt south = subst A (merid (ptSn (suc n))) pt +sphereElim (suc n) {A = A} hlev pt (merid a i) = + sphereElim n {A = λ a → PathP (λ i → A (merid a i)) pt (subst A (merid (ptSn (suc n))) pt)} + (λ a → isOfHLevelPathP' (suc n) (hlev south) _ _) + (λ i → transp (λ j → A (merid (ptSn (suc n)) (i ∧ j))) (~ i) pt) + a i + +sphereElim2 : ∀ {ℓ} (n : ℕ) {A : (S₊ (suc n)) → (S₊ (suc n)) → Type ℓ} + → ((x y : S₊ (suc n)) → isOfHLevel (suc n) (A x y)) + → A (ptSn (suc n)) (ptSn (suc n)) + → (x y : S₊ (suc n)) → A x y +sphereElim2 n hlev pt = sphereElim n (λ _ → isOfHLevelΠ (suc n) λ _ → hlev _ _) + (sphereElim n (hlev _ ) pt) + +private + compPath-lem : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : z ≡ y) + → PathP (λ i → (p ∙ sym q) i ≡ y) p q + compPath-lem {y = y} p q i j = + hcomp (λ k → λ { (i = i0) → p j + ; (i = i1) → q (~ k ∨ j) + ; (j = i1) → y }) + (p (j ∨ i)) + +sphereToPropElim : (n : ℕ) {A : (S₊ (suc n)) → Type ℓ} → ((x : S₊ (suc n)) → isProp (A x)) + → A (ptSn (suc n)) + → (x : S₊ (suc n)) → A x +sphereToPropElim zero = toPropElim +sphereToPropElim (suc n) hlev pt north = pt +sphereToPropElim (suc n) {A = A} hlev pt south = subst A (merid (ptSn (suc n))) pt +sphereToPropElim (suc n) {A = A} hlev pt (merid a i) = + isProp→PathP {B = λ i → A (merid a i)} (λ _ → hlev _) pt (subst A (merid (ptSn (suc n))) pt) i + +-- Elimination rule for fibrations (x : Sⁿ) → (y : Sᵐ) → A x y of h-Level (n + m). +-- The following principle is just the special case of the "Wedge Connectivity Lemma" +-- for spheres (See Cubical.Homotopy.WedgeConnectivity or chapter 8.6 in the HoTT book). +-- We prove it directly here for three reasons: +-- (i) it should perform better +-- (ii) we get a slightly stronger statement for spheres: one of the homotopies will, by design, be refl +-- (iii) the fact that the two homotopies only differ by (composition with) the homotopy leftFunction(base) ≡ rightFunction(base) +-- is close to trivial + +wedgeconFun : (n m : ℕ) {A : (S₊ (suc n)) → (S₊ (suc m)) → Type ℓ} + → ((x : S₊ (suc n)) (y : S₊ (suc m)) → isOfHLevel ((suc n) + (suc m)) (A x y)) + → (f : (x : _) → A (ptSn (suc n)) x) + → (g : (x : _) → A x (ptSn (suc m))) + → (g (ptSn (suc n)) ≡ f (ptSn (suc m))) + → (x : S₊ (suc n)) (y : S₊ (suc m)) → A x y +wedgeconLeft : (n m : ℕ) {A : (S₊ (suc n)) → (S₊ (suc m)) → Type ℓ} + → (hLev : ((x : S₊ (suc n)) (y : S₊ (suc m)) → isOfHLevel ((suc n) + (suc m)) (A x y))) + → (f : (x : _) → A (ptSn (suc n)) x) + → (g : (x : _) → A x (ptSn (suc m))) + → (hom : g (ptSn (suc n)) ≡ f (ptSn (suc m))) + → (x : _) → wedgeconFun n m hLev f g hom (ptSn (suc n)) x ≡ f x +wedgeconRight : (n m : ℕ) {A : (S₊ (suc n)) → (S₊ (suc m)) → Type ℓ} + → (hLev : ((x : S₊ (suc n)) (y : S₊ (suc m)) → isOfHLevel ((suc n) + (suc m)) (A x y))) + → (f : (x : _) → A (ptSn (suc n)) x) + → (g : (x : _) → A x (ptSn (suc m))) + → (hom : g (ptSn (suc n)) ≡ f (ptSn (suc m))) + → (x : _) → wedgeconFun n m hLev f g hom x (ptSn (suc m)) ≡ g x +wedgeconFun zero zero {A = A} hlev f g hom = F + where + helper : SquareP (λ i j → A (loop i) (loop j)) (cong f loop) (cong f loop) + (λ i → hcomp (λ k → λ { (i = i0) → hom k + ; (i = i1) → hom k }) + (g (loop i))) + λ i → hcomp (λ k → λ { (i = i0) → hom k + ; (i = i1) → hom k }) + (g (loop i)) + helper = toPathP (isOfHLevelPathP' 1 (hlev _ _) _ _ _ _) + + F : (x y : S¹) → A x y + F base y = f y + F (loop i) base = hcomp (λ k → λ { (i = i0) → hom k + ; (i = i1) → hom k }) + (g (loop i)) + F (loop i) (loop j) = helper i j + +wedgeconFun zero (suc m) {A = A} hlev f g hom = F₀ + module _ where + transpLemma₀ : (x : S₊ (suc m)) → transport (λ i₁ → A base (merid x i₁)) (g base) ≡ f south + transpLemma₀ x = cong (transport (λ i₁ → A base (merid x i₁))) + hom + ∙ (λ i → transp (λ j → A base (merid x (i ∨ j))) i + (f (merid x i))) + + pathOverMerid₀ : (x : S₊ (suc m)) → PathP (λ i₁ → A base (merid x i₁)) + (g base) + (transport (λ i₁ → A base (merid (ptSn (suc m)) i₁)) + (g base)) + pathOverMerid₀ x i = hcomp (λ k → λ { (i = i0) → g base + ; (i = i1) → (transpLemma₀ x ∙ sym (transpLemma₀ (ptSn (suc m)))) k}) + (transp (λ i₁ → A base (merid x (i₁ ∧ i))) (~ i) + (g base)) + + pathOverMeridId₀ : pathOverMerid₀ (ptSn (suc m)) ≡ λ i → transp (λ i₁ → A base (merid (ptSn (suc m)) (i₁ ∧ i))) (~ i) + (g base) + pathOverMeridId₀ = + (λ j i → hcomp (λ k → λ {(i = i0) → g base + ; (i = i1) → rCancel (transpLemma₀ (ptSn (suc m))) j k}) + (transp (λ i₁ → A base (merid (ptSn (suc m)) (i₁ ∧ i))) (~ i) + (g base))) + ∙ λ j i → hfill (λ k → λ { (i = i0) → g base + ; (i = i1) → transport (λ i₁ → A base (merid (ptSn (suc m)) i₁)) + (g base)}) + (inS (transp (λ i₁ → A base (merid (ptSn (suc m)) (i₁ ∧ i))) (~ i) + (g base))) (~ j) + + indStep₀ : (x : _) (a : _) → PathP (λ i → A x (merid a i)) + (g x) + (subst (λ y → A x y) (merid (ptSn (suc m))) + (g x)) + indStep₀ = wedgeconFun zero m (λ _ _ → isOfHLevelPathP' (2 + m) (hlev _ _) _ _) + pathOverMerid₀ + (λ a i → transp (λ i₁ → A a (merid (ptSn (suc m)) (i₁ ∧ i))) (~ i) + (g a)) + (sym pathOverMeridId₀) + + F₀ : (x : S¹) (y : Susp (S₊ (suc m))) → A x y + F₀ x north = g x + F₀ x south = subst (λ y → A x y) (merid (ptSn (suc m))) (g x) + F₀ x (merid a i) = indStep₀ x a i +wedgeconFun (suc n) m {A = A} hlev f g hom = F₁ + module _ where + transpLemma₁ : (x : S₊ (suc n)) → transport (λ i₁ → A (merid x i₁) (ptSn (suc m))) (f (ptSn (suc m))) ≡ g south + transpLemma₁ x = cong (transport (λ i₁ → A (merid x i₁) (ptSn (suc m)))) + (sym hom) + ∙ (λ i → transp (λ j → A (merid x (i ∨ j)) (ptSn (suc m))) i + (g (merid x i))) + + pathOverMerid₁ : (x : S₊ (suc n)) → PathP (λ i₁ → A (merid x i₁) (ptSn (suc m))) + (f (ptSn (suc m))) + (transport (λ i₁ → A (merid (ptSn (suc n)) i₁) (ptSn (suc m))) + (f (ptSn (suc m)))) + pathOverMerid₁ x i = hcomp (λ k → λ { (i = i0) → f (ptSn (suc m)) + ; (i = i1) → (transpLemma₁ x ∙ sym (transpLemma₁ (ptSn (suc n)))) k }) + (transp (λ i₁ → A (merid x (i₁ ∧ i)) (ptSn (suc m))) (~ i) + (f (ptSn (suc m)))) + + pathOverMeridId₁ : pathOverMerid₁ (ptSn (suc n)) ≡ λ i → transp (λ i₁ → A (merid (ptSn (suc n)) (i₁ ∧ i)) (ptSn (suc m))) (~ i) + (f (ptSn (suc m))) + pathOverMeridId₁ = + (λ j i → hcomp (λ k → λ { (i = i0) → f (ptSn (suc m)) + ; (i = i1) → rCancel (transpLemma₁ (ptSn (suc n))) j k }) + (transp (λ i₁ → A (merid (ptSn (suc n)) (i₁ ∧ i)) (ptSn (suc m))) (~ i) + (f (ptSn (suc m))))) + ∙ λ j i → hfill (λ k → λ { (i = i0) → f (ptSn (suc m)) + ; (i = i1) → transport (λ i₁ → A (merid (ptSn (suc n)) i₁) (ptSn (suc m))) + (f (ptSn (suc m))) }) + (inS (transp (λ i₁ → A (merid (ptSn (suc n)) (i₁ ∧ i)) (ptSn (suc m))) (~ i) + (f (ptSn (suc m))))) (~ j) + + indStep₁ : (a : _) (y : _) → PathP (λ i → A (merid a i) y) + (f y) + (subst (λ x → A x y) (merid (ptSn (suc n))) + (f y)) + indStep₁ = wedgeconFun n m (λ _ _ → isOfHLevelPathP' (suc (n + suc m)) (hlev _ _) _ _) + (λ a i → transp (λ i₁ → A (merid (ptSn (suc n)) (i₁ ∧ i)) a) (~ i) + (f a)) + pathOverMerid₁ + pathOverMeridId₁ + + F₁ : (x : Susp (S₊ (suc n))) (y : S₊ (suc m)) → A x y + F₁ north y = f y + F₁ south y = subst (λ x → A x y) (merid (ptSn (suc n))) (f y) + F₁ (merid a i) y = indStep₁ a y i +wedgeconRight zero zero {A = A} hlev f g hom = right + where + right : (x : S¹) → _ + right base = sym hom + right (loop i) j = hcomp (λ k → λ { (i = i0) → hom (~ j ∧ k) + ; (i = i1) → hom (~ j ∧ k) + ; (j = i1) → g (loop i) }) + (g (loop i)) +wedgeconRight zero (suc m) {A = A} hlev f g hom x = refl +wedgeconRight (suc n) m {A = A} hlev f g hom = right + where + lem : (x : _) → indStep₁ n m hlev f g hom x (ptSn (suc m)) ≡ _ + lem = wedgeconRight n m (λ _ _ → isOfHLevelPathP' (suc (n + suc m)) (hlev _ _) _ _) + (λ a i → transp (λ i₁ → A (merid (ptSn (suc n)) (i₁ ∧ i)) a) (~ i) + (f a)) + (pathOverMerid₁ n m hlev f g hom) + (pathOverMeridId₁ n m hlev f g hom) + + right : (x : Susp (S₊ (suc n))) → _ ≡ g x + right north = sym hom + right south = cong (subst (λ x → A x (ptSn (suc m))) + (merid (ptSn (suc n)))) + (sym hom) + ∙ λ i → transp (λ j → A (merid (ptSn (suc n)) (i ∨ j)) (ptSn (suc m))) i + (g (merid (ptSn (suc n)) i)) + right (merid a i) j = + hcomp (λ k → λ { (i = i0) → hom (~ j) + ; (i = i1) → transpLemma₁ n m hlev f g hom (ptSn (suc n)) j + ; (j = i0) → lem a (~ k) i + ; (j = i1) → g (merid a i)}) + (hcomp (λ k → λ { (i = i0) → hom (~ j) + ; (i = i1) → compPath-lem (transpLemma₁ n m hlev f g hom a) (transpLemma₁ n m hlev f g hom (ptSn (suc n))) k j + ; (j = i1) → g (merid a i)}) + (hcomp (λ k → λ { (i = i0) → hom (~ j) + ; (j = i0) → transp (λ i₂ → A (merid a (i₂ ∧ i)) (ptSn (suc m))) (~ i) + (f (ptSn (suc m))) + ; (j = i1) → transp (λ j → A (merid a (i ∧ (j ∨ k))) (ptSn (suc m))) (k ∨ ~ i) + (g (merid a (i ∧ k))) }) + (transp (λ i₂ → A (merid a (i₂ ∧ i)) (ptSn (suc m))) (~ i) + (hom (~ j))))) +wedgeconLeft zero zero {A = A} hlev f g hom x = refl +wedgeconLeft zero (suc m) {A = A} hlev f g hom = help + where + left₁ : (x : _) → indStep₀ m hlev f g hom base x ≡ _ + left₁ = wedgeconLeft zero m (λ _ _ → isOfHLevelPathP' (2 + m) (hlev _ _) _ _) + (pathOverMerid₀ m hlev f g hom) + (λ a i → transp (λ i₁ → A a (merid (ptSn (suc m)) (i₁ ∧ i))) (~ i) + (g a)) + (sym (pathOverMeridId₀ m hlev f g hom)) + + help : (x : S₊ (suc (suc m))) → _ + help north = hom + help south = cong (subst (A base) (merid (ptSn (suc m)))) hom + ∙ λ i → transp (λ j → A base (merid (ptSn (suc m)) (i ∨ j))) i + (f (merid (ptSn (suc m)) i)) + help (merid a i) j = + hcomp (λ k → λ { (i = i0) → hom j + ; (i = i1) → transpLemma₀ m hlev f g hom (ptSn (suc m)) j + ; (j = i0) → left₁ a (~ k) i + ; (j = i1) → f (merid a i)}) + (hcomp (λ k → λ { (i = i0) → hom j + ; (i = i1) → compPath-lem (transpLemma₀ m hlev f g hom a) + (transpLemma₀ m hlev f g hom (ptSn (suc m))) k j + ; (j = i1) → f (merid a i)}) + (hcomp (λ k → λ { (i = i0) → hom j + ; (j = i0) → transp (λ i₂ → A base (merid a (i₂ ∧ i))) (~ i) + (g base) + ; (j = i1) → transp (λ j → A base (merid a (i ∧ (j ∨ k)))) (k ∨ ~ i) + (f (merid a (i ∧ k)))}) + (transp (λ i₂ → A base (merid a (i₂ ∧ i))) (~ i) + (hom j)))) +wedgeconLeft (suc n) m {A = A} hlev f g hom _ = refl + +---------- Connectedness ----------- + +sphereConnected : (n : HLevel) → isConnected (suc n) (S₊ n) +sphereConnected n = ∣ ptSn n ∣ , elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + (λ a → sym (spoke ∣_∣ (ptSn n)) ∙ spoke ∣_∣ a) + +-- The fact that path spaces of Sn are connected can be proved directly for Sⁿ. +-- (Unfortunately, this does not work for higher paths) +pathIdTruncSⁿ : (n : ℕ) (x y : S₊ (suc n)) + → Path (hLevelTrunc (2 + n) (S₊ (suc n))) ∣ x ∣ ∣ y ∣ + → hLevelTrunc (suc n) (x ≡ y) +pathIdTruncSⁿ n = sphereElim n (λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelTrunc (suc n)) + (sphereElim n (λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelTrunc (suc n)) + λ _ → ∣ refl ∣) + +pathIdTruncSⁿ⁻ : (n : ℕ) (x y : S₊ (suc n)) + → hLevelTrunc (suc n) (x ≡ y) + → Path (hLevelTrunc (2 + n) (S₊ (suc n))) ∣ x ∣ ∣ y ∣ +pathIdTruncSⁿ⁻ n x y = rec (isOfHLevelTrunc (2 + n) _ _) + (J (λ y _ → Path (hLevelTrunc (2 + n) (S₊ (suc n))) ∣ x ∣ ∣ y ∣) refl) + +pathIdTruncSⁿretract : (n : ℕ) (x y : S₊ (suc n)) → (p : hLevelTrunc (suc n) (x ≡ y)) → pathIdTruncSⁿ n x y (pathIdTruncSⁿ⁻ n x y p) ≡ p +pathIdTruncSⁿretract n = + sphereElim n (λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + λ y → elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + (J (λ y p → pathIdTruncSⁿ n (ptSn (suc n)) y (pathIdTruncSⁿ⁻ n (ptSn (suc n)) y ∣ p ∣) ≡ ∣ p ∣) + (cong (pathIdTruncSⁿ n (ptSn (suc n)) (ptSn (suc n))) (transportRefl refl) ∙ pm-help n)) + where + pm-help : (n : ℕ) → pathIdTruncSⁿ n (ptSn (suc n)) (ptSn (suc n)) refl ≡ ∣ refl ∣ + pm-help zero = refl + pm-help (suc n) = refl + +isConnectedPathSⁿ : (n : ℕ) (x y : S₊ (suc n)) → isConnected (suc n) (x ≡ y) +isConnectedPathSⁿ n x y = + isContrRetract + (pathIdTruncSⁿ⁻ n x y) + (pathIdTruncSⁿ n x y) + (pathIdTruncSⁿretract n x y) + ((isContr→isProp (sphereConnected (suc n)) ∣ x ∣ ∣ y ∣) + , isProp→isSet (isContr→isProp (sphereConnected (suc n))) _ _ _) diff --git a/Cubical/HITs/Susp.agda b/Cubical/HITs/Susp.agda index deaf83c22..fe041e6f5 100644 --- a/Cubical/HITs/Susp.agda +++ b/Cubical/HITs/Susp.agda @@ -1,6 +1,6 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Susp where open import Cubical.HITs.Susp.Base public --- open import Cubical.HITs.Susp.Properties public +open import Cubical.HITs.Susp.Properties public diff --git a/Cubical/HITs/Susp/Base.agda b/Cubical/HITs/Susp/Base.agda index 253444330..5517ebc33 100644 --- a/Cubical/HITs/Susp/Base.agda +++ b/Cubical/HITs/Susp/Base.agda @@ -1,9 +1,11 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Susp.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed open import Cubical.Data.Bool open import Cubical.Data.Empty @@ -12,19 +14,31 @@ open import Cubical.HITs.S1 open import Cubical.HITs.S2 open import Cubical.HITs.S3 +open Iso + data Susp {ℓ} (A : Type ℓ) : Type ℓ where north : Susp A south : Susp A merid : (a : A) → north ≡ south +∙Susp : ∀ {ℓ} (A : Type ℓ) → Pointed ℓ +∙Susp A = Susp A , north + +-- induced function +suspFun : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) + → Susp A → Susp B +suspFun f north = north +suspFun f south = south +suspFun f (merid a i) = merid (f a) i + +BoolIsoSusp⊥ : Iso Bool (Susp ⊥) +fun BoolIsoSusp⊥ = λ {true → north; false → south} +inv BoolIsoSusp⊥ = λ {north → true; south → false} +rightInv BoolIsoSusp⊥ = λ {north → refl; south → refl} +leftInv BoolIsoSusp⊥ = λ {true → refl; false → refl} + Bool≃Susp⊥ : Bool ≃ Susp ⊥ -Bool≃Susp⊥ = - isoToEquiv - (iso - (λ {true → north; false → south}) - (λ {north → true; south → false}) - (λ {north → refl; south → refl}) - (λ {true → refl; false → refl})) +Bool≃Susp⊥ = isoToEquiv BoolIsoSusp⊥ SuspBool : Type₀ SuspBool = Susp Bool @@ -37,29 +51,34 @@ SuspBool→S¹ (merid true i) = base S¹→SuspBool : S¹ → SuspBool S¹→SuspBool base = north -S¹→SuspBool (loop i) = (merid false ∙ (sym (merid true))) i +S¹→SuspBool (loop i) = (merid false ∙ sym (merid true)) i SuspBool→S¹→SuspBool : (x : SuspBool) → Path _ (S¹→SuspBool (SuspBool→S¹ x)) x SuspBool→S¹→SuspBool north = refl SuspBool→S¹→SuspBool south = merid true -SuspBool→S¹→SuspBool (merid false i) = λ j → hcomp (λ k → (λ { (j = i1) → merid false i - ; (i = i0) → north - ; (i = i1) → merid true (j ∨ ~ k)})) - (merid false i) -SuspBool→S¹→SuspBool (merid true i) = λ j → merid true (i ∧ j) +SuspBool→S¹→SuspBool (merid false i) j = hcomp (λ k → (λ { (j = i1) → merid false i + ; (i = i0) → north + ; (i = i1) → merid true (j ∨ ~ k)})) + (merid false i) +SuspBool→S¹→SuspBool (merid true i) j = merid true (i ∧ j) S¹→SuspBool→S¹ : (x : S¹) → SuspBool→S¹ (S¹→SuspBool x) ≡ x -S¹→SuspBool→S¹ base = refl -S¹→SuspBool→S¹ (loop i) = λ j → - hfill (λ k → λ { (i = i0) → base - ; (i = i1) → base }) - (inS (loop i)) (~ j) +S¹→SuspBool→S¹ base = refl +S¹→SuspBool→S¹ (loop i) j = hfill (λ k → λ { (i = i0) → base + ; (i = i1) → base }) + (inS (loop i)) (~ j) + +S¹IsoSuspBool : Iso S¹ SuspBool +fun S¹IsoSuspBool = S¹→SuspBool +inv S¹IsoSuspBool = SuspBool→S¹ +rightInv S¹IsoSuspBool = SuspBool→S¹→SuspBool +leftInv S¹IsoSuspBool = S¹→SuspBool→S¹ S¹≃SuspBool : S¹ ≃ SuspBool -S¹≃SuspBool = isoToEquiv (iso S¹→SuspBool SuspBool→S¹ SuspBool→S¹→SuspBool S¹→SuspBool→S¹) +S¹≃SuspBool = isoToEquiv S¹IsoSuspBool S¹≡SuspBool : S¹ ≡ SuspBool -S¹≡SuspBool = isoToPath (iso S¹→SuspBool SuspBool→S¹ SuspBool→S¹→SuspBool S¹→SuspBool→S¹) +S¹≡SuspBool = ua S¹≃SuspBool -- Now the sphere @@ -93,8 +112,17 @@ SuspS¹→S²→SuspS¹ south k = merid base k SuspS¹→S²→SuspS¹ (merid base j) k = merid base (k ∧ j) SuspS¹→S²→SuspS¹ (merid (loop j) i) k = meridian-contraction i j (~ k) +S²IsoSuspS¹ : Iso S² SuspS¹ +fun S²IsoSuspS¹ = S²→SuspS¹ +inv S²IsoSuspS¹ = SuspS¹→S² +rightInv S²IsoSuspS¹ = SuspS¹→S²→SuspS¹ +leftInv S²IsoSuspS¹ = S²→SuspS¹→S² + +S²≃SuspS¹ : S² ≃ SuspS¹ +S²≃SuspS¹ = isoToEquiv S²IsoSuspS¹ + S²≡SuspS¹ : S² ≡ SuspS¹ -S²≡SuspS¹ = isoToPath (iso S²→SuspS¹ SuspS¹→S² SuspS¹→S²→SuspS¹ S²→SuspS¹→S²) +S²≡SuspS¹ = ua S²≃SuspS¹ -- And the 3-sphere @@ -130,5 +158,34 @@ SuspS²→S³→SuspS² south l = merid base l SuspS²→S³→SuspS² (merid base j) l = merid base (l ∧ j) SuspS²→S³→SuspS² (merid (surf j k) i) l = meridian-contraction-2 i j k (~ l) +S³IsoSuspS² : Iso S³ SuspS² +fun S³IsoSuspS² = S³→SuspS² +inv S³IsoSuspS² = SuspS²→S³ +rightInv S³IsoSuspS² = SuspS²→S³→SuspS² +leftInv S³IsoSuspS² = S³→SuspS²→S³ + +S³≃SuspS² : S³ ≃ SuspS² +S³≃SuspS² = isoToEquiv S³IsoSuspS² + S³≡SuspS² : S³ ≡ SuspS² -S³≡SuspS² = isoToPath (iso S³→SuspS² SuspS²→S³ SuspS²→S³→SuspS² S³→SuspS²→S³) +S³≡SuspS² = ua S³≃SuspS² + +IsoType→IsoSusp : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso A B → Iso (Susp A) (Susp B) +fun (IsoType→IsoSusp is) north = north +fun (IsoType→IsoSusp is) south = south +fun (IsoType→IsoSusp is) (merid a i) = merid (fun is a) i +inv (IsoType→IsoSusp is) north = north +inv (IsoType→IsoSusp is) south = south +inv (IsoType→IsoSusp is) (merid a i) = merid (inv is a) i +rightInv (IsoType→IsoSusp is) north = refl +rightInv (IsoType→IsoSusp is) south = refl +rightInv (IsoType→IsoSusp is) (merid a i) j = merid (rightInv is a j) i +leftInv (IsoType→IsoSusp is) north = refl +leftInv (IsoType→IsoSusp is) south = refl +leftInv (IsoType→IsoSusp is) (merid a i) j = merid (leftInv is a j) i + +IsoSuspS²SuspSuspS¹ : Iso (Susp S²) (Susp (Susp S¹)) +IsoSuspS²SuspSuspS¹ = IsoType→IsoSusp S²IsoSuspS¹ + +IsoS³S3 : Iso S³ (Susp (Susp S¹)) +IsoS³S3 = compIso S³IsoSuspS² IsoSuspS²SuspSuspS¹ diff --git a/Cubical/HITs/Susp/Properties.agda b/Cubical/HITs/Susp/Properties.agda new file mode 100644 index 000000000..af80b76f5 --- /dev/null +++ b/Cubical/HITs/Susp/Properties.agda @@ -0,0 +1,119 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Susp.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Path + +open import Cubical.Data.Bool +open import Cubical.HITs.Join +open import Cubical.HITs.Susp.Base + +open Iso + +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 +fun Susp-iso-joinBool (merid a i) = (sym (push a true) ∙ push a false) i +inv Susp-iso-joinBool (inr true ) = north +inv Susp-iso-joinBool (inr false) = south +inv Susp-iso-joinBool (inl _) = north +inv Susp-iso-joinBool (push a true i) = north +inv Susp-iso-joinBool (push a false i) = merid a i +rightInv Susp-iso-joinBool (inr true ) = refl +rightInv Susp-iso-joinBool (inr false) = refl +rightInv Susp-iso-joinBool (inl a) = sym (push a true) +rightInv Susp-iso-joinBool (push a true i) j = push a true (i ∨ ~ j) +rightInv Susp-iso-joinBool (push a false i) j + = hcomp (λ k → λ { (i = i0) → push a true (~ j) + ; (i = i1) → push a false k + ; (j = i1) → push a false (i ∧ k) }) + (push a true (~ i ∧ ~ j)) +leftInv Susp-iso-joinBool north = refl +leftInv Susp-iso-joinBool south = refl +leftInv (Susp-iso-joinBool {A = A}) (merid a i) j + = hcomp (λ k → λ { (i = i0) → transp (λ _ → Susp A) (k ∨ j) north + ; (i = i1) → transp (λ _ → Susp A) (k ∨ j) (merid a k) + ; (j = i1) → merid a (i ∧ k) }) + (transp (λ _ → Susp A) j north) + +Susp≃joinBool : ∀ {ℓ} {A : Type ℓ} → Susp A ≃ join A Bool +Susp≃joinBool = isoToEquiv Susp-iso-joinBool + +Susp≡joinBool : ∀ {ℓ} {A : Type ℓ} → Susp A ≡ join A Bool +Susp≡joinBool = isoToPath Susp-iso-joinBool + +congSuspIso : ∀ {ℓ} {A B : Type ℓ} → Iso A B → Iso (Susp A) (Susp B) +fun (congSuspIso is) = suspFun (fun is) +inv (congSuspIso is) = suspFun (inv is) +rightInv (congSuspIso is) north = refl +rightInv (congSuspIso is) south = refl +rightInv (congSuspIso is) (merid a i) j = merid (rightInv is a j) i +leftInv (congSuspIso is) north = refl +leftInv (congSuspIso is) south = refl +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)) + +suspToPropElim : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Susp A → Type ℓ'} (a : A) + → ((x : Susp A) → isProp (B x)) + → B north + → (x : Susp A) → B x +suspToPropElim a isProp Bnorth north = Bnorth +suspToPropElim {B = B} a isProp Bnorth south = subst B (merid a) Bnorth +suspToPropElim {B = B} a isProp Bnorth (merid a₁ i) = + isOfHLevel→isOfHLevelDep 1 isProp Bnorth (subst B (merid a) Bnorth) (merid a₁) i + +suspToPropElim2 : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Susp A → Susp A → Type ℓ'} (a : A) + → ((x y : Susp A) → isProp (B x y)) + → B north north + → (x y : Susp A) → B x y +suspToPropElim2 _ _ Bnorth north north = Bnorth +suspToPropElim2 {B = B} a _ Bnorth north south = subst (B north) (merid a) Bnorth +suspToPropElim2 {B = B} a isprop Bnorth north (merid x i) = + isProp→PathP (λ i → isprop north (merid x i)) + Bnorth (subst (B north) (merid a) Bnorth) i +suspToPropElim2 {B = B} a _ Bnorth south north = subst (λ x → B x north) (merid a) Bnorth +suspToPropElim2 {B = B} a _ Bnorth south south = subst (λ x → B x x) (merid a) Bnorth +suspToPropElim2 {B = B} a isprop Bnorth south (merid x i) = + isProp→PathP (λ i → isprop south (merid x i)) + (subst (λ x → B x north) (merid a) Bnorth) + (subst (λ x → B x x) (merid a) Bnorth) i +suspToPropElim2 {B = B} a isprop Bnorth (merid x i) north = + isProp→PathP (λ i → isprop (merid x i) north) + Bnorth (subst (λ x → B x north) (merid a) Bnorth) i +suspToPropElim2 {B = B} a isprop Bnorth (merid x i) south = + isProp→PathP (λ i → isprop (merid x i) south) + (subst (B north) (merid a) Bnorth) + (subst (λ x → B x x) (merid a) Bnorth) i +suspToPropElim2 {B = B} a isprop Bnorth (merid x i) (merid y j) = + isSet→SquareP (λ i j → isOfHLevelSuc 1 (isprop _ _)) + (isProp→PathP (λ i₁ → isprop north (merid y i₁)) Bnorth + (subst (B north) (merid a) Bnorth)) + (isProp→PathP (λ i₁ → isprop south (merid y i₁)) + (subst (λ x₁ → B x₁ north) (merid a) Bnorth) + (subst (λ x₁ → B x₁ x₁) (merid a) Bnorth)) + (isProp→PathP (λ i₁ → isprop (merid x i₁) north) Bnorth + (subst (λ x₁ → B x₁ north) (merid a) Bnorth)) + (isProp→PathP (λ i₁ → isprop (merid x i₁) south) + (subst (B north) (merid a) Bnorth) + (subst (λ x₁ → B x₁ x₁) (merid a) Bnorth)) i j +{- Clever proof: +suspToPropElim2 a isProp Bnorth = + suspToPropElim a (λ x → isOfHLevelΠ 1 λ y → isProp x y) + (suspToPropElim a (λ x → isProp north x) Bnorth) +-} + +funSpaceSuspIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → Iso (Σ[ x ∈ B ] Σ[ y ∈ B ] (A → x ≡ y)) (Susp A → B) +Iso.fun funSpaceSuspIso (x , y , f) north = x +Iso.fun funSpaceSuspIso (x , y , f) south = y +Iso.fun funSpaceSuspIso (x , y , f) (merid a i) = f a i +Iso.inv funSpaceSuspIso f = (f north) , (f south , (λ x → cong f (merid x))) +Iso.rightInv funSpaceSuspIso f = funExt λ {north → refl + ; south → refl + ; (merid a i) → refl} +Iso.leftInv funSpaceSuspIso _ = refl diff --git a/Cubical/HITs/Torus.agda b/Cubical/HITs/Torus.agda index be368a3b0..18acd0243 100644 --- a/Cubical/HITs/Torus.agda +++ b/Cubical/HITs/Torus.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Torus where open import Cubical.HITs.Torus.Base public diff --git a/Cubical/HITs/Torus/Base.agda b/Cubical/HITs/Torus/Base.agda index 63777f1c6..662cda683 100644 --- a/Cubical/HITs/Torus/Base.agda +++ b/Cubical/HITs/Torus/Base.agda @@ -4,7 +4,7 @@ Definition of the torus as a HIT together with a proof that it is equivalent to two circles -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Torus.Base where open import Cubical.Foundations.Prelude @@ -14,7 +14,7 @@ open import Cubical.Foundations.Univalence open import Cubical.Data.Nat open import Cubical.Data.Int -open import Cubical.Data.Prod hiding (_×_) renaming (_×Σ_ to _×_) +open import Cubical.Data.Sigma open import Cubical.HITs.S1 @@ -71,18 +71,18 @@ lemPathAnd t u = isoToPath (iso (λ tu → (λ i → tu i .fst) , λ i → tu i (λ y → refl) (λ x → refl)) -ΩTorus≡Int×Int : ΩTorus ≡ Int × Int -ΩTorus≡Int×Int = +ΩTorus≡ℤ×ℤ : ΩTorus ≡ ℤ × ℤ +ΩTorus≡ℤ×ℤ = ΩTorus ≡⟨ (λ i → Loop (point-path i)) ⟩ Loop (base , base) ≡⟨ lemPathAnd (base , base) (base , base) ⟩ ΩS¹ × ΩS¹ - ≡⟨ (λ i → ΩS¹≡Int i × ΩS¹≡Int i) ⟩ - Int × Int ∎ + ≡⟨ (λ i → ΩS¹≡ℤ i × ΩS¹≡ℤ i) ⟩ + ℤ × ℤ ∎ -- Computing the winding numbers on the torus -windingTorus : ΩTorus → Int × Int +windingTorus : ΩTorus → ℤ × ℤ windingTorus l = ( winding (λ i → t2c (l i) .fst) , winding (λ i → t2c (l i) .snd)) diff --git a/Cubical/HITs/Truncation.agda b/Cubical/HITs/Truncation.agda index 35e9c7d66..3b716c88a 100644 --- a/Cubical/HITs/Truncation.agda +++ b/Cubical/HITs/Truncation.agda @@ -1,7 +1,7 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.HITs.Truncation where open import Cubical.HITs.Truncation.Base public open import Cubical.HITs.Truncation.Properties public -import Cubical.HITs.Truncation.FromNegOne +import Cubical.HITs.Truncation.FromNegTwo diff --git a/Cubical/HITs/Truncation/Base.agda b/Cubical/HITs/Truncation/Base.agda index ef1993225..643af521b 100644 --- a/Cubical/HITs/Truncation/Base.agda +++ b/Cubical/HITs/Truncation/Base.agda @@ -1,24 +1,36 @@ -{-# OPTIONS --cubical --safe #-} +{- + +A simpler definition of truncation ∥ A ∥ n from n ≥ -1 + +Note that this uses the HoTT book's indexing, so it will be off + from `∥_∥_` in HITs.Truncation.Base by -2 + +-} +{-# OPTIONS --safe #-} module Cubical.HITs.Truncation.Base where +open import Cubical.Data.NatMinusOne open import Cubical.Foundations.Prelude -open import Cubical.Data.NatMinusOne using (ℕ₋₁; neg1; suc) -open import Cubical.Data.NatMinusTwo -open import Cubical.HITs.Nullification -open import Cubical.HITs.Sn - --- For the hub-and-spoke construction discussed in the HoTT book, which only works for n ≥ -1, see --- `HITs.Truncation.FromNegOne`. The definition of truncation here contains two more constructors --- which are redundant when n ≥ -1 but give contractibility when n = -2. - --- data ∥_∥_ {ℓ} (A : Type ℓ) (n : ℕ₋₂) : Type (ℓ-max ℓ ℓ') where --- -- the hub-and-spoke definition in `Truncation.FromNegOne` --- ∣_∣ : A → Null S A --- hub : (f : S (1+ n) → ∥ A ∥ n) → ∥ A ∥ n --- spoke : (f : S (1+ n) → ∥ A ∥ n) (s : S) → hub f ≡ f s --- -- two additional constructors needed to ensure that ∥ A ∥ -2 is contractible --- ≡hub : ∀ {x y} (p : S (1+ n) → x ≡ y) → x ≡ y --- ≡spoke : ∀ {x y} (p : S (1+ n) → x ≡ y) (s : S) → ≡hub p ≡ p s - -∥_∥_ : ∀ {ℓ} → Type ℓ → ℕ₋₂ → Type ℓ -∥ A ∥ n = Null (S (1+ n)) A +open import Cubical.Foundations.HLevels +open import Cubical.HITs.Sn.Base +open import Cubical.Data.Nat.Base +open import Cubical.Data.Unit.Base +open import Cubical.Data.Empty + +-- this definition is off by one. Use hLevelTrunc or ∥_∥ for truncations +-- (off by 2 w.r.t. the HoTT-book) +data HubAndSpoke {ℓ} (A : Type ℓ) (n : ℕ) : Type ℓ where + ∣_∣ : A → HubAndSpoke A n + hub : (f : S₊ n → HubAndSpoke A n) → HubAndSpoke A n + spoke : (f : S₊ n → HubAndSpoke A n) (x : S₊ n) → hub f ≡ f x + +hLevelTrunc : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → Type ℓ +hLevelTrunc zero A = Unit* +hLevelTrunc (suc n) A = HubAndSpoke A n + +∥_∥_ : ∀ {ℓ} (A : Type ℓ) (n : ℕ) → Type ℓ +∥ A ∥ n = hLevelTrunc n A + +∣_∣ₕ : ∀ {ℓ} {A : Type ℓ} {n : ℕ} → A → ∥ A ∥ n +∣_∣ₕ {n = zero} a = tt* +∣_∣ₕ {n = suc n} a = ∣ a ∣ diff --git a/Cubical/HITs/Truncation/FromNegOne.agda b/Cubical/HITs/Truncation/FromNegOne.agda deleted file mode 100644 index 5ef7baaff..000000000 --- a/Cubical/HITs/Truncation/FromNegOne.agda +++ /dev/null @@ -1,5 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Truncation.FromNegOne where - -open import Cubical.HITs.Truncation.FromNegOne.Base public -open import Cubical.HITs.Truncation.FromNegOne.Properties public diff --git a/Cubical/HITs/Truncation/FromNegOne/Base.agda b/Cubical/HITs/Truncation/FromNegOne/Base.agda deleted file mode 100644 index 0f029f64b..000000000 --- a/Cubical/HITs/Truncation/FromNegOne/Base.agda +++ /dev/null @@ -1,17 +0,0 @@ -{- - -An simpler definition of truncation from n ≥ -1 - --} -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Truncation.FromNegOne.Base where - -open import Cubical.Data.NatMinusOne -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.HLevels -open import Cubical.HITs.Sn - -data ∥_∥_ {ℓ} (A : Type ℓ) (n : ℕ₋₁) : Type ℓ where - ∣_∣ : A → ∥ A ∥ n - hub : (f : S (suc n) → ∥ A ∥ n) → ∥ A ∥ n - spoke : (f : S (suc n) → ∥ A ∥ n) (x : S (suc n)) → hub f ≡ f x diff --git a/Cubical/HITs/Truncation/FromNegOne/Properties.agda b/Cubical/HITs/Truncation/FromNegOne/Properties.agda deleted file mode 100644 index 2fe7e07f1..000000000 --- a/Cubical/HITs/Truncation/FromNegOne/Properties.agda +++ /dev/null @@ -1,180 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Truncation.FromNegOne.Properties where - -open import Cubical.HITs.Truncation.FromNegOne.Base -open import Cubical.Data.Nat -open import Cubical.Data.NatMinusOne -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.HLevels -open import Cubical.HITs.Sn -open import Cubical.Data.Empty -open import Cubical.HITs.Susp - -open import Cubical.HITs.PropositionalTruncation renaming (∥_∥ to ∥_∥₋₁) -open import Cubical.HITs.SetTruncation -open import Cubical.HITs.GroupoidTruncation -open import Cubical.HITs.2GroupoidTruncation - -private - variable - ℓ ℓ' : Level - A : Type ℓ - -open import Cubical.HITs.Truncation.Properties using (sphereFill; isSphereFilled) - -isSphereFilled∥∥ : {n : ℕ₋₁} → isSphereFilled (suc n) (∥ A ∥ n) -isSphereFilled∥∥ f = (hub f) , (spoke f) - -isSphereFilled→isOfHLevel : (n : ℕ₋₁) → isSphereFilled (suc n) A → isOfHLevel (1 + 1+ n) A -isSphereFilled→isOfHLevel {A = A} neg1 h x y = sym (snd (h f) north) ∙ snd (h f) south - where - f : Susp ⊥ → A - f north = x - f south = y - f (merid () i) -isSphereFilled→isOfHLevel {A = A} (suc n) h x y = isSphereFilled→isOfHLevel n (helper h x y) - where - helper : {n : ℕ₋₁} → isSphereFilled (suc n) A → (x y : A) → isSphereFilled n (x ≡ y) - helper {n = n} h x y f = l , r - where - f' : Susp (S n) → A - f' north = x - f' south = y - f' (merid u i) = f u i - - u : sphereFill (suc n) f' - u = h f' - - z : A - z = fst u - - p : z ≡ x - p = snd u north - - q : z ≡ y - q = snd u south - - l : x ≡ y - l = sym p ∙ q - - r : (s : S n) → l ≡ f s - r s i j = hcomp - (λ k → - λ { (i = i0) → compPath-filler (sym p) q k j - ; (i = i1) → snd u (merid s j) k - ; (j = i0) → p (k ∨ (~ i)) - ; (j = i1) → q k - }) - (p ((~ i) ∧ (~ j))) - -isOfHLevel→isSphereFilled : (n : ℕ₋₁) → isOfHLevel (1 + 1+ n) A → isSphereFilled (suc n) A -isOfHLevel→isSphereFilled neg1 h f = (f north) , (λ _ → h _ _) -isOfHLevel→isSphereFilled {A = A} (suc n) h = helper λ x y → isOfHLevel→isSphereFilled n (h x y) - where - helper : {n : ℕ₋₁} → ((x y : A) → isSphereFilled n (x ≡ y)) → isSphereFilled (suc n) A - helper {n = n} h f = l , r - where - l : A - l = f north - - f' : S n → f north ≡ f south - f' x i = f (merid x i) - - h' : sphereFill n f' - h' = h (f north) (f south) f' - - r : (x : S (suc n)) → l ≡ f x - r north = refl - r south = h' .fst - r (merid x i) j = hcomp (λ k → λ { (i = i0) → f north - ; (i = i1) → h' .snd x (~ k) j - ; (j = i0) → f north - ; (j = i1) → f (merid x i) }) (f (merid x (i ∧ j))) - -isOfHLevel∥∥ : (n : ℕ₋₁) → isOfHLevel (1 + 1+ n) (∥ A ∥ n) -isOfHLevel∥∥ n = isSphereFilled→isOfHLevel n isSphereFilled∥∥ - -ind : {n : ℕ₋₁} - {B : ∥ A ∥ n → Type ℓ'} - (hB : (x : ∥ A ∥ n) → isOfHLevel (1 + 1+ n) (B x)) - (g : (a : A) → B (∣ a ∣)) - (x : ∥ A ∥ n) → - B x -ind hB g (∣ a ∣ ) = g a -ind {B = B} hB g (hub f) = - isOfHLevel→isSphereFilled _ (hB (hub f)) (λ x → subst B (sym (spoke f x)) (ind hB g (f x)) ) .fst -ind {B = B} hB g (spoke f x i) = - toPathP {A = λ i → B (spoke f x (~ i))} - (sym (isOfHLevel→isSphereFilled _ (hB (hub f)) (λ x → subst B (sym (spoke f x)) (ind hB g (f x))) .snd x)) - (~ i) - -ind2 : {n : ℕ₋₁} - {B : ∥ A ∥ n → ∥ A ∥ n → Type ℓ'} - (hB : ((x y : ∥ A ∥ n) → isOfHLevel (1 + 1+ n) (B x y))) - (g : (a b : A) → B ∣ a ∣ ∣ b ∣) - (x y : ∥ A ∥ n) → - B x y -ind2 {n = n} hB g = ind (λ _ → isOfHLevelPi (1 + 1+ n) (λ _ → hB _ _)) λ a → - ind (λ _ → hB _ _) (λ b → g a b) - -ind3 : {n : ℕ₋₁} - {B : (x y z : ∥ A ∥ n) → Type ℓ'} - (hB : ((x y z : ∥ A ∥ n) → isOfHLevel (1 + 1+ n) (B x y z))) - (g : (a b c : A) → B (∣ a ∣) ∣ b ∣ ∣ c ∣) - (x y z : ∥ A ∥ n) → - B x y z -ind3 {n = n} hB g = ind2 (λ _ _ → isOfHLevelPi (1 + 1+ n) (hB _ _)) λ a b → - ind (λ _ → hB _ _ _) (λ c → g a b c) - -idemTrunc : (n : ℕ₋₁) → isOfHLevel (1 + 1+ n) A → (∥ A ∥ n) ≃ A -idemTrunc {A = A} n hA = isoToEquiv (iso f g f-g g-f) - where - f : ∥ A ∥ n → A - f = ind (λ _ → hA) λ a → a - - g : A → ∥ A ∥ n - g = ∣_∣ - - f-g : ∀ a → f (g a) ≡ a - f-g a = refl - - g-f : ∀ x → g (f x) ≡ x - g-f = ind (λ _ → isOfHLevelPath (1 + 1+ n) (isOfHLevel∥∥ n) _ _) (λ _ → refl) - -propTrunc≃Trunc-1 : ∥ A ∥₋₁ ≃ ∥ A ∥ -1 -propTrunc≃Trunc-1 = - isoToEquiv - (iso - (elimPropTrunc (λ _ → isOfHLevel∥∥ -1) ∣_∣) - (ind (λ _ → propTruncIsProp) ∣_∣) - (ind (λ _ → isOfHLevelPath 1 (isOfHLevel∥∥ -1) _ _) (λ _ → refl)) - (elimPropTrunc (λ _ → isOfHLevelPath 1 squash _ _) (λ _ → refl))) - -setTrunc≃Trunc0 : ∥ A ∥₀ ≃ ∥ A ∥ 0 -setTrunc≃Trunc0 = - isoToEquiv - (iso - (elimSetTrunc (λ _ → isOfHLevel∥∥ 0) ∣_∣) - (ind (λ _ → squash₀) ∣_∣₀) - (ind (λ _ → isOfHLevelPath 2 (isOfHLevel∥∥ 0) _ _) (λ _ → refl)) - (elimSetTrunc (λ _ → isOfHLevelPath 2 squash₀ _ _) (λ _ → refl))) - -groupoidTrunc≃Trunc1 : ∥ A ∥₁ ≃ ∥ A ∥ 1 -groupoidTrunc≃Trunc1 = - isoToEquiv - (iso - (groupoidTruncElim _ _ (λ _ → isOfHLevel∥∥ 1) ∣_∣) - (ind (λ _ → squash₁) ∣_∣₁) - (ind (λ _ → isOfHLevelPath 3 (isOfHLevel∥∥ 1) _ _) (λ _ → refl)) - (groupoidTruncElim _ _ (λ _ → isOfHLevelPath 3 squash₁ _ _) (λ _ → refl))) - -2groupoidTrunc≃Trunc2 : ∥ A ∥₂ ≃ ∥ A ∥ 2 -2groupoidTrunc≃Trunc2 = - isoToEquiv - (iso - (g2TruncElim _ _ (λ _ → isOfHLevel∥∥ 2) ∣_∣) - (ind (λ _ → squash₂) ∣_∣₂) - (ind (λ _ → isOfHLevelPath 4 (isOfHLevel∥∥ 2) _ _) (λ _ → refl)) - (g2TruncElim _ _ (λ _ → isOfHLevelPath 4 squash₂ _ _) (λ _ → refl))) diff --git a/Cubical/HITs/Truncation/FromNegTwo.agda b/Cubical/HITs/Truncation/FromNegTwo.agda new file mode 100644 index 000000000..257017b24 --- /dev/null +++ b/Cubical/HITs/Truncation/FromNegTwo.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Truncation.FromNegTwo where + +open import Cubical.HITs.Truncation.FromNegTwo.Base public +open import Cubical.HITs.Truncation.FromNegTwo.Properties public diff --git a/Cubical/HITs/Truncation/FromNegTwo/Base.agda b/Cubical/HITs/Truncation/FromNegTwo/Base.agda new file mode 100644 index 000000000..605b23d88 --- /dev/null +++ b/Cubical/HITs/Truncation/FromNegTwo/Base.agda @@ -0,0 +1,29 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Truncation.FromNegTwo.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Nat +open import Cubical.Data.NatMinusOne +open import Cubical.HITs.Nullification +open import Cubical.HITs.Sn.Base + +-- For the hub-and-spoke construction discussed in the HoTT book, which doesn't work in the base case +-- of contractibility, see `HITs.Truncation.Base`. The definition of truncation here contains +-- two more constructors which are redundant when n ≥ 1 but give contractibility when n = 0. + +-- data hLevelTrunc {ℓ} (n : HLevel) (A : Type ℓ) : Type (ℓ-max ℓ ℓ') where +-- -- the hub-and-spoke definition in `Truncation.Base` +-- ∣_∣ : A → hLevelTrunc n A +-- hub : (f : S (-1+ n) → hLevelTrunc n A) → hLevelTrunc n A +-- spoke : (f : S (-1+ n) → hLevelTrunc n A) (s : S) → hub f ≡ f s +-- -- two additional constructors needed to ensure that hLevelTrunc 0 A is contractible +-- ≡hub : ∀ {x y} (p : S (-1+ n) → x ≡ y) → x ≡ y +-- ≡spoke : ∀ {x y} (p : S (-1+ n) → x ≡ y) (s : S (-1+ n)) → ≡hub p ≡ p s + +hLevelTrunc : ∀ {ℓ} → HLevel → Type ℓ → Type ℓ +hLevelTrunc n A = Null (S (-1+ n)) A + +-- Note that relative to the HoTT book, this notation is off by +2 +∥_∥_ : ∀ {ℓ} → Type ℓ → HLevel → Type ℓ +∥ A ∥ n = hLevelTrunc n A diff --git a/Cubical/HITs/Truncation/FromNegTwo/Properties.agda b/Cubical/HITs/Truncation/FromNegTwo/Properties.agda new file mode 100644 index 000000000..a8360dfbc --- /dev/null +++ b/Cubical/HITs/Truncation/FromNegTwo/Properties.agda @@ -0,0 +1,396 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Truncation.FromNegTwo.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Equiv.PathSplit +open isPathSplitEquiv +open import Cubical.Modalities.Modality +open Modality + +open import Cubical.Data.Empty as ⊥ using (⊥) +open import Cubical.Data.Nat hiding (elim) +open import Cubical.Data.NatMinusOne as ℕ₋₁ +open import Cubical.Data.Sigma +open import Cubical.HITs.Sn.Base +open import Cubical.HITs.Susp +open import Cubical.HITs.Nullification as Null hiding (rec; elim) + +open import Cubical.HITs.Truncation.FromNegTwo.Base + +open import Cubical.HITs.PropositionalTruncation as PropTrunc + renaming (∥_∥ to ∥_∥₁; ∣_∣ to ∣_∣₁; squash to squash₁) using () +open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂; ∣_∣₂; squash₂) +open import Cubical.HITs.GroupoidTruncation as GpdTrunc using (∥_∥₃; ∣_∣₃; squash₃) +open import Cubical.HITs.2GroupoidTruncation as 2GpdTrunc using (∥_∥₄; ∣_∣₄; squash₄) + +private + variable + ℓ ℓ' : Level + A B : Type ℓ + +sphereFill : (n : ℕ₋₁) (f : S n → A) → Type _ +sphereFill {A = A} n f = Σ[ top ∈ A ] ((x : S n) → top ≡ f x) + +isSphereFilled : ℕ₋₁ → Type ℓ → Type ℓ +isSphereFilled n A = (f : S n → A) → sphereFill n f + +isSphereFilledTrunc : {n : HLevel} → isSphereFilled (-1+ n) (hLevelTrunc n A) +isSphereFilledTrunc {n = zero} f = hub f , ⊥.elim +isSphereFilledTrunc {n = suc n} f = hub f , spoke f + +isSphereFilled→isOfHLevelSuc : {n : HLevel} → isSphereFilled (ℕ→ℕ₋₁ n) A → isOfHLevel (suc n) A +isSphereFilled→isOfHLevelSuc {A = A} {zero} h x y = sym (snd (h f) north) ∙ snd (h f) south + where + f : Susp ⊥ → A + f north = x + f south = y +isSphereFilled→isOfHLevelSuc {A = A} {suc n} h x y = isSphereFilled→isOfHLevelSuc (helper h x y) + where + helper : isSphereFilled (ℕ→ℕ₋₁ (suc n)) A → (x y : A) → isSphereFilled (ℕ→ℕ₋₁ n) (x ≡ y) + helper h x y f = sym p ∙ q , r + where + f' : Susp (S (ℕ→ℕ₋₁ n)) → A + f' north = x + f' south = y + f' (merid u i) = f u i + + p = snd (h f') north + q = snd (h f') south + + r : (s : S (ℕ→ℕ₋₁ n)) → sym p ∙ q ≡ f s + r s i j = hcomp (λ k → λ { (i = i0) → compPath-filler (sym p) q k j + ; (i = i1) → snd (h f') (merid s j) k + ; (j = i0) → p (k ∨ ~ i) + ; (j = i1) → q k }) + (p (~ i ∧ ~ j)) + +isOfHLevel→isSphereFilled : {n : HLevel} → isOfHLevel n A → isSphereFilled (-1+ n) A +isOfHLevel→isSphereFilled {n = zero} h f = fst h , λ _ → snd h _ +isOfHLevel→isSphereFilled {n = suc zero} h f = f north , λ _ → h _ _ +isOfHLevel→isSphereFilled {A = A} {suc (suc n)} h = helper λ x y → isOfHLevel→isSphereFilled (h x y) + where + helper : {n : HLevel} → ((x y : A) → isSphereFilled (-1+ n) (x ≡ y)) → isSphereFilled (suc₋₁ (-1+ n)) A + helper {n} h f = l , r + where + l : A + l = f north + + f' : S (-1+ n) → f north ≡ f south + f' x i = f (merid x i) + + h' : sphereFill (-1+ n) f' + h' = h (f north) (f south) f' + + r : (x : S (suc₋₁ (-1+ n))) → l ≡ f x + r north = refl + r south = h' .fst + r (merid x i) j = hcomp (λ k → λ { (i = i0) → f north + ; (i = i1) → h' .snd x (~ k) j + ; (j = i0) → f north + ; (j = i1) → f (merid x i) }) + (f (merid x (i ∧ j))) + +-- isNull (S n) A ≃ (isSphereFilled n A) × (∀ (x y : A) → isSphereFilled n (x ≡ y)) + +isOfHLevel→isSnNull : {n : HLevel} → isOfHLevel n A → isNull (S (-1+ n)) A +fst (sec (isOfHLevel→isSnNull h)) f = fst (isOfHLevel→isSphereFilled h f) +snd (sec (isOfHLevel→isSnNull h)) f i s = snd (isOfHLevel→isSphereFilled h f) s i +fst (secCong (isOfHLevel→isSnNull h) x y) p = fst (isOfHLevel→isSphereFilled (isOfHLevelPath _ h x y) (funExt⁻ p)) +snd (secCong (isOfHLevel→isSnNull h) x y) p i j s = snd (isOfHLevel→isSphereFilled (isOfHLevelPath _ h x y) (funExt⁻ p)) s i j + +isSnNull→isOfHLevel : {n : HLevel} → isNull (S (-1+ n)) A → isOfHLevel n A +isSnNull→isOfHLevel {n = zero} nA = fst (sec nA) ⊥.rec , λ y → fst (secCong nA _ y) (funExt ⊥.elim) +isSnNull→isOfHLevel {n = suc n} nA = isSphereFilled→isOfHLevelSuc (λ f → fst (sec nA) f , λ s i → snd (sec nA) f i s) + +isOfHLevelTrunc : (n : HLevel) → isOfHLevel n (hLevelTrunc n A) +isOfHLevelTrunc zero = hub ⊥.rec , λ _ → ≡hub ⊥.rec +isOfHLevelTrunc (suc n) = isSphereFilled→isOfHLevelSuc isSphereFilledTrunc + +-- isOfHLevelTrunc n = isSnNull→isOfHLevel isNull-Null + +-- hLevelTrunc n is a modality + +rec : {n : HLevel} + {B : Type ℓ'} → + isOfHLevel n B → + (A → B) → + hLevelTrunc n A → + B +rec h = Null.rec (isOfHLevel→isSnNull h) + +elim : {n : HLevel} + {B : hLevelTrunc n A → Type ℓ'} + (hB : (x : hLevelTrunc n A) → isOfHLevel n (B x)) + (g : (a : A) → B (∣ a ∣)) + (x : hLevelTrunc n A) → + B x +elim hB = Null.elim (λ x → isOfHLevel→isSnNull (hB x)) + +elim2 : {n : HLevel} + {B : hLevelTrunc n A → hLevelTrunc n A → Type ℓ'} + (hB : ((x y : hLevelTrunc n A) → isOfHLevel n (B x y))) + (g : (a b : A) → B ∣ a ∣ ∣ b ∣) + (x y : hLevelTrunc n A) → + B x y +elim2 {n = n} hB g = + elim (λ _ → isOfHLevelΠ n (λ _ → hB _ _)) + (λ a → elim (λ _ → hB _ _) (λ b → g a b)) + +elim3 : {n : HLevel} + {B : (x y z : hLevelTrunc n A) → Type ℓ'} + (hB : ((x y z : hLevelTrunc n A) → isOfHLevel n (B x y z))) + (g : (a b c : A) → B (∣ a ∣) ∣ b ∣ ∣ c ∣) + (x y z : hLevelTrunc n A) → + B x y z +elim3 {n = n} hB g = + elim2 (λ _ _ → isOfHLevelΠ n (hB _ _)) + (λ a b → elim (λ _ → hB _ _ _) (λ c → g a b c)) + +HLevelTruncModality : ∀ {ℓ} (n : HLevel) → Modality ℓ +isModal (HLevelTruncModality n) = isOfHLevel n +isPropIsModal (HLevelTruncModality n) = isPropIsOfHLevel n +◯ (HLevelTruncModality n) = hLevelTrunc n +◯-isModal (HLevelTruncModality n) = isOfHLevelTrunc n +η (HLevelTruncModality n) = ∣_∣ +◯-elim (HLevelTruncModality n) = elim +◯-elim-β (HLevelTruncModality n) = λ _ _ _ → refl +◯-=-isModal (HLevelTruncModality n) = isOfHLevelPath n (isOfHLevelTrunc n) + +truncIdempotentIso : (n : HLevel) → isOfHLevel n A → Iso A (hLevelTrunc n A) +truncIdempotentIso n hA = isModalToIso (HLevelTruncModality n) hA + +truncIdempotent≃ : (n : HLevel) → isOfHLevel n A → A ≃ hLevelTrunc n A +truncIdempotent≃ n hA = ∣_∣ , isModalToIsEquiv (HLevelTruncModality n) hA + +truncIdempotent : (n : HLevel) → isOfHLevel n A → hLevelTrunc n A ≡ A +truncIdempotent n hA = ua (invEquiv (truncIdempotent≃ n hA)) + +-- universal property + +univTrunc : ∀ {ℓ} (n : HLevel) {B : TypeOfHLevel ℓ n} → Iso (hLevelTrunc n A → B .fst) (A → B .fst) +Iso.fun (univTrunc n {B , lev}) g a = g ∣ a ∣ +Iso.inv (univTrunc n {B , lev}) = elim λ _ → lev +Iso.rightInv (univTrunc n {B , lev}) b = refl +Iso.leftInv (univTrunc n {B , lev}) b = funExt (elim (λ x → isOfHLevelPath _ lev _ _) + λ a → refl) + +-- functorial action + +map : {n : HLevel} {B : Type ℓ'} (g : A → B) + → hLevelTrunc n A → hLevelTrunc n B +map g = rec (isOfHLevelTrunc _) (λ a → ∣ g a ∣) + +mapCompIso : {n : HLevel} {B : Type ℓ'} → (Iso A B) → Iso (hLevelTrunc n A) (hLevelTrunc n B) +Iso.fun (mapCompIso g) = map (Iso.fun g) +Iso.inv (mapCompIso g) = map (Iso.inv g) +Iso.rightInv (mapCompIso g) = elim (λ x → isOfHLevelPath _ (isOfHLevelTrunc _) _ _) λ b → cong ∣_∣ (Iso.rightInv g b) +Iso.leftInv (mapCompIso g) = elim (λ x → isOfHLevelPath _ (isOfHLevelTrunc _) _ _) λ a → cong ∣_∣ (Iso.leftInv g a) + +mapId : {n : HLevel} → ∀ t → map {n = n} (idfun A) t ≡ t +mapId {n = n} = + elim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) (λ _ → refl) + +-- equivalences to prop/set/groupoid truncations + + +propTruncTrunc1Iso : Iso ∥ A ∥₁ (∥ A ∥ 1) +Iso.fun propTruncTrunc1Iso = PropTrunc.elim (λ _ → isOfHLevelTrunc 1) ∣_∣ +Iso.inv propTruncTrunc1Iso = elim (λ _ → squash₁) ∣_∣₁ +Iso.rightInv propTruncTrunc1Iso = elim (λ _ → isOfHLevelPath 1 (isOfHLevelTrunc 1) _ _) (λ _ → refl) +Iso.leftInv propTruncTrunc1Iso = PropTrunc.elim (λ _ → isOfHLevelPath 1 squash₁ _ _) (λ _ → refl) + +propTrunc≃Trunc1 : ∥ A ∥₁ ≃ ∥ A ∥ 1 +propTrunc≃Trunc1 = isoToEquiv propTruncTrunc1Iso + +propTrunc≡Trunc1 : ∥ A ∥₁ ≡ ∥ A ∥ 1 +propTrunc≡Trunc1 = ua propTrunc≃Trunc1 + + +setTruncTrunc2Iso : Iso ∥ A ∥₂ (∥ A ∥ 2) +Iso.fun setTruncTrunc2Iso = SetTrunc.elim (λ _ → isOfHLevelTrunc 2) ∣_∣ +Iso.inv setTruncTrunc2Iso = elim (λ _ → squash₂) ∣_∣₂ +Iso.rightInv setTruncTrunc2Iso = elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) (λ _ → refl) +Iso.leftInv setTruncTrunc2Iso = SetTrunc.elim (λ _ → isOfHLevelPath 2 squash₂ _ _) (λ _ → refl) + +setTrunc≃Trunc2 : ∥ A ∥₂ ≃ ∥ A ∥ 2 +setTrunc≃Trunc2 = isoToEquiv setTruncTrunc2Iso + +propTrunc≡Trunc2 : ∥ A ∥₂ ≡ ∥ A ∥ 2 +propTrunc≡Trunc2 = ua setTrunc≃Trunc2 + +groupoidTrunc≃Trunc3Iso : Iso ∥ A ∥₃ (∥ A ∥ 3) +Iso.fun groupoidTrunc≃Trunc3Iso = GpdTrunc.elim (λ _ → isOfHLevelTrunc 3) ∣_∣ +Iso.inv groupoidTrunc≃Trunc3Iso = elim (λ _ → squash₃) ∣_∣₃ +Iso.rightInv groupoidTrunc≃Trunc3Iso = elim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) (λ _ → refl) +Iso.leftInv groupoidTrunc≃Trunc3Iso = GpdTrunc.elim (λ _ → isOfHLevelPath 3 squash₃ _ _) (λ _ → refl) + +groupoidTrunc≃Trunc3 : ∥ A ∥₃ ≃ ∥ A ∥ 3 +groupoidTrunc≃Trunc3 = isoToEquiv groupoidTrunc≃Trunc3Iso + +groupoidTrunc≡Trunc3 : ∥ A ∥₃ ≡ ∥ A ∥ 3 +groupoidTrunc≡Trunc3 = ua groupoidTrunc≃Trunc3 + + +2GroupoidTrunc≃Trunc4Iso : Iso ∥ A ∥₄ (∥ A ∥ 4) +Iso.fun 2GroupoidTrunc≃Trunc4Iso = 2GpdTrunc.elim (λ _ → isOfHLevelTrunc 4) ∣_∣ +Iso.inv 2GroupoidTrunc≃Trunc4Iso = elim (λ _ → squash₄) ∣_∣₄ +Iso.rightInv 2GroupoidTrunc≃Trunc4Iso = elim (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) (λ _ → refl) +Iso.leftInv 2GroupoidTrunc≃Trunc4Iso = 2GpdTrunc.elim (λ _ → isOfHLevelPath 4 squash₄ _ _) (λ _ → refl) + +2GroupoidTrunc≃Trunc4 : ∥ A ∥₄ ≃ ∥ A ∥ 4 +2GroupoidTrunc≃Trunc4 = + isoToEquiv + (iso + (2GpdTrunc.elim (λ _ → isOfHLevelTrunc 4) ∣_∣) + (elim (λ _ → squash₄) ∣_∣₄) + (elim (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) (λ _ → refl)) + (2GpdTrunc.elim (λ _ → isOfHLevelPath 4 squash₄ _ _) (λ _ → refl))) + +2GroupoidTrunc≡Trunc4 : ∥ A ∥₄ ≡ ∥ A ∥ 4 +2GroupoidTrunc≡Trunc4 = ua 2GroupoidTrunc≃Trunc4 + + +isContr→isContrTrunc : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A → isContr (hLevelTrunc n A) +isContr→isContrTrunc n contr = ∣ fst contr ∣ , (elim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → cong ∣_∣ (snd contr a)) + + +truncOfProdIso : (n : ℕ) → Iso (hLevelTrunc n (A × B)) (hLevelTrunc n A × hLevelTrunc n B) +Iso.fun (truncOfProdIso n) = rec (isOfHLevelΣ n (isOfHLevelTrunc n) (λ _ → isOfHLevelTrunc n)) λ {(a , b) → ∣ a ∣ , ∣ b ∣} +Iso.inv (truncOfProdIso n) (a , b) = rec (isOfHLevelTrunc n) + (λ a → rec (isOfHLevelTrunc n) + (λ b → ∣ a , b ∣) + b) + a +Iso.rightInv (truncOfProdIso n) (a , b) = + elim {B = λ a → Iso.fun (truncOfProdIso n) (Iso.inv (truncOfProdIso n) (a , b)) ≡ (a , b)} + (λ _ → isOfHLevelPath n (isOfHLevelΣ n (isOfHLevelTrunc n) (λ _ → isOfHLevelTrunc n)) _ _) + (λ a → elim {B = λ b → Iso.fun (truncOfProdIso n) (Iso.inv (truncOfProdIso n) (∣ a ∣ , b)) ≡ (∣ a ∣ , b)} + (λ _ → isOfHLevelPath n (isOfHLevelΣ n (isOfHLevelTrunc n) (λ _ → isOfHLevelTrunc n)) _ _) + (λ b → refl) b) a +Iso.leftInv (truncOfProdIso n) = elim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl + +---- ∥ Ω A ∥ ₙ ≡ Ω ∥ A ∥ₙ₊₁ ---- + + +isOfHLevelTypeOfHLevel2 : ∀ n → isOfHLevel (suc n) (TypeOfHLevel ℓ n) +isOfHLevelTypeOfHLevel2 n = isOfHLevelTypeOfHLevel n + + {- Proofs of Theorem 7.3.12. and Corollary 7.3.13. in the HoTT book -} + +module ΩTrunc where + {- We define the fibration P to show a more general result -} + P : {X : Type ℓ} {n : HLevel} → ∥ X ∥ (suc n) → ∥ X ∥ (suc n) → Type ℓ + P {n = n} x y = elim2 (λ _ _ → isOfHLevelTypeOfHLevel2 (n)) + (λ a b → ∥ a ≡ b ∥ n , isOfHLevelTrunc (n)) x y .fst + + {- We will need P to be of hLevel n + 3 -} + hLevelP : {n : HLevel} (a b : ∥ B ∥ (suc n)) → isOfHLevel ((suc n)) (P a b) + hLevelP {n = n} = + elim2 (λ x y → isProp→isOfHLevelSuc (n) (isPropIsOfHLevel (suc n))) + (λ a b → isOfHLevelSuc (n) (isOfHLevelTrunc (n))) + + {- decode function from P x y to x ≡ y -} + decode-fun : {n : HLevel} (x y : ∥ B ∥ (suc n)) → P x y → x ≡ y + decode-fun {n = n} = + elim2 (λ u v → isOfHLevelΠ (suc n) + (λ _ → isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n)) u v)) + decode* + where + decode* : ∀ {n : HLevel} (u v : B) + → P {n = n} ∣ u ∣ ∣ v ∣ → Path (∥ B ∥ (suc n)) ∣ u ∣ ∣ v ∣ + decode* {B = B} {n = zero} u v = + rec ( isOfHLevelTrunc 1 ∣ u ∣ ∣ v ∣ + , λ _ → isOfHLevelSuc 1 (isOfHLevelTrunc 1) _ _ _ _) (cong ∣_∣) + decode* {n = suc n} u v = + rec (isOfHLevelTrunc (suc (suc n)) ∣ u ∣ ∣ v ∣) (cong ∣_∣) + + {- auxiliary function r used to define encode -} + r : {m : HLevel} (u : ∥ B ∥ (suc m)) → P u u + r = elim (λ x → hLevelP x x) (λ a → ∣ refl ∣) + + {- encode function from x ≡ y to P x y -} + encode-fun : {n : HLevel} (x y : ∥ B ∥ (suc n)) → x ≡ y → P x y + encode-fun x y p = transport (λ i → P x (p i)) (r x) + + {- We need the following two lemmas on the functions behaviour for refl -} + dec-refl : {n : HLevel} (x : ∥ B ∥ (suc n)) → decode-fun x x (r x) ≡ refl + dec-refl {n = zero} = + elim (λ x → isOfHLevelSuc 1 (isOfHLevelSuc 1 (isOfHLevelTrunc 1) x x) _ _) (λ _ → refl) + dec-refl {n = suc n} = + elim (λ x → isOfHLevelSuc (suc n) + (isOfHLevelSuc (suc n) + (isOfHLevelTrunc (suc (suc n)) x x) + (decode-fun x x (r x)) refl)) + (λ _ → refl) + + enc-refl : {n : HLevel} (x : ∥ B ∥ (suc n)) → encode-fun x x refl ≡ r x + enc-refl x j = transp (λ _ → P x x) j (r x) + + {- decode-fun is a right-inverse -} + P-rinv : {n : HLevel} (u v : ∥ B ∥ (suc n)) (x : Path (∥ B ∥ (suc n)) u v) + → decode-fun u v (encode-fun u v x) ≡ x + P-rinv u v = J (λ y p → decode-fun u y (encode-fun u y p) ≡ p) + (cong (decode-fun u u) (enc-refl u) ∙ dec-refl u) + + {- decode-fun is a left-inverse -} + P-linv : {n : HLevel} (u v : ∥ B ∥ (suc n )) (x : P u v) + → encode-fun u v (decode-fun u v x) ≡ x + P-linv {n = n} = + elim2 (λ x y → isOfHLevelΠ (suc n) + (λ z → isOfHLevelSuc (suc n) (hLevelP x y) _ _)) + helper + where + helper : {n : HLevel} (a b : B) (p : P {n = n} ∣ a ∣ ∣ b ∣) + → encode-fun _ _ (decode-fun ∣ a ∣ ∣ b ∣ p) ≡ p + helper {n = zero} a b = + elim (λ x → ( sym (isOfHLevelTrunc 0 .snd _) ∙ isOfHLevelTrunc 0 .snd x + , λ y → isOfHLevelSuc 1 (isOfHLevelSuc 0 (isOfHLevelTrunc 0)) _ _ _ _)) + (J (λ y p → encode-fun ∣ a ∣ ∣ y ∣ (decode-fun _ _ ∣ p ∣) ≡ ∣ p ∣) + (enc-refl ∣ a ∣)) + helper {n = suc n} a b = + elim (λ x → hLevelP {n = suc n} ∣ a ∣ ∣ b ∣ _ _) + (J (λ y p → encode-fun {n = suc n} ∣ a ∣ ∣ y ∣ (decode-fun _ _ ∣ p ∣) ≡ ∣ p ∣) + (enc-refl ∣ a ∣)) + + {- The final Iso established -} + IsoFinal : (n : HLevel) (x y : ∥ B ∥ (suc n)) → Iso (x ≡ y) (P x y) + Iso.fun (IsoFinal _ x y) = encode-fun x y + Iso.inv (IsoFinal _ x y) = decode-fun x y + Iso.rightInv (IsoFinal _ x y) = P-linv x y + Iso.leftInv (IsoFinal _ x y) = P-rinv x y + +PathIdTrunc : {a b : A} (n : HLevel) → (Path (∥ A ∥ (suc n)) ∣ a ∣ ∣ b ∣) ≡ (∥ a ≡ b ∥ n) +PathIdTrunc n = isoToPath (ΩTrunc.IsoFinal n _ _) + +PathΩ : {a : A} (n : HLevel) → (Path (∥ A ∥ (suc n)) ∣ a ∣ ∣ a ∣) ≡ (∥ a ≡ a ∥ n) +PathΩ n = PathIdTrunc n + +{- Special case using direct defs of truncations -} +PathIdTrunc₀Iso : {a b : A} → Iso (∣ a ∣₂ ≡ ∣ b ∣₂) ∥ a ≡ b ∥₁ +PathIdTrunc₀Iso = compIso (congIso setTruncTrunc2Iso) + (compIso (ΩTrunc.IsoFinal _ ∣ _ ∣ ∣ _ ∣) + (invIso propTruncTrunc1Iso)) + +------------------------- + + +truncOfTruncIso : (n m : HLevel) → Iso (hLevelTrunc n A) (hLevelTrunc n (hLevelTrunc (m + n) A)) +Iso.fun (truncOfTruncIso n m) = elim (λ _ → isOfHLevelTrunc n) λ a → ∣ ∣ a ∣ ∣ +Iso.inv (truncOfTruncIso {A = A} n m) = + elim (λ _ → isOfHLevelTrunc n) + (elim (λ _ → (isOfHLevelPlus m (isOfHLevelTrunc n ))) + λ a → ∣ a ∣) +Iso.rightInv (truncOfTruncIso {A = A} n m) = + elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _ ) + (elim (λ x → isOfHLevelPath (m + n) (isOfHLevelPlus m (isOfHLevelTrunc n)) _ _ ) + λ a → refl) +Iso.leftInv (truncOfTruncIso n m) = elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl + +truncOfTruncEq : (n m : ℕ) → (hLevelTrunc n A) ≃ (hLevelTrunc n (hLevelTrunc (m + n) A)) +truncOfTruncEq n m = isoToEquiv (truncOfTruncIso n m) diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index a8ec17a6c..00495d4d3 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -1,373 +1,530 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} + module Cubical.HITs.Truncation.Properties where +open import Cubical.Data.NatMinusOne +open import Cubical.HITs.Truncation.Base open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.HLevels -open import Cubical.Foundations.PathSplitEquiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Equiv.PathSplit open isPathSplitEquiv -open import Cubical.Modalities.Everything +open import Cubical.Modalities.Modality open Modality -open import Cubical.Data.Empty -open import Cubical.Data.Nat -open import Cubical.Data.NatMinusOne using (ℕ₋₁; neg1; suc; ℕ→ℕ₋₁) renaming (-1+_ to -1+₋₁_ ; 1+_ to 1+₋₁_) -import Cubical.Data.NatMinusOne as ℕ₋₁ -open import Cubical.Data.NatMinusTwo -open import Cubical.HITs.Sn +open import Cubical.Data.Nat hiding (elim) +open import Cubical.Data.Sigma +open import Cubical.Data.Bool +open import Cubical.Data.Unit +open import Cubical.HITs.Sn.Base +open import Cubical.HITs.S1 open import Cubical.HITs.Susp -open import Cubical.HITs.Nullification - -open import Cubical.HITs.Truncation.Base +open import Cubical.HITs.Nullification as Null hiding (rec; elim) -open import Cubical.HITs.PropositionalTruncation renaming (∥_∥ to ∥_∥₋₁) -open import Cubical.HITs.SetTruncation -open import Cubical.HITs.GroupoidTruncation -open import Cubical.HITs.2GroupoidTruncation +open import Cubical.HITs.PropositionalTruncation as PropTrunc + renaming (∥_∥ to ∥_∥₁; ∣_∣ to ∣_∣₁; squash to squash₁) using () +open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂; ∣_∣₂; squash₂) +open import Cubical.HITs.GroupoidTruncation as GpdTrunc using (∥_∥₃; ∣_∣₃; squash₃) +open import Cubical.HITs.2GroupoidTruncation as 2GpdTrunc using (∥_∥₄; ∣_∣₄; squash₄) private variable ℓ ℓ' : Level A : Type ℓ + B : Type ℓ' -sphereFill : (n : ℕ₋₁) (f : S n → A) → Type _ -sphereFill {A = A} n f = Σ[ top ∈ A ] ((x : S n) → top ≡ f x) +sphereFill : (n : ℕ) (f : S₊ n → A) → Type _ +sphereFill {A = A} n f = Σ[ top ∈ A ] ((x : S₊ n) → top ≡ f x) -isSphereFilled : ℕ₋₁ → Type ℓ → Type ℓ -isSphereFilled n A = (f : S n → A) → sphereFill n f +isSphereFilled : ℕ → Type ℓ → Type ℓ +isSphereFilled n A = (f : S₊ n → A) → sphereFill n f -isSphereFilled∥∥ : {n : ℕ₋₂} → isSphereFilled (1+ n) (∥ A ∥ n) -isSphereFilled∥∥ {n = neg2} f = hub f , ⊥-elimDep -isSphereFilled∥∥ {n = suc _} f = hub f , spoke f +isSphereFilled∥∥ : {n : ℕ} → isSphereFilled n (HubAndSpoke A n) +isSphereFilled∥∥ f = (hub f) , (spoke f) -isSphereFilled→isOfHLevelSuc : {n : ℕ₋₂} → isSphereFilled (1+ suc n) A → isOfHLevel (2+ suc n) A -isSphereFilled→isOfHLevelSuc {A = A} {neg2} h x y = sym (snd (h f) north) ∙ snd (h f) south +isSphereFilled→isOfHLevel : (n : ℕ) → isSphereFilled n A → isOfHLevel (suc n) A +isSphereFilled→isOfHLevel {A = A} zero h x y = sym (snd (h f) true) ∙ snd (h f) false + where + f : Bool → A + f true = x + f false = y +isSphereFilled→isOfHLevel {A = A} (suc zero) h x y = + J (λ y q → (p : x ≡ y) → q ≡ p) (helper x) where - f : Susp ⊥ → A - f north = x - f south = y - f (merid () i) -isSphereFilled→isOfHLevelSuc {A = A} {suc n} h x y = isSphereFilled→isOfHLevelSuc (helper h x y) + helper : (x : A) (p : x ≡ x) → refl ≡ p + helper x p i j = + hcomp (λ k → λ { (i = i0) → h S¹→A .snd base k + ; (i = i1) → p j + ; (j = i0) → h S¹→A .snd base (i ∨ k) + ; (j = i1) → h S¹→A .snd base (i ∨ k)}) + (h S¹→A .snd (loop j) i) + where + S¹→A : S¹ → A + S¹→A base = x + S¹→A (loop i) = p i +isSphereFilled→isOfHLevel {A = A} (suc (suc n)) h x y = + isSphereFilled→isOfHLevel (suc n) (helper h x y) where - helper : {n : ℕ₋₂} → isSphereFilled (suc (1+ suc n)) A → (x y : A) → isSphereFilled (1+ suc n) (x ≡ y) - helper {n = n} h x y f = l , r + helper : {n : ℕ} → isSphereFilled (suc (suc n)) A → (x y : A) → isSphereFilled (suc n) (x ≡ y) + helper {n = n} h x y f = sym (snd (h f') north) ∙ (snd (h f') south) , r where - f' : Susp (S (1+ suc n)) → A + f' : Susp (S₊ (suc n)) → A f' north = x f' south = y f' (merid u i) = f u i - u : sphereFill (suc (1+ suc n)) f' - u = h f' - - z : A - z = fst u - - p : z ≡ x - p = snd u north - - q : z ≡ y - q = snd u south - - l : x ≡ y - l = sym p ∙ q - - r : (s : S (1+ suc n)) → l ≡ f s + r : (s : S₊ (suc n)) → sym (snd (h f') north) ∙ (snd (h f') south) ≡ f s r s i j = hcomp (λ k → - λ { (i = i0) → compPath-filler (sym p) q k j - ; (i = i1) → snd u (merid s j) k - ; (j = i0) → p (k ∨ (~ i)) - ; (j = i1) → q k + λ { (i = i1) → snd (h f') (merid s j) k + ; (j = i0) → snd (h f') north (k ∨ (~ i)) + ; (j = i1) → snd (h f') south k }) - (p ((~ i) ∧ (~ j))) + (snd (h f') north (~ i ∧ ~ j)) -isOfHLevel→isSphereFilled : {n : ℕ₋₂} → isOfHLevel (2+ n) A → isSphereFilled (1+ n) A -isOfHLevel→isSphereFilled {A = A} {neg2} h f = fst h , λ _ → snd h _ -isOfHLevel→isSphereFilled {A = A} {suc neg2} h f = f north , λ _ → h _ _ -isOfHLevel→isSphereFilled {A = A} {suc (suc n)} h = helper λ x y → isOfHLevel→isSphereFilled (h x y) +isOfHLevel→isSphereFilled : (n : ℕ) → isOfHLevel (suc n) A → isSphereFilled n A +isOfHLevel→isSphereFilled zero h f = (f true) , (λ _ → h _ _) +isOfHLevel→isSphereFilled {A = A} (suc zero) h f = (f base) , toPropElim (λ _ → h _ _) refl +isOfHLevel→isSphereFilled {A = A} (suc (suc n)) h = + helper λ x y → isOfHLevel→isSphereFilled (suc n) (h x y) where - helper : {n : ℕ₋₂} → ((x y : A) → isSphereFilled (1+ n) (x ≡ y)) → isSphereFilled (suc (1+ n)) A - helper {n = n} h f = l , r + helper : {n : ℕ} → ((x y : A) → isSphereFilled (suc n) (x ≡ y)) + → isSphereFilled (suc (suc n)) A + helper {n = n} h f = f north , r where - l : A - l = f north - - f' : S (1+ n) → f north ≡ f south - f' x i = f (merid x i) - - h' : sphereFill (1+ n) f' - h' = h (f north) (f south) f' - - r : (x : S (suc (1+ n))) → l ≡ f x + r : (x : S₊ (suc (suc n))) → f north ≡ f x r north = refl - r south = h' .fst + r south = h (f north) (f south) (λ x → cong f (merid x)) .fst r (merid x i) j = hcomp (λ k → λ { (i = i0) → f north - ; (i = i1) → h' .snd x (~ k) j - ; (j = i0) → f north - ; (j = i1) → f (merid x i) }) (f (merid x (i ∧ j))) - --- isNull (S n) A ≃ (isSphereFilled n A) × (∀ (x y : A) → isSphereFilled n (x ≡ y)) - -isOfHLevel→isSnNull : {n : ℕ₋₂} → isOfHLevel (2+ n) A → isNull (S (1+ n)) A -fst (sec (isOfHLevel→isSnNull h)) f = fst (isOfHLevel→isSphereFilled h f) -snd (sec (isOfHLevel→isSnNull h)) f i s = snd (isOfHLevel→isSphereFilled h f) s i -fst (secCong (isOfHLevel→isSnNull h) x y) p = fst (isOfHLevel→isSphereFilled (isOfHLevelPath _ h x y) (funExt⁻ p)) -snd (secCong (isOfHLevel→isSnNull h) x y) p i j s = snd (isOfHLevel→isSphereFilled (isOfHLevelPath _ h x y) (funExt⁻ p)) s i j - -isSnNull→isOfHLevel : {n : ℕ₋₂} → isNull (S (1+ n)) A → isOfHLevel (2+ n) A -isSnNull→isOfHLevel {n = neg2} nA = fst (sec nA) ⊥-elim , λ y → fst (secCong nA _ y) (funExt ⊥-elimDep) -isSnNull→isOfHLevel {n = suc n} nA = isSphereFilled→isOfHLevelSuc (λ f → fst (sec nA) f , λ s i → snd (sec nA) f i s) - -isOfHLevel∥∥ : (n : ℕ₋₂) → isOfHLevel (2+ n) (∥ A ∥ n) -isOfHLevel∥∥ neg2 = hub ⊥-elim , λ _ → ≡hub ⊥-elim -isOfHLevel∥∥ (suc n) = isSphereFilled→isOfHLevelSuc isSphereFilled∥∥ - --- isOfHLevel∥∥ n = isSnNull→isOfHLevel isNull-Null - --- ∥_∥ n is a modality - -rec : {n : ℕ₋₂} + ; (i = i1) → h (f north) (f south) (λ x → cong f (merid x)) .snd x (~ k) j + ; (j = i0) → f north + ; (j = i1) → f (merid x i) }) (f (merid x (i ∧ j))) + +isOfHLevelTrunc : (n : ℕ) → isOfHLevel n (∥ A ∥ n) +isOfHLevelTrunc zero = isOfHLevelUnit* 0 +isOfHLevelTrunc (suc n) = isSphereFilled→isOfHLevel n isSphereFilled∥∥ + +isOfHLevelTruncPath : {n : ℕ} {x y : hLevelTrunc n A} → isOfHLevel n (x ≡ y) +isOfHLevelTruncPath {n = n} = isOfHLevelPath n (isOfHLevelTrunc n) _ _ + +rec₊ : {n : HLevel} + {B : Type ℓ'} → + isOfHLevel (suc n) B → + (A → B) → + hLevelTrunc (suc n) A → + B +rec₊ h g ∣ x ∣ = g x +rec₊ {n = n} {B = B} hB g (hub f) = isOfHLevel→isSphereFilled n hB (λ x → rec₊ hB g (f x)) .fst +rec₊ {n = n} hB g (spoke f x i) = + isOfHLevel→isSphereFilled n hB (λ x → rec₊ hB g (f x)) .snd x i + +rec : {n : HLevel} {B : Type ℓ'} → - (isOfHLevel (2+ n) B) → - (g : (a : A) → B) → - (∥ A ∥ n → B) -rec {B = B} h = Null-ind {B = λ _ → B} λ x → isOfHLevel→isSnNull h - -ind : {n : ℕ₋₂} - {B : ∥ A ∥ n → Type ℓ'} - (hB : (x : ∥ A ∥ n) → isOfHLevel (2+ n) (B x)) - (g : (a : A) → B (∣ a ∣)) - (x : ∥ A ∥ n) → - B x -ind hB = Null-ind (λ x → isOfHLevel→isSnNull (hB x)) - -ind2 : {n : ℕ₋₂} - {B : ∥ A ∥ n → ∥ A ∥ n → Type ℓ'} - (hB : ((x y : ∥ A ∥ n) → isOfHLevel (2+ n) (B x y))) - (g : (a b : A) → B ∣ a ∣ ∣ b ∣) - (x y : ∥ A ∥ n) → - B x y -ind2 {n = n} hB g = ind (λ _ → isOfHLevelPi (2+ n) (λ _ → hB _ _)) λ a → - ind (λ _ → hB _ _) (λ b → g a b) - - - - -ind3 : {n : ℕ₋₂} - {B : (x y z : ∥ A ∥ n) → Type ℓ'} - (hB : ((x y z : ∥ A ∥ n) → isOfHLevel (2+ n) (B x y z))) - (g : (a b c : A) → B (∣ a ∣) ∣ b ∣ ∣ c ∣) - (x y z : ∥ A ∥ n) → - B x y z -ind3 {n = n} hB g = ind2 (λ _ _ → isOfHLevelPi (2+ n) (hB _ _)) λ a b → - ind (λ _ → hB _ _ _) (λ c → g a b c) - -TruncModality : ∀ {ℓ} (n : ℕ₋₂) → Modality ℓ -isModal (TruncModality n) = isOfHLevel (2+ n) -isModalIsProp (TruncModality n) = isPropIsOfHLevel (2+ n) -◯ (TruncModality n) = ∥_∥ n -◯-isModal (TruncModality n) = isOfHLevel∥∥ n -η (TruncModality n) = ∣_∣ -◯-elim (TruncModality n) = ind -◯-elim-β (TruncModality n) = λ _ _ _ → refl -◯-=-isModal (TruncModality n) = isOfHLevelPath (2+ n) (isOfHLevel∥∥ n) - -idemTrunc : (n : ℕ₋₂) → isOfHLevel (2+ n) A → A ≃ (∥ A ∥ n) -idemTrunc n hA = ∣_∣ , isModalToIsEquiv (TruncModality n) hA + isOfHLevel n B → + (A → B) → + hLevelTrunc n A → + B +rec {n = zero} h g t = h .fst +rec {n = suc n} = rec₊ + +rec2 : {n : HLevel} + {B : Type ℓ'} → + isOfHLevel n B → + (A → A → B) → + (t u : hLevelTrunc n A) + → B +rec2 h g = rec (isOfHLevelΠ _ (λ _ → h)) (λ a → rec h (λ b → g a b)) + +rec3 : {n : HLevel} + {B : Type ℓ'} → + isOfHLevel n B → + (A → A → A → B) → + (t u v : hLevelTrunc n A) + → B +rec3 h g = rec2 (isOfHLevelΠ _ (λ _ → h)) (λ a b → rec h (λ c → g a b c)) + +elim₊ : {n : ℕ} + {B : ∥ A ∥ (suc n) → Type ℓ'} + (hB : (x : ∥ A ∥ (suc n)) → isOfHLevel (suc n) (B x)) + (g : (a : A) → B (∣ a ∣)) + (x : ∥ A ∥ (suc n)) → + B x +elim₊ hB g (∣ a ∣ ) = g a +elim₊ {B = B} hB g (hub f) = + isOfHLevel→isSphereFilled _ (hB (hub f)) (λ x → subst B (sym (spoke f x)) (elim₊ hB g (f x)) ) .fst +elim₊ {B = B} hB g (spoke f x i) = + toPathP {A = λ i → B (spoke f x (~ i))} + (sym (isOfHLevel→isSphereFilled _ (hB (hub f)) (λ x → subst B (sym (spoke f x)) (elim₊ hB g (f x))) .snd x)) + (~ i) + +elim : {n : ℕ} + {B : ∥ A ∥ n → Type ℓ'} + (hB : (x : ∥ A ∥ n) → isOfHLevel n (B x)) + (g : (a : A) → B (∣ a ∣ₕ)) + (x : ∥ A ∥ n) → + B x +elim {n = zero} h g t = h t .fst +elim {n = suc n} = elim₊ + +elim2 : {n : ℕ} + {B : ∥ A ∥ n → ∥ A ∥ n → Type ℓ'} + (hB : ((x y : ∥ A ∥ n) → isOfHLevel n (B x y))) + (g : (a b : A) → B ∣ a ∣ₕ ∣ b ∣ₕ) + (x y : ∥ A ∥ n) → + B x y +elim2 hB g = elim (λ _ → isOfHLevelΠ _ (λ _ → hB _ _)) λ a → + elim (λ _ → hB _ _) (λ b → g a b) + +elim3 : {n : ℕ} + {B : (x y z : ∥ A ∥ n) → Type ℓ'} + (hB : ((x y z : ∥ A ∥ n) → isOfHLevel n (B x y z))) + (g : (a b c : A) → B (∣ a ∣ₕ) ∣ b ∣ₕ ∣ c ∣ₕ) + (x y z : ∥ A ∥ n) → + B x y z +elim3 hB g = elim2 (λ _ _ → isOfHLevelΠ _ (hB _ _)) λ a b → + elim (λ _ → hB _ _ _) (λ c → g a b c) + +isContr→isContr∥ : (n : ℕ) → isContr A → isContr (∥ A ∥ n) +isContr→isContr∥ zero _ = tt* , (λ _ _ → tt*) +isContr→isContr∥ (suc n) contr = ∣ fst contr ∣ , (elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + (λ a i → ∣ snd contr a i ∣)) + +isOfHLevelMin→isOfHLevel : {n m : ℕ} → isOfHLevel (min n m) A → isOfHLevel n A × isOfHLevel m A +isOfHLevelMin→isOfHLevel {n = zero} {m = m} h = h , isContr→isOfHLevel m h +isOfHLevelMin→isOfHLevel {n = suc n} {m = zero} h = (isContr→isOfHLevel (suc n) h) , h +isOfHLevelMin→isOfHLevel {A = A} {n = suc n} {m = suc m} h = + subst (λ x → isOfHLevel x A) (helper n m) + (isOfHLevelPlus (suc n ∸ (suc (min n m))) h) + , subst (λ x → isOfHLevel x A) ((λ i → m ∸ (minComm n m i) + suc (minComm n m i)) ∙ helper m n) + (isOfHLevelPlus (suc m ∸ (suc (min n m))) h) + where + helper : (n m : ℕ) → n ∸ min n m + suc (min n m) ≡ suc n + helper zero zero = refl + helper zero (suc m) = refl + helper (suc n) zero = cong suc (+-comm n 1) + helper (suc n) (suc m) = +-suc _ _ ∙ cong suc (helper n m) + +ΣTruncElim : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {n m : ℕ} + {B : (x : ∥ A ∥ n) → Type ℓ'} + {C : (Σ[ a ∈ (∥ A ∥ n) ] (∥ B a ∥ m)) → Type ℓ''} + → ((x : (Σ[ a ∈ (∥ A ∥ n) ] (∥ B a ∥ m))) → isOfHLevel (min n m) (C x)) + → ((a : A) (b : B (∣ a ∣ₕ)) → C (∣ a ∣ₕ , ∣ b ∣ₕ)) + → (x : (Σ[ a ∈ (∥ A ∥ n) ] (∥ B a ∥ m))) → C x +ΣTruncElim hB g (a , b) = + elim (λ x → isOfHLevelΠ _ λ b → isOfHLevelMin→isOfHLevel (hB (x , b)) .fst ) + (λ a → elim (λ _ → isOfHLevelMin→isOfHLevel (hB (∣ a ∣ₕ , _)) .snd) λ b → g a b) + a b + +truncIdempotentIso : (n : ℕ) → isOfHLevel n A → Iso (∥ A ∥ n) A +truncIdempotentIso zero hA = isContr→Iso (isOfHLevelUnit* 0) hA +Iso.fun (truncIdempotentIso (suc n) hA) = rec hA λ a → a +Iso.inv (truncIdempotentIso (suc n) hA) = ∣_∣ +Iso.rightInv (truncIdempotentIso (suc n) hA) _ = refl +Iso.leftInv (truncIdempotentIso (suc n) hA) = + elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) λ _ → refl + +truncIdempotent≃ : (n : ℕ) → isOfHLevel n A → ∥ A ∥ n ≃ A +truncIdempotent≃ n hA = isoToEquiv (truncIdempotentIso n hA) + +truncIdempotent : (n : ℕ) → isOfHLevel n A → ∥ A ∥ n ≡ A +truncIdempotent n hA = ua (truncIdempotent≃ n hA) + +HLevelTruncModality : ∀ {ℓ} (n : HLevel) → Modality ℓ +isModal (HLevelTruncModality n) = isOfHLevel n +isPropIsModal (HLevelTruncModality n) = isPropIsOfHLevel n +◯ (HLevelTruncModality n) = hLevelTrunc n +◯-isModal (HLevelTruncModality n) = isOfHLevelTrunc n +η (HLevelTruncModality zero) _ = tt* +η (HLevelTruncModality (suc n)) = ∣_∣ +◯-elim (HLevelTruncModality zero) cB _ tt* = cB tt* .fst +◯-elim (HLevelTruncModality (suc n)) = elim +◯-elim-β (HLevelTruncModality zero) cB f a = cB tt* .snd (f a) +◯-elim-β (HLevelTruncModality (suc n)) = λ _ _ _ → refl +◯-=-isModal (HLevelTruncModality zero) x y = (isOfHLevelUnit* 1 x y) , (isOfHLevelUnit* 2 x y _) +◯-=-isModal (HLevelTruncModality (suc n)) = isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) + +-- universal property + +univTrunc : ∀ {ℓ} (n : HLevel) {B : TypeOfHLevel ℓ n} → Iso (hLevelTrunc n A → B .fst) (A → B .fst) +univTrunc zero {B , lev} = isContr→Iso (isOfHLevelΠ 0 (λ _ → lev)) (isOfHLevelΠ 0 λ _ → lev) +Iso.fun (univTrunc (suc n) {B , lev}) g a = g ∣ a ∣ +Iso.inv (univTrunc (suc n) {B , lev}) = rec lev +Iso.rightInv (univTrunc (suc n) {B , lev}) b = refl +Iso.leftInv (univTrunc (suc n) {B , lev}) b = funExt (elim (λ x → isOfHLevelPath _ lev _ _) + λ a → refl) + +-- functorial action + +map : {n : HLevel} {B : Type ℓ'} (g : A → B) + → hLevelTrunc n A → hLevelTrunc n B +map g = rec (isOfHLevelTrunc _) (λ a → ∣ g a ∣ₕ) + +map2 : ∀ {ℓ''} {n : HLevel} {B : Type ℓ'} {C : Type ℓ''} (g : A → B → C) + → hLevelTrunc n A → hLevelTrunc n B → hLevelTrunc n C +map2 {n = zero} g = λ _ _ → tt* +map2 {n = suc n} g = rec (isOfHLevelΠ (suc n) λ _ → isOfHLevelTrunc _) + (λ a → rec (isOfHLevelTrunc _) λ b → ∣ g a b ∣) + +mapCompIso : {n : HLevel} {B : Type ℓ'} → (Iso A B) → Iso (hLevelTrunc n A) (hLevelTrunc n B) +mapCompIso {n = zero} {B} _ = isContr→Iso (isOfHLevelUnit* 0) (isOfHLevelUnit* 0) +Iso.fun (mapCompIso {n = (suc n)} g) = map (Iso.fun g) +Iso.inv (mapCompIso {n = (suc n)} g) = map (Iso.inv g) +Iso.rightInv (mapCompIso {n = (suc n)} g) = elim (λ x → isOfHLevelPath _ (isOfHLevelTrunc _) _ _) λ b → cong ∣_∣ (Iso.rightInv g b) +Iso.leftInv (mapCompIso {n = (suc n)} g) = elim (λ x → isOfHLevelPath _ (isOfHLevelTrunc _) _ _) λ a → cong ∣_∣ (Iso.leftInv g a) + +mapId : {n : HLevel} → ∀ t → map {n = n} (idfun A) t ≡ t +mapId {n = 0} tt* = refl +mapId {n = (suc n)} = + elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) (λ _ → refl) -- equivalences to prop/set/groupoid truncations -propTrunc≃Trunc-1 : ∥ A ∥₋₁ ≃ ∥ A ∥ -1 -propTrunc≃Trunc-1 = - isoToEquiv - (iso - (elimPropTrunc (λ _ → isOfHLevel∥∥ -1) ∣_∣) - (ind (λ _ → propTruncIsProp) ∣_∣) - (ind (λ _ → isOfHLevelPath 1 (isOfHLevel∥∥ -1) _ _) (λ _ → refl)) - (elimPropTrunc (λ _ → isOfHLevelPath 1 squash _ _) (λ _ → refl))) - -setTrunc≃Trunc0 : ∥ A ∥₀ ≃ ∥ A ∥ 0 -setTrunc≃Trunc0 = - isoToEquiv - (iso - (elimSetTrunc (λ _ → isOfHLevel∥∥ 0) ∣_∣) - (ind (λ _ → squash₀) ∣_∣₀) - (ind (λ _ → isOfHLevelPath 2 (isOfHLevel∥∥ 0) _ _) (λ _ → refl)) - (elimSetTrunc (λ _ → isOfHLevelPath 2 squash₀ _ _) (λ _ → refl))) - -groupoidTrunc≃Trunc1 : ∥ A ∥₁ ≃ ∥ A ∥ 1 -groupoidTrunc≃Trunc1 = - isoToEquiv - (iso - (groupoidTruncElim _ _ (λ _ → isOfHLevel∥∥ 1) ∣_∣) - (ind (λ _ → squash₁) ∣_∣₁) - (ind (λ _ → isOfHLevelPath 3 (isOfHLevel∥∥ 1) _ _) (λ _ → refl)) - (groupoidTruncElim _ _ (λ _ → isOfHLevelPath 3 squash₁ _ _) (λ _ → refl))) - -2groupoidTrunc≃Trunc2 : ∥ A ∥₂ ≃ ∥ A ∥ 2 -2groupoidTrunc≃Trunc2 = - isoToEquiv - (iso - (g2TruncElim _ _ (λ _ → isOfHLevel∥∥ 2) ∣_∣) - (ind (λ _ → squash₂) ∣_∣₂) - (ind (λ _ → isOfHLevelPath 4 (isOfHLevel∥∥ 2) _ _) (λ _ → refl)) - (g2TruncElim _ _ (λ _ → isOfHLevelPath 4 squash₂ _ _) (λ _ → refl))) +propTruncTrunc1Iso : Iso ∥ A ∥₁ (∥ A ∥ 1) +Iso.fun propTruncTrunc1Iso = PropTrunc.rec (isOfHLevelTrunc 1) ∣_∣ +Iso.inv propTruncTrunc1Iso = rec squash₁ ∣_∣₁ +Iso.rightInv propTruncTrunc1Iso = elim (λ _ → isOfHLevelPath 1 (isOfHLevelTrunc 1) _ _) (λ _ → refl) +Iso.leftInv propTruncTrunc1Iso = PropTrunc.elim (λ _ → isOfHLevelPath 1 squash₁ _ _) (λ _ → refl) + +propTrunc≃Trunc1 : ∥ A ∥₁ ≃ ∥ A ∥ 1 +propTrunc≃Trunc1 = isoToEquiv propTruncTrunc1Iso + +propTrunc≡Trunc1 : ∥ A ∥₁ ≡ ∥ A ∥ 1 +propTrunc≡Trunc1 = ua propTrunc≃Trunc1 + + +setTruncTrunc2Iso : Iso ∥ A ∥₂ (∥ A ∥ 2) +Iso.fun setTruncTrunc2Iso = SetTrunc.rec (isOfHLevelTrunc 2) ∣_∣ +Iso.inv setTruncTrunc2Iso = rec squash₂ ∣_∣₂ +Iso.rightInv setTruncTrunc2Iso = elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) (λ _ → refl) +Iso.leftInv setTruncTrunc2Iso = SetTrunc.elim (λ _ → isOfHLevelPath 2 squash₂ _ _) (λ _ → refl) + +setTrunc≃Trunc2 : ∥ A ∥₂ ≃ ∥ A ∥ 2 +setTrunc≃Trunc2 = isoToEquiv setTruncTrunc2Iso + +propTrunc≡Trunc2 : ∥ A ∥₂ ≡ ∥ A ∥ 2 +propTrunc≡Trunc2 = ua setTrunc≃Trunc2 + +groupoidTruncTrunc3Iso : Iso ∥ A ∥₃ (∥ A ∥ 3) +Iso.fun groupoidTruncTrunc3Iso = GpdTrunc.rec (isOfHLevelTrunc 3) ∣_∣ +Iso.inv groupoidTruncTrunc3Iso = rec squash₃ ∣_∣₃ +Iso.rightInv groupoidTruncTrunc3Iso = elim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) (λ _ → refl) +Iso.leftInv groupoidTruncTrunc3Iso = GpdTrunc.elim (λ _ → isOfHLevelPath 3 squash₃ _ _) (λ _ → refl) + +groupoidTrunc≃Trunc3 : ∥ A ∥₃ ≃ ∥ A ∥ 3 +groupoidTrunc≃Trunc3 = isoToEquiv groupoidTruncTrunc3Iso + +groupoidTrunc≡Trunc3 : ∥ A ∥₃ ≡ ∥ A ∥ 3 +groupoidTrunc≡Trunc3 = ua groupoidTrunc≃Trunc3 + +2GroupoidTruncTrunc4Iso : Iso ∥ A ∥₄ (∥ A ∥ 4) +Iso.fun 2GroupoidTruncTrunc4Iso = 2GpdTrunc.rec (isOfHLevelTrunc 4) ∣_∣ +Iso.inv 2GroupoidTruncTrunc4Iso = rec squash₄ ∣_∣₄ +Iso.rightInv 2GroupoidTruncTrunc4Iso = elim (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) (λ _ → refl) +Iso.leftInv 2GroupoidTruncTrunc4Iso = 2GpdTrunc.elim (λ _ → isOfHLevelPath 4 squash₄ _ _) (λ _ → refl) + +2GroupoidTrunc≃Trunc4 : ∥ A ∥₄ ≃ ∥ A ∥ 4 +2GroupoidTrunc≃Trunc4 = isoToEquiv 2GroupoidTruncTrunc4Iso + +2GroupoidTrunc≡Trunc4 : ∥ A ∥₄ ≡ ∥ A ∥ 4 +2GroupoidTrunc≡Trunc4 = ua 2GroupoidTrunc≃Trunc4 + +isContr→isContrTrunc : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A → isContr (hLevelTrunc n A) +isContr→isContrTrunc zero contr = isOfHLevelUnit* 0 +isContr→isContrTrunc (suc n) contr = + ∣ fst contr ∣ , (elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) λ a → cong ∣_∣ (snd contr a)) + +truncOfProdIso : (n : ℕ) → Iso (hLevelTrunc n (A × B)) (hLevelTrunc n A × hLevelTrunc n B) +truncOfProdIso 0 = isContr→Iso (isOfHLevelUnit* 0) (isOfHLevel× 0 (isOfHLevelUnit* 0) (isOfHLevelUnit* 0)) +Iso.fun (truncOfProdIso (suc n)) = rec (isOfHLevelΣ (suc n) (isOfHLevelTrunc (suc n)) (λ _ → isOfHLevelTrunc (suc n))) λ {(a , b) → ∣ a ∣ , ∣ b ∣} +Iso.inv (truncOfProdIso (suc n)) (a , b) = rec (isOfHLevelTrunc (suc n)) + (λ a → rec (isOfHLevelTrunc (suc n)) + (λ b → ∣ a , b ∣) + b) + a +Iso.rightInv (truncOfProdIso (suc n)) (a , b) = + elim {B = λ a → Iso.fun (truncOfProdIso (suc n)) (Iso.inv (truncOfProdIso (suc n)) (a , b)) ≡ (a , b)} + (λ _ → isOfHLevelPath (suc n) (isOfHLevelΣ (suc n) (isOfHLevelTrunc (suc n)) (λ _ → isOfHLevelTrunc (suc n))) _ _) + (λ a → elim {B = λ b → Iso.fun (truncOfProdIso (suc n)) (Iso.inv (truncOfProdIso (suc n)) (∣ a ∣ , b)) ≡ (∣ a ∣ , b)} + (λ _ → isOfHLevelPath (suc n) (isOfHLevelΣ (suc n) (isOfHLevelTrunc (suc n)) (λ _ → isOfHLevelTrunc (suc n))) _ _) + (λ b → refl) b) a +Iso.leftInv (truncOfProdIso (suc n)) = elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) λ a → refl ---- ∥ Ω A ∥ ₙ ≡ Ω ∥ A ∥ₙ₊₁ ---- + {- Proofs of Theorem 7.3.12. and Corollary 7.3.13. in the HoTT book -} -private - {- We define the fibration P to show a more general result -} - P : ∀ {ℓ} {B : Type ℓ}{n : ℕ₋₂} → ∥ B ∥ (suc n) → ∥ B ∥ (suc n) → Type ℓ - P x y = fst (P₁ x y) - - where - P₁ : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} → ∥ B ∥ (suc n) → ∥ B ∥ (suc n) → (HLevel ℓ (2+ n)) - P₁ {ℓ} {n = n} x y = ind2 (λ _ _ → isOfHLevelHLevel (2+ n)) - (λ a b → (∥ a ≡ b ∥ n , isOfHLevel∥∥ n )) - x - y - - {- We will need P to be of hLevel n + 3 -} - hLevelP : ∀{ℓ} {n : ℕ₋₂} {B : Type ℓ} - (a b : ∥ B ∥ (suc n)) → - isOfHLevel (2+ (suc n)) (P a b ) - hLevelP {n = n} {B = B} = ind2 {A = B} - {n = (suc n)} - {B = λ x y → isOfHLevel (2+ (suc n)) (P x y)} - (λ x y → isProp→isOfHLevelSuc (2+ n) (isPropIsOfHLevel (2+ suc n)) ) - λ a b → ( isOfHLevelSuc (2+ n) ) - (isOfHLevel∥∥ {A = a ≡ b} n) - - {- decode function from P x y to x ≡ y -} - decode-fun : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x y : ∥ B ∥ (suc n)) → - P x y → - _≡_ {A = ∥ B ∥ (suc n)} x y - - decode-fun {B = B} {n = n} x y = ind2 {B = λ u v → P u v → _≡_ {A = ∥ B ∥ (suc n)} u v } - (λ u v → isOfHLevelPi {A = P u v} {B = λ _ → u ≡ v} - (2+ suc n) - λ _ → (((isOfHLevelSuc (2+ suc n) (isOfHLevel∥∥ {A = B} (suc n))) u v)) ) - decode* x y - where - decode* : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂}(u v : B) → - (P {n = n} - (∣_∣ {S = (S (1+ (suc n)))} u) - (∣_∣ {S = (S (1+ (suc n)))} v)) → - _≡_ {A = ∥ B ∥ (suc n) } - ∣ u ∣ - ∣ v ∣ - decode* {B = B} {n = neg2} u v = rec {A = u ≡ v} {n = neg2} - {B = _≡_ {A = ∥ B ∥ (suc (neg2)) } (∣ u ∣) (∣ v ∣)} - ((isOfHLevel∥∥ {A = B} (suc neg2) ∣ u ∣ ∣ v ∣) , - λ y → (isOfHLevelSuc (2+ suc neg2) - (isOfHLevel∥∥ {A = B} (suc neg2)) ∣ u ∣ ∣ v ∣) - (isOfHLevel∥∥ {A = B} (suc neg2) ∣ u ∣ ∣ v ∣) y ) - (λ p → cong (λ z → ∣ z ∣) p) - decode* {B = B} {n = suc n} u v = rec {A = u ≡ v} {n = suc n} - {B = _≡_ {A = ∥ B ∥ (suc (suc n)) } (∣ u ∣) (∣ v ∣)} - (isOfHLevel∥∥ {A = B} (suc (suc n)) ∣ u ∣ ∣ v ∣) - (λ p → cong (λ z → ∣ z ∣) p) - - {- auxilliary function r used to define encode -} - r : ∀ {ℓ} {B : Type ℓ} {m : ℕ₋₂} (u : ∥ B ∥ (suc m)) → P u u - r {m = m} = ind {B = (λ u → P u u)} - (λ x → hLevelP x x) - (λ a → ∣ refl {x = a} ∣) - - {- encode function from x ≡ y to P x y -} - encode-fun : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x y : ∥ B ∥ (suc n)) → - _≡_ {A = ∥ B ∥ (suc n)} x y → - P x y - encode-fun x y p = transport (λ i → P x (p i )) (r x) - - {- We need the following two lemmas on the functions behaviour for refl -} - dec-refl : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x : ∥ B ∥ (suc n)) → - decode-fun x x (r x) ≡ refl {x = x} - dec-refl {B = B} {n = neg2} x = ind {A = B} {n = suc neg2} - {B = λ x → decode-fun x x (r x) ≡ refl {x = x} } - (λ x → (isOfHLevelSuc (2+ (suc neg2)) - (isOfHLevelSuc (2+ (suc neg2)) - (isOfHLevel∥∥ {A = B} (suc neg2)) x x)) - (decode-fun x x (r x)) refl) - (λ a → refl) x - dec-refl {B = B} {n = suc n} = ind {A = B} {n = suc (suc n)} - {B = λ x → decode-fun x x (r x) ≡ refl {x = x} } - (λ x → isOfHLevelSuc (2+ suc n) - (isOfHLevelSuc (2+ suc n) - (isOfHLevel∥∥ {A = B} (suc (suc n)) x x ) - (decode-fun x x (r x)) refl)) - λ c → refl - - enc-refl : ∀ {ℓ} {B : Type ℓ} - {n : ℕ₋₂} - (x : ∥ B ∥ (suc n)) → - encode-fun x x (refl {x = x}) ≡ r x - enc-refl x j = transp (λ i → P x (refl {x = x} i)) j (r x) - - {- decode-fun is a right-inverse -} - P-rinv : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (u v : ∥ B ∥ (suc n)) → - (x : _≡_ {A = ∥ B ∥ (suc n)} u v) → - decode-fun u v (encode-fun u v x) ≡ x - P-rinv {ℓ = ℓ} {B = B} {n = n} u v = J {ℓ} { ∥ B ∥ (suc n)} {u} {ℓ} - (λ y p → decode-fun u y (encode-fun u y p) ≡ p) - ((λ i → (decode-fun u u (enc-refl u i))) ∙ dec-refl u) - {v} - - {- decode-fun is a left-inverse -} - P-linv : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (u v : ∥ B ∥ (suc n )) → - (x : P u v) → - encode-fun u v (decode-fun u v x) ≡ x - P-linv {ℓ = ℓ} {B = B} {n = n} u v = ind2 {A = B} - {n = (suc n)} - {B = λ u v → (x : P u v) → encode-fun u v (decode-fun u v x) ≡ x} - (λ x y → isOfHLevelPi {A = P x y} - (2+ suc n) - λ z → isOfHLevelSuc (2+ suc n) - (hLevelP {n = n} x y) (encode-fun x y (decode-fun x y z)) z) - helper u v - where - helper : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} - (a b : B) - (x : P {n = n} ∣ a ∣ ∣ b ∣) → - encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x) ≡ x - helper {ℓ = ℓ} {B = B} {n = neg2} a b = ind {A = (a ≡ b)} - {n = neg2} - {B = λ x → encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x) ≡ x} - (λ x → (sym ((snd (isOfHLevel∥∥ {A = a ≡ b} neg2)) - (encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x))) - ∙ ((snd (isOfHLevel∥∥ {A = a ≡ b} neg2)) x)) - , - λ y → isOfHLevelSuc (2+ (suc neg2)) - (isOfHLevelSuc (2+ neg2) (isOfHLevel∥∥ {A = a ≡ b} (neg2))) - (encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x)) x - ((sym ((snd (isOfHLevel∥∥ {A = a ≡ b} neg2)) - (encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x))) - ∙ ((snd (isOfHLevel∥∥ {A = a ≡ b} neg2)) x))) - y) - (J {ℓ}{B}{a}{ℓ} - ((λ y p → (encode-fun {n = neg2}) ∣ a ∣ ∣ y ∣ - ((decode-fun ∣ a ∣ ∣ y ∣) ∣ p ∣) ≡ ∣ p ∣)) - (enc-refl {n = neg2} ∣ a ∣ ) - {b}) - helper {ℓ = ℓ} {B = B} {n = suc n} a b = ind {A = (a ≡ b)} - {n = suc n} - {B = λ x → encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x) ≡ x} - (λ x → (hLevelP {n = suc n} ∣ a ∣ ∣ b ∣ (encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x)) x) ) - (J {ℓ}{B}{a}{ℓ} ((λ y p → (encode-fun {n = suc n}) ∣ a ∣ ∣ y ∣ ((decode-fun ∣ a ∣ ∣ y ∣) ∣ p ∣) ≡ ∣ p ∣)) - (enc-refl {n = suc n} ∣ a ∣ ) - {b}) - - {- The final Iso established -} - IsoFinal : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x y : ∥ B ∥ (suc n)) → Iso (x ≡ y) (P x y) - IsoFinal x y = iso (encode-fun x y ) (decode-fun x y) (P-linv x y) (P-rinv x y) - -PathIdTrunc : {a b : A} (n : ℕ₋₂) → (_≡_ {A = ∥ A ∥ (suc n)} ∣ a ∣ ∣ b ∣) ≡ (∥ a ≡ b ∥ n) -PathIdTrunc {a = a} {b = b} n = isoToPath (IsoFinal {n = n} ∣ a ∣ ∣ b ∣) - -PathΩ : {a : A} (n : ℕ₋₂) → (_≡_ {A = ∥ A ∥ (suc n)} ∣ a ∣ ∣ a ∣) ≡ (∥ a ≡ a ∥ n) -PathΩ {a = a} n = PathIdTrunc {a = a} {b = a} n +module ΩTrunc {X : Type ℓ} {n : HLevel} where + {- We define the fibration P to show a more general result -} + Code : ∥ X ∥ (2 + n) → ∥ X ∥ (2 + n) → TypeOfHLevel ℓ (suc n) + Code x y = + rec2 (isOfHLevelTypeOfHLevel (suc n)) + (λ a b → ∥ a ≡ b ∥ (suc n) , isOfHLevelTrunc (suc n)) x y + + P : ∥ X ∥ (2 + n) → ∥ X ∥ (2 + n) → Type ℓ + P x y = Code x y .fst + + {- We will need P to be of hLevel n + 3 -} + hLevelP : (a b : ∥ X ∥ (2 + n)) → isOfHLevel (2 + n) (P a b) + hLevelP x y = isOfHLevelSuc (suc n) (Code x y .snd) + + {- decode function from P x y to x ≡ y -} + decode-fun : (x y : ∥ X ∥ (2 + n)) → P x y → x ≡ y + decode-fun = + elim2 (λ u v → isOfHLevelΠ (2 + n)(λ _ → isOfHLevelSuc (2 + n) (isOfHLevelTrunc (2 + n)) u v)) + decode* + where + decode* : (u v : X) → P ∣ u ∣ ∣ v ∣ → Path (∥ X ∥ (2 + n)) ∣ u ∣ ∣ v ∣ + decode* u v = + rec (isOfHLevelTrunc (2 + n) ∣ u ∣ ∣ v ∣) (cong ∣_∣) + + {- auxiliary function r used to define encode -} + r : (u : ∥ X ∥ (2 + n)) → P u u + r = elim (λ x → hLevelP x x) (λ a → ∣ refl ∣) + + encode-fun : (x y : ∥ X ∥ (2 + n)) → x ≡ y → P x y + encode-fun x y p = subst (P x) p (r x) + + encode-fill : (x y : ∥ X ∥ (2 + n)) (p : x ≡ y) + → PathP (λ i → P x (p i)) (r x) (encode-fun x y p) + encode-fill x y p = subst-filler (P x) p (r x) + + {- We need the following lemma on the functions behaviour for refl -} + dec-refl : (x : ∥ X ∥ (2 + n)) → decode-fun x x (r x) ≡ refl + dec-refl = + elim (λ x → isOfHLevelSuc (1 + n) + (isOfHLevelSuc (1 + n) + (isOfHLevelTrunc (2 + n) x x) + (decode-fun x x (r x)) refl)) + (λ _ → refl) + + {- decode-fun is a right-inverse -} + P-rinv : (u v : ∥ X ∥ (2 + n)) (x : Path (∥ X ∥ (2 + n)) u v) + → decode-fun u v (encode-fun u v x) ≡ x + P-rinv u v p = + transport + (λ i → decode-fun u (p i) (encode-fill u v p i) ≡ (λ j → p (i ∧ j))) + (dec-refl u) + + {- decode-fun is a left-inverse -} + P-linv : (u v : ∥ X ∥ (2 + n)) (x : P u v) + → encode-fun u v (decode-fun u v x) ≡ x + P-linv = + elim2 (λ x y → isOfHLevelΠ (2 + n) (λ z → isOfHLevelSuc (2 + n) (hLevelP x y) _ _)) + (λ a b → elim (λ p → hLevelP ∣ a ∣ ∣ b ∣ _ _) (helper a b)) + where + helper : (a b : X) (p : a ≡ b) → encode-fun _ _ (decode-fun ∣ a ∣ ∣ b ∣ (∣ p ∣)) ≡ ∣ p ∣ + helper a b p = + transport + (λ i → subst-filler (P ∣ a ∣) (cong ∣_∣ p) ∣ refl ∣ i ≡ ∣ (λ j → p (i ∧ j)) ∣) + refl + + {- The final Iso established -} + IsoFinal : (x y : ∥ X ∥ (2 + n)) → Iso (x ≡ y) (P x y) + Iso.fun (IsoFinal x y) = encode-fun x y + Iso.inv (IsoFinal x y) = decode-fun x y + Iso.rightInv (IsoFinal x y) = P-linv x y + Iso.leftInv (IsoFinal x y) = P-rinv x y + + +P : (x y z : ∥ X ∥ (2 + n)) → (P x y) → (P y z) → P x z + +P = + elim3 (λ x _ z → isOfHLevelΠ (2 + n) λ _ → isOfHLevelΠ (2 + n) λ _ → hLevelP x z) λ a b c → + rec (isOfHLevelΠ (suc n) λ _ → isOfHLevelTrunc (suc n)) λ p → + rec (isOfHLevelTrunc (suc n)) λ q → + ∣ p ∙ q ∣ + + +P-funct : (x y z : ∥ X ∥ (2 + n)) (p : x ≡ y) (q : y ≡ z) + → +P x y z (encode-fun x y p) (encode-fun y z q) ≡ encode-fun x z (p ∙ q) + +P-funct x y z p q = + transport + (λ i → +P x y (q i) (encode-fun x y p) (encode-fill y z q i) ≡ encode-fun x (q i) (compPath-filler p q i)) + (transport + (λ i → +P x (p i) (p i) (encode-fill x y p i) (r (p i)) ≡ encode-fill x y p i) + (elim {B = λ x → +P x x x (r x) (r x) ≡ r x} + (λ x → isOfHLevelPath (2 + n) (hLevelP x x) _ _) + (λ a i → ∣ rUnit refl (~ i) ∣) + x)) + +PathIdTruncIso : {a b : A} (n : HLevel) → Iso (Path (∥ A ∥ (suc n)) ∣ a ∣ ∣ b ∣) (∥ a ≡ b ∥ n) +PathIdTruncIso zero = + isContr→Iso + (isOfHLevelTrunc 1 _ _ , isOfHLevelPath 1 (isOfHLevelTrunc 1) ∣ _ ∣ ∣ _ ∣ _) + (isOfHLevelUnit* 0) +PathIdTruncIso (suc n) = ΩTrunc.IsoFinal ∣ _ ∣ ∣ _ ∣ + +PathIdTrunc : {a b : A} (n : HLevel) → (Path (∥ A ∥ (suc n)) ∣ a ∣ ∣ b ∣) ≡ (∥ a ≡ b ∥ n) +PathIdTrunc n = isoToPath (PathIdTruncIso n) + +PathΩ : {a : A} (n : HLevel) → (Path (∥ A ∥ (suc n)) ∣ a ∣ ∣ a ∣) ≡ (∥ a ≡ a ∥ n) +PathΩ n = PathIdTrunc n + +PathIdTruncIsoFunct : ∀ {A : Type ℓ} {a : A} (n : HLevel) → (p q : (Path (∥ A ∥ (2 + n)) ∣ a ∣ ∣ a ∣)) + → Iso.fun (PathIdTruncIso (suc n)) (p ∙ q) + ≡ map2 _∙_ (Iso.fun (PathIdTruncIso (suc n)) p) (Iso.fun (PathIdTruncIso (suc n)) q) +PathIdTruncIsoFunct {a = a} n p q = sym (ΩTrunc.+P-funct (∣ a ∣) ∣ a ∣ ∣ a ∣ p q) + +------------------------- + +truncOfTruncIso : (n m : HLevel) → Iso (hLevelTrunc n A) (hLevelTrunc n (hLevelTrunc (m + n) A)) +truncOfTruncIso zero m = isContr→Iso (isOfHLevelUnit* 0) (isOfHLevelUnit* 0) +Iso.fun (truncOfTruncIso (suc n) zero) = rec (isOfHLevelTrunc (suc n)) λ a → ∣ ∣ a ∣ ∣ +Iso.fun (truncOfTruncIso (suc n) (suc m)) = rec (isOfHLevelTrunc (suc n)) λ a → ∣ ∣ a ∣ ∣ +Iso.inv (truncOfTruncIso (suc n) zero) = rec (isOfHLevelTrunc (suc n)) + (rec (isOfHLevelTrunc (suc n)) + λ a → ∣ a ∣) +Iso.inv (truncOfTruncIso (suc n) (suc m)) = rec (isOfHLevelTrunc (suc n)) + (rec (isOfHLevelPlus (suc m) (isOfHLevelTrunc (suc n))) + λ a → ∣ a ∣) +Iso.rightInv (truncOfTruncIso (suc n) zero) = + elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _ ) + (elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _ ) + λ a → refl) +Iso.rightInv (truncOfTruncIso (suc n) (suc m)) = + elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _ ) + (elim (λ x → isOfHLevelPath ((suc m) + (suc n)) (isOfHLevelPlus (suc m) (isOfHLevelTrunc (suc n))) _ _ ) + λ a → refl) +Iso.leftInv (truncOfTruncIso (suc n) zero) = + elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + λ a → refl +Iso.leftInv (truncOfTruncIso (suc n) (suc m)) = + elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + λ a → refl + +truncOfTruncIso' : (n m : HLevel) → Iso (hLevelTrunc n A) (hLevelTrunc n (hLevelTrunc (n + m) A)) +truncOfTruncIso' zero m = isContr→Iso (isOfHLevelUnit* 0) (isOfHLevelUnit* 0) +Iso.fun (truncOfTruncIso' (suc n) m) = rec (isOfHLevelTrunc (suc n)) λ a → ∣ ∣ a ∣ ∣ +Iso.inv (truncOfTruncIso' {A = A} (suc n) m) = + rec (isOfHLevelTrunc (suc n)) + (rec (isOfHLevelPlus' {n = m} (suc n) (isOfHLevelTrunc (suc n))) ∣_∣) +Iso.rightInv (truncOfTruncIso' (suc n) m) = + elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + (elim (λ _ → isOfHLevelPath (suc n + m) (isOfHLevelPlus' {n = m} (suc n) (isOfHLevelTrunc (suc n))) _ _) + λ _ → refl) +Iso.leftInv (truncOfTruncIso' (suc n) m) = + elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + λ _ → refl + +truncOfTruncEq : (n m : ℕ) → (hLevelTrunc n A) ≃ (hLevelTrunc n (hLevelTrunc (m + n) A)) +truncOfTruncEq n m = isoToEquiv (truncOfTruncIso n m) + +truncOfTruncIso2 : (n m : HLevel) → Iso (hLevelTrunc (m + n) (hLevelTrunc n A)) (hLevelTrunc n A) +truncOfTruncIso2 n m = truncIdempotentIso (m + n) (isOfHLevelPlus m (isOfHLevelTrunc n)) + +truncOfΣIso : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} {B : A → Type ℓ'} + → Iso (hLevelTrunc n (Σ A B)) (hLevelTrunc n (Σ A λ x → hLevelTrunc n (B x))) +truncOfΣIso zero = idIso +Iso.fun (truncOfΣIso (suc n)) = map λ {(a , b) → a , ∣ b ∣} +Iso.inv (truncOfΣIso (suc n)) = + rec (isOfHLevelTrunc (suc n)) + (uncurry λ a → rec (isOfHLevelTrunc (suc n)) λ b → ∣ a , b ∣) +Iso.rightInv (truncOfΣIso (suc n)) = + elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + (uncurry λ a → elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + λ b → refl) +Iso.leftInv (truncOfΣIso (suc n)) = + elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) λ {(a , b) → refl} diff --git a/Cubical/HITs/TypeQuotients.agda b/Cubical/HITs/TypeQuotients.agda new file mode 100644 index 000000000..0059b89aa --- /dev/null +++ b/Cubical/HITs/TypeQuotients.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.TypeQuotients where + +open import Cubical.HITs.TypeQuotients.Base public +open import Cubical.HITs.TypeQuotients.Properties public diff --git a/Cubical/HITs/TypeQuotients/Base.agda b/Cubical/HITs/TypeQuotients/Base.agda new file mode 100644 index 000000000..7a8cfb216 --- /dev/null +++ b/Cubical/HITs/TypeQuotients/Base.agda @@ -0,0 +1,17 @@ +{- + +This file contains: + +- Definition of type quotients (i.e. non-truncated quotients) + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.TypeQuotients.Base where + +open import Cubical.Core.Primitives + +-- Type quotients as a higher inductive type: +data _/ₜ_ {ℓ ℓ'} (A : Type ℓ) (R : A → A → Type ℓ') : Type (ℓ-max ℓ ℓ') where + [_] : (a : A) → A /ₜ R + eq/ : (a b : A) → (r : R a b) → [ a ] ≡ [ b ] + diff --git a/Cubical/HITs/TypeQuotients/Properties.agda b/Cubical/HITs/TypeQuotients/Properties.agda new file mode 100644 index 000000000..161670061 --- /dev/null +++ b/Cubical/HITs/TypeQuotients/Properties.agda @@ -0,0 +1,60 @@ +{- + +Type quotients: + +-} + +{-# OPTIONS --safe #-} +module Cubical.HITs.TypeQuotients.Properties where + +open import Cubical.HITs.TypeQuotients.Base + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels + +open import Cubical.HITs.PropositionalTruncation as PropTrunc using (∥_∥ ; ∣_∣ ; squash) +open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂ ; ∣_∣₂ ; squash₂ + ; isSetSetTrunc) + +private + variable + ℓ ℓ' ℓ'' : Level + A : Type ℓ + R : A → A → Type ℓ' + B : A /ₜ R → Type ℓ'' + C : A /ₜ R → A /ₜ R → Type ℓ'' + + +elim : (f : (a : A) → (B [ a ])) + → ((a b : A) (r : R a b) → PathP (λ i → B (eq/ a b r i)) (f a) (f b)) + ------------------------------------------------------------------------ + → (x : A /ₜ R) → B x +elim f feq [ a ] = f a +elim f feq (eq/ a b r i) = feq a b r i + +rec : {X : Type ℓ''} + → (f : A → X) + → (∀ (a b : A) → R a b → f a ≡ f b) + ------------------------------------- + → A /ₜ R → X +rec f feq [ a ] = f a +rec f feq (eq/ a b r i) = feq a b r i + +elimProp : ((x : A /ₜ R ) → isProp (B x)) + → ((a : A) → B ( [ a ])) + --------------------------------- + → (x : A /ₜ R) → B x +elimProp Bprop f [ a ] = f a +elimProp Bprop f (eq/ a b r i) = isOfHLevel→isOfHLevelDep 1 Bprop (f a) (f b) (eq/ a b r) i + +elimProp2 : ((x y : A /ₜ R ) → isProp (C x y)) + → ((a b : A) → C [ a ] [ b ]) + -------------------------------------- + → (x y : A /ₜ R) → C x y +elimProp2 Cprop f = elimProp (λ x → isPropΠ (λ y → Cprop x y)) + (λ x → elimProp (λ y → Cprop [ x ] y) (f x)) diff --git a/Cubical/HITs/Wedge.agda b/Cubical/HITs/Wedge.agda new file mode 100644 index 000000000..1cf2e9747 --- /dev/null +++ b/Cubical/HITs/Wedge.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Wedge where + +open import Cubical.HITs.Wedge.Base public diff --git a/Cubical/HITs/Wedge/Base.agda b/Cubical/HITs/Wedge/Base.agda new file mode 100644 index 000000000..7fb25a501 --- /dev/null +++ b/Cubical/HITs/Wedge/Base.agda @@ -0,0 +1,26 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.Wedge.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.HITs.Pushout.Base +open import Cubical.Data.Unit + +_⋁_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') +_⋁_ (A , ptA) (B , ptB) = Pushout {A = Unit} {B = A} {C = B} (λ _ → ptA) (λ _ → ptB) + + +-- pointed versions + +_⋁∙ₗ_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ') +A ⋁∙ₗ B = (A ⋁ B) , (inl (snd A)) + +_⋁∙ᵣ_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ') +A ⋁∙ᵣ B = (A ⋁ B) , (inr (snd B)) + +-- Wedge sum of Units is contractible +isContr-Unit⋁Unit : isContr ((Unit , tt) ⋁ (Unit , tt)) +fst isContr-Unit⋁Unit = inl tt +snd isContr-Unit⋁Unit (inl tt) = refl +snd isContr-Unit⋁Unit (inr tt) = push tt +snd isContr-Unit⋁Unit (push tt i) j = push tt (i ∧ j) diff --git a/Cubical/Homotopy/Base.agda b/Cubical/Homotopy/Base.agda new file mode 100644 index 000000000..1faa2f9ba --- /dev/null +++ b/Cubical/Homotopy/Base.agda @@ -0,0 +1,19 @@ +{-# OPTIONS --safe #-} + +module Cubical.Homotopy.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv.Properties + +private + variable + ℓ ℓ' : Level + +_∼_ : {X : Type ℓ} {Y : X → Type ℓ'} → (f g : (x : X) → Y x) → Type (ℓ-max ℓ ℓ') +_∼_ {X = X} f g = (x : X) → f x ≡ g x + +funExt∼ : {X : Type ℓ} {Y : X → Type ℓ'} {f g : (x : X) → Y x} (H : f ∼ g) → f ≡ g +funExt∼ = funExt + +∼-refl : {X : Type ℓ} {Y : X → Type ℓ'} {f : (x : X) → Y x} → f ∼ f +∼-refl {f = f} = λ x → refl {x = f x} diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda new file mode 100644 index 000000000..a53fd226d --- /dev/null +++ b/Cubical/Homotopy/Connected.agda @@ -0,0 +1,306 @@ +{-# OPTIONS --safe #-} +module Cubical.Homotopy.Connected where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path +open import Cubical.Foundations.Univalence +open import Cubical.Functions.Fibration +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.HITs.Nullification +open import Cubical.HITs.Susp +open import Cubical.HITs.SmashProduct +open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec) +open import Cubical.Homotopy.Loopspace +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn.Base +open import Cubical.HITs.S1 +open import Cubical.Data.Bool +open import Cubical.Data.Unit + +-- Note that relative to most sources, this notation is off by +2 +isConnected : ∀ {ℓ} (n : HLevel) (A : Type ℓ) → Type ℓ +isConnected n A = isContr (hLevelTrunc n A) + +isConnectedFun : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ-max ℓ ℓ') +isConnectedFun n f = ∀ b → isConnected n (fiber f b) + +isTruncatedFun : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ-max ℓ ℓ') +isTruncatedFun n f = ∀ b → isOfHLevel n (fiber f b) + +isConnectedSubtr : ∀ {ℓ} {A : Type ℓ} (n m : HLevel) + → isConnected (m + n) A + → isConnected n A +isConnectedSubtr {A = A} n m iscon = + isOfHLevelRetractFromIso 0 (truncOfTruncIso n m) (helper n iscon) + where + helper : (n : ℕ) → isConnected (m + n) A → isContr (hLevelTrunc n (hLevelTrunc (m + n) A)) + helper zero iscon = isContrUnit* + helper (suc n) iscon = ∣ iscon .fst ∣ , (Trunc.elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) λ a → cong ∣_∣ (iscon .snd a)) + + +isConnectedFunSubtr : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : HLevel) (f : A → B) + → isConnectedFun (m + n) f + → isConnectedFun n f +isConnectedFunSubtr n m f iscon b = isConnectedSubtr n m (iscon b) + +private + typeToFiberIso : ∀ {ℓ} (A : Type ℓ) → Iso A (fiber (λ (x : A) → tt) tt) + Iso.fun (typeToFiberIso A) x = x , refl + Iso.inv (typeToFiberIso A) = fst + Iso.rightInv (typeToFiberIso A) b i = fst b , (isOfHLevelSuc 1 (isPropUnit) tt tt (snd b) refl) i + Iso.leftInv (typeToFiberIso A) a = refl + + typeToFiber : ∀ {ℓ} (A : Type ℓ) → A ≡ fiber (λ (x : A) → tt) tt + typeToFiber A = isoToPath (typeToFiberIso A) + + + +module elim {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} (f : A → B) where + private + inv : ∀ {ℓ'''} (n : HLevel) (P : B → TypeOfHLevel ℓ''' (suc n)) + → ((a : A) → P (f a) .fst) + → (b : B) + → hLevelTrunc (suc n) (fiber f b) → P b .fst + inv n P t b = + Trunc.rec + (P b .snd) + (λ {(a , p) → subst (fst ∘ P) p (t a)}) + + isIsoPrecompose : ∀ {ℓ'''} (n : ℕ) (P : B → TypeOfHLevel ℓ''' n) + → isConnectedFun n f + → Iso ((b : B) → P b .fst) ((a : A) → P (f a) .fst) + isIsoPrecompose zero P fConn = isContr→Iso (isOfHLevelΠ _ (λ b → P b .snd)) (isOfHLevelΠ _ λ a → P (f a) .snd) + Iso.fun (isIsoPrecompose (suc n) P fConn) = _∘ f + Iso.inv (isIsoPrecompose (suc n) P fConn) t b = inv n P t b (fConn b .fst) + Iso.rightInv (isIsoPrecompose (suc n) P fConn) t = + funExt λ a → cong (inv n P t (f a)) (fConn (f a) .snd ∣ a , refl ∣) + ∙ substRefl {B = fst ∘ P} (t a) + Iso.leftInv (isIsoPrecompose (suc n) P fConn) s = + funExt λ b → + Trunc.elim + {B = λ d → inv n P (s ∘ f) b d ≡ s b} + (λ _ → isOfHLevelPath (suc n) (P b .snd) _ _) + (λ {(a , p) i → transp (λ j → P (p (j ∨ i)) .fst) i (s (p i))}) + (fConn b .fst) + + isEquivPrecompose : ∀ {ℓ'''} (n : ℕ) (P : B → TypeOfHLevel ℓ''' n) + → isConnectedFun n f + → isEquiv (λ(s : (b : B) → P b .fst) → s ∘ f) + isEquivPrecompose zero P fConn = isoToIsEquiv theIso + where + theIso : Iso ((b : B) → P b .fst) ((a : A) → P (f a) .fst) + Iso.fun theIso = λ(s : (b : B) → P b .fst) → s ∘ f + Iso.inv theIso = λ _ b → P b .snd .fst + Iso.rightInv theIso g = funExt λ x → P (f x) .snd .snd (g x) + Iso.leftInv theIso g = funExt λ x → P x .snd .snd (g x) + isEquivPrecompose (suc n) P fConn = isoToIsEquiv (isIsoPrecompose (suc n) P fConn) + + isConnectedPrecompose : (n : ℕ) → ((P : B → TypeOfHLevel (ℓ-max ℓ ℓ') n) + → hasSection (λ(s : (b : B) → P b .fst) → s ∘ f)) + → isConnectedFun n f + isConnectedPrecompose zero P→sect b = isContrUnit* + isConnectedPrecompose (suc n) P→sect b = c n P→sect b , λ y → sym (fun n P→sect b y) + where + P : (n : HLevel) → ((P : B → TypeOfHLevel ℓ (suc n)) + → hasSection (λ(s : (b : B) → P b .fst) → s ∘ f)) + → B → Type _ + P n s b = hLevelTrunc (suc n) (fiber f b) + + c : (n : HLevel) → ((P : B → TypeOfHLevel (ℓ-max ℓ ℓ') (suc n)) + → hasSection (λ(s : (b : B) → P b .fst) → s ∘ f)) → (b : B) + → hLevelTrunc (suc n) (fiber f b) + c n s = (s λ b → (hLevelTrunc (suc n) (fiber f b) , isOfHLevelTrunc _)) .fst + λ a → ∣ a , refl ∣ + + fun : (n : HLevel) (P→sect : ((P : B → TypeOfHLevel (ℓ-max ℓ ℓ') (suc n)) + → hasSection λ(s : (b : B) → P b .fst) → s ∘ f)) + → (b : B) (w : (hLevelTrunc (suc n) (fiber f b))) + → w ≡ c n P→sect b + fun n P→sect b = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) + λ a → J (λ b p → ∣ (fst a) , p ∣ ≡ c n P→sect b) + (c* (fst a)) + (snd a) + where + c* : ((a : A) → ∣ (a , refl {x = f a}) ∣ ≡ c n P→sect (f a)) + c* a = sym (cong (λ x → x a) (P→sect (λ b → hLevelTrunc (suc n) (fiber f b) , isOfHLevelTrunc _) .snd λ a → ∣ a , refl ∣)) + +isOfHLevelPrecomposeConnected : ∀ {ℓ ℓ' ℓ''} (k : HLevel) (n : HLevel) + {A : Type ℓ} {B : Type ℓ'} (P : B → TypeOfHLevel ℓ'' (k + n)) (f : A → B) + → isConnectedFun n f + → isOfHLevelFun k (λ(s : (b : B) → P b .fst) → s ∘ f) +isOfHLevelPrecomposeConnected zero n P f fConn = + elim.isEquivPrecompose f n P fConn .equiv-proof +isOfHLevelPrecomposeConnected (suc k) n P f fConn t = + isOfHLevelPath'⁻ k + λ {(s₀ , p₀) (s₁ , p₁) → + isOfHLevelRetractFromIso k (invIso ΣPathIsoPathΣ) + (subst (isOfHLevel k) + (sym (fiberPath (s₀ , p₀) (s₁ , p₁))) + (isOfHLevelRetract k + (λ {(q , α) → (funExt⁻ q) , (cong funExt⁻ α)}) + (λ {(h , β) → (funExt h) , (cong funExt β)}) + (λ _ → refl) + (isOfHLevelPrecomposeConnected k n + (λ b → (s₀ b ≡ s₁ b) , isOfHLevelPath' (k + n) (P b .snd) _ _) + f fConn + (funExt⁻ (p₀ ∙∙ refl ∙∙ sym p₁)))))} + +indMapEquiv→conType : ∀ {ℓ} {A : Type ℓ} (n : HLevel) + → ((B : TypeOfHLevel ℓ n) + → isEquiv (λ (b : (fst B)) → λ (a : A) → b)) + → isConnected n A +indMapEquiv→conType {A = A} zero BEq = isContrUnit* +indMapEquiv→conType {A = A} (suc n) BEq = + isOfHLevelRetractFromIso 0 (mapCompIso {n = (suc n)} (typeToFiberIso A)) + (elim.isConnectedPrecompose (λ _ → tt) (suc n) + (λ P → ((λ a _ → a) ∘ invIsEq (BEq (P tt))) + , λ a → equiv-proof (BEq (P tt)) a .fst .snd) + tt) + +isConnectedPath : ∀ {ℓ} (n : HLevel) {A : Type ℓ} + → isConnected (suc n) A + → (a₀ a₁ : A) → isConnected n (a₀ ≡ a₁) +isConnectedPath zero connA a₀ a₁ = isContrUnit* +isConnectedPath (suc n) {A = A} connA a₀ a₁ = + isOfHLevelRetractFromIso 0 (invIso (PathIdTruncIso (suc n))) (isContr→isContrPath connA _ _) + +isConnectedPathP : ∀ {ℓ} (n : HLevel) {A : I → Type ℓ} + → isConnected (suc n) (A i1) + → (a₀ : A i0) (a₁ : A i1) → isConnected n (PathP A a₀ a₁) +isConnectedPathP n con a₀ a₁ = + subst (isConnected n) (sym (PathP≡Path _ _ _)) + (isConnectedPath n con _ _) + +isConnectedRetract : ∀ {ℓ ℓ'} (n : HLevel) + {A : Type ℓ} {B : Type ℓ'} + (f : A → B) (g : B → A) + (h : (x : A) → g (f x) ≡ x) + → isConnected n B → isConnected n A +isConnectedRetract zero _ _ _ _ = isContrUnit* +isConnectedRetract (suc n) f g h = + isContrRetract + (Trunc.map f) + (Trunc.map g) + (Trunc.elim + (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + (λ a → cong ∣_∣ (h a))) + +isConnectedRetractFromIso : ∀ {ℓ ℓ'} (n : HLevel) + {A : Type ℓ} {B : Type ℓ'} + → Iso A B + → isConnected n B → isConnected n A +isConnectedRetractFromIso n e = + isConnectedRetract n + (Iso.fun e) + (Iso.inv e) + (Iso.leftInv e) + +isConnectedPoint : ∀ {ℓ} (n : HLevel) {A : Type ℓ} + → isConnected (suc n) A + → (a : A) → isConnectedFun n (λ(_ : Unit) → a) +isConnectedPoint n connA a₀ a = + isConnectedRetract n + snd (_ ,_) (λ _ → refl) + (isConnectedPath n connA a₀ a) + +isConnectedPoint2 : ∀ {ℓ} (n : HLevel) {A : Type ℓ} (a : A) + → isConnectedFun n (λ(_ : Unit) → a) + → isConnected (suc n) A +isConnectedPoint2 n {A = A} a connMap = indMapEquiv→conType _ λ B → isoToIsEquiv (theIso B) + where + module _ {ℓ' : Level} (B : TypeOfHLevel ℓ' (suc n)) + where + helper : (f : A → fst B) → (a2 : A) → f a2 ≡ f a + helper f = equiv-proof (elim.isEquivPrecompose (λ _ → a) n (λ a2 → (f a2 ≡ f a) , + isOfHLevelPath' n (snd B) (f a2) (f a)) connMap) (λ _ → refl) .fst .fst + + theIso : Iso (fst B) (A → fst B) + Iso.fun theIso b a = b + Iso.inv theIso f = f a + Iso.rightInv theIso f = funExt λ y → sym (helper f y) + Iso.leftInv theIso _ = refl + +connectedTruncIso : ∀ {ℓ} {A B : Type ℓ} (n : HLevel) (f : A → B) + → isConnectedFun n f + → Iso (hLevelTrunc n A) (hLevelTrunc n B) +connectedTruncIso {A = A} {B = B} zero f con = isContr→Iso isContrUnit* isContrUnit* +connectedTruncIso {A = A} {B = B} (suc n) f con = g + where + back : B → hLevelTrunc (suc n) A + back y = map fst ((con y) .fst) + + backSection : (b : B) → Path (hLevelTrunc (suc n) B) + (Trunc.rec (isOfHLevelTrunc (suc n)) + (λ a → ∣ f a ∣) + (Trunc.rec (isOfHLevelTrunc (suc n)) + back ∣ b ∣)) + ∣ b ∣ + backSection b = helper (λ p → map f p ≡ ∣ b ∣) + (λ x → (isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n))) + (map f (map fst x)) ∣ b ∣) + (λ a p → cong ∣_∣ p) + (fst (con b)) + where + helper : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : A → Type ℓ'} (P : hLevelTrunc (suc n) A → Type ℓ'') + → ((x : hLevelTrunc (suc n) (Σ A B)) → isOfHLevel (suc n) (P (map fst x))) + → ((a : A) (b : B a) → P ∣ a ∣) + → (p : hLevelTrunc (suc n) (Σ A B)) + → P (map fst p) + helper P hlev pf = Trunc.elim hlev λ pair → pf (fst pair) (snd pair) + + g : Iso (hLevelTrunc (suc n) A) (hLevelTrunc (suc n) B) + Iso.fun g = map f + Iso.inv g = Trunc.rec (isOfHLevelTrunc _) back + Iso.leftInv g = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) + λ a → cong (map fst) (con (f a) .snd ∣ a , refl ∣) + Iso.rightInv g = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) + backSection + +connectedTruncIso2 : ∀ {ℓ} {A B : Type ℓ} (n m : HLevel) (f : A → B) + → Σ[ x ∈ ℕ ] x + n ≡ m + → isConnectedFun m f + → Iso (hLevelTrunc n A) (hLevelTrunc n B) +connectedTruncIso2 {A = A} {B = B} n m f (x , pf) con = + connectedTruncIso n f (isConnectedFunSubtr n x f (transport (λ i → isConnectedFun (pf (~ i)) f) con)) + +connectedTruncEquiv : ∀ {ℓ} {A B : Type ℓ} (n : HLevel) (f : A → B) + → isConnectedFun n f + → hLevelTrunc n A ≃ hLevelTrunc n B +connectedTruncEquiv {A = A} {B = B} n f con = isoToEquiv (connectedTruncIso n f con) + + +-- TODO : Reorganise the following proofs. + +inrConnected : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (n : HLevel) + → (f : C → A) (g : C → B) + → isConnectedFun n f + → isConnectedFun n {A = B} {B = Pushout f g} inr +inrConnected {A = A} {B = B} {C = C} n f g iscon = + elim.isConnectedPrecompose inr n λ P → (k P) , λ b → refl + where + module _ {ℓ : Level} (P : (Pushout f g) → TypeOfHLevel ℓ n) + (h : (b : B) → typ (P (inr b))) + where + Q : A → TypeOfHLevel _ n + Q a = (P (inl a)) + + fun : (c : C) → typ (Q (f c)) + fun c = transport (λ i → typ (P (push c (~ i)))) (h (g c)) + + k : (d : Pushout f g) → typ (P d) + k (inl x) = equiv-proof (elim.isEquivPrecompose f n Q iscon) fun .fst .fst x + k (inr x) = h x + k (push a i) = + hcomp (λ k → λ { (i = i0) → equiv-proof (elim.isEquivPrecompose f n Q iscon) fun .fst .snd i0 a + ; (i = i1) → transportTransport⁻ (λ j → typ (P (push a j))) (h (g a)) k }) + (transp (λ j → typ (P (push a (i ∧ j)))) + (~ i) + (equiv-proof (elim.isEquivPrecompose f n Q iscon) + fun .fst .snd i a)) diff --git a/Cubical/Homotopy/EilenbergSteenrod.agda b/Cubical/Homotopy/EilenbergSteenrod.agda new file mode 100644 index 000000000..cabb12e72 --- /dev/null +++ b/Cubical/Homotopy/EilenbergSteenrod.agda @@ -0,0 +1,47 @@ +{-# OPTIONS --safe #-} + +module Cubical.Homotopy.EilenbergSteenrod where + +{- +This module contains the Eilenberg-Steenrod axioms for ordinary +reduced cohomology theories with binary additivity. +The axioms are based on the ones given in Cavallo's MSc thesis +(https://www.cs.cmu.edu/~ecavallo/works/thesis15.pdf) and +Buchholtz/Favonia (2018) +-} + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.HITs.Wedge +open import Cubical.HITs.Pushout +open import Cubical.HITs.Susp + +open import Cubical.Data.Empty +open import Cubical.Relation.Nullary + +open import Cubical.Data.Nat +open import Cubical.Data.Bool +open import Cubical.Data.Sigma +open import Cubical.Data.Int + +open import Cubical.Algebra.Group hiding (ℤ ; Bool) +open import Cubical.Algebra.AbGroup + +record coHomTheory {ℓ ℓ' : Level} (H : (n : ℤ) → Pointed ℓ → AbGroup ℓ') : Type (ℓ-suc (ℓ-max ℓ ℓ')) + where + Boolℓ : Pointed ℓ + Boolℓ = Lift Bool , lift true + field + Hmap : (n : ℤ) → {A B : Pointed ℓ} (f : A →∙ B) → AbGroupHom (H n B) (H n A) + Suspension : Σ[ F ∈ ((n : ℤ) {A : Pointed ℓ} → AbGroupEquiv (H (sucℤ n) (Susp (typ A) , north)) (H n A)) ] + ({A B : Pointed ℓ} (f : A →∙ B) (n : ℤ) + → fst (Hmap (sucℤ n) (suspFun (fst f) , refl)) ∘ invEq (fst (F n {A = B})) + ≡ invEq (fst (F n {A = A})) ∘ fst (Hmap n f)) + Exactness : {A B : Pointed ℓ} (f : A →∙ B) (n : ℤ) + → Ker (Hmap n f) + ≡ Im (Hmap n {B = _ , inr (pt B)} (cfcod (fst f) , refl)) + Dimension : (n : ℤ) → ¬ n ≡ 0 → isContr (fst (H n Boolℓ)) + BinaryWedge : (n : ℤ) {A B : Pointed ℓ} → AbGroupEquiv (H n (A ⋁ B , (inl (pt A)))) (dirProdAb (H n A) (H n B)) diff --git a/Cubical/Homotopy/Everything.agda b/Cubical/Homotopy/Everything.agda new file mode 100644 index 000000000..02627af7d --- /dev/null +++ b/Cubical/Homotopy/Everything.agda @@ -0,0 +1,10 @@ +{-# OPTIONS --safe --guardedness #-} +module Cubical.Homotopy.Everything where + +import Cubical.Homotopy.Base +import Cubical.Homotopy.Connected +import Cubical.Homotopy.EilenbergSteenrod +import Cubical.Homotopy.Freudenthal +import Cubical.Homotopy.Loopspace +import Cubical.Homotopy.MayerVietorisCofiber +import Cubical.Homotopy.WedgeConnectivity diff --git a/Cubical/Homotopy/Freudenthal.agda b/Cubical/Homotopy/Freudenthal.agda new file mode 100644 index 000000000..9a097f81d --- /dev/null +++ b/Cubical/Homotopy/Freudenthal.agda @@ -0,0 +1,138 @@ +{- + +Freudenthal suspension theorem + +-} +{-# OPTIONS --safe #-} +module Cubical.Homotopy.Freudenthal where + +open import Cubical.Foundations.Everything +-- open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.HITs.Nullification +open import Cubical.HITs.Susp +open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec ; elim to trElim) +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.WedgeConnectivity +open import Cubical.Homotopy.Loopspace +open import Cubical.HITs.SmashProduct + +open import Cubical.HITs.S1 hiding (encode) +open import Cubical.HITs.Sn +open import Cubical.HITs.S2 +open import Cubical.HITs.S3 +open import Cubical.Foundations.Equiv.HalfAdjoint + +module _ {ℓ} (n : HLevel) {A : Pointed ℓ} (connA : isConnected (suc (suc n)) (typ A)) where + + σ : typ A → typ (Ω (∙Susp (typ A))) + σ a = merid a ∙ merid (pt A) ⁻¹ + + private + 2n+2 = suc n + suc n + + module WC (p : north ≡ north) = + WedgeConnectivity (suc n) (suc n) A connA A connA + (λ a b → + ( (σ b ≡ p → hLevelTrunc 2n+2 (fiber (λ x → merid x ∙ merid a ⁻¹) p)) + , isOfHLevelΠ 2n+2 λ _ → isOfHLevelTrunc 2n+2 + )) + (λ a r → ∣ a , (rCancel' (merid a) ∙ rCancel' (merid (pt A)) ⁻¹) ∙ r ∣) + (λ b r → ∣ b , r ∣) + (funExt λ r → + cong′ (λ w → ∣ pt A , w ∣) + (cong (_∙ r) (rCancel' (rCancel' (merid (pt A)))) + ∙ lUnit r ⁻¹)) + + fwd : (p : north ≡ north) (a : typ A) + → hLevelTrunc 2n+2 (fiber σ p) + → hLevelTrunc 2n+2 (fiber (λ x → merid x ∙ merid a ⁻¹) p) + fwd p a = Trunc.rec (isOfHLevelTrunc 2n+2) (uncurry (WC.extension p a)) + + isEquivFwd : (p : north ≡ north) (a : typ A) → isEquiv (fwd p a) + isEquivFwd p a .equiv-proof = + elim.isEquivPrecompose (λ _ → pt A) (suc n) + (λ a → + ( (∀ t → isContr (fiber (fwd p a) t)) + , isProp→isOfHLevelSuc n (isPropΠ λ _ → isPropIsContr) + )) + + (isConnectedPoint (suc n) connA (pt A)) + .equiv-proof + (λ _ → Trunc.elim + (λ _ → isProp→isOfHLevelSuc (n + suc n) isPropIsContr) + (λ fib → + subst (λ k → isContr (fiber k ∣ fib ∣)) + (cong (Trunc.rec (isOfHLevelTrunc 2n+2) ∘ uncurry) + (funExt (WC.right p) ⁻¹)) + (subst isEquiv + (funExt (Trunc.mapId) ⁻¹) + (idIsEquiv _) + .equiv-proof ∣ fib ∣) + )) + .fst .fst a + + interpolate : (a : typ A) + → PathP (λ i → typ A → north ≡ merid a i) (λ x → merid x ∙ merid a ⁻¹) merid + interpolate a i x j = compPath-filler (merid x) (merid a ⁻¹) (~ i) j + + Code : (y : Susp (typ A)) → north ≡ y → Type ℓ + Code north p = hLevelTrunc 2n+2 (fiber σ p) + Code south q = hLevelTrunc 2n+2 (fiber merid q) + Code (merid a i) p = + Glue + (hLevelTrunc 2n+2 (fiber (interpolate a i) p)) + (λ + { (i = i0) → _ , (fwd p a , isEquivFwd p a) + ; (i = i1) → _ , idEquiv _ + }) + + encode : (y : Susp (typ A)) (p : north ≡ y) → Code y p + encode y = J Code ∣ pt A , rCancel' (merid (pt A)) ∣ + + encodeMerid : (a : typ A) → encode south (merid a) ≡ ∣ a , refl ∣ + encodeMerid a = + cong (transport (λ i → gluePath i)) + (funExt⁻ (WC.left refl a) _ ∙ λ i → ∣ a , lem (rCancel' (merid a)) (rCancel' (merid (pt A))) i ∣) + ∙ transport (PathP≡Path gluePath _ _) + (λ i → ∣ a , (λ j k → rCancel-filler' (merid a) i j k) ∣) + where + gluePath : I → Type _ + gluePath i = hLevelTrunc 2n+2 (fiber (interpolate a i) (λ j → merid a (i ∧ j))) + + lem : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : z ≡ y) → (p ∙ q ⁻¹) ∙ q ≡ p + lem p q = assoc p (q ⁻¹) q ⁻¹ ∙∙ cong (p ∙_) (lCancel q) ∙∙ rUnit p ⁻¹ + + contractCodeSouth : (p : north ≡ south) (c : Code south p) → encode south p ≡ c + contractCodeSouth p = + Trunc.elim + (λ _ → isOfHLevelPath 2n+2 (isOfHLevelTrunc 2n+2) _ _) + (uncurry λ a → + J (λ p r → encode south p ≡ ∣ a , r ∣) (encodeMerid a)) + + isConnectedMerid : isConnectedFun 2n+2 (merid {A = typ A}) + isConnectedMerid p = encode south p , contractCodeSouth p + + isConnectedσ : isConnectedFun 2n+2 σ + isConnectedσ = + transport (λ i → isConnectedFun 2n+2 (interpolate (pt A) (~ i))) isConnectedMerid + +isConn→isConnSusp : ∀ {ℓ} {A : Pointed ℓ} → isConnected 2 (typ A) → isConnected 2 (Susp (typ A)) +isConn→isConnSusp {A = A} iscon = ∣ north ∣ + , trElim (λ _ → isOfHLevelSuc 1 (isOfHLevelTrunc 2 _ _)) + (suspToPropElim (pt A) (λ _ → isOfHLevelTrunc 2 _ _) + refl) + +FreudenthalEquiv : ∀ {ℓ} (n : HLevel) (A : Pointed ℓ) + → isConnected (2 + n) (typ A) + → hLevelTrunc ((suc n) + (suc n)) (typ A) + ≃ hLevelTrunc ((suc n) + (suc n)) (typ (Ω (Susp (typ A) , north))) +FreudenthalEquiv n A iscon = connectedTruncEquiv _ + (σ n {A = A} iscon) + (isConnectedσ _ iscon) +FreudenthalIso : ∀ {ℓ} (n : HLevel) (A : Pointed ℓ) + → isConnected (2 + n) (typ A) + → Iso (hLevelTrunc ((suc n) + (suc n)) (typ A)) + (hLevelTrunc ((suc n) + (suc n)) (typ (Ω (Susp (typ A) , north)))) +FreudenthalIso n A iscon = connectedTruncIso _ (σ n {A = A} iscon) (isConnectedσ _ iscon) diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda new file mode 100644 index 000000000..a9fc3e52d --- /dev/null +++ b/Cubical/Homotopy/Loopspace.agda @@ -0,0 +1,81 @@ +{-# OPTIONS --safe #-} + +module Cubical.Homotopy.Loopspace where + +open import Cubical.Core.Everything + +open import Cubical.Data.Nat + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.HITs.SetTruncation +open import Cubical.HITs.Truncation hiding (elim2) renaming (rec to trRec) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv.HalfAdjoint +open Iso + +{- loop space of a pointed type -} +Ω : {ℓ : Level} → Pointed ℓ → Pointed ℓ +Ω (_ , a) = ((a ≡ a) , refl) + +{- n-fold loop space of a pointed type -} +Ω^_ : ∀ {ℓ} → ℕ → Pointed ℓ → Pointed ℓ +(Ω^ 0) p = p +(Ω^ (suc n)) p = Ω ((Ω^ n) p) + +{- homotopy Group -} +π : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Type ℓ +π n A = ∥ typ ((Ω^ n) A) ∥ 2 + +{- loop space map -} +Ω→ : ∀ {ℓA ℓB} {A : Pointed ℓA} {B : Pointed ℓB} (f : A →∙ B) → (Ω A →∙ Ω B) +Ω→ (f , f∙) = (λ p → (sym f∙ ∙ cong f p) ∙ f∙) , cong (λ q → q ∙ f∙) (sym (rUnit (sym f∙))) ∙ lCancel f∙ + +ΩfunExtIso : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') → Iso (typ (Ω (A →∙ B ∙))) (A →∙ Ω B) +fst (fun (ΩfunExtIso A B) p) x = funExt⁻ (cong fst p) x +snd (fun (ΩfunExtIso A B) p) i j = snd (p j) i +fst (inv (ΩfunExtIso A B) (f , p) i) x = f x i +snd (inv (ΩfunExtIso A B) (f , p) i) j = p j i +rightInv (ΩfunExtIso A B) _ = refl +leftInv (ΩfunExtIso A B) _ = refl + +{- Commutativity of loop spaces -} +isComm∙ : ∀ {ℓ} (A : Pointed ℓ) → Type ℓ +isComm∙ A = (p q : typ (Ω A)) → p ∙ q ≡ q ∙ p + +Eckmann-Hilton : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → isComm∙ ((Ω^ (suc n)) A) +Eckmann-Hilton n α β = + transport (λ i → cong (λ x → rUnit x (~ i)) α ∙ cong (λ x → lUnit x (~ i)) β + ≡ cong (λ x → lUnit x (~ i)) β ∙ cong (λ x → rUnit x (~ i)) α) + λ i → (λ j → α (j ∧ ~ i) ∙ β (j ∧ i)) ∙ λ j → α (~ i ∨ j) ∙ β (i ∨ j) + +isCommA→isCommTrunc : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → isComm∙ A + → isOfHLevel (suc n) (typ A) + → isComm∙ (∥ typ A ∥ (suc n) , ∣ pt A ∣) +isCommA→isCommTrunc {A = (A , a)} n comm hlev p q = + ((λ i j → (leftInv (truncIdempotentIso (suc n) hlev) ((p ∙ q) j) (~ i))) + ∙∙ (λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) (cong (trRec hlev (λ x → x)) (p ∙ q))) + ∙∙ (λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) (congFunct {A = ∥ A ∥ (suc n)} {B = A} (trRec hlev (λ x → x)) p q i))) + ∙ ((λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) (comm (cong (trRec hlev (λ x → x)) p) (cong (trRec hlev (λ x → x)) q) i)) + ∙∙ (λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) (congFunct {A = ∥ A ∥ (suc n)} {B = A} (trRec hlev (λ x → x)) q p (~ i))) + ∙∙ (λ i j → (leftInv (truncIdempotentIso (suc n) hlev) ((q ∙ p) j) i))) + +ptdIso→comm : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Type ℓ'} (e : Iso (typ A) B) → isComm∙ A → isComm∙ (B , fun e (pt A)) +ptdIso→comm {A = (A , a)} {B = B} e comm p q = + sym (rightInv (congIso e) (p ∙ q)) + ∙∙ (cong (fun (congIso e)) ((invCongFunct e p q) + ∙∙ (comm (inv (congIso e) p) (inv (congIso e) q)) + ∙∙ (sym (invCongFunct e q p)))) + ∙∙ rightInv (congIso e) (q ∙ p) + +{- Homotopy group version -} +π-comp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → ∥ typ ((Ω^ (suc n)) A) ∥₂ + → ∥ typ ((Ω^ (suc n)) A) ∥₂ → ∥ typ ((Ω^ (suc n)) A) ∥₂ +π-comp n = elim2 (λ _ _ → isSetSetTrunc) λ p q → ∣ p ∙ q ∣₂ + +Eckmann-Hilton-π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p q : ∥ typ ((Ω^ (2 + n)) A) ∥₂) + → π-comp (1 + n) p q ≡ π-comp (1 + n) q p +Eckmann-Hilton-π n = elim2 (λ x y → isOfHLevelPath 2 isSetSetTrunc _ _) + λ p q → cong ∣_∣₂ (Eckmann-Hilton n p q) diff --git a/Cubical/Homotopy/MayerVietorisCofiber.agda b/Cubical/Homotopy/MayerVietorisCofiber.agda new file mode 100644 index 000000000..894352405 --- /dev/null +++ b/Cubical/Homotopy/MayerVietorisCofiber.agda @@ -0,0 +1,175 @@ +{- + +Mayer-Vietoris cofiber sequence: + + Let X be a pointed type, and let a span B ←[f]- X -[g]→ C be given. + Then the mapping cone of the canonical map (B ⋁ C) → B ⊔_X C is equivalent to Susp X. + +The sequence Susp X → (B ⋁ C) → B ⊔_X C therefore induces a long exact sequence in cohomology. +Proof is adapted from Evan Cavallo's master's thesis. + +-} +{-# OPTIONS --safe #-} +module Cubical.Homotopy.MayerVietorisCofiber where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Pointed +open import Cubical.Data.Unit +open import Cubical.HITs.MappingCones +open import Cubical.HITs.Pushout +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge + +module _ {ℓX ℓB ℓC} {X : Pointed ℓX} {B : Type ℓB} {C : Type ℓC} (f : X .fst → B) (g : X .fst → C) + where + + private + Y : Pointed _ + Y = (B , f (X .snd)) + + Z : Pointed _ + Z = (C , g (X .snd)) + + wedgeToPushout : Y ⋁ Z → Pushout f g + wedgeToPushout (inl y) = inl y + wedgeToPushout (inr z) = inr z + wedgeToPushout (push _ i) = push (pt X) i + + pushoutToSusp : Pushout f g → Susp (X .fst) + pushoutToSusp (inl y) = north + pushoutToSusp (inr z) = south + pushoutToSusp (push x i) = merid x i + + {- + Coherence lemma: + To construct a function f : (c : Cone wedgeToPushout) → D c, we can always + adjust the definition of (f (spoke (inr z) i)) so that there is a canonical + choice for (f (spoke (push tt j) i)) + -} + module Helper {ℓD} {D : Cone wedgeToPushout → Type ℓD} + (inj* : ∀ w → D (inj w)) + (hub* : D hub) + (inl* : ∀ y → PathP (λ i → D (spoke (inl y) i)) hub* (inj* (inl y))) + (inr* : ∀ z → PathP (λ i → D (spoke (inr z) i)) hub* (inj* (inr z))) + where + + cap : (i j : I) → D (spoke (push tt j) i) + cap i j = + fill + (λ i → D (spoke (push tt j) (~ i))) + (λ i → λ + { (j = i0) → inl* (Y .snd) (~ i) + ; (j = i1) → inr* (Z .snd) (~ i) + }) + (inS (inj* (push (X .snd) j))) + (~ i) + + cap0 : (j : I) → D hub + cap0 = cap i0 + + face-i≡0 : (k j : I) → D hub + face-i≡0 k j = + hfill + (λ j → λ + { (k = i0) → cap0 j + ; (k = i1) → hub* + }) + (inS hub*) + j + + inrFiller : ∀ z → (k i : I) → D (spoke (inr z) i) + inrFiller z k i = + hfill + (λ k → λ + { (i = i0) → face-i≡0 k i1 + ; (i = i1) → inj* (inr z) + }) + (inS (inr* z i)) + k + + fun : ∀ c → D c + fun (inj w) = inj* w + fun hub = hub* + fun (spoke (inl y) i) = inl* y i + fun (spoke (inr z) i) = inrFiller z i1 i + fun (spoke (push tt j) i) = + hcomp + (λ k → λ + { (i = i0) → face-i≡0 k j + ; (i = i1) → inj* (push (X .snd) j) + ; (j = i0) → inl* (Y .snd) i + }) + (cap i j) + + equiv : Cone wedgeToPushout ≃ Susp (X .fst) + equiv = isoToEquiv (iso fwd bwd fwdBwd bwdFwd) + where + fwd : Cone wedgeToPushout → Susp (X .fst) + fwd (inj w) = pushoutToSusp w + fwd hub = north + fwd (spoke (inl y) i) = north + fwd (spoke (inr z) i) = merid (X .snd) i + fwd (spoke (push tt j) i) = merid (X .snd) (i ∧ j) + + bwd : Susp (X .fst) → Cone wedgeToPushout + bwd north = hub + bwd south = hub + bwd (merid x i) = + hcomp + (λ k → λ + { (i = i0) → spoke (inl (f x)) (~ k) + ; (i = i1) → spoke (inr (g x)) (~ k) + }) + (inj (push x i)) + + bwdPushout : (w : Pushout f g) → bwd (pushoutToSusp w) ≡ inj w + bwdPushout (inl y) = spoke (inl y) + bwdPushout (inr z) = spoke (inr z) + bwdPushout (push x i) k = + hfill + (λ k → λ + { (i = i0) → spoke (inl (f x)) (~ k) + ; (i = i1) → spoke (inr (g x)) (~ k) + }) + (inS (inj (push x i))) + (~ k) + + bwdMeridPt : refl ≡ cong bwd (merid (X .snd)) + bwdMeridPt j i = + hcomp + (λ k → λ + { (i = i0) → spoke (inl (Y .snd)) (~ k) + ; (i = i1) → spoke (inr (Z .snd)) (~ k) + ; (j = i0) → spoke (push _ i) (~ k) + }) + (inj (push (X .snd) i)) + + bwdFwd : (c : Cone wedgeToPushout) → bwd (fwd c) ≡ c + bwdFwd = + Helper.fun + bwdPushout + refl + (λ y i k → spoke (inl y) (i ∧ k)) + (λ z i k → + hcomp + (λ l → λ + { (i = i0) → hub + ; (i = i1) → spoke (inr z) k + ; (k = i0) → bwdMeridPt l i + ; (k = i1) → spoke (inr z) i + }) + (spoke (inr z) (i ∧ k))) + + fwdBwd : (s : Susp (X .fst)) → fwd (bwd s) ≡ s + fwdBwd north = refl + fwdBwd south = merid (X .snd) + fwdBwd (merid a i) j = + fill + (λ _ → Susp (X .fst)) + (λ j → λ + { (i = i0) → north + ; (i = i1) → merid (X .snd) (~ j) + }) + (inS (merid a i)) + (~ j) diff --git a/Cubical/Homotopy/WedgeConnectivity.agda b/Cubical/Homotopy/WedgeConnectivity.agda new file mode 100644 index 000000000..5a50bd470 --- /dev/null +++ b/Cubical/Homotopy/WedgeConnectivity.agda @@ -0,0 +1,59 @@ +{-# OPTIONS --safe #-} +module Cubical.Homotopy.WedgeConnectivity where + +open import Cubical.Foundations.Everything +open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.HITs.Nullification +open import Cubical.HITs.Susp +open import Cubical.HITs.Truncation as Trunc +open import Cubical.Homotopy.Connected + +module WedgeConnectivity {ℓ ℓ' ℓ''} (n m : ℕ) + (A : Pointed ℓ) (connA : isConnected (suc n) (typ A)) + (B : Pointed ℓ') (connB : isConnected (suc m) (typ B)) + (P : typ A → typ B → TypeOfHLevel ℓ'' (n + m)) + (f : (a : typ A) → P a (pt B) .fst) + (g : (b : typ B) → P (pt A) b .fst) + (p : f (pt A) ≡ g (pt B)) + where + + private + Q : typ A → TypeOfHLevel _ n + Q a = + ( (Σ[ k ∈ ((b : typ B) → P a b .fst) ] k (pt B) ≡ f a) + , isOfHLevelRetract n + (λ {(h , q) → h , funExt λ _ → q}) + (λ {(h , q) → h , funExt⁻ q _}) + (λ _ → refl) + (isOfHLevelPrecomposeConnected n m (P a) (λ _ → pt B) + (isConnectedPoint m connB (pt B)) (λ _ → f a)) + ) + + main : isContr (fiber (λ s _ → s (pt A)) (λ _ → g , p ⁻¹)) + main = + elim.isEquivPrecompose (λ _ → pt A) n Q + (isConnectedPoint n connA (pt A)) + .equiv-proof (λ _ → g , p ⁻¹) + + + extension : ∀ a b → P a b .fst + extension a b = main .fst .fst a .fst b + + left : ∀ a → extension a (pt B) ≡ f a + left a = main .fst .fst a .snd + + right : ∀ b → extension (pt A) b ≡ g b + right = funExt⁻ (cong fst (funExt⁻ (main .fst .snd) _)) + + hom : left (pt A) ⁻¹ ∙ right (pt B) ≡ p + hom i j = hcomp (λ k → λ { (i = i1) → p j + ; (j = i0) → (cong snd (funExt⁻ (main .fst .snd) tt)) i (~ j) + ; (j = i1) → right (pt B) (i ∨ k)}) + (cong snd (funExt⁻ (main .fst .snd) tt) i (~ j)) + + hom' : left (pt A) ≡ right (pt B) ∙ sym p + hom' = (lUnit (left _) ∙ cong (_∙ left (pt A)) (sym (rCancel (right (pt B))))) + ∙∙ sym (assoc _ _ _) + ∙∙ cong (right (pt B) ∙_) (sym (symDistr (left (pt A) ⁻¹) (right (pt B))) ∙ (cong sym hom)) diff --git a/Cubical/Induction/Everything.agda b/Cubical/Induction/Everything.agda index b189b82cd..aeb577db2 100644 --- a/Cubical/Induction/Everything.agda +++ b/Cubical/Induction/Everything.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe --guardedness #-} module Cubical.Induction.Everything where -open import Cubical.Induction.WellFounded public +import Cubical.Induction.WellFounded diff --git a/Cubical/Induction/WellFounded.agda b/Cubical/Induction/WellFounded.agda index 3ed241ef3..561cb3ca9 100644 --- a/Cubical/Induction/WellFounded.agda +++ b/Cubical/Induction/WellFounded.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Induction.WellFounded where diff --git a/Cubical/Modalities/Everything.agda b/Cubical/Modalities/Everything.agda index f835e615d..862b98edc 100644 --- a/Cubical/Modalities/Everything.agda +++ b/Cubical/Modalities/Everything.agda @@ -1,4 +1,5 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe --guardedness #-} module Cubical.Modalities.Everything where -open import Cubical.Modalities.Modality public +import Cubical.Modalities.Lex +import Cubical.Modalities.Modality diff --git a/Cubical/Modalities/Lex.agda b/Cubical/Modalities/Lex.agda new file mode 100644 index 000000000..0c1613e8b --- /dev/null +++ b/Cubical/Modalities/Lex.agda @@ -0,0 +1,271 @@ +{-# OPTIONS --safe --postfix-projections #-} + +open import Cubical.Foundations.Everything renaming (uncurry to λ⟨,⟩_) +open import Cubical.Data.Sigma.Properties +open import Cubical.Foundations.CartesianKanOps + +module Cubical.Modalities.Lex + (◯ : ∀ {ℓ} → Type ℓ → Type ℓ) + (η : ∀ {ℓ} {A : Type ℓ} → A → ◯ A) + (isModal : ∀ {ℓ} → Type ℓ → Type ℓ) + (let isModalFam = λ {ℓ ℓ′ : Level} {A : Type ℓ} (B : A → Type ℓ′) → (x : A) → isModal (B x)) + (idemp : ∀ {ℓ} {A : Type ℓ} → isModal (◯ A)) + (≡-modal : ∀ {ℓ} {A : Type ℓ} {x y : A} (A-mod : isModal A) → isModal (x ≡ y)) + (◯-ind : ∀ {ℓ ℓ′} {A : Type ℓ} {B : ◯ A → Type ℓ′} (B-mod : isModalFam B) (f : (x : A) → B (η x)) → ([x] : ◯ A) → B [x]) + (◯-ind-β : ∀ {ℓ ℓ′} {A : Type ℓ} {B : ◯ A → Type ℓ′} (B-mod : isModalFam B) (f : (x : A) → B (η x)) (x : A) → ◯-ind B-mod f (η x) ≡ f x) + (let Type◯ = λ (ℓ : Level) → Σ (Type ℓ) isModal) + (◯-lex : ∀ {ℓ} → isModal (Type◯ ℓ)) + where + + +private + variable + ℓ ℓ′ : Level + +η-at : (A : Type ℓ) → A → ◯ A +η-at _ = η + +module _ where + private + variable + A : Type ℓ + B : Type ℓ′ + + + module ◯-rec (B-mod : isModal B) (f : A → B) where + abstract + ◯-rec : ◯ A → B + ◯-rec = ◯-ind (λ _ → B-mod) f + + ◯-rec-β : (x : A) → ◯-rec (η x) ≡ f x + ◯-rec-β = ◯-ind-β (λ _ → B-mod) f + + open ◯-rec + + module ◯-map (f : A → B) where + abstract + ◯-map : ◯ A → ◯ B + ◯-map = ◯-rec idemp λ x → η (f x) + + ◯-map-β : (x : A) → ◯-map (η x) ≡ η (f x) + ◯-map-β x = ◯-rec-β idemp _ x + + +open ◯-rec +open ◯-map + +module IsModalToUnitIsEquiv (A : Type ℓ) (A-mod : isModal A) where + abstract + inv : ◯ A → A + inv = ◯-rec A-mod λ x → x + + η-retract : retract η inv + η-retract = ◯-rec-β _ _ + + η-section : section η inv + η-section = ◯-ind (λ _ → ≡-modal idemp) λ x i → η (η-retract x i) + + η-iso : Iso A (◯ A) + Iso.fun η-iso = η + Iso.inv η-iso = inv + Iso.rightInv η-iso = η-section + Iso.leftInv η-iso = η-retract + + η-is-equiv : isEquiv (η-at A) + η-is-equiv = isoToIsEquiv η-iso + +abstract + unit-is-equiv-to-is-modal : {A : Type ℓ} → isEquiv (η-at A) → isModal A + unit-is-equiv-to-is-modal p = transport (cong isModal (sym (ua (η , p)))) idemp + + retract-is-modal + : {A : Type ℓ} {B : Type ℓ′} + → (A-mod : isModal A) (f : A → B) (g : B → A) (r : retract g f) + → isModal B + retract-is-modal {A = A} {B = B} A-mod f g r = + unit-is-equiv-to-is-modal (isoToIsEquiv (iso η η-inv η-section η-retract)) + where + η-inv : ◯ B → B + η-inv = f ∘ ◯-rec A-mod g + + η-retract : retract η η-inv + η-retract b = cong f (◯-rec-β A-mod g b) ∙ r b + + η-section : section η η-inv + η-section = ◯-ind (λ _ → ≡-modal idemp) (cong η ∘ η-retract) + + +module LiftFam {A : Type ℓ} (B : A → Type ℓ′) where + module M = IsModalToUnitIsEquiv (Type◯ ℓ′) ◯-lex + + abstract + ◯-lift-fam : ◯ A → Type◯ ℓ′ + ◯-lift-fam = M.inv ∘ ◯-map (λ a → ◯ (B a) , idemp) + + ⟨◯⟩ : ◯ A → Type ℓ′ + ⟨◯⟩ [a] = ◯-lift-fam [a] .fst + + ⟨◯⟩-modal : isModalFam ⟨◯⟩ + ⟨◯⟩-modal [a] = ◯-lift-fam [a] .snd + + ⟨◯⟩-compute : (x : A) → ⟨◯⟩ (η x) ≡ ◯ (B x) + ⟨◯⟩-compute x = + ⟨◯⟩ (η x) + ≡⟨ cong (fst ∘ M.inv) (◯-map-β _ _) ⟩ + M.inv (η (◯ (B x) , idemp)) .fst + ≡⟨ cong fst (M.η-retract _) ⟩ + ◯ (B x) ∎ + + + +open LiftFam using (⟨◯⟩; ⟨◯⟩-modal; ⟨◯⟩-compute) + +module LiftFamExtra {A : Type ℓ} {B : A → Type ℓ′} where + ⟨◯⟩←◯ : ∀ {a} → ◯ (B a) → ⟨◯⟩ B (η a) + ⟨◯⟩←◯ = transport (sym (⟨◯⟩-compute B _)) + + ⟨◯⟩→◯ : ∀ {a} → ⟨◯⟩ B (η a) → ◯ (B a) + ⟨◯⟩→◯ = transport (⟨◯⟩-compute B _) + + ⟨η⟩ : ∀ {a} → B a → ⟨◯⟩ B (η a) + ⟨η⟩ = ⟨◯⟩←◯ ∘ η + + abstract + ⟨◯⟩→◯-section : ∀ {a} → section (⟨◯⟩→◯ {a}) ⟨◯⟩←◯ + ⟨◯⟩→◯-section = transport⁻Transport (sym (⟨◯⟩-compute _ _)) + + +module Combinators where + private + variable + ℓ′′ : Level + A A′ : Type ℓ + B : A → Type ℓ′ + C : Σ A B → Type ℓ′′ + + λ/coe⟨_⟩_ : (p : A ≡ A′) → ((x : A′) → B (coe1→0 (λ i → p i) x)) → ((x : A) → B x) + λ/coe⟨_⟩_ {B = B} p f = coe1→0 (λ i → (x : p i) → B (coei→0 (λ j → p j) i x)) f + +open Combinators + + +module _ {A : Type ℓ} {B : A → Type ℓ′} where + abstract + Π-modal : isModalFam B → isModal ((x : A) → B x) + Π-modal B-mod = retract-is-modal idemp η-inv η retr + where + η-inv : ◯ ((x : A) → B x) → (x : A) → B x + η-inv [f] x = ◯-rec (B-mod x) (λ f → f x) [f] + + retr : retract η η-inv + retr f = funExt λ x → ◯-rec-β (B-mod x) _ f + + Σ-modal : isModal A → isModalFam B → isModal (Σ A B) + Σ-modal A-mod B-mod = retract-is-modal idemp η-inv η retr + where + h : ◯ (Σ A B) → A + h = ◯-rec A-mod fst + + h-β : (x : Σ A B) → h (η x) ≡ fst x + h-β = ◯-rec-β A-mod fst + + f : (i : I) (x : Σ A B) → B (h-β x i) + f i x = coe1→i (λ j → B (h-β x j)) i (snd x) + + η-inv : ◯ (Σ A B) → Σ A B + η-inv y = h y , ◯-ind (B-mod ∘ h) (f i0) y + + retr : (p : Σ A B) → η-inv (η p) ≡ p + retr p = + η-inv (η p) + ≡⟨ ΣPathP (refl , ◯-ind-β _ _ _) ⟩ + h (η p) , f i0 p + ≡⟨ ΣPathP (h-β _ , λ i → f i p) ⟩ + p ∎ + + +module Σ-commute {A : Type ℓ} (B : A → Type ℓ′) where + open LiftFamExtra + + ◯Σ = ◯ (Σ A B) + + module Σ◯ where + Σ◯ = Σ (◯ A) (⟨◯⟩ B) + abstract + Σ◯-modal : isModal Σ◯ + Σ◯-modal = Σ-modal idemp (⟨◯⟩-modal _) + + open Σ◯ + + η-Σ◯ : Σ A B → Σ◯ + η-Σ◯ (x , y) = η x , ⟨η⟩ y + + module Push where + abstract + fun : ◯Σ → Σ◯ + fun = ◯-rec Σ◯-modal η-Σ◯ + + compute : fun ∘ η ≡ η-Σ◯ + compute = funExt (◯-rec-β _ _) + + module Unpush where + abstract + fun : Σ◯ → ◯Σ + fun = + λ⟨,⟩ ◯-ind (λ _ → Π-modal λ _ → idemp) λ x → + λ/coe⟨ ⟨◯⟩-compute B x ⟩ ◯-map (x ,_) + + compute : fun ∘ η-Σ◯ ≡ η + compute = + funExt λ p → + fun (η-Σ◯ p) + ≡⟨ funExt⁻ (◯-ind-β _ _ _) _ ⟩ + transport refl (◯-map _ _) + ≡⟨ transportRefl _ ⟩ + ◯-map _ (⟨◯⟩→◯ (⟨η⟩ _)) + ≡⟨ cong (◯-map _) (⟨◯⟩→◯-section _) ⟩ + ◯-map _ (η _) + ≡⟨ ◯-map-β _ _ ⟩ + η p ∎ + + + push-unpush-compute : Push.fun ∘ Unpush.fun ∘ η-Σ◯ ≡ η-Σ◯ + push-unpush-compute = + Push.fun ∘ Unpush.fun ∘ η-Σ◯ + ≡⟨ cong (Push.fun ∘_) Unpush.compute ⟩ + Push.fun ∘ η + ≡⟨ Push.compute ⟩ + η-Σ◯ ∎ + + unpush-push-compute : Unpush.fun ∘ Push.fun ∘ η ≡ η + unpush-push-compute = + Unpush.fun ∘ Push.fun ∘ η + ≡⟨ cong (Unpush.fun ∘_) Push.compute ⟩ + Unpush.fun ∘ η-Σ◯ + ≡⟨ Unpush.compute ⟩ + η ∎ + + is-section : section Unpush.fun Push.fun + is-section = ◯-ind (λ x → ≡-modal idemp) λ x i → unpush-push-compute i x + + is-retract : retract Unpush.fun Push.fun + is-retract = + λ⟨,⟩ ◯-ind (λ _ → Π-modal λ _ → ≡-modal Σ◯-modal) λ x → + λ/coe⟨ ⟨◯⟩-compute B x ⟩ + ◯-ind + (λ _ → ≡-modal Σ◯-modal) + (λ y i → push-unpush-compute i (x , y)) + + push-sg-is-equiv : isEquiv Push.fun + push-sg-is-equiv = isoToIsEquiv (iso Push.fun Unpush.fun is-retract is-section) + + + +isConnected : Type ℓ → Type ℓ +isConnected A = isContr (◯ A) + + + +module FormalDiskBundle {A : Type ℓ} where + 𝔻 : A → Type ℓ + 𝔻 a = Σ[ x ∈ A ] η a ≡ η x diff --git a/Cubical/Modalities/Modality.agda b/Cubical/Modalities/Modality.agda index c4a80d095..fc011e597 100644 --- a/Cubical/Modalities/Modality.agda +++ b/Cubical/Modalities/Modality.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Modalities.Modality where {- @@ -11,7 +11,7 @@ open import Cubical.Foundations.Everything record Modality ℓ : Type (ℓ-suc ℓ) where field isModal : Type ℓ → Type ℓ - isModalIsProp : {A : Type ℓ} → isProp (isModal A) + isPropIsModal : {A : Type ℓ} → isProp (isModal A) ◯ : Type ℓ → Type ℓ -- \ciO ◯-isModal : {A : Type ℓ} → isModal (◯ A) @@ -29,7 +29,7 @@ record Modality ℓ : Type (ℓ-suc ℓ) where ◯-=-isModal : {A : Type ℓ} (x y : ◯ A) → isModal (x ≡ y) ◯-Types : Type (ℓ-suc ℓ) - ◯-Types = Σ[ A ∈ Type ℓ ] isModal A + ◯-Types = TypeWithStr ℓ isModal {- elimination rules -} @@ -80,18 +80,14 @@ record Modality ℓ : Type (ℓ-suc ℓ) where {- modal types and [η] being an equivalence -} - isModalToIsEquiv : {A : Type ℓ} → isModal A → isEquiv (η {A}) - isModalToIsEquiv {A} w = isoToIsEquiv (iso (η {A}) η-inv inv-l inv-r) - where η-inv : ◯ A → A - η-inv = ◯-rec w (idfun A) - - inv-r : (a : A) → η-inv (η a) ≡ a - inv-r = ◯-rec-β w (idfun A) - - inv-l : (a : ◯ A) → η (η-inv a) ≡ a - inv-l = ◯-elim (λ a₀ → ◯-=-isModal _ _) - (λ a₀ → cong η (inv-r a₀)) + isModalToIso : {A : Type ℓ} → isModal A → Iso A (◯ A) + Iso.fun (isModalToIso _) = η + Iso.inv (isModalToIso w) = ◯-rec w (idfun _) + Iso.rightInv (isModalToIso w) = ◯-elim (λ _ → ◯-=-isModal _ _) (λ a₀ → cong η (◯-rec-β w (idfun _) a₀)) + Iso.leftInv (isModalToIso w) = ◯-rec-β w (idfun _) + isModalToIsEquiv : {A : Type ℓ} → isModal A → isEquiv (η {A}) + isModalToIsEquiv {A} w = isoToIsEquiv (isModalToIso w) isEquivToIsModal : {A : Type ℓ} → isEquiv (η {A}) → isModal A isEquivToIsModal {A} eq = equivPreservesIsModal (invEquiv (η , eq)) ◯-isModal diff --git a/Cubical/Papers/Everything.agda b/Cubical/Papers/Everything.agda new file mode 100644 index 000000000..65ddd0b73 --- /dev/null +++ b/Cubical/Papers/Everything.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe --guardedness #-} +module Cubical.Papers.Everything where + +import Cubical.Papers.RepresentationIndependence +import Cubical.Papers.Synthetic +import Cubical.Papers.ZCohomology diff --git a/Cubical/Papers/RepresentationIndependence.agda b/Cubical/Papers/RepresentationIndependence.agda new file mode 100644 index 000000000..94c495dcf --- /dev/null +++ b/Cubical/Papers/RepresentationIndependence.agda @@ -0,0 +1,314 @@ +{- +Please do not move this file. Changes should only be made if necessary. + +This file contains pointers to the code examples and main results from the paper: + +Internalizing Representation Independence with Univalence + +Carlo Angiuli, Evan Cavallo, Anders Mörtberg, Max Zeuner + +Preprint: https://arxiv.org/abs/2009.05547 + +-} +{-# OPTIONS --safe #-} +module Cubical.Papers.RepresentationIndependence where + + +-- 2.1 +import Agda.Builtin.Cubical.Path as Path +import Cubical.Foundations.Prelude as Prelude +-- 2.2 +import Cubical.Foundations.Univalence as Univalence +import Cubical.Foundations.HLevels as HLevels +import Cubical.Foundations.Equiv as Equivalences +import Cubical.Data.Sigma.Properties as Sigma +-- 2.3 +import Cubical.HITs.PropositionalTruncation as PropositionalTruncation +import Cubical.HITs.Cost.Base as CostMonad +import Cubical.HITs.SetQuotients as SetQuotients +import Cubical.HITs.Rationals.QuoQ as SetQuoQ +import Cubical.HITs.Rationals.SigmaQ as SigmaQ +-- 3.1 +import Cubical.Foundations.SIP as SIP +import Cubical.Structures.Axioms as Axioms +import Cubical.Algebra.Semigroup.Base as Semigroup +open import Cubical.Data.Sigma.Base +-- 3.2 +import Cubical.Structures.Pointed as PointedStr +import Cubical.Structures.Constant as ConstantStr +import Cubical.Structures.Product as ProductStr +import Cubical.Structures.Function as FunctionStr +import Cubical.Structures.Maybe as MaybeStr +import Cubical.Foundations.Structure as Structure +import Cubical.Structures.Auto as Auto +open import Cubical.Data.Maybe.Base +-- 4.1 +import Cubical.Data.Vec.Base as Vector +import Cubical.Algebra.Matrix as Matrices +import Cubical.Data.FinData.Base as Finite +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Data.Bool.Base +-- 4.2 +import Cubical.Structures.Queue as Queues +import Cubical.Data.Queue.Truncated2List as BatchedQueues +-- 5.1 +import Cubical.Relation.Binary.Base as BinRel +import Cubical.Relation.ZigZag.Base as QER +import Cubical.Relation.ZigZag.Applications.MultiSet + as MultiSets +-- 5.2 +import Cubical.Foundations.RelationalStructure as RelStructure +import Cubical.Structures.Relational.Function as RelFunction + + + +------------------------------------------------------------------------- +-- 2. Programming in Cubical Type Theory +-- 2.1 Equalities as Paths +-- 2.2 Univalence +-- 2.3 Higher Inductive Types +------------------------------------------------------------------------- + + +-- 2.1 Equalities as Paths +open Path using (PathP) public +open Prelude using (_≡_ ; refl ; cong ; funExt + ; transport ; subst ; J) public + + +-- 2.2 Univalence +open Univalence using (ua ; uaβ) public + +-- Sets and Propositions +open Prelude using (isProp ; isSet) public +open HLevels using (isPropΠ) public +open Prelude using (isContr) public + +-- Equivalences and Isomorphisms +open Equivalences using (isEquiv ; _≃_) public +open Equivalences renaming (fiber to preim) public +open Sigma using (ΣPath≃PathΣ) public +open Equivalences renaming (propBiimpl→Equiv to prop≃) public + + +-- 2.3 Higher Inductive Types +-- Propositional Truncation +open PropositionalTruncation using (∥_∥ ; map) public +open CostMonad using (Cost ; Cost≡ ; _>>=_ ; return + ; fib ; fibTail) public +-- Computation +_ : fib 20 ≡ (6765 , PropositionalTruncation.∣ 21890 ∣) +_ = refl + +_ : fibTail 20 ≡ (6765 , PropositionalTruncation.∣ 19 ∣) +_ = refl + +-- Set Quotients +open SetQuotients using (_/_ ; setQuotUniversal) public +-- Rational Numbers +open SetQuoQ using (_∼_ ; ℚ) public +open SigmaQ renaming (ℚ to ℚ') public + + + +------------------------------------------------------------------------- +-- 3. The Structure Identity Principle +-- 3.1 Structures +-- 3.2 Building Strucutres +------------------------------------------------------------------------- + +-- 3.1 Structures +open SIP using (TypeWithStr ; StrEquiv ; _≃[_]_ + ; UnivalentStr ; SIP ; sip) public + +-- the last two terms above correspond to lemma 3.3 +-- and corollary 3.4 respectively +open Axioms using ( AxiomsStructure ; AxiomsEquivStr + ; axiomsUnivalentStr ; transferAxioms) public + +-- Monoids are defined using records and Σ-types in the library + +RawMonoidStructure : Type → Type +RawMonoidStructure X = X × (X → X → X) + +MonoidAxioms : (M : Type) → RawMonoidStructure M → Type +MonoidAxioms M (e , _·_) = Semigroup.IsSemigroup _·_ + × ((x : M) → (x · e ≡ x) × (e · x ≡ x)) + +MonoidStructure : Type → Type +MonoidStructure = AxiomsStructure RawMonoidStructure MonoidAxioms + +Monoid : Type₁ +Monoid = TypeWithStr ℓ-zero MonoidStructure + +MonoidEquiv : (M N : Monoid) → fst M ≃ fst N → Type +MonoidEquiv (_ , (εᴹ , _·_) , _) (_ , (εᴺ , _∗_) , _) (φ , _) = + (φ εᴹ ≡ εᴺ) × (∀ x y → φ (x · y) ≡ (φ x) ∗ (φ y)) + + +-- 3.2 Building Structures +-- Constant and Pointed Structures +open PointedStr using (PointedStructure ; PointedEquivStr + ; pointedUnivalentStr) public +open ConstantStr using (ConstantStructure ; ConstantEquivStr + ; constantUnivalentStr) public + +-- Product Structures +open ProductStr using (ProductStructure ; ProductEquivStr + ; productUnivalentStr) public + +-- Function Structures +open FunctionStr using (FunctionEquivStr) public + +-- Maybe Structures +open MaybeStr using (MaybeEquivStr) public + +-- Transport Structures +open Structure using (EquivAction) public +open SIP using (TransportStr ; TransportStr→UnivalentStr + ; UnivalentStr→TransportStr) public +open Structure using (EquivAction→StrEquiv) public +open FunctionStr using (FunctionEquivStr+) public + +-- Monoids Revisited + +RawMonoid : Type₁ +RawMonoid = TypeWithStr _ RawMonoidStructure + +Monoid→RawMonoid : Monoid → RawMonoid +Monoid→RawMonoid (A , r , _) = (A , r) + +RawMonoidEquivStr = Auto.AutoEquivStr RawMonoidStructure + +rawMonoidUnivalentStr : UnivalentStr _ RawMonoidEquivStr +rawMonoidUnivalentStr = Auto.autoUnivalentStr RawMonoidStructure + +isPropMonoidAxioms : (M : Type) (s : RawMonoidStructure M) → isProp (MonoidAxioms M s) +isPropMonoidAxioms M (e , _·_) = + HLevels.isPropΣ + (Semigroup.isPropIsSemigroup _·_) + (λ α → isPropΠ λ _ → + HLevels.isProp× + (Semigroup.IsSemigroup.is-set α _ _) + (Semigroup.IsSemigroup.is-set α _ _)) + +MonoidEquivStr : StrEquiv MonoidStructure ℓ-zero +MonoidEquivStr = AxiomsEquivStr RawMonoidEquivStr MonoidAxioms + +monoidUnivalentStr : UnivalentStr MonoidStructure MonoidEquivStr +monoidUnivalentStr = axiomsUnivalentStr _ isPropMonoidAxioms rawMonoidUnivalentStr + +MonoidΣPath : (M N : Monoid) → (M ≃[ MonoidEquivStr ] N) ≃ (M ≡ N) +MonoidΣPath = SIP monoidUnivalentStr + +InducedMonoid : (M : Monoid) (N : RawMonoid) (e : M .fst ≃ N .fst) + → RawMonoidEquivStr (Monoid→RawMonoid M) N e → Monoid +InducedMonoid M N e r = + Axioms.inducedStructure rawMonoidUnivalentStr M N (e , r) + +InducedMonoidPath : (M : Monoid) (N : RawMonoid) (e : M .fst ≃ N .fst) + (E : RawMonoidEquivStr (Monoid→RawMonoid M) N e) + → M ≡ InducedMonoid M N e E +InducedMonoidPath M N e E = + MonoidΣPath M (InducedMonoid M N e E) .fst (e , E) + +-- Automation +open Auto using (Transp[_] ; AutoEquivStr ; autoUnivalentStr) public +module _ (A : Type) (Aset : isSet A) where + RawQueueEquivStr = + AutoEquivStr (λ (X : Type) → X × (A → X → X) × (X → Transp[ Maybe (X × A) ])) + + + +------------------------------------------------------------------------- +-- 4. Representation Independence through the SIP +-- 4.1 Matrices +-- 4.2 Queues +------------------------------------------------------------------------- + + +-- 4.1 Matrices +open Vector using (Vec) public +open Finite using (Fin ; _==_) public +open Matrices using (VecMatrix ; FinMatrix ; FinMatrix≡VecMatrix + ; FinMatrix≃VecMatrix) public +open Matrices.FinMatrixAbGroup using (addFinMatrix ; addFinMatrixComm) public + +-- example (not in the library) +open import Cubical.Data.Int hiding (-_) + +ℤ-AbGroup : AbGroup ℓ-zero +ℤ-AbGroup = makeAbGroup {G = ℤ} 0 _+_ -_ isSetℤ +Assoc (λ x _ → x) rem +Comm + where + -_ : ℤ → ℤ + - x = 0 - x + rem : (x : ℤ) → x + (- x) ≡ 0 + rem x = +Comm x (pos 0 - x) Prelude.∙ minusPlus x 0 + +module experiment where + open Prelude + + M : FinMatrix ℤ 2 2 + M i j = if i == j then 1 else 0 + + N : FinMatrix ℤ 2 2 + N i j = if i == j then 0 else 1 + + replaceGoal : {A B : Type} {x y : A} → (e : A ≃ B) + (h : invEq e (equivFun e x) ≡ invEq e (equivFun e y)) → x ≡ y + replaceGoal e h = sym (retEq e _) ∙∙ h ∙∙ retEq e _ + + _ : addFinMatrix ℤ-AbGroup M N ≡ (λ _ _ → 1) + _ = replaceGoal (FinMatrix≃VecMatrix) refl + + +-- 4.2 Queues +open Queues.Queues-on using (RawQueueStructure ; QueueAxioms) public +open BatchedQueues.Truncated2List renaming (Q to BatchedQueueHIT) + using (Raw-1≡2 ; WithLaws) public + + + +------------------------------------------------------------------------- +-- 5. Structured Equivalences from Structured Relations +-- 5.1 Quasi-Equivalence Relations +-- 5.2 Structured Relations +------------------------------------------------------------------------- + + +-- 5.1 Quasi-Equivalence Relations +--Lemma (5.1) +open BinRel using (idPropRel ; invPropRel + ; compPropRel ; graphRel) public +-- Definitions (5.2) and (5.3) +open QER using (isZigZagComplete ; isQuasiEquivRel) public +-- Lemma (5.4) +open QER.QER→Equiv using (Thm ; bwd≡ToRel) public +-- Multisets +open MultiSets renaming (AList to AssocList) public +open MultiSets.Lists&ALists using (addIfEq ; R ; φ ; ψ + ; List/Rᴸ≃AList/Rᴬᴸ) public +open MultiSets.Lists&ALists.L using (insert ; union ; count) +open MultiSets.Lists&ALists.AL using (insert ; union ; count) + + +-- 5.2 Structured Relations +open RelStructure using (StrRel) public +-- Definition (5.6) +open RelStructure using (SuitableStrRel) public +-- Theorem (5.7) +open RelStructure using (structuredQER→structuredEquiv) public +-- Definition (5.9) +open RelStructure using (StrRelAction) public +-- Lemma (5.10) +open RelStructure using (strRelQuotientComparison) public +-- Definition (5.11) +open RelStructure using (PositiveStrRel) public +-- Theorem (5.12) +open RelFunction using (functionSuitableRel) public +-- Multisets +-- (main is applying 5.7 to the example) +open MultiSets.Lists&ALists using (multisetShape ; isStructuredR ; main ; List/Rᴸ≡AList/Rᴬᴸ) + renaming ( hasAssociativeUnion to unionAssocAxiom + ; LQassoc to LUnionAssoc + ; ALQassoc to ALUnionAssoc) public diff --git a/Cubical/Papers/Synthetic.agda b/Cubical/Papers/Synthetic.agda new file mode 100644 index 000000000..5fe73c927 --- /dev/null +++ b/Cubical/Papers/Synthetic.agda @@ -0,0 +1,214 @@ +{-# OPTIONS --safe #-} +module Cubical.Papers.Synthetic where + +-- Cubical synthetic homotopy theory +-- Mörtberg and Pujet, „Cubical synthetic homotopy theory“. +-- https://dl.acm.org/doi/abs/10.1145/3372885.3373825 + +-- 2.1 +import Agda.Builtin.Cubical.Path as Path +import Cubical.Foundations.Prelude as Prelude +-- 2.2 +import Agda.Primitive.Cubical as PrimitiveCubical +import Cubical.Data.Bool as Bool +import Cubical.Core.Primitives as CorePrimitives +-- 2.3 +import Cubical.HITs.S1 as S1 +-- 2.4 +import Cubical.Foundations.Equiv as Equiv +import Cubical.Core.Glue as CoreGlue +import Cubical.Foundations.Univalence as Univalence +-- 3. +import Cubical.HITs.Torus as T2 +-- 3.1 +import Cubical.Data.Int as Int +import Cubical.Data.Int.Properties as IntProp +import Cubical.Foundations.Isomorphism as Isomorphism +-- 4.1 +import Cubical.HITs.Susp as Suspension +import Cubical.HITs.Sn as Sn +import Agda.Builtin.Nat as BNat +import Agda.Builtin.Bool as BBool +import Cubical.Foundations.GroupoidLaws as GroupoidLaws +import Cubical.HITs.S2 as S2 +import Cubical.HITs.S3 as S3 +-- 4.2 +import Cubical.HITs.Pushout as Push +import Cubical.HITs.Pushout.Properties as PushProp +-- 4.3 +import Cubical.HITs.Join as Join +import Cubical.HITs.Join.Properties as JoinProp +-- 5. +import Cubical.HITs.Hopf as Hopf + +-------------------------------------------------------------------------------- +-- 2. Cubical Agda +-- 2.1 The Interval and Path Types +-- 2.2 Transport and Composition +-- 2.3 Higher Inductive Types +-- 2.4 Glue Types and Univalence +-------------------------------------------------------------------------------- + +-- 2.1 The Interval and Path Types +open Path using (PathP) public +open Prelude using (_≡_ ; refl ; funExt) public +open Prelude renaming (sym to _⁻¹) public + +-- 2.2 Transport and Composition +open Prelude using (transport ; subst ; J ; JRefl) public +open PrimitiveCubical using (Partial) public +open Bool using (Bool ; true ; false) public + +partialBool : ∀ i → Partial (i ∨ ~ i) Bool +partialBool i = λ {(i = i1) → true + ; (i = i0) → false} + +open CorePrimitives using (inS ; outS ; hcomp) public +open Prelude using (_∙_) public + +-- 2.3 Higher Inductive Types +open S1 using (S¹ ; base ; loop) public + +double : S¹ → S¹ +double base = base +double (loop i) = (loop ∙ loop) i + +-- 2.4 Glue Types and Univalence +open Equiv using (idEquiv) public +open CoreGlue using (Glue) public +open Univalence using (ua) public + +-------------------------------------------------------------------------------- +-- 3. The Circle and Torus +-- 3.1 The Loop Spaces of the Circle and Torus +-------------------------------------------------------------------------------- + +open T2 using ( Torus ; point ; line1 ; line2 ; square + ; t2c ; c2t ; c2t-t2c ; t2c-c2t ; Torus≡S¹×S¹) + +-- 3.1 The Loop Spaces of the Circle and Torus +open S1 using (ΩS¹) public +open T2 using (ΩTorus) public +open Int using (pos ; negsuc) renaming (ℤ to Int) public +open IntProp using (sucPathℤ) public +open S1 using (helix ; winding) public + +-- Examples computing the winding numbers of the circle +_ : winding (loop ∙ loop ∙ loop) ≡ pos 3 +_ = refl + +_ : winding ((loop ⁻¹) ∙ loop ∙ (loop ⁻¹)) ≡ negsuc 0 +_ = refl + +open S1 renaming (intLoop to loopn) public +open S1 renaming (windingℤLoop to winding-loopn) public +open S1 using (encode ; decode ; decodeEncode ; ΩS¹≡ℤ) public +open Isomorphism using (isoToPath ; iso) public + +-- Notation of the paper, current notation under ΩS¹≡Int +ΩS¹≡Int' : ΩS¹ ≡ Int +ΩS¹≡Int' = isoToPath (iso winding loopn + winding-loopn (decodeEncode base)) + +open T2 using (ΩTorus≡ℤ×ℤ ; windingTorus) public + +-- Examples at the end of section 3. +_ : windingTorus (line1 ∙ line2) ≡ (pos 1 , pos 1) +_ = refl + +_ : windingTorus ((line1 ⁻¹) ∙ line2 ∙ line1) ≡ (pos 0 , pos 1) +_ = refl + +-------------------------------------------------------------------------------- +-- 4. Suspension, Spheres and Pushouts +-- 4.1 Suspension +-- 4.2 Pushouts and the 3 × 3 Lemma +-- 4.3 The Join and S³ +-------------------------------------------------------------------------------- + +-- 4.1 Suspension +open Suspension using (Susp ; north ; south ; merid) public +open Sn using (S₊) public +open Suspension using ( SuspBool→S¹ ; S¹→SuspBool + ; SuspBool→S¹→SuspBool + ; S¹→SuspBool→S¹) public + +-- Deprecated version of S₊ +open BNat renaming (Nat to ℕ) hiding (_*_) public +open CorePrimitives renaming (Type to Set) public +open BBool using (Bool) public +-- At the time the paper was published, Set was used instead of Type +_-sphere : ℕ → Set +0 -sphere = Bool +(suc n) -sphere = Susp (n -sphere) + +-- Lemma 4.1. The (1)-sphere is equal to the circle S1. + +open BBool using (true ; false) public +-- Deprecated version of SuspBool→S¹ +s2c : 1 -sphere → S¹ +s2c north = base +s2c south = base +s2c (merid true i) = loop i +s2c (merid false i) = base -- (loop ⁻¹) i + +-- Deprecated version of S¹→SuspBool +c2s : S¹ → 1 -sphere +c2s base = north +c2s (loop i) = (merid true ∙ (merid false ⁻¹)) i + +open GroupoidLaws using (rUnit) public +-- Deprecated version of SuspBool→S¹→SuspBool +s2c-c2s : ∀ (p : S¹) → s2c (c2s p) ≡ p +s2c-c2s base = refl +s2c-c2s (loop i) j = rUnit loop (~ j) i + +h1 : I → I → 1 -sphere +h1 i j = merid false (i ∧ j) + +h2 : I → I → 1 -sphere +h2 i j = hcomp (λ k → λ { (i = i0) → north + ; (i = i1) → merid false (j ∨ ~ k) + ; (j = i1) → merid true i }) + (merid true i) + +-- Deprecated version of S¹→SuspBool→S¹ +c2s-s2c : ∀ (t : 1 -sphere) → c2s (s2c t) ≡ t +c2s-s2c north j = north +c2s-s2c south j = merid false j +c2s-s2c (merid true i) j = h2 i j +c2s-s2c (merid false i) j = merid false (i ∧ j) + +-- Notation of the paper, current notation under S¹≡SuspBool +-- Proof of Lemma 4.1 +1-sphere≡S¹ : 1 -sphere ≡ S¹ +1-sphere≡S¹ = isoToPath (iso s2c c2s s2c-c2s c2s-s2c) + +-- Definitions of S2 and S3 +open S2 using (S²) public +open S3 using (S³) public + +-- 4.2 Pushouts and the 3 × 3 Lemma +open Push using (Pushout) public +-- 3x3-span is implemented as a record +open PushProp using (3x3-span) public +open 3x3-span using (f□1) public +-- The rest of the definitions inside the 3x3-lemma +-- A□0-A○□ , A□○-A○□ ... +open 3x3-span using (3x3-lemma) public + +-- 4.3 The Join and S³ +open Join renaming (join to Join) using (S³≡joinS¹S¹) public +open JoinProp using (join-assoc) public + +-------------------------------------------------------------------------------- +-- 5. The Hopf Fibration +-------------------------------------------------------------------------------- + +-- rot (denoted by _·_ here) in the paper is substituted by a rot and rotLoop in S1 +open S1 using (_·_ ; rotLoop) public +open Hopf renaming ( HopfSuspS¹ to Hopf + ; JoinS¹S¹→TotalHopf to j2h + ; TotalHopf→JoinS¹S¹ to h2j) + using (HopfS²) public +open S1 renaming (rotInv-1 to lem-rot-inv) public diff --git a/Cubical/Papers/ZCohomology.agda b/Cubical/Papers/ZCohomology.agda new file mode 100644 index 000000000..89c48ba36 --- /dev/null +++ b/Cubical/Papers/ZCohomology.agda @@ -0,0 +1,522 @@ +{- + +Please do not move this file. Changes should only be made if +necessary. + +This file contains pointers to the code examples and main results from +the paper: + +Synthetic Cohomology Theory in Cubical Agda + +-} + +-- The "--safe" flag ensures that there are no postulates or +-- unfinished goals +{-# OPTIONS --safe #-} +module Cubical.Papers.ZCohomology where + +-- Misc. +open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Nat +open import Cubical.Foundations.Everything +open import Cubical.HITs.S1 +open import Cubical.Data.Sum +open import Cubical.Data.Sigma + +-- 2 +open import Cubical.Core.Glue as Glue +import Cubical.Foundations.Prelude as Prelude +import Cubical.Foundations.GroupoidLaws as GroupoidLaws +import Cubical.Foundations.Path as Path +import Cubical.Foundations.Pointed as Pointed + renaming (Pointed to Type∙) +open import Cubical.HITs.S1 as S1 +open import Cubical.HITs.Susp as Suspension +open import Cubical.HITs.Sn as Sn +open import Cubical.Homotopy.Loopspace as Loop +open import Cubical.Foundations.HLevels as n-types +open import Cubical.HITs.Truncation as Trunc +open import Cubical.Homotopy.Connected as Connected +import Cubical.HITs.Pushout as Push +import Cubical.HITs.Wedge as ⋁ +import Cubical.Foundations.Univalence as Unival +import Cubical.Foundations.SIP as StructIdPrinc +import Cubical.Algebra.Group as Gr +import Cubical.Algebra.Group.GroupPath as GrPath + +-- 3 +import Cubical.ZCohomology.Base as coHom + renaming (coHomK to K ; coHomK-ptd to K∙) +import Cubical.HITs.Sn.Properties as S +import Cubical.ZCohomology.GroupStructure as GroupStructure +import Cubical.ZCohomology.Properties as Properties + renaming (Kn→ΩKn+1 to σ ; ΩKn+1→Kn to σ⁻¹) +import Cubical.Experiments.ZCohomologyOld.Properties as oldCohom + +-- 4 +import Cubical.ZCohomology.RingStructure.CupProduct as Cup +import Cubical.ZCohomology.RingStructure.RingLaws as ⌣Ring +import Cubical.ZCohomology.RingStructure.GradedCommutativity as ⌣Comm +import Cubical.Foundations.Pointed.Homogeneous as Homogen + +-- 5 +import Cubical.HITs.Torus as 𝕋² + renaming (Torus to 𝕋²) +import Cubical.HITs.KleinBottle as 𝕂² + renaming (KleinBottle to 𝕂²) +import Cubical.HITs.RPn as ℝP + renaming (RP² to ℝP²) +import Cubical.ZCohomology.Groups.Sn as HⁿSⁿ + renaming (Hⁿ-Sᵐ≅0 to Hⁿ-Sᵐ≅1) +import Cubical.ZCohomology.Groups.Torus as HⁿT² +import Cubical.ZCohomology.Groups.Wedge as Hⁿ-wedge +import Cubical.ZCohomology.Groups.KleinBottle as Hⁿ𝕂² +import Cubical.ZCohomology.Groups.RP2 as HⁿℝP² + renaming (H¹-RP²≅0 to H¹-RP²≅1) +import Cubical.ZCohomology.Groups.CP2 as HⁿℂP² + renaming (CP² to ℂP² ; ℤ→HⁿCP²→ℤ to g) + +-- Appendix +import Cubical.Homotopy.EilenbergSteenrod as ES-axioms +import Cubical.ZCohomology.EilenbergSteenrodZ as satisfies-ES-axioms + renaming (coHomFunctor to H^~ ; coHomFunctor' to Ĥ) +import Cubical.ZCohomology.MayerVietorisUnreduced as MayerVietoris + +----- 2. HOMOTOPY TYPE THEORY IN CUBICAL AGDA ----- + +-- 2.1 Important notions in Cubical Agda +open Prelude using ( PathP + ; _≡_ + ; refl + ; cong + ; cong₂ + ; funExt) + +open GroupoidLaws using (_⁻¹) + +open Prelude using ( transport + ; subst + ; hcomp) + +--- 2.2 Important concepts from HoTT/UF in Cubical Agda + +-- Pointed Types +open Pointed using (Type∙) + +-- The circle, 𝕊¹ +open S1 using (S¹) + +-- Suspensions +open Suspension using (Susp) + +-- (Pointed) n-spheres, 𝕊ⁿ +open Sn using (S₊∙) + +-- Loop spaces +open Loop using (Ω^_) + +-- Eckmann-Hilton argument +open Loop using (Eckmann-Hilton) + +-- n-types Note that we start indexing from 0 in the Cubical Library +-- (so (-2)-types as referred to as 0-types, (-1) as 1-types, and so +-- on) +open n-types using (isOfHLevel) + +-- truncations +open Trunc using (hLevelTrunc) + +-- elimination principle +open Trunc using (elim) + +-- elimination principle for paths +truncPathElim : ∀ {ℓ ℓ'} {A : Type ℓ} {x y : A} (n : ℕ) + → {B : Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ y ∣ → Type ℓ'} + → ((q : _) → isOfHLevel n (B q)) + → ((p : x ≡ y) → B (cong ∣_∣ p)) + → (q : _) → B q +truncPathElim zero hlev ind q = hlev q .fst +truncPathElim (suc n) {B = B} hlev ind q = + subst B (Iso.leftInv (Trunc.PathIdTruncIso _) q) + (help (ΩTrunc.encode-fun ∣ _ ∣ ∣ _ ∣ q)) + where + help : (q : _) → B (ΩTrunc.decode-fun ∣ _ ∣ ∣ _ ∣ q) + help = Trunc.elim (λ _ → hlev _) ind + +-- Connectedness +open Connected using (isConnected) + +-- Pushouts +open Push using (Pushout) + +-- Wedge sum +open ⋁ using (_⋁_) + + +-- 2.3 Univalence + +-- Univalence and the ua function respectively +open Unival using (univalence ; ua) + +-- Glue types +open Glue using (Glue) + +-- The structure identity principle and the sip function +-- respectively +open StructIdPrinc using (SIP ; sip) + +-- Groups +open Gr using (Group) + +-- Isomorphic groups are path equal +open GrPath using (GroupPath) + + +----- 3. ℤ-COHOMOLOGY IN CUBICAL AGDA ----- + + +-- 3.1 Eilenberg-MacLane spaces + +-- Eilenberg-MacLane spaces Kₙ (unpointed and pointed respectively) +open coHom using (K ; K∙) + +-- Proposition 7 +open S using (sphereConnected) + +-- Lemma 8 +open S using (wedgeconFun; wedgeconLeft ; wedgeconRight) + +-- restated to match the formulation in the paper +wedgeConSn' : ∀ {ℓ} (n m : ℕ) {A : (S₊ (suc n)) → (S₊ (suc m)) → Type ℓ} + → ((x : S₊ (suc n)) (y : S₊ (suc m)) → isOfHLevel ((suc n) + (suc m)) (A x y)) + → (fₗ : (x : _) → A x (ptSn (suc m))) + → (fᵣ : (x : _) → A (ptSn (suc n)) x) + → (p : fₗ (ptSn (suc n)) ≡ fᵣ (ptSn (suc m))) + → Σ[ F ∈ ((x : S₊ (suc n)) (y : S₊ (suc m)) → A x y) ] + (Σ[ (left , right) ∈ ((x : S₊ (suc n)) → fₗ x ≡ F x (ptSn (suc m))) + × ((x : S₊ (suc m)) → fᵣ x ≡ F (ptSn (suc n)) x) ] + p ≡ left (ptSn (suc n)) ∙ (right (ptSn (suc m))) ⁻¹) +wedgeConSn' zero zero hlev fₗ fᵣ p = + (wedgeconFun 0 0 hlev fᵣ fₗ p) + , ((λ x → sym (wedgeconRight 0 0 hlev fᵣ fₗ p x)) + , λ _ → refl) -- right holds by refl + , rUnit _ +wedgeConSn' zero (suc m) hlev fₗ fᵣ p = + (wedgeconFun 0 (suc m) hlev fᵣ fₗ p) + , ((λ _ → refl) -- left holds by refl + , (λ x → sym (wedgeconLeft 0 (suc m) hlev fᵣ fₗ p x))) + , lUnit _ +wedgeConSn' (suc n) m hlev fₗ fᵣ p = + (wedgeconFun (suc n) m hlev fᵣ fₗ p) + , ((λ x → sym (wedgeconRight (suc n) m hlev fᵣ fₗ p x)) + , λ _ → refl) -- right holds by refl + , rUnit _ + +-- +ₖ (addition) and 0ₖ +open GroupStructure using (_+ₖ_ ; 0ₖ) + +-- The function σ : Kₙ → ΩKₙ₊₁ +open Properties using (σ) + +-- -ₖ (subtraction) +open GroupStructure using (-ₖ_) + +-- Group laws for +ₖ +open GroupStructure using ( rUnitₖ ; lUnitₖ + ; rCancelₖ ; lCancelₖ + ; commₖ + ; assocₖ) + +-- All group laws are equal to refl at 0ₖ +-- rUnitₖ (definitional) +0-rUnit≡refl : rUnitₖ 0 (0ₖ 0) ≡ refl +1-rUnit≡refl : rUnitₖ 1 (0ₖ 1) ≡ refl +0-rUnit≡refl = refl +1-rUnit≡refl = refl +n≥2-rUnit≡refl : {n : ℕ} → rUnitₖ (2 + n) (0ₖ (2 + n)) ≡ refl +n≥2-rUnit≡refl = refl + +-- rUnitₖ (definitional) +0-lUnit≡refl : lUnitₖ 0 (0ₖ 0) ≡ refl +1-lUnit≡refl : lUnitₖ 1 (0ₖ 1) ≡ refl +n≥2-lUnit≡refl : {n : ℕ} → lUnitₖ (2 + n) (0ₖ (2 + n)) ≡ refl +0-lUnit≡refl = refl +1-lUnit≡refl = refl +n≥2-lUnit≡refl = refl + +-- assocₖ (definitional) +0-assoc≡refl : assocₖ 0 (0ₖ 0) (0ₖ 0) (0ₖ 0) ≡ refl +1-assoc≡refl : assocₖ 1 (0ₖ 1) (0ₖ 1) (0ₖ 1) ≡ refl +n≥2-assoc≡refl : {n : ℕ} → assocₖ (2 + n) (0ₖ (2 + n)) (0ₖ (2 + n)) (0ₖ (2 + n)) ≡ refl +0-assoc≡refl = refl +1-assoc≡refl = refl +n≥2-assoc≡refl = refl + +-- commₖ (≡ refl ∙ refl for n ≥ 2) +0-comm≡refl : commₖ 0 (0ₖ 0) (0ₖ 0) ≡ refl +1-comm≡refl : commₖ 1 (0ₖ 1) (0ₖ 1) ≡ refl +n≥2-comm≡refl : {n : ℕ} → commₖ (2 + n) (0ₖ (2 + n)) (0ₖ (2 + n)) ≡ refl +0-comm≡refl = refl +1-comm≡refl = refl +n≥2-comm≡refl = sym (rUnit refl) + +-- lCancelₖ (definitional) +0-lCancel≡refl : lCancelₖ 0 (0ₖ 0) ≡ refl +1-lCancel≡refl : lCancelₖ 1 (0ₖ 1) ≡ refl +n≥2-lCancel≡refl : {n : ℕ} → lCancelₖ (2 + n) (0ₖ (2 + n)) ≡ refl +0-lCancel≡refl = refl +1-lCancel≡refl = refl +n≥2-lCancel≡refl = refl + +-- rCancelₖ (≡ (refl ∙ refl) ∙ refl for n ≥ 1) +0-rCancel≡refl : rCancelₖ 0 (0ₖ 0) ≡ refl +1-rCancel≡refl : rCancelₖ 1 (0ₖ 1) ≡ refl +n≥2-rCancel≡refl : {n : ℕ} → rCancelₖ (2 + n) (0ₖ (2 + n)) ≡ refl +0-rCancel≡refl = refl +1-rCancel≡refl = refl +n≥2-rCancel≡refl i = rUnit (rUnit refl (~ i)) (~ i) + +-- Proof that there is a unique h-structure on Kₙ +-- +ₖ defines an h-Structure on Kₙ +open GroupStructure using (_+ₖ_ ; 0ₖ ; rUnitₖ ; lUnitₖ ; lUnitₖ≡rUnitₖ) + +-- and so does Brunerie's addition +open oldCohom using (_+ₖ_ ; 0ₖ ; rUnitₖ ; lUnitₖ ; rUnitlUnit0) + +-- consequently both additions agree +open GroupStructure using (+ₖ-unique) +open oldCohom using (addLemma) +additionsAgree : (n : ℕ) → GroupStructure._+ₖ_ {n = n} ≡ oldCohom._+ₖ_ {n = n} +additionsAgree zero i x y = oldCohom.addLemma x y (~ i) +additionsAgree (suc n) i x y = + +ₖ-unique n (GroupStructure._+ₖ_) (oldCohom._+ₖ_) + (GroupStructure.rUnitₖ (suc n)) (GroupStructure.lUnitₖ (suc n)) + (oldCohom.rUnitₖ (suc n)) (oldCohom.lUnitₖ (suc n)) + (sym (lUnitₖ≡rUnitₖ (suc n))) + (rUnitlUnit0 (suc n)) x y i + +-- Theorem 9 (Kₙ ≃ ΩKₙ₊₁) +open Properties using (Kn≃ΩKn+1) + +-- σ and σ⁻¹ are morphisms +-- (for σ⁻¹ this is proved directly without using the fact that σ is a morphism) +open Properties using (Kn→ΩKn+1-hom ; ΩKn+1→Kn-hom) + +-- Lemma 10 (p ∙ q ≡ cong²₊(p,q)) for n = 1 and n ≥ 2 respectively +open GroupStructure using (∙≡+₁ ; ∙≡+₂) + +-- Lemma 11 (cong²₊ is commutative) and Theorem 12 respectively +open GroupStructure using (cong+ₖ-comm ; isCommΩK) + +-- 3.2 Group structure on Hⁿ(A) + +-- +ₕ (addition), -ₕ and 0ₕ +open GroupStructure using (_+ₕ_ ; -ₕ_ ; 0ₕ) + +-- Cohomology group structure +open GroupStructure using ( rUnitₕ ; lUnitₕ + ; rCancelₕ ; lCancelₕ + ; commₕ + ; assocₕ) + +-------------------------------------------------------------------- MOVE? +-- Reduced cohomology, group structure +open GroupStructure using (coHomRedGroupDir) + +-- Equality of unreduced and reduced cohmology +open Properties using (coHomGroup≡coHomRedGroup) +-------------------------------------------------------------------- + +----- 4. The Cup Product and Cohomology Ring ----- +-- 4.1 +-- Lemma 13 +open Properties using (isOfHLevel↑∙) + +-- ⌣ₖ +open Cup using (_⌣ₖ_) + +-- ⌣ₖ is pointed in both arguments +open ⌣Ring using (0ₖ-⌣ₖ ; ⌣ₖ-0ₖ) + +-- The cup product +open Cup using (_⌣_) + +-- 4.2 +-- Lemma 14 +Lem14 : (n m : ℕ) (f g : K∙ n →∙ K∙ m) → fst f ≡ fst g → f ≡ g +Lem14 n m f g p = Homogen.→∙Homogeneous≡ (Properties.isHomogeneousKn m) p + +-- Proposition 15 +open ⌣Ring using (leftDistr-⌣ₖ ; rightDistr-⌣ₖ) + +-- Lemma 16 +open ⌣Ring using (assocer-helpFun≡) + +-- Proposition 17 +open ⌣Ring using (assoc-⌣ₖ) + +-- Proposition 18 +open ⌣Comm using (gradedComm-⌣ₖ) + +-- Ring structure on ⌣ +open ⌣Ring using (leftDistr-⌣ ; rightDistr-⌣ + ; assoc-⌣ ; 1⌣ + ; rUnit⌣ ; lUnit⌣ + ; ⌣0 ; 0⌣) +open ⌣Comm using (gradedComm-⌣) + +----- 5. CHARACTERIZING ℤ-COHOMOLOGY GROUPS ----- + +-- 5.1 +-- Proposition 19 +open HⁿSⁿ using (Hⁿ-Sⁿ≅ℤ) + +-- 5.2 +-- The torus +open 𝕋² using (𝕋²) + +-- Propositions 20 and 21 respectively +open HⁿT² using (H¹-T²≅ℤ×ℤ ; H²-T²≅ℤ) + +-- 5.3 +-- The Klein bottle +open 𝕂² using (𝕂²) + +-- The real projective plane +open ℝP using (ℝP²) + +-- Proposition 22 and 23 respectively +-- ℤ/2ℤ is represented by Bool with the unique group structure +-- Lemma 23 is used implicitly in H²-𝕂²≅Bool +open Hⁿ𝕂² using (H¹-𝕂²≅ℤ ; H²-𝕂²≅Bool) + +-- First and second cohomology groups of ℝP² respectively +open HⁿℝP² using (H¹-RP²≅1 ; H²-RP²≅Bool) + +-- 5.4 +-- The complex projective plane +open HⁿℂP² using (ℂP²) + +-- Second and fourth cohomology groups ℂP² respectively +open HⁿℂP² using (H²CP²≅ℤ ; H⁴CP²≅ℤ) + +-- 6 Proving by computations in Cubical Agda +-- Proof of m = n = 1 case of graded commutativity (post truncation elimination): +-- Uncomment and give it a minute. The proof is currently not running very fast. + +{- +open ⌣Comm using (-ₖ^_·_ ) +n=m=1 : (a b : S¹) + → _⌣ₖ_ {n = 1} {m = 1} ∣ a ∣ ∣ b ∣ + ≡ (-ₖ (_⌣ₖ_ {n = 1} {m = 1} ∣ b ∣ ∣ a ∣)) +n=m=1 base base = refl +n=m=1 base (loop i) k = -ₖ (Properties.Kn→ΩKn+10ₖ _ (~ k) i) +n=m=1 (loop i) base k = Properties.Kn→ΩKn+10ₖ _ k i +n=m=1 (loop i) (loop j) k = -- This hcomp is just a simple rewriting to get paths in Ω²K₂ + hcomp (λ r → λ { (i = i0) → -ₖ Properties.Kn→ΩKn+10ₖ _ (~ k ∨ ~ r) j + ; (i = i1) → -ₖ Properties.Kn→ΩKn+10ₖ _ (~ k ∨ ~ r) j + ; (j = i0) → Properties.Kn→ΩKn+10ₖ _ (k ∨ ~ r) i + ; (j = i1) → Properties.Kn→ΩKn+10ₖ _ (k ∨ ~ r) i + ; (k = i0) → + doubleCompPath-filler + (sym (Properties.Kn→ΩKn+10ₖ _)) + (λ j i → _⌣ₖ_ {n = 1} {m = 1} ∣ loop i ∣ ∣ loop j ∣) + (Properties.Kn→ΩKn+10ₖ _) + (~ r) j i + ; (k = i1) → + -ₖ doubleCompPath-filler + (sym (Properties.Kn→ΩKn+10ₖ _)) + (λ j i → _⌣ₖ_ {n = 1} {m = 1} ∣ loop i ∣ ∣ loop j ∣) + (Properties.Kn→ΩKn+10ₖ _) + (~ r) i j}) + ((main + ∙ sym (cong-∙∙ (cong (-ₖ_)) (sym (Properties.Kn→ΩKn+10ₖ _)) + (λ j i → (_⌣ₖ_ {n = 1} {m = 1} ∣ loop i ∣ ∣ loop j ∣)) + (Properties.Kn→ΩKn+10ₖ _))) k i j) + where + open import Cubical.Foundations.Equiv.HalfAdjoint + t : Iso (typ ((Ω^ 2) (K∙ 2))) ℤ + t = compIso (congIso (invIso (Properties.Iso-Kn-ΩKn+1 1))) + (invIso (Properties.Iso-Kn-ΩKn+1 0)) + + p₁ = flipSquare ((sym (Properties.Kn→ΩKn+10ₖ _)) + ∙∙ (λ j i → _⌣ₖ_ {n = 1} {m = 1} ∣ loop i ∣ ∣ loop j ∣) + ∙∙ (Properties.Kn→ΩKn+10ₖ _)) + p₂ = (cong (cong (-ₖ_)) + ((sym (Properties.Kn→ΩKn+10ₖ _)))) + ∙∙ (λ j i → -ₖ (_⌣ₖ_ {n = 1} {m = 1} ∣ loop i ∣ ∣ loop j ∣)) + ∙∙ (cong (cong (-ₖ_)) (Properties.Kn→ΩKn+10ₖ _)) + + computation : Iso.fun t p₁ ≡ Iso.fun t p₂ + computation = refl + + main : p₁ ≡ p₂ + main = p₁ ≡⟨ sym (Iso.leftInv t p₁) ⟩ + (Iso.inv t (Iso.fun t p₁)) ≡⟨ cong (Iso.inv t) computation ⟩ + Iso.inv t (Iso.fun t p₂) ≡⟨ Iso.leftInv t p₂ ⟩ + p₂ ∎ +-} + +-- 𝕋² ≠ S² ∨ S¹ ∨ S¹ +open HⁿT² using (T²≠S²⋁S¹⋁S¹) + +-- Second "Brunerie number" +open HⁿℂP² using (g) +brunerie2 : ℤ +brunerie2 = g 1 + +----- A. Proofs ----- + +-- A.2 Proofs for Section 4 + +-- Lemma 27 +open Homogen using (→∙Homogeneous≡) + +-- Lemma 28 +open Homogen using (isHomogeneous→∙) + +-- Lemma 29 +open Properties using (isHomogeneousKn) + +-- Lemma 30, parts 1-3 respectively +open Path using (sym≡flipSquare ; sym-cong-sym≡id ; sym≡cong-sym) + +-- Lemma 31 +open ⌣Comm using (cong-ₖ-gen-inr) + + +-- A.3 Proofs for Section 5 + +-- Proposition 32 +open HⁿSⁿ using (Hⁿ-Sᵐ≅1) + +----- B. THE EILENBERG-STEENROD AXIOMS ----- + +-- B.1 The axioms in HoTT/UF + +-- The axioms are defined as a record type +open ES-axioms.coHomTheory + +-- Proof of the claim that the alternative reduced cohomology functor +-- Ĥ is the same as the usual reduced cohomology functor +_ : ∀ {ℓ} → satisfies-ES-axioms.H^~ {ℓ} ≡ satisfies-ES-axioms.Ĥ +_ = satisfies-ES-axioms.coHomFunctor≡coHomFunctor' + +-- B.2 Verifying the axioms + +-- Propositions 35 and 36. +_ : ∀ {ℓ} → ES-axioms.coHomTheory {ℓ} satisfies-ES-axioms.H^~ +_ = satisfies-ES-axioms.isCohomTheoryZ + + +-- B.3 Characterizing Z-cohomology groups using the axioms + +-- Theorem 37 +open MayerVietoris.MV using ( Ker-i⊂Im-d ; Im-d⊂Ker-i + ; Ker-Δ⊂Im-i ; Im-i⊂Ker-Δ + ; Ker-d⊂Im-Δ ; Im-Δ⊂Ker-d) + + +----- C. BENCHMARKING COMPUTATIONS WITH THE COHOMOLOGY GROUPS ----- + +import Cubical.Experiments.ZCohomology.Benchmarks diff --git a/Cubical/README.agda b/Cubical/README.agda index 4f45ff381..d8e97aea8 100644 --- a/Cubical/README.agda +++ b/Cubical/README.agda @@ -1,4 +1,5 @@ -{-# OPTIONS --cubical #-} +{-# OPTIONS --guardedness #-} + module Cubical.README where ------------------------------------------------------------------------ @@ -21,17 +22,49 @@ import Cubical.Core.Everything import Cubical.Foundations.Prelude import Cubical.Foundations.Everything +-- Kinds and properties of functions +import Cubical.Functions.Everything + -- Data types and properties import Cubical.Data.Everything --- Properties and proofs about relations -import Cubical.Relation.Everything - -- Higher-inductive types import Cubical.HITs.Everything -- Coinductive data types and properties import Cubical.Codata.Everything +-- Papers +import Cubical.Papers.Everything + +-- Properties and proofs about relations +import Cubical.Relation.Everything + +-- Category theory +import Cubical.Categories.Everything + +-- Homotopy theory +import Cubical.Homotopy.Everything + +-- Properties and kinds of Modalities +import Cubical.Modalities.Everything + -- Various experiments using Cubical Agda import Cubical.Experiments.Everything + +-- Other modules (TODO: add descriptions) +import Cubical.Induction.Everything +import Cubical.Structures.Everything +import Cubical.ZCohomology.Everything + +-- Algebra library (in development) +import Cubical.Algebra.Everything + +-- Various talks +import Cubical.Talks.Everything + +-- Reflection +import Cubical.Reflection.Everything + +-- Displayed univalent graphs +import Cubical.Displayed.Everything diff --git a/Cubical/Reflection/Base.agda b/Cubical/Reflection/Base.agda new file mode 100644 index 000000000..05896ca57 --- /dev/null +++ b/Cubical/Reflection/Base.agda @@ -0,0 +1,59 @@ +{- + +Some basic utilities for reflection + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Reflection.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.List.Base +open import Cubical.Data.Nat.Base + +import Agda.Builtin.Reflection as R +open import Agda.Builtin.String + +_>>=_ = R.bindTC +_<|>_ = R.catchTC + +_$_ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (A → B) → A → B +f $ a = f a + +_>>_ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → R.TC A → R.TC B → R.TC B +f >> g = f >>= λ _ → g + +infixl 4 _>>=_ _>>_ _<|>_ +infixr 3 _$_ + +liftTC : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (A → B) → R.TC A → R.TC B +liftTC f ta = ta >>= λ a → R.returnTC (f a) + +v : ℕ → R.Term +v n = R.var n [] + +pattern varg t = R.arg (R.arg-info R.visible (R.modality R.relevant R.quantity-ω)) t +pattern harg t = R.arg (R.arg-info R.hidden (R.modality R.relevant R.quantity-ω)) t +pattern _v∷_ a l = varg a ∷ l +pattern _h∷_ a l = harg a ∷ l + +infixr 5 _v∷_ _h∷_ + +vlam : String → R.Term → R.Term +vlam str t = R.lam R.visible (R.abs str t) + +hlam : String → R.Term → R.Term +hlam str t = R.lam R.hidden (R.abs str t) + +newMeta = R.checkType R.unknown + +extend*Context : ∀ {ℓ} {A : Type ℓ} → List (R.Arg R.Type) → R.TC A → R.TC A +extend*Context [] tac = tac +extend*Context (a ∷ as) tac = R.extendContext a (extend*Context as tac) + +makeAuxiliaryDef : String → R.Type → R.Term → R.TC R.Term +makeAuxiliaryDef s ty term = + R.freshName s >>= λ name → + R.declareDef (varg name) ty >> + R.defineFun name [ R.clause [] [] term ] >> + R.returnTC (R.def name []) + diff --git a/Cubical/Reflection/Everything.agda b/Cubical/Reflection/Everything.agda new file mode 100644 index 000000000..143dd2abe --- /dev/null +++ b/Cubical/Reflection/Everything.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe --guardedness #-} +module Cubical.Reflection.Everything where + +import Cubical.Reflection.Base +import Cubical.Reflection.RecordEquiv +import Cubical.Reflection.StrictEquiv diff --git a/Cubical/Reflection/RecordEquiv.agda b/Cubical/Reflection/RecordEquiv.agda new file mode 100644 index 000000000..9e76f9d28 --- /dev/null +++ b/Cubical/Reflection/RecordEquiv.agda @@ -0,0 +1,196 @@ +{- + + Reflection-based tools for converting between iterated record types, particularly between + record types and iterated Σ-types. + + See end of file for examples. + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Reflection.RecordEquiv where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Data.List as List +open import Cubical.Data.Nat +open import Cubical.Data.Maybe as Maybe +open import Cubical.Data.Sigma + +open import Agda.Builtin.String +import Agda.Builtin.Reflection as R +open import Cubical.Reflection.Base + +Projections = Maybe (List R.Name) + +-- Describes a correspondence between two iterated record types +RecordAssoc = List (Projections × Projections) + +-- Describes a correspondence between a record type and an iterated Σ-types; +-- more convenient than RecordAssoc for this special case +data ΣFormat : Type where + leaf : R.Name → ΣFormat + _,_ : ΣFormat → ΣFormat → ΣFormat + unit : ΣFormat + +infixr 4 _,_ + +flipRecordAssoc : RecordAssoc → RecordAssoc +flipRecordAssoc = List.map λ {p .fst → p .snd; p .snd → p .fst} + +fstIdRecordAssoc : RecordAssoc → RecordAssoc +fstIdRecordAssoc = List.map λ {p .fst → p .fst; p .snd → p .fst} + +List→ΣFormat : List R.Name → ΣFormat +List→ΣFormat [] = unit +List→ΣFormat (x ∷ []) = leaf x +List→ΣFormat (x ∷ y ∷ xs) = leaf x , List→ΣFormat (y ∷ xs) + +ΣFormat→RecordAssoc : ΣFormat → RecordAssoc +ΣFormat→RecordAssoc = go [] + where + go : List R.Name → ΣFormat → RecordAssoc + go prefix unit = [ just prefix , nothing ] + go prefix (leaf fieldName) = [ just prefix , just [ fieldName ] ] + go prefix (sig₁ , sig₂) = + go (quote fst ∷ prefix) sig₁ ++ go (quote snd ∷ prefix) sig₂ + +-- Derive the shape of the compound Σ-type +ΣFormat→Ty : ΣFormat → R.Type +ΣFormat→Ty unit = R.def (quote Unit) [] +ΣFormat→Ty (leaf _) = R.unknown +ΣFormat→Ty (sig₁ , sig₂) = + R.def (quote Σ) (ΣFormat→Ty sig₁ v∷ R.lam R.visible (R.abs "_" (ΣFormat→Ty sig₂)) v∷ []) + +recordName→isoTy : R.Name → R.Term → R.TC R.Term +recordName→isoTy name σShape = + R.inferType (R.def name []) >>= R.normalise >>= go [] + where + go : List R.ArgInfo → R.Type → R.TC R.Term + go acc (R.pi (R.arg i argTy) (R.abs s ty)) = + liftTC (λ t → R.pi (R.arg i' argTy) (R.abs s t)) (go (i ∷ acc) ty) + where + i' = R.arg-info R.hidden (R.modality R.relevant R.quantity-ω) + go acc (R.agda-sort _) = + R.returnTC (R.def (quote Iso) (R.def name (makeArgs 0 [] acc) v∷ σShape v∷ [])) + where + makeArgs : ℕ → List (R.Arg R.Term) → List R.ArgInfo → List (R.Arg R.Term) + makeArgs n acc [] = acc + makeArgs n acc (i ∷ infos) = makeArgs (suc n) (R.arg i (v n) ∷ acc) infos + go _ _ = R.typeError (R.strErr "Not a record type name: " ∷ R.nameErr name ∷ []) + +projNames→Patterns : List R.Name → List (R.Arg R.Pattern) +projNames→Patterns = go [] + where + go : List (R.Arg R.Pattern) → List R.Name → List (R.Arg R.Pattern) + go acc [] = acc + go acc (π ∷ projs) = go (varg (R.proj π) ∷ acc) projs + +projNames→Term : R.Term → List R.Name → R.Term +projNames→Term term [] = term +projNames→Term term (π ∷ projs) = R.def π [ varg (projNames→Term term projs) ] + +convertClauses : RecordAssoc → R.Term → List R.Clause +convertClauses al term = fixIfEmpty (List.filterMap makeClause al) + where + makeClause : Projections × Projections → Maybe R.Clause + makeClause (projl , just projr) = + just (R.clause [] (goPat [] projr) (Maybe.rec R.unknown goTm projl)) + where + goPat : List (R.Arg R.Pattern) → List R.Name → List (R.Arg R.Pattern) + goPat acc [] = acc + goPat acc (π ∷ projs) = goPat (varg (R.proj π) ∷ acc) projs + + goTm : List R.Name → R.Term + goTm [] = term + goTm (π ∷ projs) = R.def π [ varg (goTm projs) ] + makeClause (_ , nothing) = nothing + + fixIfEmpty : List R.Clause → List R.Clause + fixIfEmpty [] = [ R.clause [] [] R.unknown ] + fixIfEmpty (c ∷ cs) = c ∷ cs + +mapClause : + (List (String × R.Arg R.Type) → List (String × R.Arg R.Type)) + → (List (R.Arg R.Pattern) → List (R.Arg R.Pattern)) + → (R.Clause → R.Clause) +mapClause f g (R.clause tel ps t) = R.clause (f tel) (g ps) t +mapClause f g (R.absurd-clause tel ps) = R.absurd-clause (f tel) (g ps) + +recordIsoΣClauses : ΣFormat → List R.Clause +recordIsoΣClauses σ = + funClauses (quote Iso.fun) Σ↔R ++ + funClauses (quote Iso.inv) R↔Σ ++ + pathClauses (quote Iso.rightInv) R↔Σ ++ + pathClauses (quote Iso.leftInv) Σ↔R + where + R↔Σ = ΣFormat→RecordAssoc σ + Σ↔R = flipRecordAssoc R↔Σ + + funClauses : R.Name → RecordAssoc → List R.Clause + funClauses name al = + List.map + (mapClause + (("_" , varg R.unknown) ∷_) + (λ ps → R.proj name v∷ R.var 0 v∷ ps)) + (convertClauses al (v 0)) + + pathClauses : R.Name → RecordAssoc → List R.Clause + pathClauses name al = + List.map + (mapClause + (λ vs → ("_" , varg R.unknown) ∷ ("_" , varg R.unknown) ∷ vs) + (λ ps → R.proj name v∷ R.var 1 v∷ R.var 0 v∷ ps)) + (convertClauses (fstIdRecordAssoc al) (v 1)) + +recordIsoΣTerm : ΣFormat → R.Term +recordIsoΣTerm σ = R.pat-lam (recordIsoΣClauses σ) [] + +-- with a provided ΣFormat for the record +declareRecordIsoΣ' : R.Name → ΣFormat → R.Name → R.TC Unit +declareRecordIsoΣ' idName σ recordName = + let σTy = ΣFormat→Ty σ in + recordName→isoTy recordName σTy >>= λ isoTy → + R.declareDef (varg idName) isoTy >> + R.defineFun idName (recordIsoΣClauses σ) + +-- using the right-associated Σ given by the record fields +declareRecordIsoΣ : R.Name → R.Name → R.TC Unit +declareRecordIsoΣ idName recordName = + R.getDefinition recordName >>= λ where + (R.record-type _ fs) → + let σ = List→ΣFormat (List.map (λ {(R.arg _ n) → n}) fs) in + declareRecordIsoΣ' idName σ recordName + _ → + R.typeError (R.strErr "Not a record type name:" ∷ R.nameErr recordName ∷ []) + +private + module Example where + variable + ℓ ℓ' : Level + A : Type ℓ + B : A → Type ℓ' + + record Example0 {A : Type ℓ} (B : A → Type ℓ') : Type (ℓ-max ℓ ℓ') where + no-eta-equality -- works with or without eta equality + field + cool : A + fun : A + wow : B cool + + -- Declares a function `Example0IsoΣ` that gives an isomorphism between the record type and a + -- right-associated nested Σ-type (with the parameters to Example0 as implict arguments). + unquoteDecl Example0IsoΣ = declareRecordIsoΣ Example0IsoΣ (quote Example0) + + -- `Example0IsoΣ` has the type we expect + test0 : Iso (Example0 B) (Σ[ a ∈ A ] (Σ[ _ ∈ A ] B a)) + test0 = Example0IsoΣ + + -- A record with no fields is isomorphic to Unit + + record Example1 : Type where + + unquoteDecl Example1IsoΣ = declareRecordIsoΣ Example1IsoΣ (quote Example1) + + test1 : Iso Example1 Unit + test1 = Example1IsoΣ diff --git a/Cubical/Reflection/StrictEquiv.agda b/Cubical/Reflection/StrictEquiv.agda new file mode 100644 index 000000000..63a819696 --- /dev/null +++ b/Cubical/Reflection/StrictEquiv.agda @@ -0,0 +1,81 @@ +{-# OPTIONS --safe #-} +module Cubical.Reflection.StrictEquiv where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv.Base +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.List.Base +open import Cubical.Data.Unit.Base + +import Agda.Builtin.Reflection as R +open import Cubical.Reflection.Base + +strictEquivClauses : R.Term → R.Term → List R.Clause +strictEquivClauses f g = + R.clause [] + (R.proj (quote fst) v∷ []) + f + ∷ R.clause [] + (R.proj (quote snd) v∷ R.proj (quote equiv-proof) v∷ []) + (R.def (quote strictContrFibers) (g v∷ [])) + ∷ [] + +strictEquivTerm : R.Term → R.Term → R.Term +strictEquivTerm f g = R.pat-lam (strictEquivClauses f g) [] + +strictEquivMacro : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → (A → B) → (B → A) → R.Term → R.TC Unit +strictEquivMacro {A = A} {B} f g hole = + R.quoteTC (A ≃ B) >>= λ equivTy → + R.checkType hole equivTy >> + R.quoteTC f >>= λ `f` → + R.quoteTC g >>= λ `g` → + R.unify (strictEquivTerm `f` `g`) hole + +strictIsoToEquivMacro : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → Iso A B → R.Term → R.TC Unit +strictIsoToEquivMacro isom = + strictEquivMacro (Iso.fun isom) (Iso.inv isom) + +-- For use with unquoteDef + +defStrictEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → R.Name → (A → B) → (B → A) → R.TC Unit +defStrictEquiv idName f g = + R.quoteTC f >>= λ `f` → + R.quoteTC g >>= λ `g` → + R.defineFun idName (strictEquivClauses `f` `g`) + +defStrictIsoToEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → R.Name → Iso A B → R.TC Unit +defStrictIsoToEquiv idName isom = + defStrictEquiv idName (Iso.fun isom) (Iso.inv isom) + +-- For use with unquoteDef + +declStrictEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → R.Name → (A → B) → (B → A) → R.TC Unit +declStrictEquiv {A = A} {B = B} idName f g = + R.quoteTC (A ≃ B) >>= λ ty → + R.declareDef (varg idName) ty >> + defStrictEquiv idName f g + +declStrictIsoToEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → R.Name → Iso A B → R.TC Unit +declStrictIsoToEquiv idName isom = + declStrictEquiv idName (Iso.fun isom) (Iso.inv isom) + +macro + -- (f : A → B) → (g : B → A) → (A ≃ B) + -- Assumes that `f` and `g` are inverse up to definitional equality + strictEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → (A → B) → (B → A) → R.Term → R.TC Unit + strictEquiv = strictEquivMacro + + -- (isom : Iso A B) → (A ≃ B) + -- Assumes that the functions defining `isom` are inverse up to definitional equality + strictIsoToEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → Iso A B → R.Term → R.TC Unit + strictIsoToEquiv = strictIsoToEquivMacro diff --git a/Cubical/Relation/Binary.agda b/Cubical/Relation/Binary.agda index a12eef9d7..ec2f3e272 100644 --- a/Cubical/Relation/Binary.agda +++ b/Cubical/Relation/Binary.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Relation.Binary where open import Cubical.Relation.Binary.Base public diff --git a/Cubical/Relation/Binary/Base.agda b/Cubical/Relation/Binary/Base.agda index 9aaa364f7..0999f6485 100644 --- a/Cubical/Relation/Binary/Base.agda +++ b/Cubical/Relation/Binary/Base.agda @@ -1,14 +1,45 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Relation.Binary.Base where open import Cubical.Core.Everything open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels - +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Fiberwise +open import Cubical.Data.Sigma open import Cubical.HITs.SetQuotients.Base +open import Cubical.HITs.PropositionalTruncation.Base + +private + variable + ℓA ℓ≅A ℓA' ℓ≅A' : Level + +Rel : ∀ {ℓ} (A B : Type ℓ) (ℓ' : Level) → Type (ℓ-max ℓ (ℓ-suc ℓ')) +Rel A B ℓ' = A → B → Type ℓ' + +PropRel : ∀ {ℓ} (A B : Type ℓ) (ℓ' : Level) → Type (ℓ-max ℓ (ℓ-suc ℓ')) +PropRel A B ℓ' = Σ[ R ∈ Rel A B ℓ' ] ∀ a b → isProp (R a b) + +idPropRel : ∀ {ℓ} (A : Type ℓ) → PropRel A A ℓ +idPropRel A .fst a a' = ∥ a ≡ a' ∥ +idPropRel A .snd _ _ = squash + +invPropRel : ∀ {ℓ ℓ'} {A B : Type ℓ} + → PropRel A B ℓ' → PropRel B A ℓ' +invPropRel R .fst b a = R .fst a b +invPropRel R .snd b a = R .snd a b -module BinaryRelation {ℓ ℓ' : Level} {A : Type ℓ} (R : A → A → Type ℓ') where +compPropRel : ∀ {ℓ ℓ' ℓ''} {A B C : Type ℓ} + → PropRel A B ℓ' → PropRel B C ℓ'' → PropRel A C (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) +compPropRel R S .fst a c = ∥ Σ[ b ∈ _ ] (R .fst a b × S .fst b c) ∥ +compPropRel R S .snd _ _ = squash + +graphRel : ∀ {ℓ} {A B : Type ℓ} → (A → B) → Rel A B ℓ +graphRel f a b = f a ≡ b + +module BinaryRelation {ℓ ℓ' : Level} {A : Type ℓ} (R : Rel A A ℓ') where isRefl : Type (ℓ-max ℓ ℓ') isRefl = (a : A) → R a a @@ -19,7 +50,7 @@ module BinaryRelation {ℓ ℓ' : Level} {A : Type ℓ} (R : A → A → Type isTrans = (a b c : A) → R a b → R b c → R a c record isEquivRel : Type (ℓ-max ℓ ℓ') where - constructor EquivRel + constructor equivRel field reflexive : isRefl symmetric : isSym @@ -28,10 +59,87 @@ module BinaryRelation {ℓ ℓ' : Level} {A : Type ℓ} (R : A → A → Type isPropValued : Type (ℓ-max ℓ ℓ') isPropValued = (a b : A) → isProp (R a b) + isSetValued : Type (ℓ-max ℓ ℓ') + isSetValued = (a b : A) → isSet (R a b) + isEffective : Type (ℓ-max ℓ ℓ') - isEffective = (a b : A) → - let x : A / R - x = [ a ] - y : A / R - y = [ b ] - in (x ≡ y) ≃ R a b + isEffective = + (a b : A) → isEquiv (eq/ {R = R} a b) + + + impliesIdentity : Type _ + impliesIdentity = {a a' : A} → (R a a') → (a ≡ a') + + -- the total space corresponding to the binary relation w.r.t. a + relSinglAt : (a : A) → Type (ℓ-max ℓ ℓ') + relSinglAt a = Σ[ a' ∈ A ] (R a a') + + -- the statement that the total space is contractible at any a + contrRelSingl : Type (ℓ-max ℓ ℓ') + contrRelSingl = (a : A) → isContr (relSinglAt a) + + isUnivalent : Type (ℓ-max ℓ ℓ') + isUnivalent = (a a' : A) → (R a a') ≃ (a ≡ a') + + contrRelSingl→isUnivalent : isRefl → contrRelSingl → isUnivalent + contrRelSingl→isUnivalent ρ c a a' = isoToEquiv i + where + h : isProp (relSinglAt a) + h = isContr→isProp (c a) + aρa : relSinglAt a + aρa = a , ρ a + Q : (y : A) → a ≡ y → _ + Q y _ = R a y + i : Iso (R a a') (a ≡ a') + Iso.fun i r = cong fst (h aρa (a' , r)) + Iso.inv i = J Q (ρ a) + Iso.rightInv i = J (λ y p → cong fst (h aρa (y , J Q (ρ a) p)) ≡ p) + (J (λ q _ → cong fst (h aρa (a , q)) ≡ refl) + (J (λ α _ → cong fst α ≡ refl) refl + (isContr→isProp (isProp→isContrPath h aρa aρa) refl (h aρa aρa))) + (sym (JRefl Q (ρ a)))) + Iso.leftInv i r = J (λ w β → J Q (ρ a) (cong fst β) ≡ snd w) + (JRefl Q (ρ a)) (h aρa (a' , r)) + + isUnivalent→contrRelSingl : isUnivalent → contrRelSingl + isUnivalent→contrRelSingl u a = q + where + abstract + f : (x : A) → a ≡ x → R a x + f x p = invEq (u a x) p + + t : singl a → relSinglAt a + t (x , p) = x , f x p + + q : isContr (relSinglAt a) + q = isOfHLevelRespectEquiv 0 (t , totalEquiv _ _ f λ x → invEquiv (u a x) .snd) + (isContrSingl a) + +EquivRel : ∀ {ℓ} (A : Type ℓ) (ℓ' : Level) → Type (ℓ-max ℓ (ℓ-suc ℓ')) +EquivRel A ℓ' = Σ[ R ∈ Rel A A ℓ' ] BinaryRelation.isEquivRel R + +EquivPropRel : ∀ {ℓ} (A : Type ℓ) (ℓ' : Level) → Type (ℓ-max ℓ (ℓ-suc ℓ')) +EquivPropRel A ℓ' = Σ[ R ∈ PropRel A A ℓ' ] BinaryRelation.isEquivRel (R .fst) + +record RelIso {A : Type ℓA} (_≅_ : Rel A A ℓ≅A) + {A' : Type ℓA'} (_≅'_ : Rel A' A' ℓ≅A') : Type (ℓ-max (ℓ-max ℓA ℓA') (ℓ-max ℓ≅A ℓ≅A')) where + constructor reliso + field + fun : A → A' + inv : A' → A + rightInv : (a' : A') → fun (inv a') ≅' a' + leftInv : (a : A) → inv (fun a) ≅ a + +open BinaryRelation + +RelIso→Iso : {A : Type ℓA} {A' : Type ℓA'} + (_≅_ : Rel A A ℓ≅A) (_≅'_ : Rel A' A' ℓ≅A') + (uni : impliesIdentity _≅_) (uni' : impliesIdentity _≅'_) + (f : RelIso _≅_ _≅'_) + → Iso A A' +Iso.fun (RelIso→Iso _ _ _ _ f) = RelIso.fun f +Iso.inv (RelIso→Iso _ _ _ _ f) = RelIso.inv f +Iso.rightInv (RelIso→Iso _ _ uni uni' f) a' + = uni' (RelIso.rightInv f a') +Iso.leftInv (RelIso→Iso _ _ uni uni' f) a + = uni (RelIso.leftInv f a) diff --git a/Cubical/Relation/Binary/Poset.agda b/Cubical/Relation/Binary/Poset.agda new file mode 100644 index 000000000..070b66b7e --- /dev/null +++ b/Cubical/Relation/Binary/Poset.agda @@ -0,0 +1,321 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Binary.Poset where + +open import Cubical.Foundations.Prelude +open import Cubical.Functions.Logic +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv renaming (_■ to _QED) +open import Cubical.Foundations.SIP +open import Cubical.Functions.FunExtEquiv +open import Cubical.Foundations.Function +open import Cubical.Core.Primitives +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.Data.Sigma.Properties + +open import Cubical.Structures.Axioms + +-- We will adopt the convention of denoting the level of the carrier +-- set by ℓ₀ and the level of the relation result by ℓ₁. +private + variable + ℓ ℓ₀ ℓ₁ ℓ₀′ ℓ₁′ ℓ₀′′ ℓ₁′′ : Level + +Order : (ℓ₁ : Level) → Type ℓ₀ → Type (ℓ-max ℓ₀ (ℓ-suc ℓ₁)) +Order ℓ₁ A = A → A → hProp ℓ₁ + +isSetOrder : (ℓ₁ : Level) (A : Type ℓ₀) → isSet (Order ℓ₁ A) +isSetOrder ℓ₁ A = isSetΠ2 λ _ _ → isSetHProp + +-- We first start by defining what it means a for a function to be +-- order-preserving. The name "monotonic" is reserved for partial orders. + +isOrderPreserving : (M : TypeWithStr ℓ₀ (Order ℓ₁)) (N : TypeWithStr ℓ₀′ (Order ℓ₁′)) + → (fst M → fst N) → Type _ +isOrderPreserving (A , _⊑₀_) (B , _⊑₁_) f = + (x y : A) → ⟨ x ⊑₀ y ⟩ → ⟨ f x ⊑₁ f y ⟩ + +isPropIsOrderPreserving : (M : TypeWithStr ℓ₀ (Order ℓ₁)) + (N : TypeWithStr ℓ₀′ (Order ℓ₁′)) + → (f : fst M → fst N) + → isProp (isOrderPreserving M N f) +isPropIsOrderPreserving M (_ , _⊑₁_) f = isPropΠ3 λ x y p → snd (f x ⊑₁ f y) + +-- We then define what it means for an equivalence to order-preserving which is +-- nothing but the property that both directions of the equivalence are +-- order-preserving. + +isAnOrderPreservingEqv : (M : TypeWithStr ℓ₀ (Order ℓ₁)) + (N : TypeWithStr ℓ₀′ (Order ℓ₁′)) + → fst M ≃ fst N → Type _ +isAnOrderPreservingEqv M N e@(f , _) = + isOrderPreserving M N f × isOrderPreserving N M g + where + g = equivFun (invEquiv e) + +orderUnivalentStr : SNS {ℓ} (Order ℓ₁) isAnOrderPreservingEqv +orderUnivalentStr {ℓ = ℓ} {ℓ₁ = ℓ₁} {X = X} _⊑₀_ _⊑₁_ = + f , record { equiv-proof = f-equiv } + where + f : isAnOrderPreservingEqv (X , _⊑₀_) (X , _⊑₁_) (idEquiv X) → _⊑₀_ ≡ _⊑₁_ + f e@(φ , ψ) = funExt₂ λ x y → ⇔toPath (φ x y) (ψ x y) + + g : _⊑₀_ ≡ _⊑₁_ → isAnOrderPreservingEqv (X , _⊑₀_) (X , _⊑₁_) (idEquiv X) + g p = + subst + (λ _⊑_ → isAnOrderPreservingEqv (X , _⊑₀_) (X , _⊑_) (idEquiv X)) + p + ((λ _ _ x⊑₁y → x⊑₁y) , λ _ _ x⊑₁y → x⊑₁y) + + ret-f-g : retract f g + ret-f-g (φ , ψ) = + isPropΣ + (isPropIsOrderPreserving (X , _⊑₀_) (X , _⊑₁_) (idfun X)) + (λ _ → isPropIsOrderPreserving (X , _⊑₁_) (X , _⊑₀_) (idfun X)) + (g (f (φ , ψ))) (φ , ψ) + + f-equiv : (p : _⊑₀_ ≡ _⊑₁_) → isContr (fiber f p) + f-equiv p = ((to , from) , eq) , NTS + where + to : isOrderPreserving (X , _⊑₀_) (X , _⊑₁_) (idfun _) + to x y = subst (λ _⊑_ → ⟨ x ⊑₀ y ⟩ → ⟨ x ⊑ y ⟩) p (idfun _) + + from : isOrderPreserving (X , _⊑₁_) (X , _⊑₀_) (idfun _) + from x y = subst (λ _⊑_ → ⟨ x ⊑ y ⟩ → ⟨ x ⊑₀ y ⟩) p (idfun _) + + eq : f (to , from) ≡ p + eq = isSetOrder ℓ₁ X _⊑₀_ _⊑₁_ (f (to , from)) p + + NTS : (fib : fiber f p) → ((to , from) , eq) ≡ fib + NTS ((φ , ψ) , eq) = + Σ≡Prop + (λ i′ → isOfHLevelSuc 2 (isSetOrder ℓ₁ X) _⊑₀_ _⊑₁_ (f i′) p) + (Σ≡Prop + (λ _ → isPropIsOrderPreserving (X , _⊑₁_) (X , _⊑₀_) (idfun _)) + (isPropIsOrderPreserving (X , _⊑₀_) (X , _⊑₁_) (idfun _) to φ)) + +-- We now write down the axioms for a partial order and define posets on top of +-- raw ordered structures. + +isReflexive : {A : Type ℓ₀} → Order ℓ₁ A → hProp (ℓ-max ℓ₀ ℓ₁) +isReflexive {A = X} _⊑_ = ((x : X) → ⟨ x ⊑ x ⟩) , isPropΠ λ x → snd (x ⊑ x) + +isTransitive : {A : Type ℓ₀} → Order ℓ₁ A → hProp (ℓ-max ℓ₀ ℓ₁) +isTransitive {ℓ₀ = ℓ₀} {ℓ₁ = ℓ₁} {A = X} _⊑_ = φ , φ-prop + where + φ : Type (ℓ-max ℓ₀ ℓ₁) + φ = ((x y z : X) → ⟨ x ⊑ y ⇒ y ⊑ z ⇒ x ⊑ z ⟩) + φ-prop : isProp φ + φ-prop = isPropΠ3 λ x y z → snd (x ⊑ y ⇒ y ⊑ z ⇒ x ⊑ z) + +isAntisym : {A : Type ℓ₀} → isSet A → Order ℓ₁ A → hProp (ℓ-max ℓ₀ ℓ₁) +isAntisym {ℓ₀ = ℓ₀} {ℓ₁ = ℓ₁} {A = X} A-set _⊑_ = φ , φ-prop + where + φ : Type (ℓ-max ℓ₀ ℓ₁) + φ = ((x y : X) → ⟨ x ⊑ y ⟩ → ⟨ y ⊑ x ⟩ → x ≡ y) + φ-prop : isProp φ + φ-prop = isPropΠ3 λ x y z → isPropΠ λ _ → A-set x y + +-- The predicate expressing that a given order satisfies the partial order +-- axioms. +satPosetAx : (ℓ₁ : Level) (A : Type ℓ₀) → Order ℓ₁ A → hProp (ℓ-max ℓ₀ ℓ₁) +satPosetAx {ℓ₀ = ℓ₀} ℓ₁ A _⊑_ = φ , φ-prop + where + isPartial : isSet A → hProp (ℓ-max ℓ₀ ℓ₁) + isPartial A-set = isReflexive _⊑_ ⊓ isTransitive _⊑_ ⊓ isAntisym A-set _⊑_ + + φ = Σ[ A-set ∈ isSet A ] ⟨ isPartial A-set ⟩ + φ-prop = isOfHLevelΣ 1 isPropIsSet (λ x → snd (isPartial x)) + +-- The poset structure. +PosetStructure : (ℓ₁ : Level) → Type ℓ₀ → Type (ℓ-max ℓ₀ (ℓ-suc ℓ₁)) +PosetStructure ℓ₁ = AxiomsStructure (Order ℓ₁) λ A _⊑_ → ⟨ satPosetAx ℓ₁ A _⊑_ ⟩ + +isSetPosetStructure : (ℓ₁ : Level) (A : Type ℓ₀) → isSet (PosetStructure ℓ₁ A) +isSetPosetStructure ℓ₁ A = + isSetΣ + (isSetΠ2 λ _ _ → isSetHProp) λ _⊑_ → + isProp→isSet (snd (satPosetAx ℓ₁ A _⊑_)) + +Poset : (ℓ₀ ℓ₁ : Level) → Type (ℓ-max (ℓ-suc ℓ₀) (ℓ-suc ℓ₁)) +Poset ℓ₀ ℓ₁ = TypeWithStr ℓ₀ (PosetStructure ℓ₁) + +-- Some projections for syntactic convenience. + +-- Carrier set of a poset. +∣_∣ₚ : Poset ℓ₀ ℓ₁ → Type ℓ₀ +∣ X , _ ∣ₚ = X + +strₚ : (P : Poset ℓ₀ ℓ₁) → PosetStructure ℓ₁ ∣ P ∣ₚ +strₚ (_ , s) = s + +rel : (P : Poset ℓ₀ ℓ₁) → ∣ P ∣ₚ → ∣ P ∣ₚ → hProp ℓ₁ +rel (_ , _⊑_ , _) = _⊑_ + +syntax rel P x y = x ⊑[ P ] y + +⊑[_]-refl : (P : Poset ℓ₀ ℓ₁) → (x : ∣ P ∣ₚ) → ⟨ x ⊑[ P ] x ⟩ +⊑[_]-refl (_ , _ , _ , ⊑-refl , _) = ⊑-refl + +⊑[_]-trans : (P : Poset ℓ₀ ℓ₁) (x y z : ∣ P ∣ₚ) + → ⟨ x ⊑[ P ] y ⟩ → ⟨ y ⊑[ P ] z ⟩ → ⟨ x ⊑[ P ] z ⟩ +⊑[_]-trans (_ , _ , _ , _ , ⊑-trans , _) = ⊑-trans + +⊑[_]-antisym : (P : Poset ℓ₀ ℓ₁) (x y : ∣ P ∣ₚ) + → ⟨ x ⊑[ P ] y ⟩ → ⟨ y ⊑[ P ] x ⟩ → x ≡ y +⊑[_]-antisym (_ , _ , _ , _ , _ , ⊑-antisym) = ⊑-antisym + +carrier-is-set : (P : Poset ℓ₀ ℓ₁) → isSet ∣ P ∣ₚ +carrier-is-set (_ , _ , is-set , _) = is-set + +-- Definition of a monotonic map amounts to forgetting the partial order axioms. +isMonotonic : (P : Poset ℓ₀ ℓ₁) (Q : Poset ℓ₀′ ℓ₁′) → (∣ P ∣ₚ → ∣ Q ∣ₚ) → Type _ +isMonotonic (A , (_⊑₀_ , _)) (B , (_⊑₁_ , _)) = + isOrderPreserving (A , _⊑₀_) (B , _⊑₁_) + +isPropIsMonotonic : (P : Poset ℓ₀ ℓ₁) (Q : Poset ℓ₀′ ℓ₁′) + → (f : ∣ P ∣ₚ → ∣ Q ∣ₚ) + → isProp (isMonotonic P Q f) +isPropIsMonotonic (A , (_⊑₀_ , _)) (B , (_⊑₁_ , _)) f = + isPropIsOrderPreserving (A , _⊑₀_) (B , _⊑₁_) f + +-- We collect the type of monotonic maps between two posets in the following +-- type. + +_─m→_ : Poset ℓ₀ ℓ₁ → Poset ℓ₀′ ℓ₁′ → Type _ +_─m→_ P Q = Σ[ f ∈ (∣ P ∣ₚ → ∣ Q ∣ₚ) ] (isMonotonic P Q f) + +-- The identity monotonic map and composition of monotonic maps. + +𝟏m : (P : Poset ℓ₀ ℓ₁) → P ─m→ P +𝟏m P = idfun ∣ P ∣ₚ , (λ x y x⊑y → x⊑y) + +_∘m_ : {P : Poset ℓ₀ ℓ₁} {Q : Poset ℓ₀′ ℓ₁′} {R : Poset ℓ₀′′ ℓ₁′′} + → (Q ─m→ R) → (P ─m→ Q) → (P ─m→ R) +(g , pg) ∘m (f , pf) = g ∘ f , λ x y p → pg (f x) (f y) (pf x y p) + +forget-mono : (P : Poset ℓ₀ ℓ₁) (Q : Poset ℓ₀′ ℓ₁′) + ((f , f-mono) (g , g-mono) : P ─m→ Q) + → f ≡ g + → (f , f-mono) ≡ (g , g-mono) +forget-mono P Q (f , f-mono) (g , g-mono) = + Σ≡Prop (λ f → isPropΠ3 λ x y x⊑y → snd (f x ⊑[ Q ] f y)) + +module PosetReasoning (P : Poset ℓ₀ ℓ₁) where + + _⊑⟨_⟩_ : (x : ∣ P ∣ₚ) {y z : ∣ P ∣ₚ} + → ⟨ x ⊑[ P ] y ⟩ → ⟨ y ⊑[ P ] z ⟩ → ⟨ x ⊑[ P ] z ⟩ + _ ⊑⟨ p ⟩ q = ⊑[ P ]-trans _ _ _ p q + + _■ : (x : ∣ P ∣ₚ) → ⟨ x ⊑[ P ] x ⟩ + _■ = ⊑[ P ]-refl + + infixr 0 _⊑⟨_⟩_ + infix 1 _■ + +-- Univalence for posets. + +isAMonotonicEqv : (P : Poset ℓ₀ ℓ₁) (Q : Poset ℓ₀′ ℓ₁′) + → ∣ P ∣ₚ ≃ ∣ Q ∣ₚ → Type _ +isAMonotonicEqv (A , (_⊑₀_ , _)) (B , (_⊑₁_ , _)) = + isAnOrderPreservingEqv (A , _⊑₀_) (B , _⊑₁_) + +isPropIsAMonotonicEqv : (P : Poset ℓ₀ ℓ₁) (Q : Poset ℓ₀ ℓ₁′) + → (eqv : ∣ P ∣ₚ ≃ ∣ Q ∣ₚ) + → isProp (isAMonotonicEqv P Q eqv) +isPropIsAMonotonicEqv P Q e@(f , _) = + isPropΣ (isPropIsMonotonic P Q f) λ _ → isPropIsMonotonic Q P g + where + g = equivFun (invEquiv e) + +-- We denote by `_≃ₚ_` the type of monotonic poset equivalences. + +_≃ₚ_ : Poset ℓ₀ ℓ₁ → Poset ℓ₀ ℓ₁ → Type _ +_≃ₚ_ P Q = Σ[ i ∈ ∣ P ∣ₚ ≃ ∣ Q ∣ₚ ] isAMonotonicEqv P Q i + +-- From this, we can already establish that posets form an SNS and prove that +-- the category of posets is univalent. + +posetUnivalentStr : SNS {ℓ} (PosetStructure ℓ₁) isAMonotonicEqv +posetUnivalentStr {ℓ₁ = ℓ₁} = + UnivalentStr→SNS + (PosetStructure ℓ₁) + isAMonotonicEqv + (axiomsUnivalentStr _ NTS (SNS→UnivalentStr isAnOrderPreservingEqv orderUnivalentStr)) + where + NTS : (A : Type ℓ) (_⊑_ : Order ℓ₁ A) → isProp ⟨ satPosetAx ℓ₁ A _⊑_ ⟩ + NTS A _⊑_ = snd (satPosetAx ℓ₁ A _⊑_) + +poset-univ₀ : (P Q : Poset ℓ₀ ℓ₁) → (P ≃ₚ Q) ≃ (P ≡ Q) +poset-univ₀ = SIP (SNS→UnivalentStr isAMonotonicEqv posetUnivalentStr) + +-- This result is almost what we want but it is better talk directly about poset +-- _isomorphisms_ rather than equivalences. In the case when types `A` and `B` +-- are sets, the type of isomorphisms between `A` and `B` is equivalent to the +-- type of equivalences betwee them. + +-- Let us start by writing down what a poset isomorphisms is. + +isPosetIso : (P Q : Poset ℓ₀ ℓ₁) → (P ─m→ Q) → Type _ +isPosetIso P Q (f , _) = Σ[ (g , _) ∈ (Q ─m→ P) ] section f g × retract f g + +isPosetIso-prop : (P Q : Poset ℓ₀ ℓ₁) (f : P ─m→ Q) + → isProp (isPosetIso P Q f) +isPosetIso-prop P Q (f , f-mono) (g₀ , sec₀ , ret₀) (g₁ , sec₁ , ret₁) = + Σ≡Prop NTS g₀=g₁ + where + NTS : ((g , _) : Q ─m→ P) → isProp (section f g × retract f g) + NTS (g , g-mono) = isPropΣ + (isPropΠ λ x → carrier-is-set Q (f (g x)) x) λ _ → + isPropΠ λ x → carrier-is-set P (g (f x)) x + + g₀=g₁ : g₀ ≡ g₁ + g₀=g₁ = + forget-mono Q P g₀ g₁ (funExt λ x → + fst g₀ x ≡⟨ sym (cong (λ - → fst g₀ -) (sec₁ x)) ⟩ + fst g₀ (f (fst g₁ x)) ≡⟨ ret₀ (fst g₁ x) ⟩ + fst g₁ x ∎) + +-- We will denote by `P ≅ₚ Q` the type of isomorphisms between posets `P` and +-- `Q`. + +_≅ₚ_ : Poset ℓ₀ ℓ₁ → Poset ℓ₀ ℓ₁ → Type _ +P ≅ₚ Q = Σ[ f ∈ P ─m→ Q ] isPosetIso P Q f + +-- ≅ₚ is equivalent to ≃ₚ. + +≃ₚ≃≅ₚ : (P Q : Poset ℓ₀ ℓ₁) → (P ≅ₚ Q) ≃ (P ≃ₚ Q) +≃ₚ≃≅ₚ P Q = isoToEquiv (iso from to ret sec) + where + to : P ≃ₚ Q → P ≅ₚ Q + to (e@(f , _) , (f-mono , g-mono)) = + (f , f-mono) , (g , g-mono) , sec-f-g , ret-f-g + where + is = equivToIso e + g = equivFun (invEquiv e) + + sec-f-g : section f g + sec-f-g = Iso.rightInv (equivToIso e) + + ret-f-g : retract f g + ret-f-g = Iso.leftInv (equivToIso e) + + from : P ≅ₚ Q → P ≃ₚ Q + from ((f , f-mono) , ((g , g-mono) , sec , ret)) = + isoToEquiv is , f-mono , g-mono + where + is : Iso ∣ P ∣ₚ ∣ Q ∣ₚ + is = iso f g sec ret + + sec : section to from + sec (f , _) = Σ≡Prop (isPosetIso-prop P Q) refl + + ret : retract to from + ret (e , _) = Σ≡Prop (isPropIsAMonotonicEqv P Q) (Σ≡Prop isPropIsEquiv refl) + +-- Once we have this equivalence, the main result is then: the type of poset +-- isomorphisms between `P` and `Q` is equivalent to the type of identity proofs +-- between `P` and `Q` + +poset-univ : (P Q : Poset ℓ₀ ℓ₁) → (P ≅ₚ Q) ≃ (P ≡ Q) +poset-univ P Q = P ≅ₚ Q ≃⟨ ≃ₚ≃≅ₚ P Q ⟩ P ≃ₚ Q ≃⟨ poset-univ₀ P Q ⟩ P ≡ Q QED diff --git a/Cubical/Relation/Binary/Properties.agda b/Cubical/Relation/Binary/Properties.agda index f8988c437..90aa324b4 100644 --- a/Cubical/Relation/Binary/Properties.agda +++ b/Cubical/Relation/Binary/Properties.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Relation.Binary.Properties where open import Cubical.Foundations.Prelude diff --git a/Cubical/Relation/Everything.agda b/Cubical/Relation/Everything.agda index 2ad6206b9..02399d7c8 100644 --- a/Cubical/Relation/Everything.agda +++ b/Cubical/Relation/Everything.agda @@ -1,6 +1,10 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe --guardedness #-} module Cubical.Relation.Everything where -open import Cubical.Relation.Nullary public -open import Cubical.Relation.Nullary.DecidableEq public -open import Cubical.Relation.Binary public +import Cubical.Relation.Binary +import Cubical.Relation.Binary.Poset +import Cubical.Relation.Nullary +import Cubical.Relation.Nullary.DecidableEq +import Cubical.Relation.Nullary.HLevels +import Cubical.Relation.ZigZag.Applications.MultiSet +import Cubical.Relation.ZigZag.Base diff --git a/Cubical/Relation/Nullary.agda b/Cubical/Relation/Nullary.agda index 95dd0eef0..ac3c57818 100644 --- a/Cubical/Relation/Nullary.agda +++ b/Cubical/Relation/Nullary.agda @@ -1,60 +1,5 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Relation.Nullary where -open import Cubical.Foundations.Prelude - -open import Cubical.Data.Empty - -private - variable - ℓ : Level - A : Type ℓ - --- Negation -infix 3 ¬_ - -¬_ : Type ℓ → Type ℓ -¬ A = A → ⊥ - -isProp¬ : (A : Type ℓ) → isProp (¬ A) -isProp¬ A p q i x = isProp⊥ (p x) (q x) i - --- Decidable types (inspired by standard library) -data Dec (P : Type ℓ) : Type ℓ where - yes : ( p : P) → Dec P - no : (¬p : ¬ P) → Dec P - -Stable : Type ℓ → Type ℓ -Stable A = ¬ ¬ A → A - -Discrete : Type ℓ → Type ℓ -Discrete A = (x y : A) → Dec (x ≡ y) - -Stable¬ : Stable (¬ A) -Stable¬ ¬¬¬a a = ¬¬¬a ¬¬a - where - ¬¬a = λ ¬a → ¬a a - -fromYes : A → Dec A → A -fromYes _ (yes a) = a -fromYes a (no _) = a - -discreteDec : (Adis : Discrete A) → Discrete (Dec A) -discreteDec Adis (yes p) (yes p') = decideYes (Adis p p') -- TODO: monad would simply stuff - where - decideYes : Dec (p ≡ p') → Dec (yes p ≡ yes p') - decideYes (yes eq) = yes (cong yes eq) - decideYes (no ¬eq) = no λ eq → ¬eq (cong (fromYes p) eq) -discreteDec Adis (yes p) (no ¬p) = ⊥-elim (¬p p) -discreteDec Adis (no ¬p) (yes p) = ⊥-elim (¬p p) -discreteDec {A = A} Adis (no ¬p) (no ¬p') = yes (cong no (isProp¬ A ¬p ¬p')) - -isPropDec : (Aprop : isProp A) → isProp (Dec A) -isPropDec Aprop (yes a) (yes a') = cong yes (Aprop a a') -isPropDec Aprop (yes a) (no ¬a) = ⊥-elim (¬a a) -isPropDec Aprop (no ¬a) (yes a) = ⊥-elim (¬a a) -isPropDec {A = A} Aprop (no ¬a) (no ¬a') = cong no (isProp¬ A ¬a ¬a') - -mapDec : ∀ {B : Type ℓ} → (A → B) → (¬ A → ¬ B) → Dec A → Dec B -mapDec f _ (yes p) = yes (f p) -mapDec _ f (no ¬p) = no (f ¬p) +open import Cubical.Relation.Nullary.Base public +open import Cubical.Relation.Nullary.Properties public diff --git a/Cubical/Relation/Nullary/Base.agda b/Cubical/Relation/Nullary/Base.agda new file mode 100644 index 000000000..fdd1b3fa6 --- /dev/null +++ b/Cubical/Relation/Nullary/Base.agda @@ -0,0 +1,66 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Nullary.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Functions.Fixpoint + +open import Cubical.Data.Empty as ⊥ +open import Cubical.HITs.PropositionalTruncation.Base + +private + variable + ℓ : Level + A : Type ℓ + +-- Negation +infix 3 ¬_ + +¬_ : Type ℓ → Type ℓ +¬ A = A → ⊥ + +-- Decidable types (inspired by standard library) +data Dec (P : Type ℓ) : Type ℓ where + yes : ( p : P) → Dec P + no : (¬p : ¬ P) → Dec P + +NonEmpty : Type ℓ → Type ℓ +NonEmpty A = ¬ ¬ A + +Stable : Type ℓ → Type ℓ +Stable A = NonEmpty A → A + +-- reexport propositional truncation for uniformity +open Cubical.HITs.PropositionalTruncation.Base + using (∥_∥) public + +SplitSupport : Type ℓ → Type ℓ +SplitSupport A = ∥ A ∥ → A + +Collapsible : Type ℓ → Type ℓ +Collapsible A = Σ[ f ∈ (A → A) ] 2-Constant f + +Populated ⟪_⟫ : Type ℓ → Type ℓ +Populated A = (f : Collapsible A) → Fixpoint (f .fst) +⟪_⟫ = Populated + +PStable : Type ℓ → Type ℓ +PStable A = ⟪ A ⟫ → A + +onAllPaths : (Type ℓ → Type ℓ) → Type ℓ → Type ℓ +onAllPaths S A = (x y : A) → S (x ≡ y) + +Separated : Type ℓ → Type ℓ +Separated = onAllPaths Stable + +HSeparated : Type ℓ → Type ℓ +HSeparated = onAllPaths SplitSupport + +Collapsible≡ : Type ℓ → Type ℓ +Collapsible≡ = onAllPaths Collapsible + +PStable≡ : Type ℓ → Type ℓ +PStable≡ = onAllPaths PStable + +Discrete : Type ℓ → Type ℓ +Discrete = onAllPaths Dec diff --git a/Cubical/Relation/Nullary/DecidableEq.agda b/Cubical/Relation/Nullary/DecidableEq.agda index 3e3778e7b..7c66fe00a 100644 --- a/Cubical/Relation/Nullary/DecidableEq.agda +++ b/Cubical/Relation/Nullary/DecidableEq.agda @@ -1,33 +1,5 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.Relation.Nullary.DecidableEq where -open import Cubical.Core.Everything - -open import Cubical.Foundations.Prelude - -open import Cubical.Data.Empty - -open import Cubical.Relation.Nullary - --- Proof of Hedberg's theorem: a type with decidable equality is an h-set - -Dec→Stable : ∀ {ℓ} (A : Type ℓ) → Dec A → Stable A -Dec→Stable A (yes x) = λ _ → x -Dec→Stable A (no x) = λ f → ⊥-elim (f x) - -Stable≡→isSet : ∀ {ℓ} {A : Type ℓ} → (st : ∀ (a b : A) → Stable (a ≡ b)) → isSet A -Stable≡→isSet {A = A} st a b p q j i = - let f : (x : A) → a ≡ x → a ≡ x - f x p = st a x (λ h → h p) - fIsConst : (x : A) → (p q : a ≡ x) → f x p ≡ f x q - fIsConst = λ x p q i → st a x (isProp¬ _ (λ h → h p) (λ h → h q) i) - rem : (p : a ≡ b) → PathP (λ i → a ≡ p i) (f a refl) (f b p) - rem p j = f (p j) (λ i → p (i ∧ j)) - in hcomp (λ k → λ { (i = i0) → f a refl k - ; (i = i1) → fIsConst b p q j k - ; (j = i0) → rem p i k - ; (j = i1) → rem q i k }) a - --- Hedberg's theorem -Discrete→isSet : ∀ {ℓ} {A : Type ℓ} → Discrete A → isSet A -Discrete→isSet d = Stable≡→isSet (λ x y → Dec→Stable (x ≡ y) (d x y)) +open import Cubical.Relation.Nullary.Properties + using (Dec→Stable; Discrete→isSet) public diff --git a/Cubical/Relation/Nullary/HLevels.agda b/Cubical/Relation/Nullary/HLevels.agda new file mode 100644 index 000000000..467a6f2ba --- /dev/null +++ b/Cubical/Relation/Nullary/HLevels.agda @@ -0,0 +1,35 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Nullary.HLevels where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Functions.Fixpoint + +open import Cubical.Relation.Nullary + +private + variable + ℓ : Level + A : Type ℓ + +isPropPopulated : isProp (Populated A) +isPropPopulated = isPropΠ λ x → 2-Constant→isPropFixpoint (x .fst) (x .snd) + +isPropHSeparated : isProp (HSeparated A) +isPropHSeparated f g i x y a = HSeparated→isSet f x y (f x y a) (g x y a) i + +isPropCollapsible≡ : isProp (Collapsible≡ A) +isPropCollapsible≡ {A = A} f = (isPropΠ2 λ x y → isPropCollapsiblePointwise) f where + sA : isSet A + sA = Collapsible≡→isSet f + gA : isGroupoid A + gA = isSet→isGroupoid sA + isPropCollapsiblePointwise : ∀ {x y} → isProp (Collapsible (x ≡ y)) + isPropCollapsiblePointwise {x} {y} (a , ca) (b , cb) = λ i → endoFunction i , endoFunctionIsConstant i where + endoFunction : a ≡ b + endoFunction = funExt λ p → sA _ _ (a p) (b p) + isProp2-Constant : (k : I) → isProp (2-Constant (endoFunction k)) + isProp2-Constant k = isPropΠ2 λ r s → gA x y (endoFunction k r) (endoFunction k s) + endoFunctionIsConstant : PathP (λ i → 2-Constant (endoFunction i)) ca cb + endoFunctionIsConstant = isProp→PathP isProp2-Constant ca cb diff --git a/Cubical/Relation/Nullary/Properties.agda b/Cubical/Relation/Nullary/Properties.agda new file mode 100644 index 000000000..e413c967a --- /dev/null +++ b/Cubical/Relation/Nullary/Properties.agda @@ -0,0 +1,163 @@ +{-# OPTIONS --safe #-} +{- + +Properties of nullary relations, i.e. structures on types. + +Includes several results from [Notions of Anonymous Existence in Martin-Löf Type Theory](https://lmcs.episciences.org/3217) +by Altenkirch, Coquand, Escardo, Kraus. + +-} +module Cubical.Relation.Nullary.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Functions.Fixpoint + +open import Cubical.Data.Empty as ⊥ + +open import Cubical.Relation.Nullary.Base +open import Cubical.HITs.PropositionalTruncation.Base + +private + variable + ℓ : Level + A : Type ℓ + +IsoPresDiscrete : ∀ {ℓ ℓ'}{A : Type ℓ} {B : Type ℓ'} → Iso A B + → Discrete A → Discrete B +IsoPresDiscrete e dA x y with dA (Iso.inv e x) (Iso.inv e y) +... | yes p = subst Dec (λ i → Iso.rightInv e x i ≡ Iso.rightInv e y i) + (yes (cong (Iso.fun e) p)) +... | no p = subst Dec (λ i → Iso.rightInv e x i ≡ Iso.rightInv e y i) + (no λ q → p (sym (Iso.leftInv e (Iso.inv e x)) + ∙∙ cong (Iso.inv e) q + ∙∙ Iso.leftInv e (Iso.inv e y))) + +isProp¬ : (A : Type ℓ) → isProp (¬ A) +isProp¬ A p q i x = isProp⊥ (p x) (q x) i + +Stable¬ : Stable (¬ A) +Stable¬ ¬¬¬a a = ¬¬¬a ¬¬a + where + ¬¬a = λ ¬a → ¬a a + +fromYes : A → Dec A → A +fromYes _ (yes a) = a +fromYes a (no _) = a + +discreteDec : (Adis : Discrete A) → Discrete (Dec A) +discreteDec Adis (yes p) (yes p') = decideYes (Adis p p') -- TODO: monad would simply stuff + where + decideYes : Dec (p ≡ p') → Dec (yes p ≡ yes p') + decideYes (yes eq) = yes (cong yes eq) + decideYes (no ¬eq) = no λ eq → ¬eq (cong (fromYes p) eq) +discreteDec Adis (yes p) (no ¬p) = ⊥.rec (¬p p) +discreteDec Adis (no ¬p) (yes p) = ⊥.rec (¬p p) +discreteDec {A = A} Adis (no ¬p) (no ¬p') = yes (cong no (isProp¬ A ¬p ¬p')) + +isPropDec : (Aprop : isProp A) → isProp (Dec A) +isPropDec Aprop (yes a) (yes a') = cong yes (Aprop a a') +isPropDec Aprop (yes a) (no ¬a) = ⊥.rec (¬a a) +isPropDec Aprop (no ¬a) (yes a) = ⊥.rec (¬a a) +isPropDec {A = A} Aprop (no ¬a) (no ¬a') = cong no (isProp¬ A ¬a ¬a') + +mapDec : ∀ {B : Type ℓ} → (A → B) → (¬ A → ¬ B) → Dec A → Dec B +mapDec f _ (yes p) = yes (f p) +mapDec _ f (no ¬p) = no (f ¬p) + +-- we have the following implications +-- X ── ∣_∣ ─→ ∥ X ∥ +-- ∥ X ∥ ── populatedBy ─→ ⟪ X ⟫ +-- ⟪ X ⟫ ── notEmptyPopulated ─→ NonEmpty X + +-- reexport propositional truncation for uniformity +open Cubical.HITs.PropositionalTruncation.Base + using (∣_∣) public + +populatedBy : ∥ A ∥ → ⟪ A ⟫ +populatedBy {A = A} a (f , fIsConst) = h a where + h : ∥ A ∥ → Fixpoint f + h ∣ a ∣ = f a , fIsConst (f a) a + h (squash a b i) = 2-Constant→isPropFixpoint f fIsConst (h a) (h b) i + +notEmptyPopulated : ⟪ A ⟫ → NonEmpty A +notEmptyPopulated {A = A} pop u = u (fixpoint (pop (h , hIsConst))) where + h : A → A + h a = ⊥.elim (u a) + hIsConst : ∀ x y → h x ≡ h y + hIsConst x y i = ⊥.elim (isProp⊥ (u x) (u y) i) + +-- these implications induce the following for different kinds of stability, gradually weakening the assumption +Dec→Stable : Dec A → Stable A +Dec→Stable (yes x) = λ _ → x +Dec→Stable (no x) = λ f → ⊥.elim (f x) + +Stable→PStable : Stable A → PStable A +Stable→PStable st = st ∘ notEmptyPopulated + +PStable→SplitSupport : PStable A → SplitSupport A +PStable→SplitSupport pst = pst ∘ populatedBy + +-- Although SplitSupport and Collapsible are not properties, their path versions, HSeparated and Collapsible≡, are. +-- Nevertheless they are logically equivalent +SplitSupport→Collapsible : SplitSupport A → Collapsible A +SplitSupport→Collapsible {A = A} hst = h , hIsConst where + h : A → A + h p = hst ∣ p ∣ + hIsConst : 2-Constant h + hIsConst p q i = hst (squash ∣ p ∣ ∣ q ∣ i) + +Collapsible→SplitSupport : Collapsible A → SplitSupport A +Collapsible→SplitSupport f x = fixpoint (populatedBy x f) + +HSeparated→Collapsible≡ : HSeparated A → Collapsible≡ A +HSeparated→Collapsible≡ hst x y = SplitSupport→Collapsible (hst x y) + +Collapsible≡→HSeparated : Collapsible≡ A → HSeparated A +Collapsible≡→HSeparated col x y = Collapsible→SplitSupport (col x y) + +-- stability of path space under truncation or path collapsability are necessary and sufficient properties +-- for a a type to be a set. +Collapsible≡→isSet : Collapsible≡ A → isSet A +Collapsible≡→isSet {A = A} col a b p q = p≡q where + f : (x : A) → a ≡ x → a ≡ x + f x = col a x .fst + fIsConst : (x : A) → (p q : a ≡ x) → f x p ≡ f x q + fIsConst x = col a x .snd + rem : (p : a ≡ b) → PathP (λ i → a ≡ p i) (f a refl) (f b p) + rem p j = f (p j) (λ i → p (i ∧ j)) + p≡q : p ≡ q + p≡q j i = hcomp (λ k → λ { (i = i0) → f a refl k + ; (i = i1) → fIsConst b p q j k + ; (j = i0) → rem p i k + ; (j = i1) → rem q i k }) a + +HSeparated→isSet : HSeparated A → isSet A +HSeparated→isSet = Collapsible≡→isSet ∘ HSeparated→Collapsible≡ + +isSet→HSeparated : isSet A → HSeparated A +isSet→HSeparated setA x y = extract where + extract : ∥ x ≡ y ∥ → x ≡ y + extract ∣ p ∣ = p + extract (squash p q i) = setA x y (extract p) (extract q) i + +-- by the above more sufficient conditions to inhibit isSet A are given +PStable≡→HSeparated : PStable≡ A → HSeparated A +PStable≡→HSeparated pst x y = PStable→SplitSupport (pst x y) + +PStable≡→isSet : PStable≡ A → isSet A +PStable≡→isSet = HSeparated→isSet ∘ PStable≡→HSeparated + +Separated→PStable≡ : Separated A → PStable≡ A +Separated→PStable≡ st x y = Stable→PStable (st x y) + +Separated→isSet : Separated A → isSet A +Separated→isSet = PStable≡→isSet ∘ Separated→PStable≡ + +-- Proof of Hedberg's theorem: a type with decidable equality is an h-set +Discrete→Separated : Discrete A → Separated A +Discrete→Separated d x y = Dec→Stable (d x y) + +Discrete→isSet : Discrete A → isSet A +Discrete→isSet = Separated→isSet ∘ Discrete→Separated diff --git a/Cubical/Relation/ZigZag/Applications/MultiSet.agda b/Cubical/Relation/ZigZag/Applications/MultiSet.agda new file mode 100644 index 000000000..df5008247 --- /dev/null +++ b/Cubical/Relation/ZigZag/Applications/MultiSet.agda @@ -0,0 +1,358 @@ +-- We apply the theory of quasi equivalence relations (QERs) to finite multisets and association lists. +{-# OPTIONS --safe #-} +module Cubical.Relation.ZigZag.Applications.MultiSet where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.RelationalStructure +open import Cubical.Foundations.Structure +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Univalence +open import Cubical.Data.Unit +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Nat +open import Cubical.Data.List hiding ([_]) +open import Cubical.Data.Sigma +open import Cubical.HITs.SetQuotients +open import Cubical.HITs.FiniteMultiset as FMS hiding ([_] ; _++_) +open import Cubical.HITs.FiniteMultiset.CountExtensionality +open import Cubical.HITs.PropositionalTruncation +open import Cubical.Relation.Nullary +open import Cubical.Relation.ZigZag.Base + +open import Cubical.Structures.MultiSet +open import Cubical.Structures.Relational.Auto +open import Cubical.Structures.Relational.Macro + +-- we define simple association lists without any higher constructors +data AList {ℓ} (A : Type ℓ) : Type ℓ where + ⟨⟩ : AList A + ⟨_,_⟩∷_ : A → ℕ → AList A → AList A + +infixr 5 ⟨_,_⟩∷_ + +private + variable + ℓ : Level + +-- We have a CountStructure on List and AList and use these to get a QER between the two +module Lists&ALists {A : Type ℓ} (discA : Discrete A) where + + multisetShape : Type ℓ → Type ℓ + multisetShape X = X × (A → X → X) × (X → X → X) × (A → X → Const[ ℕ , isSetℕ ]) + + module S = RelMacro ℓ (autoRelDesc multisetShape) + + addIfEq : (a x : A) → ℕ → ℕ → ℕ + addIfEq a x m n with discA a x + ... | yes _ = m + n + ... | no _ = n + + module _ {a x : A} {n : ℕ} where + + addIfEq≡ : {m : ℕ} → a ≡ x → addIfEq a x m n ≡ m + n + addIfEq≡ a≡x with discA a x + ... | yes _ = refl + ... | no a≢x = ⊥.rec (a≢x a≡x) + + addIfEq≢ : {m : ℕ} → ¬ (a ≡ x) → addIfEq a x m n ≡ n + addIfEq≢ a≢x with discA a x + ... | yes a≡x = ⊥.rec (a≢x a≡x) + ... | no _ = refl + + addIfEq0 : addIfEq a x 0 n ≡ n + addIfEq0 with discA a x + ... | yes _ = refl + ... | no _ = refl + + addIfEq+ : {m : ℕ} (n' : ℕ) → addIfEq a x m (n + n') ≡ addIfEq a x m n + n' + addIfEq+ n' with discA a x + ... | yes _ = +-assoc _ n n' + ... | no _ = refl + + module L where + emp : List A + emp = [] + + insert : A → List A → List A + insert x xs = x ∷ xs + + union : List A → List A → List A + union xs ys = xs ++ ys + + count : A → List A → ℕ + count a [] = zero + count a (x ∷ xs) = addIfEq a x 1 (count a xs) + + structure : S.structure (List A) + structure = emp , insert , union , count + + countUnion : ∀ a xs ys → count a (union xs ys) ≡ count a xs + count a ys + countUnion a [] ys = refl + countUnion a (x ∷ xs) ys = + cong (addIfEq a x 1) (countUnion a xs ys) + ∙ addIfEq+ (count a ys) + + module AL where + emp : AList A + emp = ⟨⟩ + + insert* : ℕ → A → AList A → AList A + insert* m a ⟨⟩ = ⟨ a , m ⟩∷ ⟨⟩ + insert* m a (⟨ y , n ⟩∷ ys) with (discA a y) + ... | yes _ = ⟨ y , m + n ⟩∷ ys + ... | no _ = ⟨ y , n ⟩∷ insert* m a ys + + insert : A → AList A → AList A + insert = insert* 1 + + union : AList A → AList A → AList A + union ⟨⟩ ys = ys + union (⟨ x , n ⟩∷ xs) ys = insert* n x (union xs ys) + + count : A → AList A → ℕ + count a ⟨⟩ = zero + count a (⟨ y , n ⟩∷ ys) = addIfEq a y n (count a ys) + + structure : S.structure (AList A) + structure = emp , insert , union , count + + countInsert* : ∀ m a x ys → count a (insert* m x ys) ≡ addIfEq a x m (count a ys) + countInsert* m a x ⟨⟩ = refl + countInsert* m a x (⟨ y , n ⟩∷ ys) with discA a x | discA a y | discA x y + ... | yes a≡x | yes a≡y | yes x≡y = addIfEq≡ a≡y ∙ sym (+-assoc m n _) + ... | yes a≡x | yes a≡y | no x≢y = ⊥.rec (x≢y (sym a≡x ∙ a≡y)) + ... | yes a≡x | no a≢y | yes x≡y = ⊥.rec (a≢y (a≡x ∙ x≡y)) + ... | yes a≡x | no a≢y | no x≢y = addIfEq≢ a≢y ∙ countInsert* m a x ys ∙ addIfEq≡ a≡x + ... | no a≢x | yes a≡y | yes x≡y = ⊥.rec (a≢x (a≡y ∙ sym x≡y)) + ... | no a≢x | yes a≡y | no x≢y = addIfEq≡ a≡y ∙ cong (n +_) (countInsert* m a x ys ∙ addIfEq≢ a≢x) + ... | no a≢x | no a≢y | yes x≡y = addIfEq≢ a≢y + ... | no a≢x | no a≢y | no x≢y = addIfEq≢ a≢y ∙ countInsert* m a x ys ∙ addIfEq≢ a≢x + + countInsert = countInsert* 1 + + countUnion : ∀ a xs ys → count a (union xs ys) ≡ count a xs + count a ys + countUnion a ⟨⟩ ys = refl + countUnion a (⟨ x , n ⟩∷ xs) ys = + countInsert* n a x (union xs ys) + ∙ cong (addIfEq a x n) (countUnion a xs ys) + ∙ addIfEq+ (count a ys) + + -- now for the QER between List and Alist + + R : List A → AList A → Type ℓ + R xs ys = ∀ a → L.count a xs ≡ AL.count a ys + + φ : List A → AList A + φ [] = ⟨⟩ + φ (x ∷ xs) = AL.insert x (φ xs) + + ψ : AList A → List A + ψ ⟨⟩ = [] + ψ (⟨ x , zero ⟩∷ xs) = ψ xs + ψ (⟨ x , suc n ⟩∷ xs) = x ∷ ψ (⟨ x , n ⟩∷ xs) + + η : ∀ xs → R xs (φ xs) + η [] a = refl + η (x ∷ xs) a = cong (addIfEq a x 1) (η xs a) ∙ sym (AL.countInsert a x (φ xs)) + + -- for the other direction we need a little helper function + ε : ∀ y → R (ψ y) y + ε' : (x : A) (n : ℕ) (xs : AList A) (a : A) + → L.count a (ψ (⟨ x , n ⟩∷ xs)) ≡ AL.count a (⟨ x , n ⟩∷ xs) + + ε ⟨⟩ a = refl + ε (⟨ x , n ⟩∷ xs) a = ε' x n xs a + + ε' x zero xs a = ε xs a ∙ sym addIfEq0 + ε' x (suc n) xs a with discA a x + ... | yes a≡x = cong suc (ε' x n xs a ∙ addIfEq≡ a≡x) + ... | no a≢x = ε' x n xs a ∙ addIfEq≢ a≢x + + -- Induced quotients and equivalence + + open isQuasiEquivRel + + -- R is a QER + QuasiR : QuasiEquivRel _ _ ℓ + QuasiR .fst .fst = R + QuasiR .fst .snd _ _ = isPropΠ λ _ → isSetℕ _ _ + QuasiR .snd .zigzag r r' r'' a = (r a) ∙∙ sym (r' a) ∙∙ (r'' a) + QuasiR .snd .fwd a = ∣ φ a , η a ∣ + QuasiR .snd .bwd b = ∣ ψ b , ε b ∣ + + isStructuredInsert : (x : A) {xs : List A} {ys : AList A} + → R xs ys → R (L.insert x xs) (AL.insert x ys) + isStructuredInsert x {xs} {ys} r a = + cong (addIfEq a x 1) (r a) ∙ sym (AL.countInsert a x ys) + + isStructuredUnion : + {xs : List A} {ys : AList A} (r : R xs ys) + {xs' : List A} {ys' : AList A} (r' : R xs' ys') + → R (L.union xs xs') (AL.union ys ys') + isStructuredUnion {xs} {ys} r {xs'} {ys'} r' a = + L.countUnion a xs xs' ∙ cong₂ _+_ (r a) (r' a) ∙ sym (AL.countUnion a ys ys') + + -- R is structured + isStructuredR : S.relation R L.structure AL.structure + isStructuredR .fst a = refl + isStructuredR .snd .fst = isStructuredInsert + isStructuredR .snd .snd .fst {xs} {ys} = isStructuredUnion {xs} {ys} + isStructuredR .snd .snd .snd a r = r a + + module E = QER→Equiv QuasiR + open E renaming (Rᴸ to Rᴸ; Rᴿ to Rᴬᴸ) + + List/Rᴸ = (List A) / Rᴸ + AList/Rᴬᴸ = (AList A) / Rᴬᴸ + + List/Rᴸ≃AList/Rᴬᴸ : List/Rᴸ ≃ AList/Rᴬᴸ + List/Rᴸ≃AList/Rᴬᴸ = E.Thm + + main : QERDescends _ S.relation (List A , L.structure) (AList A , AL.structure) QuasiR + main = structuredQER→structuredEquiv S.suitable _ _ QuasiR isStructuredR + + open QERDescends + + LQstructure : S.structure List/Rᴸ + LQstructure = main .quoᴸ .fst + + ALQstructure : S.structure AList/Rᴬᴸ + ALQstructure = main .quoᴿ .fst + + -- We get a path between structure over the equivalence from the fact that the QER is structured + List/Rᴸ≡AList/Rᴬᴸ : + Path (TypeWithStr ℓ S.structure) (List/Rᴸ , LQstructure) (AList/Rᴬᴸ , ALQstructure) + List/Rᴸ≡AList/Rᴬᴸ = + sip S.univalent _ _ + (E.Thm , S.matches (List/Rᴸ , LQstructure) (AList/Rᴬᴸ , ALQstructure) E.Thm .fst (main .rel)) + + -- Deriving associativity of union for association list multisets + + LQunion = LQstructure .snd .snd .fst + ALQunion = ALQstructure .snd .snd .fst + + hasAssociativeUnion : TypeWithStr ℓ S.structure → Type ℓ + hasAssociativeUnion (_ , _ , _ , _⊔_ , _) = + ∀ xs ys zs → (xs ⊔ ys) ⊔ zs ≡ xs ⊔ (ys ⊔ zs) + + LQassoc : hasAssociativeUnion (List/Rᴸ , LQstructure) + LQassoc = elimProp3 (λ _ _ _ → squash/ _ _) (λ xs ys zs i → [ ++-assoc xs ys zs i ]) + + ALQassoc : hasAssociativeUnion (AList/Rᴬᴸ , ALQstructure) + ALQassoc = subst hasAssociativeUnion List/Rᴸ≡AList/Rᴬᴸ LQassoc + + -- We now show that List/Rᴸ≃FMSet + + _∷/_ : A → List/Rᴸ → List/Rᴸ + _∷/_ = LQstructure .snd .fst + + multisetShape' : Type ℓ → Type ℓ + multisetShape' X = X × (A → X → X) × (A → X → Const[ ℕ , isSetℕ ]) + + FMSstructure : S.structure (FMSet A) + FMSstructure = [] , _∷_ , FMS._++_ , FMScount discA + + infixr 5 _∷/_ + + FMSet→List/Rᴸ : FMSet A → List/Rᴸ + FMSet→List/Rᴸ = FMS.Rec.f squash/ [ [] ] _∷/_ β + where + δ : ∀ c a b xs → L.count c (a ∷ b ∷ xs) ≡ L.count c (b ∷ a ∷ xs) + δ c a b xs with discA c a | discA c b + δ c a b xs | yes _ | yes _ = refl + δ c a b xs | yes _ | no _ = refl + δ c a b xs | no _ | yes _ = refl + δ c a b xs | no _ | no _ = refl + + γ : ∀ a b xs → Rᴸ (a ∷ b ∷ xs) (b ∷ a ∷ xs) + γ a b xs = + ∣ φ (a ∷ b ∷ xs) , η (a ∷ b ∷ xs) , (λ c → δ c b a xs ∙ η (a ∷ b ∷ xs) c) ∣ + + β : ∀ a b [xs] → a ∷/ b ∷/ [xs] ≡ b ∷/ a ∷/ [xs] + β a b = elimProp (λ _ → squash/ _ _) (λ xs → eq/ _ _ (γ a b xs)) + + -- The inverse is induced by the standard projection of lists into finite multisets, + -- which is a morphism of CountStructures + -- Moreover, we need 'count-extensionality' for finite multisets + List→FMSet : List A → FMSet A + List→FMSet [] = [] + List→FMSet (x ∷ xs) = x ∷ List→FMSet xs + + List→FMSet-count : ∀ a xs → L.count a xs ≡ FMScount discA a (List→FMSet xs) + List→FMSet-count a [] = refl + List→FMSet-count a (x ∷ xs) with discA a x + ... | yes _ = cong suc (List→FMSet-count a xs) + ... | no _ = List→FMSet-count a xs + + List/Rᴸ→FMSet : List/Rᴸ → FMSet A + List/Rᴸ→FMSet [ xs ] = List→FMSet xs + List/Rᴸ→FMSet (eq/ xs ys r i) = path i + where + countsAgree : ∀ a → L.count a xs ≡ L.count a ys + countsAgree a = cong (LQstructure .snd .snd .snd a) (eq/ xs ys r) + + θ : ∀ a → FMScount discA a (List→FMSet xs) ≡ FMScount discA a (List→FMSet ys) + θ a = sym (List→FMSet-count a xs) ∙∙ countsAgree a ∙∙ List→FMSet-count a ys + + path : List→FMSet xs ≡ List→FMSet ys + path = FMScountExt.Thm discA _ _ θ + List/Rᴸ→FMSet (squash/ xs/ xs/' p q i j) = + trunc (List/Rᴸ→FMSet xs/) (List/Rᴸ→FMSet xs/') (cong List/Rᴸ→FMSet p) (cong List/Rᴸ→FMSet q) i j + + List/Rᴸ→FMSet-insert : (x : A) (ys : List/Rᴸ) → List/Rᴸ→FMSet (x ∷/ ys) ≡ x ∷ List/Rᴸ→FMSet ys + List/Rᴸ→FMSet-insert x = elimProp (λ _ → FMS.trunc _ _) λ xs → refl + + List→FMSet-union : (xs ys : List A) + → List→FMSet (xs ++ ys) ≡ FMS._++_ (List→FMSet xs) (List→FMSet ys) + List→FMSet-union [] ys = refl + List→FMSet-union (x ∷ xs) ys = cong (x ∷_) (List→FMSet-union xs ys) + + List/Rᴸ≃FMSet : List/Rᴸ ≃ FMSet A + List/Rᴸ≃FMSet = isoToEquiv (iso List/Rᴸ→FMSet FMSet→List/Rᴸ τ σ) + where + σ' : (xs : List A) → FMSet→List/Rᴸ (List/Rᴸ→FMSet [ xs ]) ≡ [ xs ] + σ' [] = refl + σ' (x ∷ xs) = cong (x ∷/_) (σ' xs) + + σ : section FMSet→List/Rᴸ List/Rᴸ→FMSet + σ = elimProp (λ _ → squash/ _ _) σ' + + τ' : ∀ x {xs} → List/Rᴸ→FMSet (FMSet→List/Rᴸ xs) ≡ xs → List/Rᴸ→FMSet (FMSet→List/Rᴸ (x ∷ xs)) ≡ x ∷ xs + τ' x {xs} p = List/Rᴸ→FMSet-insert x (FMSet→List/Rᴸ xs) ∙ cong (x ∷_) p + + τ : retract FMSet→List/Rᴸ List/Rᴸ→FMSet + τ = FMS.ElimProp.f (FMS.trunc _ _) refl τ' + + List/Rᴸ≃FMSet-EquivStr : S.equiv (List/Rᴸ , LQstructure) (FMSet A , FMSstructure) List/Rᴸ≃FMSet + List/Rᴸ≃FMSet-EquivStr .fst = refl + List/Rᴸ≃FMSet-EquivStr .snd .fst a xs = List/Rᴸ→FMSet-insert a xs + List/Rᴸ≃FMSet-EquivStr .snd .snd .fst = elimProp2 (λ _ _ → trunc _ _) List→FMSet-union + List/Rᴸ≃FMSet-EquivStr .snd .snd .snd a = + elimProp (λ _ → isSetℕ _ _) (List→FMSet-count a) + + {- + Putting everything together we get: + ≃ + List/Rᴸ ------------> AList/Rᴬᴸ + | + |≃ + | + ∨ + ≃ + FMSet A ------------> AssocList A + We thus get that AList/Rᴬᴸ≃AssocList. + Constructing such an equivalence directly requires count extensionality for association lists, + which should be even harder to prove than for finite multisets. + This strategy should work for all implementations of multisets with HITs. + We just have to show that: + ∙ The HIT is equivalent to FMSet (like AssocList) + ∙ There is a QER between lists and the basic data type of the HIT + with the higher constructors removed (like AList) + Then we get that this HIT is equivalent to the corresponding set quotient that identifies elements + that give the same count on each a : A. + TODO: Show that all the equivalences are indeed isomorphisms of multisets not only of CountStructures! + -} diff --git a/Cubical/Relation/ZigZag/Base.agda b/Cubical/Relation/ZigZag/Base.agda new file mode 100644 index 000000000..01cf6c119 --- /dev/null +++ b/Cubical/Relation/ZigZag/Base.agda @@ -0,0 +1,161 @@ +-- We define ZigZag-complete relations and prove that quasi equivalence relations +-- give rise to equivalences on the set quotients. +{-# OPTIONS --safe #-} +module Cubical.Relation.ZigZag.Base where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.HITs.SetQuotients +open import Cubical.HITs.PropositionalTruncation as Trunc +open import Cubical.Relation.Binary.Base + +open BinaryRelation +open isEquivRel + +private + variable + ℓ ℓ' : Level + +isZigZagComplete : {A B : Type ℓ} (R : A → B → Type ℓ') → Type (ℓ-max ℓ ℓ') +isZigZagComplete R = ∀ {a b a' b'} → R a b → R a' b → R a' b' → R a b' + +ZigZagRel : (A B : Type ℓ) (ℓ' : Level) → Type (ℓ-max ℓ (ℓ-suc ℓ')) +ZigZagRel A B ℓ' = Σ[ R ∈ (A → B → Type ℓ') ] (isZigZagComplete R) + +record isQuasiEquivRel {A B : Type ℓ} (R : A → B → Type ℓ') : Type (ℓ-max ℓ ℓ') where + field + zigzag : isZigZagComplete R + fwd : (a : A) → ∃[ b ∈ B ] R a b + bwd : (b : B) → ∃[ a ∈ A ] R a b + +open isQuasiEquivRel + +QuasiEquivRel : (A B : Type ℓ) (ℓ' : Level) → Type (ℓ-max ℓ (ℓ-suc ℓ')) +QuasiEquivRel A B ℓ' = + Σ[ R ∈ PropRel A B ℓ' ] isQuasiEquivRel (R .fst) + +invQER : {A B : Type ℓ} {ℓ' : Level} → QuasiEquivRel A B ℓ' → QuasiEquivRel B A ℓ' +invQER (R , qer) .fst = invPropRel R +invQER (R , qer) .snd .zigzag aRb aRb' a'Rb' = qer .zigzag a'Rb' aRb' aRb +invQER (R , qer) .snd .fwd = qer .bwd +invQER (R , qer) .snd .bwd = qer .fwd + +QER→EquivRel : {A B : Type ℓ} + → QuasiEquivRel A B ℓ' → EquivPropRel A (ℓ-max ℓ ℓ') +QER→EquivRel (R , sim) .fst = compPropRel R (invPropRel R) +QER→EquivRel (R , sim) .snd .reflexive a = Trunc.map (λ {(b , r) → b , r , r}) (sim .fwd a) +QER→EquivRel (R , sim) .snd .symmetric _ _ = Trunc.map (λ {(b , r₀ , r₁) → b , r₁ , r₀}) +QER→EquivRel (R , sim) .snd .transitive _ _ _ = + Trunc.map2 (λ {(b , r₀ , r₁) (b' , r₀' , r₁') → b , r₀ , sim .zigzag r₁' r₀' r₁}) + +-- The following result is due to Carlo Angiuli +module QER→Equiv {A B : Type ℓ} (R : QuasiEquivRel A B ℓ') where + + Rᴸ = QER→EquivRel R .fst .fst + Rᴿ = QER→EquivRel (invQER R) .fst .fst + + private + sim = R .snd + + private + f : (a : A) → ∃[ b ∈ B ] R .fst .fst a b → B / Rᴿ + f a = + Trunc.rec→Set squash/ + ([_] ∘ fst) + (λ {(b , r) (b' , r') → eq/ b b' ∣ a , r , r' ∣}) + + fPath : + (a₀ : A) (s₀ : ∃[ b ∈ B ] R .fst .fst a₀ b) + (a₁ : A) (s₁ : ∃[ b ∈ B ] R .fst .fst a₁ b) + → Rᴸ a₀ a₁ + → f a₀ s₀ ≡ f a₁ s₁ + fPath a₀ = + Trunc.elim (λ _ → isPropΠ3 λ _ _ _ → squash/ _ _) + (λ {(b₀ , r₀) a₁ → + Trunc.elim (λ _ → isPropΠ λ _ → squash/ _ _) + (λ {(b₁ , r₁) → + Trunc.elim (λ _ → squash/ _ _) + (λ {(b' , r₀' , r₁') → eq/ b₀ b₁ ∣ a₀ , r₀ , sim .zigzag r₀' r₁' r₁ ∣})})}) + + φ : A / Rᴸ → B / Rᴿ + φ [ a ] = f a (sim .fwd a) + φ (eq/ a₀ a₁ r i) = fPath a₀ (sim .fwd a₀) a₁ (sim .fwd a₁) r i + φ (squash/ _ _ p q j i) = squash/ _ _ (cong φ p) (cong φ q) j i + + + relToFwd≡ : ∀ {a b} → R .fst .fst a b → φ [ a ] ≡ [ b ] + relToFwd≡ {a} {b} r = + Trunc.elim {P = λ s → f a s ≡ [ b ]} (λ _ → squash/ _ _) + (λ {(b' , r') → eq/ b' b ∣ a , r' , r ∣}) + (sim .fwd a) + + fwd≡ToRel : ∀ {a b} → φ [ a ] ≡ [ b ] → R .fst .fst a b + fwd≡ToRel {a} {b} = + Trunc.elim {P = λ s → f a s ≡ [ b ] → R .fst .fst a b} + (λ _ → isPropΠ λ _ → R .fst .snd _ _) + (λ {(b' , r') p → + Trunc.rec (R .fst .snd _ _) + (λ {(a' , s' , s) → R .snd .zigzag r' s' s}) + (effective + (QER→EquivRel (invQER R) .fst .snd) + (QER→EquivRel (invQER R) .snd) + b' b p)}) + (sim .fwd a) + + private + g : (b : B) → ∃[ a ∈ A ] R .fst .fst a b → A / Rᴸ + g b = + Trunc.rec→Set squash/ + ([_] ∘ fst) + (λ {(a , r) (a' , r') → eq/ a a' ∣ b , r , r' ∣}) + + gPath : + (b₀ : B) (s₀ : ∃[ a ∈ A ] R .fst .fst a b₀) + (b₁ : B) (s₁ : ∃[ a ∈ A ] R .fst .fst a b₁) + → Rᴿ b₀ b₁ + → g b₀ s₀ ≡ g b₁ s₁ + gPath b₀ = + Trunc.elim (λ _ → isPropΠ3 λ _ _ _ → squash/ _ _) + (λ {(a₀ , r₀) b₁ → + Trunc.elim (λ _ → isPropΠ λ _ → squash/ _ _) + (λ {(a₁ , r₁) → + Trunc.elim (λ _ → squash/ _ _) + (λ {(a' , r₀' , r₁') → eq/ a₀ a₁ ∣ b₀ , r₀ , sim .zigzag r₁ r₁' r₀' ∣})})}) + + ψ : B / Rᴿ → A / Rᴸ + ψ [ b ] = g b (sim .bwd b) + ψ (eq/ b₀ b₁ r i) = gPath b₀ (sim .bwd b₀) b₁ (sim .bwd b₁) r i + ψ (squash/ _ _ p q j i) = squash/ _ _ (cong ψ p) (cong ψ q) j i + + relToBwd≡ : ∀ {a b} → R .fst .fst a b → ψ [ b ] ≡ [ a ] + relToBwd≡ {a} {b} r = + Trunc.elim {P = λ s → g b s ≡ [ a ]} (λ _ → squash/ _ _) + (λ {(a' , r') → eq/ a' a ∣ b , r' , r ∣}) + (sim .bwd b) + + private + η : ∀ qb → φ (ψ qb) ≡ qb + η = + elimProp (λ _ → squash/ _ _) + (λ b → + Trunc.elim {P = λ s → φ (g b s) ≡ [ b ]} (λ _ → squash/ _ _) + (λ {(a , r) → relToFwd≡ r}) + (sim .bwd b)) + + ε : ∀ qa → ψ (φ qa) ≡ qa + ε = + elimProp (λ _ → squash/ _ _) + (λ a → + Trunc.elim {P = λ s → ψ (f a s) ≡ [ a ]} (λ _ → squash/ _ _) + (λ {(b , r) → relToBwd≡ r}) + (sim .fwd a)) + + bwd≡ToRel : ∀ {a b} → ψ [ b ] ≡ [ a ] → R .fst .fst a b + bwd≡ToRel {a} {b} p = fwd≡ToRel (cong φ (sym p) ∙ η [ b ]) + + Thm : (A / Rᴸ) ≃ (B / Rᴿ) + Thm = isoToEquiv (iso φ ψ η ε) diff --git a/Cubical/Structures/Auto.agda b/Cubical/Structures/Auto.agda new file mode 100644 index 000000000..26145e992 --- /dev/null +++ b/Cubical/Structures/Auto.agda @@ -0,0 +1,250 @@ +{- + +Macros (autoDesc, AutoStructure, AutoEquivStr, autoUnivalentStr) for automatically generating structure definitions. + +For example: + + autoDesc (λ (X : Type₀) → X → X × ℕ) ↦ function+ var (var , constant ℕ) + +We prefer to use the constant structure whenever possible, e.g., [autoDesc (λ (X : Type₀) → ℕ → ℕ)] +is [constant (ℕ → ℕ)] rather than [function (constant ℕ) (constant ℕ)]. + +Writing [auto* (λ X → ⋯)] doesn't seem to work, but [auto* (λ (X : Type ℓ) → ⋯)] does. + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Auto where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.List +open import Cubical.Data.Bool +open import Cubical.Data.Maybe +open import Cubical.Structures.Macro as Macro + +import Agda.Builtin.Reflection as R +open import Cubical.Reflection.Base + +-- Magic number +private + FUEL = 10000 + +-- Mark part of a structure definition to use transport-based structured equivalences +abstract + Transp[_] : ∀ {ℓ} → Type ℓ → Type ℓ + Transp[ A ] = A + +-- Some reflection utilities +private + tLevel = R.def (quote Level) [] + + tType : R.Term → R.Term + tType ℓ = R.def (quote Type) [ varg ℓ ] + + tTranspDesc : R.Term → R.Term → R.Term + tTranspDesc ℓ ℓ' = R.def (quote TranspDesc) (ℓ v∷ ℓ' v∷ []) + + tDesc : R.Term → R.Term → R.Term → R.Term + tDesc ℓ ℓ₁ ℓ₁' = R.def (quote Desc) (ℓ v∷ ℓ₁ v∷ ℓ₁' v∷ []) + + func : (ℓ ℓ' : Level) → Type (ℓ-suc (ℓ-max ℓ ℓ')) + func ℓ ℓ' = Type ℓ → Type ℓ' + + tStruct : R.Term → R.Term → R.Term + tStruct ℓ ℓ' = R.def (quote func) (varg ℓ ∷ varg ℓ' ∷ []) + +-- We try to build a descriptor by unifying the provided structure with combinators we're aware of. We +-- redefine the structure combinators as the *Shape terms below so that we don't depend on the specific way +-- these are defined in other files (order of implicit arguments and so on); the syntactic analysis that goes +-- on here means that we would get mysterious errors if those changed. +private + constantShape : ∀ {ℓ'} (ℓ : Level) (A : Type ℓ') → (Type ℓ → Type ℓ') + constantShape _ A _ = A + + pointedShape : (ℓ : Level) → Type ℓ → Type ℓ + pointedShape _ X = X + + productShape : ∀ {ℓ₀ ℓ₁} (ℓ : Level) + → (Type ℓ → Type ℓ₀) → (Type ℓ → Type ℓ₁) → Type ℓ → Type (ℓ-max ℓ₀ ℓ₁) + productShape _ A₀ A₁ X = A₀ X × A₁ X + + functionShape : ∀ {ℓ₀ ℓ₁} (ℓ : Level) + → (Type ℓ → Type ℓ₀) → (Type ℓ → Type ℓ₁) → Type ℓ → Type (ℓ-max ℓ₀ ℓ₁) + functionShape _ A₀ A₁ X = A₀ X → A₁ X + + maybeShape : ∀ {ℓ₀} (ℓ : Level) + → (Type ℓ → Type ℓ₀) → Type ℓ → Type ℓ₀ + maybeShape _ A₀ X = Maybe (A₀ X) + + transpShape : ∀ {ℓ₀} (ℓ : Level) + → (Type ℓ → Type ℓ₀) → Type ℓ → Type ℓ₀ + transpShape _ A₀ X = Transp[ A₀ X ] + +private + -- Build transport structure descriptor from a function [t : Type ℓ → Type ℓ'] + buildTranspDesc : ℕ → R.Term → R.Term → R.Term → R.TC R.Term + buildTranspDesc zero ℓ ℓ' t = R.typeError (R.strErr "Ran out of fuel! at \n" ∷ R.termErr t ∷ []) + buildTranspDesc (suc fuel) ℓ ℓ' t = + tryConstant t <|> tryPointed t <|> tryProduct t <|> tryFunction t <|> tryMaybe t <|> + R.typeError (R.strErr "Can't automatically generate a transp structure for\n" ∷ R.termErr t ∷ []) + where + tryConstant : R.Term → R.TC R.Term + tryConstant t = + newMeta (tType ℓ') >>= λ A → + R.unify t (R.def (quote constantShape) (ℓ v∷ A v∷ [])) >> + R.returnTC (R.con (quote TranspDesc.constant) (A v∷ [])) + + tryPointed : R.Term → R.TC R.Term + tryPointed t = + R.unify t (R.def (quote pointedShape) (ℓ v∷ [])) >> + R.returnTC (R.con (quote TranspDesc.var) []) + + tryFunction : R.Term → R.TC R.Term + tryFunction t = + newMeta tLevel >>= λ ℓ₀ → + newMeta tLevel >>= λ ℓ₁ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + newMeta (tStruct ℓ ℓ₁) >>= λ A₁ → + R.unify t (R.def (quote functionShape) (ℓ v∷ A₀ v∷ A₁ v∷ [])) >> + buildTranspDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + buildTranspDesc fuel ℓ ℓ₁ A₁ >>= λ d₁ → + R.returnTC (R.con (quote TranspDesc.function) (d₀ v∷ d₁ v∷ [])) + + tryProduct : R.Term → R.TC R.Term + tryProduct t = + newMeta tLevel >>= λ ℓ₀ → + newMeta tLevel >>= λ ℓ₁ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + newMeta (tStruct ℓ ℓ₁) >>= λ A₁ → + R.unify t (R.def (quote productShape) (ℓ v∷ A₀ v∷ A₁ v∷ [])) >> + buildTranspDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + buildTranspDesc fuel ℓ ℓ₁ A₁ >>= λ d₁ → + R.returnTC (R.con (quote TranspDesc._,_) (d₀ v∷ d₁ v∷ [])) + + tryMaybe : R.Term → R.TC R.Term + tryMaybe t = + newMeta tLevel >>= λ ℓ₀ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + R.unify t (R.def (quote maybeShape) (ℓ v∷ A₀ v∷ [])) >> + buildTranspDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + R.returnTC (R.con (quote TranspDesc.maybe) (d₀ v∷ [])) + + autoTranspDesc' : R.Term → R.Term → R.TC Unit + autoTranspDesc' t hole = + R.inferType hole >>= λ H → + newMeta tLevel >>= λ ℓ → + newMeta tLevel >>= λ ℓ' → + R.unify (tTranspDesc ℓ ℓ') H >> + R.checkType t (tStruct ℓ ℓ') >> + buildTranspDesc FUEL ℓ ℓ' t >>= R.unify hole + +-- Build structure descriptor from a function [t : Type ℓ → Type ℓ'] +buildDesc : ℕ → R.Term → R.Term → R.Term → R.TC R.Term +buildDesc zero ℓ ℓ' t = R.typeError (R.strErr "Ran out of fuel! at \n" ∷ R.termErr t ∷ []) +buildDesc (suc fuel) ℓ ℓ' t = + tryConstant t <|> tryPointed t <|> tryProduct t <|> tryFunction t <|> + tryMaybe t <|> tryTransp t <|> + R.typeError (R.strErr "Can't automatically generate a structure for\n" ∷ R.termErr t ∷ []) + where + tryConstant : R.Term → R.TC R.Term + tryConstant t = + newMeta (tType ℓ') >>= λ A → + R.unify t (R.def (quote constantShape) (ℓ v∷ A v∷ [])) >> + R.returnTC (R.con (quote Desc.constant) (A v∷ [])) + + tryPointed : R.Term → R.TC R.Term + tryPointed t = + R.unify t (R.def (quote pointedShape) (ℓ v∷ [])) >> + R.returnTC (R.con (quote Desc.var) []) + + tryProduct : R.Term → R.TC R.Term + tryProduct t = + newMeta tLevel >>= λ ℓ₀ → + newMeta tLevel >>= λ ℓ₁ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + newMeta (tStruct ℓ ℓ₁) >>= λ A₁ → + R.unify t (R.def (quote productShape) (ℓ v∷ A₀ v∷ A₁ v∷ [])) >> + buildDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + buildDesc fuel ℓ ℓ₁ A₁ >>= λ d₁ → + R.returnTC (R.con (quote Desc._,_) (d₀ v∷ d₁ v∷ [])) + + tryFunction : R.Term → R.TC R.Term + tryFunction t = + newMeta tLevel >>= λ ℓ₀ → + newMeta tLevel >>= λ ℓ₁ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + newMeta (tStruct ℓ ℓ₁) >>= λ A₁ → + R.unify t (R.def (quote functionShape) (ℓ v∷ A₀ v∷ A₁ v∷ [])) >> + buildTranspDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + buildDesc fuel ℓ ℓ₁ A₁ >>= λ d₁ → + R.returnTC (R.con (quote Desc.function+) (d₀ v∷ d₁ v∷ [])) + + tryMaybe : R.Term → R.TC R.Term + tryMaybe t = + newMeta tLevel >>= λ ℓ₀ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + R.unify t (R.def (quote maybeShape) (ℓ v∷ A₀ v∷ [])) >> + buildDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + R.returnTC (R.con (quote Desc.maybe) (d₀ v∷ [])) + + tryTransp : R.Term → R.TC R.Term + tryTransp t = + newMeta (tStruct ℓ ℓ') >>= λ A₀ → + R.unify t (R.def (quote transpShape) (ℓ v∷ A₀ v∷ [])) >> + buildTranspDesc fuel ℓ ℓ' A₀ >>= λ d₀ → + R.returnTC (R.con (quote Desc.transpDesc) (d₀ v∷ [])) + +autoDesc' : R.Term → R.Term → R.TC Unit +autoDesc' t hole = + R.inferType hole >>= λ H → + newMeta tLevel >>= λ ℓ → + newMeta tLevel >>= λ ℓ' → + R.unify (tDesc ℓ ℓ' R.unknown) H >> + R.checkType t (tStruct ℓ ℓ') >> + buildDesc FUEL ℓ ℓ' t >>= R.unify hole + +macro + -- (Type ℓ → Type ℓ₁) → TranspDesc ℓ + autoTranspDesc : R.Term → R.Term → R.TC Unit + autoTranspDesc = autoTranspDesc' + + -- (S : Type ℓ → Type ℓ₁) → EquivAction (AutoStructure S) + autoEquivAction : R.Term → R.Term → R.TC Unit + autoEquivAction t hole = + newMeta (tTranspDesc R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote transpMacroAction) [ varg d ]) >> + autoTranspDesc' t d + + -- (S : Type ℓ → Type ℓ₁) → TransportStr (autoEquivAction S) + autoTransportStr : R.Term → R.Term → R.TC Unit + autoTransportStr t hole = + newMeta (tTranspDesc R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote transpMacroTransportStr) [ varg d ]) >> + autoTranspDesc' t d + + -- (S : Type ℓ → Type ℓ₁) → Desc ℓ + autoDesc : R.Term → R.Term → R.TC Unit + autoDesc = autoDesc' + + -- (S : Type ℓ → Type ℓ₁) → (Type ℓ → Type ℓ₁) + -- Removes Transp[_] annotations + AutoStructure : R.Term → R.Term → R.TC Unit + AutoStructure t hole = + newMeta (tDesc R.unknown R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote MacroStructure) [ varg d ]) >> + autoDesc' t d + + -- (S : Type ℓ → Type ℓ₁) → StrEquiv (AutoStructure S) _ + AutoEquivStr : R.Term → R.Term → R.TC Unit + AutoEquivStr t hole = + newMeta (tDesc R.unknown R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote MacroEquivStr) [ varg d ]) >> + autoDesc' t d + + -- (S : Type ℓ → Type ℓ₁) → UnivalentStr (AutoStructure S) (AutoEquivStr S) + autoUnivalentStr : R.Term → R.Term → R.TC Unit + autoUnivalentStr t hole = + newMeta (tDesc R.unknown R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote MacroUnivalentStr) [ varg d ]) >> + autoDesc' t d diff --git a/Cubical/Structures/Axioms.agda b/Cubical/Structures/Axioms.agda new file mode 100644 index 000000000..e3d2a225d --- /dev/null +++ b/Cubical/Structures/Axioms.agda @@ -0,0 +1,69 @@ +{- + +Add axioms (i.e., propositions) to a structure S without changing the definition of structured equivalence. + +X ↦ Σ[ s ∈ S X ] (P X s) where (P X s) is a proposition for all X and s. + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Axioms where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Path +open import Cubical.Foundations.SIP +open import Cubical.Data.Sigma + +private + variable + ℓ ℓ₁ ℓ₁' ℓ₂ : Level + +AxiomsStructure : (S : Type ℓ → Type ℓ₁) + (axioms : (X : Type ℓ) → S X → Type ℓ₂) + → Type ℓ → Type (ℓ-max ℓ₁ ℓ₂) +AxiomsStructure S axioms X = Σ[ s ∈ S X ] (axioms X s) + +AxiomsEquivStr : {S : Type ℓ → Type ℓ₁} (ι : StrEquiv S ℓ₁') + (axioms : (X : Type ℓ) → S X → Type ℓ₂) + → StrEquiv (AxiomsStructure S axioms) ℓ₁' +AxiomsEquivStr ι axioms (X , (s , a)) (Y , (t , b)) e = ι (X , s) (Y , t) e + +axiomsUnivalentStr : {S : Type ℓ → Type ℓ₁} + (ι : (A B : TypeWithStr ℓ S) → A .fst ≃ B .fst → Type ℓ₁') + {axioms : (X : Type ℓ) → S X → Type ℓ₂} + (axioms-are-Props : (X : Type ℓ) (s : S X) → isProp (axioms X s)) + (θ : UnivalentStr S ι) + → UnivalentStr (AxiomsStructure S axioms) (AxiomsEquivStr ι axioms) +axiomsUnivalentStr {S = S} ι {axioms = axioms} axioms-are-Props θ {X , s , a} {Y , t , b} e = + ι (X , s) (Y , t) e + ≃⟨ θ e ⟩ + PathP (λ i → S (ua e i)) s t + ≃⟨ invEquiv (Σ-contractSnd λ _ → isOfHLevelPathP' 0 (axioms-are-Props _ _) _ _) ⟩ + Σ[ p ∈ PathP (λ i → S (ua e i)) s t ] PathP (λ i → axioms (ua e i) (p i)) a b + ≃⟨ ΣPath≃PathΣ ⟩ + PathP (λ i → AxiomsStructure S axioms (ua e i)) (s , a) (t , b) + ■ + +inducedStructure : {S : Type ℓ → Type ℓ₁} + {ι : (A B : TypeWithStr ℓ S) → A .fst ≃ B .fst → Type ℓ₁'} + (θ : UnivalentStr S ι) + {axioms : (X : Type ℓ) → S X → Type ℓ₂} + (A : TypeWithStr ℓ (AxiomsStructure S axioms)) (B : TypeWithStr ℓ S) + → (typ A , str A .fst) ≃[ ι ] B + → TypeWithStr ℓ (AxiomsStructure S axioms) +inducedStructure θ {axioms} A B eqv = + B .fst , B .snd , subst (uncurry axioms) (sip θ _ _ eqv) (A .snd .snd) + +transferAxioms : {S : Type ℓ → Type ℓ₁} + {ι : (A B : TypeWithStr ℓ S) → A .fst ≃ B .fst → Type ℓ₁'} + (θ : UnivalentStr S ι) + {axioms : (X : Type ℓ) → S X → Type ℓ₂} + (A : TypeWithStr ℓ (AxiomsStructure S axioms)) (B : TypeWithStr ℓ S) + → (typ A , str A .fst) ≃[ ι ] B + → axioms (fst B) (snd B) +transferAxioms θ {axioms} A B eqv = + subst (uncurry axioms) (sip θ _ _ eqv) (A .snd .snd) diff --git a/Cubical/Structures/Constant.agda b/Cubical/Structures/Constant.agda new file mode 100644 index 000000000..ec605bba1 --- /dev/null +++ b/Cubical/Structures/Constant.agda @@ -0,0 +1,35 @@ +{- + +Constant structure: _ ↦ A + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Constant where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.Foundations.SIP + +private + variable + ℓ ℓ' : Level + +-- Structured isomorphisms + +module _ (A : Type ℓ') where + + ConstantStructure : Type ℓ → Type ℓ' + ConstantStructure _ = A + + ConstantEquivStr : StrEquiv {ℓ} ConstantStructure ℓ' + ConstantEquivStr (_ , a) (_ , a') _ = a ≡ a' + + constantUnivalentStr : UnivalentStr {ℓ} ConstantStructure ConstantEquivStr + constantUnivalentStr e = idEquiv _ + + constantEquivAction : EquivAction {ℓ} ConstantStructure + constantEquivAction e = idEquiv _ + + constantTransportStr : TransportStr {ℓ} constantEquivAction + constantTransportStr e _ = sym (transportRefl _) diff --git a/Cubical/Structures/Everything.agda b/Cubical/Structures/Everything.agda new file mode 100644 index 000000000..9b03cf61f --- /dev/null +++ b/Cubical/Structures/Everything.agda @@ -0,0 +1,27 @@ +{-# OPTIONS --safe --guardedness #-} +module Cubical.Structures.Everything where + +import Cubical.Structures.Auto +import Cubical.Structures.Axioms +import Cubical.Structures.Constant +import Cubical.Structures.Function +import Cubical.Structures.LeftAction +import Cubical.Structures.Macro +import Cubical.Structures.Maybe +import Cubical.Structures.MultiSet +import Cubical.Structures.Parameterized +import Cubical.Structures.Pointed +import Cubical.Structures.Product +import Cubical.Structures.Queue +import Cubical.Structures.Record +import Cubical.Structures.Relational.Auto +import Cubical.Structures.Relational.Constant +import Cubical.Structures.Relational.Equalizer +import Cubical.Structures.Relational.Function +import Cubical.Structures.Relational.Macro +import Cubical.Structures.Relational.Maybe +import Cubical.Structures.Relational.Parameterized +import Cubical.Structures.Relational.Pointed +import Cubical.Structures.Relational.Product +import Cubical.Structures.Transfer +import Cubical.Structures.TypeEqvTo diff --git a/Cubical/Structures/Function.agda b/Cubical/Structures/Function.agda new file mode 100644 index 000000000..4371e4c01 --- /dev/null +++ b/Cubical/Structures/Function.agda @@ -0,0 +1,82 @@ +{- + +Functions between structures S and T: X ↦ S X → T X + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Function where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence + +open import Cubical.Functions.FunExtEquiv + +open import Cubical.Data.Nat +open import Cubical.Data.Vec + +private + variable + ℓ ℓ₁ ℓ₁' ℓ₂ ℓ₂' : Level + +-- General function structure + +FunctionStructure : (S : Type ℓ → Type ℓ₁) (T : Type ℓ → Type ℓ₂) + → Type ℓ → Type (ℓ-max ℓ₁ ℓ₂) +FunctionStructure S T X = S X → T X + +FunctionEquivStr : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + → StrEquiv S ℓ₁' → StrEquiv T ℓ₂' + → StrEquiv (FunctionStructure S T) (ℓ-max ℓ₁ (ℓ-max ℓ₁' ℓ₂')) +FunctionEquivStr {S = S} {T} ι₁ ι₂ (X , f) (Y , g) e = + {s : S X} {t : S Y} → ι₁ (X , s) (Y , t) e → ι₂ (X , f s) (Y , g t) e + +functionUnivalentStr : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + (ι₁ : StrEquiv S ℓ₁') (θ₁ : UnivalentStr S ι₁) + (ι₂ : StrEquiv T ℓ₂') (θ₂ : UnivalentStr T ι₂) + → UnivalentStr (FunctionStructure S T) (FunctionEquivStr ι₁ ι₂) +functionUnivalentStr ι₁ θ₁ ι₂ θ₂ e = + compEquiv + (equivImplicitΠCod (equivImplicitΠCod (equiv→ (θ₁ e) (θ₂ e)))) + funExtDepEquiv + +functionEquivAction : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + → EquivAction S → EquivAction T + → EquivAction (FunctionStructure S T) +functionEquivAction α₁ α₂ e = equiv→ (α₁ e) (α₂ e) + +functionTransportStr : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + (α₁ : EquivAction S) (τ₁ : TransportStr α₁) + (α₂ : EquivAction T) (τ₂ : TransportStr α₂) + → TransportStr (functionEquivAction α₁ α₂) +functionTransportStr {S = S} α₁ τ₁ α₂ τ₂ e f = + funExt λ t → + cong (equivFun (α₂ e) ∘ f) (invTransportStr α₁ τ₁ e t) + ∙ τ₂ e (f (subst⁻ S (ua e) t)) + +-- Definition of structured equivalence using an action in the domain + +FunctionEquivStr+ : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + → EquivAction S → StrEquiv T ℓ₂' + → StrEquiv (FunctionStructure S T) (ℓ-max ℓ₁ ℓ₂') +FunctionEquivStr+ {S = S} {T} α₁ ι₂ (X , f) (Y , g) e = + (s : S X) → ι₂ (X , f s) (Y , g (equivFun (α₁ e) s)) e + +functionUnivalentStr+ : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + (α₁ : EquivAction S) (τ₁ : TransportStr α₁) + (ι₂ : StrEquiv T ℓ₂') (θ₂ : UnivalentStr T ι₂) + → UnivalentStr (FunctionStructure S T) (FunctionEquivStr+ α₁ ι₂) +functionUnivalentStr+ {S = S} {T} α₁ τ₁ ι₂ θ₂ {X , f} {Y , g} e = + compEquiv + (compEquiv + (equivΠCod λ s → + compEquiv + (θ₂ e) + (pathToEquiv (cong (PathP (λ i → T (ua e i)) (f s) ∘ g) (τ₁ e s)))) + (invEquiv heteroHomotopy≃Homotopy)) + funExtDepEquiv diff --git a/Cubical/Structures/LeftAction.agda b/Cubical/Structures/LeftAction.agda new file mode 100644 index 000000000..59cf68aa8 --- /dev/null +++ b/Cubical/Structures/LeftAction.agda @@ -0,0 +1,17 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.LeftAction where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.SIP + +open import Cubical.Structures.Auto + +module _ {ℓ ℓ' : Level} (A : Type ℓ') where + + LeftActionStructure : Type ℓ → Type (ℓ-max ℓ ℓ') + LeftActionStructure X = A → X → X + + LeftActionEquivStr = AutoEquivStr LeftActionStructure + + leftActionUnivalentStr : UnivalentStr _ LeftActionEquivStr + leftActionUnivalentStr = autoUnivalentStr LeftActionStructure diff --git a/Cubical/Structures/Macro.agda b/Cubical/Structures/Macro.agda new file mode 100644 index 000000000..d80e785a7 --- /dev/null +++ b/Cubical/Structures/Macro.agda @@ -0,0 +1,157 @@ +{- + +Descriptor language for easily defining structures + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Macro where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.SIP +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Sigma +open import Cubical.Data.Maybe + +open import Cubical.Structures.Axioms +open import Cubical.Structures.Constant +open import Cubical.Structures.Function +open import Cubical.Structures.Maybe +open import Cubical.Structures.Parameterized +open import Cubical.Structures.Pointed +open import Cubical.Structures.Product + +{- Transport structures -} + +data TranspDesc (ℓ : Level) : Level → Typeω where + -- constant structure: X ↦ A + constant : ∀ {ℓ₁} (A : Type ℓ₁) → TranspDesc ℓ ℓ₁ + -- pointed structure: X ↦ X + var : TranspDesc ℓ ℓ + -- product of structures S,T : X ↦ (S X × T X) + _,_ : ∀ {ℓ₁ ℓ₂} (d₀ : TranspDesc ℓ ℓ₁) (d₁ : TranspDesc ℓ ℓ₂) → TranspDesc ℓ (ℓ-max ℓ₁ ℓ₂) + -- functions between structures S,T: X ↦ (S X → T X) + function : ∀ {ℓ₁ ℓ₂} (d₀ : TranspDesc ℓ ℓ₁) (d₁ : TranspDesc ℓ ℓ₂) → TranspDesc ℓ (ℓ-max ℓ₁ ℓ₂) + -- Maybe on a structure S: X ↦ Maybe (S X) + maybe : ∀ {ℓ₁} → TranspDesc ℓ ℓ₁ → TranspDesc ℓ ℓ₁ + -- arbitrary transport structure + foreign : ∀ {ℓ₁} {S : Type ℓ → Type ℓ₁} (α : EquivAction S) → TransportStr α → TranspDesc ℓ ℓ₁ + +-- Structure defined by a transport descriptor +TranspMacroStructure : ∀ {ℓ ℓ₁} → TranspDesc ℓ ℓ₁ → Type ℓ → Type ℓ₁ +TranspMacroStructure (constant A) X = A +TranspMacroStructure var X = X +TranspMacroStructure (d₀ , d₁) X = TranspMacroStructure d₀ X × TranspMacroStructure d₁ X +TranspMacroStructure (function d₀ d₁) X = TranspMacroStructure d₀ X → TranspMacroStructure d₁ X +TranspMacroStructure (maybe d) = MaybeStructure (TranspMacroStructure d) +TranspMacroStructure (foreign {S = S} α τ) = S + +-- Action defined by a transport descriptor +transpMacroAction : ∀ {ℓ ℓ₁} (d : TranspDesc ℓ ℓ₁) → EquivAction (TranspMacroStructure d) +transpMacroAction (constant A) = constantEquivAction A +transpMacroAction var = pointedEquivAction +transpMacroAction (d₀ , d₁) = productEquivAction (transpMacroAction d₀) (transpMacroAction d₁) +transpMacroAction (function d₀ d₁) = + functionEquivAction (transpMacroAction d₀) (transpMacroAction d₁) +transpMacroAction (maybe d) = maybeEquivAction (transpMacroAction d) +transpMacroAction (foreign α _) = α + +-- Action defines a transport structure +transpMacroTransportStr : ∀ {ℓ ℓ₁} (d : TranspDesc ℓ ℓ₁) → TransportStr (transpMacroAction d) +transpMacroTransportStr (constant A) = constantTransportStr A +transpMacroTransportStr var = pointedTransportStr +transpMacroTransportStr (d₀ , d₁) = + productTransportStr + (transpMacroAction d₀) (transpMacroTransportStr d₀) + (transpMacroAction d₁) (transpMacroTransportStr d₁) +transpMacroTransportStr (function d₀ d₁) = + functionTransportStr + (transpMacroAction d₀) (transpMacroTransportStr d₀) + (transpMacroAction d₁) (transpMacroTransportStr d₁) +transpMacroTransportStr (maybe d) = + maybeTransportStr (transpMacroAction d) (transpMacroTransportStr d) +transpMacroTransportStr (foreign α τ) = τ + +{- General structures -} + +mutual + data Desc (ℓ : Level) : Level → Level → Typeω where + -- constant structure: X ↦ A + constant : ∀ {ℓ₁} (A : Type ℓ₁) → Desc ℓ ℓ₁ ℓ₁ + -- pointed structure: X ↦ X + var : Desc ℓ ℓ ℓ + -- product of structures S,T : X ↦ (S X × T X) + _,_ : ∀ {ℓ₁ ℓ₁' ℓ₂ ℓ₂'} (d₀ : Desc ℓ ℓ₁ ℓ₁') (d₁ : Desc ℓ ℓ₂ ℓ₂') + → Desc ℓ (ℓ-max ℓ₁ ℓ₂) (ℓ-max ℓ₁' ℓ₂') + -- functions between structures S,T : X ↦ (S X → T X) + function : ∀ {ℓ₁ ℓ₁' ℓ₂ ℓ₂'} (d₀ : Desc ℓ ℓ₁ ℓ₁') (d₁ : Desc ℓ ℓ₂ ℓ₂') + → Desc ℓ (ℓ-max ℓ₁ ℓ₂) (ℓ-max ℓ₁ (ℓ-max ℓ₁' ℓ₂')) + -- functions between structures S,T where S is functorial : X ↦ (S X → T X) + function+ : ∀ {ℓ₁ ℓ₂ ℓ₂'} (d₀ : TranspDesc ℓ ℓ₁) (d₁ : Desc ℓ ℓ₂ ℓ₂') → Desc ℓ (ℓ-max ℓ₁ ℓ₂) (ℓ-max ℓ₁ ℓ₂') + -- Maybe on a structure S: X ↦ Maybe (S X) + maybe : ∀ {ℓ₁ ℓ₁'} → Desc ℓ ℓ₁ ℓ₁' → Desc ℓ ℓ₁ ℓ₁' + -- add axioms to a structure + axioms : ∀ {ℓ₁ ℓ₁' ℓ₂} (d : Desc ℓ ℓ₁ ℓ₁') (ax : ∀ X → MacroStructure d X → Type ℓ₂) + → (∀ X s → isProp (ax X s)) + → Desc ℓ (ℓ-max ℓ₁ ℓ₂) ℓ₁' + -- univalent structure from transport structure + transpDesc : ∀ {ℓ₁} → TranspDesc ℓ ℓ₁ → Desc ℓ ℓ₁ ℓ₁ + -- arbitrary univalent notion of structure + foreign : ∀ {ℓ₁ ℓ₁'} {S : Type ℓ → Type ℓ₁} (ι : StrEquiv S ℓ₁') → UnivalentStr S ι → Desc ℓ ℓ₁ ℓ₁' + + infixr 4 _,_ + + -- Structure defined by a descriptor + MacroStructure : ∀ {ℓ ℓ₁ ℓ₁'} → Desc ℓ ℓ₁ ℓ₁' → Type ℓ → Type ℓ₁ + MacroStructure (constant A) X = A + MacroStructure var X = X + MacroStructure (d₀ , d₁) X = MacroStructure d₀ X × MacroStructure d₁ X + MacroStructure (function+ d₀ d₁) X = TranspMacroStructure d₀ X → MacroStructure d₁ X + MacroStructure (function d₀ d₁) X = MacroStructure d₀ X → MacroStructure d₁ X + MacroStructure (maybe d) = MaybeStructure (MacroStructure d) + MacroStructure (axioms d ax _) = AxiomsStructure (MacroStructure d) ax + MacroStructure (transpDesc d) = TranspMacroStructure d + MacroStructure (foreign {S = S} _ _) = S + +-- Notion of structured equivalence defined by a descriptor +MacroEquivStr : ∀ {ℓ ℓ₁ ℓ₁'} → (d : Desc ℓ ℓ₁ ℓ₁') → StrEquiv (MacroStructure d) ℓ₁' +MacroEquivStr (constant A) = ConstantEquivStr A +MacroEquivStr var = PointedEquivStr +MacroEquivStr (d₀ , d₁) = ProductEquivStr (MacroEquivStr d₀) (MacroEquivStr d₁) +MacroEquivStr (function+ d₀ d₁) = FunctionEquivStr+ (transpMacroAction d₀) (MacroEquivStr d₁) +MacroEquivStr (function d₀ d₁) = FunctionEquivStr (MacroEquivStr d₀) (MacroEquivStr d₁) +MacroEquivStr (maybe d) = MaybeEquivStr (MacroEquivStr d) +MacroEquivStr (axioms d ax _) = AxiomsEquivStr (MacroEquivStr d) ax +MacroEquivStr (transpDesc d) = EquivAction→StrEquiv (transpMacroAction d) +MacroEquivStr (foreign ι _) = ι + +-- Proof that structure induced by descriptor is univalent +MacroUnivalentStr : ∀ {ℓ ℓ₁ ℓ₁'} → (d : Desc ℓ ℓ₁ ℓ₁') → UnivalentStr (MacroStructure d) (MacroEquivStr d) +MacroUnivalentStr (constant A) = constantUnivalentStr A +MacroUnivalentStr var = pointedUnivalentStr +MacroUnivalentStr (d₀ , d₁) = + productUnivalentStr + (MacroEquivStr d₀) (MacroUnivalentStr d₀) + (MacroEquivStr d₁) (MacroUnivalentStr d₁) +MacroUnivalentStr (function+ d₀ d₁) = + functionUnivalentStr+ + (transpMacroAction d₀) (transpMacroTransportStr d₀) + (MacroEquivStr d₁) (MacroUnivalentStr d₁) +MacroUnivalentStr (function d₀ d₁) = + functionUnivalentStr + (MacroEquivStr d₀) (MacroUnivalentStr d₀) + (MacroEquivStr d₁) (MacroUnivalentStr d₁) +MacroUnivalentStr (maybe d) = maybeUnivalentStr (MacroEquivStr d) (MacroUnivalentStr d) +MacroUnivalentStr (axioms d _ isPropAx) = axiomsUnivalentStr (MacroEquivStr d) isPropAx (MacroUnivalentStr d) +MacroUnivalentStr (transpDesc d) = + TransportStr→UnivalentStr (transpMacroAction d) (transpMacroTransportStr d) +MacroUnivalentStr (foreign _ θ) = θ + +-- Module for easy importing +module Macro ℓ {ℓ₁ ℓ₁'} (d : Desc ℓ ℓ₁ ℓ₁') where + + structure = MacroStructure d + equiv = MacroEquivStr d + univalent = MacroUnivalentStr d diff --git a/Cubical/Structures/Maybe.agda b/Cubical/Structures/Maybe.agda new file mode 100644 index 000000000..7198a3a12 --- /dev/null +++ b/Cubical/Structures/Maybe.agda @@ -0,0 +1,108 @@ +{- + + Maybe structure: X ↦ Maybe (S X) + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Maybe where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.SIP +open import Cubical.Functions.FunExtEquiv + +open import Cubical.Data.Unit +open import Cubical.Data.Empty +open import Cubical.Data.Maybe + +private + variable + ℓ ℓ₁ ℓ₁' : Level + +MaybeRel : {A B : Type ℓ} (R : A → B → Type ℓ₁) → Maybe A → Maybe B → Type ℓ₁ +MaybeRel R nothing nothing = Lift Unit +MaybeRel R nothing (just _) = Lift ⊥ +MaybeRel R (just _) nothing = Lift ⊥ +MaybeRel R (just x) (just y) = R x y + +congMaybeRel : {A B : Type ℓ} {R : A → B → Type ℓ₁} {S : A → B → Type ℓ₁'} + → (∀ x y → R x y ≃ S x y) + → ∀ ox oy → MaybeRel R ox oy ≃ MaybeRel S ox oy +congMaybeRel e nothing nothing = Lift≃Lift (idEquiv _) +congMaybeRel e nothing (just _) = Lift≃Lift (idEquiv _) +congMaybeRel e (just _) nothing = Lift≃Lift (idEquiv _) +congMaybeRel e (just x) (just y) = e x y + +module MaybePathP where + + Code : (A : I → Type ℓ) → Maybe (A i0) → Maybe (A i1) → Type ℓ + Code A = MaybeRel (PathP A) + + encodeRefl : {A : Type ℓ} → ∀ ox → Code (λ _ → A) ox ox + encodeRefl nothing = lift tt + encodeRefl (just _) = refl + + encode : (A : I → Type ℓ) → ∀ ox oy → PathP (λ i → Maybe (A i)) ox oy → Code A ox oy + encode A ox oy p = transport (λ j → Code (λ i → A (i ∧ j)) ox (p j)) (encodeRefl ox) + + decode : {A : I → Type ℓ} → ∀ ox oy → Code A ox oy → PathP (λ i → Maybe (A i)) ox oy + decode nothing nothing p i = nothing + decode (just _) (just _) p i = just (p i) + + decodeEncodeRefl : {A : Type ℓ} (ox : Maybe A) → decode ox ox (encodeRefl ox) ≡ refl + decodeEncodeRefl nothing = refl + decodeEncodeRefl (just _) = refl + + decodeEncode : {A : I → Type ℓ} → ∀ ox oy p → decode ox oy (encode A ox oy p) ≡ p + decodeEncode {A = A} ox oy p = + transport + (λ k → + decode _ _ (transp (λ j → Code (λ i → A (i ∧ j ∧ k)) ox (p (j ∧ k))) (~ k) (encodeRefl ox)) + ≡ (λ i → p (i ∧ k))) + (decodeEncodeRefl ox) + + encodeDecode : (A : I → Type ℓ) → ∀ ox oy c → encode A ox oy (decode ox oy c) ≡ c + encodeDecode A nothing nothing c = refl + encodeDecode A (just x) (just y) c = + transport + (λ k → + encode (λ i → A (i ∧ k)) _ _ (decode (just x) (just (c k)) (λ i → c (i ∧ k))) + ≡ (λ i → c (i ∧ k))) + (transportRefl _) + + Code≃PathP : {A : I → Type ℓ} → ∀ ox oy → Code A ox oy ≃ PathP (λ i → Maybe (A i)) ox oy + Code≃PathP {A = A} ox oy = isoToEquiv isom + where + isom : Iso _ _ + isom .Iso.fun = decode ox oy + isom .Iso.inv = encode _ ox oy + isom .Iso.rightInv = decodeEncode ox oy + isom .Iso.leftInv = encodeDecode A ox oy + +-- Structured isomorphisms + +MaybeStructure : (S : Type ℓ → Type ℓ₁) → Type ℓ → Type ℓ₁ +MaybeStructure S X = Maybe (S X) + +MaybeEquivStr : {S : Type ℓ → Type ℓ₁} + → StrEquiv S ℓ₁' → StrEquiv (MaybeStructure S) ℓ₁' +MaybeEquivStr ι (X , ox) (Y , oy) e = MaybeRel (λ x y → ι (X , x) (Y , y) e) ox oy + +maybeUnivalentStr : {S : Type ℓ → Type ℓ₁} (ι : StrEquiv S ℓ₁') + → UnivalentStr S ι → UnivalentStr (MaybeStructure S) (MaybeEquivStr ι) +maybeUnivalentStr ι θ {X , ox} {Y , oy} e = + compEquiv + (congMaybeRel (λ x y → θ {X , x} {Y , y} e) ox oy) + (MaybePathP.Code≃PathP ox oy) + +maybeEquivAction : {S : Type ℓ → Type ℓ₁} + → EquivAction S → EquivAction (MaybeStructure S) +maybeEquivAction α e = congMaybeEquiv (α e) + +maybeTransportStr : {S : Type ℓ → Type ℓ₁} (α : EquivAction S) + → TransportStr α → TransportStr (maybeEquivAction α) +maybeTransportStr _ τ e nothing = refl +maybeTransportStr _ τ e (just x) = cong just (τ e x) diff --git a/Cubical/Structures/MultiSet.agda b/Cubical/Structures/MultiSet.agda new file mode 100644 index 000000000..6352a2f58 --- /dev/null +++ b/Cubical/Structures/MultiSet.agda @@ -0,0 +1,42 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.MultiSet where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.SIP +open import Cubical.Functions.FunExtEquiv + +open import Cubical.Structures.Auto + +open import Cubical.Data.Nat +open import Cubical.Data.Sigma + +private + variable + ℓ : Level + +module _ (A : Type ℓ) (Aset : isSet A) where + + CountStructure : Type ℓ → Type ℓ + CountStructure X = A → X → ℕ + + CountEquivStr = AutoEquivStr CountStructure + + countUnivalentStr : UnivalentStr _ CountEquivStr + countUnivalentStr = autoUnivalentStr CountStructure + + Count : Type (ℓ-suc ℓ) + Count = TypeWithStr ℓ CountStructure + + MultiSetStructure : Type ℓ → Type ℓ + MultiSetStructure X = X × (A → X → X) × (A → X → ℕ) + + MultiSetEquivStr = AutoEquivStr MultiSetStructure + + multiSetUnivalentStr : UnivalentStr _ MultiSetEquivStr + multiSetUnivalentStr = autoUnivalentStr MultiSetStructure + + MultiSet : Type (ℓ-suc ℓ) + MultiSet = TypeWithStr ℓ MultiSetStructure diff --git a/Cubical/Structures/Parameterized.agda b/Cubical/Structures/Parameterized.agda new file mode 100644 index 000000000..04a189d03 --- /dev/null +++ b/Cubical/Structures/Parameterized.agda @@ -0,0 +1,48 @@ +{- + +A parameterized family of structures S can be combined into a single structure: +X ↦ (a : A) → S a X + +This is more general than Structures.Function in that S can vary in A. + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Parameterized where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Functions.FunExtEquiv +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Univalence + +private + variable + ℓ ℓ₁ ℓ₁' : Level + +module _ {ℓ₀} (A : Type ℓ₀) where + + ParamStructure : (S : A → Type ℓ → Type ℓ₁) + → Type ℓ → Type (ℓ-max ℓ₀ ℓ₁) + ParamStructure S X = (a : A) → S a X + + ParamEquivStr : {S : A → Type ℓ → Type ℓ₁} + → (∀ a → StrEquiv (S a) ℓ₁') → StrEquiv (ParamStructure S) (ℓ-max ℓ₀ ℓ₁') + ParamEquivStr ι (X , l) (Y , m) e = ∀ a → ι a (X , l a) (Y , m a) e + + paramUnivalentStr : {S : A → Type ℓ → Type ℓ₁} + (ι : ∀ a → StrEquiv (S a) ℓ₁') (θ : ∀ a → UnivalentStr (S a) (ι a)) + → UnivalentStr (ParamStructure S) (ParamEquivStr ι) + paramUnivalentStr ι θ e = compEquiv (equivΠCod λ a → θ a e) funExtEquiv + + paramEquivAction : {S : A → Type ℓ → Type ℓ₁} + → (∀ a → EquivAction (S a)) → EquivAction (ParamStructure S) + paramEquivAction α e = equivΠCod (λ a → α a e) + + paramTransportStr : {S : A → Type ℓ → Type ℓ₁} + (α : ∀ a → EquivAction (S a)) (τ : ∀ a → TransportStr (α a)) + → TransportStr (paramEquivAction α) + paramTransportStr {S = S} α τ e f = + funExt λ a → + τ a e (f a) + ∙ cong (λ fib → transport (λ i → S (fib .snd (~ i)) (ua e i)) (f (fib .snd i1))) + (isContrSingl a .snd (_ , sym (transportRefl a))) diff --git a/Cubical/Structures/Pointed.agda b/Cubical/Structures/Pointed.agda new file mode 100644 index 000000000..ea6c76b1e --- /dev/null +++ b/Cubical/Structures/Pointed.agda @@ -0,0 +1,41 @@ +{- + +Pointed structure: X ↦ X + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Pointed where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.SIP + +open import Cubical.Foundations.Pointed.Base + +private + variable + ℓ : Level + +-- Structured isomorphisms + +PointedStructure : Type ℓ → Type ℓ +PointedStructure X = X + +PointedEquivStr : StrEquiv PointedStructure ℓ +PointedEquivStr A B f = equivFun f (pt A) ≡ pt B + +pointedUnivalentStr : UnivalentStr {ℓ} PointedStructure PointedEquivStr +pointedUnivalentStr f = invEquiv (ua-ungluePath-Equiv f) + +pointedSIP : (A B : Pointed ℓ) → A ≃[ PointedEquivStr ] B ≃ (A ≡ B) +pointedSIP = SIP pointedUnivalentStr + +pointed-sip : (A B : Pointed ℓ) → A ≃[ PointedEquivStr ] B → (A ≡ B) +pointed-sip A B = equivFun (pointedSIP A B) -- ≡ λ (e , p) i → ua e i , ua-gluePath e p i + +pointedEquivAction : EquivAction {ℓ} PointedStructure +pointedEquivAction e = e + +pointedTransportStr : TransportStr {ℓ} pointedEquivAction +pointedTransportStr e s = sym (transportRefl _) diff --git a/Cubical/Structures/Product.agda b/Cubical/Structures/Product.agda new file mode 100644 index 000000000..de2ddba3d --- /dev/null +++ b/Cubical/Structures/Product.agda @@ -0,0 +1,50 @@ +{- + +Product of structures S and T: X ↦ S X × T X + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Product where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Path +open import Cubical.Foundations.SIP +open import Cubical.Data.Sigma + +private + variable + ℓ ℓ₁ ℓ₁' ℓ₂ ℓ₂' : Level + +ProductStructure : (S₁ : Type ℓ → Type ℓ₁) (S₂ : Type ℓ → Type ℓ₂) + → Type ℓ → Type (ℓ-max ℓ₁ ℓ₂) +ProductStructure S₁ S₂ X = S₁ X × S₂ X + +ProductEquivStr : + {S₁ : Type ℓ → Type ℓ₁} (ι₁ : StrEquiv S₁ ℓ₁') + {S₂ : Type ℓ → Type ℓ₂} (ι₂ : StrEquiv S₂ ℓ₂') + → StrEquiv (ProductStructure S₁ S₂) (ℓ-max ℓ₁' ℓ₂') +ProductEquivStr ι₁ ι₂ (X , s₁ , s₂) (Y , t₁ , t₂) f = + (ι₁ (X , s₁) (Y , t₁) f) × (ι₂ (X , s₂) (Y , t₂) f) + +productUnivalentStr : + {S₁ : Type ℓ → Type ℓ₁} (ι₁ : StrEquiv S₁ ℓ₁') (θ₁ : UnivalentStr S₁ ι₁) + {S₂ : Type ℓ → Type ℓ₂} (ι₂ : StrEquiv S₂ ℓ₂') (θ₂ : UnivalentStr S₂ ι₂) + → UnivalentStr (ProductStructure S₁ S₂) (ProductEquivStr ι₁ ι₂) +productUnivalentStr {S₁ = S₁} ι₁ θ₁ {S₂} ι₂ θ₂ {X , s₁ , s₂} {Y , t₁ , t₂} e = + compEquiv (Σ-cong-equiv (θ₁ e) (λ _ → θ₂ e)) ΣPath≃PathΣ + +productEquivAction : + {S₁ : Type ℓ → Type ℓ₁} (α₁ : EquivAction S₁) + {S₂ : Type ℓ → Type ℓ₂} (α₂ : EquivAction S₂) + → EquivAction (ProductStructure S₁ S₂) +productEquivAction α₁ α₂ e = Σ-cong-equiv (α₁ e) (λ _ → α₂ e) + +productTransportStr : + {S₁ : Type ℓ → Type ℓ₁} (α₁ : EquivAction S₁) (τ₁ : TransportStr α₁) + {S₂ : Type ℓ → Type ℓ₂} (α₂ : EquivAction S₂) (τ₂ : TransportStr α₂) + → TransportStr (productEquivAction α₁ α₂) +productTransportStr _ τ₁ _ τ₂ e (s₁ , s₂) = ΣPathP (τ₁ e s₁ , τ₂ e s₂) diff --git a/Cubical/Structures/Queue.agda b/Cubical/Structures/Queue.agda new file mode 100644 index 000000000..567cf70da --- /dev/null +++ b/Cubical/Structures/Queue.agda @@ -0,0 +1,115 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Queue where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Functions.FunExtEquiv +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.SIP + +open import Cubical.Structures.Axioms +open import Cubical.Structures.Macro +open import Cubical.Structures.Auto + +open import Cubical.Data.Unit +open import Cubical.Data.Maybe as Maybe +open import Cubical.Data.Sigma +open import Cubical.Data.List + + +-- Developing Queues as a standard notion of structure, see +-- https://github.com/ecavallo/cubical/blob/queue/Cubical/Experiments/Queue.agda +-- for the original development + +private + variable + ℓ ℓ' : Level + +-- We start fixing a set A on which we define what it means for a type Q to have +-- a Queue structure (wrt. A) +module Queues-on (A : Type ℓ) (Aset : isSet A) where + -- A Queue structure has three components, the empty Queue, an enqueue function and a dequeue function + -- We first deal with enq and deq as separate structures + + -- deq as a structure + -- First, a few preliminary results that we will need later + deqMap : {X Y : Type ℓ} → (X → Y) → Maybe (X × A) → Maybe (Y × A) + deqMap = map-Maybe ∘ map-fst + + deqMapId : {X : Type ℓ} → ∀ r → deqMap (idfun X) r ≡ r + deqMapId = map-Maybe-id + + deqMap-∘ :{B C D : Type ℓ} + (g : C → D) (f : B → C) + → ∀ r → deqMap {X = C} g (deqMap f r) ≡ deqMap (λ b → g (f b)) r + deqMap-∘ g f nothing = refl + deqMap-∘ g f (just (b , a)) = refl + + -- Now we can do Queues: + rawQueueDesc = + autoDesc (λ (X : Type ℓ) → X × (A → X → X) × (X → Transp[ Maybe (X × A) ])) + + open Macro ℓ rawQueueDesc public renaming + ( structure to RawQueueStructure + ; equiv to RawQueueEquivStr + ; univalent to rawQueueUnivalentStr + ) + + RawQueue : Type (ℓ-suc ℓ) + RawQueue = TypeWithStr ℓ RawQueueStructure + + returnOrEnq : {Q : Type ℓ} + → RawQueueStructure Q → A → Maybe (Q × A) → Q × A + returnOrEnq (emp , enq , _) a qr = + Maybe.rec (emp , a) (λ {(q , b) → enq a q , b}) qr + + QueueAxioms : (Q : Type ℓ) → RawQueueStructure Q → Type ℓ + QueueAxioms Q S@(emp , enq , deq) = + (isSet Q) + × (deq emp ≡ nothing) + × (∀ a q → deq (enq a q) ≡ just (returnOrEnq S a (deq q))) + × (∀ a a' q q' → enq a q ≡ enq a' q' → (a ≡ a') × (q ≡ q')) + × (∀ q q' → deq q ≡ deq q' → q ≡ q') + + isPropQueueAxioms : ∀ Q S → isProp (QueueAxioms Q S) + isPropQueueAxioms Q S = + isPropΣ isPropIsSet + (λ Qset → isProp×3 (isOfHLevelDeq Qset _ _) + (isPropΠ2 λ _ _ → isOfHLevelDeq Qset _ _) + (isPropΠ3 λ _ _ _ → isPropΠ2 λ _ _ → isProp× (Aset _ _) (Qset _ _)) + (isPropΠ3 λ _ _ _ → Qset _ _)) + where + isOfHLevelDeq : isSet Q → isOfHLevel 2 (Maybe (Q × A)) + isOfHLevelDeq Qset = isOfHLevelMaybe 0 (isSet× Qset Aset) + + QueueStructure : Type ℓ → Type ℓ + QueueStructure = AxiomsStructure RawQueueStructure QueueAxioms + + Queue : Type (ℓ-suc ℓ) + Queue = TypeWithStr ℓ QueueStructure + + QueueEquivStr : StrEquiv QueueStructure ℓ + QueueEquivStr = AxiomsEquivStr RawQueueEquivStr QueueAxioms + + queueUnivalentStr : UnivalentStr QueueStructure QueueEquivStr + queueUnivalentStr = axiomsUnivalentStr RawQueueEquivStr isPropQueueAxioms rawQueueUnivalentStr + + + FiniteQueueAxioms : (Q : Type ℓ) → QueueStructure Q → Type ℓ + FiniteQueueAxioms Q ((emp , enq , _) , _) = isEquiv (foldr enq emp) + + isPropFiniteQueueAxioms : ∀ Q S → isProp (FiniteQueueAxioms Q S) + isPropFiniteQueueAxioms Q S = isPropIsEquiv _ + + FiniteQueueStructure : Type ℓ → Type ℓ + FiniteQueueStructure = AxiomsStructure QueueStructure FiniteQueueAxioms + + FiniteQueue : Type (ℓ-suc ℓ) + FiniteQueue = TypeWithStr ℓ FiniteQueueStructure + + FiniteQueueEquivStr : StrEquiv FiniteQueueStructure ℓ + FiniteQueueEquivStr = AxiomsEquivStr QueueEquivStr FiniteQueueAxioms + + finiteQueueUnivalentStr : UnivalentStr FiniteQueueStructure FiniteQueueEquivStr + finiteQueueUnivalentStr = axiomsUnivalentStr QueueEquivStr isPropFiniteQueueAxioms queueUnivalentStr diff --git a/Cubical/Structures/Record.agda b/Cubical/Structures/Record.agda new file mode 100644 index 000000000..ce26b8996 --- /dev/null +++ b/Cubical/Structures/Record.agda @@ -0,0 +1,453 @@ +{- + +Automatically generating proofs of UnivalentStr for records + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Record where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Univalence +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.List as List +open import Cubical.Data.Vec as Vec +open import Cubical.Data.Bool +open import Cubical.Data.Maybe +open import Cubical.Data.Sum +open import Cubical.Structures.Auto +import Cubical.Structures.Macro as M +import Agda.Builtin.Reflection as R +open import Cubical.Reflection.Base + +-- Magic number +private + FUEL = 10000 + +-- Types for specifying inputs to the tactics + +data AutoFieldSpec : Typeω where + autoFieldSpec : ∀ {ℓ ℓ₁ ℓ₂} (R : Type ℓ → Type ℓ₁) {S : Type ℓ → Type ℓ₂} + → ({X : Type ℓ} → R X → S X) + → AutoFieldSpec + +module _ {ℓ ℓ₁ ℓ₁'} where + mutual + data AutoFields (R : Type ℓ → Type ℓ₁) (ι : StrEquiv R ℓ₁') : Typeω + where + fields: : AutoFields R ι + _data[_∣_] : (fs : AutoFields R ι) + → ∀ {ℓ₂ ℓ₂'} {S : Type ℓ → Type ℓ₂} {ι' : StrEquiv S ℓ₂'} + → (f : {X : Type ℓ} → R X → S X) + → ({A B : TypeWithStr ℓ R} {e : typ A ≃ typ B} → ι A B e → ι' (map-snd f A) (map-snd f B) e) + → AutoFields R ι + _prop[_∣_] : (fs : AutoFields R ι) + → ∀ {ℓ₂} {P : (X : Type ℓ) → GatherFields fs X → Type ℓ₂} + → ({X : Type ℓ} (r : R X) → P X (projectFields fs r)) + → isPropProperty R ι fs P + → AutoFields R ι + + GatherFieldsLevel : {R : Type ℓ → Type ℓ₁} {ι : StrEquiv R ℓ₁'} + → AutoFields R ι + → Level + GatherFieldsLevel fields: = ℓ-zero + GatherFieldsLevel (_data[_∣_] fs {ℓ₂ = ℓ₂} _ _) = ℓ-max (GatherFieldsLevel fs) ℓ₂ + GatherFieldsLevel (_prop[_∣_] fs {ℓ₂ = ℓ₂} _ _) = ℓ-max (GatherFieldsLevel fs) ℓ₂ + + GatherFields : {R : Type ℓ → Type ℓ₁} {ι : StrEquiv R ℓ₁'} + (dat : AutoFields R ι) + → Type ℓ → Type (GatherFieldsLevel dat) + GatherFields fields: X = Unit + GatherFields (_data[_∣_] fs {S = S} _ _) X = GatherFields fs X × S X + GatherFields (_prop[_∣_] fs {P = P} _ _) X = + Σ[ s ∈ GatherFields fs X ] (P X s) + + projectFields : {R : Type ℓ → Type ℓ₁} {ι : StrEquiv R ℓ₁'} + (fs : AutoFields R ι) + → {X : Type ℓ} → R X → GatherFields fs X + projectFields fields: = _ + projectFields (fs data[ f ∣ _ ]) r = projectFields fs r , f r + projectFields (fs prop[ f ∣ _ ]) r = projectFields fs r , f r + + isPropProperty : ∀ {ℓ₂} (R : Type ℓ → Type ℓ₁) + (ι : StrEquiv R ℓ₁') + (fs : AutoFields R ι) + (P : (X : Type ℓ) → GatherFields fs X → Type ℓ₂) + → Type (ℓ-max (ℓ-suc ℓ) (ℓ-max ℓ₁ ℓ₂)) + isPropProperty R ι fs P = + {X : Type ℓ} (r : R X) → isProp (P X (projectFields fs r)) + + data AutoRecordSpec : Typeω where + autoRecordSpec : (R : Type ℓ → Type ℓ₁) (ι : StrEquiv R ℓ₁') + → AutoFields R ι + → AutoRecordSpec + +-- Some reflection utilities +private + tApply : R.Term → List (R.Arg R.Term) → R.Term + tApply t l = R.def (quote idfun) (R.unknown v∷ t v∷ l) + + tStrMap : R.Term → R.Term → R.Term + tStrMap A f = R.def (quote map-snd) (f v∷ A v∷ []) + + tStrProj : R.Term → R.Name → R.Term + tStrProj A sfield = tStrMap A (R.def sfield []) + + Fun : ∀ {ℓ ℓ'} → Type ℓ → Type ℓ' → Type (ℓ-max ℓ ℓ') + Fun A B = A → B + +-- Helper functions used in the generated univalence proof +private + pathMap : ∀ {ℓ ℓ'} {S : I → Type ℓ} {T : I → Type ℓ'} (f : {i : I} → S i → T i) + {x : S i0} {y : S i1} → PathP S x y → PathP T (f x) (f y) + pathMap f p i = f (p i) + + -- Property field helper functions + module _ + {ℓ ℓ₁ ℓ₁' ℓ₂} + (R : Type ℓ → Type ℓ₁) -- Structure record + (ι : StrEquiv R ℓ₁') -- Equivalence record + (fs : AutoFields R ι) -- Prior fields + (P : (X : Type ℓ) → GatherFields fs X → Type ℓ₂) -- Property type + (f : {X : Type ℓ} (r : R X) → P X (projectFields fs r)) -- Property projection + where + + prev = projectFields fs + Prev = GatherFields fs + + PropHelperCenterType : Type _ + PropHelperCenterType = + (A B : TypeWithStr ℓ R) (e : A .fst ≃ B .fst) + (p : PathP (λ i → Prev (ua e i)) (prev (A .snd)) (prev (B .snd))) + → PathP (λ i → P (ua e i) (p i)) (f (A .snd)) (f (B .snd)) + + PropHelperContractType : PropHelperCenterType → Type _ + PropHelperContractType c = + (A B : TypeWithStr ℓ R) (e : A .fst ≃ B .fst) + {p₀ : PathP (λ i → Prev (ua e i)) (prev (A .snd)) (prev (B .snd))} + (q : PathP (λ i → R (ua e i)) (A .snd) (B .snd)) + (p : p₀ ≡ (λ i → prev (q i))) + → PathP (λ k → PathP (λ i → P (ua e i) (p k i)) (f (A .snd)) (f (B .snd))) + (c A B e p₀) + (λ i → f (q i)) + + PropHelperType : Type _ + PropHelperType = + Σ PropHelperCenterType PropHelperContractType + + derivePropHelper : isPropProperty R ι fs P → PropHelperType + derivePropHelper propP .fst A B e p = + isOfHLevelPathP' 0 (propP _) (f (A .snd)) (f (B .snd)) .fst + derivePropHelper propP .snd A B e q p = + isOfHLevelPathP' 0 (isOfHLevelPathP 1 (propP _) _ _) _ _ .fst + + -- Build proof of univalence from an isomorphism + module _ {ℓ ℓ₁ ℓ₁'} (S : Type ℓ → Type ℓ₁) (ι : StrEquiv S ℓ₁') where + + fwdShape : Type _ + fwdShape = + (A B : TypeWithStr ℓ S) (e : typ A ≃ typ B) → ι A B e → PathP (λ i → S (ua e i)) (str A) (str B) + + bwdShape : Type _ + bwdShape = + (A B : TypeWithStr ℓ S) (e : typ A ≃ typ B) → PathP (λ i → S (ua e i)) (str A) (str B) → ι A B e + + fwdBwdShape : fwdShape → bwdShape → Type _ + fwdBwdShape fwd bwd = + (A B : TypeWithStr ℓ S) (e : typ A ≃ typ B) → ∀ p → fwd A B e (bwd A B e p) ≡ p + + bwdFwdShape : fwdShape → bwdShape → Type _ + bwdFwdShape fwd bwd = + (A B : TypeWithStr ℓ S) (e : typ A ≃ typ B) → ∀ r → bwd A B e (fwd A B e r) ≡ r + + -- The implicit arguments A,B in UnivalentStr make some things annoying so let's avoid them + ExplicitUnivalentStr : Type _ + ExplicitUnivalentStr = + (A B : TypeWithStr _ S) (e : typ A ≃ typ B) → ι A B e ≃ PathP (λ i → S (ua e i)) (str A) (str B) + + explicitUnivalentStr : (fwd : fwdShape) (bwd : bwdShape) + → fwdBwdShape fwd bwd → bwdFwdShape fwd bwd + → ExplicitUnivalentStr + explicitUnivalentStr fwd bwd fwdBwd bwdFwd A B e = isoToEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun = fwd A B e + isom .inv = bwd A B e + isom .rightInv = fwdBwd A B e + isom .leftInv = bwdFwd A B e + + ExplicitUnivalentDesc : ∀ ℓ {ℓ₁ ℓ₁'} → (d : M.Desc ℓ ℓ₁ ℓ₁') → Type _ + ExplicitUnivalentDesc _ d = + ExplicitUnivalentStr (M.MacroStructure d) (M.MacroEquivStr d) + + explicitUnivalentDesc : ∀ ℓ {ℓ₁ ℓ₁'} → (d : M.Desc ℓ ℓ₁ ℓ₁') → ExplicitUnivalentDesc ℓ d + explicitUnivalentDesc _ d A B e = M.MacroUnivalentStr d e + +-- Internal record specification type +private + record TypedTerm : Type where + field + type : R.Term + term : R.Term + + record InternalDatumField : Type where + field + sfield : R.Name -- name of structure field + efield : R.Name -- name of equivalence field + + record InternalPropField : Type where + field + sfield : R.Name -- name of structure field + + InternalField : Type + InternalField = InternalDatumField ⊎ InternalPropField + + record InternalSpec (A : Type) : Type where + field + srec : R.Term -- structure record type + erec : R.Term -- equivalence record type + fields : List (InternalField × A) -- in reverse order + + open TypedTerm + open InternalDatumField + open InternalPropField + +-- Parse a field and record specifications +private + findName : R.Term → R.TC R.Name + findName (R.def name _) = R.returnTC name + findName (R.lam R.hidden (R.abs _ t)) = findName t + findName t = R.typeError (R.strErr "Not a name + spine: " ∷ R.termErr t ∷ []) + + parseFieldSpec : R.Term → R.TC (R.Term × R.Term × R.Term × R.Term) + parseFieldSpec (R.con (quote autoFieldSpec) (ℓ h∷ ℓ₁ h∷ ℓ₂ h∷ R v∷ S h∷ f v∷ [])) = + R.reduce ℓ >>= λ ℓ → + R.returnTC (ℓ , ℓ₂ , S , f) + parseFieldSpec t = + R.typeError (R.strErr "Malformed field specification: " ∷ R.termErr t ∷ []) + + parseSpec : R.Term → R.TC (InternalSpec TypedTerm) + parseSpec (R.con (quote autoRecordSpec) (ℓ h∷ ℓ₁ h∷ ℓ₁' h∷ srecTerm v∷ erecTerm v∷ fs v∷ [])) = + parseFields fs >>= λ fs' → + R.returnTC λ { .srec → srecTerm ; .erec → erecTerm ; .fields → fs'} + where + open InternalSpec + + parseFields : R.Term → R.TC (List (InternalField × TypedTerm)) + parseFields (R.con (quote fields:) _) = R.returnTC [] + parseFields (R.con (quote _data[_∣_]) + (ℓ h∷ ℓ₁ h∷ ℓ₁' h∷ R h∷ ι h∷ fs v∷ ℓ₂ h∷ ℓ₂' h∷ S h∷ ι' h∷ sfieldTerm v∷ efieldTerm v∷ [])) + = + R.reduce ℓ >>= λ ℓ → + findName sfieldTerm >>= λ sfieldName → + findName efieldTerm >>= λ efieldName → + buildDesc FUEL ℓ ℓ₂ S >>= λ d → + let + f : InternalField × TypedTerm + f = λ + { .fst → inl λ { .sfield → sfieldName ; .efield → efieldName } + ; .snd .type → R.def (quote ExplicitUnivalentDesc) (ℓ v∷ d v∷ []) + ; .snd .term → R.def (quote explicitUnivalentDesc) (ℓ v∷ d v∷ []) + } + in + liftTC (f ∷_) (parseFields fs) + parseFields (R.con (quote _prop[_∣_]) + (ℓ h∷ ℓ₁ h∷ ℓ₁' h∷ R h∷ ι h∷ fs v∷ ℓ₂ h∷ P h∷ fieldTerm v∷ prop v∷ [])) + = + findName fieldTerm >>= λ fieldName → + let + p : InternalField × TypedTerm + p = λ + { .fst → inr λ { .sfield → fieldName } + ; .snd .type → + R.def (quote PropHelperType) (srecTerm v∷ erecTerm v∷ fs v∷ P v∷ fieldTerm v∷ []) + ; .snd .term → + R.def (quote derivePropHelper) (srecTerm v∷ erecTerm v∷ fs v∷ P v∷ fieldTerm v∷ prop v∷ []) + } + in + liftTC (p ∷_) (parseFields fs) + + parseFields t = R.typeError (R.strErr "Malformed autoRecord specification (1): " ∷ R.termErr t ∷ []) + + parseSpec t = R.typeError (R.strErr "Malformed autoRecord specification (2): " ∷ R.termErr t ∷ []) + +-- Build a proof of univalence from an InternalSpec +module _ (spec : InternalSpec ℕ) where + open InternalSpec spec + private + + fwdDatum : Vec R.Term 4 → R.Term → InternalDatumField × ℕ → R.Term + fwdDatum (A ∷ B ∷ e ∷ streq ∷ _) i (dat , n) = + R.def (quote equivFun) + (tApply (v n) (tStrProj A (dat .sfield) v∷ tStrProj B (dat .sfield) v∷ e v∷ []) + v∷ R.def (dat .efield) (streq v∷ []) + v∷ i + v∷ []) + + fwdProperty : Vec R.Term 4 → R.Term → R.Term → InternalPropField × ℕ → R.Term + fwdProperty (A ∷ B ∷ e ∷ streq ∷ _) i prevPath prop = + R.def (quote fst) (v (prop .snd) v∷ A v∷ B v∷ e v∷ prevPath v∷ i v∷ []) + + bwdClause : Vec R.Term 4 → InternalDatumField × ℕ → R.Clause + bwdClause (A ∷ B ∷ e ∷ q ∷ _) (dat , n) = + R.clause [] (R.proj (dat .efield) v∷ []) + (R.def (quote invEq) + (tApply + (v n) + (tStrProj A (dat .sfield) v∷ tStrProj B (dat .sfield) v∷ e v∷ []) + v∷ R.def (quote pathMap) (R.def (dat .sfield) [] v∷ q v∷ []) + v∷ [])) + + fwdBwdDatum : Vec R.Term 4 → R.Term → R.Term → InternalDatumField × ℕ → R.Term + fwdBwdDatum (A ∷ B ∷ e ∷ q ∷ _) j i (dat , n) = + R.def (quote secEq) + (tApply + (v n) + (tStrProj A (dat .sfield) v∷ tStrProj B (dat .sfield) v∷ e v∷ []) + v∷ R.def (quote pathMap) (R.def (dat .sfield) [] v∷ q v∷ []) + v∷ j v∷ i + v∷ []) + + fwdBwdProperty : Vec R.Term 4 → (j i prevPath : R.Term) → InternalPropField × ℕ → R.Term + fwdBwdProperty (A ∷ B ∷ e ∷ q ∷ _) j i prevPath prop = + R.def (quote snd) (v (prop .snd) v∷ A v∷ B v∷ e v∷ q v∷ prevPath v∷ j v∷ i v∷ []) + + bwdFwdClause : Vec R.Term 4 → R.Term → InternalDatumField × ℕ → R.Clause + bwdFwdClause (A ∷ B ∷ e ∷ streq ∷ _) j (dat , n) = + R.clause [] (R.proj (dat .efield) v∷ []) + (R.def (quote retEq) + (tApply + (v n) + (tStrProj A (dat .sfield) v∷ tStrProj B (dat .sfield) v∷ e v∷ []) + v∷ R.def (dat .efield) (streq v∷ []) + v∷ j + v∷ [])) + + makeVarsFrom : {n : ℕ} → ℕ → Vec R.Term n + makeVarsFrom {zero} k = [] + makeVarsFrom {suc n} k = v (n + k) ∷ (makeVarsFrom k) + + fwd : R.Term + fwd = + vlam "A" (vlam "B" (vlam "e" (vlam "streq" (vlam "i" (R.pat-lam body []))))) + where + -- input list is in reverse order + fwdClauses : ℕ → List (InternalField × ℕ) → List (R.Name × R.Term) + fwdClauses k [] = [] + fwdClauses k ((inl f , n) ∷ fs) = + fwdClauses k fs + ∷ʳ (f .sfield , fwdDatum (makeVarsFrom k) (v 0) (map-snd (4 + k +_) (f , n))) + fwdClauses k ((inr p , n) ∷ fs) = + fwdClauses k fs + ∷ʳ (p .sfield , fwdProperty (makeVarsFrom k) (v 0) prevPath (map-snd (4 + k +_) (p , n))) + where + prevPath = + vlam "i" + (List.foldl + (λ t (_ , t') → R.con (quote _,_) (t v∷ t' v∷ [])) + (R.con (quote tt) []) + (fwdClauses (suc k) fs)) + + body = + List.map (λ (n , t) → R.clause [] [ varg (R.proj n) ] t) (fwdClauses 1 fields) + + bwd : R.Term + bwd = + vlam "A" (vlam "B" (vlam "e" (vlam "q" (R.pat-lam (bwdClauses fields) [])))) + where + -- input is in reverse order + bwdClauses : List (InternalField × ℕ) → List R.Clause + bwdClauses [] = [] + bwdClauses ((inl f , n) ∷ fs) = + bwdClauses fs + ∷ʳ bwdClause (makeVarsFrom 0) (map-snd (4 +_) (f , n)) + bwdClauses ((inr p , n) ∷ fs) = bwdClauses fs + + fwdBwd : R.Term + fwdBwd = + vlam "A" (vlam "B" (vlam "e" (vlam "q" (vlam "j" (vlam "i" (R.pat-lam body [])))))) + where + -- input is in reverse order + fwdBwdClauses : ℕ → List (InternalField × ℕ) → List (R.Name × R.Term) + fwdBwdClauses k [] = [] + fwdBwdClauses k ((inl f , n) ∷ fs) = + fwdBwdClauses k fs + ∷ʳ (f .sfield , fwdBwdDatum (makeVarsFrom k) (v 1) (v 0) (map-snd (4 + k +_) (f , n))) + fwdBwdClauses k ((inr p , n) ∷ fs) = + fwdBwdClauses k fs + ∷ʳ ((p .sfield , fwdBwdProperty (makeVarsFrom k) (v 1) (v 0) prevPath (map-snd (4 + k +_) (p , n)))) + where + prevPath = + vlam "j" + (vlam "i" + (List.foldl + (λ t (_ , t') → R.con (quote _,_) (t v∷ t' v∷ [])) + (R.con (quote tt) []) + (fwdBwdClauses (2 + k) fs))) + + body = List.map (λ (n , t) → R.clause [] [ varg (R.proj n) ] t) (fwdBwdClauses 2 fields) + + bwdFwd : R.Term + bwdFwd = + vlam "A" (vlam "B" (vlam "e" (vlam "streq" (vlam "j" (R.pat-lam (bwdFwdClauses fields) []))))) + where + bwdFwdClauses : List (InternalField × ℕ) → List R.Clause + bwdFwdClauses [] = [] + bwdFwdClauses ((inl f , n) ∷ fs) = + bwdFwdClauses fs + ∷ʳ bwdFwdClause (makeVarsFrom 1) (v 0) (map-snd (5 +_) (f , n)) + bwdFwdClauses ((inr _ , n) ∷ fs) = bwdFwdClauses fs + + univalentRecord : R.Term + univalentRecord = + R.def (quote explicitUnivalentStr) + (R.unknown v∷ R.unknown v∷ fwd v∷ bwd v∷ fwdBwd v∷ bwdFwd v∷ []) + +macro + autoFieldEquiv : R.Term → R.Term → R.Term → R.Term → R.TC Unit + autoFieldEquiv spec A B hole = + (R.reduce spec >>= parseFieldSpec) >>= λ (ℓ , ℓ₂ , S , f) → + buildDesc FUEL ℓ ℓ₂ S >>= λ d → + R.unify hole (R.def (quote M.MacroEquivStr) (d v∷ tStrMap A f v∷ tStrMap B f v∷ [])) + + autoUnivalentRecord : R.Term → R.Term → R.TC Unit + autoUnivalentRecord t hole = + (R.reduce t >>= parseSpec) >>= λ spec → + -- R.typeError (R.strErr "WOW: " ∷ R.termErr (main spec) ∷ []) + R.unify (main spec) hole + where + module _ (spec : InternalSpec TypedTerm) where + open InternalSpec spec + + mapUp : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (ℕ → A → B) → ℕ → List A → List B + mapUp f _ [] = [] + mapUp f n (x ∷ xs) = f n x ∷ mapUp f (suc n) xs + + closureSpec : InternalSpec ℕ + closureSpec .InternalSpec.srec = srec + closureSpec .InternalSpec.erec = erec + closureSpec .InternalSpec.fields = mapUp (λ n → map-snd (λ _ → n)) 0 fields + + closure : R.Term + closure = + iter (List.length fields) (vlam "") (univalentRecord closureSpec) + + env : List (R.Arg R.Term) + env = List.map (varg ∘ term ∘ snd) (List.rev fields) + + closureTy : R.Term + closureTy = + List.foldr + (λ ty cod → R.def (quote Fun) (ty v∷ cod v∷ [])) + (R.def (quote ExplicitUnivalentStr) (srec v∷ erec v∷ [])) + (List.map (type ∘ snd) (List.rev fields)) + + main : R.Term + main = R.def (quote idfun) (closureTy v∷ closure v∷ env) diff --git a/Cubical/Structures/Relational/Auto.agda b/Cubical/Structures/Relational/Auto.agda new file mode 100644 index 000000000..9901bdb14 --- /dev/null +++ b/Cubical/Structures/Relational/Auto.agda @@ -0,0 +1,239 @@ +{- + +Macros (autoDesc, AutoStructure, AutoEquivStr, autoUnivalentStr) for automatically generating structure definitions. + +For example: + + autoDesc (λ (X : Type₀) → X → X × ℕ) ↦ function+ var (var , constant ℕ) + +We prefer to use the constant structure whenever possible, e.g., [autoDesc (λ (X : Type₀) → ℕ → ℕ)] +is [constant (ℕ → ℕ)] rather than [function (constant ℕ) (constant ℕ)]. + +Writing [auto* (λ X → ⋯)] doesn't seem to work, but [auto* (λ (X : Type ℓ) → ⋯)] does. + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Relational.Auto where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.List +open import Cubical.Data.Bool +open import Cubical.Data.Maybe + +open import Cubical.Structures.Relational.Macro as Macro + +import Agda.Builtin.Reflection as R +open import Cubical.Reflection.Base + +-- Magic number +private + FUEL = 10000 + +-- Mark a constant type with a proof it is a set +abstract + Const[_] : ∀ {ℓ} → hSet ℓ → Type ℓ + Const[ A ] = A .fst + +-- Some reflection utilities +private + tLevel = R.def (quote Level) [] + + tType : R.Term → R.Term + tType ℓ = R.def (quote Type) [ varg ℓ ] + + thSet : R.Term → R.Term + thSet ℓ = R.def (quote hSet) [ varg ℓ ] + + tPosRelDesc : R.Term → R.Term → R.Term + tPosRelDesc ℓ ℓ₁ = R.def (quote PosRelDesc) (ℓ v∷ ℓ₁ v∷ []) + + tRelDesc : R.Term → R.Term → R.Term → R.Term + tRelDesc ℓ ℓ₁ ℓ₁' = R.def (quote RelDesc) (ℓ v∷ ℓ₁ v∷ ℓ₁' v∷ []) + + func : (ℓ ℓ' : Level) → Type (ℓ-suc (ℓ-max ℓ ℓ')) + func ℓ ℓ' = Type ℓ → Type ℓ' + + tStruct : R.Term → R.Term → R.Term + tStruct ℓ ℓ' = R.def (quote func) (varg ℓ ∷ varg ℓ' ∷ []) + +-- We try to build a descriptor by unifying the provided structure with combinators we're aware of. We +-- redefine the structure combinators as the *Shape terms below so that we don't depend on the specific way +-- these are defined in other files (order of implicit arguments and so on); the syntactic analysis that goes +-- on here means that we would get mysterious errors if those changed. +private + constantShape : ∀ {ℓ'} (ℓ : Level) (A : hSet ℓ') → (Type ℓ → Type ℓ') + constantShape _ A _ = Const[ A ] + + pointedShape : (ℓ : Level) → Type ℓ → Type ℓ + pointedShape _ X = X + + productShape : ∀ {ℓ₀ ℓ₁} (ℓ : Level) + → (Type ℓ → Type ℓ₀) → (Type ℓ → Type ℓ₁) → Type ℓ → Type (ℓ-max ℓ₀ ℓ₁) + productShape _ A₀ A₁ X = A₀ X × A₁ X + + paramShape : ∀ {ℓ₀ ℓ'} (ℓ : Level) + → Type ℓ' → (Type ℓ → Type ℓ₀) → Type ℓ → Type (ℓ-max ℓ' ℓ₀) + paramShape _ A A₀ X = A → A₀ X + + functionShape : ∀ {ℓ₀ ℓ₁} (ℓ : Level) + → (Type ℓ → Type ℓ₀) → (Type ℓ → Type ℓ₁) → Type ℓ → Type (ℓ-max ℓ₀ ℓ₁) + functionShape _ A₀ A₁ X = A₀ X → A₁ X + + maybeShape : ∀ {ℓ₀} (ℓ : Level) + → (Type ℓ → Type ℓ₀) → Type ℓ → Type ℓ₀ + maybeShape _ A₀ X = Maybe (A₀ X) + +private + -- Build transport structure descriptor from a function [t : Type ℓ → Type ℓ'] + buildPosRelDesc : ℕ → R.Term → R.Term → R.Term → R.TC R.Term + buildPosRelDesc zero ℓ ℓ' t = R.typeError (R.strErr "Ran out of fuel! at \n" ∷ R.termErr t ∷ []) + buildPosRelDesc (suc fuel) ℓ ℓ' t = + tryConstant t <|> tryPointed t <|> tryProduct t <|> tryMaybe t <|> + R.typeError (R.strErr "Can't automatically generate a positive structure for\n" ∷ R.termErr t ∷ []) + where + tryConstant : R.Term → R.TC R.Term + tryConstant t = + newMeta (thSet ℓ') >>= λ A → + R.unify t (R.def (quote constantShape) (varg ℓ ∷ varg A ∷ [])) >> + R.returnTC (R.con (quote PosRelDesc.constant) (varg A ∷ [])) + + tryPointed : R.Term → R.TC R.Term + tryPointed t = + R.unify t (R.def (quote pointedShape) (varg ℓ ∷ [])) >> + R.returnTC (R.con (quote PosRelDesc.var) []) + + tryProduct : R.Term → R.TC R.Term + tryProduct t = + newMeta tLevel >>= λ ℓ₀ → + newMeta tLevel >>= λ ℓ₁ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + newMeta (tStruct ℓ ℓ₁) >>= λ A₁ → + R.unify t (R.def (quote productShape) (varg ℓ ∷ varg A₀ ∷ varg A₁ ∷ [])) >> + buildPosRelDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + buildPosRelDesc fuel ℓ ℓ₁ A₁ >>= λ d₁ → + R.returnTC (R.con (quote PosRelDesc._,_) (varg d₀ ∷ varg d₁ ∷ [])) + + tryMaybe : R.Term → R.TC R.Term + tryMaybe t = + newMeta tLevel >>= λ ℓ₀ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + R.unify t (R.def (quote maybeShape) (varg ℓ ∷ varg A₀ ∷ [])) >> + buildPosRelDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + R.returnTC (R.con (quote PosRelDesc.maybe) (varg d₀ ∷ [])) + + autoPosRelDesc' : R.Term → R.Term → R.TC Unit + autoPosRelDesc' t hole = + R.inferType hole >>= λ H → + newMeta tLevel >>= λ ℓ → + newMeta tLevel >>= λ ℓ' → + R.unify (tPosRelDesc ℓ ℓ') H >> + R.checkType t (tStruct ℓ ℓ') >> + buildPosRelDesc FUEL ℓ ℓ' t >>= R.unify hole + + -- Build structure descriptor from a function [t : Type ℓ → Type ℓ'] + buildRelDesc : ℕ → R.Term → R.Term → R.Term → R.TC R.Term + buildRelDesc zero ℓ ℓ' t = R.typeError (R.strErr "Ran out of fuel! at \n" ∷ R.termErr t ∷ []) + buildRelDesc (suc fuel) ℓ ℓ' t = + tryConstant t <|> tryPointed t <|> tryProduct t <|> tryParam t <|> tryFunction t <|> + tryMaybe t <|> + R.typeError (R.strErr "Can't automatically generate a structure for\n" ∷ R.termErr t ∷ []) + where + tryConstant : R.Term → R.TC R.Term + tryConstant t = + newMeta (thSet ℓ') >>= λ A → + R.unify t (R.def (quote constantShape) (varg ℓ ∷ varg A ∷ [])) >> + R.returnTC (R.con (quote RelDesc.constant) (varg A ∷ [])) + + tryPointed : R.Term → R.TC R.Term + tryPointed t = + R.unify t (R.def (quote pointedShape) (varg ℓ ∷ [])) >> + R.returnTC (R.con (quote RelDesc.var) []) + + tryProduct : R.Term → R.TC R.Term + tryProduct t = + newMeta tLevel >>= λ ℓ₀ → + newMeta tLevel >>= λ ℓ₁ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + newMeta (tStruct ℓ ℓ₁) >>= λ A₁ → + R.unify t (R.def (quote productShape) (varg ℓ ∷ varg A₀ ∷ varg A₁ ∷ [])) >> + buildRelDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + buildRelDesc fuel ℓ ℓ₁ A₁ >>= λ d₁ → + R.returnTC (R.con (quote RelDesc._,_) (varg d₀ ∷ varg d₁ ∷ [])) + + tryParam : R.Term → R.TC R.Term + tryParam t = + newMeta (tType R.unknown) >>= λ A → + newMeta tLevel >>= λ ℓ₀ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + R.unify t (R.def (quote paramShape) (varg ℓ ∷ varg A ∷ varg A₀ ∷ [])) >> + buildRelDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + R.returnTC (R.con (quote RelDesc.param) (varg A ∷ varg d₀ ∷ [])) + + tryFunction : R.Term → R.TC R.Term + tryFunction t = + newMeta tLevel >>= λ ℓ₀ → + newMeta tLevel >>= λ ℓ₁ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + newMeta (tStruct ℓ ℓ₁) >>= λ A₁ → + R.unify t (R.def (quote functionShape) (varg ℓ ∷ varg A₀ ∷ varg A₁ ∷ [])) >> + buildPosRelDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + buildRelDesc fuel ℓ ℓ₁ A₁ >>= λ d₁ → + R.returnTC (R.con (quote RelDesc.function+) (varg d₀ ∷ varg d₁ ∷ [])) + + tryMaybe : R.Term → R.TC R.Term + tryMaybe t = + newMeta tLevel >>= λ ℓ₀ → + newMeta (tStruct ℓ ℓ₀) >>= λ A₀ → + R.unify t (R.def (quote maybeShape) (varg ℓ ∷ varg A₀ ∷ [])) >> + buildRelDesc fuel ℓ ℓ₀ A₀ >>= λ d₀ → + R.returnTC (R.con (quote RelDesc.maybe) (varg d₀ ∷ [])) + + autoRelDesc' : R.Term → R.Term → R.TC Unit + autoRelDesc' t hole = + R.inferType hole >>= λ H → + newMeta tLevel >>= λ ℓ → + newMeta tLevel >>= λ ℓ' → + R.unify (tRelDesc ℓ ℓ' R.unknown) H >> + R.checkType t (tStruct ℓ ℓ') >> + buildRelDesc FUEL ℓ ℓ' t >>= R.unify hole + +macro + -- (Type ℓ → Type ℓ₁) → PosRelDesc ℓ + autoPosRelDesc : R.Term → R.Term → R.TC Unit + autoPosRelDesc = autoPosRelDesc' + + -- (S : Type ℓ → Type ℓ₁) → RelDesc ℓ + autoRelDesc : R.Term → R.Term → R.TC Unit + autoRelDesc = autoRelDesc' + + -- (S : Type ℓ → Type ℓ₁) → (Type ℓ → Type ℓ₁) + -- Sanity check: should be the identity + AutoStructure : R.Term → R.Term → R.TC Unit + AutoStructure t hole = + newMeta (tRelDesc R.unknown R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote RelMacroStructure) [ varg d ]) >> + autoRelDesc' t d + + -- (S : Type ℓ → Type ℓ₁) → StrRel S _ + AutoRelStr : R.Term → R.Term → R.TC Unit + AutoRelStr t hole = + newMeta (tRelDesc R.unknown R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote RelMacroRelStr) [ varg d ]) >> + autoRelDesc' t d + + -- (S : Type ℓ → Type ℓ₁) → SuitableStrRel S (AutoStrRel S) + autoSuitableRel : R.Term → R.Term → R.TC Unit + autoSuitableRel t hole = + newMeta (tRelDesc R.unknown R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote relMacroSuitableRel) [ varg d ]) >> + autoRelDesc' t d + + -- (S : Type ℓ → Type ℓ₁) → StrRelMatchesEquiv (AutoRelStr S) (AutoEquivStr S) + autoMatchesEquiv : R.Term → R.Term → R.TC Unit + autoMatchesEquiv t hole = + newMeta (tRelDesc R.unknown R.unknown R.unknown) >>= λ d → + R.unify hole (R.def (quote relMacroMatchesEquiv) [ varg d ]) >> + autoRelDesc' t d diff --git a/Cubical/Structures/Relational/Constant.agda b/Cubical/Structures/Relational/Constant.agda new file mode 100644 index 000000000..78394cef3 --- /dev/null +++ b/Cubical/Structures/Relational/Constant.agda @@ -0,0 +1,57 @@ +{- + +Constant structure: _ ↦ A + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Relational.Constant where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.RelationalStructure +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.SetQuotients + +open import Cubical.Structures.Constant + +private + variable + ℓ ℓ' : Level + +-- Structured relations + +module _ (A : hSet ℓ') where + + ConstantRelStr : StrRel {ℓ = ℓ} (ConstantStructure (A .fst)) ℓ' + ConstantRelStr _ a₀ a₁ = a₀ ≡ a₁ + + constantSuitableRel : SuitableStrRel {ℓ = ℓ} (ConstantStructure (A .fst)) ConstantRelStr + constantSuitableRel .quo _ _ _ = isContrSingl _ + constantSuitableRel .symmetric _ = sym + constantSuitableRel .transitive _ _ = _∙_ + constantSuitableRel .set _ = A .snd + constantSuitableRel .prop _ = A .snd + + constantRelMatchesEquiv : StrRelMatchesEquiv {ℓ = ℓ} ConstantRelStr (ConstantEquivStr (A .fst)) + constantRelMatchesEquiv _ _ _ = idEquiv _ + + constantRelAction : StrRelAction {ℓ = ℓ} ConstantRelStr + constantRelAction .actStr _ a = a + constantRelAction .actStrId _ = refl + constantRelAction .actRel _ a a' p = p + + constantPositiveRel : PositiveStrRel {ℓ = ℓ} constantSuitableRel + constantPositiveRel .act = constantRelAction + constantPositiveRel .reflexive a = refl + constantPositiveRel .detransitive R R' p = ∣ _ , p , refl ∣ + constantPositiveRel .quo R = isoToIsEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun = _ + isom .inv = [_] + isom .rightInv _ = refl + isom .leftInv = elimProp (λ _ → squash/ _ _) (λ a → refl) diff --git a/Cubical/Structures/Relational/Equalizer.agda b/Cubical/Structures/Relational/Equalizer.agda new file mode 100644 index 000000000..faeeaa38d --- /dev/null +++ b/Cubical/Structures/Relational/Equalizer.agda @@ -0,0 +1,67 @@ +{- + +Equalizer of functions f g : S X ⇉ T X such that f and g act on relation structures + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Relational.Equalizer where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.RelationalStructure +open import Cubical.Data.Sigma +open import Cubical.HITs.SetQuotients +open import Cubical.HITs.PropositionalTruncation as Trunc +open import Cubical.Relation.Binary.Base + +private + variable + ℓ ℓ₁ ℓ₁' ℓ₂ ℓ₂' : Level + +EqualizerStructure : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + (f g : ∀ X → S X → T X) + → Type ℓ → Type (ℓ-max ℓ₁ ℓ₂) +EqualizerStructure {S = S} f g X = Σ[ s ∈ S X ] f X s ≡ g X s + +EqualizerRelStr : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + {f g : ∀ X → S X → T X} + → StrRel S ℓ₁' → StrRel T ℓ₂' → StrRel (EqualizerStructure f g) ℓ₁' +EqualizerRelStr ρ₁ ρ₂ R (s , _) (s' , _) = ρ₁ R s s' + +equalizerSuitableRel : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + {f g : ∀ X → S X → T X} + {ρ₁ : StrRel S ℓ₁'} {ρ₂ : StrRel T ℓ₂'} + (αf : ∀ {X Y} {R : X → Y → Type ℓ} {s s'} → ρ₁ R s s' → ρ₂ R (f X s) (f Y s')) + (αg : ∀ {X Y} {R : X → Y → Type ℓ} {s s'} → ρ₁ R s s' → ρ₂ R (g X s) (g Y s')) + → SuitableStrRel S ρ₁ + → SuitableStrRel T ρ₂ + → SuitableStrRel (EqualizerStructure f g) (EqualizerRelStr ρ₁ ρ₂) +equalizerSuitableRel {f = f} {g} {ρ₁} {ρ₂} αf αg θ₁ θ₂ .quo (X , s , p) (R , qer) r = + ( ((quo₁ .fst .fst , sym step₀ ∙ step₁) , quo₁ .fst .snd) + , λ {((s' , _) , r') → + Σ≡Prop (λ _ → θ₁ .prop (λ _ _ → squash/ _ _) _ _) + (Σ≡Prop (λ _ → θ₂ .set squash/ _ _) + (cong fst (quo₁ .snd (s' , r'))))} + ) + where + quo₁ = θ₁ .quo (X , s) (R , qer) r + quo₂ = θ₂ .quo (X , f X s) (R , qer) (αf r) + + step₀ : quo₂ .fst .fst ≡ f (X / R .fst) (quo₁ .fst .fst) + step₀ = + cong fst + (quo₂ .snd + (f _ (quo₁ .fst .fst) , αf (quo₁ .fst .snd))) + + step₁ : quo₂ .fst .fst ≡ g (X / R .fst) (quo₁ .fst .fst) + step₁ = + cong fst + (quo₂ .snd + ( g _ (quo₁ .fst .fst) + , subst (λ t → ρ₂ (graphRel [_]) t (g _ (quo₁ .fst .fst))) (sym p) (αg (quo₁ .fst .snd)) + )) +equalizerSuitableRel _ _ θ₁ _ .symmetric R = θ₁ .symmetric R +equalizerSuitableRel _ _ θ₁ _ .transitive R R' = θ₁ .transitive R R' +equalizerSuitableRel _ _ θ₁ θ₂ .set setX = + isSetΣ (θ₁ .set setX) λ _ → isOfHLevelPath 2 (θ₂ .set setX) _ _ +equalizerSuitableRel _ _ θ₁ _ .prop propR _ _ = θ₁ .prop propR _ _ diff --git a/Cubical/Structures/Relational/Function.agda b/Cubical/Structures/Relational/Function.agda new file mode 100644 index 000000000..e618bd14b --- /dev/null +++ b/Cubical/Structures/Relational/Function.agda @@ -0,0 +1,191 @@ +{- + +Index a structure T a positive structure S: X ↦ S X → T X + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Relational.Function where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.RelationalStructure +open import Cubical.Foundations.Univalence +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Sigma +open import Cubical.Relation.Binary.Base +open import Cubical.Relation.ZigZag.Base +open import Cubical.HITs.SetQuotients +open import Cubical.HITs.PropositionalTruncation as Trunc + +open import Cubical.Structures.Function + +private + variable + ℓ ℓ₁ ℓ₁' ℓ₁'' ℓ₂ ℓ₂' ℓ₂'' : Level + +FunctionRelStr : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + → StrRel S ℓ₁' → StrRel T ℓ₂' → StrRel (FunctionStructure S T) (ℓ-max ℓ₁ (ℓ-max ℓ₁' ℓ₂')) +FunctionRelStr ρ₁ ρ₂ R f g = + ∀ {x y} → ρ₁ R x y → ρ₂ R (f x) (g y) + +open BinaryRelation +open isEquivRel + +private + composeWith[_] : {A : Type ℓ} (R : EquivPropRel A ℓ) + → compPropRel (R .fst) (quotientPropRel (R .fst .fst)) .fst ≡ graphRel [_] + composeWith[_] R = + funExt₂ λ a t → + hPropExt squash (squash/ _ _) + (Trunc.rec (squash/ _ _) (λ {(b , r , p) → eq/ a b r ∙ p })) + (λ p → ∣ a , R .snd .reflexive a , p ∣) + + [_]∙[_]⁻¹ : {A : Type ℓ} (R : EquivPropRel A ℓ) + → compPropRel (quotientPropRel (R .fst .fst)) (invPropRel (quotientPropRel (R .fst .fst))) .fst + ≡ R .fst .fst + [_]∙[_]⁻¹ R = + funExt₂ λ a b → + hPropExt squash (R .fst .snd a b) + (Trunc.rec (R .fst .snd a b) + (λ {(c , p , q) → effective (R .fst .snd) (R .snd) a b (p ∙ sym q)})) + (λ r → ∣ _ , eq/ a b r , refl ∣) + +functionSuitableRel : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + {ρ₁ : StrRel S ℓ₁'} {ρ₂ : StrRel T ℓ₂'} + (θ₁ : SuitableStrRel S ρ₁) + → PositiveStrRel θ₁ + → SuitableStrRel T ρ₂ + → SuitableStrRel (FunctionStructure S T) (FunctionRelStr ρ₁ ρ₂) +functionSuitableRel {S = S} {T = T} {ρ₁ = ρ₁} {ρ₂} θ₁ σ₁ θ₂ .quo (X , f) R h = + final + where + ref : (s : S X) → ρ₁ (R .fst .fst) s s + ref = posRelReflexive σ₁ R + + [f] : S X / ρ₁ (R .fst .fst) → T (X / R .fst .fst) + [f] [ s ] = θ₂ .quo (X , f s) R (h (ref s)) .fst .fst + [f] (eq/ s₀ s₁ r i) = + cong fst + (θ₂ .quo (X , f s₀) R (h (ref s₀)) .snd + ( [f] [ s₁ ] + , subst (λ R' → ρ₂ R' (f s₀) ([f] [ s₁ ])) + (composeWith[_] R) + (θ₂ .transitive (R .fst) (quotientPropRel (R .fst .fst)) + (h r) + (θ₂ .quo (X , f s₁) R (h (ref s₁)) .fst .snd)) + )) + i + [f] (squash/ _ _ p q j i) = + θ₂ .set squash/ _ _ (cong [f] p) (cong [f] q) j i + + relLemma : (s : S X) (t : S X) + → ρ₁ (graphRel [_]) s (funIsEq (σ₁ .quo R) [ t ]) + → ρ₂ (graphRel [_]) (f s) ([f] [ t ]) + relLemma s t r = + subst (λ R' → ρ₂ R' (f s) ([f] [ t ])) + (composeWith[_] R) + (θ₂ .transitive (R .fst) (quotientPropRel (R .fst .fst)) + (h r') + (θ₂ .quo (X , f t) R (h (ref t)) .fst .snd)) + where + r' : ρ₁ (R .fst .fst) s t + r' = + subst (λ R' → ρ₁ R' s t) ([_]∙[_]⁻¹ R) + (θ₁ .transitive + (quotientPropRel (R .fst .fst)) + (invPropRel (quotientPropRel (R .fst .fst))) + r + (θ₁ .symmetric (quotientPropRel (R .fst .fst)) + (subst + (λ t' → ρ₁ (graphRel [_]) t' (funIsEq (σ₁ .quo R) [ t ])) + (σ₁ .act .actStrId t) + (σ₁ .act .actRel eq/ t t (ref t))))) + + quoRelLemma : (s : S X) (t : S X / ρ₁ (R .fst .fst)) + → ρ₁ (graphRel [_]) s (funIsEq (σ₁ .quo R) t) + → ρ₂ (graphRel [_]) (f s) ([f] t) + quoRelLemma s = + elimProp (λ _ → isPropΠ λ _ → θ₂ .prop (λ _ _ → squash/ _ _) _ _) + (relLemma s) + + final : Σ (Σ _ _) _ + final .fst .fst = [f] ∘ invIsEq (σ₁ .quo R) + final .fst .snd {s} {t} r = + quoRelLemma s + (invIsEq (σ₁ .quo R) t) + (subst (ρ₁ (graphRel [_]) s) (sym (secIsEq (σ₁ .quo R) t)) r) + final .snd (f' , c) = + Σ≡Prop + (λ _ → isPropImplicitΠ λ s → + isPropImplicitΠ λ t → + isPropΠ λ _ → θ₂ .prop (λ _ _ → squash/ _ _) _ _) + (funExt λ s → contractorLemma (invIsEq (σ₁ .quo R) s) ∙ cong f' (secIsEq (σ₁ .quo R) s)) + where + contractorLemma : (s : S X / ρ₁ (R .fst .fst)) + → [f] s ≡ f' (funIsEq (σ₁ .quo R) s) + contractorLemma = + elimProp + (λ _ → θ₂ .set squash/ _ _) + (λ s → + cong fst + (θ₂ .quo (X , f s) R (h (ref s)) .snd + ( f' (funIsEq (σ₁ .quo R) [ s ]) + , c + (subst + (λ s' → ρ₁ (graphRel [_]) s' (funIsEq (σ₁ .quo R) [ s ])) + (σ₁ .act .actStrId s) + (σ₁ .act .actRel eq/ s s (ref s))) + ))) +functionSuitableRel {ρ₁ = ρ₁} {ρ₂} θ₁ σ θ₂ .symmetric R h r = + θ₂ .symmetric R (h (θ₁ .symmetric (invPropRel R) r)) +functionSuitableRel {ρ₁ = ρ₁} {ρ₂} θ₁ σ θ₂ .transitive R R' h h' rr' = + Trunc.rec + (θ₂ .prop (λ _ _ → squash) _ _) + (λ {(_ , r , r') → θ₂ .transitive R R' (h r) (h' r')}) + (σ .detransitive R R' rr') +functionSuitableRel {ρ₁ = ρ₁} {ρ₂} θ₁ σ θ₂ .set setX = + isSetΠ λ _ → θ₂ .set setX +functionSuitableRel {ρ₁ = ρ₁} {ρ₂} θ₁ σ θ₂ .prop propR f g = + isPropImplicitΠ λ _ → + isPropImplicitΠ λ _ → + isPropΠ λ _ → + θ₂ .prop propR _ _ + +functionRelMatchesEquiv : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + (ρ₁ : StrRel S ℓ₁') {ι₁ : StrEquiv S ℓ₁''} + (ρ₂ : StrRel T ℓ₂') {ι₂ : StrEquiv T ℓ₂''} + → StrRelMatchesEquiv ρ₁ ι₁ + → StrRelMatchesEquiv ρ₂ ι₂ + → StrRelMatchesEquiv (FunctionRelStr ρ₁ ρ₂) (FunctionEquivStr ι₁ ι₂) +functionRelMatchesEquiv ρ₁ ρ₂ μ₁ μ₂ (X , f) (Y , g) e = + equivImplicitΠCod (equivImplicitΠCod (equiv→ (μ₁ _ _ e) (μ₂ _ _ e))) + +functionRelMatchesEquiv+ : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} + (ρ₁ : StrRel S ℓ₁') (α₁ : EquivAction S) + (ρ₂ : StrRel T ℓ₂') (ι₂ : StrEquiv T ℓ₂'') + → StrRelMatchesEquiv ρ₁ (EquivAction→StrEquiv α₁) + → StrRelMatchesEquiv ρ₂ ι₂ + → StrRelMatchesEquiv (FunctionRelStr ρ₁ ρ₂) (FunctionEquivStr+ α₁ ι₂) +functionRelMatchesEquiv+ ρ₁ α₁ ρ₂ ι₂ μ₁ μ₂ (X , f) (Y , g) e = + compEquiv + (functionRelMatchesEquiv ρ₁ ρ₂ μ₁ μ₂ (X , f) (Y , g) e) + (isoToEquiv isom) + where + open Iso + isom : Iso + (FunctionEquivStr (EquivAction→StrEquiv α₁) ι₂ (X , f) (Y , g) e) + (FunctionEquivStr+ α₁ ι₂ (X , f) (Y , g) e) + isom .fun h s = h refl + isom .inv k {x} = J (λ y _ → ι₂ (X , f x) (Y , g y) e) (k x) + isom .rightInv k i x = JRefl (λ y _ → ι₂ (X , f x) (Y , g y) e) (k x) i + isom .leftInv h = + implicitFunExt λ {x} → + implicitFunExt λ {y} → + funExt λ p → + J (λ y p → isom .inv (isom .fun h) p ≡ h p) + (funExt⁻ (isom .rightInv (isom .fun h)) x) + p diff --git a/Cubical/Structures/Relational/Macro.agda b/Cubical/Structures/Relational/Macro.agda new file mode 100644 index 000000000..61228a91b --- /dev/null +++ b/Cubical/Structures/Relational/Macro.agda @@ -0,0 +1,161 @@ +{- + +Descriptor language for easily defining relational structures + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Relational.Macro where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.RelationalStructure +open import Cubical.Data.Sigma + +open import Cubical.Structures.Relational.Constant +open import Cubical.Structures.Relational.Function +open import Cubical.Structures.Relational.Maybe +open import Cubical.Structures.Relational.Parameterized +open import Cubical.Structures.Relational.Pointed +open import Cubical.Structures.Relational.Product + +open import Cubical.Structures.Macro +open import Cubical.Structures.Maybe + +data PosRelDesc (ℓ : Level) : Level → Typeω where + -- constant structure: X ↦ A + constant : ∀ {ℓ₁} → hSet ℓ₁ → PosRelDesc ℓ ℓ₁ + -- pointed structure: X ↦ X + var : PosRelDesc ℓ ℓ + -- product of structures S,T : X ↦ (S X × T X) + _,_ : ∀ {ℓ₁ ℓ₂} → PosRelDesc ℓ ℓ₁ → PosRelDesc ℓ ℓ₂ → PosRelDesc ℓ (ℓ-max ℓ₁ ℓ₂) + -- Maybe on a structure S: X ↦ Maybe (S X) + maybe : ∀ {ℓ₁} → PosRelDesc ℓ ℓ₁ → PosRelDesc ℓ ℓ₁ + +data RelDesc (ℓ : Level) : Level → Level → Typeω where + -- constant structure: X ↦ A + constant : ∀ {ℓ₁} → hSet ℓ₁ → RelDesc ℓ ℓ₁ ℓ₁ + -- pointed structure: X ↦ X + var : RelDesc ℓ ℓ ℓ + -- product of structures S,T : X ↦ (S X × T X) + _,_ : ∀ {ℓ₁ ℓ₁' ℓ₂ ℓ₂'} → RelDesc ℓ ℓ₁ ℓ₁' → RelDesc ℓ ℓ₂ ℓ₂' → RelDesc ℓ (ℓ-max ℓ₁ ℓ₂) (ℓ-max ℓ₁' ℓ₂') + -- structure S parameterized by constant A : X ↦ (A → S X) + param : ∀ {ℓ₁ ℓ₂ ℓ₂'} → (A : Type ℓ₁) → RelDesc ℓ ℓ₂ ℓ₂' → RelDesc ℓ (ℓ-max ℓ₁ ℓ₂) (ℓ-max ℓ₁ ℓ₂') + -- function from positive structure S to T: X ↦ (S X → T X) + function+ : ∀ {ℓ₁ ℓ₂ ℓ₂'} → PosRelDesc ℓ ℓ₁ → RelDesc ℓ ℓ₂ ℓ₂' → RelDesc ℓ (ℓ-max ℓ₁ ℓ₂) (ℓ-max ℓ₁ ℓ₂') + -- Maybe on a structure S: X ↦ Maybe (S X) + maybe : ∀ {ℓ₁ ℓ₁'} → RelDesc ℓ ℓ₁ ℓ₁' → RelDesc ℓ ℓ₁ ℓ₁' + +infixr 4 _,_ + +posRelDesc→TranspDesc : ∀ {ℓ ℓ₁} → PosRelDesc ℓ ℓ₁ → TranspDesc ℓ ℓ₁ +posRelDesc→TranspDesc (constant A) = constant (A .fst) +posRelDesc→TranspDesc var = var +posRelDesc→TranspDesc (d₀ , d₁) = posRelDesc→TranspDesc d₀ , posRelDesc→TranspDesc d₁ +posRelDesc→TranspDesc (maybe d) = maybe (posRelDesc→TranspDesc d) + +posRelDesc→RelDesc : ∀ {ℓ ℓ₁} → PosRelDesc ℓ ℓ₁ → RelDesc ℓ ℓ₁ ℓ₁ +posRelDesc→RelDesc (constant A) = constant A +posRelDesc→RelDesc var = var +posRelDesc→RelDesc (d₀ , d₁) = posRelDesc→RelDesc d₀ , posRelDesc→RelDesc d₁ +posRelDesc→RelDesc (maybe d) = maybe (posRelDesc→RelDesc d) + +relDesc→Desc : ∀ {ℓ ℓ₁ ℓ₁'} → RelDesc ℓ ℓ₁ ℓ₁' → Desc ℓ ℓ₁ ℓ₁' +relDesc→Desc (constant A) = constant (A .fst) +relDesc→Desc var = var +relDesc→Desc (d₀ , d₁) = relDesc→Desc d₀ , relDesc→Desc d₁ +relDesc→Desc (param A d) = function+ (constant A) (relDesc→Desc d) +relDesc→Desc (function+ d₀ d₁) = function+ (posRelDesc→TranspDesc d₀) (relDesc→Desc d₁) +relDesc→Desc (maybe d) = maybe (relDesc→Desc d) + +{- Definition of structure -} + +PosRelMacroStructure : ∀ {ℓ ℓ₁} (d : PosRelDesc ℓ ℓ₁) → Type ℓ → Type ℓ₁ +PosRelMacroStructure d = TranspMacroStructure (posRelDesc→TranspDesc d) + +RelMacroStructure : ∀ {ℓ ℓ₁ ℓ₁'} (d : RelDesc ℓ ℓ₁ ℓ₁') → Type ℓ → Type ℓ₁ +RelMacroStructure d = MacroStructure (relDesc→Desc d) + +{- Notion of structured relation defined by a descriptor -} + +PosRelMacroRelStr : ∀ {ℓ ℓ₁} (d : PosRelDesc ℓ ℓ₁) → StrRel (PosRelMacroStructure d) ℓ₁ +PosRelMacroRelStr (constant A) = ConstantRelStr A +PosRelMacroRelStr var = PointedRelStr +PosRelMacroRelStr (d₀ , d₁) = ProductRelStr (PosRelMacroRelStr d₀) (PosRelMacroRelStr d₁) +PosRelMacroRelStr (maybe d) = MaybeRelStr (PosRelMacroRelStr d) + +RelMacroRelStr : ∀ {ℓ ℓ₁ ℓ₁'} (d : RelDesc ℓ ℓ₁ ℓ₁') → StrRel (RelMacroStructure d) ℓ₁' +RelMacroRelStr (constant A) = ConstantRelStr A +RelMacroRelStr var = PointedRelStr +RelMacroRelStr (d₀ , d₁) = ProductRelStr (RelMacroRelStr d₀) (RelMacroRelStr d₁) +RelMacroRelStr (param A d) = ParamRelStr A (λ _ → RelMacroRelStr d) +RelMacroRelStr (function+ d₀ d₁) = + FunctionRelStr (PosRelMacroRelStr d₀) (RelMacroRelStr d₁) +RelMacroRelStr (maybe d) = MaybeRelStr (RelMacroRelStr d) + +{- Proof that structure induced by descriptor is suitable or positive -} + +posRelMacroSuitableRel : ∀ {ℓ ℓ₁} (d : PosRelDesc ℓ ℓ₁) → SuitableStrRel _ (PosRelMacroRelStr d) +posRelMacroSuitableRel (constant A) = constantSuitableRel A +posRelMacroSuitableRel var = pointedSuitableRel +posRelMacroSuitableRel (d₀ , d₁) = + productSuitableRel (posRelMacroSuitableRel d₀) (posRelMacroSuitableRel d₁) +posRelMacroSuitableRel (maybe d) = maybeSuitableRel (posRelMacroSuitableRel d) + +posRelMacroPositiveRel : ∀ {ℓ ℓ₁} (d : PosRelDesc ℓ ℓ₁) → PositiveStrRel (posRelMacroSuitableRel d) +posRelMacroPositiveRel (constant A) = constantPositiveRel A +posRelMacroPositiveRel var = pointedPositiveRel +posRelMacroPositiveRel (d₀ , d₁) = + productPositiveRel (posRelMacroPositiveRel d₀) (posRelMacroPositiveRel d₁) +posRelMacroPositiveRel (maybe d) = maybePositiveRel (posRelMacroPositiveRel d) + +relMacroSuitableRel : ∀ {ℓ ℓ₁ ℓ₁'} (d : RelDesc ℓ ℓ₁ ℓ₁') → SuitableStrRel _ (RelMacroRelStr d) +relMacroSuitableRel (constant A) = constantSuitableRel A +relMacroSuitableRel var = pointedSuitableRel +relMacroSuitableRel (d₀ , d₁) = productSuitableRel (relMacroSuitableRel d₀) (relMacroSuitableRel d₁) +relMacroSuitableRel (param A d) = paramSuitableRel A (λ _ → relMacroSuitableRel d) +relMacroSuitableRel (function+ d₀ d₁) = + functionSuitableRel (posRelMacroSuitableRel d₀) (posRelMacroPositiveRel d₀) (relMacroSuitableRel d₁) +relMacroSuitableRel (maybe d) = maybeSuitableRel (relMacroSuitableRel d) + +{- Proof that structured relations and equivalences agree -} + +posRelMacroMatchesEquiv : ∀ {ℓ ℓ₁} (d : PosRelDesc ℓ ℓ₁) + → StrRelMatchesEquiv (PosRelMacroRelStr d) (EquivAction→StrEquiv (transpMacroAction (posRelDesc→TranspDesc d))) +posRelMacroMatchesEquiv (constant A) _ _ _ = idEquiv _ +posRelMacroMatchesEquiv var _ _ _ = idEquiv _ +posRelMacroMatchesEquiv (d₀ , d₁) = + productRelMatchesTransp + (PosRelMacroRelStr d₀) (transpMacroAction (posRelDesc→TranspDesc d₀)) + (PosRelMacroRelStr d₁) (transpMacroAction (posRelDesc→TranspDesc d₁)) + (posRelMacroMatchesEquiv d₀) (posRelMacroMatchesEquiv d₁) +posRelMacroMatchesEquiv (maybe d) = + maybeRelMatchesTransp + (PosRelMacroRelStr d) (transpMacroAction (posRelDesc→TranspDesc d)) + (posRelMacroMatchesEquiv d) + +relMacroMatchesEquiv : ∀ {ℓ ℓ₁ ℓ₁'} (d : RelDesc ℓ ℓ₁ ℓ₁') + → StrRelMatchesEquiv (RelMacroRelStr d) (MacroEquivStr (relDesc→Desc d)) +relMacroMatchesEquiv (constant A) = constantRelMatchesEquiv A +relMacroMatchesEquiv var = pointedRelMatchesEquiv +relMacroMatchesEquiv (d₁ , d₂) = + productRelMatchesEquiv + (RelMacroRelStr d₁) (RelMacroRelStr d₂) + (relMacroMatchesEquiv d₁) (relMacroMatchesEquiv d₂) +relMacroMatchesEquiv (param A d) = + paramRelMatchesEquiv A (λ _ → RelMacroRelStr d) (λ _ → relMacroMatchesEquiv d) +relMacroMatchesEquiv (function+ d₀ d₁) = + functionRelMatchesEquiv+ + (PosRelMacroRelStr d₀) (transpMacroAction (posRelDesc→TranspDesc d₀)) + (RelMacroRelStr d₁) (MacroEquivStr (relDesc→Desc d₁)) + (posRelMacroMatchesEquiv d₀) (relMacroMatchesEquiv d₁) +relMacroMatchesEquiv (maybe d) = + maybeRelMatchesEquiv (RelMacroRelStr d) (relMacroMatchesEquiv d) + +-- Module for easy importing +module RelMacro ℓ {ℓ₁ ℓ₁'} (d : RelDesc ℓ ℓ₁ ℓ₁') where + relation = RelMacroRelStr d + suitable = relMacroSuitableRel d + matches = relMacroMatchesEquiv d + open Macro ℓ (relDesc→Desc d) public diff --git a/Cubical/Structures/Relational/Maybe.agda b/Cubical/Structures/Relational/Maybe.agda new file mode 100644 index 000000000..6c0a1a325 --- /dev/null +++ b/Cubical/Structures/Relational/Maybe.agda @@ -0,0 +1,127 @@ +{- + +Maybe structure: X ↦ Maybe (S X) + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Structures.Relational.Maybe where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Structure +open import Cubical.Foundations.RelationalStructure +open import Cubical.Data.Unit +open import Cubical.Data.Empty +open import Cubical.Data.Maybe +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation as Trunc +open import Cubical.HITs.SetQuotients + +open import Cubical.Structures.Maybe + +private + variable + ℓ ℓ₁ ℓ₁' ℓ₁'' : Level + +-- Structured relations + +MaybeRelStr : {S : Type ℓ → Type ℓ₁} {ℓ₁' : Level} + → StrRel S ℓ₁' → StrRel (λ X → Maybe (S X)) ℓ₁' +MaybeRelStr ρ R = MaybeRel (ρ R) + +maybeSuitableRel : {S : Type ℓ → Type ℓ₁} {ρ : StrRel S ℓ₁'} + → SuitableStrRel S ρ + → SuitableStrRel (MaybeStructure S) (MaybeRelStr ρ) +maybeSuitableRel θ .quo (X , nothing) R _ .fst = nothing , _ +maybeSuitableRel θ .quo (X , nothing) R _ .snd (nothing , _) = refl +maybeSuitableRel θ .quo (X , just s) R c .fst = + just (θ .quo (X , s) R c .fst .fst) , θ .quo (X , s) R c .fst .snd +maybeSuitableRel θ .quo (X , just s) R c .snd (just s' , r) = + cong (λ {(t , r') → just t , r'}) (θ .quo (X , s) R c .snd (s' , r)) +maybeSuitableRel θ .symmetric R {nothing} {nothing} r = _ +maybeSuitableRel θ .symmetric R {just s} {just t} r = θ .symmetric R r +maybeSuitableRel θ .transitive R R' {nothing} {nothing} {nothing} r r' = _ +maybeSuitableRel θ .transitive R R' {just s} {just t} {just u} r r' = θ .transitive R R' r r' +maybeSuitableRel θ .set setX = isOfHLevelMaybe 0 (θ .set setX) +maybeSuitableRel θ .prop propR nothing nothing = isOfHLevelLift 1 isPropUnit +maybeSuitableRel θ .prop propR nothing (just y) = isOfHLevelLift 1 isProp⊥ +maybeSuitableRel θ .prop propR (just x) nothing = isOfHLevelLift 1 isProp⊥ +maybeSuitableRel θ .prop propR (just x) (just y) = θ .prop propR x y + +maybeRelMatchesEquiv : {S : Type ℓ → Type ℓ₁} (ρ : StrRel S ℓ₁') {ι : StrEquiv S ℓ₁''} + → StrRelMatchesEquiv ρ ι + → StrRelMatchesEquiv (MaybeRelStr ρ) (MaybeEquivStr ι) +maybeRelMatchesEquiv ρ μ (X , nothing) (Y , nothing) _ = Lift≃Lift (idEquiv _) +maybeRelMatchesEquiv ρ μ (X , nothing) (Y , just y) _ = Lift≃Lift (idEquiv _) +maybeRelMatchesEquiv ρ μ (X , just x) (Y , nothing) _ = Lift≃Lift (idEquiv _) +maybeRelMatchesEquiv ρ μ (X , just x) (Y , just y) = μ (X , x) (Y , y) + +maybeRelAction : + {S : Type ℓ → Type ℓ₁} {ρ : StrRel S ℓ₁'} + → StrRelAction ρ + → StrRelAction (MaybeRelStr ρ) +maybeRelAction α .actStr f = map-Maybe (α .actStr f) +maybeRelAction α .actStrId s = + funExt⁻ (cong map-Maybe (funExt (α .actStrId))) s ∙ map-Maybe-id s +maybeRelAction α .actRel h nothing nothing = _ +maybeRelAction α .actRel h (just s) (just t) r = α .actRel h s t r + +maybePositiveRel : + {S : Type ℓ → Type ℓ₁} {ρ : StrRel S ℓ₁'} {θ : SuitableStrRel S ρ} + → PositiveStrRel θ + → PositiveStrRel (maybeSuitableRel θ) +maybePositiveRel σ .act = maybeRelAction (σ .act) +maybePositiveRel σ .reflexive nothing = _ +maybePositiveRel σ .reflexive (just s) = σ .reflexive s +maybePositiveRel σ .detransitive R R' {nothing} {nothing} r = ∣ nothing , _ , _ ∣ +maybePositiveRel σ .detransitive R R' {just s} {just u} rr' = + Trunc.map (λ {(t , r , r') → just t , r , r'}) (σ .detransitive R R' rr') +maybePositiveRel {S = S} {ρ = ρ} {θ = θ} σ .quo {X} R = + subst isEquiv + (funExt + (elimProp (λ _ → maybeSuitableRel θ .set squash/ _ _) + (λ {nothing → refl; (just _) → refl}))) + (compEquiv (isoToEquiv isom) (congMaybeEquiv (_ , σ .quo R)) .snd) + where + fwd : Maybe (S X) / MaybeRel (ρ (R .fst .fst)) → Maybe (S X / ρ (R .fst .fst)) + fwd [ nothing ] = nothing + fwd [ just s ] = just [ s ] + fwd (eq/ nothing nothing r i) = nothing + fwd (eq/ (just s) (just t) r i) = just (eq/ s t r i) + fwd (squash/ _ _ p q i j) = + isOfHLevelMaybe 0 squash/ _ _ (cong fwd p) (cong fwd q) i j + + bwd : Maybe (S X / ρ (R .fst .fst)) → Maybe (S X) / MaybeRel (ρ (R .fst .fst)) + bwd nothing = [ nothing ] + bwd (just [ s ]) = [ just s ] + bwd (just (eq/ s t r i)) = eq/ (just s) (just t) r i + bwd (just (squash/ _ _ p q i j)) = + squash/ _ _ (cong (bwd ∘ just) p) (cong (bwd ∘ just) q) i j + + open Iso + isom : Iso (Maybe (S X) / MaybeRel (ρ (R .fst .fst))) (Maybe (S X / ρ (R .fst .fst))) + isom .fun = fwd + isom .inv = bwd + isom .rightInv nothing = refl + isom .rightInv (just x) = + elimProp {B = λ x → fwd (bwd (just x)) ≡ just x} + (λ _ → isOfHLevelMaybe 0 squash/ _ _) + (λ _ → refl) + x + isom .leftInv = elimProp (λ _ → squash/ _ _) (λ {nothing → refl; (just _) → refl}) + +maybeRelMatchesTransp : {S : Type ℓ → Type ℓ₁} + (ρ : StrRel S ℓ₁') (α : EquivAction S) + → StrRelMatchesEquiv ρ (EquivAction→StrEquiv α) + → StrRelMatchesEquiv (MaybeRelStr ρ) (EquivAction→StrEquiv (maybeEquivAction α)) +maybeRelMatchesTransp _ _ μ (X , nothing) (Y , nothing) _ = + isContr→Equiv (isOfHLevelLift 0 isContrUnit) isContr-nothing≡nothing +maybeRelMatchesTransp _ _ μ (X , nothing) (Y , just y) _ = + uninhabEquiv lower ¬nothing≡just +maybeRelMatchesTransp _ _ μ (X , just x) (Y , nothing) _ = + uninhabEquiv lower ¬just≡nothing +maybeRelMatchesTransp _ _ μ (X , just x) (Y , just y) e = + compEquiv (μ (X , x) (Y , y) e) (_ , isEmbedding-just _ _) diff --git a/Cubical/Structures/Relational/Parameterized.agda b/Cubical/Structures/Relational/Parameterized.agda new file mode 100644 index 000000000..e051ddbbf --- /dev/null +++ b/Cubical/Structures/Relational/Parameterized.agda @@ -0,0 +1,70 @@ +{- + +A parameterized family of structures S can be combined into a single structure: +X ↦ (a : A) → S a X + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Relational.Parameterized where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.RelationalStructure +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.SetQuotients + +open import Cubical.Structures.Parameterized + +private + variable + ℓ ℓ₀ ℓ₁ ℓ₁' ℓ₁'' : Level + +-- Structured relations + +module _ (A : Type ℓ₀) where + + ParamRelStr : {S : A → Type ℓ → Type ℓ₁} + → (∀ a → StrRel (S a) ℓ₁') + → StrRel (ParamStructure A S) (ℓ-max ℓ₀ ℓ₁') + ParamRelStr ρ R s t = + (a : A) → ρ a R (s a) (t a) + + paramSuitableRel : {S : A → Type ℓ → Type ℓ₁} {ρ : ∀ a → StrRel (S a) ℓ₁'} + → (∀ a → SuitableStrRel (S a) (ρ a)) + → SuitableStrRel (ParamStructure A S) (ParamRelStr ρ) + paramSuitableRel {ρ = ρ} θ .quo (X , f) R r .fst .fst a = + θ a .quo (X , f a) R (r a) .fst .fst + paramSuitableRel {ρ = ρ} θ .quo (X , f) R r .fst .snd a = + θ a .quo (X , f a) R (r a) .fst .snd + paramSuitableRel {ρ = ρ} θ .quo (X , f) R r .snd (q , c) i .fst a = + θ a .quo (X , f a) R (r a) .snd (q a , c a) i .fst + paramSuitableRel {ρ = ρ} θ .quo (X , f) R r .snd (q , c) i .snd a = + θ a .quo (X , f a) R (r a) .snd (q a , c a) i .snd + paramSuitableRel {ρ = ρ} θ .symmetric R r a = + θ a .symmetric R (r a) + paramSuitableRel {ρ = ρ} θ .transitive R R' r r' a = + θ a .transitive R R' (r a) (r' a) + paramSuitableRel {ρ = ρ} θ .set setX = + isSetΠ λ a → θ a .set setX + paramSuitableRel {ρ = ρ} θ .prop propR s t = + isPropΠ λ a → θ a .prop propR (s a) (t a) + + paramRelMatchesEquiv : {S : A → Type ℓ → Type ℓ₁} + (ρ : ∀ a → StrRel (S a) ℓ₁') {ι : ∀ a → StrEquiv (S a) ℓ₁''} + → (∀ a → StrRelMatchesEquiv (ρ a) (ι a)) + → StrRelMatchesEquiv (ParamRelStr ρ) (ParamEquivStr A ι) + paramRelMatchesEquiv ρ μ _ _ e = equivΠCod λ a → μ a _ _ e + + paramRelAction : {S : A → Type ℓ → Type ℓ₁} {ρ : ∀ a → StrRel (S a) ℓ₁'} + → (∀ a → StrRelAction (ρ a)) + → StrRelAction (ParamRelStr ρ) + paramRelAction α .actStr f s a = α a .actStr f (s a) + paramRelAction α .actStrId s = funExt λ a → α a .actStrId (s a) + paramRelAction α .actRel h _ _ r a = α a .actRel h _ _ (r a) + + -- Detransitivity of ParamRelStr would depend on choice in general, so + -- we don't get positivity diff --git a/Cubical/Structures/Relational/Pointed.agda b/Cubical/Structures/Relational/Pointed.agda new file mode 100644 index 000000000..1531eef10 --- /dev/null +++ b/Cubical/Structures/Relational/Pointed.agda @@ -0,0 +1,56 @@ +{- + +Pointed structure: X ↦ X + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Relational.Pointed where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Structure +open import Cubical.Foundations.RelationalStructure +open import Cubical.Foundations.Univalence +open import Cubical.Relation.ZigZag.Base +open import Cubical.HITs.SetQuotients +open import Cubical.HITs.PropositionalTruncation + +open import Cubical.Structures.Pointed + +private + variable + ℓ : Level + +-- Structured relations + +PointedRelStr : StrRel PointedStructure ℓ +PointedRelStr R = R + +pointedSuitableRel : SuitableStrRel {ℓ = ℓ} PointedStructure PointedRelStr +pointedSuitableRel .quo _ _ _ = isContrSingl _ +pointedSuitableRel .symmetric _ r = r +pointedSuitableRel .transitive _ _ r r' = ∣ _ , r , r' ∣ +pointedSuitableRel .set setX = setX +pointedSuitableRel .prop propR = propR + +pointedRelMatchesEquiv : StrRelMatchesEquiv {ℓ = ℓ} PointedRelStr PointedEquivStr +pointedRelMatchesEquiv _ _ _ = idEquiv _ + +pointedRelAction : StrRelAction {ℓ = ℓ} PointedRelStr +pointedRelAction .actStr f = f +pointedRelAction .actStrId _ = refl +pointedRelAction .actRel α = α + +pointedPositiveRel : PositiveStrRel {ℓ = ℓ} pointedSuitableRel +pointedPositiveRel .act = pointedRelAction +pointedPositiveRel .reflexive x = ∣ refl ∣ +pointedPositiveRel .detransitive R R' rr' = rr' +pointedPositiveRel .quo R = isoToIsEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun = _ + isom .inv q = q + isom .rightInv = elimProp (λ _ → squash/ _ _) (λ _ → refl) + isom .leftInv = elimProp (λ _ → squash/ _ _) (λ _ → refl) diff --git a/Cubical/Structures/Relational/Product.agda b/Cubical/Structures/Relational/Product.agda new file mode 100644 index 000000000..b2148f36e --- /dev/null +++ b/Cubical/Structures/Relational/Product.agda @@ -0,0 +1,140 @@ +{- + +Product of structures S and T: X ↦ S X × T X + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Relational.Product where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Structure +open import Cubical.Foundations.RelationalStructure +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Univalence +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation as Trunc +open import Cubical.HITs.SetQuotients + +open import Cubical.Structures.Product + +private + variable + ℓ ℓ₁ ℓ₁' ℓ₁'' ℓ₂ ℓ₂' ℓ₂'' : Level + +-- Structured relations + +ProductRelStr : + {S₁ : Type ℓ → Type ℓ₁} (ρ₁ : StrRel S₁ ℓ₁') + {S₂ : Type ℓ → Type ℓ₂} (ρ₂ : StrRel S₂ ℓ₂') + → StrRel (ProductStructure S₁ S₂) (ℓ-max ℓ₁' ℓ₂') +ProductRelStr ρ₁ ρ₂ R (s₁ , s₂) (t₁ , t₂) = + ρ₁ R s₁ t₁ × ρ₂ R s₂ t₂ + +productSuitableRel : + {S₁ : Type ℓ → Type ℓ₁} {ρ₁ : StrRel S₁ ℓ₁'} + {S₂ : Type ℓ → Type ℓ₂} {ρ₂ : StrRel S₂ ℓ₂'} + → SuitableStrRel S₁ ρ₁ → SuitableStrRel S₂ ρ₂ + → SuitableStrRel (ProductStructure S₁ S₂) (ProductRelStr ρ₁ ρ₂) +productSuitableRel θ₁ θ₂ .quo (X , s₁ , s₂) R (r₁ , r₂) .fst .fst = + θ₁ .quo (X , s₁) R r₁ .fst .fst , θ₂ .quo (X , s₂) R r₂ .fst .fst +productSuitableRel θ₁ θ₂ .quo (X , s₁ , s₂) R (r₁ , r₂) .fst .snd = + θ₁ .quo (X , s₁) R r₁ .fst .snd , θ₂ .quo (X , s₂) R r₂ .fst .snd +productSuitableRel θ₁ θ₂ .quo (X , s₁ , s₂) R (r₁ , r₂) .snd ((q₁ , q₂) , (c₁ , c₂)) i .fst = + θ₁ .quo (X , s₁) R r₁ .snd (q₁ , c₁) i .fst , θ₂ .quo (X , s₂) R r₂ .snd (q₂ , c₂) i .fst +productSuitableRel θ₁ θ₂ .quo (X , s₁ , s₂) R (r₁ , r₂) .snd ((q₁ , q₂) , (c₁ , c₂)) i .snd = + θ₁ .quo (X , s₁) R r₁ .snd (q₁ , c₁) i .snd , θ₂ .quo (X , s₂) R r₂ .snd (q₂ , c₂) i .snd +productSuitableRel θ₁ θ₂ .symmetric R (r₁ , r₂) = + θ₁ .symmetric R r₁ , θ₂ .symmetric R r₂ +productSuitableRel θ₁ θ₂ .transitive R R' (r₁ , r₂) (r₁' , r₂') = + θ₁ .transitive R R' r₁ r₁' , θ₂ .transitive R R' r₂ r₂' +productSuitableRel θ₁ θ₂ .set setX = + isSet× (θ₁ .set setX) (θ₂ .set setX) +productSuitableRel θ₁ θ₂ .prop propR (s₁ , s₂) (t₁ , t₂) = + isProp× (θ₁ .prop propR s₁ t₁) (θ₂ .prop propR s₂ t₂) + +productRelMatchesEquiv : + {S₁ : Type ℓ → Type ℓ₁} (ρ₁ : StrRel S₁ ℓ₁') {ι₁ : StrEquiv S₁ ℓ₁''} + {S₂ : Type ℓ → Type ℓ₂} (ρ₂ : StrRel S₂ ℓ₂') {ι₂ : StrEquiv S₂ ℓ₂''} + → StrRelMatchesEquiv ρ₁ ι₁ → StrRelMatchesEquiv ρ₂ ι₂ + → StrRelMatchesEquiv (ProductRelStr ρ₁ ρ₂) (ProductEquivStr ι₁ ι₂) +productRelMatchesEquiv ρ₁ ρ₂ μ₁ μ₂ A B e = + Σ-cong-equiv (μ₁ _ _ e) (λ _ → μ₂ _ _ e) + +productRelAction : + {S₁ : Type ℓ → Type ℓ₁} {ρ₁ : StrRel S₁ ℓ₁'} (α₁ : StrRelAction ρ₁) + {S₂ : Type ℓ → Type ℓ₂} {ρ₂ : StrRel S₂ ℓ₂'} (α₂ : StrRelAction ρ₂) + → StrRelAction (ProductRelStr ρ₁ ρ₂) +productRelAction α₁ α₂ .actStr f (s₁ , s₂) = α₁ .actStr f s₁ , α₂ .actStr f s₂ +productRelAction α₁ α₂ .actStrId (s₁ , s₂) = ΣPathP (α₁ .actStrId s₁ , α₂ .actStrId s₂) +productRelAction α₁ α₂ .actRel h _ _ (r₁ , r₂) = α₁ .actRel h _ _ r₁ , α₂ .actRel h _ _ r₂ + +productPositiveRel : + {S₁ : Type ℓ → Type ℓ₁} {ρ₁ : StrRel S₁ ℓ₁'} {θ₁ : SuitableStrRel S₁ ρ₁} + {S₂ : Type ℓ → Type ℓ₂} {ρ₂ : StrRel S₂ ℓ₂'} {θ₂ : SuitableStrRel S₂ ρ₂} + → PositiveStrRel θ₁ + → PositiveStrRel θ₂ + → PositiveStrRel (productSuitableRel θ₁ θ₂) +productPositiveRel σ₁ σ₂ .act = productRelAction (σ₁ .act) (σ₂ .act) +productPositiveRel σ₁ σ₂ .reflexive (s₁ , s₂) = σ₁ .reflexive s₁ , σ₂ .reflexive s₂ +productPositiveRel σ₁ σ₂ .detransitive R R' (rr'₁ , rr'₂) = + Trunc.rec squash + (λ {(s₁ , r₁ , r₁') → + Trunc.rec squash + (λ {(s₂ , r₂ , r₂') → ∣ (s₁ , s₂) , (r₁ , r₂) , (r₁' , r₂') ∣}) + (σ₂ .detransitive R R' rr'₂)}) + (σ₁ .detransitive R R' rr'₁) +productPositiveRel {S₁ = S₁} {ρ₁} {θ₁} {S₂} {ρ₂} {θ₂} σ₁ σ₂ .quo {X} R = + subst isEquiv + (funExt (elimProp (λ _ → productSuitableRel θ₁ θ₂ .set squash/ _ _) (λ _ → refl))) + (compEquiv + (isoToEquiv isom) + (Σ-cong-equiv (_ , σ₁ .quo R) (λ _ → _ , σ₂ .quo R)) .snd) + where + fwd : + ProductStructure S₁ S₂ X / ProductRelStr ρ₁ ρ₂ (R .fst .fst) + → (S₁ X / ρ₁ (R .fst .fst)) × (S₂ X / ρ₂ (R .fst .fst)) + fwd [ s₁ , s₂ ] = [ s₁ ] , [ s₂ ] + fwd (eq/ (s₁ , s₂) (t₁ , t₂) (r₁ , r₂) i) = eq/ s₁ t₁ r₁ i , eq/ s₂ t₂ r₂ i + fwd (squash/ _ _ p q i j) = + isSet× squash/ squash/ _ _ (cong fwd p) (cong fwd q) i j + + bwd[] : S₁ X → S₂ X / ρ₂ (R .fst .fst) + → ProductStructure S₁ S₂ X / ProductRelStr ρ₁ ρ₂ (R .fst .fst) + bwd[] s₁ [ s₂ ] = [ s₁ , s₂ ] + bwd[] s₁ (eq/ s₂ t₂ r₂ i) = + eq/ (s₁ , s₂) (s₁ , t₂) (posRelReflexive σ₁ R s₁ , r₂) i + bwd[] s₁ (squash/ _ _ p q i j) = + squash/ _ _ (λ j → bwd[] s₁ (p j)) (λ j → bwd[] s₁ (q j)) i j + + bwd : S₁ X / ρ₁ (R .fst .fst) → S₂ X / ρ₂ (R .fst .fst) + → ProductStructure S₁ S₂ X / ProductRelStr ρ₁ ρ₂ (R .fst .fst) + bwd [ s₁ ] u = bwd[] s₁ u + bwd (eq/ s₁ t₁ r₁ i) u = path u i + where + path : ∀ u → bwd [ s₁ ] u ≡ bwd [ t₁ ] u + path = elimProp (λ _ → squash/ _ _) (λ s₂ → eq/ (s₁ , s₂) (t₁ , s₂) (r₁ , posRelReflexive σ₂ R s₂)) + bwd (squash/ _ _ p q i j) = + isSetΠ (λ _ → squash/) _ _ (cong bwd p) (cong bwd q) i j + + open Iso + isom : Iso _ _ + isom .fun = fwd + isom .inv = uncurry bwd + isom .rightInv = + uncurry + (elimProp (λ _ → isPropΠ λ _ → isSet× squash/ squash/ _ _) + (λ _ → elimProp (λ _ → isSet× squash/ squash/ _ _) (λ _ → refl))) + isom .leftInv = elimProp (λ _ → squash/ _ _) (λ _ → refl) + +productRelMatchesTransp : + {S₁ : Type ℓ → Type ℓ₁} (ρ₁ : StrRel S₁ ℓ₁') (α₁ : EquivAction S₁) + {S₂ : Type ℓ → Type ℓ₂} (ρ₂ : StrRel S₂ ℓ₂') (α₂ : EquivAction S₂) + → StrRelMatchesEquiv ρ₁ (EquivAction→StrEquiv α₁) + → StrRelMatchesEquiv ρ₂ (EquivAction→StrEquiv α₂) + → StrRelMatchesEquiv (ProductRelStr ρ₁ ρ₂) (EquivAction→StrEquiv (productEquivAction α₁ α₂)) +productRelMatchesTransp _ _ _ _ μ₁ μ₂ _ _ e = + compEquiv (Σ-cong-equiv (μ₁ _ _ e) (λ _ → μ₂ _ _ e)) ΣPath≃PathΣ diff --git a/Cubical/Structures/Transfer.agda b/Cubical/Structures/Transfer.agda new file mode 100644 index 000000000..809fbff53 --- /dev/null +++ b/Cubical/Structures/Transfer.agda @@ -0,0 +1,47 @@ +{- + +Transferring properties of terms between equivalent structures + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Transfer where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Transport +open import Cubical.Structures.Product + +private + variable + ℓ ℓ₀ ℓ₁ ℓ₂ ℓ₂' : Level + +transfer : {ℓ₂' ℓ₀ : Level} {S : Type ℓ → Type ℓ₁} {H : Type ℓ → Type ℓ₂} + (P : ∀ X → S X → H X → Type ℓ₀) + (α : EquivAction H) (τ : TransportStr α) + (ι : StrEquiv S ℓ₂') (θ : UnivalentStr S ι) + {X Y : Type ℓ} {s : S X} {t : S Y} + (e : (X , s) ≃[ ι ] (Y , t)) + → (h : H Y) → P X s (invEq (α (e .fst)) h) → P Y t h +transfer P α τ ι θ {X} {Y} {s} {t} e h = + subst (λ {(Z , u , h) → P Z u h}) + (sip + (productUnivalentStr ι θ (EquivAction→StrEquiv α) (TransportStr→UnivalentStr α τ)) + (X , s , invEq (α (e .fst)) h) + (Y , t , h) + (e .fst , e .snd , secEq (α (e .fst)) h)) + +transfer⁻ : {ℓ₂' ℓ₀ : Level} {S : Type ℓ → Type ℓ₁} {H : Type ℓ → Type ℓ₂} + (P : ∀ X → S X → H X → Type ℓ₀) + (α : EquivAction H) (τ : TransportStr α) + (ι : StrEquiv S ℓ₂') (θ : UnivalentStr S ι) + {X Y : Type ℓ} {s : S X} {t : S Y} + (e : (X , s) ≃[ ι ] (Y , t)) + → (h : H X) → P Y t (equivFun (α (e .fst)) h) → P X s h +transfer⁻ P α τ ι θ {X} {Y} {s} {t} e h = + subst⁻ (λ {(Z , u , h) → P Z u h}) + (sip + (productUnivalentStr ι θ (EquivAction→StrEquiv α) (TransportStr→UnivalentStr α τ)) + (X , s , h) + (Y , t , equivFun (α (e .fst)) h) + (e .fst , e .snd , refl)) diff --git a/Cubical/Structures/TypeEqvTo.agda b/Cubical/Structures/TypeEqvTo.agda new file mode 100644 index 000000000..e142bb913 --- /dev/null +++ b/Cubical/Structures/TypeEqvTo.agda @@ -0,0 +1,41 @@ +{-# OPTIONS --safe #-} +module Cubical.Structures.TypeEqvTo where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.HITs.PropositionalTruncation +open import Cubical.Data.Sigma + +open import Cubical.Foundations.SIP +open import Cubical.Foundations.Pointed +open import Cubical.Structures.Axioms +open import Cubical.Structures.Pointed + +private + variable + ℓ ℓ' ℓ'' : Level + +TypeEqvTo : (ℓ : Level) (X : Type ℓ') → Type (ℓ-max (ℓ-suc ℓ) ℓ') +TypeEqvTo ℓ X = TypeWithStr ℓ (λ Y → ∥ Y ≃ X ∥) + +PointedEqvTo : (ℓ : Level) (X : Type ℓ') → Type (ℓ-max (ℓ-suc ℓ) ℓ') +PointedEqvTo ℓ X = TypeWithStr ℓ (λ Y → Y × ∥ Y ≃ X ∥) + +module _ (X : Type ℓ') where + + PointedEqvToStructure : Type ℓ → Type (ℓ-max ℓ ℓ') + PointedEqvToStructure = AxiomsStructure PointedStructure (λ Y _ → ∥ Y ≃ X ∥) + + PointedEqvToEquivStr : StrEquiv PointedEqvToStructure ℓ'' + PointedEqvToEquivStr = AxiomsEquivStr PointedEquivStr (λ Y _ → ∥ Y ≃ X ∥) + + pointedEqvToUnivalentStr : UnivalentStr {ℓ} PointedEqvToStructure PointedEqvToEquivStr + pointedEqvToUnivalentStr = axiomsUnivalentStr PointedEquivStr {axioms = λ Y _ → ∥ Y ≃ X ∥} + (λ _ _ → squash) pointedUnivalentStr + + PointedEqvToSIP : (A B : PointedEqvTo ℓ X) → A ≃[ PointedEqvToEquivStr ] B ≃ (A ≡ B) + PointedEqvToSIP = SIP pointedEqvToUnivalentStr + + PointedEqvTo-sip : (A B : PointedEqvTo ℓ X) → A ≃[ PointedEqvToEquivStr ] B → (A ≡ B) + PointedEqvTo-sip A B = equivFun (PointedEqvToSIP A B) diff --git a/Cubical/Talks/EPA2020.agda b/Cubical/Talks/EPA2020.agda new file mode 100644 index 000000000..2ab4ddbf7 --- /dev/null +++ b/Cubical/Talks/EPA2020.agda @@ -0,0 +1,341 @@ +{- + +Cubical Agda - A Dependently Typed PL with Univalence and HITs +============================================================== + Anders Mörtberg + Every Proof Assistant - September 17, 2020 + + +Link to slides: https://staff.math.su.se/anders.mortberg/slides/EPA2020.pdf + +Link to video: https://vimeo.com/459020971 + +-} + +-- To make Agda cubical add the following options +{-# OPTIONS --safe --guardedness #-} +module Cubical.Talks.EPA2020 where + +-- The "Foundations" package contain a lot of important results (in +-- particular the univalence theorem) +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Int +open import Cubical.Data.Prod.Base + + +-- The interval in Cubical Agda is written I and the endpoints i0 and i1. + +apply0 : (A : Type) (p : I → A) → A +apply0 A p = p i0 + +-- We omit the universe level ℓ for simplicity in this talk. In the +-- library everything is maximally universe polymorphic. + + +-- We can write functions out of the interval using λ-abstraction: +refl' : {A : Type} (x : A) → x ≡ x +refl' x = λ i → x +-- In fact, x ≡ y is short for PathP (λ _ → A) x y + +refl'' : {A : Type} (x : A) → PathP (λ _ → A) x x +refl'' x = λ _ → x + +-- In general PathP A x y has A : I → Type, x : A i0 and y : A i1 +-- PathP = Dependent paths (Path over a Path) + +-- We cannot pattern-match on interval variables as I is not inductively defined +-- foo : {A : Type} → I → A +-- foo r = {!r!} -- Try typing C-c C-c + +-- cong has a direct proof +cong' : {A B : Type} (f : A → B) {x y : A} → x ≡ y → f x ≡ f y +cong' f p i = f (p i) + +-- function extensionality also has a direct proof. +-- It also has computational content: swap the arguments. +funExt' : {A B : Type} {f g : A → B} (p : (x : A) → f x ≡ g x) → f ≡ g +funExt' p i x = p x i + +-- Transport is more complex as ≡ isn't inductively defined (so we +-- can't define it by pattern-matching on p) +transport' : {A B : Type} → A ≡ B → A → B +transport' p a = transp (λ i → p i) i0 a + +-- This lets us define subst (which is called "transport" in the HoTT book) +subst' : {A : Type} (P : A → Type) {x y : A} (p : x ≡ y) → P x → P y +subst' P p pa = transport (λ i → P (p i)) pa + +-- The transp operation reduces differently for different types +-- formers. For paths it reduces to another primitive operation called +-- hcomp. + +-- We can also define the J eliminator (aka path induction) +J' : {A : Type} {B : A → Type} {x : A} + (P : (z : A) → x ≡ z → Type) + (d : P x refl) {y : A} (p : x ≡ y) → P y p +J' P d p = transport (λ i → P (p i) (λ j → p (i ∧ j))) d + +-- So J is provable, but it doesn't satisfy computation rule +-- definitionally. This is almost never a problem in practice as the +-- cubical primitives satisfy many new definitional equalities. + + + +-- Another key concept in HoTT/UF is the Univalence Axiom. In Cubical +-- Agda this is provable, we hence refer to it as the Univalence +-- Theorem. + +-- The univalence theorem: equivalences of types give paths of types +ua' : {A B : Type} → A ≃ B → A ≡ B +ua' = ua + +-- Any isomorphism of types gives rise to an equivalence +isoToEquiv' : {A B : Type} → Iso A B → A ≃ B +isoToEquiv' = isoToEquiv + +-- And hence to a path +isoToPath' : {A B : Type} → Iso A B → A ≡ B +isoToPath' e = ua' (isoToEquiv' e) + +-- ua satisfies the following computation rule +-- This suffices to be able to prove the standard formulation of univalence. +uaβ' : {A B : Type} (e : A ≃ B) (x : A) + → transport (ua' e) x ≡ fst e x +uaβ' e x = transportRefl (equivFun e x) + + + +-- Time for an example! + +-- Booleans +data Bool : Type where + false true : Bool + +not : Bool → Bool +not false = true +not true = false + +notPath : Bool ≡ Bool +notPath = isoToPath' (iso not not rem rem) + where + rem : (b : Bool) → not (not b) ≡ b + rem false = refl + rem true = refl + +_ : transport notPath true ≡ false +_ = refl + + +-- Another example, integers: + +sucPath : ℤ ≡ ℤ +sucPath = isoToPath' (iso sucℤ predℤ sucPred predSuc) + +_ : transport sucPath (pos 0) ≡ pos 1 +_ = refl + +_ : transport (sucPath ∙ sucPath) (pos 0) ≡ pos 2 +_ = refl + +_ : transport (sym sucPath) (pos 0) ≡ negsuc 0 +_ = refl + + + +----------------------------------------------------------------------------- +-- Higher inductive types + +-- The following definition of finite multisets is due to Vikraman +-- Choudhury and Marcelo Fiore. + +infixr 5 _∷_ + +data FMSet (A : Type) : Type where + [] : FMSet A + _∷_ : (x : A) → (xs : FMSet A) → FMSet A + comm : (x y : A) (xs : FMSet A) → x ∷ y ∷ xs ≡ y ∷ x ∷ xs +-- trunc : (xs ys : FMSet A) (p q : xs ≡ ys) → p ≡ q + +-- We need to add the trunc constructor for FMSets to be sets, omitted +-- here for simplicity. + +_++_ : ∀ {A : Type} (xs ys : FMSet A) → FMSet A +[] ++ ys = ys +(x ∷ xs) ++ ys = x ∷ xs ++ ys +comm x y xs i ++ ys = comm x y (xs ++ ys) i +-- trunc xs zs p q i j ++ ys = +-- trunc (xs ++ ys) (zs ++ ys) (cong (_++ ys) p) (cong (_++ ys) q) i j + +unitr-++ : {A : Type} (xs : FMSet A) → xs ++ [] ≡ xs +unitr-++ [] = refl +unitr-++ (x ∷ xs) = cong (x ∷_) (unitr-++ xs) +unitr-++ (comm x y xs i) j = comm x y (unitr-++ xs j) i +-- unitr-++ (trunc xs ys x y i j) = {!!} + + +-- This is a special case of set quotients! Very useful for +-- programming and set level mathematics + +data _/_ (A : Type) (R : A → A → Type) : Type where + [_] : A → A / R + eq/ : (a b : A) → R a b → [ a ] ≡ [ b ] + trunc : (a b : A / R) (p q : a ≡ b) → p ≡ q + +-- Proving that they are effective ((a b : A) → [ a ] ≡ [ b ] → R a b) +-- requires univalence for propositions. + + +------------------------------------------------------------------------- +-- Topological examples of things that are not sets + +-- We can define the circle as the following simple data declaration: +data S¹ : Type where + base : S¹ + loop : base ≡ base + +-- We can write functions on S¹ using pattern-matching equations: +double : S¹ → S¹ +double base = base +double (loop i) = (loop ∙ loop) i + +helix : S¹ → Type +helix base = ℤ +helix (loop i) = sucPathℤ i + +ΩS¹ : Type +ΩS¹ = base ≡ base + +winding : ΩS¹ → ℤ +winding p = subst helix p (pos 0) + +_ : winding (λ i → double ((loop ∙ loop) i)) ≡ pos 4 +_ = refl + + +-- We can define the Torus as: +data Torus : Type where + point : Torus + line1 : point ≡ point + line2 : point ≡ point + square : PathP (λ i → line1 i ≡ line1 i) line2 line2 + +-- And prove that it is equivalent to two circle: +t2c : Torus → S¹ × S¹ +t2c point = (base , base) +t2c (line1 i) = (loop i , base) +t2c (line2 j) = (base , loop j) +t2c (square i j) = (loop i , loop j) + +c2t : S¹ × S¹ → Torus +c2t (base , base) = point +c2t (loop i , base) = line1 i +c2t (base , loop j) = line2 j +c2t (loop i , loop j) = square i j + +c2t-t2c : (t : Torus) → c2t (t2c t) ≡ t +c2t-t2c point = refl +c2t-t2c (line1 _) = refl +c2t-t2c (line2 _) = refl +c2t-t2c (square _ _) = refl + +t2c-c2t : (p : S¹ × S¹) → t2c (c2t p) ≡ p +t2c-c2t (base , base) = refl +t2c-c2t (base , loop _) = refl +t2c-c2t (loop _ , base) = refl +t2c-c2t (loop _ , loop _) = refl + +-- Using univalence we get the following equality: +Torus≡S¹×S¹ : Torus ≡ S¹ × S¹ +Torus≡S¹×S¹ = isoToPath' (iso t2c c2t t2c-c2t c2t-t2c) + + +windingTorus : point ≡ point → ℤ × ℤ +windingTorus l = ( winding (λ i → proj₁ (t2c (l i))) + , winding (λ i → proj₂ (t2c (l i)))) + +_ : windingTorus (line1 ∙ sym line2) ≡ (pos 1 , negsuc 0) +_ = refl + +-- We have many more topological examples, including Klein bottle, RP^n, +-- higher spheres, suspensions, join, wedges, smash product: +open import Cubical.HITs.KleinBottle +open import Cubical.HITs.RPn +open import Cubical.HITs.S2 +open import Cubical.HITs.S3 +open import Cubical.HITs.Susp +open import Cubical.HITs.Join +open import Cubical.HITs.Wedge +open import Cubical.HITs.SmashProduct + +-- There's also a proof of the "3x3 lemma" for pushouts in less than +-- 200LOC. In HoTT-Agda this took about 3000LOC. For details see: +-- https://github.com/HoTT/HoTT-Agda/tree/master/theorems/homotopy/3x3 +open import Cubical.HITs.Pushout + +-- We also defined the Hopf fibration and proved that its total space +-- is S³ in about 300LOC: +open import Cubical.HITs.Hopf + +-- There is also some integer cohomology: +open import Cubical.ZCohomology.Everything +-- To compute cohomology groups of various spaces we need a bunch of +-- interesting theorems: Freudenthal suspension theorem, +-- Mayer-Vietoris sequence... +open import Cubical.Homotopy.Freudenthal +open import Cubical.ZCohomology.MayerVietorisUnreduced + + +------------------------------------------------------------------------- +-- The structure identity principle + +-- A more efficient version of finite multisets based on association lists +open import Cubical.HITs.AssocList.Base + +-- data AssocList (A : Type) : Type where +-- ⟨⟩ : AssocList A +-- ⟨_,_⟩∷_ : (a : A) (n : ℕ) (xs : AssocList A) → AssocList A +-- per : (a b : A) (m n : ℕ) (xs : AssocList A) +-- → ⟨ a , m ⟩∷ ⟨ b , n ⟩∷ xs ≡ ⟨ b , n ⟩∷ ⟨ a , m ⟩∷ xs +-- agg : (a : A) (m n : ℕ) (xs : AssocList A) +-- → ⟨ a , m ⟩∷ ⟨ a , n ⟩∷ xs ≡ ⟨ a , m + n ⟩∷ xs +-- del : (a : A) (xs : AssocList A) → ⟨ a , 0 ⟩∷ xs ≡ xs +-- trunc : (xs ys : AssocList A) (p q : xs ≡ ys) → p ≡ q + + +-- Programming and proving is more complicated with AssocList compared +-- to FMSet. This kind of example occurs everywhere in programming and +-- mathematics: one representation is easier to work with, but not +-- efficient, while another is efficient but difficult to work with. + +-- Solution: substitute using univalence +substIso : {A B : Type} (P : Type → Type) (e : Iso A B) → P A → P B +substIso P e = subst P (isoToPath e) + +-- Can transport for example Monoid structure from FMSet to AssocList +-- this way, but the achieved Monoid structure is not very efficient +-- to work with. A better solution is to prove that FMSet and +-- AssocList are equal *as monoids*, but how to do this? + +-- Solution: structure identity principle (SIP) +-- This is a very useful consequence of univalence +open import Cubical.Foundations.SIP + +sip' : {ℓ : Level} {S : Type ℓ → Type ℓ} {ι : StrEquiv S ℓ} + (θ : UnivalentStr S ι) (A B : TypeWithStr ℓ S) → A ≃[ ι ] B → A ≡ B +sip' = sip + +-- The tricky thing is to prove that (S,ι) is a univalent structure. +-- Luckily we provide automation for this in the library, see for example: +open import Cubical.Algebra.Monoid.Base + +-- Another cool application of the SIP: matrices represented as +-- functions out of pairs of Fin's and vectors are equal as abelian +-- groups: +open import Cubical.Algebra.Matrix + + +-- The end, back to slides! diff --git a/Cubical/Talks/Everything.agda b/Cubical/Talks/Everything.agda new file mode 100644 index 000000000..a0dc05fb5 --- /dev/null +++ b/Cubical/Talks/Everything.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe --guardedness #-} +module Cubical.Talks.Everything where + +import Cubical.Talks.EPA2020 diff --git a/Cubical/WithK.agda b/Cubical/WithK.agda index e5fc568d3..b0140999a 100644 --- a/Cubical/WithK.agda +++ b/Cubical/WithK.agda @@ -9,7 +9,7 @@ incompatible flags. -} -{-# OPTIONS --cubical --with-K #-} +{-# OPTIONS --with-K #-} module Cubical.WithK where @@ -21,7 +21,7 @@ open import Cubical.Data.Empty private variable ℓ : Level - A : Set ℓ + A : Type ℓ x y : A uip : (prf : x ≡p x) → prf ≡c reflp @@ -37,10 +37,10 @@ transport-not = cong (λ a → transport a true) (ptoc-ctop notEq) false-true : false ≡c true false-true = sym transport-not ∙ transport-uip (ctop notEq) -absurd : (X : Set) → X +absurd : (X : Type) → X absurd X = transport (cong sel false-true) true where - sel : Bool → Set + sel : Bool → Type sel false = Bool sel true = X diff --git a/Cubical/ZCohomology/#Benchmarks.agda# b/Cubical/ZCohomology/#Benchmarks.agda# new file mode 100644 index 000000000..8f05c6f26 --- /dev/null +++ b/Cubical/ZCohomology/#Benchmarks.agda# @@ -0,0 +1,3 @@ +{-# OPTIONS --safe #-} +module Cubical.ZCohomology.Benchmarks where +open import \ No newline at end of file diff --git a/Cubical/ZCohomology/Base.agda b/Cubical/ZCohomology/Base.agda index 1b9fc98ea..da01224fa 100644 --- a/Cubical/ZCohomology/Base.agda +++ b/Cubical/ZCohomology/Base.agda @@ -1,12 +1,11 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe #-} module Cubical.ZCohomology.Base where - -open import Cubical.Data.Int.Base +open import Cubical.Data.Int.Base hiding (_+_) open import Cubical.Data.Nat.Base -open import Cubical.Data.NatMinusTwo.Base open import Cubical.Data.Sigma +open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed.Base open import Cubical.HITs.Nullification.Base @@ -14,32 +13,39 @@ open import Cubical.HITs.SetTruncation.Base open import Cubical.HITs.Sn.Base open import Cubical.HITs.Susp.Base open import Cubical.HITs.Truncation.Base +open import Cubical.Homotopy.Loopspace private variable ℓ : Level A : Type ℓ - --- Cohomology --- -{- Types Kₙ from Brunerie 2016 -} +{- EM-spaces Kₙ from Brunerie 2016 -} coHomK : (n : ℕ) → Type₀ -coHomK zero = Int -coHomK (suc n) = ∥ S₊ (suc n) ∥ (suc (ℕ→ℕ₋₂ n)) +coHomK zero = ℤ +coHomK (suc n) = ∥ S₊ (suc n) ∥ (2 + suc n) {- Cohomology -} coHom : (n : ℕ) → Type ℓ → Type ℓ -coHom n A = ∥ (A → coHomK n) ∥₀ - +coHom n A = ∥ (A → coHomK n) ∥₂ --- Reduced cohomology --- +coHom-pt : (n : ℕ) → coHomK n +coHom-pt 0 = 0 +coHom-pt (suc n) = ∣ (ptSn (suc n)) ∣ + {- Pointed version of Kₙ -} coHomK-ptd : (n : ℕ) → Pointed (ℓ-zero) -coHomK-ptd zero = coHomK zero , (pos 0) -coHomK-ptd (suc n) = (coHomK (suc n) , ∣ north ∣) +coHomK-ptd n = coHomK n , coHom-pt n {- Reduced cohomology -} coHomRed : (n : ℕ) → (A : Pointed ℓ) → Type ℓ -coHomRed n A = ∥ (A →* (coHomK-ptd n)) ∥₀ +coHomRed n A = ∥ A →∙ coHomK-ptd n ∥₂ + +{- Kₙ, untruncated version -} +coHomKType : (n : ℕ) → Type +coHomKType zero = ℤ +coHomKType (suc n) = S₊ (suc n) diff --git a/Cubical/ZCohomology/Benchmarks.agda b/Cubical/ZCohomology/Benchmarks.agda new file mode 100644 index 000000000..959d3a13a --- /dev/null +++ b/Cubical/ZCohomology/Benchmarks.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --safe #-} +module Cubical.ZCohomology.Benchmarks where + + diff --git a/Cubical/ZCohomology/EilenbergSteenrodZ.agda b/Cubical/ZCohomology/EilenbergSteenrodZ.agda new file mode 100644 index 000000000..26b58b778 --- /dev/null +++ b/Cubical/ZCohomology/EilenbergSteenrodZ.agda @@ -0,0 +1,379 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} + +module Cubical.ZCohomology.EilenbergSteenrodZ where + +open import Cubical.Homotopy.EilenbergSteenrod + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Groups.Sn + +open import Cubical.Relation.Nullary + +open import Cubical.Data.Nat +open import Cubical.Data.Bool +open import Cubical.Data.Int +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Unit +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws renaming (assoc to assoc∙) +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +open import Cubical.HITs.SetTruncation + renaming (map to sMap ; rec to sRec ; elim2 to sElim2; elim to sElim ; rec2 to sRec2) +open import Cubical.HITs.PropositionalTruncation + renaming (map to pMap ; rec to pRec ; elim to pElim ; elim2 to pElim2) +open import Cubical.HITs.Truncation + renaming (rec to trRec ; map to trMap ; elim to trEilim ; elim2 to trElim2) +open import Cubical.HITs.Susp +open import Cubical.HITs.Pushout +open import Cubical.HITs.Wedge +open import Cubical.HITs.Sn +open import Cubical.HITs.S1 + +open import Cubical.Algebra.Group renaming (ℤ to ℤGroup) +open import Cubical.Algebra.AbGroup + +open coHomTheory +open Iso +open IsGroup +open GroupStr +open IsGroupHom + +private + suspΩFun' : ∀ {ℓ} {A : Type ℓ} (n : ℕ) (f : A → Path _ (0ₖ _) (0ₖ _)) + → Susp A → coHomK (suc n) + suspΩFun' n f north = 0ₖ _ + suspΩFun' n f south = 0ₖ _ + suspΩFun' n f (merid a i) = f a i + + suspΩFun : ∀ {ℓ} {A : Type ℓ} (n : ℕ) (f : A → coHomK n) → Susp A → coHomK (suc n) + suspΩFun n f = suspΩFun' n λ a → Kn→ΩKn+1 n (f a) + + ≡suspΩFun : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) (f : Susp (typ A) → coHomK (suc n)) (p : f north ≡ 0ₖ _) + → f ≡ suspΩFun' n λ a → sym p ∙∙ (cong f (merid a)) ∙∙ (cong f (sym (merid (pt A))) ∙ p) + ≡suspΩFun n s f p i north = p i + ≡suspΩFun n s f p i south = (cong f (sym (merid (pt s))) ∙ p) i + ≡suspΩFun n s f p i (merid a j) = + doubleCompPath-filler (sym p) (cong f (merid a)) (cong f (sym (merid (pt s))) ∙ p) i j + +-- induction principle for Hⁿ(Susp A) +SuspCohomElim : ∀ {ℓ ℓ'} {A : Pointed ℓ} (n : ℕ) {B : coHom (suc n) (Susp (typ A)) → Type ℓ'} + → ((x : coHom (suc n) (Susp (typ A))) → isProp (B x)) + → ((f : typ A → Path _ (0ₖ _) (0ₖ _)) → f (pt A) ≡ refl → B ∣ suspΩFun' n f ∣₂) + → (x : _) → B x +SuspCohomElim {A = A} n {B = B} isprop f = + coHomPointedElim _ north isprop λ g gid + → subst (B ∘ ∣_∣₂) (sym (≡suspΩFun n A g gid)) + (f _ ((λ i → (sym gid ∙∙ (λ j → g (merid (pt A) (~ i ∧ j))) ∙∙ ((λ j → g (merid (pt A) (~ i ∧ ~ j))) ∙ gid))) + ∙∙ (λ i → (λ j → gid (i ∨ ~ j)) ∙∙ refl ∙∙ lUnit (λ j → gid (i ∨ j)) (~ i)) + ∙∙ sym (rUnit refl))) + +-- (Reduced) cohomology functor +coHomFunctor : {ℓ : Level} (n : ℤ) → Pointed ℓ → AbGroup ℓ +coHomFunctor (pos n) = coHomRedGroup n +coHomFunctor (negsuc n) _ = trivialAbGroup + +-- Alternative definition with reduced groups replaced by unreduced one for n ≥ 1 +coHomFunctor' : {ℓ : Level} (n : ℤ) → Pointed ℓ → AbGroup ℓ +coHomFunctor' (pos zero) = coHomFunctor 0 +coHomFunctor' (pos (suc n)) A = coHomGroup (suc n) (typ A) +coHomFunctor' (negsuc n) = coHomFunctor (negsuc n) + +coHomFunctor≡coHomFunctor' : ∀ {ℓ} → coHomFunctor {ℓ} ≡ coHomFunctor' +coHomFunctor≡coHomFunctor' = funExt λ {(pos zero) → refl + ; (pos (suc n)) → funExt λ A → sym (coHomGroup≡coHomRedGroup n A) + ; (negsuc n) → refl} + +-- Ĥ⁰(Susp A) is contractible +H0-susp : ∀ {ℓ} {A : Pointed ℓ} → isContr (coHomRed 0 (Susp (typ A) , north)) +fst H0-susp = 0ₕ∙ _ +snd (H0-susp {A = A}) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ {(f , p) + → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) + (funExt λ {north → sym p + ; south → sym p ∙ cong f (merid (pt A)) + ; (merid a i) j → isSet→isSet' (isSetℤ) + (sym p) + (sym p ∙ cong f (merid (pt A))) + refl (cong f (merid a)) i j}))} +-- We need that Hⁿ⁺¹(Susp A) ≃ Hⁿ(A) + +private + suspFunCharacFun : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → ((Susp (typ A)) → coHomK (suc n)) + → (typ A → (coHomK n)) + suspFunCharacFun {A = A} n f x = ΩKn+1→Kn n (sym (rCancelₖ (suc n) (f north)) + ∙∙ cong (λ x → f x -[ (suc n) ]ₖ f north) ((merid x) ∙ sym (merid (pt A))) + ∙∙ rCancelₖ (suc n) (f north)) + + linvLem : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f : typ A → Path (coHomK (suc n)) (0ₖ _) (0ₖ _)) + (fId : f (pt A) ≡ refl) + → (x : _) → suspΩFun n (suspFunCharacFun {A = A} n (suspΩFun' n f)) x + ≡ suspΩFun' n f x + linvLem n f fId north = refl + linvLem n f fId south = refl + linvLem {A = A} n f fId (merid x i) j = helper n x f fId j i + where + helper : (n : ℕ) (a : typ A) (f : typ A → Path (coHomK (suc n)) (0ₖ _) (0ₖ _)) + (fId : f (pt A) ≡ refl) + → Kn→ΩKn+1 n (suspFunCharacFun {A = A} n (suspΩFun' n f) a) ≡ f a + helper zero a f fId = + Iso.rightInv (Iso-Kn-ΩKn+1 0) (sym (rCancelₖ 1 (0ₖ 1)) + ∙∙ cong (λ x → suspΩFun' 0 f x +ₖ 0ₖ 1) (merid a ∙ (sym (merid (pt A)))) + ∙∙ (rCancelₖ 1 (0ₖ 1))) + ∙∙ sym (rUnit _) + ∙∙ (λ j i → rUnitₖ 1 (congFunct (suspΩFun' 0 f) (merid a) (sym (merid (pt A))) j i) j) + ∙∙ cong (λ p → f a ∙ sym p) fId + ∙∙ sym (rUnit (f a)) + helper (suc n) a f fId = + Iso.rightInv (Iso-Kn-ΩKn+1 (suc n)) + ((sym (rCancelₖ _ (0ₖ (suc (suc n)))) + ∙∙ cong (λ x → suspΩFun' (suc n) f x +ₖ 0ₖ (suc (suc n))) (merid a ∙ (sym (merid (pt A)))) + ∙∙ rCancelₖ _ (0ₖ (suc (suc n))))) + ∙∙ (((λ i → sym (rCancel≡refl n i) + ∙∙ cong (λ x → rUnitₖ (suc (suc n)) (suspΩFun' (suc n) f x) i) (merid a ∙ (sym (merid (pt A)))) + ∙∙ rCancel≡refl n i))) + ∙∙ ((λ i → rUnit (congFunct (suspΩFun' (suc n) f) (merid a) (sym (merid (pt A))) i) (~ i))) + ∙∙ cong (λ p → f a ∙ sym p) fId + ∙∙ sym (rUnit (f a)) + + +suspFunCharac : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → Iso (coHom (suc (suc n)) (Susp (typ A))) (coHom (suc n) (typ A)) +fun (suspFunCharac {A = A} n) = + sMap λ f → suspFunCharacFun {A = A} (suc n) f +inv (suspFunCharac {A = A} n) = sMap (suspΩFun (suc n)) +rightInv (suspFunCharac {A = A} n) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f → trRec (isProp→isOfHLevelSuc n (isSetSetTrunc _ _)) + (λ fId → cong ∣_∣₂ + (funExt (λ x → cong (ΩKn+1→Kn (suc n)) + ((λ i → sym (rCancel≡refl n i) ∙∙ cong (λ x → suspΩFun (suc n) f x +ₖ 0ₖ _) + (merid x ∙ sym (merid (pt A))) ∙∙ rCancel≡refl n i) + ∙∙ sym (rUnit (cong (λ x → suspΩFun (suc n) f x +ₖ 0ₖ _) + ((merid x) ∙ sym (merid (pt A))))) + ∙∙ λ i → congFunct (λ x → rUnitₖ _ (suspΩFun (suc n) f x) i) + (merid x) (sym (merid (pt A))) i) + ∙∙ ΩKn+1→Kn-hom (suc n) (Kn→ΩKn+1 (suc n) (f x)) + (sym (Kn→ΩKn+1 (suc n) (f (pt A)))) + ∙∙ cong₂ _+ₖ_ (Iso.leftInv (Iso-Kn-ΩKn+1 (suc n)) (f x)) + (cong (λ x → ΩKn+1→Kn (suc n) (sym (Kn→ΩKn+1 (suc n) x))) fId) + ∙∙ cong (λ y → f x +ₖ ΩKn+1→Kn (suc n) (sym y)) (Kn→ΩKn+10ₖ (suc n)) + ∙∙ cong (f x +ₖ_) (ΩKn+1→Kn-refl (suc n)) + ∙ rUnitₖ _ (f x)))) + (fst (isConnectedPathKn n (f (pt A)) (0ₖ _))) +leftInv (suspFunCharac {A = A} n) = + SuspCohomElim {A = A} _ (λ _ → isSetSetTrunc _ _) + λ f fId → cong ∣_∣₂ (funExt (linvLem (suc n) f fId)) + +-- We also need that H¹(Susp A) ≃ Ĥ⁰(A) +suspFunCharac0 : ∀ {ℓ} {A : Pointed ℓ} → Iso (∥ ((Susp (typ A)) → coHomK 1) ∥₂) ∥ A →∙ (ℤ , 0) ∥₂ +fun (suspFunCharac0 {A = A}) = + sMap λ f → suspFunCharacFun {A = A} 0 f + , (cong (ΩKn+1→Kn 0) ((λ i → sym (rCancelₖ _ (f north)) + ∙∙ cong (λ x → f x -ₖ f north) (rCancel (merid (pt A)) i) + ∙∙ rCancelₖ _ (f north)) + ∙∙ (doubleCompPath-elim (sym (rCancelₖ _ (f north))) refl (rCancelₖ _ (f north))) + ∙∙ (cong (_∙ (rCancelₖ _ (f north))) (sym (rUnit (sym (rCancelₖ _ (f north)))))) + ∙ (lCancel (rCancelₖ _ (f north))))) +inv suspFunCharac0 = sMap λ f → suspΩFun 0 (fst f) +rightInv (suspFunCharac0 {A = A}) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ {(f , p) + → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) + (funExt (λ x → (λ j → transp (λ i → helix (wedgeMapS¹ (intLoop (p j) (~ i)) base)) j + ((transport (λ i → helix (trRec isGroupoidS¹ (λ x → x) + (rUnitₖ 1 ∣ intLoop (f x) i ∣ j))) + (pos 0)))) + ∙ windingℤLoop (f x))))} +leftInv (suspFunCharac0 {A = A}) = + SuspCohomElim {A = A} _ (λ _ → isSetSetTrunc _ _) + λ f fId → cong ∣_∣₂ (funExt (linvLem 0 f fId)) + +-- We now prove that the alternative definition of cohomology is a cohomology theory. +private + -- First, we need to that coHomFunctor' is contravariant + theMorph : ∀ {ℓ} (n : ℤ) {A B : Pointed ℓ} (f : A →∙ B) + → AbGroupHom (coHomFunctor' n B) (coHomFunctor' n A) + fst (theMorph (pos zero) f) = sMap λ g → (λ x → fst g (fst f x)) , cong (fst g) (snd f) ∙ snd g + snd (theMorph (pos zero) f) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f g → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl)) + theMorph (pos (suc n)) f = coHomMorph _ (fst f) + fst (theMorph (negsuc n) f) = idfun _ + snd (theMorph (negsuc n) f) = makeIsGroupHom λ _ _ → refl + + open coHomTheory + isCohomTheoryZ' : ∀ {ℓ} → coHomTheory {ℓ} coHomFunctor' + Hmap isCohomTheoryZ' = theMorph + + -------------------------- Suspension -------------------------- + -- existence of suspension isomorphism + fst (Suspension isCohomTheoryZ') (pos zero) {A = A} = + invGroupEquiv + (GroupIso→GroupEquiv + ( invIso suspFunCharac0 + , makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f g → cong ∣_∣₂ (funExt λ { north → refl + ; south → refl + ; (merid a i) j → helper a (fst f) (fst g) j i})))) + where + helper : (a : typ A) (f g : typ A → coHomK 0) + → Kn→ΩKn+1 0 (f a +[ 0 ]ₖ g a) + ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 0 (f a)) (Kn→ΩKn+1 0 (g a)) + helper a f g = Kn→ΩKn+1-hom 0 (f a) (g a) + ∙ ∙≡+₁ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a)) + fst (Suspension isCohomTheoryZ') (pos (suc n)) {A = A} = + invGroupEquiv + (GroupIso→GroupEquiv + ( invIso (suspFunCharac {A = A} n) + , makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f g → cong ∣_∣₂ (funExt λ { north → refl + ; south → refl + ; (merid a i) j → helper a f g j i})))) + where + helper : (a : typ A) (f g : typ A → coHomK (suc n)) + → Kn→ΩKn+1 (suc n) (f a +ₖ g a) + ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc n) (f a)) (Kn→ΩKn+1 (suc n) (g a)) + helper a f g = Kn→ΩKn+1-hom (suc n) (f a) (g a) + ∙ ∙≡+₂ n (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a)) + fst (Suspension isCohomTheoryZ') (negsuc zero) {A = A} = + GroupIso→GroupEquiv + ( isContr→Iso (H0-susp {A = _ , pt A}) isContrUnit* + , makeIsGroupHom λ _ _ → refl) + fst (Suspension isCohomTheoryZ') (negsuc (suc n)) = idGroupEquiv + + -- naturality of the suspension isomorphism + snd (Suspension (isCohomTheoryZ' {ℓ})) (f , p) (pos zero) = + funExt (sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ {(f , _) → cong ∣_∣₂ (funExt λ {north → refl + ; south → refl + ; (merid a i) → refl})}) + snd (Suspension (isCohomTheoryZ' {ℓ})) (f , p) (pos (suc n)) = + funExt (sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f → cong ∣_∣₂ (funExt λ {north → refl + ; south → refl + ; (merid a i) → refl})) + snd (Suspension (isCohomTheoryZ' {ℓ})) (f , p) (negsuc zero) = refl + snd (Suspension (isCohomTheoryZ' {ℓ})) (f , p) (negsuc (suc n)) = refl + + -------------------------- Exactness --------------------------- + Exactness isCohomTheoryZ' {A = A} {B = B} f n = isoToPath (exactnessIso n f) + where + exactnessIso : (n : ℤ) (f : A →∙ B) + → Iso (Ker (theMorph n f)) (Im (theMorph n (cfcod (fst f) , refl))) + fun (exactnessIso (pos zero) (f , p)) = + uncurry (sElim (λ _ → isSetΠ λ _ → isSetΣ isSetSetTrunc λ _ → isProp→isSet isPropPropTrunc) + λ {(g , q) inker → ∣ g , q ∣₂ + , pRec isPropPropTrunc + (λ gId → ∣ ∣ (λ { (inl tt) → 0 + ; (inr b) → g b + ; (push a i) → funExt⁻ (cong fst gId) a (~ i)}) , q ∣₂ + , cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl) ∣) + (Iso.fun PathIdTrunc₀Iso inker)}) + inv (exactnessIso (pos zero) (f , p)) = + uncurry (sElim (λ _ → isSetΠ λ _ → isSetΣ isSetSetTrunc λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ {(g , q) inim' + → ∣ g , q ∣₂ + , pRec (isSetSetTrunc _ _) + (uncurry + (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _)) + (λ pushmap pushId' + → pRec (isSetSetTrunc _ _) + (λ pushId + → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) + (funExt λ x → sym (funExt⁻ (cong fst pushId) (f x)) + ∙∙ cong (fst pushmap) (sym (push x) ∙ push (pt A)) + ∙∙ (cong (fst pushmap ∘ inr) p + ∙ snd pushmap)))) + (Iso.fun PathIdTrunc₀Iso pushId')))) + inim'}) + rightInv (exactnessIso (pos zero) (f , p)) = + uncurry (sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 + (isSetΣ isSetSetTrunc + (λ _ → isProp→isSet isPropPropTrunc)) _ _) + λ {(p , q) _ → Σ≡Prop (λ _ → isPropPropTrunc) refl}) + leftInv (exactnessIso (pos zero) (f , p)) = + uncurry (sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 + (isSetΣ isSetSetTrunc + (λ _ → isProp→isSet (isSetSetTrunc _ _))) _ _) + λ {(p , q) _ → Σ≡Prop (λ _ → isSetSetTrunc _ _) refl}) + fun (exactnessIso (pos (suc n)) f) ker = (fst ker) , inIm-helper (fst ker) (snd ker) + where + inIm-helper : (x : coHom (suc n) (typ B)) + → isInKer (theMorph (pos (suc n)) {A = A} {B = B} f) x + → isInIm (theMorph (pos (suc n)) {A = B} {B = _ , inr (pt B)} (cfcod (fst f) , refl)) x + inIm-helper = + coHomPointedElim _ (pt B) (λ _ → isPropΠ λ _ → isPropPropTrunc) + λ g gId inker → pRec isPropPropTrunc + (λ gIdTot → ∣ ∣ (λ { (inl tt) → 0ₖ _ + ; (inr b) → g b + ; (push a i) → funExt⁻ gIdTot a (~ i)}) ∣₂ + , cong ∣_∣₂ (funExt λ b → refl) ∣) + (Iso.fun PathIdTrunc₀Iso inker) + inv (exactnessIso (pos (suc n)) f) im = fst im , inKer-helper (fst im) (snd im) + where + inKer-helper : (x : coHom (suc n) (typ B)) + → isInIm (theMorph (pos (suc n)) {A = B} {B = _ , inr (pt B)} (cfcod (fst f) , refl)) x + → isInKer (theMorph (pos (suc n)) {A = A} {B = B} f) x + inKer-helper = + coHomPointedElim _ (pt B) (λ _ → isPropΠ λ _ → isSetSetTrunc _ _) + λ g gId → pRec (isSetSetTrunc _ _) + (uncurry λ cg p + → subst (isInKer (coHomMorph (suc n) (fst f))) + p + (helper cg)) + where + helper : (cg : _) → coHomFun (suc n) (fst f) (coHomFun (suc n) (cfcod (fst f)) cg) ≡ 0ₕ _ + helper = sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ cg → trRec (isProp→isOfHLevelSuc n (isSetSetTrunc _ _)) + (λ p → (cong ∣_∣₂ (funExt λ x → cong cg (sym (push x)) + ∙ p))) + (isConnectedPathKn _ (cg (inl tt)) (0ₖ (suc n)) .fst) + rightInv (exactnessIso (pos (suc n)) f) _ = Σ≡Prop (λ _ → isPropPropTrunc) refl + leftInv (exactnessIso (pos (suc n)) f) _ = Σ≡Prop (λ _ → isSetSetTrunc _ _) refl + exactnessIso (negsuc n) (f , p) = + isContr→Iso ((tt* , refl) + , λ {(tt* , p) → Σ≡Prop (λ _ → isOfHLevelPath 1 isPropUnit* _ _) + refl}) + ((tt* , ∣ tt* , refl ∣) + , λ {(tt* , p) → Σ≡Prop (λ _ → isPropPropTrunc) + refl}) + + -------------------------- Dimension --------------------------- + Dimension isCohomTheoryZ' (pos zero) p = ⊥-rec (p refl) + fst (Dimension isCohomTheoryZ' (pos (suc n)) _) = 0ₕ _ + snd (Dimension isCohomTheoryZ' (pos (suc n)) _) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (λ f → trRec (isProp→isOfHLevelSuc n (isSetSetTrunc _ _)) + (λ f-true → trRec (isProp→isOfHLevelSuc n (isSetSetTrunc _ _)) + (λ f-false → cong ∣_∣₂ (funExt (λ {(lift true) → f-true + ; (lift false) → f-false}))) + (isConnectedPathKn n (0ₖ _) (f (lift false)) .fst)) + (isConnectedPathKn n (0ₖ _) (f (lift true)) .fst)) + Dimension isCohomTheoryZ' (negsuc n) _ = isContrUnit* + + ------------------------ Binary wedges ------------------------- + BinaryWedge isCohomTheoryZ' (pos zero) = GroupIso→GroupEquiv (H⁰Red-⋁ _ _) + BinaryWedge isCohomTheoryZ' (pos (suc n)) = GroupIso→GroupEquiv (Hⁿ-⋁ _ _ n) + BinaryWedge isCohomTheoryZ' (negsuc n) = + GroupIso→GroupEquiv + (compGroupIso (contrGroupIsoUnit isContrUnit*) + (invGroupIso (contrGroupIsoUnit (isOfHLevel× 0 isContrUnit* isContrUnit*)))) + +-- Substituting back for our original theory, we are done +isCohomTheoryZ : ∀ {ℓ} → coHomTheory {ℓ} coHomFunctor +isCohomTheoryZ = subst coHomTheory (sym coHomFunctor≡coHomFunctor') + isCohomTheoryZ' diff --git a/Cubical/ZCohomology/Everything.agda b/Cubical/ZCohomology/Everything.agda index 160104dce..fa3417d28 100644 --- a/Cubical/ZCohomology/Everything.agda +++ b/Cubical/ZCohomology/Everything.agda @@ -1,7 +1,20 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe --guardedness #-} module Cubical.ZCohomology.Everything where -open import Cubical.ZCohomology.Base -open import Cubical.ZCohomology.Properties - -open import Cubical.ZCohomology.S1.S1 +import Cubical.ZCohomology.Base +import Cubical.ZCohomology.EilenbergSteenrodZ +import Cubical.ZCohomology.GroupStructure +import Cubical.ZCohomology.Groups.Connected +import Cubical.ZCohomology.Groups.KleinBottle +import Cubical.ZCohomology.Groups.Prelims +import Cubical.ZCohomology.Groups.RP2 +import Cubical.ZCohomology.Groups.Sn +import Cubical.ZCohomology.Groups.Torus +import Cubical.ZCohomology.Groups.Unit +import Cubical.ZCohomology.Groups.Wedge +import Cubical.ZCohomology.Groups.WedgeOfSpheres +import Cubical.ZCohomology.MayerVietorisUnreduced +import Cubical.ZCohomology.Properties +import Cubical.ZCohomology.RingStructure.CupProduct +import Cubical.ZCohomology.RingStructure.GradedCommutativity +import Cubical.ZCohomology.RingStructure.RingLaws diff --git a/Cubical/ZCohomology/FunLoopSpace.agda b/Cubical/ZCohomology/FunLoopSpace.agda deleted file mode 100644 index 215e12a0b..000000000 --- a/Cubical/ZCohomology/FunLoopSpace.agda +++ /dev/null @@ -1,205 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.FunLoopSpace where - - - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.FunExtEquiv -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Data.Nat -open import Cubical.Data.HomotopyGroup - -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - - - - - -{- Beware: Use of J -} -cancelReflMid : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) (q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q -cancelReflMid {ℓ = ℓ}{A = A} {a = a} {b = b} p q = J {ℓ} {A} {a} {ℓ} (λ b p → ((q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q)) (λ r → sym (lUnit (refl ∙ r ))) {b} p q - ---- - - -FunLR : (n : ℕ) {f : A → B} → - ((typ ((Ω^ (suc n)) ((A → B) , f))) → - (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) -LR-id : (n : ℕ) {f : A → B} → FunLR n {f} refl ≡ (λ x → refl) -FunLR zero {f} = funExt⁻ -FunLR (suc n) {f} p x = cong (λ y → y x) (sym (LR-id n)) ∙ (funExt⁻ (cong (FunLR n) p ) x) ∙ cong (λ y → y x) (LR-id n) -LR-id zero {f} = refl -LR-id (suc n) {f} = funExt λ y → (cancelReflMid (λ i → (LR-id n) (~ i) y) - (λ i → (LR-id n {f}) i y) ∙ - rCancel (λ i → (LR-id n {f}) (~ i) y)) - -{- -lemma2 : (n : ℕ) {f : A → B} {x : A} → PathP (λ i → LR-id n {f} i x ≡ λ _ → pt ((Ω^ n) (B , f x))) (cong (λ y → y x) ((LR-id n {f}))) refl -lemma2 zero = refl -lemma2 (suc n) {f} {x} j = hcomp {!!} {!(cong (λ y → y x) (funExt (λ y → cancelReflMid (λ i → LR-id n {f} (~ (i)) y) (λ i → LR-id n {f} (i ∨ j) y) ∙ rCancel (λ i → LR-id n ((~ i)) y)))) !} --} - - -FunRL : (n : ℕ) {f : A → B} → - ((x : A) → (typ ((Ω^ (suc n)) (B , f x)))) → - ((typ ((Ω^ (suc n)) ((A → B) , f)))) -RL-id : (n : ℕ) {f : A → B} → - FunRL n (λ x → pt (((Ω^ (suc n)) (B , f x)))) ≡ pt ((Ω^ (suc n)) ((A → B) , f)) -FunRL zero = funExt -FunRL (suc n) {f} = λ g → (sym (RL-id n {f}) ∙ cong (FunRL n {f}) (funExt g) ∙ (RL-id n {f})) -RL-id zero = refl -RL-id (suc n) {f} = cancelReflMid (λ i → (RL-id n {f}) (~ i)) (RL-id n) ∙ - rCancel (λ i → (RL-id n) (~ i)) - - - ----- - - - -wantToShow : (n : ℕ) (f : A → B) → - ((x : A) → (typ ((Ω^ n) (B , f x)))) ≡ (typ ((Ω^ n) ((A → B) , f))) -wantToShow zero f = refl -wantToShow (suc zero) f = ua funExtEquiv -wantToShow (suc (suc n)) f = isoToPath (iso (FunRL (suc n)) (FunLR (suc n)) (λ b → {!!}) {!!}) - where - hauptLemma : (n : ℕ) (f : A → B) (b : typ (Ω ((Ω^ n) ((A → B) , f)))) → FunRL n (FunLR n b ) ≡ b - hauptLemma zero f b = refl - hauptLemma (suc n) f b = {!!} - where - - oi : {!λ j → ((λ i → RL-id n {f} (~ i)) ∙ - (λ i → (hauptLemma n f (b i) (i))) ∙ - (λ i → RL-id n {f} i))!} - oi = {!!} - - - firstTransp : PathP (λ j → RL-id n (~ j) ≡ RL-id n (~ j)) - (FunRL (suc n) (FunLR (suc n) b )) - ((λ i → FunRL n {f} (funExt (λ x → (λ i₁ → LR-id n {f} (~ i₁) x) ∙ (λ i₁ → FunLR n {f} (b i₁) x) ∙ (λ i₁ → LR-id n {f} i₁ x)) i))) - firstTransp j = hcomp (λ k → λ{(j = i0) → FunRL (suc n) (FunLR (suc n) b ) ; - (j = i1) → (sym (lUnit ((λ i → FunRL n (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ - (λ i₁ → FunLR n {f} (b i₁) x) ∙ (λ i₁ → LR-id n {f} i₁ x)) i)) ∙ - (λ i → RL-id n {f} (i0)))) ∙ - sym (rUnit (λ i → FunRL n (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ - (λ i₁ → FunLR n {f} (b i₁) x) ∙ - (λ i₁ → LR-id n {f} i₁ x)) i)))) k }) - ((λ i → RL-id n {f} (~ i ∧ ( ~ j))) ∙ - (λ i → FunRL n (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ - (λ i₁ → FunLR n {f} (b i₁) x) ∙ - (λ i₁ → LR-id n {f} i₁ x)) i)) ∙ - (λ i → RL-id n {f} (i ∧ (~ j)))) - - SNDtransp : PathP (λ j → (FunRL n {f} (λ x → LR-id n {f} (~ j) x)) ≡ (FunRL n {f} (λ x → LR-id n {f} (~ j) x))) - ((λ i → FunRL n {f} (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ - (λ i₁ → FunLR n {f} (b i₁) x) ∙ - (λ i₁ → LR-id n {f} i₁ x)) i))) λ i → FunRL n {f} (FunLR n {f} (b i)) - SNDtransp j = hcomp (λ k → λ{ (j = i0) → - ((λ i → FunRL n (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ - (λ i₁ → FunLR n {f} (b i₁) x) ∙ - (λ i₁ → LR-id n {f} i₁ x)) i))) - ; (j = i1) → {!(FunRL n {f} (λ x → LR-id n {f} (~ i1) x))!}}) - (λ i → FunRL n {f} (funExt (λ x → (λ i₁ → LR-id n {f} (~ i₁ ∧ (~ j)) x) ∙ - (λ i₁ → FunLR n {f} (b i₁) x) ∙ - (λ i₁ → LR-id n {f} (i₁ ∧ (~ j)) x)) i)) - - 3transp : PathP (λ j → (hauptLemma n f refl j) ≡ (hauptLemma n f refl j)) (λ i → FunRL n {f} (FunLR n {f} (b i))) b - 3transp j = (λ i → (hauptLemma n f (b i) j)) - - test : (λ j₂ → RL-id n {f} (~ j₂) ≡ RL-id n {f} (~ j₂)) ∙ (λ i → FunRL n {f} (LR-id n {f} (~ i)) ≡ FunRL n {f} (LR-id n {f} (~ i))) ∙ (λ i → hauptLemma n f refl i ≡ hauptLemma n f refl i) ≡ refl - test j = {!((λ j₂ → RL-id n {f} ((~ j₂) ∨ j) ≡ RL-id n {f} ((~ j₂) ∨ j)) ∙ (λ i → FunRL n {f} (LR-id n {f} (?)) ≡ FunRL n {f} (LR-id n {f} (?))))!} - where - - - - postulate - wowiecomp : (i : I) → FunRL (n) (funExt (λ x → (λ i₁ → LR-id (n) (~ i₁) x) ∙ - (λ i₁ → FunLR (n) (b i₁) x) ∙ - (λ i₁ → LR-id (n) {f} i₁ x)) i) ≡ b i - - - - needToShow : ((λ j₁ → ((λ j₂ → RL-id n {f} (~ j₂) ≡ RL-id n {f} (~ j₂)) ∙ (λ i → FunRL n {f} (LR-id n (~ i)) ≡ FunRL n {f} (LR-id n {f} (~ i)))) - j₁) ∙ (λ i → hauptLemma n f refl i ≡ hauptLemma n f refl i)) ≡ refl - needToShow j k = {!((λ j₁ → ((λ j₂ → RL-id n {f} (~ j₂) ≡ RL-id n {f} (~ j₂)) ∙ (λ i → FunRL n {f} (LR-id n (~ i)) ≡ FunRL n {f} (LR-id n {f} (~ i))))!} - - -{- -where - - linvLemma : (n : ℕ) {f : A → B} {b : typ ((Ω^ (suc (suc (suc n)))) ((A → B) , f))} → - (λ i → RL-id (suc n) (~ i)) ∙ - (λ i → FunRL (suc n) (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x) ∙ - (λ i₁ → FunLR (suc n) (b i₁) x) ∙ - (λ i₁ → LR-id (suc n) {f} i₁ x)) i)) ∙ - RL-id (suc n) {f} - ≡ b - linvLemma n {f} {b} j = {!hcomp ? ( (λ i → RL-id (suc n) (~ i ∧ j)) ∙ - (λ i → FunRL (suc n) (funExt (λ x → (λ i₁ → LR-id (suc n) ((~ i₁)) x) ∙ - (λ i₁ → FunLR (suc n) (b i₁) x) ∙ - (λ i₁ → LR-id (suc n) {f} (i₁ ) x)) (i))) ∙ - (λ i → (RL-id (suc n) {f} (i ∧ j)))) !} - where - cunt : (i : I) → (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x) ∙ (λ i₁ → FunLR (suc n) {f} (b i₁) x) ∙ (λ i₁ → LR-id (suc n) {f} i₁ x )) i ) - ≡ (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x)) ∙ funExt (λ x → (λ i₁ → FunLR (suc n) {f} (b i₁) x)) ∙ funExt (λ x → (λ i₁ → LR-id (suc n) {f} i₁ x ))) i - cunt i = refl - - cunt2 : (i : I) → (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x) ∙ (λ i₁ → FunLR (suc n) {f} (b i₁) x) ∙ (λ i₁ → LR-id (suc n) {f} i₁ x )) i ) - ≡ (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x)) ∙ (λ j → FunLR (suc n) {f} (b j)) ∙ funExt (λ x → (λ i₁ → LR-id (suc n) {f} i₁ x ))) i - cunt2 i = refl -{- - cunt3 : (i : I) → FunRL n (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ (λ i₁ → FunLR n (b i₁) x) ∙ (λ i₁ → LR-id n {f} i₁ x)) i) ≡ FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i) - cunt3 i = refl - - cunt4 : (i : I) → FunRL (suc n) (((λ i₁ → LR-id (suc n) {f} (~ i₁)) ∙ (λ j → FunLR (suc n) {f} (b j)) ∙ (λ i₁ → LR-id (suc n) {f} i₁)) i) ≡ FunRL n {f} ( LR-id n {f} (~ i)) ∙ FunRL n (FunLR n {f} (b i)) ∙ FunRL n {f} ( LR-id n {f} i) - cunt4 i = {!!} - - -{- - main2 : (λ i → RL-id n (~ i)) ∙ (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) ∙ (λ i → RL-id n i) ≡ {!!} - main2 j i = {!FunRL n (((λ i₁ → LR-id n {f} (~ i₁ ∧ i0)) ∙ (λ k → FunLR n {f} (b (k))) ∙ (λ i₁ → LR-id n {f} (i₁ ∧ i0))) i0)!} - - main3 : PathP (λ j → (RL-id n {f} j ≡ RL-id n j)) (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) ((λ i → RL-id n (~ i)) ∙ (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) ∙ (λ i → RL-id n {f} i)) - main3 j = hcomp (λ k → λ{(j = i0) → {!((λ i → RL-id n ((~ i) ∧ j))!} ; (j = i1) → {!!}}) ((λ i → RL-id n ((~ i) ∧ j)) ∙ (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) ∙ (λ i → RL-id n (i ∧ j))) - - main4 : PathP (λ j → FunRL n (LR-id n j) ≡ FunRL n (LR-id n j)) {!LR-id n {f} i0!} (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) - main4 j = hcomp (λ k → {! LR-id n {f} (i0)!}) (λ i → FunRL n {f} (((λ i₁ → LR-id n {f} (~ i₁ ∧ j)) ∙ (λ k → FunLR n {f} (b (k))) ∙ (λ i₁ → LR-id n {f} (i₁ ∧ j))) i)) --} - - wowie : (i : I) → ( FunRL (suc n) (((λ i₁ → LR-id (suc n) {f} (~ i₁)) ∙ (λ j → FunLR (suc n) {f} (b j)) ∙ (λ i₁ → LR-id (suc n) {f} i₁)) i)) ≡ (FunRL (suc n) (((λ i₁ → (FunLR (suc n) refl)) ∙ (λ j → FunLR (suc n) {f} (b j)) ∙ (λ i₁ → (FunLR (suc n) refl))) i)) - wowie i = sym (λ j → FunRL (suc n) (((λ i₁ → LR-id (suc n) {f} (~ i₁ ∧ j)) ∙ (λ k → FunLR (suc n) {f} (b (k))) ∙ (λ i₁ → LR-id (suc n) {f} (i₁ ∧ j))) i)) - - wowie2 : (i : I) → (FunRL (suc n) (((λ i₁ → (FunLR (suc n) refl)) ∙ (λ j → FunLR (suc n) {f} (b j)) ∙ (λ i₁ → (FunLR (suc n) refl))) i)) ≡ (FunRL (suc n) ((( FunLR (suc n) {f} (b i))) )) - wowie2 i = (λ j → FunRL (suc n) {f} ((sym (lUnit ((λ j → FunLR (suc n) {f} (b j)) ∙ refl)) j) i) ) ∙ ((λ j → FunRL (suc n) {f} ((sym (rUnit ((λ j → FunLR (suc n) {f} (b j)))) j) i))) - - main4 : PathP (λ j → FunRL n (LR-id n j) ≡ FunRL n (LR-id n j)) {!LR-id n {f} i0!} (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) - main4 j = hcomp (λ k → {! LR-id n {f} (i0)!}) (λ i → FunRL n {f} (((λ i₁ → LR-id n {f} (~ i₁ ∧ j)) ∙ (λ k → FunLR n {f} (b (k))) ∙ (λ i₁ → LR-id n {f} (i₁ ∧ j))) i)) - - postulate - fixLater : (i : I) → (FunRL (suc n) ((( FunLR (suc n) {f} (b i))) )) ≡ b i - - wowie3 : - (λ i → FunRL (suc n) (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x) ∙ - (λ i₁ → FunLR (suc n) (b i₁) x) ∙ - (λ i₁ → LR-id (suc n) {f} i₁ x)) i)) - ≡ - {!b!} - wowie3 j = {!FunRL!} - - wowiecomp : (i : I) → FunRL (suc n) (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x) ∙ - (λ i₁ → FunLR (suc n) (b i₁) x) ∙ - (λ i₁ → LR-id (suc n) {f} i₁ x)) i) ≡ b i - wowiecomp i = wowie i ∙ wowie2 i ∙ fixLater {!wowiecomp i0!} - - --} - --} diff --git a/Cubical/ZCohomology/FunLoopSpace2.agda b/Cubical/ZCohomology/FunLoopSpace2.agda deleted file mode 100644 index ae60e32b6..000000000 --- a/Cubical/ZCohomology/FunLoopSpace2.agda +++ /dev/null @@ -1,34 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.FunLoopSpace2 where - - - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.FunExtEquiv -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Data.Nat -open import Cubical.Data.HomotopyGroup -open import Cubical.HITs.Nullification -open import Cubical.HITs.Sn -open import Cubical.HITs.Susp -open import Cubical.HITs.Truncation -open import Cubical.Data.NatMinusTwo - -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - - - - - -{- Beware: Use of J -} -cancelReflMid : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) (q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q -cancelReflMid {ℓ = ℓ}{A = A} {a = a} {b = b} p q = J {ℓ} {A} {a} {ℓ} (λ b p → ((q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q)) (λ r → sym (lUnit (refl ∙ r ))) {b} p q - diff --git a/Cubical/ZCohomology/GroupStructure.agda b/Cubical/ZCohomology/GroupStructure.agda new file mode 100644 index 000000000..9b8244f8d --- /dev/null +++ b/Cubical/ZCohomology/GroupStructure.agda @@ -0,0 +1,919 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.GroupStructure where + +open import Cubical.ZCohomology.Base + +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed hiding (id) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws renaming (assoc to assoc∙) +open import Cubical.Data.Sigma +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; isSetSetTrunc to §) +open import Cubical.Data.Int renaming (_+_ to _ℤ+_ ; -_ to -ℤ_) +open import Cubical.Data.Nat renaming (+-assoc to +-assocℕ ; +-comm to +-commℕ) +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3 ; map2 to trMap2) +open import Cubical.Homotopy.Loopspace +open import Cubical.Algebra.Group renaming (ℤ to ℤGroup) +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid + +open Iso renaming (inv to inv') + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + A' : Pointed ℓ + +infixr 34 _+ₖ_ +infixr 34 _+ₕ_ +infixr 34 _+ₕ∙_ + +-- Addition in the Eilenberg-Maclane spaces is uniquely determined if we require it to have left- and right-unit laws, +-- such that these agree on 0. In particular, any h-structure (see http://ericfinster.github.io/files/emhott.pdf) is unique. ++ₖ-unique : (n : ℕ) → (comp1 comp2 : coHomK (suc n) → coHomK (suc n) → coHomK (suc n)) + → (rUnit1 : (x : _) → comp1 x (coHom-pt (suc n)) ≡ x) + → (lUnit1 : (x : _) → comp1 (coHom-pt (suc n)) x ≡ x) + → (rUnit2 : (x : _) → comp2 x (coHom-pt (suc n)) ≡ x) + → (lUnit2 : (x : _) → comp2 (coHom-pt (suc n)) x ≡ x) + → (unId1 : rUnit1 (coHom-pt (suc n)) ≡ lUnit1 (coHom-pt (suc n))) + → (unId2 : rUnit2 (coHom-pt (suc n)) ≡ lUnit2 (coHom-pt (suc n))) + → (x y : _) → comp1 x y ≡ comp2 x y ++ₖ-unique n comp1 comp2 rUnit1 lUnit1 rUnit2 lUnit2 unId1 unId2 = + elim2 (λ _ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) + (wedgeconFun _ _ + (λ _ _ → help _ _) + (λ x → lUnit1 ∣ x ∣ ∙ sym (lUnit2 ∣ x ∣)) + (λ x → rUnit1 ∣ x ∣ ∙ sym (rUnit2 ∣ x ∣)) + (cong₂ _∙_ unId1 (cong sym unId2))) + where + help : isOfHLevel (2 + (n + suc n)) (coHomK (suc n)) + help = subst (λ x → isOfHLevel x (coHomK (suc n))) (+-suc n (2 + n) ∙ +-suc (suc n) (suc n)) + (isOfHLevelPlus n (isOfHLevelTrunc (3 + n))) + +wedgeConHLev : (n : ℕ) → isOfHLevel ((2 + n) + (2 + n)) (coHomK (2 + n)) +wedgeConHLev n = subst (λ x → isOfHLevel x (coHomK (2 + n))) + (sym (+-suc (2 + n) (suc n) ∙ +-suc (3 + n) n)) + (isOfHLevelPlus' {n = n} (4 + n) (isOfHLevelTrunc (4 + n))) +wedgeConHLev' : (n : ℕ) → isOfHLevel ((2 + n) + (2 + n)) (typ (Ω (coHomK-ptd (3 + n)))) +wedgeConHLev' n = subst (λ x → isOfHLevel x (typ (Ω (coHomK-ptd (3 + n))))) + (sym (+-suc (2 + n) (suc n) ∙ +-suc (3 + n) n)) + (isOfHLevelPlus' {n = n} (4 + n) (isOfHLevelTrunc (5 + n) _ _)) + +wedgeConHLevPath : (n : ℕ) → (x y : coHomK (suc n)) → isOfHLevel ((suc n) + (suc n)) (x ≡ y) +wedgeConHLevPath zero x y = isOfHLevelTrunc 3 _ _ +wedgeConHLevPath (suc n) x y = isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev n) _ _ + +-- addition for n ≥ 2 together with the left- and right-unit laws (modulo truncations) +preAdd : (n : ℕ) → (S₊ (2 + n) → S₊ (2 + n) → coHomK (2 + n)) +preAdd n = + wedgeconFun _ _ + (λ _ _ → wedgeConHLev n) + ∣_∣ + ∣_∣ + refl + +preAdd-l : (n : ℕ) → (x : (S₊ (2 + n))) → preAdd n north x ≡ ∣ x ∣ +preAdd-l n _ = refl + +preAdd-r : (n : ℕ) → (x : (S₊ (2 + n))) → preAdd n x north ≡ ∣ x ∣ +preAdd-r n = + wedgeconRight _ (suc n) + (λ _ _ → wedgeConHLev n) + ∣_∣ + ∣_∣ + refl + +-- addition for n = 1 +wedgeMapS¹ : S¹ → S¹ → S¹ +wedgeMapS¹ base y = y +wedgeMapS¹ (loop i) base = loop i +wedgeMapS¹ (loop i) (loop j) = + hcomp (λ k → λ { (i = i0) → loop j + ; (i = i1) → loop (j ∧ k) + ; (j = i0) → loop i + ; (j = i1) → loop (i ∧ k)}) + (loop (i ∨ j)) + +---------- Algebra/Group stuff -------- +0ₖ : (n : ℕ) → coHomK n +0ₖ = coHom-pt + +_+ₖ_ : {n : ℕ} → coHomK n → coHomK n → coHomK n +_+ₖ_ {n = zero} x y = x ℤ+ y +_+ₖ_ {n = suc zero} = trMap2 wedgeMapS¹ +_+ₖ_ {n = suc (suc n)} = trRec (isOfHLevelΠ (4 + n) λ _ → isOfHLevelTrunc (4 + n)) + λ x → trRec (isOfHLevelTrunc (4 + n)) (preAdd n x) + +private + isEquiv+ : (n : ℕ) → (x : coHomK (suc n)) → isEquiv (_+ₖ_ {n = (suc n)} x) + isEquiv+ zero = + trElim (λ _ → isProp→isOfHLevelSuc 2 (isPropIsEquiv _)) + (toPropElim (λ _ → isPropIsEquiv _) + (subst isEquiv (sym help) (idIsEquiv _))) + where + help : _+ₖ_ {n = 1} (coHom-pt 1) ≡ idfun _ + help = funExt (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ _ → refl) + isEquiv+ (suc n) = + trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isPropIsEquiv _)) + (suspToPropElim (ptSn (suc n)) (λ _ → isPropIsEquiv _) + (subst isEquiv (sym help) (idIsEquiv _))) + where + help : _+ₖ_ {n = (2 + n)} (coHom-pt (2 + n)) ≡ idfun _ + help = funExt (trElim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) λ _ → refl) + + + Kₙ≃Kₙ : (n : ℕ) (x : coHomK (suc n)) → coHomK (suc n) ≃ coHomK (suc n) + Kₙ≃Kₙ n x = _ , isEquiv+ n x + +-ₖ_ : {n : ℕ} → coHomK n → coHomK n +-ₖ_ {n = zero} x = 0 - x +-ₖ_ {n = suc zero} = trMap λ {base → base ; (loop i) → (loop (~ i))} +-ₖ_ {n = suc (suc n)} = trMap λ {north → north + ; south → north + ; (merid a i) → ((merid (ptSn (suc n)) ∙ sym (merid a))) i} + +_-ₖ_ : {n : ℕ} → coHomK n → coHomK n → coHomK n +_-ₖ_ {n = n} x y = _+ₖ_ {n = n} x (-ₖ_ {n = n} y) + ++ₖ-syntax : (n : ℕ) → coHomK n → coHomK n → coHomK n ++ₖ-syntax n = _+ₖ_ {n = n} + +-ₖ-syntax : (n : ℕ) → coHomK n → coHomK n +-ₖ-syntax n = -ₖ_ {n = n} + +-'ₖ-syntax : (n : ℕ) → coHomK n → coHomK n → coHomK n +-'ₖ-syntax n = _-ₖ_ {n = n} + +syntax +ₖ-syntax n x y = x +[ n ]ₖ y +syntax -ₖ-syntax n x = -[ n ]ₖ x +syntax -'ₖ-syntax n x y = x -[ n ]ₖ y + +-ₖ^2 : {n : ℕ} → (x : coHomK n) → (-ₖ (-ₖ x)) ≡ x +-ₖ^2 {n = zero} x = + +Comm (pos zero) (-ℤ (pos zero ℤ+ (-ℤ x))) ∙∙ -Dist+ (pos zero) (-ℤ x) + ∙∙ (+Comm (pos zero) (-ℤ (-ℤ x)) ∙ -Involutive x) +-ₖ^2 {n = suc zero} = + trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ { base → refl ; (loop i) → refl} +-ₖ^2 {n = suc (suc n)} = + trElim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + λ { north → refl + ; south j → ∣ merid (ptSn _) j ∣ₕ + ; (merid a i) j + → hcomp (λ k → λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ compPath-filler' (merid a) (sym (merid (ptSn _))) (~ k) (~ j) ∣ₕ + ; (j = i0) → help a (~ k) i + ; (j = i1) → ∣ merid a (i ∧ k) ∣}) + ∣ (merid a ∙ sym (merid (ptSn _))) (i ∧ ~ j) ∣ₕ} + where + help : (a : _) → cong (-ₖ_ ∘ (-ₖ_ {n = suc (suc n)})) (cong ∣_∣ₕ (merid a)) + ≡ cong ∣_∣ₕ (merid a ∙ sym (merid (ptSn _))) + help a = cong (cong ((-ₖ_ {n = suc (suc n)}))) (cong-∙ ∣_∣ₕ (merid (ptSn (suc n))) (sym (merid a))) + ∙∙ cong-∙ (-ₖ_ {n = suc (suc n)}) (cong ∣_∣ₕ (merid (ptSn (suc n)))) (cong ∣_∣ₕ (sym (merid a))) + ∙∙ (λ i → (λ j → ∣ rCancel (merid (ptSn (suc n))) i j ∣ₕ) + ∙ λ j → ∣ symDistr (merid (ptSn (suc n))) (sym (merid a)) i j ∣ₕ) + ∙ sym (lUnit _) + + +------- Groupoid Laws for Kₙ --------- +commₖ : (n : ℕ) → (x y : coHomK n) → x +[ n ]ₖ y ≡ y +[ n ]ₖ x +commₖ zero = +Comm +commₖ (suc zero) = + elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + (wedgeconFun _ _ + (λ _ _ → isOfHLevelTrunc 3 _ _) + (λ {base → refl ; (loop i) → refl}) + (λ {base → refl ; (loop i) → refl}) + refl) +commₖ (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + (wedgeconFun _ _ + (λ x y → isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev n) _ _) + (λ x → preAdd-l n x ∙ sym (preAdd-r n x)) + (λ x → preAdd-r n x ∙ sym (preAdd-l n x)) + refl) + +commₖ-base : (n : ℕ) → commₖ n (coHom-pt n) (coHom-pt n) ≡ refl +commₖ-base zero = refl +commₖ-base (suc zero) = refl +commₖ-base (suc (suc n)) = sym (rUnit _) + +rUnitₖ : (n : ℕ) → (x : coHomK n) → x +[ n ]ₖ coHom-pt n ≡ x +rUnitₖ zero x = refl +rUnitₖ (suc zero) = + trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ {base → refl + ; (loop i) → refl} +rUnitₖ (suc (suc n)) = + trElim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + (preAdd-r n) + +lUnitₖ : (n : ℕ) → (x : coHomK n) → coHom-pt n +[ n ]ₖ x ≡ x +lUnitₖ zero x = sym (pos0+ x) +lUnitₖ (suc zero) = + trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ {base → refl + ; (loop i) → refl} +lUnitₖ (suc (suc n)) = + trElim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + λ x → refl + +∙≡+₁ : (p q : typ (Ω (coHomK-ptd 1))) → p ∙ q ≡ cong₂ _+ₖ_ p q +∙≡+₁ p q = (λ i → (λ j → rUnitₖ 1 (p j) (~ i)) ∙ λ j → lUnitₖ 1 (q j) (~ i)) ∙ sym (cong₂Funct _+ₖ_ p q) + +∙≡+₂ : (n : ℕ) (p q : typ (Ω (coHomK-ptd (suc (suc n))))) → p ∙ q ≡ cong₂ _+ₖ_ p q +∙≡+₂ n p q = (λ i → (λ j → rUnitₖ (2 + n) (p j) (~ i)) ∙ λ j → lUnitₖ (2 + n) (q j) (~ i)) ∙ sym (cong₂Funct _+ₖ_ p q) + +lCancelₖ : (n : ℕ) → (x : coHomK n) → (-ₖ_ {n = n} x) +ₖ x ≡ coHom-pt n +lCancelₖ zero x = minusPlus x 0 +lCancelₖ (suc zero) = + trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ {base → refl ; (loop i) j → help j i} + where + help : cong₂ _+ₖ_ (sym (cong ∣_∣ loop)) (cong ∣_∣ loop) ≡ refl + help = sym (∙≡+₁ (sym (cong ∣_∣ loop)) (cong ∣_∣ loop)) ∙ lCancel _ +lCancelₖ (suc (suc n)) = + trElim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + λ {north → refl ; south → cong ∣_∣ (sym (merid (ptSn (suc n)))) ; (merid a i) → help a i } + where + s : (a : _) → _ ≡ _ + s x = cong₂ _+ₖ_ (sym (cong ∣_∣ (merid (ptSn (suc n)) ∙ sym (merid x)))) (cong ∣_∣ (sym (merid x))) + + help : (a : _) → PathP (λ i → (preAdd n ((merid (ptSn (suc n)) ∙ (λ i₁ → merid a (~ i₁))) i) + (merid a i)) ≡ ∣ north ∣) refl λ i₁ → ∣ merid (ptSn (suc n)) (~ i₁) ∣ + help x = + compPathR→PathP + ((sym (lCancel _) + ∙∙ (λ i → ∙≡+₂ _ (cong ∣_∣ (symDistr (merid x) (sym (merid (ptSn (suc n)))) i)) (cong ∣_∣ ((merid x) ∙ sym (merid (ptSn (suc n))))) i) + ∙∙ rUnit _) + ∙∙ (λ j → cong₂ _+ₖ_ ((cong ∣_∣ (merid (ptSn (suc n)) ∙ sym (merid x)))) + (λ i → ∣ compPath-filler ((merid x)) ((sym (merid (ptSn (suc n))))) (~ j) i ∣) + ∙ λ i → ∣ merid (ptSn (suc n)) (~ i ∧ j) ∣) + ∙∙ λ i → sym (s x) ∙ rUnit (cong ∣_∣ (sym (merid (ptSn (suc n))))) i) + +rCancelₖ : (n : ℕ) → (x : coHomK n) → x +ₖ (-ₖ_ {n = n} x) ≡ coHom-pt n +rCancelₖ zero x = +Comm x (pos 0 - x) ∙ minusPlus x 0 -- +-comm x (pos 0 - x) ∙ minusPlus x 0 +rCancelₖ (suc zero) = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ {base → refl ; (loop i) j → help j i} + where + help : (λ i → ∣ loop i ∣ₕ +ₖ (-ₖ ∣ loop i ∣ₕ)) ≡ refl + help = sym (∙≡+₁ (cong ∣_∣ₕ loop) (sym (cong ∣_∣ₕ loop))) ∙ rCancel _ +rCancelₖ (suc (suc n)) x = commₖ _ x (-ₖ x) ∙ lCancelₖ _ x + +rCancel≡refl : (n : ℕ) → rCancelₖ (2 + n) (0ₖ _) ≡ refl +rCancel≡refl n i = rUnit (rUnit refl (~ i)) (~ i) + +assocₖ : (n : ℕ) → (x y z : coHomK n) → x +[ n ]ₖ (y +[ n ]ₖ z) ≡ (x +[ n ]ₖ y) +[ n ]ₖ z +assocₖ zero = +Assoc +assocₖ (suc zero) = + trElim3 (λ _ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ x → wedgeconFun _ _ + (λ _ _ → isOfHLevelTrunc 3 _ _) + (λ y i → rUnitₖ 1 ∣ x ∣ (~ i) +ₖ ∣ y ∣) + (λ z → cong (∣ x ∣ +ₖ_) (rUnitₖ 1 ∣ z ∣) ∙ sym (rUnitₖ 1 (∣ x ∣ +ₖ ∣ z ∣))) + (helper x) + where + helper : (x : S¹) → cong (∣ x ∣ +ₖ_) (rUnitₖ 1 ∣ base ∣) ∙ sym (rUnitₖ 1 (∣ x ∣ +ₖ ∣ base ∣)) + ≡ (cong (_+ₖ ∣ base ∣) (sym (rUnitₖ 1 ∣ x ∣))) + helper = toPropElim (λ _ → isOfHLevelTrunc 3 _ _ _ _) + (sym (lUnit refl)) + +assocₖ (suc (suc n)) = + trElim3 (λ _ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + (wedgeConSn-×3 _ + (λ x z i → preAdd-r n x (~ i) +ₖ ∣ z ∣) + (λ x y → cong (∣ x ∣ +ₖ_) (rUnitₖ (2 + n) ∣ y ∣) ∙ sym (rUnitₖ (2 + n) (∣ x ∣ +ₖ ∣ y ∣))) + (lUnit (sym (rUnitₖ (2 + n) (∣ north ∣ +ₖ ∣ north ∣))))) + where + wedgeConSn-×3 : (n : ℕ) + → (f : (x z : S₊ (2 + n))→ ∣ x ∣ +ₖ ((0ₖ _) +ₖ ∣ z ∣) ≡ (∣ x ∣ +ₖ (0ₖ _)) +ₖ ∣ z ∣) + → (g : (x y : S₊ (2 + n)) → ∣ x ∣ +ₖ (∣ y ∣ +ₖ 0ₖ _) ≡ (∣ x ∣ +ₖ ∣ y ∣) +ₖ 0ₖ _) + → (f (ptSn _) (ptSn _) ≡ g (ptSn _) (ptSn _)) + → (x y z : S₊ (2 + n)) → ∣ x ∣ +ₖ (∣ y ∣ +ₖ ∣ z ∣) ≡ (∣ x ∣ +ₖ ∣ y ∣) +ₖ ∣ z ∣ + wedgeConSn-×3 n f g d x = + wedgeconFun _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev n) _ _) + (f x) + (g x) + (sphereElim _ {A = λ x → g x (ptSn (suc (suc n))) ≡ f x (ptSn (suc (suc n))) } + (λ _ → isOfHLevelTrunc (4 + n) _ _ _ _) + (sym d) x) +{- +This was the original proof for the case n ≥ 2: +For some reason it doesn't check in reasonable time anymore: +assocₖ (suc (suc n)) = + trElim3 (λ _ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + λ x → wedgeConSn _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev n) _ _) + (λ z i → preAdd n .snd .snd x (~ i) +ₖ ∣ z ∣) + (λ y → cong (∣ x ∣ +ₖ_) (rUnitₖ (2 + n) ∣ y ∣) ∙ sym (rUnitₖ (2 + n) (∣ x ∣ +ₖ ∣ y ∣))) + (helper x) .fst + where + helper : (x : S₊ (2 + n)) → cong (∣ x ∣ +ₖ_) (rUnitₖ (2 + n) ∣ north ∣) ∙ sym (rUnitₖ (2 + n) (∣ x ∣ +ₖ ∣ north ∣)) + ≡ cong (_+ₖ ∣ north ∣) (sym (preAdd n .snd .snd x)) + helper = sphereElim (suc n) (λ _ → isOfHLevelTrunc (4 + n) _ _ _ _) + (sym (lUnit (sym (rUnitₖ (2 + n) (∣ north ∣ +ₖ ∣ north ∣))))) +-} + + +lUnitₖ≡rUnitₖ : (n : ℕ) → lUnitₖ n (coHom-pt n) ≡ rUnitₖ n (coHom-pt n) +lUnitₖ≡rUnitₖ zero = isSetℤ _ _ _ _ +lUnitₖ≡rUnitₖ (suc zero) = refl +lUnitₖ≡rUnitₖ (suc (suc n)) = refl + +------ Commutativity of ΩKₙ +-- We show that p ∙ q ≡ (λ i → (p i) +ₖ (q i)) for any p q : ΩKₙ₊₁. This allows us to prove that p ∙ q ≡ q ∙ p +-- without having to use the equivalence Kₙ ≃ ΩKₙ₊₁ + + +cong+ₖ-comm : (n : ℕ) (p q : typ (Ω (coHomK-ptd (suc n)))) → cong₂ _+ₖ_ p q ≡ cong₂ _+ₖ_ q p +cong+ₖ-comm zero p q = + rUnit (cong₂ _+ₖ_ p q) + ∙∙ (λ i → (λ j → commₖ 1 ∣ base ∣ ∣ base ∣ (i ∧ j)) + ∙∙ (λ j → commₖ 1 (p j) (q j) i) + ∙∙ λ j → commₖ 1 ∣ base ∣ ∣ base ∣ (i ∧ ~ j)) + ∙∙ ((λ i → commₖ-base 1 i ∙∙ cong₂ _+ₖ_ q p ∙∙ sym (commₖ-base 1 i)) + ∙ sym (rUnit (cong₂ _+ₖ_ q p))) +cong+ₖ-comm (suc n) p q = + rUnit (cong₂ _+ₖ_ p q) + ∙∙ (λ i → (λ j → commₖ (2 + n) ∣ north ∣ ∣ north ∣ (i ∧ j)) + ∙∙ (λ j → commₖ (2 + n) (p j) (q j) i ) + ∙∙ λ j → commₖ (2 + n) ∣ north ∣ ∣ north ∣ (i ∧ ~ j)) + ∙∙ ((λ i → commₖ-base (2 + n) i ∙∙ cong₂ _+ₖ_ q p ∙∙ sym (commₖ-base (2 + n) i)) + ∙ sym (rUnit (cong₂ _+ₖ_ q p))) + +isCommΩK : (n : ℕ) → isComm∙ (coHomK-ptd n) +isCommΩK zero p q = isSetℤ _ _ (p ∙ q) (q ∙ p) +isCommΩK (suc zero) p q = ∙≡+₁ p q ∙∙ cong+ₖ-comm 0 p q ∙∙ sym (∙≡+₁ q p) +isCommΩK (suc (suc n)) p q = ∙≡+₂ n p q ∙∙ cong+ₖ-comm (suc n) p q ∙∙ sym (∙≡+₂ n q p) + +----- some other useful lemmas about algebra in Kₙ +-0ₖ : {n : ℕ} → -[ n ]ₖ (0ₖ n) ≡ (0ₖ n) +-0ₖ {n = zero} = refl +-0ₖ {n = suc zero} = refl +-0ₖ {n = suc (suc n)} = refl + +-distrₖ : (n : ℕ) (x y : coHomK n) → -[ n ]ₖ (x +[ n ]ₖ y) ≡ (-[ n ]ₖ x) +[ n ]ₖ (-[ n ]ₖ y) +-distrₖ zero x y = GroupTheory.invDistr ℤGroup x y ∙ +Comm (0 - y) (0 - x) +-distrₖ (suc zero) = + elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + (wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 3 _ _) + (λ x → sym (lUnitₖ 1 (-[ 1 ]ₖ ∣ x ∣))) + (λ x → cong (λ x → -[ 1 ]ₖ x) (rUnitₖ 1 ∣ x ∣) ∙ sym (rUnitₖ 1 (-[ 1 ]ₖ ∣ x ∣))) + (sym (rUnit refl))) +-distrₖ (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + (wedgeconFun _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev n) _ _) + (λ x → sym (lUnitₖ (2 + n) (-[ (2 + n) ]ₖ ∣ x ∣))) + (λ x → cong (λ x → -[ (2 + n) ]ₖ x) (rUnitₖ (2 + n) ∣ x ∣ ) ∙ sym (rUnitₖ (2 + n) (-[ (2 + n) ]ₖ ∣ x ∣))) + (sym (rUnit refl))) + +-cancelRₖ : (n : ℕ) (x y : coHomK n) → (y +[ n ]ₖ x) -[ n ]ₖ x ≡ y +-cancelRₖ zero x y = sym (+Assoc y x (0 - x)) + ∙∙ cong (y ℤ+_) (+Comm x (0 - x)) + ∙∙ cong (y ℤ+_) (minusPlus x (pos 0)) +-cancelRₖ (suc zero) = + elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + (wedgeconFun _ _ (λ _ _ → wedgeConHLevPath 0 _ _) + (λ x → cong (_+ₖ ∣ base ∣) (rUnitₖ 1 ∣ x ∣) ∙ rUnitₖ 1 ∣ x ∣) + (λ x → rCancelₖ 1 ∣ x ∣) + (rUnit refl)) +-cancelRₖ (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + (wedgeconFun _ _ (λ _ _ → wedgeConHLevPath (suc n) _ _) + (λ x → cong (_+ₖ ∣ north ∣) (rUnitₖ (2 + n) ∣ x ∣) ∙ rUnitₖ (2 + n) ∣ x ∣) + (λ x → rCancelₖ (2 + n) ∣ x ∣) + (sym (rUnit _))) + +-cancelLₖ : (n : ℕ) (x y : coHomK n) → (x +[ n ]ₖ y) -[ n ]ₖ x ≡ y +-cancelLₖ n x y = cong (λ z → z -[ n ]ₖ x) (commₖ n x y) ∙ -cancelRₖ n x y + +-+cancelₖ : (n : ℕ) (x y : coHomK n) → (x -[ n ]ₖ y) +[ n ]ₖ y ≡ x +-+cancelₖ zero x y = sym (+Assoc x (0 - y) y) ∙ cong (x ℤ+_) (minusPlus y (pos 0)) +-+cancelₖ (suc zero) = + elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + (wedgeconFun _ _ (λ _ _ → wedgeConHLevPath 0 _ _) + (λ x → cong (_+ₖ ∣ x ∣) (lUnitₖ 1 (-ₖ ∣ x ∣)) ∙ lCancelₖ 1 ∣ x ∣) + (λ x → cong (_+ₖ ∣ base ∣) (rUnitₖ 1 ∣ x ∣) ∙ rUnitₖ 1 ∣ x ∣) + refl) +-+cancelₖ (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + (wedgeconFun _ _ (λ _ _ → wedgeConHLevPath (suc n) _ _) + (λ x → cong (_+ₖ ∣ x ∣) (lUnitₖ (2 + n) (-ₖ ∣ x ∣)) ∙ lCancelₖ (2 + n) ∣ x ∣) + (λ x → cong (_+ₖ ∣ north ∣) (rUnitₖ (2 + n) ∣ x ∣) ∙ rUnitₖ (2 + n) ∣ x ∣) + refl) + +---- Group structure of cohomology groups +_+ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A +_+ₕ_ {n = n} = sRec2 § λ a b → ∣ (λ x → a x +[ n ]ₖ b x) ∣₂ + +-ₕ_ : {n : ℕ} → coHom n A → coHom n A +-ₕ_ {n = n} = sRec § λ a → ∣ (λ x → -[ n ]ₖ a x) ∣₂ + +_-ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A +_-ₕ_ {n = n} = sRec2 § λ a b → ∣ (λ x → a x -[ n ]ₖ b x) ∣₂ + ++ₕ-syntax : (n : ℕ) → coHom n A → coHom n A → coHom n A ++ₕ-syntax n = _+ₕ_ {n = n} + +-ₕ-syntax : (n : ℕ) → coHom n A → coHom n A +-ₕ-syntax n = -ₕ_ {n = n} + +-ₕ'-syntax : (n : ℕ) → coHom n A → coHom n A → coHom n A +-ₕ'-syntax n = _-ₕ_ {n = n} + +syntax +ₕ-syntax n x y = x +[ n ]ₕ y +syntax -ₕ-syntax n x = -[ n ]ₕ x +syntax -ₕ'-syntax n x y = x -[ n ]ₕ y + +0ₕ : (n : ℕ) → coHom n A +0ₕ n = ∣ (λ _ → (0ₖ n)) ∣₂ + +_+'ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A +_+'ₕ_ {n = n} x y = (x +ₕ 0ₕ _) +ₕ y +ₕ 0ₕ _ + +rUnitₕ : (n : ℕ) (x : coHom n A) → x +[ n ]ₕ (0ₕ n) ≡ x +rUnitₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rUnitₖ n (a x)) i ∣₂ + +lUnitₕ : (n : ℕ) (x : coHom n A) → (0ₕ n) +[ n ]ₕ x ≡ x +lUnitₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lUnitₖ n (a x)) i ∣₂ + +rCancelₕ : (n : ℕ) (x : coHom n A) → x +[ n ]ₕ (-[ n ]ₕ x) ≡ 0ₕ n +rCancelₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rCancelₖ n (a x)) i ∣₂ + +lCancelₕ : (n : ℕ) (x : coHom n A) → (-[ n ]ₕ x) +[ n ]ₕ x ≡ 0ₕ n +lCancelₕ n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lCancelₖ n (a x)) i ∣₂ + +assocₕ : (n : ℕ) (x y z : coHom n A) → (x +[ n ]ₕ (y +[ n ]ₕ z)) ≡ ((x +[ n ]ₕ y) +[ n ]ₕ z) +assocₕ n = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b c i → ∣ funExt (λ x → assocₖ n (a x) (b x) (c x)) i ∣₂ + +commₕ : (n : ℕ) (x y : coHom n A) → (x +[ n ]ₕ y) ≡ (y +[ n ]ₕ x) +commₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ funExt (λ x → commₖ n (a x) (b x)) i ∣₂ + +-cancelLₕ : (n : ℕ) (x y : coHom n A) → (x +[ n ]ₕ y) -[ n ]ₕ x ≡ y +-cancelLₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelLₖ n (a x) (b x) i) ∣₂ + +-cancelRₕ : (n : ℕ) (x y : coHom n A) → (y +[ n ]ₕ x) -[ n ]ₕ x ≡ y +-cancelRₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelRₖ n (a x) (b x) i) ∣₂ + +-+cancelₕ : (n : ℕ) (x y : coHom n A) → (x -[ n ]ₕ y) +[ n ]ₕ y ≡ x +-+cancelₕ n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -+cancelₖ n (a x) (b x) i) ∣₂ + +-- Group structure of reduced cohomology groups (in progress - might need K to compute properly first) +_+ₕ∙_ : {A : Pointed ℓ} {n : ℕ} → coHomRed n A → coHomRed n A → coHomRed n A +_+ₕ∙_ {n = zero} = sRec2 § λ { (a , pa) (b , pb) → ∣ (λ x → a x +[ zero ]ₖ b x) + , (λ i → (pa i +[ zero ]ₖ pb i)) ∣₂ } +_+ₕ∙_ {n = (suc zero)} = sRec2 § λ { (a , pa) (b , pb) → ∣ (λ x → a x +[ 1 ]ₖ b x) + , (λ i → pa i +[ 1 ]ₖ pb i) ∣₂ } +_+ₕ∙_ {n = (suc (suc n))} = + sRec2 § λ { (a , pa) (b , pb) → ∣ (λ x → a x +[ (2 + n) ]ₖ b x) + , (λ i → pa i +[ (2 + n) ]ₖ pb i) ∣₂ } + +-ₕ∙_ : {A : Pointed ℓ} {n : ℕ} → coHomRed n A → coHomRed n A +-ₕ∙_ {n = zero} = sRec § λ {(f , p) → ∣ (λ x → -[ 0 ]ₖ (f x)) + , cong (λ x → -[ 0 ]ₖ x) p ∣₂} +-ₕ∙_ {n = suc zero} = sRec § λ {(f , p) → ∣ (λ x → -ₖ (f x)) + , cong -ₖ_ p ∣₂} +-ₕ∙_ {n = suc (suc n)} = sRec § λ {(f , p) → ∣ (λ x → -ₖ (f x)) + , cong -ₖ_ p ∣₂} + +0ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A +0ₕ∙ n = ∣ (λ _ → 0ₖ n) , refl ∣₂ + ++ₕ∙-syntax : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A → coHomRed n A ++ₕ∙-syntax n = _+ₕ∙_ {n = n} + +-ₕ∙-syntax : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A +-ₕ∙-syntax n = -ₕ∙_ {n = n} + +-'ₕ∙-syntax : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A → coHomRed n A +-'ₕ∙-syntax n x y = _+ₕ∙_ {n = n} x (-ₕ∙_ {n = n} y) + +syntax +ₕ∙-syntax n x y = x +[ n ]ₕ∙ y +syntax -ₕ∙-syntax n x = -[ n ]ₕ∙ x +syntax -'ₕ∙-syntax n x y = x -[ n ]ₕ∙ y + +commₕ∙ : {A : Pointed ℓ} (n : ℕ) (x y : coHomRed n A) → x +[ n ]ₕ∙ y ≡ y +[ n ]ₕ∙ x +commₕ∙ zero = + sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) (g , q) + → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) λ i x → commₖ 0 (f x) (g x) i)} +commₕ∙ (suc zero) = + sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) (g , q) + → cong ∣_∣₂ (ΣPathP ((λ i x → commₖ 1 (f x) (g x) i) + , λ i j → commₖ 1 (p j) (q j) i))} +commₕ∙ {A = A} (suc (suc n)) = + sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) (g , q) + → cong ∣_∣₂ (ΣPathP ((λ i x → commₖ (2 + n) (f x) (g x) i) + , λ i j → hcomp (λ k → λ {(i = i0) → p j +ₖ q j + ; (i = i1) → q j +ₖ p j + ; (j = i0) → commₖ (2 + n) (f (pt A)) (g (pt A)) i + ; (j = i1) → rUnit (refl {x = 0ₖ (2 + n)}) (~ k) i}) + (commₖ (2 + n) (p j) (q j) i)))} + +rUnitₕ∙ : {A : Pointed ℓ} (n : ℕ) (x : coHomRed n A) → x +[ n ]ₕ∙ 0ₕ∙ n ≡ x +rUnitₕ∙ zero = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) λ i x → rUnitₖ zero (f x) i)} +rUnitₕ∙ (suc zero) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP ((λ i x → rUnitₖ 1 (f x) i) , λ i j → rUnitₖ 1 (p j) i))} +rUnitₕ∙ (suc (suc n)) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP ((λ i x → rUnitₖ (2 + n) (f x) i) , λ i j → rUnitₖ (2 + n) (p j) i))} + +lUnitₕ∙ : {A : Pointed ℓ} (n : ℕ) (x : coHomRed n A) → 0ₕ∙ n +[ n ]ₕ∙ x ≡ x +lUnitₕ∙ zero = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) λ i x → lUnitₖ zero (f x) i)} +lUnitₕ∙ (suc zero) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP ((λ i x → lUnitₖ 1 (f x) i) , λ i j → lUnitₖ 1 (p j) i))} +lUnitₕ∙ (suc (suc n)) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP ((λ i x → lUnitₖ (2 + n) (f x) i) , λ i j → lUnitₖ (2 + n) (p j) i))} + + +private + pp : {A : Pointed ℓ} (n : ℕ) → (f : fst A → coHomK (suc (suc n))) + → (p : f (snd A) ≡ snd (coHomK-ptd (suc (suc n)))) + → PathP (λ i → rCancelₖ (2 + n) (f (snd A)) i ≡ 0ₖ (suc (suc n))) + (λ i → (p i) +ₖ (-ₖ p i)) (λ _ → 0ₖ (suc (suc n))) + pp {A = A} n f p i j = + hcomp (λ k → λ {(i = i0) → rCancelₖ (suc (suc n)) (p j) (~ k) + ; (i = i1) → 0ₖ (suc (suc n)) + ; (j = i0) → rCancelₖ (2 + n) (f (snd A)) (i ∨ ~ k) + ; (j = i1) → rUnit (rUnit (λ _ → 0ₖ (suc (suc n))) (~ i)) (~ i) k}) + (0ₖ (suc (suc n))) + +rCancelₕ∙ : {A : Pointed ℓ} (n : ℕ) (x : coHomRed n A) → x +[ n ]ₕ∙ (-[ n ]ₕ∙ x) ≡ 0ₕ∙ n +rCancelₕ∙ zero = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) λ i x → rCancelₖ zero (f x) i)} +rCancelₕ∙ {A = A} (suc zero) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP ((λ i x → rCancelₖ 1 (f x) i) , λ i j → rCancelₖ 1 (p j) i))} +rCancelₕ∙ {A = A} (suc (suc n)) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) + → cong ∣_∣₂ (ΣPathP ((λ i x → rCancelₖ (2 + n) (f x) i) + , pp n f p))} + +lCancelₕ∙ : {A : Pointed ℓ} (n : ℕ) (x : coHomRed n A) → (-[ n ]ₕ∙ x) +[ n ]ₕ∙ x ≡ 0ₕ∙ n +lCancelₕ∙ zero = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) λ i x → lCancelₖ zero (f x) i)} +lCancelₕ∙ {A = A} (suc zero) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) + → cong ∣_∣₂ (ΣPathP ((λ i x → lCancelₖ 1 (f x) i) + , λ i j → (lCancelₖ 1 (p j) i)))} +lCancelₕ∙ {A = A} (suc (suc n)) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) + → cong ∣_∣₂ (ΣPathP ((λ i x → lCancelₖ (2 + n) (f x) i) + , λ i j → lCancelₖ (2 + n) (p j) i))} + +assocₕ∙ : {A : Pointed ℓ} (n : ℕ) (x y z : coHomRed n A) + → (x +[ n ]ₕ∙ (y +[ n ]ₕ∙ z)) ≡ ((x +[ n ]ₕ∙ y) +[ n ]ₕ∙ z) +assocₕ∙ zero = + elim3 (λ _ _ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) (g , q) (h , r) + → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) + (λ i x → assocₖ zero (f x) (g x) (h x) i))} +assocₕ∙ (suc zero) = + elim3 (λ _ _ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) (g , q) (h , r) + → cong ∣_∣₂ (ΣPathP ((λ i x → assocₖ 1 (f x) (g x) (h x) i) + , λ i j → assocₖ 1 (p j) (q j) (r j) i))} +assocₕ∙ (suc (suc n)) = + elim3 (λ _ _ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) (g , q) (h , r) + → cong ∣_∣₂ (ΣPathP ((λ i x → assocₖ (2 + n) (f x) (g x) (h x) i) + , λ i j → assocₖ (2 + n) (p j) (q j) (r j) i))} + +open IsSemigroup +open IsMonoid +open GroupStr +open IsGroupHom + +coHomGr : (n : ℕ) (A : Type ℓ) → Group ℓ +coHomGr n A = coHom n A , coHomGrnA + where + coHomGrnA : GroupStr (coHom n A) + 1g coHomGrnA = 0ₕ n + GroupStr._·_ coHomGrnA = λ x y → x +[ n ]ₕ y + inv coHomGrnA = λ x → -[ n ]ₕ x + isGroup coHomGrnA = helper + where + abstract + helper : IsGroup {G = coHom n A} (0ₕ n) (λ x y → x +[ n ]ₕ y) (λ x → -[ n ]ₕ x) + helper = makeIsGroup § (assocₕ n) (rUnitₕ n) (lUnitₕ n) (rCancelₕ n) (lCancelₕ n) + +×coHomGr : (n : ℕ) (A : Type ℓ) (B : Type ℓ') → Group _ +×coHomGr n A B = DirProd (coHomGr n A) (coHomGr n B) + +coHomGroup : (n : ℕ) (A : Type ℓ) → AbGroup ℓ +fst (coHomGroup n A) = coHom n A +AbGroupStr.0g (snd (coHomGroup n A)) = 0ₕ n +AbGroupStr._+_ (snd (coHomGroup n A)) = _+ₕ_ {n = n} +AbGroupStr.- snd (coHomGroup n A) = -ₕ_ {n = n} +IsAbGroup.isGroup (AbGroupStr.isAbGroup (snd (coHomGroup n A))) = isGroup (snd (coHomGr n A)) +IsAbGroup.comm (AbGroupStr.isAbGroup (snd (coHomGroup n A))) = commₕ n + +-- Reduced cohomology group (direct def) + +coHomRedGroupDir : (n : ℕ) (A : Pointed ℓ) → AbGroup ℓ +fst (coHomRedGroupDir n A) = coHomRed n A +AbGroupStr.0g (snd (coHomRedGroupDir n A)) = 0ₕ∙ n +AbGroupStr._+_ (snd (coHomRedGroupDir n A)) = _+ₕ∙_ {n = n} +AbGroupStr.- snd (coHomRedGroupDir n A) = -ₕ∙_ {n = n} +IsAbGroup.isGroup (AbGroupStr.isAbGroup (snd (coHomRedGroupDir n A))) = helper + where + abstract + helper : IsGroup {G = coHomRed n A} (0ₕ∙ n) (_+ₕ∙_ {n = n}) (-ₕ∙_ {n = n}) + helper = makeIsGroup § (assocₕ∙ n) (rUnitₕ∙ n) (lUnitₕ∙ n) (rCancelₕ∙ n) (lCancelₕ∙ n) +IsAbGroup.comm (AbGroupStr.isAbGroup (snd (coHomRedGroupDir n A))) = commₕ∙ n + +coHomRedGrDir : (n : ℕ) (A : Pointed ℓ) → Group ℓ +coHomRedGrDir n A = AbGroup→Group (coHomRedGroupDir n A) + +-- Induced map +coHomFun : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) (f : A → B) → coHom n B → coHom n A +coHomFun n f = sRec § λ β → ∣ β ∘ f ∣₂ + +coHomMorph : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) (f : A → B) → GroupHom (coHomGr n B) (coHomGr n A) +fst (coHomMorph n f) = coHomFun n f +snd (coHomMorph n f) = makeIsGroupHom (helper n) + where + helper : ℕ → _ + helper zero = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl + helper (suc zero) = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl + helper (suc (suc n)) = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl + +-- Alternative definition of cohomology using ΩKₙ instead. Useful for breaking proofs of group isos +-- up into smaller parts +coHomGrΩ : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → Group ℓ +coHomGrΩ n A = ∥ (A → typ (Ω (coHomK-ptd (suc n)))) ∥₂ , coHomGrnA + where + coHomGrnA : GroupStr ∥ (A → typ (Ω (coHomK-ptd (suc n)))) ∥₂ + 1g coHomGrnA = ∣ (λ _ → refl) ∣₂ + GroupStr._·_ coHomGrnA = sRec2 § λ p q → ∣ (λ x → p x ∙ q x) ∣₂ + inv coHomGrnA = map λ f x → sym (f x) + isGroup coHomGrnA = helper + where + abstract + helper : + IsGroup {G = ∥ (A → typ (Ω (coHomK-ptd (suc n)))) ∥₂} + (∣ (λ _ → refl) ∣₂) (sRec2 § λ p q → ∣ (λ x → p x ∙ q x) ∣₂) (map λ f x → sym (f x)) + helper = makeIsGroup § (elim3 (λ _ _ _ → isOfHLevelPath 2 § _ _) + (λ p q r → cong ∣_∣₂ (funExt λ x → assoc∙ (p x) (q x) (r x)))) + (sElim (λ _ → isOfHLevelPath 2 § _ _) λ p → cong ∣_∣₂ (funExt λ x → sym (rUnit (p x)))) + (sElim (λ _ → isOfHLevelPath 2 § _ _) λ p → cong ∣_∣₂ (funExt λ x → sym (lUnit (p x)))) + (sElim (λ _ → isOfHLevelPath 2 § _ _) λ p → cong ∣_∣₂ (funExt λ x → rCancel (p x))) + (sElim (λ _ → isOfHLevelPath 2 § _ _) λ p → cong ∣_∣₂ (funExt λ x → lCancel (p x))) + +--- the loopspace of Kₙ is commutative regardless of base +addIso : (n : ℕ) (x : coHomK n) → Iso (coHomK n) (coHomK n) +fun (addIso n x) y = y +[ n ]ₖ x +inv' (addIso n x) y = y -[ n ]ₖ x +rightInv (addIso n x) y = -+cancelₖ n y x +leftInv (addIso n x) y = -cancelRₖ n x y + +baseChange : (n : ℕ) (x : coHomK (suc n)) → (0ₖ (suc n) ≡ 0ₖ (suc n)) ≃ (x ≡ x) +baseChange n x = isoToEquiv is + where + f : (n : ℕ) (x : coHomK (suc n)) → (0ₖ (suc n) ≡ 0ₖ (suc n)) → (x ≡ x) + f n x p = sym (rUnitₖ _ x) ∙∙ cong (x +ₖ_) p ∙∙ rUnitₖ _ x + + g : (n : ℕ) (x : coHomK (suc n)) → (x ≡ x) → (0ₖ (suc n) ≡ 0ₖ (suc n)) + g n x p = sym (rCancelₖ _ x) ∙∙ cong (λ y → y -ₖ x) p ∙∙ rCancelₖ _ x + + f-g : (n : ℕ) (x : coHomK (suc n)) (p : x ≡ x) → f n x (g n x p) ≡ p + f-g n = + trElim (λ _ → isOfHLevelΠ (3 + n) λ _ → isOfHLevelPath (3 + n) + (isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) _ _) + (ind n) + where + ind : (n : ℕ) (a : S₊ (suc n)) (p : ∣ a ∣ₕ ≡ ∣ a ∣ₕ) → f n ∣ a ∣ₕ (g n ∣ a ∣ₕ p) ≡ p + ind zero = + toPropElim (λ _ → isPropΠ λ _ → isOfHLevelTrunc 3 _ _ _ _) + λ p → cong (f zero (0ₖ 1)) (sym (rUnit _) ∙ (λ k i → rUnitₖ _ (p i) k)) + ∙∙ sym (rUnit _) + ∙∙ λ k i → lUnitₖ _ (p i) k + ind (suc n) = + sphereElim (suc n) (λ _ → isOfHLevelΠ (2 + n) λ _ → isOfHLevelTrunc (4 + n) _ _ _ _) + λ p → cong (f (suc n) (0ₖ (2 + n))) + ((λ k → (sym (rUnit (refl ∙ refl)) ∙ sym (rUnit refl)) k + ∙∙ (λ i → p i +ₖ 0ₖ (2 + n)) ∙∙ (sym (rUnit (refl ∙ refl)) ∙ sym (rUnit refl)) k) + ∙ (λ k → rUnit (λ i → rUnitₖ _ (p i) k) (~ k))) + ∙ λ k → rUnit (λ i → lUnitₖ _ (p i) k) (~ k) + + g-f : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ 0ₖ _) → g n x (f n x p) ≡ p + g-f n = + trElim (λ _ → isOfHLevelΠ (3 + n) λ _ → isOfHLevelPath (3 + n) + (isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) _ _) + (ind n) + where + ind : (n : ℕ) (a : S₊ (suc n)) (p : 0ₖ (suc n) ≡ 0ₖ (suc n)) → g n ∣ a ∣ₕ (f n ∣ a ∣ₕ p) ≡ p + ind zero = + toPropElim (λ _ → isPropΠ λ _ → isOfHLevelTrunc 3 _ _ _ _) + λ p → cong (g zero (0ₖ 1)) (λ k → rUnit (λ i → lUnitₖ _ (p i) k) (~ k)) + ∙ (λ k → rUnit (λ i → rUnitₖ _ (p i) k) (~ k)) + ind (suc n) = + sphereElim (suc n) (λ _ → isOfHLevelΠ (2 + n) λ _ → isOfHLevelTrunc (4 + n) _ _ _ _) + λ p → cong (g (suc n) (0ₖ (2 + n))) + (λ k → rUnit (λ i → lUnitₖ _ (p i) k) (~ k)) + ∙∙ (λ k → (sym (rUnit (refl ∙ refl)) ∙ sym (rUnit refl)) k + ∙∙ (λ i → p i +ₖ 0ₖ (2 + n)) + ∙∙ (sym (rUnit (refl ∙ refl)) ∙ sym (rUnit refl)) k) + ∙∙ λ k → rUnit (λ i → rUnitₖ _ (p i) k) (~ k) + + is : Iso _ _ + fun is = f n x + inv' is = g n x + rightInv is = f-g n x + leftInv is = g-f n x + +isCommΩK-based : (n : ℕ) (x : coHomK n) → isComm∙ (coHomK n , x) +isCommΩK-based zero x p q = isSetℤ _ _ (p ∙ q) (q ∙ p) +isCommΩK-based (suc zero) x = + subst isComm∙ (λ i → coHomK 1 , lUnitₖ 1 x i) + (ptdIso→comm {A = (_ , 0ₖ 1)} (addIso 1 x) + (isCommΩK 1)) +isCommΩK-based (suc (suc n)) x = + subst isComm∙ (λ i → coHomK (suc (suc n)) , lUnitₖ (suc (suc n)) x i) + (ptdIso→comm {A = (_ , 0ₖ (suc (suc n)))} (addIso (suc (suc n)) x) + (isCommΩK (suc (suc n)))) + +-- hidden versions of cohom stuff using the "lock" hack. The locked versions can be used when proving things. +-- Swapping "key" for "tt*" will then give computing functions. +Unit' : Type₀ +Unit' = lockUnit {ℓ-zero} + +lock : ∀ {ℓ} {A : Type ℓ} → Unit' → A → A +lock unlock = λ x → x + +module lockedCohom (key : Unit') where + +K : (n : ℕ) → coHomK n → coHomK n → coHomK n + +K n = lock key (_+ₖ_ {n = n}) + + -K : (n : ℕ) → coHomK n → coHomK n + -K n = lock key (-ₖ_ {n = n}) + + -Kbin : (n : ℕ) → coHomK n → coHomK n → coHomK n + -Kbin n x y = +K n x (-K n y) + + rUnitK : (n : ℕ) (x : coHomK n) → +K n x (0ₖ n) ≡ x + rUnitK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) x (0ₖ n) ≡ x + pm unlock = rUnitₖ n x + + lUnitK : (n : ℕ) (x : coHomK n) → +K n (0ₖ n) x ≡ x + lUnitK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (0ₖ n) x ≡ x + pm unlock = lUnitₖ n x + + rCancelK : (n : ℕ) (x : coHomK n) → +K n x (-K n x) ≡ 0ₖ n + rCancelK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) x (lock t (-ₖ_ {n = n}) x) ≡ 0ₖ n + pm unlock = rCancelₖ n x + + lCancelK : (n : ℕ) (x : coHomK n) → +K n (-K n x) x ≡ 0ₖ n + lCancelK n x = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (lock t (-ₖ_ {n = n}) x) x ≡ 0ₖ n + pm unlock = lCancelₖ n x + + -cancelRK : (n : ℕ) (x y : coHomK n) → -Kbin n (+K n y x) x ≡ y + -cancelRK n x y = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (lock t (_+ₖ_ {n = n}) y x) (lock t (-ₖ_ {n = n}) x) ≡ y + pm unlock = -cancelRₖ n x y + + -cancelLK : (n : ℕ) (x y : coHomK n) → -Kbin n (+K n x y) x ≡ y + -cancelLK n x y = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (lock t (_+ₖ_ {n = n}) x y) (lock t (-ₖ_ {n = n}) x) ≡ y + pm unlock = -cancelLₖ n x y + + -+cancelK : (n : ℕ) (x y : coHomK n) → +K n (-Kbin n x y) y ≡ x + -+cancelK n x y = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) (lock t (_+ₖ_ {n = n}) x (lock t (-ₖ_ {n = n}) y)) y ≡ x + pm unlock = -+cancelₖ n x y + + assocK : (n : ℕ) (x y z : coHomK n) → +K n x (+K n y z) ≡ +K n (+K n x y) z + assocK n x y z = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) x (lock t (_+ₖ_ {n = n}) y z) + ≡ lock t (_+ₖ_ {n = n}) (lock t (_+ₖ_ {n = n}) x y) z + pm unlock = assocₖ n x y z + + commK : (n : ℕ) (x y : coHomK n) → +K n x y ≡ +K n y x + commK n x y = pm key + where + pm : (t : Unit') → lock t (_+ₖ_ {n = n}) x y ≡ lock t (_+ₖ_ {n = n}) y x + pm unlock = commₖ n x y + + -- cohom + + +H : (n : ℕ) (x y : coHom n A) → coHom n A + +H n = sRec2 § λ a b → ∣ (λ x → +K n (a x) (b x)) ∣₂ + + -H : (n : ℕ) (x : coHom n A) → coHom n A + -H n = sRec § λ a → ∣ (λ x → -K n (a x)) ∣₂ + + -Hbin : (n : ℕ) → coHom n A → coHom n A → coHom n A + -Hbin n = sRec2 § λ a b → ∣ (λ x → -Kbin n (a x) (b x)) ∣₂ + + rUnitH : (n : ℕ) (x : coHom n A) → +H n x (0ₕ n) ≡ x + rUnitH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rUnitK n (a x)) i ∣₂ + + lUnitH : (n : ℕ) (x : coHom n A) → +H n (0ₕ n) x ≡ x + lUnitH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lUnitK n (a x)) i ∣₂ + + rCancelH : (n : ℕ) (x : coHom n A) → +H n x (-H n x) ≡ 0ₕ n + rCancelH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rCancelK n (a x)) i ∣₂ + + lCancelH : (n : ℕ) (x : coHom n A) → +H n (-H n x) x ≡ 0ₕ n + lCancelH n = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lCancelK n (a x)) i ∣₂ + + assocH : (n : ℕ) (x y z : coHom n A) → (+H n x (+H n y z)) ≡ (+H n (+H n x y) z) + assocH n = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b c i → ∣ funExt (λ x → assocK n (a x) (b x) (c x)) i ∣₂ + + commH : (n : ℕ) (x y : coHom n A) → (+H n x y) ≡ (+H n y x) + commH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ funExt (λ x → commK n (a x) (b x)) i ∣₂ + + -cancelRH : (n : ℕ) (x y : coHom n A) → -Hbin n (+H n y x) x ≡ y + -cancelRH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelRK n (a x) (b x) i) ∣₂ + + -cancelLH : (n : ℕ) (x y : coHom n A) → -Hbin n (+H n x y) x ≡ y + -cancelLH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -cancelLK n (a x) (b x) i) ∣₂ + + -+cancelH : (n : ℕ) (x y : coHom n A) → +H n (-Hbin n x y) y ≡ x + -+cancelH n = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ (λ x → -+cancelK n (a x) (b x) i) ∣₂ + +lUnitK≡rUnitK : (key : Unit') (n : ℕ) → lockedCohom.lUnitK key n (0ₖ n) ≡ lockedCohom.rUnitK key n (0ₖ n) +lUnitK≡rUnitK unlock = lUnitₖ≡rUnitₖ + +open GroupStr renaming (_·_ to _+gr_) +open IsGroupHom + +-- inducedCoHom : ∀ {ℓ ℓ'} {A : Type ℓ} {G : Group {ℓ'}} {n : ℕ} +-- → GroupIso (coHomGr n A) G +-- → Group +-- inducedCoHom {A = A} {G = G} {n = n} e = +-- InducedGroup (coHomGr n A) +-- (coHom n A , λ x y → Iso.inv (isom e) (_+gr_ (snd G) (fun (isom e) x) +-- (fun (isom e) y))) +-- (idEquiv _) +-- λ x y → sym (leftInv (isom e) _) +-- ∙ cong (Iso.inv (isom e)) (isHom e x y) + +-- induced+ : ∀ {ℓ ℓ'} {A : Type ℓ} {G : Group {ℓ'}} {n : ℕ} +-- → (e : GroupIso (coHomGr n A) G) +-- → fst (inducedCoHom e) → fst (inducedCoHom e) → fst (inducedCoHom e) +-- induced+ e = _+gr_ (snd (inducedCoHom e)) + +-- inducedCoHomIso : ∀ {ℓ ℓ'} {A : Type ℓ} {G : Group {ℓ'}} {n : ℕ} +-- → (e : GroupIso (coHomGr n A) G) +-- → GroupIso (coHomGr n A) (inducedCoHom e) +-- isom (inducedCoHomIso e) = idIso +-- isHom (inducedCoHomIso e) x y = sym (leftInv (isom e) _) +-- ∙ cong (Iso.inv (isom e)) (isHom e x y) + +-- inducedCoHomPath : ∀ {ℓ ℓ'} {A : Type ℓ} {G : Group {ℓ'}} {n : ℕ} +-- → (e : GroupIso (coHomGr n A) G) +-- → coHomGr n A ≡ inducedCoHom e +-- inducedCoHomPath e = InducedGroupPath _ _ _ _ diff --git a/Cubical/ZCohomology/Groups/CP2.agda b/Cubical/ZCohomology/Groups/CP2.agda new file mode 100644 index 000000000..d6e7f7126 --- /dev/null +++ b/Cubical/ZCohomology/Groups/CP2.agda @@ -0,0 +1,125 @@ +{-# OPTIONS --safe #-} +module Cubical.ZCohomology.Groups.CP2 where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.RingStructure.CupProduct + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Sigma +open import Cubical.Data.Int +open import Cubical.Data.Nat +open import Cubical.Data.Unit +open import Cubical.Algebra.Group + renaming (ℤ to ℤGroup) hiding (Unit) + +open import Cubical.HITs.Pushout +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Hopf +open import Cubical.HITs.Join +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; elim2 to pElim2 ; ∣_∣ to ∣_∣₁) + +open IsGroupHom +open Iso + +CP² : Type +CP² = Pushout {A = TotalHopf} fst λ _ → tt + +module M = MV (S₊ 2) Unit TotalHopf fst (λ _ → tt) + +H²CP²≅ℤ : GroupIso (coHomGr 2 CP²) ℤGroup +H²CP²≅ℤ = compGroupIso (BijectionIso→GroupIso bij) + (compGroupIso (invGroupIso trivIso) (Hⁿ-Sⁿ≅ℤ 1)) + where + isContrH¹TotalHopf : isContr (coHom 1 TotalHopf) + isContrH¹TotalHopf = + isOfHLevelRetractFromIso 0 (setTruncIso (domIso (invIso (IsoS³TotalHopf)))) + (isOfHLevelRetractFromIso 0 ((fst (H¹-Sⁿ≅0 1))) isContrUnit) + + isContrH²TotalHopf : isContr (coHom 2 TotalHopf) + isContrH²TotalHopf = + isOfHLevelRetractFromIso 0 (setTruncIso (domIso (invIso (IsoS³TotalHopf)))) + ((isOfHLevelRetractFromIso 0 + (fst (Hⁿ-Sᵐ≅0 1 2 λ p → snotz (sym (cong predℕ p)))) isContrUnit)) + + trivIso : GroupIso (coHomGr 2 (Susp S¹)) (×coHomGr 2 (Susp S¹) Unit) + fun (fst trivIso) x = x , 0ₕ _ + inv (fst trivIso) = fst + rightInv (fst trivIso) (x , p) = + ΣPathP (refl , isContr→isProp (isContrHⁿ-Unit 1) _ _) + leftInv (fst trivIso) x = refl + snd trivIso = makeIsGroupHom λ _ _ → refl + + bij : BijectionIso (coHomGr 2 CP²) (×coHomGr 2 (Susp S¹) Unit) + BijectionIso.fun bij = M.i 2 + BijectionIso.inj bij x p = + pRec (squash₂ _ _) + (uncurry (λ z q + → sym q + ∙∙ cong (fst (M.d 1)) (isContr→isProp isContrH¹TotalHopf z (0ₕ _)) + ∙∙ (M.d 1) .snd .pres1)) + (M.Ker-i⊂Im-d 1 x p) + where + help : isInIm (M.d 1) x + help = M.Ker-i⊂Im-d 1 x p + BijectionIso.surj bij y = + M.Ker-Δ⊂Im-i 2 y (isContr→isProp isContrH²TotalHopf _ _) + +H⁴CP²≅ℤ : GroupIso (coHomGr 4 CP²) ℤGroup +H⁴CP²≅ℤ = compGroupIso (invGroupIso (BijectionIso→GroupIso bij)) + (compGroupIso help (Hⁿ-Sⁿ≅ℤ 2)) + where + help : GroupIso (coHomGr 3 TotalHopf) (coHomGr 3 (S₊ 3)) + help = isoType→isoCohom 3 (invIso IsoS³TotalHopf) + + bij : BijectionIso (coHomGr 3 TotalHopf) (coHomGr 4 CP²) + BijectionIso.fun bij = M.d 3 + BijectionIso.inj bij x p = + pRec (squash₂ _ _) + (uncurry (λ z q → + sym q + ∙∙ cong (M.Δ 3 .fst) + (isOfHLevelΣ 1 (isContr→isProp + (isOfHLevelRetractFromIso 0 + (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p))) isContrUnit)) + (λ _ → isContr→isProp (isContrHⁿ-Unit 2)) + z (0ₕ _ , 0ₕ _)) + ∙∙ M.Δ 3 .snd .pres1)) + (M.Ker-d⊂Im-Δ _ x p) + BijectionIso.surj bij y = + M.Ker-i⊂Im-d 3 y (isOfHLevelΣ 1 + (isContr→isProp (isOfHLevelRetractFromIso 0 + (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p))) isContrUnit)) + (λ _ → isContr→isProp (isContrHⁿ-Unit _)) _ _) + + +-- Another Brunerie number +ℤ→HⁿCP²→ℤ : ℤ → ℤ +ℤ→HⁿCP²→ℤ x = + Iso.fun (fst H⁴CP²≅ℤ) + (Iso.inv (fst H²CP²≅ℤ) x ⌣ Iso.inv (fst H²CP²≅ℤ) x) + +brunerie2 : ℤ +brunerie2 = ℤ→HⁿCP²→ℤ 1 + +{- +|brunerie2|≡1 : abs (ℤ→HⁿCP²→ℤ 1) ≡ 1 +|brunerie2|≡1 = refl +-} diff --git a/Cubical/ZCohomology/Groups/Connected.agda b/Cubical/ZCohomology/Groups/Connected.agda new file mode 100644 index 000000000..b134d252b --- /dev/null +++ b/Cubical/ZCohomology/Groups/Connected.agda @@ -0,0 +1,45 @@ +{-# OPTIONS --safe #-} +module Cubical.ZCohomology.Groups.Connected where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Unit + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; ∥_∥ to ∥_∥₁ ; ∣_∣ to ∣_∣₁) +open import Cubical.HITs.Nullification + +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.Data.Int hiding (ℤ) renaming (_+_ to _+ℤ_; +Comm to +ℤ-comm ; +Assoc to +ℤ-assoc) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation renaming (rec₊ to trRec) +open import Cubical.Algebra.Group + +open import Cubical.Homotopy.Connected +open import Cubical.Foundations.Equiv + +private + H⁰-connected-type : ∀ {ℓ} {A : Type ℓ} (a : A) → isConnected 2 A → Iso (coHom 0 A) (fst ℤ) + Iso.fun (H⁰-connected-type a con) = sRec isSetℤ λ f → f a + Iso.inv (H⁰-connected-type a con) b = ∣ (λ x → b) ∣₂ + Iso.rightInv (H⁰-connected-type a con) b = refl + Iso.leftInv (H⁰-connected-type a con) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f → cong ∣_∣₂ (funExt λ x → trRec (isSetℤ _ _) (cong f) (isConnectedPath 1 con a x .fst)) + +open IsGroupHom +open Iso + +H⁰-connected : ∀ {ℓ} {A : Type ℓ} (a : A) → ((x : A) → ∥ a ≡ x ∥₁) → GroupIso (coHomGr 0 A) ℤ +fun (fst (H⁰-connected a con)) = sRec isSetℤ (λ f → f a) +inv (fst (H⁰-connected a con)) b = ∣ (λ _ → b) ∣₂ +rightInv (fst (H⁰-connected a con)) _ = refl +leftInv (fst (H⁰-connected a con)) = + sElim (λ _ → isProp→isSet (isSetSetTrunc _ _)) + (λ f → cong ∣_∣₂ (funExt λ x → pRec (isSetℤ _ _) (cong f) (con x))) +snd (H⁰-connected a con) = makeIsGroupHom (sElim2 (λ _ _ → isProp→isSet (isSetℤ _ _)) λ x y → refl) diff --git a/Cubical/ZCohomology/Groups/KleinBottle.agda b/Cubical/ZCohomology/Groups/KleinBottle.agda new file mode 100644 index 000000000..15945515c --- /dev/null +++ b/Cubical/ZCohomology/Groups/KleinBottle.agda @@ -0,0 +1,454 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Groups.KleinBottle where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to pRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; ∣_∣ to ∣_∣₁) +open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; elim2 to trElim2) +open import Cubical.Data.Nat hiding (isEven) +open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Bool to BoolGroup ; Unit to UnitGroup) + +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Transport + +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn + +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Isomorphism +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.Foundations.Equiv +open import Cubical.Homotopy.Connected + +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Bool +open import Cubical.Data.Int renaming (+Comm to +-commℤ ; _+_ to _+ℤ_) + +open import Cubical.HITs.KleinBottle +open import Cubical.Data.Empty +open import Cubical.Foundations.Path + +open import Cubical.Homotopy.Loopspace + +open IsGroupHom +open Iso + +characFunSpace𝕂² : ∀ {ℓ} (A : Type ℓ) → + Iso (KleinBottle → A) + (Σ[ x ∈ A ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙∙ q ∙∙ p ≡ q) +fun (characFunSpace𝕂² A) f = + (f point) , + ((cong f line1) , + (cong f line2 , + fst (Square≃doubleComp + (cong f line2) (cong f line2) + (sym (cong f line1)) (cong f line1)) + (λ i j → f (square i j)))) +inv (characFunSpace𝕂² A) (x , p , q , sq) point = x +inv (characFunSpace𝕂² A) (x , p , q , sq) (line1 i) = p i +inv (characFunSpace𝕂² A) (x , p , q , sq) (line2 i) = q i +inv (characFunSpace𝕂² A) (x , p , q , sq) (square i j) = + invEq (Square≃doubleComp q q (sym p) p) sq i j +rightInv (characFunSpace𝕂² A) (x , (p , (q , sq))) = + ΣPathP (refl , (ΣPathP (refl , (ΣPathP (refl , secEq (Square≃doubleComp q q (sym p) p) sq))))) +leftInv (characFunSpace𝕂² A) f _ point = f point +leftInv (characFunSpace𝕂² A) f _ (line1 i) = f (line1 i) +leftInv (characFunSpace𝕂² A) f _ (line2 i) = f (line2 i) +leftInv (characFunSpace𝕂² A) f z (square i j) = + retEq (Square≃doubleComp + (cong f line2) (cong f line2) + (sym (cong f line1)) (cong f line1)) + (λ i j → f (square i j)) z i j +private + movePathLem : ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) → isComm∙ (A , x) + → (p ∙∙ q ∙∙ p ≡ q) ≡ ((p ∙ p) ∙ q ≡ q) + movePathLem p q comm = + cong (_≡ q) (doubleCompPath-elim' p q p ∙∙ cong (p ∙_) (comm q p) ∙∙ assoc _ _ _) + + movePathLem2 : ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) + → (((p ∙ p) ∙ q) ∙ sym q ≡ q ∙ sym q) ≡ (p ∙ p ≡ refl) + movePathLem2 p q = + cong₂ _≡_ (sym (assoc (p ∙ p) q (sym q)) ∙∙ cong ((p ∙ p) ∙_) (rCancel q) ∙∙ sym (rUnit (p ∙ p))) + (rCancel q) + + movePathIso : ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) → isComm∙ (A , x) + → Iso (p ∙∙ q ∙∙ p ≡ q) (p ∙ p ≡ refl) + movePathIso {x = x} p q comm = + compIso (pathToIso (movePathLem p q comm)) + (compIso (helper (p ∙ p)) + (pathToIso (movePathLem2 p q))) + where + helper : (p : x ≡ x) → Iso (p ∙ q ≡ q) ((p ∙ q) ∙ sym q ≡ q ∙ sym q) + helper p = congIso (equivToIso (_ , compPathr-isEquiv (sym q))) + +------ H¹(𝕂²) ≅ 0 -------------- +H⁰-𝕂² : GroupIso (coHomGr 0 KleinBottle) ℤGroup +fun (fst H⁰-𝕂²) = sRec isSetℤ λ f → f point +inv (fst H⁰-𝕂²) x = ∣ (λ _ → x) ∣₂ +rightInv (fst H⁰-𝕂²) _ = refl +leftInv (fst H⁰-𝕂²) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f → cong ∣_∣₂ (funExt (λ {point → refl + ; (line1 i) j → isSetℤ (f point) (f point) refl (cong f line1) j i + ; (line2 i) j → isSetℤ (f point) (f point) refl (cong f line2) j i + ; (square i j) z → helper f i j z})) + where + helper : (f : KleinBottle → ℤ) + → Cube (λ j z → isSetℤ (f point) (f point) refl (cong f line2) z j) + (λ j z → isSetℤ (f point) (f point) refl (cong f line2) z j) + (λ i z → isSetℤ (f point) (f point) refl (cong f line1) z (~ i)) + (λ i z → isSetℤ (f point) (f point) refl (cong f line1) z i) + refl + λ i j → f (square i j) + helper f = isGroupoid→isGroupoid' (isOfHLevelSuc 2 isSetℤ) _ _ _ _ _ _ +snd H⁰-𝕂² = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 isSetℤ _ _) λ _ _ → refl) + +------ H¹(𝕂¹) ≅ ℤ ------------ +{- +Step one : +H¹(𝕂²) := ∥ 𝕂² → K₁ ∥₂ + ≡ ∥ Σ[ x ∈ K₁ ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] (p ∙∙ q ∙∙ p ≡ q) ∥₂ (characFunSpace𝕂²) + ≡ ∥ Σ[ x ∈ K₁ ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙ p ≡ refl ∥₂ (movePathIso, using commutativity of ΩK₂) + ≡ ∥ Σ[ x ∈ K₁ ] (x ≡ x) ∥₂ (p ∙ p ≡ refl forces p ≡ refl. Also, p ∙ p ≡ refl is an hProp) +-} + +nilpotent→≡0 : (x : ℤ) → x +ℤ x ≡ 0 → x ≡ 0 +nilpotent→≡0 (pos zero) p = refl +nilpotent→≡0 (pos (suc n)) p = + ⊥-rec (negsucNotpos _ _ + (sym (cong (_- 1) (cong sucℤ (sym (helper2 n)) ∙ p)))) + where + helper2 : (n : ℕ) → pos (suc n) +pos n ≡ pos (suc (n + n)) + helper2 zero = refl + helper2 (suc n) = cong sucℤ (sym (sucℤ+pos n (pos (suc n)))) + ∙∙ cong (sucℤ ∘ sucℤ) (helper2 n) + ∙∙ cong (pos ∘ suc ∘ suc) (sym (+-suc n n)) +nilpotent→≡0 (negsuc n) p = ⊥-rec (negsucNotpos _ _ (helper2 n p)) + where + helper2 : (n : ℕ) → (negsuc n +negsuc n) ≡ pos 0 → negsuc n ≡ pos (suc n) + helper2 n p = cong (negsuc n +ℤ_) (sym (helper3 n)) + ∙ +Assoc (negsuc n) (negsuc n) (pos (suc n)) + ∙∙ cong (_+ℤ (pos (suc n))) p + ∙∙ cong sucℤ (+-commℤ (pos 0) (pos n)) + where + helper3 : (n : ℕ) → negsuc n +pos (suc n) ≡ 0 + helper3 zero = refl + helper3 (suc n) = cong sucℤ (sucℤ+pos n (negsuc (suc n))) ∙ helper3 n + +nilpotent→≡refl : (x : coHomK 1) (p : x ≡ x) → p ∙ p ≡ refl → p ≡ refl +nilpotent→≡refl = + trElim (λ _ → isGroupoidΠ2 λ _ _ → isOfHLevelPlus {n = 1} 2 (isOfHLevelTrunc 3 _ _ _ _)) + (toPropElim (λ _ → isPropΠ2 λ _ _ → isOfHLevelTrunc 3 _ _ _ _) + λ p pId → sym (rightInv (Iso-Kn-ΩKn+1 0) p) + ∙∙ cong (Kn→ΩKn+1 0) (nilpotent→≡0 (ΩKn+1→Kn 0 p) + (sym (ΩKn+1→Kn-hom 0 p p) + ∙ cong (ΩKn+1→Kn 0) pId)) + ∙∙ Kn→ΩKn+10ₖ 0) + +Iso-H¹-𝕂²₁ : Iso (Σ[ x ∈ coHomK 1 ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙ p ≡ refl) + (Σ[ x ∈ coHomK 1 ] x ≡ x) +fun Iso-H¹-𝕂²₁ (x , (_ , (q , _))) = x , q +inv Iso-H¹-𝕂²₁ (x , q) = x , (refl , (q , (sym (rUnit refl)))) +rightInv Iso-H¹-𝕂²₁ _ = refl +leftInv Iso-H¹-𝕂²₁ (x , (p , (q , P))) = + ΣPathP (refl , + (ΣPathP (sym (nilpotent→≡refl x p P) + , toPathP (Σ≡Prop (λ _ → isOfHLevelTrunc 3 _ _ _ _) + (transportRefl q))))) + +{- But this is precisely the type (minus set-truncation) of H¹(S¹) -} +Iso-H¹-𝕂²₂ : Iso (Σ[ x ∈ coHomK 1 ] x ≡ x) (S¹ → coHomK 1) +Iso-H¹-𝕂²₂ = invIso IsoFunSpaceS¹ + +H¹-𝕂²≅ℤ : GroupIso (coHomGr 1 KleinBottle) ℤGroup +H¹-𝕂²≅ℤ = compGroupIso theGroupIso (Hⁿ-Sⁿ≅ℤ 0) + where + theIso : Iso (coHom 1 KleinBottle) (coHom 1 S¹) + theIso = + setTruncIso ( + compIso (characFunSpace𝕂² (coHomK 1)) + (compIso + (Σ-cong-iso-snd (λ x → Σ-cong-iso-snd + λ p → Σ-cong-iso-snd + λ q → movePathIso p q (isCommΩK-based 1 x))) + (compIso Iso-H¹-𝕂²₁ + Iso-H¹-𝕂²₂))) + + is-hom : IsGroupHom (coHomGr 1 KleinBottle .snd) (fun theIso) (coHomGr 1 S¹ .snd) + is-hom = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f g → cong ∣_∣₂ (funExt λ {base → refl ; (loop i) → refl})) + + theGroupIso : GroupIso (coHomGr 1 KleinBottle) (coHomGr 1 S¹) + theGroupIso = (theIso , is-hom) + +------ H²(𝕂²) ≅ ℤ/2ℤ (represented here by BoolGroup) ------- +-- It suffices to show that H²(Klein) is equivalent to Bool as types + +{- +Step one : +H²(𝕂²) := ∥ 𝕂² → K₂ ∥₂ + ≡ ∥ Σ[ x ∈ K₂ ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] (p ∙∙ q ∙∙ p ≡ q) ∥₂ (characFunSpace𝕂²) + ≡ ∥ Σ[ x ∈ K₂ ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙ p ≡ refl ∥₂ (movePathIso, using commutativity of ΩK₂) + ≡ ∥ Σ[ p ∈ x ≡ x ] p ∙ p ≡ refl ∥₂ (connectedness of K₂) +-} + + +Iso-H²-𝕂²₁ : Iso ∥ Σ[ x ∈ coHomK 2 ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙ p ≡ refl ∥₂ + ∥ Σ[ p ∈ 0ₖ 2 ≡ 0ₖ 2 ] p ∙ p ≡ refl ∥₂ +fun Iso-H²-𝕂²₁ = + sRec isSetSetTrunc + (uncurry (trElim (λ _ → is2GroupoidΠ λ _ → isOfHLevelPlus {n = 2} 2 isSetSetTrunc) + (sphereElim _ (λ _ → isSetΠ λ _ → isSetSetTrunc) + λ y → ∣ fst y , snd (snd y) ∣₂))) +inv Iso-H²-𝕂²₁ = + sMap λ p → (0ₖ 2) , ((fst p) , (refl , (snd p))) +rightInv Iso-H²-𝕂²₁ = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ p → refl +leftInv Iso-H²-𝕂²₁ = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry (trElim (λ _ → is2GroupoidΠ λ _ → isOfHLevelPlus {n = 1} 3 (isSetSetTrunc _ _)) + (sphereToPropElim _ + (λ _ → isPropΠ λ _ → isSetSetTrunc _ _) + λ {(p , (q , sq)) + → trRec (isSetSetTrunc _ _) + (λ qid → cong ∣_∣₂ (ΣPathP (refl , (ΣPathP (refl , (ΣPathP (sym qid , refl))))))) + (fun (PathIdTruncIso _) + (isContr→isProp (isConnectedPathKn 1 (0ₖ 2) (0ₖ 2)) ∣ q ∣ ∣ refl ∣))}))) + +{- Step two : ∥ Σ[ p ∈ x ≡ x ] p ∙ p ≡ refl ∥₂ ≡ ∥ Σ[ x ∈ K₁ ] x + x ≡ 0 ∥₂ -} +Iso-H²-𝕂²₂ : Iso ∥ (Σ[ p ∈ 0ₖ 2 ≡ 0ₖ 2 ] p ∙ p ≡ refl) ∥₂ ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ +Iso-H²-𝕂²₂ = setTruncIso (Σ-cong-iso {B' = λ x → x +ₖ x ≡ 0ₖ 1} (invIso (Iso-Kn-ΩKn+1 1)) + λ p → compIso (congIso (invIso (Iso-Kn-ΩKn+1 1))) + (pathToIso λ i → ΩKn+1→Kn-hom 1 p p i ≡ 0ₖ 1)) + +{- Step three : +∥ Σ[ x ∈ K₁ ] x + x ≡ 0 ∥₂ ≡ Bool +We begin by defining the a map Σ[ x ∈ K₁ ] x + x ≡ 0 → Bool. For a point +(0 , p) we map it to true if winding(p) is even and false if winding(p) is odd. +We also have to show that this map respects the loop +-} + +ΣKₙNilpot→Bool : Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 → Bool +ΣKₙNilpot→Bool = uncurry (trElim (λ _ → isGroupoidΠ λ _ → isOfHLevelSuc 2 isSetBool) + λ {base p → isEven (ΩKn+1→Kn 0 p) + ; (loop i) p → hcomp (λ k → λ { (i = i0) → respectsLoop p k + ; (i = i1) → isEven (ΩKn+1→Kn 0 p)}) + (isEven (ΩKn+1→Kn 0 (transp (λ j → ∣ (loop ∙ loop) (i ∨ j) ∣ ≡ 0ₖ 1) i + p)))}) + where + isEven-2 : (x : ℤ) → isEven (-2 +ℤ x) ≡ isEven x + isEven-2 (pos zero) = refl + isEven-2 (pos (suc zero)) = refl + isEven-2 (pos (suc (suc n))) = + cong isEven (cong sucℤ (sucℤ+pos _ _) + ∙∙ sucℤ+pos _ _ + ∙∙ +-commℤ 0 (pos n)) + ∙ lossy n + where + lossy : (n : ℕ) → isEven (pos n) ≡ isEven (pos n) + lossy n = refl + isEven-2 (negsuc zero) = refl + isEven-2 (negsuc (suc n)) = + cong isEven (predℤ+negsuc n _ + ∙ +-commℤ -3 (negsuc n)) + ∙ lossy2 n + where + lossy2 : (n : ℕ) → isEven (negsuc (suc (suc (suc n)))) ≡ isEven (pos n) + lossy2 n = refl + respectsLoop : (p : 0ₖ 1 ≡ 0ₖ 1) + → isEven (ΩKn+1→Kn 0 (transport (λ i → ∣ (loop ∙ loop) i ∣ ≡ 0ₖ 1) p)) + ≡ isEven (ΩKn+1→Kn 0 p) + respectsLoop p = + cong isEven (cong (ΩKn+1→Kn 0) (cong (transport (λ i → ∣ (loop ∙ loop) i ∣ ≡ 0ₖ 1)) + (lUnit p))) + ∙∙ cong isEven (cong (ΩKn+1→Kn 0) + λ j → transp (λ i → ∣ (loop ∙ loop) (i ∨ j) ∣ ≡ 0ₖ 1) j + ((λ i → ∣ (loop ∙ loop) (~ i ∧ j) ∣) ∙ p)) + ∙∙ cong isEven (ΩKn+1→Kn-hom 0 (sym (cong ∣_∣ (loop ∙ loop))) p) + ∙ isEven-2 (ΩKn+1→Kn 0 p) + +{- +We show that for any x : ℤ we have ∣ (0ₖ 1 , Kn→ΩKn+1 0 x) ∣₂ ≡ ∣ (0ₖ 1 , refl) ∣₂ when x is even +and ∣ (0ₖ 1 , Kn→ΩKn+1 0 x) ∣₂ ≡ ∣ (0ₖ 1 , cong ∣_∣ loop) ∣₂ when x is odd + +This is done by induction on x. For the inductive step we define a multiplication _*_ on ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ +which is just ∣ (0 , p) ∣₂ * ∣ (0 , q) ∣₂ ≡ ∣ (0 , p ∙ q) ∣₂ when x is 0 +-} + +private + _*_ : ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ → ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ → ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ + _*_ = sRec (isSetΠ (λ _ → isSetSetTrunc)) λ a → sRec isSetSetTrunc λ b → *' (fst a) (fst b) (snd a) (snd b) + where + *' : (x y : coHomK 1) (p : x +ₖ x ≡ 0ₖ 1) (q : y +ₖ y ≡ 0ₖ 1) → ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ + *' = + trElim2 (λ _ _ → isGroupoidΠ2 λ _ _ → isOfHLevelSuc 2 isSetSetTrunc) + (wedgeconFun _ _ + (λ _ _ → isSetΠ2 λ _ _ → isSetSetTrunc) + (λ x p q → ∣ ∣ x ∣ , cong₂ _+ₖ_ p q ∣₂) + (λ y p q → ∣ ∣ y ∣ , sym (rUnitₖ 1 (∣ y ∣ +ₖ ∣ y ∣)) ∙ cong₂ _+ₖ_ p q ∣₂) + (funExt λ p → funExt λ q → cong ∣_∣₂ (ΣPathP (refl , (sym (lUnit _)))))) + + *=∙ : (p q : 0ₖ 1 ≡ 0ₖ 1) → ∣ 0ₖ 1 , p ∣₂ * ∣ 0ₖ 1 , q ∣₂ ≡ ∣ 0ₖ 1 , p ∙ q ∣₂ + *=∙ p q = cong ∣_∣₂ (ΣPathP (refl , sym (∙≡+₁ p q))) + +isEvenNegsuc : (n : ℕ) → isEven (pos (suc n)) ≡ true → isEven (negsuc n) ≡ true +isEvenNegsuc zero p = ⊥-rec (true≢false (sym p)) +isEvenNegsuc (suc n) p = p + +¬isEvenNegSuc : (n : ℕ) → isEven (pos (suc n)) ≡ false → isEven (negsuc n) ≡ false +¬isEvenNegSuc zero p = refl +¬isEvenNegSuc (suc n) p = p + +evenCharac : (x : ℤ) → isEven x ≡ true + → Path ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ + ∣ (0ₖ 1 , Kn→ΩKn+1 0 x) ∣₂ + ∣ (0ₖ 1 , refl) ∣₂ +evenCharac (pos zero) isisEven i = ∣ (0ₖ 1) , (rUnit refl (~ i)) ∣₂ +evenCharac (pos (suc zero)) isisEven = ⊥-rec (true≢false (sym isisEven)) +evenCharac (pos (suc (suc zero))) isisEven = + cong ∣_∣₂ ((λ i → 0ₖ 1 , rUnit (cong ∣_∣ ((lUnit loop (~ i)) ∙ loop)) (~ i)) + ∙ (ΣPathP (cong ∣_∣ loop , λ i j → ∣ (loop ∙ loop) (i ∨ j) ∣))) +evenCharac (pos (suc (suc (suc n)))) isisEven = + (λ i → ∣ 0ₖ 1 , Kn→ΩKn+1-hom 0 (pos (suc n)) 2 i ∣₂) + ∙∙ sym (*=∙ (Kn→ΩKn+1 0 (pos (suc n))) (Kn→ΩKn+1 0 (pos 2))) + ∙∙ (cong₂ _*_ (evenCharac (pos (suc n)) isisEven) (evenCharac 2 refl)) + +evenCharac (negsuc zero) isisEven = ⊥-rec (true≢false (sym isisEven)) +evenCharac (negsuc (suc zero)) isisEven = + cong ∣_∣₂ ((λ i → 0ₖ 1 + , λ i₁ → hfill (doubleComp-faces (λ i₂ → ∣ base ∣) (λ _ → ∣ base ∣) i₁) + (inS ∣ compPath≡compPath' (sym loop) (sym loop) i i₁ ∣) (~ i)) + ∙ ΣPathP ((cong ∣_∣ (sym loop)) , λ i j → ∣ (sym loop ∙' sym loop) (i ∨ j) ∣)) +evenCharac (negsuc (suc (suc n))) isisEven = + cong ∣_∣₂ (λ i → 0ₖ 1 , Kn→ΩKn+1-hom 0 (negsuc n) -2 i) + ∙∙ sym (*=∙ (Kn→ΩKn+1 0 (negsuc n)) (Kn→ΩKn+1 0 -2)) + ∙∙ cong₂ _*_ (evenCharac (negsuc n) (isEvenNegsuc n isisEven)) (evenCharac -2 refl) + +oddCharac : (x : ℤ) → isEven x ≡ false + → Path ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ + ∣ (0ₖ 1 , Kn→ΩKn+1 0 x) ∣₂ + ∣ (0ₖ 1 , cong ∣_∣ loop) ∣₂ +oddCharac (pos zero) isOdd = ⊥-rec (true≢false isOdd) +oddCharac (pos (suc zero)) isOdd i = + ∣ (0ₖ 1 , λ j → hfill (doubleComp-faces (λ i₂ → ∣ base ∣) (λ _ → ∣ base ∣) j) + (inS ∣ lUnit loop (~ i) j ∣) (~ i)) ∣₂ +oddCharac (pos (suc (suc n))) isOdd = + (λ i → ∣ 0ₖ 1 , Kn→ΩKn+1-hom 0 (pos n) 2 i ∣₂) + ∙∙ sym (*=∙ (Kn→ΩKn+1 0 (pos n)) (Kn→ΩKn+1 0 2)) + ∙∙ cong₂ _*_ (oddCharac (pos n) isOdd) (evenCharac 2 refl) +oddCharac (negsuc zero) isOdd = + cong ∣_∣₂ ((λ i → 0ₖ 1 , rUnit (sym (cong ∣_∣ loop)) (~ i)) + ∙ ΣPathP (cong ∣_∣ (sym loop) , λ i j → ∣ hcomp (λ k → λ { (i = i0) → loop (~ j ∧ k) + ; (i = i1) → loop j + ; (j = i1) → base}) + (loop (j ∨ ~ i)) ∣)) +oddCharac (negsuc (suc zero)) isOdd = ⊥-rec (true≢false isOdd) +oddCharac (negsuc (suc (suc n))) isOdd = + cong ∣_∣₂ (λ i → 0ₖ 1 , Kn→ΩKn+1-hom 0 (negsuc n) -2 i) + ∙∙ sym (*=∙ (Kn→ΩKn+1 0 (negsuc n)) (Kn→ΩKn+1 0 -2)) + ∙∙ cong₂ _*_ (oddCharac (negsuc n) (¬isEvenNegSuc n isOdd)) (evenCharac (negsuc 1) refl) + +{- We now have all we need to establish the Iso -} +Bool→ΣKₙNilpot : Bool → ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ +Bool→ΣKₙNilpot false = ∣ 0ₖ 1 , cong ∣_∣ loop ∣₂ +Bool→ΣKₙNilpot true = ∣ 0ₖ 1 , refl ∣₂ + +testIso : Iso ∥ Σ[ x ∈ coHomK 1 ] x +ₖ x ≡ 0ₖ 1 ∥₂ Bool +fun testIso = sRec isSetBool ΣKₙNilpot→Bool +inv testIso = Bool→ΣKₙNilpot +rightInv testIso false = refl +rightInv testIso true = refl +leftInv testIso = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry (trElim + (λ _ → isGroupoidΠ λ _ → isOfHLevelPlus {n = 1} 2 (isSetSetTrunc _ _)) + (toPropElim (λ _ → isPropΠ (λ _ → isSetSetTrunc _ _)) + (λ p → path p (isEven (ΩKn+1→Kn 0 p)) refl)))) + where + path : (p : 0ₖ 1 ≡ 0ₖ 1) (b : Bool) → (isEven (ΩKn+1→Kn 0 p) ≡ b) + → Bool→ΣKₙNilpot (ΣKₙNilpot→Bool (∣ base ∣ , p)) ≡ ∣ ∣ base ∣ , p ∣₂ + path p false q = + (cong Bool→ΣKₙNilpot q) + ∙∙ sym (oddCharac (ΩKn+1→Kn 0 p) q) + ∙∙ cong ∣_∣₂ λ i → 0ₖ 1 , rightInv (Iso-Kn-ΩKn+1 0) p i + path p true q = + cong Bool→ΣKₙNilpot q + ∙∙ sym (evenCharac (ΩKn+1→Kn 0 p) q) + ∙∙ cong ∣_∣₂ λ i → 0ₖ 1 , rightInv (Iso-Kn-ΩKn+1 0) p i + + +H²-𝕂²≅Bool : GroupIso (coHomGr 2 KleinBottle) BoolGroup +H²-𝕂²≅Bool = invGroupIso (≅Bool theIso) + where + theIso : Iso _ _ + theIso = + compIso (setTruncIso + (compIso (characFunSpace𝕂² (coHomK 2)) + (Σ-cong-iso-snd + λ x → Σ-cong-iso-snd + λ p → Σ-cong-iso-snd + λ q → (movePathIso p q (isCommΩK-based 2 x))))) + (compIso Iso-H²-𝕂²₁ + (compIso + Iso-H²-𝕂²₂ + testIso)) + +------ Hⁿ(𝕂²) ≅ 0 , n ≥ 3 ------ +isContrHⁿ-𝕂² : (n : ℕ) → isContr (coHom (3 + n) KleinBottle) +isContrHⁿ-𝕂² n = + isOfHLevelRetractFromIso 0 + (setTruncIso (characFunSpace𝕂² (coHomK _))) + isContrΣ-help + where + helper : (x : coHomK (3 + n))(p : x ≡ x) → (refl ≡ p) → (q : x ≡ x) → (refl ≡ q) + → (P : p ∙∙ q ∙∙ p ≡ q) + → Path ∥ (Σ[ x ∈ coHomK (3 + n) ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙∙ q ∙∙ p ≡ q) ∥₂ + ∣ x , p , q , P ∣₂ + ∣ 0ₖ _ , refl , refl , sym (rUnit refl) ∣₂ + helper = + trElim (λ _ → isProp→isOfHLevelSuc (4 + n) (isPropΠ4 λ _ _ _ _ → isPropΠ λ _ → isSetSetTrunc _ _)) + (sphereToPropElim _ (λ _ → isPropΠ4 λ _ _ _ _ → isPropΠ λ _ → isSetSetTrunc _ _) + λ p → J (λ p _ → (q : 0ₖ _ ≡ 0ₖ _) → (refl ≡ q) + → (P : p ∙∙ q ∙∙ p ≡ q) + → Path ∥ (Σ[ x ∈ coHomK (3 + n) ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙∙ q ∙∙ p ≡ q) ∥₂ + ∣ 0ₖ _ , p , q , P ∣₂ + ∣ 0ₖ _ , refl , refl , sym (rUnit refl) ∣₂) + λ q → J (λ q _ → (P : refl ∙∙ q ∙∙ refl ≡ q) + → Path ∥ (Σ[ x ∈ coHomK (3 + n) ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙∙ q ∙∙ p ≡ q) ∥₂ + ∣ 0ₖ _ , refl , q , P ∣₂ + ∣ 0ₖ _ , refl , refl , sym (rUnit refl) ∣₂) + λ P → trRec (isProp→isOfHLevelSuc n (isSetSetTrunc _ _)) + (λ P≡rUnitrefl i → ∣ 0ₖ (3 + n) , refl , refl , P≡rUnitrefl i ∣₂) + (fun (PathIdTruncIso _) + (isContr→isProp (isConnectedPath _ (isConnectedPathKn (2 + n) _ _) + (refl ∙∙ refl ∙∙ refl) refl) + ∣ P ∣ ∣ sym (rUnit refl) ∣))) + + isContrΣ-help : isContr ∥ (Σ[ x ∈ coHomK (3 + n) ] Σ[ p ∈ x ≡ x ] Σ[ q ∈ x ≡ x ] p ∙∙ q ∙∙ p ≡ q) ∥₂ + fst isContrΣ-help = ∣ 0ₖ _ , refl , refl , sym (rUnit refl) ∣₂ + snd isContrΣ-help = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ {(x , p , q , P) + → trRec (isProp→isOfHLevelSuc (suc n) (isSetSetTrunc _ _)) + (λ pId → trRec (isProp→isOfHLevelSuc (suc n) (isSetSetTrunc _ _)) + (λ qId → sym (helper x p pId q qId P)) + (fun (PathIdTruncIso (2 + n)) + (isContr→isProp (isConnectedPathKn (2 + n) _ _) ∣ refl ∣ ∣ q ∣))) + (fun (PathIdTruncIso (2 + n)) + (isContr→isProp (isConnectedPathKn (2 + n) _ _) ∣ refl ∣ ∣ p ∣))} + +Hⁿ⁺³-𝕂²≅0 : (n : ℕ) → GroupIso (coHomGr (3 + n) KleinBottle) UnitGroup +Hⁿ⁺³-𝕂²≅0 n = contrGroupIsoUnit (isContrHⁿ-𝕂² n) diff --git a/Cubical/ZCohomology/Groups/Prelims.agda b/Cubical/ZCohomology/Groups/Prelims.agda new file mode 100644 index 000000000..8db18ee7f --- /dev/null +++ b/Cubical/ZCohomology/Groups/Prelims.agda @@ -0,0 +1,280 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Groups.Prelims where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 + +open import Cubical.Homotopy.Loopspace +open import Cubical.Data.Sigma +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +Comm to +ℤ-comm ; +Assoc to +ℤ-assoc) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec) +open import Cubical.HITs.SetTruncation renaming (elim to sElim ; map to sMap ; rec to sRec) + +infixr 33 _⋄_ + +_⋄_ : _ +_⋄_ = compIso + +-- We strengthen the elimination rule for Hⁿ(S¹). We show that we only need to work with elements ∣ f ∣₂ (definitionally) sending loop to some loop p +-- and sending base to 0 +elimFunS¹ : (n : ℕ) → (p : typ (Ω (coHomK-ptd (suc n)))) → S¹ → coHomK (suc n) +elimFunS¹ n p base = ∣ ptSn (suc n) ∣ +elimFunS¹ n p (loop i) = p i + +coHomPointedElimS¹ : ∀ {ℓ} (n : ℕ) {B : coHom (suc n) S¹ → Type ℓ} + → ((x : coHom (suc n) S¹) → isProp (B x)) + → ((p : typ (Ω (coHomK-ptd (suc n)))) → B ∣ elimFunS¹ n p ∣₂) + → (x : coHom (suc n) S¹) → B x +coHomPointedElimS¹ n {B = B} x p = + coHomPointedElim n base x + λ f Id → subst B + (cong ∣_∣₂ (funExt (λ {base → sym Id ; (loop i) j → doubleCompPath-filler (sym Id) (cong f loop) Id (~ j) i}))) + (p (sym Id ∙∙ (cong f loop) ∙∙ Id)) + +coHomPointedElimS¹2 : ∀ {ℓ} (n : ℕ) {B : (x y : coHom (suc n) S¹) → Type ℓ} + → ((x y : coHom (suc n) S¹) → isProp (B x y)) + → ((p q : typ (Ω (coHomK-ptd (suc n)))) → B ∣ elimFunS¹ n p ∣₂ ∣ elimFunS¹ n q ∣₂) + → (x y : coHom (suc n) S¹) → B x y +coHomPointedElimS¹2 n {B = B} x p = + coHomPointedElim2 _ base x λ f g fId gId + → subst2 B (cong ∣_∣₂ (funExt (λ {base → sym fId ; (loop i) j → doubleCompPath-filler (sym (fId)) (cong f loop) fId (~ j) i}))) + (cong ∣_∣₂ (funExt (λ {base → sym gId ; (loop i) j → doubleCompPath-filler (sym (gId)) (cong g loop) gId (~ j) i}))) + (p (sym fId ∙∙ cong f loop ∙∙ fId) (sym gId ∙∙ cong g loop ∙∙ gId)) + +-- We do the same thing for Sⁿ, n ≥ 2. +elimFunSⁿ : (n m : ℕ) (p : S₊ (suc m) → typ (Ω (coHomK-ptd (suc n)))) + → (S₊ (2 + m)) → coHomK (suc n) +elimFunSⁿ n m p north = ∣ ptSn (suc n) ∣ +elimFunSⁿ n m p south = ∣ ptSn (suc n) ∣ +elimFunSⁿ n m p (merid a i) = p a i + +coHomPointedElimSⁿ : ∀ {ℓ} (n m : ℕ) {B : (x : coHom (suc n) (S₊ (2 + m))) → Type ℓ} + → ((x : coHom (suc n) (S₊ (2 + m))) → isProp (B x)) + → ((p : _) → B ∣ elimFunSⁿ n m p ∣₂) + → (x : coHom (suc n) (S₊ (2 + m))) → B x +coHomPointedElimSⁿ n m {B = B} isprop ind = + coHomPointedElim n north isprop + λ f fId → subst B (cong ∣_∣₂ (funExt (λ {north → sym fId + ; south → sym fId ∙' cong f (merid (ptSn (suc m))) + ; (merid a i) j → hcomp (λ k → λ {(i = i0) → fId (~ j ∧ k) + ; (i = i1) → compPath'-filler (sym fId) + (cong f (merid (ptSn (suc m)))) k j + ; (j = i1) → f (merid a i)}) + (hcomp (λ k → λ {(i = i0) → f north ; + (i = i1) → f (merid (ptSn (suc m)) (j ∨ ~ k)) ; + (j = i1) → f (merid a i)}) + (f (merid a i)))}))) + (ind λ a → sym fId ∙∙ cong f (merid a) ∙ cong f (sym (merid (ptSn (suc m)))) ∙∙ fId) + +0₀ = 0ₖ 0 +0₁ = 0ₖ 1 +0₂ = 0ₖ 2 +0₃ = 0ₖ 3 +0₄ = 0ₖ 4 + +S¹map : hLevelTrunc 3 S¹ → S¹ +S¹map = trRec isGroupoidS¹ (idfun _) + +S¹map-id : (x : hLevelTrunc 3 S¹) → Path (hLevelTrunc 3 S¹) ∣ S¹map x ∣ x +S¹map-id = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → refl + + +-- We prove that (S¹ → ∥ S¹ ∥) ≃ S¹ × ℤ (Needed for H¹(S¹)). Note that the truncation doesn't really matter, since S¹ is a groupoid. +-- Given a map f : S¹ → S¹, the idea is to send this to (f(base) , winding (f(loop))). For this to be +-- well-typed, we need to translate f(loop) into an element in Ω(S¹,base). + +S¹→S¹≡S¹×ℤ : Iso (S¹ → hLevelTrunc 3 S¹) (S¹ × ℤ) +Iso.fun S¹→S¹≡S¹×ℤ f = S¹map (f base) + , winding (basechange2⁻ (S¹map (f base)) λ i → S¹map (f (loop i))) +Iso.inv S¹→S¹≡S¹×ℤ (s , int) base = ∣ s ∣ +Iso.inv S¹→S¹≡S¹×ℤ (s , int) (loop i) = ∣ basechange2 s (intLoop int) i ∣ +Iso.rightInv S¹→S¹≡S¹×ℤ (s , int) = ΣPathP (refl , ((λ i → winding (basechange2-retr s (λ i → intLoop int i) i)) + ∙ windingℤLoop int)) +Iso.leftInv S¹→S¹≡S¹×ℤ f = funExt λ { base → S¹map-id (f base) + ; (loop i) j → helper j i} + where + helper : PathP (λ i → S¹map-id (f base) i ≡ S¹map-id (f base) i) + (λ i → ∣ basechange2 (S¹map (f base)) + (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁)))))) i ∣) + (cong f loop) + helper i j = + hcomp (λ k → λ { (i = i0) → cong ∣_∣ (basechange2 (S¹map (f base)) + (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁))))))) j + ; (i = i1) → S¹map-id (f (loop j)) k + ; (j = i0) → S¹map-id (f base) (i ∧ k) + ; (j = i1) → S¹map-id (f base) (i ∧ k)}) + (helper2 i j) + where + helper2 : Path (Path (hLevelTrunc 3 _) _ _) + (cong ∣_∣ (basechange2 (S¹map (f base)) + (intLoop + (winding + (basechange2⁻ (S¹map (f base)) + (λ i₁ → S¹map (f (loop i₁)))))))) + λ i → ∣ S¹map (f (loop i)) ∣ + helper2 i j = + ∣ ((cong (basechange2 (S¹map (f base))) + (decodeEncode base (basechange2⁻ (S¹map (f base)) + (λ i₁ → S¹map (f (loop i₁))))) + ∙ basechange2-sect (S¹map (f base)) + (λ i → S¹map (f (loop i)))) i) j ∣ + +{- Proof that (S¹ → K₁) ≃ K₁ × ℤ. Needed for H¹(T²) -} +S1→K₁≡S1×ℤ : Iso ((S₊ 1) → coHomK 1) (coHomK 1 × ℤ) +S1→K₁≡S1×ℤ = S¹→S¹≡S¹×ℤ ⋄ prodIso (invIso (truncIdempotentIso 3 (isGroupoidS¹))) idIso + +module _ (key : Unit') where + module P = lockedCohom key + private + _+K_ : {n : ℕ} → coHomK n → coHomK n → coHomK n + _+K_ {n = n} = P.+K n + + -K_ : {n : ℕ} → coHomK n → coHomK n + -K_ {n = n} = P.-K n + + _-K_ : {n : ℕ} → coHomK n → coHomK n → coHomK n + _-K_ {n = n} = P.-Kbin n + + infixr 55 _+K_ + infixr 55 -K_ + infixr 56 _-K_ + + {- Proof that S¹→K2 is isomorphic to K2×K1 (as types). Needed for H²(T²) -} + S1→K2≡K2×K1' : Iso (S₊ 1 → coHomK 2) (coHomK 2 × coHomK 1) + Iso.fun S1→K2≡K2×K1' f = f base , ΩKn+1→Kn 1 (sym (P.rCancelK 2 (f base)) + ∙∙ cong (λ x → (f x) -K f base) loop + ∙∙ P.rCancelK 2 (f base)) + Iso.inv S1→K2≡K2×K1' = invmap + where + invmap : (∥ Susp S¹ ∥ 4) × (∥ S¹ ∥ 3) → S¹ → ∥ Susp S¹ ∥ 4 + invmap (a , b) base = a +K 0₂ + invmap (a , b) (loop i) = a +K Kn→ΩKn+1 1 b i + Iso.rightInv S1→K2≡K2×K1' (a , b) = ΣPathP ((P.rUnitK 2 a) + , (cong (ΩKn+1→Kn 1) (doubleCompPath-elim' (sym (P.rCancelK 2 (a +K 0₂))) + (λ i → (a +K Kn→ΩKn+1 1 b i) -K (a +K 0₂)) + (P.rCancelK 2 (a +K 0₂))) + ∙∙ cong (ΩKn+1→Kn 1) (congHelper2 (Kn→ΩKn+1 1 b) (λ x → (a +K x) -K (a +K 0₂)) + (funExt (λ x → sym (cancelHelper a x))) + (P.rCancelK 2 (a +K 0₂))) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 1) b)) + + module _ where + cancelHelper : (a b : coHomK 2) → (a +K b) -K (a +K 0₂) ≡ b + cancelHelper a b = cong (λ x → (a +K b) -K x) (P.rUnitK 2 a) + ∙ P.-cancelLK 2 a b + + congHelper2 : (p : 0₂ ≡ 0₂) (f : coHomK 2 → coHomK 2) (id : (λ x → x) ≡ f) → (q : (f 0₂) ≡ 0₂) + → (sym q) ∙ cong f p ∙ q ≡ p + congHelper2 p f = J (λ f _ → (q : (f 0₂) ≡ 0₂) → (sym q) ∙ cong f p ∙ q ≡ p) + λ q → (cong (sym q ∙_) (isCommΩK 2 p _) ∙∙ assoc _ _ _ ∙∙ cong (_∙ p) (lCancel q)) + ∙ sym (lUnit p) + + conghelper3 : (x : coHomK 2) (p : x ≡ x) (f : coHomK 2 → coHomK 2) (id : (λ x → x) ≡ f) (q : f x ≡ x) + → (sym q) ∙ cong f p ∙ q ≡ p + conghelper3 x p f = J (λ f _ → (q : (f x) ≡ x) → (sym q) ∙ cong f p ∙ q ≡ p) + λ q → (cong (sym q ∙_) (isCommΩK-based 2 x p _) ∙∙ assoc _ _ _ ∙∙ cong (_∙ p) (lCancel q)) + ∙ sym (lUnit p) + Iso.leftInv S1→K2≡K2×K1' a i base = P.rUnitK _ (a base) i + Iso.leftInv S1→K2≡K2×K1' a i (loop j) = loop-helper i j + where + loop-helper : PathP (λ i → P.rUnitK _ (a base) i ≡ P.rUnitK _ (a base) i) + (cong (a base +K_) (Kn→ΩKn+1 1 (ΩKn+1→Kn 1 ((sym (P.rCancelK 2 (a base)) + ∙∙ (λ i → a (loop i) -K (a (base))) + ∙∙ P.rCancelK 2 (a base)))))) + (cong a loop) + loop-helper i j = + hcomp (λ k → λ { (i = i0) → (G (doubleCompPath-elim' + (sym rP) (λ i₁ → a (loop i₁) -K a base) rP (~ k))) j + ; (i = i1) → cong a loop j + ; (j = i0) → P.rUnitK 2 (a base) i + ; (j = i1) → P.rUnitK 2 (a base) i}) + (loop-helper2 i j) + + where + F : typ (Ω (coHomK-ptd 2)) → a base +K snd (coHomK-ptd 2) ≡ a base +K snd (coHomK-ptd 2) + F = cong (_+K_ {n = 2} (a base)) + + G : 0ₖ 2 ≡ 0ₖ 2 → a base +K snd (coHomK-ptd 2) ≡ a base +K snd (coHomK-ptd 2) + G p = F (Kn→ΩKn+1 1 (ΩKn+1→Kn 1 p)) + + rP : P.+K 2 (a base) (P.-K 2 (a base)) ≡ 0ₖ 2 + rP = P.rCancelK 2 (a base) + + lem : (x : coHomK 2) (p : x ≡ x) (q : 0₂ ≡ x) + → Kn→ΩKn+1 1 (ΩKn+1→Kn 1 (q ∙ p ∙ sym q)) ≡ q ∙ p ∙ sym q + lem x p q = Iso.rightInv (Iso-Kn-ΩKn+1 1) (q ∙ p ∙ sym q) + + subtr-lem : (a b : hLevelTrunc 4 (S₊ 2)) → a +K (b -K a) ≡ b + subtr-lem a b = P.commK 2 a (b -K a) ∙ P.-+cancelK 2 b a + + subtr-lem-coher : PathP (λ i → subtr-lem (a base) (a base) i ≡ subtr-lem (a base) (a base) i) + (cong (a base +K_) (λ i₁ → a (loop i₁) -K a base)) + λ i → a (loop i) + subtr-lem-coher i j = subtr-lem (a base) (a (loop j)) i + + abstract + helperFun2 : {A : Type₀} {0A a b : A} (main : 0A ≡ 0A) (start : b ≡ b) (p : a ≡ a) (q : a ≡ b) (r : b ≡ 0A) (Q : a ≡ 0A) + (R : PathP (λ i → Q i ≡ Q i) p main) + → start ≡ sym q ∙ p ∙ q + → isComm∙ (A , 0A) + → sym r ∙ start ∙ r ≡ main + helperFun2 main start p q r Q R startId comm = + sym r ∙ start ∙ r ≡[ i ]⟨ sym r ∙ startId i ∙ r ⟩ + sym r ∙ (sym q ∙ p ∙ q) ∙ r ≡[ i ]⟨ sym r ∙ assoc (sym q) (p ∙ q) r (~ i) ⟩ + sym r ∙ sym q ∙ (p ∙ q) ∙ r ≡[ i ]⟨ sym r ∙ sym q ∙ assoc p q r (~ i) ⟩ + sym r ∙ sym q ∙ p ∙ q ∙ r ≡[ i ]⟨ assoc (sym r) (rUnit (sym q) i) (p ∙ lUnit (q ∙ r) i) i ⟩ + (sym r ∙ sym q ∙ refl) ∙ p ∙ refl ∙ q ∙ r ≡[ i ]⟨ (sym r ∙ sym q ∙ λ j → Q (i ∧ j)) ∙ R i ∙ (λ j → Q ( i ∧ (~ j))) ∙ q ∙ r ⟩ + (sym r ∙ sym q ∙ Q) ∙ main ∙ sym Q ∙ q ∙ r ≡[ i ]⟨ (sym r ∙ sym q ∙ Q) ∙ main ∙ sym Q ∙ symDistr (sym r) (sym q) (~ i) ⟩ + (sym r ∙ sym q ∙ Q) ∙ main ∙ sym Q ∙ sym (sym r ∙ sym q) ≡[ i ]⟨ (assoc (sym r) (sym q) Q i) ∙ main ∙ symDistr (sym r ∙ sym q) Q (~ i) ⟩ + ((sym r ∙ sym q) ∙ Q) ∙ main ∙ sym ((sym r ∙ sym q) ∙ Q) ≡[ i ]⟨ ((sym r ∙ sym q) ∙ Q) ∙ comm main (sym ((sym r ∙ sym q) ∙ Q)) i ⟩ + ((sym r ∙ sym q) ∙ Q) ∙ sym ((sym r ∙ sym q) ∙ Q) ∙ main ≡⟨ assoc ((sym r ∙ sym q) ∙ Q) (sym ((sym r ∙ sym q) ∙ Q)) main ⟩ + (((sym r ∙ sym q) ∙ Q) ∙ sym ((sym r ∙ sym q) ∙ Q)) ∙ main ≡[ i ]⟨ rCancel (((sym r ∙ sym q) ∙ Q)) i ∙ main ⟩ + refl ∙ main ≡⟨ sym (lUnit main) ⟩ + main ∎ + + congFunct₃ : ∀ {A B : Type₀} {a b c d : A} (f : A → B) (p : a ≡ b) (q : b ≡ c) (r : c ≡ d) + → cong f (p ∙ q ∙ r) ≡ cong f p ∙ cong f q ∙ cong f r + congFunct₃ f p q r = congFunct f p (q ∙ r) ∙ cong (cong f p ∙_) (congFunct f q r) + + lem₀ : G (sym rP ∙ (λ i₁ → a (loop i₁) -K a base) ∙ rP) + ≡ F (sym rP ∙ (λ i₁ → a (loop i₁) -K a base) ∙ rP) + lem₀ = cong F (lem (a base -K (a base)) + (λ i₁ → a (loop i₁) -K a base) + (sym (P.rCancelK 2 (a base)))) + + lem₁ : G (sym rP ∙ (λ i₁ → a (loop i₁) -K a base) ∙ rP) + ≡ sym (λ i₁ → a base +K P.rCancelK 2 (a base) i₁) ∙ + cong (a base +K_) (λ i₁ → a (loop i₁) -K a base) ∙ + (λ i₁ → a base +K P.rCancelK 2 (a base) i₁) + lem₁ = lem₀ ∙ congFunct₃ (a base +K_) (sym rP) + (λ i₁ → a (loop i₁) -K a base) + rP + + loop-helper2 : PathP (λ i → P.rUnitK _ (a base) i ≡ P.rUnitK _ (a base) i) + (cong (a base +K_) (Kn→ΩKn+1 1 (ΩKn+1→Kn 1 ((sym (P.rCancelK 2 (a base)) + ∙ (λ i → a (loop i) -K (a (base))) + ∙ P.rCancelK 2 (a base)))))) + (cong a loop) + loop-helper2 = compPathL→PathP (helperFun2 (cong a loop) _ _ + (cong (a base +K_) (P.rCancelK 2 (a base))) _ _ + subtr-lem-coher + lem₁ + (isCommΩK-based 2 (a base))) + +S1→K2≡K2×K1 : Iso (S₊ 1 → coHomK 2) (coHomK 2 × coHomK 1) +S1→K2≡K2×K1 = S1→K2≡K2×K1' unlock diff --git a/Cubical/ZCohomology/Groups/RP2.agda b/Cubical/ZCohomology/Groups/RP2.agda new file mode 100644 index 000000000..0154df314 --- /dev/null +++ b/Cubical/ZCohomology/Groups/RP2.agda @@ -0,0 +1,128 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Groups.RP2 where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.Groups.KleinBottle +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to pRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim) hiding (map) +open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; elim2 to trElim2) +open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Bool to BoolGroup ; Unit to UnitGroup) + +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Transport + +open import Cubical.ZCohomology.Groups.Connected + +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Isomorphism +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.Foundations.Equiv +open import Cubical.Homotopy.Connected +open import Cubical.HITs.RPn.Base + +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Bool +open import Cubical.Data.Int + +open import Cubical.Foundations.Path + +private + variable + ℓ : Level + A : Type ℓ + +funSpaceIso-RP² : Iso (RP² → A) (Σ[ x ∈ A ] Σ[ p ∈ x ≡ x ] p ≡ sym p) +Iso.fun funSpaceIso-RP² f = f point , (cong f line , λ i j → f (square i j)) +Iso.inv funSpaceIso-RP² (x , p , P) point = x +Iso.inv funSpaceIso-RP² (x , p , P) (line i) = p i +Iso.inv funSpaceIso-RP² (x , p , P) (square i j) = P i j +Iso.rightInv funSpaceIso-RP² (x , p , P) i = x , p , P +Iso.leftInv funSpaceIso-RP² f _ point = f point +Iso.leftInv funSpaceIso-RP² f _ (line i) = f (line i) +Iso.leftInv funSpaceIso-RP² f _ (square i j) = f (square i j) + +private + pathIso : {x : A} {p : x ≡ x} → Iso (p ≡ sym p) (p ∙ p ≡ refl) + pathIso {p = p} = compIso (congIso (equivToIso (_ , compPathr-isEquiv p))) + (pathToIso (cong (p ∙ p ≡_) (lCancel p))) + +--- H⁰(RP²) ≅ ℤ ---- +H⁰-RP²≅ℤ : GroupIso (coHomGr 0 RP²) ℤGroup +H⁰-RP²≅ℤ = H⁰-connected point connectedRP¹ + where + connectedRP¹ : (x : RP²) → ∥ point ≡ x ∥ + connectedRP¹ point = ∣ refl ∣ + connectedRP¹ (line i) = + isOfHLevel→isOfHLevelDep 1 {B = λ x → ∥ point ≡ x ∥} + (λ _ → isPropPropTrunc) ∣ refl ∣ ∣ refl ∣ line i + connectedRP¹ (square i j) = helper i j + where + helper : SquareP (λ i j → ∥ point ≡ square i j ∥) + (isOfHLevel→isOfHLevelDep 1 {B = λ x → ∥ point ≡ x ∥} + (λ _ → isPropPropTrunc) ∣ refl ∣ ∣ refl ∣ line) + (symP (isOfHLevel→isOfHLevelDep 1 {B = λ x → ∥ point ≡ x ∥} + (λ _ → isPropPropTrunc) ∣ refl ∣ ∣ refl ∣ line)) + refl refl + helper = toPathP (isOfHLevelPathP 1 isPropPropTrunc _ _ _ _) + +--- H¹(RP²) ≅ 0 ---- +isContr-H¹-RP²-helper : isContr ∥ Σ[ x ∈ coHomK 1 ] Σ[ p ∈ x ≡ x ] p ∙ p ≡ refl ∥₂ +fst isContr-H¹-RP²-helper = ∣ 0ₖ 1 , refl , sym (rUnit refl) ∣₂ +snd isContr-H¹-RP²-helper = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry + (trElim (λ _ → isGroupoidΠ λ _ → isOfHLevelPlus {n = 1} 2 (isSetSetTrunc _ _)) + (toPropElim (λ _ → isPropΠ (λ _ → isSetSetTrunc _ _)) + λ {(p , nilp) + → cong ∣_∣₂ (ΣPathP (refl , Σ≡Prop (λ _ → isOfHLevelTrunc 3 _ _ _ _) + (rUnit refl + ∙∙ cong (Kn→ΩKn+1 0) (sym (nilpotent→≡0 (ΩKn+1→Kn 0 p) + (sym (ΩKn+1→Kn-hom 0 p p) + ∙ cong (ΩKn+1→Kn 0) nilp))) + ∙∙ Iso.rightInv (Iso-Kn-ΩKn+1 0) p)))}))) + +H¹-RP²≅0 : GroupIso (coHomGr 1 RP²) UnitGroup +H¹-RP²≅0 = + contrGroupIsoUnit + (isOfHLevelRetractFromIso 0 + (setTruncIso (compIso funSpaceIso-RP² + (Σ-cong-iso-snd (λ _ → Σ-cong-iso-snd λ _ → pathIso)))) + isContr-H¹-RP²-helper) + +--- H²(RP²) ≅ ℤ/2ℤ ---- + +Iso-H²-RP²₁ : Iso ∥ Σ[ x ∈ coHomK 2 ] Σ[ p ∈ x ≡ x ] p ≡ sym p ∥₂ + ∥ Σ[ p ∈ 0ₖ 2 ≡ 0ₖ 2 ] p ≡ sym p ∥₂ +Iso.fun Iso-H²-RP²₁ = + sRec isSetSetTrunc + (uncurry + (trElim (λ _ → is2GroupoidΠ λ _ → isOfHLevelPlus {n = 2} 2 isSetSetTrunc) + (sphereElim _ (λ _ → isSetΠ (λ _ → isSetSetTrunc)) + λ p → ∣ fst p , snd p ∣₂))) +Iso.inv Iso-H²-RP²₁ = sMap λ p → (0ₖ 2) , p +Iso.rightInv Iso-H²-RP²₁ = sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ _ → refl +Iso.leftInv Iso-H²-RP²₁ = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry (trElim (λ _ → is2GroupoidΠ λ _ → isOfHLevelPlus {n = 1} 3 (isSetSetTrunc _ _)) + (sphereToPropElim _ (λ _ → isPropΠ (λ _ → isSetSetTrunc _ _)) + λ p → refl))) + +Iso-H²-RP²₂ : Iso ∥ Σ[ p ∈ 0ₖ 2 ≡ 0ₖ 2 ] p ≡ sym p ∥₂ Bool +Iso-H²-RP²₂ = compIso (setTruncIso (Σ-cong-iso-snd λ _ → pathIso)) + (compIso Iso-H²-𝕂²₂ testIso) + + +H²-RP²≅Bool : GroupIso (coHomGr 2 RP²) BoolGroup +H²-RP²≅Bool = invGroupIso (≅Bool (compIso + (compIso (setTruncIso funSpaceIso-RP²) + Iso-H²-RP²₁) + Iso-H²-RP²₂)) diff --git a/Cubical/ZCohomology/Groups/Sn.agda b/Cubical/ZCohomology/Groups/Sn.agda new file mode 100644 index 000000000..d2116749e --- /dev/null +++ b/Cubical/ZCohomology/Groups/Sn.agda @@ -0,0 +1,330 @@ +{-# OPTIONS --safe #-} +module Cubical.ZCohomology.Groups.Sn where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Connected +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Prelims + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.S1 +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₁ ; ∣_∣ to ∣_∣₁) + hiding (map) + +open import Cubical.Relation.Nullary +open import Cubical.Data.Sum hiding (map) +open import Cubical.Data.Empty renaming (rec to ⊥-rec) + +open import Cubical.Data.Bool +open import Cubical.Data.Sigma +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +Comm to +ℤ-comm ; +Assoc to +ℤ-assoc) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec) + +open import Cubical.Homotopy.Connected + +open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) + +infixr 31 _□_ +_□_ : _ +_□_ = compGroupIso + +open IsGroupHom +open BijectionIso +open Iso + +Sn-connected : (n : ℕ) (x : typ (S₊∙ (suc n))) → ∥ pt (S₊∙ (suc n)) ≡ x ∥₁ +Sn-connected zero = toPropElim (λ _ → isPropPropTrunc) ∣ refl ∣₁ +Sn-connected (suc zero) = suspToPropElim base (λ _ → isPropPropTrunc) ∣ refl ∣₁ +Sn-connected (suc (suc n)) = suspToPropElim north (λ _ → isPropPropTrunc) ∣ refl ∣₁ + +suspensionAx-Sn : (n m : ℕ) → GroupIso (coHomGr (2 + n) (S₊ (2 + m))) + (coHomGr (suc n) (S₊ (suc m))) +suspensionAx-Sn n m = + compIso (setTruncIso (invIso funSpaceSuspIso)) helperIso , + makeIsGroupHom funIsHom + where + helperIso : Iso ∥ (Σ[ x ∈ coHomK (2 + n) ] + Σ[ y ∈ coHomK (2 + n) ] + (S₊ (suc m) → x ≡ y)) ∥₂ + (coHom (suc n) (S₊ (suc m))) + Iso.fun helperIso = + sRec isSetSetTrunc + (uncurry + (coHomK-elim _ + (λ _ → isOfHLevelΠ (2 + n) + λ _ → isOfHLevelPlus' {n = n} 2 isSetSetTrunc) + (uncurry + (coHomK-elim _ + (λ _ → isOfHLevelΠ (2 + n) + λ _ → isOfHLevelPlus' {n = n} 2 isSetSetTrunc) + λ f → ∣ (λ x → ΩKn+1→Kn (suc n) (f x)) ∣₂)))) + Iso.inv helperIso = + sMap λ f → (0ₖ _) , (0ₖ _ , λ x → Kn→ΩKn+1 (suc n) (f x)) + Iso.rightInv helperIso = + coHomPointedElim _ (ptSn (suc m)) (λ _ → isSetSetTrunc _ _) + λ f fId → cong ∣_∣₂ (funExt (λ x → Iso.leftInv (Iso-Kn-ΩKn+1 _) (f x))) + Iso.leftInv helperIso = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry + (coHomK-elim _ + (λ _ → isProp→isOfHLevelSuc (suc n) (isPropΠ λ _ → isSetSetTrunc _ _)) + (uncurry + (coHomK-elim _ + (λ _ → isProp→isOfHLevelSuc (suc n) (isPropΠ λ _ → isSetSetTrunc _ _)) + λ f → cong ∣_∣₂ + (ΣPathP (refl , + ΣPathP (refl , + (λ i x → Iso.rightInv (Iso-Kn-ΩKn+1 (suc n)) (f x) i)))))))) + + theFun : coHom (2 + n) (S₊ (2 + m)) → coHom (suc n) (S₊ (suc m)) + theFun = Iso.fun (compIso (setTruncIso (invIso funSpaceSuspIso)) + helperIso) + + funIsHom : (x y : coHom (2 + n) (S₊ (2 + m))) + → theFun (x +ₕ y) ≡ theFun x +ₕ theFun y + funIsHom = + coHomPointedElimSⁿ _ _ (λ _ → isPropΠ λ _ → isSetSetTrunc _ _) + λ f → coHomPointedElimSⁿ _ _ (λ _ → isSetSetTrunc _ _) + λ g → cong ∣_∣₂ (funExt λ x → cong (ΩKn+1→Kn (suc n)) (sym (∙≡+₂ n (f x) (g x))) + ∙ ΩKn+1→Kn-hom (suc n) (f x) (g x)) + + +H⁰-Sⁿ≅ℤ : (n : ℕ) → GroupIso (coHomGr 0 (S₊ (suc n))) ℤGroup +H⁰-Sⁿ≅ℤ zero = H⁰-connected base (Sn-connected 0) +H⁰-Sⁿ≅ℤ (suc n) = H⁰-connected north (Sn-connected (suc n)) + +-- -- ---------------------------------------------------------------------- + +--- We will need to switch between Sⁿ defined using suspensions and using pushouts +--- in order to apply Mayer Vietoris. + + +S1Iso : Iso S¹ (Pushout {A = Bool} (λ _ → tt) λ _ → tt) +S1Iso = S¹IsoSuspBool ⋄ invIso PushoutSuspIsoSusp + +coHomPushout≅coHomSn : (n m : ℕ) → GroupIso (coHomGr m (S₊ (suc n))) + (coHomGr m (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt)) +coHomPushout≅coHomSn zero m = + setTruncIso (domIso S1Iso) , + makeIsGroupHom (sElim2 (λ _ _ → isSet→isGroupoid isSetSetTrunc _ _) (λ _ _ → refl)) +coHomPushout≅coHomSn (suc n) m = + setTruncIso (domIso (invIso PushoutSuspIsoSusp)) , + makeIsGroupHom (sElim2 (λ _ _ → isSet→isGroupoid isSetSetTrunc _ _) (λ _ _ → refl)) + +-------------------------- H⁰(S⁰) ----------------------------- +S0→ℤ : (a : ℤ × ℤ) → S₊ 0 → ℤ +S0→ℤ a true = fst a +S0→ℤ a false = snd a + +H⁰-S⁰≅ℤ×ℤ : GroupIso (coHomGr 0 (S₊ 0)) (DirProd ℤGroup ℤGroup) +fun (fst H⁰-S⁰≅ℤ×ℤ) = sRec (isSet× isSetℤ isSetℤ) λ f → (f true) , (f false) +inv (fst H⁰-S⁰≅ℤ×ℤ) a = ∣ S0→ℤ a ∣₂ +rightInv (fst H⁰-S⁰≅ℤ×ℤ) _ = refl +leftInv (fst H⁰-S⁰≅ℤ×ℤ) = + sElim (λ _ → isSet→isGroupoid isSetSetTrunc _ _) + (λ f → cong ∣_∣₂ (funExt (λ {true → refl ; false → refl}))) +snd H⁰-S⁰≅ℤ×ℤ = + makeIsGroupHom + (sElim2 (λ _ _ → isSet→isGroupoid (isSet× isSetℤ isSetℤ) _ _) λ a b → refl) + + +------------------------- H¹(S⁰) ≅ 0 ------------------------------- + + +private + Hⁿ-S0≃Kₙ×Kₙ : (n : ℕ) → Iso (S₊ 0 → coHomK (suc n)) (coHomK (suc n) × coHomK (suc n)) + Iso.fun (Hⁿ-S0≃Kₙ×Kₙ n) f = (f true) , (f false) + Iso.inv (Hⁿ-S0≃Kₙ×Kₙ n) (a , b) true = a + Iso.inv (Hⁿ-S0≃Kₙ×Kₙ n) (a , b) false = b + Iso.rightInv (Hⁿ-S0≃Kₙ×Kₙ n) a = refl + Iso.leftInv (Hⁿ-S0≃Kₙ×Kₙ n) b = funExt λ {true → refl ; false → refl} + + isContrHⁿ-S0 : (n : ℕ) → isContr (coHom (suc n) (S₊ 0)) + isContrHⁿ-S0 n = isContrRetract (Iso.fun (setTruncIso (Hⁿ-S0≃Kₙ×Kₙ n))) + (Iso.inv (setTruncIso (Hⁿ-S0≃Kₙ×Kₙ n))) + (Iso.leftInv (setTruncIso (Hⁿ-S0≃Kₙ×Kₙ n))) + (isContrHelper n) + where + isContrHelper : (n : ℕ) → isContr (∥ (coHomK (suc n) × coHomK (suc n)) ∥₂) + fst (isContrHelper zero) = ∣ (0₁ , 0₁) ∣₂ + snd (isContrHelper zero) = sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ y → elim2 {B = λ x y → ∣ (0₁ , 0₁) ∣₂ ≡ ∣(x , y) ∣₂ } + (λ _ _ → isOfHLevelPlus {n = 2} 2 isSetSetTrunc _ _) + (toPropElim2 (λ _ _ → isSetSetTrunc _ _) refl) (fst y) (snd y) + isContrHelper (suc zero) = ∣ (0₂ , 0₂) ∣₂ + , sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ y → elim2 {B = λ x y → ∣ (0₂ , 0₂) ∣₂ ≡ ∣(x , y) ∣₂ } + (λ _ _ → isOfHLevelPlus {n = 2} 3 isSetSetTrunc _ _) + (suspToPropElim2 base (λ _ _ → isSetSetTrunc _ _) refl) (fst y) (snd y) + isContrHelper (suc (suc n)) = ∣ (0ₖ (3 + n) , 0ₖ (3 + n)) ∣₂ + , sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ y → elim2 {B = λ x y → ∣ (0ₖ (3 + n) , 0ₖ (3 + n)) ∣₂ ≡ ∣(x , y) ∣₂ } + (λ _ _ → isProp→isOfHLevelSuc (4 + n) (isSetSetTrunc _ _)) + (suspToPropElim2 north (λ _ _ → isSetSetTrunc _ _) refl) (fst y) (snd y) + +H¹-S⁰≅0 : (n : ℕ) → GroupIso (coHomGr (suc n) (S₊ 0)) UnitGroup +H¹-S⁰≅0 n = contrGroupIsoUnit (isContrHⁿ-S0 n) + +------------------------- H²(S¹) ≅ 0 ------------------------------- + +Hⁿ-S¹≅0 : (n : ℕ) → GroupIso (coHomGr (2 + n) (S₊ 1)) UnitGroup +Hⁿ-S¹≅0 n = contrGroupIsoUnit + (isOfHLevelRetractFromIso 0 helper + (_ , helper2)) + where + helper : Iso ⟨ coHomGr (2 + n) (S₊ 1)⟩ ∥ Σ (hLevelTrunc (4 + n) (S₊ (2 + n))) (λ x → ∥ x ≡ x ∥₂) ∥₂ + helper = compIso (setTruncIso IsoFunSpaceS¹) IsoSetTruncateSndΣ + + helper2 : (x : ∥ Σ (hLevelTrunc (4 + n) (S₊ (2 + n))) (λ x → ∥ x ≡ x ∥₂) ∥₂) → ∣ ∣ north ∣ , ∣ refl ∣₂ ∣₂ ≡ x + helper2 = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry + (trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isProp→isOfHLevelSuc (3 + n) (isSetSetTrunc _ _)) + (suspToPropElim (ptSn (suc n)) (λ _ → isPropΠ λ _ → isSetSetTrunc _ _) + (sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ p + → cong ∣_∣₂ (ΣPathP (refl , isContr→isProp helper3 _ _)))))) + where + helper4 : isConnected (n + 3) (hLevelTrunc (4 + n) (S₊ (2 + n))) + helper4 = subst (λ m → isConnected m (hLevelTrunc (4 + n) (S₊ (2 + n)))) (+-comm 3 n) + (isOfHLevelRetractFromIso 0 (invIso (truncOfTruncIso (3 + n) 1)) (sphereConnected (2 + n))) + + helper3 : isContr ∥ ∣ north ∣ ≡ ∣ north ∣ ∥₂ + helper3 = isOfHLevelRetractFromIso 0 setTruncTrunc2Iso + (isConnectedPath 2 (isConnectedSubtr 3 n helper4) _ _) + +-- --------------- H¹(Sⁿ), n ≥ 1 -------------------------------------------- + +H¹-Sⁿ≅0 : (n : ℕ) → GroupIso (coHomGr 1 (S₊ (2 + n))) UnitGroup +H¹-Sⁿ≅0 zero = contrGroupIsoUnit isContrH¹S² + where + isContrH¹S² : isContr ⟨ coHomGr 1 (S₊ 2) ⟩ + isContrH¹S² = ∣ (λ _ → ∣ base ∣) ∣₂ + , coHomPointedElim 0 north (λ _ → isSetSetTrunc _ _) + λ f p → cong ∣_∣₂ (funExt λ x → sym p ∙∙ sym (spoke f north) ∙∙ spoke f x) +H¹-Sⁿ≅0 (suc n) = contrGroupIsoUnit isContrH¹S³⁺ⁿ + where + anIso : Iso ⟨ coHomGr 1 (S₊ (3 + n)) ⟩ ∥ (S₊ (3 + n) → hLevelTrunc (4 + n) (coHomK 1)) ∥₂ + anIso = + setTruncIso + (codomainIso + (invIso (truncIdempotentIso (4 + n) (isOfHLevelPlus' {n = 1 + n} 3 (isOfHLevelTrunc 3))))) + + isContrH¹S³⁺ⁿ-ish : (f : (S₊ (3 + n) → hLevelTrunc (4 + n) (coHomK 1))) + → ∣ (λ _ → ∣ ∣ base ∣ ∣) ∣₂ ≡ ∣ f ∣₂ + isContrH¹S³⁺ⁿ-ish f = ind (f north) refl + where + ind : (x : hLevelTrunc (4 + n) (coHomK 1)) + → x ≡ f north + → ∣ (λ _ → ∣ ∣ base ∣ ∣) ∣₂ ≡ ∣ f ∣₂ + ind = trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelPlus' {n = (3 + n)} 1 (isSetSetTrunc _ _)) + (trElim (λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPlus {n = 1} 2 (isSetSetTrunc _ _)) + (toPropElim (λ _ → isPropΠ λ _ → isSetSetTrunc _ _) + λ p → cong ∣_∣₂ (funExt λ x → p ∙∙ sym (spoke f north) ∙∙ spoke f x))) + isContrH¹S³⁺ⁿ : isContr ⟨ coHomGr 1 (S₊ (3 + n)) ⟩ + isContrH¹S³⁺ⁿ = + isOfHLevelRetractFromIso 0 + anIso + (∣ (λ _ → ∣ ∣ base ∣ ∣) ∣₂ + , sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) isContrH¹S³⁺ⁿ-ish) + +--------- H¹(S¹) ≅ ℤ ------- +{- +Idea : +H¹(S¹) := ∥ S¹ → K₁ ∥₂ + ≃ ∥ S¹ → S¹ ∥₂ + ≃ ∥ S¹ × ℤ ∥₂ + ≃ ∥ S¹ ∥₂ × ∥ ℤ ∥₂ + ≃ ℤ +-} +coHom1S1≃ℤ : GroupIso (coHomGr 1 (S₊ 1)) ℤGroup +coHom1S1≃ℤ = theIso + where + F = Iso.fun S¹→S¹≡S¹×ℤ + F⁻ = Iso.inv S¹→S¹≡S¹×ℤ + + theIso : GroupIso (coHomGr 1 (S₊ 1)) ℤGroup + fun (fst theIso) = sRec isSetℤ (λ f → snd (F f)) + inv (fst theIso) a = ∣ (F⁻ (base , a)) ∣₂ + rightInv (fst theIso) a = cong snd (Iso.rightInv S¹→S¹≡S¹×ℤ (base , a)) + leftInv (fst theIso) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f → cong ((sRec isSetSetTrunc ∣_∣₂) + ∘ sRec isSetSetTrunc λ x → ∣ F⁻ (x , (snd (F f))) ∣₂) + (Iso.inv PathIdTrunc₀Iso (isConnectedS¹ (fst (F f)))) + ∙ cong ∣_∣₂ (Iso.leftInv S¹→S¹≡S¹×ℤ f) + snd theIso = + makeIsGroupHom + (coHomPointedElimS¹2 _ (λ _ _ → isSetℤ _ _) + λ p q → (λ i → winding (guy ∣ base ∣ (cong S¹map (help p q i)))) + ∙∙ (λ i → winding (guy ∣ base ∣ (congFunct S¹map p q i))) + ∙∙ winding-hom (guy ∣ base ∣ (cong S¹map p)) + (guy ∣ base ∣ (cong S¹map q))) + where + guy = basechange2⁻ ∘ S¹map + help : (p q : Path (coHomK 1) ∣ base ∣ ∣ base ∣) → cong₂ _+ₖ_ p q ≡ p ∙ q + help p q = cong₂Funct _+ₖ_ p q ∙ (λ i → cong (λ x → rUnitₖ 1 x i) p ∙ cong (λ x → lUnitₖ 1 x i) q) + +---------------------------- Hⁿ(Sⁿ) ≅ ℤ , n ≥ 1 ------------------- +Hⁿ-Sⁿ≅ℤ : (n : ℕ) → GroupIso (coHomGr (suc n) (S₊ (suc n))) ℤGroup +Hⁿ-Sⁿ≅ℤ zero = coHom1S1≃ℤ +Hⁿ-Sⁿ≅ℤ (suc n) = suspensionAx-Sn n n □ Hⁿ-Sⁿ≅ℤ n + +-------------- Hⁿ(Sᵐ) ≅ ℤ for n , m ≥ 1 with n ≠ m ---------------- +Hⁿ-Sᵐ≅0 : (n m : ℕ) → ¬ (n ≡ m) → GroupIso (coHomGr (suc n) (S₊ (suc m))) UnitGroup +Hⁿ-Sᵐ≅0 zero zero pf = ⊥-rec (pf refl) +Hⁿ-Sᵐ≅0 zero (suc m) pf = H¹-Sⁿ≅0 m +Hⁿ-Sᵐ≅0 (suc n) zero pf = Hⁿ-S¹≅0 n +Hⁿ-Sᵐ≅0 (suc n) (suc m) pf = suspensionAx-Sn n m + □ Hⁿ-Sᵐ≅0 n m λ p → pf (cong suc p) + + +-- Test functions +private + to₁ : coHom 1 S¹ → ℤ + to₁ = Iso.fun (fst (Hⁿ-Sⁿ≅ℤ 0)) + + to₂ : coHom 2 (S₊ 2) → ℤ + to₂ = Iso.fun (fst (Hⁿ-Sⁿ≅ℤ 1)) + + to₃ : coHom 3 (S₊ 3) → ℤ + to₃ = Iso.fun (fst (Hⁿ-Sⁿ≅ℤ 2)) + + + from₁ : ℤ → coHom 1 S¹ + from₁ = Iso.inv (fst (Hⁿ-Sⁿ≅ℤ 0)) + + from₂ : ℤ → coHom 2 (S₊ 2) + from₂ = Iso.inv (fst (Hⁿ-Sⁿ≅ℤ 1)) + + from₃ : ℤ → coHom 3 (S₊ 3) + from₃ = Iso.inv (fst (Hⁿ-Sⁿ≅ℤ 2)) + +{- +Strangely, the following won't compute +test₀ : to₂ (from₂ 1 +ₕ from₂ 1) ≡ 2 +test₀ = refl +However, the same example works for S¹ ∨ S¹ ∨ S², where the functions are essentially the same as here +(although somewhat more complicated). + +But with our new strange addition +'ₕ, it computes just fine: + +test₀ : to₂ (from₂ 1 +'ₕ from₂ 1) ≡ 2 +test₀ = refl + +-} diff --git a/Cubical/ZCohomology/Groups/Torus.agda b/Cubical/ZCohomology/Groups/Torus.agda new file mode 100644 index 000000000..cf9a73bca --- /dev/null +++ b/Cubical/ZCohomology/Groups/Torus.agda @@ -0,0 +1,309 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Groups.Torus where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Connected +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.Groups.Prelims +open import Cubical.ZCohomology.RingStructure.CupProduct + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Sigma +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +Comm to +ℤ-comm ; +Assoc to +ℤ-assoc) +open import Cubical.Data.Nat +open import Cubical.Data.Unit +open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Bool to BoolGroup ; Unit to UnitGroup) + +open import Cubical.HITs.Pushout +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) hiding (map) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim2 to pElim2 ; ∣_∣ to ∣_∣₁) hiding (map) +open import Cubical.HITs.Nullification +open import Cubical.HITs.Truncation renaming (elim to trElim ; elim2 to trElim2 ; map to trMap ; rec to trRec) +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Loopspace + +open import Cubical.ZCohomology.Groups.WedgeOfSpheres + renaming (to₂ to to₂-∨ ; from₂ to from₂-∨ ; from₁ to from₁-∨ ; to₁ to to₁-∨) hiding (to₀ ; from₀) +open import Cubical.Data.Empty +open import Cubical.HITs.Wedge + +open import Cubical.Relation.Nullary + +open IsGroupHom +open Iso + +-- The following section contains stengthened induction principles for cohomology groups of T². They are particularly useful for showing that +-- that some Isos are morphisms. They make things type-check faster, but should probably not be used for computations. + +-- We first need some functions +elimFunT² : (n : ℕ) (p q : typ (Ω (coHomK-ptd (suc n)))) + → Square q q p p + → S¹ × S¹ → coHomK (suc n) +elimFunT² n p q P (base , base) = ∣ ptSn (suc n) ∣ +elimFunT² n p q P (base , loop i) = q i +elimFunT² n p q P (loop i , base) = p i +elimFunT² n p q P (loop i , loop j) = P i j + +elimFunT²' : (n : ℕ) → Square (refl {ℓ-zero} {coHomK (suc n)} {∣ ptSn (suc n) ∣}) refl refl refl + → S¹ × S¹ → coHomK (suc n) +elimFunT²' n P (x , base) = ∣ ptSn (suc n) ∣ +elimFunT²' n P (base , loop j) = ∣ ptSn (suc n) ∣ +elimFunT²' n P (loop i , loop j) = P i j + +elimFunT²'≡elimFunT² : (n : ℕ) → (P : _) → elimFunT²' n P ≡ elimFunT² n refl refl P +elimFunT²'≡elimFunT² n P i (base , base) = ∣ ptSn (suc n) ∣ +elimFunT²'≡elimFunT² n P i (base , loop k) = ∣ ptSn (suc n) ∣ +elimFunT²'≡elimFunT² n P i (loop j , base) = ∣ ptSn (suc n) ∣ +elimFunT²'≡elimFunT² n P i (loop j , loop k) = P j k + +{- +The first induction principle says that when proving a proposition for some x : Hⁿ(T²), n ≥ 1, it suffices to show that it holds for +(elimFunT² p q P) for any paths p q : ΩKₙ, and square P : Square q q p p. This is useful because elimFunT² p q P (base , base) recudes to 0 +-} + +coHomPointedElimT² : ∀ {ℓ} (n : ℕ) {B : coHom (suc n) (S¹ × S¹) → Type ℓ} + → ((x : coHom (suc n) (S¹ × S¹)) → isProp (B x)) + → ((p q : _) (P : _) → B ∣ elimFunT² n p q P ∣₂) + → (x : coHom (suc n) (S¹ × S¹)) → B x +coHomPointedElimT² n {B = B} isprop indp = + coHomPointedElim _ (base , base) isprop + λ f fId → subst B (cong ∣_∣₂ (funExt (λ { (base , base) → sym fId + ; (base , loop i) j → helper f fId i1 i (~ j) + ; (loop i , base) j → helper f fId i i1 (~ j) + ; (loop i , loop j) k → helper f fId i j (~ k)}))) + (indp (λ i → helper f fId i i1 i1) + (λ i → helper f fId i1 i i1) + λ i j → helper f fId i j i1) + where + helper : (f : S¹ × S¹ → coHomK (suc n)) → f (base , base) ≡ ∣ ptSn (suc n) ∣ + → I → I → I → coHomK (suc n) + helper f fId i j k = + hfill (λ k → λ {(i = i0) → doubleCompPath-filler (sym fId) (cong f (λ i → (base , loop i))) fId k j + ; (i = i1) → doubleCompPath-filler (sym fId) (cong f (λ i → (base , loop i))) fId k j + ; (j = i0) → doubleCompPath-filler (sym fId) (cong f (λ i → (loop i , base))) fId k i + ; (j = i1) → doubleCompPath-filler (sym fId) (cong f (λ i → (loop i , base))) fId k i}) + (inS (f ((loop i) , (loop j)))) + k + +private + lem : ∀ {ℓ} (n : ℕ) {B : coHom (2 + n) (S¹ × S¹) → Type ℓ} + → ((P : _) → B ∣ elimFunT² (suc n) refl refl P ∣₂) + → (p : _) → (refl ≡ p) + → (q : _) → (refl ≡ q) + → (P : _) + → B ∣ elimFunT² (suc n) p q P ∣₂ + lem n {B = B} elimP p = + J (λ p _ → (q : _) → (refl ≡ q) + → (P : _) + → B ∣ elimFunT² (suc n) p q P ∣₂) + λ q → + J (λ q _ → (P : _) → B ∣ elimFunT² (suc n) refl q P ∣₂) + elimP + +{- When working with Hⁿ(T²) , n ≥ 2, we are, in the case described above, allowed to assume that any f : Hⁿ(T²) is + elimFunT² n refl refl P -} +coHomPointedElimT²' : ∀ {ℓ} (n : ℕ) {B : coHom (2 + n) (S¹ × S¹) → Type ℓ} + → ((x : coHom (2 + n) (S¹ × S¹)) → isProp (B x)) + → ((P : _) → B ∣ elimFunT² (suc n) refl refl P ∣₂) + → (x : coHom (2 + n) (S¹ × S¹)) → B x +coHomPointedElimT²' n {B = B} prop ind = + coHomPointedElimT² (suc n) prop + λ p q P → trRec (isProp→isOfHLevelSuc n (prop _)) + (λ p-refl → trRec (isProp→isOfHLevelSuc n (prop _)) + (λ q-refl → lem n {B = B} ind p (sym p-refl) q (sym q-refl) P) + (isConnectedPath _ (isConnectedPathKn (suc n) _ _) q refl .fst)) + (isConnectedPath _ (isConnectedPathKn (suc n) _ _) p refl .fst) + +{- A slight variation of the above which gives definitional equalities for all points (x , base) -} +private + coHomPointedElimT²'' : ∀ {ℓ} (n : ℕ) {B : coHom (2 + n) (S¹ × S¹) → Type ℓ} + → ((x : coHom (2 + n) (S¹ × S¹)) → isProp (B x)) + → ((P : _) → B ∣ elimFunT²' (suc n) P ∣₂) + → (x : coHom (2 + n) (S¹ × S¹)) → B x + coHomPointedElimT²'' n {B = B} prop ind = + coHomPointedElimT²' n prop λ P → subst (λ x → B ∣ x ∣₂) + (elimFunT²'≡elimFunT² (suc n) P) (ind P) + +--------- H⁰(T²) ------------ +H⁰-T²≅ℤ : GroupIso (coHomGr 0 (S₊ 1 × S₊ 1)) ℤGroup +H⁰-T²≅ℤ = + H⁰-connected (base , base) + λ (a , b) → pRec isPropPropTrunc + (λ id1 → pRec isPropPropTrunc + (λ id2 → ∣ ΣPathP (id1 , id2) ∣₁) + (Sn-connected 0 b) ) + (Sn-connected 0 a) + +--------- H¹(T²) ------------------------------- + +H¹-T²≅ℤ×ℤ : GroupIso (coHomGr 1 ((S₊ 1) × (S₊ 1))) (DirProd ℤGroup ℤGroup) +H¹-T²≅ℤ×ℤ = theIso □ GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ 0) (H⁰-Sⁿ≅ℤ 0) + where + typIso : Iso _ _ + typIso = setTruncIso (curryIso ⋄ codomainIso S1→K₁≡S1×ℤ ⋄ toProdIso) + ⋄ setTruncOfProdIso + + theIso : GroupIso _ _ + fst theIso = typIso + snd theIso = + makeIsGroupHom + (coHomPointedElimT² _ (λ _ → isPropΠ λ _ → isSet× isSetSetTrunc isSetSetTrunc _ _) + λ pf qf Pf → + coHomPointedElimT² _ (λ _ → isSet× isSetSetTrunc isSetSetTrunc _ _) + λ pg qg Pg i → ∣ funExt (helperFst pf qf pg qg Pg Pf) i ∣₂ + , ∣ funExt (helperSnd pf qf pg qg Pg Pf) i ∣₂) + where + module _ (pf qf pg qg : 0ₖ 1 ≡ 0ₖ 1) (Pg : Square qg qg pg pg) (Pf : Square qf qf pf pf) where + helperFst : (x : S¹) + → fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pf qf Pf (x , y) +ₖ elimFunT² 0 pg qg Pg (x , y)) .fst + ≡ fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pf qf Pf (x , y)) .fst + +ₖ fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pg qg Pg (x , y)) .fst + helperFst base = refl + helperFst (loop i) j = loopLem j i + where + loopLem : cong (λ x → fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pf qf Pf (x , y) +ₖ elimFunT² 0 pg qg Pg (x , y)) .fst) loop + ≡ cong (λ x → fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pf qf Pf (x , y)) .fst + +ₖ fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pg qg Pg (x , y)) .fst) loop + loopLem = (λ i j → S¹map-id (pf j +ₖ pg j) i) + ∙ (λ i j → S¹map-id (pf j) (~ i) +ₖ S¹map-id (pg j) (~ i)) + + helperSnd : (x : S¹) + → fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pf qf Pf (x , y) +ₖ elimFunT² 0 pg qg Pg (x , y)) .snd + ≡ fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pf qf Pf (x , y)) .snd +ℤ fun S1→K₁≡S1×ℤ (λ y → elimFunT² 0 pg qg Pg (x , y)) .snd + helperSnd = + toPropElim (λ _ → isSetℤ _ _) + ((λ i → winding (basechange2⁻ base λ j → S¹map (∙≡+₁ qf qg (~ i) j))) + ∙∙ cong (winding ∘ basechange2⁻ base) (congFunct S¹map qf qg) + ∙∙ (cong winding (basechange2⁻-morph base (cong S¹map qf) (cong S¹map qg)) + ∙ winding-hom (basechange2⁻ base (cong S¹map qf)) (basechange2⁻ base (cong S¹map qg)))) + +----------------------- H²(T²) ------------------------------ + +H²-T²≅ℤ : GroupIso (coHomGr 2 (S₊ 1 × S₊ 1)) ℤGroup +H²-T²≅ℤ = compGroupIso helper2 (Hⁿ-Sⁿ≅ℤ 0) + where + helper : Iso (∥ ((a : S¹) → coHomK 2) ∥₂ × ∥ ((a : S¹) → coHomK 1) ∥₂) (coHom 1 S¹) + inv helper s = 0ₕ _ , s + fun helper = snd + leftInv helper _ = + ΣPathP (isOfHLevelSuc 0 (isOfHLevelRetractFromIso 0 (fst (Hⁿ-S¹≅0 0)) (isContrUnit)) _ _ + , refl) + rightInv helper _ = refl + theIso : Iso (coHom 2 (S¹ × S¹)) (coHom 1 S¹) + theIso = setTruncIso (curryIso ⋄ codomainIso S1→K2≡K2×K1 ⋄ toProdIso) + ⋄ setTruncOfProdIso + ⋄ helper + + helper2 : GroupIso (coHomGr 2 (S¹ × S¹)) (coHomGr 1 S¹) + helper2 .fst = theIso + helper2 .snd = makeIsGroupHom ( + coHomPointedElimT²'' 0 (λ _ → isPropΠ λ _ → isSetSetTrunc _ _) + λ P → coHomPointedElimT²'' 0 (λ _ → isSetSetTrunc _ _) + λ Q → ((λ i → ∣ (λ a → ΩKn+1→Kn 1 (sym (rCancel≡refl 0 i) + ∙∙ cong (λ x → (elimFunT²' 1 P (a , x) +ₖ elimFunT²' 1 Q (a , x)) -ₖ ∣ north ∣) loop + ∙∙ rCancel≡refl 0 i)) ∣₂)) + ∙∙ (λ i → ∣ (λ a → ΩKn+1→Kn 1 (rUnit (cong (λ x → rUnitₖ 2 (elimFunT²' 1 P (a , x) +ₖ elimFunT²' 1 Q (a , x)) i) loop) (~ i))) ∣₂) + ∙∙ (λ i → ∣ (λ a → ΩKn+1→Kn 1 (∙≡+₂ 0 (cong (λ x → elimFunT²' 1 P (a , x)) loop) (cong (λ x → elimFunT²' 1 Q (a , x)) loop) (~ i))) ∣₂) + ∙∙ (λ i → ∣ (λ a → ΩKn+1→Kn-hom 1 (cong (λ x → elimFunT²' 1 P (a , x)) loop) (cong (λ x → elimFunT²' 1 Q (a , x)) loop) i) ∣₂) + ∙∙ (λ i → ∣ ((λ a → ΩKn+1→Kn 1 (rUnit (cong (λ x → rUnitₖ 2 (elimFunT²' 1 P (a , x)) (~ i)) loop) i) + +ₖ ΩKn+1→Kn 1 (rUnit (cong (λ x → rUnitₖ 2 (elimFunT²' 1 Q (a , x)) (~ i)) loop) i))) ∣₂) + ∙ (λ i → ∣ ((λ a → ΩKn+1→Kn 1 (sym (rCancel≡refl 0 (~ i)) + ∙∙ cong (λ x → elimFunT²' 1 P (a , x) +ₖ ∣ north ∣) loop + ∙∙ rCancel≡refl 0 (~ i)) + +ₖ ΩKn+1→Kn 1 (sym (rCancel≡refl 0 (~ i)) + ∙∙ cong (λ x → elimFunT²' 1 Q (a , x) +ₖ ∣ north ∣) loop + ∙∙ rCancel≡refl 0 (~ i)))) ∣₂)) + +private + to₂ : coHom 2 (S₊ 1 × S₊ 1) → ℤ + to₂ = fun (fst H²-T²≅ℤ) + + from₂ : ℤ → coHom 2 (S₊ 1 × S₊ 1) + from₂ = inv (fst H²-T²≅ℤ) + + to₁ : coHom 1 (S₊ 1 × S₊ 1) → ℤ × ℤ + to₁ = fun (fst H¹-T²≅ℤ×ℤ) + + from₁ : ℤ × ℤ → coHom 1 (S₊ 1 × S₊ 1) + from₁ = inv (fst H¹-T²≅ℤ×ℤ) + + to₀ : coHom 0 (S₊ 1 × S₊ 1) → ℤ + to₀ = fun (fst H⁰-T²≅ℤ) + + from₀ : ℤ → coHom 0 (S₊ 1 × S₊ 1) + from₀ = inv (fst H⁰-T²≅ℤ) + +{- +-- Compute fast: +test : to₁ (from₁ (0 , 1) +ₕ from₁ (1 , 0)) ≡ (1 , 1) +test = refl + +test2 : to₁ (from₁ (5 , 1) +ₕ from₁ (-2 , 3)) ≡ (3 , 4) +test2 = refl + +-- Compute pretty fast +test3 : to₂ (from₂ 1) ≡ 1 +test3 = refl + +test4 : to₂ (from₂ 2) ≡ 2 +test4 = refl + +test5 : to₂ (from₂ 3) ≡ 3 +test5 = refl + +-- Compute, but slower +test6 : to₂ (from₂ 0 +ₕ from₂ 0) ≡ 0 +test6 = refl + +test6 : to₂ (from₂ 0 +ₕ from₂ 1) ≡ 1 +test6 = refl + +-- Does not compute +test7 : to₂ (from₂ 1 +ₕ from₂ 0) ≡ 1 +test7 = refl +-} + + +-- Proof (by computation) that T² ≠ S² ∨ S¹ ∨ S¹ +private + hasTrivial⌣₁ : ∀ {ℓ} (A : Type ℓ) → Type ℓ + hasTrivial⌣₁ A = (x y : coHom 1 A) → x ⌣ y ≡ 0ₕ 2 + + hasTrivial⌣₁S²∨S¹∨S¹ : hasTrivial⌣₁ S²⋁S¹⋁S¹ + hasTrivial⌣₁S²∨S¹∨S¹ x y = + x ⌣ y ≡⟨ cong₂ _⌣_ (sym (leftInv (fst (H¹-S²⋁S¹⋁S¹)) x)) (sym (leftInv (fst (H¹-S²⋁S¹⋁S¹)) y)) ⟩ + from₁-∨ (to₁-∨ x) ⌣ from₁-∨ (to₁-∨ y) ≡⟨ sym (leftInv (fst (H²-S²⋁S¹⋁S¹)) (from₁-∨ (to₁-∨ x) ⌣ from₁-∨ (to₁-∨ y))) ⟩ + from₂-∨ (to₂-∨ (from₁-∨ (to₁-∨ x) ⌣ from₁-∨ (to₁-∨ y))) ≡⟨ refl ⟩ -- holds by computation (even with open terms in the context)! + from₂-∨ 0 ≡⟨ hom1g (snd ℤGroup) from₂-∨ (snd (coHomGr 2 S²⋁S¹⋁S¹)) + ((invGroupEquiv (GroupIso→GroupEquiv H²-S²⋁S¹⋁S¹)) .snd .pres·) ⟩ + 0ₕ 2 ∎ + + 1≠0 : ¬ (Path ℤ 1 0) + 1≠0 p = posNotnegsuc _ _ (cong predℤ p) + + ¬hasTrivial⌣₁T² : ¬ (hasTrivial⌣₁ (S¹ × S¹)) + ¬hasTrivial⌣₁T² p = 1≠0 1=0 + where + 1=0 : pos 1 ≡ pos 0 + 1=0 = + 1 ≡⟨ refl ⟩ -- holds by computation! + to₂ (from₁ (0 , 1) ⌣ from₁ (1 , 0)) ≡⟨ cong to₂ (p (from₁ (0 , 1)) (from₁ (1 , 0))) ⟩ + 0 ∎ + +T²≠S²⋁S¹⋁S¹ : ¬ S¹ × S¹ ≡ S²⋁S¹⋁S¹ +T²≠S²⋁S¹⋁S¹ p = ¬hasTrivial⌣₁T² (subst hasTrivial⌣₁ (sym p) hasTrivial⌣₁S²∨S¹∨S¹) diff --git a/Cubical/ZCohomology/Groups/Unit.agda b/Cubical/ZCohomology/Groups/Unit.agda new file mode 100644 index 000000000..6c0350f73 --- /dev/null +++ b/Cubical/ZCohomology/Groups/Unit.agda @@ -0,0 +1,85 @@ +{-# OPTIONS --safe #-} +module Cubical.ZCohomology.Groups.Unit where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.Nullification +open import Cubical.Data.Int hiding (ℤ ; _+_ ; +Comm) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation +open import Cubical.Homotopy.Connected +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.Algebra.Group renaming (Unit to UnitGroup) + +-- H⁰(Unit) +open IsGroupHom +open Iso + +H⁰-Unit≅ℤ : GroupIso (coHomGr 0 Unit) ℤ +fun (fst H⁰-Unit≅ℤ) = sRec isSetℤ (λ f → f tt) +inv (fst H⁰-Unit≅ℤ) a = ∣ (λ _ → a) ∣₂ +rightInv (fst H⁰-Unit≅ℤ) _ = refl +leftInv (fst H⁰-Unit≅ℤ) = sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ a → refl +snd H⁰-Unit≅ℤ = makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 isSetℤ _ _) λ a b → refl) + +{- Hⁿ(Unit) for n ≥ 1 -} +isContrHⁿ-Unit : (n : ℕ) → isContr (coHom (suc n) Unit) +isContrHⁿ-Unit n = subst isContr (λ i → ∥ UnitToTypePath (coHomK (suc n)) (~ i) ∥₂) (helper' n) + where + helper' : (n : ℕ) → isContr (∥ coHomK (suc n) ∥₂) + helper' n = + subst isContr + ((isoToPath (truncOfTruncIso {A = S₊ (1 + n)} 2 (1 + n))) + ∙∙ sym propTrunc≡Trunc2 + ∙∙ λ i → ∥ hLevelTrunc (suc (+-comm n 2 i)) (S₊ (1 + n)) ∥₂) + (isConnectedSubtr 2 (helper2 n .fst) + (subst (λ x → isConnected x (S₊ (suc n))) (sym (helper2 n .snd)) (sphereConnected (suc n))) ) + where + helper2 : (n : ℕ) → Σ[ m ∈ ℕ ] m + 2 ≡ 2 + n + helper2 zero = 0 , refl + helper2 (suc n) = (suc n) , λ i → suc (+-comm n 2 i) + +Hⁿ-Unit≅0 : (n : ℕ) → GroupIso (coHomGr (suc n) Unit) UnitGroup +fun (fst (Hⁿ-Unit≅0 n)) _ = _ +inv (fst (Hⁿ-Unit≅0 n)) _ = 0ₕ (suc n) +rightInv (fst (Hⁿ-Unit≅0 n)) _ = refl +leftInv (fst (Hⁿ-Unit≅0 n)) _ = isOfHLevelSuc 0 (isContrHⁿ-Unit n) _ _ +snd (Hⁿ-Unit≅0 n) = makeIsGroupHom λ _ _ → refl + + +{- Hⁿ for arbitrary contractible types -} +private + Hⁿ-contrTypeIso : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A + → Iso (coHom (suc n) A) (coHom (suc n) Unit) + Hⁿ-contrTypeIso n contr = compIso (setTruncIso (isContr→Iso2 contr)) + (setTruncIso (invIso (isContr→Iso2 isContrUnit))) + +Hⁿ-contrType≅0 : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A + → GroupIso (coHomGr (suc n) A) UnitGroup +fun (fst (Hⁿ-contrType≅0 n contr)) _ = _ +inv (fst (Hⁿ-contrType≅0 n contr)) _ = 0ₕ (suc n) +rightInv (fst (Hⁿ-contrType≅0 n contr)) _ = refl +leftInv (fst (Hⁿ-contrType≅0 {A = A} n contr)) _ = isOfHLevelSuc 0 helper _ _ + where + helper : isContr (coHom (suc n) A) + helper = (Iso.inv (Hⁿ-contrTypeIso n contr) (0ₕ (suc n))) + , λ y → cong (Iso.inv (Hⁿ-contrTypeIso n contr)) + (isOfHLevelSuc 0 (isContrHⁿ-Unit n) (0ₕ (suc n)) (Iso.fun (Hⁿ-contrTypeIso n contr) y)) + ∙ Iso.leftInv (Hⁿ-contrTypeIso n contr) y +snd (Hⁿ-contrType≅0 n contr) = makeIsGroupHom λ _ _ → refl + +isContr-HⁿRed-Unit : (n : ℕ) → isContr (coHomRed n (Unit , tt)) +fst (isContr-HⁿRed-Unit n) = 0ₕ∙ _ +snd (isContr-HⁿRed-Unit n) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP (funExt (λ _ → sym p) + , λ i j → p (~ i ∨ j)))} diff --git a/Cubical/ZCohomology/Groups/Wedge.agda b/Cubical/ZCohomology/Groups/Wedge.agda new file mode 100644 index 000000000..afc92655c --- /dev/null +++ b/Cubical/ZCohomology/Groups/Wedge.agda @@ -0,0 +1,243 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Groups.Wedge where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws renaming (assoc to assoc∙) +open import Cubical.HITs.Wedge +open import Cubical.Data.Int hiding (_+_) +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; ∣_∣ to ∣_∣₁) +open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; elim2 to trElim2) +open import Cubical.Data.Nat +open import Cubical.Algebra.Group + +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Isomorphism +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.Foundations.Equiv + +open IsGroupHom +open Iso +{- +This module proves that Hⁿ(A ⋁ B) ≅ Hⁿ(A) × Hⁿ(B) for n ≥ 1 directly (rather than by means of Mayer-Vietoris). +It also proves that Ĥⁿ(A ⋁ B) ≅ Ĥ⁰(A) × Ĥ⁰(B) (reduced groups) + +Proof sketch for n ≥ 1: + +Any ∣ f ∣₂ ∈ Hⁿ(A ⋁ B) is uniquely characterised by a pair of functions + f₁ : A → Kₙ + f₂ : B → Kₙ +together with a path + p : f₁ (pt A) ≡ f₂ (pt B) + +The map F : Hⁿ(A ⋁ B) → Hⁿ(A) × Hⁿ(B) simply forgets about p, i.e.: + F(∣ f ∣₂) := (∣ f₁ ∣₂ , ∣ f₂ ∣₂) + +The construction of its inverse is defined by + F⁻(∣ f₁ ∣₂ , ∣ f₂ ∣₂) := ∣ f₁∨f₂ ∣₂ + where + f₁∨f₂ : A ⋁ B → Kₙ is defined inductively by + f₁∨f₂ (inl x) := f₁ x + f₂ (pt B) + f₁∨f₂ (inr x) := f₁ (pt B) + f₂ x + cong f₁∨f₂ (push tt) := refl + (this is the map wedgeFun⁻ below) + +The fact that F and F⁻ cancel out is a proposition and we may thus assume for any + ∣ f ∣₂ ∈ Hⁿ(A ⋁ B) and its corresponding f₁ that + f₁ (pt A) = f₂ (pt B) = 0 (*) + and + f (inl (pt A)) = 0 (**) + +The fact that F(F⁻(∣ f₁ ∣₂ , ∣ f₂ ∣₂)) = ∣ f₁ ∣₂ , ∣ f₂ ∣₂) follows immediately from (*) + +The other way is slightly trickier. We need to construct paths + Pₗ(x) : f (inl (x)) + f (inr (pt B)) ---> f (inl (x)) + Pᵣ(x) : f (inl (pt A)) + f (inr (x)) ---> f (inr (x)) + +Together with a filler of the following square + + cong f (push tt) + -----------------> + ^ ^ + | | + | | +Pₗ(pr A) | | Pᵣ(pt B) + | | + | | + | | + -----------------> + refl + +The square is filled by first constructing Pₗ by + f (inl (x)) + f (inr (pt B)) ---[cong f (push tt)⁻¹]---> + f (inl (x)) + f (inl (pt A)) ---[(**)]---> + f (inl (x)) + 0 ---[right-unit]---> + f (inl (x)) + +and then Pᵣ by + f (inl (pt A)) + f (inr (x)) ---[(**)⁻¹]---> + 0 + f (inr (x)) ---[left-unit]---> + f (inr (x)) + +and finally by using the fact that the group laws for Kₙ are refl at its base point. +-} + +module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') where + + private + wedgeFun⁻ : ∀ n → (f : typ A → coHomK (suc n)) (g : typ B → coHomK (suc n)) + → ((A ⋁ B) → coHomK (suc n)) + wedgeFun⁻ n f g (inl x) = f x +ₖ g (pt B) + wedgeFun⁻ n f g (inr x) = f (pt A) +ₖ g x + wedgeFun⁻ n f g (push a i) = f (pt A) +ₖ g (pt B) + + Hⁿ-⋁ : (n : ℕ) → GroupIso (coHomGr (suc n) (A ⋁ B)) (×coHomGr (suc n) (typ A) (typ B)) + fun (fst (Hⁿ-⋁ zero)) = + sElim (λ _ → isSet× isSetSetTrunc isSetSetTrunc) + λ f → ∣ (λ x → f (inl x)) ∣₂ , ∣ (λ x → f (inr x)) ∣₂ + inv (fst (Hⁿ-⋁ zero)) = + uncurry (sElim2 (λ _ _ → isSetSetTrunc) + λ f g → ∣ wedgeFun⁻ 0 f g ∣₂) + rightInv (fst (Hⁿ-⋁ zero)) = + uncurry + (coHomPointedElim _ (pt A) (λ _ → isPropΠ λ _ → isSet× isSetSetTrunc isSetSetTrunc _ _) + λ f fId → coHomPointedElim _ (pt B) (λ _ → isSet× isSetSetTrunc isSetSetTrunc _ _) + λ g gId → ΣPathP ((cong ∣_∣₂ (funExt (λ x → cong (f x +ₖ_) gId ∙ rUnitₖ 1 (f x)))) + , cong ∣_∣₂ (funExt (λ x → cong (_+ₖ g x) fId ∙ lUnitₖ 1 (g x))))) + leftInv (fst (Hⁿ-⋁ zero)) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (λ f → pRec (isSetSetTrunc _ _) + (λ fId → cong ∣_∣₂ (sym fId)) + (helper f _ refl)) + where + helper : (f : A ⋁ B → coHomK 1) (x : coHomK 1) + → f (inl (pt A)) ≡ x + → ∥ f ≡ wedgeFun⁻ 0 (λ x → f (inl x)) (λ x → f (inr x)) ∥ + helper f = + trElim (λ _ → isProp→isOfHLevelSuc 2 (isPropΠ λ _ → isPropPropTrunc)) + (sphereElim 0 (λ _ → isPropΠ λ _ → isPropPropTrunc) + λ inlId → ∣ funExt (λ { (inl x) → sym (rUnitₖ 1 (f (inl x))) + ∙∙ cong ((f (inl x)) +ₖ_) (sym inlId) + ∙∙ cong ((f (inl x)) +ₖ_) (cong f (push tt)) + ; (inr x) → sym (lUnitₖ 1 (f (inr x))) + ∙ cong (_+ₖ (f (inr x))) (sym inlId) + ; (push tt i) j → helper2 (f (inl (pt A))) (sym (inlId)) + (f (inr (pt B))) (cong f (push tt)) j i} ) ∣₁) + where + helper2 : (x : coHomK 1) (r : ∣ base ∣ ≡ x) (y : coHomK 1) (p : x ≡ y) + → PathP (λ j → ((sym (rUnitₖ 1 x) ∙∙ cong (x +ₖ_) r ∙∙ cong (x +ₖ_) p)) j + ≡ (sym (lUnitₖ 1 y) ∙ cong (_+ₖ y) r) j) + p refl + helper2 x = J (λ x r → (y : coHomK 1) (p : x ≡ y) + → PathP (λ j → ((sym (rUnitₖ 1 x) ∙∙ cong (x +ₖ_) r ∙∙ cong (x +ₖ_) p)) j + ≡ (sym (lUnitₖ 1 y) ∙ cong (_+ₖ y) r) j) + p refl) + λ y → J (λ y p → PathP (λ j → ((sym (rUnitₖ 1 ∣ base ∣) ∙∙ refl ∙∙ cong (∣ base ∣ +ₖ_) p)) j + ≡ (sym (lUnitₖ 1 y) ∙ refl) j) + p refl) + λ i _ → (refl ∙ (λ _ → 0ₖ 1)) i + snd (Hⁿ-⋁ zero) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ _ _ → refl) + fun (fst (Hⁿ-⋁ (suc n))) = + sElim (λ _ → isSet× isSetSetTrunc isSetSetTrunc) + λ f → ∣ (λ x → f (inl x)) ∣₂ , ∣ (λ x → f (inr x)) ∣₂ + inv (fst (Hⁿ-⋁ (suc n))) = + uncurry (sElim2 (λ _ _ → isSetSetTrunc) + λ f g → ∣ wedgeFun⁻ (suc n) f g ∣₂) + rightInv (fst (Hⁿ-⋁ (suc n))) = + uncurry + (coHomPointedElim _ (pt A) (λ _ → isPropΠ λ _ → isSet× isSetSetTrunc isSetSetTrunc _ _) + λ f fId → coHomPointedElim _ (pt B) (λ _ → isSet× isSetSetTrunc isSetSetTrunc _ _) + λ g gId → ΣPathP ((cong ∣_∣₂ (funExt (λ x → cong (f x +ₖ_) gId ∙ rUnitₖ (2 + n) (f x)))) + , cong ∣_∣₂ (funExt (λ x → cong (_+ₖ g x) fId ∙ lUnitₖ (2 + n) (g x))))) + leftInv (fst (Hⁿ-⋁ (suc n))) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (λ f → pRec (isSetSetTrunc _ _) + (λ fId → cong ∣_∣₂ (sym fId)) + (helper f _ refl)) + where + helper : (f : A ⋁ B → coHomK (2 + n)) (x : coHomK (2 + n)) + → f (inl (pt A)) ≡ x + → ∥ f ≡ wedgeFun⁻ (suc n) (λ x → f (inl x)) (λ x → f (inr x)) ∥ + helper f = + trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isPropΠ λ _ → isPropPropTrunc)) + (sphereToPropElim (suc n) (λ _ → isPropΠ λ _ → isPropPropTrunc) + λ inlId → (∣ funExt (λ { (inl x) → sym (rUnitₖ (2 + n) (f (inl x))) + ∙∙ cong ((f (inl x)) +ₖ_) (sym inlId) + ∙∙ cong ((f (inl x)) +ₖ_) (cong f (push tt)) + ; (inr x) → sym (lUnitₖ (2 + n) (f (inr x))) + ∙ cong (_+ₖ (f (inr x))) (sym inlId) + ; (push tt i) j → helper2 (f (inl (pt A))) (sym (inlId)) + (f (inr (pt B))) (cong f (push tt)) j i}) ∣₁)) + where + helper2 : (x : coHomK (2 + n)) (r : ∣ north ∣ ≡ x) (y : coHomK (2 + n)) (p : x ≡ y) + → PathP (λ j → ((sym (rUnitₖ (2 + n) x) ∙∙ cong (x +ₖ_) r ∙∙ cong (x +ₖ_) p)) j + ≡ (sym (lUnitₖ (2 + n) y) ∙ cong (_+ₖ y) r) j) + p refl + helper2 x = J (λ x r → (y : coHomK (2 + n)) (p : x ≡ y) + → PathP (λ j → ((sym (rUnitₖ (2 + n) x) ∙∙ cong (x +ₖ_) r ∙∙ cong (x +ₖ_) p)) j + ≡ (sym (lUnitₖ (2 + n) y) ∙ cong (_+ₖ y) r) j) + p refl) + λ y → J (λ y p → PathP (λ j → ((sym (rUnitₖ (2 + n) ∣ north ∣) ∙∙ refl ∙∙ cong (∣ north ∣ +ₖ_) p)) j + ≡ (sym (lUnitₖ (2 + n) y) ∙ refl) j) + p refl) + λ i j → ((λ _ → ∣ north ∣) ∙ refl) i + snd (Hⁿ-⋁ (suc n)) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ _ _ → refl) + + H⁰Red-⋁ : GroupIso (coHomRedGrDir 0 (A ⋁ B , inl (pt A))) + (DirProd (coHomRedGrDir 0 A) (coHomRedGrDir 0 B)) + fun (fst H⁰Red-⋁) = + sRec (isSet× isSetSetTrunc isSetSetTrunc) + λ {(f , p) → ∣ (f ∘ inl) , p ∣₂ + , ∣ (f ∘ inr) , cong f (sym (push tt)) ∙ p ∣₂} + inv (fst H⁰Red-⋁) = + uncurry (sRec2 isSetSetTrunc + λ {(f , p) (g , q) → ∣ (λ {(inl a) → f a + ; (inr b) → g b + ; (push tt i) → (p ∙ sym q) i}) + , p ∣₂}) + rightInv (fst H⁰Red-⋁) = + uncurry + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ {(_ , _) (_ , _) → ΣPathP (cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl) + , cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl))}) + leftInv (fst H⁰Red-⋁) = + sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) + (funExt λ {(inl a) → refl + ; (inr b) → refl + ; (push tt i) j → (cong (p ∙_) (symDistr (cong f (sym (push tt))) p) + ∙∙ assoc∙ p (sym p) (cong f (push tt)) + ∙∙ cong (_∙ (cong f (push tt))) (rCancel p) + ∙ sym (lUnit (cong f (push tt)))) j i}))} + -- Alt. use isOfHLevel→isOfHLevelDep + snd H⁰Red-⋁ = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ {(f , p) (g , q) → ΣPathP (cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl) + , cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl))}) + + wedgeConnected : ((x : typ A) → ∥ pt A ≡ x ∥) → ((x : typ B) → ∥ pt B ≡ x ∥) → (x : A ⋁ B) → ∥ inl (pt A) ≡ x ∥ + wedgeConnected conA conB = + PushoutToProp (λ _ → isPropPropTrunc) + (λ a → pRec isPropPropTrunc (λ p → ∣ cong inl p ∣₁) (conA a)) + λ b → pRec isPropPropTrunc (λ p → ∣ push tt ∙ cong inr p ∣₁) (conB b) diff --git a/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda new file mode 100644 index 000000000..94e21aede --- /dev/null +++ b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda @@ -0,0 +1,134 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Groups.WedgeOfSpheres where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.Data.Nat +open import Cubical.Data.Int renaming (_+_ to _ℤ+_) + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Groups.Connected +open import Cubical.ZCohomology.RingStructure.CupProduct + +open import Cubical.HITs.Sn +open import Cubical.HITs.S1 +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.Pushout +open import Cubical.HITs.Truncation renaming (elim to trElim) hiding (map ; elim2) +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim) + +open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Bool to BoolGroup ; Unit to UnitGroup) + + +S¹⋁S¹ : Type₀ +S¹⋁S¹ = S₊∙ 1 ⋁ S₊∙ 1 + +S²⋁S¹⋁S¹ : Type₀ +S²⋁S¹⋁S¹ = S₊∙ 2 ⋁ (S¹⋁S¹ , inl base) + +------------- H⁰(S¹⋁S¹) ------------ +H⁰-S¹⋁S¹ : GroupIso (coHomGr 0 S¹⋁S¹) ℤGroup +H⁰-S¹⋁S¹ = H⁰-connected (inl base) (wedgeConnected _ _ (Sn-connected 0) (Sn-connected 0)) + +------------- H¹(S¹⋁S¹) ------------ +H¹-S¹⋁S¹ : GroupIso (coHomGr 1 S¹⋁S¹) (DirProd ℤGroup ℤGroup) +H¹-S¹⋁S¹ = (Hⁿ-⋁ _ _ 0) □ GroupIsoDirProd coHom1S1≃ℤ coHom1S1≃ℤ + +------------- H⁰(S²⋁S¹⋁S¹) --------- +H⁰-S²⋁S¹⋁S¹ : GroupIso (coHomGr 0 S²⋁S¹⋁S¹) ℤGroup +H⁰-S²⋁S¹⋁S¹ = H⁰-connected (inl north) + (wedgeConnected _ _ + (Sn-connected 1) + (wedgeConnected _ _ + (Sn-connected 0) + (Sn-connected 0))) + +------------- H¹(S²⋁S¹⋁S¹) --------- +H¹-S²⋁S¹⋁S¹ : GroupIso (coHomGr 1 S²⋁S¹⋁S¹) (DirProd ℤGroup ℤGroup) +H¹-S²⋁S¹⋁S¹ = + Hⁿ-⋁ (S₊∙ 2) (S¹⋁S¹ , inl base) 0 + □ GroupIsoDirProd (H¹-Sⁿ≅0 0) H¹-S¹⋁S¹ + □ lUnitGroupIso + +------------- H²(S²⋁S¹⋁S¹) --------- + +H²-S²⋁S¹⋁S¹ : GroupIso (coHomGr 2 S²⋁S¹⋁S¹) ℤGroup +H²-S²⋁S¹⋁S¹ = + compGroupIso + (Hⁿ-⋁ _ _ 1) + (GroupIsoDirProd {B = UnitGroup} + (Hⁿ-Sⁿ≅ℤ 1) + ((Hⁿ-⋁ _ _ 1) □ GroupIsoDirProd (Hⁿ-S¹≅0 0) (Hⁿ-S¹≅0 0) □ rUnitGroupIso) + □ rUnitGroupIso) + +open Iso + +to₂ : coHom 2 S²⋁S¹⋁S¹ → ℤ +to₂ = fun (fst H²-S²⋁S¹⋁S¹) +from₂ : ℤ → coHom 2 S²⋁S¹⋁S¹ +from₂ = inv (fst H²-S²⋁S¹⋁S¹) + +to₁ : coHom 1 S²⋁S¹⋁S¹ → ℤ × ℤ +to₁ = fun (fst H¹-S²⋁S¹⋁S¹) +from₁ : ℤ × ℤ → coHom 1 S²⋁S¹⋁S¹ +from₁ = inv (fst H¹-S²⋁S¹⋁S¹) + +to₀ : coHom 0 S²⋁S¹⋁S¹ → ℤ +to₀ = fun (fst H⁰-S²⋁S¹⋁S¹) +from₀ : ℤ → coHom 0 S²⋁S¹⋁S¹ +from₀ = inv (fst H⁰-S²⋁S¹⋁S¹) + +{- + +-- Compute pretty fast +test1 : to₁ (from₁ (1 , 0) +ₕ from₁ (0 , 1)) ≡ (1 , 1) +test1 = refl + +-- Computes, but only when computing some smaller numbers first +test2 : to₁ (from₁ (50 , 3) +ₕ from₁ (2 , -2)) ≡ (52 , 1) +test2 = refl + +test3 : to₂ (from₂ 0) ≡ 0 +test3 = refl + +test4 : to₂ (from₂ 3) ≡ 3 +test4 = refl + +test5 : to₂ (from₂ 1 +ₕ from₂ 1) ≡ 2 +test5 = refl +-} +{- + g : S²⋁S¹⋁S¹ → ∥ Susp S¹ ∥ 4 + g (inl x) = ∣ x ∣ + g (inr x) = 0ₖ _ + g (push a i) = 0ₖ _ + + G = ∣ g ∣₂ + +-- Does not compute: +test₀ : to₂ (G +ₕ G) ≡ 2 +test₀ = refl + +but this does: +test₀ : to₂ (G +'ₕ G) ≡ 2 +test₀ = refl +-} + + +-- Cup product is trivial +⌣-gen₁ : to₂ (from₁ (1 , 0) ⌣ from₁ (0 , 1)) ≡ 0 +⌣-gen₁ = refl + +-- Even better: +⌣-gen : (x y : ℤ × ℤ) → to₂ (from₁ x ⌣ from₁ y) ≡ 0 +⌣-gen x y = refl diff --git a/Cubical/ZCohomology/MayerVietorisUnreduced.agda b/Cubical/ZCohomology/MayerVietorisUnreduced.agda new file mode 100644 index 000000000..e14b3c812 --- /dev/null +++ b/Cubical/ZCohomology/MayerVietorisUnreduced.agda @@ -0,0 +1,249 @@ +{-# OPTIONS --safe #-} +module Cubical.ZCohomology.MayerVietorisUnreduced where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.Sigma +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.S1 +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₁ ; ∣_∣ to ∣_∣₁) +open import Cubical.Data.Nat +open import Cubical.Algebra.Group +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) + +open IsGroupHom + +module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) where + -- Proof from Brunerie 2016. + -- We first define the three morphisms involved: i, Δ and d. + + private + i* : (n : ℕ) → coHom n (Pushout f g) → coHom n A × coHom n B + i* n = sRec (isSet× isSetSetTrunc isSetSetTrunc) λ δ → ∣ (λ x → δ (inl x)) ∣₂ , ∣ (λ x → δ (inr x)) ∣₂ + + iIsHom : (n : ℕ) → IsGroupHom (coHomGr n (Pushout f g) .snd) (i* n) (×coHomGr n A B .snd) + iIsHom n = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) λ _ _ → refl) + + i : (n : ℕ) → GroupHom (coHomGr n (Pushout f g)) (×coHomGr n A B) + fst (i n) = i* n + snd (i n) = iIsHom n + + private + distrLem : (n : ℕ) (x y z w : coHomK n) → (x +[ n ]ₖ y) -[ n ]ₖ (z +[ n ]ₖ w) ≡ (x -[ n ]ₖ z) +[ n ]ₖ (y -[ n ]ₖ w) + distrLem n x y z w = cong (λ z → (x +[ n ]ₖ y) +[ n ]ₖ z) (-distrₖ n z w) + ∙∙ sym (assocₖ n x y ((-[ n ]ₖ z) +[ n ]ₖ (-[ n ]ₖ w))) + ∙∙ cong (λ y → x +[ n ]ₖ y) (commₖ n y ((-[ n ]ₖ z) +[ n ]ₖ (-[ n ]ₖ w)) ∙ sym (assocₖ n _ _ _)) + ∙∙ assocₖ n _ _ _ + ∙∙ cong (λ y → (x -[ n ]ₖ z) +[ n ]ₖ y) (commₖ n (-[ n ]ₖ w) y) + + Δ' : (n : ℕ) → coHom n A × coHom n B → coHom n C + Δ' n (α , β) = coHomFun n f α -[ n ]ₕ coHomFun n g β + + Δ'-isMorph : (n : ℕ) → IsGroupHom (×coHomGr n A B .snd) (Δ' n) (coHomGr n C .snd) + Δ'-isMorph n = + makeIsGroupHom + (prodElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _ ) + λ f' x1 g' x2 i → ∣ (λ x → distrLem n (f' (f x)) (g' (f x)) (x1 (g x)) (x2 (g x)) i) ∣₂) + + Δ : (n : ℕ) → GroupHom (×coHomGr n A B) (coHomGr n C) + fst (Δ n) = Δ' n + snd (Δ n) = Δ'-isMorph n + + d-pre : (n : ℕ) → (C → coHomK n) → Pushout f g → coHomK (suc n) + d-pre n γ (inl x) = 0ₖ (suc n) + d-pre n γ (inr x) = 0ₖ (suc n) + d-pre n γ (push a i) = Kn→ΩKn+1 n (γ a) i + + dHomHelper : (n : ℕ) (h l : C → coHomK n) (x : Pushout f g) + → d-pre n (λ x → h x +[ n ]ₖ l x) x ≡ d-pre n h x +[ suc n ]ₖ d-pre n l x + dHomHelper n h l (inl x) = sym (rUnitₖ (suc n) (0ₖ (suc n))) + dHomHelper n h l (inr x) = sym (lUnitₖ (suc n) (0ₖ (suc n))) + dHomHelper n h l (push a i) j = + hcomp (λ k → λ { (i = i0) → rUnitₖ (suc n) (0ₖ (suc n)) (~ j) + ; (i = i1) → lUnitₖ (suc n) (0ₖ (suc n)) (~ j) + ; (j = i0) → Kn→ΩKn+1-hom n (h a) (l a) (~ k) i + ; (j = i1) → cong₂Funct (λ x y → x +[ (suc n) ]ₖ y) (Kn→ΩKn+1 n (h a)) (Kn→ΩKn+1 n (l a)) (~ k) i }) + (hcomp (λ k → λ { (i = i0) → rUnitₖ (suc n) (0ₖ (suc n)) (~ j) + ; (i = i1) → lUnitₖ (suc n) (Kn→ΩKn+1 n (l a) k) (~ j)}) + (hcomp (λ k → λ { (i = i0) → rUnitₖ (suc n) (0ₖ (suc n)) (~ j) + ; (i = i1) → lUnitₖ≡rUnitₖ (suc n) (~ k) (~ j) + ; (j = i0) → Kn→ΩKn+1 n (h a) i + ; (j = i1) → (Kn→ΩKn+1 n (h a) i) +[ (suc n) ]ₖ coHom-pt (suc n)}) + (rUnitₖ (suc n) (Kn→ΩKn+1 n (h a) i) (~ j)))) + + dIsHom : (n : ℕ) → IsGroupHom (coHomGr n C .snd) (sRec isSetSetTrunc λ a → ∣ d-pre n a ∣₂) (coHomGr (suc n) (Pushout f g) .snd) + dIsHom n = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f g i → ∣ funExt (λ x → dHomHelper n f g x) i ∣₂) + + d : (n : ℕ) → GroupHom (coHomGr n C) (coHomGr (suc n) (Pushout f g)) + fst (d n) = sRec isSetSetTrunc λ a → ∣ d-pre n a ∣₂ + snd (d n) = dIsHom n + + -- The long exact sequence + Im-d⊂Ker-i : (n : ℕ) (x : ⟨ (coHomGr (suc n) (Pushout f g)) ⟩) + → isInIm (d n) x + → isInKer (i (suc n)) x + Im-d⊂Ker-i n = sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ a → pRec (isOfHLevelPath' 1 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + (sigmaElim (λ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ δ b i → sRec (isSet× isSetSetTrunc isSetSetTrunc) + (λ δ → ∣ (λ x → δ (inl x)) ∣₂ , ∣ (λ x → δ (inr x)) ∣₂ ) (b (~ i))) + + + Ker-i⊂Im-d : (n : ℕ) (x : ⟨ coHomGr (suc n) (Pushout f g) ⟩) + → isInKer (i (suc n)) x + → isInIm (d n) x + Ker-i⊂Im-d n = + sElim (λ _ → isSetΠ λ _ → isProp→isSet isPropPropTrunc) + λ a p → pRec {A = (λ x → a (inl x)) ≡ λ _ → 0ₖ (suc n)} + (isProp→ isPropPropTrunc) + (λ p1 → pRec isPropPropTrunc λ p2 → ∣ ∣ (λ c → ΩKn+1→Kn n (sym (cong (λ F → F (f c)) p1) + ∙∙ cong a (push c) + ∙∙ cong (λ F → F (g c)) p2)) ∣₂ + , cong ∣_∣₂ (funExt (λ δ → helper a p1 p2 δ)) ∣₁) + (Iso.fun PathIdTrunc₀Iso (cong fst p)) + (Iso.fun PathIdTrunc₀Iso (cong snd p)) + + where + helper : (F : (Pushout f g) → coHomK (suc n)) + (p1 : Path (_ → coHomK (suc n)) (λ a₁ → F (inl a₁)) (λ _ → coHom-pt (suc n))) + (p2 : Path (_ → coHomK (suc n)) (λ a₁ → F (inr a₁)) (λ _ → coHom-pt (suc n))) + → (δ : Pushout f g) + → d-pre n (λ c → ΩKn+1→Kn n ((λ i₁ → p1 (~ i₁) (f c)) + ∙∙ cong F (push c) + ∙∙ cong (λ F → F (g c)) p2)) δ + ≡ F δ + helper F p1 p2 (inl x) = sym (cong (λ f → f x) p1) + helper F p1 p2 (inr x) = sym (cong (λ f → f x) p2) + helper F p1 p2 (push a i) j = + hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) + ; (i = i1) → p2 (~ j) (g a) + ; (j = i0) → Iso.rightInv (Iso-Kn-ΩKn+1 n) ((λ i₁ → p1 (~ i₁) (f a)) + ∙∙ cong F (push a) + ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i + ; (j = i1) → F (push a i)}) + (doubleCompPath-filler (sym (cong (λ F → F (f a)) p1)) (cong F (push a)) (cong (λ F → F (g a)) p2) (~ j) i) + + Im-i⊂Ker-Δ : (n : ℕ) (x : ⟨ ×coHomGr n A B ⟩) + → isInIm (i n) x + → isInKer (Δ n) x + Im-i⊂Ker-Δ n (Fa , Fb) = + sElim {B = λ Fa → (Fb : _) → isInIm (i n) (Fa , Fb) + → isInKer (Δ n) (Fa , Fb)} + (λ _ → isSetΠ2 λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (λ Fa → sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ Fb → pRec (isSetSetTrunc _ _) + (sigmaElim (λ x → isProp→isSet (isSetSetTrunc _ _)) + λ Fd p → helper n Fa Fb Fd p)) + Fa + Fb + where + helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) (Fd : (Pushout f g) → coHomK n) + → (fst (i n) ∣ Fd ∣₂ ≡ (∣ Fa ∣₂ , ∣ Fb ∣₂)) + → (fst (Δ n)) (∣ Fa ∣₂ , ∣ Fb ∣₂) ≡ 0ₕ n + helper n Fa Fb Fd p = cong (fst (Δ n)) (sym p) + ∙∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₂ -[ n ]ₕ ∣ (λ x → Fd (push x (~ i))) ∣₂ ) + ∙∙ rCancelₕ n ∣ (λ x → Fd (inl (f x))) ∣₂ + + Ker-Δ⊂Im-i : (n : ℕ) (a : ⟨ ×coHomGr n A B ⟩) + → isInKer (Δ n) a + → isInIm (i n) a + Ker-Δ⊂Im-i n = prodElim (λ _ → isSetΠ (λ _ → isProp→isSet isPropPropTrunc)) + (λ Fa Fb p → pRec isPropPropTrunc + (λ q → ∣ ∣ helpFun Fa Fb q ∣₂ , refl ∣₁) + (helper Fa Fb p)) + where + helper : (Fa : A → coHomK n) (Fb : B → coHomK n) + → fst (Δ n) (∣ Fa ∣₂ , ∣ Fb ∣₂) ≡ 0ₕ n + → ∥ Path (_ → _) (λ c → Fa (f c)) (λ c → Fb (g c)) ∥₁ + helper Fa Fb p = Iso.fun PathIdTrunc₀Iso + (sym (cong ∣_∣₂ (funExt (λ x → sym (assocₖ n _ _ _) + ∙∙ cong (λ y → Fa (f x) +[ n ]ₖ y) (lCancelₖ n (Fb (g x))) + ∙∙ rUnitₖ n (Fa (f x))))) + ∙∙ cong (λ x → x +[ n ]ₕ ∣ (λ c → Fb (g c)) ∣₂) p + ∙∙ lUnitₕ n _) + + helpFun : (Fa : A → coHomK n) (Fb : B → coHomK n) + → ((λ c → Fa (f c)) ≡ (λ c → Fb (g c))) + → Pushout f g → coHomK n + helpFun Fa Fb p (inl x) = Fa x + helpFun Fa Fb p (inr x) = Fb x + helpFun Fa Fb p (push a i) = p i a + + private + distrHelper : (n : ℕ) (p q : _) + → ΩKn+1→Kn n p +[ n ]ₖ (-[ n ]ₖ ΩKn+1→Kn n q) ≡ ΩKn+1→Kn n (p ∙ sym q) + distrHelper n p q = cong (λ x → ΩKn+1→Kn n p +[ n ]ₖ x) helper ∙ sym (ΩKn+1→Kn-hom n _ _) + where + helper : -[ n ]ₖ ΩKn+1→Kn n q ≡ ΩKn+1→Kn n (sym q) + helper = + sym (rUnitₖ n _) + ∙∙ cong (λ x → (-[ n ]ₖ (ΩKn+1→Kn n q)) +[ n ]ₖ x) (sym helper2) + ∙∙ (assocₖ n _ _ _ ∙∙ cong (λ x → x +[ n ]ₖ (ΩKn+1→Kn n (sym q))) (lCancelₖ n _) ∙∙ lUnitₖ n _) + where + helper2 : ΩKn+1→Kn n q +[ n ]ₖ (ΩKn+1→Kn n (sym q)) ≡ coHom-pt n + helper2 = sym (ΩKn+1→Kn-hom n q (sym q)) ∙∙ cong (ΩKn+1→Kn n) (rCancel q) ∙∙ ΩKn+1→Kn-refl n + + Ker-d⊂Im-Δ : (n : ℕ) (a : coHom n C) + → isInKer (d n) a + → isInIm (Δ n) a + Ker-d⊂Im-Δ n = + sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 isPropPropTrunc) + λ Fc p → pRec isPropPropTrunc (λ p → ∣ (∣ (λ a → ΩKn+1→Kn n (cong (λ f → f (inl a)) p)) ∣₂ , + ∣ (λ b → ΩKn+1→Kn n (cong (λ f → f (inr b)) p)) ∣₂) , + Iso.inv (PathIdTrunc₀Iso) ∣ funExt (λ c → helper2 Fc p c) ∣₁ ∣₁) + (Iso.fun (PathIdTrunc₀Iso) p) + + where + + helper2 : (Fc : C → coHomK n) + (p : d-pre n Fc ≡ (λ _ → coHom-pt (suc n))) (c : C) + → ΩKn+1→Kn n (λ i₁ → p i₁ (inl (f c))) -[ n ]ₖ (ΩKn+1→Kn n (λ i₁ → p i₁ (inr (g c)))) ≡ Fc c + helper2 Fc p c = distrHelper n _ _ ∙∙ cong (ΩKn+1→Kn n) helper3 ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) (Fc c) + where + helper3 : (λ i₁ → p i₁ (inl (f c))) ∙ sym (λ i₁ → p i₁ (inr (g c))) ≡ Kn→ΩKn+1 n (Fc c) + helper3 = cong ((λ i₁ → p i₁ (inl (f c))) ∙_) (lUnit _) ∙ sym (PathP→compPathR (cong (λ f → cong f (push c)) p)) + + Im-Δ⊂Ker-d : (n : ℕ) (a : coHom n C) + → isInIm (Δ n) a + → isInKer (d n) a + Im-Δ⊂Ker-d n = + sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ Fc → pRec (isOfHLevelPath' 1 isSetSetTrunc _ _) + (sigmaProdElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ Fa Fb p → pRec (isOfHLevelPath' 1 isSetSetTrunc _ _) + (λ q → ((λ i → fst (d n) ∣ (q (~ i)) ∣₂) ∙ dΔ-Id n Fa Fb)) + (Iso.fun (PathIdTrunc₀Iso) p)) + + where + d-preLeftId : (n : ℕ) (Fa : A → coHomK n)(d : (Pushout f g)) + → d-pre n (Fa ∘ f) d ≡ 0ₖ (suc n) + d-preLeftId n Fa (inl x) = Kn→ΩKn+1 n (Fa x) + d-preLeftId n Fa (inr x) = refl + d-preLeftId n Fa (push a i) j = Kn→ΩKn+1 n (Fa (f a)) (j ∨ i) + + d-preRightId : (n : ℕ) (Fb : B → coHomK n) (d : (Pushout f g)) + → d-pre n (Fb ∘ g) d ≡ 0ₖ (suc n) + d-preRightId n Fb (inl x) = refl + d-preRightId n Fb (inr x) = sym (Kn→ΩKn+1 n (Fb x)) + d-preRightId n Fb (push a i) j = Kn→ΩKn+1 n (Fb (g a)) (~ j ∧ i) + + dΔ-Id : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) + → fst (d n) (fst (Δ n) (∣ Fa ∣₂ , ∣ Fb ∣₂)) ≡ 0ₕ (suc n) + dΔ-Id n Fa Fb = -distrLemma n (suc n) (d n) ∣ Fa ∘ f ∣₂ ∣ Fb ∘ g ∣₂ + ∙∙ (λ i → ∣ (λ x → d-preLeftId n Fa x i) ∣₂ -[ (suc n) ]ₕ ∣ (λ x → d-preRightId n Fb x i) ∣₂) + ∙∙ rCancelₕ (suc n) (0ₕ (suc n)) diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 5529101fb..e2392008f 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -1,63 +1,654 @@ -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.ZCohomology.Properties where +{- +This module contains: +1. direct proofs of connectedness of Kn and ΩKn +2. Induction principles for cohomology groups of pointed types +3. Equivalence between cohomology of A and reduced cohomology of (A + 1) +4. Equivalence between cohomology and reduced cohomology for dimension ≥ 1 +5. Encode-decode proof of Kₙ ≃ ΩKₙ₊₁ and proofs that this equivalence + and its inverse are morphisms +6. A proof of coHomGr ≅ coHomGrΩ +7. A locked (non-reducing) version of Kₙ ≃ ΩKₙ₊₁ +8. Some HLevel lemmas used for the cup product +-} + + open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure + +open import Cubical.HITs.S1 hiding (encode ; decode ; _·_) open import Cubical.HITs.Sn open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Everything open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Data.NatMinusTwo.Base renaming (-1+_ to -1+₋₂_ ; 1+_ to 1+₋₂_) -open import Cubical.Data.NatMinusOne.Base -open import Cubical.Data.Empty -open import Cubical.Data.Sigma -open import Cubical.Data.Prod.Base +open import Cubical.Foundations.GroupoidLaws renaming (assoc to assoc∙) +open import Cubical.Foundations.Univalence open import Cubical.HITs.Susp -open import Cubical.HITs.SetTruncation -open import Cubical.HITs.Nullification -open import Cubical.Data.Int +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; isSetSetTrunc to §) +open import Cubical.Data.Int renaming (_+_ to _ℤ+_) hiding (-_) open import Cubical.Data.Nat -open import Cubical.HITs.Truncation +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; map2 to trMap2; rec to trRec ; elim3 to trElim3) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Algebra.Group hiding (Unit ; ℤ) +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Data.Sum.Base hiding (map) +open import Cubical.Functions.Morphism +open import Cubical.Data.Sigma + +open Iso renaming (inv to inv') -open import Cubical.HITs.Pushout -open import Cubical.Data.Sum.Base -open import Cubical.Data.HomotopyGroup -open import Cubical.Data.Bool private variable ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' + +------------------- Connectedness --------------------- +is2ConnectedKn : (n : ℕ) → isConnected 2 (coHomK (suc n)) +is2ConnectedKn zero = ∣ ∣ base ∣ ∣ + , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isOfHLevelTrunc 2)) _ _) + (toPropElim (λ _ → isOfHLevelTrunc 2 _ _) refl)) +is2ConnectedKn (suc n) = ∣ ∣ north ∣ ∣ + , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isOfHLevelTrunc 2 _ _)) + (suspToPropElim (ptSn (suc n)) (λ _ → isOfHLevelTrunc 2 _ _) refl)) + +isConnectedKn : (n : ℕ) → isConnected (2 + n) (coHomK (suc n)) +isConnectedKn n = isOfHLevelRetractFromIso 0 (invIso (truncOfTruncIso (2 + n) 1)) (sphereConnected (suc n)) + +-- direct proof of connectedness of ΩKₙ₊₁ not relying on the equivalence ∥ a ≡ b ∥ₙ ≃ (∣ a ∣ₙ₊₁ ≡ ∣ b ∣ₙ₊₁) +isConnectedPathKn : (n : ℕ) (x y : (coHomK (suc n))) → isConnected (suc n) (x ≡ y) +isConnectedPathKn n = + trElim (λ _ → isProp→isOfHLevelSuc (2 + n) (isPropΠ λ _ → isPropIsContr)) + (sphereElim _ (λ _ → isProp→isOfHLevelSuc n (isPropΠ λ _ → isPropIsContr)) + λ y → isContrRetractOfConstFun + {B = (hLevelTrunc (suc n) (ptSn (suc n) ≡ ptSn (suc n)))} ∣ refl ∣ + (fun⁻ n y + , trElim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + (J (λ y p → fun⁻ n y _ ≡ _) (funExt⁻ (fun⁻Id n) ∣ refl ∣)))) + where + fun⁻ : (n : ℕ) → (y : coHomK (suc n)) → + hLevelTrunc (suc n) (ptSn (suc n) ≡ ptSn (suc n)) + → hLevelTrunc (suc n) (∣ ptSn (suc n) ∣ ≡ y) + fun⁻ n = + trElim (λ _ → isOfHLevelΠ (3 + n) λ _ → isOfHLevelSuc (2 + n) (isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n)))) + (sphereElim n (λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelTrunc (suc n)) λ _ → ∣ refl ∣) + + fun⁻Id : (n : ℕ) → fun⁻ n ∣ ptSn (suc n) ∣ ≡ λ _ → ∣ refl ∣ + fun⁻Id zero = refl + fun⁻Id (suc n) = refl + +------------------- +-- Induction principles for cohomology groups (n ≥ 1) +-- If we want to show a proposition about some x : Hⁿ(A), it suffices to show it under the +-- assumption that x = ∣ f ∣₂ for some f : A → Kₙ and that f is pointed + +coHomPointedElim : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → Type ℓ'} + → ((x : coHom (suc n) A) → isProp (B x)) + → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) + → (x : coHom (suc n) A) → B x +coHomPointedElim {ℓ' = ℓ'} {A = A} n a isprop indp = + sElim (λ _ → isOfHLevelSuc 1 (isprop _)) + λ f → helper n isprop indp f (f a) refl + where + helper : (n : ℕ) {B : coHom (suc n) A → Type ℓ'} + → ((x : coHom (suc n) A) → isProp (B x)) + → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) + → (f : A → coHomK (suc n)) + → (x : coHomK (suc n)) + → f a ≡ x → B ∣ f ∣₂ + -- pattern matching a bit extra to avoid isOfHLevelPlus' + helper zero isprop ind f = + trElim (λ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ λ _ → isprop _)) + (toPropElim (λ _ → isPropΠ λ _ → isprop _) (ind f)) + helper (suc zero) isprop ind f = + trElim (λ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ λ _ → isprop _)) + (suspToPropElim base (λ _ → isPropΠ λ _ → isprop _) (ind f)) + helper (suc (suc zero)) isprop ind f = + trElim (λ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ λ _ → isprop _)) + (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) + helper (suc (suc (suc n))) isprop ind f = + trElim (λ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ λ _ → isprop _)) + (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) + + +coHomPointedElim2 : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} + → ((x y : coHom (suc n) A) → isProp (B x y)) + → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) + → (x y : coHom (suc n) A) → B x y +coHomPointedElim2 {ℓ' = ℓ'} {A = A} n a isprop indp = sElim2 (λ _ _ → isOfHLevelSuc 1 (isprop _ _)) + λ f g → helper n a isprop indp f g (f a) (g a) refl refl + where + helper : (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} + → ((x y : coHom (suc n) A) → isProp (B x y)) + → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) + → (f g : A → coHomK (suc n)) + → (x y : coHomK (suc n)) + → f a ≡ x → g a ≡ y + → B ∣ f ∣₂ ∣ g ∣₂ + helper zero a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ2 λ _ _ → isprop _ _)) + (toPropElim2 (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + helper (suc zero) a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ2 λ _ _ → isprop _ _)) + (suspToPropElim2 base (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + helper (suc (suc zero)) a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ2 λ _ _ → isprop _ _)) + (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + helper (suc (suc (suc n))) a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ2 λ _ _ → isprop _ _)) + (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + +coHomK-elim : ∀ {ℓ} (n : ℕ) {B : coHomK (suc n) → Type ℓ} + → ((x : _) → isOfHLevel (suc n) (B x)) + → B (0ₖ (suc n)) + → (x : _) → B x +coHomK-elim n {B = B } hlev b = + trElim (λ _ → isOfHLevelPlus {n = (suc n)} 2 (hlev _)) + (sphereElim _ (hlev ∘ ∣_∣) b) {- Equivalence between cohomology of A and reduced cohomology of (A + 1) -} coHomRed+1Equiv : (n : ℕ) → - (A : Set ℓ) → - (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) -coHomRed+1Equiv zero A i = ∥ helpLemma {C = (Int , pos 0)} i ∥₀ + (A : Type ℓ) → + (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) +coHomRed+1Equiv zero A i = ∥ helpLemma {C = (ℤ , pos 0)} i ∥₂ module coHomRed+1 where - helpLemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →* C))) + helpLemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →∙ C))) helpLemma {C = C} = isoToPath (iso map1 map2 (λ b → linvPf b) (λ _ → refl)) where - map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →* C)) - map1 f = map1' , refl module helpmap where + map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →∙ C)) + map1 f = map1' , refl + module helpmap where map1' : A ⊎ Unit → fst C map1' (inl x) = f x map1' (inr x) = pt C - map2 : ((((A ⊎ Unit) , inr (tt)) →* C)) → (A → typ C) + map2 : ((((A ⊎ Unit) , inr (tt)) →∙ C)) → (A → typ C) map2 (g , pf) x = g (inl x) - linvPf : (b :((((A ⊎ Unit) , inr (tt)) →* C))) → map1 (map2 b) ≡ b + linvPf : (b :((((A ⊎ Unit) , inr (tt)) →∙ C))) → map1 (map2 b) ≡ b linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) where helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x helper (inl x) = refl helper (inr tt) = sym snd +coHomRed+1Equiv (suc zero) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK 1 , ∣ base ∣)} i ∥₂ +coHomRed+1Equiv (suc (suc n)) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK (2 + n) , ∣ north ∣)} i ∥₂ + +Iso-coHom-coHomRed : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → Iso (coHomRed (suc n) A) (coHom (suc n) (typ A)) +fun (Iso-coHom-coHomRed {A = A , a} n) = map fst +inv' (Iso-coHom-coHomRed {A = A , a} n) = map λ f → (λ x → f x -ₖ f a) , rCancelₖ _ _ +rightInv (Iso-coHom-coHomRed {A = A , a} n) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ f → trRec (isProp→isOfHLevelSuc _ (§ _ _)) + (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → f x +ₖ y) (cong -ₖ_ p ∙ -0ₖ) ∙ rUnitₖ _ (f x))) + (Iso.fun (PathIdTruncIso (suc n)) (isContr→isProp (isConnectedKn n) ∣ f a ∣ ∣ 0ₖ _ ∣)) +leftInv (Iso-coHom-coHomRed {A = A , a} n) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP (((funExt λ x → (cong (λ y → f x -ₖ y) p + ∙∙ cong (λ y → f x +ₖ y) -0ₖ + ∙∙ rUnitₖ _ (f x)) ∙ refl)) + , helper n (f a) (sym p)))} + where + path : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ x) → _ + path n x p = (cong (λ y → x -ₖ y) (sym p) ∙∙ cong (λ y → x +ₖ y) -0ₖ ∙∙ rUnitₖ _ x) ∙ refl + + helper : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ x) + → PathP (λ i → path n x p i ≡ 0ₖ _) (rCancelₖ _ x) (sym p) + helper zero x = + J (λ x p → PathP (λ i → path 0 x p i ≡ 0ₖ _) + (rCancelₖ _ x) (sym p)) + λ i j → rUnit (rUnit (λ _ → 0ₖ 1) (~ j)) (~ j) i + helper (suc n) x = + J (λ x p → PathP (λ i → path (suc n) x p i ≡ 0ₖ _) (rCancelₖ _ x) (sym p)) + λ i j → rCancelₖ (suc (suc n)) (0ₖ (suc (suc n))) (~ i ∧ ~ j) + ++∙≡+ : (n : ℕ) {A : Pointed ℓ} (x y : coHomRed (suc n) A) + → Iso.fun (Iso-coHom-coHomRed n) (x +ₕ∙ y) + ≡ Iso.fun (Iso-coHom-coHomRed n) x +ₕ Iso.fun (Iso-coHom-coHomRed n) y + ++∙≡+ zero = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl ++∙≡+ (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl + +private + homhelp : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) (x y : coHom (suc n) (typ A)) + → Iso.inv (Iso-coHom-coHomRed {A = A} n) (x +ₕ y) + ≡ Iso.inv (Iso-coHom-coHomRed n) x +ₕ∙ Iso.inv (Iso-coHom-coHomRed n) y + homhelp n A = morphLemmas.isMorphInv _+ₕ∙_ _+ₕ_ + (Iso.fun (Iso-coHom-coHomRed n)) (+∙≡+ n) _ + (Iso.rightInv (Iso-coHom-coHomRed n)) (Iso.leftInv (Iso-coHom-coHomRed n)) + +coHomGr≅coHomRedGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) + → GroupEquiv (coHomRedGrDir (suc n) A) (coHomGr (suc n) (typ A)) +fst (coHomGr≅coHomRedGr n A) = isoToEquiv (Iso-coHom-coHomRed n) +snd (coHomGr≅coHomRedGr n A) = makeIsGroupHom (+∙≡+ n) + +coHomRedGroup : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → AbGroup ℓ +coHomRedGroup zero A = coHomRedGroupDir zero A +coHomRedGroup (suc n) A = + InducedAbGroup (coHomGroup (suc n) (typ A)) + _+ₕ∙_ + (isoToEquiv (invIso (Iso-coHom-coHomRed n))) + (homhelp n A) + +abstract + coHomGroup≡coHomRedGroup : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) + → coHomGroup (suc n) (typ A) ≡ coHomRedGroup (suc n) A + coHomGroup≡coHomRedGroup n A = + InducedAbGroupPath (coHomGroup (suc n) (typ A)) + _+ₕ∙_ + (isoToEquiv (invIso (Iso-coHom-coHomRed n))) + (homhelp n A) + +------------------- Kₙ ≃ ΩKₙ₊₁ --------------------- +-- This proof uses the encode-decode method rather than Freudenthal + +-- We define the map σ : Kₙ → ΩKₙ₊₁ and prove that it is a morphism +private + module _ (n : ℕ) where + σ : {n : ℕ} → coHomK (suc n) → Path (coHomK (2 + n)) ∣ north ∣ ∣ north ∣ + σ {n = n} = trRec (isOfHLevelTrunc (4 + n) _ _) + λ a → cong ∣_∣ (merid a ∙ sym (merid (ptSn (suc n)))) + + σ-hom-helper : ∀ {ℓ} {A : Type ℓ} {a : A} (p : a ≡ a) (r : refl ≡ p) + → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r + σ-hom-helper p = J (λ p r → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r) refl + + σ-hom : {n : ℕ} (x y : coHomK (suc n)) → σ (x +ₖ y) ≡ σ x ∙ σ y + σ-hom {n = zero} = + elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 4 _ _) _ _) + (wedgeconFun _ _ + (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) + (λ x → lUnit _ + ∙ cong (_∙ σ ∣ x ∣) (cong (cong ∣_∣) (sym (rCancel (merid base))))) + (λ y → cong σ (rUnitₖ 1 ∣ y ∣) + ∙∙ rUnit _ + ∙∙ cong (σ ∣ y ∣ ∙_) (cong (cong ∣_∣) (sym (rCancel (merid base))))) + (sym (σ-hom-helper (σ ∣ base ∣) (cong (cong ∣_∣) (sym (rCancel (merid base))))))) + σ-hom {n = suc n} = + elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) _ _) _ _) + (wedgeconFun _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev' n) _ _) + (λ x → lUnit _ + ∙ cong (_∙ σ ∣ x ∣) (cong (cong ∣_∣) (sym (rCancel (merid north))))) + (λ y → cong σ (rUnitₖ (2 + n) ∣ y ∣) + ∙∙ rUnit _ + ∙∙ cong (σ ∣ y ∣ ∙_) (cong (cong ∣_∣) (sym (rCancel (merid north))))) + (sym (σ-hom-helper (σ ∣ north ∣) (cong (cong ∣_∣) (sym (rCancel (merid north))))))) + + -- We will need to following lemma + σ-minusDistr : {n : ℕ} (x y : coHomK (suc n)) → σ (x -ₖ y) ≡ σ x ∙ sym (σ y) + σ-minusDistr {n = n} = + morphLemmas.distrMinus' + _+ₖ_ _∙_ + σ σ-hom ∣ (ptSn (suc n)) ∣ refl + -ₖ_ sym + (λ x → sym (lUnit x)) (λ x → sym (rUnit x)) + (rUnitₖ (suc n)) + (lCancelₖ (suc n)) rCancel + (assocₖ (suc n)) assoc∙ + (cong (cong ∣_∣) (rCancel (merid (ptSn (suc n))))) + + -- we define the code using addIso + Code : (n : ℕ) → coHomK (2 + n) → Type₀ + Code n x = (trRec {B = TypeOfHLevel ℓ-zero (3 + n)} (isOfHLevelTypeOfHLevel (3 + n)) + λ a → Code' a , hLevCode' a) x .fst + where + Code' : (S₊ (2 + n)) → Type₀ + Code' north = coHomK (suc n) + Code' south = coHomK (suc n) + Code' (merid a i) = isoToPath (addIso (suc n) ∣ a ∣) i + + hLevCode' : (x : S₊ (2 + n)) → isOfHLevel (3 + n) (Code' x) + hLevCode' = suspToPropElim (ptSn (suc n)) (λ _ → isPropIsOfHLevel (3 + n)) (isOfHLevelTrunc (3 + n)) + + symMeridLem : (n : ℕ) → (x : S₊ (suc n)) (y : coHomK (suc n)) + → subst (Code n) (cong ∣_∣ (sym (merid x))) y ≡ y -ₖ ∣ x ∣ + symMeridLem n x = trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) + (λ y → cong (_-ₖ ∣ x ∣) (transportRefl ∣ y ∣)) + + decode : {n : ℕ} (x : coHomK (2 + n)) → Code n x → ∣ north ∣ ≡ x + decode {n = n} = trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + decode-elim + where + north≡merid : (a : S₊ (suc n)) + → Path (coHomK (2 + n)) ∣ north ∣ ∣ north ∣ + ≡ (Path (coHomK (2 + n)) ∣ north ∣ ∣ south ∣) + north≡merid a i = Path (coHomK (2 + n)) ∣ north ∣ ∣ merid a i ∣ + + decode-elim : (a : S₊ (2 + n)) → Code n ∣ a ∣ → Path (coHomK (2 + n)) ∣ north ∣ ∣ a ∣ + decode-elim north = σ + decode-elim south = trRec (isOfHLevelTrunc (4 + n) _ _) + λ a → cong ∣_∣ (merid a) + decode-elim (merid a i) = + hcomp (λ k → λ { (i = i0) → σ + ; (i = i1) → mainPath a k}) + (funTypeTransp (Code n) (λ x → ∣ north ∣ ≡ x) (cong ∣_∣ (merid a)) σ i) + where + mainPath : (a : (S₊ (suc n))) → + transport (north≡merid a) ∘ σ ∘ transport (λ i → Code n ∣ merid a (~ i) ∣) + ≡ trRec (isOfHLevelTrunc (4 + n) _ _) λ a → cong ∣_∣ (merid a) + mainPath a = funExt (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (4 + n) _ _) _ _) + (λ x → (λ i → transport (north≡merid a) (σ (symMeridLem n a ∣ x ∣ i))) + ∙∙ cong (transport (north≡merid a)) (-distrHelp x) + ∙∙ (substAbove x))) + where + -distrHelp : (x : S₊ (suc n)) → σ (∣ x ∣ -ₖ ∣ a ∣) ≡ cong ∣_∣ (merid x) ∙ cong ∣_∣ (sym (merid a)) + -distrHelp x = + σ-minusDistr ∣ x ∣ ∣ a ∣ + ∙ (λ i → (cong ∣_∣ (compPath-filler (merid x) (λ j → merid (ptSn (suc n)) (~ j ∨ i)) (~ i))) + ∙ (cong ∣_∣ (sym (compPath-filler (merid a) (λ j → merid (ptSn (suc n)) (~ j ∨ i)) (~ i))))) + + substAbove : (x : S₊ (suc n)) → transport (north≡merid a) (cong ∣_∣ (merid x) ∙ cong ∣_∣ (sym (merid a))) + ≡ cong ∣_∣ (merid x) + substAbove x i = transp (λ j → north≡merid a (i ∨ j)) i + (compPath-filler (cong ∣_∣ (merid x)) (λ j → ∣ merid a (~ j ∨ i) ∣) (~ i)) + + encode : {n : ℕ} {x : coHomK (2 + n)} → Path (coHomK (2 + n)) ∣ north ∣ x → Code n x + encode {n = n} p = transport (cong (Code n) p) ∣ (ptSn (suc n)) ∣ + + decode-encode : {n : ℕ} {x : coHomK (2 + n)} (p : Path (coHomK (2 + n)) ∣ north ∣ x) + → decode _ (encode p) ≡ p + decode-encode {n = n} = + J (λ y p → decode _ (encode p) ≡ p) + (cong (decode ∣ north ∣) (transportRefl ∣ ptSn (suc n) ∣) + ∙ cong (cong ∣_∣) (rCancel (merid (ptSn (suc n))))) + +-- We define an addition operation on Code which we can use in order to show that encode is a +-- morphism (in a very loose sense) + hLevCode : {n : ℕ} (x : coHomK (2 + n)) → isOfHLevel (3 + n) (Code n x) + hLevCode {n = n} = + trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isPropIsOfHLevel (3 + n))) + (sphereToPropElim _ + (λ _ → (isPropIsOfHLevel (3 + n))) (isOfHLevelTrunc (3 + n))) + + Code-add' : {n : ℕ} (x : _) → Code n ∣ north ∣ → Code n ∣ x ∣ → Code n ∣ x ∣ + Code-add' {n = n} north = _+ₖ_ + Code-add' {n = n} south = _+ₖ_ + Code-add' {n = n} (merid a i) = helper n a i + where + help : (n : ℕ) → (x y a : S₊ (suc n)) + → transport (λ i → Code n ∣ north ∣ → Code n ∣ merid a i ∣ → Code n ∣ merid a i ∣) + (_+ₖ_) ∣ x ∣ ∣ y ∣ + ≡ ∣ x ∣ +ₖ ∣ y ∣ + help n x y a = + (λ i → transportRefl ((∣ transportRefl x i ∣ +ₖ (∣ transportRefl y i ∣ -ₖ ∣ a ∣)) +ₖ ∣ a ∣) i) + ∙∙ cong (_+ₖ ∣ a ∣) (assocₖ _ ∣ x ∣ ∣ y ∣ (-ₖ ∣ a ∣)) + ∙∙ sym (assocₖ _ (∣ x ∣ +ₖ ∣ y ∣) (-ₖ ∣ a ∣) ∣ a ∣) + ∙∙ cong ((∣ x ∣ +ₖ ∣ y ∣) +ₖ_) (lCancelₖ _ ∣ a ∣) + ∙∙ rUnitₖ _ _ + + helper : (n : ℕ) (a : S₊ (suc n)) + → PathP (λ i → Code n ∣ north ∣ → Code n ∣ merid a i ∣ → Code n ∣ merid a i ∣) _+ₖ_ _+ₖ_ + helper n a = + toPathP (funExt + (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelΠ (3 + n) (λ _ → isOfHLevelTrunc (3 + n))) _ _) + λ x → funExt + (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) + λ y → help n x y a))) + + Code-add : {n : ℕ} (x : _) → Code n ∣ north ∣ → Code n x → Code n x + Code-add {n = n} = + trElim (λ x → isOfHLevelΠ (4 + n) λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelSuc (3 + n) (hLevCode {n = n} x)) + Code-add' + + encode-hom : {n : ℕ} {x : _} (q : 0ₖ _ ≡ 0ₖ _) (p : 0ₖ _ ≡ x) + → encode (q ∙ p) ≡ Code-add {n = n} x (encode q) (encode p) + encode-hom {n = n} q = J (λ x p → encode (q ∙ p) ≡ Code-add {n = n} x (encode q) (encode p)) + (cong encode (sym (rUnit q)) + ∙∙ sym (rUnitₖ _ (encode q)) + ∙∙ cong (encode q +ₖ_) (cong ∣_∣ (sym (transportRefl _)))) + +stabSpheres : (n : ℕ) → Iso (coHomK (suc n)) (typ (Ω (coHomK-ptd (2 + n)))) +fun (stabSpheres n) = decode _ +inv' (stabSpheres n) = encode +rightInv (stabSpheres n) p = decode-encode p +leftInv (stabSpheres n) = + trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) + λ a → cong encode (congFunct ∣_∣ (merid a) (sym (merid (ptSn (suc n))))) + ∙∙ (λ i → transport (congFunct (Code n) (cong ∣_∣ (merid a)) + (cong ∣_∣ (sym (merid (ptSn (suc n))))) i) ∣ ptSn (suc n) ∣) + ∙∙ (substComposite (λ x → x) + (cong (Code n) (cong ∣_∣ (merid a))) + (cong (Code n) (cong ∣_∣ (sym (merid (ptSn (suc n)))))) ∣ ptSn (suc n) ∣ + ∙∙ cong (transport (λ i → Code n ∣ merid (ptSn (suc n)) (~ i) ∣)) + (transportRefl (∣ (ptSn (suc n)) ∣ +ₖ ∣ a ∣) ∙ lUnitₖ (suc n) ∣ a ∣) + ∙∙ symMeridLem n (ptSn (suc n)) ∣ a ∣ + ∙∙ cong (∣ a ∣ +ₖ_) -0ₖ + ∙∙ rUnitₖ (suc n) ∣ a ∣) + +Iso-Kn-ΩKn+1 : (n : HLevel) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) +Iso-Kn-ΩKn+1 zero = invIso (compIso (congIso (truncIdempotentIso _ isGroupoidS¹)) ΩS¹Isoℤ) +Iso-Kn-ΩKn+1 (suc n) = stabSpheres n + +Kn≃ΩKn+1 : {n : ℕ} → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) +Kn≃ΩKn+1 {n = n} = isoToEquiv (Iso-Kn-ΩKn+1 n) + +-- Some properties of the Iso +Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) +Kn→ΩKn+1 n = Iso.fun (Iso-Kn-ΩKn+1 n) + +ΩKn+1→Kn : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n +ΩKn+1→Kn n = Iso.inv (Iso-Kn-ΩKn+1 n) + +Kn→ΩKn+10ₖ : (n : ℕ) → Kn→ΩKn+1 n (0ₖ n) ≡ refl +Kn→ΩKn+10ₖ zero = sym (rUnit refl) +Kn→ΩKn+10ₖ (suc n) i j = ∣ (rCancel (merid (ptSn (suc n))) i j) ∣ + +ΩKn+1→Kn-refl : (n : ℕ) → ΩKn+1→Kn n refl ≡ 0ₖ n +ΩKn+1→Kn-refl zero = refl +ΩKn+1→Kn-refl (suc zero) = refl +ΩKn+1→Kn-refl (suc (suc n)) = refl + +Kn≃ΩKn+1∙ : {n : ℕ} → coHomK-ptd n ≡ (Ω (coHomK-ptd (suc n))) +Kn≃ΩKn+1∙ {n = n} = ua∙ Kn≃ΩKn+1 (Kn→ΩKn+10ₖ n) + +Kn→ΩKn+1-hom : (n : ℕ) (x y : coHomK n) → Kn→ΩKn+1 n (x +[ n ]ₖ y) ≡ Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y +Kn→ΩKn+1-hom zero x y = (λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) + (inS (∣ intLoop (x ℤ+ y) i ∣)) (~ j)) + ∙∙ (λ j i → ∣ intLoop-hom x y (~ j) i ∣) + ∙∙ (congFunct ∣_∣ (intLoop x) (intLoop y) + ∙ cong₂ _∙_ (λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) + (inS (∣ intLoop x i ∣)) j) + λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) + (inS (∣ intLoop y i ∣)) j) +Kn→ΩKn+1-hom (suc n) = σ-hom + +ΩKn+1→Kn-hom : (n : ℕ) (x y : Path (coHomK (suc n)) (0ₖ _) (0ₖ _)) + → ΩKn+1→Kn n (x ∙ y) ≡ ΩKn+1→Kn n x +[ n ]ₖ ΩKn+1→Kn n y +ΩKn+1→Kn-hom zero p q = + cong winding (congFunct (trRec isGroupoidS¹ (λ x → x)) p q) + ∙ winding-hom (cong (trRec isGroupoidS¹ (λ x → x)) p) (cong (trRec isGroupoidS¹ (λ x → x)) q) +ΩKn+1→Kn-hom (suc n) = encode-hom + +Kn→ΩKn+1-ₖ : (n : ℕ) (x : coHomK n) → Kn→ΩKn+1 n (-ₖ x) ≡ sym (Kn→ΩKn+1 n x) +Kn→ΩKn+1-ₖ n x = + lUnit _ + ∙∙ cong (_∙ Kn→ΩKn+1 n (-ₖ x)) (sym (lCancel _)) + ∙∙ sym (assoc∙ _ _ _) + ∙∙ cong (sym (Kn→ΩKn+1 n x) ∙_) help + ∙∙ sym (rUnit _) + where + help : Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n (-ₖ x) ≡ refl + help = sym (Kn→ΩKn+1-hom n x (-ₖ x)) ∙∙ cong (Kn→ΩKn+1 n) (rCancelₖ n x) ∙∙ Kn→ΩKn+10ₖ n + +isHomogeneousKn : (n : HLevel) → isHomogeneous (coHomK-ptd n) +isHomogeneousKn n = + subst isHomogeneous + (sym (ΣPathP (ua Kn≃ΩKn+1 , ua-gluePath _ (Kn→ΩKn+10ₖ n)))) + (isHomogeneousPath _ _) + +-- With the equivalence Kn≃ΩKn+1, we get that the two definitions of cohomology groups agree +open IsGroupHom + +coHom≅coHomΩ : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → GroupIso (coHomGr n A) (coHomGrΩ n A) +fun (fst (coHom≅coHomΩ n A)) = map λ f a → Kn→ΩKn+1 n (f a) +inv' (fst (coHom≅coHomΩ n A)) = map λ f a → ΩKn+1→Kn n (f a) +rightInv (fst (coHom≅coHomΩ n A)) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ f → cong ∣_∣₂ (funExt λ x → rightInv (Iso-Kn-ΩKn+1 n) (f x)) +leftInv (fst (coHom≅coHomΩ n A)) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ f → cong ∣_∣₂ (funExt λ x → leftInv (Iso-Kn-ΩKn+1 n) (f x)) +snd (coHom≅coHomΩ n A) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) + λ f g → cong ∣_∣₂ (funExt λ x → Kn→ΩKn+1-hom n (f x) (g x))) + +module lockedKnIso (key : Unit') where + Kn→ΩKn+1' : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) + Kn→ΩKn+1' n = lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) + + ΩKn+1→Kn' : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n + ΩKn+1→Kn' n = lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) + + ΩKn+1→Kn→ΩKn+1 : (n : ℕ) → (x : typ (Ω (coHomK-ptd (suc n)))) → Kn→ΩKn+1' n (ΩKn+1→Kn' n x) ≡ x + ΩKn+1→Kn→ΩKn+1 n x = pm key + where + pm : (key : Unit') → lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) (lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) x) ≡ x + pm unlock = Iso.rightInv (Iso-Kn-ΩKn+1 n) x + + Kn→ΩKn+1→Kn : (n : ℕ) → (x : coHomK n) → ΩKn+1→Kn' n (Kn→ΩKn+1' n x) ≡ x + Kn→ΩKn+1→Kn n x = pm key + where + pm : (key : Unit') → lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) (lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) x) ≡ x + pm unlock = Iso.leftInv (Iso-Kn-ΩKn+1 n) x + +-distrLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : ℕ) (f : GroupHom (coHomGr n A) (coHomGr m B)) + (x y : coHom n A) + → fst f (x -[ n ]ₕ y) ≡ fst f x -[ m ]ₕ fst f y +-distrLemma n m f' x y = sym (-cancelRₕ m (f y) (f (x -[ n ]ₕ y))) + ∙∙ cong (λ x → x -[ m ]ₕ f y) (sym (f' .snd .pres· (x -[ n ]ₕ y) y)) + ∙∙ cong (λ x → x -[ m ]ₕ f y) ( cong f (-+cancelₕ n _ _)) + where + f = fst f' + +-- HLevel stuff for cup product +isContr-↓∙ : (n : ℕ) → isContr (coHomK-ptd (suc n) →∙ coHomK-ptd n) +fst (isContr-↓∙ zero) = (λ _ → 0) , refl +snd (isContr-↓∙ zero) (f , p) = + Σ≡Prop (λ f → isSetℤ _ _) + (funExt (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 isSetℤ) _ _) + (toPropElim (λ _ → isSetℤ _ _) (sym p)))) +fst (isContr-↓∙ (suc n)) = (λ _ → 0ₖ _) , refl +fst (snd (isContr-↓∙ (suc n)) f i) x = + trElim {B = λ x → 0ₖ (suc n) ≡ fst f x} + (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (isOfHLevelTrunc (3 + n))) _ _) + (sphereElim _ (λ _ → isOfHLevelTrunc (3 + n) _ _) + (sym (snd f))) x i +snd (snd (isContr-↓∙ (suc n)) f i) j = snd f (~ i ∨ j) + +isContr-↓∙' : (n : ℕ) → isContr (S₊∙ (suc n) →∙ coHomK-ptd n) +fst (isContr-↓∙' zero) = (λ _ → 0) , refl +snd (isContr-↓∙' zero) (f , p) = + Σ≡Prop (λ f → isSetℤ _ _) + (funExt (toPropElim (λ _ → isSetℤ _ _) (sym p))) +fst (isContr-↓∙' (suc n)) = (λ _ → 0ₖ _) , refl +fst (snd (isContr-↓∙' (suc n)) f i) x = + sphereElim _ {A = λ x → 0ₖ (suc n) ≡ fst f x} + (λ _ → isOfHLevelTrunc (3 + n) _ _) + (sym (snd f)) x i +snd (snd (isContr-↓∙' (suc n)) f i) j = snd f (~ i ∨ j) + +isOfHLevel→∙Kn : {A : Pointed₀} (n m : ℕ) + → isOfHLevel (suc m) (A →∙ coHomK-ptd n) → isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) +isOfHLevel→∙Kn {A = A} n m hlev = step₃ + where + step₁ : isOfHLevel (suc m) (A →∙ Ω (coHomK-ptd (suc n))) + step₁ = + subst (isOfHLevel (suc m)) + (λ i → A →∙ ua∙ {A = Ω (coHomK-ptd (suc n))} {B = coHomK-ptd n} + (invEquiv Kn≃ΩKn+1) + (ΩKn+1→Kn-refl n) (~ i)) hlev + + step₂ : isOfHLevel (suc m) (typ (Ω (A →∙ coHomK-ptd (suc n) ∙))) + step₂ = isOfHLevelRetractFromIso (suc m) (invIso (invIso (ΩfunExtIso _ _))) step₁ + + step₃ : isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) + step₃ = + isOfHLevelΩ→isOfHLevel m + λ f → subst (λ x → isOfHLevel (suc m) (typ (Ω x))) + (isHomogeneous→∙ (isHomogeneousKn (suc n)) f) + step₂ + +isOfHLevel↑∙ : ∀ n m → isOfHLevel (2 + n) (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m))) +isOfHLevel↑∙ zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙ m)) +isOfHLevel↑∙ (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙ n m) + +isOfHLevel↑∙' : ∀ n m → isOfHLevel (2 + n) (S₊∙ (suc m) →∙ coHomK-ptd (suc (n + m))) +isOfHLevel↑∙' zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙' m)) +isOfHLevel↑∙' (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙' n m) + +→∙KnPath : (A : Pointed₀) (n : ℕ) → Ω (A →∙ coHomK-ptd (suc n) ∙) ≡ (A →∙ coHomK-ptd n ∙) +→∙KnPath A n = + ua∙ (isoToEquiv (ΩfunExtIso A (coHomK-ptd (suc n)))) refl + ∙ (λ i → (A →∙ Kn≃ΩKn+1∙ {n = n} (~ i) ∙)) + +contr∙-lem : (n m : ℕ) + → isContr (coHomK-ptd (suc n) + →∙ (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (n + m)) ∙)) +fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl) , refl +snd (contr∙-lem n m) (f , p) = →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousKn _)) (sym (funExt help)) + where + help : (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) + help = + trElim (λ _ → isOfHLevelPath (3 + n) + (subst (λ x → isOfHLevel x (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m)))) (λ i → suc (suc (+-comm n 1 i))) + (isOfHLevelPlus' {n = 1} (2 + n) (isOfHLevel↑∙ n m))) _ _) + (sphereElim _ (λ _ → isOfHLevel↑∙ n m _ _) p) + +isOfHLevel↑∙∙ : ∀ n m l + → isOfHLevel (2 + l) (coHomK-ptd (suc n) + →∙ (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (suc (l + n + m))) ∙)) +isOfHLevel↑∙∙ n m zero = + isOfHLevelΩ→isOfHLevel 0 λ f + → subst + isProp (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) + (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) h) + where + h : isProp (coHomK-ptd (suc n) + →∙ (Ω (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (suc (n + m))) ∙))) + h = + subst isProp (λ i → coHomK-ptd (suc n) →∙ (→∙KnPath (coHomK-ptd (suc m)) (suc (n + m)) (~ i))) + (isContr→isProp (contr∙-lem n m)) +isOfHLevel↑∙∙ n m (suc l) = + isOfHLevelΩ→isOfHLevel (suc l) + λ f → + subst + (isOfHLevel (2 + l)) (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) + (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) h) + where + h : isOfHLevel (2 + l) + (coHomK-ptd (suc n) + →∙ (Ω (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (suc (suc (l + n + m)))) ∙))) + h = + subst (isOfHLevel (2 + l)) + (λ i → coHomK-ptd (suc n) →∙ →∙KnPath (coHomK-ptd (suc m)) (suc (suc (l + n + m))) (~ i)) + (isOfHLevel↑∙∙ n m l) -coHomRed+1Equiv (suc n) A i = ∥ coHomRed+1.helpLemma A i {C = ((coHomK (suc n)) , ∣ north ∣)} i ∥₀ +-- Misc. +isoType→isoCohom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) + → Iso A B + → GroupIso (coHomGr n A) (coHomGr n B) +fst (isoType→isoCohom n is) = setTruncIso (domIso is) +snd (isoType→isoCohom n is) = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + (λ _ _ → refl)) diff --git a/Cubical/ZCohomology/PropertiesTrash.agda b/Cubical/ZCohomology/PropertiesTrash.agda deleted file mode 100644 index 3319aaae9..000000000 --- a/Cubical/ZCohomology/PropertiesTrash.agda +++ /dev/null @@ -1,938 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.PropertiesTrash where - - -open import Cubical.ZCohomology.Base -open import Cubical.HITs.Sn -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Everything -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Data.NatMinusTwo.Base renaming (-1+_ to -1+₋₂_ ; 1+_ to 1+₋₂_) -open import Cubical.Data.NatMinusOne.Base -open import Cubical.Data.Empty -open import Cubical.Data.Sigma -open import Cubical.Data.Prod.Base -open import Cubical.HITs.Susp -open import Cubical.HITs.SetTruncation -open import Cubical.HITs.Nullification -open import Cubical.Data.Int hiding (_+_ ; _-_) -open import Cubical.Data.Nat -open import Cubical.Data.Unit -open import Cubical.HITs.Truncation - -open import Cubical.HITs.Pushout -open import Cubical.Data.Sum.Base -open import Cubical.Data.HomotopyGroup -open import Cubical.Data.Bool -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - -is-_-Truncated : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} - (n : ℕ₋₂) → (f : A → B) → Type (ℓ-max ℓ ℓ') -is-_-Truncated {A = A} {B = B} n f = (b : B) → isOfHLevel (2+ n) (fiber f b) - -_-_ : ℕ → ℕ → ℕ -zero - m = zero -suc n - zero = suc n -suc n - suc m = n - m - -canceller : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (p q : a ≡ b) (r : b ≡ c) → p ∙ r ≡ q ∙ r → p ≡ q -canceller {ℓ} {A} {a} {b} {c} p q r id = (rUnit p ∙ (λ j → (p ∙ (rCancel r (~ j)))) ∙ assoc p r (sym r)) ∙ (cong (λ x → x ∙ (sym r)) id) ∙ sym (assoc q r (sym r)) ∙ (λ j → q ∙ (rCancel r j)) ∙ sym (rUnit q) - -switcher : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (p : a ≡ b) (q r : a ≡ c) → p ∙ (sym p) ≡ q ∙ (sym r) → r ≡ q -switcher {ℓ} {A} {a} {b} {c} p q r id = (lUnit r) ∙ ((λ j → ((rCancel p (~ j))) ∙ r)) ∙ (cong (λ x → x ∙ r) id) ∙ sym (assoc q (sym r) r) ∙ (λ j → q ∙ (lCancel r j)) ∙ sym (rUnit q) - - -Lemma296b : ∀{ℓ ℓ' ℓ''} {X : Type ℓ} {x y : X} {A : X → Type ℓ'} {B : X → Type ℓ''} - (p : x ≡ y) - (f : (A x) → (B x)) - (g : (A y) → (B y)) → - (transport (λ i → (A (p i)) → (B (p i))) f ≡ g) ≡ ((a : A y) → transport (λ i → B (p (~ i) )) (g a) ≡ f (transport (λ i → A (p (~ i))) a)) -Lemma296b {ℓ = ℓ} {X = X} {x = x} {y = y} {A = A} {B = B} = J {ℓ} {X} {x} (λ y p → (f : (A x) → (B x)) (g : (A y) → (B y)) → - (transport (λ i → (A (p i)) → (B (p i))) f ≡ g) ≡ - ((a : A y) → transport (λ i → B (p (~ i))) (g a) ≡ f (transport (λ i → A (p (~ i))) a))) - λ f g → transport (λ i → helper1 f g (~ i)) (helper2 f g) - where - helper1 : (f : (A x) → (B x)) (g : (A x) → (B x)) → - ((transport (λ i → A (refl {x = x} i) → B (refl {x = x} i)) f ≡ g) ≡ ((a : A x) → transport (λ i → B (refl {x = x} (~ i))) (g a) ≡ f (transport (λ i → A (refl {x = x} (~ i))) a))) - ≡ ((f ≡ g) ≡ ((a : A x) → g a ≡ f a)) - helper1 f g i = ((transportRefl f i ≡ g) ≡ ((a : A x) → transportRefl (g a) i ≡ f (transportRefl a i))) - - helper2 : (f : (A x) → (B x)) (g : (A x) → (B x)) → ((f ≡ g) ≡ ((a : A x) → g a ≡ f a)) - helper2 f g = isoToPath (iso (λ p a → sym (funExt⁻ p a)) (λ p → sym (funExt p) ) (λ x → refl) λ x → refl) - -Lemma296 : ∀{ℓ ℓ' ℓ''} {X : Type ℓ} {x y : X} {A : X → Type ℓ'} {B : X → Type ℓ''} - (p : x ≡ y) - (f : (A x) → (B x)) - (g : (A y) → (B y)) → - (transport (λ i → (A (p i)) → (B (p i))) f ≡ g) ≡ ((a : A x) → transport (λ i → B (p i)) (f a) ≡ g (transport (λ i → A (p i)) a)) -Lemma296 {ℓ = ℓ} {X = X} {x = x} {y = y} {A = A} {B = B} = J {ℓ} {X} {x} - (λ y p → (f : (A x) → (B x)) (g : (A y) → (B y)) → - (transport (λ i → (A (p i)) → (B (p i))) f ≡ g) ≡ - ((a : A x) → transport (λ i → B (p i)) (f a) ≡ g (transport (λ i → A (p i)) a))) - (λ f g → transport (λ i → reflTransport f g (~ i)) (reflCase f g)) - where - reflTransport : (f g : A x → B x) → ((transport (λ i → A (refl {x = x} i) → B (refl {x = x} i)) f ≡ g) ≡ ((a : A x) → transport (λ i → B (refl {x = x} i)) (f a) ≡ g (transport (λ i → A (refl {x = x} i)) a))) ≡ ((f ≡ g) ≡ ((a : A x) → f a ≡ g a)) - reflTransport f g j = (transportRefl f j ≡ g) ≡ ((a : A x) → transportRefl (f a) j ≡ g (transportRefl a j)) - - reflCase : (f g : A x → B x) → ((f ≡ g) ≡ ((a : A x) → f a ≡ g a)) - reflCase f g = isoToPath (iso (λ p → funExt⁻ p) (λ p → funExt p) (λ x → refl) λ x → refl) - - - - - -{- TODO: Prove Kₙ ≡ Ω Kₙ₊₁ -} - -invPathCancel : {a b : A} → (p : a ≡ b) → p ∙ (sym p) ≡ refl -invPathCancel {a = a} {b = b} p = J {x = a} (λ y p → p ∙ (sym p) ≡ refl ) (sym (lUnit refl)) p -invPathCancel2 : {a b : A} → (p : a ≡ b) → (sym p) ∙ p ≡ refl -invPathCancel2 {a = a} {b = b} p = J {x = a} (λ y p → (sym p) ∙ p ≡ refl ) (sym (lUnit refl)) p - -cancelReflMid : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) (q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q -cancelReflMid {ℓ = ℓ}{A = A} {a = a} {b = b} p q = J {ℓ} {A} {a} {ℓ} (λ b p → ((q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q)) (λ r → sym (lUnit (refl ∙ r ))) {b} p q - - - -inducedFun : ∀{ℓ''} {k : ℕ₋₂} - (f : A → B) → - (P : (B → HLevel ℓ'' (2+ k))) → - (((b : B) → fst (P b)) → (a : A) → fst (P (f a)) ) -inducedFun f P = λ g a → g (f a) - -Lemma757i→ii : ∀ {ℓ} (f : A → B) (n : ℕ₋₂) → - is- n -Connected f → - (P : B → HLevel ℓ (2+ n)) → - isEquiv (inducedFun f P) -Lemma757i→ii f n isCon P = {!!} - -Lemma757i→iii : ∀ {ℓ} (f : A → B) (n : ℕ₋₂) → - is- n -Connected f → - (P : B → HLevel ℓ (2+ n)) → - hasSection (inducedFun f P) -Lemma757i→iii f n isCon P = {!!} , (λ b → {!!}) - -Lemma861* : ∀{ℓ} (n k : ℕ₋₂) (f : A → B) → - (n ≡ neg2 → ⊥) → (k ≡ neg2 → ⊥) → - (n ≡ (suc neg2) → ⊥) → (k ≡ (suc neg2) → ⊥) → - (is- n -Connected f) → - (P : B → HLevel ℓ (2+ k)) → - Σ ℕ₋₂ (λ a → (a +₋₂ n ≡ k) × ((a ≡ neg2 → ⊥) × (a ≡ (suc neg2) → ⊥))) → - is- ((k -₋₂ n) -₋₂ (suc (suc (suc (suc neg2))))) -Truncated (inducedFun {k = k } f P) -Lemma861* n neg2 f n2 k2 n1 k1 isCon P pair = ⊥-elim (k2 refl) -Lemma861* n (suc neg2) f n2 k2 n1 k1 isCon P pair = ⊥-elim (k1 refl) -Lemma861* neg2 (suc (suc k)) f n2 k2 n1 k1 isCon P pair = ⊥-elim (n2 refl) -Lemma861* (suc neg2) (suc (suc k)) f n2 k2 n1 k1 isCon P pair = ⊥-elim (n1 refl) -Lemma861* (suc (suc n)) (suc (suc k)) f n2 k2 n1 k1 isCon P (neg2 , (x , (x₁ , x₂))) = ⊥-elim (x₁ refl) -Lemma861* (suc (suc n)) (suc (suc k)) f n2 k2 n1 k1 isCon P (suc neg2 , (x , (x₁ , x₂))) = ⊥-elim (x₂ refl) -Lemma861* (suc (suc n)) (suc (suc k)) f n2 k2 n1 k1 isCon P (suc (suc neg2) , (pf , (x₁ , x₂))) = {!pf!} -Lemma861* (suc (suc n)) (suc (suc k)) f n2 k2 n1 k1 isCon P (suc (suc (suc dif)) , (pf , (x₁ , x₂))) = {!isCon!} - - -Lemma861 : ∀{ℓ} (n k : ℕ) (f : A → B) → - (is- (ℕ→ℕ₋₂ n) -Connected f) → - (P : B → HLevel ℓ (suc (suc k))) → - Σ ℕ (λ a → a + n ≡ k) → - is- (ℕ→ℕ₋₂ ((k - n) - 2)) -Truncated (inducedFun f P) -Lemma861 {A = A} {B = B} n k f isCon P (dif , pf) = {!!} where - - helper1 : (l : (a : A) → fst (P (f a))) (a b : Σ ((b : B) → fst (P b)) λ g → (a : A) → g (f a) ≡ l a ) → (a ≡ b) ≡ Σ (fst a ≡ fst b) λ r → PathP (λ i → (x : A) → (r i) (f x) ≡ l x) (snd a) (snd b) - helper1 l a b = sym (ua (Σ≡ )) - - helper2 : (l : (a : A) → fst (P (f a))) (a b : Σ ((b : B) → fst (P b)) λ g → (a : A) → g (f a) ≡ l a ) → (Σ (fst a ≡ fst b) λ r → PathP (λ i → (x : A) → (r i) (f x) ≡ l x) (snd a) (snd b)) ≡ Σ ((x : B) → (fst a) x ≡ (fst b) x) λ r → PathP (λ i → (x : A) → r (f x) i ≡ l x) (snd a) (snd b) - helper2 l (g , p) (h , q) = isoToPath (iso lrFun rlFun (λ b → refl) λ b → refl) where - lrFun : (Σ (g ≡ h) λ r → PathP (λ i → (x : A) → (r i) (f x) ≡ l x) p q) → Σ ((x : B) → g x ≡ h x) λ r → PathP (λ i → (x : A) → r (f x) i ≡ l x) p q - lrFun (a , b) = funExt⁻ a , b - rlFun : (Σ ((x : B) → g x ≡ h x) λ r → PathP (λ i → (x : A) → r (f x) i ≡ l x) p q) → (Σ (g ≡ h) λ r → PathP (λ i → (x : A) → (r i) (f x) ≡ l x) p q) - rlFun (a , b) = (funExt a) , b - abstract - helper3 : (l : (a : A) → fst (P (f a))) (a b : Σ ((b : B) → fst (P b)) λ g → (a : A) → g (f a) ≡ l a ) → (r : (x : B) → (fst a) x ≡ (fst b) x) → (PathP (λ i → (x : A) → r (f x) i ≡ l x) (snd a) (snd b)) ≡ ((x : A) → r (f x) ≡ ((snd a) x) ∙ sym ((snd b) x) ) - helper3 l (g , p) (h , q) r j = hcomp (λ k → λ{ (j = i0) → (PathP (λ i → (x : A) → r (f x) (i ∧ i1) ≡ l x) p λ x → (sym (lUnit (q x)) k )) - ; (j = i1) → throwAbout k}) - ((PathP (λ i → (x : A) → r (f x) (i ∧ ~ j) ≡ l x) p λ x → (λ i → r (f x) (~ j ∨ i)) ∙ (q x))) - where - postulate - throwAbout : (p ≡ λ x → r (f x) ∙ q x) ≡ ((x : A) → (((r (f x))) ≡ (p x) ∙ (sym (q x)))) - {- Very heavy proof - throwAbout : (p ≡ λ x → r (f x) ∙ q x) ≡ ((x : A) → (((r (f x))) ≡ (p x) ∙ (sym (q x)))) - throwAbout = isoToPath (iso (λ g x → transport (λ i → throwAboutHelper (p x) (r (f x)) (sym (q x)) i ) (funExt⁻ g x) ) - (λ g → funExt λ x → transport (λ i → throwAboutHelper (p x) (r (f x)) (sym (q x)) (~ i) ) (g x)) - (λ b → funExt λ x → (λ j → transport (λ i → throwAboutHelper (p x) (r (f x)) (λ i₁ → q x (~ i₁)) (i ∨ j)) (transport (λ i₁ → throwAboutHelper (p x) (r (f x)) (λ i₂ → q x (~ i₂)) ((~ i₁) ∨ ( j))) (b x))) ∙ λ j → transportRefl (transportRefl (b x) j) j ) - λ b → (λ j → funExt (λ x → transport (λ i → throwAboutHelper (p x) (r (f x)) (λ i₁ → q x (~ i₁)) (~ i ∧ (~ j))) - (transport (λ i → throwAboutHelper (p x) (r (f x)) (λ i₁ → q x (~ i₁)) (i ∧ (~ j))) - (λ i → b i x)))) - ∙ (λ j → funExt (λ x → transportRefl ((transportRefl (λ i → b i x)) j) j ) )) where - - - throwAboutHelper : ∀{ℓ} {A : Type ℓ} {a b c : A} → (p : a ≡ b) (q : a ≡ c) (r : b ≡ c) → (p ≡ q ∙ (sym r) ) ≡ (q ≡ p ∙ r) - throwAboutHelper {ℓ} {A} {a} {b} {c} = J {ℓ} {A} {a} (λ b p → (q : a ≡ c) (r : b ≡ c) → (p ≡ q ∙ (sym r) ) ≡ (q ≡ p ∙ r)) - (J {ℓ} {A} {a} (λ c q → (r : a ≡ c) → (refl ≡ q ∙ (sym r) ) ≡ (q ≡ refl ∙ r) ) - λ r → (λ i → refl ≡ lUnit (sym r) (~ i)) ∙ isoToPath (iso (λ id → cong (λ x → sym x) id) (λ id → cong (λ x → sym x) id ) (λ b → refl) λ b → refl) ∙ λ i → (refl ≡ lUnit r i) ) - -} - - helper4 : (l : (a : A) → fst (P (f a))) (a b : Σ ((b : B) → fst (P b)) λ g → (a : A) → g (f a) ≡ l a ) → (r : (x : B) → (fst a) x ≡ (fst b) x) → ((x : A) → r (f x) ≡ ((snd a) x) ∙ sym ((snd b) x) ) ≡ ((λ (x : A) → r (f x)) ≡ (λ (x : A) → ((snd a) x) ∙ sym ((snd b) x))) - helper4 l (h , q) (g , p) r = isoToPath (iso (λ p → funExt p) (λ p → funExt⁻ p) (λ b → refl) λ b → refl) - - - FINAL : (l : (a : A) → fst (P (f a))) (a b : Σ ((b : B) → fst (P b)) λ g → (a : A) → g (f a) ≡ l a ) → (a ≡ b) ≡ Σ ((x : B) → (fst a) x ≡ (fst b) x) (λ r → ((λ (x : A) → r (f x)) ≡ (λ (x : A) → ((snd a) x) ∙ sym ((snd b) x)))) - FINAL l (g , p) (h , q) = helper1 _ _ _ ∙ helper2 _ _ _ ∙ (λ j → Σ ((x : B) → g x ≡ h x) (λ r → helper3 l (g , p) (h , q) r j)) ∙ λ j → Σ ((x : B) → g x ≡ h x) (λ r → helper4 l (g , p) (h , q) r j) - - - -{- -Lemma862 : (A : Pointed ℓ) (B : Pointed ℓ') (n m : ℕ₋₂) → - (is- (suc (suc n)) -ConnectedType (typ A)) → - (is- (suc (suc m)) -ConnectedType (typ B)) → - (P : (typ A) → (typ B) → HLevel (ℓ-max ℓ ℓ') (2+ ((suc (suc n)) +₋₂ (suc (suc m))))) → - (f : ((a : (typ A)) → fst (P a (pt B)))) → - (g : ((b : (typ B)) → fst (P (pt A) b))) → - (p : f (pt A) ≡ g (pt B)) → - Σ ((a : typ A) → (b : typ B) → fst (P a b)) λ h → Σ (((a : typ A) → h a (pt B) ≡ f a) × ((b : typ B) → h (pt A) b ≡ g b) ) λ pair → p ≡ sym ((proj₁ pair) (pt A) ) ∙ (proj₂ pair) (pt B) -Lemma862 {ℓ = ℓ} {ℓ' = ℓ'} (A , a₀) (B , b₀) n m conA conB P f g p = (λ a b → {!!}) , ({!!} , {!!}) where - Q : A → Type (ℓ-max ℓ ℓ') - Q a = Σ ((b : B) → fst (P a b )) λ k → f a ≡ k (b₀) - - a₀-map : Unit → A - a₀-map x = a₀ - - conOfa₀-map : is- (suc n) -Connected a₀-map - conOfa₀-map = {!!} - - b₀-map : Unit → B - b₀-map x = b₀ - - conOfb₀-map : is- (suc n) -Connected b₀-map - conOfb₀-map = {!!} - - conOfQ : (a : A) → isOfHLevel (2+ (suc n)) (Q a) - conOfQ = {!!} - - typesQ : (a : A) → Σ ((a : A) → Q a) λ ℓ → ℓ a₀ ≡ (g , p) - typesQ a = (fst (Lemma757i→iii a₀-map (suc n) conOfa₀-map (λ a → (Q a , conOfQ a ))) (λ x → (g , p))) , - cong (λ f → f tt) (snd (Lemma757i→iii a₀-map (suc n) conOfa₀-map (λ a → (Q a , conOfQ a ))) (λ x → (g , p))) - - QisFib : (a : A) → Q a ≡ fiber ( λ x → inducedFun {A = Unit} {B = B} b₀-map (P a) x tt) (f a) - QisFib a = {!!} - - Ok : (a : A) → is- (suc n) -ConnectedType (fiber ( λ x → inducedFun {A = Unit} {B = B} b₀-map (P a) x tt) (f a)) - Ok a = {!!} - -σ : A → {a : A} → typ (Ω ((Susp A) , north)) -σ x {a} = merid x ∙ sym (merid a) - - -code : ∀{ℓ} {A : Type ℓ} {a : A} (n : ℕ₋₂) → is- (suc (suc n)) -ConnectedType A → (y : Susp A) → (north ≡ y → Type ℓ) -code {A = A} {a = a} n iscon north p = (∥ fiber (λ y → σ y {a}) p ∥ ((suc (suc n)) +₋₂ (suc (suc n)))) -code {ℓ} {A = A} {a = a} n iscon south q = ∥ fiber merid q ∥ ((suc (suc n)) +₋₂ (suc (suc n))) -code {ℓ} {A = A} {a = a} n iscon (merid c i) = pathToConstruct i where - pathToConstruct : PathP (λ i → north ≡ merid c i → Type ℓ) (code {a = a} n iscon north) (code {a = a} n iscon south) - pathToConstruct = toPathP (transport (λ i → transpId (~ i)) λ q → sym (ua ((RlFun q) , (RlFunEquiv q)))) - where - - transpId : (transport (λ i₁ → north ≡ merid c i₁ → Type ℓ) (code {a = a} n iscon north) ≡ code {a = a} n iscon south) ≡ - ((q : north ≡ south) → (code {a = a} n iscon south q) ≡ (code {a = a} n iscon north (q ∙ sym (merid c)))) - transpId = (transport (λ i₁ → north ≡ merid c i₁ → Type ℓ) (code {a = a} n iscon north) ≡ code {a = a} n iscon south) - ≡⟨ Lemma296b {A = λ x → north ≡ x} {B = λ _ → Type ℓ} (merid c) (code {a = a} n iscon north) (code {a = a} n iscon south) ⟩ - ((q : north ≡ south) → code {a = a} n iscon south q ≡ code {a = a} n iscon north (transport (λ i → north ≡ (merid c (~ i))) q)) - ≡⟨ (λ i → ((q : north ≡ south) → code {a = a} n iscon south q ≡ code {a = a} n iscon north (helper q (sym (merid c)) i))) ⟩ - ((q : north ≡ south) → (code {a = a} n iscon south q) ≡ (code {a = a} n iscon north (q ∙ sym (merid c)))) ∎ - - where - helper : ∀{ℓ} {A : Type ℓ} {a b c : A} (q : a ≡ b) (p : b ≡ c) → - (transport (λ i → a ≡ (p i)) q) ≡ (q ∙ p) - helper {ℓ = ℓ} {A = A} {a = a} {b = b} {c = c} q = J {ℓ} {A} {b} - (λ c p → (transport (λ i → a ≡ (p i)) q) ≡ (q ∙ p)) - (transportRefl q ∙ rUnit q) - - - - RlFun : {c : A} (q : north ≡ south) → code {a = a} n iscon north (q ∙ sym (merid c)) → code {a = a} n iscon south q - RlFun {c = c} q x = (funHelper a) x - where - funHelper : (a : A) → (∥ Σ A (λ x₁ → merid x₁ ∙ (λ i₁ → merid a (~ i₁)) ≡ q ∙ sym (merid c)) ∥ (suc (suc n) +₋₂ suc (suc n))) → ∥ Σ A (λ x → merid x ≡ q) ∥ ((suc (suc n) +₋₂ suc (suc n))) - funHelper a = ind (λ x → isOfHLevel∥∥ _) λ r → sufMap (fst r) (snd r) - where - - sufMap : (x₂ : A) → (merid x₂ ∙ sym (merid a)) ≡ (q ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ q) ∥ ((suc (suc n) +₋₂ suc (suc n))) - sufMap x₂ = fst (Lemma862 (A , a) (A , a) n n iscon iscon (λ x₂ c → (((merid x₂ ∙ sym (merid a)) ≡ (q ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ q) ∥ ((suc (suc n) +₋₂ suc (suc n))) ) , isOfHLevelPi (2+ ((suc (suc n) +₋₂ suc (suc n)))) λ _ → isOfHLevel∥∥ _)) firstFun secondFun (funExt (λ x → λ j → ∣ (refl j) , (maybe (merid a) q (canceller (merid a) q (sym (merid a)) x ) x j) ∣ ))) x₂ c - where - - - - firstFun : (a₁ : A) → merid a₁ ∙ (λ i₁ → merid a (~ i₁)) ≡ q ∙ (λ i₁ → merid (pt (A , a)) (~ i₁)) → ∥ Σ A (λ x₁ → merid x₁ ≡ q) ∥ (suc (suc n) +₋₂ suc (suc n)) - firstFun x r = ∣ x , canceller (merid x) q (sym (merid a)) r ∣ - - {-J {ℓ} {A} {b} (λ d r → ((p ∙ r) ≡ (q ∙ r)) ≡ (p ≡ q)) λ i → sym (rUnit p) i ≡ sym (rUnit q) i-} - secondFun : (b : typ (A , a)) → merid (pt (A , a)) ∙ (λ i₁ → merid a (~ i₁)) ≡ q ∙ (λ i₁ → merid b (~ i₁)) → ∥ Σ A (λ x₁ → merid x₁ ≡ q) ∥ (suc (suc n) +₋₂ suc (suc n)) - secondFun b s = ∣ b , switcher (merid a) q (merid b) s ∣ - - {- need eckmann hilton -} - - - maybe : ∀ {ℓ} {A : Type ℓ} {a b : A} → (p q : a ≡ b) → (p ≡ q) → (P : p ∙ (sym p) ≡ q ∙ (sym p)) → canceller p q (sym p) P ≡ switcher p q p P - maybe {ℓ} {A} {a} {b} p q = J {ℓ} {a ≡ b} {p} (λ q _ → (P : p ∙ (sym p) ≡ q ∙ (sym p)) → canceller p q (sym p) P ≡ switcher p q p P) (J {ℓ} {A} {a} (λ b p → (P : p ∙ (sym p) ≡ p ∙ (sym p)) → canceller p p (sym p) P ≡ switcher p p p P) (λ P → {!refl!}) p) - - - canceller2 : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (p q : a ≡ b) (r : b ≡ c) → p ∙ r ≡ q ∙ r → p ≡ q - canceller2 {ℓ} {A} {a} {b} {c} p q r id = (rUnit p ∙ (λ j → (p ∙ (rCancel r (~ j)))) ∙ assoc p r (sym r)) ∙ (cong (λ x → x ∙ (sym r)) id) ∙ sym (assoc q r (sym r)) ∙ (λ j → q ∙ (rCancel r j)) ∙ sym (rUnit q) - - RlFunEquiv : {c : A} → (q : north ≡ south) → isEquiv (RlFun {c = c} q) - RlFunEquiv {c = c} q = fst (Lemma757i→iii {A = Unit} {B = A} (λ x → a) ((suc (suc n))) (λ b → {!fst iscon!} , {!!}) (λ a → ((isEquiv (RlFun {c = a} q)) , {!!} ))) (λ t → baseCase) c where - baseCase : isEquiv (RlFun {c = a} q) - baseCase = snd (isoToEquiv (iso (RlFun {c = a} q) (ind (λ _ → isOfHLevel∥∥ _) λ a → {!!}) {!!} {!!})) - - claim : RlFun {c = a} q ≡ ind (λ _ → isOfHLevel∥∥ _) λ b → ∣ (fst b) , canceller2 _ _ _ (snd b) ∣ - claim = {!!} --cannot say before I have prev lemmas - - isCona₀ : is- (suc (suc n)) -Connected (λ (x₁ : Unit) → a) - isCona₀ = λ b → {!!} , {!!} - - -Lemma8610 : ∀{ℓ ℓ' ℓ''} {A : Type ℓ} {a1 a2 : A} {B : A → Type ℓ'} (C : (a : A) → (B a → Type ℓ'')) (p : a1 ≡ a2) (b : B a2) → transport (λ i → {!uncurry C ?!}) {!!} ≡ {!(ua ?)!} -Lemma8610 = {!!} - -Thm864 : (n : ℕ₋₂) (a : A) (iscon : is- (suc (suc n)) -ConnectedType A) {y : Susp A} → (p : north ≡ y) → isContr (code {a = a} n iscon y p) -Thm864 {ℓ} {A} n a iscon = J {ℓ} {Susp A} {north} (λ y p → (isContr (code {a = a} n iscon y p))) (∣ a , (rCancel _) ∣ , (λ y → {!refl!})) where - center : {!!} - center = {!!} - -Freudenthal : (n : ℕ₋₂ ) (A : Pointed ℓ) → is- (suc (suc n)) -ConnectedType (typ A) → ∥ typ A ∥ (((suc (suc n) +₋₂ suc (suc n))) ) ≡ ∥ typ (Ω (Susp (typ A) , north)) ∥ (((suc (suc n) +₋₂ suc (suc n))) ) -Freudenthal n A isCon = ua ({!!} , {!Lemma757i→ii ? ? ? ?!}) - - - --} - - - - - - - - - - - - - - - - - - -{- J {ℓ} {A} {a} (λ b p → ((q : a ≡ b) (P : p ∙ (sym p) ≡ q ∙ (sym p)) → canceller p q (sym p) P ≡ switcher p q p P)) λ q id → (λ i → canceller refl q ((sym symRefl ) (~ i)) id) ∙ {!!} ∙ λ j → switcher refl q (symRefl (~ j)) id where - test : {q : a ≡ a} {id : refl ∙ (λ i₁ → refl (~ i₁)) ≡ q ∙ (λ i₁ → refl {x = a} (~ i₁))} → canceller refl q refl id ≡ (sym (rCancel refl)) ∙ (cong (λ x → x ∙ refl) ((lUnit refl) ∙ id)) ∙ (sym (assoc q refl refl)) ∙ (λ j → q ∙ (rCancel refl j)) ∙ (sym (rUnit q)) - test {q = q} {id = id} i = {!(sym (rCancel refl)) ∙ (cong (λ x → x ∙ (symRefl (~ i))) ((lUnit refl) ∙ id)) ∙ (sym (assoc q refl (symRefl (~ i)))) ∙ (λ j → q ∙ (rCancel refl j)) ∙ (sym (rUnit q))!} -} - -{- - - RlFun : (q : north ≡ south) → (∥ ((fiber (λ y → σ y) (q ∙ sym (merid c))) ≡ (Σ A (λ x₁ → merid x₁ ∙ (λ i₁ → merid a (~ i₁)) ≡ q ∙ sym (merid c)))) ∥ ((suc n) +₋₂ suc (suc n))) → (∥ (fiber merid q) ≡ (Σ A (λ x₁ → merid x₁ ≡ q)) ∥ ((suc n) +₋₂ suc (suc n))) - RlFun q x = ind (λ x → isOfHLevel∥∥ _) ({!cong funHelper!}) x where - funHelper : (a : A) → merid a ∙ (sym (merid c)) ≡ q ∙ (sym (merid c)) → ∥ Σ A (λ x → merid x ≡ q) ∥ ((suc (suc n) +₋₂ suc (suc n))) - funHelper a = {!!} - - LrFun : ((q : north ≡ south) → (code {a = a} n iscon south q) → (code {a = a} n iscon north (q ∙ sym (merid c)))) - LrFun = {!!} - --} - - -- fst gp (f x) ≡ fst hq (f x) - -- fst gp (f x) ≡ fst gp (f x) - - -{- - - helper2 : (l : (x : A) → (fst (P (f x)) )) - (gp hq : Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → - Σ (fst gp ≡ fst hq) (λ a≡ → PathP (λ i → (a : A) → a≡ i (f a) ≡ l a) (snd gp) (snd hq)) - ≡ Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) - helper2 l gp hq = isoToPath (iso fun finv (λ b i → {!fst gp x!} , {!!}) {!!}) where - - - finv : Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) (λ r → (inducedFun {k = (suc (suc k))} f (λ b → (fst gp b ≡ fst hq b) , hLevelSuc (suc (suc (1+ (1+₋₂ k)))) (fst (P b)) (snd (P b)) (fst gp b) (fst hq b) ) r ≡ (λ (x : A) → (snd gp x) ∙ sym (snd hq x)))) → Σ (fst gp ≡ fst hq) (λ a≡ → PathP (λ i → (a : A) → a≡ i (f a) ≡ l a) (snd gp) (snd hq)) - finv (fs , sn) = {!!} {- (funExt fs) , {!(snd hq) !}-} where - propLemma : (a : A) → isProp (fst hq (f a) ≡ l a) - propLemma a f g = {!!} - - - fun : Σ (fst gp ≡ fst hq) (λ a≡ → PathP (λ i → (a : A) → a≡ i (f a) ≡ l a) (snd gp) (snd hq)) → Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) (λ r → (inducedFun {k = (suc (suc k))} f (λ b → (fst gp b ≡ fst hq b) , hLevelSuc (suc (suc (1+ (1+₋₂ k)))) (fst (P b)) (snd (P b)) (fst gp b) (fst hq b) ) r ≡ (λ (x : A) → (snd gp x) ∙ sym (snd hq x)))) - fun (fs , sn) = (funExt⁻ fs) , (funExt (λ x → test5)) where - - test5 : {x : A} → PathP (λ _ → fst gp (f x) ≡ fst hq (f x) ) (λ i → (fs i) (f x)) (snd gp x ∙ sym (snd hq x)) - test5 {x = x} i j = hcomp ((λ k → λ{ (i = i0) → fs j (f x) ; (i = i1) → test3 {x = x} k j ; (j = i0) → fst gp (f x) ; (j = i1) → fs (~ i ∨ k) (f x) })) (test4 {x = x} i j) where - test3 : {x : A} → PathP (λ i → fst gp (f x) ≡ fs i (f x)) (snd gp x ∙ sym (snd gp x)) (snd gp x ∙ sym (snd hq x)) - test3 {x = x} i = (snd gp x) ∙ (sym (sn i x)) - - test4 : {x : A} → PathP (λ i → fst gp (f x) ≡ fs (~ i) (f x)) (λ i → (fs i) (f x)) (snd gp x ∙ (sym (snd gp x))) - test4 {x = x} i = hcomp (λ j → λ { (i = i0) → (λ i → (fs i) (f x)) ; (i = i1) → (sym (invPathCancel (snd gp x)) j) } ) (test i) where - test : {x : A} → PathP (λ i → fst gp (f x) ≡ fs (~ i) (f x)) (λ i → (fs i) (f x)) (refl {x = (fst gp (f x))}) - test {x = x} j i = fs (i ∧ ( ~ j)) (f x) - - helper3 : (l : (x : A) → (fst (P (f x)) )) (gp hq : Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → (gp ≡ hq) ≡ Σ (fst gp ≡ fst hq) - (λ a≡ → PathP (λ i → (a : A) → a≡ i (f a) ≡ l a) (snd gp) (snd hq)) - helper3 l gp hq = sym (ua (Σ≡ {A = ((a : B) → fst (P a))} {B = (λ x → (a : A) → x (f a) ≡ l a)} {x = gp} {y = hq})) - - --} - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - -{- - - - - - - - - - - - - - - - - - - - - - ------------------- -{- Sn ≡ Ω(Sn+1) -} - - - - - - - - --------------------- Guillaume Test -------------------- - - - - - - - - - - -sphereSuc : (n : ℕ) → (S₊ n) → (S₊ (suc n)) -sphereSuc n north = north -sphereSuc n south = south -sphereSuc (suc n) (merid a i) = merid (sphereSuc n a) i - -helplemma : {A : Type ℓ} {a : A} {p q : a ≡ a} {P : p ≡ q} → (C : (a ≡ a) → Type ℓ) → (a1 a2 : C (P i0)) → (b1 b2 : C (P (i1))) → PathP (λ i → C (P i)) a1 b2 → (a1 ≡ a2) → (b1 ≡ b2) → PathP (λ i → C (P i)) a2 b2 -helplemma {ℓ = ℓ} {a = a}{p = p} {q = q} {P = P} C a1 a2 b1 b2 depP a1a2 b1b2 = {!J {ℓ} {?} !} - -pathExtender : (B : A → Type ℓ) → {!!} -pathExtender = {!!} - -promote : (n : ℕ) → (S₊ n) → typ (Ω ((S₊ (suc n)) , north) ) -promote zero north = refl {x = north} -promote zero south = refl {x = north} -promote (suc n) north = refl {x = north} -promote (suc n) south = refl {x = north} -promote (suc n) (merid a i) = (( sym (rCancel (merid (merid a i0))) ∙ (λ i → (merid (merid a i) ∙ (sym (merid (merid north i)))))) ∙ rCancel (merid (merid a i1)) ) i - -decode' : (n : ℕ) → ∥ (S₊ n) ∥ (ℕ→ℕ₋₂ n) → ∥ typ (Ω (S₊ (suc n) , north)) ∥ (ℕ→ℕ₋₂ n) -decode' n = rec {A = (S₊ n)} {B = ∥ typ (Ω (S₊ (suc n) , north)) ∥ (ℕ→ℕ₋₂ n) } (isOfHLevel∥∥ {A = typ (Ω (S₊ (suc n) , north))} (ℕ→ℕ₋₂ n)) λ x → ∣ promote n x ∣ - -CODES : (n : ℕ) → S₊ (suc n) → Type₀ -CODES n north = ∥ S₊ n ∥ (ℕ→ℕ₋₂ n) -CODES n south = ∥ S₊ n ∥ (ℕ→ℕ₋₂ n) -CODES n (merid a i) = {!(typ ((Ω^ (suc n)) (Type₀ , ∥ S₊ n ∥ (ℕ→ℕ₋₂ n))))!} - -encode' : (n : ℕ) → ∥ typ (Ω ((S₊ (suc n)) , north)) ∥ (ℕ→ℕ₋₂ n) → ∥ S₊ n ∥ (ℕ→ℕ₋₂ n) -encode' n = rec {A = typ (Ω ((S₊ (suc n)) , north))} - {B = ∥ S₊ n ∥ (ℕ→ℕ₋₂ n)} - (isOfHLevel∥∥ {A = (S₊ n )} (ℕ→ℕ₋₂ n)) - λ x → {!(cong (CODES n) x) ?!} - -{- -functions : (n : ℕ) (f : A → B) → - (typ ((Ω^ n) ((A → B) , f))) ≡ - ((x : A) → (typ ((Ω^ n) (B , f x)))) -functions zero = {!!} -functions {A = A} {B = B}(suc zero) f = isoToPath (iso (λ g x i → (g i) x) (λ g → funExt {f = f} {g = f} g ) (λ b → refl) λ b → refl) -functions (suc (suc zero)) f = isoToPath (iso {!!} {!!} {!!} {!!}) - where - helper' : {A : Type ℓ} {B : Type ℓ'} → - (f : A → B) → - typ (Ω (Ω ((A → B) , f))) ≡ ( _≡_ {A = ( _≡_ f f)} (refl {x = f}) (refl {x = f})) - helper' f = refl - helper : {A : Type ℓ} {B : Type ℓ'} → - (f : A → B) → - (typ (Ω (Ω ((A → B) , f)))) ≡ ( _≡_ {A = ((x : A) → (f x ≡ f x) )} - (λ x _ → f x) - (λ x _ → f x)) - helper f = isoToPath (iso (λ g i x j → (g i j) x) (λ x k j y → ((x k) y) j) (λ b → refl) λ b → refl) - - helper2 :{A : Type ℓ} {B : Type ℓ'} → - (f : A → B) → - ( _≡_ {A = ((x : A) → (f x ≡ f x) )} (λ x _ → f x) (λ x _ → f x)) - ≡ - ((x : A) → (_≡_ {A = f x ≡ f x} (λ _ → f x) (λ _ → f x ))) - helper2 = {!!} - - -functions (suc (suc (suc n))) f = {!!} -- isoToPath (iso {!!} (λ x → {!!}) {!!} {!!}) - where - helper : {A : Type ℓ} {B : Type ℓ'} → - (f : A → B) → - typ (Ω (Ω (Ω ((Ω^ n) ((A → B) , f))))) ≡ ( _≡_ {A = typ (Ω (Ω ((Ω^ n) ((A → B) , f))))} - (λ _ → pt (Ω ((Ω^ n) ((A → B) , f)))) - (λ _ → pt (Ω ((Ω^ n) ((A → B) , f))))) - helper f = refl - - where - Ω^n : {A : Type ℓ} {B : Type ℓ'} → - (f : A → B) → - typ (Ω (Ω ((Ω^ n) ((A → B) , f)))) ≡ ((x : A) → typ (Ω (Ω ((Ω^ n) (B , f x))))) - Ω^n f = functions (suc (suc n)) f - Ω^n' : {A : Type ℓ} {B : Type ℓ'} → - (f : A → B) → - typ (Ω (Ω (Ω ((Ω^ n) ((A → B) , f))))) ≡ typ (Ω (((x : A) → typ (Ω (Ω ((Ω^ n) (B , f x))))) , (λ _ → refl))) - Ω^n' f i = typ (Ω (( Ω^n f i) , {!!})) - - --} - -Fun1000 : (n : ℕ) {f : A → B} → - ((typ ((Ω^ (suc n)) ((A → B) , f))) → - (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) --- Fun1001 : (n : ℕ) {f : A → B} → Fun1000 n {f} refl ≡ (λ x → refl) -Fun1000 zero p x i = p i x -Fun1000 (suc n) p x = {!Fun1000 n (p j) x i!} --- Fun1001 = {!!} - - -lrFun2* : (n : ℕ) {f : A → B} → - Σ ((typ ((Ω^ (suc n)) ((A → B) , f))) → - (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) - (λ g → (g refl) ≡ (λ x → refl)) -lrFun2* zero {f = f} = funExt⁻ , refl -lrFun2* (suc n) {f = f} = (λ p x → {!(funExt⁻ (cong (fst (lrFun2* n)) p) x)!}) - , - funExt (λ y → cancelReflMid (λ i → snd (lrFun2* n) (~ i) y) - (λ i → snd (lrFun2* n {f = f}) i y) ∙ - invPathCancel (λ i → snd (lrFun2* n) (~ i) y)) - -lrFun : (n : ℕ) {f : A → B} → - ((typ ((Ω^ (suc n)) ((A → B) , f))) → - (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) -lrFun n = fst (lrFun2* n) - - - where - - -rlFun* : (n : ℕ) {f : A → B} → - Σ ((((x : A) → (typ ((Ω^ (suc n)) (B , f x))))) → - ((typ ((Ω^ (suc n)) ((A → B) , f)))) ) - (λ g → (g (λ x → snd (((Ω^ (suc n)) (B , f x)))) ≡ pt ((Ω^ (suc n)) ((A → B) , f)))) - -rlFun* zero {f = f} = (λ g → funExt g) , refl -rlFun* {A = A} {B = B}(suc n) {f = f} = (λ g → (sym (snd (rlFun* n)) ∙ cong (fst (rlFun* n)) (funExt g) ∙ snd (rlFun* n))) - , - (cancelReflMid (λ i → snd (rlFun* n) (~ i)) (snd (rlFun* n))) ∙ - invPathCancel (λ i → snd (rlFun* n) (~ i)) - -rlFun : (n : ℕ) {f : A → B} → - (((x : A) → (typ ((Ω^ (suc n)) (B , f x))))) → - ((typ ((Ω^ (suc n)) ((A → B) , f)))) -rlFun n {f = f} = fst (rlFun* n {f = f}) - where - -functions2 : (n : ℕ) (f : A → B) → - ((x : A) → (typ ((Ω^ n) (B , f x)))) ≃ (typ ((Ω^ n) ((A → B) , f))) - -functions2 {A = A} {B = B} zero f = idEquiv (A → B) -functions2 {A = A} {B = B} (suc n) f = rlFun n , record { equiv-proof = λ y → ((fst (lrFun2* n)) y , {!!}) , {!!} } where - - helper : (n : ℕ) (y : (typ ((Ω^ (suc n)) ((A → B) , f)))) → (rlFun n ((fst (lrFun2* n)) y) ≡ y) - helper zero y = refl - helper (suc n) y = (λ i → (sym (snd (rlFun* n))) ∙ ((helper5 (funExt (λ x → sym (cong (λ g → g x ) (snd (lrFun2* n))))) (cong (fst (lrFun2* n)) y) (funExt (λ x → (cong (λ g → g x ) (snd (lrFun2* n))))) (fst (rlFun* n))) i ) ∙ (snd (rlFun* n))) ∙ {!helper !} - where - - - idTest : (n : ℕ) → (rlFun n {f} (λ x → snd (lrFun2* n {f}) i1 x) ≡ rlFun n (λ x → snd (lrFun2* n {f}) i0 x)) ≡ ((fst (rlFun* n) (λ x → snd (Ω ((Ω^ n) (B , f x))))) ≡ (pt (Ω ((Ω^ n) ((A → B) , f))))) - idTest n i = refl {x = rlFun n {f} (λ x → snd (lrFun2* n {f}) i1 x)} i ≡ idTest1 i where - idTest1 : rlFun n (λ x → snd (lrFun2* n {f}) i0 x) ≡ (pt (Ω ((Ω^ n) ((A → B) , f)))) - idTest1 = helper n (pt (Ω ((Ω^ n) ((A → B) , f)))) - - - test : (n : ℕ) → PathP (λ i → {!!}) ( (cong (rlFun n {f}) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))) (snd (rlFun* n {f})) - test n i = {!!} where - - - maybe2 : {n : ℕ} (i : I) → snd (rlFun* n) (~ i) ≡ (pt (Ω ((Ω^ n) ((A → B) , f)))) - maybe2 {n} i j = snd (rlFun* n) ((~ i) ∨ j) - - maybe3 : (n : ℕ) → (pt (Ω ((Ω^ n) ((A → B) , f))) ≡ rlFun n (λ x → snd (lrFun2* n {f}) i0 x)) - maybe3 n = sym (snd (rlFun* n {f})) ∙ λ i → rlFun n (hihih ( ~ i)) where - hihih : (λ (x : A) → snd (lrFun2* n {f}) i0 x) ≡ (λ x → snd (Ω ((Ω^ n) (B , f x)))) - hihih = funExt λ x → cong (λ g → g x) (snd (lrFun2* n)) - - maybe : (n : ℕ) (y : (typ (Ω ((Ω^ (suc n)) ((A → B) , f))))) → ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x)))))) ∙ (cong (rlFun n) (funExt (λ x → (λ i → fst (lrFun2* n) (y i) x)))) ∙ ((sym (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x))))))) ≡ (snd (rlFun* n)) ∙ y ∙ (sym (snd (rlFun* n))) - maybe n y i = test n i ∙ (cong (λ x → (helper n x) i) y ) ∙ sym (test n i) - -- (λ j → (snd (rlFun* n)) (~ j ∨ i)) ∙ {!!} - - -- ∙ {!λ i → (cong (λ x → helper n x i ) y)!} ∙ {!!} -- (cong (λ x → helper n x i ) ?) - - compTest : (n : ℕ) (y : (typ (Ω ((Ω^ (suc n)) ((A → B) , f))))) → - (sym (snd (rlFun* n))) ∙ - ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x)))))) ∙ - (cong (rlFun n) (funExt (λ x → (λ i → fst (lrFun2* n) (y i) x)))) ∙ - ((sym (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x))))))) ∙ - (snd (rlFun* n)) ≡ - refl ∙ {!!} ∙ y ∙ {!!} ∙ refl - compTest n y i = (λ i₁ → snd (rlFun* n {f}) (~ i₁ ∨ i)) ∙ - {!snd (rlFun* n {f}) (~ i1 ∨ i)!} {- hcomp (λ k → λ{(i = i0) → {!maxZeuner i0!} ; (i = i1) → (maxZeuner i1)}) (maxZeuner i) -} ∙ - (cong (λ x → (helper n x) i) y ) ∙ - {!snd (rlFun* n {f}) (~ i1 ∨ i)!} - ∙ λ i₁ → snd (rlFun* n {f}) (i₁ ∨ i) - - where - - maxZeuner : (i : I) → snd (rlFun* n) i ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i - maxZeuner i j = hcomp (λ k → λ{(j = i0) → snd (rlFun* n) i ; (j = i1) → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i ∨ (~ k)) }) (snd (rlFun* n {f}) (i ∨ j)) - maxZeunerHelper : (i : I) → maxZeuner i0 ≡ {!snd (rlFun* n) i0 ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i0!} - maxZeunerHelper = {!!} - - gimmeMore : (i : I) → rlFun n (λ x → snd (lrFun2* n {f}) i0 x) ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i - gimmeMore i = (λ j → rlFun n (λ x → snd (lrFun2* n {f}) j x)) ∙ (snd (rlFun* n {f})) ∙ λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i ∨ (~ j)) - - gimmeLeft : (i : I) → rlFun n (λ x → snd (lrFun2* n {f}) i1 x) ≡ snd (rlFun* n) i - gimmeLeft i = λ j → snd (rlFun* n {f}) (j ∧ i) - - gimmei0 : ((sym (gimmeLeft i0)) ∙ ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))) ∙ (gimmeMore i0)) ≡ ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))) - gimmei0 = {!((sym (gimmeLeft i0)))!} where - - gimmei0' : ((sym (gimmeLeft i0))) ≡ refl - gimmei0' j = refl - - - - - gimmei1 : ((sym (gimmeLeft i1)) ∙ ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))) ∙ (gimmeMore i1)) ≡ refl - gimmei1 = gimmei1' ∙ more!! where - - gimmei1' : ((sym (gimmeLeft i1)) ∙ ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))) ∙ (gimmeMore i1)) ≡ (sym (gimmeLeft i1)) ∙ (snd (rlFun* n {f})) ∙ λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j)) - gimmei1' = ∙-assoc (sym (gimmeLeft i1)) (( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x))))))) ((snd (rlFun* n {f})) ∙ λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j))) where - - ∙-assoc : {A : Type ℓ}{a b c d : A} (p : a ≡ b) (q : b ≡ c) (t : b ≡ d) → p ∙ q ∙ (sym q ∙ t) ≡ p ∙ t - ∙-assoc {ℓ = ℓ} {A = A} {a = a} {b = b} {c = c} {d = d} p q t = cong (λ x → p ∙ x) {!q ∙ (sym q ∙ t) ≡ t!} - - - more!! : (sym (gimmeLeft i1)) ∙ (snd (rlFun* n {f})) ∙ (λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j))) ≡ refl - more!! = (assoc (sym (gimmeLeft i1)) (snd (rlFun* n {f})) (λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j)))) ∙ (λ j → (invPathCancel (sym (gimmeLeft i1)) j) ∙ (λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j)))) ∙ lCancel (λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j))) - - more!!! : (λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j))) ≡ refl - more!!! = refl - - - - hello : (i : I) → snd (rlFun* n) i ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i - hello i = (λ j → snd (rlFun* n {f}) (i ∨ j)) ∙ λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i ∨ (~ j)) - - kittyHelper : PathP (λ i → (snd (rlFun* n {f})) i ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i) (hello i0) refl - kittyHelper j = hcomp (λ k → λ{ (j = i0) → hello i0 ; - (j = i1) → invPathCancel refl k }) - ((λ i → snd (rlFun* n {f}) (j ∨ i)) ∙ λ i → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (j ∨ (~ i))) -{- - kittyHelperB : PathP (λ i → (snd (rlFun* n {f})) i ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i) ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x)))))) refl - kittyHelperB i j = hcomp (λ k → λ{ (j = i0) → {!!} - ; (j = i1) → {!(cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x)))))!}} ) - {!( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ j) x))))))!} --} - - -- helloKitty : hello i0 ≡ ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x)))))) - -- helloKitty k = {!!} - - -- helloKitty2 : hello i1 ≡ refl - -- helloKitty2 = {!!} - - mini : (n : ℕ) → rlFun n (λ x → snd (lrFun2* n {f}) i0 x) ≡ pt (Ω ((Ω^ n) ((A → B) , f))) - mini n = helper n refl - - mini2 : (n : ℕ) → rlFun n (λ x → snd (lrFun2* n {f}) i1 x) ≡ pt (Ω ((Ω^ n) ((A → B) , f))) - mini2 n = (λ i → rlFun n (λ x → snd (lrFun2* n {f}) (~ i) x)) ∙ helper n refl - - leftMost : PathP (λ i → (pt (Ω ((Ω^ n) ((A → B) , f))) ≡ (snd (rlFun* n)) i)) (sym (snd (rlFun* n))) refl - leftMost i = λ i₁ → snd (rlFun* n) (~ i₁ ∨ i) - - - - notLeftMost : (n : ℕ) → PathP (λ i → mini2 n i ≡ mini n i ) (( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x))))))) refl - notLeftMost zero i j = {!fst (rlFun* (suc n)) (λ x → snd (lrFun2* (suc n) {f}) (i1) x)!} - notLeftMost (suc n) i = {!(( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))))!} - where - ffs2 : (n : ℕ) (i : I) → (fst (rlFun* n {f}) (λ x → snd (lrFun2* n {f}) (~ i1 ∨ i) x)) ≡ mini n i - ffs2 n i = ((λ j → fst (rlFun* n) (λ x → snd (lrFun2* n {f}) (i ∨ j) x)) ∙ snd (rlFun* n {f}) ) ∙ λ j → mini n (i ∨ (~ j)) - - ffs : (n : ℕ) (i : I) → (fst (rlFun* n {f}) (λ x → snd (lrFun2* n {f}) (~ i0 ∨ i) x)) ≡ mini2 n i - ffs n i = snd (rlFun* n {f}) ∙ {!mini2 i1!} - - {- - - tester : ((sym (ffs i0)) ∙ (λ j → fst (rlFun* n) (snd (lrFun2* n {f}) (i0 ∨ (~ j)))) ∙ (ffs2 i0)) ≡ (( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x))))))) - tester = (({!((sym (ffs i1)) ∙ (λ j → fst (rlFun* n) (snd (lrFun2* n {f}) (i1 ∨ (~ j)))) ∙ (ffs2 i1))!}) ∙ sym (rUnit (refl ∙ (λ j → fst (rlFun* n) (snd (lrFun2* n {f}) (~ j)))))) ∙ sym (lUnit (λ j → fst (rlFun* n) (snd (lrFun2* n {f}) (i0 ∨ (~ j))))) - where - helpi : (ffs2 i0) ≡ refl - helpi = {!!} - - snLeftMost : {!(λ j → fst (rlFun* n {f}) (snd (lrFun2* n) (~ j)))!} - snLeftMost = {!!} - - -} - - - helper2 : (n : ℕ) (y : typ (Ω (Ω ((Ω^ n) ((A → B) , f))))) → funExt ((λ (x : A) → (λ i → snd (lrFun2* n) (~ i) x) ∙ (λ i → fst (lrFun2* n) (y i) x) ∙ (λ i → snd (lrFun2* n {f = f}) i x))) ≡ (funExt (λ (x : A) → (λ i → snd (lrFun2* n) (~ i) x))) ∙ funExt (λ (x : A) → (λ i → fst (lrFun2* n) (y i) x)) ∙ funExt λ (x : A) → (λ i → snd (lrFun2* n {f = f}) i x) - helper2 zero y = refl - helper2 (suc n) y = refl - - helper4 : {A : Type ℓ} {B : Type ℓ} {a b c : A} (p : a ≡ b) (q : b ≡ c) (g : A → B) → cong g (p ∙ q) ≡ ((cong g p) ∙ (cong g q)) - helper4 {ℓ = ℓ} {A} {B} {a} {b} {c} p q g = J {ℓ} {A} {a} {ℓ} (λ b1 p → ((q1 : b1 ≡ c) → cong g (p ∙ q1) ≡ ((cong g p) ∙ (cong g q1)))) - (λ q1 → J {ℓ} {A} {a} {ℓ} (λ c2 q2 → (λ i → g ((refl ∙ q2) i)) ≡ (λ i → g (refl i)) ∙ (λ i → g (q2 i))) ( (λ j → ((λ i → g (((rUnit (λ _ → a)) (~ j) ) i)))) ∙ lUnit ((λ i → g (refl i)))) {c} q1 ) - p q - - helper5 : {A : Type ℓ} {B : Type ℓ} {a b c d : A} (p : a ≡ b) (q : b ≡ c) (r : c ≡ d) (g : A → B) → cong g (p ∙ q ∙ r) ≡ ((cong g p) ∙ (cong g q) ∙ (cong g r)) - helper5 {ℓ = ℓ} {A} {B} {a} {b} {c} {d} p q r g = helper4 p (q ∙ r) g ∙ ((cong (λ x → (cong g p) ∙ x )) (helper4 q r g)) - - helper3 : (n : ℕ) (y : typ (Ω (Ω ((Ω^ n) ((A → B) , f))))) → (cong (rlFun n) (funExt (λ x → (λ i₁ → fst (lrFun2* n) (y i₁) x)))) ≡ {!!} - helper3 zero y = {!!} - helper3 (suc n) y = {!(cong (rlFun (suc n)) (funExt (λ x → (λ i₁ → fst (lrFun2* (suc n)) (y i₁) x))))!} - -{- -functions' {A = A} {B = B}(suc zero) f = isoToPath (iso (λ g x i → (g i) x) (λ g → funExt {f = f} {g = f} g ) (λ b → refl) λ b → refl) - where - helper : transport (λ i → (sym (functions' (suc zero) f)) i) (λ x → refl) ≡ (pt ((Ω ) ((A → B) , f))) - helper = {!!} -functions' (suc (suc zero)) f = isoToPath (iso {!!} {!!} {!!} {!!}) - where - helper' : {A : Type ℓ} {B : Type ℓ'} → - (f : A → B) → - typ (Ω (Ω ((A → B) , f))) ≡ ( _≡_ {A = ( _≡_ f f)} (refl {x = f}) (refl {x = f})) - helper' f = refl - helper : {A : Type ℓ} {B : Type ℓ'} → - (f : A → B) → - (typ (Ω (Ω ((A → B) , f)))) ≡ ( _≡_ {A = ((x : A) → (f x ≡ f x) )} - (λ x _ → f x) - (λ x _ → f x)) - helper f = isoToPath (iso (λ g i x j → (g i j) x) (λ x k j y → ((x k) y) j) (λ b → refl) λ b → refl) - - helper2 :{A : Type ℓ} {B : Type ℓ'} → - (f : A → B) → - ( _≡_ {A = ((x : A) → (f x ≡ f x) )} (λ x _ → f x) (λ x _ → f x)) - ≡ - ((x : A) → (_≡_ {A = f x ≡ f x} (λ _ → f x) (λ _ → f x ))) - helper2 = {!!} -functions' (suc (suc (suc n))) f = {!!} - where - helper : {A : Type ℓ} {B : Type ℓ'} → - (f : A → B) → - typ (Ω (Ω (Ω ((Ω^ n) ((A → B) , f))))) ≡ ( _≡_ {A = typ (Ω (Ω ((Ω^ n) ((A → B) , f))))} - (λ _ → pt (Ω ((Ω^ n) ((A → B) , f)))) - (λ _ → pt (Ω ((Ω^ n) ((A → B) , f))))) - helper f = refl - - where - Ω^n : {A : Type ℓ} {B : Type ℓ'} → - (f : A → B) → - typ (Ω (Ω ((Ω^ n) ((A → B) , f)))) ≡ ((x : A) → typ (Ω (Ω ((Ω^ n) (B , f x))))) - Ω^n f = functions (suc (suc n)) f - Ω^n' : {A : Type ℓ} {B : Type ℓ'} → - (f : A → B) → - (Ω (Ω (Ω ((Ω^ n) ((A → B) , f))))) ≡ (Ω (((x : A) → typ (Ω (Ω ((Ω^ n) (B , f x))))) , (λ _ → refl))) - Ω^n' f i = (Ω (( Ω^n f i) , {!(transport (λ j → ( Ω^n f j)) (refl i)) !})) where - lemma : transport (λ i → Ω^n f i) refl ≡ λ _ → refl - lemma = funExt {f = transport (λ i → Ω^n f i) refl} {g = λ _ → refl} λ x → {!!} - - - - - ---------- - -lrFun3* : (n : ℕ) {f : A → B} → - Σ ((typ ((Ω^ (suc n)) ((A → B) , f))) → - (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) (λ g → (g refl) ≡ (λ x → refl)) -lrFun3* zero {f = f} = funExt⁻ , refl -lrFun3* {A = A} {B = B} (suc n) {f = f} = (λ p x → transport (λ i → (cong (λ g → g x) (snd (lrFun3* n {f = f}))) i ≡ (cong (λ g → g x) (snd (lrFun3* n {f = f}))) i) (funExt⁻ (cong (λ x → (fst (lrFun3* n) x )) p) x)) , funExt (λ x → transpLemma x) - module lr where - helpLemma : (x : A) → ((fst (lrFun3* n) (snd (Ω ((Ω^ n) ((A → B) , f)))) x ≡ (fst (lrFun3* n) (snd (Ω ((Ω^ n) ((A → B) , f)))) x))) ≡ (typ (Ω (Ω ((Ω^ n) (B , f x))))) - helpLemma x i = helper x i ≡ helper x i where - - helper : (x : A) → fst (lrFun3* n) (snd (Ω ((Ω^ n) ((A → B) , f)))) x ≡ pt ((Ω ((Ω^ n) (B , f x)))) - helper x = cong (λ g → g x) (snd (lrFun3* n {f = f})) - - transpLemma : (x : A) → transp (λ i → helpLemma x i) i0 refl ≡ refl - transpLemma x j = transp (λ i → helpLemma x (i ∨ j)) j refl - -lrFun3 : (n : ℕ) {f : A → B} → ((typ ((Ω^ (suc n)) ((A → B) , f))) → (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) -lrFun3 n {f = f} = (fst (lrFun3* n {f = f})) - - -rlFun3* : (n : ℕ) {f : A → B} → - Σ ((((x : A) → (typ ((Ω^ (suc n)) (B , f x))))) → - ((typ ((Ω^ (suc n)) ((A → B) , f)))) ) - (λ g → (g (λ x → snd (((Ω^ (suc n)) (B , f x)))) ≡ pt ((Ω^ (suc n)) ((A → B) , f)))) - - - - -rlFun3* zero {f = f} = (λ g → funExt g) , refl -rlFun3* {A = A} {B = B}(suc n) {f = f} = (λ g → transport (λ i → (snd (rlFun3* n {f})) i ≡ (snd (rlFun3* n {f})) i) (cong (fst (rlFun3* n)) (funExt g))) , transpLemma -- (cancelReflMid (λ i → snd (rlFun* n) (~ i)) (snd (rlFun* n))) ∙ invPathCancel (λ i → snd (rlFun* n) (~ i)) - module rl where - helpLemma : (fst (rlFun3* n) (λ x → snd (Ω ((Ω^ n) (B , f x)))) ≡ fst (rlFun3* n) (λ x → snd (Ω ((Ω^ n) (B , f x))))) ≡ typ (Ω (Ω ((Ω^ n) ((A → B) , f)))) - helpLemma i = (snd (rlFun3* n {f})) i ≡ (snd (rlFun3* n {f})) i - - transpLemma : transport (λ i → helpLemma i) refl ≡ refl - transpLemma j = transp (λ i → helpLemma (i ∨ j)) j refl -- transp (λ i → helpLemma (i ∨ j)) j refl - -rlFun3 : (n : ℕ) {f : A → B} → - (((x : A) → (typ ((Ω^ (suc n)) (B , f x))))) → - ((typ ((Ω^ (suc n)) ((A → B) , f)))) -rlFun3 n {f = f} = fst (rlFun3* n {f = f}) - where - -functions3 : (n : ℕ) (f : A → B) → - ((x : A) → (typ ((Ω^ n) (B , f x)))) ≃ (typ ((Ω^ n) ((A → B) , f))) - -functions3 {A = A} {B = B} zero f = idEquiv (A → B) -functions3 {A = A} {B = B} (suc n) f = rlFun3 n , record { equiv-proof = λ y → (( (fst (lrFun3* n)) y) , λ i → {!!}) , {!!} } where - linv : (n : ℕ) (y : typ (Ω ((Ω^ n) ((A → B) , f)))) → rlFun3 n (fst (lrFun3* n {f = f}) y) ≡ y - linv zero y = refl - linv (suc n) y = {!!} where - maybe : rlFun3 (suc n) (fst (lrFun3* (suc n) {f = f}) y) ≡ λ j x → {!((cong (λ x → (fst (lrFun3* n) x )) y) (~ j)) x!} - maybe = {!!} - - -inv1 : (n : ℕ) {f : A → B} (g : ((typ ((Ω^ (suc n)) ((A → B) , f))) )) → (rlFun3 n) ((lrFun3 n) g ) ≡ g -inv1 zero {f = f} g = refl -inv1 {A = A} {B = B}(suc n) {f = f} g j = {! -- transp (λ i → (snd (rlFun3* n {f})) (i ∨ j) ≡ (snd (rlFun3* n {f})) (i ∨ j)) i0 - (cong (fst (rlFun3* n)) (funExt ( - (λ x → transp (λ i → (cong (λ g → g x) (snd (lrFun3* n {f = f}))) (i ∨ j) - ≡ (cong (λ g → g x) (snd (lrFun3* n {f = f}))) (i ∨ j)) (~ j) - (funExt⁻ (cong (λ x → (fst (lrFun3* n) x )) g) x)) ) ) )!} - - where - maybe : PathP (λ i → {! !} ≡ {!!} ) (cong (fst (rlFun3* n)) (funExt ( - (λ x → transport (λ i → (cong (λ g → g x) (snd (lrFun3* n {f = f}))) i - ≡ (cong (λ g → g x) (snd (lrFun3* n {f = f}))) i) - (funExt⁻ (cong (λ x → (fst (lrFun3* n {f}) x )) g) x)) ) ) ) - {!(cong (fst (rlFun3* n)) (funExt ( (λ x → (funExt⁻ (cong (λ x → (fst (lrFun3* n {f}) x )) g) x))))) !} - maybe j = {!fst (rlFun3* n) (λ x → snd (lrFun3* n {f}) i1 x)!} --} --} - - - - - - - - - -{- -Lemma861 : ∀{ℓ''} (n k : ℕ₋₂) → - Σ ℕ₋₂ (λ num → (suc (suc n)) +₋₂ num ≡ (suc (suc k)) ) → - (f : A → B) → - (is- (suc (suc n)) -Connected f) → - (P : (B → HLevel ℓ'' (2+ (suc (suc k))))) → - is- (((suc (suc n)) -₋₂ (suc (suc k))) -₋₂ (suc (suc neg2))) -Connected (inducedFun {k = (suc (suc k))} f P) -Lemma861 n k (neg2 , pf) f conf P b = {!!} - where - helper : (((suc (suc n)) -₋₂ (suc (suc k))) -₋₂ (suc (suc neg2))) ≡ neg2 - helper = {!!} -Lemma861 {A = A} {B = B} n k (suc j , pf) f conf P g = {!!} where - helper : (l : (x : A) → (fst (P (f x)) )) → - (fiber (inducedFun {k = (suc (suc k))} f P) l) ≡ Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a) - helper l = isoToPath (iso (fun) finv (λ b → refl) λ b → refl) where - fun : (fiber (inducedFun {k = (suc (suc k))} f P) l) → Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a) - fun (fs , sn) = fs , λ a → cong (λ x → x a) sn - - finv : (Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → (fiber (inducedFun {k = (suc (suc k))} f P) l) - finv (fs , sn) = fs , (funExt sn) - - helper'3 : (l : (x : A) → (fst (P (f x)) )) - (gp hq : Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → - (gp ≡ hq) - ≡ Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) - helper'3 l gp hq = isoToPath (iso fun finv (λ b i → ({!!})) {!!}) where - fun : (gp ≡ hq) → Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) - fun path = (λ x i → fst (path i) x ) , funExt λ x → pathFinal {x} where - pathHelper : {x : A} → PathP (λ i → fst gp (f x) ≡ fst (path i) (f x)) refl (snd gp x ∙ sym (snd hq x)) - pathHelper {x = x} i = hcomp (λ j → λ{ (i = i0) → rCancel (snd gp x) j ; (i = i1) → (snd gp x ∙ sym (snd hq x)) } ) (snd gp x ∙ sym (snd (path i) x)) - - pathHelper2 : {x : A} → PathP (λ i → fst gp (f x) ≡ fst (path (~ i)) (f x)) (λ i → fst (path i) (f x)) refl - pathHelper2 {x = x} i = λ i₁ → fst (path (i₁ ∧ (~ i))) (f x) - - pathFinal : {x : A} → (λ i → fst (path i) (f x)) ≡ snd gp x ∙ sym (snd hq x) - pathFinal {x = x} = transport (λ i → PathP (λ j₁ → (lCancel (λ i → fst gp (f x) ≡ fst (path i) (f x))) i j₁) - (λ i → fst (path i) (f x)) - (snd gp x ∙ sym (snd hq x))) - (compPathP pathHelper2 pathHelper) - - - finv : Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) → (gp ≡ hq) - finv (fs , sn) = ΣPathP {x = gp} {y = hq} (funExt (λ b → fs b) , {! !}) -- (λ b → (fs b) i) , λ a → hcomp (λ k → λ{(i = i0) → (snd gp a) ; (i = i1) → {!!}}) (transp (λ j → fs (f a) (i ∧ j) ≡ l a) (~ i) ((snd gp) a)) - where - potTerm : {a : A} → PathP (λ i → (a : A) → funExt (λ b → fs b) i (f a) ≡ l a) (snd gp) (snd hq) - potTerm {a = a} j = λ a → hcomp {!(transp (λ i → funExt (λ b → fs b) (i) (f a) ≡ l a) i0 ((snd gp a))) !} (transp (λ i → funExt (λ b → fs b) (i ∧ j) (f a) ≡ l a) i0 ((snd gp a))) - - helper2 : {a : A} → (fs (f a)) ∙ (snd hq a) ≡ snd gp a - helper2 {a = a} = (cong (λ x → x ∙ (snd hq a)) (cong (λ x → x a) sn)) ∙ sym (assoc (snd gp a) (sym (snd hq a)) (snd hq a)) ∙ (λ j → (snd gp a) ∙ (lCancel (snd hq a)) j) ∙ sym (rUnit (snd gp a)) -} diff --git a/Cubical/ZCohomology/RingStructure/CupProduct.agda b/Cubical/ZCohomology/RingStructure/CupProduct.agda new file mode 100644 index 000000000..dd9da8180 --- /dev/null +++ b/Cubical/ZCohomology/RingStructure/CupProduct.agda @@ -0,0 +1,87 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.RingStructure.CupProduct where +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties + +open import Cubical.HITs.S1 hiding (_·_) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec2 to sRec2) +open import Cubical.HITs.Truncation renaming (rec to trRec) + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Data.Int hiding (_+'_ ; +'≡+ ; _+_) +open import Cubical.Data.Nat + +infixl 30 _·₀_ +infixr 35 _⌣ₖ_ +infixr 35 _⌣_ + +--- This definition of ℕ-addition removes some unnecessary transports. +_+'_ : ℕ → ℕ → ℕ +zero +' b = b +suc a +' zero = suc a +suc a +' suc b = 2 + (a + b) + ++'≡+ : (n m : ℕ) → n +' m ≡ n + m ++'≡+ zero m = refl ++'≡+ (suc n) zero = cong suc (sym (+-comm n zero)) ++'≡+ (suc n) (suc m) = cong suc (sym (+-suc n m)) + +-- Cup product with one integer (K₀) argument +_·₀_ : {n : ℕ} (m : ℤ) → coHomK n → coHomK n +_·₀_ {n = n} (pos zero) x = 0ₖ _ +_·₀_ {n = n} (pos (suc m)) x = x +ₖ (pos m ·₀ x) +_·₀_ {n = n} (negsuc zero) x = -ₖ x +_·₀_ {n = n} (negsuc (suc m)) x = (negsuc m ·₀ x) -ₖ x + +·₀-0ₖ : {n : ℕ} (m : ℤ) → _·₀_ m (0ₖ n) ≡ 0ₖ n +·₀-0ₖ (pos zero) = refl +·₀-0ₖ (pos (suc n)) = cong (0ₖ _ +ₖ_) (·₀-0ₖ (pos n)) ∙ rUnitₖ _ (0ₖ _) +·₀-0ₖ (negsuc zero) = -0ₖ +·₀-0ₖ (negsuc (suc n)) = cong (λ x → x -ₖ (0ₖ _)) (·₀-0ₖ (negsuc n)) ∙ rCancelₖ _ (0ₖ _) + +-- Pointed version first (enables truncation elimination) +⌣ₖ∙ : (n m : ℕ) → coHomK n → coHomK-ptd m →∙ coHomK-ptd (n +' m) +fst (⌣ₖ∙ zero m a) b = a ·₀ b +snd (⌣ₖ∙ zero m a) = ·₀-0ₖ a +fst (⌣ₖ∙ (suc n) zero a) b = b ·₀ a +snd (⌣ₖ∙ (suc n) zero a) = refl +⌣ₖ∙ (suc n) (suc m) = trRec (isOfHLevel↑∙ (suc n) m) (cup n m) + where + cup : (n m : ℕ) → S₊ (suc n) → coHomK-ptd (suc m) →∙ coHomK-ptd (suc (suc (n + m))) + fst (cup zero m base) _ = 0ₖ _ + fst (cup zero m (loop i)) x = Kn→ΩKn+1 _ x i + fst (cup (suc n) m north) _ = 0ₖ _ + fst (cup (suc n) m south) _ = 0ₖ _ + fst (cup (suc n) m (merid a i)) x = Kn→ΩKn+1 _ (fst (cup n m a) x) i + snd (cup zero m base) = refl + snd (cup zero m (loop i)) k = Kn→ΩKn+10ₖ _ k i + snd (cup (suc n) m north) = refl + snd (cup (suc n) m south) = refl + snd (cup (suc n) m (merid a i)) k = (cong (Kn→ΩKn+1 _) (snd (cup n m a)) ∙ Kn→ΩKn+10ₖ _) k i + +-- Non pointed version +_⌣ₖ_ : {n m : ℕ} → coHomK n → coHomK m → coHomK (n +' m) +_⌣ₖ_ {n = n} {m = m} x y = fst (⌣ₖ∙ n m x) y + +-- Doubly pointed version +⌣ₖ∙∙ : (n m : ℕ) → coHomK-ptd n →∙ (coHomK-ptd m →∙ coHomK-ptd (n +' m) ∙) +fst (⌣ₖ∙∙ n m) = ⌣ₖ∙ n m +fst (snd (⌣ₖ∙∙ zero zero) i) x = 0 +fst (snd (⌣ₖ∙∙ zero (suc m)) i) x = 0ₖ _ +fst (snd (⌣ₖ∙∙ (suc n) zero) i) x = ·₀-0ₖ x i +fst (snd (⌣ₖ∙∙ (suc zero) (suc m)) i) x = 0ₖ _ +fst (snd (⌣ₖ∙∙ (suc (suc n)) (suc m)) i) x = 0ₖ _ +snd (snd (⌣ₖ∙∙ zero zero) i) = refl +snd (snd (⌣ₖ∙∙ zero (suc m)) i) = refl +snd (snd (⌣ₖ∙∙ (suc n) zero) i) = refl +snd (snd (⌣ₖ∙∙ (suc zero) (suc m)) i) = refl +snd (snd (⌣ₖ∙∙ (suc (suc n)) (suc m)) i) = refl + +-- Cup product +_⌣_ : ∀ {ℓ} {A : Type ℓ} {n m : ℕ} → coHom n A → coHom m A → coHom (n +' m) A +_⌣_ = sRec2 squash₂ λ f g → ∣ (λ x → f x ⌣ₖ g x) ∣₂ diff --git a/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda new file mode 100644 index 000000000..059b5af70 --- /dev/null +++ b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda @@ -0,0 +1,1075 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.RingStructure.GradedCommutativity where +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.ZCohomology.RingStructure.RingLaws +open import Cubical.ZCohomology.Properties + +open import Cubical.Homotopy.Loopspace + +open import Cubical.HITs.S1 hiding (_·_) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 + ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.Truncation + renaming (elim to trElim ; map to trMap ; map2 to trMap2 + ; rec to trRec ; elim3 to trElim3) + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.GroupoidLaws hiding (assoc) +open import Cubical.Foundations.Path +open import Cubical.Data.Int + renaming (_+_ to _ℤ+_ ; _·_ to _ℤ∙_ ; +Comm to +ℤ-comm + ; ·Comm to ∙-comm ; +Assoc to ℤ+-assoc ; -_ to -ℤ_) + hiding (_+'_ ; +'≡+) +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Sum + +private + variable + ℓ : Level + +private + natTranspLem : ∀ {A B : ℕ → Type} {n m : ℕ} (a : A n) + (f : (n : ℕ) → (a : A n) → B n) (p : n ≡ m) + → f m (subst A p a) ≡ subst B p (f n a) + natTranspLem {A = A} {B = B} a f p = sym (substCommSlice A B f p a) + ++'-comm : (n m : ℕ) → n +' m ≡ m +' n ++'-comm n m = +'≡+ n m ∙∙ +-comm n m ∙∙ sym (+'≡+ m n) + +private + transp0₁ : (n : ℕ) → subst coHomK (+'-comm 1 (suc n)) (0ₖ _) ≡ 0ₖ _ + transp0₁ zero = refl + transp0₁ (suc n) = refl + + transp0₂ : (n m : ℕ) → subst coHomK (+'-comm (suc (suc n)) (suc m)) (0ₖ _) ≡ 0ₖ _ + transp0₂ n zero = refl + transp0₂ n (suc m) = refl + +-- Recurring expressions +private + ΩKn+1→Ω²Kn+2 : {k : ℕ} → typ (Ω (coHomK-ptd k)) → typ ((Ω^ 2) (coHomK-ptd (suc k))) + ΩKn+1→Ω²Kn+2 x = sym (Kn→ΩKn+10ₖ _) ∙∙ cong (Kn→ΩKn+1 _) x ∙∙ Kn→ΩKn+10ₖ _ + + ΩKn+1→Ω²Kn+2' : {k : ℕ} → Kn→ΩKn+1 k (0ₖ k) ≡ Kn→ΩKn+1 k (0ₖ k) → typ ((Ω^ 2) (coHomK-ptd (suc k))) + ΩKn+1→Ω²Kn+2' p = sym (Kn→ΩKn+10ₖ _) ∙∙ p ∙∙ Kn→ΩKn+10ₖ _ + + Kn→Ω²Kn+2 : {k : ℕ} → coHomK k → typ ((Ω^ 2) (coHomK-ptd (2 + k))) + Kn→Ω²Kn+2 x = ΩKn+1→Ω²Kn+2 (Kn→ΩKn+1 _ x) + +-- Definition of of -ₖⁿ̇*ᵐ +-ₖ-helper : {k : ℕ} (n m : ℕ) + → isEvenT n ⊎ isOddT n → isEvenT m ⊎ isOddT m + → coHomKType k → coHomKType k +-ₖ-helper {k = zero} n m (inl x₁) q x = x +-ₖ-helper {k = zero} n m (inr x₁) (inl x₂) x = x +-ₖ-helper {k = zero} n m (inr x₁) (inr x₂) x = 0 - x +-ₖ-helper {k = suc zero} n m p q base = base +-ₖ-helper {k = suc zero} n m (inl x) q (loop i) = loop i +-ₖ-helper {k = suc zero} n m (inr x) (inl x₁) (loop i) = loop i +-ₖ-helper {k = suc zero} n m (inr x) (inr x₁) (loop i) = loop (~ i) +-ₖ-helper {k = suc (suc k)} n m p q north = north +-ₖ-helper {k = suc (suc k)} n m p q south = north +-ₖ-helper {k = suc (suc k)} n m (inl x) q (merid a i) = + (merid a ∙ sym (merid (ptSn (suc k)))) i +-ₖ-helper {k = suc (suc k)} n m (inr x) (inl x₁) (merid a i) = + (merid a ∙ sym (merid (ptSn (suc k)))) i +-ₖ-helper {k = suc (suc k)} n m (inr x) (inr x₁) (merid a i) = + (merid a ∙ sym (merid (ptSn (suc k)))) (~ i) + +-ₖ-gen : {k : ℕ} (n m : ℕ) + (p : isEvenT n ⊎ isOddT n) + (q : isEvenT m ⊎ isOddT m) + → coHomK k → coHomK k +-ₖ-gen {k = zero} = -ₖ-helper {k = zero} +-ₖ-gen {k = suc k} n m p q = trMap (-ₖ-helper {k = suc k} n m p q) + +-- -ₖⁿ̇*ᵐ +-ₖ^_·_ : {k : ℕ} (n m : ℕ) → coHomK k → coHomK k +-ₖ^_·_ {k = zero} n m = -ₖ-helper {k = zero} n m (evenOrOdd n) (evenOrOdd m) +-ₖ^_·_ {k = suc k} n m = trMap (-ₖ-helper n m (evenOrOdd n) (evenOrOdd m)) + +-- cohomology version +-ₕ^_·_ : {k : ℕ} {A : Type ℓ} (n m : ℕ) → coHom k A → coHom k A +-ₕ^_·_ n m = sMap λ f x → (-ₖ^ n · m) (f x) + +-- -ₖⁿ̇*ᵐ = -ₖ for n m odd +-ₖ-gen-inr≡-ₖ : {k : ℕ} (n m : ℕ) (p : _) (q : _) (x : coHomK k) + → -ₖ-gen n m (inr p) (inr q) x ≡ (-ₖ x) +-ₖ-gen-inr≡-ₖ {k = zero} n m p q _ = refl +-ₖ-gen-inr≡-ₖ {k = suc zero} n m p q = + trElim ((λ _ → isOfHLevelTruncPath)) + λ { base → refl + ; (loop i) → refl} +-ₖ-gen-inr≡-ₖ {k = suc (suc k)} n m p q = + trElim ((λ _ → isOfHLevelTruncPath)) + λ { north → refl + ; south → refl + ; (merid a i) k → ∣ symDistr (merid (ptSn _)) (sym (merid a)) (~ k) (~ i) ∣ₕ} + +-- -ₖⁿ̇*ᵐ x = x for n even +-ₖ-gen-inl-left : {k : ℕ} (n m : ℕ) (p : _) (q : _) (x : coHomK k) + → -ₖ-gen n m (inl p) q x ≡ x +-ₖ-gen-inl-left {k = zero} n m p q x = refl +-ₖ-gen-inl-left {k = suc zero} n m p q = + trElim (λ _ → isOfHLevelTruncPath) + λ { base → refl ; (loop i) → refl} +-ₖ-gen-inl-left {k = suc (suc k)} n m p q = + trElim (λ _ → isOfHLevelPath (4 + k) (isOfHLevelTrunc (4 + k)) _ _) + λ { north → refl + ; south → cong ∣_∣ₕ (merid (ptSn _)) + ; (merid a i) k → ∣ compPath-filler (merid a) (sym (merid (ptSn _))) (~ k) i ∣ₕ} + +-- -ₖⁿ̇*ᵐ x = x for m even +-ₖ-gen-inl-right : {k : ℕ} (n m : ℕ) (p : _) (q : _) (x : coHomK k) + → -ₖ-gen n m p (inl q) x ≡ x +-ₖ-gen-inl-right {k = zero} n m (inl x₁) q x = refl +-ₖ-gen-inl-right {k = zero} n m (inr x₁) q x = refl +-ₖ-gen-inl-right {k = suc zero} n m (inl x₁) q = + trElim (λ _ → isOfHLevelTruncPath) + λ { base → refl ; (loop i) → refl} +-ₖ-gen-inl-right {k = suc zero} n m (inr x₁) q = + trElim (λ _ → isOfHLevelTruncPath) + λ { base → refl ; (loop i) → refl} +-ₖ-gen-inl-right {k = suc (suc k)} n m (inl x₁) q = + trElim (λ _ → isOfHLevelTruncPath) + λ { north → refl + ; south → cong ∣_∣ₕ (merid (ptSn _)) + ; (merid a i) k → ∣ compPath-filler (merid a) (sym (merid (ptSn _))) (~ k) i ∣ₕ} +-ₖ-gen-inl-right {k = suc (suc k)} n m (inr x₁) q = + trElim (λ _ → isOfHLevelTruncPath) + λ { north → refl + ; south → cong ∣_∣ₕ (merid (ptSn _)) + ; (merid a i) k → ∣ compPath-filler (merid a) (sym (merid (ptSn _))) (~ k) i ∣ₕ} + +-- action of cong on -ₖⁿ̇*ᵐ +cong-ₖ-gen-inr : {k : ℕ} (n m : ℕ) (p : _) (q : _) (P : Path (coHomK (2 + k)) (0ₖ _) (0ₖ _)) + → cong (-ₖ-gen n m (inr p) (inr q)) P ≡ sym P +cong-ₖ-gen-inr {k = k} n m p q P = code≡sym (0ₖ _) P + where + code : (x : coHomK (2 + k)) → 0ₖ _ ≡ x → x ≡ 0ₖ _ + code = trElim (λ _ → isOfHLevelΠ (4 + k) λ _ → isOfHLevelTruncPath) + λ { north → cong (-ₖ-gen n m (inr p) (inr q)) + ; south P → cong ∣_∣ₕ (sym (merid (ptSn _))) ∙ (cong (-ₖ-gen n m (inr p) (inr q)) P) + ; (merid a i) → t a i} + where + t : (a : S₊ (suc k)) → PathP (λ i → 0ₖ (2 + k) ≡ ∣ merid a i ∣ₕ → ∣ merid a i ∣ₕ ≡ 0ₖ (2 + k)) + (cong (-ₖ-gen n m (inr p) (inr q))) + (λ P → cong ∣_∣ₕ (sym (merid (ptSn _))) ∙ (cong (-ₖ-gen n m (inr p) (inr q)) P)) + t a = toPathP (funExt λ P → cong (transport (λ i → ∣ merid a i ∣ ≡ 0ₖ (suc (suc k)))) + (cong (cong (-ₖ-gen n m (inr p) (inr q))) (λ i → (transp (λ j → 0ₖ (suc (suc k)) ≡ ∣ merid a (~ j ∧ ~ i) ∣) i + (compPath-filler P (λ j → ∣ merid a (~ j) ∣ₕ) i)))) + ∙∙ cong (transport (λ i → ∣ merid a i ∣ ≡ 0ₖ (suc (suc k)))) (congFunct (-ₖ-gen n m (inr p) (inr q)) P (sym (cong ∣_∣ₕ (merid a)))) + ∙∙ (λ j → transp (λ i → ∣ merid a (i ∨ j) ∣ ≡ 0ₖ (suc (suc k))) j + (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) + (cong (-ₖ-gen n m (inr p) (inr q)) P + ∙ cong (-ₖ-gen n m (inr p) (inr q)) (sym (cong ∣_∣ₕ (merid a)))) j)) + ∙∙ (λ i → sym (cong ∣_∣ₕ (merid a)) + ∙ isCommΩK (2 + k) (cong (-ₖ-gen n m (inr p) (inr q)) P) + (cong (-ₖ-gen n m (inr p) (inr q)) (sym (cong ∣_∣ₕ (merid a)))) i) + ∙∙ (λ j → (λ i → ∣ merid a (~ i ∨ j) ∣) + ∙ (cong ∣_∣ₕ (compPath-filler' (merid a) (sym (merid (ptSn _))) (~ j)) ∙ (λ i → -ₖ-gen n m (inr p) (inr q) (P i)))) + ∙ sym (lUnit _)) + + code≡sym : (x : coHomK (2 + k)) → (p : 0ₖ _ ≡ x) → code x p ≡ sym p + code≡sym x = J (λ x p → code x p ≡ sym p) refl + +cong-cong-ₖ-gen-inr : {k : ℕ} (n m : ℕ) (p : _) (q : _) + (P : Square (refl {x = 0ₖ (suc (suc k))}) refl refl refl) + → cong (cong (-ₖ-gen n m (inr p) (inr q))) P ≡ sym P +cong-cong-ₖ-gen-inr n m p q P = + rUnit _ + ∙∙ (λ k → (λ i → cong-ₖ-gen-inr n m p q refl (i ∧ k)) + ∙∙ (λ i → cong-ₖ-gen-inr n m p q (P i) k) + ∙∙ λ i → cong-ₖ-gen-inr n m p q refl (~ i ∧ k)) + ∙∙ (λ k → transportRefl refl k + ∙∙ cong sym P + ∙∙ transportRefl refl k) + ∙∙ sym (rUnit (cong sym P)) + ∙∙ sym (sym≡cong-sym P) + +Kn→ΩKn+1-ₖ' : {k : ℕ} (n m : ℕ) (p : _) (q : _) (x : coHomK k) + → Kn→ΩKn+1 k (-ₖ-gen n m (inr p) (inr q) x) ≡ sym (Kn→ΩKn+1 k x) +Kn→ΩKn+1-ₖ' n m p q x = cong (Kn→ΩKn+1 _) (-ₖ-gen-inr≡-ₖ n m p q x) ∙ Kn→ΩKn+1-ₖ _ x + +transpΩ² : {n m : ℕ} (p q : n ≡ m) → (P : _) + → transport (λ i → refl {x = 0ₖ (p i)} ≡ refl {x = 0ₖ (p i)}) P + ≡ transport (λ i → refl {x = 0ₖ (q i)} ≡ refl {x = 0ₖ (q i)}) P +transpΩ² p q P k = subst (λ n → refl {x = 0ₖ n} ≡ refl {x = 0ₖ n}) (isSetℕ _ _ p q k) P + +-- Some technical lemmas about Kn→Ω²Kn+2 and its interaction with -ₖⁿ̇*ᵐ and transports +-- TODO : Check if this can be cleaned up more by having more general lemmas +private + lem₁ : (n : ℕ) (a : _) + → (cong (cong (subst coHomK (+'-comm (suc zero) (suc (suc n))))) + (Kn→Ω²Kn+2 ∣ a ∣ₕ)) + ≡ ΩKn+1→Ω²Kn+2 + (sym (transp0₁ n) ∙∙ cong (subst coHomK (+'-comm (suc zero) (suc n))) (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ) ∙∙ transp0₁ n) + lem₁ zero a = + (λ k i j → transportRefl (Kn→Ω²Kn+2 ∣ a ∣ₕ i j) k) + ∙ cong ΩKn+1→Ω²Kn+2 λ k → rUnit (λ i → transportRefl (Kn→ΩKn+1 1 ∣ a ∣ i) (~ k)) k + lem₁ (suc n) a = + (λ k → transp (λ i → refl {x = 0ₖ (+'-comm 1 (suc (suc (suc n))) (i ∨ ~ k))} + ≡ refl {x = 0ₖ (+'-comm 1 (suc (suc (suc n))) (i ∨ ~ k))}) (~ k) + (λ i j → transp (λ i → coHomK (+'-comm 1 (suc (suc (suc n))) (i ∧ ~ k))) k + (Kn→Ω²Kn+2 ∣ a ∣ₕ i j))) + ∙∙ transpΩ² (+'-comm 1 (suc (suc (suc n)))) + (cong suc (+'-comm (suc zero) (suc (suc n)))) + (Kn→Ω²Kn+2 ∣ a ∣ₕ) + ∙∙ sym (natTranspLem {A = λ n → 0ₖ n ≡ 0ₖ n} + (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣) + (λ _ → ΩKn+1→Ω²Kn+2) + (+'-comm 1 (suc (suc n)))) + ∙∙ cong ΩKn+1→Ω²Kn+2 + (λ k → transp (λ i → 0ₖ (+'-comm (suc zero) (suc (suc n)) (i ∨ k)) + ≡ 0ₖ (+'-comm (suc zero) (suc (suc n)) (i ∨ k))) k + (λ i → transp (λ i → coHomK (+'-comm (suc zero) (suc (suc n)) (i ∧ k))) (~ k) + (Kn→ΩKn+1 _ ∣ a ∣ₕ i))) + ∙∙ cong ΩKn+1→Ω²Kn+2 (rUnit (cong (subst coHomK (+'-comm (suc zero) (suc (suc n)))) + (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣ₕ))) + + lem₂ : (n : ℕ) (a : _) (p : _) (q : _) + → (cong (cong (-ₖ-gen (suc (suc n)) (suc zero) p q + ∘ (subst coHomK (+'-comm 1 (suc (suc n)))))) + (Kn→Ω²Kn+2 (∣ a ∣ₕ))) + ≡ ΩKn+1→Ω²Kn+2 + (sym (transp0₁ n) + ∙∙ cong (subst coHomK (+'-comm (suc zero) (suc n))) + (cong (-ₖ-gen (suc (suc n)) (suc zero) p q) + (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ)) + ∙∙ transp0₁ n) + lem₂ n a (inl x) (inr y) = + (λ k i j → (-ₖ-gen-inl-left (suc (suc n)) 1 x (inr y) ( + subst coHomK (+'-comm 1 (suc (suc n))) + (Kn→Ω²Kn+2 ∣ a ∣ₕ i j))) k) + ∙∙ lem₁ n a + ∙∙ cong ΩKn+1→Ω²Kn+2 (cong (sym (transp0₁ n) ∙∙_∙∙ transp0₁ n) + λ k i → subst coHomK (+'-comm 1 (suc n)) + (-ₖ-gen-inl-left (suc (suc n)) 1 x (inr y) (Kn→ΩKn+1 (suc n) ∣ a ∣ i) (~ k))) + lem₂ n a (inr x) (inr y) = + cong-cong-ₖ-gen-inr (suc (suc n)) 1 x y + (cong + (cong + (subst coHomK (+'-comm 1 (suc (suc n))))) + (Kn→Ω²Kn+2 ∣ a ∣ₕ)) + ∙∙ cong sym (lem₁ n a) + ∙∙ λ k → ΩKn+1→Ω²Kn+2 + (sym (transp0₁ n) ∙∙ + cong (subst coHomK (+'-comm 1 (suc n))) + (cong-ₖ-gen-inr (suc (suc n)) 1 x y + (Kn→ΩKn+1 (suc n) ∣ a ∣) (~ k)) + ∙∙ transp0₁ n) + + lem₃ : (n m : ℕ) (q : _) (p : isEvenT (suc (suc n)) ⊎ isOddT (suc (suc n))) (x : _) + → (((sym (cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)) + ∙∙ (λ j → -ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (+'-comm (suc (suc m)) (suc n)) (Kn→ΩKn+1 _ x j))) + ∙∙ cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)))) + ≡ (Kn→ΩKn+1 _ (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (cong suc (+-comm (suc m) n)) x))) + lem₃ n m q p x = + sym (cong-∙∙ (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (sym (transp0₂ m n)) + (λ j → subst coHomK (+'-comm (suc (suc m)) (suc n)) (Kn→ΩKn+1 _ x j)) + (transp0₂ m n)) + ∙ h n m p q x + where + help : (n m : ℕ) (x : _) + → ((sym (transp0₂ m n)) + ∙∙ (λ j → subst coHomK (+'-comm (suc (suc m)) (suc n)) + (Kn→ΩKn+1 (suc (suc (m + n))) x j)) + ∙∙ transp0₂ m n) + ≡ Kn→ΩKn+1 (suc (n + suc m)) + (subst coHomK (cong suc (+-comm (suc m) n)) x) + help zero m x = + sym (rUnit _) + ∙∙ (λ k i → transp (λ i → coHomK (+'-comm (suc (suc m)) 1 (i ∨ k))) k + (Kn→ΩKn+1 _ + (transp (λ i → coHomK (predℕ (+'-comm (suc (suc m)) 1 (i ∧ k)))) (~ k) x) i)) + ∙∙ cong (Kn→ΩKn+1 _) + λ k → subst coHomK (isSetℕ _ _ (cong predℕ (+'-comm (suc (suc m)) 1)) + (cong suc (+-comm (suc m) zero)) k) x + help (suc n) m x = + sym (rUnit _) + ∙∙ ((λ k i → transp (λ i → coHomK (+'-comm (suc (suc m)) (suc (suc n)) (i ∨ k))) k + (Kn→ΩKn+1 _ + (transp (λ i → coHomK (predℕ (+'-comm (suc (suc m)) (suc (suc n)) (i ∧ k)))) (~ k) x) i))) + ∙∙ cong (Kn→ΩKn+1 _) + (λ k → subst coHomK (isSetℕ _ _ (cong predℕ (+'-comm (suc (suc m)) (suc (suc n)))) + (cong suc (+-comm (suc m) (suc n))) k) x) + + h : (n m : ℕ) (p : isEvenT (suc (suc n)) ⊎ isOddT (suc (suc n))) (q : _) (x : _) + → cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) + (sym (transp0₂ m n) + ∙∙ (λ j → subst coHomK (+'-comm (suc (suc m)) (suc n)) + (Kn→ΩKn+1 (suc (suc (m + n))) x j)) + ∙∙ transp0₂ m n) + ≡ Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (cong suc (+-comm (suc m) n)) x)) + h n m (inl p) (inl q) x = + (λ k → cong (-ₖ-gen (suc n) (suc (suc m)) + (isPropEvenOrOdd (suc n) (evenOrOdd (suc n)) (inr p) k) (inl q)) + (help n m x k)) + ∙∙ ((λ k i → -ₖ-gen-inl-right (suc n) (suc (suc m)) (inr p) q (help n m x i1 i) k)) + ∙∙ λ i → Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen-inl-right (suc n) (suc (suc m)) (isPropEvenOrOdd (suc n) (inr p) (evenOrOdd (suc n)) i) q + (subst coHomK (cong suc (+-comm (suc m) n)) x) (~ i)) + h n m (inl p) (inr q) x = + (λ k → cong (-ₖ-gen (suc n) (suc (suc m)) (isPropEvenOrOdd (suc n) (evenOrOdd (suc n)) (inr p) k) (inr q)) + (help n m x k)) + ∙∙ cong-ₖ-gen-inr (suc n) (suc (suc m)) p q (help n m x i1) + ∙∙ sym (Kn→ΩKn+1-ₖ' (suc n) (suc (suc m)) p q + (subst coHomK (λ i → suc (+-comm (suc m) n i)) x)) + ∙ λ k → Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen (suc n) (suc (suc m)) (isPropEvenOrOdd (suc n) (inr p) (evenOrOdd (suc n)) k) (inr q) + (subst coHomK (cong suc (+-comm (suc m) n)) x)) + h n m (inr p) (inl q) x = + (λ k → cong (-ₖ-gen (suc n) (suc (suc m)) (isPropEvenOrOdd (suc n) (evenOrOdd (suc n)) (inl p) k) (inl q)) + (help n m x k)) + ∙∙ (λ k i → -ₖ-gen-inl-left (suc n) (suc (suc m)) p (inl q) (help n m x i1 i) k) + ∙∙ λ k → Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen-inl-right (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (cong suc (+-comm (suc m) n)) x) (~ k)) + h n m (inr p) (inr q) x = + (λ k → cong (-ₖ-gen (suc n) (suc (suc m)) (isPropEvenOrOdd (suc n) (evenOrOdd (suc n)) (inl p) k) (inr q)) + (help n m x k)) + ∙∙ (λ k i → -ₖ-gen-inl-left (suc n) (suc (suc m)) p (inr q) (help n m x i1 i) k) + ∙∙ cong (Kn→ΩKn+1 (suc (n + suc m))) + (sym (-ₖ-gen-inl-left (suc n) (suc (suc m)) p (inr q) + (subst coHomK (λ i → suc (+-comm (suc m) n i)) x))) + ∙ λ k → Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen (suc n) (suc (suc m)) (isPropEvenOrOdd (suc n) (inl p) (evenOrOdd (suc n)) k) (inr q) + (subst coHomK (cong suc (+-comm (suc m) n)) x)) + + lem₄ : (n m : ℕ) (q : _) (p : isEvenT (suc (suc n)) ⊎ isOddT (suc (suc n))) (a : _) (b : _) + → cong (Kn→ΩKn+1 _) (((sym (cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)) + ∙∙ (λ j → -ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (+'-comm (suc (suc m)) (suc n)) + (_⌣ₖ_ {n = suc (suc m)} {m = (suc n)} ∣ merid b j ∣ₕ ∣ a ∣))) + ∙∙ cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)))) + ≡ cong (Kn→ΩKn+1 _) (Kn→ΩKn+1 _ (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (cong suc (+-comm (suc m) n)) + (_⌣ₖ_ {n = suc m} {m = (suc n)} ∣ b ∣ₕ ∣ a ∣)))) + lem₄ n m q p a b = cong (cong (Kn→ΩKn+1 _)) (lem₃ n m q p (_⌣ₖ_ {n = suc m} {m = (suc n)} ∣ b ∣ₕ ∣ a ∣)) + + lem₅ : (n m : ℕ) (p : _) (q : _) (a : _) (b : _) + → cong (cong (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + ∘ (subst coHomK (+'-comm (suc (suc m)) (suc (suc n)))))) + (ΩKn+1→Ω²Kn+2 (sym (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (transp0₂ n m)) + ∙∙ (λ i → -ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (+'-comm (suc (suc n)) (suc m)) + (_⌣ₖ_ {n = suc (suc n)} {m = suc m} ∣ merid a i ∣ₕ ∣ b ∣ₕ))) + ∙∙ cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (transp0₂ n m))) + ≡ Kn→Ω²Kn+2 (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (cong suc (sym (+-suc n m))) (_⌣ₖ_ {n = suc n} {m = suc m} ∣ a ∣ₕ ∣ b ∣ₕ)))) + lem₅ n m p q a b = + cong (cong (cong (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + ∘ (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))))))) + (cong (sym (Kn→ΩKn+10ₖ _) ∙∙_∙∙ Kn→ΩKn+10ₖ _) + (lem₄ m n p q b a)) + ∙ help p q (_⌣ₖ_ {n = suc n} {m = suc m} ∣ a ∣ ∣ b ∣) + where + annoying : (x : _) + → cong (cong (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))))) + (Kn→Ω²Kn+2 (subst coHomK (cong suc (+-comm (suc n) m)) x)) + ≡ Kn→Ω²Kn+2 (subst coHomK (cong suc (sym (+-suc n m))) x) + annoying x = + ((λ k → transp (λ i → refl {x = 0ₖ ((+'-comm (suc (suc m)) (suc (suc n))) (i ∨ ~ k))} + ≡ refl {x = 0ₖ ((+'-comm (suc (suc m)) (suc (suc n))) (i ∨ ~ k))}) (~ k) + + λ i j → transp (λ i → coHomK (+'-comm (suc (suc m)) (suc (suc n)) (i ∧ ~ k))) k + (Kn→Ω²Kn+2 (subst coHomK (cong suc (+-comm (suc n) m)) x) i j))) + ∙∙ cong (transport (λ i → refl {x = 0ₖ ((+'-comm (suc (suc m)) (suc (suc n))) i)} + ≡ refl {x = 0ₖ ((+'-comm (suc (suc m)) (suc (suc n))) i)})) + (natTranspLem {A = coHomK} x (λ _ → Kn→Ω²Kn+2) (cong suc (+-comm (suc n) m))) + ∙∙ sym (substComposite (λ n → refl {x = 0ₖ n} ≡ refl {x = 0ₖ n}) + (cong (suc ∘ suc ∘ suc) (+-comm (suc n) m)) (+'-comm (suc (suc m)) (suc (suc n))) + (Kn→Ω²Kn+2 x)) + ∙∙ (λ k → subst (λ n → refl {x = 0ₖ n} ≡ refl {x = 0ₖ n}) + (isSetℕ _ _ + (cong (suc ∘ suc ∘ suc) (+-comm (suc n) m) ∙ (+'-comm (suc (suc m)) (suc (suc n)))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n m))) k) + (Kn→Ω²Kn+2 x)) + ∙∙ sym (natTranspLem {A = coHomK} x (λ _ → Kn→Ω²Kn+2) (cong suc (sym (+-suc n m)))) + + help : (p : _) (q : _) (x : _) → + cong (cong (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + ∘ subst coHomK (+'-comm (suc (suc m)) (suc (suc n))))) + (Kn→Ω²Kn+2 (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (cong suc (+-comm (suc n) m)) x))) + ≡ Kn→Ω²Kn+2 + (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (cong suc (sym (+-suc n m))) x))) + help (inl x) (inl y) z = + (λ k i j → + -ₖ-gen-inl-right (suc (suc n)) (suc (suc m)) (inl x) y + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + ((ΩKn+1→Ω²Kn+2 + (Kn→ΩKn+1 _ (-ₖ-gen-inl-right (suc m) (suc (suc n)) (evenOrOdd (suc m)) x + (subst coHomK (cong suc (+-comm (suc n) m)) z) k))) i j)) k) + ∙∙ annoying z + ∙∙ cong Kn→Ω²Kn+2 + λ k → (-ₖ-gen-inl-right (suc (suc n)) (suc (suc m)) (inl x) y + (-ₖ-gen-inl-right (suc m) (suc (suc n)) (evenOrOdd (suc m)) x + (subst coHomK (cong suc (sym (+-suc n m))) z) (~ k)) (~ k)) + help (inl x) (inr y) z = + (λ k i j → + -ₖ-gen-inl-left (suc (suc n)) (suc (suc m)) x (inr y) + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + (Kn→Ω²Kn+2 (-ₖ-gen-inl-right (suc m) (suc (suc n)) (evenOrOdd (suc m)) x + (subst coHomK (cong suc (+-comm (suc n) m)) z) k) i j)) k) + ∙∙ annoying z + ∙∙ cong Kn→Ω²Kn+2 + (λ k → (-ₖ-gen-inl-left (suc (suc n)) (suc (suc m)) x (inr y) + (-ₖ-gen-inl-right (suc m) (suc (suc n)) (evenOrOdd (suc m)) x + (subst coHomK (cong suc (sym (+-suc n m))) z) (~ k)) (~ k))) + help (inr x) (inl y) z = + (λ k i j → -ₖ-gen-inl-right (suc (suc n)) (suc (suc m)) (inr x) y + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + (Kn→Ω²Kn+2 + (-ₖ-gen (suc m) (suc (suc n)) (isPropEvenOrOdd (suc m) (evenOrOdd (suc m)) (inr y) k) (inr x) + (subst coHomK (cong suc (+-comm (suc n) m)) z)) i j)) k) + ∙∙ cong (cong (cong (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))))) ∘ ΩKn+1→Ω²Kn+2) + (Kn→ΩKn+1-ₖ' (suc m) (suc (suc n)) y x + (subst coHomK (cong suc (+-comm (suc n) m)) z)) + ∙∙ cong sym (annoying z) + ∙∙ cong ΩKn+1→Ω²Kn+2 (sym (Kn→ΩKn+1-ₖ' (suc m) (suc (suc n)) y x + (subst coHomK (cong suc (sym (+-suc n m))) z))) + ∙∙ cong Kn→Ω²Kn+2 λ k → (-ₖ-gen-inl-right (suc (suc n)) (suc (suc m)) (inr x) y + (-ₖ-gen (suc m) (suc (suc n)) (isPropEvenOrOdd (suc m) (evenOrOdd (suc m)) (inr y) (~ k)) (inr x) + (subst coHomK (cong suc (sym (+-suc n m))) z)) (~ k)) + help (inr x) (inr y) z = + (λ k → cong-cong-ₖ-gen-inr (suc (suc n)) (suc (suc m)) x y + (λ i j → subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + (Kn→Ω²Kn+2 + (-ₖ-gen (suc m) (suc (suc n)) + (isPropEvenOrOdd (suc m) (evenOrOdd (suc m)) (inl y) k) (inr x) + (subst coHomK (cong suc (+-comm (suc n) m)) z)) i j)) k) + ∙∙ cong (sym ∘ cong (cong (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))))) ∘ Kn→Ω²Kn+2) + (-ₖ-gen-inl-left (suc m) (suc (suc n)) y (inr x) + (subst coHomK (cong suc (+-comm (suc n) m)) z)) + ∙∙ cong sym (annoying z) + ∙∙ cong (sym ∘ Kn→Ω²Kn+2) + (sym (-ₖ-gen-inl-left (suc m) (suc (suc n)) y (inr x) + (subst coHomK (cong suc (sym (+-suc n m))) z))) + ∙∙ cong ΩKn+1→Ω²Kn+2 + λ k → (Kn→ΩKn+1-ₖ' (suc (suc n)) (suc (suc m)) x y + (-ₖ-gen (suc m) (suc (suc n)) ( + isPropEvenOrOdd (suc m) (evenOrOdd (suc m)) (inl y) (~ k)) (inr x) + (subst coHomK (cong suc (sym (+-suc n m))) z))) (~ k) + + lem₆ : (n m : ℕ) (p : _) (q : _) (a : _) (b : _) + → flipSquare + (Kn→Ω²Kn+2 (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (cong suc (+-comm (suc m) n)) + (_⌣ₖ_ {n = suc m} {m = (suc n)} ∣ b ∣ₕ ∣ a ∣)))) + ≡ Kn→Ω²Kn+2 + (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (cong suc (sym (+-suc n m))) + (-ₖ-gen (suc n) (suc m) (evenOrOdd (suc n)) (evenOrOdd (suc m)) + (subst coHomK (+'-comm (suc m) (suc n)) (∣ b ∣ₕ ⌣ₖ ∣ a ∣ₕ)))))) + lem₆ n m p q a b = + sym (sym≡flipSquare (Kn→Ω²Kn+2 (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (cong suc (+-comm (suc m) n)) + (_⌣ₖ_ {n = suc m} {m = (suc n)} ∣ b ∣ₕ ∣ a ∣))))) + ∙ cong ΩKn+1→Ω²Kn+2 + (help₁ + (subst coHomK (cong suc (+-comm (suc m) n)) + (_⌣ₖ_ {n = suc m} {m = (suc n)} ∣ b ∣ ∣ a ∣)) p q + ∙ cong (Kn→ΩKn+1 _) (sym (help₂ (∣ b ∣ ⌣ₖ ∣ a ∣)))) + where + help₁ : (x : _) (p : _) (q : _) + → sym (Kn→ΩKn+1 (suc (n + suc m)) (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q x)) + ≡ Kn→ΩKn+1 (suc (n + suc m)) ((-ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (-ₖ-gen (suc n) (suc m) (evenOrOdd (suc n)) (evenOrOdd (suc m)) x)))) + help₁ z (inl x) (inl y) = + cong (λ x → sym (Kn→ΩKn+1 (suc (n + suc m)) x)) + (-ₖ-gen-inl-right (suc n) (suc (suc m)) (evenOrOdd (suc n)) y z) + ∙∙ sym (Kn→ΩKn+1-ₖ' (suc n) (suc m) x y z) + ∙∙ λ k → Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen-inl-right (suc (suc n)) (suc (suc m)) (inl x) y + (-ₖ-gen-inl-right (suc m) (suc (suc n)) (evenOrOdd (suc m)) x + (-ₖ-gen (suc n) (suc m) + (isPropEvenOrOdd (suc n) (inr x) (evenOrOdd (suc n)) k) + (isPropEvenOrOdd (suc m) (inr y) (evenOrOdd (suc m)) k) z) (~ k)) (~ k)) + help₁ z (inl x) (inr y) = + (λ k → sym (Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen (suc n) (suc (suc m)) + (isPropEvenOrOdd (suc n) (evenOrOdd (suc n)) (inr x) k) (inr y) z))) + ∙∙ cong sym (Kn→ΩKn+1-ₖ' (suc n) (suc (suc m)) x y z) + ∙∙ cong (Kn→ΩKn+1 (suc (n + suc m))) (sym (-ₖ-gen-inl-right (suc n) (suc m) (inr x) y z)) + ∙ λ k → Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen-inl-left (suc (suc n)) (suc (suc m)) x (inr y) + (-ₖ-gen-inl-right (suc m) (suc (suc n)) (evenOrOdd (suc m)) x + (-ₖ-gen (suc n) (suc m) + (isPropEvenOrOdd (suc n) (inr x) (evenOrOdd (suc n)) k) + (isPropEvenOrOdd (suc m) (inl y) (evenOrOdd (suc m)) k) z) (~ k)) (~ k)) + help₁ z (inr x) (inl y) = + cong (λ x → sym (Kn→ΩKn+1 (suc (n + suc m)) x)) + (-ₖ-gen-inl-right (suc n) (suc (suc m)) (evenOrOdd (suc n)) y z) + ∙∙ (λ k → Kn→ΩKn+1-ₖ' (suc m) (suc (suc n)) y x + (-ₖ-gen-inl-left (suc n) (suc m) x (inr y) z (~ k)) (~ k)) + ∙∙ λ k → Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen-inl-right (suc (suc n)) (suc (suc m)) (inr x) y + (-ₖ-gen (suc m) (suc (suc n)) + (isPropEvenOrOdd (suc m) (inr y) (evenOrOdd (suc m)) k) (inr x) + (-ₖ-gen (suc n) (suc m) + (isPropEvenOrOdd (suc n) (inl x) (evenOrOdd (suc n)) k) + (isPropEvenOrOdd (suc m) (inr y) (evenOrOdd (suc m)) k) z)) (~ k)) + help₁ z (inr x) (inr y) = + ((λ k → sym (Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen (suc n) (suc (suc m)) + (isPropEvenOrOdd (suc n) (evenOrOdd (suc n)) (inl x) k) (inr y) z)))) + ∙∙ cong sym (cong (Kn→ΩKn+1 (suc (n + suc m))) (-ₖ-gen-inl-left (suc n) (suc (suc m)) x (inr y) z)) + ∙∙ (λ k → sym (Kn→ΩKn+1 (suc (n + suc m)) + (-ₖ-gen-inl-left (suc m) (suc (suc n)) y (inr x) + (-ₖ-gen-inl-right (suc n) (suc m) (inl x) y z (~ k)) (~ k)))) + ∙ λ k → Kn→ΩKn+1-ₖ' (suc (suc n)) (suc (suc m)) x y + (-ₖ-gen (suc m) (suc (suc n)) (isPropEvenOrOdd (suc m) (inl y) (evenOrOdd (suc m)) k) (inr x) + (-ₖ-gen (suc n) (suc m) + (isPropEvenOrOdd (suc n) (inl x) (evenOrOdd (suc n)) k) + (isPropEvenOrOdd (suc m) (inl y) (evenOrOdd (suc m)) k) z)) (~ k) + + help₂ : (x : _) → + (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (cong suc (sym (+-suc n m))) + (-ₖ-gen (suc n) (suc m) (evenOrOdd (suc n)) (evenOrOdd (suc m)) + (subst coHomK (+'-comm (suc m) (suc n)) x))))) + ≡ -ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (-ₖ-gen (suc n) (suc m) (evenOrOdd (suc n)) (evenOrOdd (suc m)) + (subst coHomK (cong suc (+-comm (suc m) n)) x))) + help₂ x = + (λ k → -ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (transp (λ i → coHomK ((cong suc (sym (+-suc n m))) (i ∨ k))) k + (-ₖ-gen (suc n) (suc m) (evenOrOdd (suc n)) (evenOrOdd (suc m)) + (transp (λ i → coHomK ((cong suc (sym (+-suc n m))) (i ∧ k))) (~ k) + (subst coHomK (+'-comm (suc m) (suc n)) x)))))) + ∙ cong (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + ∘ -ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + ∘ -ₖ-gen (suc n) (suc m) (evenOrOdd (suc n)) (evenOrOdd (suc m))) + (sym (substComposite coHomK (+'-comm (suc m) (suc n)) ((cong suc (sym (+-suc n m)))) x) + ∙ λ k → subst coHomK (isSetℕ _ _ (+'-comm (suc m) (suc n) ∙ cong suc (sym (+-suc n m))) + ((cong suc (+-comm (suc m) n))) k) x) + + lem₇ : (n : ℕ) (a : _) (p : _) (q : _) + → ((λ i → Kn→ΩKn+1 _ (-ₖ-gen (suc n) (suc zero) (evenOrOdd (suc n)) (inr tt) + (transp0₁ n (~ i)))) + ∙∙ (λ i j → Kn→ΩKn+1 _ (-ₖ-gen (suc n) (suc zero) (evenOrOdd (suc n)) (inr tt) + (subst coHomK (+'-comm (suc zero) (suc n)) (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ i))) j) + ∙∙ (λ i → Kn→ΩKn+1 _ (-ₖ-gen (suc n) (suc zero) (evenOrOdd (suc n)) (inr tt) + (transp0₁ n i)))) + ≡ (cong (Kn→ΩKn+1 (suc (suc (n + zero)))) + (sym (transp0₁ n) + ∙∙ sym (cong (subst coHomK (+'-comm (suc zero) (suc n))) + (cong (-ₖ-gen (suc (suc n)) (suc zero) p q) (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ))) + ∙∙ transp0₁ n)) + lem₇ zero a (inl x) (inr tt) = + (λ k → rUnit ((cong (Kn→ΩKn+1 _) (cong-ₖ-gen-inr (suc zero) (suc zero) tt tt + (λ i → (subst coHomK (+'-comm (suc zero) (suc zero)) + (Kn→ΩKn+1 (suc zero) ∣ a ∣ₕ i))) k))) (~ k)) + ∙ λ k → ((cong (Kn→ΩKn+1 (suc (suc zero))) + (rUnit (λ i → subst coHomK (+'-comm (suc zero) (suc zero)) + (-ₖ-gen-inl-left (suc (suc zero)) (suc zero) tt (inr tt) + (Kn→ΩKn+1 (suc zero) ∣ a ∣ₕ (~ i)) (~ k))) k))) + lem₇ (suc n) a (inl x) (inr tt) = + ((λ k → rUnit (cong (Kn→ΩKn+1 _) + (λ i → -ₖ-gen (suc (suc n)) (suc zero) + (isPropEvenOrOdd n (evenOrOdd (suc (suc n))) (inr x) k) (inr tt) + (subst coHomK (+'-comm (suc zero) (suc (suc n))) + (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣ₕ i)))) (~ k))) + ∙∙ (((λ k → ((cong (Kn→ΩKn+1 _) (cong-ₖ-gen-inr (suc (suc n)) (suc zero) x tt + (λ i → (subst coHomK (+'-comm (suc zero) (suc (suc n))) + (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣ₕ i))) k)))))) + ∙∙ λ k → ((cong (Kn→ΩKn+1 (suc (suc (suc n + zero)))) + (rUnit (λ i → subst coHomK (+'-comm (suc zero) (suc (suc n))) + (-ₖ-gen-inl-left (suc (suc (suc n))) (suc zero) x (inr tt) + (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣ₕ (~ i)) (~ k))) k))) + lem₇ (suc n) a (inr x) (inr tt) = + (λ k → rUnit (λ i j → Kn→ΩKn+1 _ + (-ₖ-gen (suc (suc n)) (suc zero) + (isPropEvenOrOdd (suc (suc n)) (evenOrOdd (suc (suc n))) (inl x) k) (inr tt) + (subst coHomK (+'-comm (suc zero) (suc (suc n))) + (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣ₕ i))) j) (~ k)) + ∙∙ (λ k i j → Kn→ΩKn+1 _ (-ₖ-gen-inl-left (suc (suc n)) (suc zero) x (inr tt) + (subst coHomK (+'-comm (suc zero) (suc (suc n))) + (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣ₕ i)) k) j) + ∙∙ λ k → cong (Kn→ΩKn+1 _) + (rUnit (sym (cong (subst coHomK (+'-comm (suc zero) (suc (suc n)))) + (cong-ₖ-gen-inr (suc (suc (suc n))) (suc zero) x tt (Kn→ΩKn+1 (suc (suc n)) ∣ a ∣ₕ) (~ k)))) k) + +-- ∣ a ∣ ⌣ₖ ∣ b ∣ ≡ -ₖⁿ*ᵐ (∣ b ∣ ⌣ₖ ∣ a ∣) for n ≥ 1, m = 1 +gradedComm-elimCase-left : (n : ℕ) (p : _) (q : _) (a : S₊ (suc n)) (b : S¹) → + (_⌣ₖ_ {n = suc n} {m = (suc zero)} ∣ a ∣ₕ ∣ b ∣ₕ) + ≡ (-ₖ-gen (suc n) (suc zero) p q) + (subst coHomK (+'-comm (suc zero) (suc n)) + (_⌣ₖ_ {n = suc zero} {m = suc n} ∣ b ∣ₕ ∣ a ∣ₕ)) +gradedComm-elimCase-left zero (inr tt) (inr tt) a b = + proof a b + ∙ cong (-ₖ-gen 1 1 (inr tt) (inr tt)) + (sym (transportRefl ((_⌣ₖ_ {n = suc zero} {m = suc zero} ∣ b ∣ ∣ a ∣)))) + where + help : flipSquare (ΩKn+1→Ω²Kn+2' (λ j i → _⌣ₖ_ {n = suc zero} {m = suc zero} ∣ loop i ∣ₕ ∣ loop j ∣ₕ)) ≡ + cong (cong (-ₖ-gen 1 1 (inr tt) (inr tt))) + (ΩKn+1→Ω²Kn+2' (λ i j → _⌣ₖ_ {n = suc zero} {m = suc zero} ∣ loop j ∣ₕ ∣ loop i ∣ₕ)) + help = sym (sym≡flipSquare _) + ∙ sym (cong-cong-ₖ-gen-inr 1 1 tt tt + (ΩKn+1→Ω²Kn+2' (λ i j → _⌣ₖ_ {n = suc zero} {m = suc zero} ∣ loop j ∣ ∣ loop i ∣))) + + proof : (a b : S¹) → _⌣ₖ_ {n = suc zero} {m = suc zero} ∣ a ∣ₕ ∣ b ∣ₕ ≡ + -ₖ-gen 1 1 (inr tt) (inr tt) (_⌣ₖ_ {n = suc zero} {m = suc zero} ∣ b ∣ ∣ a ∣) + proof base base = refl + proof base (loop i) k = -ₖ-gen 1 1 (inr tt) (inr tt) (Kn→ΩKn+10ₖ _ (~ k) i) + proof (loop i) base k = Kn→ΩKn+10ₖ _ k i + proof (loop i) (loop j) k = + hcomp (λ r → λ { (i = i0) → -ₖ-gen 1 1 (inr tt) (inr tt) (Kn→ΩKn+10ₖ _ (~ k ∨ ~ r) j) + ; (i = i1) → -ₖ-gen 1 1 (inr tt) (inr tt) (Kn→ΩKn+10ₖ _ (~ k ∨ ~ r) j) + ; (j = i0) → Kn→ΩKn+10ₖ _ (k ∨ ~ r) i + ; (j = i1) → Kn→ΩKn+10ₖ _ (k ∨ ~ r) i + ; (k = i0) → doubleCompPath-filler + (sym (Kn→ΩKn+10ₖ _)) + (λ j i → _⌣ₖ_ {n = suc zero} {m = suc zero} ∣ loop i ∣ₕ ∣ loop j ∣ₕ) + (Kn→ΩKn+10ₖ _) (~ r) j i + ; (k = i1) → (-ₖ-gen 1 1 (inr tt) (inr tt) + (doubleCompPath-filler + (sym (Kn→ΩKn+10ₖ _)) + (λ i j → _⌣ₖ_ {n = suc zero} {m = suc zero} ∣ loop j ∣ₕ ∣ loop i ∣ₕ) + (Kn→ΩKn+10ₖ _) (~ r) i j))}) + (help k i j) +gradedComm-elimCase-left (suc n) p q north b = + cong (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + (sym (⌣ₖ-0ₖ _ (suc (suc n)) ∣ b ∣ₕ)) +gradedComm-elimCase-left (suc n) p q south b = + cong (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + ((sym (⌣ₖ-0ₖ _ (suc (suc n)) ∣ b ∣ₕ)) ∙ λ i → ∣ b ∣ ⌣ₖ ∣ merid (ptSn (suc n)) i ∣ₕ) +gradedComm-elimCase-left (suc n) p q (merid a i) base k = + hcomp (λ j → λ {(i = i0) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + (0ₖ _) + ; (i = i1) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + (compPath-filler (sym (⌣ₖ-0ₖ _ (suc (suc n)) ∣ base ∣ₕ)) + (λ i → ∣ base ∣ ⌣ₖ ∣ merid a i ∣ₕ) j k) + ; (k = i0) → _⌣ₖ_ {n = suc (suc n)} {m = suc zero} ∣ merid a i ∣ₕ ∣ base ∣ₕ + ; (k = i1) → -ₖ-gen (suc (suc n)) 1 p q + (subst coHomK (+'-comm 1 (suc (suc n))) + (∣ base ∣ₕ ⌣ₖ ∣ merid a i ∣ₕ))}) + (hcomp (λ j → λ {(i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (k = i0) → (sym (Kn→ΩKn+10ₖ _) + ∙ (λ j → Kn→ΩKn+1 _ + (sym (gradedComm-elimCase-left n (evenOrOdd (suc n)) (inr tt) a base + ∙ cong (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt)) (transp0₁ n)) j))) j i + ; (k = i1) → ∣ north ∣}) + ∣ north ∣) +gradedComm-elimCase-left (suc n) p q (merid a i) (loop j) k = + hcomp (λ r → + λ { (i = i0) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + (sym (⌣ₖ-0ₖ _ (suc (suc n)) ∣ (loop j) ∣ₕ) k) + ; (i = i1) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + (compPath-filler (sym (⌣ₖ-0ₖ _ (suc (suc n)) ∣ (loop j) ∣ₕ)) + (λ i → ∣ loop j ∣ ⌣ₖ ∣ merid (ptSn (suc n)) i ∣ₕ) r k) + ; (k = i0) → _⌣ₖ_ {n = suc (suc n)} {m = suc zero} ∣ merid a i ∣ₕ ∣ loop j ∣ₕ + ; (k = i1) → -ₖ-gen (suc (suc n)) 1 p q + (subst coHomK (+'-comm 1 (suc (suc n))) + (_⌣ₖ_ {n = suc zero} {m = suc (suc n)} ∣ loop j ∣ₕ + ∣ compPath-filler (merid a) (sym (merid (ptSn (suc n)))) (~ r) i ∣ₕ))}) + (hcomp (λ r → + λ { (i = i0) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + ∣ rCancel (merid (ptSn (suc (suc n)))) (~ k ∧ r) j ∣ + ; (i = i1) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + (subst coHomK (+'-comm 1 (suc (suc n))))) + ∣ rCancel (merid (ptSn (suc (suc n)))) (~ k ∧ r) j ∣ + ; (k = i0) → help₂ r i j + ; (k = i1) → -ₖ-gen (suc (suc n)) 1 p q + (subst coHomK (+'-comm 1 (suc (suc n))) + (_⌣ₖ_ {n = suc zero} {m = suc (suc n)} ∣ loop j ∣ₕ + (Kn→ΩKn+1 _ ∣ a ∣ₕ i)))}) + (-ₖ-gen (suc (suc n)) 1 p q + (subst coHomK (+'-comm 1 (suc (suc n))) + (_⌣ₖ_ {n = suc zero} {m = suc (suc n)} ∣ loop j ∣ₕ + (Kn→ΩKn+1 _ ∣ a ∣ₕ i))))) + where + P : Path _ (Kn→ΩKn+1 (suc (suc (n + 0))) (0ₖ _)) + (Kn→ΩKn+1 (suc (suc (n + 0))) (_⌣ₖ_ {n = (suc n)} {m = suc zero} ∣ a ∣ ∣ base ∣)) + P i = Kn→ΩKn+1 (suc (suc (n + 0))) + ((sym (gradedComm-elimCase-left n (evenOrOdd (suc n)) (inr tt) a base + ∙ cong (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt)) (transp0₁ n)) i)) + + help₁ : (P ∙∙ ((λ i j → _⌣ₖ_ {n = suc (suc n)} {m = suc zero} ∣ merid a j ∣ₕ ∣ loop i ∣ₕ)) ∙∙ sym P) + ≡ ((λ i → Kn→ΩKn+1 _ (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt) + (transp0₁ n (~ i)))) + ∙∙ (λ i j → Kn→ΩKn+1 _ (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt) + (subst coHomK (+'-comm 1 (suc n)) (∣ loop i ∣ₕ ⌣ₖ ∣ a ∣ₕ))) j) + ∙∙ (λ i → Kn→ΩKn+1 _ (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt) + (transp0₁ n i)))) + help₁ k i j = + ((λ i → (Kn→ΩKn+1 (suc (suc (n + 0)))) + (compPath-filler' + ((gradedComm-elimCase-left n (evenOrOdd (suc n)) (inr tt) a base)) + (cong (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt)) + (transp0₁ n)) (~ k) (~ i))) + ∙∙ (λ i j → (Kn→ΩKn+1 _ + (gradedComm-elimCase-left n (evenOrOdd (suc n)) (inr tt) a (loop i) k) j)) + ∙∙ λ i → (Kn→ΩKn+1 (suc (suc (n + 0)))) + (compPath-filler' + ((gradedComm-elimCase-left n (evenOrOdd (suc n)) (inr tt) a base)) + (cong (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt)) + (transp0₁ n)) (~ k) i)) i j + + help₂ : I → I → I → coHomK _ + help₂ r i j = + hcomp (λ k → + λ { (i = i0) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + subst coHomK (+'-comm 1 (suc (suc n)))) + ∣ rCancel (merid (ptSn (suc (suc n)))) (~ k ∨ r) j ∣ + ; (i = i1) → (-ₖ-gen (suc (suc n)) 1 p q ∘ + subst coHomK (+'-comm 1 (suc (suc n)))) + ∣ rCancel (merid (ptSn (suc (suc n)))) (~ k ∨ r) j ∣ + ; (j = i0) → compPath-filler (sym (Kn→ΩKn+10ₖ (suc (suc (n + 0))))) + P k r i + ; (j = i1) → compPath-filler (sym (Kn→ΩKn+10ₖ (suc (suc (n + 0))))) + P k r i + ; (r = i0) → -ₖ-gen (suc (suc n)) 1 p q + (subst coHomK (+'-comm 1 (suc (suc n))) + (doubleCompPath-filler (sym (Kn→ΩKn+10ₖ _)) + (λ i j → _⌣ₖ_ {n = suc zero} {m = suc (suc n)} ∣ loop j ∣ₕ + (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ i)) (Kn→ΩKn+10ₖ _) (~ k) i j)) + ; (r = i1) → doubleCompPath-filler P + (λ i j → _⌣ₖ_ {n = suc (suc n)} {m = suc zero} ∣ merid a j ∣ₕ ∣ loop i ∣ₕ) + (sym P) (~ k) j i}) + (hcomp (λ k → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (j = i0) → (Kn→ΩKn+10ₖ (suc (suc (n + 0)))) (~ r) i + ; (j = i1) → (Kn→ΩKn+10ₖ (suc (suc (n + 0)))) (~ r) i + ; (r = i0) → lem₂ n a p q (~ k) i j + ; (r = i1) → help₁ (~ k) j i}) + (hcomp (λ k → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (j = i0) → Kn→ΩKn+10ₖ (suc (suc (n + 0))) (~ r) i + ; (j = i1) → Kn→ΩKn+10ₖ (suc (suc (n + 0))) (~ r) i + ; (r = i0) → flipSquare≡cong-sym (flipSquare (ΩKn+1→Ω²Kn+2 + (sym (transp0₁ n) + ∙∙ cong (subst coHomK (+'-comm 1 (suc n))) + (cong (-ₖ-gen (suc (suc n)) 1 p q) + (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ)) + ∙∙ transp0₁ n))) (~ k) i j + ; (r = i1) → ((λ i → Kn→ΩKn+1 _ (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt) + (transp0₁ n (~ i)))) + ∙∙ (λ i j → Kn→ΩKn+1 _ (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt) + (subst coHomK (+'-comm 1 (suc n)) (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ i))) j) + ∙∙ (λ i → Kn→ΩKn+1 _ (-ₖ-gen (suc n) 1 (evenOrOdd (suc n)) (inr tt) + (transp0₁ n i)))) j i}) + (hcomp (λ k → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (j = i0) → Kn→ΩKn+10ₖ (suc (suc (n + 0))) (~ r ∧ k) i + ; (j = i1) → Kn→ΩKn+10ₖ (suc (suc (n + 0))) (~ r ∧ k) i + ; (r = i0) → doubleCompPath-filler + (sym (Kn→ΩKn+10ₖ _)) + (cong (Kn→ΩKn+1 (suc (suc (n + 0)))) + (sym (transp0₁ n) + ∙∙ sym (cong (subst coHomK (+'-comm 1 (suc n))) + (cong (-ₖ-gen (suc (suc n)) 1 p q) + (Kn→ΩKn+1 (suc n) ∣ a ∣ₕ))) + ∙∙ transp0₁ n)) + (Kn→ΩKn+10ₖ _) k j i + ; (r = i1) → lem₇ n a p q (~ k) j i}) + (lem₇ n a p q i1 j i)))) + +-- ∣ a ∣ ⌣ₖ ∣ b ∣ ≡ -ₖⁿ*ᵐ (∣ b ∣ ⌣ₖ ∣ a ∣) for all n, m ≥ 1 +gradedComm-elimCase : (k n m : ℕ) (term : n + m ≡ k) (p : _) (q : _) (a : _) (b : _) → + (_⌣ₖ_ {n = suc n} {m = (suc m)} ∣ a ∣ₕ ∣ b ∣ₕ) + ≡ (-ₖ-gen (suc n) (suc m) p q) + (subst coHomK (+'-comm (suc m) (suc n)) + (_⌣ₖ_ {n = suc m} {m = suc n} ∣ b ∣ₕ ∣ a ∣ₕ)) +gradedComm-elimCase k zero zero term p q a b = gradedComm-elimCase-left zero p q a b +gradedComm-elimCase k zero (suc m) term (inr tt) q a b = + help q + ∙ sym (cong (-ₖ-gen 1 (suc (suc m)) (inr tt) q + ∘ (subst coHomK (+'-comm (suc (suc m)) 1))) + (gradedComm-elimCase-left (suc m) q (inr tt) b a)) + where + help : (q : _) → ∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ ≡ + -ₖ-gen 1 (suc (suc m)) (inr tt) q + (subst coHomK (+'-comm (suc (suc m)) 1) + (-ₖ-gen (suc (suc m)) 1 q (inr tt) + (subst coHomK (+'-comm 1 (suc (suc m))) (∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ)))) + help (inl x) = + (sym (transportRefl _) + ∙ (λ i → subst coHomK (isSetℕ _ _ refl (+'-comm 1 (suc (suc m)) ∙ +'-comm (suc (suc m)) 1) i) + (∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ))) + ∙∙ substComposite coHomK + (+'-comm 1 (suc (suc m))) + (+'-comm (suc (suc m)) 1) + ((∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ)) + ∙∙ λ i → -ₖ-gen-inl-right (suc zero) (suc (suc m)) (inr tt) x + ((subst coHomK (+'-comm (suc (suc m)) 1) + (-ₖ-gen-inl-left (suc (suc m)) 1 x (inr tt) + (subst coHomK (+'-comm 1 (suc (suc m))) (∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ)) (~ i)))) (~ i) + help (inr x) = + (sym (transportRefl _) + ∙∙ (λ k → subst coHomK (isSetℕ _ _ refl (+'-comm 1 (suc (suc m)) ∙ +'-comm (suc (suc m)) 1) k) (∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ)) + ∙∙ sym (-ₖ^2 (subst coHomK (+'-comm 1 (suc (suc m)) ∙ +'-comm (suc (suc m)) 1) (∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ)))) + ∙∙ (λ i → -ₖ-gen-inr≡-ₖ 1 (suc (suc m)) tt x + (-ₖ-gen-inr≡-ₖ (suc (suc m)) 1 x tt + (substComposite coHomK (+'-comm 1 (suc (suc m))) (+'-comm (suc (suc m)) 1) (∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ) i) + (~ i)) (~ i)) + ∙∙ λ i → (-ₖ-gen 1 (suc (suc m)) (inr tt) (inr x) + (transp (λ j → coHomK ((+'-comm (suc (suc m)) 1) (j ∨ ~ i))) (~ i) + (-ₖ-gen (suc (suc m)) 1 (inr x) (inr tt) + (transp (λ j → coHomK ((+'-comm (suc (suc m)) 1) (j ∧ ~ i))) i + ((subst coHomK (+'-comm 1 (suc (suc m))) (∣ a ∣ₕ ⌣ₖ ∣ b ∣ₕ))))))) +gradedComm-elimCase k (suc n) zero term p q a b = + gradedComm-elimCase-left (suc n) p q a b +gradedComm-elimCase zero (suc n) (suc m) term p q a b = + ⊥-rec (snotz (sym (+-suc n m) ∙ cong predℕ term)) +gradedComm-elimCase (suc zero) (suc n) (suc m) term p q a b = + ⊥-rec (snotz (sym (+-suc n m) ∙ cong predℕ term)) +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q north north = refl +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q north south = refl +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q north (merid a i) r = + -ₖ-gen (suc (suc n)) (suc (suc m)) p q ( + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n)))) + ((sym (Kn→ΩKn+10ₖ _) + ∙ cong (Kn→ΩKn+1 _) + (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (sym (transp0₂ n m)) + ∙ sym (gradedComm-elimCase (suc k) m (suc n) (+-suc m n ∙ +-comm (suc m) n ∙ cong predℕ term) + (evenOrOdd (suc m)) p a north))) r i)) +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q south north = refl +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q south south = refl +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q south (merid a i) r = + -ₖ-gen (suc (suc n)) (suc (suc m)) p q ( + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n)))) + ((sym (Kn→ΩKn+10ₖ _) + ∙ cong (Kn→ΩKn+1 _) + (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (sym (transp0₂ n m)) + ∙ sym (gradedComm-elimCase (suc k) m (suc n) (+-suc m n ∙ +-comm (suc m) n ∙ cong predℕ term) + (evenOrOdd (suc m)) p a south))) r i)) +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q (merid a i) north r = + (cong (Kn→ΩKn+1 (suc (suc (n + suc m)))) + (gradedComm-elimCase (suc k) n (suc m) (cong predℕ term) (evenOrOdd (suc n)) q a north + ∙ cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)) + ∙' Kn→ΩKn+10ₖ _) r i +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q (merid a i) south r = + (cong (Kn→ΩKn+1 (suc (suc (n + suc m)))) + (gradedComm-elimCase (suc k) n (suc m) (cong predℕ term) (evenOrOdd (suc n)) q a south + ∙ cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)) + ∙' Kn→ΩKn+10ₖ _) r i +gradedComm-elimCase (suc (suc k)) (suc n) (suc m) term p q (merid a i) (merid b j) r = + hcomp (λ l → + λ { (i = i0) → -ₖ-gen (suc (suc n)) (suc (suc m)) p q ( + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n)))) + ((compPath-filler (sym (Kn→ΩKn+10ₖ _)) + (cong (Kn→ΩKn+1 _) + (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (sym (transp0₂ n m)) + ∙ sym (gradedComm-elimCase (suc k) m (suc n) (+-suc m n ∙ +-comm (suc m) n ∙ cong predℕ term) + (evenOrOdd (suc m)) p b north))) l r j))) + ; (i = i1) → -ₖ-gen (suc (suc n)) (suc (suc m)) p q ( + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n)))) + ((compPath-filler (sym (Kn→ΩKn+10ₖ _)) + (cong (Kn→ΩKn+1 _) + (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (sym (transp0₂ n m)) + ∙ sym (gradedComm-elimCase (suc k) m (suc n) (+-suc m n ∙ +-comm (suc m) n ∙ cong predℕ term) + (evenOrOdd (suc m)) p b south))) l r j))) + ; (r = i0) → help₂ l i j + ; (r = i1) → -ₖ-gen (suc (suc n)) (suc (suc m)) p q + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + (help₁ l i j))}) + (hcomp (λ l → + λ { (i = i0) → -ₖ-gen (suc (suc n)) (suc (suc m)) p q + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + (Kn→ΩKn+10ₖ _ (~ r ∨ ~ l) j)) + ; (i = i1) → -ₖ-gen (suc (suc n)) (suc (suc m)) p q + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + (Kn→ΩKn+10ₖ _ (~ r ∨ ~ l) j)) + ; (j = i0) → Kn→ΩKn+10ₖ _ r i + ; (j = i1) → Kn→ΩKn+10ₖ _ r i + ; (r = i0) → lem₄ n m q p a b (~ l) j i + ; (r = i1) → -ₖ-gen (suc (suc n)) (suc (suc m)) p q + (subst coHomK (+'-comm (suc (suc m)) (suc (suc n))) + (doubleCompPath-filler + (sym (Kn→ΩKn+10ₖ _)) + (λ i j → Kn→ΩKn+1 _ ((sym (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (transp0₂ n m)) + ∙∙ (λ i → -ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (+'-comm (suc (suc n)) (suc m)) + (_⌣ₖ_ {n = suc (suc n)} {m = suc m} ∣ merid a i ∣ₕ ∣ b ∣ₕ))) + ∙∙ cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (transp0₂ n m)) i) j) + (Kn→ΩKn+10ₖ _) (~ l) i j))}) + (hcomp (λ l → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (j = i0) → Kn→ΩKn+10ₖ _ r i + ; (j = i1) → Kn→ΩKn+10ₖ _ r i + ; (r = i0) → lem₄ n m q p a b i1 j i + ; (r = i1) → lem₅ n m p q a b (~ l) i j}) + (hcomp (λ l → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (j = i0) → Kn→ΩKn+10ₖ _ (r ∨ ~ l) i + ; (j = i1) → Kn→ΩKn+10ₖ _ (r ∨ ~ l) i + ; (r = i0) → doubleCompPath-filler + (sym (Kn→ΩKn+10ₖ _)) + (lem₄ n m q p a b i1) + (Kn→ΩKn+10ₖ _) (~ l) j i + ; (r = i1) → Kn→Ω²Kn+2 (-ₖ-gen (suc (suc n)) (suc (suc m)) p q + (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (cong suc (sym (+-suc n m))) + (gradedComm-elimCase k n m + (+-comm n m ∙∙ cong predℕ (+-comm (suc m) n) ∙∙ cong (predℕ ∘ predℕ) term) + (evenOrOdd (suc n)) (evenOrOdd (suc m)) a b (~ l))))) i j}) + (lem₆ n m p q a b r i j)))) + where + help₁ : I → I → I → coHomK _ + help₁ l i j = + Kn→ΩKn+1 _ + (hcomp (λ r + → λ { (i = i0) → compPath-filler' (cong ((-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p)) (sym (transp0₂ n m))) + (sym (gradedComm-elimCase (suc k) m (suc n) (+-suc m n ∙ +-comm (suc m) n ∙ cong predℕ term) + (evenOrOdd (suc m)) p b north)) r l + ; (i = i1) → compPath-filler' (cong ((-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p)) (sym (transp0₂ n m))) + (sym (gradedComm-elimCase (suc k) m (suc n) (+-suc m n ∙ +-comm (suc m) n ∙ cong predℕ term) + (evenOrOdd (suc m)) p b south)) r l + ; (l = i0) → doubleCompPath-filler (sym (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (transp0₂ n m))) + (λ i → -ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p + (subst coHomK (+'-comm (suc (suc n)) (suc m)) + (_⌣ₖ_ {n = suc (suc n)} {m = suc m} ∣ merid a i ∣ₕ ∣ b ∣ₕ))) + (cong (-ₖ-gen (suc m) (suc (suc n)) (evenOrOdd (suc m)) p) (transp0₂ n m)) r i + ; (l = i1) → _⌣ₖ_ {n = suc m} {m = suc (suc n)} ∣ b ∣ₕ ∣ merid a i ∣ₕ}) + (gradedComm-elimCase (suc k) m (suc n) (+-suc m n ∙ +-comm (suc m) n ∙ cong predℕ term) + (evenOrOdd (suc m)) p b (merid a i) (~ l))) j + + help₂ : I → I → I → coHomK _ + help₂ l i j = + hcomp (λ r → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (j = i0) → + Kn→ΩKn+1 (suc (suc (n + suc m))) + (compPath-filler (gradedComm-elimCase (suc k) n (suc m) + (cong predℕ term) (evenOrOdd (suc n)) q a north) + (cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)) r (~ l)) i + ; (j = i1) → + Kn→ΩKn+1 (suc (suc (n + suc m))) + (compPath-filler (gradedComm-elimCase (suc k) n (suc m) + (cong predℕ term) (evenOrOdd (suc n)) q a south) + (cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)) r (~ l)) i + ; (l = i0) → + Kn→ΩKn+1 _ + (doubleCompPath-filler (sym (cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n))) + (λ j → -ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q + (subst coHomK (+'-comm (suc (suc m)) (suc n)) + (_⌣ₖ_ {n = suc (suc m)} {m = (suc n)} ∣ merid b j ∣ₕ ∣ a ∣))) + (cong (-ₖ-gen (suc n) (suc (suc m)) (evenOrOdd (suc n)) q) (transp0₂ m n)) r j) i + ; (l = i1) → Kn→ΩKn+1 _ (_⌣ₖ_ {n = (suc n)} {m = suc (suc m)} ∣ a ∣ ∣ merid b j ∣ₕ) i}) + (hcomp (λ r → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → ∣ north ∣ + ; (j = i0) → Kn→ΩKn+1 (suc (suc (n + suc m))) + (gradedComm-elimCase (suc k) n (suc m) + (cong predℕ term) (evenOrOdd (suc n)) q a north (~ l ∨ ~ r)) i + ; (j = i1) → Kn→ΩKn+1 (suc (suc (n + suc m))) + (gradedComm-elimCase (suc k) n (suc m) + (cong predℕ term) (evenOrOdd (suc n)) q a south (~ l ∨ ~ r)) i + ; (l = i0) → Kn→ΩKn+1 (suc (suc (n + suc m))) + (gradedComm-elimCase (suc k) n (suc m) + (cong predℕ term) (evenOrOdd (suc n)) q a (merid b j) i1) i + ; (l = i1) → Kn→ΩKn+1 _ (gradedComm-elimCase (suc k) n (suc m) (cong predℕ term) + (evenOrOdd (suc n)) q a (merid b j) (~ r)) i}) + (Kn→ΩKn+1 (suc (suc (n + suc m))) + (gradedComm-elimCase (suc k) n (suc m) (cong predℕ term) + (evenOrOdd (suc n)) q a (merid b j) i1) i)) + +private + coherence-transp : (n m : ℕ) (p : _) (q : _) + → -ₖ-gen (suc n) (suc m) p q + (subst coHomK (+'-comm (suc m) (suc n)) (0ₖ (suc m +' suc n))) ≡ 0ₖ _ + coherence-transp zero zero p q = refl + coherence-transp zero (suc m) p q = refl + coherence-transp (suc n) zero p q = refl + coherence-transp (suc n) (suc m) p q = refl + +gradedComm-⌣ₖ∙ : (n m : ℕ) (p : _) (q : _) (a : _) + → ⌣ₖ∙ (suc n) (suc m) a + ≡ ((λ b → -ₖ-gen (suc n) (suc m) p q (subst coHomK (+'-comm (suc m) (suc n)) (b ⌣ₖ a))) + , (cong (-ₖ-gen (suc n) (suc m) p q) + (cong (subst coHomK (+'-comm (suc m) (suc n))) + (0ₖ-⌣ₖ (suc m) (suc n) a)) + ∙ coherence-transp n m p q)) +gradedComm-⌣ₖ∙ n m p q = + trElim (λ _ → isOfHLevelPath (3 + n) ((isOfHLevel↑∙ (suc n) m)) _ _) + λ a → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ b → funExt⁻ (cong fst (f₁≡f₂ b)) a) + where + f₁ : coHomK (suc m) → S₊∙ (suc n) →∙ coHomK-ptd (suc n +' suc m) + fst (f₁ b) a = _⌣ₖ_ {n = suc n} {m = suc m} ∣ a ∣ₕ b + snd (f₁ b) = 0ₖ-⌣ₖ (suc n) (suc m) b + + f₂ : coHomK (suc m) → S₊∙ (suc n) →∙ coHomK-ptd (suc n +' suc m) + fst (f₂ b) a = + -ₖ-gen (suc n) (suc m) p q (subst coHomK (+'-comm (suc m) (suc n)) + (_⌣ₖ_ {n = suc m} {m = suc n} b ∣ a ∣ₕ)) + snd (f₂ b) = + (cong (-ₖ-gen (suc n) (suc m) p q) + (cong (subst coHomK (+'-comm (suc m) (suc n))) + (⌣ₖ-0ₖ (suc m) (suc n) b)) + ∙ coherence-transp n m p q) + + f₁≡f₂ : (b : _) → f₁ b ≡ f₂ b + f₁≡f₂ = + trElim (λ _ → isOfHLevelPath (3 + m) + (subst (isOfHLevel (3 + m)) + (λ i → S₊∙ (suc n) →∙ coHomK-ptd (+'-comm (suc n) (suc m) (~ i))) + (isOfHLevel↑∙' (suc m) n)) _ _) + λ b → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ a → gradedComm-elimCase (n + m) n m refl p q a b) + +-- Finally, graded commutativity: +gradedComm-⌣ₖ : (n m : ℕ) (a : coHomK n) (b : coHomK m) + → a ⌣ₖ b ≡ (-ₖ^ n · m) (subst coHomK (+'-comm m n) (b ⌣ₖ a)) +gradedComm-⌣ₖ zero zero a b = sym (transportRefl _) ∙ cong (transport refl) (comm-·₀ a b) +gradedComm-⌣ₖ zero (suc m) a b = + sym (transportRefl _) + ∙∙ (λ k → subst coHomK (isSetℕ _ _ refl (+'-comm (suc m) zero) k) (b ⌣ₖ a)) + ∙∙ sym (-ₖ-gen-inl-left zero (suc m) tt (evenOrOdd (suc m)) + (subst coHomK (+'-comm (suc m) zero) (b ⌣ₖ a))) +gradedComm-⌣ₖ (suc n) zero a b = + sym (transportRefl _) + ∙∙ ((λ k → subst coHomK (isSetℕ _ _ refl (+'-comm zero (suc n)) k) (b ⌣ₖ a))) + ∙∙ sym (-ₖ-gen-inl-right (suc n) zero (evenOrOdd (suc n)) tt + (subst coHomK (+'-comm zero (suc n)) (b ⌣ₖ a))) +gradedComm-⌣ₖ (suc n) (suc m) a b = + funExt⁻ (cong fst (gradedComm-⌣ₖ∙ n m (evenOrOdd (suc n)) (evenOrOdd (suc m)) a)) b + +gradedComm-⌣ : {A : Type ℓ} (n m : ℕ) (a : coHom n A) (b : coHom m A) + → a ⌣ b ≡ (-ₕ^ n · m) (subst (λ n → coHom n A) (+'-comm m n) (b ⌣ a)) +gradedComm-⌣ n m = + sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → + cong ∣_∣₂ (funExt (λ x → + gradedComm-⌣ₖ n m (f x) (g x) + ∙ cong ((-ₖ^ n · m) ∘ (subst coHomK (+'-comm m n))) + λ i → g (transportRefl x (~ i)) ⌣ₖ f (transportRefl x (~ i)))) diff --git a/Cubical/ZCohomology/RingStructure/RingLaws.agda b/Cubical/ZCohomology/RingStructure/RingLaws.agda new file mode 100644 index 000000000..99b719f6f --- /dev/null +++ b/Cubical/ZCohomology/RingStructure/RingLaws.agda @@ -0,0 +1,657 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.RingStructure.RingLaws where +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.RingStructure.CupProduct + +open import Cubical.Homotopy.Loopspace + +open import Cubical.HITs.S1 hiding (_·_) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.Truncation renaming (elim to trElim) + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.GroupoidLaws hiding (assoc) +open import Cubical.Data.Int + renaming (_+_ to _ℤ+_ ; _·_ to _ℤ∙_ ; +Comm to +ℤ-comm + ; ·Comm to ∙-comm ; +Assoc to ℤ+-assoc ; -_ to -ℤ_) + hiding (_+'_ ; +'≡+) +open import Cubical.Data.Nat +open import Cubical.Data.Sigma + +private + variable + ℓ : Level + +-- Some boring lemmas +·₀≡·ℤ : (x y : ℤ) → _·₀_ {n = zero} x y ≡ x ℤ∙ y +·₀≡·ℤ (pos zero) y = refl +·₀≡·ℤ (pos (suc n)) y = cong (y ℤ+_) (·₀≡·ℤ (pos n) y) +·₀≡·ℤ (negsuc zero) (pos zero) = refl +·₀≡·ℤ (negsuc zero) (pos (suc n)) = -AntiComm 0 (pos (suc n)) +·₀≡·ℤ (negsuc zero) (negsuc n) = -AntiComm 0 (negsuc n) +·₀≡·ℤ (negsuc (suc n)) (pos m) = + (cong ((_ℤ+ (pos 0 - pos m))) (·₀≡·ℤ (negsuc n) (pos m)) + ∙ cong ((negsuc n ℤ∙ (pos m)) ℤ+_) + (-AntiComm 0 (pos m) ∙ cong (-ℤ_) (help m))) + ∙ sym (+ℤ-comm (-ℤ (pos m)) (negsuc n ℤ∙ (pos m))) + where + help : (m : ℕ) → (pos m - 0) ≡ pos m + help zero = refl + help (suc m) = refl +·₀≡·ℤ (negsuc (suc n)) (negsuc m) = + +ℤ-comm (_·₀_{n = zero} (negsuc n) (negsuc m)) (pos 0 - (negsuc m)) + ∙∙ cong ((pos 0 - negsuc m) ℤ+_) (·₀≡·ℤ (negsuc n) (negsuc m)) + ∙∙ cong (_ℤ+ (negsuc n ℤ∙ negsuc m)) (help m) + where + help : (m : ℕ) → (pos 0 - negsuc m) ≡ pos (suc m) + help zero = refl + help (suc n) = cong sucℤ (help n) + +comm-·₀ : (x y : ℤ) → _·₀_ {n = 0} x y ≡ y ·₀ x +comm-·₀ x y = ·₀≡·ℤ x y ∙∙ ∙-comm x y ∙∙ sym (·₀≡·ℤ y x) + ++'-assoc : (n m l : ℕ) → (n +' (m +' l)) ≡ ((n +' m) +' l) ++'-assoc n m l = + (λ i → +'≡+ n (+'≡+ m l i) i) + ∙∙ +-assoc n m l + ∙∙ (λ i → +'≡+ (+'≡+ n m (~ i)) l (~ i)) + +-- Zero multiplication + +⌣ₖ-0ₖ : (n m : ℕ) → (x : coHomK n) + → x ⌣ₖ (0ₖ _) ≡ 0ₖ _ +⌣ₖ-0ₖ n m x = snd (⌣ₖ∙ n m x) + +0ₖ-⌣ₖ : (n m : ℕ) → (x : coHomK m) → + (0ₖ _) ⌣ₖ x ≡ 0ₖ _ +0ₖ-⌣ₖ n m = funExt⁻ (cong fst (snd (⌣ₖ∙∙ n m))) + +⌣ₖ-0ₖ≡0ₖ-⌣ₖ : (n m : ℕ) → ⌣ₖ-0ₖ n m (0ₖ _) ≡ 0ₖ-⌣ₖ n m (0ₖ _) +⌣ₖ-0ₖ≡0ₖ-⌣ₖ zero zero = refl +⌣ₖ-0ₖ≡0ₖ-⌣ₖ zero (suc m) = refl +⌣ₖ-0ₖ≡0ₖ-⌣ₖ (suc n) zero = refl +⌣ₖ-0ₖ≡0ₖ-⌣ₖ (suc zero) (suc m) = refl +⌣ₖ-0ₖ≡0ₖ-⌣ₖ (suc (suc n)) (suc m) = refl + +-- Left distributivity + +private + ⌣ₖ-distrFun : -- z ⌣ₖ (x +ₖ y) + (n m : ℕ) → (x y : coHomK n) + → coHomK-ptd m →∙ coHomK-ptd (m +' n) + fst (⌣ₖ-distrFun n m x y) z = + z ⌣ₖ (x +ₖ y) + snd (⌣ₖ-distrFun n m x y) = + 0ₖ-⌣ₖ m n (x +ₖ y) + + ⌣ₖ-distrFun2 : -- z ⌣ₖ x +ₖ z ⌣ₖ y + (n m : ℕ) → (x y : coHomK n) + → coHomK-ptd m →∙ coHomK-ptd (m +' n) + fst (⌣ₖ-distrFun2 n m x y) z = + z ⌣ₖ x +ₖ z ⌣ₖ y + snd (⌣ₖ-distrFun2 n m x y) = + cong₂ _+ₖ_ (0ₖ-⌣ₖ m n x) (0ₖ-⌣ₖ m n y) ∙ rUnitₖ _ _ + + leftDistr-⌣ₖ· : (n m : ℕ) (x y : coHomK (suc n)) → ⌣ₖ-distrFun (suc n) (suc m) x y ≡ ⌣ₖ-distrFun2 (suc n) (suc m) x y + leftDistr-⌣ₖ· n m = + elim2 (λ _ _ → isOfHLevelSuc (2 + n) (hLevHelp n m _ _)) + main + where + hLevHelp : (n m : ℕ) (x y : _) → isOfHLevel (2 + n) ( ⌣ₖ-distrFun (suc n) (suc m) x y ≡ ⌣ₖ-distrFun2 (suc n) (suc m) x y) + hLevHelp n m x y = + (subst (isOfHLevel (3 + n)) + (λ i → (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (suc (+-comm n m i))))) + (isOfHLevel↑∙ (suc n) m)) _ _ + + left-fst : (n : ℕ) (x : S₊ (suc n)) → + fst (⌣ₖ-distrFun (suc n) (suc m) ∣ ptSn (suc n) ∣ ∣ x ∣) ≡ + fst (⌣ₖ-distrFun2 (suc n) (suc m) ∣ ptSn (suc n) ∣ ∣ x ∣) + left-fst zero y = + (funExt (λ z → sym (lUnitₖ _ (z ⌣ₖ ∣ y ∣)) ∙ λ i → ⌣ₖ-0ₖ _ _ z (~ i) +ₖ z ⌣ₖ ∣ y ∣)) + left-fst (suc n) y = + (funExt (λ z → sym (lUnitₖ _ (z ⌣ₖ ∣ y ∣)) ∙ λ i → ⌣ₖ-0ₖ _ _ z (~ i) +ₖ z ⌣ₖ ∣ y ∣)) + + right-fst : (n : ℕ) (x : S₊ (suc n)) → + fst (⌣ₖ-distrFun (suc n) (suc m) ∣ x ∣ (0ₖ _)) ≡ + fst (⌣ₖ-distrFun2 (suc n) (suc m) ∣ x ∣ (0ₖ _)) + right-fst n x = funExt (λ z → (cong (z ⌣ₖ_) (rUnitₖ _ ∣ x ∣) ∙∙ sym (rUnitₖ _ _) ∙∙ λ i → z ⌣ₖ ∣ x ∣ +ₖ ⌣ₖ-0ₖ _ (suc n) z (~ i))) + + helper : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ x) + → sym (lUnitₖ _ x) ∙ cong (_+ₖ x) p ≡ sym (rUnitₖ _ x) ∙ cong (x +ₖ_) p + helper n x = J (λ x p → sym (lUnitₖ _ x) ∙ cong (_+ₖ x) p ≡ sym (rUnitₖ _ x) ∙ cong (x +ₖ_) p) + (helper' n) + where + helper' : (n : ℕ) → (λ i → lUnitₖ (suc n) (0ₖ (suc n)) (~ i)) ∙ + (λ i → 0ₖ (suc n) +ₖ 0ₖ (suc n)) + ≡ + (λ i → rUnitₖ (suc n) (0ₖ (suc n)) (~ i)) ∙ + (λ i → 0ₖ (suc n) +ₖ 0ₖ (suc n)) + helper' zero = refl + helper' (suc n) = refl + + left-fst≡right-fst : (n : ℕ) → left-fst n (ptSn _) ≡ right-fst n (ptSn _) + left-fst≡right-fst zero i j z = helper _ (z ⌣ₖ ∣ base ∣) (sym (⌣ₖ-0ₖ _ _ z)) i j + left-fst≡right-fst (suc n) i j z = helper _ (z ⌣ₖ ∣ north ∣) (sym (⌣ₖ-0ₖ _ (suc (suc n)) z)) i j + + main : (a b : S₊ (suc n)) → + ⌣ₖ-distrFun (suc n) (suc m) ∣ a ∣ₕ ∣ b ∣ₕ ≡ + ⌣ₖ-distrFun2 (suc n) (suc m) ∣ a ∣ₕ ∣ b ∣ₕ + main = + wedgeconFun n n + (λ x y → subst (λ l → isOfHLevel l ((⌣ₖ-distrFun (suc n) (suc m) ∣ x ∣ ∣ y ∣) + ≡ ⌣ₖ-distrFun2 (suc n) (suc m) ∣ x ∣ ∣ y ∣)) + (+-suc n (suc n)) + (isOfHLevelPlus {n = 2 + n} n (hLevHelp n m ∣ x ∣ ∣ y ∣))) + (λ x → →∙Homogeneous≡ (isHomogeneousKn _) (left-fst n x)) + (λ x → →∙Homogeneous≡ (isHomogeneousKn _) (right-fst n x)) + (cong (→∙Homogeneous≡ (isHomogeneousKn _)) (sym (left-fst≡right-fst n))) + + -- Distributivity for 0 dimensional cases + leftDistr₀n : (n : ℕ) → (z : coHomK 0) (x y : coHomK n) + → z ·₀ (x +[ n ]ₖ y) ≡ z ·₀ x +[ n ]ₖ (z ·₀ y) + leftDistr₀n n (pos zero) x y = sym (rUnitₖ n (0ₖ _)) + leftDistr₀n n (pos (suc z)) x y = + cong ((x +ₖ y) +ₖ_) (leftDistr₀n n (pos z) x y) + ∙∙ sym (assocₖ n x y (pos z ·₀ x +[ n ]ₖ (pos z ·₀ y))) + ∙∙ cong (x +ₖ_) (assocₖ n y (pos z ·₀ x) (pos z ·₀ y) + ∙∙ cong (_+ₖ (pos z ·₀ y)) (commₖ n y (pos z ·₀ x)) + ∙∙ sym (assocₖ n (pos z ·₀ x) y (pos z ·₀ y))) + ∙ assocₖ n x _ _ + leftDistr₀n n (negsuc zero) x y = -distrₖ n x y + leftDistr₀n n (negsuc (suc z)) x y = + cong₂ (_+ₖ_) (leftDistr₀n n (negsuc z) x y) (-distrₖ n x y) + ∙∙ assocₖ n ((negsuc z ·₀ x) +ₖ (negsuc z ·₀ y)) (-ₖ x) (-ₖ y) + ∙∙ cong (_-ₖ y) (sym (assocₖ n (negsuc z ·₀ x) (negsuc z ·₀ y) (-ₖ x)) + ∙∙ cong ((negsuc z ·₀ x) +ₖ_) (commₖ n (negsuc z ·₀ y) (-ₖ x)) + ∙∙ assocₖ n (negsuc z ·₀ x) (-ₖ x) (negsuc z ·₀ y)) + ∙ sym (assocₖ n (negsuc (suc z) ·₀ x) _ _) + + leftDistrn₀ : (n : ℕ) → (z : coHomK n) (x y : coHomK 0) + → (x ℤ+ y) ·₀ z ≡ x ·₀ z +[ n ]ₖ (y ·₀ z) + leftDistrn₀ n z x (pos zero) = sym (rUnitₖ n (x ·₀ z)) + leftDistrn₀ n z x (pos (suc y)) = + lem (x +pos y) + ∙∙ cong (z +ₖ_) (leftDistrn₀ n z x (pos y) ∙ commₖ n _ _) + ∙∙ assocₖ n z _ _ + ∙ commₖ n _ _ + where + lem : (a : ℤ) → (sucℤ a) ·₀ z ≡ z +ₖ (a ·₀ z) + lem (pos zero) = refl + lem (pos (suc a)) = cong (z +ₖ_) (lem (pos a)) + lem (negsuc zero) = sym (rCancelₖ n z) + lem (negsuc (suc a)) = sym (-cancelLₖ n z (negsuc a ·₀ z)) ∙ sym (assocₖ n z (negsuc a ·₀ z) (-ₖ z)) + leftDistrn₀ n z x (negsuc y) = main y + where + help : (x : ℤ) → predℤ x ·₀ z ≡ (x ·₀ z -ₖ z) + help (pos zero) = sym (lUnitₖ n (-ₖ z)) + help (pos (suc x)) = sym (-cancelLₖ n z (pos x ·₀ z)) + help (negsuc x) = refl + + main : (y : _) → (x ℤ+ negsuc y) ·₀ z ≡ (x ·₀ z) +ₖ (negsuc y ·₀ z) + main zero = help x + main (suc y) = + help (x +negsuc y) + ∙∙ cong (_-ₖ z) (main y) + ∙∙ sym (assocₖ n _ _ _) + +leftDistr-⌣ₖ : (n m : ℕ) (z : coHomK n) (x y : coHomK m) + → z ⌣ₖ (x +ₖ y) ≡ (z ⌣ₖ x +ₖ z ⌣ₖ y) +leftDistr-⌣ₖ zero m z x y = leftDistr₀n m z x y +leftDistr-⌣ₖ (suc n) zero z x y = leftDistrn₀ (suc n) z x y +leftDistr-⌣ₖ (suc n) (suc m) z x y = funExt⁻ (cong fst (leftDistr-⌣ₖ· m n x y)) z + + +-- Right distributivity + +private + ⌣ₖ-distrFun-r : -- (x +ₖ y) ⌣ₖ z + (n m : ℕ) → (x y : coHomK n) + → coHomK-ptd m →∙ coHomK-ptd (n +' m) + fst (⌣ₖ-distrFun-r n m x y) z = + (x +ₖ y) ⌣ₖ z + snd (⌣ₖ-distrFun-r n m x y) = + ⌣ₖ-0ₖ n m (x +ₖ y) -- ⌣ₖ-0ₖ m n (x +ₖ y) + + ⌣ₖ-distrFun2-r : + (n m : ℕ) → (x y : coHomK n) + → coHomK-ptd m →∙ coHomK-ptd (n +' m) + fst (⌣ₖ-distrFun2-r n m x y) z = + x ⌣ₖ z +ₖ y ⌣ₖ z + snd (⌣ₖ-distrFun2-r n m x y) = + cong₂ _+ₖ_ (⌣ₖ-0ₖ n m x) (⌣ₖ-0ₖ n m y) ∙ rUnitₖ _ _ + + rightDistr-⌣ₖ· : (n m : ℕ) (x y : coHomK (suc n)) → ⌣ₖ-distrFun-r (suc n) (suc m) x y ≡ ⌣ₖ-distrFun2-r (suc n) (suc m) x y + rightDistr-⌣ₖ· n m = + elim2 (λ _ _ → isOfHLevelPath (3 + n) (isOfHLevel↑∙ (suc n) m) _ _) + main + where + fst-left : (n : ℕ) (y : S₊ (suc n)) → + fst (⌣ₖ-distrFun-r (suc n) (suc m) ∣ ptSn (suc n) ∣ ∣ y ∣) ≡ + fst (⌣ₖ-distrFun2-r (suc n) (suc m) ∣ ptSn (suc n) ∣ ∣ y ∣) + fst-left n y = + funExt (λ z + → cong (_⌣ₖ z) (lUnitₖ _ ∣ y ∣) + ∙∙ sym (lUnitₖ _ (∣ y ∣ ⌣ₖ z)) + ∙∙ cong (_+ₖ (∣ y ∣ ⌣ₖ z)) (sym (0ₖ-⌣ₖ _ _ z))) + + fst-right : (n : ℕ) (x : S₊ (suc n)) → + fst (⌣ₖ-distrFun-r (suc n) (suc m) ∣ x ∣ ∣ ptSn (suc n) ∣) ≡ + fst (⌣ₖ-distrFun2-r (suc n) (suc m) ∣ x ∣ ∣ ptSn (suc n) ∣) + fst-right n x = + funExt λ z + → cong (_⌣ₖ z) (rUnitₖ _ ∣ x ∣) + ∙∙ sym (rUnitₖ _ _) + ∙∙ cong (∣ x ∣ ⌣ₖ z +ₖ_) (sym (0ₖ-⌣ₖ _ _ z)) + + left≡right : (n : ℕ) → fst-left n (ptSn (suc n)) ≡ fst-right n (ptSn (suc n)) + left≡right zero = refl + left≡right (suc n) = refl + + main : (a b : S₊ (suc n)) → + ⌣ₖ-distrFun-r (suc n) (suc m) ∣ a ∣ₕ ∣ b ∣ₕ ≡ + ⌣ₖ-distrFun2-r (suc n) (suc m) ∣ a ∣ₕ ∣ b ∣ₕ + main = + wedgeconFun n n + (λ x y → subst (λ l → isOfHLevel l ((⌣ₖ-distrFun-r (suc n) (suc m) ∣ x ∣ ∣ y ∣) + ≡ ⌣ₖ-distrFun2-r (suc n) (suc m) ∣ x ∣ ∣ y ∣)) + (+-suc n (suc n)) + (isOfHLevelPlus {n = 2 + n} n (isOfHLevel↑∙ (suc n) m _ _))) + (λ x → →∙Homogeneous≡ (isHomogeneousKn _) (fst-left n x)) + (λ x → →∙Homogeneous≡ (isHomogeneousKn _) (fst-right n x)) + (cong (→∙Homogeneous≡ (isHomogeneousKn _)) (sym (left≡right n))) + +rightDistr-⌣ₖ : (n m : ℕ) (x y : coHomK n) (z : coHomK m) + → (x +ₖ y) ⌣ₖ z ≡ (x ⌣ₖ z +ₖ y ⌣ₖ z) +rightDistr-⌣ₖ zero zero x y z = + comm-·₀ (x ℤ+ y) z + ∙∙ leftDistr-⌣ₖ zero zero z x y + ∙∙ cong₂ _+ₖ_ (sym (comm-·₀ x z)) (sym (comm-·₀ y z)) +rightDistr-⌣ₖ zero (suc m) x y z = leftDistr-⌣ₖ _ zero z x y +rightDistr-⌣ₖ (suc n) zero x y z = leftDistr-⌣ₖ zero (suc n) z x y +rightDistr-⌣ₖ (suc n) (suc m) x y z = (funExt⁻ (cong fst (rightDistr-⌣ₖ· n m x y))) z + +-- Associativity + +private +-- We need to give the two associators as (doubly) pointed functions + assocer : (n m k : ℕ) → coHomK (suc n) + → coHomK-ptd (suc m) + →∙ (coHomK-ptd (suc k) + →∙ coHomK-ptd ((suc n) +' ((suc m) +' (suc k))) ∙) + fst (fst (assocer n m k x) y) z = x ⌣ₖ (y ⌣ₖ z) + snd (fst (assocer n m k x) y) = cong (x ⌣ₖ_) (⌣ₖ-0ₖ _ (suc k) y) ∙ ⌣ₖ-0ₖ _ _ x + snd (assocer n m k x) = + ΣPathP (funExt (λ z → cong (x ⌣ₖ_) (0ₖ-⌣ₖ (suc m) (suc k) z) ∙ ⌣ₖ-0ₖ (suc n) _ x) + , help) + where + h : (n m k : ℕ) (x : coHomK (suc n)) → cong (_⌣ₖ_ x) (⌣ₖ-0ₖ (suc m) (suc k) (0ₖ _)) ≡ + (λ i → x ⌣ₖ 0ₖ-⌣ₖ (suc m) (suc k) (0ₖ _) i) + h zero zero k x = refl + h zero (suc m) k x = refl + h (suc n) zero k x = refl + h (suc n) (suc m) k x = refl + + help : PathP (λ i → (cong (x ⌣ₖ_) (0ₖ-⌣ₖ (suc m) (suc k) (0ₖ _)) ∙ ⌣ₖ-0ₖ (suc n) _ x) i ≡ 0ₖ _) + (cong (x ⌣ₖ_) (⌣ₖ-0ₖ (suc m) (suc k) (0ₖ _)) ∙ ⌣ₖ-0ₖ (suc n) _ x) refl + help = compPathR→PathP + (cong (_∙ ⌣ₖ-0ₖ (suc n) ((suc m) +' (suc k)) x) (h n m k x) + ∙∙ rUnit _ + ∙∙ cong ((cong (x ⌣ₖ_) (0ₖ-⌣ₖ (suc m) (suc k) (0ₖ _)) ∙ ⌣ₖ-0ₖ (suc n) _ x) ∙_) (rUnit refl)) + + assoc2-sub : (n m k : ℕ) → _ → _ + assoc2-sub n m k = subst coHomK (sym (+'-assoc n m k)) + + assoc2-sub-0 : (n m k : ℕ) → assoc2-sub n m k (0ₖ _) ≡ 0ₖ _ + assoc2-sub-0 zero zero zero = refl + assoc2-sub-0 zero zero (suc zero) = refl + assoc2-sub-0 zero zero (suc (suc k)) = refl + assoc2-sub-0 zero (suc zero) zero = refl + assoc2-sub-0 zero (suc zero) (suc k) = refl + assoc2-sub-0 zero (suc (suc m)) zero = refl + assoc2-sub-0 zero (suc (suc m)) (suc k) = refl + assoc2-sub-0 (suc zero) zero zero = refl + assoc2-sub-0 (suc zero) zero (suc k) = refl + assoc2-sub-0 (suc (suc n)) zero zero = refl + assoc2-sub-0 (suc (suc n)) zero (suc k) = refl + assoc2-sub-0 (suc zero) (suc m) zero = refl + assoc2-sub-0 (suc (suc n)) (suc m) zero = refl + assoc2-sub-0 (suc zero) (suc m) (suc k) = refl + assoc2-sub-0 (suc (suc n)) (suc m) (suc k) = refl + + assocer2 : (n m k : ℕ) + → coHomK (suc n) + → coHomK-ptd (suc m) →∙ (coHomK-ptd (suc k) →∙ coHomK-ptd ((suc n) +' ((suc m) +' (suc k))) ∙) + fst (fst (assocer2 n m k x) y) z = + subst coHomK (sym (+'-assoc (suc n) (suc m) (suc k))) ((x ⌣ₖ y) ⌣ₖ z) -- + snd (fst (assocer2 zero m k x) y) = + cong (subst coHomK (sym (+'-assoc 1 (suc m) (suc k)))) (⌣ₖ-0ₖ _ _ (x ⌣ₖ y)) + snd (fst (assocer2 (suc n) m k x) y) = + cong (subst coHomK (sym (+'-assoc (2 + n) (suc m) (suc k)))) (⌣ₖ-0ₖ _ _ (x ⌣ₖ y)) + fst (snd (assocer2 zero m k x) i) z = + subst coHomK (sym (+'-assoc (suc zero) (suc m) (suc k))) (⌣ₖ-0ₖ _ _ x i ⌣ₖ z) + snd (snd (assocer2 zero m k x) i) j = + subst coHomK (sym (+'-assoc (suc zero) (suc m) (suc k))) (⌣ₖ-0ₖ _ _ (⌣ₖ-0ₖ _ _ x i) j) + fst (snd (assocer2 (suc n) m k x) i) z = + subst coHomK (sym (+'-assoc (2 + n) (suc m) (suc k))) (⌣ₖ-0ₖ _ _ x i ⌣ₖ z) + snd (snd (assocer2 (suc n) m k x) i) j = + subst coHomK (sym (+'-assoc (2 + n) (suc m) (suc k))) (⌣ₖ-0ₖ _ _ (⌣ₖ-0ₖ _ _ x i) j) + + assocer-helpFun : (n m : ℕ) → coHomK (suc n) → coHomK-ptd (suc m) →∙ Ω (coHomK-ptd (3 + (n + m))) + fst (assocer-helpFun n m a) b = Kn→ΩKn+1 _ (a ⌣ₖ b) + snd (assocer-helpFun n m a) = cong (Kn→ΩKn+1 _) (⌣ₖ-0ₖ (suc n) (suc m) a) ∙ Kn→ΩKn+10ₖ _ + + assocer-helpFun2 : (n m : ℕ) → coHomK (suc n) → coHomK-ptd (suc m) →∙ Ω (coHomK-ptd (3 + (n + m))) + fst (assocer-helpFun2 n m a) b i = (Kn→ΩKn+1 _ a i) ⌣ₖ b + snd (assocer-helpFun2 n m a) i j = ⌣ₖ-0ₖ _ (suc m) (Kn→ΩKn+1 _ a j) i + +-- Key lemma for associativity +assocer-helpFun≡ : (n m : ℕ) → (x : coHomK (suc n)) → assocer-helpFun n m x ≡ assocer-helpFun2 n m x +assocer-helpFun≡ n m = + trElim (λ _ → isOfHLevelPath (3 + n) (hLev-assocer-helpFun n m) _ _) + λ a → →∙Homogeneous≡ (subst isHomogeneous Kn≃ΩKn+1∙ (isHomogeneousKn _)) + (funExt (main n a)) + where + hLev-assocer-helpFun : (n m : ℕ) → isOfHLevel (3 + n) (coHomK-ptd (suc m) →∙ Ω (coHomK-ptd (3 + (n + m)))) + hLev-assocer-helpFun n m = + subst (isOfHLevel (3 + n)) (cong (coHomK-ptd (suc m) →∙_) + (Kn≃ΩKn+1∙)) + (isOfHLevel↑∙ (suc n) m) + + main : (n : ℕ) (a : S₊ (suc n)) (b : _) + → fst (assocer-helpFun n m ∣ a ∣) b ≡ fst (assocer-helpFun2 n m ∣ a ∣) b + main zero a b k i = + hcomp + (λ r → λ {(i = i0) → 0ₖ _ + ; (i = i1) → ∣ rCancel (merid north) r (~ k) ∣ + ; (k = i0) → Kn→ΩKn+1 _ (∣ a ∣ ⌣ₖ b) i + ; (k = i1) → (Kn→ΩKn+1 _ ∣ a ∣ i) ⌣ₖ b}) + (∣ compPath-filler (merid a) (sym (merid base)) k i ∣ ⌣ₖ b) + main (suc n) a b k i = + hcomp + (λ r → λ {(i = i0) → 0ₖ _ + ; (i = i1) → ∣ rCancel (merid north) r (~ k) ∣ + ; (k = i0) → Kn→ΩKn+1 _ (∣ a ∣ ⌣ₖ b) i + ; (k = i1) → (Kn→ΩKn+1 _ ∣ a ∣ i) ⌣ₖ b}) + (∣ compPath-filler (merid a) (sym (merid north)) k i ∣ ⌣ₖ b) + +assoc-helper : (n m : ℕ) (x : coHomK (suc n)) (y : coHomK (suc m)) + → Kn→ΩKn+1 _ (x ⌣ₖ y) ≡ λ i → (Kn→ΩKn+1 _ x i) ⌣ₖ y +assoc-helper n m x y = funExt⁻ (cong fst (assocer-helpFun≡ n m x)) y + +assoc-⌣ₖ· : (n m k : ℕ) → (x : coHomK (suc n)) → assocer n m k x ≡ assocer2 n m k x +assoc-⌣ₖ· n m k = + trElim (λ _ → isOfHLevelPath (3 + n) + (transport (λ i → isOfHLevel (3 + n) + (coHomK-ptd (suc m) →∙ (coHomK-ptd (suc k) →∙ coHomK-ptd (h (~ i)) ∙))) + (isOfHLevel↑∙∙ m k (suc n))) _ _) + λ a → →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousKn _)) + (funExt λ b → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt (main n m k a b))) + where + + h : (suc n) +' ((suc m) +' (suc k)) ≡ suc (suc (suc n + m + k)) + h = cong (2 +_) (+-suc n (m + k)) ∙ λ i → suc (suc (suc (+-assoc n m k i))) + + + main : (n m k : ℕ) (a : S₊ (suc n)) (b : coHomK (suc m)) (c : coHomK (suc k)) + → ∣ a ∣ ⌣ₖ b ⌣ₖ c ≡ + (subst coHomK (λ i → +'-assoc (suc n) (suc m) (suc k) (~ i)) + ((∣ a ∣ ⌣ₖ b) ⌣ₖ c)) + main zero m k a b c = goal a ∙ sym (funExt⁻ t ((∣ a ∣ ⌣ₖ b) ⌣ₖ c)) + where + t : subst coHomK (λ i → +'-assoc 1 (suc m) (suc k) (~ i)) ≡ idfun _ + t = cong (subst coHomK) (isSetℕ _ _ (+'-assoc 1 (suc m) (suc k)) refl) ∙ funExt transportRefl + goal : (a : _) → ∣ a ∣ ⌣ₖ b ⌣ₖ c ≡ (∣ a ∣ ⌣ₖ b) ⌣ₖ c + goal base = refl + goal (loop i) j = assoc-helper m k b c j i + main (suc n) m k north b c = refl + main (suc n) m k south b c = refl + main (suc n) m k (merid a i) b c j = help2 j i + where + transpLem : (n m : ℕ) (p : n ≡ m) → subst coHomK p (0ₖ _) ≡ 0ₖ _ + transpLem zero m = J (λ m p → subst coHomK p (0ₖ _) ≡ 0ₖ _) refl + transpLem (suc zero) m = J (λ m p → subst coHomK p (0ₖ _) ≡ 0ₖ _) refl + transpLem (suc (suc n)) m = J (λ m p → subst coHomK p (0ₖ _) ≡ 0ₖ _) refl + + transpLem-refl : transpLem (suc (suc (suc (suc (n + m + k))))) + (suc (suc n +' (suc m +' suc k))) + (λ i₃ → +'-assoc (2 + n) (suc m) (suc k) (~ i₃)) ≡ refl + transpLem-refl = transportRefl refl + moveTransports : (n m : ℕ) (x : coHomK n) (p : n ≡ m) (q : (suc n) ≡ suc m) + → PathP (λ i → transpLem _ _ q (~ i) ≡ transpLem _ _ q (~ i)) + (Kn→ΩKn+1 _ (subst coHomK p x)) + (cong (subst coHomK q) (Kn→ΩKn+1 _ x)) + moveTransports n m x = + J (λ m p → (q : (suc n) ≡ suc m) + → PathP (λ i → transpLem _ _ q (~ i) ≡ transpLem _ _ q (~ i)) + (Kn→ΩKn+1 _ (subst coHomK p x)) + (cong (subst coHomK q) (Kn→ΩKn+1 _ x))) + λ q → transport (λ j → PathP (λ i → transpLem (suc n) (suc n) (isSetℕ _ _ refl q j) (~ i) + ≡ transpLem (suc n) (suc n) (isSetℕ _ _ refl q j) (~ i)) + (Kn→ΩKn+1 _ (subst coHomK refl x)) + (cong (subst coHomK (isSetℕ _ _ refl q j)) (Kn→ΩKn+1 _ x))) + (h2 n x) + where + h2 : (n : ℕ) (x : _) + → PathP (λ i₁ → + transpLem (suc n) (suc n) (λ _ → suc n) (~ i₁) ≡ + transpLem (suc n) (suc n) (λ _ → suc n) (~ i₁)) + (Kn→ΩKn+1 n (subst coHomK (λ _ → n) x)) + (λ i₁ → subst coHomK (λ _ → suc n) (Kn→ΩKn+1 n x i₁)) + h2 zero x k i = + hcomp (λ j → λ {(i = i0) → transportRefl (refl {x = 0ₖ _}) (~ j) k + ; (i = i1) → transportRefl (refl {x = 0ₖ _}) (~ j) k + ; (k = i0) → Kn→ΩKn+1 _ (transportRefl x j) i + ; (k = i1) → transportRefl (Kn→ΩKn+1 _ x i) (~ j)}) + (Kn→ΩKn+1 _ x i) + h2 (suc n) x k i = + hcomp (λ j → λ {(i = i0) → transportRefl (refl {x = 0ₖ _}) (~ j) k + ; (i = i1) → transportRefl (refl {x = 0ₖ _}) (~ j) k + ; (k = i0) → Kn→ΩKn+1 _ (transport refl x) i + ; (k = i1) → transportRefl (Kn→ΩKn+1 _ x i) (~ j)}) + (Kn→ΩKn+1 _ (transportRefl x k) i) + + finalTransp : (n : ℕ) (a : _) → Kn→ΩKn+1 _ (subst coHomK (λ i₁ → +'-assoc (suc n) (suc m) (suc k) (~ i₁)) ((∣ a ∣ ⌣ₖ b) ⌣ₖ c)) + ≡ cong (subst coHomK (λ i₁ → +'-assoc (2 + n) (suc m) (suc k) (~ i₁))) (Kn→ΩKn+1 _ ((∣ a ∣ ⌣ₖ b) ⌣ₖ c)) + finalTransp n a = + rUnit _ + ∙∙ (λ i → (λ j → transpLem (suc (suc (suc (suc (n + m + k))))) + (suc (suc n +' (suc m +' suc k))) + (λ i₃ → +'-assoc (2 + n) (suc m) (suc k) (~ i₃)) (i ∧ j)) + ∙∙ moveTransports _ _ ((∣ a ∣ ⌣ₖ b) ⌣ₖ c) + (sym (+'-assoc (suc n) (suc m) (suc k))) + (sym (+'-assoc (2 + n) (suc m) (suc k))) i + ∙∙ λ j → transpLem (suc (suc (suc (suc (n + m + k))))) + (suc (suc n +' (suc m +' suc k))) + (λ i₃ → +'-assoc (2 + n) (suc m) (suc k) (~ i₃)) (i ∧ ~ j)) + ∙∙ (λ i → transportRefl refl i + ∙∙ cong (subst coHomK (λ i₁ → +'-assoc (2 + n) (suc m) (suc k) (~ i₁))) (Kn→ΩKn+1 _ ((∣ a ∣ ⌣ₖ b) ⌣ₖ c)) + ∙∙ transportRefl refl i) + ∙ sym (rUnit _) + + help2 : cong (λ x → (∣ x ∣) ⌣ₖ (b ⌣ₖ c)) (merid a) ≡ cong (assoc2-sub (2 + n) (suc m) (suc k)) (cong (λ x → (∣ x ∣ ⌣ₖ b) ⌣ₖ c) (merid a)) + help2 = ((λ r → Kn→ΩKn+1 _ (main n m k a b c r))) + ∙∙ finalTransp n a + ∙∙ λ r i → subst coHomK (sym (+'-assoc (2 + n) (suc m) (suc k))) (assoc-helper _ _ (∣ a ∣ ⌣ₖ b) c r i) + + +-- Some key distributivity lemmas +-Distₗ : (n m : ℕ) (x : coHomK n) (y : coHomK m) → (-ₖ (x ⌣ₖ y)) ≡ (-ₖ x) ⌣ₖ y +-Distₗ n m x y = + sym (rUnitₖ _ (-ₖ (x ⌣ₖ y))) + ∙∙ cong ((-ₖ (x ⌣ₖ y)) +ₖ_) (sym (rCancelₖ _ (x ⌣ₖ y))) + ∙∙ assocₖ _ (-ₖ (x ⌣ₖ y)) (x ⌣ₖ y) (-ₖ (x ⌣ₖ y)) + ∙∙ cong (_-ₖ (x ⌣ₖ y)) help + ∙∙ sym (assocₖ _ ((-ₖ x) ⌣ₖ y) (x ⌣ₖ y) (-ₖ (x ⌣ₖ y))) + ∙∙ cong ((-ₖ x) ⌣ₖ y +ₖ_) (rCancelₖ _ (x ⌣ₖ y)) + ∙∙ rUnitₖ _ ((-ₖ x) ⌣ₖ y) + where + help : (-ₖ (x ⌣ₖ y)) +ₖ (x ⌣ₖ y) ≡ (-ₖ x) ⌣ₖ y +ₖ (x ⌣ₖ y) + help = lCancelₖ _ _ + ∙∙ sym (0ₖ-⌣ₖ _ _ y) + ∙ cong (_⌣ₖ y) (sym (lCancelₖ _ x)) + ∙∙ rightDistr-⌣ₖ _ _ (-ₖ x) x y + +-Distᵣ : (n m : ℕ) (x : coHomK n) (y : coHomK m) → (-ₖ (x ⌣ₖ y)) ≡ x ⌣ₖ (-ₖ y) +-Distᵣ n m x y = + sym (rUnitₖ _ (-ₖ (x ⌣ₖ y))) + ∙∙ cong ((-ₖ (x ⌣ₖ y)) +ₖ_) (sym (rCancelₖ _ (x ⌣ₖ y))) + ∙∙ assocₖ _ (-ₖ (x ⌣ₖ y)) (x ⌣ₖ y) (-ₖ (x ⌣ₖ y)) + ∙∙ cong (_-ₖ (x ⌣ₖ y)) help + ∙∙ sym (assocₖ _ (x ⌣ₖ (-ₖ y)) (x ⌣ₖ y) (-ₖ (x ⌣ₖ y))) + ∙∙ cong (x ⌣ₖ (-ₖ y) +ₖ_) (rCancelₖ _ (x ⌣ₖ y)) + ∙∙ rUnitₖ _ (x ⌣ₖ (-ₖ y)) + where + help : (-ₖ (x ⌣ₖ y)) +ₖ (x ⌣ₖ y) ≡ x ⌣ₖ (-ₖ y) +ₖ (x ⌣ₖ y) + help = lCancelₖ _ _ + ∙∙ sym (⌣ₖ-0ₖ _ _ x) + ∙∙ cong (x ⌣ₖ_) (sym (lCancelₖ _ y)) + ∙ leftDistr-⌣ₖ _ _ x (-ₖ y) y + +assoc₀ : (m k : ℕ) (x : ℤ) (y : coHomK m) (z : coHomK k) → _⌣ₖ_{n = zero} x (y ⌣ₖ z) ≡ (_⌣ₖ_{n = zero} {m = m} x y) ⌣ₖ z +assoc₀ m k x y z = main x + where + h : subst coHomK (sym (+'-assoc zero m k)) ≡ idfun _ + h = cong (subst coHomK) (isSetℕ _ _ _ refl) ∙ funExt transportRefl + mainPos : (x : ℕ) → _⌣ₖ_ {n = zero} (pos x) (y ⌣ₖ z) ≡ ((_⌣ₖ_ {n = zero} {m = m} (pos x) y) ⌣ₖ z) + mainPos zero = sym (0ₖ-⌣ₖ _ _ z) ∙ cong (_⌣ₖ z) (sym (0ₖ-⌣ₖ _ _ y)) + mainPos (suc n) = + cong (y ⌣ₖ z +ₖ_) (mainPos n) + ∙∙ sym (rightDistr-⌣ₖ _ _ y (_⌣ₖ_ {n = zero} {m = m} (pos n) y) z) + ∙∙ (λ i → (y +ₖ (_⌣ₖ_ {n = zero} {m = m} (pos n) y)) ⌣ₖ z) + + main : (x : ℤ) → x ⌣ₖ y ⌣ₖ z ≡ ((_⌣ₖ_ {n = zero} {m = m} x y) ⌣ₖ z) + main (pos n) = mainPos n + main (negsuc n) = + (λ i → _⌣ₖ_ {n = zero} (+ℤ-comm (negsuc n) 0 i) (y ⌣ₖ z)) + ∙∙ sym (-Distₗ zero _ (pos (suc n)) (y ⌣ₖ z)) + ∙ cong (-ₖ_) (mainPos (suc n)) + ∙∙ -Distₗ _ _ (pos (suc n) ⌣ₖ y) z + ∙∙ cong (_⌣ₖ z) ((-Distₗ zero _ (pos (suc n)) y)) + ∙∙ λ i → (_⌣ₖ_ {n = zero} (+ℤ-comm (negsuc n) 0 (~ i)) y) ⌣ₖ z + +assoc-⌣ₖ : (n m k : ℕ) (x : coHomK n) (y : coHomK m) (z : coHomK k) + → x ⌣ₖ y ⌣ₖ z ≡ subst coHomK (sym (+'-assoc n m k)) ((x ⌣ₖ y) ⌣ₖ z) +assoc-⌣ₖ zero m k x y z = assoc₀ _ _ x y z ∙ sym (funExt⁻ h ((x ⌣ₖ y) ⌣ₖ z)) + where + h : subst coHomK (sym (+'-assoc zero m k)) ≡ idfun _ + h = cong (subst coHomK) (isSetℕ _ _ _ refl) ∙ funExt transportRefl + +assoc-⌣ₖ (suc n) zero k x y z = + help y + ∙∙ sym (transportRefl ((x ⌣ₖ y) ⌣ₖ z)) + ∙∙ λ i → transport (λ j → coHomK ((isSetℕ _ _ ((sym (+'-assoc (suc n) zero k))) refl (~ i) j))) + ((_⌣ₖ_ {m = zero} x y) ⌣ₖ z) + where + helpPos : (y : ℕ) → x ⌣ₖ (_⌣ₖ_ {n = zero} (pos y) z) ≡ + ((_⌣ₖ_ {m = zero} x (pos y)) ⌣ₖ z) + helpPos zero = + (⌣ₖ-0ₖ _ _ x) + ∙∙ (sym (0ₖ-⌣ₖ _ _ z)) + ∙∙ cong (_⌣ₖ z) (sym (⌣ₖ-0ₖ _ _ x)) + helpPos (suc y) = + leftDistr-⌣ₖ _ _ x z (_⌣ₖ_{n = zero} (pos y) z) + ∙∙ cong ((x ⌣ₖ z) +ₖ_) (helpPos y) + ∙∙ sym (rightDistr-⌣ₖ _ _ x (_⌣ₖ_{n = zero} (pos y) x) z) + + help : (y : ℤ) → x ⌣ₖ (_⌣ₖ_ {n = zero} y z) ≡ + ((_⌣ₖ_ {m = zero} x y) ⌣ₖ z) + help (pos n) = helpPos n + help (negsuc n) = + (λ i → x ⌣ₖ (_⌣ₖ_ {n = zero} (+ℤ-comm (negsuc n) 0 i) z)) + ∙∙ cong (x ⌣ₖ_) (sym (-Distₗ zero _ (pos (suc n)) z)) + ∙∙ sym (-Distᵣ _ _ x _) + ∙∙ cong -ₖ_ (helpPos (suc n)) + ∙∙ -Distₗ _ _ _ z + ∙∙ cong (_⌣ₖ z) (-Distᵣ _ zero x (pos (suc n))) + ∙∙ λ i → (_⌣ₖ_ {m = zero} x (+ℤ-comm (negsuc n) 0 (~ i))) ⌣ₖ z +assoc-⌣ₖ (suc n) (suc m) zero x y z = + (assoc-⌣ₖ (suc n) zero (suc m) x z y) + ∙∙ cong (subst coHomK (sym (+'-assoc (suc n) zero (suc m)))) + (sym (assoc₀ _ _ z x y)) + ∙∙ (funExt⁻ (sym h) (_⌣ₖ_ {n = zero} z (x ⌣ₖ y))) + where + h : subst coHomK (sym (+'-assoc (suc n) (suc m) zero)) + ≡ subst coHomK (λ i → +'-assoc (suc n) zero (suc m) (~ i)) + h = cong (subst coHomK) (isSetℕ _ _ _ _) +assoc-⌣ₖ (suc n) (suc m) (suc k) x y z = + funExt⁻ (cong fst (funExt⁻ (cong fst (assoc-⌣ₖ· n m k x)) y)) z + +-- Ring laws for ⌣ +module _ {A : Type ℓ} (n m : ℕ) where + ⌣-0ₕ : (f : coHom n A) → (f ⌣ 0ₕ m) ≡ 0ₕ _ + ⌣-0ₕ = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ f → cong ∣_∣₂ (funExt λ x → ⌣ₖ-0ₖ n m (f x)) + + 0ₕ-⌣ : (f : coHom m A) → (0ₕ n ⌣ f) ≡ 0ₕ _ + 0ₕ-⌣ = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ f → cong ∣_∣₂ (funExt λ x → 0ₖ-⌣ₖ n m (f x)) + + + leftDistr-⌣ : (f : coHom n A) (g h : coHom m A) → f ⌣ (g +ₕ h) ≡ f ⌣ g +ₕ f ⌣ h + leftDistr-⌣ = + sElim (λ _ → isSetΠ2 λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ g h → cong ∣_∣₂ (funExt λ x → leftDistr-⌣ₖ n m (f x) (g x) (h x)) + + rightDistr-⌣ : (g h : coHom n A) (f : coHom m A) → (g +ₕ h) ⌣ f ≡ g ⌣ f +ₕ h ⌣ f + rightDistr-⌣ = + sElim2 (λ _ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ g h → sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → rightDistr-⌣ₖ n m (g x) (h x) (f x)) + + assoc-⌣ : (l : ℕ) (f : coHom n A) (g : coHom m A) (h : coHom l A) + → f ⌣ g ⌣ h ≡ subst (λ x → coHom x A) (sym (+'-assoc n m l)) ((f ⌣ g) ⌣ h) + assoc-⌣ l = + sElim (λ _ → isSetΠ2 λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 squash₂ _ _) + λ g → sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ h → + cong ∣_∣₂ ((funExt (λ x → assoc-⌣ₖ n m l (f x) (g x) (h x) + ∙ cong (subst coHomK (sym (+'-assoc n m l))) + λ i → (f (transportRefl x (~ i)) ⌣ₖ g (transportRefl x (~ i))) + ⌣ₖ (h (transportRefl x (~ i)))))) + +-- Additive unit(s) +0⌣ : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) (x : coHom n A) + → x ⌣ (0ₕ m) ≡ 0ₕ _ +0⌣ n m = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → ⌣ₖ-0ₖ n m (f x)) + +⌣0 : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) (x : coHom n A) + → (0ₕ m) ⌣ x ≡ 0ₕ _ +⌣0 n m = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → 0ₖ-⌣ₖ m n (f x)) + +-- Multiplicative unit +1⌣ : ∀ {ℓ} {A : Type ℓ} → coHom 0 A +1⌣ = ∣ (λ _ → 1) ∣₂ + +private + n+'0 : (n : ℕ) → n +' 0 ≡ n + n+'0 zero = refl + n+'0 (suc n) = refl + +lUnit⌣ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) (x : coHom n A) + → x ⌣ 1⌣ ≡ subst (λ n → coHom n A) (sym (n+'0 n)) x +lUnit⌣ zero = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt (λ x → comm-·₀ (f x) (pos 1))) ∙ sym (transportRefl ∣ f ∣₂) +lUnit⌣ (suc n) = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt (λ x → cong (f x +ₖ_) (0ₖ-⌣ₖ zero _ (f x)) ∙ rUnitₖ _ (f x))) + ∙ sym (transportRefl ∣ f ∣₂) + +rUnit⌣ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) (x : coHom n A) + → 1⌣ ⌣ x ≡ x +rUnit⌣ zero = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → refl +rUnit⌣ (suc n) = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → rUnitₖ _ (f x)) + +-- TODO : Graded ring structure diff --git a/Cubical/ZCohomology/S1/S1.agda b/Cubical/ZCohomology/S1/S1.agda deleted file mode 100644 index 029a6eeed..000000000 --- a/Cubical/ZCohomology/S1/S1.agda +++ /dev/null @@ -1,36 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.S1.S1 where - -open import Cubical.ZCohomology.Base -open import Cubical.ZCohomology.Properties -open import Cubical.HITs.S1 -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Isomorphism -open import Cubical.HITs.SetTruncation -open import Cubical.HITs.Nullification -open import Cubical.Data.Int -open import Cubical.Data.Nat -open import Cubical.HITs.Truncation - ----- H⁰(S¹) = ℤ ---- - -coHom0-S1 : coHom zero S¹ ≡ Int -coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ setId isSetInt - where - helpLemma : (S¹ → Int) ≡ Int - helpLemma = isoToPath (iso fun funinv (λ _ → refl) (λ f → funExt (rinvLemma f))) - where - fun : (S¹ → Int) → Int - fun f = f base - - funinv : Int → (S¹ → Int) - funinv a base = a - funinv a (loop i) = a - - rinvLemma : (f : S¹ → Int) → (x : S¹) → funinv (fun f) x ≡ f x - rinvLemma f base = refl - rinvLemma f (loop i) j = isSetInt (f base) (f base) (λ k → f (loop k)) refl (~ j) i - -------------------------- - -{- TODO : give Hᵏ(S¹) for all k -} From d22f5ae330a0a67be0b41b762c5e6c7e8622ed14 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 6 Jul 2021 16:39:42 +0200 Subject: [PATCH 11/30] wups --- Cubical/Algebra/Everything.agda | 49 --------------------------- Cubical/Categories/Everything.agda | 20 ----------- Cubical/Data/Everything.agda | 49 --------------------------- Cubical/Displayed/Everything.agda | 16 --------- Cubical/Functions/Everything.agda | 13 ------- Cubical/HITs/Everything.agda | 46 ------------------------- Cubical/Homotopy/Everything.agda | 10 ------ Cubical/Induction/Everything.agda | 4 --- Cubical/Modalities/Everything.agda | 5 --- Cubical/Papers/Everything.agda | 6 ---- Cubical/Reflection/Everything.agda | 6 ---- Cubical/Structures/Everything.agda | 27 --------------- Cubical/Talks/Everything.agda | 4 --- Cubical/ZCohomology/#Benchmarks.agda# | 3 -- Cubical/ZCohomology/Benchmarks.agda | 4 --- Cubical/ZCohomology/Everything.agda | 20 ----------- 16 files changed, 282 deletions(-) delete mode 100644 Cubical/Algebra/Everything.agda delete mode 100644 Cubical/Categories/Everything.agda delete mode 100644 Cubical/Data/Everything.agda delete mode 100644 Cubical/Displayed/Everything.agda delete mode 100644 Cubical/Functions/Everything.agda delete mode 100644 Cubical/HITs/Everything.agda delete mode 100644 Cubical/Homotopy/Everything.agda delete mode 100644 Cubical/Induction/Everything.agda delete mode 100644 Cubical/Modalities/Everything.agda delete mode 100644 Cubical/Papers/Everything.agda delete mode 100644 Cubical/Reflection/Everything.agda delete mode 100644 Cubical/Structures/Everything.agda delete mode 100644 Cubical/Talks/Everything.agda delete mode 100644 Cubical/ZCohomology/#Benchmarks.agda# delete mode 100644 Cubical/ZCohomology/Benchmarks.agda delete mode 100644 Cubical/ZCohomology/Everything.agda diff --git a/Cubical/Algebra/Everything.agda b/Cubical/Algebra/Everything.agda deleted file mode 100644 index 1f1ed6ce3..000000000 --- a/Cubical/Algebra/Everything.agda +++ /dev/null @@ -1,49 +0,0 @@ -{-# OPTIONS --safe --guardedness #-} -module Cubical.Algebra.Everything where - -import Cubical.Algebra.AbGroup -import Cubical.Algebra.Algebra -import Cubical.Algebra.CommAlgebra -import Cubical.Algebra.CommAlgebra.Localisation -import Cubical.Algebra.CommAlgebra.Properties -import Cubical.Algebra.CommRing -import Cubical.Algebra.CommRing.BinomialThm -import Cubical.Algebra.CommRing.FGIdeal -import Cubical.Algebra.CommRing.Ideal -import Cubical.Algebra.CommRing.Instances.BiInvInt -import Cubical.Algebra.CommRing.Instances.Int -import Cubical.Algebra.CommRing.Localisation.Base -import Cubical.Algebra.CommRing.Localisation.InvertingElements -import Cubical.Algebra.CommRing.Localisation.UniversalProperty -import Cubical.Algebra.CommRing.RadicalIdeal -import Cubical.Algebra.FreeCommAlgebra -import Cubical.Algebra.Group -import Cubical.Algebra.Group.EilenbergMacLane1 -import Cubical.Algebra.Matrix -import Cubical.Algebra.Module -import Cubical.Algebra.Monoid -import Cubical.Algebra.Monoid.BigOp -import Cubical.Algebra.Ring -import Cubical.Algebra.Ring.BigOps -import Cubical.Algebra.Ring.QuotientRing -import Cubical.Algebra.RingSolver.AlgebraExpression -import Cubical.Algebra.RingSolver.AlmostRing -import Cubical.Algebra.RingSolver.CommRingAsAlmostRing -import Cubical.Algebra.RingSolver.CommRingEvalHom -import Cubical.Algebra.RingSolver.CommRingExamples -import Cubical.Algebra.RingSolver.CommRingHornerForms -import Cubical.Algebra.RingSolver.CommRingSolver -import Cubical.Algebra.RingSolver.EvaluationHomomorphism -import Cubical.Algebra.RingSolver.Examples -import Cubical.Algebra.RingSolver.HornerForms -import Cubical.Algebra.RingSolver.IntAsRawRing -import Cubical.Algebra.RingSolver.NatAsAlmostRing -import Cubical.Algebra.RingSolver.NatExamples -import Cubical.Algebra.RingSolver.RawAlgebra -import Cubical.Algebra.RingSolver.RawRing -import Cubical.Algebra.RingSolver.ReflectionSolving -import Cubical.Algebra.RingSolver.RingExpression -import Cubical.Algebra.RingSolver.Solver -import Cubical.Algebra.Semigroup -import Cubical.Algebra.SymmetricGroup -import Cubical.Algebra.ZariskiLattice.BasicOpens diff --git a/Cubical/Categories/Everything.agda b/Cubical/Categories/Everything.agda deleted file mode 100644 index 176b8ef2c..000000000 --- a/Cubical/Categories/Everything.agda +++ /dev/null @@ -1,20 +0,0 @@ -{-# OPTIONS --safe --guardedness #-} -module Cubical.Categories.Everything where - -import Cubical.Categories.Adjoint -import Cubical.Categories.Category -import Cubical.Categories.Commutativity -import Cubical.Categories.Constructions.Elements -import Cubical.Categories.Constructions.Slice -import Cubical.Categories.Equivalence -import Cubical.Categories.Functor -import Cubical.Categories.Instances.Cospan -import Cubical.Categories.Instances.Functors -import Cubical.Categories.Instances.Sets -import Cubical.Categories.Limits -import Cubical.Categories.Morphism -import Cubical.Categories.NaturalTransformation -import Cubical.Categories.Presheaf -import Cubical.Categories.Sets -import Cubical.Categories.TypesOfCategories.TypeCategory -import Cubical.Categories.Yoneda diff --git a/Cubical/Data/Everything.agda b/Cubical/Data/Everything.agda deleted file mode 100644 index bcb81412e..000000000 --- a/Cubical/Data/Everything.agda +++ /dev/null @@ -1,49 +0,0 @@ -{-# OPTIONS --safe --guardedness #-} -module Cubical.Data.Everything where - -import Cubical.Data.BinNat -import Cubical.Data.Bool -import Cubical.Data.Bool.SwitchStatement -import Cubical.Data.DescendingList -import Cubical.Data.Empty -import Cubical.Data.Equality -import Cubical.Data.Fin -import Cubical.Data.Fin.LehmerCode -import Cubical.Data.Fin.Recursive -import Cubical.Data.FinData -import Cubical.Data.FinInd -import Cubical.Data.FinSet -import Cubical.Data.FinSet.Binary.Large -import Cubical.Data.FinSet.Binary.Small -import Cubical.Data.Graph -import Cubical.Data.HomotopyGroup -import Cubical.Data.InfNat -import Cubical.Data.Int -import Cubical.Data.Int.MoreInts.BiInvInt -import Cubical.Data.Int.MoreInts.DeltaInt -import Cubical.Data.Int.MoreInts.DiffInt -import Cubical.Data.Int.MoreInts.QuoInt -import Cubical.Data.List -import Cubical.Data.Maybe -import Cubical.Data.Nat -import Cubical.Data.Nat.Algebra -import Cubical.Data.Nat.Coprime -import Cubical.Data.Nat.Divisibility -import Cubical.Data.Nat.GCD -import Cubical.Data.Nat.Literals -import Cubical.Data.Nat.Order -import Cubical.Data.Nat.Order.Recursive -import Cubical.Data.NatMinusOne -import Cubical.Data.NatPlusOne -import Cubical.Data.Prod -import Cubical.Data.Queue -import Cubical.Data.Queue.1List -import Cubical.Data.Queue.Truncated2List -import Cubical.Data.Queue.Untruncated2List -import Cubical.Data.Queue.Untruncated2ListInvariant -import Cubical.Data.Sigma -import Cubical.Data.SubFinSet -import Cubical.Data.Sum -import Cubical.Data.SumFin -import Cubical.Data.Unit -import Cubical.Data.Vec diff --git a/Cubical/Displayed/Everything.agda b/Cubical/Displayed/Everything.agda deleted file mode 100644 index d4e083945..000000000 --- a/Cubical/Displayed/Everything.agda +++ /dev/null @@ -1,16 +0,0 @@ -{-# OPTIONS --safe --guardedness #-} -module Cubical.Displayed.Everything where - -import Cubical.Displayed.Auto -import Cubical.Displayed.Base -import Cubical.Displayed.Constant -import Cubical.Displayed.Function -import Cubical.Displayed.Generic -import Cubical.Displayed.Morphism -import Cubical.Displayed.Prop -import Cubical.Displayed.Properties -import Cubical.Displayed.Record -import Cubical.Displayed.Sigma -import Cubical.Displayed.Subst -import Cubical.Displayed.Unit -import Cubical.Displayed.Universe diff --git a/Cubical/Functions/Everything.agda b/Cubical/Functions/Everything.agda deleted file mode 100644 index 08579b121..000000000 --- a/Cubical/Functions/Everything.agda +++ /dev/null @@ -1,13 +0,0 @@ -{-# OPTIONS --safe --guardedness #-} -module Cubical.Functions.Everything where - -import Cubical.Functions.Bundle -import Cubical.Functions.Embedding -import Cubical.Functions.Fibration -import Cubical.Functions.Fixpoint -import Cubical.Functions.FunExtEquiv -import Cubical.Functions.Implicit -import Cubical.Functions.Involution -import Cubical.Functions.Logic -import Cubical.Functions.Morphism -import Cubical.Functions.Surjection diff --git a/Cubical/HITs/Everything.agda b/Cubical/HITs/Everything.agda deleted file mode 100644 index 26ceba567..000000000 --- a/Cubical/HITs/Everything.agda +++ /dev/null @@ -1,46 +0,0 @@ -{-# OPTIONS --safe --guardedness #-} -module Cubical.HITs.Everything where - -import Cubical.HITs.2GroupoidTruncation -import Cubical.HITs.AssocList -import Cubical.HITs.Colimit -import Cubical.HITs.Cost -import Cubical.HITs.Cylinder -import Cubical.HITs.Delooping.Two.Base -import Cubical.HITs.Delooping.Two.Properties -import Cubical.HITs.DunceCap -import Cubical.HITs.EilenbergMacLane1 -import Cubical.HITs.FiniteMultiset -import Cubical.HITs.FiniteMultiset.CountExtensionality -import Cubical.HITs.GroupoidQuotients -import Cubical.HITs.GroupoidTruncation -import Cubical.HITs.Hopf -import Cubical.HITs.InfNat -import Cubical.HITs.Interval -import Cubical.HITs.Join -import Cubical.HITs.KleinBottle -import Cubical.HITs.ListedFiniteSet -import Cubical.HITs.Localization -import Cubical.HITs.MappingCones -import Cubical.HITs.Modulo -import Cubical.HITs.Nullification -import Cubical.HITs.PropositionalTruncation -import Cubical.HITs.PropositionalTruncation.Monad -import Cubical.HITs.Pushout -import Cubical.HITs.Pushout.Flattening -import Cubical.HITs.RPn -import Cubical.HITs.Rationals.HITQ -import Cubical.HITs.Rationals.QuoQ -import Cubical.HITs.Rationals.SigmaQ -import Cubical.HITs.S1 -import Cubical.HITs.S2 -import Cubical.HITs.S3 -import Cubical.HITs.SetQuotients -import Cubical.HITs.SetTruncation -import Cubical.HITs.SmashProduct -import Cubical.HITs.Sn -import Cubical.HITs.Susp -import Cubical.HITs.Torus -import Cubical.HITs.Truncation -import Cubical.HITs.TypeQuotients -import Cubical.HITs.Wedge diff --git a/Cubical/Homotopy/Everything.agda b/Cubical/Homotopy/Everything.agda deleted file mode 100644 index 02627af7d..000000000 --- a/Cubical/Homotopy/Everything.agda +++ /dev/null @@ -1,10 +0,0 @@ -{-# OPTIONS --safe --guardedness #-} -module Cubical.Homotopy.Everything where - -import Cubical.Homotopy.Base -import Cubical.Homotopy.Connected -import Cubical.Homotopy.EilenbergSteenrod -import Cubical.Homotopy.Freudenthal -import Cubical.Homotopy.Loopspace -import Cubical.Homotopy.MayerVietorisCofiber -import Cubical.Homotopy.WedgeConnectivity diff --git a/Cubical/Induction/Everything.agda b/Cubical/Induction/Everything.agda deleted file mode 100644 index aeb577db2..000000000 --- a/Cubical/Induction/Everything.agda +++ /dev/null @@ -1,4 +0,0 @@ -{-# OPTIONS --safe --guardedness #-} -module Cubical.Induction.Everything where - -import Cubical.Induction.WellFounded diff --git a/Cubical/Modalities/Everything.agda b/Cubical/Modalities/Everything.agda deleted file mode 100644 index 862b98edc..000000000 --- a/Cubical/Modalities/Everything.agda +++ /dev/null @@ -1,5 +0,0 @@ -{-# OPTIONS --safe --guardedness #-} -module Cubical.Modalities.Everything where - -import Cubical.Modalities.Lex -import Cubical.Modalities.Modality diff --git a/Cubical/Papers/Everything.agda b/Cubical/Papers/Everything.agda deleted file mode 100644 index 65ddd0b73..000000000 --- a/Cubical/Papers/Everything.agda +++ /dev/null @@ -1,6 +0,0 @@ -{-# OPTIONS --safe --guardedness #-} -module Cubical.Papers.Everything where - -import Cubical.Papers.RepresentationIndependence -import Cubical.Papers.Synthetic -import Cubical.Papers.ZCohomology diff --git a/Cubical/Reflection/Everything.agda b/Cubical/Reflection/Everything.agda deleted file mode 100644 index 143dd2abe..000000000 --- a/Cubical/Reflection/Everything.agda +++ /dev/null @@ -1,6 +0,0 @@ -{-# OPTIONS --safe --guardedness #-} -module Cubical.Reflection.Everything where - -import Cubical.Reflection.Base -import Cubical.Reflection.RecordEquiv -import Cubical.Reflection.StrictEquiv diff --git a/Cubical/Structures/Everything.agda b/Cubical/Structures/Everything.agda deleted file mode 100644 index 9b03cf61f..000000000 --- a/Cubical/Structures/Everything.agda +++ /dev/null @@ -1,27 +0,0 @@ -{-# OPTIONS --safe --guardedness #-} -module Cubical.Structures.Everything where - -import Cubical.Structures.Auto -import Cubical.Structures.Axioms -import Cubical.Structures.Constant -import Cubical.Structures.Function -import Cubical.Structures.LeftAction -import Cubical.Structures.Macro -import Cubical.Structures.Maybe -import Cubical.Structures.MultiSet -import Cubical.Structures.Parameterized -import Cubical.Structures.Pointed -import Cubical.Structures.Product -import Cubical.Structures.Queue -import Cubical.Structures.Record -import Cubical.Structures.Relational.Auto -import Cubical.Structures.Relational.Constant -import Cubical.Structures.Relational.Equalizer -import Cubical.Structures.Relational.Function -import Cubical.Structures.Relational.Macro -import Cubical.Structures.Relational.Maybe -import Cubical.Structures.Relational.Parameterized -import Cubical.Structures.Relational.Pointed -import Cubical.Structures.Relational.Product -import Cubical.Structures.Transfer -import Cubical.Structures.TypeEqvTo diff --git a/Cubical/Talks/Everything.agda b/Cubical/Talks/Everything.agda deleted file mode 100644 index a0dc05fb5..000000000 --- a/Cubical/Talks/Everything.agda +++ /dev/null @@ -1,4 +0,0 @@ -{-# OPTIONS --safe --guardedness #-} -module Cubical.Talks.Everything where - -import Cubical.Talks.EPA2020 diff --git a/Cubical/ZCohomology/#Benchmarks.agda# b/Cubical/ZCohomology/#Benchmarks.agda# deleted file mode 100644 index 8f05c6f26..000000000 --- a/Cubical/ZCohomology/#Benchmarks.agda# +++ /dev/null @@ -1,3 +0,0 @@ -{-# OPTIONS --safe #-} -module Cubical.ZCohomology.Benchmarks where -open import \ No newline at end of file diff --git a/Cubical/ZCohomology/Benchmarks.agda b/Cubical/ZCohomology/Benchmarks.agda deleted file mode 100644 index 959d3a13a..000000000 --- a/Cubical/ZCohomology/Benchmarks.agda +++ /dev/null @@ -1,4 +0,0 @@ -{-# OPTIONS --safe #-} -module Cubical.ZCohomology.Benchmarks where - - diff --git a/Cubical/ZCohomology/Everything.agda b/Cubical/ZCohomology/Everything.agda deleted file mode 100644 index fa3417d28..000000000 --- a/Cubical/ZCohomology/Everything.agda +++ /dev/null @@ -1,20 +0,0 @@ -{-# OPTIONS --safe --guardedness #-} -module Cubical.ZCohomology.Everything where - -import Cubical.ZCohomology.Base -import Cubical.ZCohomology.EilenbergSteenrodZ -import Cubical.ZCohomology.GroupStructure -import Cubical.ZCohomology.Groups.Connected -import Cubical.ZCohomology.Groups.KleinBottle -import Cubical.ZCohomology.Groups.Prelims -import Cubical.ZCohomology.Groups.RP2 -import Cubical.ZCohomology.Groups.Sn -import Cubical.ZCohomology.Groups.Torus -import Cubical.ZCohomology.Groups.Unit -import Cubical.ZCohomology.Groups.Wedge -import Cubical.ZCohomology.Groups.WedgeOfSpheres -import Cubical.ZCohomology.MayerVietorisUnreduced -import Cubical.ZCohomology.Properties -import Cubical.ZCohomology.RingStructure.CupProduct -import Cubical.ZCohomology.RingStructure.GradedCommutativity -import Cubical.ZCohomology.RingStructure.RingLaws From 2195d980053758eb3b1473c817bab42534d9a94f Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 6 Jul 2021 16:40:29 +0200 Subject: [PATCH 12/30] wups2 --- Cubical/Relation/Everything.agda | 10 ---------- 1 file changed, 10 deletions(-) delete mode 100644 Cubical/Relation/Everything.agda diff --git a/Cubical/Relation/Everything.agda b/Cubical/Relation/Everything.agda deleted file mode 100644 index 02399d7c8..000000000 --- a/Cubical/Relation/Everything.agda +++ /dev/null @@ -1,10 +0,0 @@ -{-# OPTIONS --safe --guardedness #-} -module Cubical.Relation.Everything where - -import Cubical.Relation.Binary -import Cubical.Relation.Binary.Poset -import Cubical.Relation.Nullary -import Cubical.Relation.Nullary.DecidableEq -import Cubical.Relation.Nullary.HLevels -import Cubical.Relation.ZigZag.Applications.MultiSet -import Cubical.Relation.ZigZag.Base From 8a92f66d978beb16a90679141c07963b6612a173 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 9 Jul 2021 03:21:55 +0200 Subject: [PATCH 13/30] Signed-off-by: aljungstrom --- Cubical/Experiments/ZCohomology/Benchmarks.agda | 2 +- Cubical/Papers/ZCohomology.agda | 6 +++--- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/Cubical/Experiments/ZCohomology/Benchmarks.agda b/Cubical/Experiments/ZCohomology/Benchmarks.agda index f1e403259..df174b3f0 100644 --- a/Cubical/Experiments/ZCohomology/Benchmarks.agda +++ b/Cubical/Experiments/ZCohomology/Benchmarks.agda @@ -380,6 +380,6 @@ module CP2-test₄ where ϕ⁻¹ = inv (fst H⁴CP²≅ℤ) {- - test₀ : ϕ (0ₕ _) ≡ true -- fails already here... + test₀ : ϕ (0ₕ _) ≡ 0 -- fails already here... test₀ = refl -} diff --git a/Cubical/Papers/ZCohomology.agda b/Cubical/Papers/ZCohomology.agda index 89c48ba36..1d557832d 100644 --- a/Cubical/Papers/ZCohomology.agda +++ b/Cubical/Papers/ZCohomology.agda @@ -236,7 +236,7 @@ open GroupStructure using ( rUnitₖ ; lUnitₖ n≥2-rUnit≡refl : {n : ℕ} → rUnitₖ (2 + n) (0ₖ (2 + n)) ≡ refl n≥2-rUnit≡refl = refl --- rUnitₖ (definitional) +-- lUnitₖ (definitional) 0-lUnit≡refl : lUnitₖ 0 (0ₖ 0) ≡ refl 1-lUnit≡refl : lUnitₖ 1 (0ₖ 1) ≡ refl n≥2-lUnit≡refl : {n : ℕ} → lUnitₖ (2 + n) (0ₖ (2 + n)) ≡ refl @@ -343,8 +343,8 @@ open Cup using (_⌣_) -- 4.2 -- Lemma 14 -Lem14 : (n m : ℕ) (f g : K∙ n →∙ K∙ m) → fst f ≡ fst g → f ≡ g -Lem14 n m f g p = Homogen.→∙Homogeneous≡ (Properties.isHomogeneousKn m) p +Lem14 : ∀ {ℓ} {A : Type∙ ℓ} (n : ℕ) (f g : A →∙ K∙ n) → fst f ≡ fst g → f ≡ g +Lem14 n f g p = Homogen.→∙Homogeneous≡ (Properties.isHomogeneousKn n) p -- Proposition 15 open ⌣Ring using (leftDistr-⌣ₖ ; rightDistr-⌣ₖ) From 66f9c88adbfa5694484295d781c8cf2b1726bf82 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sun, 11 Jul 2021 14:11:01 +0200 Subject: [PATCH 14/30] cleanup --- Cubical/Foundations/Path.agda | 53 +++++++++-------------------------- 1 file changed, 13 insertions(+), 40 deletions(-) diff --git a/Cubical/Foundations/Path.agda b/Cubical/Foundations/Path.agda index 29b7934f9..a4afb6c1b 100644 --- a/Cubical/Foundations/Path.agda +++ b/Cubical/Foundations/Path.agda @@ -191,51 +191,24 @@ Square≃doubleComp a₀₋ a₁₋ a₋₀ a₋₁ = transportEquiv (PathP≡do -- Flipping a square in Ω²A is the same as inverting it sym≡flipSquare : {x : A} (P : Square (refl {x = x}) refl refl refl) → sym P ≡ flipSquare P -sym≡flipSquare {x = x} = helper x x refl refl +sym≡flipSquare {x = x} P = sym (main refl P) where - helper : ∀ {ℓ} {A : Type ℓ} (x y : A) (p q : x ≡ y) (P : Square p q refl refl) - → PathP (λ j → PathP (λ i → Path A (p (i ∧ j)) (q (i ∨ ~ j))) - (λ k → q (~ j ∧ k)) - λ k → p (j ∨ k)) - (sym P) - (flipSquare P) - helper {A = A} x y = - J (λ y p → (q : x ≡ y) → (P : Square p q refl refl) → - PathP (λ j → PathP (λ i → Path A (p (i ∧ j)) (q (i ∨ ~ j))) - (λ k → q (~ j ∧ k)) - (λ k → p (j ∨ k))) - (sym P) - (flipSquare P)) - λ q → J (λ q P → PathP (λ j → PathP (λ i → Path A x (q (i ∨ ~ j))) - (λ k → q (~ j ∧ k)) refl) - (λ i → P (~ i)) λ i j → P j i) refl + B : (q : x ≡ x) → I → Type _ + B q i = PathP (λ j → x ≡ q (i ∨ j)) (λ k → q (i ∧ k)) refl + + main : (q : x ≡ x) (p : refl ≡ q) → PathP (λ i → B q i) (λ i j → p j i) (sym p) + main q = J (λ q p → PathP (λ i → B q i) (λ i j → p j i) (sym p)) refl -- Inverting both interval arguments of a square in Ω²A is the same as doing nothing sym-cong-sym≡id : {x : A} (P : Square (refl {x = x}) refl refl refl) → P ≡ λ i j → P (~ i) (~ j) -sym-cong-sym≡id P = - transport (λ i → doubleCompPath-filler (sym (rUnit refl)) P (lUnit refl) (~ i) - ≡ doubleCompPath-filler (sym (rUnit refl)) - (λ i j → P (~ i) (~ j)) (lUnit refl) (~ i)) (helper _ _ refl refl P) +sym-cong-sym≡id {x = x} P = sym (main refl P) where - helper : ∀ {ℓ} {A : Type ℓ} (x y : A) → (p q : x ≡ y) (P : Square p q refl refl) - → PathP (λ i → PathP (λ j → p i ≡ q (~ i)) - ((λ j → p (i ∨ j)) ∙ (λ j → q (~ i ∨ ~ j))) - ((λ j → p (i ∧ ~ j)) ∙ (λ j → q (~ i ∧ j)))) - (sym (rUnit p) ∙∙ P ∙∙ lUnit _) - (sym (lUnit (sym q)) ∙∙ (λ i j → P (~ i) (~ j)) ∙∙ rUnit (sym p)) - helper x y = - J (λ y p → (q : x ≡ y) (P : Square p q refl refl) - → PathP (λ i → PathP (λ j → p i ≡ q (~ i)) ((λ j → p (i ∨ j)) ∙ (λ j → q (~ i ∨ ~ j))) - ((λ j → p (i ∧ ~ j)) ∙ (λ j → q (~ i ∧ j)))) - (sym (rUnit p) ∙∙ P ∙∙ lUnit _) - (sym (lUnit (sym q)) ∙∙ (λ i j → P (~ i) (~ j)) ∙∙ rUnit (sym p))) - λ q → - J (λ q P → PathP (λ i → (λ j → x) ∙ (λ j → q (~ i ∨ ~ j)) ≡ - (λ j → x) ∙ (λ j → q (~ i ∧ j))) - ((λ i → rUnit refl (~ i)) ∙∙ P ∙∙ lUnit q) - ((λ i → lUnit (λ i₁ → q (~ i₁)) (~ i)) ∙∙ (λ i j → P (~ i) (~ j)) ∙∙ rUnit refl)) - refl + B : (q : x ≡ x) → I → Type _ + B q i = Path (x ≡ q i) (λ j → q (i ∨ ~ j)) λ j → q (i ∧ j) + + main : (q : x ≡ x) (p : refl ≡ q) → PathP (λ i → B q i) (λ i j → p (~ i) (~ j)) p + main q = J (λ q p → PathP (λ i → B q i) (λ i j → p (~ i) (~ j)) p) refl -- Applying cong sym is the same as flipping a square in Ω²A flipSquare≡cong-sym : ∀ {ℓ} {A : Type ℓ} {x : A} (P : Square (refl {x = x}) refl refl refl) @@ -245,7 +218,7 @@ flipSquare≡cong-sym P = sym (sym≡flipSquare P) ∙ sym (sym-cong-sym≡id (c -- Applying cong sym is the same as inverting a square in Ω²A sym≡cong-sym : ∀ {ℓ} {A : Type ℓ} {x : A} (P : Square (refl {x = x}) refl refl refl) → sym P ≡ cong sym P -sym≡cong-sym P = sym≡flipSquare _ ∙ flipSquare≡cong-sym P +sym≡cong-sym P = sym-cong-sym≡id (sym P) -- sym induces an equivalence on identity types of paths symIso : {a b : A} (p q : a ≡ b) → Iso (p ≡ q) (q ≡ p) From 538c4a877c7ed8fafff24cd015b9f5ffac41afb0 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sun, 11 Jul 2021 14:19:45 +0200 Subject: [PATCH 15/30] merge --- Cubical/Experiments/ZCohomology/Benchmarks.agda | 5 ----- 1 file changed, 5 deletions(-) diff --git a/Cubical/Experiments/ZCohomology/Benchmarks.agda b/Cubical/Experiments/ZCohomology/Benchmarks.agda index dae87fedb..de982fc14 100644 --- a/Cubical/Experiments/ZCohomology/Benchmarks.agda +++ b/Cubical/Experiments/ZCohomology/Benchmarks.agda @@ -29,14 +29,9 @@ open import Cubical.Algebra.Group hiding (ℤ ; Bool) open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.Properties open import Cubical.ZCohomology.GroupStructure hiding (_+ₕ_) renaming (_+'ₕ_ to _+ₕ_) -<<<<<<< HEAD --- _+'ₕ_ is just (λ x y → (x +ₕ 0ₕ) +ₕ (x +ₕ 0ₕ)) --- For technical reason, this gives nicer reductions and computes better -======= {- _+'ₕ_ is just (λ x y → (x +ₕ 0ₕ) +ₕ (y +ₕ 0ₕ)) For technical reason, this gives nicer reductions and computes better in higher dimensions. -} ->>>>>>> ff94da5ddbae97b810c2d8d67802dcde6792e9de open import Cubical.ZCohomology.Groups.Sn open import Cubical.ZCohomology.Groups.Wedge open import Cubical.ZCohomology.Groups.Torus From b2e69762b7f790c517432067109cb04d734cda0c Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 13 Jul 2021 03:58:35 +0200 Subject: [PATCH 16/30] stuff --- Cubical/HITs/EilenbergMacLane1/Base.agda | 18 + .../HITs/EilenbergMacLane1/Properties.agda | 616 +++++++++++++++++- 2 files changed, 633 insertions(+), 1 deletion(-) diff --git a/Cubical/HITs/EilenbergMacLane1/Base.agda b/Cubical/HITs/EilenbergMacLane1/Base.agda index b28a48d27..957d4bd86 100644 --- a/Cubical/HITs/EilenbergMacLane1/Base.agda +++ b/Cubical/HITs/EilenbergMacLane1/Base.agda @@ -9,6 +9,8 @@ This file contains: module Cubical.HITs.EilenbergMacLane1.Base where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws + renaming (assoc to assoc∙) open import Cubical.Algebra.Group.Base private @@ -41,3 +43,19 @@ module _ (Group@(G , str) : Group ℓ) where emloop-comp g h i = compPath-unique refl (emloop g) (emloop h) (emloop (g · h) , emcomp g h) (emloop g ∙ emloop h , compPath-filler (emloop g) (emloop h)) i .fst + + emloop-1g : emloop 1g ≡ refl + emloop-1g = + lUnit (emloop 1g) + ∙∙ cong (_∙ emloop 1g) (sym (lCancel (emloop 1g)) ) + ∙∙ sym (assoc∙ _ _ _) + ∙∙ cong (sym (emloop 1g) ∙_) (sym (emloop-comp 1g 1g) ∙ cong emloop (lid 1g)) + ∙∙ rCancel _ + + emloop-sym : (g : G) → emloop (inv g) ≡ sym (emloop g) + emloop-sym g = + rUnit _ + ∙∙ cong (emloop (inv g) ∙_) (sym (rCancel (emloop g))) + ∙∙ assoc∙ _ _ _ + ∙∙ cong (_∙ sym (emloop g)) (sym (emloop-comp (inv g) g) ∙∙ cong emloop (invl g) ∙∙ emloop-1g) + ∙∙ sym (lUnit _) diff --git a/Cubical/HITs/EilenbergMacLane1/Properties.agda b/Cubical/HITs/EilenbergMacLane1/Properties.agda index 7af601e5d..cdd312941 100644 --- a/Cubical/HITs/EilenbergMacLane1/Properties.agda +++ b/Cubical/HITs/EilenbergMacLane1/Properties.agda @@ -4,7 +4,7 @@ Eilenberg–Mac Lane type K(G, 1) -} -{-# OPTIONS --cubical --no-import-sorts --safe #-} +{-# OPTIONS --cubical --no-import-sorts --safe --experimental-lossy-unification #-} module Cubical.HITs.EilenbergMacLane1.Properties where open import Cubical.HITs.EilenbergMacLane1.Base @@ -15,10 +15,16 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws open import Cubical.Data.Sigma +open import Cubical.Data.Empty renaming (rec to ⊥-rec) hiding (elim) + open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Properties + +open import Cubical.Algebra.AbGroup.Base open import Cubical.HITs.PropositionalTruncation as PropTrunc using (∥_∥; ∣_∣; squash) open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂; ∣_∣₂; squash₂) @@ -31,6 +37,29 @@ module _ ((G , str) : Group ℓG) where open GroupStr str + elimGroupoid : + {B : EM₁ (G , str) → Type ℓ} + → ((x : EM₁ (G , str)) → isGroupoid (B x)) + → (b : B embase) + → (bloop : ((g : G) → PathP (λ i → B (emloop g i)) b b)) + → ((g h : G) → PathP (λ i → PathP (λ j → B (emcomp g h j i)) + (bloop g i) (bloop (g · h) i)) (λ _ → b) (bloop h)) + → (x : EM₁ (G , str)) + → B x + elimGroupoid Bgroup b bloop bcomp embase = b + elimGroupoid Bgroup b bloop bcomp (emloop x i) = bloop x i + elimGroupoid Bgroup b bloop bcomp (emcomp g h j i) = bcomp g h i j + elimGroupoid {B = B} Bgroup b bloop bcomp (emsquash g h p q r s i j k) = help i j k + where + help : PathP (λ i → PathP (λ j → PathP (λ k → B (emsquash g h p q r s i j k)) + (elimGroupoid Bgroup b bloop bcomp g) + (elimGroupoid Bgroup b bloop bcomp h)) + (λ k → elimGroupoid Bgroup b bloop bcomp (p k)) + λ k → elimGroupoid Bgroup b bloop bcomp (q k)) + (λ j k → elimGroupoid Bgroup b bloop bcomp (r j k)) + λ j k → elimGroupoid Bgroup b bloop bcomp (s j k) + help = toPathP (isOfHLevelPathP' 1 (isOfHLevelPathP' 2 (Bgroup _) _ _) _ _ _ _) + elimSet : {B : EM₁ (G , str) → Type ℓ} → ((x : EM₁ (G , str)) → isSet (B x)) → (b : B embase) @@ -98,3 +127,588 @@ module _ ((G , str) : Group ℓG) where → (x : EM₁ (G , str)) → B rec Bgpd = elim (λ _ → Bgpd) + + +open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) +open import Cubical.Data.Nat +open import Cubical.HITs.Susp +open import Cubical.Foundations.Pointed + +Susp̂ : (n : ℕ) → Type ℓG → Type ℓG +Susp̂ zero A = A +Susp̂ (suc n) A = Susp (Susp̂ n A) + +pt-Susp̂ : (n : ℕ) (A : Pointed ℓG) → Susp̂ n (typ A) +pt-Susp̂ zero A = snd A +pt-Susp̂ (suc n) A = north + +ptS : {n : ℕ} {G : Group ℓG} → Susp̂ n (EM₁ G) +ptS {n = n} {G = G} = pt-Susp̂ n (EM₁ G , embase) + +EM-raw : (G : AbGroup ℓG) (n : ℕ) → Type ℓG +EM-raw G zero = fst G +EM-raw G (suc n) = Susp̂ n (EM₁ (AbGroup→Group G)) + +EM-raw-elim : (G : AbGroup ℓG) (n : ℕ) {A : EM-raw G (suc n) → Type ℓ} + → ((x : _) → isOfHLevel (suc n) (A x) ) + → A ptS + → (x : _) → A x +EM-raw-elim G zero hlev b = elimProp _ hlev b +EM-raw-elim G (suc n) hlev b north = b +EM-raw-elim G (suc n) {A = A} hlev b south = subst A (merid ptS) b +EM-raw-elim G (suc n) {A = A} hlev b (merid a i) = help a i + where + help : (a : _) → PathP (λ i → A (merid a i)) b (subst A (merid ptS) b) + help = EM-raw-elim G n (λ _ → isOfHLevelPathP' (suc n) (hlev _) _ _) + λ i → transp (λ j → A (merid ptS (i ∧ j))) (~ i) b + +EM∙ : (G : AbGroup ℓG) (n : ℕ) → Pointed ℓG +EM∙ G zero = fst G , AbGroupStr.0g (snd G) +EM∙ G (suc zero) = EM₁ (AbGroup→Group G) , embase +EM∙ G (suc (suc n)) = hLevelTrunc (4 + n) (Susp̂ (suc n) (EM₁ (AbGroup→Group G))) , ∣ north ∣ + +EM : (G : AbGroup ℓG) (n : ℕ) → Type ℓG +EM G zero = EM-raw G zero +EM G (suc zero) = EM-raw G 1 +EM G (suc (suc n)) = hLevelTrunc (4 + n) (EM-raw G (suc (suc n))) + +wedgeConFun' : (G H : AbGroup ℓG) (n : ℕ) → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ} + → ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y)) + → (f : (x : _) → A ptS x) + → (g : (x : _) → A x ptS) + → f ptS ≡ g ptS + → (x : _) (y : _) → A x y +wedgeConFun'ᵣ : (G H : AbGroup ℓG) (n : ℕ) → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ} + → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y))) + → (f : (x : _) → A ptS x) + → (g : (x : _) → A x ptS) + → (p : f ptS ≡ g ptS) + → (x : _) → wedgeConFun' G H n hLev f g p x ptS ≡ g x +wedgeConFun' G H zero {A = A} hlev f g p = + elimSet _ (λ _ → isSetΠ λ _ → hlev _ _) f k + where + I→A : (x : fst G) → (i : I) → A (emloop x i) embase + I→A x i = + hcomp (λ k → λ {(i = i0) → p (~ k) ; (i = i1) → p (~ k)}) + (g (emloop x i)) + + SquareP2 : (x : _) (z : _) + → SquareP (λ i j → A (emloop x i) (emloop z j)) + (cong f (emloop z)) (cong f (emloop z)) + (λ i → I→A x i) λ i → I→A x i + SquareP2 x z = + toPathP (isOfHLevelPathP' 1 (λ _ _ → hlev _ _ _ _) _ _ _ _) + + CubeP2 : (x : _) (g h : _) + → PathP (λ k → PathP (λ j → PathP (λ i → A (emloop x i) (emcomp g h j k)) + (f (emcomp g h j k)) (f (emcomp g h j k))) + (λ i → SquareP2 x g i k) λ i → SquareP2 x ((snd (AbGroup→Group H) GroupStr.· g) h) i k) + (λ _ i → I→A x i) λ j i → SquareP2 x h i j + CubeP2 x g h = toPathP (isOfHLevelPathP' 1 (isOfHLevelPathP 2 (hlev _ _) _ _) _ _ _ _) + + k : (x : _) → PathP (λ i → (y : EM₁ (AbGroup→Group H)) → A (emloop x i) y) f f + k x i embase = I→A x i + k x i (emloop z j) = SquareP2 x z i j + k x i (emcomp g h j k) = CubeP2 x g h k j i + k x i (emsquash y z p q r s j k' l) = mega i j k' l + where + mega : + PathP (λ i → + PathP (λ j → + PathP (λ k' → + PathP (λ l → A (emloop x i) (emsquash y z p q r s j k' l)) + (k x i y) + (k x i z)) + (λ l → k x i (p l)) + λ l → k x i (q l)) + (λ k' l → k x i (r k' l)) + λ k' l → k x i (s k' l)) + (λ j k l → f (emsquash y z p q r s j k l)) + λ j k l → f (emsquash y z p q r s j k l) + mega = + toPathP (isOfHLevelPathP' 1 + (isOfHLevelPathP 2 (isOfHLevelPathP 2 (hlev _ _) _ _) _ _) _ _ _ _) +wedgeConFun' G H (suc n) {A = A} hLev f g p north y = f y +wedgeConFun' G H (suc n) {A = A} hLev f g p south y = subst (λ x → A x y) (merid ptS) (f y) +wedgeConFun' G H (suc n) {A = A} hLev f g p (merid a i) y = mainₗ a y i + module _ where + llem₁ : g south ≡ subst (λ x₁ → A x₁ ptS) (merid ptS) (f ptS) + llem₁ = (λ i → transp (λ j → A (merid ptS (j ∨ ~ i)) ptS) (~ i) (g (merid ptS (~ i)))) + ∙ cong (subst (λ x₁ → A x₁ ptS) (merid ptS)) (sym p) + + llem₂' : + Square + (λ i → transp (λ j → A (merid ptS (i ∨ j)) ptS) i (g (merid ptS i))) refl + (cong (subst (λ x → A x ptS) (merid ptS)) (sym p)) llem₁ + llem₂' i j = + hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptS (j ∨ z)) ptS) j + (g (merid ptS j)) + ; (i = i1) → subst (λ x₁ → A x₁ ptS) (merid ptS) (p (~ k)) + ; (j = i0) → (subst (λ x → A x ptS) (merid ptS)) (p (~ k ∨ ~ i))}) + (transp (λ k → A (merid ptS (k ∨ ~ i ∧ j)) ptS) (~ i ∧ j) (g (merid ptS (~ i ∧ j)))) + + llem₂ : (λ i₁ → transp (λ j → A (merid ptS (i₁ ∧ j)) ptS) (~ i₁) (f ptS)) + ≡ (λ i₁ → hcomp (λ k → λ { (i₁ = i0) → p (~ k) ; (i₁ = i1) → llem₁ k }) + (g (merid ptS i₁))) + llem₂ i j = + hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptS (j ∧ j₁)) ptS) (~ j) (p (~ k)) + ; (j = i0) → p (~ k) + ; (j = i1) → llem₂' k i}) + (transp (λ k → A (merid ptS ((i ∨ k) ∧ j)) ptS) (i ∨ ~ j) (g (merid ptS (i ∧ j)))) + + mainₗ : (a : _) (y : _) + → PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) + mainₗ = + wedgeConFun' G H n + (λ _ _ → isOfHLevelPathP' (suc (n + 1)) (hLev _ _) _ _) + (λ x i → transp (λ j → A (merid ptS (i ∧ j)) x) (~ i) (f x)) + (λ x i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → llem₁ k}) + (g (merid x i))) + llem₂ + + mainP : (y : _) + → mainₗ y ptS + ≡ λ i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → llem₁ k}) + (g (merid y i)) + mainP y = + wedgeConFun'ᵣ G H n + (λ _ _ → isOfHLevelPathP' (suc (n + 1)) (hLev _ _) _ _) + (λ x i → transp (λ j → A (merid ptS (i ∧ j)) x) (~ i) (f x)) + (λ x i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → llem₁ k}) + (g (merid x i))) + llem₂ y +wedgeConFun'ᵣ G H zero {A = A} hLev f g p = + elimProp _ (λ _ → hLev _ _ _ _) p +wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p north = p +wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p south = sym (llem₁ G H n hLev f g p ptS i0 ptS) +wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p (merid a i) k = P k i + where + lem : _ + lem i j = + hcomp (λ k → λ { (i = i1) → g (merid a j) + ; (j = i0) → p (i ∨ ~ k) + ; (j = i1) → llem₁ G H n hLev f g p ptS i0 ptS (~ i ∧ k)}) + (g (merid a j)) + + P : PathP (λ k → PathP (λ i → A (merid a i) ptS) + (p k) (llem₁ G H n hLev f g p ptS i0 ptS (~ k))) + (λ i → mainₗ G H n hLev f g p a i ptS a ptS i) λ i → g (merid a i) + P = mainP G H n hLev f g p a i0 ptS a ◁ lem + +private + wedgeConFun : (G H : AbGroup ℓG) (k n m : ℕ) → (n + m ≡ k) → {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} + → ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y)) + → (f : (x : _) → A ptS x) + → (g : (x : _) → A x ptS) + → f ptS ≡ g ptS + → (x : _) (y : _) → A x y + wedgeconLeft : (G H : AbGroup ℓG) (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} + → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) + → (f : (x : _) → A ptS x) + → (g : (x : _) → A x ptS) + → (p : f ptS ≡ g ptS) + → (x : _) → wedgeConFun G H k n m P hLev f g p ptS x ≡ f x + wedgeconRight : (G H : AbGroup ℓG) (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} + → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) + → (f : (x : _) → A ptS x) + → (g : (x : _) → A x ptS) + → (p : f ptS ≡ g ptS) + → (x : _) → wedgeConFun G H k n m P hLev f g p x ptS ≡ g x + wedgeConFun G H k n zero P {A = A} hLev f g p = wedgeConFun' G H n hLev f g p + wedgeConFun G H k zero (suc m) P {A = A} hLev f g p x y = + wedgeConFun' H G (suc m) {A = λ x y → A y x} + (λ x y → subst (λ n → isOfHLevel (2 + n) (A y x)) (+-comm 1 m) (hLev y x)) + g f (sym p) y x + wedgeConFun G H l (suc n) (suc m) P {A = A} hlev f g p north y = f y + wedgeConFun G H l (suc n) (suc m) P {A = A} hlev f g p south y = subst (λ x → A x y) (merid ptS) (f y) + wedgeConFun G H zero (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = main** i + module _ where + main** : PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) + main** = ⊥-rec (snotz P) + wedgeConFun G H (suc l) (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = main* a y i + module _ where + lem₁* : g south ≡ (subst (λ x → A x ptS) (merid ptS) (f ptS)) + lem₁* = ((λ i → transp (λ j → A (merid ptS (~ i ∨ j)) ptS) (~ i) (g (merid ptS (~ i))))) + ∙ cong (subst (λ x → A x ptS) (merid ptS)) (sym p) + + lem₁'* : + Square + (λ i → transp (λ j → A (merid ptS (~ i ∨ j)) ptS) (~ i) (g (merid ptS (~ i)))) + (refl {x = subst (λ x → A x ptS) (merid ptS) (f ptS)}) + lem₁* + ((cong (transport (λ z → A (merid ptS z) ptS)) (sym p))) + lem₁'* i j = + hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptS (~ j ∨ j₁)) ptS) (~ j) (g (merid ptS (~ j))) + ; (i = i1) → subst (λ x → A x ptS) (merid ptS) (p (~ k)) + ; (j = i1) → cong (transport (λ z → A (merid ptS z) ptS)) (sym p) (i ∧ k)}) + (transp (λ j₁ → A (merid ptS ((~ j ∧ ~ i) ∨ j₁)) ptS) (~ j ∧ ~ i) (g (merid ptS (~ j ∧ ~ i)))) + + + lem₂* : (λ i₁ → transp (λ j → A (merid ptS (j ∧ i₁)) ptS) (~ i₁) (f ptS)) + ≡ (λ i₁ → hcomp (λ k → λ { (i₁ = i0) → p (~ k) ; (i₁ = i1) → lem₁* k }) + (g (merid ptS i₁))) + lem₂* i j = + hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptS (z ∧ j)) ptS) (~ j) (p (~ k)) + ; (j = i0) → p (~ k) + ; (j = i1) → lem₁'* k (~ i)}) + (transp (λ z → A (merid ptS ((i ∨ z) ∧ j)) ptS) (i ∨ ~ j) (g (merid ptS (i ∧ j)))) + + main* : (a : _) (y : _) → PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) + main* = + wedgeConFun G H l n (suc m) (cong predℕ P) + (λ _ _ → isOfHLevelPathP' (suc (n + (suc (suc m)))) (hlev _ _) _ _) + (λ x i → transp (λ j → A (merid ptS (j ∧ i)) x) (~ i) (f x)) + (λ y i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → lem₁* k}) + (g (merid y i))) + lem₂* + + mainR : (x : _) → main* x ptS + ≡ λ i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → lem₁* k}) + (g (merid x i)) + mainR x = + wedgeconRight G H l n (suc m) (cong predℕ P) + (λ _ _ → isOfHLevelPathP' (suc (n + (suc (suc m)))) (hlev _ _) _ _) + (λ x i → transp (λ j → A (merid ptS (j ∧ i)) x) (~ i) (f x)) + (λ y i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → lem₁* k}) + (g (merid y i))) + lem₂* x + wedgeconLeft G H k zero zero P {A = A} hLev f g p _ = refl + wedgeconLeft G H k (suc n) zero P {A = A} hLev f g p _ = refl + wedgeconLeft G H k zero (suc m) P {A = A} hLev f g p x = + wedgeConFun'ᵣ H G (suc m) + (λ x₁ y → + subst (λ n → (x₂ y₁ : A y x₁) → isOfHLevel (suc n) (x₂ ≡ y₁)) + (+-comm 1 m) (hLev y x₁)) + g f (λ i → p (~ i)) x + wedgeconLeft G H k (suc n) (suc m) P {A = A} hLev f g p _ = refl + wedgeconRight G H k n zero P {A = A} hLev f g p = wedgeConFun'ᵣ G H n hLev f g p + wedgeconRight G H k zero (suc m) P {A = A} hLev f g p _ = refl + wedgeconRight G H zero (suc n) (suc m) P {A = A} hLev f g p x = ⊥-rec (snotz P) + wedgeconRight G H l (suc n) (suc m) P {A = A} hLev f g p north = p + wedgeconRight G H l (suc n) (suc m) P {A = A} hLev f g p south = + sym (lem₁* G H _ n m refl hLev f g p ptS i0 ptS) + wedgeconRight G H (suc l) (suc n) (suc m) P {A = A} hLev f g p (merid a i) k = + help k i + where + lem : _ + lem i j = + hcomp (λ k → λ { (i = i1) → g (merid a j) + ; (j = i0) → p (i ∨ ~ k) + ; (j = i1) → lem₁* G H l n m P hLev f g p ptS i0 ptS (~ i ∧ k)}) + (g (merid a j)) + + help : PathP (λ k → PathP (λ i → A (merid a i) ptS) + (p k) (lem₁* G H l n m P hLev f g p ptS i0 ptS (~ k))) + (λ i → main* G H l n m P hLev f g p a i north a north i) (cong g (merid a)) + help = mainR G H l n m P hLev f g p a i0 ptS a ◁ lem + + +module wedgeConEM (G H : AbGroup ℓG) (n m : ℕ) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} + (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) + (f : (x : _) → A ptS x) + (g : (x : _) → A x ptS) + (p : f ptS ≡ g ptS) where + fun : (x : EM-raw G (suc n)) (y : EM-raw H (suc m)) → A x y + fun = wedgeConFun G H (n + m) n m refl hLev f g p + + left : (x : EM-raw H (suc m)) → fun ptS x ≡ f x + left = wedgeconLeft G H (n + m) n m refl hLev f g p + + right : (x : EM-raw G (suc n)) → fun x ptS ≡ g x + right = wedgeconRight G H (n + m) n m refl hLev f g p + +module _ {G : AbGroup ℓG} where + infixr 34 _+ₖ_ + infixr 34 _-ₖ_ + open AbGroupStr (snd G) renaming (_+_ to _+G_ ; -_ to -G_ ; assoc to assocG) + + private + help : (n : ℕ) → n + (4 + n) ≡ (2 + n) + (2 + n) + help n = +-suc n (3 + n) ∙ cong suc (+-suc n (suc (suc n))) + + hLevHelp : (n : ℕ) → isOfHLevel ((2 + n) + (2 + n)) (EM G (2 + n)) + hLevHelp n = + transport (λ i → isOfHLevel (help n i) (EM G (2 + n))) + (isOfHLevelPlus {n = 4 + n} n (isOfHLevelTrunc (4 + n))) + + helper : (g h : fst G) + → PathP (λ i → Path (EM₁ (AbGroup→Group G)) + (emloop h i) (emloop h i)) (emloop g) (emloop g) + helper g h = + compPathL→PathP + (cong (sym (emloop h) ∙_) + (sym (emloop-comp _ g h) + ∙∙ cong emloop (comm g h) + ∙∙ emloop-comp _ h g) + ∙∙ assoc _ _ _ + ∙∙ cong (_∙ emloop g) (lCancel _) + ∙ sym (lUnit _)) + open import Cubical.Foundations.Path + _+ₖ_ : {n : ℕ} → EM G n → EM G n → EM G n + _+ₖ_ {n = zero} = _+G_ + _+ₖ_ {n = suc zero} = + rec _ (isGroupoidΠ (λ _ → emsquash)) + (λ x → x) + hi! + λ g h i j x → el g h x i j + where + lol : (g h : fst G) + → Square (emloop g) (emloop (g +G h)) refl (emloop h) + lol g h = compPath-filler (emloop g) (emloop h) ▷ sym (emloop-comp _ g h) + + hi! : fst G → (λ x → x) ≡ (λ x → x) + hi! g = funExt (elimSet _ (λ _ → emsquash _ _) + (emloop g) + (helper g)) + el : (g h : fst G) (x : EM₁ (AbGroup→Group G)) + → Square (λ j → hi! g j x) (λ j → hi! ((snd (AbGroup→Group G) GroupStr.· g) h) j x) + refl λ j → hi! h j x + el g h = + elimProp _ (λ _ → isOfHLevelPathP' 1 (emsquash _ _) _ _) + (lol g h) + _+ₖ_ {n = suc (suc n)} = + trRec2 (isOfHLevelTrunc (4 + n)) + (wedgeConEM.fun G G (suc n) (suc n) + (λ _ _ → hLevHelp n) + ∣_∣ ∣_∣ refl) + + σ-EM : (n : ℕ) → EM-raw G (suc n) → Path (EM-raw G (2 + n)) ptS ptS + σ-EM n x i = (merid x ∙ sym (merid ptS)) i + + -ₖ_ : {n : ℕ} → EM G n → EM G n + -ₖ_ {n = zero} x = -G x + -ₖ_ {n = suc zero} = + rec _ emsquash + ptS + (λ g → sym (emloop g)) + λ g h → sym (emloop-sym _ g) + ◁ (flipSquare + (flipSquare (emcomp (-G g) (-G h)) + ▷ emloop-sym _ h) + ▷ (cong emloop (comm (-G g) (-G h) + ∙ sym (GroupTheory.invDistr (AbGroup→Group G) g h)) + ∙ emloop-sym _ (g +G h))) + -ₖ_ {n = suc (suc n)} = + map λ { north → north + ; south → north + ; (merid a i) → σ-EM n a (~ i)} + + _-ₖ_ : {n : ℕ} → EM G n → EM G n → EM G n + _-ₖ_ {n = n} x y = _+ₖ_ {n = n} x (-ₖ_ {n = n} y) + + +ₖ-syntax : (n : ℕ) → EM G n → EM G n → EM G n + +ₖ-syntax n = _+ₖ_ {n = n} + + -ₖ-syntax : (n : ℕ) → EM G n → EM G n + -ₖ-syntax n = -ₖ_ {n = n} + + -'ₖ-syntax : (n : ℕ) → EM G n → EM G n → EM G n + -'ₖ-syntax n = _-ₖ_ {n = n} + + syntax +ₖ-syntax n x y = x +[ n ]ₖ y + syntax -ₖ-syntax n x = -[ n ]ₖ x + syntax -'ₖ-syntax n x y = x -[ n ]ₖ y + + 0ₖ : (n : ℕ) → EM G n + 0ₖ zero = 0g + 0ₖ (suc zero) = ptS + 0ₖ (suc (suc n)) = ∣ ptS ∣ + + lUnitₖ : (n : ℕ) (x : EM G n) → 0ₖ n +[ n ]ₖ x ≡ x + lUnitₖ zero x = lid x + lUnitₖ (suc zero) _ = refl + lUnitₖ (suc (suc n)) = + trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ _ → refl + + rUnitₖ : (n : ℕ) (x : EM G n) → x +[ n ]ₖ 0ₖ n ≡ x + rUnitₖ zero x = rid x + rUnitₖ (suc zero) = + elimSet _ (λ _ → emsquash _ _) + refl + λ _ _ → refl + rUnitₖ (suc (suc n)) = + trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + (wedgeConEM.right G G (suc n) (suc n) + (λ _ _ → hLevHelp n) + ∣_∣ ∣_∣ refl) + + commₖ : (n : ℕ) (x y : EM G n) → x +[ n ]ₖ y ≡ y +[ n ]ₖ x + commₖ zero = comm + commₖ (suc zero) = + wedgeConEM.fun G G 0 0 (λ _ _ → emsquash _ _) + (λ x → sym (rUnitₖ 1 x)) + (rUnitₖ 1) + refl + commₖ (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelTruncPath {n = 4 + n}) + (wedgeConEM.fun G G _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (hLevHelp n) _ _) + (λ x → sym (rUnitₖ (2 + n) ∣ x ∣)) + (λ x → rUnitₖ (2 + n) ∣ x ∣) + refl) + + open import Cubical.Homotopy.Loopspace + cong₂+₁ : (p q : typ (Ω (EM∙ G 1))) → cong₂ (λ x y → x +[ 1 ]ₖ y) p q ≡ p ∙ q + cong₂+₁ p q = + (cong₂Funct (λ x y → x +[ 1 ]ₖ y) p q) + ∙ (λ i → (cong (λ x → rUnitₖ 1 x i) p) ∙ (cong (λ x → lUnitₖ 1 x i) q)) + + cong₂+₂ : (n : ℕ) (p q : typ (Ω (EM∙ G (2 + n)))) → cong₂ (λ x y → x +[ (2 + n) ]ₖ y) p q ≡ p ∙ q + cong₂+₂ n p q = + (cong₂Funct (λ x y → x +[ (2 + n) ]ₖ y) p q) + ∙ (λ i → (cong (λ x → rUnitₖ (2 + n) x i) p) ∙ (cong (λ x → lUnitₖ (2 + n) x i) q)) + + isCommΩEM : (n : ℕ) (p q : typ (Ω (EM∙ G (suc n)))) → p ∙ q ≡ q ∙ p + isCommΩEM zero p q = + sym (cong₂+₁ p q) + ∙∙ (λ i j → commₖ 1 (p j) (q j) i) + ∙∙ cong₂+₁ q p + isCommΩEM (suc n) p q = + (sym (cong₂+₂ n p q) + ∙∙ (λ i j → commₖ (suc (suc n)) (p j) (q j) i) + ∙∙ cong₂+₂ n q p) + + cong-₁ : (p : typ (Ω (EM∙ G 1))) → cong (λ x → -[ 1 ]ₖ x) p ≡ sym p + cong-₁ p = main ptS p + where + decoder : (x : EM G 1) → ptS ≡ x → x ≡ ptS + decoder = + elimSet _ + (λ _ → isSetΠ λ _ → emsquash _ _) + (λ p i → -[ 1 ]ₖ (p i)) + λ g → toPathP + (funExt λ x → + (λ i → transport (λ i → Path (EM G 1) (emloop g i) ptS) + (cong (-ₖ_ {n = 1}) + (transp (λ j → Path (EM G 1) ptS (emloop g (~ j ∧ ~ i))) i + (compPath-filler x (sym (emloop g)) i) ))) + ∙∙ (λ i → transp (λ j → Path (EM G 1) (emloop g (i ∨ j)) ptS) i + (compPath-filler' + (sym (emloop g)) + (cong-∙ (-ₖ_ {n = 1}) + x (sym (emloop g)) i) i)) + ∙∙ (cong (sym (emloop g) ∙_) (isCommΩEM 0 (cong (-ₖ_ {n = 1}) x) (emloop g))) + ∙∙ assoc _ _ _ + ∙∙ cong (_∙ (cong (-ₖ_ {n = 1}) x)) (lCancel (emloop g)) + ∙ sym (lUnit _)) + + main : (x : EM G 1) (p : ptS ≡ x) → decoder x p ≡ sym p + main x = J (λ x p → decoder x p ≡ sym p) refl + + cong-₂ : (n : ℕ) (p : typ (Ω (EM∙ G (2 + n)))) → cong (λ x → -[ 2 + n ]ₖ x) p ≡ sym p + cong-₂ n p = main _ p + where + pp : (a : _) → PathP (λ i → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ → ∣ merid a i ∣ₕ ≡ 0ₖ (2 + n)) + (cong (λ x → -[ 2 + n ]ₖ x)) λ p → cong ∣_∣ₕ (sym (merid ptS)) ∙ cong (λ x → -[ 2 + n ]ₖ x) p + pp a = + toPathP + (funExt λ x → + (λ k → transp (λ i → Path (EM G (2 + n)) ∣ merid a (i ∨ k) ∣ ∣ ptS ∣) k (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) + (cong (-ₖ-syntax (suc (suc n))) + (transp (λ j → Path (EM G (2 + n)) ∣ ptS ∣ ∣ merid a (~ j ∧ ~ k) ∣) k (compPath-filler x (sym (cong ∣_∣ₕ (merid a))) k))) k)) + ∙∙ cong (cong ∣_∣ₕ (sym (merid a)) ∙_) (cong-∙ (λ x → -[ 2 + n ]ₖ x) x (sym (cong ∣_∣ₕ (merid a))) + ∙ isCommΩEM (suc n) (cong (λ x → -[ 2 + n ]ₖ x) x) (cong ∣_∣ₕ (σ-EM n a))) + ∙∙ (λ k → (λ i → ∣ merid a (~ i ∨ k) ∣) ∙ (λ i → ∣ compPath-filler' (merid a) (sym (merid ptS)) (~ k) i ∣) + ∙ cong (λ x → -ₖ-syntax (suc (suc n)) x) x) + ∙ sym (lUnit _)) + + decoder : (x : EM G (2 + n)) → 0ₖ (2 + n) ≡ x → x ≡ 0ₖ (2 + n) + decoder = + trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ { north → pp ptS i0 + ; south → pp ptS i1 + ; (merid a i) → pp a i} + + main : (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → decoder x p ≡ sym p + main x = J (λ x p → decoder x p ≡ sym p) refl + + rCancelₖ : (n : ℕ) (x : EM G n) → x +[ n ]ₖ (-[ n ]ₖ x) ≡ 0ₖ n + rCancelₖ zero x = invr x + rCancelₖ (suc zero) = + elimSet _ (λ _ → emsquash _ _) + refl + λ g → flipSquare (cong₂+₁ (emloop g) (λ i → -ₖ-syntax 1 (emloop g i)) + ∙ rCancel (emloop g)) + rCancelₖ (suc (suc n)) = + trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ { north → refl + ; south i → +ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣ (-ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣) + ; (merid a i) j + → hcomp (λ r → λ { (i = i0) → 0ₖ (2 + n) + ; (i = i1) → ∣ merid ptS (~ j ∧ r) ∣ₕ -[ 2 + n ]ₖ ∣ merid ptS (~ j ∧ r) ∣ + ; (j = i0) → ∣ compPath-filler (merid a) (sym (merid ptS)) (~ r) i ∣ + -[ 2 + n ]ₖ ∣ compPath-filler (merid a) (sym (merid ptS)) (~ r) i ∣ + ; (j = i1) → 0ₖ (2 + n)}) + (help' a j i) } + where + help' : (a : _) → cong₂ (λ x y → ∣ x ∣ -[ suc (suc n) ]ₖ ∣ y ∣) (σ-EM n a) (σ-EM n a) ≡ refl + help' a = + cong₂+₂ n (cong ∣_∣ₕ (σ-EM n a)) (cong (λ x → -[ 2 + n ]ₖ ∣ x ∣) (σ-EM n a)) + ∙∙ cong (cong ∣_∣ₕ (σ-EM n a) ∙_) (cong-₂ n (cong ∣_∣ₕ (σ-EM n a))) + ∙∙ rCancel _ + + lCancelₖ : (n : ℕ) (x : EM G n) → (-[ n ]ₖ x) +[ n ]ₖ x ≡ 0ₖ n + lCancelₖ n x = commₖ n (-[ n ]ₖ x) x ∙ rCancelₖ n x + + assocₖ : (n : ℕ) (x y z : EM G n) + → (x +[ n ]ₖ (y +[ n ]ₖ z) ≡ (x +[ n ]ₖ y) +[ n ]ₖ z) + assocₖ zero = assocG + assocₖ (suc zero) = + elimSet _ (λ _ → isSetΠ2 λ _ _ → emsquash _ _) + (λ _ _ → refl) + λ g i y z k → lem g y z k i + where + lem : (g : fst G) (y z : _) → cong (λ x → x +[ suc zero ]ₖ (y +[ suc zero ]ₖ z)) (emloop g) + ≡ cong (λ x → (x +[ suc zero ]ₖ y) +[ suc zero ]ₖ z) (emloop g) + lem g = + elimProp _ (λ _ → isPropΠ λ _ → emsquash _ _ _ _) + (elimProp _ (λ _ → emsquash _ _ _ _) + refl) + assocₖ (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ a b → trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + (λ c → main c a b) + where + lem : (c : _) (a b : _) + → PathP (λ i → (∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ merid c i ∣ₕ) + ≡ (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ merid c i ∣ₕ)) + (cong (λ x → ∣ a ∣ₕ +[ suc (suc n) ]ₖ x) (rUnitₖ (suc (suc n)) ∣ b ∣) + ∙ sym (rUnitₖ (suc (suc n)) (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ))) + ((λ i → ∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ merid ptS (~ i) ∣ₕ)) + ∙∙ cong (λ x → ∣ a ∣ₕ +[ suc (suc n) ]ₖ x) (rUnitₖ (suc (suc n)) ∣ b ∣) + ∙ sym (rUnitₖ (suc (suc n)) (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ)) + ∙∙ λ i → (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ merid ptS i ∣ₕ) + lem c = + EM-raw-elim G (suc n) + (λ _ → isOfHLevelΠ (2 + n) (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _)) + (EM-raw-elim G (suc n) (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _) + ((sym (rUnit refl) + ◁ λ _ → refl) + ▷ (sym (lCancel (cong ∣_∣ₕ (merid ptS))) + ∙ λ i → (λ j → ∣ merid ptS (~ j ∨ ~ i) ∣ₕ) + ∙∙ lUnit (λ j → ∣ merid ptS (~ j ∧ ~ i) ∣ₕ) i + ∙∙ cong ∣_∣ₕ (merid ptS)))) + main : (c a b : _) → (∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ c ∣ₕ) ≡ (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ c ∣ₕ) + main north a b = lem ptS a b i0 + main south a b = lem ptS a b i1 + main (merid c i) a b = lem c a b i + + addIso : (n : ℕ) (x : EM G n) → Iso (EM G n) (EM G n) + Iso.fun (addIso n x) y = y +[ n ]ₖ x + Iso.inv (addIso n x) y = y -[ n ]ₖ x + Iso.rightInv (addIso n x) y = + sym (assocₖ n y (-[ n ]ₖ x) x) + ∙∙ cong (λ x → y +[ n ]ₖ x) (lCancelₖ n x) + ∙∙ rUnitₖ n y + Iso.leftInv (addIso n x) y = + {!!} + + + CODE : {!!} + CODE = {!!} From 920178b96038035ffdfe3daa09320c480648a83d Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 14 Jul 2021 04:39:26 +0200 Subject: [PATCH 17/30] stuff --- .../HITs/EilenbergMacLane1/Properties.agda | 323 +++++++++++++++++- 1 file changed, 320 insertions(+), 3 deletions(-) diff --git a/Cubical/HITs/EilenbergMacLane1/Properties.agda b/Cubical/HITs/EilenbergMacLane1/Properties.agda index cdd312941..de286334f 100644 --- a/Cubical/HITs/EilenbergMacLane1/Properties.agda +++ b/Cubical/HITs/EilenbergMacLane1/Properties.agda @@ -16,6 +16,7 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence open import Cubical.Data.Sigma open import Cubical.Data.Empty renaming (rec to ⊥-rec) hiding (elim) @@ -26,6 +27,8 @@ open import Cubical.Algebra.Group.Properties open import Cubical.Algebra.AbGroup.Base +open import Cubical.Functions.Morphism + open import Cubical.HITs.PropositionalTruncation as PropTrunc using (∥_∥; ∣_∣; squash) open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂; ∣_∣₂; squash₂) @@ -172,6 +175,24 @@ EM G zero = EM-raw G zero EM G (suc zero) = EM-raw G 1 EM G (suc (suc n)) = hLevelTrunc (4 + n) (EM-raw G (suc (suc n))) +hLevelEM : (G : AbGroup ℓG) (n : ℕ) → isOfHLevel (2 + n) (EM G n) +hLevelEM G zero = AbGroupStr.is-set (snd G) +hLevelEM G (suc zero) = emsquash +hLevelEM G (suc (suc n)) = isOfHLevelTrunc (4 + n) + +EM-raw→EM : (G : AbGroup ℓG) (n : ℕ) → EM-raw G n → EM G n +EM-raw→EM G zero x = x +EM-raw→EM G (suc zero) x = x +EM-raw→EM G (suc (suc n)) = ∣_∣ + +EM-elim : {G : AbGroup ℓG} (n : ℕ) {A : EM G n → Type ℓ} + → ((x : _) → isOfHLevel (2 + n) (A x)) + → ((x : EM-raw G n) → A (EM-raw→EM G n x)) + → (x : _) → A x +EM-elim zero hlev hyp x = hyp x +EM-elim (suc zero) hlev hyp x = hyp x +EM-elim (suc (suc n)) hlev hyp = trElim (λ _ → hlev _) hyp + wedgeConFun' : (G H : AbGroup ℓG) (n : ℕ) → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ} → ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y)) → (f : (x : _) → A ptS x) @@ -699,6 +720,93 @@ module _ {G : AbGroup ℓG} where main south a b = lem ptS a b i1 main (merid c i) a b = lem c a b i + σ-EM' : (n : ℕ) (x : EM G (suc n)) + → Path (EM G (suc (suc n))) + (0ₖ (suc (suc n))) + (0ₖ (suc (suc n))) + σ-EM' zero x = cong ∣_∣ₕ (σ-EM zero x) + σ-EM' (suc n) = + trElim (λ _ → isOfHLevelTrunc (5 + n) _ _) + λ x → cong ∣_∣ₕ (σ-EM (suc n) x) + + σ-EM'-0ₖ : (n : ℕ) → σ-EM' n (0ₖ (suc n)) ≡ refl + σ-EM'-0ₖ zero = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + σ-EM'-0ₖ (suc n) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + + private + P : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (r : refl ≡ p) + → lUnit p ∙ cong (_∙ p) r + ≡ rUnit p ∙ cong (p ∙_) r + P p = J (λ p r → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r) refl + + σ-EM'-hom : (n : ℕ) → (a b : _) → σ-EM' n (a +ₖ b) ≡ σ-EM' n a ∙ σ-EM' n b + σ-EM'-hom zero = + wedgeConEM.fun G G 0 0 (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) l r p + where + l : _ + l x = cong (σ-EM' zero) (lUnitₖ 1 x) + ∙∙ lUnit (σ-EM' zero x) + ∙∙ cong (_∙ σ-EM' zero x) (sym (σ-EM'-0ₖ zero)) + + r : _ + r x = + cong (σ-EM' zero) (rUnitₖ 1 x) + ∙∙ rUnit (σ-EM' zero x) + ∙∙ cong (σ-EM' zero x ∙_) (sym (σ-EM'-0ₖ zero)) + + p : _ + p = P (σ-EM' zero embase) (sym (σ-EM'-0ₖ zero)) + σ-EM'-hom (suc n) = + elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) _ _) _ _) + (wedgeConEM.fun G G _ _ + (λ x y → transport (λ i → isOfHLevel (help n i) + ((σ-EM' (suc n) (∣ x ∣ₕ +ₖ ∣ y ∣ₕ) + ≡ σ-EM' (suc n) ∣ x ∣ₕ ∙ σ-EM' (suc n) ∣ y ∣ₕ))) + (isOfHLevelPlus {n = 4 + n} n + (isOfHLevelPath (4 + n) + (isOfHLevelTrunc (5 + n) _ _) _ _))) + (λ x → cong (σ-EM' (suc n)) (lUnitₖ (suc (suc n)) ∣ x ∣) + ∙∙ lUnit (σ-EM' (suc n) ∣ x ∣) + ∙∙ cong (_∙ σ-EM' (suc n) ∣ x ∣) (sym (σ-EM'-0ₖ (suc n)))) + (λ x → cong (σ-EM' (suc n)) (rUnitₖ (2 + n) ∣ x ∣) + ∙∙ rUnit (σ-EM' (suc n) ∣ x ∣) + ∙∙ cong (σ-EM' (suc n) ∣ x ∣ ∙_) (sym (σ-EM'-0ₖ (suc n)))) + (P (σ-EM' (suc n) (0ₖ (2 + n))) (sym (σ-EM'-0ₖ (suc n))))) + + σ-EM'-ₖ : (n : ℕ) → (a : _) → σ-EM' n (-ₖ a) ≡ sym (σ-EM' n a) + σ-EM'-ₖ n = + morphLemmas.distrMinus + (λ x y → x +[ suc n ]ₖ y) (_∙_) + (σ-EM' n) (σ-EM'-hom n) + (0ₖ (suc n)) refl + (λ x → -ₖ x) sym + (λ x → sym (lUnit x)) (λ x → sym (rUnit x)) + (lCancelₖ (suc n)) rCancel + assoc (σ-EM'-0ₖ n) + + -Dist : (n : ℕ) (x y : EM G n) → -[ n ]ₖ (x +[ n ]ₖ y) ≡ (-[ n ]ₖ x) +[ n ]ₖ (-[ n ]ₖ y) + -Dist zero x y = (GroupTheory.invDistr (AbGroup→Group G) x y) ∙ commₖ zero _ _ + -Dist (suc zero) = k + where -- useless where clause. Needed for fast type checking for some reason. + l : _ + l x = refl + + r : _ + r x = cong (λ z → -[ 1 ]ₖ z) (rUnitₖ 1 x) ∙ sym (rUnitₖ 1 (-[ 1 ]ₖ x)) + + p : r ptS ≡ l ptS + p = sym (rUnit refl) + + k = wedgeConEM.fun G G 0 0 (λ _ _ → emsquash _ _) l r (sym p) + + -Dist (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelTruncPath {n = 4 + n}) + (wedgeConEM.fun G G (suc n) (suc n) + (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (hLevHelp n) _ _) + (λ x → refl) + (λ x → cong (λ z → -[ (suc (suc n)) ]ₖ z) (rUnitₖ (suc (suc n)) ∣ x ∣ₕ) ∙ sym (rUnitₖ (suc (suc n)) (-[ (suc (suc n)) ]ₖ ∣ x ∣ₕ))) + (rUnit refl)) + addIso : (n : ℕ) (x : EM G n) → Iso (EM G n) (EM G n) Iso.fun (addIso n x) y = y +[ n ]ₖ x Iso.inv (addIso n x) y = y -[ n ]ₖ x @@ -707,8 +815,217 @@ module _ {G : AbGroup ℓG} where ∙∙ cong (λ x → y +[ n ]ₖ x) (lCancelₖ n x) ∙∙ rUnitₖ n y Iso.leftInv (addIso n x) y = - {!!} + sym (assocₖ n y x (-[ n ]ₖ x)) + ∙∙ cong (λ x → y +[ n ]ₖ x) (rCancelₖ n x) + ∙∙ rUnitₖ n y + CODE : (n : ℕ) → EM G (suc n) → TypeOfHLevel ℓG (2 + n) + CODE zero = + rec _ (isOfHLevelTypeOfHLevel 2) + (typ G , is-set) + (λ g → Σ≡Prop (λ _ → isPropIsOfHLevel 2) + (isoToPath (addIso 0 g))) + (λ g h → ΣSquareSet {B = isOfHLevel 2} + (λ _ → isSetΠ2 λ _ _ → isProp→isOfHLevelSuc 1 isPropIsProp) + (without g h)) + where + ΣSquareSet : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} (prop : (x : A) → isSet (B x)) + → {x y z w : Σ A B} {p : x ≡ y} {q : y ≡ z} {r : w ≡ z} {s : x ≡ w} + → Square (cong fst s) (cong fst q) (cong fst p) (cong fst r) + → Square s q p r + fst (ΣSquareSet prop sq i j) = sq i j + snd (ΣSquareSet {B = B} prop {p = p} {q = q} {r = r} {s = s} sq i j) = sqP i j + where + sqP : SquareP (λ i j → B (sq i j)) + (cong snd s) (cong snd q) (cong snd p) (cong snd r) + sqP = toPathP (isOfHLevelPathP' 1 (prop _) _ _ _ _) + + lem : (g h : fst G) → compEquiv (isoToEquiv (addIso 0 g)) (isoToEquiv (addIso 0 h)) + ≡ isoToEquiv (addIso 0 (snd G .AbGroupStr._+_ g h)) + lem g h = + Σ≡Prop (λ _ → isPropIsEquiv _) + (funExt λ x → sym (assocₖ 0 x g h)) + + without : (g h : fst G) → + Square (isoToPath (addIso 0 g)) + (isoToPath (addIso 0 ((snd (AbGroup→Group G) GroupStr.· g) h))) + refl (isoToPath (addIso 0 h)) + without g h = + compPathL→PathP + (sym (lUnit _) + ∙∙ sym (uaCompEquiv (isoToEquiv (addIso 0 g)) (isoToEquiv (addIso 0 h))) + ∙∙ cong ua (lem g h)) + + CODE (suc n) = + trElim (λ _ → isOfHLevelTypeOfHLevel (3 + n)) + λ { north → EM G (suc n) , hLevelEM G (suc n) + ; south → EM G (suc n) , hLevelEM G (suc n) + ; (merid a i) → fib n a i} + where + fib : (n : ℕ) → (a : EM-raw G (suc n)) + → Path (TypeOfHLevel _ (3 + n)) (EM G (suc n) , hLevelEM G (suc n)) (EM G (suc n) , hLevelEM G (suc n)) + fib zero a = Σ≡Prop (λ _ → isPropIsOfHLevel 3) + (isoToPath (addIso 1 a)) + fib (suc n) a = Σ≡Prop (λ _ → isPropIsOfHLevel (4 + n)) + (isoToPath (addIso (suc (suc n)) ∣ a ∣)) + + decode : (n : ℕ) (x : EM G (suc n)) → CODE n x .fst → 0ₖ (suc n) ≡ x + decode zero = + elimSet _ (λ _ → isSetΠ (λ _ → emsquash _ _)) + emloop main + where + main : (g : _) → PathP (λ i → CODE zero (emloop g i) .fst → 0ₖ 1 ≡ emloop g i) + emloop emloop + main g = + toPathP + (funExt λ x + → (λ j → transp (λ i → Path (EM-raw G 1) ptS (emloop g (i ∨ j))) j + (compPath-filler (emloop ((transportRefl x j) +G -G g)) (emloop g) j)) + ∙∙ sym (emloop-comp _ (x +G -G g) g) + ∙∙ cong emloop (sym (assocG x (-G g) g) ∙∙ cong (x +G_) (invl g) ∙∙ rid x)) + decode (suc n) = + trElim (λ _ → isOfHLevelΠ (4 + n) + λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + λ { north → f n + ; south → g n + ; (merid a i) → main a i} + where + f : (n : ℕ) → _ + f n = σ-EM' n + + g : (n : ℕ) → EM G (suc n) → ∣ ptS ∣ ≡ ∣ south ∣ + g n x = σ-EM' n x ∙ cong ∣_∣ₕ (merid ptS) + + lem₂ : (n : ℕ) (a x : _) + → Path (Path (EM G (suc (suc n))) _ _) + ((λ i → cong ∣_∣ₕ (σ-EM n x) i) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) ∙ (λ i → ∣ merid a i ∣ₕ)) + (g n (EM-raw→EM G (suc n) x)) + lem₂ zero a x = + cong (f zero x ∙_) + (cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid ptS))) + ∙ cong-∙ ∣_∣ₕ (merid (ptS)) (sym (merid a))) + ∙∙ sym (assoc _ _ _) + ∙∙ cong (cong ∣_∣ₕ (merid ptS) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) + ∙ sym (rUnit _)) + lem₂ (suc n) a x = + cong (f (suc n) ∣ x ∣ ∙_) + ((cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid ptS))) + ∙ cong-∙ ∣_∣ₕ (merid (ptS)) (sym (merid a))) + ∙∙ sym (assoc _ _ _) + ∙∙ cong (cong ∣_∣ₕ (merid ptS) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) + ∙ sym (rUnit _))) + + lem : (n : ℕ) (x a : EM-raw G (suc n)) + → f n (transport (sym (cong (λ x → CODE (suc n) x .fst) (cong ∣_∣ₕ (merid a)))) (EM-raw→EM G (suc n) x)) + ≡ cong ∣_∣ₕ (σ-EM n x) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) + lem zero x a = (λ i → cong ∣_∣ₕ (merid (transportRefl x i -[ 1 ]ₖ a) ∙ sym (merid embase))) + ∙∙ σ-EM'-hom zero x (-ₖ a) + ∙∙ cong (f zero x ∙_) (σ-EM'-ₖ zero a) + lem (suc n) x a = + cong (f (suc n)) (λ i → transportRefl ∣ x ∣ i -[ 2 + n ]ₖ ∣ a ∣) + ∙∙ σ-EM'-hom (suc n) ∣ x ∣ (-ₖ ∣ a ∣) + ∙∙ cong (f (suc n) ∣ x ∣ ∙_) (σ-EM'-ₖ (suc n) ∣ a ∣) + + main : (a : _) + → PathP (λ i → CODE (suc n) ∣ merid a i ∣ₕ .fst + → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ) (f n) (g n) + main a = + toPathP (funExt + (EM-elim _ (λ _ → isOfHLevelPathP (2 + (suc n)) (isOfHLevelTrunc (4 + n) _ _) _ _) + λ x → + (λ i → transp (λ j → Path (EM G (2 + n)) ∣ ptS ∣ ∣ merid a (i ∨ j) ∣) + i (compPath-filler (lem n x a i) (cong ∣_∣ₕ (merid a)) i) ) + ∙∙ sym (assoc _ _ _) + ∙∙ lem₂ n a x)) - CODE : {!!} - CODE = {!!} + private + ptCode0 : (n : ℕ) → CODE n (0ₖ (suc n)) .fst + ptCode0 zero = 0g + ptCode0 (suc n) = 0ₖ (suc n) + + encode : (n : ℕ) (x : EM G (suc n)) → 0ₖ (suc n) ≡ x → CODE n x .fst + encode n x p = subst (λ x → CODE n x .fst) p (ptCode0 n) + + open import Cubical.Foundations.Transport + + encode-morph : (n : ℕ) → (p q : 0ₖ (suc (suc n)) ≡ 0ₖ (suc (suc n))) + → (encode (suc n) _ (p ∙ q)) ≡ (encode (suc n) _ p) +[ suc n ]ₖ (encode (suc n) _ q) + encode-morph n p q = + substComposite (λ x → CODE (suc n) x .fst) p q (0ₖ (suc n)) + ∙∙ {!!} + ∙∙ {!!} + + decode-encode : (n : ℕ) (x : EM G (suc n)) (p : 0ₖ (suc n) ≡ x) → decode n x (encode n x p) ≡ p + decode-encode zero x = + J (λ x p → decode zero x (encode zero x p) ≡ p) + (cong emloop (transportRefl 0g) ∙ emloop-1g (AbGroup→Group G)) + decode-encode (suc zero) x = + J (λ x p → decode 1 x (encode 1 x p) ≡ p) + (σ-EM'-0ₖ 0) + decode-encode (suc (suc n)) x = + J (λ x p → decode (2 + n) x (encode (2 + n) x p) ≡ p) + (σ-EM'-0ₖ (suc n)) + + encode-decode : (n : ℕ) (x : _) → encode n (0ₖ (suc n)) (decode n (0ₖ (suc n)) x) ≡ x + encode-decode zero x i = transportRefl (lid x i) i + encode-decode (suc zero) x = + cong (encode 1 (0ₖ 2)) (cong-∙ ∣_∣ₕ (merid x) (sym (merid ptS))) + ∙∙ substComposite (λ x → CODE (suc zero) x .fst) + (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptS))) ptS + ∙∙ cong (subst (λ x₁ → CODE 1 x₁ .fst) (λ i → ∣ merid ptS (~ i) ∣ₕ)) + (λ i → lUnitₖ 1 (transportRefl x i) i) + ∙ (λ i → rUnitₖ 1 (transportRefl x i) i) + encode-decode (suc (suc n)) = + trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ x → cong (encode (2 + n) (0ₖ (3 + n))) (cong-∙ ∣_∣ₕ (merid x) (sym (merid ptS))) + ∙∙ substComposite (λ x → CODE (suc (suc n)) x .fst) + (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptS))) (0ₖ (2 + n)) + ∙∙ cong (subst (λ x₁ → CODE (2 + n) x₁ .fst) (λ i → ∣ merid ptS (~ i) ∣ₕ)) + (λ i → lUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) + ∙ (λ i → rUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) + + EM∙' : (n : ℕ) → Pointed _ + EM∙' n = EM G n , 0ₖ n + + Iso-EM-ΩEM+1 : (n : ℕ) → Iso (EM G n) (typ (Ω (EM∙' (suc n)))) + Iso.fun (Iso-EM-ΩEM+1 zero) = decode zero (0ₖ 1) + Iso.inv (Iso-EM-ΩEM+1 zero) = encode zero (0ₖ 1) + Iso.rightInv (Iso-EM-ΩEM+1 zero) = decode-encode zero (0ₖ 1) + Iso.leftInv (Iso-EM-ΩEM+1 zero) = encode-decode zero + Iso.fun (Iso-EM-ΩEM+1 (suc zero)) = decode 1 (0ₖ 2) + Iso.inv (Iso-EM-ΩEM+1 (suc zero)) = encode 1 (0ₖ 2) + Iso.rightInv (Iso-EM-ΩEM+1 (suc zero)) = decode-encode 1 (0ₖ 2) + Iso.leftInv (Iso-EM-ΩEM+1 (suc zero)) = encode-decode 1 + Iso.fun (Iso-EM-ΩEM+1 (suc (suc n))) = decode (2 + n) (0ₖ (3 + n)) + Iso.inv (Iso-EM-ΩEM+1 (suc (suc n))) = encode (2 + n) (0ₖ (3 + n)) + Iso.rightInv (Iso-EM-ΩEM+1 (suc (suc n))) = decode-encode (2 + n) (0ₖ (3 + n)) + Iso.leftInv (Iso-EM-ΩEM+1 (suc (suc n))) = encode-decode (2 + n) + + EM→ΩEM+1 : (n : ℕ) → EM G n → typ (Ω (EM∙' (suc n))) + EM→ΩEM+1 n = Iso.fun (Iso-EM-ΩEM+1 n) + + ΩEM+1→EM : (n : ℕ) → typ (Ω (EM∙' (suc n))) → EM G n + ΩEM+1→EM n = Iso.inv (Iso-EM-ΩEM+1 n) + + EM→ΩEM+1-hom : (n : ℕ) (x y : EM G n) + → EM→ΩEM+1 n (x +[ n ]ₖ y) ≡ EM→ΩEM+1 n x ∙ EM→ΩEM+1 n y + EM→ΩEM+1-hom zero x y = emloop-comp _ x y + EM→ΩEM+1-hom (suc zero) x y = σ-EM'-hom zero x y + EM→ΩEM+1-hom (suc (suc n)) x y = σ-EM'-hom (suc n) x y + + ΩEM+1→EM-hom : (n : ℕ) → (p q : typ (Ω (EM∙' (suc n)))) + → ΩEM+1→EM n (p ∙ q) ≡ (ΩEM+1→EM n p) +[ n ]ₖ (ΩEM+1→EM n q) + ΩEM+1→EM-hom n = + morphLemmas.isMorphInv (λ x y → x +[ n ]ₖ y) (_∙_) (EM→ΩEM+1 n) + (EM→ΩEM+1-hom n) (ΩEM+1→EM n) + (Iso.rightInv (Iso-EM-ΩEM+1 n)) (Iso.leftInv (Iso-EM-ΩEM+1 n)) + + EM→ΩEM+1-0ₖ : (n : ℕ) → EM→ΩEM+1 n (0ₖ n) ≡ refl + EM→ΩEM+1-0ₖ zero = emloop-1g _ + EM→ΩEM+1-0ₖ (suc zero) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + EM→ΩEM+1-0ₖ (suc (suc n)) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + + ΩEM+1→EM-refl : (n : ℕ) → ΩEM+1→EM n refl ≡ 0ₖ n + ΩEM+1→EM-refl zero = transportRefl 0g + ΩEM+1→EM-refl (suc zero) = refl + ΩEM+1→EM-refl (suc (suc n)) = refl From 7403d267f99b580d8fcef5618904478f9cfacffb Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sat, 17 Jul 2021 14:27:19 +0200 Subject: [PATCH 18/30] stuff --- Cubical/Algebra/AbGroup/TensorProduct.agda | 854 +++++++++++++++++ .../Algebra/Group/EilenbergMacLane/Base.agda | 90 ++ .../EilenbergMacLane/GroupStructure.agda | 405 ++++++++ .../Group/EilenbergMacLane/Properties.agda | 454 +++++++++ .../EilenbergMacLane/WedgeConnectivity.agda | 287 ++++++ Cubical/Algebra/Group/EilenbergMacLane1.agda | 844 ++++++++++++++++ Cubical/Foundations/HLevels.agda | 16 + .../HITs/EilenbergMacLane1/Properties.agda | 899 ------------------ 8 files changed, 2950 insertions(+), 899 deletions(-) create mode 100644 Cubical/Algebra/AbGroup/TensorProduct.agda create mode 100644 Cubical/Algebra/Group/EilenbergMacLane/Base.agda create mode 100644 Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda create mode 100644 Cubical/Algebra/Group/EilenbergMacLane/Properties.agda create mode 100644 Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda diff --git a/Cubical/Algebra/AbGroup/TensorProduct.agda b/Cubical/Algebra/AbGroup/TensorProduct.agda new file mode 100644 index 000000000..74be6bf8f --- /dev/null +++ b/Cubical/Algebra/AbGroup/TensorProduct.agda @@ -0,0 +1,854 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.AbGroup.TensorProduct where + +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Function hiding (flip) +open import Cubical.Foundations.Isomorphism + +open import Cubical.Relation.Binary + +open import Cubical.Data.Sigma + +open import Cubical.Data.List +open import Cubical.Data.Nat as ℕ +open import Cubical.Foundations.HLevels +open import Cubical.Data.Int +open import Cubical.Data.Sum hiding (map) + +open import Cubical.HITs.PropositionalTruncation renaming (map to pMap ; rec to pRec ; elim to pElim ; elim2 to pElim2) + +open import Cubical.Algebra.Group hiding (ℤ) + +-- open import Cubical.HITs.SetQuotients + +private + variable + ℓ ℓ' : Level + + +module _ (AGr : AbGroup ℓ) (BGr : AbGroup ℓ') where + private + open AbGroupStr renaming (_+_ to _+G_ ; -_ to -G_) + + strA = snd AGr + strB = snd BGr + + A = fst AGr + B = fst BGr + + 0A = 0g strA + 0B = 0g strB + + _+A_ = _+G_ strA + _+B_ = _+G_ strB + + -A_ = -G_ strA + -B_ = -G_ strB + + data _⨂₁_ : Type (ℓ-max ℓ ℓ') where + _⊗_ : (a : A) → (b : B) → _⨂₁_ + _+⊗_ : (w : _⨂₁_) → (u : _⨂₁_) → _⨂₁_ + + ⊗comm : (x y : _) → x +⊗ y ≡ y +⊗ x + ⊗assoc : (x y z : _) → x +⊗ (y +⊗ z) ≡ (x +⊗ y) +⊗ z + ⊗lUnit : (x : _) → (0A ⊗ 0B) +⊗ x ≡ x + + linl : (x y : A) (z : B) → (x +A y) ⊗ z ≡ ((x ⊗ z) +⊗ (y ⊗ z)) + linr : (x : A) (y z : B) → x ⊗ (y +B z) ≡ ((x ⊗ y) +⊗ (x ⊗ z)) + + flip : (x : A) (b : B) → x ⊗ (-B b) ≡ ((-A x) ⊗ b) -- needed ? + + ⊗squash : isSet _⨂₁_ + + +move4 : ∀ {ℓ} {A : Type ℓ} (x y z w : A) (_+_ : A → A → A) + → ((x y z : A) → x + (y + z) ≡ (x + y) + z) + → ((x y : A) → x + y ≡ y + x) + → (x + y) + (z + w) ≡ ((x + z) + (y + w)) +move4 x y z w _+_ assoc comm = + sym (assoc x y (z + w)) + ∙∙ cong (x +_) (assoc y z w ∙∙ cong (_+ w) (comm y z) ∙∙ sym (assoc z y w)) + ∙∙ assoc x z (y + w) + +module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where + private + open AbGroupStr renaming (_+_ to _+G_ ; -_ to -G_) + + strA = snd AGr + strB = snd BGr + + A = fst AGr + B = fst BGr + + 0A = 0g strA + 0B = 0g strB + + _+A_ = _+G_ strA + _+B_ = _+G_ strB + + -A_ = -G_ strA + -B_ = -G_ strB + + A⊗B = AGr ⨂₁ BGr + + 0⊗ : AGr ⨂₁ BGr + 0⊗ = 0A ⊗ 0B + + ⊗elimProp : ∀ {ℓ} {C : AGr ⨂₁ BGr → Type ℓ} + → ((x : _) → isProp (C x)) + → (f : (x : A) (b : B) → C (x ⊗ b)) + → (g : ((x y : _) → C x → C y → C (x +⊗ y))) + → (x : _) → C x + ⊗elimProp p f g (a ⊗ b) = f a b + ⊗elimProp p f g (x +⊗ y) = g x y (⊗elimProp p f g x) (⊗elimProp p f g y) + ⊗elimProp {C = C} p f g (⊗comm x y j) = + isOfHLevel→isOfHLevelDep 1 {B = C} p + (g x y (⊗elimProp p f g x) (⊗elimProp p f g y)) + (g y x (⊗elimProp p f g y) (⊗elimProp p f g x)) (⊗comm x y) j + ⊗elimProp {C = C} p f g (⊗assoc x y z j) = + isOfHLevel→isOfHLevelDep 1 {B = C} p + (g x (y +⊗ z) (⊗elimProp p f g x) + (g y z (⊗elimProp p f g y) (⊗elimProp p f g z))) + (g (x +⊗ y) z (g x y (⊗elimProp p f g x) + (⊗elimProp p f g y)) (⊗elimProp p f g z)) (⊗assoc x y z) j + ⊗elimProp {C = C} p f g (⊗lUnit x i) = + isOfHLevel→isOfHLevelDep 1 {B = C} p + (g (0A ⊗ 0B) x (f 0A 0B) (⊗elimProp p f g x)) + (⊗elimProp p f g x) + (⊗lUnit x) i + ⊗elimProp {C = C} p f g (linl x y z i) = + isOfHLevel→isOfHLevelDep 1 {B = C} p + (f (x +A y) z) (g (x ⊗ z) (y ⊗ z) (f x z) (f y z)) (linl x y z) i + ⊗elimProp {C = C} p f g (linr x y z i) = + isOfHLevel→isOfHLevelDep 1 {B = C} p + (f x (y +B z)) (g (x ⊗ y) (x ⊗ z) (f x y) (f x z)) (linr x y z) i + ⊗elimProp {C = C} p f g (flip x b i) = + isOfHLevel→isOfHLevelDep 1 {B = C} p + (f x (-B b)) (f (-A x) b) (flip x b) i + ⊗elimProp {C = C} p f g (⊗squash x y q r i j) = + isOfHLevel→isOfHLevelDep 2 {B = C} (λ x → isProp→isSet (p x)) + _ _ (λ j → ⊗elimProp p f g (q j)) (λ j → ⊗elimProp p f g (r j)) (⊗squash x y q r) i j + + -⊗ : AGr ⨂₁ BGr → AGr ⨂₁ BGr + -⊗ (a ⊗ b) = (-A a) ⊗ b + -⊗ (x +⊗ x₁) = -⊗ x +⊗ -⊗ x₁ + -⊗ (⊗comm x y i) = ⊗comm (-⊗ x) (-⊗ y) i + -⊗ (⊗assoc x y z i) = ⊗assoc (-⊗ x) (-⊗ y) (-⊗ z) i + -⊗ (⊗lUnit x i) = + ((λ i → (GroupTheory.inv1g (AbGroup→Group AGr) i ⊗ 0B) +⊗ -⊗ x) + ∙ ⊗lUnit (-⊗ x)) i + -⊗ (linl x y z i) = + ((λ i → (GroupTheory.invDistr (AbGroup→Group AGr) x y + ∙ comm strA (-A y) (-A x)) i ⊗ z) + ∙ linl (-A x) (-A y) z) i + -⊗ (linr x y z i) = linr (-A x) y z i + -⊗ (flip x b i) = flip (-A x) b i + -⊗ (⊗squash x y p q i j) = + ⊗squash _ _ (λ j → -⊗ (p j)) (λ j → -⊗ (q j)) i j + + ⊗rUnit : (x : A⊗B) → x +⊗ 0⊗ ≡ x + ⊗rUnit x = ⊗comm x 0⊗ ∙ ⊗lUnit x + + listify : List (A × B) → AGr ⨂₁ BGr + listify [] = 0A ⊗ 0B + listify (x ∷ x₁) = (fst x ⊗ snd x) +⊗ listify x₁ + + listify++ : (x y : List (A × B)) → listify (x ++ y) ≡ (listify x +⊗ listify y) + listify++ [] y = sym (⊗lUnit (listify y)) + listify++ (x ∷ x₁) y = + (λ i → (fst x ⊗ snd x) +⊗ (listify++ x₁ y i)) + ∙ ⊗assoc (fst x ⊗ snd x) (listify x₁) (listify y) + + slick : (x : AGr ⨂₁ BGr) → ∃[ l ∈ List (A × B) ] listify l ≡ x + slick = + ⊗elimProp (λ _ → squash) + (λ a b → ∣ [ a , b ] , ⊗rUnit (a ⊗ b) ∣) + λ x y → rec2 squash λ {(l1 , p) (l2 , q) → ∣ (l1 ++ l2) , listify++ l1 l2 ∙ cong₂ _+⊗_ p q ∣} + + + ⊗elimPropCool : ∀ {ℓ} {C : AGr ⨂₁ BGr → Type ℓ} + → ((x : _) → isProp (C x)) + → ((x : _) → C (listify x)) + → (x : _) → C x + ⊗elimPropCool {C = C} p f = + ⊗elimProp p (λ x y → subst C (⊗rUnit (x ⊗ y)) (f [ x , y ])) + λ x y → pRec (isPropΠ2 λ _ _ → p _) + (pRec (isPropΠ3 λ _ _ _ → p _) + (λ {(l1 , p) (l2 , q) ind1 ind2 → subst C (listify++ l2 l1 ∙ cong₂ _+⊗_ q p) (f (l2 ++ l1))}) + (slick y)) + (slick x) + + rCancelPrim : (x : B) → (0A ⊗ x) ≡ 0⊗ + rCancelPrim x = + (λ i → rid strA 0A (~ i) ⊗ x) + ∙∙ linl 0A 0A x + ∙∙ cong ((0A ⊗ x) +⊗_) (cong (_⊗ x) (sym (GroupTheory.inv1g (AbGroup→Group AGr))) + ∙ sym (flip 0A x)) + ∙∙ sym (linr 0A x (-B x)) + ∙∙ (λ i → 0A ⊗ (invr strB x i)) + + ⊗rCancel : (x : AGr ⨂₁ BGr) → (x +⊗ -⊗ x) ≡ 0⊗ + ⊗rCancel = + ⊗elimPropCool (λ _ → ⊗squash _ _) h + where + h : (x : List (A × B)) → (listify x +⊗ -⊗ (listify x)) ≡ 0⊗ + h [] = sym (linl 0A (-A 0A) (0B)) + ∙ cong (λ x → _⊗_ {AGr = AGr} {BGr = BGr} x 0B) (invr strA 0A) + h (x ∷ x₁) = + move4 (fst x ⊗ snd x) (listify x₁) ((-A fst x) ⊗ snd x) (-⊗ (listify x₁)) + _+⊗_ ⊗assoc ⊗comm + ∙∙ cong₂ _+⊗_ (sym (linl (fst x) (-A (fst x)) (snd x)) + ∙∙ (λ i → invr strA (fst x) i ⊗ (snd x)) + ∙∙ rCancelPrim (snd x)) + (h x₁) + ∙∙ ⊗rUnit 0⊗ + + ⊗lCancel : (x : AGr ⨂₁ BGr) → (-⊗ x +⊗ x) ≡ 0⊗ + ⊗lCancel x = ⊗comm _ _ ∙ ⊗rCancel x + +module _ where + open import Cubical.Algebra.Monoid + open import Cubical.Algebra.Semigroup + open AbGroupStr + open IsAbGroup + open IsGroup + open IsMonoid + open IsSemigroup + _⨂_ : AbGroup ℓ → AbGroup ℓ' → AbGroup (ℓ-max ℓ ℓ') + fst (A ⨂ B) = A ⨂₁ B + 0g (snd (A ⨂ B)) = 0⊗ + AbGroupStr._+_ (snd (A ⨂ B)) = _+⊗_ + AbGroupStr.- snd (A ⨂ B) = -⊗ + is-set (isSemigroup (isMonoid (isGroup (isAbGroup (snd (A ⨂ B)))))) = ⊗squash + assoc (isSemigroup (isMonoid (isGroup (isAbGroup (snd (A ⨂ B)))))) = ⊗assoc + fst (identity (isMonoid (isGroup (isAbGroup (snd (A ⨂ B))))) x) = ⊗rUnit x + snd (identity (isMonoid (isGroup (isAbGroup (snd (A ⨂ B))))) x) = ⊗lUnit x + fst (inverse (isGroup (isAbGroup (snd (A ⨂ B)))) x) = ⊗rCancel x + snd (inverse (isGroup (isAbGroup (snd (A ⨂ B)))) x) = ⊗lCancel x + comm (isAbGroup (snd (A ⨂ B))) = ⊗comm + +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties + +_* : AbGroup ℓ → Group ℓ +_* = AbGroup→Group + +module _ (AGr : Group ℓ) (BGr : AbGroup ℓ') where + private + strA = snd AGr + strB = snd BGr + + A = fst AGr + B = fst BGr + open IsGroupHom + + open AbGroupStr strB renaming (_+_ to _+B_ ; -_ to -B_ ; 0g to 0B ; rid to ridB ; lid to lidB ; assoc to assocB ; comm to commB ; invr to invrB ; invl to invlB) + open GroupStr strA renaming (_·_ to _∙A_ ; inv to -A_ ; 1g to 1A ; rid to ridA) + + trivGroupHom : GroupHom AGr (BGr *) + fst trivGroupHom x = 0B + snd trivGroupHom = makeIsGroupHom λ _ _ → sym (ridB 0B) + + compHom : GroupHom AGr (BGr *) → GroupHom AGr (BGr *) → GroupHom AGr (BGr *) + fst (compHom f g) x = fst f x +B fst g x + snd (compHom f g) = + makeIsGroupHom λ x y + → cong₂ _+B_ (pres· (snd f) x y) (pres· (snd g) x y) + ∙ move4 (fst f x) (fst f y) (fst g x) (fst g y) + _+B_ assocB commB + + invHom : GroupHom AGr (BGr *) → GroupHom AGr (BGr *) + fst (invHom (f , p)) x = -B f x + snd (invHom (f , p)) = + makeIsGroupHom + λ x y → cong -B_ (pres· p x y) + ∙∙ GroupTheory.invDistr (BGr *) (f x) (f y) + ∙∙ commB _ _ + + open import Cubical.Algebra.Monoid + open import Cubical.Algebra.Semigroup + open AbGroupStr + open IsAbGroup + open IsGroup + open IsMonoid + open IsSemigroup + + HomGroup : AbGroup (ℓ-max ℓ ℓ') + fst HomGroup = GroupHom AGr (BGr *) + 0g (snd HomGroup) = trivGroupHom + AbGroupStr._+_ (snd HomGroup) = compHom + AbGroupStr.- snd HomGroup = invHom + is-set (isSemigroup (isMonoid (isGroup (isAbGroup (snd HomGroup))))) = + isSetGroupHom + assoc (isSemigroup (isMonoid (isGroup (isAbGroup (snd HomGroup))))) (f , p) (g , q) (h , r) = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ x → assocB _ _ _) + fst (identity (isMonoid (isGroup (isAbGroup (snd HomGroup)))) (f , p)) = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ y → ridB _) + snd (identity (isMonoid (isGroup (isAbGroup (snd HomGroup)))) (f , p)) = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ x → lidB _) + fst (inverse (isGroup (isAbGroup (snd HomGroup))) (f , p)) = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ x → invrB (f x)) + snd (inverse (isGroup (isAbGroup (snd HomGroup))) (f , p)) = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ x → invlB (f x)) + comm (isAbGroup (snd HomGroup)) (f , p) (g , q) = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ x → commB _ _) + +tensorFun : (A : Group ℓ) (B : Group ℓ') (T C : AbGroup (ℓ-max ℓ ℓ')) + (f : GroupHom A (HomGroup B T *)) + → GroupHom (T *) (C *) + → GroupHom A (HomGroup B C *) +fst (fst (tensorFun A B T C (f , p) (g , q)) a) b = g (fst (f a) b) +snd (fst (tensorFun A B T C (f , p) (g , q)) a) = + makeIsGroupHom λ x y + → cong g (IsGroupHom.pres· (snd (f a)) x y) + ∙ IsGroupHom.pres· q _ _ +snd (tensorFun A B T C (f , p) (g , q)) = + makeIsGroupHom λ x y + → Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ b + → cong g (funExt⁻ (cong fst (IsGroupHom.pres· p x y)) b) + ∙ IsGroupHom.pres· q _ _) +{- + makeIsGroupHom λ x y + → cong g (IsGroupHom.pres· (snd (f a)) x y) + ∙ IsGroupHom.pres· q _ _ +snd (tensorFun A B T C (f , p) (g , q)) = + makeIsGroupHom λ x y + → cong g (IsGroupHom.pres· (snd (f a)) x y) + ∙ IsGroupHom.pres· q _ _ +snd (h (f , p) C (g , q)) = + makeIsGroupHom λ x y + → Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ b + → cong g (funExt⁻ (cong fst (IsGroupHom.pres· p x y)) b) + ∙ IsGroupHom.pres· q _ _) -} + +isTensorProductOf_and_ : AbGroup ℓ → AbGroup ℓ' → AbGroup (ℓ-max ℓ ℓ')→ Type _ +isTensorProductOf_and_ {ℓ} {ℓ'} A B T = + Σ[ f ∈ GroupHom (A *) ((HomGroup (B *) T) *) ] + ((C : AbGroup (ℓ-max ℓ ℓ')) + → isEquiv {A = GroupHom (T *) (C *)} + {B = GroupHom (A *) ((HomGroup (B *) C) *)} + (tensorFun (A *) (B *) T C f)) + +module UP (AGr : AbGroup ℓ) (BGr : AbGroup ℓ') where + private + open AbGroupStr renaming (_+_ to _+G_ ; -_ to -G_) + + strA = snd AGr + strB = snd BGr + + A = fst AGr + B = fst BGr + + 0A = 0g strA + 0B = 0g strB + + _+A_ = _+G_ strA + _+B_ = _+G_ strB + + -A_ = -G_ strA + -B_ = -G_ strB + + mainF : GroupHom (AGr *) (HomGroup (BGr *) (AGr ⨂ BGr) *) + fst (fst mainF a) b = a ⊗ b + snd (fst mainF a) = + makeIsGroupHom (linr a) + snd mainF = + makeIsGroupHom λ x y → Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (linl x y)) + + isTensorProduct⨂ : (isTensorProductOf AGr and BGr) (AGr ⨂ BGr) + fst isTensorProduct⨂ = mainF + snd isTensorProduct⨂ C = + isoToIsEquiv mainIso + where + invF : GroupHom (AGr *) (HomGroup (BGr *) C *) → GroupHom ((AGr ⨂ BGr) *) (C *) + fst (invF (f , p)) = F + where + lem : f (0g (snd AGr)) .fst (0g (snd BGr)) ≡ 0g (snd C) + lem = funExt⁻ (cong fst (IsGroupHom.pres1 p)) (0g (snd BGr)) + F : ((AGr ⨂ BGr) *) .fst → (C *) .fst + F (a ⊗ b) = f a .fst b + F (x +⊗ x₁) = _+G_ (snd C) (F x) (F x₁) + F (⊗comm x x₁ i) = comm (snd C) (F x) (F x₁) i + F (⊗assoc x y z i) = assoc (snd C) (F x) (F y) (F z) i + F (⊗lUnit x i) = (cong (λ y → _+G_ (snd C) y (F x)) lem ∙ lid (snd C) (F x)) i + F (linl x y z i) = fst (IsGroupHom.pres· p x y i) z + F (linr x y z i) = IsGroupHom.pres· (f x .snd) y z i + F (flip x b i) = (IsGroupHom.presinv (f x .snd) b + ∙ sym (funExt⁻ (cong fst (IsGroupHom.presinv p x)) b)) i + F (⊗squash x x₁ x₂ y i i₁) = + is-set (snd C) (F x) (F x₁) + (λ i → F (x₂ i)) (λ i → F (y i)) i i₁ + snd (invF (f , p)) = + makeIsGroupHom λ x y → refl + + mainIso : Iso (GroupHom ((AGr ⨂ BGr) *) (C *)) + (GroupHom (AGr *) (HomGroup (BGr *) C *)) + Iso.fun mainIso = _ + Iso.inv mainIso = invF + Iso.rightInv mainIso (f , p) = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ a → Σ≡Prop (λ _ → isPropIsGroupHom _ _) refl) + Iso.leftInv mainIso (f , p) = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (⊗elimProp (λ _ → is-set (snd C) _ _) + (λ _ _ → refl) + λ x y ind1 ind2 → cong₂ (_+G_ (snd C)) ind1 ind2 ∙ sym (IsGroupHom.pres· p x y))) + +-- CODE : (z x : FMSet (A × B)) (y : _⨂₁_) → TypeOfHLevel (ℓ-max ℓ ℓ') 1 +-- JFun : (z x : {!!}) → {!!} → {!!} +-- JFun = {!!} +-- CODE z x (inc x₁) = Path _⨂₁_ (inc (z ++ x)) (inc (z ++ x₁)) , ⊗squash _ _ +-- CODE z x (unit i) = {!!} +-- CODE z x (lin-l x₁ y b i) = {!!} +-- where +-- h : (z : _) → Path (TypeOfHLevel (ℓ-max ℓ ℓ') 1) +-- (Path _⨂₁_ (inc (z ++ x)) (inc (z ++ [ (x₁ +A y) , b ])) , +-- ⊗squash (inc (z ++ x)) (inc (z ++ [ (x₁ +A y) , b ]))) +-- (Path _⨂₁_ (inc (z ++ x)) +-- (inc (z ++ ((x₁ , b) ∷ [ y , b ]))) +-- , ⊗squash (inc (z ++ x)) (inc (z ++ ((x₁ , b) ∷ [ y , b ])))) +-- h = ElimProp.f {!λ xs → ?!} +-- (Σ≡Prop {!!} {!!}) +-- λ m {xs} p → Σ≡Prop {!!} (cong (Path _⨂₁_ (inc ((m ∷ xs) ++ x))) {!p!}) -- λ m xss → Σ≡Prop {!!} (cong (Path _⨂₁_ (inc ((m ∷ xs) ++ x))) {!!}) -- Σ≡Prop {!!} {!!} +-- CODE z x (lin-r x₁ y z₁ i) = {!!} +-- CODE z x (⊗squash y y₁ x₁ y₂ i i₁) = {!!} + + +-- inclem : (x y z : FMSet (A × B)) → Path _⨂₁_ (inc x) (inc y) → Path _⨂₁_ (inc (z ++ x)) (inc (z ++ y)) +-- inclem x y z = {!!} + +-- _+⊗_ : _⨂₁_ → _⨂₁_ → _⨂₁_ +-- inc x +⊗ inc x₁ = inc (x ++ x₁) +-- inc x +⊗ unit i = {!!} +-- inc x +⊗ lin-l y z b i = {!!} +-- where +-- lem : (x : _) → inc (x ++ [ (y +A z) , b ]) ≡ inc (x ++ ((y , b) ∷ [ z , b ])) +-- lem = +-- ElimProp.f {!!} +-- (lin-l y z b) +-- λ x {xs} p → inclem _ _ [ x ] p -- inclem _ _ {!x ∷ xs!} p +-- inc x +⊗ lin-r x₁ y z i = {!!} +-- inc x +⊗ ⊗squash y y₁ x₁ y₂ i i₁ = {!!} +-- unit i +⊗ y = {!!} +-- lin-l x y₁ b i +⊗ y = {!!} +-- lin-r x y₁ z i +⊗ y = {!!} +-- ⊗squash x x₁ x₂ y₁ i i₁ +⊗ y = {!!} + +-- -- Word : (A : Type ℓ) → Type ℓ +-- -- Word A = List (A ⊎ A) + +-- -- module _ {A : Type ℓ} where +-- -- pm = A ⊎ A + +-- -- flip : pm → pm +-- -- flip (inl x) = inr x +-- -- flip (inr x) = inl x + +-- -- flip-flip : (x : pm) → flip (flip x) ≡ x +-- -- flip-flip (inl x) = refl +-- -- flip-flip (inr x) = refl + +-- -- wordFlip : Word A → Word A +-- -- wordFlip = map flip + +-- -- wordInverse : Word A → Word A +-- -- wordInverse = rev ∘ wordFlip + +-- -- map++ : {B : Type ℓ} (v w : Word A) +-- -- → (f : _ → B) → map f (v ++ w) +-- -- ≡ map f v ++ map f w +-- -- map++ [] w f = refl +-- -- map++ (x ∷ v) w f = cong (f x ∷_) (map++ v w f) + +-- -- wordInverseInverse : (x : Word A) → wordInverse (wordInverse x) ≡ x +-- -- wordInverseInverse [] = refl +-- -- wordInverseInverse (x ∷ x₁) = +-- -- cong rev (map++ (rev (map flip x₁)) (flip x ∷ []) flip) +-- -- ∙∙ rev-++ (map flip (rev (map flip x₁))) (flip (flip x) ∷ []) +-- -- ∙∙ cong₂ _∷_ (flip-flip x) (wordInverseInverse x₁) + +-- -- wordExp : A → ℤ → Word A +-- -- wordExp a (pos zero) = [] +-- -- wordExp a (pos (suc n)) = inl a ∷ wordExp a (pos n) +-- -- wordExp a (negsuc zero) = [ inr a ] +-- -- wordExp a (negsuc (suc n)) = inr a ∷ wordExp a (negsuc n) + +-- -- module _ {A : Type ℓ} (G : Group ℓ') (f : A → fst G) where +-- -- private +-- -- str = snd G +-- -- open GroupStr str renaming (_·_ to _·G_) + +-- -- PlusMinus-extendᴳ : A ⊎ A → fst G +-- -- PlusMinus-extendᴳ (inl x) = f x +-- -- PlusMinus-extendᴳ (inr x) = inv (f x) + +-- -- PlusMinus-extendᴳ-flip : ∀ x → PlusMinus-extendᴳ (flip x) +-- -- ≡ inv (PlusMinus-extendᴳ x) +-- -- PlusMinus-extendᴳ-flip (inl x) = refl +-- -- PlusMinus-extendᴳ-flip (inr x) = sym (GroupTheory.invInv G (f x)) + +-- -- Word-extendᴳ : Word A → fst G +-- -- Word-extendᴳ = foldr _·G_ 1g ∘ map PlusMinus-extendᴳ + +-- -- Word-extendᴳ-++ : ∀ x y → Word-extendᴳ (x ++ y) +-- -- ≡ (Word-extendᴳ x ·G Word-extendᴳ y) +-- -- Word-extendᴳ-++ [] y = sym (lid (Word-extendᴳ y)) +-- -- Word-extendᴳ-++ (x ∷ x₁) y = +-- -- cong (PlusMinus-extendᴳ x ·G_) (Word-extendᴳ-++ x₁ y) +-- -- ∙∙ assoc _ _ _ +-- -- ∙∙ cong (_·G Word-extendᴳ y) (help x x₁) +-- -- where +-- -- help : (x : _) (y : _) +-- -- → PlusMinus-extendᴳ x ·G Word-extendᴳ y ≡ Word-extendᴳ (x ∷ y) +-- -- help x [] = refl +-- -- help x (z ∷ y) = refl + + +-- -- module _ {A : Type ℓ} (G : AbGroup ℓ') where +-- -- private +-- -- str = snd G +-- -- open AbGroupStr str renaming (_+_ to _+G_ ; -_ to -G_) +-- -- G' = AbGroup→Group G + +-- -- PlusMinus-extendᴳ-comp : (f g : A → fst G) (x : _) +-- -- → PlusMinus-extendᴳ G' (λ a → f a +G g a) x +-- -- ≡ (PlusMinus-extendᴳ G' f x +G PlusMinus-extendᴳ G' g x) +-- -- PlusMinus-extendᴳ-comp f g (inl x) = refl +-- -- PlusMinus-extendᴳ-comp f g (inr x) = +-- -- GroupTheory.invDistr G' (f x) (g x) ∙ comm _ _ + +-- -- Word-extendᴳ-comp : (f g : A → fst G) (x : _) → Word-extendᴳ G' (λ a → f a +G g a) x +-- -- ≡ (Word-extendᴳ G' f x +G Word-extendᴳ G' g x) +-- -- Word-extendᴳ-comp f g [] = sym (lid 0g) +-- -- Word-extendᴳ-comp f g (x ∷ x₁) = +-- -- cong₂ _+G_ (PlusMinus-extendᴳ-comp f g x) +-- -- (Word-extendᴳ-comp f g x₁) +-- -- ∙∙ sym (assoc pm-fx pm-gx _) +-- -- ∙∙ cong (pm-fx +G_) (assoc _ _ _ +-- -- ∙∙ cong (_+G we-gx) (comm _ _) +-- -- ∙∙ sym (assoc _ _ _)) +-- -- ∙∙ assoc pm-fx we-fx _ +-- -- ∙∙ cong₂ _+G_ (main x x₁ f) (main x x₁ g) +-- -- where +-- -- main : (x : _) (y : _) (f : _) +-- -- → (PlusMinus-extendᴳ G' f x +G Word-extendᴳ G' f y) +-- -- ≡ Word-extendᴳ G' f (x ∷ y) +-- -- main x [] f = refl +-- -- main x (x₁ ∷ y) f = refl + +-- -- pm-fx = PlusMinus-extendᴳ G' f x +-- -- pm-gx = PlusMinus-extendᴳ G' g x +-- -- we-fx = Word-extendᴳ G' f x₁ +-- -- we-gx = Word-extendᴳ G' g x₁ + +-- -- module GeneratedGroup (A : Type ℓ) (R : Rel (Word A) (Word A) ℓ') where +-- -- data QuotWordRel : Word A → Word A → Type (ℓ-max ℓ ℓ') where +-- -- qwr-refl : ∀ {x} → QuotWordRel x x +-- -- qwr-trans : ∀ {x y z} → QuotWordRel x y → QuotWordRel y z → QuotWordRel x z +-- -- qwr-sym : ∀ {x y} → QuotWordRel x y → QuotWordRel y x +-- -- qwr-cong : ∀ {x y z w} → QuotWordRel x y → QuotWordRel z w → QuotWordRel (x ++ z) (y ++ w) +-- -- qwr-flip-r : ∀ x → QuotWordRel (x ∷ flip x ∷ []) [] +-- -- qwr-rel : ∀ {x y} → R x y → QuotWordRel x y + +-- -- qwr-cong-refl : {!!} +-- -- qwr-cong-refl = {!!} + +-- -- qwr-cong-l : {!!} +-- -- qwr-cong-l = {!!} + + + + + +-- -- {- +-- -- module _ {ℓ : Level} where +-- -- ℤmult : {A : Type ℓ} (_+_ : A → A → A) (-A_ : A → A) (0A : A) (n : ℤ) → A → A +-- -- ℤmult _+_ -A_ 0A (pos zero) a = 0A +-- -- ℤmult _+_ -A_ 0A (pos (suc n)) a = a + ℤmult _+_ -A_ 0A (pos n) a +-- -- ℤmult _+_ -A_ 0A (negsuc zero) a = -A a +-- -- ℤmult _+_ -A_ 0A (negsuc (suc n)) a = (-A a) + ℤmult _+_ -A_ 0A (negsuc n) a + +-- -- data commList (A : Type ℓ) : Type ℓ where +-- -- [] : commList A +-- -- _∷_ : A → commList A → commList A +-- -- isComm : (a b : A) (w : commList A) → a ∷ b ∷ w ≡ b ∷ a ∷ w +-- -- squashCommList : isSet (commList A) + +-- -- _++'_ : {A : Type ℓ} → commList A → commList A → commList A +-- -- [] ++' y = y +-- -- (x ∷ x₁) ++' y = x ∷ (x₁ ++' y) +-- -- isComm a b w i ++' y = isComm a b (w ++' y) i +-- -- squashCommList x y p q i j ++' z = +-- -- squashCommList _ _ (λ j → p j ++' z) (λ j → q j ++' z) i j + +-- -- propElimCommList : ∀ {ℓ'} {A : Type ℓ} {B : (x : commList A) → Type ℓ'} +-- -- → ((x : _) → isProp (B x)) +-- -- → B [] +-- -- → ((x : A) (y : commList A) → B y → B (x ∷ y)) +-- -- → (x : _) → B x +-- -- propElimCommList {B = B} prop c l [] = c +-- -- propElimCommList {B = B} prop c l (x ∷ x₁) = l x x₁ (propElimCommList {B = B} prop c l x₁) +-- -- propElimCommList {B = B} prop c l (isComm a b x i) = +-- -- isOfHLevel→isOfHLevelDep 1 prop +-- -- (l a (b ∷ x) (l b x (propElimCommList prop c l x))) (l b (a ∷ x) (l a x (propElimCommList prop c l x))) +-- -- (isComm a b x) i +-- -- propElimCommList {B = B} prop c l (squashCommList x y p q i j) = +-- -- isOfHLevel→isOfHLevelDep 2 {B = B} (λ _ → isProp→isSet (prop _)) +-- -- _ _ (λ j → propElimCommList prop c l (p j)) +-- -- (λ j → propElimCommList prop c l (q j)) +-- -- (squashCommList x y p q) i j + +-- -- propElimCommList' : +-- -- ∀ {ℓ'} {A : Type ℓ} {B : (x : commList A) → Type ℓ'} +-- -- → ((x : _) → isProp (B x)) +-- -- → B [] +-- -- → ((x : A) → B (x ∷ [])) +-- -- → ((x : A) (y : commList A) → B (x ∷ []) → B y → B (x ∷ y)) +-- -- → (x : _) → B x +-- -- propElimCommList' {B = B} prop b ind1 ind2 = +-- -- propElimCommList prop b h +-- -- where +-- -- h : (x : _) (y : commList _) → B y → B (x ∷ y) +-- -- h x [] _ = ind1 x +-- -- h x (x₁ ∷ y) l = ind2 x (x₁ ∷ y) (h x [] b) l +-- -- h x (isComm a b y i) = K i +-- -- where +-- -- K : PathP (λ i → B (isComm a b y i) → B (x ∷ isComm a b y i)) +-- -- (ind2 x (a ∷ (b ∷ y)) (ind1 x)) (ind2 x (b ∷ (a ∷ y)) (ind1 x)) +-- -- K = isProp→PathP (λ _ → isPropΠ λ _ → prop _) _ _ +-- -- h x (squashCommList w z p q i j) = K i j +-- -- where +-- -- K : SquareP (λ i j → B (squashCommList w z p q i j) → B (x ∷ squashCommList w z p q i j)) +-- -- (λ j → h x (p j)) +-- -- (λ j → h x (q j)) +-- -- (λ _ → h x w) +-- -- λ _ → h x z +-- -- K = toPathP (isOfHLevelPathP' 1 (isProp→isSet (isPropΠ (λ _ → prop _))) _ _ _ _) + +-- -- ++'-assoc : {A : Type ℓ} → (x y z : commList A) → x ++' (y ++' z) ≡ ((x ++' y) ++' z) +-- -- ++'-assoc = +-- -- propElimCommList' +-- -- (λ _ → isPropΠ2 λ _ _ → squashCommList _ _) +-- -- (λ _ _ → refl) +-- -- (λ x → propElimCommList' (λ _ → isPropΠ λ _ → squashCommList _ _) +-- -- (λ _ → refl) +-- -- (λ _ _ → refl) +-- -- λ x y p q → propElimCommList' (λ _ → squashCommList _ _) +-- -- refl +-- -- (λ _ → refl) +-- -- λ z w P Q → refl) +-- -- λ x y p q z w → cong (x ∷_) (q z w) + +-- -- ++'-comm : {A : Type ℓ} → (x y : commList A) → x ++' y ≡ y ++' x +-- -- ++'-comm = +-- -- propElimCommList' (λ _ → isPropΠ λ _ → squashCommList _ _) +-- -- (propElimCommList' (λ _ → squashCommList _ _) +-- -- refl +-- -- (λ _ → refl) +-- -- (λ x y p q → cong (x ∷_) q)) +-- -- (λ x → propElimCommList' +-- -- (λ _ → squashCommList _ _) +-- -- refl +-- -- (λ _ → isComm _ _ _) +-- -- λ y w p q → isComm x y w ∙ cong (y ∷_) q) +-- -- λ x y p q r → p (y ++' r) +-- -- ∙∙ cong (_++' (x ∷ [])) (q r) +-- -- ∙∙ (sym (++'-assoc r y (x ∷ [])) ∙ cong (r ++'_) (q (x ∷ []))) + +-- -- len : {A : Type ℓ} → (x : commList A) → ℕ +-- -- len [] = 0 +-- -- len (x ∷ x₁) = ℕ._+_ 1 (len x₁) +-- -- len (isComm a b x i) = suc (suc (len x)) +-- -- len (squashCommList x x₁ x₂ y i i₁) = +-- -- isSetℕ (len x) (len x₁) (λ i → len (x₂ i)) (λ i → len (y i)) i i₁ + +-- -- module _ (A : AbGroup ℓ) (B : AbGroup ℓ') where +-- -- private +-- -- strA = snd A +-- -- strB = snd B +-- -- open AbGroupStr renaming (_+_ to +G ; -_ to -G) + +-- -- _⨂₁-raw2_ : Type _ +-- -- _⨂₁-raw2_ = List (fst A × fst B) + +-- -- A' = fst A +-- -- B' = fst B + +-- -- data _⨂₁_ : Type (ℓ-max ℓ ℓ') where +-- -- inc : commList (A' × B') → _⨂₁_ +-- -- bilinl : (x y : A') (b : B') (w : _) +-- -- → inc ((+G strA x y , b) ∷ w) ≡ inc (((x , b) ∷ w) ++' ((y , b) ∷ w)) +-- -- bilinr : (x : A') (y z : B') (w : _) +-- -- → inc ((x , +G strB y z) ∷ w) ≡ inc (((x , y) ∷ w) ++' ((x , z) ∷ w)) +-- -- l : (x : A') (y : B') → inc ((x , y) ∷ []) ≡ inc [] +-- -- ⊗squash : isSet _⨂₁_ + +-- -- +⊗ : _⨂₁_ → _⨂₁_ → _⨂₁_ +-- -- +⊗ x y = {!!} + +-- -- -- (inc x) (inc y) = inc (x ++' y) +-- -- -- +⊗ (inc x) (bilinl y z w r i) = commListInd x i +-- -- -- where +-- -- -- commListInd : (x : _) → inc (x ++' ((+G strA y z , w) ∷ r)) ≡ inc (x ++' ((y , w) ∷ (r ++' ((z , w) ∷ r)))) +-- -- -- commListInd = +-- -- -- propElimCommList' (λ _ → ⊗squash _ _) +-- -- -- (bilinl y z w r) +-- -- -- (λ l → (λ i → inc ((l ∷ []) ++' (((+G strA y z , w) ∷ []) ++' r))) ∙∙ {!((+G strA y z , w₁) ∷ r)!} ∙∙ {!!}) +-- -- -- λ x y p q → {!!} +-- -- -- ∙∙ {!!} +-- -- -- ∙∙ {!!} + +-- -- -- lem : ((r ++' x) ++' ((z , w) ∷ (r ++' x))) ≡ ((r ++' ((z , w) ∷ r)) ++' x) +-- -- -- lem = {!!} +-- -- -- ∙∙ {!!} +-- -- -- ∙∙ ++'-assoc r ((z , w) ∷ r) x +-- -- -- help : inc (x ++' ((+G strA y z , w) ∷ r)) ≡ inc (x ++' ((y , w) ∷ (r ++' ((z , w) ∷ r)))) +-- -- -- help = cong inc (++'-comm x _) +-- -- -- ∙∙ bilinl y z w (r ++' x) +-- -- -- ∙∙ (λ i → inc (((y , w) ∷ []) ++' ((r ++' x) ++' (((z , w) ∷ []) ++' (r ++' x))))) +-- -- -- ∙∙ cong inc (cong (((y , w) ∷ []) ++'_) {!!}) +-- -- -- ∙∙ {!!} +-- -- -- ∙∙ {!!} +-- -- -- ∙∙ cong inc (cong (((y , w) ∷ []) ++'_) {!!}) +-- -- -- ∙∙ (λ i → inc (((y , w) ∷ []) ++' ((r ++' (((z , w) ∷ []) ++' r)) ++' x))) +-- -- -- ∙∙ cong inc (sym (++'-comm x ((y , w) ∷ (r ++' ((z , w) ∷ r))))) +-- -- -- +⊗ (inc x) (bilinr y z w r i) = {!!} +-- -- -- +⊗ (inc x) (⊗squash y y₁ x₁ y₂ i i₁) = {!!} +-- -- -- +⊗ (bilinl x y z w i) r = {!!} +-- -- -- +⊗ (bilinr x y₁ z w i) y = {!!} +-- -- -- +⊗ (⊗squash x x₁ x₂ y₁ i i₁) y = {!!} + + + +-- -- -- -- _⊗_ : A' → B' → _⨂₁-raw_ +-- -- -- -- a ⊗ b = inc [ a , b ] + +-- -- -- -- invList : _⨂₁-raw2_ → _⨂₁-raw2_ +-- -- -- -- invList [] = [] +-- -- -- -- invList (x ∷ x₁) = +-- -- -- -- ((-G strA (fst x)) , (snd x)) ∷ invList x₁ + + + +-- -- -- -- _ℤ∙_ : ℤ → _⨂₁-raw2_ → _⨂₁-raw2_ +-- -- -- -- pos zero ℤ∙ y = [] +-- -- -- -- pos (suc n) ℤ∙ y = y ++ (pos n ℤ∙ y) +-- -- -- -- negsuc zero ℤ∙ y = invList y +-- -- -- -- negsuc (suc n) ℤ∙ y = (invList y) ++ (negsuc n ℤ∙ y) + +-- -- -- -- 0⊗ₗ : (b : B') → 0g strA ⊗ b ≡ (0g strA ⊗ 0g strB) +-- -- -- -- 0⊗ₗ b = {!!} +-- -- -- -- where +-- -- -- -- lem : (0g strA ⊗ b) ≡ inc ([ 0g strA , b ] ++ [ 0g strA , b ]) +-- -- -- -- lem = cong (_⊗ b) (sym (rid strA (0g strA))) ∙ bilinl (0g strA) (0g strA) b + +-- -- -- -- module _ {A : AbGroup ℓ} {B : AbGroup ℓ'} where +-- -- -- -- private +-- -- -- -- strA = snd A +-- -- -- -- strB = snd B +-- -- -- -- open AbGroupStr renaming (_+_ to +G ; -_ to -G) + +-- -- -- -- rec⨂₁ : ∀ {ℓ'} {C : Type ℓ'} +-- -- -- -- → isSet C +-- -- -- -- → (c : C) +-- -- -- -- → (f : fst A → fst B → C → C) +-- -- -- -- → (((x y : fst A) (z : fst B) → f (+G strA x y) z c ≡ f x z (f y z c))) +-- -- -- -- → ((x : fst A) (y z : fst B) → f x (+G strB y z) c ≡ f x y (f x z c)) +-- -- -- -- → (x : A ⨂₁-raw B) → C +-- -- -- -- rec⨂₁ {ℓ'} {C} set c f bil bir x = {!!} +-- -- -- -- -- set _ _ (λ j → rec⨂₁ set c f bil bir (p j)) (λ j → rec⨂₁ set c f bil bir (q j)) i j + +-- -- -- -- elimProp : ∀ {ℓ'} {C : A ⨂₁-raw B → Type ℓ'} +-- -- -- -- → ((x : _) → isProp (C x)) +-- -- -- -- → ((x : A ⨂₁-raw2 B) → C (inc x)) +-- -- -- -- → (x : _) → C x +-- -- -- -- elimProp {C = C} prop ind (inc x) = ind x +-- -- -- -- elimProp {C = C} prop ind (bilinl x y z i) = +-- -- -- -- isOfHLevel→isOfHLevelDep 1 {B = C} prop +-- -- -- -- (ind (((+G strA x y) , z) ∷ [])) +-- -- -- -- (ind ((x , z) ∷ (y , z) ∷ [])) (bilinl x y z) i +-- -- -- -- elimProp {C = C} prop ind (bilinr x y z i) = +-- -- -- -- isOfHLevel→isOfHLevelDep 1 {B = C} prop +-- -- -- -- (ind ((x , (+G strB y z)) ∷ [])) +-- -- -- -- (ind ((x , y) ∷ (x , z) ∷ [])) (bilinr x y z) i +-- -- -- -- elimProp {C = C} prop ind (⊗squash x y p q i j) = +-- -- -- -- isOfHLevel→isOfHLevelDep 2 {B = C} (λ x → isProp→isSet (prop x)) _ _ +-- -- -- -- (λ j → elimProp prop ind (p j)) +-- -- -- -- (λ j → elimProp prop ind (q j)) +-- -- -- -- (⊗squash x y p q) i j + +-- -- -- -- AA = fst A +-- -- -- -- BB = fst B + +-- -- -- -- bilinRGen : (x : fst A) (y z : fst B) (w : _) → Path (A ⨂₁-raw B) (inc ((x , +G strB y z) ∷ w)) (inc ((x , y) ∷ (x , z) ∷ w)) +-- -- -- -- bilinRGen x y z [] = bilinr x y z +-- -- -- -- bilinRGen x y z (x₁ ∷ w) = {!bilinRGen x y z w!} + +-- -- -- -- pre-comm : (x y : A ⨂₁-raw2 B) → Path (A ⨂₁-raw B) (inc (x ++ y)) (inc (y ++ x)) +-- -- -- -- pre-comm [] [] = refl +-- -- -- -- pre-comm [] (x ∷ y) = {!!} -- cong inc (cong (x ∷_) {!pre-comm [] y!}) +-- -- -- -- pre-comm (x ∷ x₁) y = {!!} + +-- -- -- -- +⊗-mere : (x y : A ⨂₁-raw2 B) (z : _) → Path (A ⨂₁-raw B) (inc x) (inc y) → Path (A ⨂₁-raw B) (inc (z ∷ x)) (inc (z ∷ y)) +-- -- -- -- +⊗-mere [] [] z p = refl +-- -- -- -- +⊗-mere [] (x ∷ y) z p = {!!} +-- -- -- -- +⊗-mere (x ∷ x₁) y z = {!!} + +-- -- -- -- +⊗-merecomm' : (x y : fst A × fst B) → Path (A ⨂₁-raw B) (inc ([ x ] ++ [ y ])) (inc ([ y ] ++ [ x ])) +-- -- -- -- +⊗-merecomm' (a , b) (c , d) = +-- -- -- -- cong inc (cong (_∷ (c , d) ∷ []) (cong (_, b) (sym (rid strA a) +-- -- -- -- ∙∙ cong (+G strA a) (sym (invr strA c)) +-- -- -- -- ∙∙ {!!}))) +-- -- -- -- ∙∙ {!!} +-- -- -- -- ∙∙ {!!} + + +-- -- -- -- _+⊗_ : A ⨂₁-raw B → A ⨂₁-raw B → A ⨂₁-raw B +-- -- -- -- inc x +⊗ inc x₁ = inc (x ++ x₁) +-- -- -- -- inc [] +⊗ bilinl x₁ y z i = bilinl x₁ y z i +-- -- -- -- inc ((a , b) ∷ []) +⊗ bilinl x₁ y z i = {!!} -- help i +-- -- -- -- where +-- -- -- -- help : inc +-- -- -- -- ((a , b) ∷ +-- -- -- -- (+G strA x₁ y , z) ∷ +-- -- -- -- []) ≡ inc ((a , b) ∷ (x₁ , z) ∷ (y , z) ∷ []) +-- -- -- -- help = {!sym (bilinl a _ b)!} ∙ {!!} +-- -- -- -- inc (x ∷ y ∷ z) +⊗ bilinl a b c i = {!!} -- inc (x ∷ x₃) +⊗ {!!} +-- -- -- -- inc x +⊗ bilinr x₁ y z i = {!!} +-- -- -- -- inc x +⊗ ⊗squash y y₁ x₁ y₂ i i₁ = {!!} +-- -- -- -- bilinl x y₁ z i +⊗ y = {!!} +-- -- -- -- bilinr x y₁ z i +⊗ y = {!!} +-- -- -- -- ⊗squash x x₁ x₂ y₁ i i₁ +⊗ y = {!!} + +-- -- -- -- -- module _ (A : Type ℓ) (R : Rel (Word A) (Word A) ℓ') where +-- -- -- -- -- data AbGroupRel : Word A → Word A → Type (ℓ-max ℓ ℓ') where +-- -- -- -- -- agr-commutes : ∀ x y → AbGroupRel (x ++ y) (y ++ x) +-- -- -- -- -- agr-refl : ∀ {x y} → R x y → AbGroupRel x y + +-- -- -- -- -- agr-rev : (w : Word A) → {!⨂₁!} -- QuotWordRel (rev w) w +-- -- -- -- -- agr-rev = {!!} + +-- -- -- -- -- module _ (G H : AbGroup ℓ) where +-- -- -- -- -- TensorCarrier : Type _ +-- -- -- -- -- TensorCarrier = +-- -- -- -- -- Σ[ T ∈ AbGroup ℓ ] +-- -- -- -- -- (Σ[ t ∈ (fst G → fst H → fst T) ] +-- -- -- -- -- ((C : AbGroup ℓ) +-- -- -- -- -- → isEquiv {A = fst T → fst C} {B = fst G → fst H → fst C} +-- -- -- -- -- λ f a b → f (t a b))) +-- -- -- -- -- anIso : {!!} +-- -- -- -- -- anIso = {!!} + +-- -- -- -- -- -- 0⊗ : TensorCarrier +-- -- -- -- -- -- fst 0⊗ = dirProdAb G H +-- -- -- -- -- -- fst (snd 0⊗) x y = x , y +-- -- -- -- -- -- snd (snd 0⊗) C = isoToIsEquiv {!0⊗!} +-- -- -} diff --git a/Cubical/Algebra/Group/EilenbergMacLane/Base.agda b/Cubical/Algebra/Group/EilenbergMacLane/Base.agda new file mode 100644 index 000000000..687c1fe4a --- /dev/null +++ b/Cubical/Algebra/Group/EilenbergMacLane/Base.agda @@ -0,0 +1,90 @@ +{-# OPTIONS --cubical --no-import-sorts --safe --experimental-lossy-unification #-} + +module Cubical.Algebra.Group.EilenbergMacLane.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Properties +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec; elim to trElim) +open import Cubical.HITs.EilenbergMacLane1 +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Data.Empty + renaming (rec to ⊥-rec) +open import Cubical.HITs.Truncation + renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) +open import Cubical.Data.Nat hiding (_·_) +open import Cubical.HITs.Susp +open import Cubical.Functions.Morphism +open import Cubical.Foundations.Path + +private + variable ℓ : Level + + _* = AbGroup→Group + +EM-raw : (G : AbGroup ℓ) (n : ℕ) → Type ℓ +EM-raw G zero = fst G +EM-raw G (suc zero) = EM₁ (G *) +EM-raw G (suc (suc n)) = Susp (EM-raw G (suc n)) + +ptS : {n : ℕ} {G : AbGroup ℓ} → EM-raw G n +ptS {n = zero} {G = G} = AbGroupStr.0g (snd G) +ptS {n = suc zero} {G = G} = embase +ptS {n = suc (suc n)} {G = G} = north + +EM-raw-elim : (G : AbGroup ℓ) (n : ℕ) {A : EM-raw G (suc n) → Type ℓ} + → ((x : _) → isOfHLevel (suc n) (A x) ) + → A ptS + → (x : _) → A x +EM-raw-elim G zero hlev b = elimProp _ hlev b +EM-raw-elim G (suc n) hlev b north = b +EM-raw-elim G (suc n) {A = A} hlev b south = subst A (merid ptS) b +EM-raw-elim G (suc n) {A = A} hlev b (merid a i) = help a i + where + help : (a : _) → PathP (λ i → A (merid a i)) b (subst A (merid ptS) b) + help = EM-raw-elim G n (λ _ → isOfHLevelPathP' (suc n) (hlev _) _ _) + λ i → transp (λ j → A (merid ptS (j ∧ i))) (~ i) b + +EM : (G : AbGroup ℓ) (n : ℕ) → Type ℓ +EM G zero = EM-raw G zero +EM G (suc zero) = EM-raw G 1 +EM G (suc (suc n)) = hLevelTrunc (4 + n) (EM-raw G (suc (suc n))) + +0ₖ : {G : AbGroup ℓ} (n : ℕ) → EM G n +0ₖ {G = G} zero = AbGroupStr.0g (snd G) +0ₖ (suc zero) = embase +0ₖ (suc (suc n)) = ∣ ptS ∣ + +EM∙ : (G : AbGroup ℓ) (n : ℕ) → Pointed ℓ +EM∙ G n = EM G n , (0ₖ n) + +hLevelEM : (G : AbGroup ℓ) (n : ℕ) → isOfHLevel (2 + n) (EM G n) +hLevelEM G zero = AbGroupStr.is-set (snd G) +hLevelEM G (suc zero) = emsquash +hLevelEM G (suc (suc n)) = isOfHLevelTrunc (4 + n) + +EM-raw→EM : (G : AbGroup ℓ) (n : ℕ) → EM-raw G n → EM G n +EM-raw→EM G zero x = x +EM-raw→EM G (suc zero) x = x +EM-raw→EM G (suc (suc n)) = ∣_∣ + +EM-elim : {G : AbGroup ℓ} (n : ℕ) {A : EM G n → Type ℓ} + → ((x : _) → isOfHLevel (2 + n) (A x)) + → ((x : EM-raw G n) → A (EM-raw→EM G n x)) + → (x : _) → A x +EM-elim zero hlev hyp x = hyp x +EM-elim (suc zero) hlev hyp x = hyp x +EM-elim (suc (suc n)) hlev hyp = trElim (λ _ → hlev _) hyp diff --git a/Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda b/Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda new file mode 100644 index 000000000..bfbdbf532 --- /dev/null +++ b/Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda @@ -0,0 +1,405 @@ +{-# OPTIONS --cubical --no-import-sorts --safe --experimental-lossy-unification #-} + +module Cubical.Algebra.Group.EilenbergMacLane.GroupStructure where + +open import Cubical.Algebra.Group.EilenbergMacLane.Base +open import Cubical.Algebra.Group.EilenbergMacLane.WedgeConnectivity +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Properties +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec; elim to trElim) +open import Cubical.HITs.EilenbergMacLane1 +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Data.Empty + renaming (rec to ⊥-rec) +open import Cubical.HITs.Truncation + renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) +open import Cubical.Data.Nat hiding (_·_) +open import Cubical.HITs.Susp +open import Cubical.Functions.Morphism +open import Cubical.Foundations.Path + +private + variable ℓ : Level + +module _ {G : AbGroup ℓ} where + infixr 34 _+ₖ_ + infixr 34 _-ₖ_ + open AbGroupStr (snd G) renaming (_+_ to _+G_ ; -_ to -G_ ; assoc to assocG) + + private + help : (n : ℕ) → n + (4 + n) ≡ (2 + n) + (2 + n) + help n = +-suc n (3 + n) ∙ cong suc (+-suc n (suc (suc n))) + + hLevHelp : (n : ℕ) → isOfHLevel ((2 + n) + (2 + n)) (EM G (2 + n)) + hLevHelp n = + transport (λ i → isOfHLevel (help n i) (EM G (2 + n))) + (isOfHLevelPlus {n = 4 + n} n (isOfHLevelTrunc (4 + n))) + + helper : (g h : fst G) + → PathP (λ i → Path (EM₁ (AbGroup→Group G)) + (emloop h i) (emloop h i)) (emloop g) (emloop g) + helper g h = + compPathL→PathP + (cong (sym (emloop h) ∙_) + (sym (emloop-comp _ g h) + ∙∙ cong emloop (comm g h) + ∙∙ emloop-comp _ h g) + ∙∙ ∙assoc _ _ _ + ∙∙ cong (_∙ emloop g) (lCancel _) + ∙ sym (lUnit _)) + _+ₖ_ : {n : ℕ} → EM G n → EM G n → EM G n + _+ₖ_ {n = zero} = _+G_ + _+ₖ_ {n = suc zero} = + rec _ (isGroupoidΠ (λ _ → emsquash)) + (λ x → x) + hi! + λ g h i j x → el g h x i j + where + lol : (g h : fst G) + → Square (emloop g) (emloop (g +G h)) refl (emloop h) + lol g h = compPath-filler (emloop g) (emloop h) ▷ sym (emloop-comp _ g h) + + hi! : fst G → (λ x → x) ≡ (λ x → x) + hi! g = funExt (elimSet _ (λ _ → emsquash _ _) + (emloop g) + (helper g)) + el : (g h : fst G) (x : EM₁ (AbGroup→Group G)) + → Square (λ j → hi! g j x) (λ j → hi! ((snd (AbGroup→Group G) GroupStr.· g) h) j x) + refl λ j → hi! h j x + el g h = + elimProp _ (λ _ → isOfHLevelPathP' 1 (emsquash _ _) _ _) + (lol g h) + _+ₖ_ {n = suc (suc n)} = + trRec2 (isOfHLevelTrunc (4 + n)) + (wedgeConEM.fun G G (suc n) (suc n) + (λ _ _ → hLevHelp n) + ∣_∣ ∣_∣ refl) + + σ-EM : (n : ℕ) → EM-raw G (suc n) → Path (EM-raw G (2 + n)) ptS ptS + σ-EM n x i = (merid x ∙ sym (merid ptS)) i + + -ₖ_ : {n : ℕ} → EM G n → EM G n + -ₖ_ {n = zero} x = -G x + -ₖ_ {n = suc zero} = + rec _ emsquash + embase + (λ g → sym (emloop g)) + λ g h → sym (emloop-sym _ g) + ◁ (flipSquare + (flipSquare (emcomp (-G g) (-G h)) + ▷ emloop-sym _ h) + ▷ (cong emloop (comm (-G g) (-G h) + ∙ sym (GroupTheory.invDistr (AbGroup→Group G) g h)) + ∙ emloop-sym _ (g +G h))) + -ₖ_ {n = suc (suc n)} = + map λ { north → north + ; south → north + ; (merid a i) → σ-EM n a (~ i)} + + _-ₖ_ : {n : ℕ} → EM G n → EM G n → EM G n + _-ₖ_ {n = n} x y = _+ₖ_ {n = n} x (-ₖ_ {n = n} y) + + +ₖ-syntax : (n : ℕ) → EM G n → EM G n → EM G n + +ₖ-syntax n = _+ₖ_ {n = n} + + -ₖ-syntax : (n : ℕ) → EM G n → EM G n + -ₖ-syntax n = -ₖ_ {n = n} + + -'ₖ-syntax : (n : ℕ) → EM G n → EM G n → EM G n + -'ₖ-syntax n = _-ₖ_ {n = n} + + syntax +ₖ-syntax n x y = x +[ n ]ₖ y + syntax -ₖ-syntax n x = -[ n ]ₖ x + syntax -'ₖ-syntax n x y = x -[ n ]ₖ y + + lUnitₖ : (n : ℕ) (x : EM G n) → 0ₖ n +[ n ]ₖ x ≡ x + lUnitₖ zero x = lid x + lUnitₖ (suc zero) _ = refl + lUnitₖ (suc (suc n)) = + trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ _ → refl + + rUnitₖ : (n : ℕ) (x : EM G n) → x +[ n ]ₖ 0ₖ n ≡ x + rUnitₖ zero x = rid x + rUnitₖ (suc zero) = + elimSet _ (λ _ → emsquash _ _) + refl + λ _ _ → refl + rUnitₖ (suc (suc n)) = + trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + (wedgeConEM.right G G (suc n) (suc n) + (λ _ _ → hLevHelp n) + ∣_∣ ∣_∣ refl) + + commₖ : (n : ℕ) (x y : EM G n) → x +[ n ]ₖ y ≡ y +[ n ]ₖ x + commₖ zero = comm + commₖ (suc zero) = + wedgeConEM.fun G G 0 0 (λ _ _ → emsquash _ _) + (λ x → sym (rUnitₖ 1 x)) + (rUnitₖ 1) + refl + commₖ (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelTruncPath {n = 4 + n}) + (wedgeConEM.fun G G _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (hLevHelp n) _ _) + (λ x → sym (rUnitₖ (2 + n) ∣ x ∣)) + (λ x → rUnitₖ (2 + n) ∣ x ∣) + refl) + + open import Cubical.Homotopy.Loopspace + cong₂+₁ : (p q : typ (Ω (EM∙ G 1))) → cong₂ (λ x y → x +[ 1 ]ₖ y) p q ≡ p ∙ q + cong₂+₁ p q = + (cong₂Funct (λ x y → x +[ 1 ]ₖ y) p q) + ∙ (λ i → (cong (λ x → rUnitₖ 1 x i) p) ∙ (cong (λ x → lUnitₖ 1 x i) q)) + + cong₂+₂ : (n : ℕ) (p q : typ (Ω (EM∙ G (suc (suc n))))) → cong₂ (λ x y → x +[ (2 + n) ]ₖ y) p q ≡ p ∙ q + cong₂+₂ n p q = + (cong₂Funct (λ x y → x +[ (2 + n) ]ₖ y) p q) + ∙ (λ i → (cong (λ x → rUnitₖ (2 + n) x i) p) ∙ (cong (λ x → lUnitₖ (2 + n) x i) q)) + + isCommΩEM : (n : ℕ) (p q : typ (Ω (EM∙ G (suc n)))) → p ∙ q ≡ q ∙ p + isCommΩEM zero p q = + sym (cong₂+₁ p q) + ∙∙ (λ i j → commₖ 1 (p j) (q j) i) + ∙∙ cong₂+₁ q p + isCommΩEM (suc n) p q = + (sym (cong₂+₂ n p q) + ∙∙ (λ i j → commₖ (suc (suc n)) (p j) (q j) i) + ∙∙ cong₂+₂ n q p) + + cong-₁ : (p : typ (Ω (EM∙ G 1))) → cong (λ x → -[ 1 ]ₖ x) p ≡ sym p + cong-₁ p = main embase p + where + decoder : (x : EM G 1) → embase ≡ x → x ≡ embase + decoder = + elimSet _ + (λ _ → isSetΠ λ _ → emsquash _ _) + (λ p i → -[ 1 ]ₖ (p i)) + λ g → toPathP + (funExt λ x → + (λ i → transport (λ i → Path (EM G 1) (emloop g i) embase) + (cong (-ₖ_ {n = 1}) + (transp (λ j → Path (EM G 1) embase (emloop g (~ j ∧ ~ i))) i + (compPath-filler x (sym (emloop g)) i) ))) + ∙∙ (λ i → transp (λ j → Path (EM G 1) (emloop g (i ∨ j)) embase) i + (compPath-filler' + (sym (emloop g)) + (cong-∙ (-ₖ_ {n = 1}) + x (sym (emloop g)) i) i)) + ∙∙ (cong (sym (emloop g) ∙_) (isCommΩEM 0 (cong (-ₖ_ {n = 1}) x) (emloop g))) + ∙∙ ∙assoc _ _ _ + ∙∙ cong (_∙ (cong (-ₖ_ {n = 1}) x)) (lCancel (emloop g)) + ∙ sym (lUnit _)) + + main : (x : EM G 1) (p : embase ≡ x) → decoder x p ≡ sym p + main x = J (λ x p → decoder x p ≡ sym p) refl + + cong-₂ : (n : ℕ) (p : typ (Ω (EM∙ G (2 + n)))) → cong (λ x → -[ 2 + n ]ₖ x) p ≡ sym p + cong-₂ n p = main _ p + where + pp : (a : _) → PathP (λ i → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ → ∣ merid a i ∣ₕ ≡ 0ₖ (2 + n)) + (cong (λ x → -[ 2 + n ]ₖ x)) λ p → cong ∣_∣ₕ (sym (merid ptS)) ∙ cong (λ x → -[ 2 + n ]ₖ x) p + pp a = + toPathP + (funExt λ x → + (λ k → transp (λ i → Path (EM G (2 + n)) ∣ merid a (i ∨ k) ∣ ∣ ptS ∣) k (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) + (cong (-ₖ-syntax (suc (suc n))) + (transp (λ j → Path (EM G (2 + n)) ∣ ptS ∣ ∣ merid a (~ j ∧ ~ k) ∣) k (compPath-filler x (sym (cong ∣_∣ₕ (merid a))) k))) k)) + ∙∙ cong (cong ∣_∣ₕ (sym (merid a)) ∙_) (cong-∙ (λ x → -[ 2 + n ]ₖ x) x (sym (cong ∣_∣ₕ (merid a))) + ∙ isCommΩEM (suc n) (cong (λ x → -[ 2 + n ]ₖ x) x) (cong ∣_∣ₕ (σ-EM n a))) + ∙∙ (λ k → (λ i → ∣ merid a (~ i ∨ k) ∣) ∙ (λ i → ∣ compPath-filler' (merid a) (sym (merid ptS)) (~ k) i ∣) + ∙ cong (λ x → -ₖ-syntax (suc (suc n)) x) x) + ∙ sym (lUnit _)) + + decoder : (x : EM G (2 + n)) → 0ₖ (2 + n) ≡ x → x ≡ 0ₖ (2 + n) + decoder = + trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ { north → pp ptS i0 + ; south → pp ptS i1 + ; (merid a i) → pp a i} + + main : (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → decoder x p ≡ sym p + main x = J (λ x p → decoder x p ≡ sym p) refl + + rCancelₖ : (n : ℕ) (x : EM G n) → x +[ n ]ₖ (-[ n ]ₖ x) ≡ 0ₖ n + rCancelₖ zero x = invr x + rCancelₖ (suc zero) = + elimSet _ (λ _ → emsquash _ _) + refl + λ g → flipSquare (cong₂+₁ (emloop g) (λ i → -ₖ-syntax 1 (emloop g i)) + ∙ rCancel (emloop g)) + rCancelₖ (suc (suc n)) = + trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ { north → refl + ; south i → +ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣ (-ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣) + ; (merid a i) j + → hcomp (λ r → λ { (i = i0) → 0ₖ (2 + n) + ; (i = i1) → ∣ merid ptS (~ j ∧ r) ∣ₕ -[ 2 + n ]ₖ ∣ merid ptS (~ j ∧ r) ∣ + ; (j = i0) → ∣ compPath-filler (merid a) (sym (merid ptS)) (~ r) i ∣ + -[ 2 + n ]ₖ ∣ compPath-filler (merid a) (sym (merid ptS)) (~ r) i ∣ + ; (j = i1) → 0ₖ (2 + n)}) + (help' a j i) } + where + help' : (a : _) → cong₂ (λ x y → ∣ x ∣ -[ suc (suc n) ]ₖ ∣ y ∣) (σ-EM n a) (σ-EM n a) ≡ refl + help' a = + cong₂+₂ n (cong ∣_∣ₕ (σ-EM n a)) (cong (λ x → -[ 2 + n ]ₖ ∣ x ∣) (σ-EM n a)) + ∙∙ cong (cong ∣_∣ₕ (σ-EM n a) ∙_) (cong-₂ n (cong ∣_∣ₕ (σ-EM n a))) + ∙∙ rCancel _ + + lCancelₖ : (n : ℕ) (x : EM G n) → (-[ n ]ₖ x) +[ n ]ₖ x ≡ 0ₖ n + lCancelₖ n x = commₖ n (-[ n ]ₖ x) x ∙ rCancelₖ n x + + assocₖ : (n : ℕ) (x y z : EM G n) + → (x +[ n ]ₖ (y +[ n ]ₖ z) ≡ (x +[ n ]ₖ y) +[ n ]ₖ z) + assocₖ zero = assocG + assocₖ (suc zero) = + elimSet _ (λ _ → isSetΠ2 λ _ _ → emsquash _ _) + (λ _ _ → refl) + λ g i y z k → lem g y z k i + where + lem : (g : fst G) (y z : _) → cong (λ x → x +[ suc zero ]ₖ (y +[ suc zero ]ₖ z)) (emloop g) + ≡ cong (λ x → (x +[ suc zero ]ₖ y) +[ suc zero ]ₖ z) (emloop g) + lem g = + elimProp _ (λ _ → isPropΠ λ _ → emsquash _ _ _ _) + (elimProp _ (λ _ → emsquash _ _ _ _) + refl) + assocₖ (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ a b → trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + (λ c → main c a b) + where + lem : (c : _) (a b : _) + → PathP (λ i → (∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ merid c i ∣ₕ) + ≡ (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ merid c i ∣ₕ)) + (cong (λ x → ∣ a ∣ₕ +[ suc (suc n) ]ₖ x) (rUnitₖ (suc (suc n)) ∣ b ∣) + ∙ sym (rUnitₖ (suc (suc n)) (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ))) + ((λ i → ∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ merid ptS (~ i) ∣ₕ)) + ∙∙ cong (λ x → ∣ a ∣ₕ +[ suc (suc n) ]ₖ x) (rUnitₖ (suc (suc n)) ∣ b ∣) + ∙ sym (rUnitₖ (suc (suc n)) (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ)) + ∙∙ λ i → (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ merid ptS i ∣ₕ) + lem c = + EM-raw-elim G (suc n) + (λ _ → isOfHLevelΠ (2 + n) (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _)) + (EM-raw-elim G (suc n) (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _) + ((sym (rUnit refl) + ◁ λ _ → refl) + ▷ (sym (lCancel (cong ∣_∣ₕ (merid ptS))) + ∙ λ i → (λ j → ∣ merid ptS (~ j ∨ ~ i) ∣ₕ) + ∙∙ lUnit (λ j → ∣ merid ptS (~ j ∧ ~ i) ∣ₕ) i + ∙∙ cong ∣_∣ₕ (merid ptS)))) + main : (c a b : _) → (∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ c ∣ₕ) ≡ (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ c ∣ₕ) + main north a b = lem ptS a b i0 + main south a b = lem ptS a b i1 + main (merid c i) a b = lem c a b i + + σ-EM' : (n : ℕ) (x : EM G (suc n)) + → Path (EM G (suc (suc n))) + (0ₖ (suc (suc n))) + (0ₖ (suc (suc n))) + σ-EM' zero x = cong ∣_∣ₕ (σ-EM zero x) + σ-EM' (suc n) = + trElim (λ _ → isOfHLevelTrunc (5 + n) _ _) + λ x → cong ∣_∣ₕ (σ-EM (suc n) x) + + σ-EM'-0ₖ : (n : ℕ) → σ-EM' n (0ₖ (suc n)) ≡ refl + σ-EM'-0ₖ zero = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + σ-EM'-0ₖ (suc n) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + + private + P : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (r : refl ≡ p) + → lUnit p ∙ cong (_∙ p) r + ≡ rUnit p ∙ cong (p ∙_) r + P p = J (λ p r → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r) refl + + σ-EM'-hom : (n : ℕ) → (a b : _) → σ-EM' n (a +ₖ b) ≡ σ-EM' n a ∙ σ-EM' n b + σ-EM'-hom zero = + wedgeConEM.fun G G 0 0 (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) l r p + where + l : _ + l x = cong (σ-EM' zero) (lUnitₖ 1 x) + ∙∙ lUnit (σ-EM' zero x) + ∙∙ cong (_∙ σ-EM' zero x) (sym (σ-EM'-0ₖ zero)) + + r : _ + r x = + cong (σ-EM' zero) (rUnitₖ 1 x) + ∙∙ rUnit (σ-EM' zero x) + ∙∙ cong (σ-EM' zero x ∙_) (sym (σ-EM'-0ₖ zero)) + + p : _ + p = P (σ-EM' zero embase) (sym (σ-EM'-0ₖ zero)) + σ-EM'-hom (suc n) = + elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) _ _) _ _) + (wedgeConEM.fun G G _ _ + (λ x y → transport (λ i → isOfHLevel (help n i) + ((σ-EM' (suc n) (∣ x ∣ₕ +ₖ ∣ y ∣ₕ) + ≡ σ-EM' (suc n) ∣ x ∣ₕ ∙ σ-EM' (suc n) ∣ y ∣ₕ))) + (isOfHLevelPlus {n = 4 + n} n + (isOfHLevelPath (4 + n) + (isOfHLevelTrunc (5 + n) _ _) _ _))) + (λ x → cong (σ-EM' (suc n)) (lUnitₖ (suc (suc n)) ∣ x ∣) + ∙∙ lUnit (σ-EM' (suc n) ∣ x ∣) + ∙∙ cong (_∙ σ-EM' (suc n) ∣ x ∣) (sym (σ-EM'-0ₖ (suc n)))) + (λ x → cong (σ-EM' (suc n)) (rUnitₖ (2 + n) ∣ x ∣) + ∙∙ rUnit (σ-EM' (suc n) ∣ x ∣) + ∙∙ cong (σ-EM' (suc n) ∣ x ∣ ∙_) (sym (σ-EM'-0ₖ (suc n)))) + (P (σ-EM' (suc n) (0ₖ (2 + n))) (sym (σ-EM'-0ₖ (suc n))))) + + σ-EM'-ₖ : (n : ℕ) → (a : _) → σ-EM' n (-ₖ a) ≡ sym (σ-EM' n a) + σ-EM'-ₖ n = + morphLemmas.distrMinus + (λ x y → x +[ suc n ]ₖ y) (_∙_) + (σ-EM' n) (σ-EM'-hom n) + (0ₖ (suc n)) refl + (λ x → -ₖ x) sym + (λ x → sym (lUnit x)) (λ x → sym (rUnit x)) + (lCancelₖ (suc n)) rCancel + ∙assoc (σ-EM'-0ₖ n) + + -Dist : (n : ℕ) (x y : EM G n) → -[ n ]ₖ (x +[ n ]ₖ y) ≡ (-[ n ]ₖ x) +[ n ]ₖ (-[ n ]ₖ y) + -Dist zero x y = (GroupTheory.invDistr (AbGroup→Group G) x y) ∙ commₖ zero _ _ + -Dist (suc zero) = k + where -- useless where clause. Needed for fast type checking for some reason. + l : _ + l x = refl + + r : _ + r x = cong (λ z → -[ 1 ]ₖ z) (rUnitₖ 1 x) ∙ sym (rUnitₖ 1 (-[ 1 ]ₖ x)) + + p : r ptS ≡ l ptS + p = sym (rUnit refl) + + k = wedgeConEM.fun G G 0 0 (λ _ _ → emsquash _ _) l r (sym p) + + -Dist (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelTruncPath {n = 4 + n}) + (wedgeConEM.fun G G (suc n) (suc n) + (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (hLevHelp n) _ _) + (λ x → refl) + (λ x → cong (λ z → -[ (suc (suc n)) ]ₖ z) (rUnitₖ (suc (suc n)) ∣ x ∣ₕ) ∙ sym (rUnitₖ (suc (suc n)) (-[ (suc (suc n)) ]ₖ ∣ x ∣ₕ))) + (rUnit refl)) + + addIso : (n : ℕ) (x : EM G n) → Iso (EM G n) (EM G n) + Iso.fun (addIso n x) y = y +[ n ]ₖ x + Iso.inv (addIso n x) y = y -[ n ]ₖ x + Iso.rightInv (addIso n x) y = + sym (assocₖ n y (-[ n ]ₖ x) x) + ∙∙ cong (λ x → y +[ n ]ₖ x) (lCancelₖ n x) + ∙∙ rUnitₖ n y + Iso.leftInv (addIso n x) y = + sym (assocₖ n y x (-[ n ]ₖ x)) + ∙∙ cong (λ x → y +[ n ]ₖ x) (rCancelₖ n x) + ∙∙ rUnitₖ n y diff --git a/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda b/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda new file mode 100644 index 000000000..f31f865dc --- /dev/null +++ b/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda @@ -0,0 +1,454 @@ +{-# OPTIONS --cubical --no-import-sorts --safe --experimental-lossy-unification #-} + +module Cubical.Algebra.Group.EilenbergMacLane.Properties where + +open import Cubical.Algebra.Group.EilenbergMacLane.Base +open import Cubical.Algebra.Group.EilenbergMacLane.WedgeConnectivity +open import Cubical.Algebra.Group.EilenbergMacLane.GroupStructure +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport +open import Cubical.Homotopy.Loopspace + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Properties +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec; elim to trElim) +open import Cubical.HITs.EilenbergMacLane1 +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Data.Empty + renaming (rec to ⊥-rec) +open import Cubical.HITs.Truncation + renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) +open import Cubical.Data.Nat hiding (_·_) +open import Cubical.HITs.Susp +open import Cubical.Functions.Morphism +open import Cubical.Foundations.Path + +open import Cubical.Foundations.Pointed.Homogeneous + +private + variable ℓ : Level + + +isConnectedEM₁ : (G : Group ℓ) → isConnected 2 (EM₁ G) +isConnectedEM₁ G = ∣ embase ∣ , h + where + h : (y : hLevelTrunc 2 (EM₁ G)) → ∣ embase ∣ ≡ y + h = trElim (λ y → isOfHLevelSuc 1 (isOfHLevelTrunc 2 ∣ embase ∣ y)) + (elimProp G (λ x → isOfHLevelTrunc 2 ∣ embase ∣ ∣ x ∣) refl) + +module _ {G' : AbGroup ℓ} where + private + G = fst G' + open AbGroupStr (snd G') renaming (_+_ to _+G_) + + + isConnectedEM : (n : ℕ) → isConnected (suc n) (EM G' n) + fst (isConnectedEM zero) = ∣ 0g ∣ + snd (isConnectedEM zero) y = isOfHLevelTrunc 1 _ _ + isConnectedEM (suc zero) = isConnectedEM₁ (AbGroup→Group G') + fst (isConnectedEM (suc (suc n))) = ∣ 0ₖ _ ∣ + snd (isConnectedEM (suc (suc n))) = + trElim (λ _ → isOfHLevelTruncPath) + (trElim (λ _ → isOfHLevelSuc (3 + n) (isOfHLevelTruncPath {n = (3 + n)})) + (EM-raw-elim G' + (suc n) + (λ _ → isOfHLevelTrunc (3 + n) _ _) + refl)) + + +module _ (Ĝ : Group ℓ) where + private + G = fst Ĝ + open GroupStr (snd Ĝ) + + emloop-id : emloop 1g ≡ refl + emloop-id = + emloop 1g ≡⟨ rUnit (emloop 1g) ⟩ + emloop 1g ∙ refl ≡⟨ cong (emloop 1g ∙_) (rCancel (emloop 1g) ⁻¹) ⟩ + emloop 1g ∙ (emloop 1g ∙ (emloop 1g) ⁻¹) ≡⟨ ∙assoc _ _ _ ⟩ + (emloop 1g ∙ emloop 1g) ∙ (emloop 1g) ⁻¹ ≡⟨ cong (_∙ emloop 1g ⁻¹) + ((emloop-comp Ĝ 1g 1g) ⁻¹) ⟩ + emloop (1g · 1g) ∙ (emloop 1g) ⁻¹ ≡⟨ cong (λ g → emloop {Group = Ĝ} g + ∙ (emloop 1g) ⁻¹) + (rid 1g) ⟩ + emloop 1g ∙ (emloop 1g) ⁻¹ ≡⟨ rCancel (emloop 1g) ⟩ + refl ∎ + + emloop-inv : (g : G) → emloop (inv g) ≡ (emloop g) ⁻¹ + emloop-inv g = + emloop (inv g) ≡⟨ rUnit (emloop (inv g)) ⟩ + emloop (inv g) ∙ refl ≡⟨ cong (emloop (inv g) ∙_) + (rCancel (emloop g) ⁻¹) ⟩ + emloop (inv g) ∙ (emloop g ∙ (emloop g) ⁻¹) ≡⟨ ∙assoc _ _ _ ⟩ + (emloop (inv g) ∙ emloop g) ∙ (emloop g) ⁻¹ ≡⟨ cong (_∙ emloop g ⁻¹) + ((emloop-comp Ĝ (inv g) g) ⁻¹) ⟩ + emloop (inv g · g) ∙ (emloop g) ⁻¹ ≡⟨ cong (λ h → emloop {Group = Ĝ} h + ∙ (emloop g) ⁻¹) + (invl g) ⟩ + emloop 1g ∙ (emloop g) ⁻¹ ≡⟨ cong (_∙ (emloop g) ⁻¹) emloop-id ⟩ + refl ∙ (emloop g) ⁻¹ ≡⟨ (lUnit ((emloop g) ⁻¹)) ⁻¹ ⟩ + (emloop g) ⁻¹ ∎ + + isGroupoidEM₁ : isGroupoid (EM₁ Ĝ) + isGroupoidEM₁ = emsquash + + {- since we write composition in diagrammatic order, + and function composition in the other order, + we need right multiplication here -} + rightEquiv : (g : G) → G ≃ G + rightEquiv g = isoToEquiv isom + where + isom : Iso G G + isom .Iso.fun = _· g + isom .Iso.inv = _· inv g + isom .Iso.rightInv h = + (h · inv g) · g ≡⟨ (assoc h (inv g) g) ⁻¹ ⟩ + h · inv g · g ≡⟨ cong (h ·_) (invl g) ⟩ + h · 1g ≡⟨ rid h ⟩ h ∎ + isom .Iso.leftInv h = + (h · g) · inv g ≡⟨ (assoc h g (inv g)) ⁻¹ ⟩ + h · g · inv g ≡⟨ cong (h ·_) (invr g) ⟩ + h · 1g ≡⟨ rid h ⟩ h ∎ + + compRightEquiv : (g h : G) + → compEquiv (rightEquiv g) (rightEquiv h) ≡ rightEquiv (g · h) + compRightEquiv g h = equivEq (funExt (λ x → (assoc x g h) ⁻¹)) + + CodesSet : EM₁ Ĝ → hSet ℓ + CodesSet = rec Ĝ (isOfHLevelTypeOfHLevel 2) (G , is-set) RE REComp + where + RE : (g : G) → Path (hSet ℓ) (G , is-set) (G , is-set) + RE g = Σ≡Prop (λ X → isPropIsOfHLevel {A = X} 2) (ua (rightEquiv g)) + + lemma₁ : (g h : G) → Square + (ua (rightEquiv g)) (ua (rightEquiv (g · h))) + refl (ua (rightEquiv h)) + lemma₁ g h = invEq + (Square≃doubleComp (ua (rightEquiv g)) (ua (rightEquiv (g · h))) + refl (ua (rightEquiv h))) + (ua (rightEquiv g) ∙ ua (rightEquiv h) + ≡⟨ (uaCompEquiv (rightEquiv g) (rightEquiv h)) ⁻¹ ⟩ + ua (compEquiv (rightEquiv g) (rightEquiv h)) + ≡⟨ cong ua (compRightEquiv g h) ⟩ + ua (rightEquiv (g · h)) ∎) + + REComp : (g h : G) → Square (RE g) (RE (g · h)) refl (RE h) + REComp g h = Σ≡Set (λ _ → isProp→isSet isPropIsSet) + (RE g) (RE (g · h)) refl (RE h) (lemma₁ g h) + Codes : EM₁ Ĝ → Type ℓ + Codes x = CodesSet x .fst + + encode : (x : EM₁ Ĝ) → embase ≡ x → Codes x + encode x p = subst Codes p 1g + + decode : (x : EM₁ Ĝ) → Codes x → embase ≡ x + decode = elimSet Ĝ (λ x → isOfHLevelΠ 2 (λ c → isGroupoidEM₁ (embase) x)) + emloop λ g → ua→ λ h → emcomp h g + + decode-encode : (x : EM₁ Ĝ) (p : embase ≡ x) → decode x (encode x p) ≡ p + decode-encode x p = J (λ y q → decode y (encode y q) ≡ q) + (emloop (transport refl 1g) ≡⟨ cong emloop (transportRefl 1g) ⟩ + emloop 1g ≡⟨ emloop-id ⟩ refl ∎) p + + encode-decode : (x : EM₁ Ĝ) (c : Codes x) → encode x (decode x c) ≡ c + encode-decode = elimProp Ĝ (λ x → isOfHLevelΠ 1 (λ c → CodesSet x .snd _ _)) + λ g → encode embase (decode embase g) ≡⟨ refl ⟩ + encode embase (emloop g) ≡⟨ refl ⟩ + transport (ua (rightEquiv g)) 1g ≡⟨ uaβ (rightEquiv g) 1g ⟩ + 1g · g ≡⟨ lid g ⟩ + g ∎ + + ΩEM₁Iso : Iso (Path (EM₁ Ĝ) embase embase) G + Iso.fun ΩEM₁Iso = encode embase + Iso.inv ΩEM₁Iso = emloop + Iso.rightInv ΩEM₁Iso = encode-decode embase + Iso.leftInv ΩEM₁Iso = decode-encode embase + + ΩEM₁≡ : (Path (EM₁ Ĝ) embase embase) ≡ G + ΩEM₁≡ = isoToPath ΩEM₁Iso + +module _ {G : AbGroup ℓ} where + open AbGroupStr (snd G) + renaming (_+_ to _+G_ ; -_ to -G_ ; assoc to assocG) + + CODE : (n : ℕ) → EM G (suc (suc n)) → TypeOfHLevel ℓ (3 + n) + CODE n = + trElim (λ _ → isOfHLevelTypeOfHLevel (3 + n)) + λ { north → EM G (suc n) , hLevelEM G (suc n) + ; south → EM G (suc n) , hLevelEM G (suc n) + ; (merid a i) → fib n a i} + where + fib : (n : ℕ) → (a : EM-raw G (suc n)) + → Path (TypeOfHLevel _ (3 + n)) (EM G (suc n) , hLevelEM G (suc n)) (EM G (suc n) , hLevelEM G (suc n)) + fib zero a = Σ≡Prop (λ _ → isPropIsOfHLevel 3) + (isoToPath (addIso 1 a)) + fib (suc n) a = Σ≡Prop (λ _ → isPropIsOfHLevel (4 + n)) + (isoToPath (addIso (suc (suc n)) ∣ a ∣)) + + decode' : (n : ℕ) (x : EM G (suc (suc n))) → CODE n x .fst → 0ₖ (suc (suc n)) ≡ x + decode' n = + trElim (λ _ → isOfHLevelΠ (4 + n) + λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + λ { north → f n + ; south → g n + ; (merid a i) → main a i} + where + f : (n : ℕ) → _ + f n = σ-EM' n + + g : (n : ℕ) → EM G (suc n) → ∣ ptS ∣ ≡ ∣ south ∣ + g n x = σ-EM' n x ∙ cong ∣_∣ₕ (merid ptS) + + lem₂ : (n : ℕ) (a x : _) + → Path (Path (EM G (suc (suc n))) _ _) + ((λ i → cong ∣_∣ₕ (σ-EM n x) i) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) ∙ (λ i → ∣ merid a i ∣ₕ)) + (g n (EM-raw→EM G (suc n) x)) + lem₂ zero a x = + cong (f zero x ∙_) + (cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid embase))) + ∙ cong-∙ ∣_∣ₕ (merid embase) (sym (merid a))) + ∙∙ sym (∙assoc _ _ _) + ∙∙ cong (cong ∣_∣ₕ (merid embase) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) + ∙ sym (rUnit _)) + lem₂ (suc n) a x = + cong (f (suc n) ∣ x ∣ ∙_) + ((cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid north))) + ∙ cong-∙ ∣_∣ₕ (merid north) (sym (merid a))) + ∙∙ sym (∙assoc _ _ _) + ∙∙ cong (cong ∣_∣ₕ (merid north) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) + ∙ sym (rUnit _))) + + lem : (n : ℕ) (x a : EM-raw G (suc n)) + → f n (transport (sym (cong (λ x → CODE n x .fst) (cong ∣_∣ₕ (merid a)))) (EM-raw→EM G (suc n) x)) + ≡ cong ∣_∣ₕ (σ-EM n x) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) + lem zero x a = (λ i → cong ∣_∣ₕ (merid (transportRefl x i -[ 1 ]ₖ a) ∙ sym (merid embase))) + ∙∙ σ-EM'-hom zero x (-ₖ a) + ∙∙ cong (f zero x ∙_) (σ-EM'-ₖ zero a) + lem (suc n) x a = + cong (f (suc n)) (λ i → transportRefl ∣ x ∣ i -[ 2 + n ]ₖ ∣ a ∣) + ∙∙ σ-EM'-hom (suc n) ∣ x ∣ (-ₖ ∣ a ∣) + ∙∙ cong (f (suc n) ∣ x ∣ ∙_) (σ-EM'-ₖ (suc n) ∣ a ∣) + + main : (a : _) + → PathP (λ i → CODE n ∣ merid a i ∣ₕ .fst + → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ) (f n) (g n) + main a = + toPathP (funExt + (EM-elim _ (λ _ → isOfHLevelPathP (2 + (suc n)) (isOfHLevelTrunc (4 + n) _ _) _ _) + λ x → + ((λ i → transp (λ j → Path (EM G (2 + n)) ∣ ptS ∣ ∣ merid a (i ∨ j) ∣) + i (compPath-filler (lem n x a i) (cong ∣_∣ₕ (merid a)) i) )) + ∙∙ sym (∙assoc _ _ _) + ∙∙ lem₂ n a x)) + + encode' : (n : ℕ) (x : EM G (suc (suc n))) → 0ₖ (suc (suc n)) ≡ x → CODE n x .fst + encode' n x p = subst (λ x → CODE n x .fst) p (0ₖ (suc n)) + + decode'-encode' : (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → decode' n x (encode' n x p) ≡ p + decode'-encode' zero x = + J (λ x p → decode' 0 x (encode' 0 x p) ≡ p) + (σ-EM'-0ₖ 0) + decode'-encode' (suc n) x = + J (λ x p → decode' (suc n) x (encode' (suc n) x p) ≡ p) + (σ-EM'-0ₖ (suc n)) + + encode'-decode' : (n : ℕ) (x : _) → encode' n (0ₖ (suc (suc n))) (decode' n (0ₖ (suc (suc n))) x) ≡ x + encode'-decode' zero x = + cong (encode' zero (0ₖ 2)) (cong-∙ ∣_∣ₕ (merid x) (sym (merid embase))) + ∙∙ substComposite (λ x → CODE zero x .fst) + (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptS))) embase + ∙∙ (cong (subst (λ x₁ → CODE zero x₁ .fst) (λ i → ∣ merid embase (~ i) ∣ₕ)) + (λ i → lUnitₖ 1 (transportRefl x i) i) + ∙ (λ i → rUnitₖ 1 (transportRefl x i) i)) + encode'-decode' (suc n) = + trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ x → cong (encode' (suc n) (0ₖ (3 + n))) (cong-∙ ∣_∣ₕ (merid x) (sym (merid north))) + ∙∙ substComposite (λ x → CODE (suc n) x .fst) + (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptS))) (0ₖ (2 + n)) + ∙∙ cong (subst (λ x₁ → CODE (suc n) x₁ .fst) (λ i → ∣ merid ptS (~ i) ∣ₕ)) + (λ i → lUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) + ∙ (λ i → rUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) + + EM∙' : (n : ℕ) → Pointed _ + EM∙' n = EM G n , 0ₖ n + + Iso-EM-ΩEM+1 : (n : ℕ) → Iso (EM G n) (typ (Ω (EM∙' (suc n)))) + Iso-EM-ΩEM+1 zero = invIso (ΩEM₁Iso (AbGroup→Group G)) + Iso.fun (Iso-EM-ΩEM+1 (suc zero)) = decode' 0 (0ₖ 2) + Iso.inv (Iso-EM-ΩEM+1 (suc zero)) = encode' 0 (0ₖ 2) + Iso.rightInv (Iso-EM-ΩEM+1 (suc zero)) = decode'-encode' 0 (0ₖ 2) + Iso.leftInv (Iso-EM-ΩEM+1 (suc zero)) = encode'-decode' 0 + Iso.fun (Iso-EM-ΩEM+1 (suc (suc n))) = decode' (suc n) (0ₖ (3 + n)) + Iso.inv (Iso-EM-ΩEM+1 (suc (suc n))) = encode' (suc n) (0ₖ (3 + n)) + Iso.rightInv (Iso-EM-ΩEM+1 (suc (suc n))) = decode'-encode' (suc n) (0ₖ (3 + n)) + Iso.leftInv (Iso-EM-ΩEM+1 (suc (suc n))) = encode'-decode' (suc n) + + EM→ΩEM+1 : (n : ℕ) → EM G n → typ (Ω (EM∙' (suc n))) + EM→ΩEM+1 n = Iso.fun (Iso-EM-ΩEM+1 n) + + ΩEM+1→EM : (n : ℕ) → typ (Ω (EM∙' (suc n))) → EM G n + ΩEM+1→EM n = Iso.inv (Iso-EM-ΩEM+1 n) + + EM→ΩEM+1-hom : (n : ℕ) (x y : EM G n) + → EM→ΩEM+1 n (x +[ n ]ₖ y) ≡ EM→ΩEM+1 n x ∙ EM→ΩEM+1 n y + EM→ΩEM+1-hom zero x y = emloop-comp _ x y + EM→ΩEM+1-hom (suc zero) x y = σ-EM'-hom zero x y + EM→ΩEM+1-hom (suc (suc n)) x y = σ-EM'-hom (suc n) x y + + ΩEM+1→EM-hom : (n : ℕ) → (p q : typ (Ω (EM∙' (suc n)))) + → ΩEM+1→EM n (p ∙ q) ≡ (ΩEM+1→EM n p) +[ n ]ₖ (ΩEM+1→EM n q) + ΩEM+1→EM-hom n = + morphLemmas.isMorphInv (λ x y → x +[ n ]ₖ y) (_∙_) (EM→ΩEM+1 n) + (EM→ΩEM+1-hom n) (ΩEM+1→EM n) + (Iso.rightInv (Iso-EM-ΩEM+1 n)) (Iso.leftInv (Iso-EM-ΩEM+1 n)) + + EM→ΩEM+1-0ₖ : (n : ℕ) → EM→ΩEM+1 n (0ₖ n) ≡ refl + EM→ΩEM+1-0ₖ zero = emloop-1g _ + EM→ΩEM+1-0ₖ (suc zero) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + EM→ΩEM+1-0ₖ (suc (suc n)) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + + ΩEM+1→EM-refl : (n : ℕ) → ΩEM+1→EM n refl ≡ 0ₖ n + ΩEM+1→EM-refl zero = transportRefl 0g + ΩEM+1→EM-refl (suc zero) = refl + ΩEM+1→EM-refl (suc (suc n)) = refl + + isHomogeneousEM : (n : ℕ) → isHomogeneous (EM∙ G n) + isHomogeneousEM n x = + ua∙ (isoToEquiv (addIso n x)) (lUnitₖ n x) + + isContr-↓∙ : (n : ℕ) → isContr (EM∙ G (suc n) →∙ EM∙ G n) + fst (isContr-↓∙ zero) = (λ _ → 0g) , refl + snd (isContr-↓∙ zero) (f , p) = + Σ≡Prop (λ x → is-set _ _) + (funExt (EM-raw-elim G 0 (λ _ → is-set _ _) + (sym p))) + fst (isContr-↓∙ (suc n)) = (λ _ → 0ₖ (suc n)) , refl + snd (isContr-↓∙ (suc n)) (f , p) = + →∙Homogeneous≡ (isHomogeneousEM _) + (funExt + (trElim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (hLevelEM G (suc n))) _ _) + (EM-raw-elim _ _ (λ _ → hLevelEM G (suc n) _ _) + (sym p)))) + + +{- +-- HLevel stuff for cup product +isContr-↓∙ : (n : ℕ) → isContr (coHomK-ptd (suc n) →∙ coHomK-ptd n) +fst (isContr-↓∙ zero) = (λ _ → 0) , refl +snd (isContr-↓∙ zero) (f , p) = + Σ≡Prop (λ f → isSetℤ _ _) + (funExt (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 isSetℤ) _ _) + (toPropElim (λ _ → isSetℤ _ _) (sym p)))) +fst (isContr-↓∙ (suc n)) = (λ _ → 0ₖ _) , refl +fst (snd (isContr-↓∙ (suc n)) f i) x = + trElim {B = λ x → 0ₖ (suc n) ≡ fst f x} + (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (isOfHLevelTrunc (3 + n))) _ _) + (sphereElim _ (λ _ → isOfHLevelTrunc (3 + n) _ _) + (sym (snd f))) x i +snd (snd (isContr-↓∙ (suc n)) f i) j = snd f (~ i ∨ j) + +isContr-↓∙' : (n : ℕ) → isContr (S₊∙ (suc n) →∙ coHomK-ptd n) +fst (isContr-↓∙' zero) = (λ _ → 0) , refl +snd (isContr-↓∙' zero) (f , p) = + Σ≡Prop (λ f → isSetℤ _ _) + (funExt (toPropElim (λ _ → isSetℤ _ _) (sym p))) +fst (isContr-↓∙' (suc n)) = (λ _ → 0ₖ _) , refl +fst (snd (isContr-↓∙' (suc n)) f i) x = + sphereElim _ {A = λ x → 0ₖ (suc n) ≡ fst f x} + (λ _ → isOfHLevelTrunc (3 + n) _ _) + (sym (snd f)) x i +snd (snd (isContr-↓∙' (suc n)) f i) j = snd f (~ i ∨ j) + +isOfHLevel→∙Kn : {A : Pointed₀} (n m : ℕ) + → isOfHLevel (suc m) (A →∙ coHomK-ptd n) → isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) +isOfHLevel→∙Kn {A = A} n m hlev = step₃ + where + step₁ : isOfHLevel (suc m) (A →∙ Ω (coHomK-ptd (suc n))) + step₁ = + subst (isOfHLevel (suc m)) + (λ i → A →∙ ua∙ {A = Ω (coHomK-ptd (suc n))} {B = coHomK-ptd n} + (invEquiv Kn≃ΩKn+1) + (ΩKn+1→Kn-refl n) (~ i)) hlev + + step₂ : isOfHLevel (suc m) (typ (Ω (A →∙ coHomK-ptd (suc n) ∙))) + step₂ = isOfHLevelRetractFromIso (suc m) (invIso (invIso (ΩfunExtIso _ _))) step₁ + + step₃ : isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) + step₃ = + isOfHLevelΩ→isOfHLevel m + λ f → subst (λ x → isOfHLevel (suc m) (typ (Ω x))) + (isHomogeneous→∙ (isHomogeneousKn (suc n)) f) + step₂ + +isOfHLevel↑∙ : ∀ n m → isOfHLevel (2 + n) (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m))) +isOfHLevel↑∙ zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙ m)) +isOfHLevel↑∙ (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙ n m) + +isOfHLevel↑∙' : ∀ n m → isOfHLevel (2 + n) (S₊∙ (suc m) →∙ coHomK-ptd (suc (n + m))) +isOfHLevel↑∙' zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙' m)) +isOfHLevel↑∙' (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙' n m) + +→∙KnPath : (A : Pointed₀) (n : ℕ) → Ω (A →∙ coHomK-ptd (suc n) ∙) ≡ (A →∙ coHomK-ptd n ∙) +→∙KnPath A n = + ua∙ (isoToEquiv (ΩfunExtIso A (coHomK-ptd (suc n)))) refl + ∙ (λ i → (A →∙ Kn≃ΩKn+1∙ {n = n} (~ i) ∙)) + +contr∙-lem : (n m : ℕ) + → isContr (coHomK-ptd (suc n) + →∙ (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (n + m)) ∙)) +fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl) , refl +snd (contr∙-lem n m) (f , p) = →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousKn _)) (sym (funExt help)) + where + help : (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) + help = + trElim (λ _ → isOfHLevelPath (3 + n) + (subst (λ x → isOfHLevel x (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m)))) (λ i → suc (suc (+-comm n 1 i))) + (isOfHLevelPlus' {n = 1} (2 + n) (isOfHLevel↑∙ n m))) _ _) + (sphereElim _ (λ _ → isOfHLevel↑∙ n m _ _) p) + +isOfHLevel↑∙∙ : ∀ n m l + → isOfHLevel (2 + l) (coHomK-ptd (suc n) + →∙ (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (suc (l + n + m))) ∙)) +isOfHLevel↑∙∙ n m zero = + isOfHLevelΩ→isOfHLevel 0 λ f + → subst + isProp (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) + (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) h) + where + h : isProp (coHomK-ptd (suc n) + →∙ (Ω (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (suc (n + m))) ∙))) + h = + subst isProp (λ i → coHomK-ptd (suc n) →∙ (→∙KnPath (coHomK-ptd (suc m)) (suc (n + m)) (~ i))) + (isContr→isProp (contr∙-lem n m)) +isOfHLevel↑∙∙ n m (suc l) = + isOfHLevelΩ→isOfHLevel (suc l) + λ f → + subst + (isOfHLevel (2 + l)) (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) + (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) h) + where + h : isOfHLevel (2 + l) + (coHomK-ptd (suc n) + →∙ (Ω (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (suc (suc (l + n + m)))) ∙))) + h = + subst (isOfHLevel (2 + l)) + (λ i → coHomK-ptd (suc n) →∙ →∙KnPath (coHomK-ptd (suc m)) (suc (suc (l + n + m))) (~ i)) + (isOfHLevel↑∙∙ n m l) +-} diff --git a/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda b/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda new file mode 100644 index 000000000..945df9935 --- /dev/null +++ b/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda @@ -0,0 +1,287 @@ +{-# OPTIONS --cubical --no-import-sorts --safe --experimental-lossy-unification #-} + +module Cubical.Algebra.Group.EilenbergMacLane.WedgeConnectivity where + +open import Cubical.Algebra.Group.EilenbergMacLane.Base +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Properties +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec; elim to trElim) +open import Cubical.HITs.EilenbergMacLane1 +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Data.Empty + renaming (rec to ⊥-rec) +open import Cubical.HITs.Truncation + renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) +open import Cubical.Data.Nat hiding (_·_) +open import Cubical.HITs.Susp +open import Cubical.Functions.Morphism +open import Cubical.Foundations.Path + +private + variable ℓ : Level + + wedgeConFun' : (G H : AbGroup ℓ) (n : ℕ) + → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ} + → ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y)) + → (f : (x : _) → A ptS x) + → (g : (x : _) → A x ptS) + → f ptS ≡ g ptS + → (x : _) (y : _) → A x y + wedgeConFun'ᵣ : (G H : AbGroup ℓ) (n : ℕ) + → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ} + → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y))) + → (f : (x : _) → A ptS x) + → (g : (x : _) → A x ptS) + → (p : f ptS ≡ g ptS) + → (x : _) → wedgeConFun' G H n hLev f g p x ptS ≡ g x + wedgeConFun' G H zero {A = A} hlev f g p = + elimSet _ (λ _ → isSetΠ λ _ → hlev _ _) f k + where + I→A : (x : fst G) → (i : I) → A (emloop x i) embase + I→A x i = + hcomp (λ k → λ {(i = i0) → p (~ k) ; (i = i1) → p (~ k)}) + (g (emloop x i)) + + SquareP2 : (x : _) (z : _) + → SquareP (λ i j → A (emloop x i) (emloop z j)) + (cong f (emloop z)) (cong f (emloop z)) + (λ i → I→A x i) λ i → I→A x i + SquareP2 x z = + toPathP (isOfHLevelPathP' 1 (λ _ _ → hlev _ _ _ _) _ _ _ _) + + CubeP2 : (x : _) (g h : _) + → PathP (λ k → PathP (λ j → PathP (λ i → A (emloop x i) (emcomp g h j k)) + (f (emcomp g h j k)) (f (emcomp g h j k))) + (λ i → SquareP2 x g i k) λ i → SquareP2 x ((snd (AbGroup→Group H) GroupStr.· g) h) i k) + (λ _ i → I→A x i) λ j i → SquareP2 x h i j + CubeP2 x g h = toPathP (isOfHLevelPathP' 1 (isOfHLevelPathP 2 (hlev _ _) _ _) _ _ _ _) + + k : (x : _) → PathP (λ i → (y : EM₁ (AbGroup→Group H)) → A (emloop x i) y) f f + k x i embase = I→A x i + k x i (emloop z j) = SquareP2 x z i j + k x i (emcomp g h j k) = CubeP2 x g h k j i + k x i (emsquash y z p q r s j k' l) = mega i j k' l + where + mega : + PathP (λ i → + PathP (λ j → + PathP (λ k' → + PathP (λ l → A (emloop x i) (emsquash y z p q r s j k' l)) + (k x i y) + (k x i z)) + (λ l → k x i (p l)) + λ l → k x i (q l)) + (λ k' l → k x i (r k' l)) + λ k' l → k x i (s k' l)) + (λ j k l → f (emsquash y z p q r s j k l)) + λ j k l → f (emsquash y z p q r s j k l) + mega = + toPathP (isOfHLevelPathP' 1 + (isOfHLevelPathP 2 (isOfHLevelPathP 2 (hlev _ _) _ _) _ _) _ _ _ _) + wedgeConFun' G H (suc n) {A = A} hLev f g p north y = f y + wedgeConFun' G H (suc n) {A = A} hLev f g p south y = subst (λ x → A x y) (merid ptS) (f y) + wedgeConFun' G H (suc n) {A = A} hLev f g p (merid a i) y = mainₗ a y i + module _ where + llem₁ : g south ≡ subst (λ x₁ → A x₁ ptS) (merid ptS) (f ptS) + llem₁ = (λ i → transp (λ j → A (merid ptS (j ∨ ~ i)) ptS) (~ i) (g (merid ptS (~ i)))) + ∙ cong (subst (λ x₁ → A x₁ ptS) (merid ptS)) (sym p) + + llem₂' : + Square + (λ i → transp (λ j → A (merid ptS (i ∨ j)) ptS) i (g (merid ptS i))) refl + (cong (subst (λ x → A x ptS) (merid ptS)) (sym p)) llem₁ + llem₂' i j = + hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptS (j ∨ z)) ptS) j + (g (merid ptS j)) + ; (i = i1) → subst (λ x₁ → A x₁ ptS) (merid ptS) (p (~ k)) + ; (j = i0) → (subst (λ x → A x ptS) (merid ptS)) (p (~ k ∨ ~ i))}) + (transp (λ k → A (merid ptS (k ∨ ~ i ∧ j)) ptS) (~ i ∧ j) (g (merid ptS (~ i ∧ j)))) + + llem₂ : (λ i₁ → transp (λ j → A (merid ptS (i₁ ∧ j)) ptS) (~ i₁) (f ptS)) + ≡ (λ i₁ → hcomp (λ k → λ { (i₁ = i0) → p (~ k) ; (i₁ = i1) → llem₁ k }) + (g (merid ptS i₁))) + llem₂ i j = + hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptS (j ∧ j₁)) ptS) (~ j) (p (~ k)) + ; (j = i0) → p (~ k) + ; (j = i1) → llem₂' k i}) + (transp (λ k → A (merid ptS ((i ∨ k) ∧ j)) ptS) (i ∨ ~ j) (g (merid ptS (i ∧ j)))) + + mainₗ : (a : _) (y : _) + → PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) + mainₗ = + wedgeConFun' G H n + (λ _ _ → isOfHLevelPathP' (suc (n + 1)) (hLev _ _) _ _) + (λ x i → transp (λ j → A (merid ptS (i ∧ j)) x) (~ i) (f x)) + (λ x i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → llem₁ k}) + (g (merid x i))) + llem₂ + + mainP : (y : _) + → mainₗ y ptS + ≡ λ i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → llem₁ k}) + (g (merid y i)) + mainP y = + wedgeConFun'ᵣ G H n + (λ _ _ → isOfHLevelPathP' (suc (n + 1)) (hLev _ _) _ _) + (λ x i → transp (λ j → A (merid ptS (i ∧ j)) x) (~ i) (f x)) + (λ x i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → llem₁ k}) + (g (merid x i))) + llem₂ y + wedgeConFun'ᵣ G H zero {A = A} hLev f g p = + elimProp _ (λ _ → hLev _ _ _ _) p + wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p north = p + wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p south = sym (llem₁ G H n hLev f g p ptS i0 ptS) + wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p (merid a i) k = P k i + where + lem : _ + lem i j = + hcomp (λ k → λ { (i = i1) → g (merid a j) + ; (j = i0) → p (i ∨ ~ k) + ; (j = i1) → llem₁ G H n hLev f g p ptS i0 ptS (~ i ∧ k)}) + (g (merid a j)) + + P : PathP (λ k → PathP (λ i → A (merid a i) ptS) + (p k) (llem₁ G H n hLev f g p ptS i0 ptS (~ k))) + (λ i → mainₗ G H n hLev f g p a i ptS a ptS i) λ i → g (merid a i) + P = mainP G H n hLev f g p a i0 ptS a ◁ lem + +private + wedgeConFun : (G H : AbGroup ℓ) (k n m : ℕ) → (n + m ≡ k) → {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} + → ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y)) + → (f : (x : _) → A ptS x) + → (g : (x : _) → A x ptS) + → f ptS ≡ g ptS + → (x : _) (y : _) → A x y + wedgeconLeft : (G H : AbGroup ℓ) (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} + → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) + → (f : (x : _) → A ptS x) + → (g : (x : _) → A x ptS) + → (p : f ptS ≡ g ptS) + → (x : _) → wedgeConFun G H k n m P hLev f g p ptS x ≡ f x + wedgeconRight : (G H : AbGroup ℓ) (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} + → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) + → (f : (x : _) → A ptS x) + → (g : (x : _) → A x ptS) + → (p : f ptS ≡ g ptS) + → (x : _) → wedgeConFun G H k n m P hLev f g p x ptS ≡ g x + wedgeConFun G H k n zero P {A = A} hLev f g p = wedgeConFun' G H n hLev f g p + wedgeConFun G H k zero (suc m) P {A = A} hLev f g p x y = + wedgeConFun' H G (suc m) {A = λ x y → A y x} + (λ x y → subst (λ n → isOfHLevel (2 + n) (A y x)) (+-comm 1 m) (hLev y x)) + g f (sym p) y x + wedgeConFun G H l (suc n) (suc m) P {A = A} hlev f g p north y = f y + wedgeConFun G H l (suc n) (suc m) P {A = A} hlev f g p south y = subst (λ x → A x y) (merid ptS) (f y) + wedgeConFun G H zero (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = main** i + module _ where + main** : PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) + main** = ⊥-rec (snotz P) + wedgeConFun G H (suc l) (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = main* a y i + module _ where + lem₁* : g south ≡ (subst (λ x → A x ptS) (merid ptS) (f ptS)) + lem₁* = ((λ i → transp (λ j → A (merid ptS (~ i ∨ j)) ptS) (~ i) (g (merid ptS (~ i))))) + ∙ cong (subst (λ x → A x ptS) (merid ptS)) (sym p) + + lem₁'* : + Square + (λ i → transp (λ j → A (merid ptS (~ i ∨ j)) ptS) (~ i) (g (merid ptS (~ i)))) + (refl {x = subst (λ x → A x ptS) (merid ptS) (f ptS)}) + lem₁* + ((cong (transport (λ z → A (merid ptS z) ptS)) (sym p))) + lem₁'* i j = + hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptS (~ j ∨ j₁)) ptS) (~ j) (g (merid ptS (~ j))) + ; (i = i1) → subst (λ x → A x ptS) (merid ptS) (p (~ k)) + ; (j = i1) → cong (transport (λ z → A (merid ptS z) ptS)) (sym p) (i ∧ k)}) + (transp (λ j₁ → A (merid ptS ((~ j ∧ ~ i) ∨ j₁)) ptS) (~ j ∧ ~ i) (g (merid ptS (~ j ∧ ~ i)))) + + + lem₂* : (λ i₁ → transp (λ j → A (merid ptS (j ∧ i₁)) ptS) (~ i₁) (f ptS)) + ≡ (λ i₁ → hcomp (λ k → λ { (i₁ = i0) → p (~ k) ; (i₁ = i1) → lem₁* k }) + (g (merid ptS i₁))) + lem₂* i j = + hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptS (z ∧ j)) ptS) (~ j) (p (~ k)) + ; (j = i0) → p (~ k) + ; (j = i1) → lem₁'* k (~ i)}) + (transp (λ z → A (merid ptS ((i ∨ z) ∧ j)) ptS) (i ∨ ~ j) (g (merid ptS (i ∧ j)))) + + main* : (a : _) (y : _) → PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) + main* = + wedgeConFun G H l n (suc m) (cong predℕ P) + (λ _ _ → isOfHLevelPathP' (suc (n + (suc (suc m)))) (hlev _ _) _ _) + (λ x i → transp (λ j → A (merid ptS (j ∧ i)) x) (~ i) (f x)) + (λ y i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → lem₁* k}) + (g (merid y i))) + lem₂* + + mainR : (x : _) → main* x ptS + ≡ λ i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → lem₁* k}) + (g (merid x i)) + mainR x = + wedgeconRight G H l n (suc m) (cong predℕ P) + (λ _ _ → isOfHLevelPathP' (suc (n + (suc (suc m)))) (hlev _ _) _ _) + (λ x i → transp (λ j → A (merid ptS (j ∧ i)) x) (~ i) (f x)) + (λ y i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → lem₁* k}) + (g (merid y i))) + lem₂* x + wedgeconLeft G H k zero zero P {A = A} hLev f g p _ = refl + wedgeconLeft G H k (suc n) zero P {A = A} hLev f g p _ = refl + wedgeconLeft G H k zero (suc m) P {A = A} hLev f g p x = + wedgeConFun'ᵣ H G (suc m) + (λ x₁ y → + subst (λ n → (x₂ y₁ : A y x₁) → isOfHLevel (suc n) (x₂ ≡ y₁)) + (+-comm 1 m) (hLev y x₁)) + g f (λ i → p (~ i)) x + wedgeconLeft G H k (suc n) (suc m) P {A = A} hLev f g p _ = refl + wedgeconRight G H k n zero P {A = A} hLev f g p = wedgeConFun'ᵣ G H n hLev f g p + wedgeconRight G H k zero (suc m) P {A = A} hLev f g p _ = refl + wedgeconRight G H zero (suc n) (suc m) P {A = A} hLev f g p x = ⊥-rec (snotz P) + wedgeconRight G H l (suc n) (suc m) P {A = A} hLev f g p north = p + wedgeconRight G H l (suc n) (suc m) P {A = A} hLev f g p south = + sym (lem₁* G H _ n m refl hLev f g p ptS i0 ptS) + wedgeconRight G H (suc l) (suc n) (suc m) P {A = A} hLev f g p (merid a i) k = + help k i + where + lem : _ + lem i j = + hcomp (λ k → λ { (i = i1) → g (merid a j) + ; (j = i0) → p (i ∨ ~ k) + ; (j = i1) → lem₁* G H l n m P hLev f g p ptS i0 ptS (~ i ∧ k)}) + (g (merid a j)) + + help : PathP (λ k → PathP (λ i → A (merid a i) ptS) + (p k) (lem₁* G H l n m P hLev f g p ptS i0 ptS (~ k))) + (λ i → main* G H l n m P hLev f g p a i north a north i) (cong g (merid a)) + help = mainR G H l n m P hLev f g p a i0 ptS a ◁ lem + +module wedgeConEM (G H : AbGroup ℓ) (n m : ℕ) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} + (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) + (f : (x : _) → A ptS x) + (g : (x : _) → A x ptS) + (p : f ptS ≡ g ptS) where + fun : (x : EM-raw G (suc n)) (y : EM-raw H (suc m)) → A x y + fun = wedgeConFun G H (n + m) n m refl hLev f g p + + left : (x : EM-raw H (suc m)) → fun ptS x ≡ f x + left = wedgeconLeft G H (n + m) n m refl hLev f g p + + right : (x : EM-raw G (suc n)) → fun x ptS ≡ g x + right = wedgeconRight G H (n + m) n m refl hLev f g p diff --git a/Cubical/Algebra/Group/EilenbergMacLane1.agda b/Cubical/Algebra/Group/EilenbergMacLane1.agda index 93f0566b5..fd97d5c5e 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane1.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane1.agda @@ -10,6 +10,9 @@ open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) open import Cubical.Foundations.Path open import Cubical.Foundations.HLevels open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport + open import Cubical.Data.Unit open import Cubical.Data.Sigma open import Cubical.Algebra.Group.Base @@ -17,6 +20,15 @@ open import Cubical.Algebra.Group.Properties open import Cubical.Homotopy.Connected open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec; elim to trElim) open import Cubical.HITs.EilenbergMacLane1 +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Data.Empty + renaming (rec to ⊥-rec) +open import Cubical.HITs.Truncation + renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) +open import Cubical.Data.Nat hiding (_·_) +open import Cubical.HITs.Susp +open import Cubical.Functions.Morphism +open import Cubical.Foundations.Path private variable ℓ : Level @@ -149,3 +161,835 @@ module _ (Ĝ : Group ℓ) where ΩEM₁≡ : (Path (EM₁ Ĝ) embase embase) ≡ G ΩEM₁≡ = isoToPath ΩEM₁Iso + +private + variable + ℓG : Level + + +Susp̂ : (n : ℕ) → Type ℓG → Type ℓG +Susp̂ zero A = A +Susp̂ (suc n) A = Susp (Susp̂ n A) + +pt-Susp̂ : (n : ℕ) (A : Pointed ℓG) → Susp̂ n (typ A) +pt-Susp̂ zero A = snd A +pt-Susp̂ (suc n) A = north + + + +_* = AbGroup→Group + +EM-raw : (G : AbGroup ℓG) (n : ℕ) → Type ℓG +EM-raw G zero = fst G +EM-raw G (suc zero) = EM₁ (G *) +EM-raw G (suc (suc n)) = Susp (EM-raw G (suc n)) + +ptS : {n : ℕ} {G : AbGroup ℓG} → EM-raw G n +ptS {n = zero} {G = G} = AbGroupStr.0g (snd G) +ptS {n = suc zero} {G = G} = embase +ptS {n = suc (suc n)} {G = G} = north + +EM-raw-elim : (G : AbGroup ℓG) (n : ℕ) {A : EM-raw G (suc n) → Type ℓ} + → ((x : _) → isOfHLevel (suc n) (A x) ) + → A ptS + → (x : _) → A x +EM-raw-elim G zero hlev b = elimProp _ hlev b +EM-raw-elim G (suc n) hlev b north = b +EM-raw-elim G (suc n) {A = A} hlev b south = subst A (merid ptS) b +EM-raw-elim G (suc n) {A = A} hlev b (merid a i) = help a i + where + help : (a : _) → PathP (λ i → A (merid a i)) b (subst A (merid ptS) b) + help = EM-raw-elim G n (λ _ → isOfHLevelPathP' (suc n) (hlev _) _ _) + λ i → transp (λ j → A (merid ptS (j ∧ i))) (~ i) b + +EM : (G : AbGroup ℓG) (n : ℕ) → Type ℓG +EM G zero = EM-raw G zero +EM G (suc zero) = EM-raw G 1 +EM G (suc (suc n)) = hLevelTrunc (4 + n) (EM-raw G (suc (suc n))) + +0ₖ : {G : AbGroup ℓG} (n : ℕ) → EM G n +0ₖ {G = G} zero = AbGroupStr.0g (snd G) +0ₖ (suc zero) = embase +0ₖ (suc (suc n)) = ∣ ptS ∣ + +EM∙ : (G : AbGroup ℓG) (n : ℕ) → Pointed ℓG +EM∙ G n = EM G n , (0ₖ n) + +hLevelEM : (G : AbGroup ℓG) (n : ℕ) → isOfHLevel (2 + n) (EM G n) +hLevelEM G zero = AbGroupStr.is-set (snd G) +hLevelEM G (suc zero) = emsquash +hLevelEM G (suc (suc n)) = isOfHLevelTrunc (4 + n) + +EM-raw→EM : (G : AbGroup ℓG) (n : ℕ) → EM-raw G n → EM G n +EM-raw→EM G zero x = x +EM-raw→EM G (suc zero) x = x +EM-raw→EM G (suc (suc n)) = ∣_∣ + +EM-elim : {G : AbGroup ℓG} (n : ℕ) {A : EM G n → Type ℓ} + → ((x : _) → isOfHLevel (2 + n) (A x)) + → ((x : EM-raw G n) → A (EM-raw→EM G n x)) + → (x : _) → A x +EM-elim zero hlev hyp x = hyp x +EM-elim (suc zero) hlev hyp x = hyp x +EM-elim (suc (suc n)) hlev hyp = trElim (λ _ → hlev _) hyp + +wedgeConFun' : (G H : AbGroup ℓG) (n : ℕ) → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ} + → ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y)) + → (f : (x : _) → A ptS x) + → (g : (x : _) → A x ptS) + → f ptS ≡ g ptS + → (x : _) (y : _) → A x y +wedgeConFun'ᵣ : (G H : AbGroup ℓG) (n : ℕ) → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ} + → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y))) + → (f : (x : _) → A ptS x) + → (g : (x : _) → A x ptS) + → (p : f ptS ≡ g ptS) + → (x : _) → wedgeConFun' G H n hLev f g p x ptS ≡ g x +wedgeConFun' G H zero {A = A} hlev f g p = + elimSet _ (λ _ → isSetΠ λ _ → hlev _ _) f k + where + I→A : (x : fst G) → (i : I) → A (emloop x i) embase + I→A x i = + hcomp (λ k → λ {(i = i0) → p (~ k) ; (i = i1) → p (~ k)}) + (g (emloop x i)) + + SquareP2 : (x : _) (z : _) + → SquareP (λ i j → A (emloop x i) (emloop z j)) + (cong f (emloop z)) (cong f (emloop z)) + (λ i → I→A x i) λ i → I→A x i + SquareP2 x z = + toPathP (isOfHLevelPathP' 1 (λ _ _ → hlev _ _ _ _) _ _ _ _) + + CubeP2 : (x : _) (g h : _) + → PathP (λ k → PathP (λ j → PathP (λ i → A (emloop x i) (emcomp g h j k)) + (f (emcomp g h j k)) (f (emcomp g h j k))) + (λ i → SquareP2 x g i k) λ i → SquareP2 x ((snd (AbGroup→Group H) GroupStr.· g) h) i k) + (λ _ i → I→A x i) λ j i → SquareP2 x h i j + CubeP2 x g h = toPathP (isOfHLevelPathP' 1 (isOfHLevelPathP 2 (hlev _ _) _ _) _ _ _ _) + + k : (x : _) → PathP (λ i → (y : EM₁ (AbGroup→Group H)) → A (emloop x i) y) f f + k x i embase = I→A x i + k x i (emloop z j) = SquareP2 x z i j + k x i (emcomp g h j k) = CubeP2 x g h k j i + k x i (emsquash y z p q r s j k' l) = mega i j k' l + where + mega : + PathP (λ i → + PathP (λ j → + PathP (λ k' → + PathP (λ l → A (emloop x i) (emsquash y z p q r s j k' l)) + (k x i y) + (k x i z)) + (λ l → k x i (p l)) + λ l → k x i (q l)) + (λ k' l → k x i (r k' l)) + λ k' l → k x i (s k' l)) + (λ j k l → f (emsquash y z p q r s j k l)) + λ j k l → f (emsquash y z p q r s j k l) + mega = + toPathP (isOfHLevelPathP' 1 + (isOfHLevelPathP 2 (isOfHLevelPathP 2 (hlev _ _) _ _) _ _) _ _ _ _) +wedgeConFun' G H (suc n) {A = A} hLev f g p north y = f y +wedgeConFun' G H (suc n) {A = A} hLev f g p south y = subst (λ x → A x y) (merid ptS) (f y) +wedgeConFun' G H (suc n) {A = A} hLev f g p (merid a i) y = mainₗ a y i + module _ where + llem₁ : g south ≡ subst (λ x₁ → A x₁ ptS) (merid ptS) (f ptS) + llem₁ = (λ i → transp (λ j → A (merid ptS (j ∨ ~ i)) ptS) (~ i) (g (merid ptS (~ i)))) + ∙ cong (subst (λ x₁ → A x₁ ptS) (merid ptS)) (sym p) + + llem₂' : + Square + (λ i → transp (λ j → A (merid ptS (i ∨ j)) ptS) i (g (merid ptS i))) refl + (cong (subst (λ x → A x ptS) (merid ptS)) (sym p)) llem₁ + llem₂' i j = + hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptS (j ∨ z)) ptS) j + (g (merid ptS j)) + ; (i = i1) → subst (λ x₁ → A x₁ ptS) (merid ptS) (p (~ k)) + ; (j = i0) → (subst (λ x → A x ptS) (merid ptS)) (p (~ k ∨ ~ i))}) + (transp (λ k → A (merid ptS (k ∨ ~ i ∧ j)) ptS) (~ i ∧ j) (g (merid ptS (~ i ∧ j)))) + + llem₂ : (λ i₁ → transp (λ j → A (merid ptS (i₁ ∧ j)) ptS) (~ i₁) (f ptS)) + ≡ (λ i₁ → hcomp (λ k → λ { (i₁ = i0) → p (~ k) ; (i₁ = i1) → llem₁ k }) + (g (merid ptS i₁))) + llem₂ i j = + hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptS (j ∧ j₁)) ptS) (~ j) (p (~ k)) + ; (j = i0) → p (~ k) + ; (j = i1) → llem₂' k i}) + (transp (λ k → A (merid ptS ((i ∨ k) ∧ j)) ptS) (i ∨ ~ j) (g (merid ptS (i ∧ j)))) + + mainₗ : (a : _) (y : _) + → PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) + mainₗ = + wedgeConFun' G H n + (λ _ _ → isOfHLevelPathP' (suc (n + 1)) (hLev _ _) _ _) + (λ x i → transp (λ j → A (merid ptS (i ∧ j)) x) (~ i) (f x)) + (λ x i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → llem₁ k}) + (g (merid x i))) + llem₂ + + mainP : (y : _) + → mainₗ y ptS + ≡ λ i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → llem₁ k}) + (g (merid y i)) + mainP y = + wedgeConFun'ᵣ G H n + (λ _ _ → isOfHLevelPathP' (suc (n + 1)) (hLev _ _) _ _) + (λ x i → transp (λ j → A (merid ptS (i ∧ j)) x) (~ i) (f x)) + (λ x i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → llem₁ k}) + (g (merid x i))) + llem₂ y +wedgeConFun'ᵣ G H zero {A = A} hLev f g p = + elimProp _ (λ _ → hLev _ _ _ _) p +wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p north = p +wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p south = sym (llem₁ G H n hLev f g p ptS i0 ptS) +wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p (merid a i) k = P k i + where + lem : _ + lem i j = + hcomp (λ k → λ { (i = i1) → g (merid a j) + ; (j = i0) → p (i ∨ ~ k) + ; (j = i1) → llem₁ G H n hLev f g p ptS i0 ptS (~ i ∧ k)}) + (g (merid a j)) + + P : PathP (λ k → PathP (λ i → A (merid a i) ptS) + (p k) (llem₁ G H n hLev f g p ptS i0 ptS (~ k))) + (λ i → mainₗ G H n hLev f g p a i ptS a ptS i) λ i → g (merid a i) + P = mainP G H n hLev f g p a i0 ptS a ◁ lem + +private + wedgeConFun : (G H : AbGroup ℓG) (k n m : ℕ) → (n + m ≡ k) → {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} + → ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y)) + → (f : (x : _) → A ptS x) + → (g : (x : _) → A x ptS) + → f ptS ≡ g ptS + → (x : _) (y : _) → A x y + wedgeconLeft : (G H : AbGroup ℓG) (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} + → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) + → (f : (x : _) → A ptS x) + → (g : (x : _) → A x ptS) + → (p : f ptS ≡ g ptS) + → (x : _) → wedgeConFun G H k n m P hLev f g p ptS x ≡ f x + wedgeconRight : (G H : AbGroup ℓG) (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} + → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) + → (f : (x : _) → A ptS x) + → (g : (x : _) → A x ptS) + → (p : f ptS ≡ g ptS) + → (x : _) → wedgeConFun G H k n m P hLev f g p x ptS ≡ g x + wedgeConFun G H k n zero P {A = A} hLev f g p = wedgeConFun' G H n hLev f g p + wedgeConFun G H k zero (suc m) P {A = A} hLev f g p x y = + wedgeConFun' H G (suc m) {A = λ x y → A y x} + (λ x y → subst (λ n → isOfHLevel (2 + n) (A y x)) (+-comm 1 m) (hLev y x)) + g f (sym p) y x + wedgeConFun G H l (suc n) (suc m) P {A = A} hlev f g p north y = f y + wedgeConFun G H l (suc n) (suc m) P {A = A} hlev f g p south y = subst (λ x → A x y) (merid ptS) (f y) + wedgeConFun G H zero (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = main** i + module _ where + main** : PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) + main** = ⊥-rec (snotz P) + wedgeConFun G H (suc l) (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = main* a y i + module _ where + lem₁* : g south ≡ (subst (λ x → A x ptS) (merid ptS) (f ptS)) + lem₁* = ((λ i → transp (λ j → A (merid ptS (~ i ∨ j)) ptS) (~ i) (g (merid ptS (~ i))))) + ∙ cong (subst (λ x → A x ptS) (merid ptS)) (sym p) + + lem₁'* : + Square + (λ i → transp (λ j → A (merid ptS (~ i ∨ j)) ptS) (~ i) (g (merid ptS (~ i)))) + (refl {x = subst (λ x → A x ptS) (merid ptS) (f ptS)}) + lem₁* + ((cong (transport (λ z → A (merid ptS z) ptS)) (sym p))) + lem₁'* i j = + hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptS (~ j ∨ j₁)) ptS) (~ j) (g (merid ptS (~ j))) + ; (i = i1) → subst (λ x → A x ptS) (merid ptS) (p (~ k)) + ; (j = i1) → cong (transport (λ z → A (merid ptS z) ptS)) (sym p) (i ∧ k)}) + (transp (λ j₁ → A (merid ptS ((~ j ∧ ~ i) ∨ j₁)) ptS) (~ j ∧ ~ i) (g (merid ptS (~ j ∧ ~ i)))) + + + lem₂* : (λ i₁ → transp (λ j → A (merid ptS (j ∧ i₁)) ptS) (~ i₁) (f ptS)) + ≡ (λ i₁ → hcomp (λ k → λ { (i₁ = i0) → p (~ k) ; (i₁ = i1) → lem₁* k }) + (g (merid ptS i₁))) + lem₂* i j = + hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptS (z ∧ j)) ptS) (~ j) (p (~ k)) + ; (j = i0) → p (~ k) + ; (j = i1) → lem₁'* k (~ i)}) + (transp (λ z → A (merid ptS ((i ∨ z) ∧ j)) ptS) (i ∨ ~ j) (g (merid ptS (i ∧ j)))) + + main* : (a : _) (y : _) → PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) + main* = + wedgeConFun G H l n (suc m) (cong predℕ P) + (λ _ _ → isOfHLevelPathP' (suc (n + (suc (suc m)))) (hlev _ _) _ _) + (λ x i → transp (λ j → A (merid ptS (j ∧ i)) x) (~ i) (f x)) + (λ y i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → lem₁* k}) + (g (merid y i))) + lem₂* + + mainR : (x : _) → main* x ptS + ≡ λ i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → lem₁* k}) + (g (merid x i)) + mainR x = + wedgeconRight G H l n (suc m) (cong predℕ P) + (λ _ _ → isOfHLevelPathP' (suc (n + (suc (suc m)))) (hlev _ _) _ _) + (λ x i → transp (λ j → A (merid ptS (j ∧ i)) x) (~ i) (f x)) + (λ y i → hcomp (λ k → λ { (i = i0) → p (~ k) + ; (i = i1) → lem₁* k}) + (g (merid y i))) + lem₂* x + wedgeconLeft G H k zero zero P {A = A} hLev f g p _ = refl + wedgeconLeft G H k (suc n) zero P {A = A} hLev f g p _ = refl + wedgeconLeft G H k zero (suc m) P {A = A} hLev f g p x = + wedgeConFun'ᵣ H G (suc m) + (λ x₁ y → + subst (λ n → (x₂ y₁ : A y x₁) → isOfHLevel (suc n) (x₂ ≡ y₁)) + (+-comm 1 m) (hLev y x₁)) + g f (λ i → p (~ i)) x + wedgeconLeft G H k (suc n) (suc m) P {A = A} hLev f g p _ = refl + wedgeconRight G H k n zero P {A = A} hLev f g p = wedgeConFun'ᵣ G H n hLev f g p + wedgeconRight G H k zero (suc m) P {A = A} hLev f g p _ = refl + wedgeconRight G H zero (suc n) (suc m) P {A = A} hLev f g p x = ⊥-rec (snotz P) + wedgeconRight G H l (suc n) (suc m) P {A = A} hLev f g p north = p + wedgeconRight G H l (suc n) (suc m) P {A = A} hLev f g p south = + sym (lem₁* G H _ n m refl hLev f g p ptS i0 ptS) + wedgeconRight G H (suc l) (suc n) (suc m) P {A = A} hLev f g p (merid a i) k = + help k i + where + lem : _ + lem i j = + hcomp (λ k → λ { (i = i1) → g (merid a j) + ; (j = i0) → p (i ∨ ~ k) + ; (j = i1) → lem₁* G H l n m P hLev f g p ptS i0 ptS (~ i ∧ k)}) + (g (merid a j)) + + help : PathP (λ k → PathP (λ i → A (merid a i) ptS) + (p k) (lem₁* G H l n m P hLev f g p ptS i0 ptS (~ k))) + (λ i → main* G H l n m P hLev f g p a i north a north i) (cong g (merid a)) + help = mainR G H l n m P hLev f g p a i0 ptS a ◁ lem + + +module wedgeConEM (G H : AbGroup ℓG) (n m : ℕ) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} + (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) + (f : (x : _) → A ptS x) + (g : (x : _) → A x ptS) + (p : f ptS ≡ g ptS) where + fun : (x : EM-raw G (suc n)) (y : EM-raw H (suc m)) → A x y + fun = wedgeConFun G H (n + m) n m refl hLev f g p + + left : (x : EM-raw H (suc m)) → fun ptS x ≡ f x + left = wedgeconLeft G H (n + m) n m refl hLev f g p + + right : (x : EM-raw G (suc n)) → fun x ptS ≡ g x + right = wedgeconRight G H (n + m) n m refl hLev f g p + +module _ {G : AbGroup ℓG} where + infixr 34 _+ₖ_ + infixr 34 _-ₖ_ + open AbGroupStr (snd G) renaming (_+_ to _+G_ ; -_ to -G_ ; assoc to assocG) + + private + help : (n : ℕ) → n + (4 + n) ≡ (2 + n) + (2 + n) + help n = +-suc n (3 + n) ∙ cong suc (+-suc n (suc (suc n))) + + hLevHelp : (n : ℕ) → isOfHLevel ((2 + n) + (2 + n)) (EM G (2 + n)) + hLevHelp n = + transport (λ i → isOfHLevel (help n i) (EM G (2 + n))) + (isOfHLevelPlus {n = 4 + n} n (isOfHLevelTrunc (4 + n))) + + helper : (g h : fst G) + → PathP (λ i → Path (EM₁ (AbGroup→Group G)) + (emloop h i) (emloop h i)) (emloop g) (emloop g) + helper g h = + compPathL→PathP + (cong (sym (emloop h) ∙_) + (sym (emloop-comp _ g h) + ∙∙ cong emloop (comm g h) + ∙∙ emloop-comp _ h g) + ∙∙ ∙assoc _ _ _ + ∙∙ cong (_∙ emloop g) (lCancel _) + ∙ sym (lUnit _)) + _+ₖ_ : {n : ℕ} → EM G n → EM G n → EM G n + _+ₖ_ {n = zero} = _+G_ + _+ₖ_ {n = suc zero} = + rec _ (isGroupoidΠ (λ _ → emsquash)) + (λ x → x) + hi! + λ g h i j x → el g h x i j + where + lol : (g h : fst G) + → Square (emloop g) (emloop (g +G h)) refl (emloop h) + lol g h = compPath-filler (emloop g) (emloop h) ▷ sym (emloop-comp _ g h) + + hi! : fst G → (λ x → x) ≡ (λ x → x) + hi! g = funExt (elimSet _ (λ _ → emsquash _ _) + (emloop g) + (helper g)) + el : (g h : fst G) (x : EM₁ (AbGroup→Group G)) + → Square (λ j → hi! g j x) (λ j → hi! ((snd (AbGroup→Group G) GroupStr.· g) h) j x) + refl λ j → hi! h j x + el g h = + elimProp _ (λ _ → isOfHLevelPathP' 1 (emsquash _ _) _ _) + (lol g h) + _+ₖ_ {n = suc (suc n)} = + trRec2 (isOfHLevelTrunc (4 + n)) + (wedgeConEM.fun G G (suc n) (suc n) + (λ _ _ → hLevHelp n) + ∣_∣ ∣_∣ refl) + + σ-EM : (n : ℕ) → EM-raw G (suc n) → Path (EM-raw G (2 + n)) ptS ptS + σ-EM n x i = (merid x ∙ sym (merid ptS)) i + + -ₖ_ : {n : ℕ} → EM G n → EM G n + -ₖ_ {n = zero} x = -G x + -ₖ_ {n = suc zero} = + rec _ emsquash + embase + (λ g → sym (emloop g)) + λ g h → sym (emloop-sym _ g) + ◁ (flipSquare + (flipSquare (emcomp (-G g) (-G h)) + ▷ emloop-sym _ h) + ▷ (cong emloop (comm (-G g) (-G h) + ∙ sym (GroupTheory.invDistr (AbGroup→Group G) g h)) + ∙ emloop-sym _ (g +G h))) + -ₖ_ {n = suc (suc n)} = + map λ { north → north + ; south → north + ; (merid a i) → σ-EM n a (~ i)} + + _-ₖ_ : {n : ℕ} → EM G n → EM G n → EM G n + _-ₖ_ {n = n} x y = _+ₖ_ {n = n} x (-ₖ_ {n = n} y) + + +ₖ-syntax : (n : ℕ) → EM G n → EM G n → EM G n + +ₖ-syntax n = _+ₖ_ {n = n} + + -ₖ-syntax : (n : ℕ) → EM G n → EM G n + -ₖ-syntax n = -ₖ_ {n = n} + + -'ₖ-syntax : (n : ℕ) → EM G n → EM G n → EM G n + -'ₖ-syntax n = _-ₖ_ {n = n} + + syntax +ₖ-syntax n x y = x +[ n ]ₖ y + syntax -ₖ-syntax n x = -[ n ]ₖ x + syntax -'ₖ-syntax n x y = x -[ n ]ₖ y + + lUnitₖ : (n : ℕ) (x : EM G n) → 0ₖ n +[ n ]ₖ x ≡ x + lUnitₖ zero x = lid x + lUnitₖ (suc zero) _ = refl + lUnitₖ (suc (suc n)) = + trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ _ → refl + + rUnitₖ : (n : ℕ) (x : EM G n) → x +[ n ]ₖ 0ₖ n ≡ x + rUnitₖ zero x = rid x + rUnitₖ (suc zero) = + elimSet _ (λ _ → emsquash _ _) + refl + λ _ _ → refl + rUnitₖ (suc (suc n)) = + trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + (wedgeConEM.right G G (suc n) (suc n) + (λ _ _ → hLevHelp n) + ∣_∣ ∣_∣ refl) + + commₖ : (n : ℕ) (x y : EM G n) → x +[ n ]ₖ y ≡ y +[ n ]ₖ x + commₖ zero = comm + commₖ (suc zero) = + wedgeConEM.fun G G 0 0 (λ _ _ → emsquash _ _) + (λ x → sym (rUnitₖ 1 x)) + (rUnitₖ 1) + refl + commₖ (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelTruncPath {n = 4 + n}) + (wedgeConEM.fun G G _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (hLevHelp n) _ _) + (λ x → sym (rUnitₖ (2 + n) ∣ x ∣)) + (λ x → rUnitₖ (2 + n) ∣ x ∣) + refl) + + open import Cubical.Homotopy.Loopspace + cong₂+₁ : (p q : typ (Ω (EM∙ G 1))) → cong₂ (λ x y → x +[ 1 ]ₖ y) p q ≡ p ∙ q + cong₂+₁ p q = + (cong₂Funct (λ x y → x +[ 1 ]ₖ y) p q) + ∙ (λ i → (cong (λ x → rUnitₖ 1 x i) p) ∙ (cong (λ x → lUnitₖ 1 x i) q)) + + cong₂+₂ : (n : ℕ) (p q : typ (Ω (EM∙ G (suc (suc n))))) → cong₂ (λ x y → x +[ (2 + n) ]ₖ y) p q ≡ p ∙ q + cong₂+₂ n p q = + (cong₂Funct (λ x y → x +[ (2 + n) ]ₖ y) p q) + ∙ (λ i → (cong (λ x → rUnitₖ (2 + n) x i) p) ∙ (cong (λ x → lUnitₖ (2 + n) x i) q)) + + isCommΩEM : (n : ℕ) (p q : typ (Ω (EM∙ G (suc n)))) → p ∙ q ≡ q ∙ p + isCommΩEM zero p q = + sym (cong₂+₁ p q) + ∙∙ (λ i j → commₖ 1 (p j) (q j) i) + ∙∙ cong₂+₁ q p + isCommΩEM (suc n) p q = + (sym (cong₂+₂ n p q) + ∙∙ (λ i j → commₖ (suc (suc n)) (p j) (q j) i) + ∙∙ cong₂+₂ n q p) + + cong-₁ : (p : typ (Ω (EM∙ G 1))) → cong (λ x → -[ 1 ]ₖ x) p ≡ sym p + cong-₁ p = main embase p + where + decoder : (x : EM G 1) → embase ≡ x → x ≡ embase + decoder = + elimSet _ + (λ _ → isSetΠ λ _ → emsquash _ _) + (λ p i → -[ 1 ]ₖ (p i)) + λ g → toPathP + (funExt λ x → + (λ i → transport (λ i → Path (EM G 1) (emloop g i) embase) + (cong (-ₖ_ {n = 1}) + (transp (λ j → Path (EM G 1) embase (emloop g (~ j ∧ ~ i))) i + (compPath-filler x (sym (emloop g)) i) ))) + ∙∙ (λ i → transp (λ j → Path (EM G 1) (emloop g (i ∨ j)) embase) i + (compPath-filler' + (sym (emloop g)) + (cong-∙ (-ₖ_ {n = 1}) + x (sym (emloop g)) i) i)) + ∙∙ (cong (sym (emloop g) ∙_) (isCommΩEM 0 (cong (-ₖ_ {n = 1}) x) (emloop g))) + ∙∙ ∙assoc _ _ _ + ∙∙ cong (_∙ (cong (-ₖ_ {n = 1}) x)) (lCancel (emloop g)) + ∙ sym (lUnit _)) + + main : (x : EM G 1) (p : embase ≡ x) → decoder x p ≡ sym p + main x = J (λ x p → decoder x p ≡ sym p) refl + + cong-₂ : (n : ℕ) (p : typ (Ω (EM∙ G (2 + n)))) → cong (λ x → -[ 2 + n ]ₖ x) p ≡ sym p + cong-₂ n p = main _ p + where + pp : (a : _) → PathP (λ i → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ → ∣ merid a i ∣ₕ ≡ 0ₖ (2 + n)) + (cong (λ x → -[ 2 + n ]ₖ x)) λ p → cong ∣_∣ₕ (sym (merid ptS)) ∙ cong (λ x → -[ 2 + n ]ₖ x) p + pp a = + toPathP + (funExt λ x → + (λ k → transp (λ i → Path (EM G (2 + n)) ∣ merid a (i ∨ k) ∣ ∣ ptS ∣) k (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) + (cong (-ₖ-syntax (suc (suc n))) + (transp (λ j → Path (EM G (2 + n)) ∣ ptS ∣ ∣ merid a (~ j ∧ ~ k) ∣) k (compPath-filler x (sym (cong ∣_∣ₕ (merid a))) k))) k)) + ∙∙ cong (cong ∣_∣ₕ (sym (merid a)) ∙_) (cong-∙ (λ x → -[ 2 + n ]ₖ x) x (sym (cong ∣_∣ₕ (merid a))) + ∙ isCommΩEM (suc n) (cong (λ x → -[ 2 + n ]ₖ x) x) (cong ∣_∣ₕ (σ-EM n a))) + ∙∙ (λ k → (λ i → ∣ merid a (~ i ∨ k) ∣) ∙ (λ i → ∣ compPath-filler' (merid a) (sym (merid ptS)) (~ k) i ∣) + ∙ cong (λ x → -ₖ-syntax (suc (suc n)) x) x) + ∙ sym (lUnit _)) + + decoder : (x : EM G (2 + n)) → 0ₖ (2 + n) ≡ x → x ≡ 0ₖ (2 + n) + decoder = + trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ { north → pp ptS i0 + ; south → pp ptS i1 + ; (merid a i) → pp a i} + + main : (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → decoder x p ≡ sym p + main x = J (λ x p → decoder x p ≡ sym p) refl + + rCancelₖ : (n : ℕ) (x : EM G n) → x +[ n ]ₖ (-[ n ]ₖ x) ≡ 0ₖ n + rCancelₖ zero x = invr x + rCancelₖ (suc zero) = + elimSet _ (λ _ → emsquash _ _) + refl + λ g → flipSquare (cong₂+₁ (emloop g) (λ i → -ₖ-syntax 1 (emloop g i)) + ∙ rCancel (emloop g)) + rCancelₖ (suc (suc n)) = + trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ { north → refl + ; south i → +ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣ (-ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣) + ; (merid a i) j + → hcomp (λ r → λ { (i = i0) → 0ₖ (2 + n) + ; (i = i1) → ∣ merid ptS (~ j ∧ r) ∣ₕ -[ 2 + n ]ₖ ∣ merid ptS (~ j ∧ r) ∣ + ; (j = i0) → ∣ compPath-filler (merid a) (sym (merid ptS)) (~ r) i ∣ + -[ 2 + n ]ₖ ∣ compPath-filler (merid a) (sym (merid ptS)) (~ r) i ∣ + ; (j = i1) → 0ₖ (2 + n)}) + (help' a j i) } + where + help' : (a : _) → cong₂ (λ x y → ∣ x ∣ -[ suc (suc n) ]ₖ ∣ y ∣) (σ-EM n a) (σ-EM n a) ≡ refl + help' a = + cong₂+₂ n (cong ∣_∣ₕ (σ-EM n a)) (cong (λ x → -[ 2 + n ]ₖ ∣ x ∣) (σ-EM n a)) + ∙∙ cong (cong ∣_∣ₕ (σ-EM n a) ∙_) (cong-₂ n (cong ∣_∣ₕ (σ-EM n a))) + ∙∙ rCancel _ + + lCancelₖ : (n : ℕ) (x : EM G n) → (-[ n ]ₖ x) +[ n ]ₖ x ≡ 0ₖ n + lCancelₖ n x = commₖ n (-[ n ]ₖ x) x ∙ rCancelₖ n x + + assocₖ : (n : ℕ) (x y z : EM G n) + → (x +[ n ]ₖ (y +[ n ]ₖ z) ≡ (x +[ n ]ₖ y) +[ n ]ₖ z) + assocₖ zero = assocG + assocₖ (suc zero) = + elimSet _ (λ _ → isSetΠ2 λ _ _ → emsquash _ _) + (λ _ _ → refl) + λ g i y z k → lem g y z k i + where + lem : (g : fst G) (y z : _) → cong (λ x → x +[ suc zero ]ₖ (y +[ suc zero ]ₖ z)) (emloop g) + ≡ cong (λ x → (x +[ suc zero ]ₖ y) +[ suc zero ]ₖ z) (emloop g) + lem g = + elimProp _ (λ _ → isPropΠ λ _ → emsquash _ _ _ _) + (elimProp _ (λ _ → emsquash _ _ _ _) + refl) + assocₖ (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ a b → trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + (λ c → main c a b) + where + lem : (c : _) (a b : _) + → PathP (λ i → (∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ merid c i ∣ₕ) + ≡ (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ merid c i ∣ₕ)) + (cong (λ x → ∣ a ∣ₕ +[ suc (suc n) ]ₖ x) (rUnitₖ (suc (suc n)) ∣ b ∣) + ∙ sym (rUnitₖ (suc (suc n)) (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ))) + ((λ i → ∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ merid ptS (~ i) ∣ₕ)) + ∙∙ cong (λ x → ∣ a ∣ₕ +[ suc (suc n) ]ₖ x) (rUnitₖ (suc (suc n)) ∣ b ∣) + ∙ sym (rUnitₖ (suc (suc n)) (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ)) + ∙∙ λ i → (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ merid ptS i ∣ₕ) + lem c = + EM-raw-elim G (suc n) + (λ _ → isOfHLevelΠ (2 + n) (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _)) + (EM-raw-elim G (suc n) (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _) + ((sym (rUnit refl) + ◁ λ _ → refl) + ▷ (sym (lCancel (cong ∣_∣ₕ (merid ptS))) + ∙ λ i → (λ j → ∣ merid ptS (~ j ∨ ~ i) ∣ₕ) + ∙∙ lUnit (λ j → ∣ merid ptS (~ j ∧ ~ i) ∣ₕ) i + ∙∙ cong ∣_∣ₕ (merid ptS)))) + main : (c a b : _) → (∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ c ∣ₕ) ≡ (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ c ∣ₕ) + main north a b = lem ptS a b i0 + main south a b = lem ptS a b i1 + main (merid c i) a b = lem c a b i + + σ-EM' : (n : ℕ) (x : EM G (suc n)) + → Path (EM G (suc (suc n))) + (0ₖ (suc (suc n))) + (0ₖ (suc (suc n))) + σ-EM' zero x = cong ∣_∣ₕ (σ-EM zero x) + σ-EM' (suc n) = + trElim (λ _ → isOfHLevelTrunc (5 + n) _ _) + λ x → cong ∣_∣ₕ (σ-EM (suc n) x) + + σ-EM'-0ₖ : (n : ℕ) → σ-EM' n (0ₖ (suc n)) ≡ refl + σ-EM'-0ₖ zero = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + σ-EM'-0ₖ (suc n) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + + private + P : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (r : refl ≡ p) + → lUnit p ∙ cong (_∙ p) r + ≡ rUnit p ∙ cong (p ∙_) r + P p = J (λ p r → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r) refl + + σ-EM'-hom : (n : ℕ) → (a b : _) → σ-EM' n (a +ₖ b) ≡ σ-EM' n a ∙ σ-EM' n b + σ-EM'-hom zero = + wedgeConEM.fun G G 0 0 (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) l r p + where + l : _ + l x = cong (σ-EM' zero) (lUnitₖ 1 x) + ∙∙ lUnit (σ-EM' zero x) + ∙∙ cong (_∙ σ-EM' zero x) (sym (σ-EM'-0ₖ zero)) + + r : _ + r x = + cong (σ-EM' zero) (rUnitₖ 1 x) + ∙∙ rUnit (σ-EM' zero x) + ∙∙ cong (σ-EM' zero x ∙_) (sym (σ-EM'-0ₖ zero)) + + p : _ + p = P (σ-EM' zero embase) (sym (σ-EM'-0ₖ zero)) + σ-EM'-hom (suc n) = + elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) _ _) _ _) + (wedgeConEM.fun G G _ _ + (λ x y → transport (λ i → isOfHLevel (help n i) + ((σ-EM' (suc n) (∣ x ∣ₕ +ₖ ∣ y ∣ₕ) + ≡ σ-EM' (suc n) ∣ x ∣ₕ ∙ σ-EM' (suc n) ∣ y ∣ₕ))) + (isOfHLevelPlus {n = 4 + n} n + (isOfHLevelPath (4 + n) + (isOfHLevelTrunc (5 + n) _ _) _ _))) + (λ x → cong (σ-EM' (suc n)) (lUnitₖ (suc (suc n)) ∣ x ∣) + ∙∙ lUnit (σ-EM' (suc n) ∣ x ∣) + ∙∙ cong (_∙ σ-EM' (suc n) ∣ x ∣) (sym (σ-EM'-0ₖ (suc n)))) + (λ x → cong (σ-EM' (suc n)) (rUnitₖ (2 + n) ∣ x ∣) + ∙∙ rUnit (σ-EM' (suc n) ∣ x ∣) + ∙∙ cong (σ-EM' (suc n) ∣ x ∣ ∙_) (sym (σ-EM'-0ₖ (suc n)))) + (P (σ-EM' (suc n) (0ₖ (2 + n))) (sym (σ-EM'-0ₖ (suc n))))) + + σ-EM'-ₖ : (n : ℕ) → (a : _) → σ-EM' n (-ₖ a) ≡ sym (σ-EM' n a) + σ-EM'-ₖ n = + morphLemmas.distrMinus + (λ x y → x +[ suc n ]ₖ y) (_∙_) + (σ-EM' n) (σ-EM'-hom n) + (0ₖ (suc n)) refl + (λ x → -ₖ x) sym + (λ x → sym (lUnit x)) (λ x → sym (rUnit x)) + (lCancelₖ (suc n)) rCancel + ∙assoc (σ-EM'-0ₖ n) + + -Dist : (n : ℕ) (x y : EM G n) → -[ n ]ₖ (x +[ n ]ₖ y) ≡ (-[ n ]ₖ x) +[ n ]ₖ (-[ n ]ₖ y) + -Dist zero x y = (GroupTheory.invDistr (AbGroup→Group G) x y) ∙ commₖ zero _ _ + -Dist (suc zero) = k + where -- useless where clause. Needed for fast type checking for some reason. + l : _ + l x = refl + + r : _ + r x = cong (λ z → -[ 1 ]ₖ z) (rUnitₖ 1 x) ∙ sym (rUnitₖ 1 (-[ 1 ]ₖ x)) + + p : r ptS ≡ l ptS + p = sym (rUnit refl) + + k = wedgeConEM.fun G G 0 0 (λ _ _ → emsquash _ _) l r (sym p) + + -Dist (suc (suc n)) = + elim2 (λ _ _ → isOfHLevelTruncPath {n = 4 + n}) + (wedgeConEM.fun G G (suc n) (suc n) + (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (hLevHelp n) _ _) + (λ x → refl) + (λ x → cong (λ z → -[ (suc (suc n)) ]ₖ z) (rUnitₖ (suc (suc n)) ∣ x ∣ₕ) ∙ sym (rUnitₖ (suc (suc n)) (-[ (suc (suc n)) ]ₖ ∣ x ∣ₕ))) + (rUnit refl)) + + addIso : (n : ℕ) (x : EM G n) → Iso (EM G n) (EM G n) + Iso.fun (addIso n x) y = y +[ n ]ₖ x + Iso.inv (addIso n x) y = y -[ n ]ₖ x + Iso.rightInv (addIso n x) y = + sym (assocₖ n y (-[ n ]ₖ x) x) + ∙∙ cong (λ x → y +[ n ]ₖ x) (lCancelₖ n x) + ∙∙ rUnitₖ n y + Iso.leftInv (addIso n x) y = + sym (assocₖ n y x (-[ n ]ₖ x)) + ∙∙ cong (λ x → y +[ n ]ₖ x) (rCancelₖ n x) + ∙∙ rUnitₖ n y + + CODE : (n : ℕ) → EM G (suc (suc n)) → TypeOfHLevel ℓG (3 + n) + CODE n = + trElim (λ _ → isOfHLevelTypeOfHLevel (3 + n)) + λ { north → EM G (suc n) , hLevelEM G (suc n) + ; south → EM G (suc n) , hLevelEM G (suc n) + ; (merid a i) → fib n a i} + where + fib : (n : ℕ) → (a : EM-raw G (suc n)) + → Path (TypeOfHLevel _ (3 + n)) (EM G (suc n) , hLevelEM G (suc n)) (EM G (suc n) , hLevelEM G (suc n)) + fib zero a = Σ≡Prop (λ _ → isPropIsOfHLevel 3) + (isoToPath (addIso 1 a)) + fib (suc n) a = Σ≡Prop (λ _ → isPropIsOfHLevel (4 + n)) + (isoToPath (addIso (suc (suc n)) ∣ a ∣)) + + decode' : (n : ℕ) (x : EM G (suc (suc n))) → CODE n x .fst → 0ₖ (suc (suc n)) ≡ x + decode' n = + trElim (λ _ → isOfHLevelΠ (4 + n) + λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + λ { north → f n + ; south → g n + ; (merid a i) → main a i} + where + f : (n : ℕ) → _ + f n = σ-EM' n + + g : (n : ℕ) → EM G (suc n) → ∣ ptS ∣ ≡ ∣ south ∣ + g n x = σ-EM' n x ∙ cong ∣_∣ₕ (merid ptS) + + lem₂ : (n : ℕ) (a x : _) + → Path (Path (EM G (suc (suc n))) _ _) + ((λ i → cong ∣_∣ₕ (σ-EM n x) i) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) ∙ (λ i → ∣ merid a i ∣ₕ)) + (g n (EM-raw→EM G (suc n) x)) + lem₂ zero a x = + cong (f zero x ∙_) + (cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid embase))) + ∙ cong-∙ ∣_∣ₕ (merid embase) (sym (merid a))) + ∙∙ sym (∙assoc _ _ _) + ∙∙ cong (cong ∣_∣ₕ (merid embase) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) + ∙ sym (rUnit _)) + lem₂ (suc n) a x = + cong (f (suc n) ∣ x ∣ ∙_) + ((cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid north))) + ∙ cong-∙ ∣_∣ₕ (merid north) (sym (merid a))) + ∙∙ sym (∙assoc _ _ _) + ∙∙ cong (cong ∣_∣ₕ (merid north) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) + ∙ sym (rUnit _))) + + lem : (n : ℕ) (x a : EM-raw G (suc n)) + → f n (transport (sym (cong (λ x → CODE n x .fst) (cong ∣_∣ₕ (merid a)))) (EM-raw→EM G (suc n) x)) + ≡ cong ∣_∣ₕ (σ-EM n x) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) + lem zero x a = (λ i → cong ∣_∣ₕ (merid (transportRefl x i -[ 1 ]ₖ a) ∙ sym (merid embase))) + ∙∙ σ-EM'-hom zero x (-ₖ a) + ∙∙ cong (f zero x ∙_) (σ-EM'-ₖ zero a) + lem (suc n) x a = + cong (f (suc n)) (λ i → transportRefl ∣ x ∣ i -[ 2 + n ]ₖ ∣ a ∣) + ∙∙ σ-EM'-hom (suc n) ∣ x ∣ (-ₖ ∣ a ∣) + ∙∙ cong (f (suc n) ∣ x ∣ ∙_) (σ-EM'-ₖ (suc n) ∣ a ∣) + + main : (a : _) + → PathP (λ i → CODE n ∣ merid a i ∣ₕ .fst + → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ) (f n) (g n) + main a = + toPathP (funExt + (EM-elim _ (λ _ → isOfHLevelPathP (2 + (suc n)) (isOfHLevelTrunc (4 + n) _ _) _ _) + λ x → + ((λ i → transp (λ j → Path (EM G (2 + n)) ∣ ptS ∣ ∣ merid a (i ∨ j) ∣) + i (compPath-filler (lem n x a i) (cong ∣_∣ₕ (merid a)) i) )) + ∙∙ sym (∙assoc _ _ _) + ∙∙ lem₂ n a x)) + + encode' : (n : ℕ) (x : EM G (suc (suc n))) → 0ₖ (suc (suc n)) ≡ x → CODE n x .fst + encode' n x p = subst (λ x → CODE n x .fst) p (0ₖ (suc n)) + + decode'-encode' : (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → decode' n x (encode' n x p) ≡ p + decode'-encode' zero x = + J (λ x p → decode' 0 x (encode' 0 x p) ≡ p) + (σ-EM'-0ₖ 0) + decode'-encode' (suc n) x = + J (λ x p → decode' (suc n) x (encode' (suc n) x p) ≡ p) + (σ-EM'-0ₖ (suc n)) + + encode'-decode' : (n : ℕ) (x : _) → encode' n (0ₖ (suc (suc n))) (decode' n (0ₖ (suc (suc n))) x) ≡ x + encode'-decode' zero x = + cong (encode' zero (0ₖ 2)) (cong-∙ ∣_∣ₕ (merid x) (sym (merid embase))) + ∙∙ substComposite (λ x → CODE zero x .fst) + (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptS))) embase + ∙∙ (cong (subst (λ x₁ → CODE zero x₁ .fst) (λ i → ∣ merid embase (~ i) ∣ₕ)) + (λ i → lUnitₖ 1 (transportRefl x i) i) + ∙ (λ i → rUnitₖ 1 (transportRefl x i) i)) + encode'-decode' (suc n) = + trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) + λ x → cong (encode' (suc n) (0ₖ (3 + n))) (cong-∙ ∣_∣ₕ (merid x) (sym (merid north))) + ∙∙ substComposite (λ x → CODE (suc n) x .fst) + (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptS))) (0ₖ (2 + n)) + ∙∙ cong (subst (λ x₁ → CODE (suc n) x₁ .fst) (λ i → ∣ merid ptS (~ i) ∣ₕ)) + (λ i → lUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) + ∙ (λ i → rUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) + + EM∙' : (n : ℕ) → Pointed _ + EM∙' n = EM G n , 0ₖ n + + Iso-EM-ΩEM+1 : (n : ℕ) → Iso (EM G n) (typ (Ω (EM∙' (suc n)))) + Iso-EM-ΩEM+1 zero = invIso (ΩEM₁Iso (G *)) + Iso.fun (Iso-EM-ΩEM+1 (suc zero)) = decode' 0 (0ₖ 2) + Iso.inv (Iso-EM-ΩEM+1 (suc zero)) = encode' 0 (0ₖ 2) + Iso.rightInv (Iso-EM-ΩEM+1 (suc zero)) = decode'-encode' 0 (0ₖ 2) + Iso.leftInv (Iso-EM-ΩEM+1 (suc zero)) = encode'-decode' 0 + Iso.fun (Iso-EM-ΩEM+1 (suc (suc n))) = decode' (suc n) (0ₖ (3 + n)) + Iso.inv (Iso-EM-ΩEM+1 (suc (suc n))) = encode' (suc n) (0ₖ (3 + n)) + Iso.rightInv (Iso-EM-ΩEM+1 (suc (suc n))) = decode'-encode' (suc n) (0ₖ (3 + n)) + Iso.leftInv (Iso-EM-ΩEM+1 (suc (suc n))) = encode'-decode' (suc n) + + EM→ΩEM+1 : (n : ℕ) → EM G n → typ (Ω (EM∙' (suc n))) + EM→ΩEM+1 n = Iso.fun (Iso-EM-ΩEM+1 n) + + ΩEM+1→EM : (n : ℕ) → typ (Ω (EM∙' (suc n))) → EM G n + ΩEM+1→EM n = Iso.inv (Iso-EM-ΩEM+1 n) + + EM→ΩEM+1-hom : (n : ℕ) (x y : EM G n) + → EM→ΩEM+1 n (x +[ n ]ₖ y) ≡ EM→ΩEM+1 n x ∙ EM→ΩEM+1 n y + EM→ΩEM+1-hom zero x y = emloop-comp _ x y + EM→ΩEM+1-hom (suc zero) x y = σ-EM'-hom zero x y + EM→ΩEM+1-hom (suc (suc n)) x y = σ-EM'-hom (suc n) x y + + ΩEM+1→EM-hom : (n : ℕ) → (p q : typ (Ω (EM∙' (suc n)))) + → ΩEM+1→EM n (p ∙ q) ≡ (ΩEM+1→EM n p) +[ n ]ₖ (ΩEM+1→EM n q) + ΩEM+1→EM-hom n = + morphLemmas.isMorphInv (λ x y → x +[ n ]ₖ y) (_∙_) (EM→ΩEM+1 n) + (EM→ΩEM+1-hom n) (ΩEM+1→EM n) + (Iso.rightInv (Iso-EM-ΩEM+1 n)) (Iso.leftInv (Iso-EM-ΩEM+1 n)) + + EM→ΩEM+1-0ₖ : (n : ℕ) → EM→ΩEM+1 n (0ₖ n) ≡ refl + EM→ΩEM+1-0ₖ zero = emloop-1g _ + EM→ΩEM+1-0ₖ (suc zero) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + EM→ΩEM+1-0ₖ (suc (suc n)) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + + ΩEM+1→EM-refl : (n : ℕ) → ΩEM+1→EM n refl ≡ 0ₖ n + ΩEM+1→EM-refl zero = transportRefl 0g + ΩEM+1→EM-refl (suc zero) = refl + ΩEM+1→EM-refl (suc (suc n)) = refl diff --git a/Cubical/Foundations/HLevels.agda b/Cubical/Foundations/HLevels.agda index 9a128fbfd..07b41e1fc 100644 --- a/Cubical/Foundations/HLevels.agda +++ b/Cubical/Foundations/HLevels.agda @@ -637,3 +637,19 @@ isOfHLevelΣ' {A = A} {B = B} (suc (suc n)) Alvl Blvl (w , y) (x , z) ΣPathP (λ x → refl) (isOfHLevelΣ' (suc n) (Alvl w x) (Blvl y z)) + +Σ≡Set : ((x : A) → isSet (B x)) + → {a₀₀ a₀₁ : Σ A B} (p₀₋ : a₀₀ ≡ a₀₁) + → {a₁₀ a₁₁ : Σ A B} (p₁₋ : a₁₀ ≡ a₁₁) + → (p₋₀ : a₀₀ ≡ a₁₀) (p₋₁ : a₀₁ ≡ a₁₁) + → Square (cong fst p₀₋) (cong fst p₁₋) (cong fst p₋₀) (cong fst p₋₁) + → Square p₀₋ p₁₋ p₋₀ p₋₁ +fst (Σ≡Set set p₀₋ p₁₋ p₋₀ p₋₁ sq i j) = sq i j +snd (Σ≡Set {B = B} set p₀₋ p₁₋ p₋₀ p₋₁ sq i j) = sq2 i j + where + sq2 : SquareP (λ i j → B (sq i j)) + (λ j → snd (p₀₋ j)) + (λ j → snd (p₁₋ j)) + (λ i → snd (p₋₀ i)) + (λ i → snd (p₋₁ i)) + sq2 = toPathP (isOfHLevelPathP' 1 (set _) _ _ _ _) diff --git a/Cubical/HITs/EilenbergMacLane1/Properties.agda b/Cubical/HITs/EilenbergMacLane1/Properties.agda index de286334f..de5d1f51b 100644 --- a/Cubical/HITs/EilenbergMacLane1/Properties.agda +++ b/Cubical/HITs/EilenbergMacLane1/Properties.agda @@ -130,902 +130,3 @@ module _ ((G , str) : Group ℓG) where → (x : EM₁ (G , str)) → B rec Bgpd = elim (λ _ → Bgpd) - - -open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) -open import Cubical.Data.Nat -open import Cubical.HITs.Susp -open import Cubical.Foundations.Pointed - -Susp̂ : (n : ℕ) → Type ℓG → Type ℓG -Susp̂ zero A = A -Susp̂ (suc n) A = Susp (Susp̂ n A) - -pt-Susp̂ : (n : ℕ) (A : Pointed ℓG) → Susp̂ n (typ A) -pt-Susp̂ zero A = snd A -pt-Susp̂ (suc n) A = north - -ptS : {n : ℕ} {G : Group ℓG} → Susp̂ n (EM₁ G) -ptS {n = n} {G = G} = pt-Susp̂ n (EM₁ G , embase) - -EM-raw : (G : AbGroup ℓG) (n : ℕ) → Type ℓG -EM-raw G zero = fst G -EM-raw G (suc n) = Susp̂ n (EM₁ (AbGroup→Group G)) - -EM-raw-elim : (G : AbGroup ℓG) (n : ℕ) {A : EM-raw G (suc n) → Type ℓ} - → ((x : _) → isOfHLevel (suc n) (A x) ) - → A ptS - → (x : _) → A x -EM-raw-elim G zero hlev b = elimProp _ hlev b -EM-raw-elim G (suc n) hlev b north = b -EM-raw-elim G (suc n) {A = A} hlev b south = subst A (merid ptS) b -EM-raw-elim G (suc n) {A = A} hlev b (merid a i) = help a i - where - help : (a : _) → PathP (λ i → A (merid a i)) b (subst A (merid ptS) b) - help = EM-raw-elim G n (λ _ → isOfHLevelPathP' (suc n) (hlev _) _ _) - λ i → transp (λ j → A (merid ptS (i ∧ j))) (~ i) b - -EM∙ : (G : AbGroup ℓG) (n : ℕ) → Pointed ℓG -EM∙ G zero = fst G , AbGroupStr.0g (snd G) -EM∙ G (suc zero) = EM₁ (AbGroup→Group G) , embase -EM∙ G (suc (suc n)) = hLevelTrunc (4 + n) (Susp̂ (suc n) (EM₁ (AbGroup→Group G))) , ∣ north ∣ - -EM : (G : AbGroup ℓG) (n : ℕ) → Type ℓG -EM G zero = EM-raw G zero -EM G (suc zero) = EM-raw G 1 -EM G (suc (suc n)) = hLevelTrunc (4 + n) (EM-raw G (suc (suc n))) - -hLevelEM : (G : AbGroup ℓG) (n : ℕ) → isOfHLevel (2 + n) (EM G n) -hLevelEM G zero = AbGroupStr.is-set (snd G) -hLevelEM G (suc zero) = emsquash -hLevelEM G (suc (suc n)) = isOfHLevelTrunc (4 + n) - -EM-raw→EM : (G : AbGroup ℓG) (n : ℕ) → EM-raw G n → EM G n -EM-raw→EM G zero x = x -EM-raw→EM G (suc zero) x = x -EM-raw→EM G (suc (suc n)) = ∣_∣ - -EM-elim : {G : AbGroup ℓG} (n : ℕ) {A : EM G n → Type ℓ} - → ((x : _) → isOfHLevel (2 + n) (A x)) - → ((x : EM-raw G n) → A (EM-raw→EM G n x)) - → (x : _) → A x -EM-elim zero hlev hyp x = hyp x -EM-elim (suc zero) hlev hyp x = hyp x -EM-elim (suc (suc n)) hlev hyp = trElim (λ _ → hlev _) hyp - -wedgeConFun' : (G H : AbGroup ℓG) (n : ℕ) → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ} - → ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y)) - → (f : (x : _) → A ptS x) - → (g : (x : _) → A x ptS) - → f ptS ≡ g ptS - → (x : _) (y : _) → A x y -wedgeConFun'ᵣ : (G H : AbGroup ℓG) (n : ℕ) → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ} - → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y))) - → (f : (x : _) → A ptS x) - → (g : (x : _) → A x ptS) - → (p : f ptS ≡ g ptS) - → (x : _) → wedgeConFun' G H n hLev f g p x ptS ≡ g x -wedgeConFun' G H zero {A = A} hlev f g p = - elimSet _ (λ _ → isSetΠ λ _ → hlev _ _) f k - where - I→A : (x : fst G) → (i : I) → A (emloop x i) embase - I→A x i = - hcomp (λ k → λ {(i = i0) → p (~ k) ; (i = i1) → p (~ k)}) - (g (emloop x i)) - - SquareP2 : (x : _) (z : _) - → SquareP (λ i j → A (emloop x i) (emloop z j)) - (cong f (emloop z)) (cong f (emloop z)) - (λ i → I→A x i) λ i → I→A x i - SquareP2 x z = - toPathP (isOfHLevelPathP' 1 (λ _ _ → hlev _ _ _ _) _ _ _ _) - - CubeP2 : (x : _) (g h : _) - → PathP (λ k → PathP (λ j → PathP (λ i → A (emloop x i) (emcomp g h j k)) - (f (emcomp g h j k)) (f (emcomp g h j k))) - (λ i → SquareP2 x g i k) λ i → SquareP2 x ((snd (AbGroup→Group H) GroupStr.· g) h) i k) - (λ _ i → I→A x i) λ j i → SquareP2 x h i j - CubeP2 x g h = toPathP (isOfHLevelPathP' 1 (isOfHLevelPathP 2 (hlev _ _) _ _) _ _ _ _) - - k : (x : _) → PathP (λ i → (y : EM₁ (AbGroup→Group H)) → A (emloop x i) y) f f - k x i embase = I→A x i - k x i (emloop z j) = SquareP2 x z i j - k x i (emcomp g h j k) = CubeP2 x g h k j i - k x i (emsquash y z p q r s j k' l) = mega i j k' l - where - mega : - PathP (λ i → - PathP (λ j → - PathP (λ k' → - PathP (λ l → A (emloop x i) (emsquash y z p q r s j k' l)) - (k x i y) - (k x i z)) - (λ l → k x i (p l)) - λ l → k x i (q l)) - (λ k' l → k x i (r k' l)) - λ k' l → k x i (s k' l)) - (λ j k l → f (emsquash y z p q r s j k l)) - λ j k l → f (emsquash y z p q r s j k l) - mega = - toPathP (isOfHLevelPathP' 1 - (isOfHLevelPathP 2 (isOfHLevelPathP 2 (hlev _ _) _ _) _ _) _ _ _ _) -wedgeConFun' G H (suc n) {A = A} hLev f g p north y = f y -wedgeConFun' G H (suc n) {A = A} hLev f g p south y = subst (λ x → A x y) (merid ptS) (f y) -wedgeConFun' G H (suc n) {A = A} hLev f g p (merid a i) y = mainₗ a y i - module _ where - llem₁ : g south ≡ subst (λ x₁ → A x₁ ptS) (merid ptS) (f ptS) - llem₁ = (λ i → transp (λ j → A (merid ptS (j ∨ ~ i)) ptS) (~ i) (g (merid ptS (~ i)))) - ∙ cong (subst (λ x₁ → A x₁ ptS) (merid ptS)) (sym p) - - llem₂' : - Square - (λ i → transp (λ j → A (merid ptS (i ∨ j)) ptS) i (g (merid ptS i))) refl - (cong (subst (λ x → A x ptS) (merid ptS)) (sym p)) llem₁ - llem₂' i j = - hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptS (j ∨ z)) ptS) j - (g (merid ptS j)) - ; (i = i1) → subst (λ x₁ → A x₁ ptS) (merid ptS) (p (~ k)) - ; (j = i0) → (subst (λ x → A x ptS) (merid ptS)) (p (~ k ∨ ~ i))}) - (transp (λ k → A (merid ptS (k ∨ ~ i ∧ j)) ptS) (~ i ∧ j) (g (merid ptS (~ i ∧ j)))) - - llem₂ : (λ i₁ → transp (λ j → A (merid ptS (i₁ ∧ j)) ptS) (~ i₁) (f ptS)) - ≡ (λ i₁ → hcomp (λ k → λ { (i₁ = i0) → p (~ k) ; (i₁ = i1) → llem₁ k }) - (g (merid ptS i₁))) - llem₂ i j = - hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptS (j ∧ j₁)) ptS) (~ j) (p (~ k)) - ; (j = i0) → p (~ k) - ; (j = i1) → llem₂' k i}) - (transp (λ k → A (merid ptS ((i ∨ k) ∧ j)) ptS) (i ∨ ~ j) (g (merid ptS (i ∧ j)))) - - mainₗ : (a : _) (y : _) - → PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) - mainₗ = - wedgeConFun' G H n - (λ _ _ → isOfHLevelPathP' (suc (n + 1)) (hLev _ _) _ _) - (λ x i → transp (λ j → A (merid ptS (i ∧ j)) x) (~ i) (f x)) - (λ x i → hcomp (λ k → λ { (i = i0) → p (~ k) - ; (i = i1) → llem₁ k}) - (g (merid x i))) - llem₂ - - mainP : (y : _) - → mainₗ y ptS - ≡ λ i → hcomp (λ k → λ { (i = i0) → p (~ k) - ; (i = i1) → llem₁ k}) - (g (merid y i)) - mainP y = - wedgeConFun'ᵣ G H n - (λ _ _ → isOfHLevelPathP' (suc (n + 1)) (hLev _ _) _ _) - (λ x i → transp (λ j → A (merid ptS (i ∧ j)) x) (~ i) (f x)) - (λ x i → hcomp (λ k → λ { (i = i0) → p (~ k) - ; (i = i1) → llem₁ k}) - (g (merid x i))) - llem₂ y -wedgeConFun'ᵣ G H zero {A = A} hLev f g p = - elimProp _ (λ _ → hLev _ _ _ _) p -wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p north = p -wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p south = sym (llem₁ G H n hLev f g p ptS i0 ptS) -wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p (merid a i) k = P k i - where - lem : _ - lem i j = - hcomp (λ k → λ { (i = i1) → g (merid a j) - ; (j = i0) → p (i ∨ ~ k) - ; (j = i1) → llem₁ G H n hLev f g p ptS i0 ptS (~ i ∧ k)}) - (g (merid a j)) - - P : PathP (λ k → PathP (λ i → A (merid a i) ptS) - (p k) (llem₁ G H n hLev f g p ptS i0 ptS (~ k))) - (λ i → mainₗ G H n hLev f g p a i ptS a ptS i) λ i → g (merid a i) - P = mainP G H n hLev f g p a i0 ptS a ◁ lem - -private - wedgeConFun : (G H : AbGroup ℓG) (k n m : ℕ) → (n + m ≡ k) → {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} - → ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y)) - → (f : (x : _) → A ptS x) - → (g : (x : _) → A x ptS) - → f ptS ≡ g ptS - → (x : _) (y : _) → A x y - wedgeconLeft : (G H : AbGroup ℓG) (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} - → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) - → (f : (x : _) → A ptS x) - → (g : (x : _) → A x ptS) - → (p : f ptS ≡ g ptS) - → (x : _) → wedgeConFun G H k n m P hLev f g p ptS x ≡ f x - wedgeconRight : (G H : AbGroup ℓG) (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} - → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) - → (f : (x : _) → A ptS x) - → (g : (x : _) → A x ptS) - → (p : f ptS ≡ g ptS) - → (x : _) → wedgeConFun G H k n m P hLev f g p x ptS ≡ g x - wedgeConFun G H k n zero P {A = A} hLev f g p = wedgeConFun' G H n hLev f g p - wedgeConFun G H k zero (suc m) P {A = A} hLev f g p x y = - wedgeConFun' H G (suc m) {A = λ x y → A y x} - (λ x y → subst (λ n → isOfHLevel (2 + n) (A y x)) (+-comm 1 m) (hLev y x)) - g f (sym p) y x - wedgeConFun G H l (suc n) (suc m) P {A = A} hlev f g p north y = f y - wedgeConFun G H l (suc n) (suc m) P {A = A} hlev f g p south y = subst (λ x → A x y) (merid ptS) (f y) - wedgeConFun G H zero (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = main** i - module _ where - main** : PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) - main** = ⊥-rec (snotz P) - wedgeConFun G H (suc l) (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = main* a y i - module _ where - lem₁* : g south ≡ (subst (λ x → A x ptS) (merid ptS) (f ptS)) - lem₁* = ((λ i → transp (λ j → A (merid ptS (~ i ∨ j)) ptS) (~ i) (g (merid ptS (~ i))))) - ∙ cong (subst (λ x → A x ptS) (merid ptS)) (sym p) - - lem₁'* : - Square - (λ i → transp (λ j → A (merid ptS (~ i ∨ j)) ptS) (~ i) (g (merid ptS (~ i)))) - (refl {x = subst (λ x → A x ptS) (merid ptS) (f ptS)}) - lem₁* - ((cong (transport (λ z → A (merid ptS z) ptS)) (sym p))) - lem₁'* i j = - hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptS (~ j ∨ j₁)) ptS) (~ j) (g (merid ptS (~ j))) - ; (i = i1) → subst (λ x → A x ptS) (merid ptS) (p (~ k)) - ; (j = i1) → cong (transport (λ z → A (merid ptS z) ptS)) (sym p) (i ∧ k)}) - (transp (λ j₁ → A (merid ptS ((~ j ∧ ~ i) ∨ j₁)) ptS) (~ j ∧ ~ i) (g (merid ptS (~ j ∧ ~ i)))) - - - lem₂* : (λ i₁ → transp (λ j → A (merid ptS (j ∧ i₁)) ptS) (~ i₁) (f ptS)) - ≡ (λ i₁ → hcomp (λ k → λ { (i₁ = i0) → p (~ k) ; (i₁ = i1) → lem₁* k }) - (g (merid ptS i₁))) - lem₂* i j = - hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptS (z ∧ j)) ptS) (~ j) (p (~ k)) - ; (j = i0) → p (~ k) - ; (j = i1) → lem₁'* k (~ i)}) - (transp (λ z → A (merid ptS ((i ∨ z) ∧ j)) ptS) (i ∨ ~ j) (g (merid ptS (i ∧ j)))) - - main* : (a : _) (y : _) → PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) - main* = - wedgeConFun G H l n (suc m) (cong predℕ P) - (λ _ _ → isOfHLevelPathP' (suc (n + (suc (suc m)))) (hlev _ _) _ _) - (λ x i → transp (λ j → A (merid ptS (j ∧ i)) x) (~ i) (f x)) - (λ y i → hcomp (λ k → λ { (i = i0) → p (~ k) - ; (i = i1) → lem₁* k}) - (g (merid y i))) - lem₂* - - mainR : (x : _) → main* x ptS - ≡ λ i → hcomp (λ k → λ { (i = i0) → p (~ k) - ; (i = i1) → lem₁* k}) - (g (merid x i)) - mainR x = - wedgeconRight G H l n (suc m) (cong predℕ P) - (λ _ _ → isOfHLevelPathP' (suc (n + (suc (suc m)))) (hlev _ _) _ _) - (λ x i → transp (λ j → A (merid ptS (j ∧ i)) x) (~ i) (f x)) - (λ y i → hcomp (λ k → λ { (i = i0) → p (~ k) - ; (i = i1) → lem₁* k}) - (g (merid y i))) - lem₂* x - wedgeconLeft G H k zero zero P {A = A} hLev f g p _ = refl - wedgeconLeft G H k (suc n) zero P {A = A} hLev f g p _ = refl - wedgeconLeft G H k zero (suc m) P {A = A} hLev f g p x = - wedgeConFun'ᵣ H G (suc m) - (λ x₁ y → - subst (λ n → (x₂ y₁ : A y x₁) → isOfHLevel (suc n) (x₂ ≡ y₁)) - (+-comm 1 m) (hLev y x₁)) - g f (λ i → p (~ i)) x - wedgeconLeft G H k (suc n) (suc m) P {A = A} hLev f g p _ = refl - wedgeconRight G H k n zero P {A = A} hLev f g p = wedgeConFun'ᵣ G H n hLev f g p - wedgeconRight G H k zero (suc m) P {A = A} hLev f g p _ = refl - wedgeconRight G H zero (suc n) (suc m) P {A = A} hLev f g p x = ⊥-rec (snotz P) - wedgeconRight G H l (suc n) (suc m) P {A = A} hLev f g p north = p - wedgeconRight G H l (suc n) (suc m) P {A = A} hLev f g p south = - sym (lem₁* G H _ n m refl hLev f g p ptS i0 ptS) - wedgeconRight G H (suc l) (suc n) (suc m) P {A = A} hLev f g p (merid a i) k = - help k i - where - lem : _ - lem i j = - hcomp (λ k → λ { (i = i1) → g (merid a j) - ; (j = i0) → p (i ∨ ~ k) - ; (j = i1) → lem₁* G H l n m P hLev f g p ptS i0 ptS (~ i ∧ k)}) - (g (merid a j)) - - help : PathP (λ k → PathP (λ i → A (merid a i) ptS) - (p k) (lem₁* G H l n m P hLev f g p ptS i0 ptS (~ k))) - (λ i → main* G H l n m P hLev f g p a i north a north i) (cong g (merid a)) - help = mainR G H l n m P hLev f g p a i0 ptS a ◁ lem - - -module wedgeConEM (G H : AbGroup ℓG) (n m : ℕ) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} - (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) - (f : (x : _) → A ptS x) - (g : (x : _) → A x ptS) - (p : f ptS ≡ g ptS) where - fun : (x : EM-raw G (suc n)) (y : EM-raw H (suc m)) → A x y - fun = wedgeConFun G H (n + m) n m refl hLev f g p - - left : (x : EM-raw H (suc m)) → fun ptS x ≡ f x - left = wedgeconLeft G H (n + m) n m refl hLev f g p - - right : (x : EM-raw G (suc n)) → fun x ptS ≡ g x - right = wedgeconRight G H (n + m) n m refl hLev f g p - -module _ {G : AbGroup ℓG} where - infixr 34 _+ₖ_ - infixr 34 _-ₖ_ - open AbGroupStr (snd G) renaming (_+_ to _+G_ ; -_ to -G_ ; assoc to assocG) - - private - help : (n : ℕ) → n + (4 + n) ≡ (2 + n) + (2 + n) - help n = +-suc n (3 + n) ∙ cong suc (+-suc n (suc (suc n))) - - hLevHelp : (n : ℕ) → isOfHLevel ((2 + n) + (2 + n)) (EM G (2 + n)) - hLevHelp n = - transport (λ i → isOfHLevel (help n i) (EM G (2 + n))) - (isOfHLevelPlus {n = 4 + n} n (isOfHLevelTrunc (4 + n))) - - helper : (g h : fst G) - → PathP (λ i → Path (EM₁ (AbGroup→Group G)) - (emloop h i) (emloop h i)) (emloop g) (emloop g) - helper g h = - compPathL→PathP - (cong (sym (emloop h) ∙_) - (sym (emloop-comp _ g h) - ∙∙ cong emloop (comm g h) - ∙∙ emloop-comp _ h g) - ∙∙ assoc _ _ _ - ∙∙ cong (_∙ emloop g) (lCancel _) - ∙ sym (lUnit _)) - open import Cubical.Foundations.Path - _+ₖ_ : {n : ℕ} → EM G n → EM G n → EM G n - _+ₖ_ {n = zero} = _+G_ - _+ₖ_ {n = suc zero} = - rec _ (isGroupoidΠ (λ _ → emsquash)) - (λ x → x) - hi! - λ g h i j x → el g h x i j - where - lol : (g h : fst G) - → Square (emloop g) (emloop (g +G h)) refl (emloop h) - lol g h = compPath-filler (emloop g) (emloop h) ▷ sym (emloop-comp _ g h) - - hi! : fst G → (λ x → x) ≡ (λ x → x) - hi! g = funExt (elimSet _ (λ _ → emsquash _ _) - (emloop g) - (helper g)) - el : (g h : fst G) (x : EM₁ (AbGroup→Group G)) - → Square (λ j → hi! g j x) (λ j → hi! ((snd (AbGroup→Group G) GroupStr.· g) h) j x) - refl λ j → hi! h j x - el g h = - elimProp _ (λ _ → isOfHLevelPathP' 1 (emsquash _ _) _ _) - (lol g h) - _+ₖ_ {n = suc (suc n)} = - trRec2 (isOfHLevelTrunc (4 + n)) - (wedgeConEM.fun G G (suc n) (suc n) - (λ _ _ → hLevHelp n) - ∣_∣ ∣_∣ refl) - - σ-EM : (n : ℕ) → EM-raw G (suc n) → Path (EM-raw G (2 + n)) ptS ptS - σ-EM n x i = (merid x ∙ sym (merid ptS)) i - - -ₖ_ : {n : ℕ} → EM G n → EM G n - -ₖ_ {n = zero} x = -G x - -ₖ_ {n = suc zero} = - rec _ emsquash - ptS - (λ g → sym (emloop g)) - λ g h → sym (emloop-sym _ g) - ◁ (flipSquare - (flipSquare (emcomp (-G g) (-G h)) - ▷ emloop-sym _ h) - ▷ (cong emloop (comm (-G g) (-G h) - ∙ sym (GroupTheory.invDistr (AbGroup→Group G) g h)) - ∙ emloop-sym _ (g +G h))) - -ₖ_ {n = suc (suc n)} = - map λ { north → north - ; south → north - ; (merid a i) → σ-EM n a (~ i)} - - _-ₖ_ : {n : ℕ} → EM G n → EM G n → EM G n - _-ₖ_ {n = n} x y = _+ₖ_ {n = n} x (-ₖ_ {n = n} y) - - +ₖ-syntax : (n : ℕ) → EM G n → EM G n → EM G n - +ₖ-syntax n = _+ₖ_ {n = n} - - -ₖ-syntax : (n : ℕ) → EM G n → EM G n - -ₖ-syntax n = -ₖ_ {n = n} - - -'ₖ-syntax : (n : ℕ) → EM G n → EM G n → EM G n - -'ₖ-syntax n = _-ₖ_ {n = n} - - syntax +ₖ-syntax n x y = x +[ n ]ₖ y - syntax -ₖ-syntax n x = -[ n ]ₖ x - syntax -'ₖ-syntax n x y = x -[ n ]ₖ y - - 0ₖ : (n : ℕ) → EM G n - 0ₖ zero = 0g - 0ₖ (suc zero) = ptS - 0ₖ (suc (suc n)) = ∣ ptS ∣ - - lUnitₖ : (n : ℕ) (x : EM G n) → 0ₖ n +[ n ]ₖ x ≡ x - lUnitₖ zero x = lid x - lUnitₖ (suc zero) _ = refl - lUnitₖ (suc (suc n)) = - trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) - λ _ → refl - - rUnitₖ : (n : ℕ) (x : EM G n) → x +[ n ]ₖ 0ₖ n ≡ x - rUnitₖ zero x = rid x - rUnitₖ (suc zero) = - elimSet _ (λ _ → emsquash _ _) - refl - λ _ _ → refl - rUnitₖ (suc (suc n)) = - trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) - (wedgeConEM.right G G (suc n) (suc n) - (λ _ _ → hLevHelp n) - ∣_∣ ∣_∣ refl) - - commₖ : (n : ℕ) (x y : EM G n) → x +[ n ]ₖ y ≡ y +[ n ]ₖ x - commₖ zero = comm - commₖ (suc zero) = - wedgeConEM.fun G G 0 0 (λ _ _ → emsquash _ _) - (λ x → sym (rUnitₖ 1 x)) - (rUnitₖ 1) - refl - commₖ (suc (suc n)) = - elim2 (λ _ _ → isOfHLevelTruncPath {n = 4 + n}) - (wedgeConEM.fun G G _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (hLevHelp n) _ _) - (λ x → sym (rUnitₖ (2 + n) ∣ x ∣)) - (λ x → rUnitₖ (2 + n) ∣ x ∣) - refl) - - open import Cubical.Homotopy.Loopspace - cong₂+₁ : (p q : typ (Ω (EM∙ G 1))) → cong₂ (λ x y → x +[ 1 ]ₖ y) p q ≡ p ∙ q - cong₂+₁ p q = - (cong₂Funct (λ x y → x +[ 1 ]ₖ y) p q) - ∙ (λ i → (cong (λ x → rUnitₖ 1 x i) p) ∙ (cong (λ x → lUnitₖ 1 x i) q)) - - cong₂+₂ : (n : ℕ) (p q : typ (Ω (EM∙ G (2 + n)))) → cong₂ (λ x y → x +[ (2 + n) ]ₖ y) p q ≡ p ∙ q - cong₂+₂ n p q = - (cong₂Funct (λ x y → x +[ (2 + n) ]ₖ y) p q) - ∙ (λ i → (cong (λ x → rUnitₖ (2 + n) x i) p) ∙ (cong (λ x → lUnitₖ (2 + n) x i) q)) - - isCommΩEM : (n : ℕ) (p q : typ (Ω (EM∙ G (suc n)))) → p ∙ q ≡ q ∙ p - isCommΩEM zero p q = - sym (cong₂+₁ p q) - ∙∙ (λ i j → commₖ 1 (p j) (q j) i) - ∙∙ cong₂+₁ q p - isCommΩEM (suc n) p q = - (sym (cong₂+₂ n p q) - ∙∙ (λ i j → commₖ (suc (suc n)) (p j) (q j) i) - ∙∙ cong₂+₂ n q p) - - cong-₁ : (p : typ (Ω (EM∙ G 1))) → cong (λ x → -[ 1 ]ₖ x) p ≡ sym p - cong-₁ p = main ptS p - where - decoder : (x : EM G 1) → ptS ≡ x → x ≡ ptS - decoder = - elimSet _ - (λ _ → isSetΠ λ _ → emsquash _ _) - (λ p i → -[ 1 ]ₖ (p i)) - λ g → toPathP - (funExt λ x → - (λ i → transport (λ i → Path (EM G 1) (emloop g i) ptS) - (cong (-ₖ_ {n = 1}) - (transp (λ j → Path (EM G 1) ptS (emloop g (~ j ∧ ~ i))) i - (compPath-filler x (sym (emloop g)) i) ))) - ∙∙ (λ i → transp (λ j → Path (EM G 1) (emloop g (i ∨ j)) ptS) i - (compPath-filler' - (sym (emloop g)) - (cong-∙ (-ₖ_ {n = 1}) - x (sym (emloop g)) i) i)) - ∙∙ (cong (sym (emloop g) ∙_) (isCommΩEM 0 (cong (-ₖ_ {n = 1}) x) (emloop g))) - ∙∙ assoc _ _ _ - ∙∙ cong (_∙ (cong (-ₖ_ {n = 1}) x)) (lCancel (emloop g)) - ∙ sym (lUnit _)) - - main : (x : EM G 1) (p : ptS ≡ x) → decoder x p ≡ sym p - main x = J (λ x p → decoder x p ≡ sym p) refl - - cong-₂ : (n : ℕ) (p : typ (Ω (EM∙ G (2 + n)))) → cong (λ x → -[ 2 + n ]ₖ x) p ≡ sym p - cong-₂ n p = main _ p - where - pp : (a : _) → PathP (λ i → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ → ∣ merid a i ∣ₕ ≡ 0ₖ (2 + n)) - (cong (λ x → -[ 2 + n ]ₖ x)) λ p → cong ∣_∣ₕ (sym (merid ptS)) ∙ cong (λ x → -[ 2 + n ]ₖ x) p - pp a = - toPathP - (funExt λ x → - (λ k → transp (λ i → Path (EM G (2 + n)) ∣ merid a (i ∨ k) ∣ ∣ ptS ∣) k (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) - (cong (-ₖ-syntax (suc (suc n))) - (transp (λ j → Path (EM G (2 + n)) ∣ ptS ∣ ∣ merid a (~ j ∧ ~ k) ∣) k (compPath-filler x (sym (cong ∣_∣ₕ (merid a))) k))) k)) - ∙∙ cong (cong ∣_∣ₕ (sym (merid a)) ∙_) (cong-∙ (λ x → -[ 2 + n ]ₖ x) x (sym (cong ∣_∣ₕ (merid a))) - ∙ isCommΩEM (suc n) (cong (λ x → -[ 2 + n ]ₖ x) x) (cong ∣_∣ₕ (σ-EM n a))) - ∙∙ (λ k → (λ i → ∣ merid a (~ i ∨ k) ∣) ∙ (λ i → ∣ compPath-filler' (merid a) (sym (merid ptS)) (~ k) i ∣) - ∙ cong (λ x → -ₖ-syntax (suc (suc n)) x) x) - ∙ sym (lUnit _)) - - decoder : (x : EM G (2 + n)) → 0ₖ (2 + n) ≡ x → x ≡ 0ₖ (2 + n) - decoder = - trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelTruncPath {n = 4 + n}) - λ { north → pp ptS i0 - ; south → pp ptS i1 - ; (merid a i) → pp a i} - - main : (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → decoder x p ≡ sym p - main x = J (λ x p → decoder x p ≡ sym p) refl - - rCancelₖ : (n : ℕ) (x : EM G n) → x +[ n ]ₖ (-[ n ]ₖ x) ≡ 0ₖ n - rCancelₖ zero x = invr x - rCancelₖ (suc zero) = - elimSet _ (λ _ → emsquash _ _) - refl - λ g → flipSquare (cong₂+₁ (emloop g) (λ i → -ₖ-syntax 1 (emloop g i)) - ∙ rCancel (emloop g)) - rCancelₖ (suc (suc n)) = - trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) - λ { north → refl - ; south i → +ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣ (-ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣) - ; (merid a i) j - → hcomp (λ r → λ { (i = i0) → 0ₖ (2 + n) - ; (i = i1) → ∣ merid ptS (~ j ∧ r) ∣ₕ -[ 2 + n ]ₖ ∣ merid ptS (~ j ∧ r) ∣ - ; (j = i0) → ∣ compPath-filler (merid a) (sym (merid ptS)) (~ r) i ∣ - -[ 2 + n ]ₖ ∣ compPath-filler (merid a) (sym (merid ptS)) (~ r) i ∣ - ; (j = i1) → 0ₖ (2 + n)}) - (help' a j i) } - where - help' : (a : _) → cong₂ (λ x y → ∣ x ∣ -[ suc (suc n) ]ₖ ∣ y ∣) (σ-EM n a) (σ-EM n a) ≡ refl - help' a = - cong₂+₂ n (cong ∣_∣ₕ (σ-EM n a)) (cong (λ x → -[ 2 + n ]ₖ ∣ x ∣) (σ-EM n a)) - ∙∙ cong (cong ∣_∣ₕ (σ-EM n a) ∙_) (cong-₂ n (cong ∣_∣ₕ (σ-EM n a))) - ∙∙ rCancel _ - - lCancelₖ : (n : ℕ) (x : EM G n) → (-[ n ]ₖ x) +[ n ]ₖ x ≡ 0ₖ n - lCancelₖ n x = commₖ n (-[ n ]ₖ x) x ∙ rCancelₖ n x - - assocₖ : (n : ℕ) (x y z : EM G n) - → (x +[ n ]ₖ (y +[ n ]ₖ z) ≡ (x +[ n ]ₖ y) +[ n ]ₖ z) - assocₖ zero = assocG - assocₖ (suc zero) = - elimSet _ (λ _ → isSetΠ2 λ _ _ → emsquash _ _) - (λ _ _ → refl) - λ g i y z k → lem g y z k i - where - lem : (g : fst G) (y z : _) → cong (λ x → x +[ suc zero ]ₖ (y +[ suc zero ]ₖ z)) (emloop g) - ≡ cong (λ x → (x +[ suc zero ]ₖ y) +[ suc zero ]ₖ z) (emloop g) - lem g = - elimProp _ (λ _ → isPropΠ λ _ → emsquash _ _ _ _) - (elimProp _ (λ _ → emsquash _ _ _ _) - refl) - assocₖ (suc (suc n)) = - elim2 (λ _ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelTruncPath {n = 4 + n}) - λ a b → trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) - (λ c → main c a b) - where - lem : (c : _) (a b : _) - → PathP (λ i → (∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ merid c i ∣ₕ) - ≡ (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ merid c i ∣ₕ)) - (cong (λ x → ∣ a ∣ₕ +[ suc (suc n) ]ₖ x) (rUnitₖ (suc (suc n)) ∣ b ∣) - ∙ sym (rUnitₖ (suc (suc n)) (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ))) - ((λ i → ∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ merid ptS (~ i) ∣ₕ)) - ∙∙ cong (λ x → ∣ a ∣ₕ +[ suc (suc n) ]ₖ x) (rUnitₖ (suc (suc n)) ∣ b ∣) - ∙ sym (rUnitₖ (suc (suc n)) (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ)) - ∙∙ λ i → (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ merid ptS i ∣ₕ) - lem c = - EM-raw-elim G (suc n) - (λ _ → isOfHLevelΠ (2 + n) (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _)) - (EM-raw-elim G (suc n) (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _) - ((sym (rUnit refl) - ◁ λ _ → refl) - ▷ (sym (lCancel (cong ∣_∣ₕ (merid ptS))) - ∙ λ i → (λ j → ∣ merid ptS (~ j ∨ ~ i) ∣ₕ) - ∙∙ lUnit (λ j → ∣ merid ptS (~ j ∧ ~ i) ∣ₕ) i - ∙∙ cong ∣_∣ₕ (merid ptS)))) - main : (c a b : _) → (∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ c ∣ₕ) ≡ (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ c ∣ₕ) - main north a b = lem ptS a b i0 - main south a b = lem ptS a b i1 - main (merid c i) a b = lem c a b i - - σ-EM' : (n : ℕ) (x : EM G (suc n)) - → Path (EM G (suc (suc n))) - (0ₖ (suc (suc n))) - (0ₖ (suc (suc n))) - σ-EM' zero x = cong ∣_∣ₕ (σ-EM zero x) - σ-EM' (suc n) = - trElim (λ _ → isOfHLevelTrunc (5 + n) _ _) - λ x → cong ∣_∣ₕ (σ-EM (suc n) x) - - σ-EM'-0ₖ : (n : ℕ) → σ-EM' n (0ₖ (suc n)) ≡ refl - σ-EM'-0ₖ zero = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) - σ-EM'-0ₖ (suc n) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) - - private - P : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (r : refl ≡ p) - → lUnit p ∙ cong (_∙ p) r - ≡ rUnit p ∙ cong (p ∙_) r - P p = J (λ p r → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r) refl - - σ-EM'-hom : (n : ℕ) → (a b : _) → σ-EM' n (a +ₖ b) ≡ σ-EM' n a ∙ σ-EM' n b - σ-EM'-hom zero = - wedgeConEM.fun G G 0 0 (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) l r p - where - l : _ - l x = cong (σ-EM' zero) (lUnitₖ 1 x) - ∙∙ lUnit (σ-EM' zero x) - ∙∙ cong (_∙ σ-EM' zero x) (sym (σ-EM'-0ₖ zero)) - - r : _ - r x = - cong (σ-EM' zero) (rUnitₖ 1 x) - ∙∙ rUnit (σ-EM' zero x) - ∙∙ cong (σ-EM' zero x ∙_) (sym (σ-EM'-0ₖ zero)) - - p : _ - p = P (σ-EM' zero embase) (sym (σ-EM'-0ₖ zero)) - σ-EM'-hom (suc n) = - elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) _ _) _ _) - (wedgeConEM.fun G G _ _ - (λ x y → transport (λ i → isOfHLevel (help n i) - ((σ-EM' (suc n) (∣ x ∣ₕ +ₖ ∣ y ∣ₕ) - ≡ σ-EM' (suc n) ∣ x ∣ₕ ∙ σ-EM' (suc n) ∣ y ∣ₕ))) - (isOfHLevelPlus {n = 4 + n} n - (isOfHLevelPath (4 + n) - (isOfHLevelTrunc (5 + n) _ _) _ _))) - (λ x → cong (σ-EM' (suc n)) (lUnitₖ (suc (suc n)) ∣ x ∣) - ∙∙ lUnit (σ-EM' (suc n) ∣ x ∣) - ∙∙ cong (_∙ σ-EM' (suc n) ∣ x ∣) (sym (σ-EM'-0ₖ (suc n)))) - (λ x → cong (σ-EM' (suc n)) (rUnitₖ (2 + n) ∣ x ∣) - ∙∙ rUnit (σ-EM' (suc n) ∣ x ∣) - ∙∙ cong (σ-EM' (suc n) ∣ x ∣ ∙_) (sym (σ-EM'-0ₖ (suc n)))) - (P (σ-EM' (suc n) (0ₖ (2 + n))) (sym (σ-EM'-0ₖ (suc n))))) - - σ-EM'-ₖ : (n : ℕ) → (a : _) → σ-EM' n (-ₖ a) ≡ sym (σ-EM' n a) - σ-EM'-ₖ n = - morphLemmas.distrMinus - (λ x y → x +[ suc n ]ₖ y) (_∙_) - (σ-EM' n) (σ-EM'-hom n) - (0ₖ (suc n)) refl - (λ x → -ₖ x) sym - (λ x → sym (lUnit x)) (λ x → sym (rUnit x)) - (lCancelₖ (suc n)) rCancel - assoc (σ-EM'-0ₖ n) - - -Dist : (n : ℕ) (x y : EM G n) → -[ n ]ₖ (x +[ n ]ₖ y) ≡ (-[ n ]ₖ x) +[ n ]ₖ (-[ n ]ₖ y) - -Dist zero x y = (GroupTheory.invDistr (AbGroup→Group G) x y) ∙ commₖ zero _ _ - -Dist (suc zero) = k - where -- useless where clause. Needed for fast type checking for some reason. - l : _ - l x = refl - - r : _ - r x = cong (λ z → -[ 1 ]ₖ z) (rUnitₖ 1 x) ∙ sym (rUnitₖ 1 (-[ 1 ]ₖ x)) - - p : r ptS ≡ l ptS - p = sym (rUnit refl) - - k = wedgeConEM.fun G G 0 0 (λ _ _ → emsquash _ _) l r (sym p) - - -Dist (suc (suc n)) = - elim2 (λ _ _ → isOfHLevelTruncPath {n = 4 + n}) - (wedgeConEM.fun G G (suc n) (suc n) - (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (hLevHelp n) _ _) - (λ x → refl) - (λ x → cong (λ z → -[ (suc (suc n)) ]ₖ z) (rUnitₖ (suc (suc n)) ∣ x ∣ₕ) ∙ sym (rUnitₖ (suc (suc n)) (-[ (suc (suc n)) ]ₖ ∣ x ∣ₕ))) - (rUnit refl)) - - addIso : (n : ℕ) (x : EM G n) → Iso (EM G n) (EM G n) - Iso.fun (addIso n x) y = y +[ n ]ₖ x - Iso.inv (addIso n x) y = y -[ n ]ₖ x - Iso.rightInv (addIso n x) y = - sym (assocₖ n y (-[ n ]ₖ x) x) - ∙∙ cong (λ x → y +[ n ]ₖ x) (lCancelₖ n x) - ∙∙ rUnitₖ n y - Iso.leftInv (addIso n x) y = - sym (assocₖ n y x (-[ n ]ₖ x)) - ∙∙ cong (λ x → y +[ n ]ₖ x) (rCancelₖ n x) - ∙∙ rUnitₖ n y - - CODE : (n : ℕ) → EM G (suc n) → TypeOfHLevel ℓG (2 + n) - CODE zero = - rec _ (isOfHLevelTypeOfHLevel 2) - (typ G , is-set) - (λ g → Σ≡Prop (λ _ → isPropIsOfHLevel 2) - (isoToPath (addIso 0 g))) - (λ g h → ΣSquareSet {B = isOfHLevel 2} - (λ _ → isSetΠ2 λ _ _ → isProp→isOfHLevelSuc 1 isPropIsProp) - (without g h)) - where - ΣSquareSet : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} (prop : (x : A) → isSet (B x)) - → {x y z w : Σ A B} {p : x ≡ y} {q : y ≡ z} {r : w ≡ z} {s : x ≡ w} - → Square (cong fst s) (cong fst q) (cong fst p) (cong fst r) - → Square s q p r - fst (ΣSquareSet prop sq i j) = sq i j - snd (ΣSquareSet {B = B} prop {p = p} {q = q} {r = r} {s = s} sq i j) = sqP i j - where - sqP : SquareP (λ i j → B (sq i j)) - (cong snd s) (cong snd q) (cong snd p) (cong snd r) - sqP = toPathP (isOfHLevelPathP' 1 (prop _) _ _ _ _) - - lem : (g h : fst G) → compEquiv (isoToEquiv (addIso 0 g)) (isoToEquiv (addIso 0 h)) - ≡ isoToEquiv (addIso 0 (snd G .AbGroupStr._+_ g h)) - lem g h = - Σ≡Prop (λ _ → isPropIsEquiv _) - (funExt λ x → sym (assocₖ 0 x g h)) - - without : (g h : fst G) → - Square (isoToPath (addIso 0 g)) - (isoToPath (addIso 0 ((snd (AbGroup→Group G) GroupStr.· g) h))) - refl (isoToPath (addIso 0 h)) - without g h = - compPathL→PathP - (sym (lUnit _) - ∙∙ sym (uaCompEquiv (isoToEquiv (addIso 0 g)) (isoToEquiv (addIso 0 h))) - ∙∙ cong ua (lem g h)) - - CODE (suc n) = - trElim (λ _ → isOfHLevelTypeOfHLevel (3 + n)) - λ { north → EM G (suc n) , hLevelEM G (suc n) - ; south → EM G (suc n) , hLevelEM G (suc n) - ; (merid a i) → fib n a i} - where - fib : (n : ℕ) → (a : EM-raw G (suc n)) - → Path (TypeOfHLevel _ (3 + n)) (EM G (suc n) , hLevelEM G (suc n)) (EM G (suc n) , hLevelEM G (suc n)) - fib zero a = Σ≡Prop (λ _ → isPropIsOfHLevel 3) - (isoToPath (addIso 1 a)) - fib (suc n) a = Σ≡Prop (λ _ → isPropIsOfHLevel (4 + n)) - (isoToPath (addIso (suc (suc n)) ∣ a ∣)) - - decode : (n : ℕ) (x : EM G (suc n)) → CODE n x .fst → 0ₖ (suc n) ≡ x - decode zero = - elimSet _ (λ _ → isSetΠ (λ _ → emsquash _ _)) - emloop main - where - main : (g : _) → PathP (λ i → CODE zero (emloop g i) .fst → 0ₖ 1 ≡ emloop g i) - emloop emloop - main g = - toPathP - (funExt λ x - → (λ j → transp (λ i → Path (EM-raw G 1) ptS (emloop g (i ∨ j))) j - (compPath-filler (emloop ((transportRefl x j) +G -G g)) (emloop g) j)) - ∙∙ sym (emloop-comp _ (x +G -G g) g) - ∙∙ cong emloop (sym (assocG x (-G g) g) ∙∙ cong (x +G_) (invl g) ∙∙ rid x)) - decode (suc n) = - trElim (λ _ → isOfHLevelΠ (4 + n) - λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) - λ { north → f n - ; south → g n - ; (merid a i) → main a i} - where - f : (n : ℕ) → _ - f n = σ-EM' n - - g : (n : ℕ) → EM G (suc n) → ∣ ptS ∣ ≡ ∣ south ∣ - g n x = σ-EM' n x ∙ cong ∣_∣ₕ (merid ptS) - - lem₂ : (n : ℕ) (a x : _) - → Path (Path (EM G (suc (suc n))) _ _) - ((λ i → cong ∣_∣ₕ (σ-EM n x) i) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) ∙ (λ i → ∣ merid a i ∣ₕ)) - (g n (EM-raw→EM G (suc n) x)) - lem₂ zero a x = - cong (f zero x ∙_) - (cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid ptS))) - ∙ cong-∙ ∣_∣ₕ (merid (ptS)) (sym (merid a))) - ∙∙ sym (assoc _ _ _) - ∙∙ cong (cong ∣_∣ₕ (merid ptS) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) - ∙ sym (rUnit _)) - lem₂ (suc n) a x = - cong (f (suc n) ∣ x ∣ ∙_) - ((cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid ptS))) - ∙ cong-∙ ∣_∣ₕ (merid (ptS)) (sym (merid a))) - ∙∙ sym (assoc _ _ _) - ∙∙ cong (cong ∣_∣ₕ (merid ptS) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) - ∙ sym (rUnit _))) - - lem : (n : ℕ) (x a : EM-raw G (suc n)) - → f n (transport (sym (cong (λ x → CODE (suc n) x .fst) (cong ∣_∣ₕ (merid a)))) (EM-raw→EM G (suc n) x)) - ≡ cong ∣_∣ₕ (σ-EM n x) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) - lem zero x a = (λ i → cong ∣_∣ₕ (merid (transportRefl x i -[ 1 ]ₖ a) ∙ sym (merid embase))) - ∙∙ σ-EM'-hom zero x (-ₖ a) - ∙∙ cong (f zero x ∙_) (σ-EM'-ₖ zero a) - lem (suc n) x a = - cong (f (suc n)) (λ i → transportRefl ∣ x ∣ i -[ 2 + n ]ₖ ∣ a ∣) - ∙∙ σ-EM'-hom (suc n) ∣ x ∣ (-ₖ ∣ a ∣) - ∙∙ cong (f (suc n) ∣ x ∣ ∙_) (σ-EM'-ₖ (suc n) ∣ a ∣) - - main : (a : _) - → PathP (λ i → CODE (suc n) ∣ merid a i ∣ₕ .fst - → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ) (f n) (g n) - main a = - toPathP (funExt - (EM-elim _ (λ _ → isOfHLevelPathP (2 + (suc n)) (isOfHLevelTrunc (4 + n) _ _) _ _) - λ x → - (λ i → transp (λ j → Path (EM G (2 + n)) ∣ ptS ∣ ∣ merid a (i ∨ j) ∣) - i (compPath-filler (lem n x a i) (cong ∣_∣ₕ (merid a)) i) ) - ∙∙ sym (assoc _ _ _) - ∙∙ lem₂ n a x)) - - private - ptCode0 : (n : ℕ) → CODE n (0ₖ (suc n)) .fst - ptCode0 zero = 0g - ptCode0 (suc n) = 0ₖ (suc n) - - encode : (n : ℕ) (x : EM G (suc n)) → 0ₖ (suc n) ≡ x → CODE n x .fst - encode n x p = subst (λ x → CODE n x .fst) p (ptCode0 n) - - open import Cubical.Foundations.Transport - - encode-morph : (n : ℕ) → (p q : 0ₖ (suc (suc n)) ≡ 0ₖ (suc (suc n))) - → (encode (suc n) _ (p ∙ q)) ≡ (encode (suc n) _ p) +[ suc n ]ₖ (encode (suc n) _ q) - encode-morph n p q = - substComposite (λ x → CODE (suc n) x .fst) p q (0ₖ (suc n)) - ∙∙ {!!} - ∙∙ {!!} - - decode-encode : (n : ℕ) (x : EM G (suc n)) (p : 0ₖ (suc n) ≡ x) → decode n x (encode n x p) ≡ p - decode-encode zero x = - J (λ x p → decode zero x (encode zero x p) ≡ p) - (cong emloop (transportRefl 0g) ∙ emloop-1g (AbGroup→Group G)) - decode-encode (suc zero) x = - J (λ x p → decode 1 x (encode 1 x p) ≡ p) - (σ-EM'-0ₖ 0) - decode-encode (suc (suc n)) x = - J (λ x p → decode (2 + n) x (encode (2 + n) x p) ≡ p) - (σ-EM'-0ₖ (suc n)) - - encode-decode : (n : ℕ) (x : _) → encode n (0ₖ (suc n)) (decode n (0ₖ (suc n)) x) ≡ x - encode-decode zero x i = transportRefl (lid x i) i - encode-decode (suc zero) x = - cong (encode 1 (0ₖ 2)) (cong-∙ ∣_∣ₕ (merid x) (sym (merid ptS))) - ∙∙ substComposite (λ x → CODE (suc zero) x .fst) - (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptS))) ptS - ∙∙ cong (subst (λ x₁ → CODE 1 x₁ .fst) (λ i → ∣ merid ptS (~ i) ∣ₕ)) - (λ i → lUnitₖ 1 (transportRefl x i) i) - ∙ (λ i → rUnitₖ 1 (transportRefl x i) i) - encode-decode (suc (suc n)) = - trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) - λ x → cong (encode (2 + n) (0ₖ (3 + n))) (cong-∙ ∣_∣ₕ (merid x) (sym (merid ptS))) - ∙∙ substComposite (λ x → CODE (suc (suc n)) x .fst) - (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptS))) (0ₖ (2 + n)) - ∙∙ cong (subst (λ x₁ → CODE (2 + n) x₁ .fst) (λ i → ∣ merid ptS (~ i) ∣ₕ)) - (λ i → lUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) - ∙ (λ i → rUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) - - EM∙' : (n : ℕ) → Pointed _ - EM∙' n = EM G n , 0ₖ n - - Iso-EM-ΩEM+1 : (n : ℕ) → Iso (EM G n) (typ (Ω (EM∙' (suc n)))) - Iso.fun (Iso-EM-ΩEM+1 zero) = decode zero (0ₖ 1) - Iso.inv (Iso-EM-ΩEM+1 zero) = encode zero (0ₖ 1) - Iso.rightInv (Iso-EM-ΩEM+1 zero) = decode-encode zero (0ₖ 1) - Iso.leftInv (Iso-EM-ΩEM+1 zero) = encode-decode zero - Iso.fun (Iso-EM-ΩEM+1 (suc zero)) = decode 1 (0ₖ 2) - Iso.inv (Iso-EM-ΩEM+1 (suc zero)) = encode 1 (0ₖ 2) - Iso.rightInv (Iso-EM-ΩEM+1 (suc zero)) = decode-encode 1 (0ₖ 2) - Iso.leftInv (Iso-EM-ΩEM+1 (suc zero)) = encode-decode 1 - Iso.fun (Iso-EM-ΩEM+1 (suc (suc n))) = decode (2 + n) (0ₖ (3 + n)) - Iso.inv (Iso-EM-ΩEM+1 (suc (suc n))) = encode (2 + n) (0ₖ (3 + n)) - Iso.rightInv (Iso-EM-ΩEM+1 (suc (suc n))) = decode-encode (2 + n) (0ₖ (3 + n)) - Iso.leftInv (Iso-EM-ΩEM+1 (suc (suc n))) = encode-decode (2 + n) - - EM→ΩEM+1 : (n : ℕ) → EM G n → typ (Ω (EM∙' (suc n))) - EM→ΩEM+1 n = Iso.fun (Iso-EM-ΩEM+1 n) - - ΩEM+1→EM : (n : ℕ) → typ (Ω (EM∙' (suc n))) → EM G n - ΩEM+1→EM n = Iso.inv (Iso-EM-ΩEM+1 n) - - EM→ΩEM+1-hom : (n : ℕ) (x y : EM G n) - → EM→ΩEM+1 n (x +[ n ]ₖ y) ≡ EM→ΩEM+1 n x ∙ EM→ΩEM+1 n y - EM→ΩEM+1-hom zero x y = emloop-comp _ x y - EM→ΩEM+1-hom (suc zero) x y = σ-EM'-hom zero x y - EM→ΩEM+1-hom (suc (suc n)) x y = σ-EM'-hom (suc n) x y - - ΩEM+1→EM-hom : (n : ℕ) → (p q : typ (Ω (EM∙' (suc n)))) - → ΩEM+1→EM n (p ∙ q) ≡ (ΩEM+1→EM n p) +[ n ]ₖ (ΩEM+1→EM n q) - ΩEM+1→EM-hom n = - morphLemmas.isMorphInv (λ x y → x +[ n ]ₖ y) (_∙_) (EM→ΩEM+1 n) - (EM→ΩEM+1-hom n) (ΩEM+1→EM n) - (Iso.rightInv (Iso-EM-ΩEM+1 n)) (Iso.leftInv (Iso-EM-ΩEM+1 n)) - - EM→ΩEM+1-0ₖ : (n : ℕ) → EM→ΩEM+1 n (0ₖ n) ≡ refl - EM→ΩEM+1-0ₖ zero = emloop-1g _ - EM→ΩEM+1-0ₖ (suc zero) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) - EM→ΩEM+1-0ₖ (suc (suc n)) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) - - ΩEM+1→EM-refl : (n : ℕ) → ΩEM+1→EM n refl ≡ 0ₖ n - ΩEM+1→EM-refl zero = transportRefl 0g - ΩEM+1→EM-refl (suc zero) = refl - ΩEM+1→EM-refl (suc (suc n)) = refl From 7c8e10f3b807492836e4e776bd1d2c22da3ee717 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 23 Jul 2021 13:23:05 +0200 Subject: [PATCH 19/30] stuff --- Cubical/Algebra/AbGroup/TensorProduct.agda | 634 +++++------------- .../Algebra/Group/EilenbergMacLane/Base.agda | 9 +- .../Group/EilenbergMacLane/CupProduct.agda | 412 ++++++++++++ .../Group/EilenbergMacLane/Properties.agda | 264 ++++---- .../EilenbergMacLane/WedgeConnectivity.agda | 20 +- Cubical/Algebra/Group/MorphismProperties.agda | 37 +- Cubical/Foundations/GroupoidLaws.agda | 9 + Cubical/Foundations/Pointed/Homogeneous.agda | 51 +- 8 files changed, 813 insertions(+), 623 deletions(-) create mode 100644 Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda diff --git a/Cubical/Algebra/AbGroup/TensorProduct.agda b/Cubical/Algebra/AbGroup/TensorProduct.agda index 74be6bf8f..949167fea 100644 --- a/Cubical/Algebra/AbGroup/TensorProduct.agda +++ b/Cubical/Algebra/AbGroup/TensorProduct.agda @@ -181,8 +181,8 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where (slick y)) (slick x) - rCancelPrim : (x : B) → (0A ⊗ x) ≡ 0⊗ - rCancelPrim x = + lCancelPrim : (x : B) → (0A ⊗ x) ≡ 0⊗ + lCancelPrim x = (λ i → rid strA 0A (~ i) ⊗ x) ∙∙ linl 0A 0A x ∙∙ cong ((0A ⊗ x) +⊗_) (cong (_⊗ x) (sym (GroupTheory.inv1g (AbGroup→Group AGr))) @@ -190,6 +190,14 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where ∙∙ sym (linr 0A x (-B x)) ∙∙ (λ i → 0A ⊗ (invr strB x i)) + rCancelPrim : (x : A) → (x ⊗ 0B) ≡ 0⊗ + rCancelPrim x = + (λ i → x ⊗ rid strB 0B (~ i)) + ∙∙ linr x 0B 0B + ∙∙ cong ((x ⊗ 0B) +⊗_) (cong (x ⊗_) (sym (GroupTheory.inv1g (AbGroup→Group BGr))) ∙ flip x 0B) + ∙∙ sym (linl x (-A x) 0B) + ∙∙ (λ i → (invr strA x i) ⊗ 0B) + ⊗rCancel : (x : AGr ⨂₁ BGr) → (x +⊗ -⊗ x) ≡ 0⊗ ⊗rCancel = ⊗elimPropCool (λ _ → ⊗squash _ _) h @@ -202,7 +210,7 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where _+⊗_ ⊗assoc ⊗comm ∙∙ cong₂ _+⊗_ (sym (linl (fst x) (-A (fst x)) (snd x)) ∙∙ (λ i → invr strA (fst x) i ⊗ (snd x)) - ∙∙ rCancelPrim (snd x)) + ∙∙ lCancelPrim (snd x)) (h x₁) ∙∙ ⊗rUnit 0⊗ @@ -236,7 +244,7 @@ open import Cubical.Algebra.Group.MorphismProperties _* : AbGroup ℓ → Group ℓ _* = AbGroup→Group -module _ (AGr : Group ℓ) (BGr : AbGroup ℓ') where +module _ {ℓ ℓ' : Level} (AGr : Group ℓ) (BGr : AbGroup ℓ') where private strA = snd AGr strB = snd BGr @@ -401,454 +409,170 @@ module UP (AGr : AbGroup ℓ) (BGr : AbGroup ℓ') where (funExt (⊗elimProp (λ _ → is-set (snd C) _ _) (λ _ _ → refl) λ x y ind1 ind2 → cong₂ (_+G_ (snd C)) ind1 ind2 ∙ sym (IsGroupHom.pres· p x y))) - --- CODE : (z x : FMSet (A × B)) (y : _⨂₁_) → TypeOfHLevel (ℓ-max ℓ ℓ') 1 --- JFun : (z x : {!!}) → {!!} → {!!} --- JFun = {!!} --- CODE z x (inc x₁) = Path _⨂₁_ (inc (z ++ x)) (inc (z ++ x₁)) , ⊗squash _ _ --- CODE z x (unit i) = {!!} --- CODE z x (lin-l x₁ y b i) = {!!} --- where --- h : (z : _) → Path (TypeOfHLevel (ℓ-max ℓ ℓ') 1) --- (Path _⨂₁_ (inc (z ++ x)) (inc (z ++ [ (x₁ +A y) , b ])) , --- ⊗squash (inc (z ++ x)) (inc (z ++ [ (x₁ +A y) , b ]))) --- (Path _⨂₁_ (inc (z ++ x)) --- (inc (z ++ ((x₁ , b) ∷ [ y , b ]))) --- , ⊗squash (inc (z ++ x)) (inc (z ++ ((x₁ , b) ∷ [ y , b ])))) --- h = ElimProp.f {!λ xs → ?!} --- (Σ≡Prop {!!} {!!}) --- λ m {xs} p → Σ≡Prop {!!} (cong (Path _⨂₁_ (inc ((m ∷ xs) ++ x))) {!p!}) -- λ m xss → Σ≡Prop {!!} (cong (Path _⨂₁_ (inc ((m ∷ xs) ++ x))) {!!}) -- Σ≡Prop {!!} {!!} --- CODE z x (lin-r x₁ y z₁ i) = {!!} --- CODE z x (⊗squash y y₁ x₁ y₂ i i₁) = {!!} - - --- inclem : (x y z : FMSet (A × B)) → Path _⨂₁_ (inc x) (inc y) → Path _⨂₁_ (inc (z ++ x)) (inc (z ++ y)) --- inclem x y z = {!!} - --- _+⊗_ : _⨂₁_ → _⨂₁_ → _⨂₁_ --- inc x +⊗ inc x₁ = inc (x ++ x₁) --- inc x +⊗ unit i = {!!} --- inc x +⊗ lin-l y z b i = {!!} --- where --- lem : (x : _) → inc (x ++ [ (y +A z) , b ]) ≡ inc (x ++ ((y , b) ∷ [ z , b ])) --- lem = --- ElimProp.f {!!} --- (lin-l y z b) --- λ x {xs} p → inclem _ _ [ x ] p -- inclem _ _ {!x ∷ xs!} p --- inc x +⊗ lin-r x₁ y z i = {!!} --- inc x +⊗ ⊗squash y y₁ x₁ y₂ i i₁ = {!!} --- unit i +⊗ y = {!!} --- lin-l x y₁ b i +⊗ y = {!!} --- lin-r x y₁ z i +⊗ y = {!!} --- ⊗squash x x₁ x₂ y₁ i i₁ +⊗ y = {!!} - --- -- Word : (A : Type ℓ) → Type ℓ --- -- Word A = List (A ⊎ A) - --- -- module _ {A : Type ℓ} where --- -- pm = A ⊎ A - --- -- flip : pm → pm --- -- flip (inl x) = inr x --- -- flip (inr x) = inl x - --- -- flip-flip : (x : pm) → flip (flip x) ≡ x --- -- flip-flip (inl x) = refl --- -- flip-flip (inr x) = refl - --- -- wordFlip : Word A → Word A --- -- wordFlip = map flip - --- -- wordInverse : Word A → Word A --- -- wordInverse = rev ∘ wordFlip - --- -- map++ : {B : Type ℓ} (v w : Word A) --- -- → (f : _ → B) → map f (v ++ w) --- -- ≡ map f v ++ map f w --- -- map++ [] w f = refl --- -- map++ (x ∷ v) w f = cong (f x ∷_) (map++ v w f) - --- -- wordInverseInverse : (x : Word A) → wordInverse (wordInverse x) ≡ x --- -- wordInverseInverse [] = refl --- -- wordInverseInverse (x ∷ x₁) = --- -- cong rev (map++ (rev (map flip x₁)) (flip x ∷ []) flip) --- -- ∙∙ rev-++ (map flip (rev (map flip x₁))) (flip (flip x) ∷ []) --- -- ∙∙ cong₂ _∷_ (flip-flip x) (wordInverseInverse x₁) - --- -- wordExp : A → ℤ → Word A --- -- wordExp a (pos zero) = [] --- -- wordExp a (pos (suc n)) = inl a ∷ wordExp a (pos n) --- -- wordExp a (negsuc zero) = [ inr a ] --- -- wordExp a (negsuc (suc n)) = inr a ∷ wordExp a (negsuc n) - --- -- module _ {A : Type ℓ} (G : Group ℓ') (f : A → fst G) where --- -- private --- -- str = snd G --- -- open GroupStr str renaming (_·_ to _·G_) - --- -- PlusMinus-extendᴳ : A ⊎ A → fst G --- -- PlusMinus-extendᴳ (inl x) = f x --- -- PlusMinus-extendᴳ (inr x) = inv (f x) - --- -- PlusMinus-extendᴳ-flip : ∀ x → PlusMinus-extendᴳ (flip x) --- -- ≡ inv (PlusMinus-extendᴳ x) --- -- PlusMinus-extendᴳ-flip (inl x) = refl --- -- PlusMinus-extendᴳ-flip (inr x) = sym (GroupTheory.invInv G (f x)) - --- -- Word-extendᴳ : Word A → fst G --- -- Word-extendᴳ = foldr _·G_ 1g ∘ map PlusMinus-extendᴳ - --- -- Word-extendᴳ-++ : ∀ x y → Word-extendᴳ (x ++ y) --- -- ≡ (Word-extendᴳ x ·G Word-extendᴳ y) --- -- Word-extendᴳ-++ [] y = sym (lid (Word-extendᴳ y)) --- -- Word-extendᴳ-++ (x ∷ x₁) y = --- -- cong (PlusMinus-extendᴳ x ·G_) (Word-extendᴳ-++ x₁ y) --- -- ∙∙ assoc _ _ _ --- -- ∙∙ cong (_·G Word-extendᴳ y) (help x x₁) --- -- where --- -- help : (x : _) (y : _) --- -- → PlusMinus-extendᴳ x ·G Word-extendᴳ y ≡ Word-extendᴳ (x ∷ y) --- -- help x [] = refl --- -- help x (z ∷ y) = refl - - --- -- module _ {A : Type ℓ} (G : AbGroup ℓ') where --- -- private --- -- str = snd G --- -- open AbGroupStr str renaming (_+_ to _+G_ ; -_ to -G_) --- -- G' = AbGroup→Group G - --- -- PlusMinus-extendᴳ-comp : (f g : A → fst G) (x : _) --- -- → PlusMinus-extendᴳ G' (λ a → f a +G g a) x --- -- ≡ (PlusMinus-extendᴳ G' f x +G PlusMinus-extendᴳ G' g x) --- -- PlusMinus-extendᴳ-comp f g (inl x) = refl --- -- PlusMinus-extendᴳ-comp f g (inr x) = --- -- GroupTheory.invDistr G' (f x) (g x) ∙ comm _ _ - --- -- Word-extendᴳ-comp : (f g : A → fst G) (x : _) → Word-extendᴳ G' (λ a → f a +G g a) x --- -- ≡ (Word-extendᴳ G' f x +G Word-extendᴳ G' g x) --- -- Word-extendᴳ-comp f g [] = sym (lid 0g) --- -- Word-extendᴳ-comp f g (x ∷ x₁) = --- -- cong₂ _+G_ (PlusMinus-extendᴳ-comp f g x) --- -- (Word-extendᴳ-comp f g x₁) --- -- ∙∙ sym (assoc pm-fx pm-gx _) --- -- ∙∙ cong (pm-fx +G_) (assoc _ _ _ --- -- ∙∙ cong (_+G we-gx) (comm _ _) --- -- ∙∙ sym (assoc _ _ _)) --- -- ∙∙ assoc pm-fx we-fx _ --- -- ∙∙ cong₂ _+G_ (main x x₁ f) (main x x₁ g) --- -- where --- -- main : (x : _) (y : _) (f : _) --- -- → (PlusMinus-extendᴳ G' f x +G Word-extendᴳ G' f y) --- -- ≡ Word-extendᴳ G' f (x ∷ y) --- -- main x [] f = refl --- -- main x (x₁ ∷ y) f = refl - --- -- pm-fx = PlusMinus-extendᴳ G' f x --- -- pm-gx = PlusMinus-extendᴳ G' g x --- -- we-fx = Word-extendᴳ G' f x₁ --- -- we-gx = Word-extendᴳ G' g x₁ - --- -- module GeneratedGroup (A : Type ℓ) (R : Rel (Word A) (Word A) ℓ') where --- -- data QuotWordRel : Word A → Word A → Type (ℓ-max ℓ ℓ') where --- -- qwr-refl : ∀ {x} → QuotWordRel x x --- -- qwr-trans : ∀ {x y z} → QuotWordRel x y → QuotWordRel y z → QuotWordRel x z --- -- qwr-sym : ∀ {x y} → QuotWordRel x y → QuotWordRel y x --- -- qwr-cong : ∀ {x y z w} → QuotWordRel x y → QuotWordRel z w → QuotWordRel (x ++ z) (y ++ w) --- -- qwr-flip-r : ∀ x → QuotWordRel (x ∷ flip x ∷ []) [] --- -- qwr-rel : ∀ {x y} → R x y → QuotWordRel x y - --- -- qwr-cong-refl : {!!} --- -- qwr-cong-refl = {!!} - --- -- qwr-cong-l : {!!} --- -- qwr-cong-l = {!!} - - - - - --- -- {- --- -- module _ {ℓ : Level} where --- -- ℤmult : {A : Type ℓ} (_+_ : A → A → A) (-A_ : A → A) (0A : A) (n : ℤ) → A → A --- -- ℤmult _+_ -A_ 0A (pos zero) a = 0A --- -- ℤmult _+_ -A_ 0A (pos (suc n)) a = a + ℤmult _+_ -A_ 0A (pos n) a --- -- ℤmult _+_ -A_ 0A (negsuc zero) a = -A a --- -- ℤmult _+_ -A_ 0A (negsuc (suc n)) a = (-A a) + ℤmult _+_ -A_ 0A (negsuc n) a - --- -- data commList (A : Type ℓ) : Type ℓ where --- -- [] : commList A --- -- _∷_ : A → commList A → commList A --- -- isComm : (a b : A) (w : commList A) → a ∷ b ∷ w ≡ b ∷ a ∷ w --- -- squashCommList : isSet (commList A) - --- -- _++'_ : {A : Type ℓ} → commList A → commList A → commList A --- -- [] ++' y = y --- -- (x ∷ x₁) ++' y = x ∷ (x₁ ++' y) --- -- isComm a b w i ++' y = isComm a b (w ++' y) i --- -- squashCommList x y p q i j ++' z = --- -- squashCommList _ _ (λ j → p j ++' z) (λ j → q j ++' z) i j - --- -- propElimCommList : ∀ {ℓ'} {A : Type ℓ} {B : (x : commList A) → Type ℓ'} --- -- → ((x : _) → isProp (B x)) --- -- → B [] --- -- → ((x : A) (y : commList A) → B y → B (x ∷ y)) --- -- → (x : _) → B x --- -- propElimCommList {B = B} prop c l [] = c --- -- propElimCommList {B = B} prop c l (x ∷ x₁) = l x x₁ (propElimCommList {B = B} prop c l x₁) --- -- propElimCommList {B = B} prop c l (isComm a b x i) = --- -- isOfHLevel→isOfHLevelDep 1 prop --- -- (l a (b ∷ x) (l b x (propElimCommList prop c l x))) (l b (a ∷ x) (l a x (propElimCommList prop c l x))) --- -- (isComm a b x) i --- -- propElimCommList {B = B} prop c l (squashCommList x y p q i j) = --- -- isOfHLevel→isOfHLevelDep 2 {B = B} (λ _ → isProp→isSet (prop _)) --- -- _ _ (λ j → propElimCommList prop c l (p j)) --- -- (λ j → propElimCommList prop c l (q j)) --- -- (squashCommList x y p q) i j - --- -- propElimCommList' : --- -- ∀ {ℓ'} {A : Type ℓ} {B : (x : commList A) → Type ℓ'} --- -- → ((x : _) → isProp (B x)) --- -- → B [] --- -- → ((x : A) → B (x ∷ [])) --- -- → ((x : A) (y : commList A) → B (x ∷ []) → B y → B (x ∷ y)) --- -- → (x : _) → B x --- -- propElimCommList' {B = B} prop b ind1 ind2 = --- -- propElimCommList prop b h --- -- where --- -- h : (x : _) (y : commList _) → B y → B (x ∷ y) --- -- h x [] _ = ind1 x --- -- h x (x₁ ∷ y) l = ind2 x (x₁ ∷ y) (h x [] b) l --- -- h x (isComm a b y i) = K i --- -- where --- -- K : PathP (λ i → B (isComm a b y i) → B (x ∷ isComm a b y i)) --- -- (ind2 x (a ∷ (b ∷ y)) (ind1 x)) (ind2 x (b ∷ (a ∷ y)) (ind1 x)) --- -- K = isProp→PathP (λ _ → isPropΠ λ _ → prop _) _ _ --- -- h x (squashCommList w z p q i j) = K i j --- -- where --- -- K : SquareP (λ i j → B (squashCommList w z p q i j) → B (x ∷ squashCommList w z p q i j)) --- -- (λ j → h x (p j)) --- -- (λ j → h x (q j)) --- -- (λ _ → h x w) --- -- λ _ → h x z --- -- K = toPathP (isOfHLevelPathP' 1 (isProp→isSet (isPropΠ (λ _ → prop _))) _ _ _ _) - --- -- ++'-assoc : {A : Type ℓ} → (x y z : commList A) → x ++' (y ++' z) ≡ ((x ++' y) ++' z) --- -- ++'-assoc = --- -- propElimCommList' --- -- (λ _ → isPropΠ2 λ _ _ → squashCommList _ _) --- -- (λ _ _ → refl) --- -- (λ x → propElimCommList' (λ _ → isPropΠ λ _ → squashCommList _ _) --- -- (λ _ → refl) --- -- (λ _ _ → refl) --- -- λ x y p q → propElimCommList' (λ _ → squashCommList _ _) --- -- refl --- -- (λ _ → refl) --- -- λ z w P Q → refl) --- -- λ x y p q z w → cong (x ∷_) (q z w) - --- -- ++'-comm : {A : Type ℓ} → (x y : commList A) → x ++' y ≡ y ++' x --- -- ++'-comm = --- -- propElimCommList' (λ _ → isPropΠ λ _ → squashCommList _ _) --- -- (propElimCommList' (λ _ → squashCommList _ _) --- -- refl --- -- (λ _ → refl) --- -- (λ x y p q → cong (x ∷_) q)) --- -- (λ x → propElimCommList' --- -- (λ _ → squashCommList _ _) --- -- refl --- -- (λ _ → isComm _ _ _) --- -- λ y w p q → isComm x y w ∙ cong (y ∷_) q) --- -- λ x y p q r → p (y ++' r) --- -- ∙∙ cong (_++' (x ∷ [])) (q r) --- -- ∙∙ (sym (++'-assoc r y (x ∷ [])) ∙ cong (r ++'_) (q (x ∷ []))) - --- -- len : {A : Type ℓ} → (x : commList A) → ℕ --- -- len [] = 0 --- -- len (x ∷ x₁) = ℕ._+_ 1 (len x₁) --- -- len (isComm a b x i) = suc (suc (len x)) --- -- len (squashCommList x x₁ x₂ y i i₁) = --- -- isSetℕ (len x) (len x₁) (λ i → len (x₂ i)) (λ i → len (y i)) i i₁ - --- -- module _ (A : AbGroup ℓ) (B : AbGroup ℓ') where --- -- private --- -- strA = snd A --- -- strB = snd B --- -- open AbGroupStr renaming (_+_ to +G ; -_ to -G) - --- -- _⨂₁-raw2_ : Type _ --- -- _⨂₁-raw2_ = List (fst A × fst B) - --- -- A' = fst A --- -- B' = fst B - --- -- data _⨂₁_ : Type (ℓ-max ℓ ℓ') where --- -- inc : commList (A' × B') → _⨂₁_ --- -- bilinl : (x y : A') (b : B') (w : _) --- -- → inc ((+G strA x y , b) ∷ w) ≡ inc (((x , b) ∷ w) ++' ((y , b) ∷ w)) --- -- bilinr : (x : A') (y z : B') (w : _) --- -- → inc ((x , +G strB y z) ∷ w) ≡ inc (((x , y) ∷ w) ++' ((x , z) ∷ w)) --- -- l : (x : A') (y : B') → inc ((x , y) ∷ []) ≡ inc [] --- -- ⊗squash : isSet _⨂₁_ - --- -- +⊗ : _⨂₁_ → _⨂₁_ → _⨂₁_ --- -- +⊗ x y = {!!} - --- -- -- (inc x) (inc y) = inc (x ++' y) --- -- -- +⊗ (inc x) (bilinl y z w r i) = commListInd x i --- -- -- where --- -- -- commListInd : (x : _) → inc (x ++' ((+G strA y z , w) ∷ r)) ≡ inc (x ++' ((y , w) ∷ (r ++' ((z , w) ∷ r)))) --- -- -- commListInd = --- -- -- propElimCommList' (λ _ → ⊗squash _ _) --- -- -- (bilinl y z w r) --- -- -- (λ l → (λ i → inc ((l ∷ []) ++' (((+G strA y z , w) ∷ []) ++' r))) ∙∙ {!((+G strA y z , w₁) ∷ r)!} ∙∙ {!!}) --- -- -- λ x y p q → {!!} --- -- -- ∙∙ {!!} --- -- -- ∙∙ {!!} - --- -- -- lem : ((r ++' x) ++' ((z , w) ∷ (r ++' x))) ≡ ((r ++' ((z , w) ∷ r)) ++' x) --- -- -- lem = {!!} --- -- -- ∙∙ {!!} --- -- -- ∙∙ ++'-assoc r ((z , w) ∷ r) x --- -- -- help : inc (x ++' ((+G strA y z , w) ∷ r)) ≡ inc (x ++' ((y , w) ∷ (r ++' ((z , w) ∷ r)))) --- -- -- help = cong inc (++'-comm x _) --- -- -- ∙∙ bilinl y z w (r ++' x) --- -- -- ∙∙ (λ i → inc (((y , w) ∷ []) ++' ((r ++' x) ++' (((z , w) ∷ []) ++' (r ++' x))))) --- -- -- ∙∙ cong inc (cong (((y , w) ∷ []) ++'_) {!!}) --- -- -- ∙∙ {!!} --- -- -- ∙∙ {!!} --- -- -- ∙∙ cong inc (cong (((y , w) ∷ []) ++'_) {!!}) --- -- -- ∙∙ (λ i → inc (((y , w) ∷ []) ++' ((r ++' (((z , w) ∷ []) ++' r)) ++' x))) --- -- -- ∙∙ cong inc (sym (++'-comm x ((y , w) ∷ (r ++' ((z , w) ∷ r))))) --- -- -- +⊗ (inc x) (bilinr y z w r i) = {!!} --- -- -- +⊗ (inc x) (⊗squash y y₁ x₁ y₂ i i₁) = {!!} --- -- -- +⊗ (bilinl x y z w i) r = {!!} --- -- -- +⊗ (bilinr x y₁ z w i) y = {!!} --- -- -- +⊗ (⊗squash x x₁ x₂ y₁ i i₁) y = {!!} - - - --- -- -- -- _⊗_ : A' → B' → _⨂₁-raw_ --- -- -- -- a ⊗ b = inc [ a , b ] - --- -- -- -- invList : _⨂₁-raw2_ → _⨂₁-raw2_ --- -- -- -- invList [] = [] --- -- -- -- invList (x ∷ x₁) = --- -- -- -- ((-G strA (fst x)) , (snd x)) ∷ invList x₁ - - - --- -- -- -- _ℤ∙_ : ℤ → _⨂₁-raw2_ → _⨂₁-raw2_ --- -- -- -- pos zero ℤ∙ y = [] --- -- -- -- pos (suc n) ℤ∙ y = y ++ (pos n ℤ∙ y) --- -- -- -- negsuc zero ℤ∙ y = invList y --- -- -- -- negsuc (suc n) ℤ∙ y = (invList y) ++ (negsuc n ℤ∙ y) - --- -- -- -- 0⊗ₗ : (b : B') → 0g strA ⊗ b ≡ (0g strA ⊗ 0g strB) --- -- -- -- 0⊗ₗ b = {!!} --- -- -- -- where --- -- -- -- lem : (0g strA ⊗ b) ≡ inc ([ 0g strA , b ] ++ [ 0g strA , b ]) --- -- -- -- lem = cong (_⊗ b) (sym (rid strA (0g strA))) ∙ bilinl (0g strA) (0g strA) b - --- -- -- -- module _ {A : AbGroup ℓ} {B : AbGroup ℓ'} where --- -- -- -- private --- -- -- -- strA = snd A --- -- -- -- strB = snd B --- -- -- -- open AbGroupStr renaming (_+_ to +G ; -_ to -G) - --- -- -- -- rec⨂₁ : ∀ {ℓ'} {C : Type ℓ'} --- -- -- -- → isSet C --- -- -- -- → (c : C) --- -- -- -- → (f : fst A → fst B → C → C) --- -- -- -- → (((x y : fst A) (z : fst B) → f (+G strA x y) z c ≡ f x z (f y z c))) --- -- -- -- → ((x : fst A) (y z : fst B) → f x (+G strB y z) c ≡ f x y (f x z c)) --- -- -- -- → (x : A ⨂₁-raw B) → C --- -- -- -- rec⨂₁ {ℓ'} {C} set c f bil bir x = {!!} --- -- -- -- -- set _ _ (λ j → rec⨂₁ set c f bil bir (p j)) (λ j → rec⨂₁ set c f bil bir (q j)) i j - --- -- -- -- elimProp : ∀ {ℓ'} {C : A ⨂₁-raw B → Type ℓ'} --- -- -- -- → ((x : _) → isProp (C x)) --- -- -- -- → ((x : A ⨂₁-raw2 B) → C (inc x)) --- -- -- -- → (x : _) → C x --- -- -- -- elimProp {C = C} prop ind (inc x) = ind x --- -- -- -- elimProp {C = C} prop ind (bilinl x y z i) = --- -- -- -- isOfHLevel→isOfHLevelDep 1 {B = C} prop --- -- -- -- (ind (((+G strA x y) , z) ∷ [])) --- -- -- -- (ind ((x , z) ∷ (y , z) ∷ [])) (bilinl x y z) i --- -- -- -- elimProp {C = C} prop ind (bilinr x y z i) = --- -- -- -- isOfHLevel→isOfHLevelDep 1 {B = C} prop --- -- -- -- (ind ((x , (+G strB y z)) ∷ [])) --- -- -- -- (ind ((x , y) ∷ (x , z) ∷ [])) (bilinr x y z) i --- -- -- -- elimProp {C = C} prop ind (⊗squash x y p q i j) = --- -- -- -- isOfHLevel→isOfHLevelDep 2 {B = C} (λ x → isProp→isSet (prop x)) _ _ --- -- -- -- (λ j → elimProp prop ind (p j)) --- -- -- -- (λ j → elimProp prop ind (q j)) --- -- -- -- (⊗squash x y p q) i j - --- -- -- -- AA = fst A --- -- -- -- BB = fst B - --- -- -- -- bilinRGen : (x : fst A) (y z : fst B) (w : _) → Path (A ⨂₁-raw B) (inc ((x , +G strB y z) ∷ w)) (inc ((x , y) ∷ (x , z) ∷ w)) --- -- -- -- bilinRGen x y z [] = bilinr x y z --- -- -- -- bilinRGen x y z (x₁ ∷ w) = {!bilinRGen x y z w!} - --- -- -- -- pre-comm : (x y : A ⨂₁-raw2 B) → Path (A ⨂₁-raw B) (inc (x ++ y)) (inc (y ++ x)) --- -- -- -- pre-comm [] [] = refl --- -- -- -- pre-comm [] (x ∷ y) = {!!} -- cong inc (cong (x ∷_) {!pre-comm [] y!}) --- -- -- -- pre-comm (x ∷ x₁) y = {!!} - --- -- -- -- +⊗-mere : (x y : A ⨂₁-raw2 B) (z : _) → Path (A ⨂₁-raw B) (inc x) (inc y) → Path (A ⨂₁-raw B) (inc (z ∷ x)) (inc (z ∷ y)) --- -- -- -- +⊗-mere [] [] z p = refl --- -- -- -- +⊗-mere [] (x ∷ y) z p = {!!} --- -- -- -- +⊗-mere (x ∷ x₁) y z = {!!} - --- -- -- -- +⊗-merecomm' : (x y : fst A × fst B) → Path (A ⨂₁-raw B) (inc ([ x ] ++ [ y ])) (inc ([ y ] ++ [ x ])) --- -- -- -- +⊗-merecomm' (a , b) (c , d) = --- -- -- -- cong inc (cong (_∷ (c , d) ∷ []) (cong (_, b) (sym (rid strA a) --- -- -- -- ∙∙ cong (+G strA a) (sym (invr strA c)) --- -- -- -- ∙∙ {!!}))) --- -- -- -- ∙∙ {!!} --- -- -- -- ∙∙ {!!} - - --- -- -- -- _+⊗_ : A ⨂₁-raw B → A ⨂₁-raw B → A ⨂₁-raw B --- -- -- -- inc x +⊗ inc x₁ = inc (x ++ x₁) --- -- -- -- inc [] +⊗ bilinl x₁ y z i = bilinl x₁ y z i --- -- -- -- inc ((a , b) ∷ []) +⊗ bilinl x₁ y z i = {!!} -- help i --- -- -- -- where --- -- -- -- help : inc --- -- -- -- ((a , b) ∷ --- -- -- -- (+G strA x₁ y , z) ∷ --- -- -- -- []) ≡ inc ((a , b) ∷ (x₁ , z) ∷ (y , z) ∷ []) --- -- -- -- help = {!sym (bilinl a _ b)!} ∙ {!!} --- -- -- -- inc (x ∷ y ∷ z) +⊗ bilinl a b c i = {!!} -- inc (x ∷ x₃) +⊗ {!!} --- -- -- -- inc x +⊗ bilinr x₁ y z i = {!!} --- -- -- -- inc x +⊗ ⊗squash y y₁ x₁ y₂ i i₁ = {!!} --- -- -- -- bilinl x y₁ z i +⊗ y = {!!} --- -- -- -- bilinr x y₁ z i +⊗ y = {!!} --- -- -- -- ⊗squash x x₁ x₂ y₁ i i₁ +⊗ y = {!!} - --- -- -- -- -- module _ (A : Type ℓ) (R : Rel (Word A) (Word A) ℓ') where --- -- -- -- -- data AbGroupRel : Word A → Word A → Type (ℓ-max ℓ ℓ') where --- -- -- -- -- agr-commutes : ∀ x y → AbGroupRel (x ++ y) (y ++ x) --- -- -- -- -- agr-refl : ∀ {x y} → R x y → AbGroupRel x y - --- -- -- -- -- agr-rev : (w : Word A) → {!⨂₁!} -- QuotWordRel (rev w) w --- -- -- -- -- agr-rev = {!!} - --- -- -- -- -- module _ (G H : AbGroup ℓ) where --- -- -- -- -- TensorCarrier : Type _ --- -- -- -- -- TensorCarrier = --- -- -- -- -- Σ[ T ∈ AbGroup ℓ ] --- -- -- -- -- (Σ[ t ∈ (fst G → fst H → fst T) ] --- -- -- -- -- ((C : AbGroup ℓ) --- -- -- -- -- → isEquiv {A = fst T → fst C} {B = fst G → fst H → fst C} --- -- -- -- -- λ f a b → f (t a b))) --- -- -- -- -- anIso : {!!} --- -- -- -- -- anIso = {!!} - --- -- -- -- -- -- 0⊗ : TensorCarrier --- -- -- -- -- -- fst 0⊗ = dirProdAb G H --- -- -- -- -- -- fst (snd 0⊗) x y = x , y --- -- -- -- -- -- snd (snd 0⊗) C = isoToIsEquiv {!0⊗!} --- -- -} + +commFun : ∀ {ℓ ℓ'} {A : AbGroup ℓ} {B : AbGroup ℓ'} → A ⨂₁ B → B ⨂₁ A +commFun (a ⊗ b) = b ⊗ a +commFun (x +⊗ x₁) = commFun x +⊗ commFun x₁ +commFun (⊗comm x x₁ i) = ⊗comm (commFun x) (commFun x₁) i +commFun (⊗assoc x x₁ x₂ i) = ⊗assoc (commFun x) (commFun x₁) (commFun x₂) i +commFun (⊗lUnit x i) = ⊗lUnit (commFun x) i +commFun (linl x y z i) = linr z x y i +commFun (linr x y z i) = linl y z x i +commFun (flip x b i) = flip b x (~ i) +commFun (⊗squash x x₁ x₂ y i i₁) = + ⊗squash (commFun x) (commFun x₁) + (λ i → commFun (x₂ i)) (λ i → commFun (y i)) i i₁ + +commFunx2 : ∀ {ℓ ℓ'} {A : AbGroup ℓ} {B : AbGroup ℓ'} (x : A ⨂₁ B) → commFun (commFun x) ≡ x +commFunx2 = + ⊗elimProp (λ _ → ⊗squash _ _) + (λ _ _ → refl) + λ _ _ p q → cong₂ _+⊗_ p q + + +⨂-commIso : ∀ {ℓ ℓ'} {A : AbGroup ℓ} {B : AbGroup ℓ'} → GroupIso ((A ⨂ B) *) ((B ⨂ A) *) +Iso.fun (fst ⨂-commIso) = commFun +Iso.inv (fst ⨂-commIso) = commFun +Iso.rightInv (fst ⨂-commIso) = commFunx2 +Iso.leftInv (fst ⨂-commIso) = commFunx2 +snd ⨂-commIso = makeIsGroupHom λ x y → refl + +⨂-comm : ∀ {ℓ ℓ'} {A : AbGroup ℓ} {B : AbGroup ℓ'} → AbGroupEquiv (A ⨂ B) (B ⨂ A) +fst ⨂-comm = isoToEquiv (fst (⨂-commIso)) +snd ⨂-comm = snd ⨂-commIso + +module _ {ℓ ℓ' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} where + open AbGroupStr renaming (_+_ to _+G_ ; -_ to -G) + _+A_ = _+G_ (snd A) + _+B_ = _+G_ (snd B) + + 0A = 0g (snd A) + 0B = 0g (snd B) + + + ⨂→AbGroup-elim : ∀ {ℓ} (C : AbGroup ℓ) + → (f : (fst A × fst B) → fst C) + → (f (0A , 0B) ≡ 0g (snd C)) + → (linL : (x y : fst A) (b : fst B) → f (x +A y , b) ≡ _+G_ (snd C) (f (x , b)) (f (y , b))) + → (linR : (a : fst A) (x y : fst B) → f (a , x +B y) ≡ _+G_ (snd C) (f (a , x)) (f (a , y))) + → (fl : (x : fst A) (y : fst B) → f (x , -G (snd B) y) ≡ f ((-G (snd A) x) , y)) + → A ⨂₁ B → fst C + ⨂→AbGroup-elim C f p linL linR fl (a ⊗ b) = f (a , b) + ⨂→AbGroup-elim C f p linL linR fl (x +⊗ x₁) = + _+G_ (snd C) (⨂→AbGroup-elim C f p linL linR fl x) (⨂→AbGroup-elim C f p linL linR fl x₁) + ⨂→AbGroup-elim C f p linL linR fl (⊗comm x x₁ i) = + comm (snd C) (⨂→AbGroup-elim C f p linL linR fl x) (⨂→AbGroup-elim C f p linL linR fl x₁) i + ⨂→AbGroup-elim C f p linL linR fl (⊗assoc x x₁ x₂ i) = + assoc (snd C) (⨂→AbGroup-elim C f p linL linR fl x) (⨂→AbGroup-elim C f p linL linR fl x₁) (⨂→AbGroup-elim C f p linL linR fl x₂) i + ⨂→AbGroup-elim C f p linL linR fl (⊗lUnit x i) = + (cong (λ y → (snd C +G y) (⨂→AbGroup-elim C f p linL linR fl x)) p + ∙ (lid (snd C) (⨂→AbGroup-elim C f p linL linR fl x))) i + ⨂→AbGroup-elim C f p linL linR fl (linl x y z i) = linL x y z i + ⨂→AbGroup-elim C f p linL linR fl (linr x y z i) = linR x y z i + ⨂→AbGroup-elim C f p linL linR fl (flip x b i) = fl x b i + ⨂→AbGroup-elim C f p linL linR fl (⊗squash x x₁ x₂ y i i₁) = + is-set (snd C) _ _ (λ i → ⨂→AbGroup-elim C f p linL linR fl (x₂ i)) (λ i → ⨂→AbGroup-elim C f p linL linR fl (y i)) i i₁ + + ⨂→AbGroup-elim-hom : ∀ {ℓ} (C : AbGroup ℓ) → (f : (fst A × fst B) → fst C) (linL : _) (linR : _) (fl : _) (p : _) + → AbGroupHom (A ⨂ B) C + fst (⨂→AbGroup-elim-hom C f linL linR fl p) = ⨂→AbGroup-elim C f p linL linR fl + snd (⨂→AbGroup-elim-hom C f linL linR fl p) = + makeIsGroupHom + (λ x y → refl) + +module _ {ℓ ℓ' ℓ'' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} {C : AbGroup ℓ''} where + private + open AbGroupStr renaming (_+_ to +G ; -_ to -G) + _+A'_ = +G (snd A) + _+B'_ = +G (snd B) + _+C'_ = +G (snd C) + -A = -G (snd A) + -B = -G (snd B) + -C = -G (snd C) + + f : (c : fst C) → AbGroupHom (A ⨂ B) (A ⨂ (B ⨂ C)) + f c = ⨂→AbGroup-elim-hom (A ⨂ (B ⨂ C)) ((λ ab → fst ab ⊗ (snd ab ⊗ c))) + (λ x y b → linl x y (b ⊗ c)) + (λ x y b → (λ i → x ⊗ (linl y b c i)) ∙ linr x (y ⊗ c) (b ⊗ c)) + (λ _ _ → flip _ _) + (λ i → 0g (snd A) ⊗ (lCancelPrim c i)) + + assocHom : AbGroupHom ((A ⨂ B) ⨂ C) (A ⨂ (B ⨂ C)) + assocHom = + ⨂→AbGroup-elim-hom _ (λ x → f (snd x) .fst (fst x)) + (λ x y b → IsGroupHom.pres· (snd (f b)) x y) + (⊗elimProp (λ _ → isPropΠ2 λ _ _ → ⊗squash _ _) + (λ a b x y → (λ i → a ⊗ (linr b x y i)) + ∙∙ linr a (b ⊗ x) (b ⊗ y) + ∙∙ refl) + λ a b ind1 ind2 x y → cong₂ _+⊗_ (ind1 x y) (ind2 x y) + ∙∙ move4 (f x .fst a) (f y .fst a) (f x .fst b) (f y .fst b) _+⊗_ ⊗assoc ⊗comm + ∙∙ cong₂ _+⊗_ (sym (IsGroupHom.pres· (snd (f x)) a b)) (IsGroupHom.pres· (snd (f y)) a b)) + (⊗elimProp (λ _ → isPropΠ λ _ → ⊗squash _ _) + (λ a b c → (λ i → a ⊗ (flip b c i)) + ∙ flip a (b ⊗ c)) + (λ x y ind1 ind2 c → cong₂ _+⊗_ (ind1 c) (ind2 c) + ∙ IsGroupHom.pres· (snd (f c)) (-⊗ x) (-⊗ y))) + refl + +module _ {ℓ ℓ' ℓ'' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} {C : AbGroup ℓ''} where + private + open AbGroupStr renaming (_+_ to +G ; -_ to -G) + _+A'_ = +G (snd A) + _+B'_ = +G (snd B) + _+C'_ = +G (snd C) + -A = -G (snd A) + -B = -G (snd B) + -C = -G (snd C) + + f' : (a : fst A) → AbGroupHom (B ⨂ C) ((A ⨂ B) ⨂ C) + f' a = ⨂→AbGroup-elim-hom ((A ⨂ B) ⨂ C) + (λ bc → (a ⊗ fst bc) ⊗ snd bc) + (λ x y b → (λ i → (linr a x y i) ⊗ b) ∙ linl (a ⊗ x) (a ⊗ y) b) + (λ x y b → linr (a ⊗ x) y b) + (λ b c → flip (a ⊗ b) c ∙ λ i → flip a b (~ i) ⊗ c) + λ i → rCancelPrim a i ⊗ (0g (snd C)) + + assocHom⁻ : AbGroupHom (A ⨂ (B ⨂ C)) ((A ⨂ B) ⨂ C) + assocHom⁻ = ⨂→AbGroup-elim-hom _ (λ abc → f' (fst abc) .fst (snd abc)) + (λ x y → ⊗elimProp (λ _ → ⊗squash _ _) + (λ b c → (λ i → linl x y b i ⊗ c) ∙ linl (x ⊗ b) (y ⊗ b) c) + λ a b ind1 ind2 → cong₂ _+⊗_ ind1 ind2 + ∙∙ move4 _ _ _ _ _+⊗_ ⊗assoc ⊗comm + ∙∙ cong₂ _+⊗_ (IsGroupHom.pres· (snd (f' x)) a b) + (IsGroupHom.pres· (snd (f' y)) a b)) + (λ a → IsGroupHom.pres· (snd (f' a))) + (λ a → ⊗elimProp (λ _ → ⊗squash _ _) + (λ b c → λ i → flip a b i ⊗ c) + λ x y ind1 ind2 → IsGroupHom.pres· (snd (f' a)) (-⊗ x) (-⊗ y) + ∙ cong₂ _+⊗_ ind1 ind2) + refl + + ⨂AssocIso : Iso (A ⨂₁ (B ⨂ C)) ((A ⨂ B) ⨂₁ C) + Iso.fun ⨂AssocIso = fst assocHom⁻ + Iso.inv ⨂AssocIso = fst assocHom + Iso.rightInv ⨂AssocIso = + ⊗elimProp (λ _ → ⊗squash _ _) + (⊗elimProp (λ _ → isPropΠ (λ _ → ⊗squash _ _)) + (λ a b c → refl) + λ a b ind1 ind2 c → cong₂ _+⊗_ (ind1 c) (ind2 c) ∙ sym (linl a b c)) + λ x y p q → cong (fst assocHom⁻) (IsGroupHom.pres· (snd assocHom) x y) + ∙∙ IsGroupHom.pres· (snd assocHom⁻) (fst assocHom x) (fst assocHom y) + ∙∙ cong₂ _+⊗_ p q + Iso.leftInv ⨂AssocIso = + ⊗elimProp (λ _ → ⊗squash _ _) + (λ a → ⊗elimProp (λ _ → ⊗squash _ _) + (λ b c → refl) + λ x y ind1 ind2 → + cong (fst assocHom ∘ fst assocHom⁻) (linr a x y) + ∙∙ cong (fst assocHom) (IsGroupHom.pres· (snd assocHom⁻) (a ⊗ x) (a ⊗ y)) + ∙∙ IsGroupHom.pres· (snd assocHom) (fst assocHom⁻ (a ⊗ x)) (fst assocHom⁻ (a ⊗ y)) + ∙∙ cong₂ _+⊗_ ind1 ind2 + ∙∙ sym (linr a x y)) + λ x y p q → cong (fst assocHom) (IsGroupHom.pres· (snd assocHom⁻) x y) + ∙∙ IsGroupHom.pres· (snd assocHom) (fst assocHom⁻ x) (fst assocHom⁻ y) + ∙∙ cong₂ _+⊗_ p q + + ⨂assoc : AbGroupEquiv (A ⨂ (B ⨂ C)) ((A ⨂ B) ⨂ C) + fst ⨂assoc = isoToEquiv ⨂AssocIso + snd ⨂assoc = snd assocHom⁻ diff --git a/Cubical/Algebra/Group/EilenbergMacLane/Base.agda b/Cubical/Algebra/Group/EilenbergMacLane/Base.agda index 687c1fe4a..0fde35c93 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/Base.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/Base.agda @@ -31,7 +31,7 @@ open import Cubical.Functions.Morphism open import Cubical.Foundations.Path private - variable ℓ : Level + variable ℓ ℓ' : Level _* = AbGroup→Group @@ -45,7 +45,7 @@ ptS {n = zero} {G = G} = AbGroupStr.0g (snd G) ptS {n = suc zero} {G = G} = embase ptS {n = suc (suc n)} {G = G} = north -EM-raw-elim : (G : AbGroup ℓ) (n : ℕ) {A : EM-raw G (suc n) → Type ℓ} +EM-raw-elim : (G : AbGroup ℓ) (n : ℕ) {A : EM-raw G (suc n) → Type ℓ'} → ((x : _) → isOfHLevel (suc n) (A x) ) → A ptS → (x : _) → A x @@ -71,6 +71,9 @@ EM G (suc (suc n)) = hLevelTrunc (4 + n) (EM-raw G (suc (suc n))) EM∙ : (G : AbGroup ℓ) (n : ℕ) → Pointed ℓ EM∙ G n = EM G n , (0ₖ n) +EM-raw∙ : (G : AbGroup ℓ) (n : ℕ) → Pointed ℓ +EM-raw∙ G n = EM-raw G n , ptS + hLevelEM : (G : AbGroup ℓ) (n : ℕ) → isOfHLevel (2 + n) (EM G n) hLevelEM G zero = AbGroupStr.is-set (snd G) hLevelEM G (suc zero) = emsquash @@ -81,7 +84,7 @@ EM-raw→EM G zero x = x EM-raw→EM G (suc zero) x = x EM-raw→EM G (suc (suc n)) = ∣_∣ -EM-elim : {G : AbGroup ℓ} (n : ℕ) {A : EM G n → Type ℓ} +EM-elim : {G : AbGroup ℓ} (n : ℕ) {A : EM G n → Type ℓ'} → ((x : _) → isOfHLevel (2 + n) (A x)) → ((x : EM-raw G n) → A (EM-raw→EM G n x)) → (x : _) → A x diff --git a/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda b/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda new file mode 100644 index 000000000..59d77b9ca --- /dev/null +++ b/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda @@ -0,0 +1,412 @@ +{-# OPTIONS --safe --experimental-lossy-unification --no-forcing #-} + +module Cubical.Algebra.Group.EilenbergMacLane.CupProduct where + +open import Cubical.Algebra.Group.EilenbergMacLane.Base +open import Cubical.Algebra.Group.EilenbergMacLane.WedgeConnectivity +open import Cubical.Algebra.Group.EilenbergMacLane.GroupStructure +open import Cubical.Algebra.Group.EilenbergMacLane.Properties +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport +open import Cubical.Homotopy.Loopspace + +open import Cubical.Algebra.Group.MorphismProperties + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Properties +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec; elim to trElim) +open import Cubical.HITs.EilenbergMacLane1 renaming (rec to EMrec) +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Data.Empty + renaming (rec to ⊥-rec) +open import Cubical.HITs.Truncation + renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) +open import Cubical.Data.Nat hiding (_·_) +open import Cubical.HITs.Susp +open import Cubical.Functions.Morphism +open import Cubical.Foundations.Path + +open import Cubical.Algebra.AbGroup.TensorProduct +open import Cubical.Algebra.Group + +open import Cubical.Foundations.Pointed.Homogeneous + + +private + variable + ℓ ℓ' : Level + +stupidInd : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 → A 1 + → ((n : ℕ) → (A (suc n) → A (suc (suc n)))) + → (n : ℕ) → A n +stupidInd a0 a1 ind zero = a0 +stupidInd a0 a1 ind (suc zero) = a1 +stupidInd {A = A} a0 a1 ind (suc (suc n)) = ind n (stupidInd {A = A} a0 a1 ind (suc n)) + +stupidInd' : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 + → ((n : ℕ) → (A n → A (suc n))) + → (n : ℕ) → A n +stupidInd' a0 ind zero = a0 +stupidInd' a0 ind (suc n) = ind n (stupidInd' a0 ind n) + +_+'_ : ℕ → ℕ → ℕ +zero +' m = m +suc n +' zero = suc n +suc n +' suc m = suc (suc (n + m)) + ++'-comm : (n m : ℕ) → n +' m ≡ m +' n ++'-comm zero zero = refl ++'-comm zero (suc m) = refl ++'-comm (suc n) zero = refl ++'-comm (suc n) (suc m) = cong (2 +_) (+-comm n m) + ++'≡+ : (n m : ℕ) → n +' m ≡ n + m ++'≡+ zero zero = refl ++'≡+ zero (suc m) = refl ++'≡+ (suc n) zero = cong suc (+-comm zero n) ++'≡+ (suc n) (suc m) = cong suc (sym (+-suc n m)) + +open AbGroupStr renaming (_+_ to _+Gr_ ; -_ to -Gr_) + +inducedFun-EM-raw : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} + → AbGroupHom G' H' + → ∀ n + → EM-raw G' n → EM-raw H' n +inducedFun-EM-raw f = + stupidInd (fst f) + (EMrec _ emsquash embase + (λ g → emloop (fst f g)) + λ g h → compPathR→PathP (sym + (sym (lUnit _) + ∙∙ cong (_∙ (sym (emloop (fst f h)))) (cong emloop (IsGroupHom.pres· (snd f) g h) + ∙ emloop-comp _ (fst f g) (fst f h)) + ∙∙ sym (∙assoc _ _ _) + ∙∙ cong (emloop (fst f g) ∙_) (rCancel _) + ∙∙ sym (rUnit _)))) + (λ n ind → λ { north → north + ; south → south + ; (merid a i) → merid (ind a) i} ) + +inducedFun-EM-raw-sect : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} + → (f : AbGroupEquiv G' H') + → ∀ n + → (x : EM-raw G' n) → inducedFun-EM-raw (invEq (fst f) , isGroupHomInv f) n + (inducedFun-EM-raw (fst (fst f) , snd f) n x) + ≡ x +inducedFun-EM-raw-sect f = + stupidInd (λ x → {!!}) + {!!} + {!!} + +inducedFun-EM-rawIso : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} + → AbGroupEquiv G' H' + → ∀ n + → Iso (EM-raw G' n) (EM-raw H' n) +inducedFun-EM-rawIso e n = {!!} + +EM⊗-raw-commFun : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} → ∀ n → EM-raw (G' ⨂ H') n → EM-raw (H' ⨂ G') n +EM⊗-raw-commFun = + stupidInd commFun + (EMrec _ emsquash embase + (λ g → emloop (commFun g)) + λ g h → compPathR→PathP (sym (sym (lUnit _) + ∙∙ cong (_∙ sym (emloop (commFun h))) (emloop-comp _ (commFun g) (commFun h)) + ∙∙ sym (∙assoc _ _ _) + ∙∙ cong (emloop (commFun g) ∙_) (rCancel _) + ∙∙ sym (rUnit _)))) + (λ n ind → λ { north → north + ; south → south + ; (merid a i) → merid (ind a) i} ) + +EM⊗-raw-comm-sect : {!!} +EM⊗-raw-comm-sect = {!!} + +module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where + G = fst G' + H = fst H' + + strG = snd G' + strH = snd H' + + 0G = 0g strG + 0H = 0g strH + + _+G_ = _+Gr_ strG + _+H_ = _+Gr_ strH + + -H_ = -Gr_ strH + -G_ = -Gr_ strG + + + + precup : ∀ n m → EM-raw G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m) + precup zero m = {!!} + precup (suc zero) m = {!!} + precup (suc (suc n)) m = {!!} + + + + ·₀' : H → (m : ℕ) → EM G' m → EM (G' ⨂ H') m + ·₀' h = + stupidInd + (_⊗ h) + (elimGroupoid _ (λ _ → emsquash) + embase + (λ g → emloop (g ⊗ h)) + λ g l → + compPathR→PathP + (sym (∙assoc _ _ _ + ∙∙ cong₂ _∙_ (sym (emloop-comp _ _ _) ∙ cong emloop (sym (linl g l h))) refl + ∙∙ rCancel _))) + λ n f → trRec (isOfHLevelTrunc (4 + n)) + λ { north → 0ₖ (suc (suc n)) + ; south → 0ₖ (suc (suc n)) + ; (merid a i) → EM→ΩEM+1 (suc n) (f (EM-raw→EM _ _ a)) i} + + ·₀ : G → (m : ℕ) → EM H' m → EM (G' ⨂ H') m + ·₀ g = + stupidInd (λ h → g ⊗ h) + (elimGroupoid _ (λ _ → emsquash) + embase + (λ h → emloop (g ⊗ h)) + λ h l → compPathR→PathP + (sym (∙assoc _ _ _ + ∙∙ cong₂ _∙_ (sym (emloop-comp _ _ _) ∙ cong emloop (sym (linr g h l))) refl + ∙∙ rCancel _))) + λ n f + → trRec (isOfHLevelTrunc (4 + n)) + λ { north → 0ₖ (suc (suc n)) + ; south → 0ₖ (suc (suc n)) + ; (merid a i) → EM→ΩEM+1 (suc n) (f (EM-raw→EM _ _ a)) i} + + ·₀-distr : (g h : G) → (m : ℕ) (x : EM H' m) → ·₀ (g +G h) m x ≡ ·₀ g m x +ₖ ·₀ h m x + ·₀-distr g h = + stupidInd + (linl g h) + (elimSet _ (λ _ → emsquash _ _) + refl + (λ w → compPathR→PathP (sym ((λ i → emloop (linl g h w i) + ∙ (lUnit (sym (cong₂+₁ (emloop (g ⊗ w)) (emloop (h ⊗ w)) i)) (~ i))) + ∙∙ cong₂ _∙_ (emloop-comp _ (g ⊗ w) (h ⊗ w)) refl + ∙∙ rCancel _)))) + λ m ind → + trElim (λ _ → isOfHLevelTruncPath) + λ { north → refl + ; south → refl + ; (merid a i) k → z m ind a k i} + + where + z : (m : ℕ) → ((x : EM H' (suc m)) → ·₀ (g +G h) (suc m) x ≡ ·₀ g (suc m) x +ₖ ·₀ h (suc m) x) → (a : EM-raw H' (suc m)) + → cong (·₀ (g +G h) (suc (suc m))) (cong ∣_∣ₕ (merid a)) ≡ + cong₂ _+ₖ_ + (cong (·₀ g (suc (suc m))) (cong ∣_∣ₕ (merid a))) + (cong (·₀ h (suc (suc m))) (cong ∣_∣ₕ (merid a))) + z m ind a = (λ i → EM→ΩEM+1 _ (ind (EM-raw→EM _ _ a) i)) + ∙∙ EM→ΩEM+1-hom _ (·₀ g (suc m) (EM-raw→EM H' (suc m) a)) (·₀ h (suc m) (EM-raw→EM H' (suc m) a)) + ∙∙ sym (cong₂+₂ m (cong (·₀ g (suc (suc m))) (cong ∣_∣ₕ (merid a))) + (cong (·₀ h (suc (suc m))) (cong ∣_∣ₕ (merid a)))) + + ·₀0 : (m : ℕ) → (g : G) → ·₀ g m (0ₖ m) ≡ 0ₖ m + ·₀0 zero = rCancelPrim + ·₀0 (suc zero) g = refl + ·₀0 (suc (suc m)) g = refl + + ·₀'0 : (m : ℕ) (h : H) → ·₀' h m (0ₖ m) ≡ 0ₖ m + ·₀'0 zero = lCancelPrim + ·₀'0 (suc zero) g = refl + ·₀'0 (suc (suc m)) g = refl + + 0·₀ : (m : ℕ) → (x : _) → ·₀ 0G m x ≡ 0ₖ m + 0·₀ = + stupidInd lCancelPrim + (elimSet _ (λ _ → emsquash _ _) + refl + λ g → compPathR→PathP ((sym (emloop-1g _) + ∙ cong emloop (sym (lCancelPrim g))) + ∙∙ (λ i → rUnit (rUnit (cong (·₀ 0G 1) (emloop g)) i) i) + ∙∙ sym (∙assoc _ _ _))) + λ n f → trElim (λ _ → isOfHLevelTruncPath) + λ { north → refl + ; south → refl + ; (merid a i) j → (cong (EM→ΩEM+1 (suc n)) (f (EM-raw→EM _ _ a)) + ∙ EM→ΩEM+1-0ₖ _) j i} + + 0·₀' : (m : ℕ) (g : _) → ·₀' 0H m g ≡ 0ₖ m + 0·₀' = + stupidInd + rCancelPrim + (elimSet _ (λ _ → emsquash _ _) + refl + λ g → compPathR→PathP (sym (∙assoc _ _ _ + ∙∙ sym (rUnit _) ∙ sym (rUnit _) + ∙∙ (cong emloop (rCancelPrim g) + ∙ emloop-1g _)))) + λ n f → trElim (λ _ → isOfHLevelTruncPath) + λ { north → refl + ; south → refl + ; (merid a i) j → (cong (EM→ΩEM+1 (suc n)) (f (EM-raw→EM _ _ a)) + ∙ EM→ΩEM+1-0ₖ _) j i} + + cup∙ : ∀ n m → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m) + cup∙ = + stupidInd' + (λ m g → (·₀ g m) , ·₀0 m g) + λ n f → + stupidInd' + (λ g → (λ h → ·₀' h (suc n) g) , 0·₀' (suc n) g) + λ m _ → main n m f + + where + main : (n m : ℕ) (ind : ((m : ℕ) → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m))) + → EM G' (suc n) → EM∙ H' (suc m) →∙ EM∙ (G' ⨂ H') (suc (suc (n + m))) + main zero m ind = + elimGroupoid _ (λ _ → isOfHLevel↑∙ _ _) + ((λ _ → 0ₖ (2 + m)) , refl) + (f m) + λ n h → finalpp m n h + where + f : (m : ℕ) → G → typ (Ω (EM∙ H' (suc m) →∙ EM∙ (G' ⨂ H') (suc (suc m)) ∙)) + fst (f m g i) x = EM→ΩEM+1 _ (·₀ g _ x) i + snd (f zero g i) j = EM→ΩEM+1-0ₖ (suc zero) j i + snd (f (suc m) g i) j = EM→ΩEM+1-0ₖ (suc (suc m)) j i + + f-hom-fst : (m : ℕ) (g h : G) → cong fst (f m (g +G h)) ≡ cong fst (f m g ∙ f m h) + f-hom-fst m g h = + (λ i j x → EM→ΩEM+1 _ (·₀-distr g h (suc m) x i) j) + ∙∙ (λ i j x → EM→ΩEM+1-hom _ (·₀ g (suc m) x) (·₀ h (suc m) x) i j) + ∙∙ sym (cong-∙ fst (f m g) (f m h)) + + f-hom : (m : ℕ) (g h : G) → f m (g +G h) ≡ f m g ∙ f m h + f-hom m g h = →∙Homogeneous≡Path (isHomogeneousEM _) _ _ (f-hom-fst m g h) + + finalpp : (m : ℕ) (g h : G) → PathP (λ i → f m g i ≡ f m (g +G h) i) refl (f m h) + finalpp m g h = + compPathR→PathP (sym (rCancel _) + ∙∙ cong (_∙ sym (f m (g +G h))) (f-hom m g h) + ∙∙ sym (∙assoc _ _ _)) + + main (suc n) m ind = + trElim (λ _ → isOfHLevel↑∙ (2 + n) m) + λ { north → (λ _ → 0ₖ (3 + (n + m))) , refl + ; south → (λ _ → 0ₖ (3 + (n + m))) , refl + ; (merid a i) → (λ x → EM→ΩEM+1 _ (ind _ (EM-raw→EM _ _ a) .fst x) i) + , λ j → (cong (EM→ΩEM+1 (suc (suc (n + m)))) (ind (suc m) (EM-raw→EM G' (suc n) a) .snd) + ∙ EM→ΩEM+1-0ₖ _) j i} + + _⌣ₖ_ : {n m : ℕ} (x : EM G' n) (y : EM H' m) → EM (G' ⨂ H') (n +' m) + _⌣ₖ_ x y = cup∙ _ _ x .fst y + + ⌣ₖ-0ₖ : (n m : ℕ) (x : EM G' n) → (x ⌣ₖ 0ₖ m) ≡ 0ₖ (n +' m) + ⌣ₖ-0ₖ n m x = cup∙ n m x .snd + + 0ₖ-⌣ₖ : (n m : ℕ) (x : EM H' m) → ((0ₖ n) ⌣ₖ x) ≡ 0ₖ (n +' m) + 0ₖ-⌣ₖ zero m = 0·₀ _ + 0ₖ-⌣ₖ (suc zero) zero x = refl + 0ₖ-⌣ₖ (suc (suc n)) zero x = refl + 0ₖ-⌣ₖ (suc zero) (suc m) x = refl + 0ₖ-⌣ₖ (suc (suc n)) (suc m) x = refl + + private + distrl1 : (n m : ℕ) → EM H' m → EM H' m → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) + fst (distrl1 n m x y) z = z ⌣ₖ (x +ₖ y) + snd (distrl1 n m x y) = 0ₖ-⌣ₖ n m _ + + distrl2 : (n m : ℕ) → EM H' m → EM H' m → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) + fst (distrl2 n m x y) z = (z ⌣ₖ x) +ₖ (z ⌣ₖ y) + snd (distrl2 n m x y) = cong₂ _+ₖ_ (0ₖ-⌣ₖ n m x) (0ₖ-⌣ₖ n m y) ∙ rUnitₖ _ (0ₖ (n +' m)) + + hLevLem : (n m : ℕ) → isOfHLevel (suc (suc m)) (EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') ((suc n) +' m)) + hLevLem n m = subst (isOfHLevel (suc (suc m))) (λ i → EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') + ((cong suc (+-comm m n) ∙ sym (+'≡+ (suc n) m)) i)) (isOfHLevel↑∙ m n) + + + pathType : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → Type ℓ + pathType n x p = sym (rUnitₖ (2 + n) x) ∙ (λ i → x +ₖ p i) ≡ sym (lUnitₖ (2 + n) x) ∙ λ i → p i +ₖ x + + lem : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) + → pathType n x p + lem n x = J (λ x p → pathType n x p) refl + + mainDistrL : (n m : ℕ) (x y : EM H' (suc m)) → distrl1 (suc n) (suc m) x y ≡ distrl2 (suc n) (suc m) x y + mainDistrL n zero = + wedgeConEM.fun H' H' 0 0 + (λ _ _ → hLevLem _ _ _ _) + (λ x → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt λ z → l x z)) + (λ y → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt λ z → r y z )) + λ i → →∙Homogeneous≡ (isHomogeneousEM (suc (suc (n + 0)))) + (funExt (λ z → l≡r z i)) + where + l : (x : EM H' 1) (z : _) → (distrl1 (suc n) 1 embase x .fst z) ≡ (distrl2 (suc n) 1 embase x .fst z) + l x z = cong (z ⌣ₖ_) (lUnitₖ _ x) + ∙∙ sym (lUnitₖ _ (z ⌣ₖ x)) + ∙∙ λ i → (⌣ₖ-0ₖ (suc n) (suc zero) z (~ i)) +ₖ (z ⌣ₖ x) + + r : (z : EM H' 1) (x : EM G' (suc n)) → (distrl1 (suc n) 1 z embase .fst x) ≡ (distrl2 (suc n) 1 z embase .fst x) + r y z = cong (z ⌣ₖ_) (rUnitₖ _ y) + ∙∙ sym (rUnitₖ _ (z ⌣ₖ y)) + ∙∙ λ i → (z ⌣ₖ y) +ₖ (⌣ₖ-0ₖ (suc n) (suc zero) z (~ i)) + + l≡r : (z : EM G' (suc n)) → l embase z ≡ r embase z + l≡r z = sym (lem _ _ (sym (⌣ₖ-0ₖ (suc n) (suc zero) z))) + + mainDistrL n (suc m) = + elim2 (λ _ _ → isOfHLevelPath (4 + m) (hLevLem _ _) _ _) + (wedgeConEM.fun H' H' (suc m) (suc m) + (λ x y p q → isOfHLevelPlus {n = suc (suc m)} (suc m) + (hLevLem n (suc (suc m)) + (distrl1 (suc n) (suc (suc m)) ∣ x ∣ ∣ y ∣) + (distrl2 (suc n) (suc (suc m)) ∣ x ∣ ∣ y ∣) p q)) + (λ x → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt (l x))) + (λ x → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt (r x))) + λ i → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt (λ z → l≡r z i))) + where + l : (x : EM-raw H' (suc (suc m))) (z : EM G' (suc n)) + → (distrl1 (suc n) (suc (suc m)) (0ₖ _) ∣ x ∣ₕ .fst z) ≡ (distrl2 (suc n) (suc (suc m)) (0ₖ _) ∣ x ∣ₕ .fst z) + l x z = cong (z ⌣ₖ_) (lUnitₖ (suc (suc m)) ∣ x ∣) + ∙∙ sym (lUnitₖ _ (z ⌣ₖ ∣ x ∣)) + ∙∙ λ i → (⌣ₖ-0ₖ (suc n) (suc (suc m)) z (~ i)) +ₖ (z ⌣ₖ ∣ x ∣) + + r : (x : EM-raw H' (suc (suc m))) (z : EM G' (suc n)) + → (distrl1 (suc n) (suc (suc m)) ∣ x ∣ₕ (0ₖ _) .fst z) ≡ (distrl2 (suc n) (suc (suc m)) ∣ x ∣ₕ (0ₖ _) .fst z) + r x z = cong (z ⌣ₖ_) (rUnitₖ (suc (suc m)) ∣ x ∣) + ∙∙ sym (rUnitₖ _ (z ⌣ₖ ∣ x ∣)) + ∙∙ λ i → (z ⌣ₖ ∣ x ∣) +ₖ (⌣ₖ-0ₖ (suc n) (suc (suc m)) z (~ i)) + + l≡r : (z : EM G' (suc n)) → l north z ≡ r north z + l≡r z = sym (lem _ _ (sym (⌣ₖ-0ₖ (suc n) (suc (suc m)) z))) + + + commer : {!!} + commer = {!!} + + private + distrr1 : (n m : ℕ) → EM H' n → EM H' n → EM∙ G' m →∙ EM∙ (H' ⨂ G') (n +' m) + fst (distrr1 n m x y) z = {!!} -- (x +ₖ y) ⌣ₖ ? + snd (distrr1 n m x y) = {!!} -- ⌣ₖ-0ₖ n m _ + + distrr2 : (n m : ℕ) → EM H' m → EM H' m → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) + fst (distrr2 n m x y) z = (z ⌣ₖ x) +ₖ (z ⌣ₖ y) + snd (distrr2 n m x y) = cong₂ _+ₖ_ (0ₖ-⌣ₖ n m x) (0ₖ-⌣ₖ n m y) ∙ rUnitₖ _ (0ₖ (n +' m)) + + hLevLem2 : (n m : ℕ) → isOfHLevel (suc (suc m)) (EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') ((suc n) +' m)) + hLevLem2 n m = subst (isOfHLevel (suc (suc m))) (λ i → EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') + ((cong suc (+-comm m n) ∙ sym (+'≡+ (suc n) m)) i)) (isOfHLevel↑∙ m n) + + mainDistrR : {!(n m : ℕ) (x y : EM H' (suc m)) → distr1 (suc n) (suc m) x y ≡ distrl2 (suc n) (suc m) x y!} + mainDistrR = {!!} diff --git a/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda b/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda index f31f865dc..e49e2ea88 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda @@ -37,8 +37,7 @@ open import Cubical.Foundations.Path open import Cubical.Foundations.Pointed.Homogeneous private - variable ℓ : Level - + variable ℓ ℓ' ℓ'' : Level isConnectedEM₁ : (G : Group ℓ) → isConnected 2 (EM₁ G) isConnectedEM₁ G = ∣ embase ∣ , h @@ -280,10 +279,7 @@ module _ {G : AbGroup ℓ} where (λ i → lUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) ∙ (λ i → rUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) - EM∙' : (n : ℕ) → Pointed _ - EM∙' n = EM G n , 0ₖ n - - Iso-EM-ΩEM+1 : (n : ℕ) → Iso (EM G n) (typ (Ω (EM∙' (suc n)))) + Iso-EM-ΩEM+1 : (n : ℕ) → Iso (EM G n) (typ (Ω (EM∙ G (suc n)))) Iso-EM-ΩEM+1 zero = invIso (ΩEM₁Iso (AbGroup→Group G)) Iso.fun (Iso-EM-ΩEM+1 (suc zero)) = decode' 0 (0ₖ 2) Iso.inv (Iso-EM-ΩEM+1 (suc zero)) = encode' 0 (0ₖ 2) @@ -294,161 +290,159 @@ module _ {G : AbGroup ℓ} where Iso.rightInv (Iso-EM-ΩEM+1 (suc (suc n))) = decode'-encode' (suc n) (0ₖ (3 + n)) Iso.leftInv (Iso-EM-ΩEM+1 (suc (suc n))) = encode'-decode' (suc n) - EM→ΩEM+1 : (n : ℕ) → EM G n → typ (Ω (EM∙' (suc n))) + EM→ΩEM+1 : (n : ℕ) → EM G n → typ (Ω (EM∙ G (suc n))) EM→ΩEM+1 n = Iso.fun (Iso-EM-ΩEM+1 n) - ΩEM+1→EM : (n : ℕ) → typ (Ω (EM∙' (suc n))) → EM G n + ΩEM+1→EM : (n : ℕ) → typ (Ω (EM∙ G (suc n))) → EM G n ΩEM+1→EM n = Iso.inv (Iso-EM-ΩEM+1 n) + EM→ΩEM+1-0ₖ : (n : ℕ) → EM→ΩEM+1 n (0ₖ n) ≡ refl + EM→ΩEM+1-0ₖ zero = emloop-1g _ + EM→ΩEM+1-0ₖ (suc zero) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + EM→ΩEM+1-0ₖ (suc (suc n)) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + EM→ΩEM+1-hom : (n : ℕ) (x y : EM G n) → EM→ΩEM+1 n (x +[ n ]ₖ y) ≡ EM→ΩEM+1 n x ∙ EM→ΩEM+1 n y EM→ΩEM+1-hom zero x y = emloop-comp _ x y EM→ΩEM+1-hom (suc zero) x y = σ-EM'-hom zero x y EM→ΩEM+1-hom (suc (suc n)) x y = σ-EM'-hom (suc n) x y - ΩEM+1→EM-hom : (n : ℕ) → (p q : typ (Ω (EM∙' (suc n)))) + ΩEM+1→EM-hom : (n : ℕ) → (p q : typ (Ω (EM∙ G (suc n)))) → ΩEM+1→EM n (p ∙ q) ≡ (ΩEM+1→EM n p) +[ n ]ₖ (ΩEM+1→EM n q) ΩEM+1→EM-hom n = morphLemmas.isMorphInv (λ x y → x +[ n ]ₖ y) (_∙_) (EM→ΩEM+1 n) (EM→ΩEM+1-hom n) (ΩEM+1→EM n) (Iso.rightInv (Iso-EM-ΩEM+1 n)) (Iso.leftInv (Iso-EM-ΩEM+1 n)) - EM→ΩEM+1-0ₖ : (n : ℕ) → EM→ΩEM+1 n (0ₖ n) ≡ refl - EM→ΩEM+1-0ₖ zero = emloop-1g _ - EM→ΩEM+1-0ₖ (suc zero) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) - EM→ΩEM+1-0ₖ (suc (suc n)) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + ΩEM+1→EM-refl : (n : ℕ) → ΩEM+1→EM n refl ≡ 0ₖ n ΩEM+1→EM-refl zero = transportRefl 0g ΩEM+1→EM-refl (suc zero) = refl ΩEM+1→EM-refl (suc (suc n)) = refl + + EM≃ΩEM+1 : (n : ℕ) → EM G n ≃ typ (Ω (EM∙ G (suc n))) + EM≃ΩEM+1 n = isoToEquiv (Iso-EM-ΩEM+1 n) + + EM≃ΩEM+1∙ : (n : ℕ) → EM∙ G n ≡ Ω (EM∙ G (suc n)) + EM≃ΩEM+1∙ n = ua∙ (EM≃ΩEM+1 n) (EM→ΩEM+1-0ₖ n) + isHomogeneousEM : (n : ℕ) → isHomogeneous (EM∙ G n) isHomogeneousEM n x = ua∙ (isoToEquiv (addIso n x)) (lUnitₖ n x) - isContr-↓∙ : (n : ℕ) → isContr (EM∙ G (suc n) →∙ EM∙ G n) - fst (isContr-↓∙ zero) = (λ _ → 0g) , refl - snd (isContr-↓∙ zero) (f , p) = - Σ≡Prop (λ x → is-set _ _) - (funExt (EM-raw-elim G 0 (λ _ → is-set _ _) - (sym p))) - fst (isContr-↓∙ (suc n)) = (λ _ → 0ₖ (suc n)) , refl - snd (isContr-↓∙ (suc n)) (f , p) = - →∙Homogeneous≡ (isHomogeneousEM _) - (funExt - (trElim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (hLevelEM G (suc n))) _ _) - (EM-raw-elim _ _ (λ _ → hLevelEM G (suc n) _ _) - (sym p)))) - - -{- --- HLevel stuff for cup product -isContr-↓∙ : (n : ℕ) → isContr (coHomK-ptd (suc n) →∙ coHomK-ptd n) -fst (isContr-↓∙ zero) = (λ _ → 0) , refl -snd (isContr-↓∙ zero) (f , p) = - Σ≡Prop (λ f → isSetℤ _ _) - (funExt (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 isSetℤ) _ _) - (toPropElim (λ _ → isSetℤ _ _) (sym p)))) -fst (isContr-↓∙ (suc n)) = (λ _ → 0ₖ _) , refl -fst (snd (isContr-↓∙ (suc n)) f i) x = - trElim {B = λ x → 0ₖ (suc n) ≡ fst f x} - (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (isOfHLevelTrunc (3 + n))) _ _) - (sphereElim _ (λ _ → isOfHLevelTrunc (3 + n) _ _) - (sym (snd f))) x i -snd (snd (isContr-↓∙ (suc n)) f i) j = snd f (~ i ∨ j) - -isContr-↓∙' : (n : ℕ) → isContr (S₊∙ (suc n) →∙ coHomK-ptd n) -fst (isContr-↓∙' zero) = (λ _ → 0) , refl -snd (isContr-↓∙' zero) (f , p) = - Σ≡Prop (λ f → isSetℤ _ _) - (funExt (toPropElim (λ _ → isSetℤ _ _) (sym p))) -fst (isContr-↓∙' (suc n)) = (λ _ → 0ₖ _) , refl -fst (snd (isContr-↓∙' (suc n)) f i) x = - sphereElim _ {A = λ x → 0ₖ (suc n) ≡ fst f x} - (λ _ → isOfHLevelTrunc (3 + n) _ _) - (sym (snd f)) x i -snd (snd (isContr-↓∙' (suc n)) f i) j = snd f (~ i ∨ j) - -isOfHLevel→∙Kn : {A : Pointed₀} (n m : ℕ) - → isOfHLevel (suc m) (A →∙ coHomK-ptd n) → isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) -isOfHLevel→∙Kn {A = A} n m hlev = step₃ - where - step₁ : isOfHLevel (suc m) (A →∙ Ω (coHomK-ptd (suc n))) - step₁ = - subst (isOfHLevel (suc m)) - (λ i → A →∙ ua∙ {A = Ω (coHomK-ptd (suc n))} {B = coHomK-ptd n} - (invEquiv Kn≃ΩKn+1) - (ΩKn+1→Kn-refl n) (~ i)) hlev - - step₂ : isOfHLevel (suc m) (typ (Ω (A →∙ coHomK-ptd (suc n) ∙))) - step₂ = isOfHLevelRetractFromIso (suc m) (invIso (invIso (ΩfunExtIso _ _))) step₁ - - step₃ : isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) - step₃ = - isOfHLevelΩ→isOfHLevel m + + +module _ where + open AbGroupStr renaming (_+_ to comp) + + isContr-↓∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} (n : ℕ) → isContr (EM∙ G (suc n) →∙ EM∙ H n) + fst (isContr-↓∙ {G = G} {H = H} zero) = (λ _ → 0g (snd H)) , refl + snd (isContr-↓∙{G = G} {H = H} zero) (f , p) = + Σ≡Prop (λ x → is-set (snd H) _ _) + (funExt (EM-raw-elim G 0 (λ _ → is-set (snd H) _ _) (sym p))) + fst (isContr-↓∙ {G = G} {H = H} (suc n)) = (λ _ → 0ₖ (suc n)) , refl + fst (snd (isContr-↓∙ {G = G} {H = H} (suc n)) f i) x = + trElim {B = λ x → 0ₖ (suc n) ≡ fst f x} + (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (hLevelEM H (suc n))) _ _) + (EM-raw-elim _ _ (λ _ → hLevelEM H (suc n) _ _) (sym (snd f))) + x i + snd (snd (isContr-↓∙ (suc n)) f i) j = snd f (~ i ∨ j) + + isContr-↓∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} (n : ℕ) → isContr ((EM-raw G (suc n) , ptS) →∙ EM∙ H n) + isContr-↓∙' zero = isContr-↓∙ zero + fst (isContr-↓∙' (suc n)) = (λ _ → 0ₖ (suc n)) , refl + fst (snd (isContr-↓∙' {H = H} (suc n)) f i) x = + (EM-raw-elim _ _ {A = λ x → 0ₖ (suc n) ≡ fst f x} (λ _ → hLevelEM H (suc n) _ _) (sym (snd f))) x i + snd (snd (isContr-↓∙' (suc n)) f i) j = snd f (~ i ∨ j) + + isOfHLevel→∙EM : ∀ {ℓ} {A : Pointed ℓ} {G : AbGroup ℓ'} (n m : ℕ) + → isOfHLevel (suc m) (A →∙ EM∙ G n) → isOfHLevel (suc (suc m)) (A →∙ EM∙ G (suc n)) + isOfHLevel→∙EM {A = A} {G = G} n m hlev = step₃ + where + step₁ : isOfHLevel (suc m) (A →∙ Ω (EM∙ G (suc n))) + step₁ = + subst (isOfHLevel (suc m)) + (λ i → A →∙ ua∙ {A = Ω (EM∙ G (suc n))} {B = EM∙ G n} + (invEquiv (EM≃ΩEM+1 n)) + (ΩEM+1→EM-refl n) (~ i)) hlev + + step₂ : isOfHLevel (suc m) (typ (Ω (A →∙ EM∙ G (suc n) ∙))) + step₂ = isOfHLevelRetractFromIso (suc m) (invIso (invIso (ΩfunExtIso _ _))) step₁ + + step₃ : isOfHLevel (suc (suc m)) (A →∙ EM∙ G (suc n)) + step₃ = isOfHLevelΩ→isOfHLevel m λ f → subst (λ x → isOfHLevel (suc m) (typ (Ω x))) - (isHomogeneous→∙ (isHomogeneousKn (suc n)) f) + (isHomogeneous→∙ (isHomogeneousEM (suc n)) f) step₂ -isOfHLevel↑∙ : ∀ n m → isOfHLevel (2 + n) (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m))) -isOfHLevel↑∙ zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙ m)) -isOfHLevel↑∙ (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙ n m) - -isOfHLevel↑∙' : ∀ n m → isOfHLevel (2 + n) (S₊∙ (suc m) →∙ coHomK-ptd (suc (n + m))) -isOfHLevel↑∙' zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙' m)) -isOfHLevel↑∙' (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙' n m) - -→∙KnPath : (A : Pointed₀) (n : ℕ) → Ω (A →∙ coHomK-ptd (suc n) ∙) ≡ (A →∙ coHomK-ptd n ∙) -→∙KnPath A n = - ua∙ (isoToEquiv (ΩfunExtIso A (coHomK-ptd (suc n)))) refl - ∙ (λ i → (A →∙ Kn≃ΩKn+1∙ {n = n} (~ i) ∙)) - -contr∙-lem : (n m : ℕ) - → isContr (coHomK-ptd (suc n) - →∙ (coHomK-ptd (suc m) - →∙ coHomK-ptd (suc (n + m)) ∙)) -fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl) , refl -snd (contr∙-lem n m) (f , p) = →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousKn _)) (sym (funExt help)) - where - help : (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) - help = - trElim (λ _ → isOfHLevelPath (3 + n) - (subst (λ x → isOfHLevel x (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m)))) (λ i → suc (suc (+-comm n 1 i))) - (isOfHLevelPlus' {n = 1} (2 + n) (isOfHLevel↑∙ n m))) _ _) - (sphereElim _ (λ _ → isOfHLevel↑∙ n m _ _) p) - -isOfHLevel↑∙∙ : ∀ n m l - → isOfHLevel (2 + l) (coHomK-ptd (suc n) - →∙ (coHomK-ptd (suc m) - →∙ coHomK-ptd (suc (suc (l + n + m))) ∙)) -isOfHLevel↑∙∙ n m zero = - isOfHLevelΩ→isOfHLevel 0 λ f - → subst - isProp (cong (λ x → typ (Ω x)) - (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) - (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) h) - where - h : isProp (coHomK-ptd (suc n) - →∙ (Ω (coHomK-ptd (suc m) - →∙ coHomK-ptd (suc (suc (n + m))) ∙))) - h = - subst isProp (λ i → coHomK-ptd (suc n) →∙ (→∙KnPath (coHomK-ptd (suc m)) (suc (n + m)) (~ i))) - (isContr→isProp (contr∙-lem n m)) -isOfHLevel↑∙∙ n m (suc l) = - isOfHLevelΩ→isOfHLevel (suc l) - λ f → - subst - (isOfHLevel (2 + l)) (cong (λ x → typ (Ω x)) - (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) - (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) h) - where - h : isOfHLevel (2 + l) - (coHomK-ptd (suc n) - →∙ (Ω (coHomK-ptd (suc m) - →∙ coHomK-ptd (suc (suc (suc (l + n + m)))) ∙))) - h = - subst (isOfHLevel (2 + l)) - (λ i → coHomK-ptd (suc n) →∙ →∙KnPath (coHomK-ptd (suc m)) (suc (suc (l + n + m))) (~ i)) - (isOfHLevel↑∙∙ n m l) --} + isOfHLevel↑∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} + → ∀ n m → isOfHLevel (2 + n) (EM∙ G (suc m) →∙ EM∙ H (suc (n + m))) + isOfHLevel↑∙ zero m = isOfHLevel→∙EM m 0 (isContr→isProp (isContr-↓∙ m)) + isOfHLevel↑∙ (suc n) m = isOfHLevel→∙EM (suc (n + m)) (suc n) (isOfHLevel↑∙ n m) + + isOfHLevel↑∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} + → ∀ n m → isOfHLevel (2 + n) (EM-raw∙ G (suc m) →∙ EM∙ H (suc (n + m))) + isOfHLevel↑∙' zero m = isOfHLevel→∙EM m 0 (isContr→isProp (isContr-↓∙' m)) + isOfHLevel↑∙' (suc n) m = isOfHLevel→∙EM (suc (n + m)) (suc n) (isOfHLevel↑∙' n m) + + →∙EMPath : ∀ {ℓ} {G : AbGroup ℓ} (A : Pointed ℓ') (n : ℕ) → Ω (A →∙ EM∙ G (suc n) ∙) ≡ (A →∙ EM∙ G n ∙) + →∙EMPath {G = G} A n = + ua∙ (isoToEquiv (ΩfunExtIso A (EM∙ G (suc n)))) refl + ∙ (λ i → (A →∙ EM≃ΩEM+1∙ {G = G} n (~ i) ∙)) + + contr∙-lem : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) + → isContr (EM∙ G (suc n) →∙ (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) + fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl), refl + snd (contr∙-lem {G = G} {H = H} {L = L} n m) (f , p) = + →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) + (sym (funExt (help n f p))) + where + help : (n : ℕ) → (f : EM G (suc n) → EM∙ H (suc m) →∙ EM∙ L (suc (n + m))) + → f (snd (EM∙ G (suc n))) ≡ snd (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) + → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) + help zero f p = + EM-raw-elim G zero + (λ _ → isOfHLevel↑∙ zero m _ _) p + help (suc n) f p = + trElim (λ _ → isOfHLevelPath (4 + n) + (subst (λ x → isOfHLevel x (EM∙ H (suc m) →∙ EM∙ L (suc ((suc n) + m)))) + (λ i → suc (suc (+-comm (suc n) 1 i))) + (isOfHLevelPlus' {n = 1} (3 + n) (isOfHLevel↑∙ (suc n) m))) _ _) + (EM-raw-elim G (suc n) + (λ _ → isOfHLevel↑∙ (suc n) m _ _) p) + + isOfHLevel↑∙∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} + → ∀ n m l → isOfHLevel (2 + l) (EM∙ G (suc n) + →∙ (EM∙ H (suc m) + →∙ EM∙ L (suc (suc (l + n + m))) ∙)) + isOfHLevel↑∙∙ {G = G} {H = H} {L = L} n m zero = + isOfHLevelΩ→isOfHLevel 0 λ f + → subst + isProp (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) + (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) h) + where + h : isProp (EM∙ G (suc n) + →∙ (Ω (EM∙ H (suc m) + →∙ EM∙ L (suc (suc (n + m))) ∙))) + h = + subst isProp (λ i → EM∙ G (suc n) →∙ (→∙EMPath {G = L} (EM∙ H (suc m)) (suc (n + m)) (~ i))) + (isContr→isProp (contr∙-lem n m)) + isOfHLevel↑∙∙ {G = G} {H = H} {L = L} n m (suc l) = + isOfHLevelΩ→isOfHLevel (suc l) + λ f → + subst + (isOfHLevel (2 + l)) (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) + (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) h) + where + h : isOfHLevel (2 + l) + (EM∙ G (suc n) + →∙ (Ω (EM∙ H (suc m) + →∙ EM∙ L (suc (suc (suc (l + n + m)))) ∙))) + h = + subst (isOfHLevel (2 + l)) + (λ i → EM∙ G (suc n) →∙ →∙EMPath {G = L} (EM∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) + (isOfHLevel↑∙∙ n m l) diff --git a/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda b/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda index 945df9935..953e3e4f7 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda @@ -32,17 +32,17 @@ open import Cubical.Functions.Morphism open import Cubical.Foundations.Path private - variable ℓ : Level + variable ℓ ℓ' ℓ'' : Level - wedgeConFun' : (G H : AbGroup ℓ) (n : ℕ) - → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ} + wedgeConFun' : (G : AbGroup ℓ) (H : AbGroup ℓ') (n : ℕ) + → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ''} → ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y)) → (f : (x : _) → A ptS x) → (g : (x : _) → A x ptS) → f ptS ≡ g ptS → (x : _) (y : _) → A x y - wedgeConFun'ᵣ : (G H : AbGroup ℓ) (n : ℕ) - → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ} + wedgeConFun'ᵣ : (G : AbGroup ℓ) (H : AbGroup ℓ') (n : ℕ) + → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ''} → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y))) → (f : (x : _) → A ptS x) → (g : (x : _) → A x ptS) @@ -163,19 +163,21 @@ private P = mainP G H n hLev f g p a i0 ptS a ◁ lem private - wedgeConFun : (G H : AbGroup ℓ) (k n m : ℕ) → (n + m ≡ k) → {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} + wedgeConFun : (G : AbGroup ℓ) (H : AbGroup ℓ') + (k n m : ℕ) → (n + m ≡ k) → {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ''} → ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y)) → (f : (x : _) → A ptS x) → (g : (x : _) → A x ptS) → f ptS ≡ g ptS → (x : _) (y : _) → A x y - wedgeconLeft : (G H : AbGroup ℓ) (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} + wedgeconLeft : (G : AbGroup ℓ) (H : AbGroup ℓ') (k n m : ℕ) (P : n + m ≡ k) + {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ''} → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) → (f : (x : _) → A ptS x) → (g : (x : _) → A x ptS) → (p : f ptS ≡ g ptS) → (x : _) → wedgeConFun G H k n m P hLev f g p ptS x ≡ f x - wedgeconRight : (G H : AbGroup ℓ) (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} + wedgeconRight : (G : AbGroup ℓ) (H : AbGroup ℓ') (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ''} → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) → (f : (x : _) → A ptS x) → (g : (x : _) → A x ptS) @@ -272,7 +274,7 @@ private (λ i → main* G H l n m P hLev f g p a i north a north i) (cong g (merid a)) help = mainR G H l n m P hLev f g p a i0 ptS a ◁ lem -module wedgeConEM (G H : AbGroup ℓ) (n m : ℕ) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} +module wedgeConEM (G : AbGroup ℓ) (H : AbGroup ℓ') (n m : ℕ) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ''} (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) (f : (x : _) → A ptS x) (g : (x : _) → A x ptS) diff --git a/Cubical/Algebra/Group/MorphismProperties.agda b/Cubical/Algebra/Group/MorphismProperties.agda index 373da7b87..fae146ec9 100644 --- a/Cubical/Algebra/Group/MorphismProperties.agda +++ b/Cubical/Algebra/Group/MorphismProperties.agda @@ -77,6 +77,21 @@ module _ {A : Type ℓ} {B : Type ℓ'} {G : GroupStr A} {f : A → B} {H : Grou makeIsGroupHom .pres1 = hom1g G f H pres makeIsGroupHom .presinv = homInv G f H pres +isGroupHomInv : (f : GroupEquiv G H) → IsGroupHom (H .snd) (invEq (fst f)) (G .snd) +isGroupHomInv {G = G} {H = H} f = makeIsGroupHom λ h h' → + isInj-f _ _ + (f' (g (h ⋆² h')) ≡⟨ secEq (fst f) _ ⟩ + (h ⋆² h') ≡⟨ sym (cong₂ _⋆²_ (secEq (fst f) h) (secEq (fst f) h')) ⟩ + (f' (g h) ⋆² f' (g h')) ≡⟨ sym (pres· (snd f) _ _) ⟩ + f' (g h ⋆¹ g h') ∎) + where + f' = fst (fst f) + _⋆¹_ = _·_ (snd G) + _⋆²_ = _·_ (snd H) + g = invEq (fst f) + + isInj-f : (x y : ⟨ G ⟩) → f' x ≡ f' y → x ≡ y + isInj-f x y = invEq (_ , isEquiv→isEmbedding (snd (fst f)) x y) -- H-level results isPropIsGroupHom : (G : Group ℓ) (H : Group ℓ') {f : ⟨ G ⟩ → ⟨ H ⟩} @@ -187,22 +202,6 @@ snd (compGroupEquiv f g) = isGroupHomComp (_ , f .snd) (_ , g .snd) invGroupEquiv : GroupEquiv G H → GroupEquiv H G fst (invGroupEquiv f) = invEquiv (fst f) snd (invGroupEquiv f) = isGroupHomInv f - where - isGroupHomInv : (f : GroupEquiv G H) → IsGroupHom (H .snd) (invEq (fst f)) (G .snd) - isGroupHomInv {G = G} {H = H} f = makeIsGroupHom λ h h' → - isInj-f _ _ - (f' (g (h ⋆² h')) ≡⟨ secEq (fst f) _ ⟩ - (h ⋆² h') ≡⟨ sym (cong₂ _⋆²_ (secEq (fst f) h) (secEq (fst f) h')) ⟩ - (f' (g h) ⋆² f' (g h')) ≡⟨ sym (pres· (snd f) _ _) ⟩ - f' (g h ⋆¹ g h') ∎) - where - f' = fst (fst f) - _⋆¹_ = _·_ (snd G) - _⋆²_ = _·_ (snd H) - g = invEq (fst f) - - isInj-f : (x y : ⟨ G ⟩) → f' x ≡ f' y → x ≡ y - isInj-f x y = invEq (_ , isEquiv→isEmbedding (snd (fst f)) x y) GroupEquivDirProd : {A : Group ℓ} {B : Group ℓ'} {C : Group ℓ''} {D : Group ℓ'''} → GroupEquiv A C → GroupEquiv B D @@ -229,10 +228,10 @@ snd (compGroupIso iso1 iso2) = isGroupHomComp (_ , snd iso1) (_ , snd iso2) invGroupIso : GroupIso G H → GroupIso H G fst (invGroupIso iso1) = invIso (fst iso1) -snd (invGroupIso iso1) = isGroupHomInv iso1 +snd (invGroupIso iso1) = isGroupHomInv' iso1 where - isGroupHomInv : (f : GroupIso G H) → IsGroupHom (H .snd) (inv (fst f)) (G .snd) - isGroupHomInv {G = G} {H = H} f = makeIsGroupHom λ h h' → + isGroupHomInv' : (f : GroupIso G H) → IsGroupHom (H .snd) (inv (fst f)) (G .snd) + isGroupHomInv' {G = G} {H = H} f = makeIsGroupHom λ h h' → isInj-f _ _ (f' (g (h ⋆² h')) ≡⟨ (rightInv (fst f)) _ ⟩ (h ⋆² h') ≡⟨ sym (cong₂ _⋆²_ (rightInv (fst f) h) (rightInv (fst f) h')) ⟩ diff --git a/Cubical/Foundations/GroupoidLaws.agda b/Cubical/Foundations/GroupoidLaws.agda index 8440e8e10..7da8f33c6 100644 --- a/Cubical/Foundations/GroupoidLaws.agda +++ b/Cubical/Foundations/GroupoidLaws.agda @@ -201,6 +201,15 @@ cong-∙ : ∀ {B : Type ℓ} (f : A → B) (p : x ≡ y) (q : y ≡ z) → cong f (p ∙ q) ≡ (cong f p) ∙ (cong f q) cong-∙ f p q = cong-∙∙ f refl p q +cong-compPathP : ∀ {B : A → Type ℓ} (f : (x : A) → B x) (p : x ≡ y) (q : y ≡ z) + → cong f (p ∙ q) ≡ compPathP' {B = B} (cong f p) (cong f q) +cong-compPathP {x = x} {B = B} f p q i j = + comp (λ k → B (compPath-filler p q k j)) + (λ k → λ { (j = i0) → f x + ; (j = i1) → f (q k) + ; (i = i0) → f (compPath-filler p q k j) }) + (f (p j)) + hcomp-unique : ∀ {ℓ} {A : Type ℓ} {φ} → (u : I → Partial φ A) → (u0 : A [ φ ↦ u i0 ]) → (h2 : ∀ i → A [ (φ ∨ ~ i) ↦ (\ { (φ = i1) → u i 1=1; (i = i0) → outS u0}) ]) → (hcomp u (outS u0) ≡ outS (h2 i1)) [ φ ↦ (\ { (φ = i1) → (\ i → u i1 1=1)}) ] diff --git a/Cubical/Foundations/Pointed/Homogeneous.agda b/Cubical/Foundations/Pointed/Homogeneous.agda index 5315e2685..a6de4f062 100644 --- a/Cubical/Foundations/Pointed/Homogeneous.agda +++ b/Cubical/Foundations/Pointed/Homogeneous.agda @@ -6,13 +6,14 @@ Portions of this file adapted from Nicolai Kraus' code here: https://bitbucket.org/nicolaikraus/agda/src/e30d70c72c6af8e62b72eefabcc57623dd921f04/trunc-inverse.lagda -} -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.Foundations.Pointed.Homogeneous where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport open import Cubical.Foundations.Path open import Cubical.Data.Sigma open import Cubical.Data.Empty as ⊥ @@ -26,11 +27,14 @@ open import Cubical.Structures.Pointed isHomogeneous : ∀ {ℓ} → Pointed ℓ → Type (ℓ-suc ℓ) isHomogeneous {ℓ} (A , x) = ∀ y → Path (Pointed ℓ) (A , x) (A , y) +_→∙Dep_ : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : typ A → Pointed ℓ') → Type _ +A →∙Dep B = Σ[ f ∈ ((x : typ A) → B x .fst) ] f (snd A) ≡ snd (B (snd A)) + -- Pointed functions into a homogeneous type are equal as soon as they are equal -- as unpointed functions →∙Homogeneous≡ : ∀ {ℓ ℓ'} {A∙ : Pointed ℓ} {B∙ : Pointed ℓ'} {f∙ g∙ : A∙ →∙ B∙} (h : isHomogeneous B∙) → f∙ .fst ≡ g∙ .fst → f∙ ≡ g∙ -→∙Homogeneous≡ {A∙ = A∙@(_ , a₀)} {B∙@(B , _)} {f∙@(_ , f₀)} {g∙@(_ , g₀)} h p = +→∙Homogeneous≡ {A∙ = A∙@(A , a₀)} {B∙@(B , b)} {f∙@(_ , f₀)} {g∙@(_ , g₀)} h p = subst (λ Q∙ → PathP (λ i → A∙ →∙ Q∙ i) f∙ g∙) (sym (flipSquare fix)) badPath where badPath : PathP (λ i → A∙ →∙ (B , (sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) f∙ g∙ @@ -46,6 +50,49 @@ isHomogeneous {ℓ} (A , x) = ∀ y → Path (Pointed ℓ) (A , x) (A , y) }) (sym (h (pt B∙)) ∙ h ((sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) +→∙Homogeneous≡Path : ∀ {ℓ ℓ'} {A∙ : Pointed ℓ} {B∙ : Pointed ℓ'} {f∙ g∙ : A∙ →∙ B∙} + (h : isHomogeneous B∙) → (p q : f∙ ≡ g∙) → cong fst p ≡ cong fst q → p ≡ q +→∙Homogeneous≡Path {A∙ = A∙@(A , a₀)} {B∙@(B , b)} {f∙@(f , f₀)} {g∙@(g , g₀)} h p q r = + transport (λ k + → PathP (λ i + → PathP (λ j → (A , a₀) →∙ newPath-refl p q r i j (~ k)) + (f , f₀) (g , g₀)) p q) + (badPath p q r) + where + newPath : (p q : f∙ ≡ g∙) (r : cong fst p ≡ cong fst q) + → Square (refl {x = b}) refl refl refl + newPath p q r i j = + hcomp (λ k → λ {(i = i0) → cong snd p j k + ; (i = i1) → cong snd q j k + ; (j = i0) → f₀ k + ; (j = i1) → g₀ k}) + (r i j a₀) + + newPath-refl : (p q : f∙ ≡ g∙) (r : cong fst p ≡ cong fst q) + → PathP (λ i → (PathP (λ j → B∙ ≡ (B , newPath p q r i j))) refl refl) refl refl + newPath-refl p q r i j k = + hcomp (λ w → λ { (i = i0) → lCancel (h b) w k + ; (i = i1) → lCancel (h b) w k + ; (j = i0) → lCancel (h b) w k + ; (j = i1) → lCancel (h b) w k + ; (k = i0) → lCancel (h b) w k + ; (k = i1) → B , newPath p q r i j}) + ((sym (h b) ∙ h (newPath p q r i j)) k) + + badPath : (p q : f∙ ≡ g∙) (r : cong fst p ≡ cong fst q) + → PathP (λ i → + PathP (λ j → A∙ →∙ (B , newPath p q r i j)) + (f , f₀) (g , g₀)) + p q + fst (badPath p q r i j) = r i j + snd (badPath p q s i j) k = + hcomp (λ r → λ { (i = i0) → snd (p j) (r ∧ k) + ; (i = i1) → snd (q j) (r ∧ k) + ; (j = i0) → f₀ (k ∧ r) + ; (j = i1) → g₀ (k ∧ r) + ; (k = i0) → s i j a₀}) + (s i j a₀) + isHomogeneousPi : ∀ {ℓ ℓ'} {A : Type ℓ} {B∙ : A → Pointed ℓ'} → (∀ a → isHomogeneous (B∙ a)) → isHomogeneous (Πᵘ∙ A B∙) isHomogeneousPi h f i .fst = ∀ a → typ (h a (f a) i) From cd8630f948881b2c68bf910f6b2ba31f7c14b171 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 12 Aug 2021 02:31:43 +0200 Subject: [PATCH 20/30] stuff --- .../Group/EilenbergMacLane/CupProduct.agda | 832 ++++++++++-------- Cubical/HITs/SmashProduct/Base.agda | 314 ++++++- 2 files changed, 780 insertions(+), 366 deletions(-) diff --git a/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda b/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda index 59d77b9ca..e77da3c20 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda @@ -45,368 +45,470 @@ open import Cubical.Foundations.Pointed.Homogeneous private variable - ℓ ℓ' : Level - -stupidInd : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 → A 1 - → ((n : ℕ) → (A (suc n) → A (suc (suc n)))) - → (n : ℕ) → A n -stupidInd a0 a1 ind zero = a0 -stupidInd a0 a1 ind (suc zero) = a1 -stupidInd {A = A} a0 a1 ind (suc (suc n)) = ind n (stupidInd {A = A} a0 a1 ind (suc n)) - -stupidInd' : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 - → ((n : ℕ) → (A n → A (suc n))) - → (n : ℕ) → A n -stupidInd' a0 ind zero = a0 -stupidInd' a0 ind (suc n) = ind n (stupidInd' a0 ind n) - -_+'_ : ℕ → ℕ → ℕ -zero +' m = m -suc n +' zero = suc n -suc n +' suc m = suc (suc (n + m)) - -+'-comm : (n m : ℕ) → n +' m ≡ m +' n -+'-comm zero zero = refl -+'-comm zero (suc m) = refl -+'-comm (suc n) zero = refl -+'-comm (suc n) (suc m) = cong (2 +_) (+-comm n m) - -+'≡+ : (n m : ℕ) → n +' m ≡ n + m -+'≡+ zero zero = refl -+'≡+ zero (suc m) = refl -+'≡+ (suc n) zero = cong suc (+-comm zero n) -+'≡+ (suc n) (suc m) = cong suc (sym (+-suc n m)) - -open AbGroupStr renaming (_+_ to _+Gr_ ; -_ to -Gr_) - -inducedFun-EM-raw : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} - → AbGroupHom G' H' - → ∀ n - → EM-raw G' n → EM-raw H' n -inducedFun-EM-raw f = - stupidInd (fst f) - (EMrec _ emsquash embase - (λ g → emloop (fst f g)) - λ g h → compPathR→PathP (sym - (sym (lUnit _) - ∙∙ cong (_∙ (sym (emloop (fst f h)))) (cong emloop (IsGroupHom.pres· (snd f) g h) - ∙ emloop-comp _ (fst f g) (fst f h)) - ∙∙ sym (∙assoc _ _ _) - ∙∙ cong (emloop (fst f g) ∙_) (rCancel _) - ∙∙ sym (rUnit _)))) - (λ n ind → λ { north → north - ; south → south - ; (merid a i) → merid (ind a) i} ) - -inducedFun-EM-raw-sect : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} - → (f : AbGroupEquiv G' H') - → ∀ n - → (x : EM-raw G' n) → inducedFun-EM-raw (invEq (fst f) , isGroupHomInv f) n - (inducedFun-EM-raw (fst (fst f) , snd f) n x) - ≡ x -inducedFun-EM-raw-sect f = - stupidInd (λ x → {!!}) - {!!} - {!!} - -inducedFun-EM-rawIso : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} - → AbGroupEquiv G' H' - → ∀ n - → Iso (EM-raw G' n) (EM-raw H' n) -inducedFun-EM-rawIso e n = {!!} - -EM⊗-raw-commFun : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} → ∀ n → EM-raw (G' ⨂ H') n → EM-raw (H' ⨂ G') n -EM⊗-raw-commFun = - stupidInd commFun - (EMrec _ emsquash embase - (λ g → emloop (commFun g)) - λ g h → compPathR→PathP (sym (sym (lUnit _) - ∙∙ cong (_∙ sym (emloop (commFun h))) (emloop-comp _ (commFun g) (commFun h)) - ∙∙ sym (∙assoc _ _ _) - ∙∙ cong (emloop (commFun g) ∙_) (rCancel _) - ∙∙ sym (rUnit _)))) - (λ n ind → λ { north → north - ; south → south - ; (merid a i) → merid (ind a) i} ) - -EM⊗-raw-comm-sect : {!!} -EM⊗-raw-comm-sect = {!!} - -module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where - G = fst G' - H = fst H' - - strG = snd G' - strH = snd H' - - 0G = 0g strG - 0H = 0g strH - - _+G_ = _+Gr_ strG - _+H_ = _+Gr_ strH - - -H_ = -Gr_ strH - -G_ = -Gr_ strG - - - - precup : ∀ n m → EM-raw G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m) - precup zero m = {!!} - precup (suc zero) m = {!!} - precup (suc (suc n)) m = {!!} - - - - ·₀' : H → (m : ℕ) → EM G' m → EM (G' ⨂ H') m - ·₀' h = - stupidInd - (_⊗ h) - (elimGroupoid _ (λ _ → emsquash) - embase - (λ g → emloop (g ⊗ h)) - λ g l → - compPathR→PathP - (sym (∙assoc _ _ _ - ∙∙ cong₂ _∙_ (sym (emloop-comp _ _ _) ∙ cong emloop (sym (linl g l h))) refl - ∙∙ rCancel _))) - λ n f → trRec (isOfHLevelTrunc (4 + n)) - λ { north → 0ₖ (suc (suc n)) - ; south → 0ₖ (suc (suc n)) - ; (merid a i) → EM→ΩEM+1 (suc n) (f (EM-raw→EM _ _ a)) i} - - ·₀ : G → (m : ℕ) → EM H' m → EM (G' ⨂ H') m - ·₀ g = - stupidInd (λ h → g ⊗ h) - (elimGroupoid _ (λ _ → emsquash) - embase - (λ h → emloop (g ⊗ h)) - λ h l → compPathR→PathP - (sym (∙assoc _ _ _ - ∙∙ cong₂ _∙_ (sym (emloop-comp _ _ _) ∙ cong emloop (sym (linr g h l))) refl - ∙∙ rCancel _))) - λ n f - → trRec (isOfHLevelTrunc (4 + n)) - λ { north → 0ₖ (suc (suc n)) - ; south → 0ₖ (suc (suc n)) - ; (merid a i) → EM→ΩEM+1 (suc n) (f (EM-raw→EM _ _ a)) i} - - ·₀-distr : (g h : G) → (m : ℕ) (x : EM H' m) → ·₀ (g +G h) m x ≡ ·₀ g m x +ₖ ·₀ h m x - ·₀-distr g h = - stupidInd - (linl g h) - (elimSet _ (λ _ → emsquash _ _) - refl - (λ w → compPathR→PathP (sym ((λ i → emloop (linl g h w i) - ∙ (lUnit (sym (cong₂+₁ (emloop (g ⊗ w)) (emloop (h ⊗ w)) i)) (~ i))) - ∙∙ cong₂ _∙_ (emloop-comp _ (g ⊗ w) (h ⊗ w)) refl - ∙∙ rCancel _)))) - λ m ind → - trElim (λ _ → isOfHLevelTruncPath) - λ { north → refl - ; south → refl - ; (merid a i) k → z m ind a k i} - - where - z : (m : ℕ) → ((x : EM H' (suc m)) → ·₀ (g +G h) (suc m) x ≡ ·₀ g (suc m) x +ₖ ·₀ h (suc m) x) → (a : EM-raw H' (suc m)) - → cong (·₀ (g +G h) (suc (suc m))) (cong ∣_∣ₕ (merid a)) ≡ - cong₂ _+ₖ_ - (cong (·₀ g (suc (suc m))) (cong ∣_∣ₕ (merid a))) - (cong (·₀ h (suc (suc m))) (cong ∣_∣ₕ (merid a))) - z m ind a = (λ i → EM→ΩEM+1 _ (ind (EM-raw→EM _ _ a) i)) - ∙∙ EM→ΩEM+1-hom _ (·₀ g (suc m) (EM-raw→EM H' (suc m) a)) (·₀ h (suc m) (EM-raw→EM H' (suc m) a)) - ∙∙ sym (cong₂+₂ m (cong (·₀ g (suc (suc m))) (cong ∣_∣ₕ (merid a))) - (cong (·₀ h (suc (suc m))) (cong ∣_∣ₕ (merid a)))) - - ·₀0 : (m : ℕ) → (g : G) → ·₀ g m (0ₖ m) ≡ 0ₖ m - ·₀0 zero = rCancelPrim - ·₀0 (suc zero) g = refl - ·₀0 (suc (suc m)) g = refl - - ·₀'0 : (m : ℕ) (h : H) → ·₀' h m (0ₖ m) ≡ 0ₖ m - ·₀'0 zero = lCancelPrim - ·₀'0 (suc zero) g = refl - ·₀'0 (suc (suc m)) g = refl - - 0·₀ : (m : ℕ) → (x : _) → ·₀ 0G m x ≡ 0ₖ m - 0·₀ = - stupidInd lCancelPrim - (elimSet _ (λ _ → emsquash _ _) - refl - λ g → compPathR→PathP ((sym (emloop-1g _) - ∙ cong emloop (sym (lCancelPrim g))) - ∙∙ (λ i → rUnit (rUnit (cong (·₀ 0G 1) (emloop g)) i) i) - ∙∙ sym (∙assoc _ _ _))) - λ n f → trElim (λ _ → isOfHLevelTruncPath) - λ { north → refl - ; south → refl - ; (merid a i) j → (cong (EM→ΩEM+1 (suc n)) (f (EM-raw→EM _ _ a)) - ∙ EM→ΩEM+1-0ₖ _) j i} - - 0·₀' : (m : ℕ) (g : _) → ·₀' 0H m g ≡ 0ₖ m - 0·₀' = - stupidInd - rCancelPrim - (elimSet _ (λ _ → emsquash _ _) - refl - λ g → compPathR→PathP (sym (∙assoc _ _ _ - ∙∙ sym (rUnit _) ∙ sym (rUnit _) - ∙∙ (cong emloop (rCancelPrim g) - ∙ emloop-1g _)))) - λ n f → trElim (λ _ → isOfHLevelTruncPath) - λ { north → refl - ; south → refl - ; (merid a i) j → (cong (EM→ΩEM+1 (suc n)) (f (EM-raw→EM _ _ a)) - ∙ EM→ΩEM+1-0ₖ _) j i} - - cup∙ : ∀ n m → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m) - cup∙ = - stupidInd' - (λ m g → (·₀ g m) , ·₀0 m g) - λ n f → - stupidInd' - (λ g → (λ h → ·₀' h (suc n) g) , 0·₀' (suc n) g) - λ m _ → main n m f - - where - main : (n m : ℕ) (ind : ((m : ℕ) → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m))) - → EM G' (suc n) → EM∙ H' (suc m) →∙ EM∙ (G' ⨂ H') (suc (suc (n + m))) - main zero m ind = - elimGroupoid _ (λ _ → isOfHLevel↑∙ _ _) - ((λ _ → 0ₖ (2 + m)) , refl) - (f m) - λ n h → finalpp m n h - where - f : (m : ℕ) → G → typ (Ω (EM∙ H' (suc m) →∙ EM∙ (G' ⨂ H') (suc (suc m)) ∙)) - fst (f m g i) x = EM→ΩEM+1 _ (·₀ g _ x) i - snd (f zero g i) j = EM→ΩEM+1-0ₖ (suc zero) j i - snd (f (suc m) g i) j = EM→ΩEM+1-0ₖ (suc (suc m)) j i - - f-hom-fst : (m : ℕ) (g h : G) → cong fst (f m (g +G h)) ≡ cong fst (f m g ∙ f m h) - f-hom-fst m g h = - (λ i j x → EM→ΩEM+1 _ (·₀-distr g h (suc m) x i) j) - ∙∙ (λ i j x → EM→ΩEM+1-hom _ (·₀ g (suc m) x) (·₀ h (suc m) x) i j) - ∙∙ sym (cong-∙ fst (f m g) (f m h)) - - f-hom : (m : ℕ) (g h : G) → f m (g +G h) ≡ f m g ∙ f m h - f-hom m g h = →∙Homogeneous≡Path (isHomogeneousEM _) _ _ (f-hom-fst m g h) - - finalpp : (m : ℕ) (g h : G) → PathP (λ i → f m g i ≡ f m (g +G h) i) refl (f m h) - finalpp m g h = - compPathR→PathP (sym (rCancel _) - ∙∙ cong (_∙ sym (f m (g +G h))) (f-hom m g h) - ∙∙ sym (∙assoc _ _ _)) - - main (suc n) m ind = - trElim (λ _ → isOfHLevel↑∙ (2 + n) m) - λ { north → (λ _ → 0ₖ (3 + (n + m))) , refl - ; south → (λ _ → 0ₖ (3 + (n + m))) , refl - ; (merid a i) → (λ x → EM→ΩEM+1 _ (ind _ (EM-raw→EM _ _ a) .fst x) i) - , λ j → (cong (EM→ΩEM+1 (suc (suc (n + m)))) (ind (suc m) (EM-raw→EM G' (suc n) a) .snd) - ∙ EM→ΩEM+1-0ₖ _) j i} - - _⌣ₖ_ : {n m : ℕ} (x : EM G' n) (y : EM H' m) → EM (G' ⨂ H') (n +' m) - _⌣ₖ_ x y = cup∙ _ _ x .fst y - - ⌣ₖ-0ₖ : (n m : ℕ) (x : EM G' n) → (x ⌣ₖ 0ₖ m) ≡ 0ₖ (n +' m) - ⌣ₖ-0ₖ n m x = cup∙ n m x .snd - - 0ₖ-⌣ₖ : (n m : ℕ) (x : EM H' m) → ((0ₖ n) ⌣ₖ x) ≡ 0ₖ (n +' m) - 0ₖ-⌣ₖ zero m = 0·₀ _ - 0ₖ-⌣ₖ (suc zero) zero x = refl - 0ₖ-⌣ₖ (suc (suc n)) zero x = refl - 0ₖ-⌣ₖ (suc zero) (suc m) x = refl - 0ₖ-⌣ₖ (suc (suc n)) (suc m) x = refl - - private - distrl1 : (n m : ℕ) → EM H' m → EM H' m → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) - fst (distrl1 n m x y) z = z ⌣ₖ (x +ₖ y) - snd (distrl1 n m x y) = 0ₖ-⌣ₖ n m _ - - distrl2 : (n m : ℕ) → EM H' m → EM H' m → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) - fst (distrl2 n m x y) z = (z ⌣ₖ x) +ₖ (z ⌣ₖ y) - snd (distrl2 n m x y) = cong₂ _+ₖ_ (0ₖ-⌣ₖ n m x) (0ₖ-⌣ₖ n m y) ∙ rUnitₖ _ (0ₖ (n +' m)) - - hLevLem : (n m : ℕ) → isOfHLevel (suc (suc m)) (EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') ((suc n) +' m)) - hLevLem n m = subst (isOfHLevel (suc (suc m))) (λ i → EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') - ((cong suc (+-comm m n) ∙ sym (+'≡+ (suc n) m)) i)) (isOfHLevel↑∙ m n) - - - pathType : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → Type ℓ - pathType n x p = sym (rUnitₖ (2 + n) x) ∙ (λ i → x +ₖ p i) ≡ sym (lUnitₖ (2 + n) x) ∙ λ i → p i +ₖ x - - lem : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) - → pathType n x p - lem n x = J (λ x p → pathType n x p) refl - - mainDistrL : (n m : ℕ) (x y : EM H' (suc m)) → distrl1 (suc n) (suc m) x y ≡ distrl2 (suc n) (suc m) x y - mainDistrL n zero = - wedgeConEM.fun H' H' 0 0 - (λ _ _ → hLevLem _ _ _ _) - (λ x → →∙Homogeneous≡ (isHomogeneousEM _) - (funExt λ z → l x z)) - (λ y → →∙Homogeneous≡ (isHomogeneousEM _) - (funExt λ z → r y z )) - λ i → →∙Homogeneous≡ (isHomogeneousEM (suc (suc (n + 0)))) - (funExt (λ z → l≡r z i)) - where - l : (x : EM H' 1) (z : _) → (distrl1 (suc n) 1 embase x .fst z) ≡ (distrl2 (suc n) 1 embase x .fst z) - l x z = cong (z ⌣ₖ_) (lUnitₖ _ x) - ∙∙ sym (lUnitₖ _ (z ⌣ₖ x)) - ∙∙ λ i → (⌣ₖ-0ₖ (suc n) (suc zero) z (~ i)) +ₖ (z ⌣ₖ x) - - r : (z : EM H' 1) (x : EM G' (suc n)) → (distrl1 (suc n) 1 z embase .fst x) ≡ (distrl2 (suc n) 1 z embase .fst x) - r y z = cong (z ⌣ₖ_) (rUnitₖ _ y) - ∙∙ sym (rUnitₖ _ (z ⌣ₖ y)) - ∙∙ λ i → (z ⌣ₖ y) +ₖ (⌣ₖ-0ₖ (suc n) (suc zero) z (~ i)) - - l≡r : (z : EM G' (suc n)) → l embase z ≡ r embase z - l≡r z = sym (lem _ _ (sym (⌣ₖ-0ₖ (suc n) (suc zero) z))) - - mainDistrL n (suc m) = - elim2 (λ _ _ → isOfHLevelPath (4 + m) (hLevLem _ _) _ _) - (wedgeConEM.fun H' H' (suc m) (suc m) - (λ x y p q → isOfHLevelPlus {n = suc (suc m)} (suc m) - (hLevLem n (suc (suc m)) - (distrl1 (suc n) (suc (suc m)) ∣ x ∣ ∣ y ∣) - (distrl2 (suc n) (suc (suc m)) ∣ x ∣ ∣ y ∣) p q)) - (λ x → →∙Homogeneous≡ (isHomogeneousEM _) - (funExt (l x))) - (λ x → →∙Homogeneous≡ (isHomogeneousEM _) - (funExt (r x))) - λ i → →∙Homogeneous≡ (isHomogeneousEM _) - (funExt (λ z → l≡r z i))) - where - l : (x : EM-raw H' (suc (suc m))) (z : EM G' (suc n)) - → (distrl1 (suc n) (suc (suc m)) (0ₖ _) ∣ x ∣ₕ .fst z) ≡ (distrl2 (suc n) (suc (suc m)) (0ₖ _) ∣ x ∣ₕ .fst z) - l x z = cong (z ⌣ₖ_) (lUnitₖ (suc (suc m)) ∣ x ∣) - ∙∙ sym (lUnitₖ _ (z ⌣ₖ ∣ x ∣)) - ∙∙ λ i → (⌣ₖ-0ₖ (suc n) (suc (suc m)) z (~ i)) +ₖ (z ⌣ₖ ∣ x ∣) - - r : (x : EM-raw H' (suc (suc m))) (z : EM G' (suc n)) - → (distrl1 (suc n) (suc (suc m)) ∣ x ∣ₕ (0ₖ _) .fst z) ≡ (distrl2 (suc n) (suc (suc m)) ∣ x ∣ₕ (0ₖ _) .fst z) - r x z = cong (z ⌣ₖ_) (rUnitₖ (suc (suc m)) ∣ x ∣) - ∙∙ sym (rUnitₖ _ (z ⌣ₖ ∣ x ∣)) - ∙∙ λ i → (z ⌣ₖ ∣ x ∣) +ₖ (⌣ₖ-0ₖ (suc n) (suc (suc m)) z (~ i)) - - l≡r : (z : EM G' (suc n)) → l north z ≡ r north z - l≡r z = sym (lem _ _ (sym (⌣ₖ-0ₖ (suc n) (suc (suc m)) z))) - - - commer : {!!} - commer = {!!} - - private - distrr1 : (n m : ℕ) → EM H' n → EM H' n → EM∙ G' m →∙ EM∙ (H' ⨂ G') (n +' m) - fst (distrr1 n m x y) z = {!!} -- (x +ₖ y) ⌣ₖ ? - snd (distrr1 n m x y) = {!!} -- ⌣ₖ-0ₖ n m _ - - distrr2 : (n m : ℕ) → EM H' m → EM H' m → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) - fst (distrr2 n m x y) z = (z ⌣ₖ x) +ₖ (z ⌣ₖ y) - snd (distrr2 n m x y) = cong₂ _+ₖ_ (0ₖ-⌣ₖ n m x) (0ₖ-⌣ₖ n m y) ∙ rUnitₖ _ (0ₖ (n +' m)) - - hLevLem2 : (n m : ℕ) → isOfHLevel (suc (suc m)) (EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') ((suc n) +' m)) - hLevLem2 n m = subst (isOfHLevel (suc (suc m))) (λ i → EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') - ((cong suc (+-comm m n) ∙ sym (+'≡+ (suc n) m)) i)) (isOfHLevel↑∙ m n) - - mainDistrR : {!(n m : ℕ) (x y : EM H' (suc m)) → distr1 (suc n) (suc m) x y ≡ distrl2 (suc n) (suc m) x y!} - mainDistrR = {!!} + ℓ ℓ' ℓ'' : Level + +open import Cubical.Foundations.SIP +univP : ∀ {ℓ} (A B : Type ℓ) → Type _ +univP {ℓ = ℓ} A B = (C : Type ℓ) → Iso (A → C) (B → C) + +univP' : ∀ {ℓ} (A B : Type ℓ) → Type _ +univP' {ℓ = ℓ} A B = Σ[ f ∈ ((C : Type ℓ) → Iso (A → C) (B → C)) ] Iso.fun (f B) ≡ λ _ b → b + +univP'' : ∀ {ℓ} (A B : Type ℓ) → Type _ +univP'' {ℓ = ℓ} A B = (P : (C : Type ℓ) → Type ℓ) → Iso ((x : A) → (P A)) ((x : B) → (P B)) + +univP→Id : ∀ {ℓ} (A B : Pointed ℓ) → univP'' (typ A) (typ B) → Iso (typ A) (typ B) +univP→Id A B is = Iso.inv (is (λ C → Iso C (typ B))) (λ _ → idIso) (pt A) + +open import Cubical.HITs.SmashProduct + +test : ∀ {ℓ} (A B C : Pointed ℓ) → univP'' (Smash (SmashPt A B) C) (Smash A (SmashPt B C)) +test A B C P = {!!} + where + h' : (Smash (SmashPt A B) C → P (Smash (SmashPt A B) C)) + h' basel = {!!} + h' baser = {!!} + h' (proj x y) = {!!} + h' (gluel a i) = {!!} + h' (gluer b i) = {!!} + + h : (Smash (SmashPt A B) C → P (Smash (SmashPt A B) C)) → Σ[ l ∈ P (Smash (SmashPt A B) C) ] {!!} + h = {!!} + + +-- stupidInd : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 → A 1 +-- → ((n : ℕ) → (A (suc n) → A (suc (suc n)))) +-- → (n : ℕ) → A n +-- stupidInd a0 a1 ind zero = a0 +-- stupidInd a0 a1 ind (suc zero) = a1 +-- stupidInd {A = A} a0 a1 ind (suc (suc n)) = ind n (stupidInd {A = A} a0 a1 ind (suc n)) + +-- stupidInd' : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 +-- → ((n : ℕ) → (A n → A (suc n))) +-- → (n : ℕ) → A n +-- stupidInd' a0 ind zero = a0 +-- stupidInd' a0 ind (suc n) = ind n (stupidInd' a0 ind n) + +-- _+'_ : ℕ → ℕ → ℕ +-- zero +' m = m +-- suc n +' zero = suc n +-- suc n +' suc m = suc (suc (n + m)) + +-- +'-comm : (n m : ℕ) → n +' m ≡ m +' n +-- +'-comm zero zero = refl +-- +'-comm zero (suc m) = refl +-- +'-comm (suc n) zero = refl +-- +'-comm (suc n) (suc m) = cong (2 +_) (+-comm n m) + +-- +'≡+ : (n m : ℕ) → n +' m ≡ n + m +-- +'≡+ zero zero = refl +-- +'≡+ zero (suc m) = refl +-- +'≡+ (suc n) zero = cong suc (+-comm zero n) +-- +'≡+ (suc n) (suc m) = cong suc (sym (+-suc n m)) + +-- open AbGroupStr renaming (_+_ to _+Gr_ ; -_ to -Gr_) + +-- inducedFun-EM-raw : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} +-- → AbGroupHom G' H' +-- → ∀ n +-- → EM-raw G' n → EM-raw H' n +-- inducedFun-EM-raw f = +-- stupidInd (fst f) +-- (EMrec _ emsquash embase +-- (λ g → emloop (fst f g)) +-- λ g h → compPathR→PathP (sym +-- (sym (lUnit _) +-- ∙∙ cong (_∙ (sym (emloop (fst f h)))) (cong emloop (IsGroupHom.pres· (snd f) g h) +-- ∙ emloop-comp _ (fst f g) (fst f h)) +-- ∙∙ sym (∙assoc _ _ _) +-- ∙∙ cong (emloop (fst f g) ∙_) (rCancel _) +-- ∙∙ sym (rUnit _)))) +-- (λ n ind → λ { north → north +-- ; south → south +-- ; (merid a i) → merid (ind a) i} ) + +-- inducedFun-EM-rawIso : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} +-- → AbGroupEquiv G' H' +-- → ∀ n → Iso (EM-raw G' n) (EM-raw H' n) +-- Iso.fun (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , (snd e)) n +-- Iso.inv (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , isGroupHomInv e) n +-- Iso.rightInv (inducedFun-EM-rawIso e n) = h n +-- where +-- h : (n : ℕ) → section (inducedFun-EM-raw (fst e .fst , snd e) n) +-- (inducedFun-EM-raw (invEq (fst e) , isGroupHomInv e) n) +-- h = stupidInd +-- (secEq (fst e)) +-- (elimSet _ (λ _ → emsquash _ _) refl +-- (λ g → compPathR→PathP +-- ((sym (cong₂ _∙_ (cong emloop (secEq (fst e) g)) (sym (lUnit _)) +-- ∙ rCancel _))))) +-- λ n p → λ { north → refl +-- ; south → refl +-- ; (merid a i) k → merid (p a k) i} +-- Iso.leftInv (inducedFun-EM-rawIso e n) = h n +-- where +-- h : (n : ℕ) → retract (Iso.fun (inducedFun-EM-rawIso e n)) +-- (Iso.inv (inducedFun-EM-rawIso e n)) +-- h = stupidInd +-- (retEq (fst e)) +-- (elimSet _ (λ _ → emsquash _ _) refl +-- (λ g → compPathR→PathP +-- ((sym (cong₂ _∙_ (cong emloop (retEq (fst e) g)) (sym (lUnit _)) +-- ∙ rCancel _))))) +-- λ n p → λ { north → refl +-- ; south → refl +-- ; (merid a i) k → merid (p a k) i} + +-- Iso→EMIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} → AbGroupEquiv G H → ∀ n → Iso (EM G n) (EM H n) +-- Iso→EMIso is zero = inducedFun-EM-rawIso is zero +-- Iso→EMIso is (suc zero) = inducedFun-EM-rawIso is 1 +-- Iso→EMIso is (suc (suc n)) = mapCompIso (inducedFun-EM-rawIso is (suc (suc n))) + +-- EM⊗-commFun : {G : AbGroup ℓ} {H : AbGroup ℓ'} → ∀ n → Iso (EM (G ⨂ H) n) (EM (H ⨂ G) n) +-- EM⊗-commFun {G = G} {H = H} = Iso→EMIso (GroupIso→GroupEquiv ⨂-commIso) + +-- EM⊗-assocIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} → ∀ n → Iso (EM (G ⨂ (H ⨂ L)) n) (EM ((G ⨂ H) ⨂ L) n) +-- EM⊗-assocIso = Iso→EMIso (GroupIso→GroupEquiv (GroupEquiv→GroupIso ⨂assoc)) + +-- EM⊗-raw-comm-sect : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} → ∀ n → Iso (EM (G ⨂ H) n) (EM (H ⨂ G) n) +-- EM⊗-raw-comm-sect = {!!} + +-- pathType : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → Type ℓ +-- pathType n x p = sym (rUnitₖ (2 + n) x) ∙ (λ i → x +ₖ p i) ≡ sym (lUnitₖ (2 + n) x) ∙ λ i → p i +ₖ x + +-- lem : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) +-- → pathType n x p +-- lem n x = J (λ x p → pathType n x p) refl + +-- module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where +-- private +-- G = fst G' +-- H = fst H' + +-- strG = snd G' +-- strH = snd H' + +-- 0G = 0g strG +-- 0H = 0g strH + +-- _+G_ = _+Gr_ strG +-- _+H_ = _+Gr_ strH + +-- -H_ = -Gr_ strH +-- -G_ = -Gr_ strG + +-- ·₀' : H → (m : ℕ) → EM G' m → EM (G' ⨂ H') m +-- ·₀' h = +-- stupidInd +-- (_⊗ h) +-- (elimGroupoid _ (λ _ → emsquash) +-- embase +-- (λ g → emloop (g ⊗ h)) +-- λ g l → +-- compPathR→PathP +-- (sym (∙assoc _ _ _ +-- ∙∙ cong₂ _∙_ (sym (emloop-comp _ _ _) ∙ cong emloop (sym (linl g l h))) refl +-- ∙∙ rCancel _))) +-- λ n f → trRec (isOfHLevelTrunc (4 + n)) +-- λ { north → 0ₖ (suc (suc n)) +-- ; south → 0ₖ (suc (suc n)) +-- ; (merid a i) → EM→ΩEM+1 (suc n) (f (EM-raw→EM _ _ a)) i} + +-- ·₀ : G → (m : ℕ) → EM H' m → EM (G' ⨂ H') m +-- ·₀ g = +-- stupidInd (λ h → g ⊗ h) +-- (elimGroupoid _ (λ _ → emsquash) +-- embase +-- (λ h → emloop (g ⊗ h)) +-- λ h l → compPathR→PathP +-- (sym (∙assoc _ _ _ +-- ∙∙ cong₂ _∙_ (sym (emloop-comp _ _ _) ∙ cong emloop (sym (linr g h l))) refl +-- ∙∙ rCancel _))) +-- λ n f +-- → trRec (isOfHLevelTrunc (4 + n)) +-- λ { north → 0ₖ (suc (suc n)) +-- ; south → 0ₖ (suc (suc n)) +-- ; (merid a i) → EM→ΩEM+1 (suc n) (f (EM-raw→EM _ _ a)) i} + +-- ·₀-distr : (g h : G) → (m : ℕ) (x : EM H' m) → ·₀ (g +G h) m x ≡ ·₀ g m x +ₖ ·₀ h m x +-- ·₀-distr g h = +-- stupidInd +-- (linl g h) +-- (elimSet _ (λ _ → emsquash _ _) +-- refl +-- (λ w → compPathR→PathP (sym ((λ i → emloop (linl g h w i) +-- ∙ (lUnit (sym (cong₂+₁ (emloop (g ⊗ w)) (emloop (h ⊗ w)) i)) (~ i))) +-- ∙∙ cong₂ _∙_ (emloop-comp _ (g ⊗ w) (h ⊗ w)) refl +-- ∙∙ rCancel _)))) +-- λ m ind → +-- trElim (λ _ → isOfHLevelTruncPath) +-- λ { north → refl +-- ; south → refl +-- ; (merid a i) k → z m ind a k i} + +-- where +-- z : (m : ℕ) → ((x : EM H' (suc m)) → ·₀ (g +G h) (suc m) x ≡ ·₀ g (suc m) x +ₖ ·₀ h (suc m) x) → (a : EM-raw H' (suc m)) +-- → cong (·₀ (g +G h) (suc (suc m))) (cong ∣_∣ₕ (merid a)) ≡ +-- cong₂ _+ₖ_ +-- (cong (·₀ g (suc (suc m))) (cong ∣_∣ₕ (merid a))) +-- (cong (·₀ h (suc (suc m))) (cong ∣_∣ₕ (merid a))) +-- z m ind a = (λ i → EM→ΩEM+1 _ (ind (EM-raw→EM _ _ a) i)) +-- ∙∙ EM→ΩEM+1-hom _ (·₀ g (suc m) (EM-raw→EM H' (suc m) a)) (·₀ h (suc m) (EM-raw→EM H' (suc m) a)) +-- ∙∙ sym (cong₂+₂ m (cong (·₀ g (suc (suc m))) (cong ∣_∣ₕ (merid a))) +-- (cong (·₀ h (suc (suc m))) (cong ∣_∣ₕ (merid a)))) + +-- ·₀0 : (m : ℕ) → (g : G) → ·₀ g m (0ₖ m) ≡ 0ₖ m +-- ·₀0 zero = rCancelPrim +-- ·₀0 (suc zero) g = refl +-- ·₀0 (suc (suc m)) g = refl + +-- ·₀'0 : (m : ℕ) (h : H) → ·₀' h m (0ₖ m) ≡ 0ₖ m +-- ·₀'0 zero = lCancelPrim +-- ·₀'0 (suc zero) g = refl +-- ·₀'0 (suc (suc m)) g = refl + +-- 0·₀ : (m : ℕ) → (x : _) → ·₀ 0G m x ≡ 0ₖ m +-- 0·₀ = +-- stupidInd lCancelPrim +-- (elimSet _ (λ _ → emsquash _ _) +-- refl +-- λ g → compPathR→PathP ((sym (emloop-1g _) +-- ∙ cong emloop (sym (lCancelPrim g))) +-- ∙∙ (λ i → rUnit (rUnit (cong (·₀ 0G 1) (emloop g)) i) i) +-- ∙∙ sym (∙assoc _ _ _))) +-- λ n f → trElim (λ _ → isOfHLevelTruncPath) +-- λ { north → refl +-- ; south → refl +-- ; (merid a i) j → (cong (EM→ΩEM+1 (suc n)) (f (EM-raw→EM _ _ a)) +-- ∙ EM→ΩEM+1-0ₖ _) j i} + +-- 0·₀' : (m : ℕ) (g : _) → ·₀' 0H m g ≡ 0ₖ m +-- 0·₀' = +-- stupidInd +-- rCancelPrim +-- (elimSet _ (λ _ → emsquash _ _) +-- refl +-- λ g → compPathR→PathP (sym (∙assoc _ _ _ +-- ∙∙ sym (rUnit _) ∙ sym (rUnit _) +-- ∙∙ (cong emloop (rCancelPrim g) +-- ∙ emloop-1g _)))) +-- λ n f → trElim (λ _ → isOfHLevelTruncPath) +-- λ { north → refl +-- ; south → refl +-- ; (merid a i) j → (cong (EM→ΩEM+1 (suc n)) (f (EM-raw→EM _ _ a)) +-- ∙ EM→ΩEM+1-0ₖ _) j i} + +-- cup∙ : ∀ n m → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m) +-- cup∙ = +-- stupidInd' +-- (λ m g → (·₀ g m) , ·₀0 m g) +-- λ n f → +-- stupidInd' +-- (λ g → (λ h → ·₀' h (suc n) g) , 0·₀' (suc n) g) +-- λ m _ → main n m f + +-- where +-- main : (n m : ℕ) (ind : ((m : ℕ) → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m))) +-- → EM G' (suc n) → EM∙ H' (suc m) →∙ EM∙ (G' ⨂ H') (suc (suc (n + m))) +-- main zero m ind = +-- elimGroupoid _ (λ _ → isOfHLevel↑∙ _ _) +-- ((λ _ → 0ₖ (2 + m)) , refl) +-- (f m) +-- λ n h → finalpp m n h +-- where +-- f : (m : ℕ) → G → typ (Ω (EM∙ H' (suc m) →∙ EM∙ (G' ⨂ H') (suc (suc m)) ∙)) +-- fst (f m g i) x = EM→ΩEM+1 _ (·₀ g _ x) i +-- snd (f zero g i) j = EM→ΩEM+1-0ₖ (suc zero) j i +-- snd (f (suc m) g i) j = EM→ΩEM+1-0ₖ (suc (suc m)) j i + +-- f-hom-fst : (m : ℕ) (g h : G) → cong fst (f m (g +G h)) ≡ cong fst (f m g ∙ f m h) +-- f-hom-fst m g h = +-- (λ i j x → EM→ΩEM+1 _ (·₀-distr g h (suc m) x i) j) +-- ∙∙ (λ i j x → EM→ΩEM+1-hom _ (·₀ g (suc m) x) (·₀ h (suc m) x) i j) +-- ∙∙ sym (cong-∙ fst (f m g) (f m h)) + +-- f-hom : (m : ℕ) (g h : G) → f m (g +G h) ≡ f m g ∙ f m h +-- f-hom m g h = →∙Homogeneous≡Path (isHomogeneousEM _) _ _ (f-hom-fst m g h) + +-- finalpp : (m : ℕ) (g h : G) → PathP (λ i → f m g i ≡ f m (g +G h) i) refl (f m h) +-- finalpp m g h = +-- compPathR→PathP (sym (rCancel _) +-- ∙∙ cong (_∙ sym (f m (g +G h))) (f-hom m g h) +-- ∙∙ sym (∙assoc _ _ _)) + +-- main (suc n) m ind = +-- trElim (λ _ → isOfHLevel↑∙ (2 + n) m) +-- λ { north → (λ _ → 0ₖ (3 + (n + m))) , refl +-- ; south → (λ _ → 0ₖ (3 + (n + m))) , refl +-- ; (merid a i) → (λ x → EM→ΩEM+1 _ (ind _ (EM-raw→EM _ _ a) .fst x) i) +-- , λ j → (cong (EM→ΩEM+1 (suc (suc (n + m)))) (ind (suc m) (EM-raw→EM G' (suc n) a) .snd) +-- ∙ EM→ΩEM+1-0ₖ _) j i} + +-- _⌣ₖ_ : {n m : ℕ} (x : EM G' n) (y : EM H' m) → EM (G' ⨂ H') (n +' m) +-- _⌣ₖ_ x y = cup∙ _ _ x .fst y + +-- ⌣ₖ-0ₖ : (n m : ℕ) (x : EM G' n) → (x ⌣ₖ 0ₖ m) ≡ 0ₖ (n +' m) +-- ⌣ₖ-0ₖ n m x = cup∙ n m x .snd + +-- 0ₖ-⌣ₖ : (n m : ℕ) (x : EM H' m) → ((0ₖ n) ⌣ₖ x) ≡ 0ₖ (n +' m) +-- 0ₖ-⌣ₖ zero m = 0·₀ _ +-- 0ₖ-⌣ₖ (suc zero) zero x = refl +-- 0ₖ-⌣ₖ (suc (suc n)) zero x = refl +-- 0ₖ-⌣ₖ (suc zero) (suc m) x = refl +-- 0ₖ-⌣ₖ (suc (suc n)) (suc m) x = refl + +-- private +-- distrl1 : (n m : ℕ) → EM H' m → EM H' m → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) +-- fst (distrl1 n m x y) z = z ⌣ₖ (x +ₖ y) +-- snd (distrl1 n m x y) = 0ₖ-⌣ₖ n m _ + +-- distrl2 : (n m : ℕ) → EM H' m → EM H' m → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) +-- fst (distrl2 n m x y) z = (z ⌣ₖ x) +ₖ (z ⌣ₖ y) +-- snd (distrl2 n m x y) = cong₂ _+ₖ_ (0ₖ-⌣ₖ n m x) (0ₖ-⌣ₖ n m y) ∙ rUnitₖ _ (0ₖ (n +' m)) + +-- hLevLem : (n m : ℕ) → isOfHLevel (suc (suc m)) (EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') ((suc n) +' m)) +-- hLevLem n m = subst (isOfHLevel (suc (suc m))) (λ i → EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') +-- ((cong suc (+-comm m n) ∙ sym (+'≡+ (suc n) m)) i)) (isOfHLevel↑∙ m n) + + + + +-- mainDistrL : (n m : ℕ) (x y : EM H' (suc m)) → distrl1 (suc n) (suc m) x y ≡ distrl2 (suc n) (suc m) x y +-- mainDistrL n zero = +-- wedgeConEM.fun H' H' 0 0 +-- (λ _ _ → hLevLem _ _ _ _) +-- (λ x → →∙Homogeneous≡ (isHomogeneousEM _) +-- (funExt λ z → l x z)) +-- (λ y → →∙Homogeneous≡ (isHomogeneousEM _) +-- (funExt λ z → r y z )) +-- λ i → →∙Homogeneous≡ (isHomogeneousEM (suc (suc (n + 0)))) +-- (funExt (λ z → l≡r z i)) +-- where +-- l : (x : EM H' 1) (z : _) → (distrl1 (suc n) 1 embase x .fst z) ≡ (distrl2 (suc n) 1 embase x .fst z) +-- l x z = cong (z ⌣ₖ_) (lUnitₖ _ x) +-- ∙∙ sym (lUnitₖ _ (z ⌣ₖ x)) +-- ∙∙ λ i → (⌣ₖ-0ₖ (suc n) (suc zero) z (~ i)) +ₖ (z ⌣ₖ x) + +-- r : (z : EM H' 1) (x : EM G' (suc n)) → (distrl1 (suc n) 1 z embase .fst x) ≡ (distrl2 (suc n) 1 z embase .fst x) +-- r y z = cong (z ⌣ₖ_) (rUnitₖ _ y) +-- ∙∙ sym (rUnitₖ _ (z ⌣ₖ y)) +-- ∙∙ λ i → (z ⌣ₖ y) +ₖ (⌣ₖ-0ₖ (suc n) (suc zero) z (~ i)) + +-- l≡r : (z : EM G' (suc n)) → l embase z ≡ r embase z +-- l≡r z = sym (lem _ _ (sym (⌣ₖ-0ₖ (suc n) (suc zero) z))) + +-- mainDistrL n (suc m) = +-- elim2 (λ _ _ → isOfHLevelPath (4 + m) (hLevLem _ _) _ _) +-- (wedgeConEM.fun H' H' (suc m) (suc m) +-- (λ x y p q → isOfHLevelPlus {n = suc (suc m)} (suc m) +-- (hLevLem n (suc (suc m)) +-- (distrl1 (suc n) (suc (suc m)) ∣ x ∣ ∣ y ∣) +-- (distrl2 (suc n) (suc (suc m)) ∣ x ∣ ∣ y ∣) p q)) +-- (λ x → →∙Homogeneous≡ (isHomogeneousEM _) +-- (funExt (l x))) +-- (λ x → →∙Homogeneous≡ (isHomogeneousEM _) +-- (funExt (r x))) +-- λ i → →∙Homogeneous≡ (isHomogeneousEM _) +-- (funExt (λ z → l≡r z i))) +-- where +-- l : (x : EM-raw H' (suc (suc m))) (z : EM G' (suc n)) +-- → (distrl1 (suc n) (suc (suc m)) (0ₖ _) ∣ x ∣ₕ .fst z) ≡ (distrl2 (suc n) (suc (suc m)) (0ₖ _) ∣ x ∣ₕ .fst z) +-- l x z = cong (z ⌣ₖ_) (lUnitₖ (suc (suc m)) ∣ x ∣) +-- ∙∙ sym (lUnitₖ _ (z ⌣ₖ ∣ x ∣)) +-- ∙∙ λ i → (⌣ₖ-0ₖ (suc n) (suc (suc m)) z (~ i)) +ₖ (z ⌣ₖ ∣ x ∣) + +-- r : (x : EM-raw H' (suc (suc m))) (z : EM G' (suc n)) +-- → (distrl1 (suc n) (suc (suc m)) ∣ x ∣ₕ (0ₖ _) .fst z) ≡ (distrl2 (suc n) (suc (suc m)) ∣ x ∣ₕ (0ₖ _) .fst z) +-- r x z = cong (z ⌣ₖ_) (rUnitₖ (suc (suc m)) ∣ x ∣) +-- ∙∙ sym (rUnitₖ _ (z ⌣ₖ ∣ x ∣)) +-- ∙∙ λ i → (z ⌣ₖ ∣ x ∣) +ₖ (⌣ₖ-0ₖ (suc n) (suc (suc m)) z (~ i)) + +-- l≡r : (z : EM G' (suc n)) → l north z ≡ r north z +-- l≡r z = sym (lem _ _ (sym (⌣ₖ-0ₖ (suc n) (suc (suc m)) z))) + + +-- commer : {!!} +-- commer = {!!} + +-- module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where +-- private +-- G = fst G' +-- H = fst H' + +-- strG = snd G' +-- strH = snd H' + +-- 0G = 0g strG +-- 0H = 0g strH + +-- _+G_ = _+Gr_ strG +-- _+H_ = _+Gr_ strH + +-- -H_ = -Gr_ strH +-- -G_ = -Gr_ strG + + +-- distrr1 : (n m : ℕ) → EM G' n → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m) +-- fst (distrr1 n m x y) z = (x +ₖ y) ⌣ₖ z +-- snd (distrr1 n m x y) = ⌣ₖ-0ₖ n m _ + +-- distrr2 : (n m : ℕ) → EM G' n → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m) +-- fst (distrr2 n m x y) z = (x ⌣ₖ z) +ₖ (y ⌣ₖ z) +-- snd (distrr2 n m x y) = cong₂ _+ₖ_ (⌣ₖ-0ₖ n m x) (⌣ₖ-0ₖ n m y) ∙ rUnitₖ _ (0ₖ (n +' m)) + +-- hLevLem2 : (n m : ℕ) → isOfHLevel (suc (suc m)) (EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') ((suc n) +' m)) +-- hLevLem2 n m = subst (isOfHLevel (suc (suc m))) (λ i → EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') +-- ((cong suc (+-comm m n) ∙ sym (+'≡+ (suc n) m)) i)) (isOfHLevel↑∙ m n) + +-- mainDistrR : (n m : ℕ) (x y : EM G' (suc n)) → distrr1 (suc n) (suc m) x y ≡ distrr2 (suc n) (suc m) x y +-- mainDistrR zero m = +-- wedgeConEM.fun G' G' 0 0 +-- (λ _ _ → isOfHLevel↑∙ 1 m _ _) +-- (λ x → →∙Homogeneous≡ (isHomogeneousEM _) +-- (funExt (l x))) +-- (λ x → →∙Homogeneous≡ (isHomogeneousEM _) +-- (funExt (r x))) +-- λ i → →∙Homogeneous≡ (isHomogeneousEM _) +-- (funExt λ z → l≡r z i) +-- where +-- l : (x : _) (z : _) → _ ≡ _ +-- l x z = +-- (λ i → (lUnitₖ 1 x i) ⌣ₖ z) +-- ∙∙ sym (lUnitₖ _ (x ⌣ₖ z)) +-- ∙∙ λ i → 0ₖ-⌣ₖ _ _ z (~ i) +ₖ (x ⌣ₖ z) + +-- r : (x : _) (z : _) → _ ≡ _ +-- r x z = +-- ((λ i → (rUnitₖ 1 x i) ⌣ₖ z)) +-- ∙∙ sym (rUnitₖ _ _) +-- ∙∙ λ i → (_⌣ₖ_ {n = 1} {m = suc m} x z) +ₖ 0ₖ-⌣ₖ (suc zero) (suc m) z (~ i) + +-- l≡r : (z : _) → l embase z ≡ r embase z +-- l≡r z = lem _ _ _ + +-- mainDistrR (suc n) m = +-- elim2 (λ _ _ → isOfHLevelPath (4 + n) +-- (isOfHLevel↑∙ (2 + n) m) _ _) +-- (wedgeConEM.fun _ _ _ _ +-- (λ x y → isOfHLevelPath ((2 + n) + (2 + n)) +-- {!subst !} _ _) +-- (λ x → →∙Homogeneous≡ (isHomogeneousEM _) +-- (funExt (l x))) +-- (λ x → →∙Homogeneous≡ (isHomogeneousEM _) +-- (funExt (r x))) +-- λ i → →∙Homogeneous≡ (isHomogeneousEM _) +-- (funExt λ z → r≡l z i)) +-- where + +-- l : (x : _) (z : _) → _ ≡ _ +-- l x z = (λ i → (lUnitₖ _ ∣ x ∣ i) ⌣ₖ z) +-- ∙∙ sym (lUnitₖ _ (∣ x ∣ ⌣ₖ z)) +-- ∙∙ λ i → 0ₖ-⌣ₖ _ _ z (~ i) +ₖ (∣ x ∣ ⌣ₖ z) + +-- r : (x : _) (z : _) → _ ≡ _ +-- r x z = (λ i → (rUnitₖ _ ∣ x ∣ i) ⌣ₖ z) +-- ∙∙ sym (rUnitₖ _ (∣ x ∣ ⌣ₖ z)) +-- ∙∙ λ i → (∣ x ∣ ⌣ₖ z) +ₖ 0ₖ-⌣ₖ _ _ z (~ i) + +-- r≡l : (z : _) → l north z ≡ r north z +-- r≡l z = lem _ _ _ diff --git a/Cubical/HITs/SmashProduct/Base.agda b/Cubical/HITs/SmashProduct/Base.agda index 59cd41698..86baa52d9 100644 --- a/Cubical/HITs/SmashProduct/Base.agda +++ b/Cubical/HITs/SmashProduct/Base.agda @@ -53,6 +53,38 @@ SmashPt A B = (Smash A B , basel) SmashPtProj : (A : Pointed ℓ) (B : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') SmashPtProj A B = Smash A B , (proj (snd A) (snd B)) +private + rCancelPathP : ∀ {ℓ} {A B : Type ℓ} {x y : A} (p : x ≡ y) (f : A → B) + → PathP (λ i → cong-∙ f p (sym p) i ≡ refl) + (cong (cong f) (rCancel p)) + (rCancel (cong f p)) + rCancelPathP {x = x} {y = y} p f i j k = + hcomp (λ r → λ { (i = i0) → f (rCancel-filler p r j k) + ; (j = i1) → f x + ; (k = i0) → f x + ; (k = i1) → f (p (~ r ∧ ~ j))}) + (f (p (k ∧ ~ j))) + + mainGen : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) + → (P : refl ≡ p) + → Cube (λ j k → (cong (p ∙_) (cong sym (sym P)) ∙ sym (rUnit _)) k j) + refl + refl + refl + (rCancel p) + (sym P) + mainGen {x = x} p = J (λ p P → Cube (λ j k → (cong (p ∙_) (cong sym (sym P)) ∙ sym (rUnit _)) k j) + refl + refl + refl + (rCancel p) + (sym P)) + (transport (λ i → Cube (λ j k → (lUnit (sym (rUnit (refl {x = x}))) i) k j) + l l l (rCancel (λ _ → x)) l) + λ i j k → rCancel (λ _ → x) (k ∨ i) j) + where + l = refl {x = refl {x = x}} + --- Alternative definition i∧ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → A ⋁ B → (typ A) × (typ B) @@ -99,7 +131,8 @@ Smash→⋀ (proj x y) = inr (x , y) Smash→⋀ (gluel a i) = push (inl a) (~ i) Smash→⋀ (gluer b i) = push (inr b) (~ i) -{- associativity maps for smash produts. Proof pretty much direcly translated from https://github.com/ecavallo/redtt/blob/master/library/pointed/smash.red -} +{- associativity maps for smash produts. Proof pretty much direcly translated from + https://github.com/ecavallo/redtt/blob/master/library/pointed/smash.red -} private pivotl : (b b' : typ B) → Path (Smash A B) (proj (snd A) b) (proj (snd A) b') @@ -211,3 +244,282 @@ SmashAssociate⁻ = rearrange ∘ comm ∘ Smash-map (idfun∙ _) (comm , refl) ⋀-associate⁻ : A ⋀ (B ⋀∙ C) → (A ⋀∙ B) ⋀ C ⋀-associate⁻ = (SmashPtProj→⋀∙ ⋀→ idfun∙ _) ∘ Smash→⋀ ∘ SmashAssociate⁻ ∘ ⋀→Smash ∘ (idfun∙ _ ⋀→ ⋀∙→SmashPtProj) + + +data _⋀'_ {ℓ ℓ' : Level} (A : Pointed ℓ) (B : Pointed ℓ') : Type (ℓ-max ℓ ℓ') where + incl : typ A → typ B → A ⋀' B + pathl : (a : typ A) → incl a (pt B) ≡ incl (pt A) (pt B) + pathr : (b : typ B) → incl (pt A) b ≡ incl (pt A) (pt B) + pathlCoh : pathl (pt A) ≡ refl + pathrCoh : pathr (pt B) ≡ refl + +module _ {ℓ'' : Level} (P P' : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ')) + (d : ∀ {ℓ ℓ' ℓ''} (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') → Iso (typ (P (P A B) C)) (typ (P A (P B C)))) + (d' : ∀ {ℓ ℓ' ℓ''} (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') → Iso (typ (P' (P' A B) C)) (typ (P' A (P' B C)))) + (∧-map : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → typ (P A B) → typ (P' A B)) + (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') + (is1 : Iso (typ (P A ((Pushout {A = typ (P B C)} ∧-map λ _ → tt) , inl (pt (P' B C))))) + (typ (P ((Pushout {A = typ (P A B)} ∧-map λ _ → tt) , inl (pt (P' A B))) C))) + (is2 : Iso (typ (P' A ((Pushout {A = typ (P B C)} ∧-map λ _ → tt) , inl (pt (P' B C))))) + (typ (P' ((Pushout {A = typ (P A B)} ∧-map λ _ → tt) , inl (pt (P' A B))) C))) + + (coherence1 : (a : _) → Iso.fun is2 (∧-map a) ≡ ∧-map (Iso.fun is1 a)) + (coherence2 : (a : _) → Iso.inv is2 (∧-map a) ≡ ∧-map (Iso.inv is1 a)) + where + + _⊚_ : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') → Type _ + A ⊚ B = Pushout {A = typ (P A B)} ∧-map λ _ → tt + + typL = A ⊚ ((B ⊚ C) , inl (pt (P' B C))) + typR = ((A ⊚ B) , inl (pt (P' A B))) ⊚ C + + f : typL → typR + f (inl x) = inl (Iso.fun is2 x) -- inl (Iso.fun is1 x ) + f (inr x) = inr x + f (push a i) = s i + where + s : inl (Iso.fun is2 (∧-map a)) ≡ inr tt + s = cong inl (coherence1 a) ∙ push (Iso.fun is1 a) + + f⁻ : typR → typL + f⁻ (inl x) = inl (Iso.inv is2 x) + f⁻ (inr x) = inr x + f⁻ (push a i) = s i + where + s : inl (Iso.inv is2 (∧-map a)) ≡ inr tt + s = cong inl (coherence2 a) + ∙ push (Iso.inv is1 a) + + module _ (pp1 : (a : _) → PathP (λ i → inl (Iso.rightInv is2 (∧-map a) i) ≡ inr tt) + (cong f (cong inl (coherence2 a)) ∙ cong f (push (Iso.inv is1 a))) (push a)) + + (pp2 : (a : _) → PathP (λ i → inl (Iso.leftInv is2 (∧-map a) i) ≡ inr tt) + (cong f⁻ (cong inl (coherence1 a)) ∙ cong f⁻ (push (Iso.fun is1 a))) + (push a)) + where + f→f⁻→f : (a : _) → f (f⁻ a) ≡ a + f→f⁻→f (inl x) = cong inl (Iso.rightInv is2 x) + f→f⁻→f (inr x) = refl + f→f⁻→f (push a i) j = + hcomp (λ k → λ { (i = i0) → inl (Iso.rightInv is2 (∧-map a) j) + ; (i = i1) → inr tt + ; (j = i0) → cong-∙ f (cong inl (coherence2 a)) (push (Iso.inv is1 a)) (~ k) i + ; (j = i1) → push a i}) + (pp1 a j i) + + f⁻→f→f⁻ : (a : _) → f⁻ (f a) ≡ a + f⁻→f→f⁻ (inl x) = cong inl (Iso.leftInv is2 x) + f⁻→f→f⁻ (inr x) = refl + f⁻→f→f⁻ (push a i) j = + hcomp (λ k → λ { (i = i0) → inl (Iso.leftInv is2 (∧-map a) j) + ; (i = i1) → inr tt + ; (j = i0) → cong-∙ f⁻ (cong inl (coherence1 a)) (push (Iso.fun is1 a)) (~ k) i + ; (j = i1) → push a i}) + (pp2 a j i) + + mainIso : Iso typL typR + Iso.fun mainIso = f + Iso.inv mainIso = f⁻ + Iso.rightInv mainIso = f→f⁻→f + Iso.leftInv mainIso = f⁻→f→f⁻ + + +wedge : ∀ {ℓ ℓ'} → {A : Pointed ℓ} → {B : Pointed ℓ'} → A ⋁ B → fst A × fst B +wedge {B = B} (inl x) = x , pt B +wedge {A = A} (inr x) = (pt A) , x +wedge {A = A} {B = B} (push a i) = (pt A) , (pt B) + +_Sm_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') +A Sm B = Pushout {A = A ⋁ B} wedge λ _ → tt + +SmashIso : ∀ {ℓ ℓ' ℓ''} (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') + → Iso (A Sm (B Sm C , inl (pt B , pt C))) ((A Sm B , inl (pt A , pt B)) Sm C) +SmashIso A B C = + mainIso (λ A B → A ⋁ B , inl (pt A)) + (λ A B → (typ A × typ B) , ((pt A) , (pt B))) + ∨-assoc + (λ _ _ _ → invIso ×Assoc) + wedge + A + B + C + is1 + {!!} + {!!} + {!!} + {!!} + {!!} + where + is1 : Iso _ _ + Iso.fun is1 (inl x) = {!!} + Iso.fun is1 (inr x) = {!!} + Iso.fun is1 (push a i) = {!!} + Iso.inv is1 = {!!} + Iso.rightInv is1 = {!!} + Iso.leftInv is1 = {!!} + ×Assoc : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso (A × (B × C)) ((A × B) × C) + Iso.fun ×Assoc (x , (y , z)) = (x , y) , z + Iso.inv ×Assoc ((x , y) , z) = x , (y , z) + Iso.rightInv ×Assoc ((x , x₂) , x₁) = refl + Iso.leftInv ×Assoc (x , (y , z)) = refl + + FF : ∀ {ℓ₁ ℓ''' ℓ'''' : Level} + {A : Pointed ℓ₁} {B : Pointed ℓ'''} {C : Pointed ℓ''''} + → ((A ⋁ B) , inl (pt A)) ⋁ C → + A ⋁ ((B ⋁ C) , inl (pt B)) + FF (inl (inl x)) = inl x + FF (inl (inr x)) = inr (inl x) + FF (inl (push a i)) = push a i + FF (inr x) = inr (inr x) + FF (push a i) = (push a ∙ cong inr (push tt)) i + + FF⁻ : ∀ {ℓ₁ ℓ''' ℓ'''' : Level} + {A : Pointed ℓ₁} {B : Pointed ℓ'''} {C : Pointed ℓ''''} + → A ⋁ ((B ⋁ C) , inl (pt B)) + → ((A ⋁ B) , inl (pt A)) ⋁ C + FF⁻ (inl x) = inl (inl x) + FF⁻ (inr (inl x)) = inl (inr x) + FF⁻ (inr (inr x)) = inr x + FF⁻ (inr (push a i)) = (cong inl (sym (push a)) ∙ push a) i + FF⁻ (push a i) = inl (push a i) + + FF→FF⁻→FF : ∀ {ℓ₁ ℓ''' ℓ'''' : Level} + {A : Pointed ℓ₁} {B : Pointed ℓ'''} {C : Pointed ℓ''''} + (x : A ⋁ ((B ⋁ C) , inl (pt B))) → FF (FF⁻ x) ≡ x + FF→FF⁻→FF (inl x) = refl + FF→FF⁻→FF (inr (inl x)) = refl + FF→FF⁻→FF (inr (inr x)) = refl + FF→FF⁻→FF (inr (push a i)) k = h k i + where + h : cong FF ((cong inl (sym (push a)) ∙ push a)) ≡ cong inr (push a) + h = cong-∙ FF (cong inl (sym (push a))) (push a) + ∙∙ assoc _ _ _ + ∙∙ cong (_∙ cong inr (push a)) (lCancel _) + ∙ sym (lUnit _) + FF→FF⁻→FF (push a i) = refl + + FF⁻→FF→FF⁻ : ∀ {ℓ₁ ℓ''' ℓ'''' : Level} + {A : Pointed ℓ₁} {B : Pointed ℓ'''} {C : Pointed ℓ''''} + (x : ((A ⋁ B) , inl (pt A)) ⋁ C) → FF⁻ (FF x) ≡ x + FF⁻→FF→FF⁻ (inl (inl x)) = refl + FF⁻→FF→FF⁻ (inl (inr x)) = refl + FF⁻→FF→FF⁻ (inl (push a i)) = refl + FF⁻→FF→FF⁻ (inr x) = refl + FF⁻→FF→FF⁻ (push a i) k = h k i + where + h : cong FF⁻ (push a ∙ cong inr (push tt)) ≡ push a + h = cong-∙ FF⁻ (push a) (cong inr (push a)) + ∙∙ assoc _ _ _ + ∙∙ cong (_∙ push a) (rCancel _) + ∙ sym (lUnit _) + + ∨-assoc : ∀ {ℓ₁ ℓ''' ℓ'''' : Level} + (A : Pointed ℓ₁) (B : Pointed ℓ''') (C : Pointed ℓ'''') → + Iso (((A ⋁ B) , inl (pt A)) ⋁ C) + (A ⋁ ((B ⋁ C) , inl (pt B))) + Iso.fun (∨-assoc A B C) = FF + Iso.inv (∨-assoc A B C) = FF⁻ + Iso.rightInv (∨-assoc A B C) = FF→FF⁻→FF + Iso.leftInv (∨-assoc A B C) = FF⁻→FF→FF⁻ + + {- + i = i0 ⊢ cong f (cong inl (coherence2 a) ∙ push (Iso.inv is1 a)) j +i = i1 ⊢ push a j +j = i0 ⊢ inl (Iso.rightInv is2 (∧-map a) i) +j = i1 ⊢ inr tt + -} + + +-- s : Iso (A ⊚ ((B ⊚ C) , inl tt)) (((A ⊚ B) , inl tt) ⊚ C) +-- Iso.fun s = {!!} +-- Iso.inv s = {!!} +-- Iso.rightInv s = {!!} +-- Iso.leftInv s = {!!} + + +-- -- module _ {ℓ'' : Level} (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') where +-- -- assocFun : ((A ⋀' B) , (incl (pt A) (pt B))) ⋀' C +-- -- → A ⋀' (B ⋀' C , incl (pt B) (pt C)) +-- -- assocFun (incl (incl x x₁) y) = incl x (incl x₁ y) +-- -- assocFun (incl (pathl a i) y) = (cong (incl a) (pathr y) ∙∙ pathl a ∙∙ sym (pathr (incl (pt B) y))) i +-- -- assocFun (incl (pathr b i) y) = (pathr (incl b y) ∙ cong (incl (pt A)) (sym (pathr y))) i +-- -- assocFun (incl (pathlCoh i j) y) = {!!} +-- -- assocFun (incl (pathrCoh i j) y) = {!!} +-- -- assocFun (pathl (incl x x₁) i) = {!!} +-- -- assocFun (pathl (pathl a i₁) i) = {!!} +-- -- assocFun (pathl (pathr b i₁) i) = {!!} +-- -- assocFun (pathl (pathlCoh i₁ i₂) i) = {!!} +-- -- assocFun (pathl (pathrCoh i₁ i₂) i) = {!!} +-- -- assocFun (pathr b i) = +-- -- {!Goal: A ⋀' ((B ⋀' C) , incl (pt B) (pt C)) +-- -- ———— Boundary —————————————————————————————————————————————— +-- -- i = i0 ⊢ incl (pt A) (incl b y) +-- -- i = i1 ⊢ incl (pt A) (incl (pt B) y)!} +-- -- assocFun (pathlCoh i i₁) = {!!} +-- -- assocFun (pathrCoh i i₁) = {!!} + +-- -- module _ (A : Pointed ℓ) (B : Pointed ℓ') where +-- -- smashFun : (A ⋀' B) → (Smash A B) +-- -- smashFun (incl x y) = proj x y +-- -- smashFun (pathl a i) = pivotr a (pt A) i +-- -- smashFun (pathr b i) = pivotl b (pt B) i +-- -- smashFun (pathlCoh i j) = rCancel (gluel (pt A)) i j +-- -- smashFun (pathrCoh i j) = rCancel (gluer (pt B)) i j + +-- -- smashFun⁻ : (Smash A B) → A ⋀' B +-- -- smashFun⁻ basel = incl (pt A) (pt B) +-- -- smashFun⁻ baser = incl (pt A) (pt B) +-- -- smashFun⁻ (proj x y) = incl x y +-- -- smashFun⁻ (gluel a i) = pathl a i +-- -- smashFun⁻ (gluer b i) = pathr b i + +-- -- sectSmashFun : section smashFun smashFun⁻ +-- -- sectSmashFun basel = gluel (pt A) +-- -- sectSmashFun baser = gluer (pt B) +-- -- sectSmashFun (proj x y) = refl +-- -- sectSmashFun (gluel a i) j = compPath-filler (gluel a) (sym (gluel (pt A))) (~ j) i +-- -- sectSmashFun (gluer b i) j = compPath-filler (gluer b) (sym (gluer (pt B))) (~ j) i + +-- -- retrSmashFun : retract smashFun smashFun⁻ +-- -- retrSmashFun (incl x x₁) = refl +-- -- retrSmashFun (pathl a i) k = help k i +-- -- module _ where +-- -- help : cong smashFun⁻ (cong smashFun (pathl a)) ≡ pathl a +-- -- help = congFunct smashFun⁻ (gluel a) (sym (gluel (pt A))) ∙ cong (cong smashFun⁻ (gluel a) ∙_) (cong sym pathlCoh) ∙ sym (rUnit _) +-- -- retrSmashFun (pathr b i) k = help2 k i +-- -- module _ where +-- -- help2 : cong smashFun⁻ (cong smashFun (pathr b)) ≡ pathr b +-- -- help2 = congFunct smashFun⁻ (gluer b) (sym (gluer (pt B))) +-- -- ∙ cong (cong smashFun⁻ (gluer b) ∙_) (cong sym pathrCoh) +-- -- ∙ sym (rUnit _) +-- -- retrSmashFun (pathlCoh i j) k = main i j k +-- -- where + +-- -- h = help (pt A) i0 i0 +-- -- main : PathP (λ i → PathP (λ j → smashFun⁻ (smashFun (pathlCoh i j)) ≡ pathlCoh i j) +-- -- refl refl) +-- -- (λ j k → h k j) refl +-- -- main i j k = +-- -- hcomp (λ r → λ { (i = i0) → compPath-filler' (congFunct smashFun⁻ (gluel (pt A)) (sym (gluel (pt A)))) +-- -- (cong (cong smashFun⁻ (gluel (pt A)) ∙_) (cong sym pathlCoh) ∙ sym (rUnit _)) r k j +-- -- ; (i = i1) → incl (pt A) (pt B) +-- -- ; (j = i0) → incl (pt A) (pt B) +-- -- ; (j = i1) → incl (pt A) (pt B) +-- -- ; (k = i0) → rCancelPathP (gluel (pt A)) smashFun⁻ (~ r) i j +-- -- ; (k = i1) → pathlCoh i j}) +-- -- (mainGen (pathl (pt A)) (sym (pathlCoh)) i j k) +-- -- retrSmashFun (pathrCoh i j) k = +-- -- hcomp (λ r → λ {(i = i0) → compPath-filler' (congFunct smashFun⁻ (gluer (pt B)) (sym (gluer (pt B)))) +-- -- (cong (cong smashFun⁻ (gluer (pt B)) ∙_) (cong sym pathrCoh) ∙ sym (rUnit _)) r k j +-- -- ; (i = i1) → incl (pt A) (pt B) +-- -- ; (j = i0) → incl (pt A) (pt B) +-- -- ; (j = i1) → incl (pt A) (pt B) +-- -- ; (k = i0) → rCancelPathP (gluer (pt B)) smashFun⁻ (~ r) i j +-- -- ; (k = i1) → pathrCoh i j}) +-- -- (mainGen (pathr (pt B)) (sym (pathrCoh)) i j k) + +-- -- smashIso : Iso (A ⋀' B) (Smash A B) +-- -- Iso.fun smashIso = smashFun +-- -- Iso.inv smashIso = smashFun⁻ +-- -- Iso.rightInv smashIso = sectSmashFun +-- -- Iso.leftInv smashIso = retrSmashFun From 347f349c48fe8ff9d2646011028e823b5eba163d Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 19 Aug 2021 17:53:25 +0200 Subject: [PATCH 21/30] trim --- Cubical/Algebra/AbGroup/Base.agda | 69 ++ Cubical/Algebra/AbGroup/TensorProduct.agda | 332 +++---- .../Group/EilenbergMacLane/CupProduct.agda | 894 +++++++++--------- .../EilenbergMacLane/GroupStructure.agda | 155 +-- .../Group/EilenbergMacLane/Properties.agda | 85 +- Cubical/Foundations/Prelude.agda | 10 + Cubical/HITs/SmashProduct/Base.agda | 279 ------ 7 files changed, 777 insertions(+), 1047 deletions(-) diff --git a/Cubical/Algebra/AbGroup/Base.agda b/Cubical/Algebra/AbGroup/Base.agda index 84dc2c07b..70f9b7d69 100644 --- a/Cubical/Algebra/AbGroup/Base.agda +++ b/Cubical/Algebra/AbGroup/Base.agda @@ -190,3 +190,72 @@ assoc (isSemigroup (isMonoid (isGroup (isAbGroup (snd trivialAbGroup))))) _ _ _ identity (isMonoid (isGroup (isAbGroup (snd trivialAbGroup)))) _ = refl , refl inverse (isGroup (isAbGroup (snd trivialAbGroup))) _ = refl , refl comm (isAbGroup (snd trivialAbGroup)) _ _ = refl + +---- The type of homomorphisms A → B is an AbGroup if B is ----- +module _ {ℓ ℓ' : Level} (AGr : Group ℓ) (BGr : AbGroup ℓ') where + private + strA = snd AGr + strB = snd BGr + + _* = AbGroup→Group + + A = fst AGr + B = fst BGr + open IsGroupHom + + open AbGroupStr strB + renaming (_+_ to _+B_ ; -_ to -B_ ; 0g to 0B + ; rid to ridB ; lid to lidB + ; assoc to assocB ; comm to commB + ; invr to invrB ; invl to invlB) + open GroupStr strA + renaming (_·_ to _∙A_ ; inv to -A_ + ; 1g to 1A ; rid to ridA) + + trivGroupHom : GroupHom AGr (BGr *) + fst trivGroupHom x = 0B + snd trivGroupHom = makeIsGroupHom λ _ _ → sym (ridB 0B) + + compHom : GroupHom AGr (BGr *) → GroupHom AGr (BGr *) → GroupHom AGr (BGr *) + fst (compHom f g) x = fst f x +B fst g x + snd (compHom f g) = + makeIsGroupHom λ x y + → cong₂ _+B_ (pres· (snd f) x y) (pres· (snd g) x y) + ∙ move4 (fst f x) (fst f y) (fst g x) (fst g y) + _+B_ assocB commB + + invHom : GroupHom AGr (BGr *) → GroupHom AGr (BGr *) + fst (invHom (f , p)) x = -B f x + snd (invHom (f , p)) = + makeIsGroupHom + λ x y → cong -B_ (pres· p x y) + ∙∙ GroupTheory.invDistr (BGr *) (f x) (f y) + ∙∙ commB _ _ + + open AbGroupStr + open IsAbGroup + open IsGroup + open IsMonoid + open IsSemigroup + + HomGroup : AbGroup (ℓ-max ℓ ℓ') + fst HomGroup = GroupHom AGr (BGr *) + 0g (snd HomGroup) = trivGroupHom + AbGroupStr._+_ (snd HomGroup) = compHom + AbGroupStr.- snd HomGroup = invHom + is-set (isSemigroup (isMonoid (isGroup (isAbGroup (snd HomGroup))))) = + isSetGroupHom + assoc (isSemigroup (isMonoid (isGroup (isAbGroup (snd HomGroup))))) (f , p) (g , q) (h , r) = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ x → assocB _ _ _) + fst (identity (isMonoid (isGroup (isAbGroup (snd HomGroup)))) (f , p)) = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ y → ridB _) + snd (identity (isMonoid (isGroup (isAbGroup (snd HomGroup)))) (f , p)) = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ x → lidB _) + fst (inverse (isGroup (isAbGroup (snd HomGroup))) (f , p)) = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ x → invrB (f x)) + snd (inverse (isGroup (isAbGroup (snd HomGroup))) (f , p)) = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ x → invlB (f x)) + comm (isAbGroup (snd HomGroup)) (f , p) (g , q) = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ x → commB _ _) diff --git a/Cubical/Algebra/AbGroup/TensorProduct.agda b/Cubical/Algebra/AbGroup/TensorProduct.agda index 949167fea..0e5417616 100644 --- a/Cubical/Algebra/AbGroup/TensorProduct.agda +++ b/Cubical/Algebra/AbGroup/TensorProduct.agda @@ -18,11 +18,14 @@ open import Cubical.Foundations.HLevels open import Cubical.Data.Int open import Cubical.Data.Sum hiding (map) -open import Cubical.HITs.PropositionalTruncation renaming (map to pMap ; rec to pRec ; elim to pElim ; elim2 to pElim2) +open import Cubical.HITs.PropositionalTruncation + renaming (map to pMap ; rec to pRec ; elim to pElim ; elim2 to pElim2) open import Cubical.Algebra.Group hiding (ℤ) - --- open import Cubical.HITs.SetQuotients +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.Semigroup private variable @@ -63,16 +66,6 @@ module _ (AGr : AbGroup ℓ) (BGr : AbGroup ℓ') where ⊗squash : isSet _⨂₁_ - -move4 : ∀ {ℓ} {A : Type ℓ} (x y z w : A) (_+_ : A → A → A) - → ((x y z : A) → x + (y + z) ≡ (x + y) + z) - → ((x y : A) → x + y ≡ y + x) - → (x + y) + (z + w) ≡ ((x + z) + (y + w)) -move4 x y z w _+_ assoc comm = - sym (assoc x y (z + w)) - ∙∙ cong (x +_) (assoc y z w ∙∙ cong (_+ w) (comm y z) ∙∙ sym (assoc z y w)) - ∙∙ assoc x z (y + w) - module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where private open AbGroupStr renaming (_+_ to _+G_ ; -_ to -G_) @@ -94,9 +87,6 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where A⊗B = AGr ⨂₁ BGr - 0⊗ : AGr ⨂₁ BGr - 0⊗ = 0A ⊗ 0B - ⊗elimProp : ∀ {ℓ} {C : AGr ⨂₁ BGr → Type ℓ} → ((x : _) → isProp (C x)) → (f : (x : A) (b : B) → C (x ⊗ b)) @@ -130,7 +120,12 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where (f x (-B b)) (f (-A x) b) (flip x b) i ⊗elimProp {C = C} p f g (⊗squash x y q r i j) = isOfHLevel→isOfHLevelDep 2 {B = C} (λ x → isProp→isSet (p x)) - _ _ (λ j → ⊗elimProp p f g (q j)) (λ j → ⊗elimProp p f g (r j)) (⊗squash x y q r) i j + _ _ (λ j → ⊗elimProp p f g (q j)) (λ j → ⊗elimProp p f g (r j)) + (⊗squash x y q r) i j + + -- Group structure + 0⊗ : AGr ⨂₁ BGr + 0⊗ = 0A ⊗ 0B -⊗ : AGr ⨂₁ BGr → AGr ⨂₁ BGr -⊗ (a ⊗ b) = (-A a) ⊗ b @@ -152,22 +147,26 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where ⊗rUnit : (x : A⊗B) → x +⊗ 0⊗ ≡ x ⊗rUnit x = ⊗comm x 0⊗ ∙ ⊗lUnit x + ------------------------------------------------------------------------------- + -- Useful induction principle, which lets us view elements of A ⨂ B as lists + -- over (A × B). Used for the right cancellation law listify : List (A × B) → AGr ⨂₁ BGr listify [] = 0A ⊗ 0B listify (x ∷ x₁) = (fst x ⊗ snd x) +⊗ listify x₁ - listify++ : (x y : List (A × B)) → listify (x ++ y) ≡ (listify x +⊗ listify y) + listify++ : (x y : List (A × B)) + → listify (x ++ y) ≡ (listify x +⊗ listify y) listify++ [] y = sym (⊗lUnit (listify y)) listify++ (x ∷ x₁) y = (λ i → (fst x ⊗ snd x) +⊗ (listify++ x₁ y i)) ∙ ⊗assoc (fst x ⊗ snd x) (listify x₁) (listify y) - slick : (x : AGr ⨂₁ BGr) → ∃[ l ∈ List (A × B) ] listify l ≡ x + slick : (x : AGr ⨂₁ BGr) → ∃[ l ∈ List (A × B) ] (listify l ≡ x) slick = ⊗elimProp (λ _ → squash) (λ a b → ∣ [ a , b ] , ⊗rUnit (a ⊗ b) ∣) - λ x y → rec2 squash λ {(l1 , p) (l2 , q) → ∣ (l1 ++ l2) , listify++ l1 l2 ∙ cong₂ _+⊗_ p q ∣} - + λ x y → rec2 squash λ {(l1 , p) (l2 , q) + → ∣ (l1 ++ l2) , listify++ l1 l2 ∙ cong₂ _+⊗_ p q ∣} ⊗elimPropCool : ∀ {ℓ} {C : AGr ⨂₁ BGr → Type ℓ} → ((x : _) → isProp (C x)) @@ -177,9 +176,11 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where ⊗elimProp p (λ x y → subst C (⊗rUnit (x ⊗ y)) (f [ x , y ])) λ x y → pRec (isPropΠ2 λ _ _ → p _) (pRec (isPropΠ3 λ _ _ _ → p _) - (λ {(l1 , p) (l2 , q) ind1 ind2 → subst C (listify++ l2 l1 ∙ cong₂ _+⊗_ q p) (f (l2 ++ l1))}) + (λ {(l1 , p) (l2 , q) ind1 ind2 + → subst C (listify++ l2 l1 ∙ cong₂ _+⊗_ q p) (f (l2 ++ l1))}) (slick y)) (slick x) + ----------------------------------------------------------------------------------- lCancelPrim : (x : B) → (0A ⊗ x) ≡ 0⊗ lCancelPrim x = @@ -194,7 +195,9 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where rCancelPrim x = (λ i → x ⊗ rid strB 0B (~ i)) ∙∙ linr x 0B 0B - ∙∙ cong ((x ⊗ 0B) +⊗_) (cong (x ⊗_) (sym (GroupTheory.inv1g (AbGroup→Group BGr))) ∙ flip x 0B) + ∙∙ cong ((x ⊗ 0B) +⊗_) + (cong (x ⊗_) + (sym (GroupTheory.inv1g (AbGroup→Group BGr))) ∙ flip x 0B) ∙∙ sym (linl x (-A x) 0B) ∙∙ (λ i → (invr strA x i) ⊗ 0B) @@ -218,13 +221,12 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where ⊗lCancel x = ⊗comm _ _ ∙ ⊗rCancel x module _ where - open import Cubical.Algebra.Monoid - open import Cubical.Algebra.Semigroup open AbGroupStr open IsAbGroup open IsGroup open IsMonoid open IsSemigroup + _⨂_ : AbGroup ℓ → AbGroup ℓ' → AbGroup (ℓ-max ℓ ℓ') fst (A ⨂ B) = A ⨂₁ B 0g (snd (A ⨂ B)) = 0⊗ @@ -238,74 +240,61 @@ module _ where snd (inverse (isGroup (isAbGroup (snd (A ⨂ B)))) x) = ⊗lCancel x comm (isAbGroup (snd (A ⨂ B))) = ⊗comm -open import Cubical.Algebra.Group.Morphisms -open import Cubical.Algebra.Group.MorphismProperties +-------------- Elimination principle into AbGroups -------------- +module _ {ℓ ℓ' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} where + private + open AbGroupStr renaming (_+_ to _+G_ ; -_ to -G) + _+A_ = _+G_ (snd A) + _+B_ = _+G_ (snd B) -_* : AbGroup ℓ → Group ℓ -_* = AbGroup→Group + 0A = 0g (snd A) + 0B = 0g (snd B) -module _ {ℓ ℓ' : Level} (AGr : Group ℓ) (BGr : AbGroup ℓ') where - private - strA = snd AGr - strB = snd BGr + ⨂→AbGroup-elim : ∀ {ℓ} (C : AbGroup ℓ) + → (f : (fst A × fst B) → fst C) + → (f (0A , 0B) ≡ 0g (snd C)) + → (linL : (x y : fst A) (b : fst B) + → f (x +A y , b) ≡ _+G_ (snd C) (f (x , b)) (f (y , b))) + → (linR : (a : fst A) (x y : fst B) + → f (a , x +B y) ≡ _+G_ (snd C) (f (a , x)) (f (a , y))) + → (fl : (x : fst A) (y : fst B) + → f (x , -G (snd B) y) ≡ f ((-G (snd A) x) , y)) + → A ⨂₁ B → fst C + ⨂→AbGroup-elim C f p linL linR fl (a ⊗ b) = f (a , b) + ⨂→AbGroup-elim C f p linL linR fl (x +⊗ x₁) = + _+G_ (snd C) (⨂→AbGroup-elim C f p linL linR fl x) + (⨂→AbGroup-elim C f p linL linR fl x₁) + ⨂→AbGroup-elim C f p linL linR fl (⊗comm x x₁ i) = + comm (snd C) (⨂→AbGroup-elim C f p linL linR fl x) + (⨂→AbGroup-elim C f p linL linR fl x₁) i + ⨂→AbGroup-elim C f p linL linR fl (⊗assoc x x₁ x₂ i) = + assoc (snd C) (⨂→AbGroup-elim C f p linL linR fl x) + (⨂→AbGroup-elim C f p linL linR fl x₁) + (⨂→AbGroup-elim C f p linL linR fl x₂) i + ⨂→AbGroup-elim C f p linL linR fl (⊗lUnit x i) = + (cong (λ y → (snd C +G y) (⨂→AbGroup-elim C f p linL linR fl x)) p + ∙ (lid (snd C) (⨂→AbGroup-elim C f p linL linR fl x))) i + ⨂→AbGroup-elim C f p linL linR fl (linl x y z i) = linL x y z i + ⨂→AbGroup-elim C f p linL linR fl (linr x y z i) = linR x y z i + ⨂→AbGroup-elim C f p linL linR fl (flip x b i) = fl x b i + ⨂→AbGroup-elim C f p linL linR fl (⊗squash x x₁ x₂ y i i₁) = + is-set (snd C) _ _ + (λ i → ⨂→AbGroup-elim C f p linL linR fl (x₂ i)) + (λ i → ⨂→AbGroup-elim C f p linL linR fl (y i)) i i₁ - A = fst AGr - B = fst BGr - open IsGroupHom - - open AbGroupStr strB renaming (_+_ to _+B_ ; -_ to -B_ ; 0g to 0B ; rid to ridB ; lid to lidB ; assoc to assocB ; comm to commB ; invr to invrB ; invl to invlB) - open GroupStr strA renaming (_·_ to _∙A_ ; inv to -A_ ; 1g to 1A ; rid to ridA) - - trivGroupHom : GroupHom AGr (BGr *) - fst trivGroupHom x = 0B - snd trivGroupHom = makeIsGroupHom λ _ _ → sym (ridB 0B) - - compHom : GroupHom AGr (BGr *) → GroupHom AGr (BGr *) → GroupHom AGr (BGr *) - fst (compHom f g) x = fst f x +B fst g x - snd (compHom f g) = - makeIsGroupHom λ x y - → cong₂ _+B_ (pres· (snd f) x y) (pres· (snd g) x y) - ∙ move4 (fst f x) (fst f y) (fst g x) (fst g y) - _+B_ assocB commB - - invHom : GroupHom AGr (BGr *) → GroupHom AGr (BGr *) - fst (invHom (f , p)) x = -B f x - snd (invHom (f , p)) = + ⨂→AbGroup-elim-hom : ∀ {ℓ} (C : AbGroup ℓ) + → (f : (fst A × fst B) → fst C) (linL : _) (linR : _) (fl : _) (p : _) + → AbGroupHom (A ⨂ B) C + fst (⨂→AbGroup-elim-hom C f linL linR fl p) = ⨂→AbGroup-elim C f p linL linR fl + snd (⨂→AbGroup-elim-hom C f linL linR fl p) = makeIsGroupHom - λ x y → cong -B_ (pres· p x y) - ∙∙ GroupTheory.invDistr (BGr *) (f x) (f y) - ∙∙ commB _ _ + (λ x y → refl) - open import Cubical.Algebra.Monoid - open import Cubical.Algebra.Semigroup - open AbGroupStr - open IsAbGroup - open IsGroup - open IsMonoid - open IsSemigroup +private + _* = AbGroup→Group - HomGroup : AbGroup (ℓ-max ℓ ℓ') - fst HomGroup = GroupHom AGr (BGr *) - 0g (snd HomGroup) = trivGroupHom - AbGroupStr._+_ (snd HomGroup) = compHom - AbGroupStr.- snd HomGroup = invHom - is-set (isSemigroup (isMonoid (isGroup (isAbGroup (snd HomGroup))))) = - isSetGroupHom - assoc (isSemigroup (isMonoid (isGroup (isAbGroup (snd HomGroup))))) (f , p) (g , q) (h , r) = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt λ x → assocB _ _ _) - fst (identity (isMonoid (isGroup (isAbGroup (snd HomGroup)))) (f , p)) = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ y → ridB _) - snd (identity (isMonoid (isGroup (isAbGroup (snd HomGroup)))) (f , p)) = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ x → lidB _) - fst (inverse (isGroup (isAbGroup (snd HomGroup))) (f , p)) = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ x → invrB (f x)) - snd (inverse (isGroup (isAbGroup (snd HomGroup))) (f , p)) = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ x → invlB (f x)) - comm (isAbGroup (snd HomGroup)) (f , p) (g , q) = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt λ x → commB _ _) +----------- Definition of universal property ------------ tensorFun : (A : Group ℓ) (B : Group ℓ') (T C : AbGroup (ℓ-max ℓ ℓ')) (f : GroupHom A (HomGroup B T *)) → GroupHom (T *) (C *) @@ -320,21 +309,7 @@ snd (tensorFun A B T C (f , p) (g , q)) = → Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ b → cong g (funExt⁻ (cong fst (IsGroupHom.pres· p x y)) b) - ∙ IsGroupHom.pres· q _ _) -{- - makeIsGroupHom λ x y - → cong g (IsGroupHom.pres· (snd (f a)) x y) - ∙ IsGroupHom.pres· q _ _ -snd (tensorFun A B T C (f , p) (g , q)) = - makeIsGroupHom λ x y - → cong g (IsGroupHom.pres· (snd (f a)) x y) - ∙ IsGroupHom.pres· q _ _ -snd (h (f , p) C (g , q)) = - makeIsGroupHom λ x y - → Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt λ b - → cong g (funExt⁻ (cong fst (IsGroupHom.pres· p x y)) b) - ∙ IsGroupHom.pres· q _ _) -} + ∙ IsGroupHom.pres· q _ _) isTensorProductOf_and_ : AbGroup ℓ → AbGroup ℓ' → AbGroup (ℓ-max ℓ ℓ')→ Type _ isTensorProductOf_and_ {ℓ} {ℓ'} A B T = @@ -344,6 +319,7 @@ isTensorProductOf_and_ {ℓ} {ℓ'} A B T = {B = GroupHom (A *) ((HomGroup (B *) C) *)} (tensorFun (A *) (B *) T C f)) +------ _⨂_ satisfies the universal property -------- module UP (AGr : AbGroup ℓ) (BGr : AbGroup ℓ') where private open AbGroupStr renaming (_+_ to _+G_ ; -_ to -G_) @@ -410,6 +386,7 @@ module UP (AGr : AbGroup ℓ) (BGr : AbGroup ℓ') where (λ _ _ → refl) λ x y ind1 ind2 → cong₂ (_+G_ (snd C)) ind1 ind2 ∙ sym (IsGroupHom.pres· p x y))) +-------------------- Commutativity ------------------------ commFun : ∀ {ℓ ℓ'} {A : AbGroup ℓ} {B : AbGroup ℓ'} → A ⨂₁ B → B ⨂₁ A commFun (a ⊗ b) = b ⊗ a commFun (x +⊗ x₁) = commFun x +⊗ commFun x₁ @@ -423,73 +400,30 @@ commFun (⊗squash x x₁ x₂ y i i₁) = ⊗squash (commFun x) (commFun x₁) (λ i → commFun (x₂ i)) (λ i → commFun (y i)) i i₁ -commFunx2 : ∀ {ℓ ℓ'} {A : AbGroup ℓ} {B : AbGroup ℓ'} (x : A ⨂₁ B) → commFun (commFun x) ≡ x -commFunx2 = +commFun²≡id : ∀ {ℓ ℓ'} {A : AbGroup ℓ} {B : AbGroup ℓ'} (x : A ⨂₁ B) + → commFun (commFun x) ≡ x +commFun²≡id = ⊗elimProp (λ _ → ⊗squash _ _) (λ _ _ → refl) λ _ _ p q → cong₂ _+⊗_ p q - -⨂-commIso : ∀ {ℓ ℓ'} {A : AbGroup ℓ} {B : AbGroup ℓ'} → GroupIso ((A ⨂ B) *) ((B ⨂ A) *) +⨂-commIso : ∀ {ℓ ℓ'} {A : AbGroup ℓ} {B : AbGroup ℓ'} + → GroupIso ((A ⨂ B) *) ((B ⨂ A) *) Iso.fun (fst ⨂-commIso) = commFun Iso.inv (fst ⨂-commIso) = commFun -Iso.rightInv (fst ⨂-commIso) = commFunx2 -Iso.leftInv (fst ⨂-commIso) = commFunx2 +Iso.rightInv (fst ⨂-commIso) = commFun²≡id +Iso.leftInv (fst ⨂-commIso) = commFun²≡id snd ⨂-commIso = makeIsGroupHom λ x y → refl ⨂-comm : ∀ {ℓ ℓ'} {A : AbGroup ℓ} {B : AbGroup ℓ'} → AbGroupEquiv (A ⨂ B) (B ⨂ A) fst ⨂-comm = isoToEquiv (fst (⨂-commIso)) snd ⨂-comm = snd ⨂-commIso -module _ {ℓ ℓ' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} where - open AbGroupStr renaming (_+_ to _+G_ ; -_ to -G) - _+A_ = _+G_ (snd A) - _+B_ = _+G_ (snd B) - - 0A = 0g (snd A) - 0B = 0g (snd B) - - - ⨂→AbGroup-elim : ∀ {ℓ} (C : AbGroup ℓ) - → (f : (fst A × fst B) → fst C) - → (f (0A , 0B) ≡ 0g (snd C)) - → (linL : (x y : fst A) (b : fst B) → f (x +A y , b) ≡ _+G_ (snd C) (f (x , b)) (f (y , b))) - → (linR : (a : fst A) (x y : fst B) → f (a , x +B y) ≡ _+G_ (snd C) (f (a , x)) (f (a , y))) - → (fl : (x : fst A) (y : fst B) → f (x , -G (snd B) y) ≡ f ((-G (snd A) x) , y)) - → A ⨂₁ B → fst C - ⨂→AbGroup-elim C f p linL linR fl (a ⊗ b) = f (a , b) - ⨂→AbGroup-elim C f p linL linR fl (x +⊗ x₁) = - _+G_ (snd C) (⨂→AbGroup-elim C f p linL linR fl x) (⨂→AbGroup-elim C f p linL linR fl x₁) - ⨂→AbGroup-elim C f p linL linR fl (⊗comm x x₁ i) = - comm (snd C) (⨂→AbGroup-elim C f p linL linR fl x) (⨂→AbGroup-elim C f p linL linR fl x₁) i - ⨂→AbGroup-elim C f p linL linR fl (⊗assoc x x₁ x₂ i) = - assoc (snd C) (⨂→AbGroup-elim C f p linL linR fl x) (⨂→AbGroup-elim C f p linL linR fl x₁) (⨂→AbGroup-elim C f p linL linR fl x₂) i - ⨂→AbGroup-elim C f p linL linR fl (⊗lUnit x i) = - (cong (λ y → (snd C +G y) (⨂→AbGroup-elim C f p linL linR fl x)) p - ∙ (lid (snd C) (⨂→AbGroup-elim C f p linL linR fl x))) i - ⨂→AbGroup-elim C f p linL linR fl (linl x y z i) = linL x y z i - ⨂→AbGroup-elim C f p linL linR fl (linr x y z i) = linR x y z i - ⨂→AbGroup-elim C f p linL linR fl (flip x b i) = fl x b i - ⨂→AbGroup-elim C f p linL linR fl (⊗squash x x₁ x₂ y i i₁) = - is-set (snd C) _ _ (λ i → ⨂→AbGroup-elim C f p linL linR fl (x₂ i)) (λ i → ⨂→AbGroup-elim C f p linL linR fl (y i)) i i₁ - - ⨂→AbGroup-elim-hom : ∀ {ℓ} (C : AbGroup ℓ) → (f : (fst A × fst B) → fst C) (linL : _) (linR : _) (fl : _) (p : _) - → AbGroupHom (A ⨂ B) C - fst (⨂→AbGroup-elim-hom C f linL linR fl p) = ⨂→AbGroup-elim C f p linL linR fl - snd (⨂→AbGroup-elim-hom C f linL linR fl p) = - makeIsGroupHom - (λ x y → refl) +open AbGroupStr +-------------------- Associativity ------------------------ module _ {ℓ ℓ' ℓ'' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} {C : AbGroup ℓ''} where private - open AbGroupStr renaming (_+_ to +G ; -_ to -G) - _+A'_ = +G (snd A) - _+B'_ = +G (snd B) - _+C'_ = +G (snd C) - -A = -G (snd A) - -B = -G (snd B) - -C = -G (snd C) - f : (c : fst C) → AbGroupHom (A ⨂ B) (A ⨂ (B ⨂ C)) f c = ⨂→AbGroup-elim-hom (A ⨂ (B ⨂ C)) ((λ ab → fst ab ⊗ (snd ab ⊗ c))) (λ x y b → linl x y (b ⊗ c)) @@ -500,58 +434,54 @@ module _ {ℓ ℓ' ℓ'' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} {C : AbGr assocHom : AbGroupHom ((A ⨂ B) ⨂ C) (A ⨂ (B ⨂ C)) assocHom = ⨂→AbGroup-elim-hom _ (λ x → f (snd x) .fst (fst x)) - (λ x y b → IsGroupHom.pres· (snd (f b)) x y) - (⊗elimProp (λ _ → isPropΠ2 λ _ _ → ⊗squash _ _) - (λ a b x y → (λ i → a ⊗ (linr b x y i)) - ∙∙ linr a (b ⊗ x) (b ⊗ y) - ∙∙ refl) - λ a b ind1 ind2 x y → cong₂ _+⊗_ (ind1 x y) (ind2 x y) - ∙∙ move4 (f x .fst a) (f y .fst a) (f x .fst b) (f y .fst b) _+⊗_ ⊗assoc ⊗comm - ∙∙ cong₂ _+⊗_ (sym (IsGroupHom.pres· (snd (f x)) a b)) (IsGroupHom.pres· (snd (f y)) a b)) - (⊗elimProp (λ _ → isPropΠ λ _ → ⊗squash _ _) - (λ a b c → (λ i → a ⊗ (flip b c i)) - ∙ flip a (b ⊗ c)) - (λ x y ind1 ind2 c → cong₂ _+⊗_ (ind1 c) (ind2 c) - ∙ IsGroupHom.pres· (snd (f c)) (-⊗ x) (-⊗ y))) - refl + (λ x y b → IsGroupHom.pres· (snd (f b)) x y) + (⊗elimProp (λ _ → isPropΠ2 λ _ _ → ⊗squash _ _) + (λ a b x y → (λ i → a ⊗ (linr b x y i)) + ∙∙ linr a (b ⊗ x) (b ⊗ y) + ∙∙ refl) + λ a b ind1 ind2 x y → cong₂ _+⊗_ (ind1 x y) (ind2 x y) + ∙∙ move4 (f x .fst a) (f y .fst a) (f x .fst b) (f y .fst b) + _+⊗_ ⊗assoc ⊗comm + ∙∙ cong₂ _+⊗_ (sym (IsGroupHom.pres· (snd (f x)) a b)) + (IsGroupHom.pres· (snd (f y)) a b)) + (⊗elimProp (λ _ → isPropΠ λ _ → ⊗squash _ _) + (λ a b c → (λ i → a ⊗ (flip b c i)) + ∙ flip a (b ⊗ c)) + (λ x y ind1 ind2 c → cong₂ _+⊗_ (ind1 c) (ind2 c) + ∙ IsGroupHom.pres· (snd (f c)) (-⊗ x) (-⊗ y))) + refl module _ {ℓ ℓ' ℓ'' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} {C : AbGroup ℓ''} where private - open AbGroupStr renaming (_+_ to +G ; -_ to -G) - _+A'_ = +G (snd A) - _+B'_ = +G (snd B) - _+C'_ = +G (snd C) - -A = -G (snd A) - -B = -G (snd B) - -C = -G (snd C) - f' : (a : fst A) → AbGroupHom (B ⨂ C) ((A ⨂ B) ⨂ C) - f' a = ⨂→AbGroup-elim-hom ((A ⨂ B) ⨂ C) - (λ bc → (a ⊗ fst bc) ⊗ snd bc) - (λ x y b → (λ i → (linr a x y i) ⊗ b) ∙ linl (a ⊗ x) (a ⊗ y) b) - (λ x y b → linr (a ⊗ x) y b) - (λ b c → flip (a ⊗ b) c ∙ λ i → flip a b (~ i) ⊗ c) - λ i → rCancelPrim a i ⊗ (0g (snd C)) + f' a = + ⨂→AbGroup-elim-hom ((A ⨂ B) ⨂ C) + (λ bc → (a ⊗ fst bc) ⊗ snd bc) + (λ x y b → (λ i → (linr a x y i) ⊗ b) ∙ linl (a ⊗ x) (a ⊗ y) b) + (λ x y b → linr (a ⊗ x) y b) + (λ b c → flip (a ⊗ b) c ∙ λ i → flip a b (~ i) ⊗ c) + λ i → rCancelPrim a i ⊗ (0g (snd C)) assocHom⁻ : AbGroupHom (A ⨂ (B ⨂ C)) ((A ⨂ B) ⨂ C) - assocHom⁻ = ⨂→AbGroup-elim-hom _ (λ abc → f' (fst abc) .fst (snd abc)) - (λ x y → ⊗elimProp (λ _ → ⊗squash _ _) - (λ b c → (λ i → linl x y b i ⊗ c) ∙ linl (x ⊗ b) (y ⊗ b) c) - λ a b ind1 ind2 → cong₂ _+⊗_ ind1 ind2 - ∙∙ move4 _ _ _ _ _+⊗_ ⊗assoc ⊗comm - ∙∙ cong₂ _+⊗_ (IsGroupHom.pres· (snd (f' x)) a b) - (IsGroupHom.pres· (snd (f' y)) a b)) - (λ a → IsGroupHom.pres· (snd (f' a))) - (λ a → ⊗elimProp (λ _ → ⊗squash _ _) - (λ b c → λ i → flip a b i ⊗ c) - λ x y ind1 ind2 → IsGroupHom.pres· (snd (f' a)) (-⊗ x) (-⊗ y) - ∙ cong₂ _+⊗_ ind1 ind2) - refl - - ⨂AssocIso : Iso (A ⨂₁ (B ⨂ C)) ((A ⨂ B) ⨂₁ C) - Iso.fun ⨂AssocIso = fst assocHom⁻ - Iso.inv ⨂AssocIso = fst assocHom - Iso.rightInv ⨂AssocIso = + assocHom⁻ = + ⨂→AbGroup-elim-hom _ (λ abc → f' (fst abc) .fst (snd abc)) + (λ x y → ⊗elimProp (λ _ → ⊗squash _ _) + (λ b c → (λ i → linl x y b i ⊗ c) ∙ linl (x ⊗ b) (y ⊗ b) c) + λ a b ind1 ind2 → cong₂ _+⊗_ ind1 ind2 + ∙∙ move4 _ _ _ _ _+⊗_ ⊗assoc ⊗comm + ∙∙ cong₂ _+⊗_ (IsGroupHom.pres· (snd (f' x)) a b) + (IsGroupHom.pres· (snd (f' y)) a b)) + (λ a → IsGroupHom.pres· (snd (f' a))) + (λ a → ⊗elimProp (λ _ → ⊗squash _ _) + (λ b c → λ i → flip a b i ⊗ c) + λ x y ind1 ind2 → IsGroupHom.pres· (snd (f' a)) (-⊗ x) (-⊗ y) + ∙ cong₂ _+⊗_ ind1 ind2) + refl + + ⨂assocIso : Iso (A ⨂₁ (B ⨂ C)) ((A ⨂ B) ⨂₁ C) + Iso.fun ⨂assocIso = fst assocHom⁻ + Iso.inv ⨂assocIso = fst assocHom + Iso.rightInv ⨂assocIso = ⊗elimProp (λ _ → ⊗squash _ _) (⊗elimProp (λ _ → isPropΠ (λ _ → ⊗squash _ _)) (λ a b c → refl) @@ -559,7 +489,7 @@ module _ {ℓ ℓ' ℓ'' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} {C : AbGr λ x y p q → cong (fst assocHom⁻) (IsGroupHom.pres· (snd assocHom) x y) ∙∙ IsGroupHom.pres· (snd assocHom⁻) (fst assocHom x) (fst assocHom y) ∙∙ cong₂ _+⊗_ p q - Iso.leftInv ⨂AssocIso = + Iso.leftInv ⨂assocIso = ⊗elimProp (λ _ → ⊗squash _ _) (λ a → ⊗elimProp (λ _ → ⊗squash _ _) (λ b c → refl) @@ -574,5 +504,5 @@ module _ {ℓ ℓ' ℓ'' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} {C : AbGr ∙∙ cong₂ _+⊗_ p q ⨂assoc : AbGroupEquiv (A ⨂ (B ⨂ C)) ((A ⨂ B) ⨂ C) - fst ⨂assoc = isoToEquiv ⨂AssocIso + fst ⨂assoc = isoToEquiv ⨂assocIso snd ⨂assoc = snd assocHom⁻ diff --git a/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda b/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda index e77da3c20..986b75964 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda @@ -47,468 +47,436 @@ private variable ℓ ℓ' ℓ'' : Level -open import Cubical.Foundations.SIP -univP : ∀ {ℓ} (A B : Type ℓ) → Type _ -univP {ℓ = ℓ} A B = (C : Type ℓ) → Iso (A → C) (B → C) - -univP' : ∀ {ℓ} (A B : Type ℓ) → Type _ -univP' {ℓ = ℓ} A B = Σ[ f ∈ ((C : Type ℓ) → Iso (A → C) (B → C)) ] Iso.fun (f B) ≡ λ _ b → b - -univP'' : ∀ {ℓ} (A B : Type ℓ) → Type _ -univP'' {ℓ = ℓ} A B = (P : (C : Type ℓ) → Type ℓ) → Iso ((x : A) → (P A)) ((x : B) → (P B)) - -univP→Id : ∀ {ℓ} (A B : Pointed ℓ) → univP'' (typ A) (typ B) → Iso (typ A) (typ B) -univP→Id A B is = Iso.inv (is (λ C → Iso C (typ B))) (λ _ → idIso) (pt A) - -open import Cubical.HITs.SmashProduct - -test : ∀ {ℓ} (A B C : Pointed ℓ) → univP'' (Smash (SmashPt A B) C) (Smash A (SmashPt B C)) -test A B C P = {!!} +stupidInd : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 → A 1 + → ((n : ℕ) → (A (suc n) → A (suc (suc n)))) + → (n : ℕ) → A n +stupidInd a0 a1 ind zero = a0 +stupidInd a0 a1 ind (suc zero) = a1 +stupidInd {A = A} a0 a1 ind (suc (suc n)) = + ind n (stupidInd {A = A} a0 a1 ind (suc n)) + +stupidInd' : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 + → ((n : ℕ) → (A n → A (suc n))) + → (n : ℕ) → A n +stupidInd' a0 ind zero = a0 +stupidInd' a0 ind (suc n) = ind n (stupidInd' a0 ind n) + +_+'_ : ℕ → ℕ → ℕ +zero +' m = m +suc n +' zero = suc n +suc n +' suc m = suc (suc (n + m)) + ++'-comm : (n m : ℕ) → n +' m ≡ m +' n ++'-comm zero zero = refl ++'-comm zero (suc m) = refl ++'-comm (suc n) zero = refl ++'-comm (suc n) (suc m) = cong (2 +_) (+-comm n m) + ++'≡+ : (n m : ℕ) → n +' m ≡ n + m ++'≡+ zero zero = refl ++'≡+ zero (suc m) = refl ++'≡+ (suc n) zero = cong suc (+-comm zero n) ++'≡+ (suc n) (suc m) = cong suc (sym (+-suc n m)) + +open AbGroupStr renaming (_+_ to _+Gr_ ; -_ to -Gr_) + +inducedFun-EM-raw : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} + → AbGroupHom G' H' + → ∀ n + → EM-raw G' n → EM-raw H' n +inducedFun-EM-raw f = + stupidInd (fst f) + (EMrec _ emsquash embase + (λ g → emloop (fst f g)) + λ g h → compPathR→PathP (sym + (sym (lUnit _) + ∙∙ cong (_∙ (sym (emloop (fst f h)))) (cong emloop (IsGroupHom.pres· (snd f) g h) + ∙ emloop-comp _ (fst f g) (fst f h)) + ∙∙ sym (∙assoc _ _ _) + ∙∙ cong (emloop (fst f g) ∙_) (rCancel _) + ∙∙ sym (rUnit _)))) + (λ n ind → λ { north → north + ; south → south + ; (merid a i) → merid (ind a) i} ) + +inducedFun-EM-rawIso : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} + → AbGroupEquiv G' H' + → ∀ n → Iso (EM-raw G' n) (EM-raw H' n) +Iso.fun (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , (snd e)) n +Iso.inv (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , isGroupHomInv e) n +Iso.rightInv (inducedFun-EM-rawIso e n) = h n + where + h : (n : ℕ) → section (inducedFun-EM-raw (fst e .fst , snd e) n) + (inducedFun-EM-raw (invEq (fst e) , isGroupHomInv e) n) + h = stupidInd + (secEq (fst e)) + (elimSet _ (λ _ → emsquash _ _) refl + (λ g → compPathR→PathP + ((sym (cong₂ _∙_ (cong emloop (secEq (fst e) g)) (sym (lUnit _)) + ∙ rCancel _))))) + λ n p → λ { north → refl + ; south → refl + ; (merid a i) k → merid (p a k) i} +Iso.leftInv (inducedFun-EM-rawIso e n) = h n where - h' : (Smash (SmashPt A B) C → P (Smash (SmashPt A B) C)) - h' basel = {!!} - h' baser = {!!} - h' (proj x y) = {!!} - h' (gluel a i) = {!!} - h' (gluer b i) = {!!} - - h : (Smash (SmashPt A B) C → P (Smash (SmashPt A B) C)) → Σ[ l ∈ P (Smash (SmashPt A B) C) ] {!!} - h = {!!} - - --- stupidInd : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 → A 1 --- → ((n : ℕ) → (A (suc n) → A (suc (suc n)))) --- → (n : ℕ) → A n --- stupidInd a0 a1 ind zero = a0 --- stupidInd a0 a1 ind (suc zero) = a1 --- stupidInd {A = A} a0 a1 ind (suc (suc n)) = ind n (stupidInd {A = A} a0 a1 ind (suc n)) - --- stupidInd' : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 --- → ((n : ℕ) → (A n → A (suc n))) --- → (n : ℕ) → A n --- stupidInd' a0 ind zero = a0 --- stupidInd' a0 ind (suc n) = ind n (stupidInd' a0 ind n) - --- _+'_ : ℕ → ℕ → ℕ --- zero +' m = m --- suc n +' zero = suc n --- suc n +' suc m = suc (suc (n + m)) - --- +'-comm : (n m : ℕ) → n +' m ≡ m +' n --- +'-comm zero zero = refl --- +'-comm zero (suc m) = refl --- +'-comm (suc n) zero = refl --- +'-comm (suc n) (suc m) = cong (2 +_) (+-comm n m) - --- +'≡+ : (n m : ℕ) → n +' m ≡ n + m --- +'≡+ zero zero = refl --- +'≡+ zero (suc m) = refl --- +'≡+ (suc n) zero = cong suc (+-comm zero n) --- +'≡+ (suc n) (suc m) = cong suc (sym (+-suc n m)) - --- open AbGroupStr renaming (_+_ to _+Gr_ ; -_ to -Gr_) - --- inducedFun-EM-raw : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} --- → AbGroupHom G' H' --- → ∀ n --- → EM-raw G' n → EM-raw H' n --- inducedFun-EM-raw f = --- stupidInd (fst f) --- (EMrec _ emsquash embase --- (λ g → emloop (fst f g)) --- λ g h → compPathR→PathP (sym --- (sym (lUnit _) --- ∙∙ cong (_∙ (sym (emloop (fst f h)))) (cong emloop (IsGroupHom.pres· (snd f) g h) --- ∙ emloop-comp _ (fst f g) (fst f h)) --- ∙∙ sym (∙assoc _ _ _) --- ∙∙ cong (emloop (fst f g) ∙_) (rCancel _) --- ∙∙ sym (rUnit _)))) --- (λ n ind → λ { north → north --- ; south → south --- ; (merid a i) → merid (ind a) i} ) - --- inducedFun-EM-rawIso : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} --- → AbGroupEquiv G' H' --- → ∀ n → Iso (EM-raw G' n) (EM-raw H' n) --- Iso.fun (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , (snd e)) n --- Iso.inv (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , isGroupHomInv e) n --- Iso.rightInv (inducedFun-EM-rawIso e n) = h n --- where --- h : (n : ℕ) → section (inducedFun-EM-raw (fst e .fst , snd e) n) --- (inducedFun-EM-raw (invEq (fst e) , isGroupHomInv e) n) --- h = stupidInd --- (secEq (fst e)) --- (elimSet _ (λ _ → emsquash _ _) refl --- (λ g → compPathR→PathP --- ((sym (cong₂ _∙_ (cong emloop (secEq (fst e) g)) (sym (lUnit _)) --- ∙ rCancel _))))) --- λ n p → λ { north → refl --- ; south → refl --- ; (merid a i) k → merid (p a k) i} --- Iso.leftInv (inducedFun-EM-rawIso e n) = h n --- where --- h : (n : ℕ) → retract (Iso.fun (inducedFun-EM-rawIso e n)) --- (Iso.inv (inducedFun-EM-rawIso e n)) --- h = stupidInd --- (retEq (fst e)) --- (elimSet _ (λ _ → emsquash _ _) refl --- (λ g → compPathR→PathP --- ((sym (cong₂ _∙_ (cong emloop (retEq (fst e) g)) (sym (lUnit _)) --- ∙ rCancel _))))) --- λ n p → λ { north → refl --- ; south → refl --- ; (merid a i) k → merid (p a k) i} - --- Iso→EMIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} → AbGroupEquiv G H → ∀ n → Iso (EM G n) (EM H n) --- Iso→EMIso is zero = inducedFun-EM-rawIso is zero --- Iso→EMIso is (suc zero) = inducedFun-EM-rawIso is 1 --- Iso→EMIso is (suc (suc n)) = mapCompIso (inducedFun-EM-rawIso is (suc (suc n))) - --- EM⊗-commFun : {G : AbGroup ℓ} {H : AbGroup ℓ'} → ∀ n → Iso (EM (G ⨂ H) n) (EM (H ⨂ G) n) --- EM⊗-commFun {G = G} {H = H} = Iso→EMIso (GroupIso→GroupEquiv ⨂-commIso) - --- EM⊗-assocIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} → ∀ n → Iso (EM (G ⨂ (H ⨂ L)) n) (EM ((G ⨂ H) ⨂ L) n) --- EM⊗-assocIso = Iso→EMIso (GroupIso→GroupEquiv (GroupEquiv→GroupIso ⨂assoc)) - --- EM⊗-raw-comm-sect : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} → ∀ n → Iso (EM (G ⨂ H) n) (EM (H ⨂ G) n) --- EM⊗-raw-comm-sect = {!!} - --- pathType : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → Type ℓ --- pathType n x p = sym (rUnitₖ (2 + n) x) ∙ (λ i → x +ₖ p i) ≡ sym (lUnitₖ (2 + n) x) ∙ λ i → p i +ₖ x - --- lem : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) --- → pathType n x p --- lem n x = J (λ x p → pathType n x p) refl - --- module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where --- private --- G = fst G' --- H = fst H' - --- strG = snd G' --- strH = snd H' - --- 0G = 0g strG --- 0H = 0g strH - --- _+G_ = _+Gr_ strG --- _+H_ = _+Gr_ strH - --- -H_ = -Gr_ strH --- -G_ = -Gr_ strG - --- ·₀' : H → (m : ℕ) → EM G' m → EM (G' ⨂ H') m --- ·₀' h = --- stupidInd --- (_⊗ h) --- (elimGroupoid _ (λ _ → emsquash) --- embase --- (λ g → emloop (g ⊗ h)) --- λ g l → --- compPathR→PathP --- (sym (∙assoc _ _ _ --- ∙∙ cong₂ _∙_ (sym (emloop-comp _ _ _) ∙ cong emloop (sym (linl g l h))) refl --- ∙∙ rCancel _))) --- λ n f → trRec (isOfHLevelTrunc (4 + n)) --- λ { north → 0ₖ (suc (suc n)) --- ; south → 0ₖ (suc (suc n)) --- ; (merid a i) → EM→ΩEM+1 (suc n) (f (EM-raw→EM _ _ a)) i} - --- ·₀ : G → (m : ℕ) → EM H' m → EM (G' ⨂ H') m --- ·₀ g = --- stupidInd (λ h → g ⊗ h) --- (elimGroupoid _ (λ _ → emsquash) --- embase --- (λ h → emloop (g ⊗ h)) --- λ h l → compPathR→PathP --- (sym (∙assoc _ _ _ --- ∙∙ cong₂ _∙_ (sym (emloop-comp _ _ _) ∙ cong emloop (sym (linr g h l))) refl --- ∙∙ rCancel _))) --- λ n f --- → trRec (isOfHLevelTrunc (4 + n)) --- λ { north → 0ₖ (suc (suc n)) --- ; south → 0ₖ (suc (suc n)) --- ; (merid a i) → EM→ΩEM+1 (suc n) (f (EM-raw→EM _ _ a)) i} - --- ·₀-distr : (g h : G) → (m : ℕ) (x : EM H' m) → ·₀ (g +G h) m x ≡ ·₀ g m x +ₖ ·₀ h m x --- ·₀-distr g h = --- stupidInd --- (linl g h) --- (elimSet _ (λ _ → emsquash _ _) --- refl --- (λ w → compPathR→PathP (sym ((λ i → emloop (linl g h w i) --- ∙ (lUnit (sym (cong₂+₁ (emloop (g ⊗ w)) (emloop (h ⊗ w)) i)) (~ i))) --- ∙∙ cong₂ _∙_ (emloop-comp _ (g ⊗ w) (h ⊗ w)) refl --- ∙∙ rCancel _)))) --- λ m ind → --- trElim (λ _ → isOfHLevelTruncPath) --- λ { north → refl --- ; south → refl --- ; (merid a i) k → z m ind a k i} - --- where --- z : (m : ℕ) → ((x : EM H' (suc m)) → ·₀ (g +G h) (suc m) x ≡ ·₀ g (suc m) x +ₖ ·₀ h (suc m) x) → (a : EM-raw H' (suc m)) --- → cong (·₀ (g +G h) (suc (suc m))) (cong ∣_∣ₕ (merid a)) ≡ --- cong₂ _+ₖ_ --- (cong (·₀ g (suc (suc m))) (cong ∣_∣ₕ (merid a))) --- (cong (·₀ h (suc (suc m))) (cong ∣_∣ₕ (merid a))) --- z m ind a = (λ i → EM→ΩEM+1 _ (ind (EM-raw→EM _ _ a) i)) --- ∙∙ EM→ΩEM+1-hom _ (·₀ g (suc m) (EM-raw→EM H' (suc m) a)) (·₀ h (suc m) (EM-raw→EM H' (suc m) a)) --- ∙∙ sym (cong₂+₂ m (cong (·₀ g (suc (suc m))) (cong ∣_∣ₕ (merid a))) --- (cong (·₀ h (suc (suc m))) (cong ∣_∣ₕ (merid a)))) - --- ·₀0 : (m : ℕ) → (g : G) → ·₀ g m (0ₖ m) ≡ 0ₖ m --- ·₀0 zero = rCancelPrim --- ·₀0 (suc zero) g = refl --- ·₀0 (suc (suc m)) g = refl - --- ·₀'0 : (m : ℕ) (h : H) → ·₀' h m (0ₖ m) ≡ 0ₖ m --- ·₀'0 zero = lCancelPrim --- ·₀'0 (suc zero) g = refl --- ·₀'0 (suc (suc m)) g = refl - --- 0·₀ : (m : ℕ) → (x : _) → ·₀ 0G m x ≡ 0ₖ m --- 0·₀ = --- stupidInd lCancelPrim --- (elimSet _ (λ _ → emsquash _ _) --- refl --- λ g → compPathR→PathP ((sym (emloop-1g _) --- ∙ cong emloop (sym (lCancelPrim g))) --- ∙∙ (λ i → rUnit (rUnit (cong (·₀ 0G 1) (emloop g)) i) i) --- ∙∙ sym (∙assoc _ _ _))) --- λ n f → trElim (λ _ → isOfHLevelTruncPath) --- λ { north → refl --- ; south → refl --- ; (merid a i) j → (cong (EM→ΩEM+1 (suc n)) (f (EM-raw→EM _ _ a)) --- ∙ EM→ΩEM+1-0ₖ _) j i} - --- 0·₀' : (m : ℕ) (g : _) → ·₀' 0H m g ≡ 0ₖ m --- 0·₀' = --- stupidInd --- rCancelPrim --- (elimSet _ (λ _ → emsquash _ _) --- refl --- λ g → compPathR→PathP (sym (∙assoc _ _ _ --- ∙∙ sym (rUnit _) ∙ sym (rUnit _) --- ∙∙ (cong emloop (rCancelPrim g) --- ∙ emloop-1g _)))) --- λ n f → trElim (λ _ → isOfHLevelTruncPath) --- λ { north → refl --- ; south → refl --- ; (merid a i) j → (cong (EM→ΩEM+1 (suc n)) (f (EM-raw→EM _ _ a)) --- ∙ EM→ΩEM+1-0ₖ _) j i} - --- cup∙ : ∀ n m → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m) --- cup∙ = --- stupidInd' --- (λ m g → (·₀ g m) , ·₀0 m g) --- λ n f → --- stupidInd' --- (λ g → (λ h → ·₀' h (suc n) g) , 0·₀' (suc n) g) --- λ m _ → main n m f - --- where --- main : (n m : ℕ) (ind : ((m : ℕ) → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m))) --- → EM G' (suc n) → EM∙ H' (suc m) →∙ EM∙ (G' ⨂ H') (suc (suc (n + m))) --- main zero m ind = --- elimGroupoid _ (λ _ → isOfHLevel↑∙ _ _) --- ((λ _ → 0ₖ (2 + m)) , refl) --- (f m) --- λ n h → finalpp m n h --- where --- f : (m : ℕ) → G → typ (Ω (EM∙ H' (suc m) →∙ EM∙ (G' ⨂ H') (suc (suc m)) ∙)) --- fst (f m g i) x = EM→ΩEM+1 _ (·₀ g _ x) i --- snd (f zero g i) j = EM→ΩEM+1-0ₖ (suc zero) j i --- snd (f (suc m) g i) j = EM→ΩEM+1-0ₖ (suc (suc m)) j i - --- f-hom-fst : (m : ℕ) (g h : G) → cong fst (f m (g +G h)) ≡ cong fst (f m g ∙ f m h) --- f-hom-fst m g h = --- (λ i j x → EM→ΩEM+1 _ (·₀-distr g h (suc m) x i) j) --- ∙∙ (λ i j x → EM→ΩEM+1-hom _ (·₀ g (suc m) x) (·₀ h (suc m) x) i j) --- ∙∙ sym (cong-∙ fst (f m g) (f m h)) - --- f-hom : (m : ℕ) (g h : G) → f m (g +G h) ≡ f m g ∙ f m h --- f-hom m g h = →∙Homogeneous≡Path (isHomogeneousEM _) _ _ (f-hom-fst m g h) - --- finalpp : (m : ℕ) (g h : G) → PathP (λ i → f m g i ≡ f m (g +G h) i) refl (f m h) --- finalpp m g h = --- compPathR→PathP (sym (rCancel _) --- ∙∙ cong (_∙ sym (f m (g +G h))) (f-hom m g h) --- ∙∙ sym (∙assoc _ _ _)) - --- main (suc n) m ind = --- trElim (λ _ → isOfHLevel↑∙ (2 + n) m) --- λ { north → (λ _ → 0ₖ (3 + (n + m))) , refl --- ; south → (λ _ → 0ₖ (3 + (n + m))) , refl --- ; (merid a i) → (λ x → EM→ΩEM+1 _ (ind _ (EM-raw→EM _ _ a) .fst x) i) --- , λ j → (cong (EM→ΩEM+1 (suc (suc (n + m)))) (ind (suc m) (EM-raw→EM G' (suc n) a) .snd) --- ∙ EM→ΩEM+1-0ₖ _) j i} - --- _⌣ₖ_ : {n m : ℕ} (x : EM G' n) (y : EM H' m) → EM (G' ⨂ H') (n +' m) --- _⌣ₖ_ x y = cup∙ _ _ x .fst y - --- ⌣ₖ-0ₖ : (n m : ℕ) (x : EM G' n) → (x ⌣ₖ 0ₖ m) ≡ 0ₖ (n +' m) --- ⌣ₖ-0ₖ n m x = cup∙ n m x .snd - --- 0ₖ-⌣ₖ : (n m : ℕ) (x : EM H' m) → ((0ₖ n) ⌣ₖ x) ≡ 0ₖ (n +' m) --- 0ₖ-⌣ₖ zero m = 0·₀ _ --- 0ₖ-⌣ₖ (suc zero) zero x = refl --- 0ₖ-⌣ₖ (suc (suc n)) zero x = refl --- 0ₖ-⌣ₖ (suc zero) (suc m) x = refl --- 0ₖ-⌣ₖ (suc (suc n)) (suc m) x = refl - --- private --- distrl1 : (n m : ℕ) → EM H' m → EM H' m → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) --- fst (distrl1 n m x y) z = z ⌣ₖ (x +ₖ y) --- snd (distrl1 n m x y) = 0ₖ-⌣ₖ n m _ - --- distrl2 : (n m : ℕ) → EM H' m → EM H' m → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) --- fst (distrl2 n m x y) z = (z ⌣ₖ x) +ₖ (z ⌣ₖ y) --- snd (distrl2 n m x y) = cong₂ _+ₖ_ (0ₖ-⌣ₖ n m x) (0ₖ-⌣ₖ n m y) ∙ rUnitₖ _ (0ₖ (n +' m)) - --- hLevLem : (n m : ℕ) → isOfHLevel (suc (suc m)) (EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') ((suc n) +' m)) --- hLevLem n m = subst (isOfHLevel (suc (suc m))) (λ i → EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') --- ((cong suc (+-comm m n) ∙ sym (+'≡+ (suc n) m)) i)) (isOfHLevel↑∙ m n) - - - - --- mainDistrL : (n m : ℕ) (x y : EM H' (suc m)) → distrl1 (suc n) (suc m) x y ≡ distrl2 (suc n) (suc m) x y --- mainDistrL n zero = --- wedgeConEM.fun H' H' 0 0 --- (λ _ _ → hLevLem _ _ _ _) --- (λ x → →∙Homogeneous≡ (isHomogeneousEM _) --- (funExt λ z → l x z)) --- (λ y → →∙Homogeneous≡ (isHomogeneousEM _) --- (funExt λ z → r y z )) --- λ i → →∙Homogeneous≡ (isHomogeneousEM (suc (suc (n + 0)))) --- (funExt (λ z → l≡r z i)) --- where --- l : (x : EM H' 1) (z : _) → (distrl1 (suc n) 1 embase x .fst z) ≡ (distrl2 (suc n) 1 embase x .fst z) --- l x z = cong (z ⌣ₖ_) (lUnitₖ _ x) --- ∙∙ sym (lUnitₖ _ (z ⌣ₖ x)) --- ∙∙ λ i → (⌣ₖ-0ₖ (suc n) (suc zero) z (~ i)) +ₖ (z ⌣ₖ x) - --- r : (z : EM H' 1) (x : EM G' (suc n)) → (distrl1 (suc n) 1 z embase .fst x) ≡ (distrl2 (suc n) 1 z embase .fst x) --- r y z = cong (z ⌣ₖ_) (rUnitₖ _ y) --- ∙∙ sym (rUnitₖ _ (z ⌣ₖ y)) --- ∙∙ λ i → (z ⌣ₖ y) +ₖ (⌣ₖ-0ₖ (suc n) (suc zero) z (~ i)) - --- l≡r : (z : EM G' (suc n)) → l embase z ≡ r embase z --- l≡r z = sym (lem _ _ (sym (⌣ₖ-0ₖ (suc n) (suc zero) z))) - --- mainDistrL n (suc m) = --- elim2 (λ _ _ → isOfHLevelPath (4 + m) (hLevLem _ _) _ _) --- (wedgeConEM.fun H' H' (suc m) (suc m) --- (λ x y p q → isOfHLevelPlus {n = suc (suc m)} (suc m) --- (hLevLem n (suc (suc m)) --- (distrl1 (suc n) (suc (suc m)) ∣ x ∣ ∣ y ∣) --- (distrl2 (suc n) (suc (suc m)) ∣ x ∣ ∣ y ∣) p q)) --- (λ x → →∙Homogeneous≡ (isHomogeneousEM _) --- (funExt (l x))) --- (λ x → →∙Homogeneous≡ (isHomogeneousEM _) --- (funExt (r x))) --- λ i → →∙Homogeneous≡ (isHomogeneousEM _) --- (funExt (λ z → l≡r z i))) --- where --- l : (x : EM-raw H' (suc (suc m))) (z : EM G' (suc n)) --- → (distrl1 (suc n) (suc (suc m)) (0ₖ _) ∣ x ∣ₕ .fst z) ≡ (distrl2 (suc n) (suc (suc m)) (0ₖ _) ∣ x ∣ₕ .fst z) --- l x z = cong (z ⌣ₖ_) (lUnitₖ (suc (suc m)) ∣ x ∣) --- ∙∙ sym (lUnitₖ _ (z ⌣ₖ ∣ x ∣)) --- ∙∙ λ i → (⌣ₖ-0ₖ (suc n) (suc (suc m)) z (~ i)) +ₖ (z ⌣ₖ ∣ x ∣) - --- r : (x : EM-raw H' (suc (suc m))) (z : EM G' (suc n)) --- → (distrl1 (suc n) (suc (suc m)) ∣ x ∣ₕ (0ₖ _) .fst z) ≡ (distrl2 (suc n) (suc (suc m)) ∣ x ∣ₕ (0ₖ _) .fst z) --- r x z = cong (z ⌣ₖ_) (rUnitₖ (suc (suc m)) ∣ x ∣) --- ∙∙ sym (rUnitₖ _ (z ⌣ₖ ∣ x ∣)) --- ∙∙ λ i → (z ⌣ₖ ∣ x ∣) +ₖ (⌣ₖ-0ₖ (suc n) (suc (suc m)) z (~ i)) - --- l≡r : (z : EM G' (suc n)) → l north z ≡ r north z --- l≡r z = sym (lem _ _ (sym (⌣ₖ-0ₖ (suc n) (suc (suc m)) z))) - - --- commer : {!!} --- commer = {!!} - --- module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where --- private --- G = fst G' --- H = fst H' - --- strG = snd G' --- strH = snd H' - --- 0G = 0g strG --- 0H = 0g strH - --- _+G_ = _+Gr_ strG --- _+H_ = _+Gr_ strH - --- -H_ = -Gr_ strH --- -G_ = -Gr_ strG - - --- distrr1 : (n m : ℕ) → EM G' n → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m) --- fst (distrr1 n m x y) z = (x +ₖ y) ⌣ₖ z --- snd (distrr1 n m x y) = ⌣ₖ-0ₖ n m _ - --- distrr2 : (n m : ℕ) → EM G' n → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m) --- fst (distrr2 n m x y) z = (x ⌣ₖ z) +ₖ (y ⌣ₖ z) --- snd (distrr2 n m x y) = cong₂ _+ₖ_ (⌣ₖ-0ₖ n m x) (⌣ₖ-0ₖ n m y) ∙ rUnitₖ _ (0ₖ (n +' m)) - --- hLevLem2 : (n m : ℕ) → isOfHLevel (suc (suc m)) (EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') ((suc n) +' m)) --- hLevLem2 n m = subst (isOfHLevel (suc (suc m))) (λ i → EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') --- ((cong suc (+-comm m n) ∙ sym (+'≡+ (suc n) m)) i)) (isOfHLevel↑∙ m n) - --- mainDistrR : (n m : ℕ) (x y : EM G' (suc n)) → distrr1 (suc n) (suc m) x y ≡ distrr2 (suc n) (suc m) x y --- mainDistrR zero m = --- wedgeConEM.fun G' G' 0 0 --- (λ _ _ → isOfHLevel↑∙ 1 m _ _) --- (λ x → →∙Homogeneous≡ (isHomogeneousEM _) --- (funExt (l x))) --- (λ x → →∙Homogeneous≡ (isHomogeneousEM _) --- (funExt (r x))) --- λ i → →∙Homogeneous≡ (isHomogeneousEM _) --- (funExt λ z → l≡r z i) --- where --- l : (x : _) (z : _) → _ ≡ _ --- l x z = --- (λ i → (lUnitₖ 1 x i) ⌣ₖ z) --- ∙∙ sym (lUnitₖ _ (x ⌣ₖ z)) --- ∙∙ λ i → 0ₖ-⌣ₖ _ _ z (~ i) +ₖ (x ⌣ₖ z) - --- r : (x : _) (z : _) → _ ≡ _ --- r x z = --- ((λ i → (rUnitₖ 1 x i) ⌣ₖ z)) --- ∙∙ sym (rUnitₖ _ _) --- ∙∙ λ i → (_⌣ₖ_ {n = 1} {m = suc m} x z) +ₖ 0ₖ-⌣ₖ (suc zero) (suc m) z (~ i) - --- l≡r : (z : _) → l embase z ≡ r embase z --- l≡r z = lem _ _ _ - --- mainDistrR (suc n) m = --- elim2 (λ _ _ → isOfHLevelPath (4 + n) --- (isOfHLevel↑∙ (2 + n) m) _ _) --- (wedgeConEM.fun _ _ _ _ --- (λ x y → isOfHLevelPath ((2 + n) + (2 + n)) --- {!subst !} _ _) --- (λ x → →∙Homogeneous≡ (isHomogeneousEM _) --- (funExt (l x))) --- (λ x → →∙Homogeneous≡ (isHomogeneousEM _) --- (funExt (r x))) --- λ i → →∙Homogeneous≡ (isHomogeneousEM _) --- (funExt λ z → r≡l z i)) --- where + h : (n : ℕ) → retract (Iso.fun (inducedFun-EM-rawIso e n)) + (Iso.inv (inducedFun-EM-rawIso e n)) + h = stupidInd + (retEq (fst e)) + (elimSet _ (λ _ → emsquash _ _) refl + (λ g → compPathR→PathP + ((sym (cong₂ _∙_ (cong emloop (retEq (fst e) g)) (sym (lUnit _)) + ∙ rCancel _))))) + λ n p → λ { north → refl + ; south → refl + ; (merid a i) k → merid (p a k) i} + +Iso→EMIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} + → AbGroupEquiv G H → ∀ n → Iso (EM G n) (EM H n) +Iso→EMIso is zero = inducedFun-EM-rawIso is zero +Iso→EMIso is (suc zero) = inducedFun-EM-rawIso is 1 +Iso→EMIso is (suc (suc n)) = mapCompIso (inducedFun-EM-rawIso is (suc (suc n))) + +EM⊗-commFun : {G : AbGroup ℓ} {H : AbGroup ℓ'} + → ∀ n → Iso (EM (G ⨂ H) n) (EM (H ⨂ G) n) +EM⊗-commFun {G = G} {H = H} = Iso→EMIso (GroupIso→GroupEquiv ⨂-commIso) + +EM⊗-assocIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} + → ∀ n → Iso (EM (G ⨂ (H ⨂ L)) n) (EM ((G ⨂ H) ⨂ L) n) +EM⊗-assocIso = Iso→EMIso (GroupIso→GroupEquiv (GroupEquiv→GroupIso ⨂assoc)) + +pathType : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → Type ℓ +pathType n x p = sym (rUnitₖ (2 + n) x) ∙ (λ i → x +ₖ p i) ≡ sym (lUnitₖ (2 + n) x) ∙ λ i → p i +ₖ x + +lem : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) + → pathType n x p +lem n x = J (λ x p → pathType n x p) refl + +module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where + private + G = fst G' + H = fst H' + + strG = snd G' + strH = snd H' + + 0G = 0g strG + 0H = 0g strH + + _+G_ = _+Gr_ strG + _+H_ = _+Gr_ strH + + -H_ = -Gr_ strH + -G_ = -Gr_ strG + + ·₀' : H → (m : ℕ) → EM G' m → EM (G' ⨂ H') m + ·₀' h = + stupidInd + (_⊗ h) + (elimGroupoid _ (λ _ → emsquash) + embase + (λ g → emloop (g ⊗ h)) + λ g l → + compPathR→PathP + (sym (∙assoc _ _ _ + ∙∙ cong₂ _∙_ (sym (emloop-comp _ _ _) ∙ cong emloop (sym (linl g l h))) refl + ∙∙ rCancel _))) + λ n f → trRec (isOfHLevelTrunc (4 + n)) + λ { north → 0ₖ (suc (suc n)) + ; south → 0ₖ (suc (suc n)) + ; (merid a i) → EM→ΩEM+1 (suc n) (f (EM-raw→EM _ _ a)) i} + + ·₀ : G → (m : ℕ) → EM H' m → EM (G' ⨂ H') m + ·₀ g = + stupidInd (λ h → g ⊗ h) + (elimGroupoid _ (λ _ → emsquash) + embase + (λ h → emloop (g ⊗ h)) + λ h l → compPathR→PathP + (sym (∙assoc _ _ _ + ∙∙ cong₂ _∙_ (sym (emloop-comp _ _ _) ∙ cong emloop (sym (linr g h l))) refl + ∙∙ rCancel _))) + λ n f + → trRec (isOfHLevelTrunc (4 + n)) + λ { north → 0ₖ (suc (suc n)) + ; south → 0ₖ (suc (suc n)) + ; (merid a i) → EM→ΩEM+1 (suc n) (f (EM-raw→EM _ _ a)) i} + + ·₀-distr : (g h : G) → (m : ℕ) (x : EM H' m) → ·₀ (g +G h) m x ≡ ·₀ g m x +ₖ ·₀ h m x + ·₀-distr g h = + stupidInd + (linl g h) + (elimSet _ (λ _ → emsquash _ _) + refl + (λ w → compPathR→PathP (sym ((λ i → emloop (linl g h w i) + ∙ (lUnit (sym (cong₂+₁ (emloop (g ⊗ w)) (emloop (h ⊗ w)) i)) (~ i))) + ∙∙ cong₂ _∙_ (emloop-comp _ (g ⊗ w) (h ⊗ w)) refl + ∙∙ rCancel _)))) + λ m ind → + trElim (λ _ → isOfHLevelTruncPath) + λ { north → refl + ; south → refl + ; (merid a i) k → z m ind a k i} + + where + z : (m : ℕ) → ((x : EM H' (suc m)) → ·₀ (g +G h) (suc m) x ≡ ·₀ g (suc m) x +ₖ ·₀ h (suc m) x) → (a : EM-raw H' (suc m)) + → cong (·₀ (g +G h) (suc (suc m))) (cong ∣_∣ₕ (merid a)) ≡ + cong₂ _+ₖ_ + (cong (·₀ g (suc (suc m))) (cong ∣_∣ₕ (merid a))) + (cong (·₀ h (suc (suc m))) (cong ∣_∣ₕ (merid a))) + z m ind a = (λ i → EM→ΩEM+1 _ (ind (EM-raw→EM _ _ a) i)) + ∙∙ EM→ΩEM+1-hom _ (·₀ g (suc m) (EM-raw→EM H' (suc m) a)) (·₀ h (suc m) (EM-raw→EM H' (suc m) a)) + ∙∙ sym (cong₂+₂ m (cong (·₀ g (suc (suc m))) (cong ∣_∣ₕ (merid a))) + (cong (·₀ h (suc (suc m))) (cong ∣_∣ₕ (merid a)))) + + ·₀0 : (m : ℕ) → (g : G) → ·₀ g m (0ₖ m) ≡ 0ₖ m + ·₀0 zero = rCancelPrim + ·₀0 (suc zero) g = refl + ·₀0 (suc (suc m)) g = refl + + ·₀'0 : (m : ℕ) (h : H) → ·₀' h m (0ₖ m) ≡ 0ₖ m + ·₀'0 zero = lCancelPrim + ·₀'0 (suc zero) g = refl + ·₀'0 (suc (suc m)) g = refl + + 0·₀ : (m : ℕ) → (x : _) → ·₀ 0G m x ≡ 0ₖ m + 0·₀ = + stupidInd lCancelPrim + (elimSet _ (λ _ → emsquash _ _) + refl + λ g → compPathR→PathP ((sym (emloop-1g _) + ∙ cong emloop (sym (lCancelPrim g))) + ∙∙ (λ i → rUnit (rUnit (cong (·₀ 0G 1) (emloop g)) i) i) + ∙∙ sym (∙assoc _ _ _))) + λ n f → trElim (λ _ → isOfHLevelTruncPath) + λ { north → refl + ; south → refl + ; (merid a i) j → (cong (EM→ΩEM+1 (suc n)) (f (EM-raw→EM _ _ a)) + ∙ EM→ΩEM+1-0ₖ _) j i} + + 0·₀' : (m : ℕ) (g : _) → ·₀' 0H m g ≡ 0ₖ m + 0·₀' = + stupidInd + rCancelPrim + (elimSet _ (λ _ → emsquash _ _) + refl + λ g → compPathR→PathP (sym (∙assoc _ _ _ + ∙∙ sym (rUnit _) ∙ sym (rUnit _) + ∙∙ (cong emloop (rCancelPrim g) + ∙ emloop-1g _)))) + λ n f → trElim (λ _ → isOfHLevelTruncPath) + λ { north → refl + ; south → refl + ; (merid a i) j → (cong (EM→ΩEM+1 (suc n)) (f (EM-raw→EM _ _ a)) + ∙ EM→ΩEM+1-0ₖ _) j i} + + cup∙ : ∀ n m → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m) + cup∙ = + stupidInd' + (λ m g → (·₀ g m) , ·₀0 m g) + λ n f → + stupidInd' + (λ g → (λ h → ·₀' h (suc n) g) , 0·₀' (suc n) g) + λ m _ → main n m f + + where + main : (n m : ℕ) (ind : ((m : ℕ) → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m))) + → EM G' (suc n) → EM∙ H' (suc m) →∙ EM∙ (G' ⨂ H') (suc (suc (n + m))) + main zero m ind = + elimGroupoid _ (λ _ → isOfHLevel↑∙ _ _) + ((λ _ → 0ₖ (2 + m)) , refl) + (f m) + λ n h → finalpp m n h + where + f : (m : ℕ) → G → typ (Ω (EM∙ H' (suc m) →∙ EM∙ (G' ⨂ H') (suc (suc m)) ∙)) + fst (f m g i) x = EM→ΩEM+1 _ (·₀ g _ x) i + snd (f zero g i) j = EM→ΩEM+1-0ₖ (suc zero) j i + snd (f (suc m) g i) j = EM→ΩEM+1-0ₖ (suc (suc m)) j i + + f-hom-fst : (m : ℕ) (g h : G) → cong fst (f m (g +G h)) ≡ cong fst (f m g ∙ f m h) + f-hom-fst m g h = + (λ i j x → EM→ΩEM+1 _ (·₀-distr g h (suc m) x i) j) + ∙∙ (λ i j x → EM→ΩEM+1-hom _ (·₀ g (suc m) x) (·₀ h (suc m) x) i j) + ∙∙ sym (cong-∙ fst (f m g) (f m h)) + + f-hom : (m : ℕ) (g h : G) → f m (g +G h) ≡ f m g ∙ f m h + f-hom m g h = →∙Homogeneous≡Path (isHomogeneousEM _) _ _ (f-hom-fst m g h) + + finalpp : (m : ℕ) (g h : G) → PathP (λ i → f m g i ≡ f m (g +G h) i) refl (f m h) + finalpp m g h = + compPathR→PathP (sym (rCancel _) + ∙∙ cong (_∙ sym (f m (g +G h))) (f-hom m g h) + ∙∙ sym (∙assoc _ _ _)) + + main (suc n) m ind = + trElim (λ _ → isOfHLevel↑∙ (2 + n) m) + λ { north → (λ _ → 0ₖ (3 + (n + m))) , refl + ; south → (λ _ → 0ₖ (3 + (n + m))) , refl + ; (merid a i) → (λ x → EM→ΩEM+1 _ (ind _ (EM-raw→EM _ _ a) .fst x) i) + , λ j → (cong (EM→ΩEM+1 (suc (suc (n + m)))) (ind (suc m) (EM-raw→EM G' (suc n) a) .snd) + ∙ EM→ΩEM+1-0ₖ _) j i} + + _⌣ₖ_ : {n m : ℕ} (x : EM G' n) (y : EM H' m) → EM (G' ⨂ H') (n +' m) + _⌣ₖ_ x y = cup∙ _ _ x .fst y + + ⌣ₖ-0ₖ : (n m : ℕ) (x : EM G' n) → (x ⌣ₖ 0ₖ m) ≡ 0ₖ (n +' m) + ⌣ₖ-0ₖ n m x = cup∙ n m x .snd + + 0ₖ-⌣ₖ : (n m : ℕ) (x : EM H' m) → ((0ₖ n) ⌣ₖ x) ≡ 0ₖ (n +' m) + 0ₖ-⌣ₖ zero m = 0·₀ _ + 0ₖ-⌣ₖ (suc zero) zero x = refl + 0ₖ-⌣ₖ (suc (suc n)) zero x = refl + 0ₖ-⌣ₖ (suc zero) (suc m) x = refl + 0ₖ-⌣ₖ (suc (suc n)) (suc m) x = refl + + private + distrl1 : (n m : ℕ) → EM H' m → EM H' m → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) + fst (distrl1 n m x y) z = z ⌣ₖ (x +ₖ y) + snd (distrl1 n m x y) = 0ₖ-⌣ₖ n m _ + + distrl2 : (n m : ℕ) → EM H' m → EM H' m → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) + fst (distrl2 n m x y) z = (z ⌣ₖ x) +ₖ (z ⌣ₖ y) + snd (distrl2 n m x y) = cong₂ _+ₖ_ (0ₖ-⌣ₖ n m x) (0ₖ-⌣ₖ n m y) ∙ rUnitₖ _ (0ₖ (n +' m)) + + hLevLem : (n m : ℕ) → isOfHLevel (suc (suc m)) (EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') ((suc n) +' m)) + hLevLem n m = subst (isOfHLevel (suc (suc m))) (λ i → EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') + ((cong suc (+-comm m n) ∙ sym (+'≡+ (suc n) m)) i)) (isOfHLevel↑∙ m n) + + + + + mainDistrL : (n m : ℕ) (x y : EM H' (suc m)) → distrl1 (suc n) (suc m) x y ≡ distrl2 (suc n) (suc m) x y + mainDistrL n zero = + wedgeConEM.fun H' H' 0 0 + (λ _ _ → hLevLem _ _ _ _) + (λ x → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt λ z → l x z)) + (λ y → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt λ z → r y z )) + λ i → →∙Homogeneous≡ (isHomogeneousEM (suc (suc (n + 0)))) + (funExt (λ z → l≡r z i)) + where + l : (x : EM H' 1) (z : _) → (distrl1 (suc n) 1 embase x .fst z) ≡ (distrl2 (suc n) 1 embase x .fst z) + l x z = cong (z ⌣ₖ_) (lUnitₖ _ x) + ∙∙ sym (lUnitₖ _ (z ⌣ₖ x)) + ∙∙ λ i → (⌣ₖ-0ₖ (suc n) (suc zero) z (~ i)) +ₖ (z ⌣ₖ x) + + r : (z : EM H' 1) (x : EM G' (suc n)) → (distrl1 (suc n) 1 z embase .fst x) ≡ (distrl2 (suc n) 1 z embase .fst x) + r y z = cong (z ⌣ₖ_) (rUnitₖ _ y) + ∙∙ sym (rUnitₖ _ (z ⌣ₖ y)) + ∙∙ λ i → (z ⌣ₖ y) +ₖ (⌣ₖ-0ₖ (suc n) (suc zero) z (~ i)) + + l≡r : (z : EM G' (suc n)) → l embase z ≡ r embase z + l≡r z = sym (lem _ _ (sym (⌣ₖ-0ₖ (suc n) (suc zero) z))) + + mainDistrL n (suc m) = + elim2 (λ _ _ → isOfHLevelPath (4 + m) (hLevLem _ _) _ _) + (wedgeConEM.fun H' H' (suc m) (suc m) + (λ x y p q → isOfHLevelPlus {n = suc (suc m)} (suc m) + (hLevLem n (suc (suc m)) + (distrl1 (suc n) (suc (suc m)) ∣ x ∣ ∣ y ∣) + (distrl2 (suc n) (suc (suc m)) ∣ x ∣ ∣ y ∣) p q)) + (λ x → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt (l x))) + (λ x → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt (r x))) + λ i → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt (λ z → l≡r z i))) + where + l : (x : EM-raw H' (suc (suc m))) (z : EM G' (suc n)) + → (distrl1 (suc n) (suc (suc m)) (0ₖ _) ∣ x ∣ₕ .fst z) ≡ (distrl2 (suc n) (suc (suc m)) (0ₖ _) ∣ x ∣ₕ .fst z) + l x z = cong (z ⌣ₖ_) (lUnitₖ (suc (suc m)) ∣ x ∣) + ∙∙ sym (lUnitₖ _ (z ⌣ₖ ∣ x ∣)) + ∙∙ λ i → (⌣ₖ-0ₖ (suc n) (suc (suc m)) z (~ i)) +ₖ (z ⌣ₖ ∣ x ∣) + + r : (x : EM-raw H' (suc (suc m))) (z : EM G' (suc n)) + → (distrl1 (suc n) (suc (suc m)) ∣ x ∣ₕ (0ₖ _) .fst z) ≡ (distrl2 (suc n) (suc (suc m)) ∣ x ∣ₕ (0ₖ _) .fst z) + r x z = cong (z ⌣ₖ_) (rUnitₖ (suc (suc m)) ∣ x ∣) + ∙∙ sym (rUnitₖ _ (z ⌣ₖ ∣ x ∣)) + ∙∙ λ i → (z ⌣ₖ ∣ x ∣) +ₖ (⌣ₖ-0ₖ (suc n) (suc (suc m)) z (~ i)) + + l≡r : (z : EM G' (suc n)) → l north z ≡ r north z + l≡r z = sym (lem _ _ (sym (⌣ₖ-0ₖ (suc n) (suc (suc m)) z))) + +module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where + private + G = fst G' + H = fst H' + + strG = snd G' + strH = snd H' + + 0G = 0g strG + 0H = 0g strH + + _+G_ = _+Gr_ strG + _+H_ = _+Gr_ strH + + -H_ = -Gr_ strH + -G_ = -Gr_ strG + + + distrr1 : (n m : ℕ) → EM G' n → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m) + fst (distrr1 n m x y) z = (x +ₖ y) ⌣ₖ z + snd (distrr1 n m x y) = ⌣ₖ-0ₖ n m _ + + distrr2 : (n m : ℕ) → EM G' n → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m) + fst (distrr2 n m x y) z = (x ⌣ₖ z) +ₖ (y ⌣ₖ z) + snd (distrr2 n m x y) = cong₂ _+ₖ_ (⌣ₖ-0ₖ n m x) (⌣ₖ-0ₖ n m y) ∙ rUnitₖ _ (0ₖ (n +' m)) + + hLevLem2 : (n m : ℕ) → isOfHLevel (suc (suc m)) (EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') ((suc n) +' m)) + hLevLem2 n m = subst (isOfHLevel (suc (suc m))) (λ i → EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') + ((cong suc (+-comm m n) ∙ sym (+'≡+ (suc n) m)) i)) (isOfHLevel↑∙ m n) + + mainDistrR : (n m : ℕ) (x y : EM G' (suc n)) → distrr1 (suc n) (suc m) x y ≡ distrr2 (suc n) (suc m) x y + mainDistrR zero m = + wedgeConEM.fun G' G' 0 0 + (λ _ _ → isOfHLevel↑∙ 1 m _ _) + (λ x → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt (l x))) + (λ x → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt (r x))) + λ i → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt λ z → l≡r z i) + where + l : (x : _) (z : _) → _ ≡ _ + l x z = + (λ i → (lUnitₖ 1 x i) ⌣ₖ z) + ∙∙ sym (lUnitₖ _ (x ⌣ₖ z)) + ∙∙ λ i → 0ₖ-⌣ₖ _ _ z (~ i) +ₖ (x ⌣ₖ z) + + r : (x : _) (z : _) → _ ≡ _ + r x z = + ((λ i → (rUnitₖ 1 x i) ⌣ₖ z)) + ∙∙ sym (rUnitₖ _ _) + ∙∙ λ i → (_⌣ₖ_ {n = 1} {m = suc m} x z) +ₖ 0ₖ-⌣ₖ (suc zero) (suc m) z (~ i) + + l≡r : (z : _) → l embase z ≡ r embase z + l≡r z = lem _ _ _ + + mainDistrR (suc n) m = + elim2 (λ _ _ → isOfHLevelPath (4 + n) + (isOfHLevel↑∙ (2 + n) m) _ _) + (wedgeConEM.fun _ _ _ _ + (λ x y → isOfHLevelPath ((2 + n) + (2 + n)) + {!subst !} _ _) + (λ x → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt (l x))) + (λ x → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt (r x))) + λ i → →∙Homogeneous≡ (isHomogeneousEM _) + (funExt λ z → r≡l z i)) + where --- l : (x : _) (z : _) → _ ≡ _ --- l x z = (λ i → (lUnitₖ _ ∣ x ∣ i) ⌣ₖ z) --- ∙∙ sym (lUnitₖ _ (∣ x ∣ ⌣ₖ z)) --- ∙∙ λ i → 0ₖ-⌣ₖ _ _ z (~ i) +ₖ (∣ x ∣ ⌣ₖ z) - --- r : (x : _) (z : _) → _ ≡ _ --- r x z = (λ i → (rUnitₖ _ ∣ x ∣ i) ⌣ₖ z) --- ∙∙ sym (rUnitₖ _ (∣ x ∣ ⌣ₖ z)) --- ∙∙ λ i → (∣ x ∣ ⌣ₖ z) +ₖ 0ₖ-⌣ₖ _ _ z (~ i) - --- r≡l : (z : _) → l north z ≡ r north z --- r≡l z = lem _ _ _ + l : (x : _) (z : _) → _ ≡ _ + l x z = (λ i → (lUnitₖ _ ∣ x ∣ i) ⌣ₖ z) + ∙∙ sym (lUnitₖ _ (∣ x ∣ ⌣ₖ z)) + ∙∙ λ i → 0ₖ-⌣ₖ _ _ z (~ i) +ₖ (∣ x ∣ ⌣ₖ z) + + r : (x : _) (z : _) → _ ≡ _ + r x z = (λ i → (rUnitₖ _ ∣ x ∣ i) ⌣ₖ z) + ∙∙ sym (rUnitₖ _ (∣ x ∣ ⌣ₖ z)) + ∙∙ λ i → (∣ x ∣ ⌣ₖ z) +ₖ 0ₖ-⌣ₖ _ _ z (~ i) + + r≡l : (z : _) → l north z ≡ r north z + r≡l z = lem _ _ _ diff --git a/Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda b/Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda index bfbdbf532..a1688dc55 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda @@ -1,34 +1,30 @@ -{-# OPTIONS --cubical --no-import-sorts --safe --experimental-lossy-unification #-} +{-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.Algebra.Group.EilenbergMacLane.GroupStructure where open import Cubical.Algebra.Group.EilenbergMacLane.Base open import Cubical.Algebra.Group.EilenbergMacLane.WedgeConnectivity +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Properties +open import Cubical.Algebra.AbGroup.Base + +open import Cubical.Data.Nat hiding (_·_) + open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) -open import Cubical.Foundations.Path open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Univalence open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Transport -open import Cubical.Data.Unit -open import Cubical.Data.Sigma -open import Cubical.Algebra.Group.Base -open import Cubical.Algebra.Group.Properties open import Cubical.Homotopy.Connected + open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec; elim to trElim) open import Cubical.HITs.EilenbergMacLane1 -open import Cubical.Algebra.AbGroup.Base -open import Cubical.Data.Empty - renaming (rec to ⊥-rec) open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) -open import Cubical.Data.Nat hiding (_·_) open import Cubical.HITs.Susp + open import Cubical.Functions.Morphism open import Cubical.Foundations.Path @@ -41,48 +37,52 @@ module _ {G : AbGroup ℓ} where open AbGroupStr (snd G) renaming (_+_ to _+G_ ; -_ to -G_ ; assoc to assocG) private - help : (n : ℕ) → n + (4 + n) ≡ (2 + n) + (2 + n) - help n = +-suc n (3 + n) ∙ cong suc (+-suc n (suc (suc n))) - - hLevHelp : (n : ℕ) → isOfHLevel ((2 + n) + (2 + n)) (EM G (2 + n)) - hLevHelp n = - transport (λ i → isOfHLevel (help n i) (EM G (2 + n))) - (isOfHLevelPlus {n = 4 + n} n (isOfHLevelTrunc (4 + n))) - - helper : (g h : fst G) - → PathP (λ i → Path (EM₁ (AbGroup→Group G)) - (emloop h i) (emloop h i)) (emloop g) (emloop g) - helper g h = - compPathL→PathP - (cong (sym (emloop h) ∙_) - (sym (emloop-comp _ g h) - ∙∙ cong emloop (comm g h) - ∙∙ emloop-comp _ h g) - ∙∙ ∙assoc _ _ _ - ∙∙ cong (_∙ emloop g) (lCancel _) - ∙ sym (lUnit _)) + help : (n : ℕ) → n + (4 + n) ≡ (2 + n) + (2 + n) + help n = +-suc n (3 + n) ∙ cong suc (+-suc n (suc (suc n))) + + hLevHelp : (n : ℕ) → isOfHLevel ((2 + n) + (2 + n)) (EM G (2 + n)) + hLevHelp n = + transport (λ i → isOfHLevel (help n i) (EM G (2 + n))) + (isOfHLevelPlus {n = 4 + n} n (isOfHLevelTrunc (4 + n))) + + helper : (g h : fst G) + → PathP (λ i → Path (EM₁ (AbGroup→Group G)) + (emloop h i) (emloop h i)) (emloop g) (emloop g) + helper g h = + compPathL→PathP + (cong (sym (emloop h) ∙_) + (sym (emloop-comp _ g h) + ∙∙ cong emloop (comm g h) + ∙∙ emloop-comp _ h g) + ∙∙ ∙assoc _ _ _ + ∙∙ cong (_∙ emloop g) (lCancel _) + ∙ sym (lUnit _)) + _+ₖ_ : {n : ℕ} → EM G n → EM G n → EM G n _+ₖ_ {n = zero} = _+G_ _+ₖ_ {n = suc zero} = rec _ (isGroupoidΠ (λ _ → emsquash)) (λ x → x) - hi! + looper λ g h i j x → el g h x i j where + looper : fst G → (λ x → x) ≡ (λ x → x) + looper g = funExt (elimSet _ (λ _ → emsquash _ _) + (emloop g) + (helper g)) + lol : (g h : fst G) → Square (emloop g) (emloop (g +G h)) refl (emloop h) lol g h = compPath-filler (emloop g) (emloop h) ▷ sym (emloop-comp _ g h) - hi! : fst G → (λ x → x) ≡ (λ x → x) - hi! g = funExt (elimSet _ (λ _ → emsquash _ _) - (emloop g) - (helper g)) el : (g h : fst G) (x : EM₁ (AbGroup→Group G)) - → Square (λ j → hi! g j x) (λ j → hi! ((snd (AbGroup→Group G) GroupStr.· g) h) j x) - refl λ j → hi! h j x + → Square (λ j → looper g j x) + (λ j → looper ((snd (AbGroup→Group G) GroupStr.· g) h) j x) + refl λ j → looper h j x el g h = elimProp _ (λ _ → isOfHLevelPathP' 1 (emsquash _ _) _ _) (lol g h) + _+ₖ_ {n = suc (suc n)} = trRec2 (isOfHLevelTrunc (4 + n)) (wedgeConEM.fun G G (suc n) (suc n) @@ -160,12 +160,14 @@ module _ {G : AbGroup ℓ} where refl) open import Cubical.Homotopy.Loopspace - cong₂+₁ : (p q : typ (Ω (EM∙ G 1))) → cong₂ (λ x y → x +[ 1 ]ₖ y) p q ≡ p ∙ q + cong₂+₁ : (p q : typ (Ω (EM∙ G 1))) + → cong₂ (λ x y → x +[ 1 ]ₖ y) p q ≡ p ∙ q cong₂+₁ p q = (cong₂Funct (λ x y → x +[ 1 ]ₖ y) p q) ∙ (λ i → (cong (λ x → rUnitₖ 1 x i) p) ∙ (cong (λ x → lUnitₖ 1 x i) q)) - cong₂+₂ : (n : ℕ) (p q : typ (Ω (EM∙ G (suc (suc n))))) → cong₂ (λ x y → x +[ (2 + n) ]ₖ y) p q ≡ p ∙ q + cong₂+₂ : (n : ℕ) (p q : typ (Ω (EM∙ G (suc (suc n))))) + → cong₂ (λ x y → x +[ (2 + n) ]ₖ y) p q ≡ p ∙ q cong₂+₂ n p q = (cong₂Funct (λ x y → x +[ (2 + n) ]ₖ y) p q) ∙ (λ i → (cong (λ x → rUnitₖ (2 + n) x i) p) ∙ (cong (λ x → lUnitₖ (2 + n) x i) q)) @@ -207,24 +209,32 @@ module _ {G : AbGroup ℓ} where main : (x : EM G 1) (p : embase ≡ x) → decoder x p ≡ sym p main x = J (λ x p → decoder x p ≡ sym p) refl - cong-₂ : (n : ℕ) (p : typ (Ω (EM∙ G (2 + n)))) → cong (λ x → -[ 2 + n ]ₖ x) p ≡ sym p + cong-₂ : (n : ℕ) (p : typ (Ω (EM∙ G (2 + n)))) + → cong (λ x → -[ 2 + n ]ₖ x) p ≡ sym p cong-₂ n p = main _ p where - pp : (a : _) → PathP (λ i → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ → ∣ merid a i ∣ₕ ≡ 0ₖ (2 + n)) - (cong (λ x → -[ 2 + n ]ₖ x)) λ p → cong ∣_∣ₕ (sym (merid ptS)) ∙ cong (λ x → -[ 2 + n ]ₖ x) p + pp : (a : _) + → PathP (λ i → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ → ∣ merid a i ∣ₕ ≡ 0ₖ (2 + n)) + (cong (λ x → -[ 2 + n ]ₖ x)) + λ p → cong ∣_∣ₕ (sym (merid ptS)) ∙ cong (λ x → -[ 2 + n ]ₖ x) p pp a = toPathP (funExt λ x → - (λ k → transp (λ i → Path (EM G (2 + n)) ∣ merid a (i ∨ k) ∣ ∣ ptS ∣) k (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) - (cong (-ₖ-syntax (suc (suc n))) - (transp (λ j → Path (EM G (2 + n)) ∣ ptS ∣ ∣ merid a (~ j ∧ ~ k) ∣) k (compPath-filler x (sym (cong ∣_∣ₕ (merid a))) k))) k)) - ∙∙ cong (cong ∣_∣ₕ (sym (merid a)) ∙_) (cong-∙ (λ x → -[ 2 + n ]ₖ x) x (sym (cong ∣_∣ₕ (merid a))) + (λ k → transp (λ i → Path (EM G (2 + n)) ∣ merid a (i ∨ k) ∣ ∣ ptS ∣) k + (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) + (cong (-ₖ-syntax (suc (suc n))) + (transp (λ j → Path (EM G (2 + n)) ∣ ptS ∣ ∣ merid a (~ j ∧ ~ k) ∣) k + (compPath-filler x (sym (cong ∣_∣ₕ (merid a))) k))) k)) + ∙∙ cong (cong ∣_∣ₕ (sym (merid a)) ∙_) + (cong-∙ (λ x → -[ 2 + n ]ₖ x) x (sym (cong ∣_∣ₕ (merid a))) ∙ isCommΩEM (suc n) (cong (λ x → -[ 2 + n ]ₖ x) x) (cong ∣_∣ₕ (σ-EM n a))) - ∙∙ (λ k → (λ i → ∣ merid a (~ i ∨ k) ∣) ∙ (λ i → ∣ compPath-filler' (merid a) (sym (merid ptS)) (~ k) i ∣) + ∙∙ (λ k → (λ i → ∣ merid a (~ i ∨ k) ∣) + ∙ (λ i → ∣ compPath-filler' (merid a) (sym (merid ptS)) (~ k) i ∣) ∙ cong (λ x → -ₖ-syntax (suc (suc n)) x) x) ∙ sym (lUnit _)) - decoder : (x : EM G (2 + n)) → 0ₖ (2 + n) ≡ x → x ≡ 0ₖ (2 + n) + decoder : (x : EM G (2 + n)) + → 0ₖ (2 + n) ≡ x → x ≡ 0ₖ (2 + n) decoder = trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelTruncPath {n = 4 + n}) λ { north → pp ptS i0 @@ -244,7 +254,8 @@ module _ {G : AbGroup ℓ} where rCancelₖ (suc (suc n)) = trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) λ { north → refl - ; south i → +ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣ (-ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣) + ; south i → +ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣ + (-ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣) ; (merid a i) j → hcomp (λ r → λ { (i = i0) → 0ₖ (2 + n) ; (i = i1) → ∣ merid ptS (~ j ∧ r) ∣ₕ -[ 2 + n ]ₖ ∣ merid ptS (~ j ∧ r) ∣ @@ -253,7 +264,8 @@ module _ {G : AbGroup ℓ} where ; (j = i1) → 0ₖ (2 + n)}) (help' a j i) } where - help' : (a : _) → cong₂ (λ x y → ∣ x ∣ -[ suc (suc n) ]ₖ ∣ y ∣) (σ-EM n a) (σ-EM n a) ≡ refl + help' : (a : _) + → cong₂ (λ x y → ∣ x ∣ -[ suc (suc n) ]ₖ ∣ y ∣) (σ-EM n a) (σ-EM n a) ≡ refl help' a = cong₂+₂ n (cong ∣_∣ₕ (σ-EM n a)) (cong (λ x → -[ 2 + n ]ₖ ∣ x ∣) (σ-EM n a)) ∙∙ cong (cong ∣_∣ₕ (σ-EM n a) ∙_) (cong-₂ n (cong ∣_∣ₕ (σ-EM n a))) @@ -270,8 +282,9 @@ module _ {G : AbGroup ℓ} where (λ _ _ → refl) λ g i y z k → lem g y z k i where - lem : (g : fst G) (y z : _) → cong (λ x → x +[ suc zero ]ₖ (y +[ suc zero ]ₖ z)) (emloop g) - ≡ cong (λ x → (x +[ suc zero ]ₖ y) +[ suc zero ]ₖ z) (emloop g) + lem : (g : fst G) (y z : _) + → cong (λ x → x +[ suc zero ]ₖ (y +[ suc zero ]ₖ z)) (emloop g) + ≡ cong (λ x → (x +[ suc zero ]ₖ y) +[ suc zero ]ₖ z) (emloop g) lem g = elimProp _ (λ _ → isPropΠ λ _ → emsquash _ _ _ _) (elimProp _ (λ _ → emsquash _ _ _ _) @@ -292,15 +305,19 @@ module _ {G : AbGroup ℓ} where ∙∙ λ i → (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ merid ptS i ∣ₕ) lem c = EM-raw-elim G (suc n) - (λ _ → isOfHLevelΠ (2 + n) (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _)) - (EM-raw-elim G (suc n) (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _) - ((sym (rUnit refl) - ◁ λ _ → refl) - ▷ (sym (lCancel (cong ∣_∣ₕ (merid ptS))) - ∙ λ i → (λ j → ∣ merid ptS (~ j ∨ ~ i) ∣ₕ) - ∙∙ lUnit (λ j → ∣ merid ptS (~ j ∧ ~ i) ∣ₕ) i - ∙∙ cong ∣_∣ₕ (merid ptS)))) - main : (c a b : _) → (∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ c ∣ₕ) ≡ (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ c ∣ₕ) + (λ _ → isOfHLevelΠ (2 + n) + (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _)) + (EM-raw-elim G (suc n) + (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _) + ((sym (rUnit refl) + ◁ λ _ → refl) + ▷ (sym (lCancel (cong ∣_∣ₕ (merid ptS))) + ∙ λ i → (λ j → ∣ merid ptS (~ j ∨ ~ i) ∣ₕ) + ∙∙ lUnit (λ j → ∣ merid ptS (~ j ∧ ~ i) ∣ₕ) i + ∙∙ cong ∣_∣ₕ (merid ptS)))) + main : (c a b : _) + → (∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ c ∣ₕ) + ≡ (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ c ∣ₕ) main north a b = lem ptS a b i0 main south a b = lem ptS a b i1 main (merid c i) a b = lem c a b i @@ -335,7 +352,7 @@ module _ {G : AbGroup ℓ} where r : _ r x = - cong (σ-EM' zero) (rUnitₖ 1 x) + cong (σ-EM' zero) (rUnitₖ 1 x) ∙∙ rUnit (σ-EM' zero x) ∙∙ cong (σ-EM' zero x ∙_) (sym (σ-EM'-0ₖ zero)) @@ -353,13 +370,13 @@ module _ {G : AbGroup ℓ} where (λ x → cong (σ-EM' (suc n)) (lUnitₖ (suc (suc n)) ∣ x ∣) ∙∙ lUnit (σ-EM' (suc n) ∣ x ∣) ∙∙ cong (_∙ σ-EM' (suc n) ∣ x ∣) (sym (σ-EM'-0ₖ (suc n)))) - (λ x → cong (σ-EM' (suc n)) (rUnitₖ (2 + n) ∣ x ∣) + (λ x → cong (σ-EM' (suc n)) (rUnitₖ (2 + n) ∣ x ∣) ∙∙ rUnit (σ-EM' (suc n) ∣ x ∣) ∙∙ cong (σ-EM' (suc n) ∣ x ∣ ∙_) (sym (σ-EM'-0ₖ (suc n)))) (P (σ-EM' (suc n) (0ₖ (2 + n))) (sym (σ-EM'-0ₖ (suc n))))) σ-EM'-ₖ : (n : ℕ) → (a : _) → σ-EM' n (-ₖ a) ≡ sym (σ-EM' n a) - σ-EM'-ₖ n = + σ-EM'-ₖ n = morphLemmas.distrMinus (λ x y → x +[ suc n ]ₖ y) (_∙_) (σ-EM' n) (σ-EM'-hom n) @@ -389,7 +406,9 @@ module _ {G : AbGroup ℓ} where (wedgeConEM.fun G G (suc n) (suc n) (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (hLevHelp n) _ _) (λ x → refl) - (λ x → cong (λ z → -[ (suc (suc n)) ]ₖ z) (rUnitₖ (suc (suc n)) ∣ x ∣ₕ) ∙ sym (rUnitₖ (suc (suc n)) (-[ (suc (suc n)) ]ₖ ∣ x ∣ₕ))) + (λ x → cong (λ z → -[ (suc (suc n)) ]ₖ z) + (rUnitₖ (suc (suc n)) ∣ x ∣ₕ) + ∙ sym (rUnitₖ (suc (suc n)) (-[ (suc (suc n)) ]ₖ ∣ x ∣ₕ))) (rUnit refl)) addIso : (n : ℕ) (x : EM G n) → Iso (EM G n) (EM G n) diff --git a/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda b/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda index e49e2ea88..7c1665958 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda @@ -1,40 +1,39 @@ -{-# OPTIONS --cubical --no-import-sorts --safe --experimental-lossy-unification #-} +{-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.Algebra.Group.EilenbergMacLane.Properties where open import Cubical.Algebra.Group.EilenbergMacLane.Base open import Cubical.Algebra.Group.EilenbergMacLane.WedgeConnectivity open import Cubical.Algebra.Group.EilenbergMacLane.GroupStructure +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Properties +open import Cubical.Algebra.AbGroup.Base + open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) open import Cubical.Foundations.Path open import Cubical.Foundations.HLevels open import Cubical.Foundations.Univalence open import Cubical.Foundations.Pointed open import Cubical.Foundations.Transport +open import Cubical.Foundations.Pointed.Homogeneous + +open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Loopspace -open import Cubical.Data.Unit +open import Cubical.Functions.Morphism + open import Cubical.Data.Sigma -open import Cubical.Algebra.Group.Base -open import Cubical.Algebra.Group.Properties -open import Cubical.Homotopy.Connected -open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec; elim to trElim) +open import Cubical.Data.Nat hiding (_·_) + +open import Cubical.HITs.Truncation as Trunc + renaming (rec to trRec; elim to trElim) open import Cubical.HITs.EilenbergMacLane1 -open import Cubical.Algebra.AbGroup.Base -open import Cubical.Data.Empty - renaming (rec to ⊥-rec) open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) -open import Cubical.Data.Nat hiding (_·_) open import Cubical.HITs.Susp -open import Cubical.Functions.Morphism -open import Cubical.Foundations.Path - -open import Cubical.Foundations.Pointed.Homogeneous private variable ℓ ℓ' ℓ'' : Level @@ -102,6 +101,8 @@ module _ (Ĝ : Group ℓ) where isGroupoidEM₁ : isGroupoid (EM₁ Ĝ) isGroupoidEM₁ = emsquash + --------- Ω (EM₁ G) ≃ G --------- + {- since we write composition in diagrammatic order, and function composition in the other order, we need right multiplication here -} @@ -177,6 +178,7 @@ module _ (Ĝ : Group ℓ) where ΩEM₁≡ : (Path (EM₁ Ĝ) embase embase) ≡ G ΩEM₁≡ = isoToPath ΩEM₁Iso +--------- Ω (EMₙ₊₁ G) ≃ EMₙ G --------- module _ {G : AbGroup ℓ} where open AbGroupStr (snd G) renaming (_+_ to _+G_ ; -_ to -G_ ; assoc to assocG) @@ -189,7 +191,9 @@ module _ {G : AbGroup ℓ} where ; (merid a i) → fib n a i} where fib : (n : ℕ) → (a : EM-raw G (suc n)) - → Path (TypeOfHLevel _ (3 + n)) (EM G (suc n) , hLevelEM G (suc n)) (EM G (suc n) , hLevelEM G (suc n)) + → Path (TypeOfHLevel _ (3 + n)) + (EM G (suc n) , hLevelEM G (suc n)) + (EM G (suc n) , hLevelEM G (suc n)) fib zero a = Σ≡Prop (λ _ → isPropIsOfHLevel 3) (isoToPath (addIso 1 a)) fib (suc n) a = Σ≡Prop (λ _ → isPropIsOfHLevel (4 + n)) @@ -229,7 +233,8 @@ module _ {G : AbGroup ℓ} where ∙ sym (rUnit _))) lem : (n : ℕ) (x a : EM-raw G (suc n)) - → f n (transport (sym (cong (λ x → CODE n x .fst) (cong ∣_∣ₕ (merid a)))) (EM-raw→EM G (suc n) x)) + → f n (transport (sym (cong (λ x → CODE n x .fst) (cong ∣_∣ₕ (merid a)))) + (EM-raw→EM G (suc n) x)) ≡ cong ∣_∣ₕ (σ-EM n x) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) lem zero x a = (λ i → cong ∣_∣ₕ (merid (transportRefl x i -[ 1 ]ₖ a) ∙ sym (merid embase))) ∙∙ σ-EM'-hom zero x (-ₖ a) @@ -254,7 +259,8 @@ module _ {G : AbGroup ℓ} where encode' : (n : ℕ) (x : EM G (suc (suc n))) → 0ₖ (suc (suc n)) ≡ x → CODE n x .fst encode' n x p = subst (λ x → CODE n x .fst) p (0ₖ (suc n)) - decode'-encode' : (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → decode' n x (encode' n x p) ≡ p + decode'-encode' : (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) + → decode' n x (encode' n x p) ≡ p decode'-encode' zero x = J (λ x p → decode' 0 x (encode' 0 x p) ≡ p) (σ-EM'-0ₖ 0) @@ -262,7 +268,8 @@ module _ {G : AbGroup ℓ} where J (λ x p → decode' (suc n) x (encode' (suc n) x p) ≡ p) (σ-EM'-0ₖ (suc n)) - encode'-decode' : (n : ℕ) (x : _) → encode' n (0ₖ (suc (suc n))) (decode' n (0ₖ (suc (suc n))) x) ≡ x + encode'-decode' : (n : ℕ) (x : _) + → encode' n (0ₖ (suc (suc n))) (decode' n (0ₖ (suc (suc n))) x) ≡ x encode'-decode' zero x = cong (encode' zero (0ₖ 2)) (cong-∙ ∣_∣ₕ (merid x) (sym (merid embase))) ∙∙ substComposite (λ x → CODE zero x .fst) @@ -290,6 +297,10 @@ module _ {G : AbGroup ℓ} where Iso.rightInv (Iso-EM-ΩEM+1 (suc (suc n))) = decode'-encode' (suc n) (0ₖ (3 + n)) Iso.leftInv (Iso-EM-ΩEM+1 (suc (suc n))) = encode'-decode' (suc n) + EM≃ΩEM+1 : (n : ℕ) → EM G n ≃ typ (Ω (EM∙ G (suc n))) + EM≃ΩEM+1 n = isoToEquiv (Iso-EM-ΩEM+1 n) + + -- Some properties of the isomorphism EM→ΩEM+1 : (n : ℕ) → EM G n → typ (Ω (EM∙ G (suc n))) EM→ΩEM+1 n = Iso.fun (Iso-EM-ΩEM+1 n) @@ -314,17 +325,11 @@ module _ {G : AbGroup ℓ} where (EM→ΩEM+1-hom n) (ΩEM+1→EM n) (Iso.rightInv (Iso-EM-ΩEM+1 n)) (Iso.leftInv (Iso-EM-ΩEM+1 n)) - - ΩEM+1→EM-refl : (n : ℕ) → ΩEM+1→EM n refl ≡ 0ₖ n ΩEM+1→EM-refl zero = transportRefl 0g ΩEM+1→EM-refl (suc zero) = refl ΩEM+1→EM-refl (suc (suc n)) = refl - - EM≃ΩEM+1 : (n : ℕ) → EM G n ≃ typ (Ω (EM∙ G (suc n))) - EM≃ΩEM+1 n = isoToEquiv (Iso-EM-ΩEM+1 n) - EM≃ΩEM+1∙ : (n : ℕ) → EM∙ G n ≡ Ω (EM∙ G (suc n)) EM≃ΩEM+1∙ n = ua∙ (EM≃ΩEM+1 n) (EM→ΩEM+1-0ₖ n) @@ -333,7 +338,8 @@ module _ {G : AbGroup ℓ} where ua∙ (isoToEquiv (addIso n x)) (lUnitₖ n x) - +-- Some HLevel lemmas about function spaces (EM∙ G n →∙ EM∙ H m), mainly used for +-- the cup product module _ where open AbGroupStr renaming (_+_ to comp) @@ -350,15 +356,18 @@ module _ where x i snd (snd (isContr-↓∙ (suc n)) f i) j = snd f (~ i ∨ j) - isContr-↓∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} (n : ℕ) → isContr ((EM-raw G (suc n) , ptS) →∙ EM∙ H n) + isContr-↓∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} (n : ℕ) + → isContr ((EM-raw G (suc n) , ptS) →∙ EM∙ H n) isContr-↓∙' zero = isContr-↓∙ zero fst (isContr-↓∙' (suc n)) = (λ _ → 0ₖ (suc n)) , refl fst (snd (isContr-↓∙' {H = H} (suc n)) f i) x = - (EM-raw-elim _ _ {A = λ x → 0ₖ (suc n) ≡ fst f x} (λ _ → hLevelEM H (suc n) _ _) (sym (snd f))) x i + EM-raw-elim _ _ {A = λ x → 0ₖ (suc n) ≡ fst f x} + (λ _ → hLevelEM H (suc n) _ _) (sym (snd f)) x i snd (snd (isContr-↓∙' (suc n)) f i) j = snd f (~ i ∨ j) isOfHLevel→∙EM : ∀ {ℓ} {A : Pointed ℓ} {G : AbGroup ℓ'} (n m : ℕ) - → isOfHLevel (suc m) (A →∙ EM∙ G n) → isOfHLevel (suc (suc m)) (A →∙ EM∙ G (suc n)) + → isOfHLevel (suc m) (A →∙ EM∙ G n) + → isOfHLevel (suc (suc m)) (A →∙ EM∙ G (suc n)) isOfHLevel→∙EM {A = A} {G = G} n m hlev = step₃ where step₁ : isOfHLevel (suc m) (A →∙ Ω (EM∙ G (suc n))) @@ -387,20 +396,21 @@ module _ where isOfHLevel↑∙' zero m = isOfHLevel→∙EM m 0 (isContr→isProp (isContr-↓∙' m)) isOfHLevel↑∙' (suc n) m = isOfHLevel→∙EM (suc (n + m)) (suc n) (isOfHLevel↑∙' n m) - →∙EMPath : ∀ {ℓ} {G : AbGroup ℓ} (A : Pointed ℓ') (n : ℕ) → Ω (A →∙ EM∙ G (suc n) ∙) ≡ (A →∙ EM∙ G n ∙) + →∙EMPath : ∀ {ℓ} {G : AbGroup ℓ} (A : Pointed ℓ') (n : ℕ) + → Ω (A →∙ EM∙ G (suc n) ∙) ≡ (A →∙ EM∙ G n ∙) →∙EMPath {G = G} A n = ua∙ (isoToEquiv (ΩfunExtIso A (EM∙ G (suc n)))) refl ∙ (λ i → (A →∙ EM≃ΩEM+1∙ {G = G} n (~ i) ∙)) contr∙-lem : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) → isContr (EM∙ G (suc n) →∙ (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) - fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl), refl + fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl), refl snd (contr∙-lem {G = G} {H = H} {L = L} n m) (f , p) = →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) (sym (funExt (help n f p))) where help : (n : ℕ) → (f : EM G (suc n) → EM∙ H (suc m) →∙ EM∙ L (suc (n + m))) - → f (snd (EM∙ G (suc n))) ≡ snd (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) + → f (snd (EM∙ G (suc n))) ≡ snd (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) help zero f p = EM-raw-elim G zero @@ -428,13 +438,15 @@ module _ where →∙ (Ω (EM∙ H (suc m) →∙ EM∙ L (suc (suc (n + m))) ∙))) h = - subst isProp (λ i → EM∙ G (suc n) →∙ (→∙EMPath {G = L} (EM∙ H (suc m)) (suc (n + m)) (~ i))) + subst isProp + (λ i → EM∙ G (suc n) + →∙ (→∙EMPath {G = L} (EM∙ H (suc m)) (suc (n + m)) (~ i))) (isContr→isProp (contr∙-lem n m)) isOfHLevel↑∙∙ {G = G} {H = H} {L = L} n m (suc l) = isOfHLevelΩ→isOfHLevel (suc l) λ f → - subst - (isOfHLevel (2 + l)) (cong (λ x → typ (Ω x)) + subst (isOfHLevel (2 + l)) + (cong (λ x → typ (Ω x)) (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousEM _)) f)) (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) h) where @@ -444,5 +456,6 @@ module _ where →∙ EM∙ L (suc (suc (suc (l + n + m)))) ∙))) h = subst (isOfHLevel (2 + l)) - (λ i → EM∙ G (suc n) →∙ →∙EMPath {G = L} (EM∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) + (λ i → EM∙ G (suc n) + →∙ →∙EMPath {G = L} (EM∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) (isOfHLevel↑∙∙ n m l) diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 8d5731cf4..0c5ec87bb 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -517,3 +517,13 @@ compPathR→PathP∙∙ {p = p} {q = q} {r = r} {s = s} P j i = ; (j = i0) → r i ; (j = i1) → doubleCompPath-filler p s (sym q) (~ k) i}) (P j i) + +-- useful lemma... TODO: move +move4 : ∀ {ℓ} {A : Type ℓ} (x y z w : A) (_+_ : A → A → A) + → ((x y z : A) → x + (y + z) ≡ (x + y) + z) + → ((x y : A) → x + y ≡ y + x) + → (x + y) + (z + w) ≡ ((x + z) + (y + w)) +move4 x y z w _+_ assoc comm = + sym (assoc x y (z + w)) + ∙∙ cong (x +_) (assoc y z w ∙∙ cong (_+ w) (comm y z) ∙∙ sym (assoc z y w)) + ∙∙ assoc x z (y + w) diff --git a/Cubical/HITs/SmashProduct/Base.agda b/Cubical/HITs/SmashProduct/Base.agda index 86baa52d9..7dfe128f7 100644 --- a/Cubical/HITs/SmashProduct/Base.agda +++ b/Cubical/HITs/SmashProduct/Base.agda @@ -244,282 +244,3 @@ SmashAssociate⁻ = rearrange ∘ comm ∘ Smash-map (idfun∙ _) (comm , refl) ⋀-associate⁻ : A ⋀ (B ⋀∙ C) → (A ⋀∙ B) ⋀ C ⋀-associate⁻ = (SmashPtProj→⋀∙ ⋀→ idfun∙ _) ∘ Smash→⋀ ∘ SmashAssociate⁻ ∘ ⋀→Smash ∘ (idfun∙ _ ⋀→ ⋀∙→SmashPtProj) - - -data _⋀'_ {ℓ ℓ' : Level} (A : Pointed ℓ) (B : Pointed ℓ') : Type (ℓ-max ℓ ℓ') where - incl : typ A → typ B → A ⋀' B - pathl : (a : typ A) → incl a (pt B) ≡ incl (pt A) (pt B) - pathr : (b : typ B) → incl (pt A) b ≡ incl (pt A) (pt B) - pathlCoh : pathl (pt A) ≡ refl - pathrCoh : pathr (pt B) ≡ refl - -module _ {ℓ'' : Level} (P P' : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ')) - (d : ∀ {ℓ ℓ' ℓ''} (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') → Iso (typ (P (P A B) C)) (typ (P A (P B C)))) - (d' : ∀ {ℓ ℓ' ℓ''} (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') → Iso (typ (P' (P' A B) C)) (typ (P' A (P' B C)))) - (∧-map : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → typ (P A B) → typ (P' A B)) - (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') - (is1 : Iso (typ (P A ((Pushout {A = typ (P B C)} ∧-map λ _ → tt) , inl (pt (P' B C))))) - (typ (P ((Pushout {A = typ (P A B)} ∧-map λ _ → tt) , inl (pt (P' A B))) C))) - (is2 : Iso (typ (P' A ((Pushout {A = typ (P B C)} ∧-map λ _ → tt) , inl (pt (P' B C))))) - (typ (P' ((Pushout {A = typ (P A B)} ∧-map λ _ → tt) , inl (pt (P' A B))) C))) - - (coherence1 : (a : _) → Iso.fun is2 (∧-map a) ≡ ∧-map (Iso.fun is1 a)) - (coherence2 : (a : _) → Iso.inv is2 (∧-map a) ≡ ∧-map (Iso.inv is1 a)) - where - - _⊚_ : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') → Type _ - A ⊚ B = Pushout {A = typ (P A B)} ∧-map λ _ → tt - - typL = A ⊚ ((B ⊚ C) , inl (pt (P' B C))) - typR = ((A ⊚ B) , inl (pt (P' A B))) ⊚ C - - f : typL → typR - f (inl x) = inl (Iso.fun is2 x) -- inl (Iso.fun is1 x ) - f (inr x) = inr x - f (push a i) = s i - where - s : inl (Iso.fun is2 (∧-map a)) ≡ inr tt - s = cong inl (coherence1 a) ∙ push (Iso.fun is1 a) - - f⁻ : typR → typL - f⁻ (inl x) = inl (Iso.inv is2 x) - f⁻ (inr x) = inr x - f⁻ (push a i) = s i - where - s : inl (Iso.inv is2 (∧-map a)) ≡ inr tt - s = cong inl (coherence2 a) - ∙ push (Iso.inv is1 a) - - module _ (pp1 : (a : _) → PathP (λ i → inl (Iso.rightInv is2 (∧-map a) i) ≡ inr tt) - (cong f (cong inl (coherence2 a)) ∙ cong f (push (Iso.inv is1 a))) (push a)) - - (pp2 : (a : _) → PathP (λ i → inl (Iso.leftInv is2 (∧-map a) i) ≡ inr tt) - (cong f⁻ (cong inl (coherence1 a)) ∙ cong f⁻ (push (Iso.fun is1 a))) - (push a)) - where - f→f⁻→f : (a : _) → f (f⁻ a) ≡ a - f→f⁻→f (inl x) = cong inl (Iso.rightInv is2 x) - f→f⁻→f (inr x) = refl - f→f⁻→f (push a i) j = - hcomp (λ k → λ { (i = i0) → inl (Iso.rightInv is2 (∧-map a) j) - ; (i = i1) → inr tt - ; (j = i0) → cong-∙ f (cong inl (coherence2 a)) (push (Iso.inv is1 a)) (~ k) i - ; (j = i1) → push a i}) - (pp1 a j i) - - f⁻→f→f⁻ : (a : _) → f⁻ (f a) ≡ a - f⁻→f→f⁻ (inl x) = cong inl (Iso.leftInv is2 x) - f⁻→f→f⁻ (inr x) = refl - f⁻→f→f⁻ (push a i) j = - hcomp (λ k → λ { (i = i0) → inl (Iso.leftInv is2 (∧-map a) j) - ; (i = i1) → inr tt - ; (j = i0) → cong-∙ f⁻ (cong inl (coherence1 a)) (push (Iso.fun is1 a)) (~ k) i - ; (j = i1) → push a i}) - (pp2 a j i) - - mainIso : Iso typL typR - Iso.fun mainIso = f - Iso.inv mainIso = f⁻ - Iso.rightInv mainIso = f→f⁻→f - Iso.leftInv mainIso = f⁻→f→f⁻ - - -wedge : ∀ {ℓ ℓ'} → {A : Pointed ℓ} → {B : Pointed ℓ'} → A ⋁ B → fst A × fst B -wedge {B = B} (inl x) = x , pt B -wedge {A = A} (inr x) = (pt A) , x -wedge {A = A} {B = B} (push a i) = (pt A) , (pt B) - -_Sm_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') -A Sm B = Pushout {A = A ⋁ B} wedge λ _ → tt - -SmashIso : ∀ {ℓ ℓ' ℓ''} (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') - → Iso (A Sm (B Sm C , inl (pt B , pt C))) ((A Sm B , inl (pt A , pt B)) Sm C) -SmashIso A B C = - mainIso (λ A B → A ⋁ B , inl (pt A)) - (λ A B → (typ A × typ B) , ((pt A) , (pt B))) - ∨-assoc - (λ _ _ _ → invIso ×Assoc) - wedge - A - B - C - is1 - {!!} - {!!} - {!!} - {!!} - {!!} - where - is1 : Iso _ _ - Iso.fun is1 (inl x) = {!!} - Iso.fun is1 (inr x) = {!!} - Iso.fun is1 (push a i) = {!!} - Iso.inv is1 = {!!} - Iso.rightInv is1 = {!!} - Iso.leftInv is1 = {!!} - ×Assoc : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso (A × (B × C)) ((A × B) × C) - Iso.fun ×Assoc (x , (y , z)) = (x , y) , z - Iso.inv ×Assoc ((x , y) , z) = x , (y , z) - Iso.rightInv ×Assoc ((x , x₂) , x₁) = refl - Iso.leftInv ×Assoc (x , (y , z)) = refl - - FF : ∀ {ℓ₁ ℓ''' ℓ'''' : Level} - {A : Pointed ℓ₁} {B : Pointed ℓ'''} {C : Pointed ℓ''''} - → ((A ⋁ B) , inl (pt A)) ⋁ C → - A ⋁ ((B ⋁ C) , inl (pt B)) - FF (inl (inl x)) = inl x - FF (inl (inr x)) = inr (inl x) - FF (inl (push a i)) = push a i - FF (inr x) = inr (inr x) - FF (push a i) = (push a ∙ cong inr (push tt)) i - - FF⁻ : ∀ {ℓ₁ ℓ''' ℓ'''' : Level} - {A : Pointed ℓ₁} {B : Pointed ℓ'''} {C : Pointed ℓ''''} - → A ⋁ ((B ⋁ C) , inl (pt B)) - → ((A ⋁ B) , inl (pt A)) ⋁ C - FF⁻ (inl x) = inl (inl x) - FF⁻ (inr (inl x)) = inl (inr x) - FF⁻ (inr (inr x)) = inr x - FF⁻ (inr (push a i)) = (cong inl (sym (push a)) ∙ push a) i - FF⁻ (push a i) = inl (push a i) - - FF→FF⁻→FF : ∀ {ℓ₁ ℓ''' ℓ'''' : Level} - {A : Pointed ℓ₁} {B : Pointed ℓ'''} {C : Pointed ℓ''''} - (x : A ⋁ ((B ⋁ C) , inl (pt B))) → FF (FF⁻ x) ≡ x - FF→FF⁻→FF (inl x) = refl - FF→FF⁻→FF (inr (inl x)) = refl - FF→FF⁻→FF (inr (inr x)) = refl - FF→FF⁻→FF (inr (push a i)) k = h k i - where - h : cong FF ((cong inl (sym (push a)) ∙ push a)) ≡ cong inr (push a) - h = cong-∙ FF (cong inl (sym (push a))) (push a) - ∙∙ assoc _ _ _ - ∙∙ cong (_∙ cong inr (push a)) (lCancel _) - ∙ sym (lUnit _) - FF→FF⁻→FF (push a i) = refl - - FF⁻→FF→FF⁻ : ∀ {ℓ₁ ℓ''' ℓ'''' : Level} - {A : Pointed ℓ₁} {B : Pointed ℓ'''} {C : Pointed ℓ''''} - (x : ((A ⋁ B) , inl (pt A)) ⋁ C) → FF⁻ (FF x) ≡ x - FF⁻→FF→FF⁻ (inl (inl x)) = refl - FF⁻→FF→FF⁻ (inl (inr x)) = refl - FF⁻→FF→FF⁻ (inl (push a i)) = refl - FF⁻→FF→FF⁻ (inr x) = refl - FF⁻→FF→FF⁻ (push a i) k = h k i - where - h : cong FF⁻ (push a ∙ cong inr (push tt)) ≡ push a - h = cong-∙ FF⁻ (push a) (cong inr (push a)) - ∙∙ assoc _ _ _ - ∙∙ cong (_∙ push a) (rCancel _) - ∙ sym (lUnit _) - - ∨-assoc : ∀ {ℓ₁ ℓ''' ℓ'''' : Level} - (A : Pointed ℓ₁) (B : Pointed ℓ''') (C : Pointed ℓ'''') → - Iso (((A ⋁ B) , inl (pt A)) ⋁ C) - (A ⋁ ((B ⋁ C) , inl (pt B))) - Iso.fun (∨-assoc A B C) = FF - Iso.inv (∨-assoc A B C) = FF⁻ - Iso.rightInv (∨-assoc A B C) = FF→FF⁻→FF - Iso.leftInv (∨-assoc A B C) = FF⁻→FF→FF⁻ - - {- - i = i0 ⊢ cong f (cong inl (coherence2 a) ∙ push (Iso.inv is1 a)) j -i = i1 ⊢ push a j -j = i0 ⊢ inl (Iso.rightInv is2 (∧-map a) i) -j = i1 ⊢ inr tt - -} - - --- s : Iso (A ⊚ ((B ⊚ C) , inl tt)) (((A ⊚ B) , inl tt) ⊚ C) --- Iso.fun s = {!!} --- Iso.inv s = {!!} --- Iso.rightInv s = {!!} --- Iso.leftInv s = {!!} - - --- -- module _ {ℓ'' : Level} (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') where --- -- assocFun : ((A ⋀' B) , (incl (pt A) (pt B))) ⋀' C --- -- → A ⋀' (B ⋀' C , incl (pt B) (pt C)) --- -- assocFun (incl (incl x x₁) y) = incl x (incl x₁ y) --- -- assocFun (incl (pathl a i) y) = (cong (incl a) (pathr y) ∙∙ pathl a ∙∙ sym (pathr (incl (pt B) y))) i --- -- assocFun (incl (pathr b i) y) = (pathr (incl b y) ∙ cong (incl (pt A)) (sym (pathr y))) i --- -- assocFun (incl (pathlCoh i j) y) = {!!} --- -- assocFun (incl (pathrCoh i j) y) = {!!} --- -- assocFun (pathl (incl x x₁) i) = {!!} --- -- assocFun (pathl (pathl a i₁) i) = {!!} --- -- assocFun (pathl (pathr b i₁) i) = {!!} --- -- assocFun (pathl (pathlCoh i₁ i₂) i) = {!!} --- -- assocFun (pathl (pathrCoh i₁ i₂) i) = {!!} --- -- assocFun (pathr b i) = --- -- {!Goal: A ⋀' ((B ⋀' C) , incl (pt B) (pt C)) --- -- ———— Boundary —————————————————————————————————————————————— --- -- i = i0 ⊢ incl (pt A) (incl b y) --- -- i = i1 ⊢ incl (pt A) (incl (pt B) y)!} --- -- assocFun (pathlCoh i i₁) = {!!} --- -- assocFun (pathrCoh i i₁) = {!!} - --- -- module _ (A : Pointed ℓ) (B : Pointed ℓ') where --- -- smashFun : (A ⋀' B) → (Smash A B) --- -- smashFun (incl x y) = proj x y --- -- smashFun (pathl a i) = pivotr a (pt A) i --- -- smashFun (pathr b i) = pivotl b (pt B) i --- -- smashFun (pathlCoh i j) = rCancel (gluel (pt A)) i j --- -- smashFun (pathrCoh i j) = rCancel (gluer (pt B)) i j - --- -- smashFun⁻ : (Smash A B) → A ⋀' B --- -- smashFun⁻ basel = incl (pt A) (pt B) --- -- smashFun⁻ baser = incl (pt A) (pt B) --- -- smashFun⁻ (proj x y) = incl x y --- -- smashFun⁻ (gluel a i) = pathl a i --- -- smashFun⁻ (gluer b i) = pathr b i - --- -- sectSmashFun : section smashFun smashFun⁻ --- -- sectSmashFun basel = gluel (pt A) --- -- sectSmashFun baser = gluer (pt B) --- -- sectSmashFun (proj x y) = refl --- -- sectSmashFun (gluel a i) j = compPath-filler (gluel a) (sym (gluel (pt A))) (~ j) i --- -- sectSmashFun (gluer b i) j = compPath-filler (gluer b) (sym (gluer (pt B))) (~ j) i - --- -- retrSmashFun : retract smashFun smashFun⁻ --- -- retrSmashFun (incl x x₁) = refl --- -- retrSmashFun (pathl a i) k = help k i --- -- module _ where --- -- help : cong smashFun⁻ (cong smashFun (pathl a)) ≡ pathl a --- -- help = congFunct smashFun⁻ (gluel a) (sym (gluel (pt A))) ∙ cong (cong smashFun⁻ (gluel a) ∙_) (cong sym pathlCoh) ∙ sym (rUnit _) --- -- retrSmashFun (pathr b i) k = help2 k i --- -- module _ where --- -- help2 : cong smashFun⁻ (cong smashFun (pathr b)) ≡ pathr b --- -- help2 = congFunct smashFun⁻ (gluer b) (sym (gluer (pt B))) --- -- ∙ cong (cong smashFun⁻ (gluer b) ∙_) (cong sym pathrCoh) --- -- ∙ sym (rUnit _) --- -- retrSmashFun (pathlCoh i j) k = main i j k --- -- where - --- -- h = help (pt A) i0 i0 --- -- main : PathP (λ i → PathP (λ j → smashFun⁻ (smashFun (pathlCoh i j)) ≡ pathlCoh i j) --- -- refl refl) --- -- (λ j k → h k j) refl --- -- main i j k = --- -- hcomp (λ r → λ { (i = i0) → compPath-filler' (congFunct smashFun⁻ (gluel (pt A)) (sym (gluel (pt A)))) --- -- (cong (cong smashFun⁻ (gluel (pt A)) ∙_) (cong sym pathlCoh) ∙ sym (rUnit _)) r k j --- -- ; (i = i1) → incl (pt A) (pt B) --- -- ; (j = i0) → incl (pt A) (pt B) --- -- ; (j = i1) → incl (pt A) (pt B) --- -- ; (k = i0) → rCancelPathP (gluel (pt A)) smashFun⁻ (~ r) i j --- -- ; (k = i1) → pathlCoh i j}) --- -- (mainGen (pathl (pt A)) (sym (pathlCoh)) i j k) --- -- retrSmashFun (pathrCoh i j) k = --- -- hcomp (λ r → λ {(i = i0) → compPath-filler' (congFunct smashFun⁻ (gluer (pt B)) (sym (gluer (pt B)))) --- -- (cong (cong smashFun⁻ (gluer (pt B)) ∙_) (cong sym pathrCoh) ∙ sym (rUnit _)) r k j --- -- ; (i = i1) → incl (pt A) (pt B) --- -- ; (j = i0) → incl (pt A) (pt B) --- -- ; (j = i1) → incl (pt A) (pt B) --- -- ; (k = i0) → rCancelPathP (gluer (pt B)) smashFun⁻ (~ r) i j --- -- ; (k = i1) → pathrCoh i j}) --- -- (mainGen (pathr (pt B)) (sym (pathrCoh)) i j k) - --- -- smashIso : Iso (A ⋀' B) (Smash A B) --- -- Iso.fun smashIso = smashFun --- -- Iso.inv smashIso = smashFun⁻ --- -- Iso.rightInv smashIso = sectSmashFun --- -- Iso.leftInv smashIso = retrSmashFun From 42202ea3c1f6fe98b4711c0221a77445b5ceb1e8 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 20 Aug 2021 14:07:38 +0200 Subject: [PATCH 22/30] cleanup --- Cubical/Algebra/AbGroup/TensorProduct.agda | 17 +- .../Group/EilenbergMacLane/CupProduct.agda | 271 ++--- .../EilenbergMacLane/WedgeConnectivity.agda | 142 ++- Cubical/Algebra/Group/EilenbergMacLane1.agda | 995 ------------------ .../ZCohomology/RingStructure/CupProduct.agda | 3 + .../RingStructure/GradedCommutativity.agda | 4 - .../ZCohomology/RingStructure/RingLaws.agda | 9 +- 7 files changed, 226 insertions(+), 1215 deletions(-) delete mode 100644 Cubical/Algebra/Group/EilenbergMacLane1.agda diff --git a/Cubical/Algebra/AbGroup/TensorProduct.agda b/Cubical/Algebra/AbGroup/TensorProduct.agda index 0e5417616..48e494438 100644 --- a/Cubical/Algebra/AbGroup/TensorProduct.agda +++ b/Cubical/Algebra/AbGroup/TensorProduct.agda @@ -161,25 +161,25 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where (λ i → (fst x ⊗ snd x) +⊗ (listify++ x₁ y i)) ∙ ⊗assoc (fst x ⊗ snd x) (listify x₁) (listify y) - slick : (x : AGr ⨂₁ BGr) → ∃[ l ∈ List (A × B) ] (listify l ≡ x) - slick = + ∃List : (x : AGr ⨂₁ BGr) → ∃[ l ∈ List (A × B) ] (listify l ≡ x) + ∃List = ⊗elimProp (λ _ → squash) (λ a b → ∣ [ a , b ] , ⊗rUnit (a ⊗ b) ∣) λ x y → rec2 squash λ {(l1 , p) (l2 , q) → ∣ (l1 ++ l2) , listify++ l1 l2 ∙ cong₂ _+⊗_ p q ∣} - ⊗elimPropCool : ∀ {ℓ} {C : AGr ⨂₁ BGr → Type ℓ} + ⊗elimPropListify : ∀ {ℓ} {C : AGr ⨂₁ BGr → Type ℓ} → ((x : _) → isProp (C x)) → ((x : _) → C (listify x)) → (x : _) → C x - ⊗elimPropCool {C = C} p f = + ⊗elimPropListify {C = C} p f = ⊗elimProp p (λ x y → subst C (⊗rUnit (x ⊗ y)) (f [ x , y ])) λ x y → pRec (isPropΠ2 λ _ _ → p _) (pRec (isPropΠ3 λ _ _ _ → p _) (λ {(l1 , p) (l2 , q) ind1 ind2 → subst C (listify++ l2 l1 ∙ cong₂ _+⊗_ q p) (f (l2 ++ l1))}) - (slick y)) - (slick x) + (∃List y)) + (∃List x) ----------------------------------------------------------------------------------- lCancelPrim : (x : B) → (0A ⊗ x) ≡ 0⊗ @@ -203,7 +203,7 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where ⊗rCancel : (x : AGr ⨂₁ BGr) → (x +⊗ -⊗ x) ≡ 0⊗ ⊗rCancel = - ⊗elimPropCool (λ _ → ⊗squash _ _) h + ⊗elimPropListify (λ _ → ⊗squash _ _) h where h : (x : List (A × B)) → (listify x +⊗ -⊗ (listify x)) ≡ 0⊗ h [] = sym (linl 0A (-A 0A) (0B)) @@ -349,8 +349,7 @@ module UP (AGr : AbGroup ℓ) (BGr : AbGroup ℓ') where isTensorProduct⨂ : (isTensorProductOf AGr and BGr) (AGr ⨂ BGr) fst isTensorProduct⨂ = mainF - snd isTensorProduct⨂ C = - isoToIsEquiv mainIso + snd isTensorProduct⨂ C = isoToIsEquiv mainIso where invF : GroupHom (AGr *) (HomGroup (BGr *) C *) → GroupHom ((AGr ⨂ BGr) *) (C *) fst (invF (f , p)) = F diff --git a/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda b/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda index 986b75964..6c3d3e3c8 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --safe --experimental-lossy-unification --no-forcing #-} +{-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.Algebra.Group.EilenbergMacLane.CupProduct where @@ -6,25 +6,24 @@ open import Cubical.Algebra.Group.EilenbergMacLane.Base open import Cubical.Algebra.Group.EilenbergMacLane.WedgeConnectivity open import Cubical.Algebra.Group.EilenbergMacLane.GroupStructure open import Cubical.Algebra.Group.EilenbergMacLane.Properties +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Properties + open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) open import Cubical.Foundations.Path open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Univalence open import Cubical.Foundations.Pointed open import Cubical.Foundations.Transport -open import Cubical.Homotopy.Loopspace -open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Foundations.Pointed.Homogeneous + +open import Cubical.Functions.Morphism + +open import Cubical.Homotopy.Loopspace -open import Cubical.Data.Unit -open import Cubical.Data.Sigma -open import Cubical.Algebra.Group.Base -open import Cubical.Algebra.Group.Properties -open import Cubical.Homotopy.Connected open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec; elim to trElim) open import Cubical.HITs.EilenbergMacLane1 renaming (rec to EMrec) open import Cubical.Algebra.AbGroup.Base @@ -34,64 +33,50 @@ open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) open import Cubical.Data.Nat hiding (_·_) open import Cubical.HITs.Susp -open import Cubical.Functions.Morphism -open import Cubical.Foundations.Path open import Cubical.Algebra.AbGroup.TensorProduct open import Cubical.Algebra.Group -open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.ZCohomology.RingStructure.CupProduct + using (_+'_ ; +'≡+ ; +'-comm) +open AbGroupStr renaming (_+_ to _+Gr_ ; -_ to -Gr_) private variable ℓ ℓ' ℓ'' : Level -stupidInd : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 → A 1 - → ((n : ℕ) → (A (suc n) → A (suc (suc n)))) - → (n : ℕ) → A n -stupidInd a0 a1 ind zero = a0 -stupidInd a0 a1 ind (suc zero) = a1 -stupidInd {A = A} a0 a1 ind (suc (suc n)) = - ind n (stupidInd {A = A} a0 a1 ind (suc n)) - -stupidInd' : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 - → ((n : ℕ) → (A n → A (suc n))) - → (n : ℕ) → A n -stupidInd' a0 ind zero = a0 -stupidInd' a0 ind (suc n) = ind n (stupidInd' a0 ind n) - -_+'_ : ℕ → ℕ → ℕ -zero +' m = m -suc n +' zero = suc n -suc n +' suc m = suc (suc (n + m)) - -+'-comm : (n m : ℕ) → n +' m ≡ m +' n -+'-comm zero zero = refl -+'-comm zero (suc m) = refl -+'-comm (suc n) zero = refl -+'-comm (suc n) (suc m) = cong (2 +_) (+-comm n m) - -+'≡+ : (n m : ℕ) → n +' m ≡ n + m -+'≡+ zero zero = refl -+'≡+ zero (suc m) = refl -+'≡+ (suc n) zero = cong suc (+-comm zero n) -+'≡+ (suc n) (suc m) = cong suc (sym (+-suc n m)) - -open AbGroupStr renaming (_+_ to _+Gr_ ; -_ to -Gr_) - +private + -- ℕ induction. For some reason, this helps the termination checker + ind' : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 + → ((n : ℕ) → (A n → A (suc n))) + → (n : ℕ) → A n + ind' a0 ind zero = a0 + ind' a0 ind (suc n) = ind n (ind' a0 ind n) + + ind+2 : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 → A 1 + → ((n : ℕ) → (A (suc n) → A (suc (suc n)))) + → (n : ℕ) → A n + ind+2 a0 a1 ind zero = a0 + ind+2 a0 a1 ind (suc zero) = a1 + ind+2 {A = A} a0 a1 ind (suc (suc n)) = + ind n (ind+2 {A = A} a0 a1 ind (suc n)) + +-- A homomorphism φ : G → H of AbGroups induces a homomorphism +-- φ' : K(G,n) → K(H,n) inducedFun-EM-raw : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} → AbGroupHom G' H' → ∀ n - → EM-raw G' n → EM-raw H' n + → EM-raw G' n → EM-raw H' n inducedFun-EM-raw f = - stupidInd (fst f) + ind+2 (fst f) (EMrec _ emsquash embase (λ g → emloop (fst f g)) λ g h → compPathR→PathP (sym (sym (lUnit _) - ∙∙ cong (_∙ (sym (emloop (fst f h)))) (cong emloop (IsGroupHom.pres· (snd f) g h) - ∙ emloop-comp _ (fst f g) (fst f h)) + ∙∙ cong (_∙ (sym (emloop (fst f h)))) + (cong emloop (IsGroupHom.pres· (snd f) g h) + ∙ emloop-comp _ (fst f g) (fst f h)) ∙∙ sym (∙assoc _ _ _) ∙∙ cong (emloop (fst f g) ∙_) (rCancel _) ∙∙ sym (rUnit _)))) @@ -99,59 +84,17 @@ inducedFun-EM-raw f = ; south → south ; (merid a i) → merid (ind a) i} ) -inducedFun-EM-rawIso : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} - → AbGroupEquiv G' H' - → ∀ n → Iso (EM-raw G' n) (EM-raw H' n) -Iso.fun (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , (snd e)) n -Iso.inv (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , isGroupHomInv e) n -Iso.rightInv (inducedFun-EM-rawIso e n) = h n - where - h : (n : ℕ) → section (inducedFun-EM-raw (fst e .fst , snd e) n) - (inducedFun-EM-raw (invEq (fst e) , isGroupHomInv e) n) - h = stupidInd - (secEq (fst e)) - (elimSet _ (λ _ → emsquash _ _) refl - (λ g → compPathR→PathP - ((sym (cong₂ _∙_ (cong emloop (secEq (fst e) g)) (sym (lUnit _)) - ∙ rCancel _))))) - λ n p → λ { north → refl - ; south → refl - ; (merid a i) k → merid (p a k) i} -Iso.leftInv (inducedFun-EM-rawIso e n) = h n - where - h : (n : ℕ) → retract (Iso.fun (inducedFun-EM-rawIso e n)) - (Iso.inv (inducedFun-EM-rawIso e n)) - h = stupidInd - (retEq (fst e)) - (elimSet _ (λ _ → emsquash _ _) refl - (λ g → compPathR→PathP - ((sym (cong₂ _∙_ (cong emloop (retEq (fst e) g)) (sym (lUnit _)) - ∙ rCancel _))))) - λ n p → λ { north → refl - ; south → refl - ; (merid a i) k → merid (p a k) i} - -Iso→EMIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} - → AbGroupEquiv G H → ∀ n → Iso (EM G n) (EM H n) -Iso→EMIso is zero = inducedFun-EM-rawIso is zero -Iso→EMIso is (suc zero) = inducedFun-EM-rawIso is 1 -Iso→EMIso is (suc (suc n)) = mapCompIso (inducedFun-EM-rawIso is (suc (suc n))) - -EM⊗-commFun : {G : AbGroup ℓ} {H : AbGroup ℓ'} - → ∀ n → Iso (EM (G ⨂ H) n) (EM (H ⨂ G) n) -EM⊗-commFun {G = G} {H = H} = Iso→EMIso (GroupIso→GroupEquiv ⨂-commIso) - -EM⊗-assocIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} - → ∀ n → Iso (EM (G ⨂ (H ⨂ L)) n) (EM ((G ⨂ H) ⨂ L) n) -EM⊗-assocIso = Iso→EMIso (GroupIso→GroupEquiv (GroupEquiv→GroupIso ⨂assoc)) - +-- Lemma for distributativity of cup product (used later) pathType : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → Type ℓ -pathType n x p = sym (rUnitₖ (2 + n) x) ∙ (λ i → x +ₖ p i) ≡ sym (lUnitₖ (2 + n) x) ∙ λ i → p i +ₖ x +pathType n x p = sym (rUnitₖ (2 + n) x) ∙ (λ i → x +ₖ p i) + ≡ sym (lUnitₖ (2 + n) x) ∙ λ i → p i +ₖ x lem : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → pathType n x p lem n x = J (λ x p → pathType n x p) refl + +-- Definition of cup product (⌣ₖ, given by ·₀ when first argument is in K(G,0)) module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where private G = fst G' @@ -171,7 +114,7 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where ·₀' : H → (m : ℕ) → EM G' m → EM (G' ⨂ H') m ·₀' h = - stupidInd + ind+2 (_⊗ h) (elimGroupoid _ (λ _ → emsquash) embase @@ -179,7 +122,8 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where λ g l → compPathR→PathP (sym (∙assoc _ _ _ - ∙∙ cong₂ _∙_ (sym (emloop-comp _ _ _) ∙ cong emloop (sym (linl g l h))) refl + ∙∙ cong₂ _∙_ (sym (emloop-comp _ _ _) + ∙ cong emloop (sym (linl g l h))) refl ∙∙ rCancel _))) λ n f → trRec (isOfHLevelTrunc (4 + n)) λ { north → 0ₖ (suc (suc n)) @@ -188,7 +132,7 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where ·₀ : G → (m : ℕ) → EM H' m → EM (G' ⨂ H') m ·₀ g = - stupidInd (λ h → g ⊗ h) + ind+2 (λ h → g ⊗ h) (elimGroupoid _ (λ _ → emsquash) embase (λ h → emloop (g ⊗ h)) @@ -204,7 +148,7 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where ·₀-distr : (g h : G) → (m : ℕ) (x : EM H' m) → ·₀ (g +G h) m x ≡ ·₀ g m x +ₖ ·₀ h m x ·₀-distr g h = - stupidInd + ind+2 (linl g h) (elimSet _ (λ _ → emsquash _ _) refl @@ -219,13 +163,16 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where ; (merid a i) k → z m ind a k i} where - z : (m : ℕ) → ((x : EM H' (suc m)) → ·₀ (g +G h) (suc m) x ≡ ·₀ g (suc m) x +ₖ ·₀ h (suc m) x) → (a : EM-raw H' (suc m)) + z : (m : ℕ) → ((x : EM H' (suc m)) + → ·₀ (g +G h) (suc m) x + ≡ ·₀ g (suc m) x +ₖ ·₀ h (suc m) x) → (a : EM-raw H' (suc m)) → cong (·₀ (g +G h) (suc (suc m))) (cong ∣_∣ₕ (merid a)) ≡ cong₂ _+ₖ_ (cong (·₀ g (suc (suc m))) (cong ∣_∣ₕ (merid a))) (cong (·₀ h (suc (suc m))) (cong ∣_∣ₕ (merid a))) z m ind a = (λ i → EM→ΩEM+1 _ (ind (EM-raw→EM _ _ a) i)) - ∙∙ EM→ΩEM+1-hom _ (·₀ g (suc m) (EM-raw→EM H' (suc m) a)) (·₀ h (suc m) (EM-raw→EM H' (suc m) a)) + ∙∙ EM→ΩEM+1-hom _ (·₀ g (suc m) (EM-raw→EM H' (suc m) a)) + (·₀ h (suc m) (EM-raw→EM H' (suc m) a)) ∙∙ sym (cong₂+₂ m (cong (·₀ g (suc (suc m))) (cong ∣_∣ₕ (merid a))) (cong (·₀ h (suc (suc m))) (cong ∣_∣ₕ (merid a)))) @@ -241,7 +188,7 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where 0·₀ : (m : ℕ) → (x : _) → ·₀ 0G m x ≡ 0ₖ m 0·₀ = - stupidInd lCancelPrim + ind+2 lCancelPrim (elimSet _ (λ _ → emsquash _ _) refl λ g → compPathR→PathP ((sym (emloop-1g _) @@ -256,7 +203,7 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where 0·₀' : (m : ℕ) (g : _) → ·₀' 0H m g ≡ 0ₖ m 0·₀' = - stupidInd + ind+2 rCancelPrim (elimSet _ (λ _ → emsquash _ _) refl @@ -268,14 +215,16 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where λ { north → refl ; south → refl ; (merid a i) j → (cong (EM→ΩEM+1 (suc n)) (f (EM-raw→EM _ _ a)) - ∙ EM→ΩEM+1-0ₖ _) j i} + ∙ EM→ΩEM+1-0ₖ _) j i} + +-- Definition of the cup product cup∙ : ∀ n m → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m) cup∙ = - stupidInd' + ind' (λ m g → (·₀ g m) , ·₀0 m g) λ n f → - stupidInd' + ind' (λ g → (λ h → ·₀' h (suc n) g) , 0·₀' (suc n) g) λ m _ → main n m f @@ -313,7 +262,8 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where λ { north → (λ _ → 0ₖ (3 + (n + m))) , refl ; south → (λ _ → 0ₖ (3 + (n + m))) , refl ; (merid a i) → (λ x → EM→ΩEM+1 _ (ind _ (EM-raw→EM _ _ a) .fst x) i) - , λ j → (cong (EM→ΩEM+1 (suc (suc (n + m)))) (ind (suc m) (EM-raw→EM G' (suc n) a) .snd) + , λ j → (cong (EM→ΩEM+1 (suc (suc (n + m)))) + (ind (suc m) (EM-raw→EM G' (suc n) a) .snd) ∙ EM→ΩEM+1-0ₖ _) j i} _⌣ₖ_ : {n m : ℕ} (x : EM G' n) (y : EM H' m) → EM (G' ⨂ H') (n +' m) @@ -329,23 +279,27 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where 0ₖ-⌣ₖ (suc zero) (suc m) x = refl 0ₖ-⌣ₖ (suc (suc n)) (suc m) x = refl +module LeftDistributivity {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where private - distrl1 : (n m : ℕ) → EM H' m → EM H' m → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) + distrl1 : (n m : ℕ) → EM H' m → EM H' m + → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) fst (distrl1 n m x y) z = z ⌣ₖ (x +ₖ y) snd (distrl1 n m x y) = 0ₖ-⌣ₖ n m _ - distrl2 : (n m : ℕ) → EM H' m → EM H' m → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) + distrl2 : (n m : ℕ) → EM H' m → EM H' m + → EM∙ G' n →∙ EM∙ (G' ⨂ H') (n +' m) fst (distrl2 n m x y) z = (z ⌣ₖ x) +ₖ (z ⌣ₖ y) - snd (distrl2 n m x y) = cong₂ _+ₖ_ (0ₖ-⌣ₖ n m x) (0ₖ-⌣ₖ n m y) ∙ rUnitₖ _ (0ₖ (n +' m)) + snd (distrl2 n m x y) = + cong₂ _+ₖ_ (0ₖ-⌣ₖ n m x) (0ₖ-⌣ₖ n m y) ∙ rUnitₖ _ (0ₖ (n +' m)) - hLevLem : (n m : ℕ) → isOfHLevel (suc (suc m)) (EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') ((suc n) +' m)) - hLevLem n m = subst (isOfHLevel (suc (suc m))) (λ i → EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') - ((cong suc (+-comm m n) ∙ sym (+'≡+ (suc n) m)) i)) (isOfHLevel↑∙ m n) + hLevLem : (n m : ℕ) → isOfHLevel (suc (suc m)) (EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') ((suc n) +' m)) + hLevLem n m = + subst (isOfHLevel (suc (suc m))) + (λ i → EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') + ((cong suc (+-comm m n) ∙ sym (+'≡+ (suc n) m)) i)) (isOfHLevel↑∙ m n) - - - - mainDistrL : (n m : ℕ) (x y : EM H' (suc m)) → distrl1 (suc n) (suc m) x y ≡ distrl2 (suc n) (suc m) x y + mainDistrL : (n m : ℕ) (x y : EM H' (suc m)) + → distrl1 (suc n) (suc m) x y ≡ distrl2 (suc n) (suc m) x y mainDistrL n zero = wedgeConEM.fun H' H' 0 0 (λ _ _ → hLevLem _ _ _ _) @@ -356,12 +310,14 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where λ i → →∙Homogeneous≡ (isHomogeneousEM (suc (suc (n + 0)))) (funExt (λ z → l≡r z i)) where - l : (x : EM H' 1) (z : _) → (distrl1 (suc n) 1 embase x .fst z) ≡ (distrl2 (suc n) 1 embase x .fst z) + l : (x : EM H' 1) (z : _) + → (distrl1 (suc n) 1 embase x .fst z) ≡ (distrl2 (suc n) 1 embase x .fst z) l x z = cong (z ⌣ₖ_) (lUnitₖ _ x) ∙∙ sym (lUnitₖ _ (z ⌣ₖ x)) ∙∙ λ i → (⌣ₖ-0ₖ (suc n) (suc zero) z (~ i)) +ₖ (z ⌣ₖ x) - r : (z : EM H' 1) (x : EM G' (suc n)) → (distrl1 (suc n) 1 z embase .fst x) ≡ (distrl2 (suc n) 1 z embase .fst x) + r : (z : EM H' 1) (x : EM G' (suc n)) + → (distrl1 (suc n) 1 z embase .fst x) ≡ (distrl2 (suc n) 1 z embase .fst x) r y z = cong (z ⌣ₖ_) (rUnitₖ _ y) ∙∙ sym (rUnitₖ _ (z ⌣ₖ y)) ∙∙ λ i → (z ⌣ₖ y) +ₖ (⌣ₖ-0ₖ (suc n) (suc zero) z (~ i)) @@ -384,13 +340,15 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where (funExt (λ z → l≡r z i))) where l : (x : EM-raw H' (suc (suc m))) (z : EM G' (suc n)) - → (distrl1 (suc n) (suc (suc m)) (0ₖ _) ∣ x ∣ₕ .fst z) ≡ (distrl2 (suc n) (suc (suc m)) (0ₖ _) ∣ x ∣ₕ .fst z) + → (distrl1 (suc n) (suc (suc m)) (0ₖ _) ∣ x ∣ₕ .fst z) + ≡ (distrl2 (suc n) (suc (suc m)) (0ₖ _) ∣ x ∣ₕ .fst z) l x z = cong (z ⌣ₖ_) (lUnitₖ (suc (suc m)) ∣ x ∣) ∙∙ sym (lUnitₖ _ (z ⌣ₖ ∣ x ∣)) ∙∙ λ i → (⌣ₖ-0ₖ (suc n) (suc (suc m)) z (~ i)) +ₖ (z ⌣ₖ ∣ x ∣) r : (x : EM-raw H' (suc (suc m))) (z : EM G' (suc n)) - → (distrl1 (suc n) (suc (suc m)) ∣ x ∣ₕ (0ₖ _) .fst z) ≡ (distrl2 (suc n) (suc (suc m)) ∣ x ∣ₕ (0ₖ _) .fst z) + → (distrl1 (suc n) (suc (suc m)) ∣ x ∣ₕ (0ₖ _) .fst z) + ≡ (distrl2 (suc n) (suc (suc m)) ∣ x ∣ₕ (0ₖ _) .fst z) r x z = cong (z ⌣ₖ_) (rUnitₖ (suc (suc m)) ∣ x ∣) ∙∙ sym (rUnitₖ _ (z ⌣ₖ ∣ x ∣)) ∙∙ λ i → (z ⌣ₖ ∣ x ∣) +ₖ (⌣ₖ-0ₖ (suc n) (suc (suc m)) z (~ i)) @@ -398,7 +356,7 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where l≡r : (z : EM G' (suc n)) → l north z ≡ r north z l≡r z = sym (lem _ _ (sym (⌣ₖ-0ₖ (suc n) (suc (suc m)) z))) -module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where +module RightDistributivity {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where private G = fst G' H = fst H' @@ -424,11 +382,8 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where fst (distrr2 n m x y) z = (x ⌣ₖ z) +ₖ (y ⌣ₖ z) snd (distrr2 n m x y) = cong₂ _+ₖ_ (⌣ₖ-0ₖ n m x) (⌣ₖ-0ₖ n m y) ∙ rUnitₖ _ (0ₖ (n +' m)) - hLevLem2 : (n m : ℕ) → isOfHLevel (suc (suc m)) (EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') ((suc n) +' m)) - hLevLem2 n m = subst (isOfHLevel (suc (suc m))) (λ i → EM∙ G' (suc n) →∙ EM∙ (G' ⨂ H') - ((cong suc (+-comm m n) ∙ sym (+'≡+ (suc n) m)) i)) (isOfHLevel↑∙ m n) - - mainDistrR : (n m : ℕ) (x y : EM G' (suc n)) → distrr1 (suc n) (suc m) x y ≡ distrr2 (suc n) (suc m) x y + mainDistrR : (n m : ℕ) (x y : EM G' (suc n)) + → distrr1 (suc n) (suc m) x y ≡ distrr2 (suc n) (suc m) x y mainDistrR zero m = wedgeConEM.fun G' G' 0 0 (λ _ _ → isOfHLevel↑∙ 1 m _ _) @@ -459,7 +414,12 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where (isOfHLevel↑∙ (2 + n) m) _ _) (wedgeConEM.fun _ _ _ _ (λ x y → isOfHLevelPath ((2 + n) + (2 + n)) - {!subst !} _ _) + (transport (λ i → isOfHLevel (((λ i → (+-comm n 2 (~ i) + (2 + n))) + ∙ sym (+-assoc n 2 (2 + n))) (~ i)) + (EM∙ H' (suc m) →∙ EM∙ ((fst (AbGroupPath (G' ⨂ H') (H' ⨂ G'))) ⨂-comm (~ i)) + ((+'-comm (suc m) (suc (suc n))) i))) + (isOfHLevelPlus n + (LeftDistributivity.hLevLem m (suc (suc n))))) _ _) (λ x → →∙Homogeneous≡ (isHomogeneousEM _) (funExt (l x))) (λ x → →∙Homogeneous≡ (isHomogeneousEM _) @@ -467,7 +427,6 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where λ i → →∙Homogeneous≡ (isHomogeneousEM _) (funExt λ z → r≡l z i)) where - l : (x : _) (z : _) → _ ≡ _ l x z = (λ i → (lUnitₖ _ ∣ x ∣ i) ⌣ₖ z) ∙∙ sym (lUnitₖ _ (∣ x ∣ ⌣ₖ z)) @@ -480,3 +439,53 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where r≡l : (z : _) → l north z ≡ r north z r≡l z = lem _ _ _ + +-- TODO: Summarise distributivity proofs +-- TODO: Associativity and graded commutativity, following Cubical.ZCohomology.RingStructure +-- The following lemmas will be needed to make the types match up. + +inducedFun-EM-rawIso : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} + → AbGroupEquiv G' H' + → ∀ n → Iso (EM-raw G' n) (EM-raw H' n) +Iso.fun (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , (snd e)) n +Iso.inv (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , isGroupHomInv e) n +Iso.rightInv (inducedFun-EM-rawIso e n) = h n + where + h : (n : ℕ) → section (inducedFun-EM-raw (fst e .fst , snd e) n) + (inducedFun-EM-raw (invEq (fst e) , isGroupHomInv e) n) + h = ind+2 + (secEq (fst e)) + (elimSet _ (λ _ → emsquash _ _) refl + (λ g → compPathR→PathP + ((sym (cong₂ _∙_ (cong emloop (secEq (fst e) g)) (sym (lUnit _)) + ∙ rCancel _))))) + λ n p → λ { north → refl + ; south → refl + ; (merid a i) k → merid (p a k) i} +Iso.leftInv (inducedFun-EM-rawIso e n) = h n + where + h : (n : ℕ) → retract (Iso.fun (inducedFun-EM-rawIso e n)) + (Iso.inv (inducedFun-EM-rawIso e n)) + h = ind+2 + (retEq (fst e)) + (elimSet _ (λ _ → emsquash _ _) refl + (λ g → compPathR→PathP + ((sym (cong₂ _∙_ (cong emloop (retEq (fst e) g)) (sym (lUnit _)) + ∙ rCancel _))))) + λ n p → λ { north → refl + ; south → refl + ; (merid a i) k → merid (p a k) i} + +Iso→EMIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} + → AbGroupEquiv G H → ∀ n → Iso (EM G n) (EM H n) +Iso→EMIso is zero = inducedFun-EM-rawIso is zero +Iso→EMIso is (suc zero) = inducedFun-EM-rawIso is 1 +Iso→EMIso is (suc (suc n)) = mapCompIso (inducedFun-EM-rawIso is (suc (suc n))) + +EM⊗-commFun : {G : AbGroup ℓ} {H : AbGroup ℓ'} + → ∀ n → Iso (EM (G ⨂ H) n) (EM (H ⨂ G) n) +EM⊗-commFun {G = G} {H = H} = Iso→EMIso (GroupIso→GroupEquiv ⨂-commIso) + +EM⊗-assocIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} + → ∀ n → Iso (EM (G ⨂ (H ⨂ L)) n) (EM ((G ⨂ H) ⨂ L) n) +EM⊗-assocIso = Iso→EMIso (GroupIso→GroupEquiv (GroupEquiv→GroupIso ⨂assoc)) diff --git a/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda b/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda index 953e3e4f7..576baefd9 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda @@ -1,39 +1,31 @@ -{-# OPTIONS --cubical --no-import-sorts --safe --experimental-lossy-unification #-} +{-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.Algebra.Group.EilenbergMacLane.WedgeConnectivity where open import Cubical.Algebra.Group.EilenbergMacLane.Base open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Equiv.HalfAdjoint -open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) -open import Cubical.Foundations.Path open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Transport -open import Cubical.Data.Unit -open import Cubical.Data.Sigma open import Cubical.Algebra.Group.Base -open import Cubical.Algebra.Group.Properties -open import Cubical.Homotopy.Connected open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec; elim to trElim) open import Cubical.HITs.EilenbergMacLane1 open import Cubical.Algebra.AbGroup.Base open import Cubical.Data.Empty renaming (rec to ⊥-rec) -open import Cubical.HITs.Truncation - renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) -open import Cubical.Data.Nat hiding (_·_) +open import Cubical.Data.Nat open import Cubical.HITs.Susp -open import Cubical.Functions.Morphism -open import Cubical.Foundations.Path + +{- +This file contains a direct construction of the wedge connectivity lemma for EM +spaces. This direct construction gives nicer reductions and computes better than +the more general theorem. The main results are in the module "wedgeConEM" at the +end of this file. +-} private variable ℓ ℓ' ℓ'' : Level +-- One of the base cases, stated separately to avoid termination issues wedgeConFun' : (G : AbGroup ℓ) (H : AbGroup ℓ') (n : ℕ) → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ''} → ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y)) @@ -49,7 +41,7 @@ private → (p : f ptS ≡ g ptS) → (x : _) → wedgeConFun' G H n hLev f g p x ptS ≡ g x wedgeConFun' G H zero {A = A} hlev f g p = - elimSet _ (λ _ → isSetΠ λ _ → hlev _ _) f k + elimSet _ (λ _ → isSetΠ λ _ → hlev _ _) f mainpath where I→A : (x : fst G) → (i : I) → A (emloop x i) embase I→A x i = @@ -70,25 +62,23 @@ private (λ _ i → I→A x i) λ j i → SquareP2 x h i j CubeP2 x g h = toPathP (isOfHLevelPathP' 1 (isOfHLevelPathP 2 (hlev _ _) _ _) _ _ _ _) - k : (x : _) → PathP (λ i → (y : EM₁ (AbGroup→Group H)) → A (emloop x i) y) f f - k x i embase = I→A x i - k x i (emloop z j) = SquareP2 x z i j - k x i (emcomp g h j k) = CubeP2 x g h k j i - k x i (emsquash y z p q r s j k' l) = mega i j k' l + mainpath : (x : _) → PathP (λ i → (y : EM₁ (AbGroup→Group H)) → A (emloop x i) y) f f + mainpath x i embase = I→A x i + mainpath x i (emloop z j) = SquareP2 x z i j + mainpath x i (emcomp g h j k) = CubeP2 x g h k j i + mainpath x i (emsquash y z p q r s j k' l) = mega i j k' l where mega : - PathP (λ i → - PathP (λ j → - PathP (λ k' → - PathP (λ l → A (emloop x i) (emsquash y z p q r s j k' l)) - (k x i y) - (k x i z)) - (λ l → k x i (p l)) - λ l → k x i (q l)) - (λ k' l → k x i (r k' l)) - λ k' l → k x i (s k' l)) - (λ j k l → f (emsquash y z p q r s j k l)) - λ j k l → f (emsquash y z p q r s j k l) + PathP (λ i → PathP (λ j → PathP (λ k → + PathP (λ l → A (emloop x i) (emsquash y z p q r s j k l)) + (mainpath x i y) + (mainpath x i z)) + (λ l → mainpath x i (p l)) + λ l → mainpath x i (q l)) + (λ k l → mainpath x i (r k l)) + λ k l → mainpath x i (s k l)) + (λ j mainpath l → f (emsquash y z p q r s j mainpath l)) + λ j mainpath l → f (emsquash y z p q r s j mainpath l) mega = toPathP (isOfHLevelPathP' 1 (isOfHLevelPathP 2 (isOfHLevelPathP 2 (hlev _ _) _ _) _ _) _ _ _ _) @@ -100,11 +90,11 @@ private llem₁ = (λ i → transp (λ j → A (merid ptS (j ∨ ~ i)) ptS) (~ i) (g (merid ptS (~ i)))) ∙ cong (subst (λ x₁ → A x₁ ptS) (merid ptS)) (sym p) - llem₂' : + llem₁' : Square (λ i → transp (λ j → A (merid ptS (i ∨ j)) ptS) i (g (merid ptS i))) refl (cong (subst (λ x → A x ptS) (merid ptS)) (sym p)) llem₁ - llem₂' i j = + llem₁' i j = hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptS (j ∨ z)) ptS) j (g (merid ptS j)) ; (i = i1) → subst (λ x₁ → A x₁ ptS) (merid ptS) (p (~ k)) @@ -117,7 +107,7 @@ private llem₂ i j = hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptS (j ∧ j₁)) ptS) (~ j) (p (~ k)) ; (j = i0) → p (~ k) - ; (j = i1) → llem₂' k i}) + ; (j = i1) → llem₁' k i}) (transp (λ k → A (merid ptS ((i ∨ k) ∧ j)) ptS) (i ∨ ~ j) (g (merid ptS (i ∧ j)))) mainₗ : (a : _) (y : _) @@ -150,8 +140,8 @@ private wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p south = sym (llem₁ G H n hLev f g p ptS i0 ptS) wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p (merid a i) k = P k i where - lem : _ - lem i j = + llem : _ + llem i j = hcomp (λ k → λ { (i = i1) → g (merid a j) ; (j = i0) → p (i ∨ ~ k) ; (j = i1) → llem₁ G H n hLev f g p ptS i0 ptS (~ i ∧ k)}) @@ -160,8 +150,10 @@ private P : PathP (λ k → PathP (λ i → A (merid a i) ptS) (p k) (llem₁ G H n hLev f g p ptS i0 ptS (~ k))) (λ i → mainₗ G H n hLev f g p a i ptS a ptS i) λ i → g (merid a i) - P = mainP G H n hLev f g p a i0 ptS a ◁ lem + P = mainP G H n hLev f g p a i0 ptS a ◁ llem +-- Here, the actual stuff gets proved. However an additional natural number is stuck into the context +-- to convince the termination checker private wedgeConFun : (G : AbGroup ℓ) (H : AbGroup ℓ') (k n m : ℕ) → (n + m ≡ k) → {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ''} @@ -177,7 +169,8 @@ private → (g : (x : _) → A x ptS) → (p : f ptS ≡ g ptS) → (x : _) → wedgeConFun G H k n m P hLev f g p ptS x ≡ f x - wedgeconRight : (G : AbGroup ℓ) (H : AbGroup ℓ') (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ''} + wedgeconRight : (G : AbGroup ℓ) (H : AbGroup ℓ') + (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ''} → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) → (f : (x : _) → A ptS x) → (g : (x : _) → A x ptS) @@ -190,60 +183,62 @@ private g f (sym p) y x wedgeConFun G H l (suc n) (suc m) P {A = A} hlev f g p north y = f y wedgeConFun G H l (suc n) (suc m) P {A = A} hlev f g p south y = subst (λ x → A x y) (merid ptS) (f y) - wedgeConFun G H zero (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = main** i - module _ where - main** : PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) - main** = ⊥-rec (snotz P) - wedgeConFun G H (suc l) (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = main* a y i + wedgeConFun G H zero (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = ⊥-path i + where + ⊥-path : PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) + ⊥-path = ⊥-rec (snotz P) + wedgeConFun G H (suc l) (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = mmain a y i module _ where - lem₁* : g south ≡ (subst (λ x → A x ptS) (merid ptS) (f ptS)) - lem₁* = ((λ i → transp (λ j → A (merid ptS (~ i ∨ j)) ptS) (~ i) (g (merid ptS (~ i))))) + llem₃ : g south ≡ (subst (λ x → A x ptS) (merid ptS) (f ptS)) + llem₃ = ((λ i → transp (λ j → A (merid ptS (~ i ∨ j)) ptS) (~ i) (g (merid ptS (~ i))))) ∙ cong (subst (λ x → A x ptS) (merid ptS)) (sym p) - lem₁'* : + llem₃' : Square (λ i → transp (λ j → A (merid ptS (~ i ∨ j)) ptS) (~ i) (g (merid ptS (~ i)))) (refl {x = subst (λ x → A x ptS) (merid ptS) (f ptS)}) - lem₁* + llem₃ ((cong (transport (λ z → A (merid ptS z) ptS)) (sym p))) - lem₁'* i j = + llem₃' i j = hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptS (~ j ∨ j₁)) ptS) (~ j) (g (merid ptS (~ j))) ; (i = i1) → subst (λ x → A x ptS) (merid ptS) (p (~ k)) ; (j = i1) → cong (transport (λ z → A (merid ptS z) ptS)) (sym p) (i ∧ k)}) (transp (λ j₁ → A (merid ptS ((~ j ∧ ~ i) ∨ j₁)) ptS) (~ j ∧ ~ i) (g (merid ptS (~ j ∧ ~ i)))) - lem₂* : (λ i₁ → transp (λ j → A (merid ptS (j ∧ i₁)) ptS) (~ i₁) (f ptS)) - ≡ (λ i₁ → hcomp (λ k → λ { (i₁ = i0) → p (~ k) ; (i₁ = i1) → lem₁* k }) + llem₄ : (λ i₁ → transp (λ j → A (merid ptS (j ∧ i₁)) ptS) (~ i₁) (f ptS)) + ≡ (λ i₁ → hcomp (λ k → λ { (i₁ = i0) → p (~ k) ; (i₁ = i1) → llem₃ k }) (g (merid ptS i₁))) - lem₂* i j = + llem₄ i j = hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptS (z ∧ j)) ptS) (~ j) (p (~ k)) ; (j = i0) → p (~ k) - ; (j = i1) → lem₁'* k (~ i)}) + ; (j = i1) → llem₃' k (~ i)}) (transp (λ z → A (merid ptS ((i ∨ z) ∧ j)) ptS) (i ∨ ~ j) (g (merid ptS (i ∧ j)))) - main* : (a : _) (y : _) → PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) - main* = + mmain : (a : _) (y : _) + → PathP (λ i → A (merid a i) y) (f y) + (subst (λ x → A x y) (merid ptS) (f y)) + mmain = wedgeConFun G H l n (suc m) (cong predℕ P) (λ _ _ → isOfHLevelPathP' (suc (n + (suc (suc m)))) (hlev _ _) _ _) (λ x i → transp (λ j → A (merid ptS (j ∧ i)) x) (~ i) (f x)) (λ y i → hcomp (λ k → λ { (i = i0) → p (~ k) - ; (i = i1) → lem₁* k}) + ; (i = i1) → llem₃ k}) (g (merid y i))) - lem₂* + llem₄ - mainR : (x : _) → main* x ptS + mainR : (x : _) → mmain x ptS ≡ λ i → hcomp (λ k → λ { (i = i0) → p (~ k) - ; (i = i1) → lem₁* k}) + ; (i = i1) → llem₃ k}) (g (merid x i)) mainR x = wedgeconRight G H l n (suc m) (cong predℕ P) (λ _ _ → isOfHLevelPathP' (suc (n + (suc (suc m)))) (hlev _ _) _ _) (λ x i → transp (λ j → A (merid ptS (j ∧ i)) x) (~ i) (f x)) (λ y i → hcomp (λ k → λ { (i = i0) → p (~ k) - ; (i = i1) → lem₁* k}) + ; (i = i1) → llem₃ k}) (g (merid y i))) - lem₂* x + llem₄ x wedgeconLeft G H k zero zero P {A = A} hLev f g p _ = refl wedgeconLeft G H k (suc n) zero P {A = A} hLev f g p _ = refl wedgeconLeft G H k zero (suc m) P {A = A} hLev f g p x = @@ -258,23 +253,24 @@ private wedgeconRight G H zero (suc n) (suc m) P {A = A} hLev f g p x = ⊥-rec (snotz P) wedgeconRight G H l (suc n) (suc m) P {A = A} hLev f g p north = p wedgeconRight G H l (suc n) (suc m) P {A = A} hLev f g p south = - sym (lem₁* G H _ n m refl hLev f g p ptS i0 ptS) + sym (llem₃ G H _ n m refl hLev f g p ptS i0 ptS) wedgeconRight G H (suc l) (suc n) (suc m) P {A = A} hLev f g p (merid a i) k = help k i where - lem : _ - lem i j = + llem : _ + llem i j = hcomp (λ k → λ { (i = i1) → g (merid a j) ; (j = i0) → p (i ∨ ~ k) - ; (j = i1) → lem₁* G H l n m P hLev f g p ptS i0 ptS (~ i ∧ k)}) + ; (j = i1) → llem₃ G H l n m P hLev f g p ptS i0 ptS (~ i ∧ k)}) (g (merid a j)) help : PathP (λ k → PathP (λ i → A (merid a i) ptS) - (p k) (lem₁* G H l n m P hLev f g p ptS i0 ptS (~ k))) - (λ i → main* G H l n m P hLev f g p a i north a north i) (cong g (merid a)) - help = mainR G H l n m P hLev f g p a i0 ptS a ◁ lem + (p k) (llem₃ G H l n m P hLev f g p ptS i0 ptS (~ k))) + (λ i → mmain G H l n m P hLev f g p a i north a north i) (cong g (merid a)) + help = mainR G H l n m P hLev f g p a i0 ptS a ◁ llem -module wedgeConEM (G : AbGroup ℓ) (H : AbGroup ℓ') (n m : ℕ) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ''} +module wedgeConEM (G : AbGroup ℓ) (H : AbGroup ℓ') (n m : ℕ) + {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ''} (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) (f : (x : _) → A ptS x) (g : (x : _) → A x ptS) diff --git a/Cubical/Algebra/Group/EilenbergMacLane1.agda b/Cubical/Algebra/Group/EilenbergMacLane1.agda deleted file mode 100644 index fd97d5c5e..000000000 --- a/Cubical/Algebra/Group/EilenbergMacLane1.agda +++ /dev/null @@ -1,995 +0,0 @@ -{-# OPTIONS --cubical --no-import-sorts --safe --experimental-lossy-unification #-} - -module Cubical.Algebra.Group.EilenbergMacLane1 where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Equiv.HalfAdjoint -open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) -open import Cubical.Foundations.Path -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Transport - -open import Cubical.Data.Unit -open import Cubical.Data.Sigma -open import Cubical.Algebra.Group.Base -open import Cubical.Algebra.Group.Properties -open import Cubical.Homotopy.Connected -open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec; elim to trElim) -open import Cubical.HITs.EilenbergMacLane1 -open import Cubical.Algebra.AbGroup.Base -open import Cubical.Data.Empty - renaming (rec to ⊥-rec) -open import Cubical.HITs.Truncation - renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) -open import Cubical.Data.Nat hiding (_·_) -open import Cubical.HITs.Susp -open import Cubical.Functions.Morphism -open import Cubical.Foundations.Path - -private - variable ℓ : Level - -module _ (Ĝ : Group ℓ) where - private - G = fst Ĝ - open GroupStr (snd Ĝ) - - emloop-id : emloop 1g ≡ refl - emloop-id = - emloop 1g ≡⟨ rUnit (emloop 1g) ⟩ - emloop 1g ∙ refl ≡⟨ cong (emloop 1g ∙_) (rCancel (emloop 1g) ⁻¹) ⟩ - emloop 1g ∙ (emloop 1g ∙ (emloop 1g) ⁻¹) ≡⟨ ∙assoc _ _ _ ⟩ - (emloop 1g ∙ emloop 1g) ∙ (emloop 1g) ⁻¹ ≡⟨ cong (_∙ emloop 1g ⁻¹) - ((emloop-comp Ĝ 1g 1g) ⁻¹) ⟩ - emloop (1g · 1g) ∙ (emloop 1g) ⁻¹ ≡⟨ cong (λ g → emloop {Group = Ĝ} g - ∙ (emloop 1g) ⁻¹) - (rid 1g) ⟩ - emloop 1g ∙ (emloop 1g) ⁻¹ ≡⟨ rCancel (emloop 1g) ⟩ - refl ∎ - - emloop-inv : (g : G) → emloop (inv g) ≡ (emloop g) ⁻¹ - emloop-inv g = - emloop (inv g) ≡⟨ rUnit (emloop (inv g)) ⟩ - emloop (inv g) ∙ refl ≡⟨ cong (emloop (inv g) ∙_) - (rCancel (emloop g) ⁻¹) ⟩ - emloop (inv g) ∙ (emloop g ∙ (emloop g) ⁻¹) ≡⟨ ∙assoc _ _ _ ⟩ - (emloop (inv g) ∙ emloop g) ∙ (emloop g) ⁻¹ ≡⟨ cong (_∙ emloop g ⁻¹) - ((emloop-comp Ĝ (inv g) g) ⁻¹) ⟩ - emloop (inv g · g) ∙ (emloop g) ⁻¹ ≡⟨ cong (λ h → emloop {Group = Ĝ} h - ∙ (emloop g) ⁻¹) - (invl g) ⟩ - emloop 1g ∙ (emloop g) ⁻¹ ≡⟨ cong (_∙ (emloop g) ⁻¹) emloop-id ⟩ - refl ∙ (emloop g) ⁻¹ ≡⟨ (lUnit ((emloop g) ⁻¹)) ⁻¹ ⟩ - (emloop g) ⁻¹ ∎ - - isGroupoidEM₁ : isGroupoid (EM₁ Ĝ) - isGroupoidEM₁ = emsquash - - isConnectedEM₁ : isConnected 2 (EM₁ Ĝ) - isConnectedEM₁ = ∣ embase ∣ , h - where - h : (y : hLevelTrunc 2 (EM₁ Ĝ)) → ∣ embase ∣ ≡ y - h = trElim (λ y → isOfHLevelSuc 1 (isOfHLevelTrunc 2 ∣ embase ∣ y)) - (elimProp Ĝ (λ x → isOfHLevelTrunc 2 ∣ embase ∣ ∣ x ∣) refl) - - {- since we write composition in diagrammatic order, - and function composition in the other order, - we need right multiplication here -} - rightEquiv : (g : G) → G ≃ G - rightEquiv g = isoToEquiv isom - where - isom : Iso G G - isom .Iso.fun = _· g - isom .Iso.inv = _· inv g - isom .Iso.rightInv h = - (h · inv g) · g ≡⟨ (assoc h (inv g) g) ⁻¹ ⟩ - h · inv g · g ≡⟨ cong (h ·_) (invl g) ⟩ - h · 1g ≡⟨ rid h ⟩ h ∎ - isom .Iso.leftInv h = - (h · g) · inv g ≡⟨ (assoc h g (inv g)) ⁻¹ ⟩ - h · g · inv g ≡⟨ cong (h ·_) (invr g) ⟩ - h · 1g ≡⟨ rid h ⟩ h ∎ - - compRightEquiv : (g h : G) - → compEquiv (rightEquiv g) (rightEquiv h) ≡ rightEquiv (g · h) - compRightEquiv g h = equivEq (funExt (λ x → (assoc x g h) ⁻¹)) - - CodesSet : EM₁ Ĝ → hSet ℓ - CodesSet = rec Ĝ (isOfHLevelTypeOfHLevel 2) (G , is-set) RE REComp - where - RE : (g : G) → Path (hSet ℓ) (G , is-set) (G , is-set) - RE g = Σ≡Prop (λ X → isPropIsOfHLevel {A = X} 2) (ua (rightEquiv g)) - - lemma₁ : (g h : G) → Square - (ua (rightEquiv g)) (ua (rightEquiv (g · h))) - refl (ua (rightEquiv h)) - lemma₁ g h = invEq - (Square≃doubleComp (ua (rightEquiv g)) (ua (rightEquiv (g · h))) - refl (ua (rightEquiv h))) - (ua (rightEquiv g) ∙ ua (rightEquiv h) - ≡⟨ (uaCompEquiv (rightEquiv g) (rightEquiv h)) ⁻¹ ⟩ - ua (compEquiv (rightEquiv g) (rightEquiv h)) - ≡⟨ cong ua (compRightEquiv g h) ⟩ - ua (rightEquiv (g · h)) ∎) - - lemma₂ : {A₀₀ A₀₁ : hSet ℓ} (p₀₋ : A₀₀ ≡ A₀₁) - {A₁₀ A₁₁ : hSet ℓ} (p₁₋ : A₁₀ ≡ A₁₁) - (p₋₀ : A₀₀ ≡ A₁₀) (p₋₁ : A₀₁ ≡ A₁₁) - (s : Square (cong fst p₀₋) (cong fst p₁₋) (cong fst p₋₀) (cong fst p₋₁)) - → Square p₀₋ p₁₋ p₋₀ p₋₁ - fst (lemma₂ p₀₋ p₁₋ p₋₀ p₋₁ s i j) = s i j - snd (lemma₂ p₀₋ p₁₋ p₋₀ p₋₁ s i j) = - isSet→SquareP {A = (λ i j → isSet (s i j))} - (λ i j → isProp→isSet (isPropIsOfHLevel 2)) - (cong snd p₀₋) (cong snd p₁₋) (cong snd p₋₀) (cong snd p₋₁) i j - - REComp : (g h : G) → Square (RE g) (RE (g · h)) refl (RE h) - REComp g h = lemma₂ (RE g) (RE (g · h)) refl (RE h) (lemma₁ g h) - Codes : EM₁ Ĝ → Type ℓ - Codes x = CodesSet x .fst - - encode : (x : EM₁ Ĝ) → embase ≡ x → Codes x - encode x p = subst Codes p 1g - - decode : (x : EM₁ Ĝ) → Codes x → embase ≡ x - decode = elimSet Ĝ (λ x → isOfHLevelΠ 2 (λ c → isGroupoidEM₁ (embase) x)) - emloop λ g → ua→ λ h → emcomp h g - - - decode-encode : (x : EM₁ Ĝ) (p : embase ≡ x) → decode x (encode x p) ≡ p - decode-encode x p = J (λ y q → decode y (encode y q) ≡ q) - (emloop (transport refl 1g) ≡⟨ cong emloop (transportRefl 1g) ⟩ - emloop 1g ≡⟨ emloop-id ⟩ refl ∎) p - - encode-decode : (x : EM₁ Ĝ) (c : Codes x) → encode x (decode x c) ≡ c - encode-decode = elimProp Ĝ (λ x → isOfHLevelΠ 1 (λ c → CodesSet x .snd _ _)) - λ g → encode embase (decode embase g) ≡⟨ refl ⟩ - encode embase (emloop g) ≡⟨ refl ⟩ - transport (ua (rightEquiv g)) 1g ≡⟨ uaβ (rightEquiv g) 1g ⟩ - 1g · g ≡⟨ lid g ⟩ - g ∎ - - ΩEM₁Iso : Iso (Path (EM₁ Ĝ) embase embase) G - Iso.fun ΩEM₁Iso = encode embase - Iso.inv ΩEM₁Iso = emloop - Iso.rightInv ΩEM₁Iso = encode-decode embase - Iso.leftInv ΩEM₁Iso = decode-encode embase - - ΩEM₁≡ : (Path (EM₁ Ĝ) embase embase) ≡ G - ΩEM₁≡ = isoToPath ΩEM₁Iso - -private - variable - ℓG : Level - - -Susp̂ : (n : ℕ) → Type ℓG → Type ℓG -Susp̂ zero A = A -Susp̂ (suc n) A = Susp (Susp̂ n A) - -pt-Susp̂ : (n : ℕ) (A : Pointed ℓG) → Susp̂ n (typ A) -pt-Susp̂ zero A = snd A -pt-Susp̂ (suc n) A = north - - - -_* = AbGroup→Group - -EM-raw : (G : AbGroup ℓG) (n : ℕ) → Type ℓG -EM-raw G zero = fst G -EM-raw G (suc zero) = EM₁ (G *) -EM-raw G (suc (suc n)) = Susp (EM-raw G (suc n)) - -ptS : {n : ℕ} {G : AbGroup ℓG} → EM-raw G n -ptS {n = zero} {G = G} = AbGroupStr.0g (snd G) -ptS {n = suc zero} {G = G} = embase -ptS {n = suc (suc n)} {G = G} = north - -EM-raw-elim : (G : AbGroup ℓG) (n : ℕ) {A : EM-raw G (suc n) → Type ℓ} - → ((x : _) → isOfHLevel (suc n) (A x) ) - → A ptS - → (x : _) → A x -EM-raw-elim G zero hlev b = elimProp _ hlev b -EM-raw-elim G (suc n) hlev b north = b -EM-raw-elim G (suc n) {A = A} hlev b south = subst A (merid ptS) b -EM-raw-elim G (suc n) {A = A} hlev b (merid a i) = help a i - where - help : (a : _) → PathP (λ i → A (merid a i)) b (subst A (merid ptS) b) - help = EM-raw-elim G n (λ _ → isOfHLevelPathP' (suc n) (hlev _) _ _) - λ i → transp (λ j → A (merid ptS (j ∧ i))) (~ i) b - -EM : (G : AbGroup ℓG) (n : ℕ) → Type ℓG -EM G zero = EM-raw G zero -EM G (suc zero) = EM-raw G 1 -EM G (suc (suc n)) = hLevelTrunc (4 + n) (EM-raw G (suc (suc n))) - -0ₖ : {G : AbGroup ℓG} (n : ℕ) → EM G n -0ₖ {G = G} zero = AbGroupStr.0g (snd G) -0ₖ (suc zero) = embase -0ₖ (suc (suc n)) = ∣ ptS ∣ - -EM∙ : (G : AbGroup ℓG) (n : ℕ) → Pointed ℓG -EM∙ G n = EM G n , (0ₖ n) - -hLevelEM : (G : AbGroup ℓG) (n : ℕ) → isOfHLevel (2 + n) (EM G n) -hLevelEM G zero = AbGroupStr.is-set (snd G) -hLevelEM G (suc zero) = emsquash -hLevelEM G (suc (suc n)) = isOfHLevelTrunc (4 + n) - -EM-raw→EM : (G : AbGroup ℓG) (n : ℕ) → EM-raw G n → EM G n -EM-raw→EM G zero x = x -EM-raw→EM G (suc zero) x = x -EM-raw→EM G (suc (suc n)) = ∣_∣ - -EM-elim : {G : AbGroup ℓG} (n : ℕ) {A : EM G n → Type ℓ} - → ((x : _) → isOfHLevel (2 + n) (A x)) - → ((x : EM-raw G n) → A (EM-raw→EM G n x)) - → (x : _) → A x -EM-elim zero hlev hyp x = hyp x -EM-elim (suc zero) hlev hyp x = hyp x -EM-elim (suc (suc n)) hlev hyp = trElim (λ _ → hlev _) hyp - -wedgeConFun' : (G H : AbGroup ℓG) (n : ℕ) → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ} - → ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y)) - → (f : (x : _) → A ptS x) - → (g : (x : _) → A x ptS) - → f ptS ≡ g ptS - → (x : _) (y : _) → A x y -wedgeConFun'ᵣ : (G H : AbGroup ℓG) (n : ℕ) → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ} - → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y))) - → (f : (x : _) → A ptS x) - → (g : (x : _) → A x ptS) - → (p : f ptS ≡ g ptS) - → (x : _) → wedgeConFun' G H n hLev f g p x ptS ≡ g x -wedgeConFun' G H zero {A = A} hlev f g p = - elimSet _ (λ _ → isSetΠ λ _ → hlev _ _) f k - where - I→A : (x : fst G) → (i : I) → A (emloop x i) embase - I→A x i = - hcomp (λ k → λ {(i = i0) → p (~ k) ; (i = i1) → p (~ k)}) - (g (emloop x i)) - - SquareP2 : (x : _) (z : _) - → SquareP (λ i j → A (emloop x i) (emloop z j)) - (cong f (emloop z)) (cong f (emloop z)) - (λ i → I→A x i) λ i → I→A x i - SquareP2 x z = - toPathP (isOfHLevelPathP' 1 (λ _ _ → hlev _ _ _ _) _ _ _ _) - - CubeP2 : (x : _) (g h : _) - → PathP (λ k → PathP (λ j → PathP (λ i → A (emloop x i) (emcomp g h j k)) - (f (emcomp g h j k)) (f (emcomp g h j k))) - (λ i → SquareP2 x g i k) λ i → SquareP2 x ((snd (AbGroup→Group H) GroupStr.· g) h) i k) - (λ _ i → I→A x i) λ j i → SquareP2 x h i j - CubeP2 x g h = toPathP (isOfHLevelPathP' 1 (isOfHLevelPathP 2 (hlev _ _) _ _) _ _ _ _) - - k : (x : _) → PathP (λ i → (y : EM₁ (AbGroup→Group H)) → A (emloop x i) y) f f - k x i embase = I→A x i - k x i (emloop z j) = SquareP2 x z i j - k x i (emcomp g h j k) = CubeP2 x g h k j i - k x i (emsquash y z p q r s j k' l) = mega i j k' l - where - mega : - PathP (λ i → - PathP (λ j → - PathP (λ k' → - PathP (λ l → A (emloop x i) (emsquash y z p q r s j k' l)) - (k x i y) - (k x i z)) - (λ l → k x i (p l)) - λ l → k x i (q l)) - (λ k' l → k x i (r k' l)) - λ k' l → k x i (s k' l)) - (λ j k l → f (emsquash y z p q r s j k l)) - λ j k l → f (emsquash y z p q r s j k l) - mega = - toPathP (isOfHLevelPathP' 1 - (isOfHLevelPathP 2 (isOfHLevelPathP 2 (hlev _ _) _ _) _ _) _ _ _ _) -wedgeConFun' G H (suc n) {A = A} hLev f g p north y = f y -wedgeConFun' G H (suc n) {A = A} hLev f g p south y = subst (λ x → A x y) (merid ptS) (f y) -wedgeConFun' G H (suc n) {A = A} hLev f g p (merid a i) y = mainₗ a y i - module _ where - llem₁ : g south ≡ subst (λ x₁ → A x₁ ptS) (merid ptS) (f ptS) - llem₁ = (λ i → transp (λ j → A (merid ptS (j ∨ ~ i)) ptS) (~ i) (g (merid ptS (~ i)))) - ∙ cong (subst (λ x₁ → A x₁ ptS) (merid ptS)) (sym p) - - llem₂' : - Square - (λ i → transp (λ j → A (merid ptS (i ∨ j)) ptS) i (g (merid ptS i))) refl - (cong (subst (λ x → A x ptS) (merid ptS)) (sym p)) llem₁ - llem₂' i j = - hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptS (j ∨ z)) ptS) j - (g (merid ptS j)) - ; (i = i1) → subst (λ x₁ → A x₁ ptS) (merid ptS) (p (~ k)) - ; (j = i0) → (subst (λ x → A x ptS) (merid ptS)) (p (~ k ∨ ~ i))}) - (transp (λ k → A (merid ptS (k ∨ ~ i ∧ j)) ptS) (~ i ∧ j) (g (merid ptS (~ i ∧ j)))) - - llem₂ : (λ i₁ → transp (λ j → A (merid ptS (i₁ ∧ j)) ptS) (~ i₁) (f ptS)) - ≡ (λ i₁ → hcomp (λ k → λ { (i₁ = i0) → p (~ k) ; (i₁ = i1) → llem₁ k }) - (g (merid ptS i₁))) - llem₂ i j = - hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptS (j ∧ j₁)) ptS) (~ j) (p (~ k)) - ; (j = i0) → p (~ k) - ; (j = i1) → llem₂' k i}) - (transp (λ k → A (merid ptS ((i ∨ k) ∧ j)) ptS) (i ∨ ~ j) (g (merid ptS (i ∧ j)))) - - mainₗ : (a : _) (y : _) - → PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) - mainₗ = - wedgeConFun' G H n - (λ _ _ → isOfHLevelPathP' (suc (n + 1)) (hLev _ _) _ _) - (λ x i → transp (λ j → A (merid ptS (i ∧ j)) x) (~ i) (f x)) - (λ x i → hcomp (λ k → λ { (i = i0) → p (~ k) - ; (i = i1) → llem₁ k}) - (g (merid x i))) - llem₂ - - mainP : (y : _) - → mainₗ y ptS - ≡ λ i → hcomp (λ k → λ { (i = i0) → p (~ k) - ; (i = i1) → llem₁ k}) - (g (merid y i)) - mainP y = - wedgeConFun'ᵣ G H n - (λ _ _ → isOfHLevelPathP' (suc (n + 1)) (hLev _ _) _ _) - (λ x i → transp (λ j → A (merid ptS (i ∧ j)) x) (~ i) (f x)) - (λ x i → hcomp (λ k → λ { (i = i0) → p (~ k) - ; (i = i1) → llem₁ k}) - (g (merid x i))) - llem₂ y -wedgeConFun'ᵣ G H zero {A = A} hLev f g p = - elimProp _ (λ _ → hLev _ _ _ _) p -wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p north = p -wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p south = sym (llem₁ G H n hLev f g p ptS i0 ptS) -wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p (merid a i) k = P k i - where - lem : _ - lem i j = - hcomp (λ k → λ { (i = i1) → g (merid a j) - ; (j = i0) → p (i ∨ ~ k) - ; (j = i1) → llem₁ G H n hLev f g p ptS i0 ptS (~ i ∧ k)}) - (g (merid a j)) - - P : PathP (λ k → PathP (λ i → A (merid a i) ptS) - (p k) (llem₁ G H n hLev f g p ptS i0 ptS (~ k))) - (λ i → mainₗ G H n hLev f g p a i ptS a ptS i) λ i → g (merid a i) - P = mainP G H n hLev f g p a i0 ptS a ◁ lem - -private - wedgeConFun : (G H : AbGroup ℓG) (k n m : ℕ) → (n + m ≡ k) → {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} - → ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y)) - → (f : (x : _) → A ptS x) - → (g : (x : _) → A x ptS) - → f ptS ≡ g ptS - → (x : _) (y : _) → A x y - wedgeconLeft : (G H : AbGroup ℓG) (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} - → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) - → (f : (x : _) → A ptS x) - → (g : (x : _) → A x ptS) - → (p : f ptS ≡ g ptS) - → (x : _) → wedgeConFun G H k n m P hLev f g p ptS x ≡ f x - wedgeconRight : (G H : AbGroup ℓG) (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} - → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) - → (f : (x : _) → A ptS x) - → (g : (x : _) → A x ptS) - → (p : f ptS ≡ g ptS) - → (x : _) → wedgeConFun G H k n m P hLev f g p x ptS ≡ g x - wedgeConFun G H k n zero P {A = A} hLev f g p = wedgeConFun' G H n hLev f g p - wedgeConFun G H k zero (suc m) P {A = A} hLev f g p x y = - wedgeConFun' H G (suc m) {A = λ x y → A y x} - (λ x y → subst (λ n → isOfHLevel (2 + n) (A y x)) (+-comm 1 m) (hLev y x)) - g f (sym p) y x - wedgeConFun G H l (suc n) (suc m) P {A = A} hlev f g p north y = f y - wedgeConFun G H l (suc n) (suc m) P {A = A} hlev f g p south y = subst (λ x → A x y) (merid ptS) (f y) - wedgeConFun G H zero (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = main** i - module _ where - main** : PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) - main** = ⊥-rec (snotz P) - wedgeConFun G H (suc l) (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = main* a y i - module _ where - lem₁* : g south ≡ (subst (λ x → A x ptS) (merid ptS) (f ptS)) - lem₁* = ((λ i → transp (λ j → A (merid ptS (~ i ∨ j)) ptS) (~ i) (g (merid ptS (~ i))))) - ∙ cong (subst (λ x → A x ptS) (merid ptS)) (sym p) - - lem₁'* : - Square - (λ i → transp (λ j → A (merid ptS (~ i ∨ j)) ptS) (~ i) (g (merid ptS (~ i)))) - (refl {x = subst (λ x → A x ptS) (merid ptS) (f ptS)}) - lem₁* - ((cong (transport (λ z → A (merid ptS z) ptS)) (sym p))) - lem₁'* i j = - hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptS (~ j ∨ j₁)) ptS) (~ j) (g (merid ptS (~ j))) - ; (i = i1) → subst (λ x → A x ptS) (merid ptS) (p (~ k)) - ; (j = i1) → cong (transport (λ z → A (merid ptS z) ptS)) (sym p) (i ∧ k)}) - (transp (λ j₁ → A (merid ptS ((~ j ∧ ~ i) ∨ j₁)) ptS) (~ j ∧ ~ i) (g (merid ptS (~ j ∧ ~ i)))) - - - lem₂* : (λ i₁ → transp (λ j → A (merid ptS (j ∧ i₁)) ptS) (~ i₁) (f ptS)) - ≡ (λ i₁ → hcomp (λ k → λ { (i₁ = i0) → p (~ k) ; (i₁ = i1) → lem₁* k }) - (g (merid ptS i₁))) - lem₂* i j = - hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptS (z ∧ j)) ptS) (~ j) (p (~ k)) - ; (j = i0) → p (~ k) - ; (j = i1) → lem₁'* k (~ i)}) - (transp (λ z → A (merid ptS ((i ∨ z) ∧ j)) ptS) (i ∨ ~ j) (g (merid ptS (i ∧ j)))) - - main* : (a : _) (y : _) → PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) - main* = - wedgeConFun G H l n (suc m) (cong predℕ P) - (λ _ _ → isOfHLevelPathP' (suc (n + (suc (suc m)))) (hlev _ _) _ _) - (λ x i → transp (λ j → A (merid ptS (j ∧ i)) x) (~ i) (f x)) - (λ y i → hcomp (λ k → λ { (i = i0) → p (~ k) - ; (i = i1) → lem₁* k}) - (g (merid y i))) - lem₂* - - mainR : (x : _) → main* x ptS - ≡ λ i → hcomp (λ k → λ { (i = i0) → p (~ k) - ; (i = i1) → lem₁* k}) - (g (merid x i)) - mainR x = - wedgeconRight G H l n (suc m) (cong predℕ P) - (λ _ _ → isOfHLevelPathP' (suc (n + (suc (suc m)))) (hlev _ _) _ _) - (λ x i → transp (λ j → A (merid ptS (j ∧ i)) x) (~ i) (f x)) - (λ y i → hcomp (λ k → λ { (i = i0) → p (~ k) - ; (i = i1) → lem₁* k}) - (g (merid y i))) - lem₂* x - wedgeconLeft G H k zero zero P {A = A} hLev f g p _ = refl - wedgeconLeft G H k (suc n) zero P {A = A} hLev f g p _ = refl - wedgeconLeft G H k zero (suc m) P {A = A} hLev f g p x = - wedgeConFun'ᵣ H G (suc m) - (λ x₁ y → - subst (λ n → (x₂ y₁ : A y x₁) → isOfHLevel (suc n) (x₂ ≡ y₁)) - (+-comm 1 m) (hLev y x₁)) - g f (λ i → p (~ i)) x - wedgeconLeft G H k (suc n) (suc m) P {A = A} hLev f g p _ = refl - wedgeconRight G H k n zero P {A = A} hLev f g p = wedgeConFun'ᵣ G H n hLev f g p - wedgeconRight G H k zero (suc m) P {A = A} hLev f g p _ = refl - wedgeconRight G H zero (suc n) (suc m) P {A = A} hLev f g p x = ⊥-rec (snotz P) - wedgeconRight G H l (suc n) (suc m) P {A = A} hLev f g p north = p - wedgeconRight G H l (suc n) (suc m) P {A = A} hLev f g p south = - sym (lem₁* G H _ n m refl hLev f g p ptS i0 ptS) - wedgeconRight G H (suc l) (suc n) (suc m) P {A = A} hLev f g p (merid a i) k = - help k i - where - lem : _ - lem i j = - hcomp (λ k → λ { (i = i1) → g (merid a j) - ; (j = i0) → p (i ∨ ~ k) - ; (j = i1) → lem₁* G H l n m P hLev f g p ptS i0 ptS (~ i ∧ k)}) - (g (merid a j)) - - help : PathP (λ k → PathP (λ i → A (merid a i) ptS) - (p k) (lem₁* G H l n m P hLev f g p ptS i0 ptS (~ k))) - (λ i → main* G H l n m P hLev f g p a i north a north i) (cong g (merid a)) - help = mainR G H l n m P hLev f g p a i0 ptS a ◁ lem - - -module wedgeConEM (G H : AbGroup ℓG) (n m : ℕ) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ} - (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) - (f : (x : _) → A ptS x) - (g : (x : _) → A x ptS) - (p : f ptS ≡ g ptS) where - fun : (x : EM-raw G (suc n)) (y : EM-raw H (suc m)) → A x y - fun = wedgeConFun G H (n + m) n m refl hLev f g p - - left : (x : EM-raw H (suc m)) → fun ptS x ≡ f x - left = wedgeconLeft G H (n + m) n m refl hLev f g p - - right : (x : EM-raw G (suc n)) → fun x ptS ≡ g x - right = wedgeconRight G H (n + m) n m refl hLev f g p - -module _ {G : AbGroup ℓG} where - infixr 34 _+ₖ_ - infixr 34 _-ₖ_ - open AbGroupStr (snd G) renaming (_+_ to _+G_ ; -_ to -G_ ; assoc to assocG) - - private - help : (n : ℕ) → n + (4 + n) ≡ (2 + n) + (2 + n) - help n = +-suc n (3 + n) ∙ cong suc (+-suc n (suc (suc n))) - - hLevHelp : (n : ℕ) → isOfHLevel ((2 + n) + (2 + n)) (EM G (2 + n)) - hLevHelp n = - transport (λ i → isOfHLevel (help n i) (EM G (2 + n))) - (isOfHLevelPlus {n = 4 + n} n (isOfHLevelTrunc (4 + n))) - - helper : (g h : fst G) - → PathP (λ i → Path (EM₁ (AbGroup→Group G)) - (emloop h i) (emloop h i)) (emloop g) (emloop g) - helper g h = - compPathL→PathP - (cong (sym (emloop h) ∙_) - (sym (emloop-comp _ g h) - ∙∙ cong emloop (comm g h) - ∙∙ emloop-comp _ h g) - ∙∙ ∙assoc _ _ _ - ∙∙ cong (_∙ emloop g) (lCancel _) - ∙ sym (lUnit _)) - _+ₖ_ : {n : ℕ} → EM G n → EM G n → EM G n - _+ₖ_ {n = zero} = _+G_ - _+ₖ_ {n = suc zero} = - rec _ (isGroupoidΠ (λ _ → emsquash)) - (λ x → x) - hi! - λ g h i j x → el g h x i j - where - lol : (g h : fst G) - → Square (emloop g) (emloop (g +G h)) refl (emloop h) - lol g h = compPath-filler (emloop g) (emloop h) ▷ sym (emloop-comp _ g h) - - hi! : fst G → (λ x → x) ≡ (λ x → x) - hi! g = funExt (elimSet _ (λ _ → emsquash _ _) - (emloop g) - (helper g)) - el : (g h : fst G) (x : EM₁ (AbGroup→Group G)) - → Square (λ j → hi! g j x) (λ j → hi! ((snd (AbGroup→Group G) GroupStr.· g) h) j x) - refl λ j → hi! h j x - el g h = - elimProp _ (λ _ → isOfHLevelPathP' 1 (emsquash _ _) _ _) - (lol g h) - _+ₖ_ {n = suc (suc n)} = - trRec2 (isOfHLevelTrunc (4 + n)) - (wedgeConEM.fun G G (suc n) (suc n) - (λ _ _ → hLevHelp n) - ∣_∣ ∣_∣ refl) - - σ-EM : (n : ℕ) → EM-raw G (suc n) → Path (EM-raw G (2 + n)) ptS ptS - σ-EM n x i = (merid x ∙ sym (merid ptS)) i - - -ₖ_ : {n : ℕ} → EM G n → EM G n - -ₖ_ {n = zero} x = -G x - -ₖ_ {n = suc zero} = - rec _ emsquash - embase - (λ g → sym (emloop g)) - λ g h → sym (emloop-sym _ g) - ◁ (flipSquare - (flipSquare (emcomp (-G g) (-G h)) - ▷ emloop-sym _ h) - ▷ (cong emloop (comm (-G g) (-G h) - ∙ sym (GroupTheory.invDistr (AbGroup→Group G) g h)) - ∙ emloop-sym _ (g +G h))) - -ₖ_ {n = suc (suc n)} = - map λ { north → north - ; south → north - ; (merid a i) → σ-EM n a (~ i)} - - _-ₖ_ : {n : ℕ} → EM G n → EM G n → EM G n - _-ₖ_ {n = n} x y = _+ₖ_ {n = n} x (-ₖ_ {n = n} y) - - +ₖ-syntax : (n : ℕ) → EM G n → EM G n → EM G n - +ₖ-syntax n = _+ₖ_ {n = n} - - -ₖ-syntax : (n : ℕ) → EM G n → EM G n - -ₖ-syntax n = -ₖ_ {n = n} - - -'ₖ-syntax : (n : ℕ) → EM G n → EM G n → EM G n - -'ₖ-syntax n = _-ₖ_ {n = n} - - syntax +ₖ-syntax n x y = x +[ n ]ₖ y - syntax -ₖ-syntax n x = -[ n ]ₖ x - syntax -'ₖ-syntax n x y = x -[ n ]ₖ y - - lUnitₖ : (n : ℕ) (x : EM G n) → 0ₖ n +[ n ]ₖ x ≡ x - lUnitₖ zero x = lid x - lUnitₖ (suc zero) _ = refl - lUnitₖ (suc (suc n)) = - trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) - λ _ → refl - - rUnitₖ : (n : ℕ) (x : EM G n) → x +[ n ]ₖ 0ₖ n ≡ x - rUnitₖ zero x = rid x - rUnitₖ (suc zero) = - elimSet _ (λ _ → emsquash _ _) - refl - λ _ _ → refl - rUnitₖ (suc (suc n)) = - trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) - (wedgeConEM.right G G (suc n) (suc n) - (λ _ _ → hLevHelp n) - ∣_∣ ∣_∣ refl) - - commₖ : (n : ℕ) (x y : EM G n) → x +[ n ]ₖ y ≡ y +[ n ]ₖ x - commₖ zero = comm - commₖ (suc zero) = - wedgeConEM.fun G G 0 0 (λ _ _ → emsquash _ _) - (λ x → sym (rUnitₖ 1 x)) - (rUnitₖ 1) - refl - commₖ (suc (suc n)) = - elim2 (λ _ _ → isOfHLevelTruncPath {n = 4 + n}) - (wedgeConEM.fun G G _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (hLevHelp n) _ _) - (λ x → sym (rUnitₖ (2 + n) ∣ x ∣)) - (λ x → rUnitₖ (2 + n) ∣ x ∣) - refl) - - open import Cubical.Homotopy.Loopspace - cong₂+₁ : (p q : typ (Ω (EM∙ G 1))) → cong₂ (λ x y → x +[ 1 ]ₖ y) p q ≡ p ∙ q - cong₂+₁ p q = - (cong₂Funct (λ x y → x +[ 1 ]ₖ y) p q) - ∙ (λ i → (cong (λ x → rUnitₖ 1 x i) p) ∙ (cong (λ x → lUnitₖ 1 x i) q)) - - cong₂+₂ : (n : ℕ) (p q : typ (Ω (EM∙ G (suc (suc n))))) → cong₂ (λ x y → x +[ (2 + n) ]ₖ y) p q ≡ p ∙ q - cong₂+₂ n p q = - (cong₂Funct (λ x y → x +[ (2 + n) ]ₖ y) p q) - ∙ (λ i → (cong (λ x → rUnitₖ (2 + n) x i) p) ∙ (cong (λ x → lUnitₖ (2 + n) x i) q)) - - isCommΩEM : (n : ℕ) (p q : typ (Ω (EM∙ G (suc n)))) → p ∙ q ≡ q ∙ p - isCommΩEM zero p q = - sym (cong₂+₁ p q) - ∙∙ (λ i j → commₖ 1 (p j) (q j) i) - ∙∙ cong₂+₁ q p - isCommΩEM (suc n) p q = - (sym (cong₂+₂ n p q) - ∙∙ (λ i j → commₖ (suc (suc n)) (p j) (q j) i) - ∙∙ cong₂+₂ n q p) - - cong-₁ : (p : typ (Ω (EM∙ G 1))) → cong (λ x → -[ 1 ]ₖ x) p ≡ sym p - cong-₁ p = main embase p - where - decoder : (x : EM G 1) → embase ≡ x → x ≡ embase - decoder = - elimSet _ - (λ _ → isSetΠ λ _ → emsquash _ _) - (λ p i → -[ 1 ]ₖ (p i)) - λ g → toPathP - (funExt λ x → - (λ i → transport (λ i → Path (EM G 1) (emloop g i) embase) - (cong (-ₖ_ {n = 1}) - (transp (λ j → Path (EM G 1) embase (emloop g (~ j ∧ ~ i))) i - (compPath-filler x (sym (emloop g)) i) ))) - ∙∙ (λ i → transp (λ j → Path (EM G 1) (emloop g (i ∨ j)) embase) i - (compPath-filler' - (sym (emloop g)) - (cong-∙ (-ₖ_ {n = 1}) - x (sym (emloop g)) i) i)) - ∙∙ (cong (sym (emloop g) ∙_) (isCommΩEM 0 (cong (-ₖ_ {n = 1}) x) (emloop g))) - ∙∙ ∙assoc _ _ _ - ∙∙ cong (_∙ (cong (-ₖ_ {n = 1}) x)) (lCancel (emloop g)) - ∙ sym (lUnit _)) - - main : (x : EM G 1) (p : embase ≡ x) → decoder x p ≡ sym p - main x = J (λ x p → decoder x p ≡ sym p) refl - - cong-₂ : (n : ℕ) (p : typ (Ω (EM∙ G (2 + n)))) → cong (λ x → -[ 2 + n ]ₖ x) p ≡ sym p - cong-₂ n p = main _ p - where - pp : (a : _) → PathP (λ i → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ → ∣ merid a i ∣ₕ ≡ 0ₖ (2 + n)) - (cong (λ x → -[ 2 + n ]ₖ x)) λ p → cong ∣_∣ₕ (sym (merid ptS)) ∙ cong (λ x → -[ 2 + n ]ₖ x) p - pp a = - toPathP - (funExt λ x → - (λ k → transp (λ i → Path (EM G (2 + n)) ∣ merid a (i ∨ k) ∣ ∣ ptS ∣) k (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) - (cong (-ₖ-syntax (suc (suc n))) - (transp (λ j → Path (EM G (2 + n)) ∣ ptS ∣ ∣ merid a (~ j ∧ ~ k) ∣) k (compPath-filler x (sym (cong ∣_∣ₕ (merid a))) k))) k)) - ∙∙ cong (cong ∣_∣ₕ (sym (merid a)) ∙_) (cong-∙ (λ x → -[ 2 + n ]ₖ x) x (sym (cong ∣_∣ₕ (merid a))) - ∙ isCommΩEM (suc n) (cong (λ x → -[ 2 + n ]ₖ x) x) (cong ∣_∣ₕ (σ-EM n a))) - ∙∙ (λ k → (λ i → ∣ merid a (~ i ∨ k) ∣) ∙ (λ i → ∣ compPath-filler' (merid a) (sym (merid ptS)) (~ k) i ∣) - ∙ cong (λ x → -ₖ-syntax (suc (suc n)) x) x) - ∙ sym (lUnit _)) - - decoder : (x : EM G (2 + n)) → 0ₖ (2 + n) ≡ x → x ≡ 0ₖ (2 + n) - decoder = - trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelTruncPath {n = 4 + n}) - λ { north → pp ptS i0 - ; south → pp ptS i1 - ; (merid a i) → pp a i} - - main : (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → decoder x p ≡ sym p - main x = J (λ x p → decoder x p ≡ sym p) refl - - rCancelₖ : (n : ℕ) (x : EM G n) → x +[ n ]ₖ (-[ n ]ₖ x) ≡ 0ₖ n - rCancelₖ zero x = invr x - rCancelₖ (suc zero) = - elimSet _ (λ _ → emsquash _ _) - refl - λ g → flipSquare (cong₂+₁ (emloop g) (λ i → -ₖ-syntax 1 (emloop g i)) - ∙ rCancel (emloop g)) - rCancelₖ (suc (suc n)) = - trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) - λ { north → refl - ; south i → +ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣ (-ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣) - ; (merid a i) j - → hcomp (λ r → λ { (i = i0) → 0ₖ (2 + n) - ; (i = i1) → ∣ merid ptS (~ j ∧ r) ∣ₕ -[ 2 + n ]ₖ ∣ merid ptS (~ j ∧ r) ∣ - ; (j = i0) → ∣ compPath-filler (merid a) (sym (merid ptS)) (~ r) i ∣ - -[ 2 + n ]ₖ ∣ compPath-filler (merid a) (sym (merid ptS)) (~ r) i ∣ - ; (j = i1) → 0ₖ (2 + n)}) - (help' a j i) } - where - help' : (a : _) → cong₂ (λ x y → ∣ x ∣ -[ suc (suc n) ]ₖ ∣ y ∣) (σ-EM n a) (σ-EM n a) ≡ refl - help' a = - cong₂+₂ n (cong ∣_∣ₕ (σ-EM n a)) (cong (λ x → -[ 2 + n ]ₖ ∣ x ∣) (σ-EM n a)) - ∙∙ cong (cong ∣_∣ₕ (σ-EM n a) ∙_) (cong-₂ n (cong ∣_∣ₕ (σ-EM n a))) - ∙∙ rCancel _ - - lCancelₖ : (n : ℕ) (x : EM G n) → (-[ n ]ₖ x) +[ n ]ₖ x ≡ 0ₖ n - lCancelₖ n x = commₖ n (-[ n ]ₖ x) x ∙ rCancelₖ n x - - assocₖ : (n : ℕ) (x y z : EM G n) - → (x +[ n ]ₖ (y +[ n ]ₖ z) ≡ (x +[ n ]ₖ y) +[ n ]ₖ z) - assocₖ zero = assocG - assocₖ (suc zero) = - elimSet _ (λ _ → isSetΠ2 λ _ _ → emsquash _ _) - (λ _ _ → refl) - λ g i y z k → lem g y z k i - where - lem : (g : fst G) (y z : _) → cong (λ x → x +[ suc zero ]ₖ (y +[ suc zero ]ₖ z)) (emloop g) - ≡ cong (λ x → (x +[ suc zero ]ₖ y) +[ suc zero ]ₖ z) (emloop g) - lem g = - elimProp _ (λ _ → isPropΠ λ _ → emsquash _ _ _ _) - (elimProp _ (λ _ → emsquash _ _ _ _) - refl) - assocₖ (suc (suc n)) = - elim2 (λ _ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelTruncPath {n = 4 + n}) - λ a b → trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) - (λ c → main c a b) - where - lem : (c : _) (a b : _) - → PathP (λ i → (∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ merid c i ∣ₕ) - ≡ (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ merid c i ∣ₕ)) - (cong (λ x → ∣ a ∣ₕ +[ suc (suc n) ]ₖ x) (rUnitₖ (suc (suc n)) ∣ b ∣) - ∙ sym (rUnitₖ (suc (suc n)) (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ))) - ((λ i → ∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ merid ptS (~ i) ∣ₕ)) - ∙∙ cong (λ x → ∣ a ∣ₕ +[ suc (suc n) ]ₖ x) (rUnitₖ (suc (suc n)) ∣ b ∣) - ∙ sym (rUnitₖ (suc (suc n)) (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ)) - ∙∙ λ i → (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ merid ptS i ∣ₕ) - lem c = - EM-raw-elim G (suc n) - (λ _ → isOfHLevelΠ (2 + n) (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _)) - (EM-raw-elim G (suc n) (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _) - ((sym (rUnit refl) - ◁ λ _ → refl) - ▷ (sym (lCancel (cong ∣_∣ₕ (merid ptS))) - ∙ λ i → (λ j → ∣ merid ptS (~ j ∨ ~ i) ∣ₕ) - ∙∙ lUnit (λ j → ∣ merid ptS (~ j ∧ ~ i) ∣ₕ) i - ∙∙ cong ∣_∣ₕ (merid ptS)))) - main : (c a b : _) → (∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ c ∣ₕ) ≡ (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ c ∣ₕ) - main north a b = lem ptS a b i0 - main south a b = lem ptS a b i1 - main (merid c i) a b = lem c a b i - - σ-EM' : (n : ℕ) (x : EM G (suc n)) - → Path (EM G (suc (suc n))) - (0ₖ (suc (suc n))) - (0ₖ (suc (suc n))) - σ-EM' zero x = cong ∣_∣ₕ (σ-EM zero x) - σ-EM' (suc n) = - trElim (λ _ → isOfHLevelTrunc (5 + n) _ _) - λ x → cong ∣_∣ₕ (σ-EM (suc n) x) - - σ-EM'-0ₖ : (n : ℕ) → σ-EM' n (0ₖ (suc n)) ≡ refl - σ-EM'-0ₖ zero = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) - σ-EM'-0ₖ (suc n) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) - - private - P : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (r : refl ≡ p) - → lUnit p ∙ cong (_∙ p) r - ≡ rUnit p ∙ cong (p ∙_) r - P p = J (λ p r → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r) refl - - σ-EM'-hom : (n : ℕ) → (a b : _) → σ-EM' n (a +ₖ b) ≡ σ-EM' n a ∙ σ-EM' n b - σ-EM'-hom zero = - wedgeConEM.fun G G 0 0 (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) l r p - where - l : _ - l x = cong (σ-EM' zero) (lUnitₖ 1 x) - ∙∙ lUnit (σ-EM' zero x) - ∙∙ cong (_∙ σ-EM' zero x) (sym (σ-EM'-0ₖ zero)) - - r : _ - r x = - cong (σ-EM' zero) (rUnitₖ 1 x) - ∙∙ rUnit (σ-EM' zero x) - ∙∙ cong (σ-EM' zero x ∙_) (sym (σ-EM'-0ₖ zero)) - - p : _ - p = P (σ-EM' zero embase) (sym (σ-EM'-0ₖ zero)) - σ-EM'-hom (suc n) = - elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) _ _) _ _) - (wedgeConEM.fun G G _ _ - (λ x y → transport (λ i → isOfHLevel (help n i) - ((σ-EM' (suc n) (∣ x ∣ₕ +ₖ ∣ y ∣ₕ) - ≡ σ-EM' (suc n) ∣ x ∣ₕ ∙ σ-EM' (suc n) ∣ y ∣ₕ))) - (isOfHLevelPlus {n = 4 + n} n - (isOfHLevelPath (4 + n) - (isOfHLevelTrunc (5 + n) _ _) _ _))) - (λ x → cong (σ-EM' (suc n)) (lUnitₖ (suc (suc n)) ∣ x ∣) - ∙∙ lUnit (σ-EM' (suc n) ∣ x ∣) - ∙∙ cong (_∙ σ-EM' (suc n) ∣ x ∣) (sym (σ-EM'-0ₖ (suc n)))) - (λ x → cong (σ-EM' (suc n)) (rUnitₖ (2 + n) ∣ x ∣) - ∙∙ rUnit (σ-EM' (suc n) ∣ x ∣) - ∙∙ cong (σ-EM' (suc n) ∣ x ∣ ∙_) (sym (σ-EM'-0ₖ (suc n)))) - (P (σ-EM' (suc n) (0ₖ (2 + n))) (sym (σ-EM'-0ₖ (suc n))))) - - σ-EM'-ₖ : (n : ℕ) → (a : _) → σ-EM' n (-ₖ a) ≡ sym (σ-EM' n a) - σ-EM'-ₖ n = - morphLemmas.distrMinus - (λ x y → x +[ suc n ]ₖ y) (_∙_) - (σ-EM' n) (σ-EM'-hom n) - (0ₖ (suc n)) refl - (λ x → -ₖ x) sym - (λ x → sym (lUnit x)) (λ x → sym (rUnit x)) - (lCancelₖ (suc n)) rCancel - ∙assoc (σ-EM'-0ₖ n) - - -Dist : (n : ℕ) (x y : EM G n) → -[ n ]ₖ (x +[ n ]ₖ y) ≡ (-[ n ]ₖ x) +[ n ]ₖ (-[ n ]ₖ y) - -Dist zero x y = (GroupTheory.invDistr (AbGroup→Group G) x y) ∙ commₖ zero _ _ - -Dist (suc zero) = k - where -- useless where clause. Needed for fast type checking for some reason. - l : _ - l x = refl - - r : _ - r x = cong (λ z → -[ 1 ]ₖ z) (rUnitₖ 1 x) ∙ sym (rUnitₖ 1 (-[ 1 ]ₖ x)) - - p : r ptS ≡ l ptS - p = sym (rUnit refl) - - k = wedgeConEM.fun G G 0 0 (λ _ _ → emsquash _ _) l r (sym p) - - -Dist (suc (suc n)) = - elim2 (λ _ _ → isOfHLevelTruncPath {n = 4 + n}) - (wedgeConEM.fun G G (suc n) (suc n) - (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (hLevHelp n) _ _) - (λ x → refl) - (λ x → cong (λ z → -[ (suc (suc n)) ]ₖ z) (rUnitₖ (suc (suc n)) ∣ x ∣ₕ) ∙ sym (rUnitₖ (suc (suc n)) (-[ (suc (suc n)) ]ₖ ∣ x ∣ₕ))) - (rUnit refl)) - - addIso : (n : ℕ) (x : EM G n) → Iso (EM G n) (EM G n) - Iso.fun (addIso n x) y = y +[ n ]ₖ x - Iso.inv (addIso n x) y = y -[ n ]ₖ x - Iso.rightInv (addIso n x) y = - sym (assocₖ n y (-[ n ]ₖ x) x) - ∙∙ cong (λ x → y +[ n ]ₖ x) (lCancelₖ n x) - ∙∙ rUnitₖ n y - Iso.leftInv (addIso n x) y = - sym (assocₖ n y x (-[ n ]ₖ x)) - ∙∙ cong (λ x → y +[ n ]ₖ x) (rCancelₖ n x) - ∙∙ rUnitₖ n y - - CODE : (n : ℕ) → EM G (suc (suc n)) → TypeOfHLevel ℓG (3 + n) - CODE n = - trElim (λ _ → isOfHLevelTypeOfHLevel (3 + n)) - λ { north → EM G (suc n) , hLevelEM G (suc n) - ; south → EM G (suc n) , hLevelEM G (suc n) - ; (merid a i) → fib n a i} - where - fib : (n : ℕ) → (a : EM-raw G (suc n)) - → Path (TypeOfHLevel _ (3 + n)) (EM G (suc n) , hLevelEM G (suc n)) (EM G (suc n) , hLevelEM G (suc n)) - fib zero a = Σ≡Prop (λ _ → isPropIsOfHLevel 3) - (isoToPath (addIso 1 a)) - fib (suc n) a = Σ≡Prop (λ _ → isPropIsOfHLevel (4 + n)) - (isoToPath (addIso (suc (suc n)) ∣ a ∣)) - - decode' : (n : ℕ) (x : EM G (suc (suc n))) → CODE n x .fst → 0ₖ (suc (suc n)) ≡ x - decode' n = - trElim (λ _ → isOfHLevelΠ (4 + n) - λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) - λ { north → f n - ; south → g n - ; (merid a i) → main a i} - where - f : (n : ℕ) → _ - f n = σ-EM' n - - g : (n : ℕ) → EM G (suc n) → ∣ ptS ∣ ≡ ∣ south ∣ - g n x = σ-EM' n x ∙ cong ∣_∣ₕ (merid ptS) - - lem₂ : (n : ℕ) (a x : _) - → Path (Path (EM G (suc (suc n))) _ _) - ((λ i → cong ∣_∣ₕ (σ-EM n x) i) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) ∙ (λ i → ∣ merid a i ∣ₕ)) - (g n (EM-raw→EM G (suc n) x)) - lem₂ zero a x = - cong (f zero x ∙_) - (cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid embase))) - ∙ cong-∙ ∣_∣ₕ (merid embase) (sym (merid a))) - ∙∙ sym (∙assoc _ _ _) - ∙∙ cong (cong ∣_∣ₕ (merid embase) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) - ∙ sym (rUnit _)) - lem₂ (suc n) a x = - cong (f (suc n) ∣ x ∣ ∙_) - ((cong (_∙ cong ∣_∣ₕ (merid a)) (cong (cong ∣_∣ₕ) (symDistr (merid a) (sym (merid north))) - ∙ cong-∙ ∣_∣ₕ (merid north) (sym (merid a))) - ∙∙ sym (∙assoc _ _ _) - ∙∙ cong (cong ∣_∣ₕ (merid north) ∙_) (lCancel (cong ∣_∣ₕ (merid a))) - ∙ sym (rUnit _))) - - lem : (n : ℕ) (x a : EM-raw G (suc n)) - → f n (transport (sym (cong (λ x → CODE n x .fst) (cong ∣_∣ₕ (merid a)))) (EM-raw→EM G (suc n) x)) - ≡ cong ∣_∣ₕ (σ-EM n x) ∙ sym (cong ∣_∣ₕ (σ-EM n a)) - lem zero x a = (λ i → cong ∣_∣ₕ (merid (transportRefl x i -[ 1 ]ₖ a) ∙ sym (merid embase))) - ∙∙ σ-EM'-hom zero x (-ₖ a) - ∙∙ cong (f zero x ∙_) (σ-EM'-ₖ zero a) - lem (suc n) x a = - cong (f (suc n)) (λ i → transportRefl ∣ x ∣ i -[ 2 + n ]ₖ ∣ a ∣) - ∙∙ σ-EM'-hom (suc n) ∣ x ∣ (-ₖ ∣ a ∣) - ∙∙ cong (f (suc n) ∣ x ∣ ∙_) (σ-EM'-ₖ (suc n) ∣ a ∣) - - main : (a : _) - → PathP (λ i → CODE n ∣ merid a i ∣ₕ .fst - → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ) (f n) (g n) - main a = - toPathP (funExt - (EM-elim _ (λ _ → isOfHLevelPathP (2 + (suc n)) (isOfHLevelTrunc (4 + n) _ _) _ _) - λ x → - ((λ i → transp (λ j → Path (EM G (2 + n)) ∣ ptS ∣ ∣ merid a (i ∨ j) ∣) - i (compPath-filler (lem n x a i) (cong ∣_∣ₕ (merid a)) i) )) - ∙∙ sym (∙assoc _ _ _) - ∙∙ lem₂ n a x)) - - encode' : (n : ℕ) (x : EM G (suc (suc n))) → 0ₖ (suc (suc n)) ≡ x → CODE n x .fst - encode' n x p = subst (λ x → CODE n x .fst) p (0ₖ (suc n)) - - decode'-encode' : (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → decode' n x (encode' n x p) ≡ p - decode'-encode' zero x = - J (λ x p → decode' 0 x (encode' 0 x p) ≡ p) - (σ-EM'-0ₖ 0) - decode'-encode' (suc n) x = - J (λ x p → decode' (suc n) x (encode' (suc n) x p) ≡ p) - (σ-EM'-0ₖ (suc n)) - - encode'-decode' : (n : ℕ) (x : _) → encode' n (0ₖ (suc (suc n))) (decode' n (0ₖ (suc (suc n))) x) ≡ x - encode'-decode' zero x = - cong (encode' zero (0ₖ 2)) (cong-∙ ∣_∣ₕ (merid x) (sym (merid embase))) - ∙∙ substComposite (λ x → CODE zero x .fst) - (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptS))) embase - ∙∙ (cong (subst (λ x₁ → CODE zero x₁ .fst) (λ i → ∣ merid embase (~ i) ∣ₕ)) - (λ i → lUnitₖ 1 (transportRefl x i) i) - ∙ (λ i → rUnitₖ 1 (transportRefl x i) i)) - encode'-decode' (suc n) = - trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) - λ x → cong (encode' (suc n) (0ₖ (3 + n))) (cong-∙ ∣_∣ₕ (merid x) (sym (merid north))) - ∙∙ substComposite (λ x → CODE (suc n) x .fst) - (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptS))) (0ₖ (2 + n)) - ∙∙ cong (subst (λ x₁ → CODE (suc n) x₁ .fst) (λ i → ∣ merid ptS (~ i) ∣ₕ)) - (λ i → lUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) - ∙ (λ i → rUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) - - EM∙' : (n : ℕ) → Pointed _ - EM∙' n = EM G n , 0ₖ n - - Iso-EM-ΩEM+1 : (n : ℕ) → Iso (EM G n) (typ (Ω (EM∙' (suc n)))) - Iso-EM-ΩEM+1 zero = invIso (ΩEM₁Iso (G *)) - Iso.fun (Iso-EM-ΩEM+1 (suc zero)) = decode' 0 (0ₖ 2) - Iso.inv (Iso-EM-ΩEM+1 (suc zero)) = encode' 0 (0ₖ 2) - Iso.rightInv (Iso-EM-ΩEM+1 (suc zero)) = decode'-encode' 0 (0ₖ 2) - Iso.leftInv (Iso-EM-ΩEM+1 (suc zero)) = encode'-decode' 0 - Iso.fun (Iso-EM-ΩEM+1 (suc (suc n))) = decode' (suc n) (0ₖ (3 + n)) - Iso.inv (Iso-EM-ΩEM+1 (suc (suc n))) = encode' (suc n) (0ₖ (3 + n)) - Iso.rightInv (Iso-EM-ΩEM+1 (suc (suc n))) = decode'-encode' (suc n) (0ₖ (3 + n)) - Iso.leftInv (Iso-EM-ΩEM+1 (suc (suc n))) = encode'-decode' (suc n) - - EM→ΩEM+1 : (n : ℕ) → EM G n → typ (Ω (EM∙' (suc n))) - EM→ΩEM+1 n = Iso.fun (Iso-EM-ΩEM+1 n) - - ΩEM+1→EM : (n : ℕ) → typ (Ω (EM∙' (suc n))) → EM G n - ΩEM+1→EM n = Iso.inv (Iso-EM-ΩEM+1 n) - - EM→ΩEM+1-hom : (n : ℕ) (x y : EM G n) - → EM→ΩEM+1 n (x +[ n ]ₖ y) ≡ EM→ΩEM+1 n x ∙ EM→ΩEM+1 n y - EM→ΩEM+1-hom zero x y = emloop-comp _ x y - EM→ΩEM+1-hom (suc zero) x y = σ-EM'-hom zero x y - EM→ΩEM+1-hom (suc (suc n)) x y = σ-EM'-hom (suc n) x y - - ΩEM+1→EM-hom : (n : ℕ) → (p q : typ (Ω (EM∙' (suc n)))) - → ΩEM+1→EM n (p ∙ q) ≡ (ΩEM+1→EM n p) +[ n ]ₖ (ΩEM+1→EM n q) - ΩEM+1→EM-hom n = - morphLemmas.isMorphInv (λ x y → x +[ n ]ₖ y) (_∙_) (EM→ΩEM+1 n) - (EM→ΩEM+1-hom n) (ΩEM+1→EM n) - (Iso.rightInv (Iso-EM-ΩEM+1 n)) (Iso.leftInv (Iso-EM-ΩEM+1 n)) - - EM→ΩEM+1-0ₖ : (n : ℕ) → EM→ΩEM+1 n (0ₖ n) ≡ refl - EM→ΩEM+1-0ₖ zero = emloop-1g _ - EM→ΩEM+1-0ₖ (suc zero) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) - EM→ΩEM+1-0ₖ (suc (suc n)) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) - - ΩEM+1→EM-refl : (n : ℕ) → ΩEM+1→EM n refl ≡ 0ₖ n - ΩEM+1→EM-refl zero = transportRefl 0g - ΩEM+1→EM-refl (suc zero) = refl - ΩEM+1→EM-refl (suc (suc n)) = refl diff --git a/Cubical/ZCohomology/RingStructure/CupProduct.agda b/Cubical/ZCohomology/RingStructure/CupProduct.agda index dd9da8180..b2a0f2697 100644 --- a/Cubical/ZCohomology/RingStructure/CupProduct.agda +++ b/Cubical/ZCohomology/RingStructure/CupProduct.agda @@ -31,6 +31,9 @@ suc a +' suc b = 2 + (a + b) +'≡+ (suc n) zero = cong suc (sym (+-comm n zero)) +'≡+ (suc n) (suc m) = cong suc (sym (+-suc n m)) ++'-comm : (n m : ℕ) → n +' m ≡ m +' n ++'-comm n m = +'≡+ n m ∙∙ +-comm n m ∙∙ sym (+'≡+ m n) + -- Cup product with one integer (K₀) argument _·₀_ : {n : ℕ} (m : ℤ) → coHomK n → coHomK n _·₀_ {n = n} (pos zero) x = 0ₖ _ diff --git a/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda index 059b5af70..053d805b5 100644 --- a/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda +++ b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda @@ -45,10 +45,6 @@ private → f m (subst A p a) ≡ subst B p (f n a) natTranspLem {A = A} {B = B} a f p = sym (substCommSlice A B f p a) -+'-comm : (n m : ℕ) → n +' m ≡ m +' n -+'-comm n m = +'≡+ n m ∙∙ +-comm n m ∙∙ sym (+'≡+ m n) - -private transp0₁ : (n : ℕ) → subst coHomK (+'-comm 1 (suc n)) (0ₖ _) ≡ 0ₖ _ transp0₁ zero = refl transp0₁ (suc n) = refl diff --git a/Cubical/ZCohomology/RingStructure/RingLaws.agda b/Cubical/ZCohomology/RingStructure/RingLaws.agda index 99b719f6f..93d62cc9c 100644 --- a/Cubical/ZCohomology/RingStructure/RingLaws.agda +++ b/Cubical/ZCohomology/RingStructure/RingLaws.agda @@ -100,12 +100,14 @@ private snd (⌣ₖ-distrFun2 n m x y) = cong₂ _+ₖ_ (0ₖ-⌣ₖ m n x) (0ₖ-⌣ₖ m n y) ∙ rUnitₖ _ _ - leftDistr-⌣ₖ· : (n m : ℕ) (x y : coHomK (suc n)) → ⌣ₖ-distrFun (suc n) (suc m) x y ≡ ⌣ₖ-distrFun2 (suc n) (suc m) x y + leftDistr-⌣ₖ· : (n m : ℕ) (x y : coHomK (suc n)) + → ⌣ₖ-distrFun (suc n) (suc m) x y ≡ ⌣ₖ-distrFun2 (suc n) (suc m) x y leftDistr-⌣ₖ· n m = elim2 (λ _ _ → isOfHLevelSuc (2 + n) (hLevHelp n m _ _)) main where - hLevHelp : (n m : ℕ) (x y : _) → isOfHLevel (2 + n) ( ⌣ₖ-distrFun (suc n) (suc m) x y ≡ ⌣ₖ-distrFun2 (suc n) (suc m) x y) + hLevHelp : (n m : ℕ) (x y : _) + → isOfHLevel (2 + n) ( ⌣ₖ-distrFun (suc n) (suc m) x y ≡ ⌣ₖ-distrFun2 (suc n) (suc m) x y) hLevHelp n m x y = (subst (isOfHLevel (3 + n)) (λ i → (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (suc (+-comm n m i))))) @@ -228,7 +230,8 @@ private snd (⌣ₖ-distrFun2-r n m x y) = cong₂ _+ₖ_ (⌣ₖ-0ₖ n m x) (⌣ₖ-0ₖ n m y) ∙ rUnitₖ _ _ - rightDistr-⌣ₖ· : (n m : ℕ) (x y : coHomK (suc n)) → ⌣ₖ-distrFun-r (suc n) (suc m) x y ≡ ⌣ₖ-distrFun2-r (suc n) (suc m) x y + rightDistr-⌣ₖ· : (n m : ℕ) (x y : coHomK (suc n)) + → ⌣ₖ-distrFun-r (suc n) (suc m) x y ≡ ⌣ₖ-distrFun2-r (suc n) (suc m) x y rightDistr-⌣ₖ· n m = elim2 (λ _ _ → isOfHLevelPath (3 + n) (isOfHLevel↑∙ (suc n) m) _ _) main From ad7687af7dab29754e7f7e9d9c3ba072ad9e4bd5 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 20 Aug 2021 14:12:58 +0200 Subject: [PATCH 23/30] stuff --- Cubical/Foundations/Pointed/Homogeneous.agda | 8 ++--- Cubical/HITs/SmashProduct/Base.agda | 35 +------------------- 2 files changed, 3 insertions(+), 40 deletions(-) diff --git a/Cubical/Foundations/Pointed/Homogeneous.agda b/Cubical/Foundations/Pointed/Homogeneous.agda index a6de4f062..7bb614d26 100644 --- a/Cubical/Foundations/Pointed/Homogeneous.agda +++ b/Cubical/Foundations/Pointed/Homogeneous.agda @@ -6,14 +6,13 @@ Portions of this file adapted from Nicolai Kraus' code here: https://bitbucket.org/nicolaikraus/agda/src/e30d70c72c6af8e62b72eefabcc57623dd921f04/trunc-inverse.lagda -} -{-# OPTIONS --safe --experimental-lossy-unification #-} +{-# OPTIONS --safe #-} module Cubical.Foundations.Pointed.Homogeneous where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Transport open import Cubical.Foundations.Path open import Cubical.Data.Sigma open import Cubical.Data.Empty as ⊥ @@ -27,14 +26,11 @@ open import Cubical.Structures.Pointed isHomogeneous : ∀ {ℓ} → Pointed ℓ → Type (ℓ-suc ℓ) isHomogeneous {ℓ} (A , x) = ∀ y → Path (Pointed ℓ) (A , x) (A , y) -_→∙Dep_ : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : typ A → Pointed ℓ') → Type _ -A →∙Dep B = Σ[ f ∈ ((x : typ A) → B x .fst) ] f (snd A) ≡ snd (B (snd A)) - -- Pointed functions into a homogeneous type are equal as soon as they are equal -- as unpointed functions →∙Homogeneous≡ : ∀ {ℓ ℓ'} {A∙ : Pointed ℓ} {B∙ : Pointed ℓ'} {f∙ g∙ : A∙ →∙ B∙} (h : isHomogeneous B∙) → f∙ .fst ≡ g∙ .fst → f∙ ≡ g∙ -→∙Homogeneous≡ {A∙ = A∙@(A , a₀)} {B∙@(B , b)} {f∙@(_ , f₀)} {g∙@(_ , g₀)} h p = +→∙Homogeneous≡ {A∙ = A∙@(_ , a₀)} {B∙@(B , _)} {f∙@(_ , f₀)} {g∙@(_ , g₀)} h p = subst (λ Q∙ → PathP (λ i → A∙ →∙ Q∙ i) f∙ g∙) (sym (flipSquare fix)) badPath where badPath : PathP (λ i → A∙ →∙ (B , (sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) f∙ g∙ diff --git a/Cubical/HITs/SmashProduct/Base.agda b/Cubical/HITs/SmashProduct/Base.agda index 7dfe128f7..59cd41698 100644 --- a/Cubical/HITs/SmashProduct/Base.agda +++ b/Cubical/HITs/SmashProduct/Base.agda @@ -53,38 +53,6 @@ SmashPt A B = (Smash A B , basel) SmashPtProj : (A : Pointed ℓ) (B : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') SmashPtProj A B = Smash A B , (proj (snd A) (snd B)) -private - rCancelPathP : ∀ {ℓ} {A B : Type ℓ} {x y : A} (p : x ≡ y) (f : A → B) - → PathP (λ i → cong-∙ f p (sym p) i ≡ refl) - (cong (cong f) (rCancel p)) - (rCancel (cong f p)) - rCancelPathP {x = x} {y = y} p f i j k = - hcomp (λ r → λ { (i = i0) → f (rCancel-filler p r j k) - ; (j = i1) → f x - ; (k = i0) → f x - ; (k = i1) → f (p (~ r ∧ ~ j))}) - (f (p (k ∧ ~ j))) - - mainGen : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) - → (P : refl ≡ p) - → Cube (λ j k → (cong (p ∙_) (cong sym (sym P)) ∙ sym (rUnit _)) k j) - refl - refl - refl - (rCancel p) - (sym P) - mainGen {x = x} p = J (λ p P → Cube (λ j k → (cong (p ∙_) (cong sym (sym P)) ∙ sym (rUnit _)) k j) - refl - refl - refl - (rCancel p) - (sym P)) - (transport (λ i → Cube (λ j k → (lUnit (sym (rUnit (refl {x = x}))) i) k j) - l l l (rCancel (λ _ → x)) l) - λ i j k → rCancel (λ _ → x) (k ∨ i) j) - where - l = refl {x = refl {x = x}} - --- Alternative definition i∧ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → A ⋁ B → (typ A) × (typ B) @@ -131,8 +99,7 @@ Smash→⋀ (proj x y) = inr (x , y) Smash→⋀ (gluel a i) = push (inl a) (~ i) Smash→⋀ (gluer b i) = push (inr b) (~ i) -{- associativity maps for smash produts. Proof pretty much direcly translated from - https://github.com/ecavallo/redtt/blob/master/library/pointed/smash.red -} +{- associativity maps for smash produts. Proof pretty much direcly translated from https://github.com/ecavallo/redtt/blob/master/library/pointed/smash.red -} private pivotl : (b b' : typ B) → Path (Smash A B) (proj (snd A) b) (proj (snd A) b') From c503745b00a57bb9e276feae9240a458b38ab3a0 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 20 Aug 2021 14:40:35 +0200 Subject: [PATCH 24/30] clean --- .../Algebra/Group/EilenbergMacLane/GroupStructure.agda | 7 +++---- Cubical/Foundations/GroupoidLaws.agda | 9 --------- 2 files changed, 3 insertions(+), 13 deletions(-) diff --git a/Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda b/Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda index a1688dc55..bac2bbcb0 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda @@ -8,7 +8,7 @@ open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.Properties open import Cubical.Algebra.AbGroup.Base -open import Cubical.Data.Nat hiding (_·_) +open import Cubical.Data.Nat open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism @@ -16,8 +16,9 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) open import Cubical.Foundations.HLevels open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Path -open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Loopspace open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec; elim to trElim) open import Cubical.HITs.EilenbergMacLane1 @@ -26,7 +27,6 @@ open import Cubical.HITs.Truncation open import Cubical.HITs.Susp open import Cubical.Functions.Morphism -open import Cubical.Foundations.Path private variable ℓ : Level @@ -159,7 +159,6 @@ module _ {G : AbGroup ℓ} where (λ x → rUnitₖ (2 + n) ∣ x ∣) refl) - open import Cubical.Homotopy.Loopspace cong₂+₁ : (p q : typ (Ω (EM∙ G 1))) → cong₂ (λ x y → x +[ 1 ]ₖ y) p q ≡ p ∙ q cong₂+₁ p q = diff --git a/Cubical/Foundations/GroupoidLaws.agda b/Cubical/Foundations/GroupoidLaws.agda index 7da8f33c6..8440e8e10 100644 --- a/Cubical/Foundations/GroupoidLaws.agda +++ b/Cubical/Foundations/GroupoidLaws.agda @@ -201,15 +201,6 @@ cong-∙ : ∀ {B : Type ℓ} (f : A → B) (p : x ≡ y) (q : y ≡ z) → cong f (p ∙ q) ≡ (cong f p) ∙ (cong f q) cong-∙ f p q = cong-∙∙ f refl p q -cong-compPathP : ∀ {B : A → Type ℓ} (f : (x : A) → B x) (p : x ≡ y) (q : y ≡ z) - → cong f (p ∙ q) ≡ compPathP' {B = B} (cong f p) (cong f q) -cong-compPathP {x = x} {B = B} f p q i j = - comp (λ k → B (compPath-filler p q k j)) - (λ k → λ { (j = i0) → f x - ; (j = i1) → f (q k) - ; (i = i0) → f (compPath-filler p q k j) }) - (f (p j)) - hcomp-unique : ∀ {ℓ} {A : Type ℓ} {φ} → (u : I → Partial φ A) → (u0 : A [ φ ↦ u i0 ]) → (h2 : ∀ i → A [ (φ ∨ ~ i) ↦ (\ { (φ = i1) → u i 1=1; (i = i0) → outS u0}) ]) → (hcomp u (outS u0) ≡ outS (h2 i1)) [ φ ↦ (\ { (φ = i1) → (\ i → u i1 1=1)}) ] From 312622b0e0a430bd37b65006a01cc0c91476c768 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 20 Aug 2021 14:47:28 +0200 Subject: [PATCH 25/30] whitespace;) --- Cubical/Algebra/Group/EilenbergMacLane/Base.agda | 4 ++-- Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda | 2 +- Cubical/Foundations/HLevels.agda | 2 +- Cubical/HITs/EilenbergMacLane1/Base.agda | 2 +- 4 files changed, 5 insertions(+), 5 deletions(-) diff --git a/Cubical/Algebra/Group/EilenbergMacLane/Base.agda b/Cubical/Algebra/Group/EilenbergMacLane/Base.agda index 0fde35c93..87b24a57e 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/Base.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --no-import-sorts --safe --experimental-lossy-unification #-} +{-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.Algebra.Group.EilenbergMacLane.Base where @@ -72,7 +72,7 @@ EM∙ : (G : AbGroup ℓ) (n : ℕ) → Pointed ℓ EM∙ G n = EM G n , (0ₖ n) EM-raw∙ : (G : AbGroup ℓ) (n : ℕ) → Pointed ℓ -EM-raw∙ G n = EM-raw G n , ptS +EM-raw∙ G n = EM-raw G n , ptS hLevelEM : (G : AbGroup ℓ) (n : ℕ) → isOfHLevel (2 + n) (EM G n) hLevelEM G zero = AbGroupStr.is-set (snd G) diff --git a/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda b/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda index 576baefd9..afd688a16 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda @@ -126,7 +126,7 @@ private ≡ λ i → hcomp (λ k → λ { (i = i0) → p (~ k) ; (i = i1) → llem₁ k}) (g (merid y i)) - mainP y = + mainP y = wedgeConFun'ᵣ G H n (λ _ _ → isOfHLevelPathP' (suc (n + 1)) (hLev _ _) _ _) (λ x i → transp (λ j → A (merid ptS (i ∧ j)) x) (~ i) (f x)) diff --git a/Cubical/Foundations/HLevels.agda b/Cubical/Foundations/HLevels.agda index 07b41e1fc..c1d7d730e 100644 --- a/Cubical/Foundations/HLevels.agda +++ b/Cubical/Foundations/HLevels.agda @@ -643,7 +643,7 @@ isOfHLevelΣ' {A = A} {B = B} (suc (suc n)) Alvl Blvl (w , y) (x , z) → {a₁₀ a₁₁ : Σ A B} (p₁₋ : a₁₀ ≡ a₁₁) → (p₋₀ : a₀₀ ≡ a₁₀) (p₋₁ : a₀₁ ≡ a₁₁) → Square (cong fst p₀₋) (cong fst p₁₋) (cong fst p₋₀) (cong fst p₋₁) - → Square p₀₋ p₁₋ p₋₀ p₋₁ + → Square p₀₋ p₁₋ p₋₀ p₋₁ fst (Σ≡Set set p₀₋ p₁₋ p₋₀ p₋₁ sq i j) = sq i j snd (Σ≡Set {B = B} set p₀₋ p₁₋ p₋₀ p₋₁ sq i j) = sq2 i j where diff --git a/Cubical/HITs/EilenbergMacLane1/Base.agda b/Cubical/HITs/EilenbergMacLane1/Base.agda index 957d4bd86..c9c76557a 100644 --- a/Cubical/HITs/EilenbergMacLane1/Base.agda +++ b/Cubical/HITs/EilenbergMacLane1/Base.agda @@ -10,7 +10,7 @@ module Cubical.HITs.EilenbergMacLane1.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.GroupoidLaws - renaming (assoc to assoc∙) + renaming (assoc to assoc∙) open import Cubical.Algebra.Group.Base private From 4ca13fd862452d5e93a40c7cd81ad056084dfe07 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sun, 9 Jan 2022 22:42:55 +0100 Subject: [PATCH 26/30] wups --- Cubical/Algebra/Group/MorphismProperties.agda | 16 ---------------- 1 file changed, 16 deletions(-) diff --git a/Cubical/Algebra/Group/MorphismProperties.agda b/Cubical/Algebra/Group/MorphismProperties.agda index c5eff5082..f8a145b1c 100644 --- a/Cubical/Algebra/Group/MorphismProperties.agda +++ b/Cubical/Algebra/Group/MorphismProperties.agda @@ -199,22 +199,6 @@ compGroupEquiv : GroupEquiv F G → GroupEquiv G H → GroupEquiv F H fst (compGroupEquiv f g) = compEquiv (fst f) (fst g) snd (compGroupEquiv f g) = isGroupHomComp (_ , f .snd) (_ , g .snd) -isGroupHomInv : (f : GroupEquiv G H) → IsGroupHom (H .snd) (invEq (fst f)) (G .snd) -isGroupHomInv {G = G} {H = H} f = makeIsGroupHom λ h h' → - isInj-f _ _ - (f' (g (h ⋆² h')) ≡⟨ secEq (fst f) _ ⟩ - (h ⋆² h') ≡⟨ sym (cong₂ _⋆²_ (secEq (fst f) h) (secEq (fst f) h')) ⟩ - (f' (g h) ⋆² f' (g h')) ≡⟨ sym (pres· (snd f) _ _) ⟩ - f' (g h ⋆¹ g h') ∎) - where - f' = fst (fst f) - _⋆¹_ = _·_ (snd G) - _⋆²_ = _·_ (snd H) - g = invEq (fst f) - - isInj-f : (x y : ⟨ G ⟩) → f' x ≡ f' y → x ≡ y - isInj-f x y = invEq (_ , isEquiv→isEmbedding (snd (fst f)) x y) - invGroupEquiv : GroupEquiv G H → GroupEquiv H G fst (invGroupEquiv f) = invEquiv (fst f) snd (invGroupEquiv f) = isGroupHomInv f From c4b751701da1a9b13f5dfc0983a9d8358c5d8e0f Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sun, 9 Jan 2022 23:55:50 +0100 Subject: [PATCH 27/30] wups2 --- .../RingStructure/GradedCommutativity.agda | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda index 00d86a4dc..1257f8cf0 100644 --- a/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda +++ b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda @@ -45,13 +45,13 @@ natTranspLem : ∀ {ℓ} {A B : ℕ → Type ℓ} {n m : ℕ} (a : A n) → f m (subst A p a) ≡ subst B p (f n a) natTranspLem {A = A} {B = B} a f p = sym (substCommSlice A B f p a) - transp0₁ : (n : ℕ) → subst coHomK (+'-comm 1 (suc n)) (0ₖ _) ≡ 0ₖ _ - transp0₁ zero = refl - transp0₁ (suc n) = refl +transp0₁ : (n : ℕ) → subst coHomK (+'-comm 1 (suc n)) (0ₖ _) ≡ 0ₖ _ +transp0₁ zero = refl +transp0₁ (suc n) = refl - transp0₂ : (n m : ℕ) → subst coHomK (+'-comm (suc (suc n)) (suc m)) (0ₖ _) ≡ 0ₖ _ - transp0₂ n zero = refl - transp0₂ n (suc m) = refl +transp0₂ : (n m : ℕ) → subst coHomK (+'-comm (suc (suc n)) (suc m)) (0ₖ _) ≡ 0ₖ _ +transp0₂ n zero = refl +transp0₂ n (suc m) = refl -- Recurring expressions private From cb38b9ffb0b383a73ba7fdf5d2be880b5f426959 Mon Sep 17 00:00:00 2001 From: ecavallo Date: Mon, 10 Jan 2022 17:02:26 +0100 Subject: [PATCH 28/30] derive flip --- Cubical/Algebra/AbGroup/TensorProduct.agda | 116 ++++++++++----------- 1 file changed, 58 insertions(+), 58 deletions(-) diff --git a/Cubical/Algebra/AbGroup/TensorProduct.agda b/Cubical/Algebra/AbGroup/TensorProduct.agda index 48e494438..dc8cec005 100644 --- a/Cubical/Algebra/AbGroup/TensorProduct.agda +++ b/Cubical/Algebra/AbGroup/TensorProduct.agda @@ -62,8 +62,6 @@ module _ (AGr : AbGroup ℓ) (BGr : AbGroup ℓ') where linl : (x y : A) (z : B) → (x +A y) ⊗ z ≡ ((x ⊗ z) +⊗ (y ⊗ z)) linr : (x : A) (y z : B) → x ⊗ (y +B z) ≡ ((x ⊗ y) +⊗ (x ⊗ z)) - flip : (x : A) (b : B) → x ⊗ (-B b) ≡ ((-A x) ⊗ b) -- needed ? - ⊗squash : isSet _⨂₁_ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where @@ -115,9 +113,6 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where ⊗elimProp {C = C} p f g (linr x y z i) = isOfHLevel→isOfHLevelDep 1 {B = C} p (f x (y +B z)) (g (x ⊗ y) (x ⊗ z) (f x y) (f x z)) (linr x y z) i - ⊗elimProp {C = C} p f g (flip x b i) = - isOfHLevel→isOfHLevelDep 1 {B = C} p - (f x (-B b)) (f (-A x) b) (flip x b) i ⊗elimProp {C = C} p f g (⊗squash x y q r i j) = isOfHLevel→isOfHLevelDep 2 {B = C} (λ x → isProp→isSet (p x)) _ _ (λ j → ⊗elimProp p f g (q j)) (λ j → ⊗elimProp p f g (r j)) @@ -140,7 +135,6 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where ∙ comm strA (-A y) (-A x)) i ⊗ z) ∙ linl (-A x) (-A y) z) i -⊗ (linr x y z i) = linr (-A x) y z i - -⊗ (flip x b i) = flip (-A x) b i -⊗ (⊗squash x y p q i j) = ⊗squash _ _ (λ j → -⊗ (p j)) (λ j → -⊗ (q j)) i j @@ -184,22 +178,45 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where lCancelPrim : (x : B) → (0A ⊗ x) ≡ 0⊗ lCancelPrim x = - (λ i → rid strA 0A (~ i) ⊗ x) - ∙∙ linl 0A 0A x - ∙∙ cong ((0A ⊗ x) +⊗_) (cong (_⊗ x) (sym (GroupTheory.inv1g (AbGroup→Group AGr))) - ∙ sym (flip 0A x)) - ∙∙ sym (linr 0A x (-B x)) - ∙∙ (λ i → 0A ⊗ (invr strB x i)) + sym (⊗rUnit _) + ∙∙ cong ((0A ⊗ x) +⊗_) (sym cancelLem) + ∙∙ ⊗assoc (0A ⊗ x) (0A ⊗ x) (0A ⊗ (-B x)) + ∙∙ cong (_+⊗ (0A ⊗ (-B x))) inner + ∙∙ cancelLem + where + cancelLem : (0A ⊗ x) +⊗ (0A ⊗ (-B x)) ≡ 0⊗ + cancelLem = sym (linr 0A x (-B x)) ∙ (λ i → 0A ⊗ invr strB x i) + + inner : (0A ⊗ x) +⊗ (0A ⊗ x) ≡ 0A ⊗ x + inner = sym (linl 0A 0A x) ∙ (λ i → rid strA 0A i ⊗ x) rCancelPrim : (x : A) → (x ⊗ 0B) ≡ 0⊗ rCancelPrim x = - (λ i → x ⊗ rid strB 0B (~ i)) - ∙∙ linr x 0B 0B - ∙∙ cong ((x ⊗ 0B) +⊗_) - (cong (x ⊗_) - (sym (GroupTheory.inv1g (AbGroup→Group BGr))) ∙ flip x 0B) - ∙∙ sym (linl x (-A x) 0B) - ∙∙ (λ i → (invr strA x i) ⊗ 0B) + sym (⊗rUnit _) + ∙∙ cong ((x ⊗ 0B) +⊗_) (sym cancelLem) + ∙∙ ⊗assoc (x ⊗ 0B) (x ⊗ 0B) ((-A x) ⊗ 0B) + ∙∙ cong (_+⊗ ((-A x) ⊗ 0B)) inner + ∙∙ cancelLem + where + cancelLem : (x ⊗ 0B) +⊗ ((-A x) ⊗ 0B) ≡ 0⊗ + cancelLem = sym (linl x (-A x) 0B) ∙ (λ i → invr strA x i ⊗ 0B) + + inner : (x ⊗ 0B) +⊗ (x ⊗ 0B) ≡ x ⊗ 0B + inner = sym (linr x 0B 0B) ∙ (λ i → x ⊗ rid strB 0B i) + + flip : (x : A) (b : B) → x ⊗ (-B b) ≡ ((-A x) ⊗ b) + flip x b = + sym (⊗rUnit _) + ∙ cong ((x ⊗ (-B b)) +⊗_) (sym cancelLemA) + ∙ ⊗assoc (x ⊗ (-B b)) (x ⊗ b) ((-A x) ⊗ b) + ∙ cong (_+⊗ ((-A x) ⊗ b)) cancelLemB + ∙ ⊗lUnit _ + where + cancelLemA : (x ⊗ b) +⊗ ((-A x) ⊗ b) ≡ 0⊗ + cancelLemA = sym (linl x (-A x) b) ∙ (λ i → invr strA x i ⊗ b) ∙ lCancelPrim b + + cancelLemB : (x ⊗ (-B b)) +⊗ (x ⊗ b) ≡ 0⊗ + cancelLemB = sym (linr x (-B b) b) ∙ (λ i → x ⊗ invl strB b i) ∙ rCancelPrim x ⊗rCancel : (x : AGr ⨂₁ BGr) → (x +⊗ -⊗ x) ≡ 0⊗ ⊗rCancel = @@ -257,36 +274,33 @@ module _ {ℓ ℓ' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} where → f (x +A y , b) ≡ _+G_ (snd C) (f (x , b)) (f (y , b))) → (linR : (a : fst A) (x y : fst B) → f (a , x +B y) ≡ _+G_ (snd C) (f (a , x)) (f (a , y))) - → (fl : (x : fst A) (y : fst B) - → f (x , -G (snd B) y) ≡ f ((-G (snd A) x) , y)) → A ⨂₁ B → fst C - ⨂→AbGroup-elim C f p linL linR fl (a ⊗ b) = f (a , b) - ⨂→AbGroup-elim C f p linL linR fl (x +⊗ x₁) = - _+G_ (snd C) (⨂→AbGroup-elim C f p linL linR fl x) - (⨂→AbGroup-elim C f p linL linR fl x₁) - ⨂→AbGroup-elim C f p linL linR fl (⊗comm x x₁ i) = - comm (snd C) (⨂→AbGroup-elim C f p linL linR fl x) - (⨂→AbGroup-elim C f p linL linR fl x₁) i - ⨂→AbGroup-elim C f p linL linR fl (⊗assoc x x₁ x₂ i) = - assoc (snd C) (⨂→AbGroup-elim C f p linL linR fl x) - (⨂→AbGroup-elim C f p linL linR fl x₁) - (⨂→AbGroup-elim C f p linL linR fl x₂) i - ⨂→AbGroup-elim C f p linL linR fl (⊗lUnit x i) = - (cong (λ y → (snd C +G y) (⨂→AbGroup-elim C f p linL linR fl x)) p - ∙ (lid (snd C) (⨂→AbGroup-elim C f p linL linR fl x))) i - ⨂→AbGroup-elim C f p linL linR fl (linl x y z i) = linL x y z i - ⨂→AbGroup-elim C f p linL linR fl (linr x y z i) = linR x y z i - ⨂→AbGroup-elim C f p linL linR fl (flip x b i) = fl x b i - ⨂→AbGroup-elim C f p linL linR fl (⊗squash x x₁ x₂ y i i₁) = + ⨂→AbGroup-elim C f p linL linR (a ⊗ b) = f (a , b) + ⨂→AbGroup-elim C f p linL linR (x +⊗ x₁) = + _+G_ (snd C) (⨂→AbGroup-elim C f p linL linR x) + (⨂→AbGroup-elim C f p linL linR x₁) + ⨂→AbGroup-elim C f p linL linR (⊗comm x x₁ i) = + comm (snd C) (⨂→AbGroup-elim C f p linL linR x) + (⨂→AbGroup-elim C f p linL linR x₁) i + ⨂→AbGroup-elim C f p linL linR (⊗assoc x x₁ x₂ i) = + assoc (snd C) (⨂→AbGroup-elim C f p linL linR x) + (⨂→AbGroup-elim C f p linL linR x₁) + (⨂→AbGroup-elim C f p linL linR x₂) i + ⨂→AbGroup-elim C f p linL linR (⊗lUnit x i) = + (cong (λ y → (snd C +G y) (⨂→AbGroup-elim C f p linL linR x)) p + ∙ (lid (snd C) (⨂→AbGroup-elim C f p linL linR x))) i + ⨂→AbGroup-elim C f p linL linR (linl x y z i) = linL x y z i + ⨂→AbGroup-elim C f p linL linR (linr x y z i) = linR x y z i + ⨂→AbGroup-elim C f p linL linR (⊗squash x x₁ x₂ y i i₁) = is-set (snd C) _ _ - (λ i → ⨂→AbGroup-elim C f p linL linR fl (x₂ i)) - (λ i → ⨂→AbGroup-elim C f p linL linR fl (y i)) i i₁ + (λ i → ⨂→AbGroup-elim C f p linL linR (x₂ i)) + (λ i → ⨂→AbGroup-elim C f p linL linR (y i)) i i₁ ⨂→AbGroup-elim-hom : ∀ {ℓ} (C : AbGroup ℓ) - → (f : (fst A × fst B) → fst C) (linL : _) (linR : _) (fl : _) (p : _) + → (f : (fst A × fst B) → fst C) (linL : _) (linR : _) (p : _) → AbGroupHom (A ⨂ B) C - fst (⨂→AbGroup-elim-hom C f linL linR fl p) = ⨂→AbGroup-elim C f p linL linR fl - snd (⨂→AbGroup-elim-hom C f linL linR fl p) = + fst (⨂→AbGroup-elim-hom C f linL linR p) = ⨂→AbGroup-elim C f p linL linR + snd (⨂→AbGroup-elim-hom C f linL linR p) = makeIsGroupHom (λ x y → refl) @@ -364,8 +378,6 @@ module UP (AGr : AbGroup ℓ) (BGr : AbGroup ℓ') where F (⊗lUnit x i) = (cong (λ y → _+G_ (snd C) y (F x)) lem ∙ lid (snd C) (F x)) i F (linl x y z i) = fst (IsGroupHom.pres· p x y i) z F (linr x y z i) = IsGroupHom.pres· (f x .snd) y z i - F (flip x b i) = (IsGroupHom.presinv (f x .snd) b - ∙ sym (funExt⁻ (cong fst (IsGroupHom.presinv p x)) b)) i F (⊗squash x x₁ x₂ y i i₁) = is-set (snd C) (F x) (F x₁) (λ i → F (x₂ i)) (λ i → F (y i)) i i₁ @@ -394,7 +406,6 @@ commFun (⊗assoc x x₁ x₂ i) = ⊗assoc (commFun x) (commFun x₁) (commFun commFun (⊗lUnit x i) = ⊗lUnit (commFun x) i commFun (linl x y z i) = linr z x y i commFun (linr x y z i) = linl y z x i -commFun (flip x b i) = flip b x (~ i) commFun (⊗squash x x₁ x₂ y i i₁) = ⊗squash (commFun x) (commFun x₁) (λ i → commFun (x₂ i)) (λ i → commFun (y i)) i i₁ @@ -427,7 +438,6 @@ module _ {ℓ ℓ' ℓ'' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} {C : AbGr f c = ⨂→AbGroup-elim-hom (A ⨂ (B ⨂ C)) ((λ ab → fst ab ⊗ (snd ab ⊗ c))) (λ x y b → linl x y (b ⊗ c)) (λ x y b → (λ i → x ⊗ (linl y b c i)) ∙ linr x (y ⊗ c) (b ⊗ c)) - (λ _ _ → flip _ _) (λ i → 0g (snd A) ⊗ (lCancelPrim c i)) assocHom : AbGroupHom ((A ⨂ B) ⨂ C) (A ⨂ (B ⨂ C)) @@ -443,11 +453,6 @@ module _ {ℓ ℓ' ℓ'' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} {C : AbGr _+⊗_ ⊗assoc ⊗comm ∙∙ cong₂ _+⊗_ (sym (IsGroupHom.pres· (snd (f x)) a b)) (IsGroupHom.pres· (snd (f y)) a b)) - (⊗elimProp (λ _ → isPropΠ λ _ → ⊗squash _ _) - (λ a b c → (λ i → a ⊗ (flip b c i)) - ∙ flip a (b ⊗ c)) - (λ x y ind1 ind2 c → cong₂ _+⊗_ (ind1 c) (ind2 c) - ∙ IsGroupHom.pres· (snd (f c)) (-⊗ x) (-⊗ y))) refl module _ {ℓ ℓ' ℓ'' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} {C : AbGroup ℓ''} where @@ -458,7 +463,6 @@ module _ {ℓ ℓ' ℓ'' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} {C : AbGr (λ bc → (a ⊗ fst bc) ⊗ snd bc) (λ x y b → (λ i → (linr a x y i) ⊗ b) ∙ linl (a ⊗ x) (a ⊗ y) b) (λ x y b → linr (a ⊗ x) y b) - (λ b c → flip (a ⊗ b) c ∙ λ i → flip a b (~ i) ⊗ c) λ i → rCancelPrim a i ⊗ (0g (snd C)) assocHom⁻ : AbGroupHom (A ⨂ (B ⨂ C)) ((A ⨂ B) ⨂ C) @@ -471,10 +475,6 @@ module _ {ℓ ℓ' ℓ'' : Level} {A : AbGroup ℓ} {B : AbGroup ℓ'} {C : AbGr ∙∙ cong₂ _+⊗_ (IsGroupHom.pres· (snd (f' x)) a b) (IsGroupHom.pres· (snd (f' y)) a b)) (λ a → IsGroupHom.pres· (snd (f' a))) - (λ a → ⊗elimProp (λ _ → ⊗squash _ _) - (λ b c → λ i → flip a b i ⊗ c) - λ x y ind1 ind2 → IsGroupHom.pres· (snd (f' a)) (-⊗ x) (-⊗ y) - ∙ cong₂ _+⊗_ ind1 ind2) refl ⨂assocIso : Iso (A ⨂₁ (B ⨂ C)) ((A ⨂ B) ⨂₁ C) From 57f091aa5dbebba595761a898b2b9402f93082f1 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 2 May 2022 17:46:27 +0200 Subject: [PATCH 29/30] fixes --- Cubical/Algebra/AbGroup/Base.agda | 10 ++ Cubical/Algebra/AbGroup/TensorProduct.agda | 40 ++--- .../Algebra/Group/EilenbergMacLane/Base.agda | 44 ++--- .../Group/EilenbergMacLane/CupProduct.agda | 120 ++------------ .../EilenbergMacLane/GroupStructure.agda | 92 +++++------ .../Group/EilenbergMacLane/Properties.agda | 151 +++++++++++++----- .../EilenbergMacLane/WedgeConnectivity.agda | 146 ++++++++--------- Cubical/Data/Nat/Base.agda | 8 + Cubical/Foundations/Path.agda | 13 ++ Cubical/Foundations/Prelude.agda | 10 -- 10 files changed, 320 insertions(+), 314 deletions(-) diff --git a/Cubical/Algebra/AbGroup/Base.agda b/Cubical/Algebra/AbGroup/Base.agda index 70f9b7d69..8d9e2d56e 100644 --- a/Cubical/Algebra/AbGroup/Base.agda +++ b/Cubical/Algebra/AbGroup/Base.agda @@ -191,6 +191,16 @@ identity (isMonoid (isGroup (isAbGroup (snd trivialAbGroup)))) _ = refl , refl inverse (isGroup (isAbGroup (snd trivialAbGroup))) _ = refl , refl comm (isAbGroup (snd trivialAbGroup)) _ _ = refl +-- useful lemma +move4 : ∀ {ℓ} {A : Type ℓ} (x y z w : A) (_+_ : A → A → A) + → ((x y z : A) → x + (y + z) ≡ (x + y) + z) + → ((x y : A) → x + y ≡ y + x) + → (x + y) + (z + w) ≡ ((x + z) + (y + w)) +move4 x y z w _+_ assoc comm = + sym (assoc x y (z + w)) + ∙∙ cong (x +_) (assoc y z w ∙∙ cong (_+ w) (comm y z) ∙∙ sym (assoc z y w)) + ∙∙ assoc x z (y + w) + ---- The type of homomorphisms A → B is an AbGroup if B is ----- module _ {ℓ ℓ' : Level} (AGr : Group ℓ) (BGr : AbGroup ℓ') where private diff --git a/Cubical/Algebra/AbGroup/TensorProduct.agda b/Cubical/Algebra/AbGroup/TensorProduct.agda index dc8cec005..3ad6aa1ee 100644 --- a/Cubical/Algebra/AbGroup/TensorProduct.agda +++ b/Cubical/Algebra/AbGroup/TensorProduct.agda @@ -144,34 +144,34 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where ------------------------------------------------------------------------------- -- Useful induction principle, which lets us view elements of A ⨂ B as lists -- over (A × B). Used for the right cancellation law - listify : List (A × B) → AGr ⨂₁ BGr - listify [] = 0A ⊗ 0B - listify (x ∷ x₁) = (fst x ⊗ snd x) +⊗ listify x₁ - - listify++ : (x y : List (A × B)) - → listify (x ++ y) ≡ (listify x +⊗ listify y) - listify++ [] y = sym (⊗lUnit (listify y)) - listify++ (x ∷ x₁) y = - (λ i → (fst x ⊗ snd x) +⊗ (listify++ x₁ y i)) - ∙ ⊗assoc (fst x ⊗ snd x) (listify x₁) (listify y) - - ∃List : (x : AGr ⨂₁ BGr) → ∃[ l ∈ List (A × B) ] (listify l ≡ x) + unlist : List (A × B) → AGr ⨂₁ BGr + unlist [] = 0A ⊗ 0B + unlist (x ∷ x₁) = (fst x ⊗ snd x) +⊗ unlist x₁ + + unlist++ : (x y : List (A × B)) + → unlist (x ++ y) ≡ (unlist x +⊗ unlist y) + unlist++ [] y = sym (⊗lUnit (unlist y)) + unlist++ (x ∷ x₁) y = + (λ i → (fst x ⊗ snd x) +⊗ (unlist++ x₁ y i)) + ∙ ⊗assoc (fst x ⊗ snd x) (unlist x₁) (unlist y) + + ∃List : (x : AGr ⨂₁ BGr) → ∃[ l ∈ List (A × B) ] (unlist l ≡ x) ∃List = ⊗elimProp (λ _ → squash) (λ a b → ∣ [ a , b ] , ⊗rUnit (a ⊗ b) ∣) λ x y → rec2 squash λ {(l1 , p) (l2 , q) - → ∣ (l1 ++ l2) , listify++ l1 l2 ∙ cong₂ _+⊗_ p q ∣} + → ∣ (l1 ++ l2) , unlist++ l1 l2 ∙ cong₂ _+⊗_ p q ∣} - ⊗elimPropListify : ∀ {ℓ} {C : AGr ⨂₁ BGr → Type ℓ} + ⊗elimPropUnlist : ∀ {ℓ} {C : AGr ⨂₁ BGr → Type ℓ} → ((x : _) → isProp (C x)) - → ((x : _) → C (listify x)) + → ((x : _) → C (unlist x)) → (x : _) → C x - ⊗elimPropListify {C = C} p f = + ⊗elimPropUnlist {C = C} p f = ⊗elimProp p (λ x y → subst C (⊗rUnit (x ⊗ y)) (f [ x , y ])) λ x y → pRec (isPropΠ2 λ _ _ → p _) (pRec (isPropΠ3 λ _ _ _ → p _) (λ {(l1 , p) (l2 , q) ind1 ind2 - → subst C (listify++ l2 l1 ∙ cong₂ _+⊗_ q p) (f (l2 ++ l1))}) + → subst C (unlist++ l2 l1 ∙ cong₂ _+⊗_ q p) (f (l2 ++ l1))}) (∃List y)) (∃List x) ----------------------------------------------------------------------------------- @@ -220,13 +220,13 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where ⊗rCancel : (x : AGr ⨂₁ BGr) → (x +⊗ -⊗ x) ≡ 0⊗ ⊗rCancel = - ⊗elimPropListify (λ _ → ⊗squash _ _) h + ⊗elimPropUnlist (λ _ → ⊗squash _ _) h where - h : (x : List (A × B)) → (listify x +⊗ -⊗ (listify x)) ≡ 0⊗ + h : (x : List (A × B)) → (unlist x +⊗ -⊗ (unlist x)) ≡ 0⊗ h [] = sym (linl 0A (-A 0A) (0B)) ∙ cong (λ x → _⊗_ {AGr = AGr} {BGr = BGr} x 0B) (invr strA 0A) h (x ∷ x₁) = - move4 (fst x ⊗ snd x) (listify x₁) ((-A fst x) ⊗ snd x) (-⊗ (listify x₁)) + move4 (fst x ⊗ snd x) (unlist x₁) ((-A fst x) ⊗ snd x) (-⊗ (unlist x₁)) _+⊗_ ⊗assoc ⊗comm ∙∙ cong₂ _+⊗_ (sym (linl (fst x) (-A (fst x)) (snd x)) ∙∙ (λ i → invr strA (fst x) i ⊗ (snd x)) diff --git a/Cubical/Algebra/Group/EilenbergMacLane/Base.agda b/Cubical/Algebra/Group/EilenbergMacLane/Base.agda index 87b24a57e..f4d56a499 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/Base.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/Base.agda @@ -19,13 +19,13 @@ open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.Properties open import Cubical.Homotopy.Connected open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec; elim to trElim) -open import Cubical.HITs.EilenbergMacLane1 +open import Cubical.HITs.EilenbergMacLane1 hiding (elim) open import Cubical.Algebra.AbGroup.Base open import Cubical.Data.Empty - renaming (rec to ⊥-rec) + renaming (rec to ⊥-rec) hiding (elim) open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) -open import Cubical.Data.Nat hiding (_·_) +open import Cubical.Data.Nat hiding (_·_ ; elim) open import Cubical.HITs.Susp open import Cubical.Functions.Morphism open import Cubical.Foundations.Path @@ -40,23 +40,23 @@ EM-raw G zero = fst G EM-raw G (suc zero) = EM₁ (G *) EM-raw G (suc (suc n)) = Susp (EM-raw G (suc n)) -ptS : {n : ℕ} {G : AbGroup ℓ} → EM-raw G n -ptS {n = zero} {G = G} = AbGroupStr.0g (snd G) -ptS {n = suc zero} {G = G} = embase -ptS {n = suc (suc n)} {G = G} = north +ptEM-raw : {n : ℕ} {G : AbGroup ℓ} → EM-raw G n +ptEM-raw {n = zero} {G = G} = AbGroupStr.0g (snd G) +ptEM-raw {n = suc zero} {G = G} = embase +ptEM-raw {n = suc (suc n)} {G = G} = north -EM-raw-elim : (G : AbGroup ℓ) (n : ℕ) {A : EM-raw G (suc n) → Type ℓ'} +raw-elim : (G : AbGroup ℓ) (n : ℕ) {A : EM-raw G (suc n) → Type ℓ'} → ((x : _) → isOfHLevel (suc n) (A x) ) - → A ptS + → A ptEM-raw → (x : _) → A x -EM-raw-elim G zero hlev b = elimProp _ hlev b -EM-raw-elim G (suc n) hlev b north = b -EM-raw-elim G (suc n) {A = A} hlev b south = subst A (merid ptS) b -EM-raw-elim G (suc n) {A = A} hlev b (merid a i) = help a i +raw-elim G zero hlev b = elimProp _ hlev b +raw-elim G (suc n) hlev b north = b +raw-elim G (suc n) {A = A} hlev b south = subst A (merid ptEM-raw) b +raw-elim G (suc n) {A = A} hlev b (merid a i) = help a i where - help : (a : _) → PathP (λ i → A (merid a i)) b (subst A (merid ptS) b) - help = EM-raw-elim G n (λ _ → isOfHLevelPathP' (suc n) (hlev _) _ _) - λ i → transp (λ j → A (merid ptS (j ∧ i))) (~ i) b + help : (a : _) → PathP (λ i → A (merid a i)) b (subst A (merid ptEM-raw) b) + help = raw-elim G n (λ _ → isOfHLevelPathP' (suc n) (hlev _) _ _) + λ i → transp (λ j → A (merid ptEM-raw (j ∧ i))) (~ i) b EM : (G : AbGroup ℓ) (n : ℕ) → Type ℓ EM G zero = EM-raw G zero @@ -66,13 +66,13 @@ EM G (suc (suc n)) = hLevelTrunc (4 + n) (EM-raw G (suc (suc n))) 0ₖ : {G : AbGroup ℓ} (n : ℕ) → EM G n 0ₖ {G = G} zero = AbGroupStr.0g (snd G) 0ₖ (suc zero) = embase -0ₖ (suc (suc n)) = ∣ ptS ∣ +0ₖ (suc (suc n)) = ∣ ptEM-raw ∣ EM∙ : (G : AbGroup ℓ) (n : ℕ) → Pointed ℓ EM∙ G n = EM G n , (0ₖ n) EM-raw∙ : (G : AbGroup ℓ) (n : ℕ) → Pointed ℓ -EM-raw∙ G n = EM-raw G n , ptS +EM-raw∙ G n = EM-raw G n , ptEM-raw hLevelEM : (G : AbGroup ℓ) (n : ℕ) → isOfHLevel (2 + n) (EM G n) hLevelEM G zero = AbGroupStr.is-set (snd G) @@ -84,10 +84,10 @@ EM-raw→EM G zero x = x EM-raw→EM G (suc zero) x = x EM-raw→EM G (suc (suc n)) = ∣_∣ -EM-elim : {G : AbGroup ℓ} (n : ℕ) {A : EM G n → Type ℓ'} +elim : {G : AbGroup ℓ} (n : ℕ) {A : EM G n → Type ℓ'} → ((x : _) → isOfHLevel (2 + n) (A x)) → ((x : EM-raw G n) → A (EM-raw→EM G n x)) → (x : _) → A x -EM-elim zero hlev hyp x = hyp x -EM-elim (suc zero) hlev hyp x = hyp x -EM-elim (suc (suc n)) hlev hyp = trElim (λ _ → hlev _) hyp +elim zero hlev hyp x = hyp x +elim (suc zero) hlev hyp x = hyp x +elim (suc (suc n)) hlev hyp = trElim (λ _ → hlev _) hyp diff --git a/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda b/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda index 6c3d3e3c8..f6339e928 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/CupProduct.agda @@ -2,7 +2,7 @@ module Cubical.Algebra.Group.EilenbergMacLane.CupProduct where -open import Cubical.Algebra.Group.EilenbergMacLane.Base +open import Cubical.Algebra.Group.EilenbergMacLane.Base renaming (elim to EM-elim) open import Cubical.Algebra.Group.EilenbergMacLane.WedgeConnectivity open import Cubical.Algebra.Group.EilenbergMacLane.GroupStructure open import Cubical.Algebra.Group.EilenbergMacLane.Properties @@ -31,7 +31,7 @@ open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) -open import Cubical.Data.Nat hiding (_·_) +open import Cubical.Data.Nat hiding (_·_) renaming (elim to ℕelim) open import Cubical.HITs.Susp open import Cubical.Algebra.AbGroup.TensorProduct @@ -46,52 +46,14 @@ private variable ℓ ℓ' ℓ'' : Level -private - -- ℕ induction. For some reason, this helps the termination checker - ind' : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 - → ((n : ℕ) → (A n → A (suc n))) - → (n : ℕ) → A n - ind' a0 ind zero = a0 - ind' a0 ind (suc n) = ind n (ind' a0 ind n) - - ind+2 : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 → A 1 - → ((n : ℕ) → (A (suc n) → A (suc (suc n)))) - → (n : ℕ) → A n - ind+2 a0 a1 ind zero = a0 - ind+2 a0 a1 ind (suc zero) = a1 - ind+2 {A = A} a0 a1 ind (suc (suc n)) = - ind n (ind+2 {A = A} a0 a1 ind (suc n)) - --- A homomorphism φ : G → H of AbGroups induces a homomorphism --- φ' : K(G,n) → K(H,n) -inducedFun-EM-raw : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} - → AbGroupHom G' H' - → ∀ n - → EM-raw G' n → EM-raw H' n -inducedFun-EM-raw f = - ind+2 (fst f) - (EMrec _ emsquash embase - (λ g → emloop (fst f g)) - λ g h → compPathR→PathP (sym - (sym (lUnit _) - ∙∙ cong (_∙ (sym (emloop (fst f h)))) - (cong emloop (IsGroupHom.pres· (snd f) g h) - ∙ emloop-comp _ (fst f g) (fst f h)) - ∙∙ sym (∙assoc _ _ _) - ∙∙ cong (emloop (fst f g) ∙_) (rCancel _) - ∙∙ sym (rUnit _)))) - (λ n ind → λ { north → north - ; south → south - ; (merid a i) → merid (ind a) i} ) - -- Lemma for distributativity of cup product (used later) pathType : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → Type ℓ pathType n x p = sym (rUnitₖ (2 + n) x) ∙ (λ i → x +ₖ p i) ≡ sym (lUnitₖ (2 + n) x) ∙ λ i → p i +ₖ x -lem : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) +pathTypeMake : ∀ {ℓ} {G : AbGroup ℓ} (n : ℕ) (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → pathType n x p -lem n x = J (λ x p → pathType n x p) refl +pathTypeMake n x = J (λ x p → pathType n x p) refl -- Definition of cup product (⌣ₖ, given by ·₀ when first argument is in K(G,0)) @@ -114,7 +76,7 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where ·₀' : H → (m : ℕ) → EM G' m → EM (G' ⨂ H') m ·₀' h = - ind+2 + elim+2 (_⊗ h) (elimGroupoid _ (λ _ → emsquash) embase @@ -132,7 +94,7 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where ·₀ : G → (m : ℕ) → EM H' m → EM (G' ⨂ H') m ·₀ g = - ind+2 (λ h → g ⊗ h) + elim+2 (λ h → g ⊗ h) (elimGroupoid _ (λ _ → emsquash) embase (λ h → emloop (g ⊗ h)) @@ -148,7 +110,7 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where ·₀-distr : (g h : G) → (m : ℕ) (x : EM H' m) → ·₀ (g +G h) m x ≡ ·₀ g m x +ₖ ·₀ h m x ·₀-distr g h = - ind+2 + elim+2 (linl g h) (elimSet _ (λ _ → emsquash _ _) refl @@ -188,7 +150,7 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where 0·₀ : (m : ℕ) → (x : _) → ·₀ 0G m x ≡ 0ₖ m 0·₀ = - ind+2 lCancelPrim + elim+2 lCancelPrim (elimSet _ (λ _ → emsquash _ _) refl λ g → compPathR→PathP ((sym (emloop-1g _) @@ -203,7 +165,7 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where 0·₀' : (m : ℕ) (g : _) → ·₀' 0H m g ≡ 0ₖ m 0·₀' = - ind+2 + elim+2 rCancelPrim (elimSet _ (λ _ → emsquash _ _) refl @@ -221,10 +183,10 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where -- Definition of the cup product cup∙ : ∀ n m → EM G' n → EM∙ H' m →∙ EM∙ (G' ⨂ H') (n +' m) cup∙ = - ind' + ℕelim (λ m g → (·₀ g m) , ·₀0 m g) λ n f → - ind' + ℕelim (λ g → (λ h → ·₀' h (suc n) g) , 0·₀' (suc n) g) λ m _ → main n m f @@ -261,10 +223,8 @@ module _ {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where trElim (λ _ → isOfHLevel↑∙ (2 + n) m) λ { north → (λ _ → 0ₖ (3 + (n + m))) , refl ; south → (λ _ → 0ₖ (3 + (n + m))) , refl - ; (merid a i) → (λ x → EM→ΩEM+1 _ (ind _ (EM-raw→EM _ _ a) .fst x) i) - , λ j → (cong (EM→ΩEM+1 (suc (suc (n + m)))) - (ind (suc m) (EM-raw→EM G' (suc n) a) .snd) - ∙ EM→ΩEM+1-0ₖ _) j i} + ; (merid a i) → Iso.inv (ΩfunExtIso _ _) + (EM→ΩEM+1∙ _ ∘∙ ind (suc m) (EM-raw→EM _ _ a)) i} _⌣ₖ_ : {n m : ℕ} (x : EM G' n) (y : EM H' m) → EM (G' ⨂ H') (n +' m) _⌣ₖ_ x y = cup∙ _ _ x .fst y @@ -323,7 +283,7 @@ module LeftDistributivity {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where ∙∙ λ i → (z ⌣ₖ y) +ₖ (⌣ₖ-0ₖ (suc n) (suc zero) z (~ i)) l≡r : (z : EM G' (suc n)) → l embase z ≡ r embase z - l≡r z = sym (lem _ _ (sym (⌣ₖ-0ₖ (suc n) (suc zero) z))) + l≡r z = sym (pathTypeMake _ _ (sym (⌣ₖ-0ₖ (suc n) (suc zero) z))) mainDistrL n (suc m) = elim2 (λ _ _ → isOfHLevelPath (4 + m) (hLevLem _ _) _ _) @@ -354,7 +314,7 @@ module LeftDistributivity {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where ∙∙ λ i → (z ⌣ₖ ∣ x ∣) +ₖ (⌣ₖ-0ₖ (suc n) (suc (suc m)) z (~ i)) l≡r : (z : EM G' (suc n)) → l north z ≡ r north z - l≡r z = sym (lem _ _ (sym (⌣ₖ-0ₖ (suc n) (suc (suc m)) z))) + l≡r z = sym (pathTypeMake _ _ (sym (⌣ₖ-0ₖ (suc n) (suc (suc m)) z))) module RightDistributivity {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where private @@ -407,7 +367,7 @@ module RightDistributivity {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where ∙∙ λ i → (_⌣ₖ_ {n = 1} {m = suc m} x z) +ₖ 0ₖ-⌣ₖ (suc zero) (suc m) z (~ i) l≡r : (z : _) → l embase z ≡ r embase z - l≡r z = lem _ _ _ + l≡r z = pathTypeMake _ _ _ mainDistrR (suc n) m = elim2 (λ _ _ → isOfHLevelPath (4 + n) @@ -438,54 +398,8 @@ module RightDistributivity {G' : AbGroup ℓ} {H' : AbGroup ℓ'} where ∙∙ λ i → (∣ x ∣ ⌣ₖ z) +ₖ 0ₖ-⌣ₖ _ _ z (~ i) r≡l : (z : _) → l north z ≡ r north z - r≡l z = lem _ _ _ + r≡l z = pathTypeMake _ _ _ -- TODO: Summarise distributivity proofs -- TODO: Associativity and graded commutativity, following Cubical.ZCohomology.RingStructure -- The following lemmas will be needed to make the types match up. - -inducedFun-EM-rawIso : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} - → AbGroupEquiv G' H' - → ∀ n → Iso (EM-raw G' n) (EM-raw H' n) -Iso.fun (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , (snd e)) n -Iso.inv (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , isGroupHomInv e) n -Iso.rightInv (inducedFun-EM-rawIso e n) = h n - where - h : (n : ℕ) → section (inducedFun-EM-raw (fst e .fst , snd e) n) - (inducedFun-EM-raw (invEq (fst e) , isGroupHomInv e) n) - h = ind+2 - (secEq (fst e)) - (elimSet _ (λ _ → emsquash _ _) refl - (λ g → compPathR→PathP - ((sym (cong₂ _∙_ (cong emloop (secEq (fst e) g)) (sym (lUnit _)) - ∙ rCancel _))))) - λ n p → λ { north → refl - ; south → refl - ; (merid a i) k → merid (p a k) i} -Iso.leftInv (inducedFun-EM-rawIso e n) = h n - where - h : (n : ℕ) → retract (Iso.fun (inducedFun-EM-rawIso e n)) - (Iso.inv (inducedFun-EM-rawIso e n)) - h = ind+2 - (retEq (fst e)) - (elimSet _ (λ _ → emsquash _ _) refl - (λ g → compPathR→PathP - ((sym (cong₂ _∙_ (cong emloop (retEq (fst e) g)) (sym (lUnit _)) - ∙ rCancel _))))) - λ n p → λ { north → refl - ; south → refl - ; (merid a i) k → merid (p a k) i} - -Iso→EMIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} - → AbGroupEquiv G H → ∀ n → Iso (EM G n) (EM H n) -Iso→EMIso is zero = inducedFun-EM-rawIso is zero -Iso→EMIso is (suc zero) = inducedFun-EM-rawIso is 1 -Iso→EMIso is (suc (suc n)) = mapCompIso (inducedFun-EM-rawIso is (suc (suc n))) - -EM⊗-commFun : {G : AbGroup ℓ} {H : AbGroup ℓ'} - → ∀ n → Iso (EM (G ⨂ H) n) (EM (H ⨂ G) n) -EM⊗-commFun {G = G} {H = H} = Iso→EMIso (GroupIso→GroupEquiv ⨂-commIso) - -EM⊗-assocIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} - → ∀ n → Iso (EM (G ⨂ (H ⨂ L)) n) (EM ((G ⨂ H) ⨂ L) n) -EM⊗-assocIso = Iso→EMIso (GroupIso→GroupEquiv (GroupEquiv→GroupIso ⨂assoc)) diff --git a/Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda b/Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda index bac2bbcb0..0eed521de 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/GroupStructure.agda @@ -20,7 +20,6 @@ open import Cubical.Foundations.Path open import Cubical.Homotopy.Loopspace -open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec; elim to trElim) open import Cubical.HITs.EilenbergMacLane1 open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) @@ -49,39 +48,31 @@ module _ {G : AbGroup ℓ} where → PathP (λ i → Path (EM₁ (AbGroup→Group G)) (emloop h i) (emloop h i)) (emloop g) (emloop g) helper g h = - compPathL→PathP - (cong (sym (emloop h) ∙_) - (sym (emloop-comp _ g h) - ∙∙ cong emloop (comm g h) - ∙∙ emloop-comp _ h g) - ∙∙ ∙assoc _ _ _ - ∙∙ cong (_∙ emloop g) (lCancel _) - ∙ sym (lUnit _)) + comm→PathP + ((sym (emloop-comp _ h g) + ∙∙ cong emloop (comm h g) + ∙∙ emloop-comp _ g h)) _+ₖ_ : {n : ℕ} → EM G n → EM G n → EM G n _+ₖ_ {n = zero} = _+G_ _+ₖ_ {n = suc zero} = rec _ (isGroupoidΠ (λ _ → emsquash)) (λ x → x) - looper + (λ x → funExt (looper x)) λ g h i j x → el g h x i j where - looper : fst G → (λ x → x) ≡ (λ x → x) - looper g = funExt (elimSet _ (λ _ → emsquash _ _) + looper : fst G → (x : _) → x ≡ x + looper g = (elimSet _ (λ _ → emsquash _ _) (emloop g) (helper g)) - lol : (g h : fst G) - → Square (emloop g) (emloop (g +G h)) refl (emloop h) - lol g h = compPath-filler (emloop g) (emloop h) ▷ sym (emloop-comp _ g h) - el : (g h : fst G) (x : EM₁ (AbGroup→Group G)) - → Square (λ j → looper g j x) - (λ j → looper ((snd (AbGroup→Group G) GroupStr.· g) h) j x) - refl λ j → looper h j x + → Square (looper g x) + (looper (g +G h) x) + refl (looper h x) el g h = elimProp _ (λ _ → isOfHLevelPathP' 1 (emsquash _ _) _ _) - (lol g h) + (emcomp g h) _+ₖ_ {n = suc (suc n)} = trRec2 (isOfHLevelTrunc (4 + n)) @@ -89,8 +80,8 @@ module _ {G : AbGroup ℓ} where (λ _ _ → hLevHelp n) ∣_∣ ∣_∣ refl) - σ-EM : (n : ℕ) → EM-raw G (suc n) → Path (EM-raw G (2 + n)) ptS ptS - σ-EM n x i = (merid x ∙ sym (merid ptS)) i + σ-EM : (n : ℕ) → EM-raw G (suc n) → Path (EM-raw G (2 + n)) ptEM-raw ptEM-raw + σ-EM n x = merid x ∙ sym (merid ptEM-raw) -ₖ_ : {n : ℕ} → EM G n → EM G n -ₖ_ {n = zero} x = -G x @@ -215,20 +206,20 @@ module _ {G : AbGroup ℓ} where pp : (a : _) → PathP (λ i → 0ₖ (suc (suc n)) ≡ ∣ merid a i ∣ₕ → ∣ merid a i ∣ₕ ≡ 0ₖ (2 + n)) (cong (λ x → -[ 2 + n ]ₖ x)) - λ p → cong ∣_∣ₕ (sym (merid ptS)) ∙ cong (λ x → -[ 2 + n ]ₖ x) p + λ p → cong ∣_∣ₕ (sym (merid ptEM-raw)) ∙ cong (λ x → -[ 2 + n ]ₖ x) p pp a = toPathP (funExt λ x → - (λ k → transp (λ i → Path (EM G (2 + n)) ∣ merid a (i ∨ k) ∣ ∣ ptS ∣) k + (λ k → transp (λ i → Path (EM G (2 + n)) ∣ merid a (i ∨ k) ∣ ∣ ptEM-raw ∣) k (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) (cong (-ₖ-syntax (suc (suc n))) - (transp (λ j → Path (EM G (2 + n)) ∣ ptS ∣ ∣ merid a (~ j ∧ ~ k) ∣) k + (transp (λ j → Path (EM G (2 + n)) ∣ ptEM-raw ∣ ∣ merid a (~ j ∧ ~ k) ∣) k (compPath-filler x (sym (cong ∣_∣ₕ (merid a))) k))) k)) ∙∙ cong (cong ∣_∣ₕ (sym (merid a)) ∙_) (cong-∙ (λ x → -[ 2 + n ]ₖ x) x (sym (cong ∣_∣ₕ (merid a))) ∙ isCommΩEM (suc n) (cong (λ x → -[ 2 + n ]ₖ x) x) (cong ∣_∣ₕ (σ-EM n a))) ∙∙ (λ k → (λ i → ∣ merid a (~ i ∨ k) ∣) - ∙ (λ i → ∣ compPath-filler' (merid a) (sym (merid ptS)) (~ k) i ∣) + ∙ (λ i → ∣ compPath-filler' (merid a) (sym (merid ptEM-raw)) (~ k) i ∣) ∙ cong (λ x → -ₖ-syntax (suc (suc n)) x) x) ∙ sym (lUnit _)) @@ -236,8 +227,8 @@ module _ {G : AbGroup ℓ} where → 0ₖ (2 + n) ≡ x → x ≡ 0ₖ (2 + n) decoder = trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelTruncPath {n = 4 + n}) - λ { north → pp ptS i0 - ; south → pp ptS i1 + λ { north → pp ptEM-raw i0 + ; south → pp ptEM-raw i1 ; (merid a i) → pp a i} main : (x : EM G (2 + n)) (p : 0ₖ (2 + n) ≡ x) → decoder x p ≡ sym p @@ -253,13 +244,13 @@ module _ {G : AbGroup ℓ} where rCancelₖ (suc (suc n)) = trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) λ { north → refl - ; south i → +ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣ - (-ₖ-syntax (suc (suc n)) ∣ merid ptS (~ i) ∣) + ; south i → +ₖ-syntax (suc (suc n)) ∣ merid ptEM-raw (~ i) ∣ + (-ₖ-syntax (suc (suc n)) ∣ merid ptEM-raw (~ i) ∣) ; (merid a i) j → hcomp (λ r → λ { (i = i0) → 0ₖ (2 + n) - ; (i = i1) → ∣ merid ptS (~ j ∧ r) ∣ₕ -[ 2 + n ]ₖ ∣ merid ptS (~ j ∧ r) ∣ - ; (j = i0) → ∣ compPath-filler (merid a) (sym (merid ptS)) (~ r) i ∣ - -[ 2 + n ]ₖ ∣ compPath-filler (merid a) (sym (merid ptS)) (~ r) i ∣ + ; (i = i1) → ∣ merid ptEM-raw (~ j ∧ r) ∣ₕ -[ 2 + n ]ₖ ∣ merid ptEM-raw (~ j ∧ r) ∣ + ; (j = i0) → ∣ compPath-filler (merid a) (sym (merid ptEM-raw)) (~ r) i ∣ + -[ 2 + n ]ₖ ∣ compPath-filler (merid a) (sym (merid ptEM-raw)) (~ r) i ∣ ; (j = i1) → 0ₖ (2 + n)}) (help' a j i) } where @@ -298,27 +289,27 @@ module _ {G : AbGroup ℓ} where ≡ (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ merid c i ∣ₕ)) (cong (λ x → ∣ a ∣ₕ +[ suc (suc n) ]ₖ x) (rUnitₖ (suc (suc n)) ∣ b ∣) ∙ sym (rUnitₖ (suc (suc n)) (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ))) - ((λ i → ∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ merid ptS (~ i) ∣ₕ)) + ((λ i → ∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ merid ptEM-raw (~ i) ∣ₕ)) ∙∙ cong (λ x → ∣ a ∣ₕ +[ suc (suc n) ]ₖ x) (rUnitₖ (suc (suc n)) ∣ b ∣) ∙ sym (rUnitₖ (suc (suc n)) (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ)) - ∙∙ λ i → (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ merid ptS i ∣ₕ) + ∙∙ λ i → (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ merid ptEM-raw i ∣ₕ) lem c = - EM-raw-elim G (suc n) + raw-elim G (suc n) (λ _ → isOfHLevelΠ (2 + n) (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _)) - (EM-raw-elim G (suc n) + (raw-elim G (suc n) (λ _ → isOfHLevelPathP' (2 + n) (isOfHLevelTrunc (4 + n) _ _) _ _) ((sym (rUnit refl) ◁ λ _ → refl) - ▷ (sym (lCancel (cong ∣_∣ₕ (merid ptS))) - ∙ λ i → (λ j → ∣ merid ptS (~ j ∨ ~ i) ∣ₕ) - ∙∙ lUnit (λ j → ∣ merid ptS (~ j ∧ ~ i) ∣ₕ) i - ∙∙ cong ∣_∣ₕ (merid ptS)))) + ▷ (sym (lCancel (cong ∣_∣ₕ (merid ptEM-raw))) + ∙ λ i → (λ j → ∣ merid ptEM-raw (~ j ∨ ~ i) ∣ₕ) + ∙∙ lUnit (λ j → ∣ merid ptEM-raw (~ j ∧ ~ i) ∣ₕ) i + ∙∙ cong ∣_∣ₕ (merid ptEM-raw)))) main : (c a b : _) → (∣ a ∣ₕ +[ suc (suc n) ]ₖ (∣ b ∣ₕ +[ suc (suc n) ]ₖ ∣ c ∣ₕ) ≡ (∣ a ∣ₕ +[ suc (suc n) ]ₖ ∣ b ∣ₕ) +[ suc (suc n) ]ₖ ∣ c ∣ₕ) - main north a b = lem ptS a b i0 - main south a b = lem ptS a b i1 + main north a b = lem ptEM-raw a b i0 + main south a b = lem ptEM-raw a b i1 main (merid c i) a b = lem c a b i σ-EM' : (n : ℕ) (x : EM G (suc n)) @@ -331,14 +322,15 @@ module _ {G : AbGroup ℓ} where λ x → cong ∣_∣ₕ (σ-EM (suc n) x) σ-EM'-0ₖ : (n : ℕ) → σ-EM' n (0ₖ (suc n)) ≡ refl - σ-EM'-0ₖ zero = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) - σ-EM'-0ₖ (suc n) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + σ-EM'-0ₖ zero = cong (cong ∣_∣ₕ) (rCancel (merid ptEM-raw)) + σ-EM'-0ₖ (suc n) = cong (cong ∣_∣ₕ) (rCancel (merid ptEM-raw)) private - P : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (r : refl ≡ p) + lUnit-rUnit-coh : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (r : refl ≡ p) → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r - P p = J (λ p r → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r) refl + lUnit-rUnit-coh p = + J (λ p r → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r) refl σ-EM'-hom : (n : ℕ) → (a b : _) → σ-EM' n (a +ₖ b) ≡ σ-EM' n a ∙ σ-EM' n b σ-EM'-hom zero = @@ -356,7 +348,7 @@ module _ {G : AbGroup ℓ} where ∙∙ cong (σ-EM' zero x ∙_) (sym (σ-EM'-0ₖ zero)) p : _ - p = P (σ-EM' zero embase) (sym (σ-EM'-0ₖ zero)) + p = lUnit-rUnit-coh (σ-EM' zero embase) (sym (σ-EM'-0ₖ zero)) σ-EM'-hom (suc n) = elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) _ _) _ _) (wedgeConEM.fun G G _ _ @@ -372,7 +364,7 @@ module _ {G : AbGroup ℓ} where (λ x → cong (σ-EM' (suc n)) (rUnitₖ (2 + n) ∣ x ∣) ∙∙ rUnit (σ-EM' (suc n) ∣ x ∣) ∙∙ cong (σ-EM' (suc n) ∣ x ∣ ∙_) (sym (σ-EM'-0ₖ (suc n)))) - (P (σ-EM' (suc n) (0ₖ (2 + n))) (sym (σ-EM'-0ₖ (suc n))))) + (lUnit-rUnit-coh (σ-EM' (suc n) (0ₖ (2 + n))) (sym (σ-EM'-0ₖ (suc n))))) σ-EM'-ₖ : (n : ℕ) → (a : _) → σ-EM' n (-ₖ a) ≡ sym (σ-EM' n a) σ-EM'-ₖ n = @@ -395,7 +387,7 @@ module _ {G : AbGroup ℓ} where r : _ r x = cong (λ z → -[ 1 ]ₖ z) (rUnitₖ 1 x) ∙ sym (rUnitₖ 1 (-[ 1 ]ₖ x)) - p : r ptS ≡ l ptS + p : r ptEM-raw ≡ l ptEM-raw p = sym (rUnit refl) k = wedgeConEM.fun G G 0 0 (λ _ _ → emsquash _ _) l r (sym p) diff --git a/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda b/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda index 7c1665958..305bc9085 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda @@ -2,11 +2,14 @@ module Cubical.Algebra.Group.EilenbergMacLane.Properties where -open import Cubical.Algebra.Group.EilenbergMacLane.Base +open import Cubical.Algebra.Group.EilenbergMacLane.Base renaming (elim to EM-elim) open import Cubical.Algebra.Group.EilenbergMacLane.WedgeConnectivity open import Cubical.Algebra.Group.EilenbergMacLane.GroupStructure open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.Properties +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.AbGroup.TensorProduct open import Cubical.Algebra.AbGroup.Base open import Cubical.Foundations.Prelude @@ -30,7 +33,7 @@ open import Cubical.Data.Nat hiding (_·_) open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec; elim to trElim) -open import Cubical.HITs.EilenbergMacLane1 +open import Cubical.HITs.EilenbergMacLane1 renaming (rec to EMrec) open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; rec2 to trRec2) open import Cubical.HITs.Susp @@ -59,7 +62,7 @@ module _ {G' : AbGroup ℓ} where snd (isConnectedEM (suc (suc n))) = trElim (λ _ → isOfHLevelTruncPath) (trElim (λ _ → isOfHLevelSuc (3 + n) (isOfHLevelTruncPath {n = (3 + n)})) - (EM-raw-elim G' + (raw-elim G' (suc n) (λ _ → isOfHLevelTrunc (3 + n) _ _) refl)) @@ -126,7 +129,7 @@ module _ (Ĝ : Group ℓ) where compRightEquiv g h = equivEq (funExt (λ x → (assoc x g h) ⁻¹)) CodesSet : EM₁ Ĝ → hSet ℓ - CodesSet = rec Ĝ (isOfHLevelTypeOfHLevel 2) (G , is-set) RE REComp + CodesSet = EMrec Ĝ (isOfHLevelTypeOfHLevel 2) (G , is-set) RE REComp where RE : (g : G) → Path (hSet ℓ) (G , is-set) (G , is-set) RE g = Σ≡Prop (λ X → isPropIsOfHLevel {A = X} 2) (ua (rightEquiv g)) @@ -210,8 +213,8 @@ module _ {G : AbGroup ℓ} where f : (n : ℕ) → _ f n = σ-EM' n - g : (n : ℕ) → EM G (suc n) → ∣ ptS ∣ ≡ ∣ south ∣ - g n x = σ-EM' n x ∙ cong ∣_∣ₕ (merid ptS) + g : (n : ℕ) → EM G (suc n) → ∣ ptEM-raw ∣ ≡ ∣ south ∣ + g n x = σ-EM' n x ∙ cong ∣_∣ₕ (merid ptEM-raw) lem₂ : (n : ℕ) (a x : _) → Path (Path (EM G (suc (suc n))) _ _) @@ -251,7 +254,7 @@ module _ {G : AbGroup ℓ} where toPathP (funExt (EM-elim _ (λ _ → isOfHLevelPathP (2 + (suc n)) (isOfHLevelTrunc (4 + n) _ _) _ _) λ x → - ((λ i → transp (λ j → Path (EM G (2 + n)) ∣ ptS ∣ ∣ merid a (i ∨ j) ∣) + ((λ i → transp (λ j → Path (EM G (2 + n)) ∣ ptEM-raw ∣ ∣ merid a (i ∨ j) ∣) i (compPath-filler (lem n x a i) (cong ∣_∣ₕ (merid a)) i) )) ∙∙ sym (∙assoc _ _ _) ∙∙ lem₂ n a x)) @@ -273,7 +276,7 @@ module _ {G : AbGroup ℓ} where encode'-decode' zero x = cong (encode' zero (0ₖ 2)) (cong-∙ ∣_∣ₕ (merid x) (sym (merid embase))) ∙∙ substComposite (λ x → CODE zero x .fst) - (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptS))) embase + (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptEM-raw))) embase ∙∙ (cong (subst (λ x₁ → CODE zero x₁ .fst) (λ i → ∣ merid embase (~ i) ∣ₕ)) (λ i → lUnitₖ 1 (transportRefl x i) i) ∙ (λ i → rUnitₖ 1 (transportRefl x i) i)) @@ -281,8 +284,8 @@ module _ {G : AbGroup ℓ} where trElim (λ _ → isOfHLevelTruncPath {n = 4 + n}) λ x → cong (encode' (suc n) (0ₖ (3 + n))) (cong-∙ ∣_∣ₕ (merid x) (sym (merid north))) ∙∙ substComposite (λ x → CODE (suc n) x .fst) - (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptS))) (0ₖ (2 + n)) - ∙∙ cong (subst (λ x₁ → CODE (suc n) x₁ .fst) (λ i → ∣ merid ptS (~ i) ∣ₕ)) + (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid ptEM-raw))) (0ₖ (2 + n)) + ∙∙ cong (subst (λ x₁ → CODE (suc n) x₁ .fst) (λ i → ∣ merid ptEM-raw (~ i) ∣ₕ)) (λ i → lUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) ∙ (λ i → rUnitₖ (2 + n) (transportRefl ∣ x ∣ₕ i) i) @@ -309,8 +312,8 @@ module _ {G : AbGroup ℓ} where EM→ΩEM+1-0ₖ : (n : ℕ) → EM→ΩEM+1 n (0ₖ n) ≡ refl EM→ΩEM+1-0ₖ zero = emloop-1g _ - EM→ΩEM+1-0ₖ (suc zero) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) - EM→ΩEM+1-0ₖ (suc (suc n)) = cong (cong ∣_∣ₕ) (rCancel (merid ptS)) + EM→ΩEM+1-0ₖ (suc zero) = cong (cong ∣_∣ₕ) (rCancel (merid ptEM-raw)) + EM→ΩEM+1-0ₖ (suc (suc n)) = cong (cong ∣_∣ₕ) (rCancel (merid ptEM-raw)) EM→ΩEM+1-hom : (n : ℕ) (x y : EM G n) → EM→ΩEM+1 n (x +[ n ]ₖ y) ≡ EM→ΩEM+1 n x ∙ EM→ΩEM+1 n y @@ -330,6 +333,12 @@ module _ {G : AbGroup ℓ} where ΩEM+1→EM-refl (suc zero) = refl ΩEM+1→EM-refl (suc (suc n)) = refl + EM→ΩEM+1∙ : (n : ℕ) → EM∙ G n →∙ Ω (EM∙ G (suc n)) + EM→ΩEM+1∙ n .fst = EM→ΩEM+1 n + EM→ΩEM+1∙ zero .snd = emloop-1g (AbGroup→Group G) + EM→ΩEM+1∙ (suc zero) .snd = cong (cong ∣_∣ₕ) (rCancel (merid embase)) + EM→ΩEM+1∙ (suc (suc n)) .snd = cong (cong ∣_∣ₕ) (rCancel (merid north)) + EM≃ΩEM+1∙ : (n : ℕ) → EM∙ G n ≡ Ω (EM∙ G (suc n)) EM≃ΩEM+1∙ n = ua∙ (EM≃ΩEM+1 n) (EM→ΩEM+1-0ₖ n) @@ -347,21 +356,21 @@ module _ where fst (isContr-↓∙ {G = G} {H = H} zero) = (λ _ → 0g (snd H)) , refl snd (isContr-↓∙{G = G} {H = H} zero) (f , p) = Σ≡Prop (λ x → is-set (snd H) _ _) - (funExt (EM-raw-elim G 0 (λ _ → is-set (snd H) _ _) (sym p))) + (funExt (raw-elim G 0 (λ _ → is-set (snd H) _ _) (sym p))) fst (isContr-↓∙ {G = G} {H = H} (suc n)) = (λ _ → 0ₖ (suc n)) , refl fst (snd (isContr-↓∙ {G = G} {H = H} (suc n)) f i) x = trElim {B = λ x → 0ₖ (suc n) ≡ fst f x} (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (hLevelEM H (suc n))) _ _) - (EM-raw-elim _ _ (λ _ → hLevelEM H (suc n) _ _) (sym (snd f))) + (raw-elim _ _ (λ _ → hLevelEM H (suc n) _ _) (sym (snd f))) x i snd (snd (isContr-↓∙ (suc n)) f i) j = snd f (~ i ∨ j) isContr-↓∙' : {G : AbGroup ℓ} {H : AbGroup ℓ'} (n : ℕ) - → isContr ((EM-raw G (suc n) , ptS) →∙ EM∙ H n) + → isContr ((EM-raw G (suc n) , ptEM-raw) →∙ EM∙ H n) isContr-↓∙' zero = isContr-↓∙ zero fst (isContr-↓∙' (suc n)) = (λ _ → 0ₖ (suc n)) , refl fst (snd (isContr-↓∙' {H = H} (suc n)) f i) x = - EM-raw-elim _ _ {A = λ x → 0ₖ (suc n) ≡ fst f x} + raw-elim _ _ {A = λ x → 0ₖ (suc n) ≡ fst f x} (λ _ → hLevelEM H (suc n) _ _) (sym (snd f)) x i snd (snd (isContr-↓∙' (suc n)) f i) j = snd f (~ i ∨ j) @@ -402,26 +411,27 @@ module _ where ua∙ (isoToEquiv (ΩfunExtIso A (EM∙ G (suc n)))) refl ∙ (λ i → (A →∙ EM≃ΩEM+1∙ {G = G} n (~ i) ∙)) - contr∙-lem : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) - → isContr (EM∙ G (suc n) →∙ (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) - fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl), refl - snd (contr∙-lem {G = G} {H = H} {L = L} n m) (f , p) = - →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) - (sym (funExt (help n f p))) - where - help : (n : ℕ) → (f : EM G (suc n) → EM∙ H (suc m) →∙ EM∙ L (suc (n + m))) - → f (snd (EM∙ G (suc n))) ≡ snd (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) - → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) - help zero f p = - EM-raw-elim G zero - (λ _ → isOfHLevel↑∙ zero m _ _) p - help (suc n) f p = - trElim (λ _ → isOfHLevelPath (4 + n) - (subst (λ x → isOfHLevel x (EM∙ H (suc m) →∙ EM∙ L (suc ((suc n) + m)))) - (λ i → suc (suc (+-comm (suc n) 1 i))) - (isOfHLevelPlus' {n = 1} (3 + n) (isOfHLevel↑∙ (suc n) m))) _ _) - (EM-raw-elim G (suc n) - (λ _ → isOfHLevel↑∙ (suc n) m _ _) p) + private + contr∙-lem : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} (n m : ℕ) + → isContr (EM∙ G (suc n) →∙ (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙)) + fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl), refl + snd (contr∙-lem {G = G} {H = H} {L = L} n m) (f , p) = + →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousEM _)) + (sym (funExt (help n f p))) + where + help : (n : ℕ) → (f : EM G (suc n) → EM∙ H (suc m) →∙ EM∙ L (suc (n + m))) + → f (snd (EM∙ G (suc n))) ≡ snd (EM∙ H (suc m) →∙ EM∙ L (suc (n + m)) ∙) + → (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) + help zero f p = + raw-elim G zero + (λ _ → isOfHLevel↑∙ zero m _ _) p + help (suc n) f p = + trElim (λ _ → isOfHLevelPath (4 + n) + (subst (λ x → isOfHLevel x (EM∙ H (suc m) →∙ EM∙ L (suc ((suc n) + m)))) + (λ i → suc (suc (+-comm (suc n) 1 i))) + (isOfHLevelPlus' {n = 1} (3 + n) (isOfHLevel↑∙ (suc n) m))) _ _) + (raw-elim G (suc n) + (λ _ → isOfHLevel↑∙ (suc n) m _ _) p) isOfHLevel↑∙∙ : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} → ∀ n m l → isOfHLevel (2 + l) (EM∙ G (suc n) @@ -459,3 +469,72 @@ module _ where (λ i → EM∙ G (suc n) →∙ →∙EMPath {G = L} (EM∙ H (suc m)) (suc (suc (l + n + m))) (~ i)) (isOfHLevel↑∙∙ n m l) + + +-- A homomorphism φ : G → H of AbGroups induces a homomorphism +-- φ' : K(G,n) → K(H,n) +inducedFun-EM-raw : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} + → AbGroupHom G' H' + → ∀ n + → EM-raw G' n → EM-raw H' n +inducedFun-EM-raw f = + elim+2 (fst f) + (EMrec _ emsquash embase + (λ g → emloop (fst f g)) + λ g h → compPathR→PathP (sym + (sym (lUnit _) + ∙∙ cong (_∙ (sym (emloop (fst f h)))) + (cong emloop (IsGroupHom.pres· (snd f) g h) + ∙ emloop-comp _ (fst f g) (fst f h)) + ∙∙ sym (∙assoc _ _ _) + ∙∙ cong (emloop (fst f g) ∙_) (rCancel _) + ∙∙ sym (rUnit _)))) + (λ n ind → λ { north → north + ; south → south + ; (merid a i) → merid (ind a) i} ) + +inducedFun-EM-rawIso : {G' : AbGroup ℓ} {H' : AbGroup ℓ'} + → AbGroupEquiv G' H' + → ∀ n → Iso (EM-raw G' n) (EM-raw H' n) +Iso.fun (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , (snd e)) n +Iso.inv (inducedFun-EM-rawIso e n) = inducedFun-EM-raw (_ , isGroupHomInv e) n +Iso.rightInv (inducedFun-EM-rawIso e n) = h n + where + h : (n : ℕ) → section (inducedFun-EM-raw (fst e .fst , snd e) n) + (inducedFun-EM-raw (invEq (fst e) , isGroupHomInv e) n) + h = elim+2 + (secEq (fst e)) + (elimSet _ (λ _ → emsquash _ _) refl + (λ g → compPathR→PathP + ((sym (cong₂ _∙_ (cong emloop (secEq (fst e) g)) (sym (lUnit _)) + ∙ rCancel _))))) + λ n p → λ { north → refl + ; south → refl + ; (merid a i) k → merid (p a k) i} +Iso.leftInv (inducedFun-EM-rawIso e n) = h n + where + h : (n : ℕ) → retract (Iso.fun (inducedFun-EM-rawIso e n)) + (Iso.inv (inducedFun-EM-rawIso e n)) + h = elim+2 + (retEq (fst e)) + (elimSet _ (λ _ → emsquash _ _) refl + (λ g → compPathR→PathP + ((sym (cong₂ _∙_ (cong emloop (retEq (fst e) g)) (sym (lUnit _)) + ∙ rCancel _))))) + λ n p → λ { north → refl + ; south → refl + ; (merid a i) k → merid (p a k) i} + +Iso→EMIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} + → AbGroupEquiv G H → ∀ n → Iso (EM G n) (EM H n) +Iso→EMIso is zero = inducedFun-EM-rawIso is zero +Iso→EMIso is (suc zero) = inducedFun-EM-rawIso is 1 +Iso→EMIso is (suc (suc n)) = mapCompIso (inducedFun-EM-rawIso is (suc (suc n))) + +EM⊗-commIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} + → ∀ n → Iso (EM (G ⨂ H) n) (EM (H ⨂ G) n) +EM⊗-commIso {G = G} {H = H} = Iso→EMIso (GroupIso→GroupEquiv ⨂-commIso) + +EM⊗-assocIso : {G : AbGroup ℓ} {H : AbGroup ℓ'} {L : AbGroup ℓ''} + → ∀ n → Iso (EM (G ⨂ (H ⨂ L)) n) (EM ((G ⨂ H) ⨂ L) n) +EM⊗-assocIso = Iso→EMIso (GroupIso→GroupEquiv (GroupEquiv→GroupIso ⨂assoc)) diff --git a/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda b/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda index afd688a16..1a9ae3a27 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/WedgeConnectivity.agda @@ -29,17 +29,17 @@ private wedgeConFun' : (G : AbGroup ℓ) (H : AbGroup ℓ') (n : ℕ) → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ''} → ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y)) - → (f : (x : _) → A ptS x) - → (g : (x : _) → A x ptS) - → f ptS ≡ g ptS + → (f : (x : _) → A ptEM-raw x) + → (g : (x : _) → A x ptEM-raw) + → f ptEM-raw ≡ g ptEM-raw → (x : _) (y : _) → A x y wedgeConFun'ᵣ : (G : AbGroup ℓ) (H : AbGroup ℓ') (n : ℕ) → {A : EM-raw G (suc n) → EM-raw H (suc zero) → Type ℓ''} → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc zero) (A x y))) - → (f : (x : _) → A ptS x) - → (g : (x : _) → A x ptS) - → (p : f ptS ≡ g ptS) - → (x : _) → wedgeConFun' G H n hLev f g p x ptS ≡ g x + → (f : (x : _) → A ptEM-raw x) + → (g : (x : _) → A x ptEM-raw) + → (p : f ptEM-raw ≡ g ptEM-raw) + → (x : _) → wedgeConFun' G H n hLev f g p x ptEM-raw ≡ g x wedgeConFun' G H zero {A = A} hlev f g p = elimSet _ (λ _ → isSetΠ λ _ → hlev _ _) f mainpath where @@ -83,53 +83,53 @@ private toPathP (isOfHLevelPathP' 1 (isOfHLevelPathP 2 (isOfHLevelPathP 2 (hlev _ _) _ _) _ _) _ _ _ _) wedgeConFun' G H (suc n) {A = A} hLev f g p north y = f y - wedgeConFun' G H (suc n) {A = A} hLev f g p south y = subst (λ x → A x y) (merid ptS) (f y) + wedgeConFun' G H (suc n) {A = A} hLev f g p south y = subst (λ x → A x y) (merid ptEM-raw) (f y) wedgeConFun' G H (suc n) {A = A} hLev f g p (merid a i) y = mainₗ a y i module _ where - llem₁ : g south ≡ subst (λ x₁ → A x₁ ptS) (merid ptS) (f ptS) - llem₁ = (λ i → transp (λ j → A (merid ptS (j ∨ ~ i)) ptS) (~ i) (g (merid ptS (~ i)))) - ∙ cong (subst (λ x₁ → A x₁ ptS) (merid ptS)) (sym p) + llem₁ : g south ≡ subst (λ x₁ → A x₁ ptEM-raw) (merid ptEM-raw) (f ptEM-raw) + llem₁ = (λ i → transp (λ j → A (merid ptEM-raw (j ∨ ~ i)) ptEM-raw) (~ i) (g (merid ptEM-raw (~ i)))) + ∙ cong (subst (λ x₁ → A x₁ ptEM-raw) (merid ptEM-raw)) (sym p) llem₁' : Square - (λ i → transp (λ j → A (merid ptS (i ∨ j)) ptS) i (g (merid ptS i))) refl - (cong (subst (λ x → A x ptS) (merid ptS)) (sym p)) llem₁ + (λ i → transp (λ j → A (merid ptEM-raw (i ∨ j)) ptEM-raw) i (g (merid ptEM-raw i))) refl + (cong (subst (λ x → A x ptEM-raw) (merid ptEM-raw)) (sym p)) llem₁ llem₁' i j = - hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptS (j ∨ z)) ptS) j - (g (merid ptS j)) - ; (i = i1) → subst (λ x₁ → A x₁ ptS) (merid ptS) (p (~ k)) - ; (j = i0) → (subst (λ x → A x ptS) (merid ptS)) (p (~ k ∨ ~ i))}) - (transp (λ k → A (merid ptS (k ∨ ~ i ∧ j)) ptS) (~ i ∧ j) (g (merid ptS (~ i ∧ j)))) + hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptEM-raw (j ∨ z)) ptEM-raw) j + (g (merid ptEM-raw j)) + ; (i = i1) → subst (λ x₁ → A x₁ ptEM-raw) (merid ptEM-raw) (p (~ k)) + ; (j = i0) → (subst (λ x → A x ptEM-raw) (merid ptEM-raw)) (p (~ k ∨ ~ i))}) + (transp (λ k → A (merid ptEM-raw (k ∨ ~ i ∧ j)) ptEM-raw) (~ i ∧ j) (g (merid ptEM-raw (~ i ∧ j)))) - llem₂ : (λ i₁ → transp (λ j → A (merid ptS (i₁ ∧ j)) ptS) (~ i₁) (f ptS)) + llem₂ : (λ i₁ → transp (λ j → A (merid ptEM-raw (i₁ ∧ j)) ptEM-raw) (~ i₁) (f ptEM-raw)) ≡ (λ i₁ → hcomp (λ k → λ { (i₁ = i0) → p (~ k) ; (i₁ = i1) → llem₁ k }) - (g (merid ptS i₁))) + (g (merid ptEM-raw i₁))) llem₂ i j = - hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptS (j ∧ j₁)) ptS) (~ j) (p (~ k)) + hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptEM-raw (j ∧ j₁)) ptEM-raw) (~ j) (p (~ k)) ; (j = i0) → p (~ k) ; (j = i1) → llem₁' k i}) - (transp (λ k → A (merid ptS ((i ∨ k) ∧ j)) ptS) (i ∨ ~ j) (g (merid ptS (i ∧ j)))) + (transp (λ k → A (merid ptEM-raw ((i ∨ k) ∧ j)) ptEM-raw) (i ∨ ~ j) (g (merid ptEM-raw (i ∧ j)))) mainₗ : (a : _) (y : _) - → PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) + → PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptEM-raw) (f y)) mainₗ = wedgeConFun' G H n (λ _ _ → isOfHLevelPathP' (suc (n + 1)) (hLev _ _) _ _) - (λ x i → transp (λ j → A (merid ptS (i ∧ j)) x) (~ i) (f x)) + (λ x i → transp (λ j → A (merid ptEM-raw (i ∧ j)) x) (~ i) (f x)) (λ x i → hcomp (λ k → λ { (i = i0) → p (~ k) ; (i = i1) → llem₁ k}) (g (merid x i))) llem₂ mainP : (y : _) - → mainₗ y ptS + → mainₗ y ptEM-raw ≡ λ i → hcomp (λ k → λ { (i = i0) → p (~ k) ; (i = i1) → llem₁ k}) (g (merid y i)) mainP y = wedgeConFun'ᵣ G H n (λ _ _ → isOfHLevelPathP' (suc (n + 1)) (hLev _ _) _ _) - (λ x i → transp (λ j → A (merid ptS (i ∧ j)) x) (~ i) (f x)) + (λ x i → transp (λ j → A (merid ptEM-raw (i ∧ j)) x) (~ i) (f x)) (λ x i → hcomp (λ k → λ { (i = i0) → p (~ k) ; (i = i1) → llem₁ k}) (g (merid x i))) @@ -137,20 +137,20 @@ private wedgeConFun'ᵣ G H zero {A = A} hLev f g p = elimProp _ (λ _ → hLev _ _ _ _) p wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p north = p - wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p south = sym (llem₁ G H n hLev f g p ptS i0 ptS) + wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p south = sym (llem₁ G H n hLev f g p ptEM-raw i0 ptEM-raw) wedgeConFun'ᵣ G H (suc n) {A = A} hLev f g p (merid a i) k = P k i where llem : _ llem i j = hcomp (λ k → λ { (i = i1) → g (merid a j) ; (j = i0) → p (i ∨ ~ k) - ; (j = i1) → llem₁ G H n hLev f g p ptS i0 ptS (~ i ∧ k)}) + ; (j = i1) → llem₁ G H n hLev f g p ptEM-raw i0 ptEM-raw (~ i ∧ k)}) (g (merid a j)) - P : PathP (λ k → PathP (λ i → A (merid a i) ptS) - (p k) (llem₁ G H n hLev f g p ptS i0 ptS (~ k))) - (λ i → mainₗ G H n hLev f g p a i ptS a ptS i) λ i → g (merid a i) - P = mainP G H n hLev f g p a i0 ptS a ◁ llem + P : PathP (λ k → PathP (λ i → A (merid a i) ptEM-raw) + (p k) (llem₁ G H n hLev f g p ptEM-raw i0 ptEM-raw (~ k))) + (λ i → mainₗ G H n hLev f g p a i ptEM-raw a ptEM-raw i) λ i → g (merid a i) + P = mainP G H n hLev f g p a i0 ptEM-raw a ◁ llem -- Here, the actual stuff gets proved. However an additional natural number is stuck into the context -- to convince the termination checker @@ -158,83 +158,83 @@ private wedgeConFun : (G : AbGroup ℓ) (H : AbGroup ℓ') (k n m : ℕ) → (n + m ≡ k) → {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ''} → ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y)) - → (f : (x : _) → A ptS x) - → (g : (x : _) → A x ptS) - → f ptS ≡ g ptS + → (f : (x : _) → A ptEM-raw x) + → (g : (x : _) → A x ptEM-raw) + → f ptEM-raw ≡ g ptEM-raw → (x : _) (y : _) → A x y wedgeconLeft : (G : AbGroup ℓ) (H : AbGroup ℓ') (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ''} → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) - → (f : (x : _) → A ptS x) - → (g : (x : _) → A x ptS) - → (p : f ptS ≡ g ptS) - → (x : _) → wedgeConFun G H k n m P hLev f g p ptS x ≡ f x + → (f : (x : _) → A ptEM-raw x) + → (g : (x : _) → A x ptEM-raw) + → (p : f ptEM-raw ≡ g ptEM-raw) + → (x : _) → wedgeConFun G H k n m P hLev f g p ptEM-raw x ≡ f x wedgeconRight : (G : AbGroup ℓ) (H : AbGroup ℓ') (k n m : ℕ) (P : n + m ≡ k) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ''} → (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) - → (f : (x : _) → A ptS x) - → (g : (x : _) → A x ptS) - → (p : f ptS ≡ g ptS) - → (x : _) → wedgeConFun G H k n m P hLev f g p x ptS ≡ g x + → (f : (x : _) → A ptEM-raw x) + → (g : (x : _) → A x ptEM-raw) + → (p : f ptEM-raw ≡ g ptEM-raw) + → (x : _) → wedgeConFun G H k n m P hLev f g p x ptEM-raw ≡ g x wedgeConFun G H k n zero P {A = A} hLev f g p = wedgeConFun' G H n hLev f g p wedgeConFun G H k zero (suc m) P {A = A} hLev f g p x y = wedgeConFun' H G (suc m) {A = λ x y → A y x} (λ x y → subst (λ n → isOfHLevel (2 + n) (A y x)) (+-comm 1 m) (hLev y x)) g f (sym p) y x wedgeConFun G H l (suc n) (suc m) P {A = A} hlev f g p north y = f y - wedgeConFun G H l (suc n) (suc m) P {A = A} hlev f g p south y = subst (λ x → A x y) (merid ptS) (f y) + wedgeConFun G H l (suc n) (suc m) P {A = A} hlev f g p south y = subst (λ x → A x y) (merid ptEM-raw) (f y) wedgeConFun G H zero (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = ⊥-path i where - ⊥-path : PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptS) (f y)) + ⊥-path : PathP (λ i → A (merid a i) y) (f y) (subst (λ x → A x y) (merid ptEM-raw) (f y)) ⊥-path = ⊥-rec (snotz P) wedgeConFun G H (suc l) (suc n) (suc m) P {A = A} hlev f g p (merid a i) y = mmain a y i module _ where - llem₃ : g south ≡ (subst (λ x → A x ptS) (merid ptS) (f ptS)) - llem₃ = ((λ i → transp (λ j → A (merid ptS (~ i ∨ j)) ptS) (~ i) (g (merid ptS (~ i))))) - ∙ cong (subst (λ x → A x ptS) (merid ptS)) (sym p) + llem₃ : g south ≡ (subst (λ x → A x ptEM-raw) (merid ptEM-raw) (f ptEM-raw)) + llem₃ = ((λ i → transp (λ j → A (merid ptEM-raw (~ i ∨ j)) ptEM-raw) (~ i) (g (merid ptEM-raw (~ i))))) + ∙ cong (subst (λ x → A x ptEM-raw) (merid ptEM-raw)) (sym p) llem₃' : Square - (λ i → transp (λ j → A (merid ptS (~ i ∨ j)) ptS) (~ i) (g (merid ptS (~ i)))) - (refl {x = subst (λ x → A x ptS) (merid ptS) (f ptS)}) + (λ i → transp (λ j → A (merid ptEM-raw (~ i ∨ j)) ptEM-raw) (~ i) (g (merid ptEM-raw (~ i)))) + (refl {x = subst (λ x → A x ptEM-raw) (merid ptEM-raw) (f ptEM-raw)}) llem₃ - ((cong (transport (λ z → A (merid ptS z) ptS)) (sym p))) + ((cong (transport (λ z → A (merid ptEM-raw z) ptEM-raw)) (sym p))) llem₃' i j = - hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptS (~ j ∨ j₁)) ptS) (~ j) (g (merid ptS (~ j))) - ; (i = i1) → subst (λ x → A x ptS) (merid ptS) (p (~ k)) - ; (j = i1) → cong (transport (λ z → A (merid ptS z) ptS)) (sym p) (i ∧ k)}) - (transp (λ j₁ → A (merid ptS ((~ j ∧ ~ i) ∨ j₁)) ptS) (~ j ∧ ~ i) (g (merid ptS (~ j ∧ ~ i)))) + hcomp (λ k → λ { (i = i0) → transp (λ j₁ → A (merid ptEM-raw (~ j ∨ j₁)) ptEM-raw) (~ j) (g (merid ptEM-raw (~ j))) + ; (i = i1) → subst (λ x → A x ptEM-raw) (merid ptEM-raw) (p (~ k)) + ; (j = i1) → cong (transport (λ z → A (merid ptEM-raw z) ptEM-raw)) (sym p) (i ∧ k)}) + (transp (λ j₁ → A (merid ptEM-raw ((~ j ∧ ~ i) ∨ j₁)) ptEM-raw) (~ j ∧ ~ i) (g (merid ptEM-raw (~ j ∧ ~ i)))) - llem₄ : (λ i₁ → transp (λ j → A (merid ptS (j ∧ i₁)) ptS) (~ i₁) (f ptS)) + llem₄ : (λ i₁ → transp (λ j → A (merid ptEM-raw (j ∧ i₁)) ptEM-raw) (~ i₁) (f ptEM-raw)) ≡ (λ i₁ → hcomp (λ k → λ { (i₁ = i0) → p (~ k) ; (i₁ = i1) → llem₃ k }) - (g (merid ptS i₁))) + (g (merid ptEM-raw i₁))) llem₄ i j = - hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptS (z ∧ j)) ptS) (~ j) (p (~ k)) + hcomp (λ k → λ { (i = i0) → transp (λ z → A (merid ptEM-raw (z ∧ j)) ptEM-raw) (~ j) (p (~ k)) ; (j = i0) → p (~ k) ; (j = i1) → llem₃' k (~ i)}) - (transp (λ z → A (merid ptS ((i ∨ z) ∧ j)) ptS) (i ∨ ~ j) (g (merid ptS (i ∧ j)))) + (transp (λ z → A (merid ptEM-raw ((i ∨ z) ∧ j)) ptEM-raw) (i ∨ ~ j) (g (merid ptEM-raw (i ∧ j)))) mmain : (a : _) (y : _) → PathP (λ i → A (merid a i) y) (f y) - (subst (λ x → A x y) (merid ptS) (f y)) + (subst (λ x → A x y) (merid ptEM-raw) (f y)) mmain = wedgeConFun G H l n (suc m) (cong predℕ P) (λ _ _ → isOfHLevelPathP' (suc (n + (suc (suc m)))) (hlev _ _) _ _) - (λ x i → transp (λ j → A (merid ptS (j ∧ i)) x) (~ i) (f x)) + (λ x i → transp (λ j → A (merid ptEM-raw (j ∧ i)) x) (~ i) (f x)) (λ y i → hcomp (λ k → λ { (i = i0) → p (~ k) ; (i = i1) → llem₃ k}) (g (merid y i))) llem₄ - mainR : (x : _) → mmain x ptS + mainR : (x : _) → mmain x ptEM-raw ≡ λ i → hcomp (λ k → λ { (i = i0) → p (~ k) ; (i = i1) → llem₃ k}) (g (merid x i)) mainR x = wedgeconRight G H l n (suc m) (cong predℕ P) (λ _ _ → isOfHLevelPathP' (suc (n + (suc (suc m)))) (hlev _ _) _ _) - (λ x i → transp (λ j → A (merid ptS (j ∧ i)) x) (~ i) (f x)) + (λ x i → transp (λ j → A (merid ptEM-raw (j ∧ i)) x) (~ i) (f x)) (λ y i → hcomp (λ k → λ { (i = i0) → p (~ k) ; (i = i1) → llem₃ k}) (g (merid y i))) @@ -253,7 +253,7 @@ private wedgeconRight G H zero (suc n) (suc m) P {A = A} hLev f g p x = ⊥-rec (snotz P) wedgeconRight G H l (suc n) (suc m) P {A = A} hLev f g p north = p wedgeconRight G H l (suc n) (suc m) P {A = A} hLev f g p south = - sym (llem₃ G H _ n m refl hLev f g p ptS i0 ptS) + sym (llem₃ G H _ n m refl hLev f g p ptEM-raw i0 ptEM-raw) wedgeconRight G H (suc l) (suc n) (suc m) P {A = A} hLev f g p (merid a i) k = help k i where @@ -261,25 +261,25 @@ private llem i j = hcomp (λ k → λ { (i = i1) → g (merid a j) ; (j = i0) → p (i ∨ ~ k) - ; (j = i1) → llem₃ G H l n m P hLev f g p ptS i0 ptS (~ i ∧ k)}) + ; (j = i1) → llem₃ G H l n m P hLev f g p ptEM-raw i0 ptEM-raw (~ i ∧ k)}) (g (merid a j)) - help : PathP (λ k → PathP (λ i → A (merid a i) ptS) - (p k) (llem₃ G H l n m P hLev f g p ptS i0 ptS (~ k))) + help : PathP (λ k → PathP (λ i → A (merid a i) ptEM-raw) + (p k) (llem₃ G H l n m P hLev f g p ptEM-raw i0 ptEM-raw (~ k))) (λ i → mmain G H l n m P hLev f g p a i north a north i) (cong g (merid a)) - help = mainR G H l n m P hLev f g p a i0 ptS a ◁ llem + help = mainR G H l n m P hLev f g p a i0 ptEM-raw a ◁ llem module wedgeConEM (G : AbGroup ℓ) (H : AbGroup ℓ') (n m : ℕ) {A : EM-raw G (suc n) → EM-raw H (suc m) → Type ℓ''} (hLev : ((x : _) (y : _) → isOfHLevel (suc n + suc m) (A x y))) - (f : (x : _) → A ptS x) - (g : (x : _) → A x ptS) - (p : f ptS ≡ g ptS) where + (f : (x : _) → A ptEM-raw x) + (g : (x : _) → A x ptEM-raw) + (p : f ptEM-raw ≡ g ptEM-raw) where fun : (x : EM-raw G (suc n)) (y : EM-raw H (suc m)) → A x y fun = wedgeConFun G H (n + m) n m refl hLev f g p - left : (x : EM-raw H (suc m)) → fun ptS x ≡ f x + left : (x : EM-raw H (suc m)) → fun ptEM-raw x ≡ f x left = wedgeconLeft G H (n + m) n m refl hLev f g p - right : (x : EM-raw G (suc n)) → fun x ptS ≡ g x + right : (x : EM-raw G (suc n)) → fun x ptEM-raw ≡ g x right = wedgeconRight G H (n + m) n m refl hLev f g p diff --git a/Cubical/Data/Nat/Base.agda b/Cubical/Data/Nat/Base.agda index 798728362..b18b84087 100644 --- a/Cubical/Data/Nat/Base.agda +++ b/Cubical/Data/Nat/Base.agda @@ -42,6 +42,14 @@ elim : ∀ {ℓ} {A : ℕ → Type ℓ} elim a₀ _ zero = a₀ elim a₀ f (suc n) = f n (elim a₀ f n) +elim+2 : ∀ {ℓ} {A : ℕ → Type ℓ} → A 0 → A 1 + → ((n : ℕ) → (A (suc n) → A (suc (suc n)))) + → (n : ℕ) → A n +elim+2 a0 a1 ind zero = a0 +elim+2 a0 a1 ind (suc zero) = a1 +elim+2 {A = A} a0 a1 ind (suc (suc n)) = + ind n (elim+2 {A = A} a0 a1 ind (suc n)) + isEven isOdd : ℕ → Bool isEven zero = true isEven (suc n) = isOdd n diff --git a/Cubical/Foundations/Path.agda b/Cubical/Foundations/Path.agda index 4452fc272..86b1c3d08 100644 --- a/Cubical/Foundations/Path.agda +++ b/Cubical/Foundations/Path.agda @@ -351,3 +351,16 @@ compPathR→PathP∙∙ {p = p} {q = q} {r = r} {s = s} P j i = ; (j = i0) → r i ; (j = i1) → doubleCompPath-filler p s (sym q) (~ k) i}) (P j i) + +comm→PathP : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} + → p ∙ s ≡ r ∙ q + → PathP (λ i → p i ≡ q i) r s +comm→PathP {p = p} {q = q} {r = r} {s = s} P i j = + hcomp + (λ k → λ + { (i = i0) → r (j ∧ k) + ; (i = i1) → s (j ∨ ~ k) + ; (j = i0) → compPath-filler p s (~ k) i + ; (j = i1) → compPath-filler' r q (~ k) i + }) + (P j i) diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 567a9857c..94da7d6ee 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -493,13 +493,3 @@ open Lift public liftExt : ∀ {A : Type ℓ} {a b : Lift {ℓ} {ℓ'} A} → (lower a ≡ lower b) → a ≡ b liftExt x i = lift (x i) - --- useful lemma... TODO: move (but where?) -move4 : ∀ {ℓ} {A : Type ℓ} (x y z w : A) (_+_ : A → A → A) - → ((x y z : A) → x + (y + z) ≡ (x + y) + z) - → ((x y : A) → x + y ≡ y + x) - → (x + y) + (z + w) ≡ ((x + z) + (y + w)) -move4 x y z w _+_ assoc comm = - sym (assoc x y (z + w)) - ∙∙ cong (x +_) (assoc y z w ∙∙ cong (_+ w) (comm y z) ∙∙ sym (assoc z y w)) - ∙∙ assoc x z (y + w) From cdf715f6fcd09eb41f57d17439f39c964a630a82 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 3 May 2022 00:54:06 +0200 Subject: [PATCH 30/30] fix --- Cubical/Algebra/Group/EilenbergMacLane/Properties.agda | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda b/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda index 305bc9085..cbc712a7c 100644 --- a/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda +++ b/Cubical/Algebra/Group/EilenbergMacLane/Properties.agda @@ -147,8 +147,7 @@ module _ (Ĝ : Group ℓ) where ua (rightEquiv (g · h)) ∎) REComp : (g h : G) → Square (RE g) (RE (g · h)) refl (RE h) - REComp g h = Σ≡Set (λ _ → isProp→isSet isPropIsSet) - (RE g) (RE (g · h)) refl (RE h) (lemma₁ g h) + REComp g h = ΣSquareSet (λ _ → isProp→isSet isPropIsSet) (lemma₁ g h) Codes : EM₁ Ĝ → Type ℓ Codes x = CodesSet x .fst