diff --git a/CONTRIBUTING.md b/CONTRIBUTING.md index 0f5d4a836b..ea24b2d457 100644 --- a/CONTRIBUTING.md +++ b/CONTRIBUTING.md @@ -111,6 +111,8 @@ When preparing a PR here are some general guidelines: - Avoid importing `Foundations.Everything`; import only the modules in `Foundations` you are using. Be reasonably specific in general when importing. + In particular, avoid importing useless files or useless renaming + and try to group them by folder like `Foundations` or `Data` - Avoid `public` imports, except in modules that are specifically meant to collect and re-export results from several modules. diff --git a/Cubical/Algebra/AbGroup/Instances/Direct-Sum.agda b/Cubical/Algebra/AbGroup/Instances/DirectSum.agda similarity index 78% rename from Cubical/Algebra/AbGroup/Instances/Direct-Sum.agda rename to Cubical/Algebra/AbGroup/Instances/DirectSum.agda index 42c4d24bfb..b852136639 100644 --- a/Cubical/Algebra/AbGroup/Instances/Direct-Sum.agda +++ b/Cubical/Algebra/AbGroup/Instances/DirectSum.agda @@ -1,12 +1,12 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.AbGroup.Instances.Direct-Sum where +module Cubical.Algebra.AbGroup.Instances.DirectSum where open import Cubical.Foundations.Prelude open import Cubical.Algebra.AbGroup -open import Cubical.Algebra.Direct-Sum.Base -open import Cubical.Algebra.Direct-Sum.Properties +open import Cubical.Algebra.DirectSum.Base +open import Cubical.Algebra.DirectSum.Properties open import Cubical.Algebra.Polynomials.Multivariate.Base diff --git a/Cubical/Algebra/AbGroup/TensorProduct.agda b/Cubical/Algebra/AbGroup/TensorProduct.agda index f23aef11fc..94a1f6a204 100644 --- a/Cubical/Algebra/AbGroup/TensorProduct.agda +++ b/Cubical/Algebra/AbGroup/TensorProduct.agda @@ -157,10 +157,10 @@ module _ {AGr : AbGroup ℓ} {BGr : AbGroup ℓ'} where ∃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) , unlist++ l1 l2 ∙ cong₂ _+⊗_ p q ∣} + ⊗elimProp (λ _ → squash₁) + (λ a b → ∣ [ a , b ] , ⊗rUnit (a ⊗ b) ∣₁) + λ x y → rec2 squash₁ λ {(l1 , p) (l2 , q) + → ∣ (l1 ++ l2) , unlist++ l1 l2 ∙ cong₂ _+⊗_ p q ∣₁} ⊗elimPropUnlist : ∀ {ℓ} {C : AGr ⨂₁ BGr → Type ℓ} → ((x : _) → isProp (C x)) diff --git a/Cubical/Algebra/CommAlgebra/FPAlgebra.agda b/Cubical/Algebra/CommAlgebra/FPAlgebra.agda index f9b1c0104a..7e96312546 100644 --- a/Cubical/Algebra/CommAlgebra/FPAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra/FPAlgebra.agda @@ -241,7 +241,7 @@ module _ {R : CommRing ℓ} where equiv : CommAlgebraEquiv (FPAlgebra n relations) A isFPAlgebra : (A : CommAlgebra R ℓ) → Type _ - isFPAlgebra A = ∥ FinitePresentation A ∥ + isFPAlgebra A = ∥ FinitePresentation A ∥₁ isFPAlgebraIsProp : {A : CommAlgebra R ℓ} → isProp (isFPAlgebra A) isFPAlgebraIsProp = isPropPropTrunc diff --git a/Cubical/Algebra/CommAlgebra/Localisation.agda b/Cubical/Algebra/CommAlgebra/Localisation.agda index 2d4f4fb062..11d8ac2d58 100644 --- a/Cubical/Algebra/CommAlgebra/Localisation.agda +++ b/Cubical/Algebra/CommAlgebra/Localisation.agda @@ -302,7 +302,7 @@ module DoubleAlgLoc (R : CommRing ℓ) (f g : (fst R)) where helper2 : (α : FinVec (fst R) 1) → x ^ n ≡ linearCombination R α (replicateFinVec 1 y) → x ∈ᵢ √ ⟨ y · x ⟩ - helper2 α p = ∣ (suc n) , ∣ α , cong (x ·_) p ∙ useSolver x y (α zero) ∣ ∣ + helper2 α p = ∣ (suc n) , ∣ α , cong (x ·_) p ∙ useSolver x y (α zero) ∣₁ ∣₁ where useSolver : ∀ x y a → x · (a · y + 0r) ≡ a · (y · x) + 0r useSolver = solve R diff --git a/Cubical/Algebra/CommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebra/Properties.agda index ca51c36ef5..e6ef4ed706 100644 --- a/Cubical/Algebra/CommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/Properties.agda @@ -300,7 +300,7 @@ recPT→CommAlgebra : {R : CommRing ℓ} {A : Type ℓ'} (𝓕 : A → CommAlge → (σ : ∀ x y → CommAlgebraEquiv (𝓕 x) (𝓕 y)) → (∀ x y z → σ x z ≡ compCommAlgebraEquiv (σ x y) (σ y z)) ------------------------------------------------------ - → ∥ A ∥ → CommAlgebra R ℓ'' + → ∥ A ∥₁ → CommAlgebra R ℓ'' recPT→CommAlgebra 𝓕 σ compCoh = GpdElim.rec→Gpd isGroupoidCommAlgebra 𝓕 (3-ConstantCompChar 𝓕 (λ x y → uaCommAlgebra (σ x y)) λ x y z → sym ( cong uaCommAlgebra (compCoh x y z) diff --git a/Cubical/Algebra/CommRing/FGIdeal.agda b/Cubical/Algebra/CommRing/FGIdeal.agda index 08fe088cf1..60a2d6e308 100644 --- a/Cubical/Algebra/CommRing/FGIdeal.agda +++ b/Cubical/Algebra/CommRing/FGIdeal.agda @@ -84,7 +84,7 @@ module _ (Ring@(R , str) : CommRing ℓ) where {- 0r is the trivial linear Combination -} isLinearCombination0 : {n : ℕ} (V : FinVec R n) → isLinearCombination V 0r - isLinearCombination0 V = ∣ _ , sym (dist0 V) ∣ + isLinearCombination0 V = ∣ _ , sym (dist0 V) ∣₁ {- Linear combinations are stable under left multiplication -} isLinearCombinationL· : {n : ℕ} (V : FinVec R n) (r : R) {x : R} @@ -217,7 +217,7 @@ module _ (R' : CommRing ℓ) where λ i → ·Closed (I .snd) _ (∀i→Vi∈I i)) indInIdeal : ∀ {n : ℕ} (U : FinVec R n) (i : Fin n) → U i ∈ ⟨ U ⟩ - indInIdeal U i = ∣ (δ i) , sym (∑Mul1r _ U i) ∣ + indInIdeal U i = ∣ (δ i) , sym (∑Mul1r _ U i) ∣₁ sucIncl : ∀ {n : ℕ} (U : FinVec R (ℕsuc n)) → ⟨ U ∘ suc ⟩ ⊆ ⟨ U ⟩ sucIncl U x = PT.map λ (α , x≡∑αUsuc) → (λ { zero → 0r ; (suc i) → α i }) , x≡∑αUsuc ∙ path _ _ @@ -227,7 +227,7 @@ module _ (R' : CommRing ℓ) where emptyFGIdeal : ∀ (V : FinVec R 0) → ⟨ V ⟩ ≡ 0Ideal emptyFGIdeal V = CommIdeal≡Char (λ _ → PT.rec (is-set _ _) snd) - (λ _ x≡0 → ∣ (λ ()) , x≡0 ∣) + (λ _ x≡0 → ∣ (λ ()) , x≡0 ∣₁) 0FGIdealLIncl : {n : ℕ} → ⟨ replicateFinVec n 0r ⟩ ⊆ 0Ideal 0FGIdealLIncl x = PT.elim (λ _ → is-set _ _) @@ -244,7 +244,7 @@ module _ (R' : CommRing ℓ) where FGIdealAddLemmaLIncl : {n m : ℕ} (U : FinVec R n) (V : FinVec R m) → ⟨ U ++Fin V ⟩ ⊆ (⟨ U ⟩ +i ⟨ V ⟩) FGIdealAddLemmaLIncl {n = ℕzero} U V x x∈⟨V⟩ = - ∣ (0r , x) , ⟨ U ⟩ .snd .contains0 , x∈⟨V⟩ , sym (+Lid x) ∣ + ∣ (0r , x) , ⟨ U ⟩ .snd .contains0 , x∈⟨V⟩ , sym (+Lid x) ∣₁ FGIdealAddLemmaLIncl {n = ℕsuc n} U V x = PT.rec isPropPropTrunc helperΣ where helperΣ : Σ[ α ∈ FinVec R _ ] (x ≡ ∑ λ i → α i · (U ++Fin V) i) → x ∈ (⟨ U ⟩ +i ⟨ V ⟩) @@ -258,7 +258,7 @@ module _ (R' : CommRing ℓ) where sumIncl : (∑ λ i → (α ∘ suc) i · ((U ∘ suc) ++Fin V) i) ∈ (⟨ U ⟩ +i ⟨ V ⟩) sumIncl = let sum = ∑ λ i → (α ∘ suc) i · ((U ∘ suc) ++Fin V) i in +iRespLincl ⟨ U ∘ suc ⟩ ⟨ U ⟩ ⟨ V ⟩ (sucIncl U) sum - (FGIdealAddLemmaLIncl (U ∘ suc) V _ ∣ (α ∘ suc) , refl ∣) + (FGIdealAddLemmaLIncl (U ∘ suc) V _ ∣ (α ∘ suc) , refl ∣₁) FGIdealAddLemmaRIncl : {n m : ℕ} (U : FinVec R n) (V : FinVec R m) → (⟨ U ⟩ +i ⟨ V ⟩) ⊆ ⟨ U ++Fin V ⟩ @@ -269,7 +269,7 @@ module _ (R' : CommRing ℓ) where → Σ[ β ∈ FinVec R _ ] (z ≡ ∑ λ i → β i · V i) → x ≡ y + z → x ∈ ⟨ U ++Fin V ⟩ - helperΣ (y , z) (α , y≡∑αU) (β , z≡∑βV) x≡y+z = ∣ (α ++Fin β) , path ∣ + helperΣ (y , z) (α , y≡∑αU) (β , z≡∑βV) x≡y+z = ∣ (α ++Fin β) , path ∣₁ where path : x ≡ ∑ λ i → (α ++Fin β) i · (U ++Fin V) i path = x ≡⟨ x≡y+z ⟩ @@ -359,7 +359,7 @@ module GeneratingPowers (R' : CommRing ℓ) (n : ℕ) where lemma : (m : ℕ) (α U : FinVec R (ℕsuc m)) → (linearCombination R' α U) ^ ((ℕsuc m) ·ℕ n) ∈ ⟨ U ⁿ ⟩ - lemma ℕzero α U = ∣ α ⁿ , path ∣ + lemma ℕzero α U = ∣ α ⁿ , path ∣₁ where path : (α zero · U zero + 0r) ^ (n +ℕ 0) ≡ α zero ^ n · U zero ^ n + 0r path = (α zero · U zero + 0r) ^ (n +ℕ 0) ≡⟨ cong (_^ (n +ℕ 0)) (+Rid _) ⟩ diff --git a/Cubical/Algebra/CommRing/Ideal.agda b/Cubical/Algebra/CommRing/Ideal.agda index 54f8308b90..661caec6c4 100644 --- a/Cubical/Algebra/CommRing/Ideal.agda +++ b/Cubical/Algebra/CommRing/Ideal.agda @@ -113,7 +113,7 @@ module CommIdeal (R' : CommRing ℓ) where +ClosedΣ ((y₁ , z₁) , y₁∈I , z₁∈J , x₁≡y₁+z₁) ((y₂ , z₂) , y₂∈I , z₂∈J , x₂≡y₂+z₂) = (y₁ + y₂ , z₁ + z₂) , +Closed (snd I) y₁∈I y₂∈I , +Closed (snd J) z₁∈J z₂∈J , cong₂ (_+_) x₁≡y₁+z₁ x₂≡y₂+z₂ ∙ +ShufflePairs _ _ _ _ - contains0 (snd (I +i J)) = ∣ (0r , 0r) , contains0 (snd I) , contains0 (snd J) , sym (+Rid _) ∣ + contains0 (snd (I +i J)) = ∣ (0r , 0r) , contains0 (snd I) , contains0 (snd J) , sym (+Rid _) ∣₁ ·Closed (snd (I +i J)) {x = x} r = map ·ClosedΣ where ·ClosedΣ : Σ[ (y₁ , z₁) ∈ (R × R) ] ((y₁ ∈ I) × (z₁ ∈ J) × (x ≡ y₁ + z₁)) @@ -135,16 +135,16 @@ module CommIdeal (R' : CommRing ℓ) where → subst-∈ I (sym (x≡y+z ∙∙ cong (_+ z) y≡0 ∙∙ +Lid z)) z∈I +iLidRIncl : ∀ (I : CommIdeal) → I ⊆ (0Ideal +i I) - +iLidRIncl I x x∈I = ∣ (0r , x) , refl , x∈I , sym (+Lid _) ∣ + +iLidRIncl I x x∈I = ∣ (0r , x) , refl , x∈I , sym (+Lid _) ∣₁ +iLid : ∀ (I : CommIdeal) → 0Ideal +i I ≡ I +iLid I = CommIdeal≡Char (+iLidLIncl I) (+iLidRIncl I) +iLincl : ∀ (I J : CommIdeal) → I ⊆ (I +i J) - +iLincl I J x x∈I = ∣ (x , 0r) , x∈I , J .snd .contains0 , sym (+Rid x) ∣ + +iLincl I J x x∈I = ∣ (x , 0r) , x∈I , J .snd .contains0 , sym (+Rid x) ∣₁ +iRincl : ∀ (I J : CommIdeal) → J ⊆ (I +i J) - +iRincl I J x x∈J = ∣ (0r , x) , I .snd .contains0 , x∈J , sym (+Lid x) ∣ + +iRincl I J x x∈J = ∣ (0r , x) , I .snd .contains0 , x∈J , sym (+Lid x) ∣₁ +iRespLincl : ∀ (I J K : CommIdeal) → I ⊆ J → (I +i K) ⊆ (J +i K) +iRespLincl I J K I⊆J x = map λ ((y , z) , y∈I , z∈K , x≡y+z) → ((y , z) , I⊆J y y∈I , z∈K , x≡y+z) @@ -153,15 +153,15 @@ module CommIdeal (R' : CommRing ℓ) where +iAssocLIncl I J K x = elim (λ _ → ((I +i J) +i K) .fst x .snd) (uncurry3 λ (y , z) y∈I → elim (λ _ → isPropΠ λ _ → ((I +i J) +i K) .fst x .snd) λ ((u , v) , u∈J , v∈K , z≡u+v) x≡y+z - → ∣ (y + u , v) , ∣ _ , y∈I , u∈J , refl ∣ , v∈K - , x≡y+z ∙∙ cong (y +_) z≡u+v ∙∙ +Assoc _ _ _ ∣) + → ∣ (y + u , v) , ∣ _ , y∈I , u∈J , refl ∣₁ , v∈K + , x≡y+z ∙∙ cong (y +_) z≡u+v ∙∙ +Assoc _ _ _ ∣₁) +iAssocRIncl : ∀ (I J K : CommIdeal) → ((I +i J) +i K) ⊆ (I +i (J +i K)) +iAssocRIncl I J K x = elim (λ _ → (I +i (J +i K)) .fst x .snd) (uncurry3 λ (y , z) → elim (λ _ → isPropΠ2 λ _ _ → (I +i (J +i K)) .fst x .snd) λ ((u , v) , u∈I , v∈J , y≡u+v) z∈K x≡y+z - → ∣ (u , v + z) , u∈I , ∣ _ , v∈J , z∈K , refl ∣ - , x≡y+z ∙∙ cong (_+ z) y≡u+v ∙∙ sym (+Assoc _ _ _) ∣) + → ∣ (u , v + z) , u∈I , ∣ _ , v∈J , z∈K , refl ∣₁ + , x≡y+z ∙∙ cong (_+ z) y≡u+v ∙∙ sym (+Assoc _ _ _) ∣₁) +iAssoc : ∀ (I J K : CommIdeal) → I +i (J +i K) ≡ (I +i J) +i K +iAssoc I J K = CommIdeal≡Char (+iAssocLIncl I J K) (+iAssocRIncl I J K) @@ -171,7 +171,7 @@ module CommIdeal (R' : CommRing ℓ) where → subst-∈ I (sym x≡y+z) (I .snd .+Closed y∈I z∈I) +iIdemRIncl : ∀ (I : CommIdeal) → I ⊆ (I +i I) - +iIdemRIncl I x x∈I = ∣ (0r , x) , I .snd .contains0 , x∈I , sym (+Lid _) ∣ + +iIdemRIncl I x x∈I = ∣ (0r , x) , I .snd .contains0 , x∈I , sym (+Lid _) ∣₁ +iIdem : ∀ (I : CommIdeal) → I +i I ≡ I +iIdem I = CommIdeal≡Char (+iIdemLIncl I) (+iIdemRIncl I) @@ -195,7 +195,7 @@ module CommIdeal (R' : CommRing ℓ) where , ++FinPres∈ (J .fst) ∀βi∈J ∀δi∈J , cong₂ (_+_) x≡∑αβ y≡∑γδ ∙∙ sym (∑Split++ (λ i → α i · β i) (λ i → γ i · δ i)) ∙∙ ∑Ext (mul++dist α β γ δ) - contains0 (snd (I ·i J)) = ∣ 0 , ((λ ()) , (λ ())) , (λ ()) , (λ ()) , refl ∣ + contains0 (snd (I ·i J)) = ∣ 0 , ((λ ()) , (λ ())) , (λ ()) , (λ ()) , refl ∣₁ ·Closed (snd (I ·i J)) r = map λ (n , (α , β) , ∀αi∈I , ∀βi∈J , x≡∑αβ) → n , ((λ i → r · α i) , β) , (λ i → I .snd .·Closed r (∀αi∈I i)) , ∀βi∈J @@ -205,7 +205,7 @@ module CommIdeal (R' : CommRing ℓ) where prodInProd : ∀ (I J : CommIdeal) (x y : R) → x ∈ I → y ∈ J → (x · y) ∈ (I ·i J) prodInProd _ _ x y x∈I y∈J = - ∣ 1 , ((λ _ → x) , λ _ → y) , (λ _ → x∈I) , (λ _ → y∈J) , sym (+Rid _) ∣ + ∣ 1 , ((λ _ → x) , λ _ → y) , (λ _ → x∈I) , (λ _ → y∈J) , sym (+Rid _) ∣₁ ·iLincl : ∀ (I J : CommIdeal) → (I ·i J) ⊆ I ·iLincl I J x = elim (λ _ → I .fst x .snd) @@ -220,7 +220,7 @@ module CommIdeal (R' : CommRing ℓ) where ·iComm I J = CommIdeal≡Char (·iComm⊆ I J) (·iComm⊆ J I) I⊆I1 : ∀ (I : CommIdeal) → I ⊆ (I ·i 1Ideal) - I⊆I1 I x x∈I = ∣ 1 , ((λ _ → x) , λ _ → 1r) , (λ _ → x∈I) , (λ _ → lift tt) , useSolver x ∣ + I⊆I1 I x x∈I = ∣ 1 , ((λ _ → x) , λ _ → 1r) , (λ _ → x∈I) , (λ _ → lift tt) , useSolver x ∣₁ where useSolver : ∀ x → x ≡ x · 1r + 0r useSolver = solve R' @@ -273,7 +273,7 @@ module CommIdeal (R' : CommRing ℓ) where (λ ((γi , δi) , γi∈J , δi∈K , βi≡γi+δi) → ∣ (α i · γi , α i · δi) , prodInProd I J _ _ (α∈I i) γi∈J , prodInProd I K _ _ (α∈I i) δi∈K - , cong (α i ·_) βi≡γi+δi ∙ ·Rdist+ _ _ _ ∣) + , cong (α i ·_) βi≡γi+δi ∙ ·Rdist+ _ _ _ ∣₁) (β∈J+K i)) ·iRdist+iRIncl : ∀ (I J K : CommIdeal) → ((I ·i J) +i (I ·i K)) ⊆ (I ·i (J +i K)) diff --git a/Cubical/Algebra/CommRing/Instances/Int.agda b/Cubical/Algebra/CommRing/Instances/Int.agda index ee4009bd43..cda3242b49 100644 --- a/Cubical/Algebra/CommRing/Instances/Int.agda +++ b/Cubical/Algebra/CommRing/Instances/Int.agda @@ -5,18 +5,18 @@ open import Cubical.Foundations.Prelude open import Cubical.Algebra.CommRing open import Cubical.Data.Int as Int - renaming ( ℤ to ℤType ; _+_ to _+ℤ_; _·_ to _·ℤ_; -_ to -ℤ_) + renaming ( ℤ to ℤ ; _+_ to _+ℤ_; _·_ to _·ℤ_; -_ to -ℤ_) open CommRingStr -ℤ : CommRing ℓ-zero -fst ℤ = ℤType -0r (snd ℤ) = 0 -1r (snd ℤ) = 1 -_+_ (snd ℤ) = _+ℤ_ -_·_ (snd ℤ) = _·ℤ_ -- snd ℤ = -ℤ_ -isCommRing (snd ℤ) = isCommRingℤ +ℤCommRing : CommRing ℓ-zero +fst ℤCommRing = ℤ +0r (snd ℤCommRing) = 0 +1r (snd ℤCommRing) = 1 +_+_ (snd ℤCommRing) = _+ℤ_ +_·_ (snd ℤCommRing) = _·ℤ_ +- snd ℤCommRing = -ℤ_ +isCommRing (snd ℤCommRing) = isCommRingℤ where abstract isCommRingℤ : IsCommRing 0 1 _+ℤ_ _·ℤ_ -ℤ_ diff --git a/Cubical/Algebra/CommRing/Instances/MultivariatePoly-Quotient.agda b/Cubical/Algebra/CommRing/Instances/MultivariatePoly-Quotient.agda index 31009e2bf8..37f0888eb4 100644 --- a/Cubical/Algebra/CommRing/Instances/MultivariatePoly-Quotient.agda +++ b/Cubical/Algebra/CommRing/Instances/MultivariatePoly-Quotient.agda @@ -14,7 +14,7 @@ open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.QuotientRing -open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤ to ℤCR) +open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤCommRing to ℤCR) open import Cubical.Algebra.Polynomials.Multivariate.Base open import Cubical.Algebra.CommRing.Instances.MultivariatePoly diff --git a/Cubical/Algebra/CommRing/Instances/MultivariatePoly-notationZ.agda b/Cubical/Algebra/CommRing/Instances/MultivariatePoly-notationZ.agda index 503216bd0e..17040f67ba 100644 --- a/Cubical/Algebra/CommRing/Instances/MultivariatePoly-notationZ.agda +++ b/Cubical/Algebra/CommRing/Instances/MultivariatePoly-notationZ.agda @@ -23,25 +23,25 @@ open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-Quotient -- Notations for ℤ polynomial rings ℤ[X] : CommRing ℓ-zero -ℤ[X] = PolyCommRing ℤ 1 +ℤ[X] = PolyCommRing ℤCommRing 1 ℤ[x] : Type ℓ-zero ℤ[x] = fst ℤ[X] ℤ[X,Y] : CommRing ℓ-zero -ℤ[X,Y] = PolyCommRing ℤ 2 +ℤ[X,Y] = PolyCommRing ℤCommRing 2 ℤ[x,y] : Type ℓ-zero ℤ[x,y] = fst ℤ[X,Y] ℤ[X,Y,Z] : CommRing ℓ-zero -ℤ[X,Y,Z] = PolyCommRing ℤ 3 +ℤ[X,Y,Z] = PolyCommRing ℤCommRing 3 ℤ[x,y,z] : Type ℓ-zero ℤ[x,y,z] = fst ℤ[X,Y,Z] ℤ[X1,···,Xn] : (n : ℕ) → CommRing ℓ-zero -ℤ[X1,···,Xn] n = A[X1,···,Xn] ℤ n +ℤ[X1,···,Xn] n = A[X1,···,Xn] ℤCommRing n ℤ[x1,···,xn] : (n : ℕ) → Type ℓ-zero ℤ[x1,···,xn] n = fst (ℤ[X1,···,Xn] n) @@ -51,37 +51,37 @@ open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-Quotient -- Notation for quotiented ℤ polynomial ring : FinVec ℤ[x] 1 - = ℤ 1 0 1 + = ℤCommRing 1 0 1 : FinVec ℤ[x] 1 - = ℤ 1 0 2 + = ℤCommRing 1 0 2 : FinVec ℤ[x] 1 - = ℤ 1 0 3 + = ℤCommRing 1 0 3 : (k : ℕ) → FinVec ℤ[x] 1 - k = ℤ 1 0 k + k = ℤCommRing 1 0 k ℤ[X]/X : CommRing ℓ-zero -ℤ[X]/X = A[X1,···,Xn]/ ℤ 1 0 1 +ℤ[X]/X = A[X1,···,Xn]/ ℤCommRing 1 0 1 ℤ[x]/x : Type ℓ-zero ℤ[x]/x = fst ℤ[X]/X ℤ[X]/X² : CommRing ℓ-zero -ℤ[X]/X² = A[X1,···,Xn]/ ℤ 1 0 2 +ℤ[X]/X² = A[X1,···,Xn]/ ℤCommRing 1 0 2 ℤ[x]/x² : Type ℓ-zero ℤ[x]/x² = fst ℤ[X]/X² ℤ[X]/X³ : CommRing ℓ-zero -ℤ[X]/X³ = A[X1,···,Xn]/ ℤ 1 0 3 +ℤ[X]/X³ = A[X1,···,Xn]/ ℤCommRing 1 0 3 ℤ[x]/x³ : Type ℓ-zero ℤ[x]/x³ = fst ℤ[X]/X³ ℤ[X1,···,Xn]/ : (n : ℕ) → CommRing ℓ-zero -ℤ[X1,···,Xn]/ n = A[X1,···,Xn]/ ℤ n +ℤ[X1,···,Xn]/ n = A[X1,···,Xn]/ ℤCommRing n ℤ[x1,···,xn]/ : (n : ℕ) → Type ℓ-zero ℤ[x1,···,xn]/ n = fst (ℤ[X1,···,Xn]/ n) @@ -91,8 +91,8 @@ open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-Quotient -- they only holds up to a path ℤ'[X]/X : CommRing ℓ-zero -ℤ'[X]/X = A[X1,···,Xn]/ ℤ 1 +ℤ'[X]/X = A[X1,···,Xn]/ ℤCommRing 1 equivℤ[X] : ℤ'[X]/X ≡ ℤ[X]/X -equivℤ[X] = cong (λ X → (A[X1,···,Xn] ℤ 1) / (genIdeal ((A[X1,···,Xn] ℤ 1)) X)) +equivℤ[X] = cong (λ X → (A[X1,···,Xn] ℤCommRing 1) / (genIdeal ((A[X1,···,Xn] ℤCommRing 1)) X)) (funExt (λ {zero → refl })) diff --git a/Cubical/Algebra/CommRing/Instances/MultivariatePoly.agda b/Cubical/Algebra/CommRing/Instances/MultivariatePoly.agda index 849f2461e7..212fc7ba5e 100644 --- a/Cubical/Algebra/CommRing/Instances/MultivariatePoly.agda +++ b/Cubical/Algebra/CommRing/Instances/MultivariatePoly.agda @@ -31,20 +31,19 @@ module _ fst PolyCommRing = Poly A n 0r (snd PolyCommRing) = 0P 1r (snd PolyCommRing) = 1P - _+_ (snd PolyCommRing) = _Poly+_ - _·_ (snd PolyCommRing) = _Poly*_ - - snd PolyCommRing = Poly-inv + _+_ (snd PolyCommRing) = _poly+_ + _·_ (snd PolyCommRing) = _poly*_ + - snd PolyCommRing = polyInv isCommRing (snd PolyCommRing) = makeIsCommRing trunc - Poly+-assoc - Poly+-Rid - Poly+-rinv - Poly+-comm - Poly*-assoc - Poly*-Rid - Poly*-Rdist - Poly*-comm - + poly+Assoc + poly+IdR + poly+InvR + poly+Comm + poly*Assoc + poly*IdR + poly*DistR + poly*Comm ----------------------------------------------------------------------------- diff --git a/Cubical/Algebra/CommRing/Instances/UnivariatePoly.agda b/Cubical/Algebra/CommRing/Instances/UnivariatePoly.agda index 227a3fffe9..bdba5b10af 100644 --- a/Cubical/Algebra/CommRing/Instances/UnivariatePoly.agda +++ b/Cubical/Algebra/CommRing/Instances/UnivariatePoly.agda @@ -3,6 +3,8 @@ module Cubical.Algebra.CommRing.Instances.UnivariatePoly where open import Cubical.Foundations.Prelude +open import Cubical.Data.Nat using (ℕ ; zero ; suc) + open import Cubical.Algebra.CommRing open import Cubical.Algebra.Polynomials.Univariate.Base open import Cubical.Algebra.Polynomials.Univariate.Properties @@ -31,3 +33,8 @@ UnivariatePoly R = (PolyMod.Poly R) , str (PolyModTheory.Poly*Rid R) (PolyModTheory.Poly*LDistrPoly+ R) (PolyModTheory.Poly*Commutative R) + + +nUnivariatePoly : (A' : CommRing ℓ) → (n : ℕ) → CommRing ℓ +nUnivariatePoly A' zero = A' +nUnivariatePoly A' (suc n) = UnivariatePoly (nUnivariatePoly A' n) diff --git a/Cubical/Algebra/CommRing/Localisation/InvertingElements.agda b/Cubical/Algebra/CommRing/Localisation/InvertingElements.agda index d4afbd6b18..8c3fe77ab7 100644 --- a/Cubical/Algebra/CommRing/Localisation/InvertingElements.agda +++ b/Cubical/Algebra/CommRing/Localisation/InvertingElements.agda @@ -69,7 +69,7 @@ module InvertingElementsBase (R' : CommRing ℓ) where -- (n,s≡fⁿ,p) (m,s≡fᵐ,q) then n≤m by p and m≤n by q => n≡m powersFormMultClosedSubset : (f : R) → isMultClosedSubset R' [ f ⁿ|n≥0] - powersFormMultClosedSubset f .containsOne = PT.∣ zero , refl ∣ + 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 @@ -79,9 +79,9 @@ module InvertingElementsBase (R' : CommRing ℓ) where -- a quick fact isContrR[1/0] : isContr R[1/ 0r ] - fst isContrR[1/0] = [ 1r , 0r , ∣ 1 , sym (·Rid 0r) ∣ ] -- everything is equal to 1/0 + fst isContrR[1/0] = [ 1r , 0r , ∣ 1 , sym (·Rid 0r) ∣₁ ] -- everything is equal to 1/0 snd isContrR[1/0] = elimProp (λ _ → squash/ _ _) - λ _ → eq/ _ _ ((0r , ∣ 1 , sym (·Rid 0r) ∣) , useSolver _ _) + λ _ → eq/ _ _ ((0r , ∣ 1 , sym (·Rid 0r) ∣₁) , useSolver _ _) where useSolver : ∀ s r → 0r · 1r · s ≡ 0r · r · 0r useSolver = solve R' @@ -90,7 +90,7 @@ module InvertingElementsBase (R' : CommRing ℓ) where 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) ∣ ] ≡ + ^-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) @@ -101,7 +101,7 @@ module InvertingElementsBase (R' : CommRing ℓ) where -- 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ⁿ) + → (∀ (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 @@ -115,31 +115,31 @@ module InvertingElementsBase (R' : CommRing ℓ) where invElPropElim2 : {f g : R} {P : R[1/ f ] → R[1/ g ] → Type ℓ'} → (∀ x y → isProp (P x y)) - → (∀ (r s : R) (n : ℕ) → P [ r , (f ^ n) , PT.∣ n , refl ∣ ] - [ s , (g ^ n) , PT.∣ n , refl ∣ ]) + → (∀ (r s : R) (n : ℕ) → P [ r , (f ^ n) , PT.∣ n , refl ∣₁ ] + [ s , (g ^ n) , PT.∣ n , refl ∣₁ ]) ---------------------------------------------------------- → ∀ x y → P x y invElPropElim2 {f = f} {g = g} {P = P} PisProp base = invElPropElim (λ _ → isPropΠ (λ _ → PisProp _ _)) reduce1 where - reduce1 : ∀ (r : R) (n : ℕ) (y : R[1/ g ]) → P [ r , f ^ n , ∣ n , refl ∣ ] y + reduce1 : ∀ (r : R) (n : ℕ) (y : R[1/ g ]) → P [ r , f ^ n , ∣ n , refl ∣₁ ] y reduce1 r n = invElPropElim (λ _ → PisProp _ _) reduce2 where - reduce2 : (s : R) (m : ℕ) → P [ r , f ^ n , ∣ n , refl ∣ ] [ s , g ^ m , ∣ m , refl ∣ ] + reduce2 : (s : R) (m : ℕ) → P [ r , f ^ n , ∣ n , refl ∣₁ ] [ s , g ^ m , ∣ m , refl ∣₁ ] reduce2 s m = subst2 P p q (base _ _ l) where l = max m n x : R[1/ f ] - x = [ r , f ^ n , ∣ n , refl ∣ ] + x = [ r , f ^ n , ∣ n , refl ∣₁ ] y : R[1/ g ] - y = [ s , g ^ m , ∣ m , refl ∣ ] + y = [ s , g ^ m , ∣ m , refl ∣₁ ] x' : R[1/ f ] - x' = [ r · f ^ (l ∸ n) , f ^ l , ∣ l , refl ∣ ] + x' = [ r · f ^ (l ∸ n) , f ^ l , ∣ l , refl ∣₁ ] y' : R[1/ g ] - y' = [ s · g ^ (l ∸ m) , g ^ l , ∣ l , refl ∣ ] + y' = [ s · g ^ (l ∸ m) , g ^ l , ∣ l , refl ∣₁ ] p : x' ≡ x - p = eq/ _ _ ((1r , ∣ 0 , refl ∣) , path) + p = eq/ _ _ ((1r , ∣ 0 , refl ∣₁) , path) where useSolver1 : ∀ a b c → 1r · (a · b) · c ≡ a · (b · c) useSolver1 = solve R' @@ -153,7 +153,7 @@ module InvertingElementsBase (R' : CommRing ℓ) where 1r · r · f ^ l ∎ q : y' ≡ y - q = eq/ _ _ ((1r , ∣ 0 , refl ∣) , path) + q = eq/ _ _ ((1r , ∣ 0 , refl ∣₁) , path) where useSolver1 : ∀ a b c → 1r · (a · b) · c ≡ a · (b · c) useSolver1 = solve R' @@ -205,7 +205,7 @@ module RadicalLemma (R' : CommRing ℓ) (f g : (fst R')) where ℕhelper : (n : ℕ) → f ^ n ∈ᵢ ⟨ g ⟩ → (g /1) ∈ R[1/ f ]AsCommRing ˣ ℕhelper n = PT.rec isPropGoal -- fⁿ≡αg → g⁻¹≡α/fⁿ - λ (α , p) → [ (α zero) , (f ^ n) , ∣ n , refl ∣ ] + λ (α , p) → [ (α zero) , (f ^ n) , ∣ n , refl ∣₁ ] , eq/ _ _ ((1r , powersFormMultClosedSubset f .containsOne) , useSolver1 _ _ ∙ sym p ∙ useSolver2 _) where @@ -249,7 +249,7 @@ module DoubleLoc (R' : CommRing ℓ) (f g : (fst R')) where _/1/1 : R → R[1/f][1/g] - r /1/1 = [ [ r , 1r , PT.∣ 0 , refl ∣ ] , 1ᶠ , PT.∣ 0 , refl ∣ ] + 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 @@ -280,9 +280,9 @@ module DoubleLoc (R' : CommRing ℓ) (f g : (fst R')) where 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 ∣ ] + ℕ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 @@ -328,22 +328,22 @@ module DoubleLoc (R' : CommRing ℓ) (f g : (fst R')) where ∥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] : ∃[ 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 ∣ + → ∃[ 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 ∣ ] + 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 + [ g ^ n , 1r , PT.∣ 0 , refl ∣₁ ] ·ᶠ r/1 ≡⟨ cong (_·ᶠ r/1) (^-respects-/1 _ n) ⟩ @@ -362,11 +362,11 @@ module DoubleLoc (R' : CommRing ℓ) (f g : (fst R')) where 0ᶠ ∎ - toGoal : ∃[ n ∈ ℕ ] [ g ^ n · r , 1r , PT.∣ 0 , refl ∣ ] ≡ 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ᶠ + Σ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 @@ -382,7 +382,7 @@ module DoubleLoc (R' : CommRing ℓ) (f g : (fst R')) where 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 ∣ + 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)) @@ -390,7 +390,7 @@ module DoubleLoc (R' : CommRing ℓ) (f g : (fst R')) where Σ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 + Σhelper2 (m , q') = (((f · g) ^ l) , PT.∣ l , refl ∣₁) , path where l = max m n @@ -440,12 +440,12 @@ module DoubleLoc (R' : CommRing ℓ) (f g : (fst R')) where 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 , 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)) ∣ + , (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 @@ -503,28 +503,28 @@ module DoubleLoc (R' : CommRing ℓ) (f g : (fst R')) where 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) + ≡ [ [ 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)) + ≡ [ [ 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) + (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 ∣ ] + ≡ [ 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 .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 ∣ ] + ≡ [ 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 @@ -551,15 +551,15 @@ module DoubleLoc (R' : CommRing ℓ) (f g : (fst R')) 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 (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 ∣) , + coh (n , s≡fg^n) (m , s≡fg^m) = eq/ _ _ ((1ᶠ , PT.∣ 0 , refl ∣₁) , eq/ _ _ ( (1r , powersFormMultClosedSubset R' f .containsOne) , path)) where @@ -606,11 +606,11 @@ module DoubleLoc (R' : CommRing ℓ) (f g : (fst R')) where → (α : Σ[ n ∈ ℕ ] s ≡ (f · g) ^ n) → (β : Σ[ m ∈ ℕ ] s' ≡ (f · g) ^ m) → (γ : Σ[ l ∈ ℕ ] u ≡ (f · g) ^ l) - → ϕ (r , s , PT.∣ α ∣) ≡ ϕ (r' , s' , PT.∣ β ∣) + → ϕ (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)) + , 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) @@ -705,7 +705,7 @@ module DoubleLoc (R' : CommRing ℓ) (f g : (fst R')) where φ≡χ = invElPropElim _ (λ _ → squash/ _ _) ℕcase where ℕcase : (r : R) (n : ℕ) - → φ [ r , (f · g) ^ n , PT.∣ n , refl ∣ ] ≡ χ [ r , (f · g) ^ n , PT.∣ n , refl ∣ ] + → φ [ 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 diff --git a/Cubical/Algebra/CommRing/Localisation/PullbackSquare.agda b/Cubical/Algebra/CommRing/Localisation/PullbackSquare.agda index 65fca53951..a4f89d059e 100644 --- a/Cubical/Algebra/CommRing/Localisation/PullbackSquare.agda +++ b/Cubical/Algebra/CommRing/Localisation/PullbackSquare.agda @@ -130,7 +130,7 @@ module _ (R' : CommRing ℓ) (f g : (fst R')) where where unitHelper : ∀ s → s ∈ₚ [ f ⁿ|n≥0] → s /1ᶠᵍ ∈ₚ (R[1/ (f · g) ]AsCommRing) ˣ unitHelper = powersPropElim (λ s → Units.inverseUniqueness _ (s /1ᶠᵍ)) - λ n → [ g ^ n , (f · g) ^ n , ∣ n , refl ∣ ] + λ n → [ g ^ n , (f · g) ^ n , ∣ n , refl ∣₁ ] , eq/ _ _ ((1r , powersFormMultClosedSubset (f · g) .containsOne) , path n) where @@ -147,7 +147,7 @@ module _ (R' : CommRing ℓ) (f g : (fst R')) where where unitHelper : ∀ s → s ∈ₚ [ g ⁿ|n≥0] → s /1ᶠᵍ ∈ₚ (R[1/ (f · g) ]AsCommRing) ˣ unitHelper = powersPropElim (λ s → Units.inverseUniqueness _ (s /1ᶠᵍ)) - λ n → [ f ^ n , (f · g) ^ n , ∣ n , refl ∣ ] + λ n → [ f ^ n , (f · g) ^ n , ∣ n , refl ∣₁ ] , eq/ _ _ ((1r , powersFormMultClosedSubset (f · g) .containsOne) , path n) where @@ -239,9 +239,9 @@ module _ (R' : CommRing ℓ) (f g : (fst R')) where equalizerLemma 1∈⟨f,g⟩ = invElPropElim2 (λ _ _ → isPropΠ (λ _ → isProp∃!)) baseCase where baseCase : ∀ (x y : R) (n : ℕ) - → fst χ₁ ([ x , f ^ n , ∣ n , refl ∣ ]) ≡ fst χ₂ ([ y , g ^ n , ∣ n , refl ∣ ]) - → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣ ]) - × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣ ])) + → fst χ₁ ([ x , f ^ n , ∣ n , refl ∣₁ ]) ≡ fst χ₂ ([ y , g ^ n , ∣ n , refl ∣₁ ]) + → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣₁ ]) + × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣₁ ])) baseCase x y n χ₁[x/fⁿ]≡χ₂[y/gⁿ] = PT.rec isProp∃! annihilatorHelper exAnnihilator where -- doesn't compute that well but at least it computes... @@ -253,8 +253,8 @@ module _ (R' : CommRing ℓ) (f g : (fst R')) where annihilatorHelper : Σ[ s ∈ Sᶠᵍ ] (fst s · (x · transport refl (g ^ n)) · (1r · transport refl ((f · g) ^ n)) ≡ fst s · (y · transport refl (f ^ n)) · (1r · transport refl ((f · g) ^ n))) - → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣ ]) - × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣ ])) + → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣₁ ]) + × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣₁ ])) annihilatorHelper ((s , s∈[fgⁿ]) , p) = PT.rec isProp∃! exponentHelper s∈[fgⁿ] where sxgⁿ[fg]ⁿ≡syfⁿ[fg]ⁿ : s · x · g ^ n · (f · g) ^ n ≡ s · y · f ^ n · (f · g) ^ n @@ -290,8 +290,8 @@ module _ (R' : CommRing ℓ) (f g : (fst R')) where exponentHelper : Σ[ m ∈ ℕ ] s ≡ (f · g) ^ m - → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣ ]) - × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣ ])) + → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣₁ ]) + × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣₁ ])) exponentHelper (m , s≡[fg]ᵐ) = PT.rec isProp∃! Σhelper (GeneratingPowers.thm R' 2n+m _ fgVec 1∈⟨f,g⟩) where @@ -336,8 +336,8 @@ module _ (R' : CommRing ℓ) (f g : (fst R')) where 2n+m = n +ℕ (n +ℕ m) -- extracting information from the fact that R=⟨f,g⟩ Σhelper : Σ[ α ∈ FinVec R 2 ] 1r ≡ linearCombination R' α (fⁿgⁿVec 2n+m) - → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣ ]) - × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣ ])) + → ∃![ z ∈ R ] ((z /1ᶠ ≡ [ x , f ^ n , ∣ n , refl ∣₁ ]) + × (z /1ᵍ ≡ [ y , g ^ n , ∣ n , refl ∣₁ ])) Σhelper (α , linCombi) = uniqueExists z (z/1≡x/fⁿ , z/1≡y/gⁿ) (λ _ → isProp× (is-set _ _) (is-set _ _)) uniqueness @@ -351,8 +351,8 @@ module _ (R' : CommRing ℓ) (f g : (fst R')) where -- definition of the element z = α₀ · x · f ^ (n +ℕ m) + α₁ · y · g ^ (n +ℕ m) - z/1≡x/fⁿ : (z /1ᶠ) ≡ [ x , f ^ n , ∣ n , refl ∣ ] - z/1≡x/fⁿ = eq/ _ _ ((f ^ (n +ℕ m) , ∣ n +ℕ m , refl ∣) , path) + z/1≡x/fⁿ : (z /1ᶠ) ≡ [ x , f ^ n , ∣ n , refl ∣₁ ] + z/1≡x/fⁿ = eq/ _ _ ((f ^ (n +ℕ m) , ∣ n +ℕ m , refl ∣₁) , path) where useSolver1 : ∀ x y α₀ α₁ fⁿ⁺ᵐ gⁿ⁺ᵐ fⁿ → fⁿ⁺ᵐ · (α₀ · x · fⁿ⁺ᵐ + α₁ · y · gⁿ⁺ᵐ) · fⁿ @@ -403,8 +403,8 @@ module _ (R' : CommRing ℓ) (f g : (fst R')) where f ^ (n +ℕ m) · x · 1r ∎ - z/1≡y/gⁿ : (z /1ᵍ) ≡ [ y , g ^ n , ∣ n , refl ∣ ] - z/1≡y/gⁿ = eq/ _ _ ((g ^ (n +ℕ m) , ∣ n +ℕ m , refl ∣) , path) + z/1≡y/gⁿ : (z /1ᵍ) ≡ [ y , g ^ n , ∣ n , refl ∣₁ ] + z/1≡y/gⁿ = eq/ _ _ ((g ^ (n +ℕ m) , ∣ n +ℕ m , refl ∣₁) , path) where useSolver1 : ∀ x y α₀ α₁ fⁿ⁺ᵐ gⁿ⁺ᵐ gⁿ → gⁿ⁺ᵐ · (α₀ · x · fⁿ⁺ᵐ + α₁ · y · gⁿ⁺ᵐ) · gⁿ @@ -456,8 +456,8 @@ module _ (R' : CommRing ℓ) (f g : (fst R')) where g ^ (n +ℕ m) · y · 1r ∎ - uniqueness : ∀ a → ((a /1ᶠ) ≡ [ x , f ^ n , ∣ n , refl ∣ ]) - × ((a /1ᵍ) ≡ [ y , g ^ n , ∣ n , refl ∣ ]) + uniqueness : ∀ a → ((a /1ᶠ) ≡ [ x , f ^ n , ∣ n , refl ∣₁ ]) + × ((a /1ᵍ) ≡ [ y , g ^ n , ∣ n , refl ∣₁ ]) → z ≡ a uniqueness a (a/1≡x/fⁿ , a/1≡y/gⁿ) = equalByDifference _ _ (injectivityLemma 1∈⟨f,g⟩ (z - a) [z-a]/1≡0overF [z-a]/1≡0overG) @@ -471,9 +471,9 @@ module _ (R' : CommRing ℓ) (f g : (fst R')) where ≡⟨ cong₂ _-_ z/1≡x/fⁿ a/1≡x/fⁿ ⟩ - [ x , f ^ n , ∣ n , refl ∣ ] - [ x , f ^ n , ∣ n , refl ∣ ] + [ x , f ^ n , ∣ n , refl ∣₁ ] - [ x , f ^ n , ∣ n , refl ∣₁ ] - ≡⟨ +Rinv ([ x , f ^ n , ∣ n , refl ∣ ]) ⟩ + ≡⟨ +Rinv ([ x , f ^ n , ∣ n , refl ∣₁ ]) ⟩ 0r ∎ @@ -486,9 +486,9 @@ module _ (R' : CommRing ℓ) (f g : (fst R')) where ≡⟨ cong₂ _-_ z/1≡y/gⁿ a/1≡y/gⁿ ⟩ - [ y , g ^ n , ∣ n , refl ∣ ] - [ y , g ^ n , ∣ n , refl ∣ ] + [ y , g ^ n , ∣ n , refl ∣₁ ] - [ y , g ^ n , ∣ n , refl ∣₁ ] - ≡⟨ +Rinv ([ y , g ^ n , ∣ n , refl ∣ ]) ⟩ + ≡⟨ +Rinv ([ y , g ^ n , ∣ n , refl ∣₁ ]) ⟩ 0r ∎ diff --git a/Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda b/Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda index 0142433771..b9c4e0c513 100644 --- a/Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda +++ b/Cubical/Algebra/CommRing/Localisation/UniversalProperty.agda @@ -413,7 +413,7 @@ module _ (R' : CommRing ℓ) (S' : ℙ (fst R')) (SMultClosedSubset : isMultClos ∙∙ cong (_· s) (sym (0RightAnnihilates _)) Surχ : isSurjection (fst χ) - Surχ a = PT.rec isPropPropTrunc (λ x → PT.∣ [ x .fst ] , x .snd ∣) (surχ a) + Surχ a = PT.rec isPropPropTrunc (λ x → PT.∣ [ x .fst ] , x .snd ∣₁) (surχ a) S⁻¹RChar : (A' : CommRing ℓ) (φ : CommRingHom R' A') diff --git a/Cubical/Algebra/CommRing/Properties.agda b/Cubical/Algebra/CommRing/Properties.agda index 0dbcbcf515..7f0ed26100 100644 --- a/Cubical/Algebra/CommRing/Properties.agda +++ b/Cubical/Algebra/CommRing/Properties.agda @@ -342,7 +342,7 @@ recPT→CommRing : {A : Type ℓ'} (𝓕 : A → CommRing ℓ) → (σ : ∀ x y → CommRingEquiv (𝓕 x) (𝓕 y)) → (∀ x y z → σ x z ≡ compCommRingEquiv (σ x y) (σ y z)) ------------------------------------------------------ - → ∥ A ∥ → CommRing ℓ + → ∥ A ∥₁ → CommRing ℓ recPT→CommRing 𝓕 σ compCoh = GpdElim.rec→Gpd isGroupoidCommRing 𝓕 (3-ConstantCompChar 𝓕 (λ x y → uaCommRing (σ x y)) λ x y z → sym ( cong uaCommRing (compCoh x y z) diff --git a/Cubical/Algebra/CommRing/RadicalIdeal.agda b/Cubical/Algebra/CommRing/RadicalIdeal.agda index 5c5dd5b885..45b433be8d 100644 --- a/Cubical/Algebra/CommRing/RadicalIdeal.agda +++ b/Cubical/Algebra/CommRing/RadicalIdeal.agda @@ -84,7 +84,7 @@ module RadicalIdeal (R' : CommRing ℓ) where ∑Binomial∈I : ∑ (BinomialVec (n +ℕ m) x y) ∈ I ∑Binomial∈I = ∑Closed I (BinomialVec (n +ℕ m) _ _) binomialCoeff∈I contains0 (snd (√ I)) = - ∣ 1 , subst-∈ I (sym (0LeftAnnihilates 1r)) ((I .snd) .contains0) ∣ + ∣ 1 , subst-∈ I (sym (0LeftAnnihilates 1r)) ((I .snd) .contains0) ∣₁ ·Closed (snd (√ I)) r = map λ { (n , xⁿ∈I) → n , subst-∈ I (sym (^-ldist-· r _ n)) ((I .snd) .·Closed (r ^ n) xⁿ∈I) } @@ -94,7 +94,7 @@ module RadicalIdeal (R' : CommRing ℓ) where map (λ { (m , [xⁿ]ᵐ∈I) → (n ·ℕ m) , subst-∈ I (sym (^-rdist-·ℕ x n m)) [xⁿ]ᵐ∈I }) ∈→∈√ : ∀ (I : CommIdeal) (x : R) → x ∈ I → x ∈ √ I - ∈→∈√ I _ x∈I = ∣ 1 , subst-∈ I (sym (·Rid _)) x∈I ∣ + ∈→∈√ I _ x∈I = ∣ 1 , subst-∈ I (sym (·Rid _)) x∈I ∣₁ @@ -103,7 +103,7 @@ module RadicalIdeal (R' : CommRing ℓ) where √FGIdealCharLImpl : {n : ℕ} (V : FinVec R n) (I : CommIdeal) → √ ⟨ V ⟩[ R' ] ⊆ √ I → (∀ i → V i ∈ √ I) √FGIdealCharLImpl V I √⟨V⟩⊆√I i = √⟨V⟩⊆√I _ (∈→∈√ ⟨ V ⟩[ R' ] (V i) - ∣ (λ j → δ i j) , sym (∑Mul1r _ _ i) ∣) + ∣ (λ j → δ i j) , sym (∑Mul1r _ _ i) ∣₁) √FGIdealCharRImpl : {n : ℕ} (V : FinVec R n) (I : CommIdeal) → (∀ i → V i ∈ √ I) → √ ⟨ V ⟩[ R' ] ⊆ √ I @@ -159,14 +159,14 @@ module RadicalIdeal (R' : CommRing ℓ) where Σhelper : Σ[ m ∈ ℕ ] (z ^ m ∈ J) → x ^ n ≡ y + z → x ∈ √ (I +i J) Σhelper (m , z^m∈J) x^n≡y+z = - ∣ n ·ℕ m , ∣ (∑ (yVec m) , z ^ m) , ∑yVec∈I m , z^m∈J , path m x^n≡y+z ∣ ∣ + ∣ n ·ℕ m , ∣ (∑ (yVec m) , z ^ m) , ∑yVec∈I m , z^m∈J , path m x^n≡y+z ∣₁ ∣₁ √+RContrRIncl : (I J : CommIdeal) → √ (I +i J) ⊆ √ (I +i √ J) √+RContrRIncl I J x = PT.elim (λ _ → isPropPropTrunc) (uncurry curriedIncl2) where curriedIncl2 : (n : ℕ) → (x ^ n ∈ (I +i J)) → x ∈ √ ((I +i √ J)) curriedIncl2 n = map λ ((y , z) , y∈I , z∈J , x≡y+z) - → n , ∣ (y , z) , y∈I , ∈→∈√ J z z∈J , x≡y+z ∣ + → n , ∣ (y , z) , y∈I , ∈→∈√ J z z∈J , x≡y+z ∣₁ √+RContr : (I J : CommIdeal) → √ (I +i √ J) ≡ √ (I +i J) √+RContr I J = CommIdeal≡Char (√+RContrLIncl I J) (√+RContrRIncl I J) @@ -194,7 +194,7 @@ module RadicalIdeal (R' : CommRing ℓ) where (∈→∈√ I _ (·RClosed (I .snd) y x∈I)) curriedHelper x y x∈I (suc l) y^l+1∈J = -- (xy)^l+1 ≡ x^l · x (∈I) · y^l+1 (∈J) ∣ suc l , subst-∈ (I ·i J) (sym (^-ldist-· _ _ (suc l))) - (prodInProd I J _ _ (subst-∈ I (·Comm _ _) (I .snd .·Closed (x ^ l) x∈I)) y^l+1∈J) ∣ + (prodInProd I J _ _ (subst-∈ I (·Comm _ _) (I .snd .·Closed (x ^ l) x∈I)) y^l+1∈J) ∣₁ prodHelper : ∀ i → β i ∈ √ J → α i · β i ∈ √ (I ·i J) prodHelper i = PT.elim (λ _ → isPropPropTrunc) (uncurry (curriedHelper (α i) (β i) (α∈I i))) @@ -204,7 +204,7 @@ module RadicalIdeal (R' : CommRing ℓ) where where curriedIncl2 : (n : ℕ) → x ^ n ∈ (I ·i J) → x ∈ √ (I ·i √ J) curriedIncl2 n = map λ (m , (α , β) , α∈I , β∈J , xⁿ≡∑αβ) - → n , ∣ m , (α , β) , α∈I , (λ i → ∈→∈√ J (β i) (β∈J i)) , xⁿ≡∑αβ ∣ + → n , ∣ m , (α , β) , α∈I , (λ i → ∈→∈√ J (β i) (β∈J i)) , xⁿ≡∑αβ ∣₁ √·RContr : (I J : CommIdeal) → √ (I ·i √ J) ≡ √ (I ·i J) √·RContr I J = CommIdeal≡Char (√·RContrLIncl I J) (√·RContrRIncl I J) diff --git a/Cubical/Algebra/CommRingSolver/Examples.agda b/Cubical/Algebra/CommRingSolver/Examples.agda index 8d5668cb84..28bedad24f 100644 --- a/Cubical/Algebra/CommRingSolver/Examples.agda +++ b/Cubical/Algebra/CommRingSolver/Examples.agda @@ -30,10 +30,10 @@ module TestErrors (R : CommRing ℓ) where -} module TestWithℤ where - open CommRingStr (ℤ .snd) + open CommRingStr (ℤCommRing .snd) - _ : (a b : fst ℤ) → a + b ≡ b + a - _ = solve ℤ + _ : (a b : fst ℤCommRing) → a + b ≡ b + a + _ = solve ℤCommRing module Test (R : CommRing ℓ) where open CommRingStr (snd R) diff --git a/Cubical/Algebra/CommRingSolver/Reflection.agda b/Cubical/Algebra/CommRingSolver/Reflection.agda index f5e1fa5067..13e50aa6a5 100644 --- a/Cubical/Algebra/CommRingSolver/Reflection.agda +++ b/Cubical/Algebra/CommRingSolver/Reflection.agda @@ -31,7 +31,6 @@ open import Cubical.Data.Vec using (Vec) renaming ([] to emptyVec; _∷_ to _∷ open import Cubical.Algebra.CommRingSolver.AlgebraExpression open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommRing.Instances.Int using () renaming (ℤ to ℤRing) open import Cubical.Algebra.CommRingSolver.RawAlgebra open import Cubical.Algebra.CommRingSolver.IntAsRawRing open import Cubical.Algebra.CommRingSolver.Solver renaming (solve to ringSolve) diff --git a/Cubical/Algebra/Direct-Sum/Base.agda b/Cubical/Algebra/DirectSum/Base.agda similarity index 99% rename from Cubical/Algebra/Direct-Sum/Base.agda rename to Cubical/Algebra/DirectSum/Base.agda index 11bcb8ca8b..16f7f7c6ab 100644 --- a/Cubical/Algebra/Direct-Sum/Base.agda +++ b/Cubical/Algebra/DirectSum/Base.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.Direct-Sum.Base where +module Cubical.Algebra.DirectSum.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels diff --git a/Cubical/Algebra/Direct-Sum/Properties.agda b/Cubical/Algebra/DirectSum/Properties.agda similarity index 96% rename from Cubical/Algebra/Direct-Sum/Properties.agda rename to Cubical/Algebra/DirectSum/Properties.agda index 68beece653..57b9cbf386 100644 --- a/Cubical/Algebra/Direct-Sum/Properties.agda +++ b/Cubical/Algebra/DirectSum/Properties.agda @@ -1,12 +1,12 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.Direct-Sum.Properties where +module Cubical.Algebra.DirectSum.Properties where open import Cubical.Foundations.Prelude open import Cubical.Algebra.Group open import Cubical.Algebra.AbGroup -open import Cubical.Algebra.Direct-Sum.Base +open import Cubical.Algebra.DirectSum.Base private variable ℓ ℓ' : Level diff --git a/Cubical/Algebra/Group/Exact.agda b/Cubical/Algebra/Group/Exact.agda index 22fb72be65..0111d20215 100644 --- a/Cubical/Algebra/Group/Exact.agda +++ b/Cubical/Algebra/Group/Exact.agda @@ -80,9 +80,9 @@ ImG→H⊂KerH→L (extendExact4Surjective G H L R S G→H H→L L→R R→S sur pRec (GroupStr.is-set (snd R) _ _) (uncurry λ g → J (λ x _ → isInKer L→R x) (ImG→H⊂KerH→L ex (fst H→L (fst G→H g)) - ∣ (fst G→H g) , refl ∣)) + ∣ (fst G→H g) , refl ∣₁)) KerH→L⊂ImG→H (extendExact4Surjective G H L R S G→H H→L L→R R→S surj ex) x ker = - pRec squash + pRec squash₁ (uncurry λ y → J (λ x _ → isInIm (compGroupHom G→H H→L) x) (pMap (uncurry (λ y → J (λ y _ → Σ[ g ∈ fst G ] fst H→L (fst G→H g) ≡ H→L .fst y) diff --git a/Cubical/Algebra/Group/Instances/Bool.agda b/Cubical/Algebra/Group/Instances/Bool.agda index b673971889..6e0408fb31 100644 --- a/Cubical/Algebra/Group/Instances/Bool.agda +++ b/Cubical/Algebra/Group/Instances/Bool.agda @@ -4,48 +4,53 @@ module Cubical.Algebra.Group.Instances.Bool where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Structure -open import Cubical.Data.Bool renaming (Bool to BoolType) -open import Cubical.Data.Empty renaming (rec to ⊥-rec) + +open import Cubical.Relation.Nullary + +open import Cubical.Data.Bool +open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Sum hiding (map ; rec) + open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Monoid open import Cubical.Algebra.Semigroup -open import Cubical.Relation.Nullary open GroupStr open IsGroup open IsMonoid open IsSemigroup renaming (assoc to assoc') -Bool : Group₀ -fst Bool = BoolType -1g (snd Bool) = true -(snd Bool GroupStr.· false) false = true -(snd Bool GroupStr.· false) true = false -(snd Bool GroupStr.· true) y = y -(inv (snd Bool)) x = x -is-set (isSemigroup (isMonoid (isGroup (snd Bool)))) = isSetBool -assoc' (isSemigroup (isMonoid (isGroup (snd Bool)))) false false false = refl -assoc' (isSemigroup (isMonoid (isGroup (snd Bool)))) false false true = refl -assoc' (isSemigroup (isMonoid (isGroup (snd Bool)))) false true false = refl -assoc' (isSemigroup (isMonoid (isGroup (snd Bool)))) false true true = refl -assoc' (isSemigroup (isMonoid (isGroup (snd Bool)))) true false false = refl -assoc' (isSemigroup (isMonoid (isGroup (snd Bool)))) true false true = refl -assoc' (isSemigroup (isMonoid (isGroup (snd Bool)))) true true false = refl -assoc' (isSemigroup (isMonoid (isGroup (snd Bool)))) true true true = refl -identity (IsGroup.isMonoid (isGroup (snd Bool))) false = refl , refl -identity (IsGroup.isMonoid (isGroup (snd Bool))) true = refl , refl -inverse (isGroup (snd Bool)) false = refl , refl -inverse (isGroup (snd Bool)) true = refl , refl + + +BoolGroup : Group₀ +fst BoolGroup = Bool +1g (snd BoolGroup) = true +(snd BoolGroup GroupStr.· false) false = true +(snd BoolGroup GroupStr.· false) true = false +(snd BoolGroup GroupStr.· true) y = y +(inv (snd BoolGroup)) x = x +is-set (isSemigroup (isMonoid (isGroup (snd BoolGroup)))) = isSetBool +assoc' (isSemigroup (isMonoid (isGroup (snd BoolGroup)))) false false false = refl +assoc' (isSemigroup (isMonoid (isGroup (snd BoolGroup)))) false false true = refl +assoc' (isSemigroup (isMonoid (isGroup (snd BoolGroup)))) false true false = refl +assoc' (isSemigroup (isMonoid (isGroup (snd BoolGroup)))) false true true = refl +assoc' (isSemigroup (isMonoid (isGroup (snd BoolGroup)))) true false false = refl +assoc' (isSemigroup (isMonoid (isGroup (snd BoolGroup)))) true false true = refl +assoc' (isSemigroup (isMonoid (isGroup (snd BoolGroup)))) true true false = refl +assoc' (isSemigroup (isMonoid (isGroup (snd BoolGroup)))) true true true = refl +identity (IsGroup.isMonoid (isGroup (snd BoolGroup))) false = refl , refl +identity (IsGroup.isMonoid (isGroup (snd BoolGroup))) true = refl , refl +inverse (isGroup (snd BoolGroup)) false = refl , refl +inverse (isGroup (snd BoolGroup)) true = refl , refl -- Proof that any Group equivalent to Bool as types is also isomorphic to Bool as groups. open GroupStr renaming (assoc to assocG) -module _ {ℓ : Level} {A : Group ℓ} (e : Iso (fst A) BoolType) where +module _ {ℓ : Level} {A : Group ℓ} (e : Iso (fst A) Bool) where private discreteA : Discrete (typ A) discreteA = isoPresDiscrete (invIso e) _≟_ @@ -53,7 +58,7 @@ module _ {ℓ : Level} {A : Group ℓ} (e : Iso (fst A) BoolType) where _·A_ = GroupStr._·_ (snd A) -A_ = GroupStr.inv (snd A) - IsoABool : Iso BoolType (typ A) + IsoABool : Iso Bool (typ A) IsoABool with (Iso.fun e (1g (snd A))) ≟ true ... | yes p = invIso e ... | no p = compIso notIso (invIso e) @@ -68,12 +73,12 @@ module _ {ℓ : Level} {A : Group ℓ} (e : Iso (fst A) BoolType) where decA x with (Iso.inv IsoABool x) ≟ false | discreteA x (1g (snd A)) ... | yes p | yes q = inl q ... | yes p | no q = inr (sym (Iso.rightInv IsoABool x) ∙ cong (Iso.fun (IsoABool)) p) - ... | no p | no q = inr (⊥-rec (q (sym (Iso.rightInv IsoABool x) + ... | no p | no q = inr (⊥.rec (q (sym (Iso.rightInv IsoABool x) ∙∙ cong (Iso.fun IsoABool) (¬false→true _ p) ∙∙ true→1))) ... | no p | yes q = inl q - ≅Bool : GroupIso Bool A + ≅Bool : GroupIso BoolGroup A ≅Bool .fst = IsoABool ≅Bool .snd = makeIsGroupHom homHelp where diff --git a/Cubical/Algebra/Group/Instances/DiffInt.agda b/Cubical/Algebra/Group/Instances/DiffInt.agda index f9fd62d140..e5276df7ae 100644 --- a/Cubical/Algebra/Group/Instances/DiffInt.agda +++ b/Cubical/Algebra/Group/Instances/DiffInt.agda @@ -1,24 +1,28 @@ {-# OPTIONS --safe #-} module Cubical.Algebra.Group.Instances.DiffInt where -open import Cubical.HITs.SetQuotients open import Cubical.Foundations.Prelude -open import Cubical.Data.Int.MoreInts.DiffInt renaming (ℤ to ℤType ; _+_ to _+ℤ_ ; _-_ to _-ℤ_) + +open import Cubical.Data.Int.MoreInts.DiffInt renaming ( _+_ to _+ℤ_ ; _-_ to _-ℤ_) + open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Semigroup.Base open import Cubical.Algebra.Monoid.Base +open import Cubical.HITs.SetQuotients + open GroupStr -ℤ-isGroup : IsGroup {G = ℤType} ([ 0 , 0 ]) (_+ℤ_) (-ℤ_) + +ℤ-isGroup : IsGroup {G = ℤ} ([ 0 , 0 ]) (_+ℤ_) (-ℤ_) IsSemigroup.is-set (IsMonoid.isSemigroup (IsGroup.isMonoid ℤ-isGroup)) = ℤ-isSet IsSemigroup.assoc (IsMonoid.isSemigroup (IsGroup.isMonoid ℤ-isGroup)) = +ℤ-assoc IsMonoid.identity (IsGroup.isMonoid ℤ-isGroup) = λ x → (zero-identityʳ 0 x , zero-identityˡ 0 x) IsGroup.inverse ℤ-isGroup = λ x → (-ℤ-invʳ x , -ℤ-invˡ x) -ℤ : Group₀ -fst ℤ = ℤType -1g (snd ℤ) = [ 0 , 0 ] -_·_ (snd ℤ) = _+ℤ_ -inv (snd ℤ) = -ℤ_ -isGroup (snd ℤ) = ℤ-isGroup +ℤGroup : Group₀ +fst ℤGroup = ℤ +1g (snd ℤGroup) = [ 0 , 0 ] +_·_ (snd ℤGroup) = _+ℤ_ +inv (snd ℤGroup) = -ℤ_ +isGroup (snd ℤGroup) = ℤ-isGroup diff --git a/Cubical/Algebra/Group/Instances/Int.agda b/Cubical/Algebra/Group/Instances/Int.agda index 4d0a8104af..131a3ec875 100644 --- a/Cubical/Algebra/Group/Instances/Int.agda +++ b/Cubical/Algebra/Group/Instances/Int.agda @@ -3,21 +3,24 @@ module Cubical.Algebra.Group.Instances.Int where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism + open import Cubical.Data.Int - renaming (ℤ to ℤType ; _+_ to _+ℤ_ ; _-_ to _-ℤ_; -_ to -ℤ_ ; _·_ to _·ℤ_) + renaming (_+_ to _+ℤ_ ; _-_ to _-ℤ_; -_ to -ℤ_ ; _·_ to _·ℤ_) + open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Properties open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Properties + open GroupStr -ℤ : Group₀ -fst ℤ = ℤType -1g (snd ℤ) = 0 -_·_ (snd ℤ) = _+ℤ_ -inv (snd ℤ) = -ℤ_ -isGroup (snd ℤ) = isGroupℤ +ℤGroup : Group₀ +fst ℤGroup = ℤ +1g (snd ℤGroup) = 0 +_·_ (snd ℤGroup) = _+ℤ_ +inv (snd ℤGroup) = -ℤ_ +isGroup (snd ℤGroup) = isGroupℤ where abstract isGroupℤ : IsGroup (pos 0) (_+ℤ_) (-ℤ_) @@ -25,17 +28,17 @@ isGroup (snd ℤ) = isGroupℤ +Assoc (λ _ → refl) (+Comm 0) -Cancel -Cancel' -ℤHom : (n : ℤType) → GroupHom ℤ ℤ +ℤHom : (n : ℤ) → GroupHom ℤGroup ℤGroup fst (ℤHom n) x = n ·ℤ x snd (ℤHom n) = makeIsGroupHom λ x y → ·DistR+ n x y -negEquivℤ : GroupEquiv ℤ ℤ +negEquivℤ : GroupEquiv ℤGroup ℤGroup fst negEquivℤ = isoToEquiv - (iso (GroupStr.inv (snd ℤ)) - (GroupStr.inv (snd ℤ)) - (GroupTheory.invInv ℤ) - (GroupTheory.invInv ℤ)) + (iso (GroupStr.inv (snd ℤGroup)) + (GroupStr.inv (snd ℤGroup)) + (GroupTheory.invInv ℤGroup) + (GroupTheory.invInv ℤGroup)) snd negEquivℤ = makeIsGroupHom -Dist+ diff --git a/Cubical/Algebra/Group/Instances/IntMod.agda b/Cubical/Algebra/Group/Instances/IntMod.agda index 32b70aaa3c..4d34edb6a1 100644 --- a/Cubical/Algebra/Group/Instances/IntMod.agda +++ b/Cubical/Algebra/Group/Instances/IntMod.agda @@ -4,91 +4,90 @@ module Cubical.Algebra.Group.Instances.IntMod where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.HLevels -open import Cubical.Algebra.Group.Instances.Int -open import Cubical.Algebra.Group.Base -open import Cubical.Algebra.Monoid.Base -open import Cubical.Algebra.Semigroup.Base -open import Cubical.Data.Empty renaming (rec to ⊥-rec) + +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Unit open import Cubical.Data.Bool hiding (isProp≤) -open import Cubical.Data.Fin -open import Cubical.Data.Fin.Arithmetic -open import Cubical.Data.Int renaming (_+_ to _+ℤ_ ; ℤ to ℤType) open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) -open import Cubical.Data.Unit open import Cubical.Data.Nat.Mod open import Cubical.Data.Nat.Order -open import Cubical.Algebra.Group.Instances.Int - renaming (ℤ to ℤGroup) -open import Cubical.Algebra.Group.Instances.Unit -open import Cubical.Algebra.Group.Instances.Bool - renaming (Bool to BoolGroup) +open import Cubical.Data.Int renaming (_+_ to _+ℤ_) +open import Cubical.Data.Fin +open import Cubical.Data.Fin.Arithmetic +open import Cubical.Data.Sigma + +open import Cubical.Algebra.Monoid.Base +open import Cubical.Algebra.Semigroup.Base +open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.Properties -open import Cubical.Foundations.HLevels -open import Cubical.Data.Sigma +open import Cubical.Algebra.Group.Instances.Unit +open import Cubical.Algebra.Group.Instances.Bool +open import Cubical.Algebra.Group.Instances.Int open GroupStr open IsGroup open IsMonoid -ℤ/_ : ℕ → Group₀ -ℤ/ zero = ℤGroup -fst (ℤ/ suc n) = Fin (suc n) -1g (snd (ℤ/ suc n)) = 0 -GroupStr._·_ (snd (ℤ/ suc n)) = _+ₘ_ -inv (snd (ℤ/ suc n)) = -ₘ_ -IsSemigroup.is-set (isSemigroup (isMonoid (isGroup (snd (ℤ/ suc n))))) = + +ℤGroup/_ : ℕ → Group₀ +ℤGroup/ zero = ℤGroup +fst (ℤGroup/ suc n) = Fin (suc n) +1g (snd (ℤGroup/ suc n)) = 0 +GroupStr._·_ (snd (ℤGroup/ suc n)) = _+ₘ_ +inv (snd (ℤGroup/ suc n)) = -ₘ_ +IsSemigroup.is-set (isSemigroup (isMonoid (isGroup (snd (ℤGroup/ suc n))))) = isSetFin -IsSemigroup.assoc (isSemigroup (isMonoid (isGroup (snd (ℤ/ suc n))))) = +IsSemigroup.assoc (isSemigroup (isMonoid (isGroup (snd (ℤGroup/ suc n))))) = λ x y z → sym (+ₘ-assoc x y z) -fst (identity (isMonoid (isGroup (snd (ℤ/ suc n)))) x) = +ₘ-rUnit x -snd (identity (isMonoid (isGroup (snd (ℤ/ suc n)))) x) = +ₘ-lUnit x -fst (inverse (isGroup (snd (ℤ/ suc n))) x) = +ₘ-rCancel x -snd (inverse (isGroup (snd (ℤ/ suc n))) x) = +ₘ-lCancel x - -ℤ/1≅Unit : GroupIso (ℤ/ 1) UnitGroup₀ -ℤ/1≅Unit = contrGroupIsoUnit isContrFin1 - -Bool≅ℤ/2 : GroupIso BoolGroup (ℤ/ 2) -Iso.fun (fst Bool≅ℤ/2) false = 1 -Iso.fun (fst Bool≅ℤ/2) true = 0 -Iso.inv (fst Bool≅ℤ/2) (zero , p) = true -Iso.inv (fst Bool≅ℤ/2) (suc zero , p) = false -Iso.inv (fst Bool≅ℤ/2) (suc (suc x) , p) = - ⊥-rec (¬-<-zero (predℕ-≤-predℕ (predℕ-≤-predℕ p))) -Iso.rightInv (fst Bool≅ℤ/2) (zero , p) = +fst (identity (isMonoid (isGroup (snd (ℤGroup/ suc n)))) x) = +ₘ-rUnit x +snd (identity (isMonoid (isGroup (snd (ℤGroup/ suc n)))) x) = +ₘ-lUnit x +fst (inverse (isGroup (snd (ℤGroup/ suc n))) x) = +ₘ-rCancel x +snd (inverse (isGroup (snd (ℤGroup/ suc n))) x) = +ₘ-lCancel x + +ℤGroup/1≅Unit : GroupIso (ℤGroup/ 1) UnitGroup₀ +ℤGroup/1≅Unit = contrGroupIsoUnit isContrFin1 + +Bool≅ℤGroup/2 : GroupIso BoolGroup (ℤGroup/ 2) +Iso.fun (fst Bool≅ℤGroup/2) false = 1 +Iso.fun (fst Bool≅ℤGroup/2) true = 0 +Iso.inv (fst Bool≅ℤGroup/2) (zero , p) = true +Iso.inv (fst Bool≅ℤGroup/2) (suc zero , p) = false +Iso.inv (fst Bool≅ℤGroup/2) (suc (suc x) , p) = + ⊥.rec (¬-<-zero (predℕ-≤-predℕ (predℕ-≤-predℕ p))) +Iso.rightInv (fst Bool≅ℤGroup/2) (zero , p) = Σ≡Prop (λ _ → isProp≤) refl -Iso.rightInv (fst Bool≅ℤ/2) (suc zero , p) = +Iso.rightInv (fst Bool≅ℤGroup/2) (suc zero , p) = Σ≡Prop (λ _ → isProp≤) refl -Iso.rightInv (fst Bool≅ℤ/2) (suc (suc x) , p) = - ⊥-rec (¬-<-zero (predℕ-≤-predℕ (predℕ-≤-predℕ p))) -Iso.leftInv (fst Bool≅ℤ/2) false = refl -Iso.leftInv (fst Bool≅ℤ/2) true = refl -snd Bool≅ℤ/2 = +Iso.rightInv (fst Bool≅ℤGroup/2) (suc (suc x) , p) = + ⊥.rec (¬-<-zero (predℕ-≤-predℕ (predℕ-≤-predℕ p))) +Iso.leftInv (fst Bool≅ℤGroup/2) false = refl +Iso.leftInv (fst Bool≅ℤGroup/2) true = refl +snd Bool≅ℤGroup/2 = makeIsGroupHom λ { false false → refl ; false true → refl ; true false → refl ; true true → refl} -ℤ/2≅Bool : GroupIso (ℤ/ 2) BoolGroup -ℤ/2≅Bool = invGroupIso Bool≅ℤ/2 +ℤGroup/2≅Bool : GroupIso (ℤGroup/ 2) BoolGroup +ℤGroup/2≅Bool = invGroupIso Bool≅ℤGroup/2 --- Definition of the quotient map homomorphism ℤ → ℤ/ (suc n) +-- Definition of the quotient map homomorphism ℤ → ℤGroup/ (suc n) -- as a group homomorphism. -ℤ→Fin : (n : ℕ) → ℤType → Fin (suc n) +ℤ→Fin : (n : ℕ) → ℤ → Fin (suc n) ℤ→Fin n (pos x) = x mod (suc n) , mod< n x ℤ→Fin n (negsuc x) = -ₘ (suc x mod suc n , mod< n (suc x)) -ℤ→Fin-presinv : (n : ℕ) (x : ℤType) → ℤ→Fin n (- x) ≡ -ₘ ℤ→Fin n x +ℤ→Fin-presinv : (n : ℕ) (x : ℤ) → ℤ→Fin n (- x) ≡ -ₘ ℤ→Fin n x ℤ→Fin-presinv n (pos zero) = Σ≡Prop (λ _ → isProp≤) ((λ _ → zero) ∙ sym (cong fst help)) where help : (-ₘ_ {n = n} 0) ≡ 0 - help = GroupTheory.inv1g (ℤ/ (suc n)) + help = GroupTheory.inv1g (ℤGroup/ (suc n)) ℤ→Fin-presinv n (pos (suc x)) = Σ≡Prop (λ _ → isProp≤) refl ℤ→Fin-presinv n (negsuc x) = - sym (GroupTheory.invInv (ℤ/ (suc n)) _) + sym (GroupTheory.invInv (ℤGroup/ (suc n)) _) -ₘ1-id : (n : ℕ) @@ -108,7 +107,7 @@ suc-ₘ1 : (n y : ℕ) ≡ (y mod suc n , mod< n y) suc-ₘ1 zero y = isContr→isProp - (isOfHLevelRetractFromIso 0 (fst ℤ/1≅Unit) isContrUnit) _ _ + (isOfHLevelRetractFromIso 0 (fst ℤGroup/1≅Unit) isContrUnit) _ _ suc-ₘ1 (suc n) y = (λ i → ((suc y mod suc (suc n)) , mod< (suc n) (suc y)) +ₘ (-ₘ1-id (suc n) i)) @@ -140,15 +139,15 @@ suc-ₘ1 (suc n) y = +ₘ (-ₘ (((suc y mod suc n) , mod< n (suc y))))) ≡ -ₘ ((y mod suc n) , mod< n y) 1-ₘsuc n y = - sym (GroupTheory.invInv (ℤ/ (suc n)) _) + sym (GroupTheory.invInv (ℤGroup/ (suc n)) _) ∙ cong -ₘ_ - (GroupTheory.invDistr (ℤ/ (suc n)) + (GroupTheory.invDistr (ℤGroup/ (suc n)) (modInd n 1 , mod< n 1) (-ₘ (modInd n (suc y) , mod< n (suc y))) ∙ cong (_-ₘ (modInd n 1 , mod< n 1)) - (GroupTheory.invInv (ℤ/ (suc n)) (modInd n (suc y) , mod< n (suc y))) + (GroupTheory.invInv (ℤGroup/ (suc n)) (modInd n (suc y) , mod< n (suc y))) ∙ suc-ₘ1 n y) -isHomℤ→Fin : (n : ℕ) → IsGroupHom (snd ℤGroup) (ℤ→Fin n) (snd (ℤ/ (suc n))) +isHomℤ→Fin : (n : ℕ) → IsGroupHom (snd ℤGroup) (ℤ→Fin n) (snd (ℤGroup/ (suc n))) isHomℤ→Fin n = makeIsGroupHom λ { (pos x) y → pos+case x y @@ -160,19 +159,19 @@ isHomℤ→Fin n = sym (cong (ℤ→Fin n) (-Dist+ (pos (suc x)) (pos (suc y)))) ∙∙ ℤ→Fin-presinv n (pos (suc x) +ℤ (pos (suc y))) ∙∙ cong -ₘ_ (pos+case (suc x) (pos (suc y))) - ∙∙ GroupTheory.invDistr (ℤ/ (suc n)) + ∙∙ GroupTheory.invDistr (ℤGroup/ (suc n)) (modInd n (suc x) , mod< n (suc x)) (modInd n (suc y) , mod< n (suc y)) ∙∙ +ₘ-comm (ℤ→Fin n (negsuc y)) (ℤ→Fin n (negsuc x))} where - +1case : (y : ℤType) → ℤ→Fin n (1 +ℤ y) ≡ ℤ→Fin n 1 +ₘ ℤ→Fin n y - +1case (pos zero) = sym (GroupStr.rid (snd (ℤ/ (suc n))) _) + +1case : (y : ℤ) → ℤ→Fin n (1 +ℤ y) ≡ ℤ→Fin n 1 +ₘ ℤ→Fin n y + +1case (pos zero) = sym (GroupStr.rid (snd (ℤGroup/ (suc n))) _) +1case (pos (suc y)) = cong (ℤ→Fin n) (+Comm 1 (pos (suc y))) ∙ Σ≡Prop (λ _ → isProp≤) (mod+mod≡mod (suc n) 1 (suc y)) +1case (negsuc zero) = Σ≡Prop (λ _ → isProp≤) refl - ∙ sym (GroupStr.invr (snd (ℤ/ (suc n))) (modInd n 1 , mod< n 1)) + ∙ sym (GroupStr.invr (snd (ℤGroup/ (suc n))) (modInd n 1 , mod< n 1)) +1case (negsuc (suc y)) = Σ≡Prop (λ _ → isProp≤) (cong fst (cong (ℤ→Fin n) (+Comm 1 (negsuc (suc y)))) @@ -181,11 +180,11 @@ isHomℤ→Fin n = ∙ λ i → fst ((1 mod (suc n) , mod< n 1) +ₘ (-ₘ (((suc (suc y) mod suc n) , mod< n (suc (suc y))))))) - pos+case : (x : ℕ) (y : ℤType) + pos+case : (x : ℕ) (y : ℤ) → ℤ→Fin n (pos x +ℤ y) ≡ ℤ→Fin n (pos x) +ₘ ℤ→Fin n y pos+case zero y = cong (ℤ→Fin n) (+Comm 0 y) - ∙ sym (GroupStr.lid (snd (ℤ/ (suc n))) (ℤ→Fin n y)) + ∙ sym (GroupStr.lid (snd (ℤGroup/ (suc n))) (ℤ→Fin n y)) pos+case (suc zero) y = +1case y pos+case (suc (suc x)) y = cong (ℤ→Fin n) (cong (_+ℤ y) (+Comm (pos (suc x)) 1) diff --git a/Cubical/Algebra/Group/Instances/Unit.agda b/Cubical/Algebra/Group/Instances/Unit.agda index 5add299a72..e0a237cf5a 100644 --- a/Cubical/Algebra/Group/Instances/Unit.agda +++ b/Cubical/Algebra/Group/Instances/Unit.agda @@ -6,7 +6,9 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Structure open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv -open import Cubical.Data.Unit renaming (Unit to UnitType; Unit* to UnitType*) + +open import Cubical.Data.Unit + open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.DirProd open import Cubical.Algebra.Group.Morphisms @@ -21,7 +23,7 @@ private ℓ : Level UnitGroup₀ : Group₀ -fst UnitGroup₀ = UnitType +fst UnitGroup₀ = Unit 1g (snd UnitGroup₀) = tt _·_ (snd UnitGroup₀) = λ _ _ → tt inv (snd UnitGroup₀) = λ _ → tt @@ -30,7 +32,7 @@ isGroup (snd UnitGroup₀) = (λ _ → refl) (λ _ → refl) UnitGroup : Group ℓ -fst UnitGroup = UnitType* +fst UnitGroup = Unit* 1g (snd UnitGroup) = tt* _·_ (snd UnitGroup) = λ _ _ → tt* inv (snd UnitGroup) = λ _ → tt* diff --git a/Cubical/Algebra/Group/IsomorphismTheorems.agda b/Cubical/Algebra/Group/IsomorphismTheorems.agda index 02c342db0b..b876a0bcf6 100644 --- a/Cubical/Algebra/Group/IsomorphismTheorems.agda +++ b/Cubical/Algebra/Group/IsomorphismTheorems.agda @@ -72,8 +72,8 @@ module _ {G H : Group ℓ} (ϕ : GroupHom G H) where H.1g ∎ f2 : ⟨ G / kerNormalSubgroup ⟩ → ⟨ imϕ ⟩ - f2 = recS imG.is-set (λ y → ϕ .fst y , ∣ y , refl ∣) - (λ x y r → Σ≡Prop (λ _ → squash) + f2 = recS imG.is-set (λ y → ϕ .fst y , ∣ y , refl ∣₁) + (λ x y r → Σ≡Prop (λ _ → squash₁) (rem x y r)) where rem : (x y : ⟨ G ⟩) → ϕ .fst (x G.· G.inv y) ≡ H.1g → ϕ .fst x ≡ ϕ .fst y @@ -93,7 +93,7 @@ module _ {G H : Group ℓ} (ϕ : GroupHom G H) where f21 : (x : ⟨ imϕ ⟩) → f2 (f1 x) ≡ x f21 (x , hx) = elim {P = λ hx → f2 (f1 (x , hx)) ≡ (x , hx)} (λ _ → imG.is-set _ _) - (λ {(x , hx) → Σ≡Prop (λ _ → squash) hx}) + (λ {(x , hx) → Σ≡Prop (λ _ → squash₁) hx}) hx f1-isHom : (x y : ⟨ imϕ ⟩) → f1 (x imG.· y) ≡ f1 x kerG.· f1 y @@ -120,8 +120,8 @@ module _ {G H : Group ℓ} (ϕ : GroupHom G H) where (uncurry (λ h → elim (λ _ → isPropΠ (λ _ → imG.is-set _ _)) (uncurry λ g y - → λ inker → Σ≡Prop (λ _ → squash) inker))) - λ b → map (λ x → (b , ∣ x ∣) , refl) (surj b)) + → λ inker → Σ≡Prop (λ _ → squash₁) inker))) + λ b → map (λ x → (b , ∣ x ∣₁) , refl) (surj b)) where indHom : GroupHom imϕ H fst indHom = fst diff --git a/Cubical/Algebra/Group/MorphismProperties.agda b/Cubical/Algebra/Group/MorphismProperties.agda index c3897279a2..00de48a8f2 100644 --- a/Cubical/Algebra/Group/MorphismProperties.agda +++ b/Cubical/Algebra/Group/MorphismProperties.agda @@ -109,7 +109,7 @@ isSetGroupHom {G = G} {H = H} = isSetΣ (isSetΠ λ _ → is-set (snd H)) λ _ → isProp→isSet (isPropIsGroupHom G H) isPropIsInIm : (f : GroupHom G H) (x : ⟨ H ⟩) → isProp (isInIm f x) -isPropIsInIm f x = squash +isPropIsInIm f x = squash₁ isSetIm : (f : GroupHom G H) → isSet (Im f) isSetIm {H = H} f = isSetΣ (is-set (snd H)) λ x → isProp→isSet (isPropIsInIm f x) @@ -195,7 +195,7 @@ compSurjective : ∀ {ℓ ℓ' ℓ''} {G : Group ℓ} {H : Group ℓ'} {L : Grou → isSurjective G→H → isSurjective H→L → isSurjective (compGroupHom G→H H→L) compSurjective G→H H→L surj1 surj2 l = - rec squash + rec squash₁ (λ {(h , p) → pMap (λ {(g , q) → g , (cong (fst H→L) q ∙ p)}) (surj1 h)}) @@ -284,7 +284,7 @@ BijectionIso→GroupIso {G = G} {H = H} i = grIso inv (fst grIso) b = rec (helper b) (λ a → a) (surj i b) .fst rightInv (fst grIso) b = rec (helper b) (λ a → a) (surj i b) .snd leftInv (fst grIso) b j = rec (helper (f b)) (λ a → a) - (isPropPropTrunc (surj i (f b)) ∣ b , refl ∣ j) .fst + (isPropPropTrunc (surj i (f b)) ∣ b , refl ∣₁ j) .fst snd grIso = snd (fun i) BijectionIsoToGroupEquiv : BijectionIso G H → GroupEquiv G H diff --git a/Cubical/Algebra/Group/Subgroup.agda b/Cubical/Algebra/Group/Subgroup.agda index da17604f26..6017266fe0 100644 --- a/Cubical/Algebra/Group/Subgroup.agda +++ b/Cubical/Algebra/Group/Subgroup.agda @@ -162,7 +162,7 @@ module _ {G H : Group ℓ} (ϕ : GroupHom G H) where imSubset x = isInIm ϕ x , isPropIsInIm ϕ x isSubgroupIm : isSubgroup H imSubset - id-closed isSubgroupIm = ∣ G.1g , ϕ.pres1 ∣ + id-closed isSubgroupIm = ∣ G.1g , ϕ.pres1 ∣₁ op-closed isSubgroupIm = map2 λ { (x , hx) (y , hy) → x G.· y , ϕ.pres· x y ∙ λ i → hx i H.· hy i } inv-closed isSubgroupIm = map λ { (x , hx) → G.inv x , ϕ.presinv x ∙ cong H.inv hx } diff --git a/Cubical/Algebra/Group/ZAction.agda b/Cubical/Algebra/Group/ZAction.agda index 491c834caa..81f351a35d 100644 --- a/Cubical/Algebra/Group/ZAction.agda +++ b/Cubical/Algebra/Group/ZAction.agda @@ -31,7 +31,7 @@ open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.Properties open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Int open import Cubical.Algebra.Group.Instances.Unit open import Cubical.Algebra.Group.Instances.IntMod open import Cubical.Algebra.Group.DirProd @@ -492,7 +492,7 @@ module _ (f : GroupHom ℤGroup ℤGroup) where ∙ GroupTheory.invInv ℤGroup y} ℤHom→ℤ/im≅ℤ/im1 : (n : ℕ) → fst f 1 ≡ pos (suc n) - → BijectionIso (ℤGroup / imℤHomSubGroup f) (ℤ/ (suc n)) + → BijectionIso (ℤGroup / imℤHomSubGroup f) (ℤGroup/ (suc n)) fst (BijectionIso.fun (ℤHom→ℤ/im≅ℤ/im1 n p)) = sRec isSetFin (ℤ→Fin n) λ a b → rec (isSetFin _ _) (uncurry λ x q @@ -501,7 +501,7 @@ module _ (f : GroupHom ℤGroup ℤGroup) where ∙ pres· (isHomℤ→Fin n) (pos (suc n) * x) b ∙ cong (_+ₘ ℤ→Fin n b) (lem x) - ∙ GroupStr.lid (snd (ℤ/ (suc n))) (ℤ→Fin n b)) + ∙ GroupStr.lid (snd (ℤGroup/ (suc n))) (ℤ→Fin n b)) where lem : (x : ℤ) → ℤ→Fin n (pos (suc n) * x) ≡ 0 lem (pos x) = cong (ℤ→Fin n) (sym (pos· (suc n) x)) @@ -514,7 +514,7 @@ module _ (f : GroupHom ℤGroup ℤGroup) where ∙∙ cong -ₘ_ (Σ≡Prop (λ _ → isProp≤) (cong (_mod (suc n)) (·-comm (suc n) (suc x)) ∙ zero-charac-gen (suc n) (suc x))) - ∙∙ GroupTheory.inv1g (ℤ/ (suc n)) + ∙∙ GroupTheory.inv1g (ℤGroup/ (suc n)) cancel-lem : (a b x : ℤ) → a +ℤ (- b) ≡ x → a ≡ x +ℤ b cancel-lem a b x p = sym (minusPlus b a) ∙ cong (_+ℤ b) p @@ -528,7 +528,7 @@ module _ (f : GroupHom ℤGroup ℤGroup) where (funExt⁻ (cong fst (characGroupHomℤ f ∙ cong ℤHom p)) ((pos (quotient x / (suc n)))) ∙ sym (pos· (suc n) (quotient x / (suc n))) ∙ cong pos ((λ i → q (~ i) .fst +ℕ suc n ·ℕ (quotient x / suc n))) - ∙ cong pos (≡remainder+quotient (suc n) x))) ∣ ; + ∙ cong pos (≡remainder+quotient (suc n) x))) ∣₁ ; (negsuc x) q → eq/ (negsuc x) 0 ∣ (((- pos (quotient suc x / (suc n)))) , presinv (snd f) (pos (quotient suc x / (suc n))) @@ -536,27 +536,27 @@ module _ (f : GroupHom ℤGroup ℤGroup) where (pos (quotient (suc x) / (suc n)))) ∙∙ cong -_ (sym (pos· (suc n) (quotient suc x / (suc n))) ∙ (λ i → pos (fst ((sym (GroupTheory.invInv - (ℤ/ (suc n)) + (ℤGroup/ (suc n)) ((suc x mod suc n) , mod< n (suc x))) ∙ cong -ₘ_ q - ∙ GroupTheory.inv1g (ℤ/ (suc n))) (~ i)) + ∙ GroupTheory.inv1g (ℤGroup/ (suc n))) (~ i)) +ℕ suc n ·ℕ quotient (suc x) / suc n))) - ∙∙ cong -_ (cong pos (≡remainder+quotient (suc n) (suc x))))) ∣}) + ∙∙ cong -_ (cong pos (≡remainder+quotient (suc n) (suc x))))) ∣₁}) BijectionIso.surj (ℤHom→ℤ/im≅ℤ/im1 n p) x = ∣ [ pos (fst x) ] - , (Σ≡Prop (λ _ → isProp≤) (modIndBase n (fst x) (snd x))) ∣ + , (Σ≡Prop (λ _ → isProp≤) (modIndBase n (fst x) (snd x))) ∣₁ -- main result ℤ/imIso : (f : GroupHom ℤGroup ℤGroup) - → GroupIso (ℤGroup / imℤHomSubGroup f) (ℤ/ abs (fst f 1)) + → GroupIso (ℤGroup / imℤHomSubGroup f) (ℤGroup/ abs (fst f 1)) ℤ/imIso f = helpIso _ refl where helpIso : (a : ℤ) - → fst f 1 ≡ a → GroupIso (ℤGroup / imℤHomSubGroup f) (ℤ/ abs a) + → fst f 1 ≡ a → GroupIso (ℤGroup / imℤHomSubGroup f) (ℤGroup/ abs a) helpIso (pos zero) p = invGroupIso (trivHom→ℤ≅ℤ/im f p) helpIso (pos (suc n)) p = BijectionIso→GroupIso (ℤHom→ℤ/im≅ℤ/im1 f n p) helpIso (negsuc n) p = - subst ((λ x → GroupIso (ℤGroup / x) (ℤ/ abs (negsuc n)))) + subst ((λ x → GroupIso (ℤGroup / x) (ℤGroup/ abs (negsuc n)))) (sym lem1) (BijectionIso→GroupIso (ℤHom→ℤ/im≅ℤ/im1 extendHom n (cong -_ p))) @@ -573,8 +573,8 @@ module _ (f : GroupHom ℤGroup ℤGroup) where ∙ GroupTheory.invInv ℤGroup (fst f x) ∙ q })) (Prop.map (λ { (x , q) → (- x) , (presinv (snd f) x ∙ q) })) - ((λ _ → squash _ _)) - (λ _ → squash _ _))))) + ((λ _ → squash₁ _ _)) + (λ _ → squash₁ _ _))))) -- Goal: given G -ᶠ→ H → L → Unit exact, with G ≅ H ≅ ℤ, we get -- an iso ℤ/abs (f 1) ≅ H, where f 1 and 1 are viewed as integers @@ -600,7 +600,7 @@ module _ (f : GroupHom ℤGroup ℤGroup) (G : Group₀) ℤ/im≅ℤ/ker = GroupEquiv→GroupIso (invEq (GroupPath _ _) (cong (ℤGroup /_) imf≡kerg)) - GroupIsoℤ/abs : GroupIso (ℤ/ abs (fst f (pos 1))) G + GroupIsoℤ/abs : GroupIso (ℤGroup/ abs (fst f (pos 1))) G GroupIsoℤ/abs = compGroupIso (invGroupIso (ℤ/imIso f)) @@ -616,29 +616,29 @@ GroupEquivℤ/abs-gen : (G H L : Group₀) → (r : GroupEquiv ℤGroup H) → (f : GroupHom G H) (g : GroupHom H L) → Exact4 G H L UnitGroup₀ f g (→UnitHom L) - → GroupEquiv (ℤ/ abs (invEq (fst r) (fst f (fst (fst e) 1)))) L + → GroupEquiv (ℤGroup/ abs (invEq (fst r) (fst f (fst (fst e) 1)))) L GroupEquivℤ/abs-gen G H L = GroupEquivJ (λ G e → (r : GroupEquiv ℤGroup H) → (f : GroupHom G H) (g : GroupHom H L) → Exact4 G H L UnitGroup₀ f g (→UnitHom L) - → GroupEquiv (ℤ/ abs (invEq (fst r) (fst f (fst (fst e) 1)))) L) + → GroupEquiv (ℤGroup/ abs (invEq (fst r) (fst f (fst (fst e) 1)))) L) (GroupEquivJ (λ H r → (f : GroupHom ℤGroup H) (g : GroupHom H L) → Exact4 ℤGroup H L UnitGroup₀ f g (→UnitHom L) → GroupEquiv - (ℤ/ abs (invEq (fst r) (fst f 1))) L) + (ℤGroup/ abs (invEq (fst r) (fst f 1))) L) λ f g ex → GroupIso→GroupEquiv (GroupIsoℤ/abs f L g ex)) -- for type checking reasons, let's also do it with an abstract type abstract - abstractℤ/_ : ℕ → Group₀ - abstractℤ/_ n = ℤ/ n + abstractℤGroup/_ : ℕ → Group₀ + abstractℤGroup/_ n = ℤGroup/ n - abstractℤ/≡ℤ : abstractℤ/_ ≡ ℤ/_ + abstractℤ/≡ℤ : abstractℤGroup/_ ≡ ℤGroup/_ abstractℤ/≡ℤ = refl - abstractℤ/≅ℤ : (n : ℕ) → GroupEquiv (abstractℤ/ n) (ℤ/ n) + abstractℤ/≅ℤ : (n : ℕ) → GroupEquiv (abstractℤGroup/_ n) (ℤGroup/ n) abstractℤ/≅ℤ n = idGroupEquiv GroupEquiv-abstractℤ/abs-gen : (G H L : Group₀) @@ -648,11 +648,11 @@ GroupEquiv-abstractℤ/abs-gen : (G H L : Group₀) → Exact4 G H L UnitGroup₀ f g (→UnitHom L) → (n : ℕ) → abs (invEq (fst r) (fst f (fst (fst e) 1))) ≡ n - → GroupEquiv (abstractℤ/ n) L + → GroupEquiv (abstractℤGroup/_ n) L GroupEquiv-abstractℤ/abs-gen G H L e r f g ex n p = main where abstract - main : GroupEquiv (abstractℤ/ n) L + main : GroupEquiv (abstractℤGroup/_ n) L main = transport (λ i → GroupEquiv (abstractℤ/≡ℤ (~ i) (p i)) L) diff --git a/Cubical/Algebra/IntegerMatrix/Base.agda b/Cubical/Algebra/IntegerMatrix/Base.agda index 77d1441c37..def0752117 100644 --- a/Cubical/Algebra/IntegerMatrix/Base.agda +++ b/Cubical/Algebra/IntegerMatrix/Base.agda @@ -10,20 +10,12 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function hiding (const) open import Cubical.Foundations.Powerset -open import Cubical.Data.Nat - hiding (_·_) - renaming (_+_ to _+ℕ_ ; +-assoc to +Assocℕ) +open import Cubical.Data.Nat hiding (_·_) renaming (_+_ to _+ℕ_ ; +-assoc to +Assocℕ) open import Cubical.Data.Nat.Order -open import Cubical.Data.Nat.Divisibility - using (m∣n→m≤n) - renaming (∣-trans to ∣ℕ-trans) -open import Cubical.Data.Int - hiding (_+_ ; _·_ ; _-_ ; -_ ; addEq) +open import Cubical.Data.Nat.Divisibility using (m∣n→m≤n) renaming (∣-trans to ∣ℕ-trans) +open import Cubical.Data.Int hiding (_+_ ; _·_ ; _-_ ; -_ ; addEq) open import Cubical.Data.Int.Divisibility open import Cubical.Data.FinData - -open import Cubical.Data.Empty as Empty -open import Cubical.Data.Unit as Unit open import Cubical.Data.Sum open import Cubical.Data.Sigma @@ -32,7 +24,6 @@ open import Cubical.Algebra.Matrix.CommRingCoefficient open import Cubical.Algebra.Ring.BigOps open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Instances.Int - renaming (ℤ to ℤRing) open import Cubical.Relation.Nullary open import Cubical.Induction.WellFounded @@ -41,9 +32,10 @@ private variable m n k : ℕ -open CommRingStr (ℤRing .snd) -open Coefficient ℤRing -open Sum (CommRing→Ring ℤRing) +open CommRingStr (ℤCommRing .snd) +open Coefficient ℤCommRing +open Sum (CommRing→Ring ℤCommRing) + -- Using well-foundedness to do induction diff --git a/Cubical/Algebra/IntegerMatrix/Diagonalization.agda b/Cubical/Algebra/IntegerMatrix/Diagonalization.agda index 2ebb999f04..5cb4dd4bba 100644 --- a/Cubical/Algebra/IntegerMatrix/Diagonalization.agda +++ b/Cubical/Algebra/IntegerMatrix/Diagonalization.agda @@ -40,7 +40,6 @@ open import Cubical.Algebra.IntegerMatrix.Elementaries open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Instances.Int - renaming (ℤ to ℤRing) open import Cubical.Relation.Nullary open import Cubical.Induction.WellFounded @@ -49,12 +48,12 @@ private variable m n k : ℕ -open CommRingStr (ℤRing .snd) +open CommRingStr (ℤCommRing .snd) -open Coefficient ℤRing +open Coefficient ℤCommRing open Sim -open ElemTransformation ℤRing +open ElemTransformation ℤCommRing open ElemTransformationℤ open SwapPivot open RowsImproved diff --git a/Cubical/Algebra/IntegerMatrix/Elementaries.agda b/Cubical/Algebra/IntegerMatrix/Elementaries.agda index b2f8b95c69..9c8bbadcfb 100644 --- a/Cubical/Algebra/IntegerMatrix/Elementaries.agda +++ b/Cubical/Algebra/IntegerMatrix/Elementaries.agda @@ -17,7 +17,7 @@ open import Cubical.Relation.Nullary open import Cubical.Algebra.CommRingSolver.Reflection open import Cubical.Algebra.Ring.BigOps open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤ to Ringℤ) +open import Cubical.Algebra.CommRing.Instances.Int open import Cubical.Algebra.Matrix open import Cubical.Algebra.Matrix.CommRingCoefficient @@ -31,7 +31,7 @@ private -- It seems there are bugs when applying ring solver to integers. -- The following is a work-around. private - module Helper {ℓ : Level}(𝓡 : CommRing ℓ) where + module Helper {ℓ : Level} (𝓡 : CommRing ℓ) where open CommRingStr (𝓡 .snd) helper1 : (a b x y g : 𝓡 .fst) → (a · x - b · - y) · g ≡ a · (x · g) + b · (y · g) @@ -40,7 +40,7 @@ private helper2 : (a b : 𝓡 .fst) → a ≡ 1r · a + 0r · b helper2 = solve 𝓡 -open Helper Ringℤ +open Helper ℤCommRing module ElemTransformationℤ where @@ -50,13 +50,13 @@ module ElemTransformationℤ where open import Cubical.Data.Int.Divisibility private - ℤ = Ringℤ .fst + ℤ = ℤCommRing .fst - open CommRingStr (Ringℤ .snd) - open Sum (CommRing→Ring Ringℤ) + open CommRingStr (ℤCommRing .snd) + open Sum (CommRing→Ring ℤCommRing) - open Coefficient Ringℤ - open LinearTransformation Ringℤ + open Coefficient ℤCommRing + open LinearTransformation ℤCommRing open Bézout open SimRel @@ -79,7 +79,7 @@ module ElemTransformationℤ where module _ (p : ¬ x ≡ 0) where - open Units Ringℤ + open Units ℤCommRing private detEq : det2×2 bézout2Mat · b .gcd ≡ b .gcd diff --git a/Cubical/Algebra/IntegerMatrix/Smith/NormalForm.agda b/Cubical/Algebra/IntegerMatrix/Smith/NormalForm.agda index 0a42763c94..ee2722a9e4 100644 --- a/Cubical/Algebra/IntegerMatrix/Smith/NormalForm.agda +++ b/Cubical/Algebra/IntegerMatrix/Smith/NormalForm.agda @@ -27,13 +27,12 @@ open import Cubical.Relation.Nullary open import Cubical.Algebra.Matrix.CommRingCoefficient open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Instances.Int - renaming (ℤ to ℤRing) private variable m n k : ℕ -open Coefficient ℤRing +open Coefficient ℤCommRing -- Sequence of consecutively divisible integers diff --git a/Cubical/Algebra/IntegerMatrix/Smith/Normalization.agda b/Cubical/Algebra/IntegerMatrix/Smith/Normalization.agda index a3a2a33d1b..ee872632a1 100644 --- a/Cubical/Algebra/IntegerMatrix/Smith/Normalization.agda +++ b/Cubical/Algebra/IntegerMatrix/Smith/Normalization.agda @@ -32,7 +32,6 @@ open import Cubical.Algebra.IntegerMatrix.Smith.NormalForm open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Instances.Int - renaming (ℤ to ℤRing) open import Cubical.Relation.Nullary open import Cubical.Induction.WellFounded @@ -41,13 +40,13 @@ private variable m n k : ℕ -open CommRingStr (ℤRing .snd) -open Coefficient ℤRing +open CommRingStr (ℤCommRing .snd) +open Coefficient ℤCommRing open Sim -- The elementary transformations needed -open ElemTransformation ℤRing +open ElemTransformation ℤCommRing open ElemTransformationℤ open SwapFirstRow open SwapPivot diff --git a/Cubical/Algebra/Polynomials/Multivariate/Base.agda b/Cubical/Algebra/Polynomials/Multivariate/Base.agda index d94f9f5239..b22849d2fb 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/Base.agda +++ b/Cubical/Algebra/Polynomials/Multivariate/Base.agda @@ -25,14 +25,14 @@ module _ (A' : CommRing ℓ) where -- elements 0P : Poly n base : (v : Vec ℕ n) → (a : A) → Poly n - _Poly+_ : (P : Poly n) → (Q : Poly n) → Poly n + _poly+_ : (P : Poly n) → (Q : Poly n) → Poly n -- AbGroup eq - Poly+-assoc : (P Q R : Poly n) → P Poly+ (Q Poly+ R) ≡ (P Poly+ Q) Poly+ R - Poly+-Rid : (P : Poly n) → P Poly+ 0P ≡ P - Poly+-comm : (P Q : Poly n) → P Poly+ Q ≡ Q Poly+ P + poly+Assoc : (P Q R : Poly n) → P poly+ (Q poly+ R) ≡ (P poly+ Q) poly+ R + poly+IdR : (P : Poly n) → P poly+ 0P ≡ P + poly+Comm : (P Q : Poly n) → P poly+ Q ≡ Q poly+ P -- Base eq base-0P : (v : Vec ℕ n) → base v 0r ≡ 0P - base-Poly+ : (v : Vec ℕ n) → (a b : A) → (base v a) Poly+ (base v b) ≡ base v (a + b) + base-poly+ : (v : Vec ℕ n) → (a b : A) → (base v a) poly+ (base v b) ≡ base v (a + b) -- Set Trunc trunc : isSet(Poly n) @@ -52,29 +52,29 @@ module _ (A' : CommRing ℓ) where -- elements (0P* : F 0P) (base* : (v : Vec ℕ n) → (a : A) → F (base v a)) - (_Poly+*_ : {P Q : Poly A' n} → (PS : F P) → (QS : F Q) → F (P Poly+ Q)) + (_poly+*_ : {P Q : Poly A' n} → (PS : F P) → (QS : F Q) → F (P poly+ Q)) -- AbGroup eq - (Poly+-assoc* : {P Q R : Poly A' n} → (PS : F P) → (QS : F Q) → (RS : F R) - → PathP (λ i → F (Poly+-assoc P Q R i)) (PS Poly+* (QS Poly+* RS)) ((PS Poly+* QS) Poly+* RS)) - (Poly+-Rid* : {P : Poly A' n} → (PS : F P) → - PathP (λ i → F (Poly+-Rid P i)) (PS Poly+* 0P*) PS) - (Poly+-comm* : {P Q : Poly A' n} → (PS : F P) → (QS : F Q) - → PathP (λ i → F (Poly+-comm P Q i)) (PS Poly+* QS) (QS Poly+* PS)) + (poly+Assoc* : {P Q R : Poly A' n} → (PS : F P) → (QS : F Q) → (RS : F R) + → PathP (λ i → F (poly+Assoc P Q R i)) (PS poly+* (QS poly+* RS)) ((PS poly+* QS) poly+* RS)) + (poly+IdR* : {P : Poly A' n} → (PS : F P) → + PathP (λ i → F (poly+IdR P i)) (PS poly+* 0P*) PS) + (poly+Comm* : {P Q : Poly A' n} → (PS : F P) → (QS : F Q) + → PathP (λ i → F (poly+Comm P Q i)) (PS poly+* QS) (QS poly+* PS)) -- Base eq (base-0P* : (v : Vec ℕ n) → PathP (λ i → F (base-0P v i)) (base* v 0r) 0P*) - (base-Poly+* : (v : Vec ℕ n) → (a b : A) - → PathP (λ i → F (base-Poly+ v a b i)) ((base* v a) Poly+* (base* v b)) (base* v (a + b))) + (base-poly+* : (v : Vec ℕ n) → (a b : A) + → PathP (λ i → F (base-poly+ v a b i)) ((base* v a) poly+* (base* v b)) (base* v (a + b))) where f : (P : Poly A' n) → F P f 0P = 0P* f (base v a) = base* v a - f (P Poly+ Q) = (f P) Poly+* (f Q) - f (Poly+-assoc P Q R i) = Poly+-assoc* (f P) (f Q) (f R) i - f (Poly+-Rid P i) = Poly+-Rid* (f P) i - f (Poly+-comm P Q i) = Poly+-comm* (f P) (f Q) i + f (P poly+ Q) = (f P) poly+* (f Q) + f (poly+Assoc P Q R i) = poly+Assoc* (f P) (f Q) (f R) i + f (poly+IdR P i) = poly+IdR* (f P) i + f (poly+Comm P Q i) = poly+Comm* (f P) (f Q) i f (base-0P v i) = base-0P* v i - f (base-Poly+ v a b i) = base-Poly+* v a b i + f (base-poly+ v a b i) = base-poly+* v a b i f (trunc P Q p q i j) = isOfHLevel→isOfHLevelDep 2 issd (f P) (f Q) (cong f p) (cong f q) (trunc P Q p q) i j @@ -86,18 +86,18 @@ module _ (A' : CommRing ℓ) where -- elements (0P* : B) (base* : (v : Vec ℕ n) → (a : A) → B) - (_Poly+*_ : B → B → B) + (_poly+*_ : B → B → B) -- AbGroup eq - (Poly+-assoc* : (PS QS RS : B) → (PS Poly+* (QS Poly+* RS)) ≡ ((PS Poly+* QS) Poly+* RS)) - (Poly+-Rid* : (PS : B) → (PS Poly+* 0P*) ≡ PS) - (Poly+-comm* : (PS QS : B) → (PS Poly+* QS) ≡ (QS Poly+* PS)) + (poly+Assoc* : (PS QS RS : B) → (PS poly+* (QS poly+* RS)) ≡ ((PS poly+* QS) poly+* RS)) + (poly+IdR* : (PS : B) → (PS poly+* 0P*) ≡ PS) + (poly+Comm* : (PS QS : B) → (PS poly+* QS) ≡ (QS poly+* PS)) -- Base eq (base-0P* : (v : Vec ℕ n) → (base* v 0r) ≡ 0P*) - (base-Poly+* : (v : Vec ℕ n) → (a b : A) → ((base* v a) Poly+* (base* v b)) ≡ (base* v (a + b))) + (base-poly+* : (v : Vec ℕ n) → (a b : A) → ((base* v a) poly+* (base* v b)) ≡ (base* v (a + b))) where f : Poly A' n → B - f = Poly-Ind-Set.f n (λ _ → B) (λ _ → iss) 0P* base* _Poly+*_ Poly+-assoc* Poly+-Rid* Poly+-comm* base-0P* base-Poly+* + f = Poly-Ind-Set.f n (λ _ → B) (λ _ → iss) 0P* base* _poly+*_ poly+Assoc* poly+IdR* poly+Comm* base-0P* base-poly+* module Poly-Ind-Prop @@ -108,16 +108,16 @@ module _ (A' : CommRing ℓ) where -- elements (0P* : F 0P) (base* : (v : Vec ℕ n) → (a : A) → F (base v a)) - (_Poly+*_ : {P Q : Poly A' n} → (PS : F P) → (QS : F Q) → F (P Poly+ Q)) + (_poly+*_ : {P Q : Poly A' n} → (PS : F P) → (QS : F Q) → F (P poly+ Q)) where f : (P : Poly A' n) → F P - f = Poly-Ind-Set.f n F (λ P → isProp→isSet (ispd P)) 0P* base* _Poly+*_ - (λ {P Q R} PS QS RQ → toPathP (ispd _ (transport (λ i → F (Poly+-assoc P Q R i)) _) _)) - (λ {P} PS → toPathP (ispd _ (transport (λ i → F (Poly+-Rid P i)) _) _)) - (λ {P Q} PS QS → toPathP (ispd _ (transport (λ i → F (Poly+-comm P Q i)) _) _)) + f = Poly-Ind-Set.f n F (λ P → isProp→isSet (ispd P)) 0P* base* _poly+*_ + (λ {P Q R} PS QS RQ → toPathP (ispd _ (transport (λ i → F (poly+Assoc P Q R i)) _) _)) + (λ {P} PS → toPathP (ispd _ (transport (λ i → F (poly+IdR P i)) _) _)) + (λ {P Q} PS QS → toPathP (ispd _ (transport (λ i → F (poly+Comm P Q i)) _) _)) (λ v → toPathP (ispd _ (transport (λ i → F (base-0P v i)) _) _)) - (λ v a b → toPathP (ispd _ (transport (λ i → F (base-Poly+ v a b i)) _) _)) + (λ v a b → toPathP (ispd _ (transport (λ i → F (base-poly+ v a b i)) _) _)) module Poly-Rec-Prop @@ -128,8 +128,8 @@ module _ (A' : CommRing ℓ) where -- elements (0P* : B) (base* : (v : Vec ℕ n) → (a : A) → B) - (_Poly+*_ : B → B → B) + (_poly+*_ : B → B → B) where f : Poly A' n → B - f = Poly-Ind-Prop.f n (λ _ → B) (λ _ → isp) 0P* base* _Poly+*_ + f = Poly-Ind-Prop.f n (λ _ → B) (λ _ → isp) 0P* base* _poly+*_ diff --git a/Cubical/Algebra/Polynomials/Multivariate/Equiv-Polyn-nPoly.agda b/Cubical/Algebra/Polynomials/Multivariate/Equiv-Polyn-nPoly.agda deleted file mode 100644 index 02dda272d5..0000000000 --- a/Cubical/Algebra/Polynomials/Multivariate/Equiv-Polyn-nPoly.agda +++ /dev/null @@ -1,43 +0,0 @@ -{-# OPTIONS --safe --experimental-lossy-unification #-} -module Cubical.Algebra.Polynomials.Multivariate.Equiv-Polyn-nPoly where - -open import Cubical.Foundations.Everything - -open import Cubical.Data.Nat renaming (_+_ to _+n_; _·_ to _·n_) -open import Cubical.Data.Vec -open import Cubical.Data.Sigma - -open import Cubical.Algebra.Ring -open import Cubical.Algebra.CommRing - -open import Cubical.Algebra.Polynomials.Univariate.Base -open import Cubical.Algebra.CommRing.Instances.UnivariatePoly - -open import Cubical.Algebra.Polynomials.Multivariate.Base -open import Cubical.Algebra.Polynomials.Multivariate.Properties -open import Cubical.Algebra.CommRing.Instances.MultivariatePoly -open import Cubical.Algebra.Polynomials.Multivariate.Equiv-Polyn-nPoly.Poly0-A -open import Cubical.Algebra.Polynomials.Multivariate.Equiv-Polyn-nPoly.Poly1-Poly -open import Cubical.Algebra.Polynomials.Multivariate.Equiv-Polyn-nPoly.Comp-Poly -open import Cubical.Algebra.Polynomials.Multivariate.Equiv-Polyn-nPoly.Induced-Poly - - -open Nth-Poly-structure -open CommRingEquivs renaming (compCommRingEquiv to _∘-ecr_ ; invCommRingEquiv to inv-ecr) - -private variable - ℓ : Level - - ------------------------------------------------------------------------------ --- Definition - -nPoly : (A' : CommRing ℓ) → (n : ℕ) → CommRing ℓ -nPoly A' zero = A' -nPoly A' (suc n) = UnivariatePoly (nPoly A' n) - -Equiv-Polyn-nPoly : (A' : CommRing ℓ) → (n : ℕ) → CommRingEquiv (PolyCommRing A' n) (nPoly A' n) -Equiv-Polyn-nPoly A' zero = CRE-Poly0-A A' -Equiv-Polyn-nPoly A' (suc n) = inv-ecr _ _ (CRE-PolyN∘M-PolyN+M A' 1 n) - ∘-ecr (lift-equiv-poly _ _ (Equiv-Polyn-nPoly A' n) 1 - ∘-ecr CRE-Poly1-Poly: (nPoly A' n)) diff --git a/Cubical/Algebra/Polynomials/Multivariate/Equiv-Polyn-nPoly/Induced-Poly.agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/AB-An[X]Bn[X].agda similarity index 84% rename from Cubical/Algebra/Polynomials/Multivariate/Equiv-Polyn-nPoly/Induced-Poly.agda rename to Cubical/Algebra/Polynomials/Multivariate/EquivCarac/AB-An[X]Bn[X].agda index 1b6c5558da..1995a0d548 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/Equiv-Polyn-nPoly/Induced-Poly.agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/AB-An[X]Bn[X].agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --experimental-lossy-unification #-} -module Cubical.Algebra.Polynomials.Multivariate.Equiv-Polyn-nPoly.Induced-Poly where +module Cubical.Algebra.Polynomials.Multivariate.EquivCarac.AB-An[X]Bn[X] where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -33,12 +33,12 @@ makeCommRingHomPoly : (A' B' : CommRing ℓ) → (f : CommRingHom A' B') → (n fst (makeCommRingHomPoly A' B' (f , fcrhom) n) = Poly-Rec-Set.f A' n (Poly B' n) trunc 0P (λ v a → base v (f a)) - _Poly+_ - Poly+-assoc - Poly+-Rid - Poly+-comm + _poly+_ + poly+Assoc + poly+IdR + poly+Comm (λ v → (cong (base v) (pres0 fcrhom)) ∙ (base-0P v)) - λ v a b → (base-Poly+ v (f a) (f b)) ∙ (cong (base v) (sym (pres+ fcrhom a b))) + λ v a b → (base-poly+ v (f a) (f b)) ∙ (cong (base v) (sym (pres+ fcrhom a b))) snd (makeCommRingHomPoly A' B' (f , fcrhom) n) = makeIsRingHom (cong (base (replicate zero)) (pres1 fcrhom)) (λ P Q → refl) @@ -47,8 +47,8 @@ snd (makeCommRingHomPoly A' B' (f , fcrhom) n) = makeIsRingHom (λ v a → Poly-Ind-Prop.f A' n _ (λ _ → trunc _ _) refl (λ v' a' → cong (base (v +n-vec v')) (pres· fcrhom a a')) - λ {U V} ind-U ind-V → cong₂ _Poly+_ ind-U ind-V) - λ {U V} ind-U ind-V Q → cong₂ _Poly+_ (ind-U Q) (ind-V Q)) + λ {U V} ind-U ind-V → cong₂ _poly+_ ind-U ind-V) + λ {U V} ind-U ind-V Q → cong₂ _poly+_ (ind-U Q) (ind-V Q)) @@ -73,9 +73,9 @@ fst (lift-equiv-poly A' B' e n) = isoToEquiv is Iso.rightInv is = (Poly-Ind-Prop.f B' n _ (λ _ → trunc _ _) refl (λ v a → cong (base v) (secEq et a)) - λ {U V} ind-U ind-V → cong₂ _Poly+_ ind-U ind-V) + λ {U V} ind-U ind-V → cong₂ _poly+_ ind-U ind-V) Iso.leftInv is = (Poly-Ind-Prop.f A' n _ (λ _ → trunc _ _) refl (λ v a → cong (base v) (retEq et a)) - λ {U V} ind-U ind-V → cong₂ _Poly+_ ind-U ind-V) + λ {U V} ind-U ind-V → cong₂ _poly+_ ind-U ind-V) snd (lift-equiv-poly A' B' e n) = snd (makeCommRingHomPoly A' B' (fst (fst e) , snd e) n) diff --git a/Cubical/Algebra/Polynomials/Multivariate/Equiv-A[X]X-A.agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/A[X]X-A.agda similarity index 96% rename from Cubical/Algebra/Polynomials/Multivariate/Equiv-A[X]X-A.agda rename to Cubical/Algebra/Polynomials/Multivariate/EquivCarac/A[X]X-A.agda index b7ea410f75..9b73569714 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/Equiv-A[X]X-A.agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/A[X]X-A.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --experimental-lossy-unification #-} -module Cubical.Algebra.Polynomials.Multivariate.Equiv-A[X]X-A where +module Cubical.Algebra.Polynomials.Multivariate.EquivCarac.A[X]X-A where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function @@ -17,7 +17,7 @@ open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.QuotientRing open import Cubical.Algebra.Polynomials.Multivariate.Base -open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤ to ℤCR) +open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤCommRing to ℤCR) open import Cubical.Algebra.CommRing.Instances.MultivariatePoly open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-Quotient open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-notationZ @@ -26,7 +26,6 @@ open import Cubical.Relation.Nullary open import Cubical.HITs.SetQuotients as SQ open import Cubical.HITs.PropositionalTruncation as PT - renaming (∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) private variable ℓ : Level @@ -195,7 +194,7 @@ module Properties-Equiv-QuotientXn-A A→A[x] a = base (0 ∷ []) a A→A[x]-pres+ : (a a' : A) → A→A[x] (a +A a') ≡ A→A[x] a +PA A→A[x] a' - A→A[x]-pres+ a a' = sym (base-Poly+ (0 ∷ []) a a') + A→A[x]-pres+ a a' = sym (base-poly+ (0 ∷ []) a a') A→A[x]/x : A → A[x]/x A→A[x]/x = [_] ∘ A→A[x] @@ -226,10 +225,10 @@ module Properties-Equiv-QuotientXn-A where base-eq : (a : A) → (v : Vec ℕ 1) → A→A[x]/x (A[x]/x→A [ (base v a) ]) ≡ [ (base v a) ] base-eq a (zero ∷ []) = cong [_] refl - base-eq a (suc k ∷ []) = eq/ (base (0 ∷ []) 0A) (base (suc k ∷ []) a) ∣ ((λ x → base (k ∷ []) (-A a)) , helper) ∣₋₁ + base-eq a (suc k ∷ []) = eq/ (base (0 ∷ []) 0A) (base (suc k ∷ []) a) ∣ ((λ x → base (k ∷ []) (-A a)) , helper) ∣₁ where helper : _ - helper = cong (λ X → X Poly+ base (suc k ∷ []) (-A a)) (base-0P _) + helper = cong (λ X → X poly+ base (suc k ∷ []) (-A a)) (base-0P _) ∙ +PALid _ ∙ sym (+PARid _ ∙ cong₂ base diff --git a/Cubical/Algebra/Polynomials/Multivariate/Equiv-Polyn-nPoly/Comp-Poly.agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[Am[X]]-Anm[X].agda similarity index 74% rename from Cubical/Algebra/Polynomials/Multivariate/Equiv-Polyn-nPoly/Comp-Poly.agda rename to Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[Am[X]]-Anm[X].agda index 98d0976ca2..21a990f92a 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/Equiv-Polyn-nPoly/Comp-Poly.agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[Am[X]]-Anm[X].agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --experimental-lossy-unification #-} -module Cubical.Algebra.Polynomials.Multivariate.Equiv-Polyn-nPoly.Comp-Poly where +module Cubical.Algebra.Polynomials.Multivariate.EquivCarac.An[Am[X]]-Anm[X] where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism @@ -38,21 +38,21 @@ module Comp-Poly-nm (A' : CommRing ℓ) (n m : ℕ) where N∘M→N+M-b v = Poly-Rec-Set.f A' m (Poly A' (n +n m)) trunc 0P (λ v' a → base (v ++ v') a) - _Poly+_ - Poly+-assoc - Poly+-Rid - Poly+-comm + _poly+_ + poly+Assoc + poly+IdR + poly+Comm (λ v' → base-0P (v ++ v')) - (λ v' a b → base-Poly+ (v ++ v') a b) + (λ v' a b → base-poly+ (v ++ v') a b) N∘M→N+M : Poly (PolyCommRing A' m) n → Poly A' (n +n m) N∘M→N+M = Poly-Rec-Set.f (PolyCommRing A' m) n (Poly A' (n +n m)) trunc 0P N∘M→N+M-b - _Poly+_ - Poly+-assoc - Poly+-Rid - Poly+-comm + _poly+_ + poly+Assoc + poly+IdR + poly+Comm (λ _ → refl) (λ v a b → refl) @@ -65,13 +65,13 @@ module Comp-Poly-nm (A' : CommRing ℓ) (n m : ℕ) where 0P (λ v a → let v , v' = sep-vec n m v in base v (base v' a)) - _Poly+_ - Poly+-assoc - Poly+-Rid - Poly+-comm + _poly+_ + poly+Assoc + poly+IdR + poly+Comm (λ v → (cong (base (fst (sep-vec n m v))) (base-0P (snd (sep-vec n m v)))) ∙ (base-0P (fst (sep-vec n m v)))) - λ v a b → base-Poly+ (fst (sep-vec n m v)) (base (snd (sep-vec n m v)) a) (base (snd (sep-vec n m v)) b) - ∙ cong (base (fst (sep-vec n m v))) (base-Poly+ (snd (sep-vec n m v)) a b) + λ v a b → base-poly+ (fst (sep-vec n m v)) (base (snd (sep-vec n m v)) a) (base (snd (sep-vec n m v)) b) + ∙ cong (base (fst (sep-vec n m v))) (base-poly+ (snd (sep-vec n m v)) a b) ----------------------------------------------------------------------------- @@ -83,7 +83,7 @@ module Comp-Poly-nm (A' : CommRing ℓ) (n m : ℕ) where (λ _ → trunc _ _) refl (λ v a → cong (λ X → base X a) (sep-vec-id n m v)) - (λ {U V} ind-U ind-V → cong₂ _Poly+_ ind-U ind-V) + (λ {U V} ind-U ind-V → cong₂ _poly+_ ind-U ind-V) ----------------------------------------------------------------------------- @@ -99,8 +99,8 @@ module Comp-Poly-nm (A' : CommRing ℓ) (n m : ℕ) where (λ _ → trunc _ _) (sym (base-0P v)) (λ v' a → cong₂ base (sep-vec-fst n m v v') (cong (λ X → base X a) (sep-vec-snd n m v v'))) - (λ {U V} ind-U ind-V → (cong₂ _Poly+_ ind-U ind-V) ∙ (base-Poly+ v U V))) - (λ {U V} ind-U ind-V → cong₂ _Poly+_ ind-U ind-V ) + (λ {U V} ind-U ind-V → (cong₂ _poly+_ ind-U ind-V) ∙ (base-poly+ v U V))) + (λ {U V} ind-U ind-V → cong₂ _poly+_ ind-U ind-V ) ----------------------------------------------------------------------------- @@ -109,46 +109,46 @@ module Comp-Poly-nm (A' : CommRing ℓ) (n m : ℕ) where map-0P : N∘M→N+M (0P) ≡ 0P map-0P = refl - N∘M→N+M-gmorph : (P Q : Poly (PolyCommRing A' m) n) → N∘M→N+M ( P Poly+ Q) ≡ N∘M→N+M P Poly+ N∘M→N+M Q + N∘M→N+M-gmorph : (P Q : Poly (PolyCommRing A' m) n) → N∘M→N+M ( P poly+ Q) ≡ N∘M→N+M P poly+ N∘M→N+M Q N∘M→N+M-gmorph = λ P Q → refl map-1P : N∘M→N+M (N∘Mr.1P) ≡ N+Mr.1P map-1P = cong (λ X → base X 1r) (rep-concat n m 0 ) - N∘M→N+M-rmorph : (P Q : Poly (PolyCommRing A' m) n) → N∘M→N+M ( P N∘Mr.Poly* Q) ≡ N∘M→N+M P N+Mr.Poly* N∘M→N+M Q + N∘M→N+M-rmorph : (P Q : Poly (PolyCommRing A' m) n) → N∘M→N+M ( P N∘Mr.poly* Q) ≡ N∘M→N+M P N+Mr.poly* N∘M→N+M Q N∘M→N+M-rmorph = -- Ind P Poly-Ind-Prop.f (PolyCommRing A' m) n - (λ P → (Q : Poly (PolyCommRing A' m) n) → N∘M→N+M (P N∘Mr.Poly* Q) ≡ (N∘M→N+M P N+Mr.Poly* N∘M→N+M Q)) + (λ P → (Q : Poly (PolyCommRing A' m) n) → N∘M→N+M (P N∘Mr.poly* Q) ≡ (N∘M→N+M P N+Mr.poly* N∘M→N+M Q)) (λ P p q i Q j → trunc _ _ (p Q) (q Q) i j) (λ Q → refl) (λ v → -- Ind Base P Poly-Ind-Prop.f A' m - (λ P → (Q : Poly (PolyCommRing A' m) n) → N∘M→N+M (base v P N∘Mr.Poly* Q) ≡ (N∘M→N+M (base v P) N+Mr.Poly* N∘M→N+M Q)) + (λ P → (Q : Poly (PolyCommRing A' m) n) → N∘M→N+M (base v P N∘Mr.poly* Q) ≡ (N∘M→N+M (base v P) N+Mr.poly* N∘M→N+M Q)) (λ P p q i Q j → trunc _ _ (p Q) (q Q) i j) - (λ Q → cong (λ X → N∘M→N+M (X N∘Mr.Poly* Q)) (base-0P v)) + (λ Q → cong (λ X → N∘M→N+M (X N∘Mr.poly* Q)) (base-0P v)) (λ v' a → -- Ind Q Poly-Ind-Prop.f (PolyCommRing A' m) n - (λ Q → N∘M→N+M (base v (base v' a) N∘Mr.Poly* Q) ≡ (N∘M→N+M (base v (base v' a)) N+Mr.Poly* N∘M→N+M Q)) + (λ Q → N∘M→N+M (base v (base v' a) N∘Mr.poly* Q) ≡ (N∘M→N+M (base v (base v' a)) N+Mr.poly* N∘M→N+M Q)) (λ _ → trunc _ _) - (sym (N+Mr.0PRightAnnihilatesPoly (N∘M→N+M (base v (base v' a))))) + (sym (N+Mr.poly*AnnihilR (N∘M→N+M (base v (base v' a))))) (λ w → -- Ind base Q Poly-Ind-Prop.f A' m _ (λ _ → trunc _ _) - (sym (N+Mr.0PRightAnnihilatesPoly (N∘M→N+M (base v (base v' a))))) + (sym (N+Mr.poly*AnnihilR (N∘M→N+M (base v (base v' a))))) (λ w' b → cong (λ X → base X (a · b)) (+n-vec-concat n m v w v' w')) - λ {U V} ind-U ind-V → cong (λ X → N∘M→N+M (base v (base v' a) N∘Mr.Poly* X)) (sym (base-Poly+ w U V)) - ∙ cong₂ (_Poly+_ ) ind-U ind-V - ∙ sym (cong (λ X → N∘M→N+M (base v (base v' a)) N+Mr.Poly* N∘M→N+M X) (base-Poly+ w U V)) ) + λ {U V} ind-U ind-V → cong (λ X → N∘M→N+M (base v (base v' a) N∘Mr.poly* X)) (sym (base-poly+ w U V)) + ∙ cong₂ (_poly+_ ) ind-U ind-V + ∙ sym (cong (λ X → N∘M→N+M (base v (base v' a)) N+Mr.poly* N∘M→N+M X) (base-poly+ w U V)) ) -- End Ind base Q - λ {U V} ind-U ind-V → cong₂ _Poly+_ ind-U ind-V) + λ {U V} ind-U ind-V → cong₂ _poly+_ ind-U ind-V) -- End Ind Q - λ {U V} ind-U ind-V Q → cong (λ X → N∘M→N+M (X N∘Mr.Poly* Q)) (sym (base-Poly+ v U V)) - ∙ cong₂ _Poly+_ (ind-U Q) (ind-V Q) - ∙ sym (cong (λ X → (N∘M→N+M X) N+Mr.Poly* (N∘M→N+M Q)) (sym (base-Poly+ v U V)) )) + λ {U V} ind-U ind-V Q → cong (λ X → N∘M→N+M (X N∘Mr.poly* Q)) (sym (base-poly+ v U V)) + ∙ cong₂ _poly+_ (ind-U Q) (ind-V Q) + ∙ sym (cong (λ X → (N∘M→N+M X) N+Mr.poly* (N∘M→N+M Q)) (sym (base-poly+ v U V)) )) -- End Ind base P - λ {U V} ind-U ind-V Q → cong₂ _Poly+_ (ind-U Q) (ind-V Q) + λ {U V} ind-U ind-V Q → cong₂ _poly+_ (ind-U Q) (ind-V Q) -- End Ind P diff --git a/Cubical/Algebra/Polynomials/Multivariate/Equiv-PolynQuotient-A.agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[X]X-A.agda similarity index 96% rename from Cubical/Algebra/Polynomials/Multivariate/Equiv-PolynQuotient-A.agda rename to Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[X]X-A.agda index edc68430c4..bc79ade974 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/Equiv-PolynQuotient-A.agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[X]X-A.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --experimental-lossy-unification #-} -module Cubical.Algebra.Polynomials.Multivariate.Equiv-PolynQuotient-A where +module Cubical.Algebra.Polynomials.Multivariate.EquivCarac.An[X]X-A where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function @@ -20,7 +20,7 @@ open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.QuotientRing open import Cubical.Algebra.Polynomials.Multivariate.Base -open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤ to ℤCR) +open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤCommRing to ℤCR) open import Cubical.Algebra.CommRing.Instances.MultivariatePoly open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-Quotient @@ -28,7 +28,6 @@ open import Cubical.Relation.Nullary open import Cubical.HITs.SetQuotients as SQ open import Cubical.HITs.PropositionalTruncation as PT - renaming (∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) private variable ℓ : Level @@ -205,7 +204,7 @@ module Properties-Equiv-QuotientXn-A A→PA-pres1 = refl A→PA-pres+ : (a a' : A) → A→PA (a +A a') ≡ (A→PA a) +PA (A→PA a') - A→PA-pres+ a a' = sym (base-Poly+ _ _ _) + A→PA-pres+ a a' = sym (base-poly+ _ _ _) A→PA-pres· : (a a' : A) → A→PA (a ·A a') ≡ (A→PA a) ·PA (A→PA a') A→PA-pres· a a' = cong (λ X → base X (a ·A a')) (sym (+n-vec-rid _)) @@ -246,10 +245,10 @@ module Properties-Equiv-QuotientXn-A ... | yes p = cong [_] (cong (λ X → base X a) (sym p)) ... | no ¬p with (pred-vec-≢0 v ¬p) ... | k , v' , infkn , eqvv' = eq/ (base (replicate 0) 0A) - (base v a) ∣ ((genδ-FinVec n k (base v' (-A a)) 0PA) , helper) ∣₋₁ + (base v a) ∣ ((genδ-FinVec n k (base v' (-A a)) 0PA) , helper) ∣₁ where helper : _ - helper = cong (λ X → X Poly+ base v (-A a)) (base-0P (replicate 0)) + helper = cong (λ X → X poly+ base v (-A a)) (base-0P (replicate 0)) ∙ +PALid (base v (-A a)) ∙ sym ( genδ-FinVec-ℕLinearCombi ((A[X1,···,Xn] Ar n)) n k infkn (base v' (-A a)) ( Ar n) diff --git a/Cubical/Algebra/Polynomials/Multivariate/Equiv-Polyn-nPoly/Poly0-A.agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/Poly0-A.agda similarity index 85% rename from Cubical/Algebra/Polynomials/Multivariate/Equiv-Polyn-nPoly/Poly0-A.agda rename to Cubical/Algebra/Polynomials/Multivariate/EquivCarac/Poly0-A.agda index 1771f74509..375cf455f4 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/Equiv-Polyn-nPoly/Poly0-A.agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/Poly0-A.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --experimental-lossy-unification #-} -module Cubical.Algebra.Polynomials.Multivariate.Equiv-Polyn-nPoly.Poly0-A where +module Cubical.Algebra.Polynomials.Multivariate.EquivCarac.Poly0-A where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism @@ -53,7 +53,7 @@ module Equiv-Poly0-A (A' : CommRing ℓ) where (λ _ → trunc _ _) (base-0P []) (λ { [] a → refl }) - λ {U V} ind-U ind-V → (sym (base-Poly+ [] (Poly0→A U) (Poly0→A V))) ∙ (cong₂ _Poly+_ ind-U ind-V) + λ {U V} ind-U ind-V → (sym (base-poly+ [] (Poly0→A U) (Poly0→A V))) ∙ (cong₂ _poly+_ ind-U ind-V) ----------------------------------------------------------------------------- @@ -62,19 +62,19 @@ module Equiv-Poly0-A (A' : CommRing ℓ) where map-0P : Poly0→A 0P ≡ 0r map-0P = refl - Poly0→A-gmorph : (P Q : Poly A' 0) → Poly0→A ( P Poly+ Q) ≡ Poly0→A P + Poly0→A Q + Poly0→A-gmorph : (P Q : Poly A' 0) → Poly0→A ( P poly+ Q) ≡ Poly0→A P + Poly0→A Q Poly0→A-gmorph P Q = refl map-1P : Poly0→A 1P ≡ 1r map-1P = refl - Poly0→A-rmorph : (P Q : Poly A' 0) → Poly0→A ( P Poly* Q) ≡ Poly0→A P · Poly0→A Q + Poly0→A-rmorph : (P Q : Poly A' 0) → Poly0→A ( P poly* Q) ≡ Poly0→A P · Poly0→A Q Poly0→A-rmorph = Poly-Ind-Prop.f A' 0 - (λ P → (Q : Poly A' 0) → Poly0→A (P Poly* Q) ≡ Poly0→A P · Poly0→A Q) - (λ P p q i Q j → isSetA (Poly0→A (P Poly* Q)) (Poly0→A P · Poly0→A Q) (p Q) (q Q) i j) + (λ P → (Q : Poly A' 0) → Poly0→A (P poly* Q) ≡ Poly0→A P · Poly0→A Q) + (λ P p q i Q j → isSetA (Poly0→A (P poly* Q)) (Poly0→A P · Poly0→A Q) (p Q) (q Q) i j) (λ Q → sym (RingTheory.0LeftAnnihilates (CommRing→Ring A') (Poly0→A Q))) (λ v a → Poly-Ind-Prop.f A' 0 - (λ P → Poly0→A (base v a Poly* P) ≡ Poly0→A (base v a) · Poly0→A P) + (λ P → Poly0→A (base v a poly* P) ≡ Poly0→A (base v a) · Poly0→A P) (λ _ → isSetA _ _) (sym (RingTheory.0RightAnnihilates (CommRing→Ring A') (Poly0→A (base v a)))) (λ v' a' → refl) diff --git a/Cubical/Algebra/Polynomials/Multivariate/Equiv-Polyn-nPoly/Poly1-Poly.agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/Poly1-1Poly.agda similarity index 65% rename from Cubical/Algebra/Polynomials/Multivariate/Equiv-Polyn-nPoly/Poly1-Poly.agda rename to Cubical/Algebra/Polynomials/Multivariate/EquivCarac/Poly1-1Poly.agda index 27b6d15e67..63dc564ada 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/Equiv-Polyn-nPoly/Poly1-Poly.agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/Poly1-1Poly.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --experimental-lossy-unification #-} -module Cubical.Algebra.Polynomials.Multivariate.Equiv-Polyn-nPoly.Poly1-Poly where +module Cubical.Algebra.Polynomials.Multivariate.EquivCarac.Poly1-1Poly where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism @@ -35,21 +35,21 @@ module Equiv-Poly1-Poly: (A' : CommRing ℓ) where open PolyModTheory A' renaming ( 0P to 0P: - ; Poly- to Poly:- - ; _Poly+_ to _Poly:+_ - ; Poly+Lid to Poly:+Lid - ; Poly+Rid to Poly:+Rid - ; Poly+Assoc to Poly:+Assoc - ; Poly+Inverses to Poly:+Inverses - ; Poly+Comm to Poly:+Comm - ; _Poly*_ to _Poly:*_ + ; Poly- to poly:- + ; _Poly+_ to _poly:+_ + ; Poly+Lid to poly:+IdL + ; Poly+Rid to poly:+IdR + ; Poly+Assoc to poly:+Assoc + ; Poly+Inverses to poly:+Inverses + ; Poly+Comm to poly:+Comm + ; _Poly*_ to _poly:*_ ; 1P to 1P: - ; 0PRightAnnihilates to 0PRightAnnihilates: - ; 0PLeftAnnihilates to 0PLeftAnnihilates: - ; Poly*Lid to Poly:*Lid - ; Poly*Rid to Poly:*Rid - ; Poly*Associative to Poly:*Assoc - ; Poly*Commutative to Poly:*Comm + ; 0PRightAnnihilates to poly:*AnnihilR + ; 0PLeftAnnihilates to poly:*AnnihilL + ; Poly*Lid to poly:*IdL + ; Poly*Rid to poly:*IdR + ; Poly*Associative to poly:*Assoc + ; Poly*Commutative to poly:*Comm ; prod-Xn-0P to prod-Xn-0P: ) @@ -69,22 +69,22 @@ module Equiv-Poly1-Poly: (A' : CommRing ℓ) where trad-base-neutral : (v : Vec ℕ 1) → trad-base v 0r ≡ [] trad-base-neutral (n :: <>) = cong (prod-Xn n) drop0 ∙ prod-Xn-0P: n - trad-base-add : (v : Vec ℕ 1) → (a b : A) → (trad-base v a) Poly:+ (trad-base v b) ≡ trad-base v (a + b) + trad-base-add : (v : Vec ℕ 1) → (a b : A) → (trad-base v a) poly:+ (trad-base v b) ≡ trad-base v (a + b) trad-base-add (n :: <>) a b = prod-Xn-sum n (a ∷ []) (b ∷ []) Poly1→Poly: : Poly A' 1 → Poly: Poly1→Poly: = Poly-Rec-Set.f A' 1 Poly: isSetPoly: [] trad-base - (λ x y → x Poly:+ y) - Poly:+Assoc - Poly:+Rid - Poly:+Comm + (λ x y → x poly:+ y) + poly:+Assoc + poly:+IdR + poly:+Comm trad-base-neutral trad-base-add - Poly1→Poly:-gmorph : (P Q : Poly A' 1) → Poly1→Poly: (P Poly+ Q) ≡ (Poly1→Poly: P) Poly:+ (Poly1→Poly: Q) - Poly1→Poly:-gmorph P Q = refl + Poly1→Poly:-pres+ : (P Q : Poly A' 1) → Poly1→Poly: (P poly+ Q) ≡ (Poly1→Poly: P) poly:+ (Poly1→Poly: Q) + Poly1→Poly:-pres+ P Q = refl @@ -93,31 +93,31 @@ module Equiv-Poly1-Poly: (A' : CommRing ℓ) where Poly:→Poly1-int : (n : ℕ) → Poly: → Poly A' 1 Poly:→Poly1-int n [] = 0P - Poly:→Poly1-int n (a ∷ x) = (base (n :: <>) a) Poly+ Poly:→Poly1-int (suc n) x - Poly:→Poly1-int n (drop0 i) = ((cong (λ X → X Poly+ 0P) (base-0P (n :: <>))) ∙ (Poly+-Rid 0P)) i + Poly:→Poly1-int n (a ∷ x) = (base (n :: <>) a) poly+ Poly:→Poly1-int (suc n) x + Poly:→Poly1-int n (drop0 i) = ((cong (λ X → X poly+ 0P) (base-0P (n :: <>))) ∙ (poly+IdR 0P)) i Poly:→Poly1 : Poly: → Poly A' 1 Poly:→Poly1 x = Poly:→Poly1-int 0 x - Poly:→Poly1-int-gmorph : (x y : Poly:) → (n : ℕ) → Poly:→Poly1-int n (x Poly:+ y) ≡ (Poly:→Poly1-int n x) Poly+ (Poly:→Poly1-int n y) - Poly:→Poly1-int-gmorph = ElimProp.f - (λ x → (y : Poly:) → (n : ℕ) → Poly:→Poly1-int n (x Poly:+ y) ≡ (Poly:→Poly1-int n x Poly+ Poly:→Poly1-int n y)) - (λ y n → (cong (Poly:→Poly1-int n) (Poly:+Lid y)) ∙ (sym (Poly+-Lid (Poly:→Poly1-int n y)))) + Poly:→Poly1-int-pres+ : (x y : Poly:) → (n : ℕ) → Poly:→Poly1-int n (x poly:+ y) ≡ (Poly:→Poly1-int n x) poly+ (Poly:→Poly1-int n y) + Poly:→Poly1-int-pres+ = ElimProp.f + (λ x → (y : Poly:) → (n : ℕ) → Poly:→Poly1-int n (x poly:+ y) ≡ (Poly:→Poly1-int n x poly+ Poly:→Poly1-int n y)) + (λ y n → (cong (Poly:→Poly1-int n) (poly:+IdL y)) ∙ (sym (poly+IdL (Poly:→Poly1-int n y)))) (λ a x ind-x → ElimProp.f - (λ y → (n : ℕ) → Poly:→Poly1-int n ((a ∷ x) Poly:+ y) ≡ (Poly:→Poly1-int n (a ∷ x) Poly+ Poly:→Poly1-int n y)) - (λ n → sym (Poly+-Rid (Poly:→Poly1-int n (a ∷ x)))) + (λ y → (n : ℕ) → Poly:→Poly1-int n ((a ∷ x) poly:+ y) ≡ (Poly:→Poly1-int n (a ∷ x) poly+ Poly:→Poly1-int n y)) + (λ n → sym (poly+IdR (Poly:→Poly1-int n (a ∷ x)))) (λ b y ind-y n → sym ( - (Poly-com-adv (base (n :: <>) a) (Poly:→Poly1-int (suc n) x) (base (n :: <>) b) (Poly:→Poly1-int (suc n) y)) + (poly-com-adv (base (n :: <>) a) (Poly:→Poly1-int (suc n) x) (base (n :: <>) b) (Poly:→Poly1-int (suc n) y)) ∙ - (cong₂ _Poly+_ (base-Poly+ (n :: <>) a b) (sym (ind-x y (suc n))))) ) + (cong₂ _poly+_ (base-poly+ (n :: <>) a b) (sym (ind-x y (suc n))))) ) λ {y} p q i n j → trunc - (Poly:→Poly1-int n ((a ∷ x) Poly:+ y)) - (Poly:→Poly1-int n (a ∷ x) Poly+ Poly:→Poly1-int n y) + (Poly:→Poly1-int n ((a ∷ x) poly:+ y)) + (Poly:→Poly1-int n (a ∷ x) poly+ Poly:→Poly1-int n y) (p n) (q n) i j ) - λ {x} p q i y n j → trunc (Poly:→Poly1-int n (x Poly:+ y)) (Poly:→Poly1-int n x Poly+ Poly:→Poly1-int n y) (p y n) (q y n) i j + λ {x} p q i y n j → trunc (Poly:→Poly1-int n (x poly:+ y)) (Poly:→Poly1-int n x poly+ Poly:→Poly1-int n y) (p y n) (q y n) i j - Poly:→Poly1-gmorph : (x y : Poly:) → Poly:→Poly1 (x Poly:+ y) ≡ (Poly:→Poly1 x) Poly+ (Poly:→Poly1 y) - Poly:→Poly1-gmorph x y = Poly:→Poly1-int-gmorph x y 0 + Poly:→Poly1-pres+ : (x y : Poly:) → Poly:→Poly1 (x poly:+ y) ≡ (Poly:→Poly1 x) poly+ (Poly:→Poly1 y) + Poly:→Poly1-pres+ x y = Poly:→Poly1-int-pres+ x y 0 ----------------------------------------------------------------------------- -- section @@ -126,9 +126,9 @@ module Equiv-Poly1-Poly: (A' : CommRing ℓ) where e-sect-int = ElimProp.f (λ z → (n : ℕ) → Poly1→Poly: (Poly:→Poly1-int n z) ≡ prod-Xn n z) (λ n → sym (prod-Xn-0P: n)) - (λ a x ind-x n → ((prod-Xn n [ a ] ) Poly:+ (Poly1→Poly: (Poly:→Poly1-int (suc n) x))) - ≡⟨ cong (λ X → prod-Xn n [ a ] Poly:+ X) (ind-x (suc n)) ⟩ - (prod-Xn n (a ∷ []) Poly:+ ( 0r ∷ prod-Xn n x)) + (λ a x ind-x n → ((prod-Xn n [ a ] ) poly:+ (Poly1→Poly: (Poly:→Poly1-int (suc n) x))) + ≡⟨ cong (λ X → prod-Xn n [ a ] poly:+ X) (ind-x (suc n)) ⟩ + (prod-Xn n (a ∷ []) poly:+ ( 0r ∷ prod-Xn n x)) ≡⟨ prod-Xn-∷ n a x ⟩ prod-Xn n (a ∷ x) ∎) (λ {x} p q i n j → isSetPoly: (Poly1→Poly: (Poly:→Poly1-int n x)) (prod-Xn n x) (p n) (q n) i j) @@ -142,9 +142,9 @@ module Equiv-Poly1-Poly: (A' : CommRing ℓ) where -- retraction idde : (m n : ℕ) → (a : A) → Poly:→Poly1-int n (prod-Xn m [ a ]) ≡ base ((n +n m) :: <>) a - idde zero n a = Poly+-Rid (base (n :: <>) a) ∙ cong (λ X → base (X :: <>) a) (sym (+-zero n)) - idde (suc m) n a = cong (λ X → X Poly+ Poly:→Poly1-int (suc n) (prod-Xn m (a ∷ []))) (base-0P (n :: <>)) - ∙ Poly+-Lid (Poly:→Poly1-int (suc n) (prod-Xn m (a ∷ []))) + idde zero n a = poly+IdR (base (n :: <>) a) ∙ cong (λ X → base (X :: <>) a) (sym (+-zero n)) + idde (suc m) n a = cong (λ X → X poly+ Poly:→Poly1-int (suc n) (prod-Xn m (a ∷ []))) (base-0P (n :: <>)) + ∙ poly+IdL (Poly:→Poly1-int (suc n) (prod-Xn m (a ∷ []))) ∙ idde m (suc n) a ∙ cong (λ X → base (X :: <>) a) (sym (+-suc n m)) @@ -158,37 +158,37 @@ module Equiv-Poly1-Poly: (A' : CommRing ℓ) where (λ _ → trunc _ _) refl (λ v a → idde-v v a) - λ {P Q} ind-P ind-Q → cong Poly:→Poly1 (Poly1→Poly:-gmorph P Q) + λ {P Q} ind-P ind-Q → cong Poly:→Poly1 (Poly1→Poly:-pres+ P Q) ∙ - Poly:→Poly1-gmorph (Poly1→Poly: P) (Poly1→Poly: Q) + Poly:→Poly1-pres+ (Poly1→Poly: P) (Poly1→Poly: Q) ∙ - cong₂ _Poly+_ ind-P ind-Q + cong₂ _poly+_ ind-P ind-Q ----------------------------------------------------------------------------- -- Ring morphism - map-1P : Poly1→Poly: 1P ≡ 1P: - map-1P = refl + Poly1→Poly:-pres1 : Poly1→Poly: 1P ≡ 1P: + Poly1→Poly:-pres1 = refl trad-base-prod : (v v' : Vec ℕ 1) → (a a' : A) → trad-base (v +n-vec v') (a · a') ≡ - (trad-base v a Poly:* trad-base v' a') + (trad-base v a poly:* trad-base v' a') trad-base-prod (k :: <>) (l :: <>) a a' = sym ((prod-Xn-prod k l [ a ] [ a' ]) ∙ cong (λ X → prod-Xn (k +n l) [ X ]) (+Rid (a · a'))) - Poly1→Poly:-rmorph : (P Q : Poly A' 1) → Poly1→Poly: (P Poly* Q) ≡ (Poly1→Poly: P) Poly:* (Poly1→Poly: Q) - Poly1→Poly:-rmorph = Poly-Ind-Prop.f A' 1 - (λ P → (Q : Poly A' 1) → Poly1→Poly: (P Poly* Q) ≡ (Poly1→Poly: P Poly:* Poly1→Poly: Q)) - (λ P p q i Q j → isSetPoly: (Poly1→Poly: (P Poly* Q)) (Poly1→Poly: P Poly:* Poly1→Poly: Q) (p Q) (q Q) i j) + Poly1→Poly:-pres· : (P Q : Poly A' 1) → Poly1→Poly: (P poly* Q) ≡ (Poly1→Poly: P) poly:* (Poly1→Poly: Q) + Poly1→Poly:-pres· = Poly-Ind-Prop.f A' 1 + (λ P → (Q : Poly A' 1) → Poly1→Poly: (P poly* Q) ≡ (Poly1→Poly: P poly:* Poly1→Poly: Q)) + (λ P p q i Q j → isSetPoly: (Poly1→Poly: (P poly* Q)) (Poly1→Poly: P poly:* Poly1→Poly: Q) (p Q) (q Q) i j) (λ Q → refl) (λ v a → Poly-Ind-Prop.f A' 1 - (λ P → Poly1→Poly: (base v a Poly* P) ≡ (Poly1→Poly: (base v a) Poly:* Poly1→Poly: P)) + (λ P → Poly1→Poly: (base v a poly* P) ≡ (Poly1→Poly: (base v a) poly:* Poly1→Poly: P)) (λ _ → isSetPoly: _ _) - (sym (0PLeftAnnihilates: (trad-base v a))) + (sym (poly:*AnnihilL (trad-base v a))) (λ v' a' → trad-base-prod v v' a a') - λ {U V} ind-U ind-V → (cong₂ _Poly:+_ ind-U ind-V) + λ {U V} ind-U ind-V → (cong₂ _poly:+_ ind-U ind-V) ∙ sym (Poly*LDistrPoly+ (Poly1→Poly: (base v a)) (Poly1→Poly: U) (Poly1→Poly: V))) - λ {U V} ind-U ind-V Q → (cong₂ _Poly:+_ (ind-U Q) (ind-V Q)) + λ {U V} ind-U ind-V Q → (cong₂ _poly:+_ (ind-U Q) (ind-V Q)) ∙ sym (Poly*RDistrPoly+ (Poly1→Poly: U) (Poly1→Poly: V) (Poly1→Poly: Q)) @@ -208,4 +208,7 @@ module _ (A' : CommRing ℓ) where Iso.inv is = Poly:→Poly1 Iso.rightInv is = e-sect Iso.leftInv is = e-retr - snd CRE-Poly1-Poly: = makeIsRingHom map-1P Poly1→Poly:-gmorph Poly1→Poly:-rmorph + snd CRE-Poly1-Poly: = makeIsRingHom + Poly1→Poly:-pres1 + Poly1→Poly:-pres+ + Poly1→Poly:-pres· diff --git a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/Polyn-nPoly.agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/Polyn-nPoly.agda new file mode 100644 index 0000000000..d9b39329fc --- /dev/null +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/Polyn-nPoly.agda @@ -0,0 +1,35 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.Polynomials.Multivariate.EquivCarac.Polyn-nPoly where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Nat renaming (_+_ to _+n_; _·_ to _·n_) +open import Cubical.Data.Vec +open import Cubical.Data.Sigma + +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommRing + +open import Cubical.Algebra.CommRing.Instances.UnivariatePoly +open import Cubical.Algebra.CommRing.Instances.MultivariatePoly + +open import Cubical.Algebra.Polynomials.Multivariate.EquivCarac.Poly0-A +open import Cubical.Algebra.Polynomials.Multivariate.EquivCarac.Poly1-1Poly +open import Cubical.Algebra.Polynomials.Multivariate.EquivCarac.An[Am[X]]-Anm[X] +open import Cubical.Algebra.Polynomials.Multivariate.EquivCarac.AB-An[X]Bn[X] + +open CommRingEquivs renaming (compCommRingEquiv to _∘-ecr_ ; invCommRingEquiv to inv-ecr) + +private variable + ℓ : Level + + +----------------------------------------------------------------------------- +-- Definition + +Equiv-Polyn-nPoly : (A' : CommRing ℓ) → (n : ℕ) → CommRingEquiv (PolyCommRing A' n) (nUnivariatePoly A' n) +Equiv-Polyn-nPoly A' zero = CRE-Poly0-A A' +Equiv-Polyn-nPoly A' (suc n) = inv-ecr _ _ (CRE-PolyN∘M-PolyN+M A' 1 n) + ∘-ecr (lift-equiv-poly _ _ (Equiv-Polyn-nPoly A' n) 1 + ∘-ecr CRE-Poly1-Poly: (nUnivariatePoly A' n)) diff --git a/Cubical/Algebra/Polynomials/Multivariate/Properties.agda b/Cubical/Algebra/Polynomials/Multivariate/Properties.agda index ba0e7c4c17..c115bd3c51 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/Properties.agda +++ b/Cubical/Algebra/Polynomials/Multivariate/Properties.agda @@ -12,6 +12,7 @@ open import Cubical.Algebra.CommRing open import Cubical.Algebra.Polynomials.Multivariate.Base + private variable ℓ ℓ' : Level @@ -26,61 +27,63 @@ module Nth-Poly-structure (A' : CommRing ℓ) (n : ℕ) where ----------------------------------------------------------------------------- - Poly-com-adv : (P Q R S : Poly A' n) → ((P Poly+ Q) Poly+ (R Poly+ S) ≡ (P Poly+ R) Poly+ (Q Poly+ S)) - Poly-com-adv P Q R S = ((P Poly+ Q) Poly+ (R Poly+ S) ≡⟨ Poly+-assoc (P Poly+ Q) R S ⟩ - (((P Poly+ Q) Poly+ R) Poly+ S) ≡⟨ cong (λ X → X Poly+ S) (sym (Poly+-assoc P Q R)) ⟩ - ((P Poly+ (Q Poly+ R)) Poly+ S) ≡⟨ cong (λ X → (P Poly+ X) Poly+ S) (Poly+-comm Q R) ⟩ - ((P Poly+ (R Poly+ Q)) Poly+ S) ≡⟨ cong (λ X → X Poly+ S) (Poly+-assoc P R Q) ⟩ - (((P Poly+ R) Poly+ Q) Poly+ S) ≡⟨ sym (Poly+-assoc (P Poly+ R) Q S) ⟩ - ((P Poly+ R) Poly+ (Q Poly+ S)) ∎) - - - Poly+-Lid : (P : Poly A' n) → 0P Poly+ P ≡ P - Poly+-Lid P = (Poly+-comm 0P P) ∙ (Poly+-Rid P) - - - Poly-inv : Poly A' n → Poly A' n - Poly-inv = Poly-Rec-Set.f A' n (Poly A' n) trunc - 0P - (λ v a → base v (- a)) - (λ PS RS → PS Poly+ RS) - Poly+-assoc - Poly+-Rid - Poly+-comm - (λ v → base v (- 0r) ≡⟨ cong (base v) 0Selfinverse ⟩ base v 0r ≡⟨ base-0P v ⟩ 0P ∎) - λ v a b → (base-Poly+ v (- a) (- b)) ∙ (cong (base v) (-Dist a b)) - - - Poly-invinv : (P : Poly A' n) → Poly-inv (Poly-inv P) ≡ P - Poly-invinv = Poly-Ind-Prop.f A' n (λ P → Poly-inv (Poly-inv P) ≡ P) (λ _ → trunc _ _) - refl - (λ v a → cong (base v) (-Idempotent a)) - λ {P Q} ind-P ind-Q → cong₂ _Poly+_ ind-P ind-Q - - - Poly+-rinv : (P : Poly A' n ) → P Poly+ (Poly-inv P) ≡ 0P - Poly+-rinv = Poly-Ind-Prop.f A' n (λ P → (P Poly+ Poly-inv P) ≡ 0P) (λ _ → trunc _ _) - (Poly+-Rid 0P) - (λ v a → (base-Poly+ v a (- a)) ∙ cong (base v) (+Rinv a) ∙ base-0P v) - λ {P Q} ind-P ind-Q → ((P Poly+ Q) Poly+ ((Poly-inv P) Poly+ (Poly-inv Q))) - ≡⟨ Poly-com-adv P Q (Poly-inv P) (Poly-inv Q) ⟩ - ((P Poly+ Poly-inv P) Poly+ (Q Poly+ Poly-inv Q)) - ≡⟨ cong₂ _Poly+_ ind-P ind-Q ⟩ - (0P Poly+ 0P) - ≡⟨ Poly+-Rid 0P ⟩ - 0P ∎ - - Poly+-linv : (P : Poly A' n) → (Poly-inv P) Poly+ P ≡ 0P - Poly+-linv = Poly-Ind-Prop.f A' n _ (λ _ → trunc _ _) - (Poly+-Rid 0P) - (λ v a → (base-Poly+ v (- a) a) ∙ cong (base v) (snd (+Inv a)) ∙ base-0P v) - λ {U V} ind-U ind-V → Poly-com-adv (Poly-inv U) (Poly-inv V) U V ∙ cong₂ _Poly+_ ind-U ind-V ∙ Poly+-Rid 0P + poly-com-adv : (P Q R S : Poly A' n) → ((P poly+ Q) poly+ (R poly+ S) ≡ (P poly+ R) poly+ (Q poly+ S)) + poly-com-adv P Q R S = ((P poly+ Q) poly+ (R poly+ S) ≡⟨ poly+Assoc (P poly+ Q) R S ⟩ + (((P poly+ Q) poly+ R) poly+ S) ≡⟨ cong (λ X → X poly+ S) (sym (poly+Assoc P Q R)) ⟩ + ((P poly+ (Q poly+ R)) poly+ S) ≡⟨ cong (λ X → (P poly+ X) poly+ S) (poly+Comm Q R) ⟩ + ((P poly+ (R poly+ Q)) poly+ S) ≡⟨ cong (λ X → X poly+ S) (poly+Assoc P R Q) ⟩ + (((P poly+ R) poly+ Q) poly+ S) ≡⟨ sym (poly+Assoc (P poly+ R) Q S) ⟩ + ((P poly+ R) poly+ (Q poly+ S)) ∎) + + + poly+IdL : (P : Poly A' n) → 0P poly+ P ≡ P + poly+IdL P = (poly+Comm 0P P) ∙ (poly+IdR P) + + + polyInv : Poly A' n → Poly A' n + polyInv = Poly-Rec-Set.f A' n (Poly A' n) trunc + 0P + (λ v a → base v (- a)) + (λ PS RS → PS poly+ RS) + poly+Assoc + poly+IdR + poly+Comm + (λ v → base v (- 0r) ≡⟨ cong (base v) 0Selfinverse ⟩ base v 0r ≡⟨ base-0P v ⟩ 0P ∎) + λ v a b → (base-poly+ v (- a) (- b)) ∙ (cong (base v) (-Dist a b)) + + + polyInvol : (P : Poly A' n) → polyInv (polyInv P) ≡ P + polyInvol = Poly-Ind-Prop.f A' n (λ P → polyInv (polyInv P) ≡ P) (λ _ → trunc _ _) + refl + (λ v a → cong (base v) (-Idempotent a)) + λ {P Q} ind-P ind-Q → cong₂ _poly+_ ind-P ind-Q + + + poly+InvR : (P : Poly A' n ) → P poly+ (polyInv P) ≡ 0P + poly+InvR = Poly-Ind-Prop.f A' n (λ P → (P poly+ polyInv P) ≡ 0P) (λ _ → trunc _ _) + (poly+IdR 0P) + (λ v a → (base-poly+ v a (- a)) ∙ cong (base v) (+Rinv a) ∙ base-0P v) + λ {P Q} ind-P ind-Q → ((P poly+ Q) poly+ ((polyInv P) poly+ (polyInv Q))) + ≡⟨ poly-com-adv P Q (polyInv P) (polyInv Q) ⟩ + ((P poly+ polyInv P) poly+ (Q poly+ polyInv Q)) + ≡⟨ cong₂ _poly+_ ind-P ind-Q ⟩ + (0P poly+ 0P) + ≡⟨ poly+IdR 0P ⟩ + 0P ∎ + + poly+InvL : (P : Poly A' n) → (polyInv P) poly+ P ≡ 0P + poly+InvL = Poly-Ind-Prop.f A' n _ (λ _ → trunc _ _) + (poly+IdR 0P) + (λ v a → (base-poly+ v (- a) a) ∙ cong (base v) (snd (+Inv a)) ∙ base-0P v) + λ {U V} ind-U ind-V → poly-com-adv (polyInv U) (polyInv V) U V + ∙ cong₂ _poly+_ ind-U ind-V + ∙ poly+IdR 0P ----------------------------------------------------------------------------- - _Poly*_ : Poly A' n → Poly A' n → Poly A' n - _Poly*_ = -- Induction Left Argument + _poly*_ : Poly A' n → Poly A' n → Poly A' n + _poly*_ = -- Induction Left Argument Poly-Rec-Set.f A' n (Poly A' n → Poly A' n) (λ f g p q i j Q → trunc (f Q) (g Q) (λ X → p X Q) (λ X → q X Q) i j ) (λ Q → 0P) @@ -88,96 +91,98 @@ module Nth-Poly-structure (A' : CommRing ℓ) (n : ℕ) where Poly-Rec-Set.f A' n (Poly A' n) trunc 0P (λ v' a' → base (v +n-vec v') (a · a')) - _Poly+_ - Poly+-assoc - Poly+-Rid - Poly+-comm + _poly+_ + poly+Assoc + poly+IdR + poly+Comm (λ v' → (cong (base (v +n-vec v')) (0RightAnnihilates a)) ∙ (base-0P (v +n-vec v'))) - λ v' b c → (base-Poly+ (v +n-vec v') (a · b) (a · c)) ∙ (cong (base (v +n-vec v')) (sym (·Rdist+ a b c)))) + λ v' b c → (base-poly+ (v +n-vec v') (a · b) (a · c)) + ∙ (cong (base (v +n-vec v')) (sym (·Rdist+ a b c)))) -- End Right induction - (λ PS QS Q → (PS Q) Poly+ (QS Q) ) - (λ PS QS RS i Q → Poly+-assoc (PS Q) (QS Q) (RS Q) i) - (λ PS i Q → Poly+-Rid (PS Q) i) - (λ PS QS i Q → Poly+-comm (PS Q) (QS Q) i) + (λ PS QS Q → (PS Q) poly+ (QS Q) ) + (λ PS QS RS i Q → poly+Assoc (PS Q) (QS Q) (RS Q) i) + (λ PS i Q → poly+IdR (PS Q) i) + (λ PS QS i Q → poly+Comm (PS Q) (QS Q) i) (λ v → funExt ( -- Induction Right Argument Poly-Ind-Prop.f A' n _ (λ _ → trunc _ _) refl (λ v' a' → (cong (base (v +n-vec v')) (0LeftAnnihilates a')) ∙ (base-0P (v +n-vec v'))) - λ {P Q} ind-P ind-Q → (cong₂ _Poly+_ ind-P ind-Q) ∙ (Poly+-Rid 0P) )) + λ {P Q} ind-P ind-Q → (cong₂ _poly+_ ind-P ind-Q) ∙ (poly+IdR 0P) )) -- End Right Induction λ v a b → funExt ( -- Induction Right Argument Poly-Ind-Prop.f A' n _ (λ _ → trunc _ _) - (Poly+-Rid 0P) - (λ v' c → (base-Poly+ (v +n-vec v') (a · c) (b · c)) ∙ (cong (base (v +n-vec v')) (sym (·Ldist+ a b c)))) - λ {P Q} ind-P ind-Q → (Poly-com-adv _ _ _ _) ∙ (cong₂ _Poly+_ ind-P ind-Q)) + (poly+IdR 0P) + (λ v' c → (base-poly+ (v +n-vec v') (a · c) (b · c)) + ∙ (cong (base (v +n-vec v')) (sym (·Ldist+ a b c)))) + λ {P Q} ind-P ind-Q → (poly-com-adv _ _ _ _) ∙ (cong₂ _poly+_ ind-P ind-Q)) -- End Right Induction -- End Left Induction - Poly*-assoc : (P Q R : Poly A' n) → P Poly* (Q Poly* R) ≡ (P Poly* Q) Poly* R - Poly*-assoc = Poly-Ind-Prop.f A' n _ - (λ P p q i Q R j → trunc (P Poly* (Q Poly* R)) ((P Poly* Q) Poly* R) (p Q R) (q Q R) i j) + poly*Assoc : (P Q R : Poly A' n) → P poly* (Q poly* R) ≡ (P poly* Q) poly* R + poly*Assoc = Poly-Ind-Prop.f A' n _ + (λ P p q i Q R j → trunc (P poly* (Q poly* R)) ((P poly* Q) poly* R) (p Q R) (q Q R) i j) (λ _ _ → refl) (λ v a → Poly-Ind-Prop.f A' n _ - (λ Q p q i R j → trunc (base v a Poly* (Q Poly* R)) ((base v a Poly* Q) Poly* R) (p R) (q R) i j) + (λ Q p q i R j → trunc (base v a poly* (Q poly* R)) ((base v a poly* Q) poly* R) (p R) (q R) i j) (λ _ → refl) (λ v' a' → Poly-Ind-Prop.f A' n _ (λ _ → trunc _ _) refl (λ v'' a'' → cong₂ base (+n-vec-assoc v v' v'') (·Assoc a a' a'')) - (λ {U V} ind-U ind-V → cong₂ _Poly+_ ind-U ind-V)) - (λ {U V} ind-U ind-V R → cong₂ _Poly+_ (ind-U R) (ind-V R))) - λ {U V} ind-U ind-V Q R → cong₂ _Poly+_ (ind-U Q R) (ind-V Q R) + (λ {U V} ind-U ind-V → cong₂ _poly+_ ind-U ind-V)) + (λ {U V} ind-U ind-V R → cong₂ _poly+_ (ind-U R) (ind-V R))) + λ {U V} ind-U ind-V Q R → cong₂ _poly+_ (ind-U Q R) (ind-V Q R) - 0PLeftAnnihilatesPoly : (P : Poly A' n) → 0P Poly* P ≡ 0P - 0PLeftAnnihilatesPoly P = refl + poly*AnnihilL : (P : Poly A' n) → 0P poly* P ≡ 0P + poly*AnnihilL P = refl - 0PRightAnnihilatesPoly : (P : Poly A' n) → P Poly* 0P ≡ 0P - 0PRightAnnihilatesPoly = Poly-Ind-Prop.f A' n (λ P → (P Poly* 0P) ≡ 0P) (λ _ → trunc _ _) + poly*AnnihilR : (P : Poly A' n) → P poly* 0P ≡ 0P + poly*AnnihilR = Poly-Ind-Prop.f A' n (λ P → (P poly* 0P) ≡ 0P) (λ _ → trunc _ _) refl (λ _ _ → refl) - λ {P Q} ind-P ind-Q → (cong₂ _Poly+_ ind-P ind-Q) ∙ (Poly+-Rid 0P) + λ {P Q} ind-P ind-Q → (cong₂ _poly+_ ind-P ind-Q) ∙ (poly+IdR 0P) 1P : Poly A' n 1P = base (replicate zero) 1r - Poly*-Rid : (P : Poly A' n) → P Poly* 1P ≡ P - Poly*-Rid = Poly-Ind-Prop.f A' n (λ P → (P Poly* 1P) ≡ P) (λ _ → trunc _ _) + poly*IdR : (P : Poly A' n) → P poly* 1P ≡ P + poly*IdR = Poly-Ind-Prop.f A' n (λ P → (P poly* 1P) ≡ P) (λ _ → trunc _ _) refl (λ v a → cong₂ base (+n-vec-rid v) (·Rid a)) - (λ {P Q} ind-P ind-Q → cong₂ _Poly+_ ind-P ind-Q) + (λ {P Q} ind-P ind-Q → cong₂ _poly+_ ind-P ind-Q) - Poly*-Lid : (P : Poly A' n) → 1P Poly* P ≡ P - Poly*-Lid = Poly-Ind-Prop.f A' n (λ P → (1P Poly* P) ≡ P) (λ _ → trunc _ _) + poly*IdL : (P : Poly A' n) → 1P poly* P ≡ P + poly*IdL = Poly-Ind-Prop.f A' n (λ P → (1P poly* P) ≡ P) (λ _ → trunc _ _) refl (λ v a → cong₂ base (+n-vec-lid v) (·Lid a)) - λ {P Q} ind-P ind-Q → cong₂ _Poly+_ ind-P ind-Q + λ {P Q} ind-P ind-Q → cong₂ _poly+_ ind-P ind-Q - Poly*-Rdist : (P Q R : Poly A' n) → P Poly* (Q Poly+ R) ≡ (P Poly* Q) Poly+ (P Poly* R) - Poly*-Rdist = Poly-Ind-Prop.f A' n _ - (λ P p q i Q R j → trunc (P Poly* (Q Poly+ R)) ((P Poly* Q) Poly+ (P Poly* R)) (p Q R) (q Q R) i j) - (λ _ _ → sym (Poly+-Rid 0P)) - (λ v a → λ Q R → refl) - λ {U V} ind-U ind-V Q R → (cong₂ _Poly+_ (ind-U Q R) (ind-V Q R)) ∙ Poly-com-adv (U Poly* Q) (U Poly* R) (V Poly* Q) (V Poly* R) + poly*DistR : (P Q R : Poly A' n) → P poly* (Q poly+ R) ≡ (P poly* Q) poly+ (P poly* R) + poly*DistR = Poly-Ind-Prop.f A' n _ + (λ P p q i Q R j → trunc (P poly* (Q poly+ R)) ((P poly* Q) poly+ (P poly* R)) (p Q R) (q Q R) i j) + (λ _ _ → sym (poly+IdR 0P)) + (λ v a → λ Q R → refl) + λ {U V} ind-U ind-V Q R → (cong₂ _poly+_ (ind-U Q R) (ind-V Q R)) ∙ poly-com-adv (U poly* Q) (U poly* R) (V poly* Q) (V poly* R) - Poly*-Ldist : (P Q R : Poly A' n) → (P Poly+ Q) Poly* R ≡ (P Poly* R) Poly+ (Q Poly* R) - Poly*-Ldist P Q R = refl + poly*DistL : (P Q R : Poly A' n) → (P poly+ Q) poly* R ≡ (P poly* R) poly+ (Q poly* R) + poly*DistL P Q R = refl - Poly*-comm : (P Q : Poly A' n) → P Poly* Q ≡ Q Poly* P - Poly*-comm = Poly-Ind-Prop.f A' n _ - (λ P p q i Q j → trunc (P Poly* Q) (Q Poly* P) (p Q) (q Q) i j) - (λ Q → sym (0PRightAnnihilatesPoly Q)) - (λ v a → Poly-Ind-Prop.f A' n _ (λ _ → trunc _ _) - refl - (λ v' a' → cong₂ base (+n-vec-comm v v') (·Comm a a')) - (λ {U V} ind-U ind-V → cong₂ _Poly+_ ind-U ind-V)) - λ {U V} ind-U ind-V Q → ((cong₂ _Poly+_ (ind-U Q) (ind-V Q)) ∙ sym (Poly*-Rdist Q U V)) + poly*Comm : (P Q : Poly A' n) → P poly* Q ≡ Q poly* P + poly*Comm = Poly-Ind-Prop.f A' n _ + (λ P p q i Q j → trunc (P poly* Q) (Q poly* P) (p Q) (q Q) i j) + (λ Q → sym (poly*AnnihilR Q)) + (λ v a → Poly-Ind-Prop.f A' n _ (λ _ → trunc _ _) + refl + (λ v' a' → cong₂ base (+n-vec-comm v v') (·Comm a a')) + (λ {U V} ind-U ind-V → cong₂ _poly+_ ind-U ind-V)) + λ {U V} ind-U ind-V Q → ((cong₂ _poly+_ (ind-U Q) (ind-V Q)) ∙ sym (poly*DistR Q U V)) diff --git a/Cubical/Algebra/Polynomials/Univariate/Base.agda b/Cubical/Algebra/Polynomials/Univariate/Base.agda index 01d72f0f70..ce3ab6b393 100644 --- a/Cubical/Algebra/Polynomials/Univariate/Base.agda +++ b/Cubical/Algebra/Polynomials/Univariate/Base.agda @@ -100,7 +100,7 @@ module PolyMod (R' : CommRing ℓ) where isSetPolyFun : isSet PolyFun - isSetPolyFun = isSetΣSndProp (isSetΠ (λ x → isSetCommRing R')) λ f x y → squash x y + isSetPolyFun = isSetΣSndProp (isSetΠ (λ x → isSetCommRing R')) λ f x y → squash₁ x y --construction of the function that represents the polynomial @@ -116,14 +116,14 @@ module PolyMod (R' : CommRing ℓ) where Poly→Prf : (p : Poly) → ∃[ n ∈ ℕ ] ((m : ℕ) → n ≤ m → (Poly→Fun p m ≡ 0r)) Poly→Prf = ElimProp.f (λ p → ∃[ n ∈ ℕ ] ((m : ℕ) → n ≤ m → (Poly→Fun p m ≡ 0r))) - ∣ 0 , (λ m ineq → refl) ∣ + ∣ 0 , (λ m ineq → refl) ∣₁ (λ r p → map ( λ (n , ineq) → (suc n) , λ { zero h → ⊥rec (znots (sym (≤0→≡0 h))) ; (suc m) h → ineq m (pred-≤-pred h) } ) ) - squash + squash₁ Poly→PolyFun : Poly → PolyFun Poly→PolyFun p = (Poly→Fun p) , (Poly→Prf p) diff --git a/Cubical/Algebra/Ring/Properties.agda b/Cubical/Algebra/Ring/Properties.agda index a66b40cf16..b785dd004d 100644 --- a/Cubical/Algebra/Ring/Properties.agda +++ b/Cubical/Algebra/Ring/Properties.agda @@ -325,7 +325,7 @@ recPT→Ring : {A : Type ℓ'} (𝓕 : A → Ring ℓ) → (σ : ∀ x y → RingEquiv (𝓕 x) (𝓕 y)) → (∀ x y z → σ x z ≡ compRingEquiv (σ x y) (σ y z)) ------------------------------------------------------ - → ∥ A ∥ → Ring ℓ + → ∥ A ∥₁ → Ring ℓ recPT→Ring 𝓕 σ compCoh = rec→Gpd isGroupoidRing 𝓕 (3-ConstantCompChar 𝓕 (λ x y → uaRing (σ x y)) λ x y z → sym ( cong uaRing (compCoh x y z) diff --git a/Cubical/Algebra/ZariskiLattice/StructureSheaf.agda b/Cubical/Algebra/ZariskiLattice/StructureSheaf.agda index 93f8df6adc..682e47da64 100644 --- a/Cubical/Algebra/ZariskiLattice/StructureSheaf.agda +++ b/Cubical/Algebra/ZariskiLattice/StructureSheaf.agda @@ -113,14 +113,14 @@ module _ (R' : CommRing ℓ) where BO = Σ[ 𝔞 ∈ ZL ] (𝔞 ∈ₚ BasicOpens) basicOpensAreBasis : IsBasis ZariskiLattice BasicOpens - contains1 basicOpensAreBasis = ∣ 1r , isZarMapD .pres1 ∣ + contains1 basicOpensAreBasis = ∣ 1r , isZarMapD .pres1 ∣₁ ∧lClosed basicOpensAreBasis 𝔞 𝔟 = map2 λ (f , Df≡𝔞) (g , Dg≡𝔟) → (f · g) , isZarMapD .·≡∧ f g ∙ cong₂ (_∧z_) Df≡𝔞 Dg≡𝔟 ⋁Basis basicOpensAreBasis = elimProp (λ _ → isPropPropTrunc) Σhelper where Σhelper : (a : Σ[ n ∈ ℕ ] FinVec R n) → ∃[ n ∈ ℕ ] Σ[ α ∈ FinVec ZL n ] (∀ i → α i ∈ₚ BasicOpens) × (⋁ α ≡ [ a ]) - Σhelper (n , α) = ∣ n , (D ∘ α) , (λ i → ∣ α i , refl ∣) , path ∣ + Σhelper (n , α) = ∣ n , (D ∘ α) , (λ i → ∣ α i , refl ∣₁) , path ∣₁ where path : ⋁ (D ∘ α) ≡ [ n , α ] path = funExt⁻ (cong fst ZLUniversalPropCorollary) _ @@ -189,7 +189,7 @@ module _ (R' : CommRing ℓ) where _ = BasisStructurePShf canonical0∈BO : 0z ∈ₚ BasicOpens - canonical0∈BO = ∣ 0r , isZarMapD .pres0 ∣ + canonical0∈BO = ∣ 0r , isZarMapD .pres0 ∣₁ canonical0∈BO≡0∈BO : canonical0∈BO ≡ 0∈BO canonical0∈BO≡0∈BO = BasicOpens 0z .snd _ _ @@ -236,29 +236,29 @@ module _ (R' : CommRing ℓ) where -- write everything explicitly so things can type-check thePShfCospan : (a : Σ[ f ∈ R ] D f ≡ 𝔞) (b : Σ[ g ∈ R ] D g ≡ 𝔟) → Cospan (CommAlgebrasCategory R') - Cospan.l (thePShfCospan (f , Df≡𝔞) (g , Dg≡𝔟)) = BasisStructurePShf .Functor.F-ob (𝔟 , ∣ g , Dg≡𝔟 ∣) + Cospan.l (thePShfCospan (f , Df≡𝔞) (g , Dg≡𝔟)) = BasisStructurePShf .Functor.F-ob (𝔟 , ∣ g , Dg≡𝔟 ∣₁) Cospan.m (thePShfCospan (f , Df≡𝔞) (g , Dg≡𝔟)) = BasisStructurePShf .Functor.F-ob - (𝔞 ∧z 𝔟 , basicOpensAreBasis .∧lClosed 𝔞 𝔟 ∣ f , Df≡𝔞 ∣ ∣ g , Dg≡𝔟 ∣) - Cospan.r (thePShfCospan (f , Df≡𝔞) (g , Dg≡𝔟)) = BasisStructurePShf .Functor.F-ob (𝔞 , ∣ f , Df≡𝔞 ∣) + (𝔞 ∧z 𝔟 , basicOpensAreBasis .∧lClosed 𝔞 𝔟 ∣ f , Df≡𝔞 ∣₁ ∣ g , Dg≡𝔟 ∣₁) + Cospan.r (thePShfCospan (f , Df≡𝔞) (g , Dg≡𝔟)) = BasisStructurePShf .Functor.F-ob (𝔞 , ∣ f , Df≡𝔞 ∣₁) Cospan.s₁ (thePShfCospan (f , Df≡𝔞) (g , Dg≡𝔟)) = BasisStructurePShf .Functor.F-hom - {x = (𝔟 , ∣ g , Dg≡𝔟 ∣)} - {y = (𝔞 ∧z 𝔟 , basicOpensAreBasis .∧lClosed 𝔞 𝔟 ∣ f , Df≡𝔞 ∣ ∣ g , Dg≡𝔟 ∣)} + {x = (𝔟 , ∣ g , Dg≡𝔟 ∣₁)} + {y = (𝔞 ∧z 𝔟 , basicOpensAreBasis .∧lClosed 𝔞 𝔟 ∣ f , Df≡𝔞 ∣₁ ∣ g , Dg≡𝔟 ∣₁)} (hom-∧₂ ZariskiLattice (CommAlgebrasCategory R' {ℓ' = ℓ}) (TerminalCommAlgebra R') 𝔞 𝔟) Cospan.s₂ (thePShfCospan (f , Df≡𝔞) (g , Dg≡𝔟)) = BasisStructurePShf .Functor.F-hom - {x = (𝔞 , ∣ f , Df≡𝔞 ∣)} - {y = (𝔞 ∧z 𝔟 , basicOpensAreBasis .∧lClosed 𝔞 𝔟 ∣ f , Df≡𝔞 ∣ ∣ g , Dg≡𝔟 ∣)} + {x = (𝔞 , ∣ f , Df≡𝔞 ∣₁)} + {y = (𝔞 ∧z 𝔟 , basicOpensAreBasis .∧lClosed 𝔞 𝔟 ∣ f , Df≡𝔞 ∣₁ ∣ g , Dg≡𝔟 ∣₁)} (hom-∧₁ ZariskiLattice (CommAlgebrasCategory R' {ℓ' = ℓ}) (TerminalCommAlgebra R') 𝔞 𝔟) Σhelper : (a : Σ[ f ∈ R ] D f ≡ 𝔞) (b : Σ[ g ∈ R ] D g ≡ 𝔟) (c : Σ[ h ∈ R ] D h ≡ 𝔞 ∨z 𝔟) → isPullback (CommAlgebrasCategory R') (thePShfCospan a b) _ _ - (BFsq (𝔞 , ∣ a ∣) (𝔟 , ∣ b ∣) ∣ c ∣ BasisStructurePShf) + (BFsq (𝔞 , ∣ a ∣₁) (𝔟 , ∣ b ∣₁) ∣ c ∣₁ BasisStructurePShf) Σhelper (f , Df≡𝔞) (g , Dg≡𝔟) (h , Dh≡𝔞∨𝔟) = toSheaf.lemma - (𝔞 ∨z 𝔟 , ∣ h , Dh≡𝔞∨𝔟 ∣) - (𝔞 , ∣ f , Df≡𝔞 ∣) - (𝔟 , ∣ g , Dg≡𝔟 ∣) - (𝔞 ∧z 𝔟 , basicOpensAreBasis .∧lClosed 𝔞 𝔟 ∣ f , Df≡𝔞 ∣ ∣ g , Dg≡𝔟 ∣) - (Bsq (𝔞 , ∣ f , Df≡𝔞 ∣) (𝔟 , ∣ g , Dg≡𝔟 ∣) ∣ h , Dh≡𝔞∨𝔟 ∣) + (𝔞 ∨z 𝔟 , ∣ h , Dh≡𝔞∨𝔟 ∣₁) + (𝔞 , ∣ f , Df≡𝔞 ∣₁) + (𝔟 , ∣ g , Dg≡𝔟 ∣₁) + (𝔞 ∧z 𝔟 , basicOpensAreBasis .∧lClosed 𝔞 𝔟 ∣ f , Df≡𝔞 ∣₁ ∣ g , Dg≡𝔟 ∣₁) + (Bsq (𝔞 , ∣ f , Df≡𝔞 ∣₁) (𝔟 , ∣ g , Dg≡𝔟 ∣₁) ∣ h , Dh≡𝔞∨𝔟 ∣₁) theAlgebraCospan theAlgebraPullback refl gPath fPath fgPath where open Exponentiation R' @@ -310,12 +310,12 @@ module _ (R' : CommRing ℓ) where helper3 (α , p) = β , path where β : FinVec R[1/ h ] 2 - β zero = [ α zero , h ^ n , ∣ n , refl ∣ ] - β (suc zero) = [ α (suc zero) , h ^ n , ∣ n , refl ∣ ] + β zero = [ α zero , h ^ n , ∣ n , refl ∣₁ ] + β (suc zero) = [ α (suc zero) , h ^ n , ∣ n , refl ∣₁ ] path : 1r ≡ linearCombination R[1/ h ]AsCommRing β λ { zero → f /1 ; (suc zero) → g /1 } - path = eq/ _ _ ((1r , ∣ 0 , refl ∣) , bigPath) + path = eq/ _ _ ((1r , ∣ 0 , refl ∣₁) , bigPath) ∙ cong (β zero · (f /1) +_) (sym (+Rid (β (suc zero) · (g /1)))) where useSolver1 : ∀ hn → 1r · 1r · ((hn · 1r) · (hn · 1r)) ≡ hn · hn @@ -371,7 +371,7 @@ module _ (R' : CommRing ℓ) where open IsRingHom /1/1AsCommRingHomFG : CommRingHom R' R[1/h][1/fg]AsCommRing - fst /1/1AsCommRingHomFG r = [ [ r , 1r , ∣ 0 , refl ∣ ] , 1r , ∣ 0 , refl ∣ ] + fst /1/1AsCommRingHomFG r = [ [ r , 1r , ∣ 0 , refl ∣₁ ] , 1r , ∣ 0 , refl ∣₁ ] pres0 (snd /1/1AsCommRingHomFG) = refl pres1 (snd /1/1AsCommRingHomFG) = refl pres+ (snd /1/1AsCommRingHomFG) x y = cong [_] (≡-× (cong [_] (≡-× diff --git a/Cubical/Axiom/Omniscience.agda b/Cubical/Axiom/Omniscience.agda index 65af8906fd..b12801aa93 100644 --- a/Cubical/Axiom/Omniscience.agda +++ b/Cubical/Axiom/Omniscience.agda @@ -30,10 +30,10 @@ LLPO : Type ℓ → Type ℓ LLPO A = ∀(P Q : A → 𝟚) → (∀ x y → ⟨ P x ⟩ → ⟨ Q y ⟩ → ⊥) - → ∥ (∀ x → ¬ ⟨ P x ⟩) ⊎ (∀ y → ¬ ⟨ Q y ⟩) ∥ + → ∥ (∀ x → ¬ ⟨ P x ⟩) ⊎ (∀ y → ¬ ⟨ Q y ⟩) ∥₁ isPropLLPO : isProp (LLPO A) -isPropLLPO = isPropΠ3 λ _ _ _ → squash +isPropLLPO = isPropΠ3 λ _ _ _ → squash₁ -- As above, but without ensuring propositionality LLPO∞ : Type ℓ → Type ℓ @@ -43,7 +43,7 @@ LLPO∞ A → (∀ x → ¬ ⟨ P x ⟩) ⊎ (∀ y → ¬ ⟨ Q y ⟩) LLPO∞→LLPO : LLPO∞ A → LLPO A -LLPO∞→LLPO llpo' P Q ¬both = ∣ llpo' P Q ¬both ∣ +LLPO∞→LLPO llpo' P Q ¬both = ∣ llpo' P Q ¬both ∣₁ -- Weak limited principle of omniscience -- @@ -92,7 +92,7 @@ WLPO→LLPO∞ {A = A} womn P Q ¬both with womn P -- -- Either a decidable predicate never holds, or it does LPO : Type ℓ → Type ℓ -LPO A = ∀(P : A → 𝟚) → (∀ x → ¬ ⟨ P x ⟩) ⊎ ∥ Σ[ x ∈ A ] ⟨ P x ⟩ ∥ +LPO A = ∀(P : A → 𝟚) → (∀ x → ¬ ⟨ P x ⟩) ⊎ ∥ Σ[ x ∈ A ] ⟨ P x ⟩ ∥₁ LPO→WLPO : LPO A → WLPO A LPO→WLPO omn P with omn P @@ -104,4 +104,4 @@ LPO∞ : Type ℓ → Type ℓ LPO∞ A = ∀(P : A → 𝟚) → (∀ x → ¬ ⟨ P x ⟩) ⊎ (Σ[ x ∈ A ] ⟨ P x ⟩) LPO∞→LPO : LPO∞ A → LPO A -LPO∞→LPO omn P = Sum.map (idfun _) ∣_∣ (omn P) +LPO∞→LPO omn P = Sum.map (idfun _) ∣_∣₁ (omn P) diff --git a/Cubical/Categories/Equivalence/Properties.agda b/Cubical/Categories/Equivalence/Properties.agda index 6e05df1276..c1e73b8deb 100644 --- a/Cubical/Categories/Equivalence/Properties.agda +++ b/Cubical/Categories/Equivalence/Properties.agda @@ -52,7 +52,7 @@ module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where isEquiv→Full {F} eq@record { invFunc = G ; η = η ; ε = _ } - c c' g = ∣ h , isEquiv→Faithful (symEquiv eq) _ _ _ _ GFh≡Gg ∣ -- apply faithfulness of 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) diff --git a/Cubical/Categories/Instances/CommAlgebras.agda b/Cubical/Categories/Instances/CommAlgebras.agda index 8bb81ec3ec..d1c177ab3b 100644 --- a/Cubical/Categories/Instances/CommAlgebras.agda +++ b/Cubical/Categories/Instances/CommAlgebras.agda @@ -272,19 +272,19 @@ module PreSheafFromUniversalProp (C : Category ℓ ℓ') (P : ob C → Type ℓ) private ∥P∥ : ℙ (ob C) - ∥P∥ x = ∥ P x ∥ , isPropPropTrunc + ∥P∥ x = ∥ P x ∥₁ , isPropPropTrunc ΣC∥P∥Cat = ΣPropCat C ∥P∥ CommAlgCat = CommAlgebrasCategory {ℓ = ℓ''} R {ℓ' = ℓ''} 𝓕UniqueEquiv : (x : ob C) (p q : P x) → isContr (CommAlgebraEquiv (𝓕 (x , p)) (𝓕 (x , q))) 𝓕UniqueEquiv x = contrCommAlgebraHom→contrCommAlgebraEquiv (curry 𝓕 x) λ p q → uniqueHom _ _ (id C) - theMap : (x : ob C) → ∥ P x ∥ → CommAlgebra R ℓ'' + theMap : (x : ob C) → ∥ P x ∥₁ → CommAlgebra R ℓ'' theMap x = recPT→CommAlgebra (curry 𝓕 x) (λ p q → 𝓕UniqueEquiv x p q .fst) λ p q r → 𝓕UniqueEquiv x p r .snd _ theAction : (x y : ob C) → C [ x , y ] - → (p : ∥ P x ∥) (q : ∥ P y ∥) → isContr (CommAlgebraHom (theMap y q) (theMap x p)) + → (p : ∥ P x ∥₁) (q : ∥ P y ∥₁) → isContr (CommAlgebraHom (theMap y q) (theMap x p)) theAction _ _ f = elim2 (λ _ _ → isPropIsContr) λ _ _ → uniqueHom _ _ f open Functor diff --git a/Cubical/Categories/Limits/BinCoproduct.agda b/Cubical/Categories/Limits/BinCoproduct.agda index 352f0b57a9..68e65661fa 100644 --- a/Cubical/Categories/Limits/BinCoproduct.agda +++ b/Cubical/Categories/Limits/BinCoproduct.agda @@ -40,4 +40,4 @@ module _ (C : Category ℓ ℓ') where BinCoproducts = (x y : ob) → BinCoproduct x y hasBinCoproducts : Type (ℓ-max ℓ ℓ') - hasBinCoproducts = ∥ BinCoproducts ∥ + hasBinCoproducts = ∥ BinCoproducts ∥₁ diff --git a/Cubical/Categories/Limits/BinProduct.agda b/Cubical/Categories/Limits/BinProduct.agda index 02cc35ba44..b76f8b5d89 100644 --- a/Cubical/Categories/Limits/BinProduct.agda +++ b/Cubical/Categories/Limits/BinProduct.agda @@ -40,4 +40,4 @@ module _ (C : Category ℓ ℓ') where BinProducts = (x y : ob) → BinProduct x y hasBinProducts : Type (ℓ-max ℓ ℓ') - hasBinProducts = ∥ BinProducts ∥ + hasBinProducts = ∥ BinProducts ∥₁ diff --git a/Cubical/Categories/Limits/Initial.agda b/Cubical/Categories/Limits/Initial.agda index b61c4f02c2..a5c06b6a42 100644 --- a/Cubical/Categories/Limits/Initial.agda +++ b/Cubical/Categories/Limits/Initial.agda @@ -41,7 +41,7 @@ module _ (C : Category ℓ ℓ') where initialEndoIsId T f = isContr→isProp (T .snd (initialOb T)) f id hasInitial : Type (ℓ-max ℓ ℓ') - hasInitial = ∥ Initial ∥ + hasInitial = ∥ Initial ∥₁ -- Initiality of an object is a proposition. isPropIsInitial : (x : ob) → isProp (isInitial x) diff --git a/Cubical/Categories/Limits/Limits.agda b/Cubical/Categories/Limits/Limits.agda index 8ba2e7db0b..6572fdff53 100644 --- a/Cubical/Categories/Limits/Limits.agda +++ b/Cubical/Categories/Limits/Limits.agda @@ -123,7 +123,7 @@ Limits : {ℓJ ℓJ' ℓC ℓC' : Level} → Category ℓC ℓC' → Type _ Limits {ℓJ} {ℓJ'} C = (J : Category ℓJ ℓJ') → (D : Functor J C) → LimCone D hasLimits : {ℓJ ℓJ' ℓC ℓC' : Level} → Category ℓC ℓC' → Type _ -hasLimits {ℓJ} {ℓJ'} C = (J : Category ℓJ ℓJ') → (D : Functor J C) → ∥ LimCone D ∥ +hasLimits {ℓJ} {ℓJ'} C = (J : Category ℓJ ℓJ') → (D : Functor J C) → ∥ LimCone D ∥₁ -- Limits of a specific shape J in a category C LimitsOfShape : {ℓJ ℓJ' ℓC ℓC' : Level} → Category ℓJ ℓJ' → Category ℓC ℓC' → Type _ diff --git a/Cubical/Categories/Limits/Pullback.agda b/Cubical/Categories/Limits/Pullback.agda index ae1e7a333f..a872212b80 100644 --- a/Cubical/Categories/Limits/Pullback.agda +++ b/Cubical/Categories/Limits/Pullback.agda @@ -89,7 +89,7 @@ module _ (C : Category ℓ ℓ') where Pullbacks = (cspn : Cospan) → Pullback cspn hasPullbacks : Type (ℓ-max ℓ ℓ') - hasPullbacks = ∥ Pullbacks ∥ + hasPullbacks = ∥ Pullbacks ∥₁ -- Pullbacks from limits diff --git a/Cubical/Categories/Limits/Terminal.agda b/Cubical/Categories/Limits/Terminal.agda index e81d9b6d04..b9a472ddaa 100644 --- a/Cubical/Categories/Limits/Terminal.agda +++ b/Cubical/Categories/Limits/Terminal.agda @@ -37,7 +37,7 @@ module _ (C : Category ℓ ℓ') where terminalEndoIsId T f = isContr→isProp (T .snd (terminalOb T)) f id hasTerminal : Type (ℓ-max ℓ ℓ') - hasTerminal = ∥ Terminal ∥ + hasTerminal = ∥ Terminal ∥₁ -- Terminality of an object is a proposition. isPropIsTerminal : (x : ob) → isProp (isTerminal x) diff --git a/Cubical/Categories/Yoneda.agda b/Cubical/Categories/Yoneda.agda index 83f03f0419..150db9c4db 100644 --- a/Cubical/Categories/Yoneda.agda +++ b/Cubical/Categories/Yoneda.agda @@ -170,7 +170,7 @@ module _ {C : Category ℓ ℓ} where isFullYO : isFull YO - isFullYO x y F[f] = ∣ yo-yo-yo _ F[f] , yoIso {x} (yo y) .Iso.leftInv F[f] ∣ + 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 = diff --git a/Cubical/Codata/M/Bisimilarity.agda b/Cubical/Codata/M/Bisimilarity.agda index b88086d6e1..602d2f8989 100644 --- a/Cubical/Codata/M/Bisimilarity.agda +++ b/Cubical/Codata/M/Bisimilarity.agda @@ -1,10 +1,16 @@ {-# OPTIONS --postfix-projections --guardedness #-} module Cubical.Codata.M.Bisimilarity where -open import Cubical.Core.Everything -open import Cubical.Codata.M +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv open import Cubical.Foundations.Equiv.Fiberwise -open import Cubical.Foundations.Everything +open import Cubical.Foundations.Transport +open import Cubical.Foundations.HLevels + + +open import Cubical.Codata.M + + open Helpers using (J') -- Bisimilarity as a coinductive record type. diff --git a/Cubical/Data/DescendingList/Base.agda b/Cubical/Data/DescendingList/Base.agda index 4ef8e9220d..af6d168085 100644 --- a/Cubical/Data/DescendingList/Base.agda +++ b/Cubical/Data/DescendingList/Base.agda @@ -4,7 +4,7 @@ {-# OPTIONS --safe #-} -open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude module Cubical.Data.DescendingList.Base diff --git a/Cubical/Data/DescendingList/Examples.agda b/Cubical/Data/DescendingList/Examples.agda index bc5894d725..45a4efb214 100644 --- a/Cubical/Data/DescendingList/Examples.agda +++ b/Cubical/Data/DescendingList/Examples.agda @@ -13,7 +13,7 @@ module Cubical.Data.DescendingList.Examples where -open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Nat diff --git a/Cubical/Data/DescendingList/Properties.agda b/Cubical/Data/DescendingList/Properties.agda index 1992c02cc3..b1f4e8eceb 100644 --- a/Cubical/Data/DescendingList/Properties.agda +++ b/Cubical/Data/DescendingList/Properties.agda @@ -12,7 +12,8 @@ {-# OPTIONS --safe #-} -open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Unit diff --git a/Cubical/Data/DescendingList/Strict/Properties.agda b/Cubical/Data/DescendingList/Strict/Properties.agda index ed23dc0009..f61950b8b3 100644 --- a/Cubical/Data/DescendingList/Strict/Properties.agda +++ b/Cubical/Data/DescendingList/Strict/Properties.agda @@ -64,7 +64,7 @@ module IsoToLFSet >ᴴ-trans x y (cons z zs _) x>y (>ᴴcons y>z) = >ᴴcons (>-trans x>y y>z) ≡ₚ-sym : ∀ {A : Type} {x y : A} → ⟨ x ≡ₚ y ⟩ → ⟨ y ≡ₚ x ⟩ - ≡ₚ-sym p = PropTrunc.rec squash (λ p → ∣ sym p ∣) p + ≡ₚ-sym p = PropTrunc.rec squash₁ (λ p → ∣ sym p ∣₁) p >-all : ∀ x l → x >ᴴ l → ∀ a → ⟨ a ∈ˡ l ⟩ → x > a >-all x (cons y zs y>zs) (>ᴴcons x>y) a a∈l = @@ -153,14 +153,14 @@ module IsoToLFSet Memˡ-inj : ∀ l₁ l₂ → Memˡ l₁ ≡ Memˡ l₂ → l₁ ≡ l₂ Memˡ-inj [] [] eq = 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 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 _ _) (PropTrunc.rec (SDL-isSet _ _) with-x≡y) (⊥.rec ∘ x∉ys) - (transport (λ i → ⟨ e i x ⟩) (inl ∣ refl ∣)) where + (transport (λ i → ⟨ e i x ⟩) (inl ∣ refl ∣₁)) where xxs = cons x xs x>xs @@ -170,7 +170,7 @@ module IsoToLFSet 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 = (transport (λ i → ⟨ e (~ i) y ⟩) (inl ∣ refl ∣₁)) y>y : y > y y>y = >-all y xxs (>ᴴcons y>x) y y∈xxs diff --git a/Cubical/Data/Equality.agda b/Cubical/Data/Equality.agda index 33ec8eb7b3..4e8609af5e 100644 --- a/Cubical/Data/Equality.agda +++ b/Cubical/Data/Equality.agda @@ -32,8 +32,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence hiding (univalence) open import Cubical.HITs.PropositionalTruncation public - renaming ( squash to squashPath - ; rec to recPropTruncPath + renaming (rec to recPropTruncPath ; elim to elimPropTruncPath ) open import Cubical.HITs.S1 as S1 renaming (loop to loopPath) @@ -257,14 +256,14 @@ univalenceEq {A = A} {B = B} = equivPathToEquiv rem -- Propositional truncation using ≡ with paths under the hood -∥∥-isProp : ∀ (x y : ∥ A ∥) → x ≡ y -∥∥-isProp x y = pathToEq (squashPath x y) +∥∥-isProp : ∀ (x y : ∥ A ∥₁) → x ≡ y +∥∥-isProp x y = pathToEq (squash₁ x y) -∥∥-recursion : ∀ {A : Type ℓ} {P : Type ℓ} → isProp P → (A → P) → ∥ A ∥ → P +∥∥-recursion : ∀ {A : Type ℓ} {P : Type ℓ} → isProp P → (A → P) → ∥ A ∥₁ → P ∥∥-recursion Pprop = recPropTruncPath (isPropToIsPropPath Pprop) -∥∥-induction : ∀ {A : Type ℓ} {P : ∥ A ∥ → Type ℓ} → ((a : ∥ A ∥) → isProp (P a)) → - ((x : A) → P ∣ x ∣) → (a : ∥ A ∥) → P a +∥∥-induction : ∀ {A : Type ℓ} {P : ∥ A ∥₁ → Type ℓ} → ((a : ∥ A ∥₁) → isProp (P a)) → + ((x : A) → P ∣ x ∣₁) → (a : ∥ A ∥₁) → P a ∥∥-induction Pprop = elimPropTruncPath (λ a → isPropToIsPropPath (Pprop a)) @@ -276,9 +275,6 @@ loop = pathToEq loopPath S¹-rec : {A : Type ℓ} (b : A) (l : b ≡ b) → S¹ → A S¹-rec b l = S1.rec b (eqToPath l) -S¹-elim : (C : S¹ → Type ℓ) (b : C base) (l : transport C loop b ≡ b) → (x : S¹) → C x -S¹-elim C b l x = S1.ind C b (eqToPath (substPath≡transport C b loopPath ∙ l)) x - S¹-recβ : (b : A) (l : b ≡ b) → ap (S¹-rec b l) loop ≡ l S¹-recβ b l = ap (S¹-rec b l) loop ≡⟨ ap≡congPath (S¹-rec b l) loopPath ⟩ @@ -329,4 +325,3 @@ private -- test-winding-neg5 : winding (loop^ (negsuc 2)) ≡ negsuc 2 -- test-winding-neg5 = refl - diff --git a/Cubical/Data/Fin/LehmerCode.agda b/Cubical/Data/Fin/LehmerCode.agda index ece75c97b9..eede1283b9 100644 --- a/Cubical/Data/Fin/LehmerCode.agda +++ b/Cubical/Data/Fin/LehmerCode.agda @@ -108,7 +108,7 @@ isEquivPunchOut {i = i} = isEmbedding×isSurjection→isEquiv (isEmbPunchOut , i isEmbPunchOut : isEmbedding (punchOut i) isEmbPunchOut = injEmbedding isSetFinExcept isSetFin λ {_} {_} → punchOut-injective i _ _ isSurPunchOut : isSurjection (punchOut i) - isSurPunchOut b = ∣ _ , punchOut∘In i b ∣ + isSurPunchOut b = ∥_∥₁.∣ _ , (punchOut∘In i b) ∣₁ punchOutEquiv : {i : Fin (suc n)} → FinExcept i ≃ Fin n punchOutEquiv = _ , isEquivPunchOut diff --git a/Cubical/Data/Fin/Properties.agda b/Cubical/Data/Fin/Properties.agda index 3267fab2e0..7e736212ef 100644 --- a/Cubical/Data/Fin/Properties.agda +++ b/Cubical/Data/Fin/Properties.agda @@ -516,7 +516,7 @@ factorEquiv {suc n} {m} = intro , isEmbedding×isSurjection→isEquiv (isEmbeddi mm0→isInhab : card X > 0 → ∥ X .fst ∥ + card>0→isInhab : card X > 0 → ∥ X .fst ∥₁ card>0→isInhab p = Prop.map (λ e → invEq e (Fin>0→isInhab _ p)) (∣≃card∣ X) - card>1→hasNonEqualTerm : card X > 1 → ∥ Σ[ a ∈ X .fst ] Σ[ b ∈ X .fst ] ¬ a ≡ b ∥ + card>1→hasNonEqualTerm : card X > 1 → ∥ Σ[ a ∈ X .fst ] Σ[ b ∈ X .fst ] ¬ a ≡ b ∥₁ card>1→hasNonEqualTerm p = Prop.map (λ e → @@ -111,7 +111,7 @@ module _ card≤1→isProp p = Prop.rec isPropIsProp (λ e → isOfHLevelRespectEquiv 1 (invEquiv e) (Fin≤1→isProp (card X) p)) (∣≃card∣ X) - card≡n : card X ≡ n → ∥ X ≡ 𝔽in n ∥ + card≡n : card X ≡ n → ∥ X ≡ 𝔽in n ∥₁ card≡n {n = n} p = Prop.map (λ e → @@ -119,7 +119,7 @@ module _ ua e i , isProp→PathP {B = λ j → isFinSet (ua e j)} (λ _ → isPropIsFinSet) (X .snd) (𝔽in n .snd) i )) - (∣≃card∣ X ⋆̂ ∣ pathToEquiv (cong Fin p) ⋆ invEquiv (𝔽in≃Fin n) ∣) + (∣≃card∣ X ⋆̂ ∣ pathToEquiv (cong Fin p) ⋆ invEquiv (𝔽in≃Fin n) ∣₁) card≡0 : card X ≡ 0 → X ≡ 𝟘 card≡0 p = @@ -146,7 +146,7 @@ module _ isEmpty→card≡0 p = Prop.rec (isSetℕ _ _) (λ e → sym (isEmpty→Fin≡0 _ (p ∘ invEq e))) (∣≃card∣ X) - isInhab→card>0 : ∥ X .fst ∥ → card X > 0 + isInhab→card>0 : ∥ X .fst ∥₁ → card X > 0 isInhab→card>0 = Prop.rec2 isProp≤ (λ p x → isInhab→Fin>0 _ (p .fst x)) (∣≃card∣ X) hasNonEqualTerm→card>1 : {a b : X. fst} → ¬ a ≡ b → card X > 1 @@ -154,7 +154,7 @@ module _ Prop.rec isProp≤ (λ p → hasNonEqualTerm→Fin>1 _ (p .fst a) (p .fst b) (q ∘ invEq (congEquiv p))) (∣≃card∣ X) isContr→card≡1 : isContr (X .fst) → card X ≡ 1 - isContr→card≡1 p = cardEquiv X (_ , isFinSetUnit) ∣ isContr→≃Unit p ∣ + isContr→card≡1 p = cardEquiv X (_ , isFinSetUnit) ∣ isContr→≃Unit p ∣₁ isProp→card≤1 : isProp (X .fst) → card X ≤ 1 isProp→card≤1 p = isProp→Fin≤1 (card X) (Prop.rec isPropIsProp (λ e → isOfHLevelRespectEquiv 1 e p) (∣≃card∣ X)) @@ -170,7 +170,7 @@ card𝟙 : card (𝟙 {ℓ}) ≡ 1 card𝟙 {ℓ = ℓ} = isContr→card≡1 (𝟙 {ℓ}) isContrUnit* card𝔽in : (n : ℕ) → card (𝔽in {ℓ} n) ≡ n -card𝔽in {ℓ = ℓ} n = cardEquiv (𝔽in {ℓ} n) (_ , isFinSetFin) ∣ 𝔽in≃Fin n ∣ +card𝔽in {ℓ = ℓ} n = cardEquiv (𝔽in {ℓ} n) (_ , isFinSetFin) ∣ 𝔽in≃Fin n ∣₁ -- addition/product formula @@ -215,12 +215,12 @@ module _ sum𝟙 : sum 𝟙 f ≡ f tt* sum𝟙 = cardEquiv (_ , isFinSetΣ 𝟙 (λ x → Fin (f x) , isFinSetFin)) - (Fin (f tt*) , isFinSetFin) ∣ Σ-contractFst isContrUnit* ∣ + (Fin (f tt*) , isFinSetFin) ∣ Σ-contractFst isContrUnit* ∣₁ prod𝟙 : prod 𝟙 f ≡ f tt* prod𝟙 = cardEquiv (_ , isFinSetΠ 𝟙 (λ x → Fin (f x) , isFinSetFin)) - (Fin (f tt*) , isFinSetFin) ∣ ΠUnit* _ ∣ + (Fin (f tt*) , isFinSetFin) ∣ ΠUnit* _ ∣₁ module _ (X : FinSet ℓ ) @@ -231,7 +231,7 @@ module _ sum⊎ = cardEquiv (_ , isFinSetΣ (_ , isFinSet⊎ X Y) (λ x → Fin (f x) , isFinSetFin)) (_ , isFinSet⊎ (_ , isFinSetΣ X (λ x → Fin (f (inl x)) , isFinSetFin)) - (_ , isFinSetΣ Y (λ y → Fin (f (inr y)) , isFinSetFin))) ∣ Σ⊎≃ ∣ + (_ , isFinSetΣ Y (λ y → Fin (f (inr y)) , isFinSetFin))) ∣ Σ⊎≃ ∣₁ ∙ card+ (_ , isFinSetΣ X (λ x → Fin (f (inl x)) , isFinSetFin)) (_ , isFinSetΣ Y (λ y → Fin (f (inr y)) , isFinSetFin)) @@ -239,7 +239,7 @@ module _ prod⊎ = cardEquiv (_ , isFinSetΠ (_ , isFinSet⊎ X Y) (λ x → Fin (f x) , isFinSetFin)) (_ , isFinSet× (_ , isFinSetΠ X (λ x → Fin (f (inl x)) , isFinSetFin)) - (_ , isFinSetΠ Y (λ y → Fin (f (inr y)) , isFinSetFin))) ∣ Π⊎≃ ∣ + (_ , isFinSetΠ Y (λ y → Fin (f (inr y)) , isFinSetFin))) ∣ Π⊎≃ ∣₁ ∙ card× (_ , isFinSetΠ X (λ x → Fin (f (inl x)) , isFinSetFin)) (_ , isFinSetΠ Y (λ y → Fin (f (inr y)) , isFinSetFin)) @@ -297,7 +297,7 @@ sum≤𝔽in 0 f g _ = subst2 (_≤_) (sym (sum𝟘 f)) (sym (sum𝟘 g)) ≤-re sum≤𝔽in (suc n) f g h = ≡≤ (h (inl tt*)) (sum≤𝔽in n (f ∘ inr) (g ∘ inr) (h ∘ inr)) (sum𝔽in1+n n f) (sum𝔽in1+n n g) -sum<𝔽in : (n : ℕ)(f g : 𝔽in {ℓ} n .fst → ℕ)(t : ∥ 𝔽in {ℓ} n .fst ∥)(h : (x : 𝔽in n .fst) → f x < g x) +sum<𝔽in : (n : ℕ)(f g : 𝔽in {ℓ} n .fst → ℕ)(t : ∥ 𝔽in {ℓ} n .fst ∥₁)(h : (x : 𝔽in n .fst) → f x < g x) → sum (𝔽in n) f < sum (𝔽in n) g sum<𝔽in {ℓ = ℓ} 0 _ _ t _ = Empty.rec (<→≢ (isInhab→card>0 (𝔽in 0) t) (card𝟘 {ℓ = ℓ})) sum<𝔽in (suc n) f g t h = @@ -326,13 +326,13 @@ module _ (λ X → isPropΠ3 (λ _ _ _ → isProp≤)) sum≤𝔽in X f g h module _ - (t : ∥ X .fst ∥) + (t : ∥ X .fst ∥₁) (h : (x : X .fst) → f x < g x) where sum< : sum X f < sum X g sum< = elimProp - (λ X → (f g : X .fst → ℕ)(t : ∥ X .fst ∥)(h : (x : X .fst) → f x < g x) → sum X f < sum X g) + (λ X → (f g : X .fst → ℕ)(t : ∥ X .fst ∥₁)(h : (x : X .fst) → f x < g x) → sum X f < sum X g) (λ X → isPropΠ4 (λ _ _ _ _ → isProp≤)) sum<𝔽in X f g t h module _ @@ -389,7 +389,7 @@ module _ sumCardFiber : card X ≡ sum Y (λ y → card (_ , isFinSetFiber X Y f y)) sumCardFiber = - cardEquiv X (_ , isFinSetΣ Y (λ y → _ , isFinSetFiber X Y f y)) ∣ totalEquiv f ∣ + cardEquiv X (_ , isFinSetΣ Y (λ y → _ , isFinSetFiber X Y f y)) ∣ totalEquiv f ∣₁ ∙ cardΣ Y (λ y → _ , isFinSetFiber X Y f y) -- the pigeonhole priniple @@ -420,7 +420,7 @@ module _ ¬ΠQ→¬¬ΣP (Y .fst) (λ y → _ > n) (λ y → _ ≤ n) (λ y → <-asym') (λ h → <-asym p (fiberCount h)) - pigeonHole : ∥ Σ[ y ∈ Y .fst ] card (_ , isFinSetFiber X Y f y) > n ∥ + pigeonHole : ∥ Σ[ y ∈ Y .fst ] card (_ , isFinSetFiber X Y f y) > n ∥₁ pigeonHole = PeirceLaw (isFinSetΣ Y (λ _ → _ , isDecProp→isFinSet isProp≤ (≤Dec _ _))) ¬¬pigeonHole -- a special case, proved in Cubical.Data.Fin.Properties @@ -428,7 +428,7 @@ module _ -- a technical lemma private Σ∥P∥→∥ΣP∥ : (X : Type ℓ)(P : X → Type ℓ') - → Σ X (λ x → ∥ P x ∥) → ∥ Σ X P ∥ + → Σ X (λ x → ∥ P x ∥₁) → ∥ Σ X P ∥₁ Σ∥P∥→∥ΣP∥ _ _ (x , p) = Prop.map (λ q → x , q) p module _ @@ -436,7 +436,7 @@ module _ (p : card X > card Y) where fiberNonEqualTerm : Σ[ y ∈ Y .fst ] card (_ , isFinSetFiber X Y f y) > 1 - → ∥ Σ[ y ∈ Y .fst ] Σ[ a ∈ fiber f y ] Σ[ b ∈ fiber f y ] ¬ a ≡ b ∥ + → ∥ Σ[ y ∈ Y .fst ] Σ[ a ∈ fiber f y ] Σ[ b ∈ fiber f y ] ¬ a ≡ b ∥₁ fiberNonEqualTerm (y , p) = Σ∥P∥→∥ΣP∥ _ _ (y , card>1→hasNonEqualTerm {X = _ , isFinSetFiber X Y f y} p) nonInj : Σ[ y ∈ Y .fst ] Σ[ a ∈ fiber f y ] Σ[ b ∈ fiber f y ] ¬ a ≡ b @@ -447,7 +447,7 @@ module _ t (λ i → u i , isSet→SquareP (λ i j → isFinSet→isSet (Y .snd)) p q (cong f u) refl i) nonInj (y , (x , p) , (x' , q) , t) .snd .snd .snd = p ∙ sym q - pigeonHole' : ∥ Σ[ x ∈ X .fst ] Σ[ x' ∈ X .fst ] (¬ x ≡ x') × (f x ≡ f x') ∥ + pigeonHole' : ∥ Σ[ x ∈ X .fst ] Σ[ x' ∈ X .fst ] (¬ x ≡ x') × (f x ≡ f x') ∥₁ pigeonHole' = Prop.map nonInj (Prop.rec isPropPropTrunc fiberNonEqualTerm @@ -479,10 +479,10 @@ module _ (sumBoundedBelow Y (λ y → card (_ , isFinSetFiber X Y f y)) 1 (λ y → isInhab→card>0 (_ , isFinSetFiber X Y f y) (p y))) - card↪Inequality : ∥ X .fst ↪ Y .fst ∥ → card X ≤ card Y + card↪Inequality : ∥ X .fst ↪ Y .fst ∥₁ → card X ≤ card Y card↪Inequality = Prop.rec isProp≤ (λ (f , p) → card↪Inequality' f p) - card↠Inequality : ∥ X .fst ↠ Y .fst ∥ → card X ≥ card Y + card↠Inequality : ∥ X .fst ↠ Y .fst ∥₁ → card X ≥ card Y card↠Inequality = Prop.rec isProp≤ (λ (f , p) → card↠Inequality' f p) -- maximal value of numerical functions @@ -507,7 +507,7 @@ module _ ΣMax = Σ[ x ∈ X ] isMax x ∃Max : Type ℓ - ∃Max = ∥ ΣMax ∥ + ∃Max = ∥ ΣMax ∥₁ ∃Max→maxValue : ∃Max → ℕ ∃Max→maxValue = @@ -541,24 +541,24 @@ module _ ΣMax𝟙 f .snd x = _ , cong f (sym (isContrUnit* .snd x)) ∃Max𝟙 : (f : 𝟙 {ℓ} .fst → ℕ) → ∃Max _ f -∃Max𝟙 f = ∣ ΣMax𝟙 f ∣ +∃Max𝟙 f = ∣ ΣMax𝟙 f ∣₁ -∃Max𝔽in : (n : ℕ)(f : 𝔽in {ℓ} n .fst → ℕ)(x : ∥ 𝔽in {ℓ} n .fst ∥) → ∃Max _ f +∃Max𝔽in : (n : ℕ)(f : 𝔽in {ℓ} n .fst → ℕ)(x : ∥ 𝔽in {ℓ} n .fst ∥₁) → ∃Max _ f ∃Max𝔽in {ℓ = ℓ} 0 _ x = Empty.rec (<→≢ (isInhab→card>0 (𝔽in 0) x) (card𝟘 {ℓ = ℓ})) ∃Max𝔽in 1 f _ = subst (λ X → (f : X .fst → ℕ) → ∃Max _ f) (sym 𝔽in1≡𝟙) ∃Max𝟙 f ∃Max𝔽in (suc (suc n)) f _ = - ∃Max⊎ (𝟙 .fst) (𝔽in (suc n) .fst) f (∃Max𝟙 (f ∘ inl)) (∃Max𝔽in (suc n) (f ∘ inr) ∣ * {n = n} ∣) + ∃Max⊎ (𝟙 .fst) (𝔽in (suc n) .fst) f (∃Max𝟙 (f ∘ inl)) (∃Max𝔽in (suc n) (f ∘ inr) ∣ * {n = n} ∣₁) module _ (X : FinSet ℓ) (f : X .fst → ℕ) - (x : ∥ X .fst ∥) where + (x : ∥ X .fst ∥₁) where ∃MaxFinSet : ∃Max _ f ∃MaxFinSet = elimProp - (λ X → (f : X .fst → ℕ)(x : ∥ X .fst ∥) → ∃Max _ f) + (λ X → (f : X .fst → ℕ)(x : ∥ X .fst ∥₁) → ∃Max _ f) (λ X → isPropΠ2 (λ _ _ → isPropPropTrunc)) ∃Max𝔽in X f x maxValue : ℕ @@ -606,7 +606,7 @@ card-case P {n = suc (suc n)} p = Empty.rec (¬-<-zero (pred-≤-pred (subst (λ a → a ≤ 1) p (isProp→card≤1 (P .fst) (P .snd))))) isSurjectionBool→FinProp : isSurjection (Bool→FinProp {ℓ = ℓ}) -isSurjectionBool→FinProp P = ∣ card-case P refl ∣ +isSurjectionBool→FinProp P = ∣ card-case P refl ∣₁ FinProp≃Bool : FinProp ℓ ≃ Bool FinProp≃Bool = @@ -628,9 +628,9 @@ module _ (λ p1 p2 → ∣ equivComp (p1 ⋆ pathToEquiv (cong Fin p) ⋆ SumFin≃Fin _) (p2 ⋆ SumFin≃Fin _) ⋆ lehmerEquiv ⋆ lehmerFinEquiv - ⋆ invEquiv (SumFin≃Fin _) ∣) + ⋆ invEquiv (SumFin≃Fin _) ∣₁) (∣≃card∣ X) (∣≃card∣ Y) - isFinSet≃Eff' (no ¬p) = 0 , ∣ uninhabEquiv (¬p ∘ cardEquiv X Y ∘ ∣_∣) (idfun _) ∣ + isFinSet≃Eff' (no ¬p) = 0 , ∣ uninhabEquiv (¬p ∘ cardEquiv X Y ∘ ∣_∣₁) (idfun _) ∣₁ isFinSet≃Eff : isFinSet (X .fst ≃ Y .fst) isFinSet≃Eff = isFinSet≃Eff' (discreteℕ _ _) diff --git a/Cubical/Data/FinSet/Constructors.agda b/Cubical/Data/FinSet/Constructors.agda index 332138f50a..e4b95450dc 100644 --- a/Cubical/Data/FinSet/Constructors.agda +++ b/Cubical/Data/FinSet/Constructors.agda @@ -40,7 +40,7 @@ private module _ (X : Type ℓ)(p : isFinOrd X) where - isFinOrd∥∥ : isFinOrd ∥ X ∥ + isFinOrd∥∥ : isFinOrd ∥ X ∥₁ isFinOrd∥∥ = _ , propTrunc≃ (p .snd) ⋆ SumFin∥∥≃ _ isFinOrd≃ : isFinOrd (X ≃ X) @@ -117,7 +117,7 @@ module _ isFinSet≡ : (a b : X .fst) → isFinSet (a ≡ b) isFinSet≡ a b = isDecProp→isFinSet (isFinSet→isSet (X .snd) a b) (isFinSet→Discrete (X .snd) a b) - isFinSet∥∥ : isFinSet ∥ X .fst ∥ + isFinSet∥∥ : isFinSet ∥ X .fst ∥₁ isFinSet∥∥ = Prop.rec isPropIsFinSet (λ p → isFinOrd→isFinSet (isFinOrd∥∥ (X .fst) (_ , p))) (X .snd .snd) isFinSetIsContr : isFinSet (isContr (X .fst)) diff --git a/Cubical/Data/FinSet/DecidablePredicate.agda b/Cubical/Data/FinSet/DecidablePredicate.agda index 65025ff805..8ca52411be 100644 --- a/Cubical/Data/FinSet/DecidablePredicate.agda +++ b/Cubical/Data/FinSet/DecidablePredicate.agda @@ -40,7 +40,7 @@ module _ isDecProp¬' : isDecProp (¬ X) isDecProp¬' = _ , invEquiv (preCompEquiv (p .snd)) ⋆ SumFin¬ _ - isDecProp∥∥' : isDecProp ∥ X ∥ + isDecProp∥∥' : isDecProp ∥ X ∥₁ isDecProp∥∥' = _ , propTrunc≃ (p .snd) ⋆ SumFin∥∥DecProp _ module _ @@ -57,7 +57,7 @@ module _ ⋆ Σ-cong-equiv-snd (λ x → dec (invEq e x) .snd) ⋆ SumFinSub≃ _ (fst ∘ dec ∘ invEq e) - isDecProp∃' : isDecProp ∥ Σ X P ∥ + isDecProp∃' : isDecProp ∥ Σ X P ∥₁ isDecProp∃' = _ , Prop.propTrunc≃ ( Σ-cong-equiv {B' = λ x → P (invEq e x)} e (transpFamily p) @@ -90,7 +90,7 @@ module _ (λ p → isFinOrd→isFinSet (isFinOrdSub (X .fst) (_ , p) (λ x → P x .fst) (λ x → P x .snd))) (X .snd .snd) - isDecProp∃ : isDecProp ∥ Σ (X .fst) (λ x → P x .fst) ∥ + isDecProp∃ : isDecProp ∥ Σ (X .fst) (λ x → P x .fst) ∥₁ isDecProp∃ = Prop.rec isPropIsDecProp (λ p → isDecProp∃' (X .fst) (_ , p) (λ x → P x .fst) (λ x → P x .snd)) (X .snd .snd) @@ -137,6 +137,6 @@ module _ isDecProp¬ = Prop.rec isPropIsDecProp (λ p → isDecProp¬' (X .fst) (_ , p)) (X .snd .snd) - isDecProp∥∥ : isDecProp ∥ X .fst ∥ + isDecProp∥∥ : isDecProp ∥ X .fst ∥₁ isDecProp∥∥ = Prop.rec isPropIsDecProp (λ p → isDecProp∥∥' (X .fst) (_ , p)) (X .snd .snd) diff --git a/Cubical/Data/FinSet/FiniteChoice.agda b/Cubical/Data/FinSet/FiniteChoice.agda index eb5657278b..a02940c023 100644 --- a/Cubical/Data/FinSet/FiniteChoice.agda +++ b/Cubical/Data/FinSet/FiniteChoice.agda @@ -29,7 +29,7 @@ private ℓ ℓ' : Level choice≃Fin : - {n : ℕ}(Y : Fin n → Type ℓ) → ((x : Fin n) → ∥ Y x ∥) ≃ ∥ ((x : Fin n) → Y x) ∥ + {n : ℕ}(Y : Fin n → Type ℓ) → ((x : Fin n) → ∥ Y x ∥₁) ≃ ∥ ((x : Fin n) → Y x) ∥₁ choice≃Fin {n = 0} Y = isContr→≃Unit (isContrΠ⊥) ⋆ Unit≃Unit* @@ -37,7 +37,7 @@ choice≃Fin {n = 0} Y = ⋆ propTrunc≃ (invEquiv (isContr→≃Unit* (isContrΠ⊥ {A = Y}))) choice≃Fin {n = suc n} Y = Π⊎≃ - ⋆ Σ-cong-equiv-fst (ΠUnit (λ x → ∥ Y (inl x) ∥)) + ⋆ Σ-cong-equiv-fst (ΠUnit (λ x → ∥ Y (inl x) ∥₁)) ⋆ Σ-cong-equiv-snd (λ _ → choice≃Fin {n = n} (λ x → Y (inr x))) ⋆ Σ-cong-equiv-fst (propTrunc≃ (invEquiv (ΠUnit (λ x → Y (inl x))))) ⋆ ∥∥-×-≃ @@ -50,9 +50,9 @@ module _ private e = p .snd - choice≃' : ((x : X) → ∥ Y x ∥) ≃ ∥ ((x : X) → Y x) ∥ + choice≃' : ((x : X) → ∥ Y x ∥₁) ≃ ∥ ((x : X) → Y x) ∥₁ choice≃' = - equivΠ {B' = λ x → ∥ Y (invEq e x) ∥} e (transpFamily p) + equivΠ {B' = λ x → ∥ Y (invEq e x) ∥₁} e (transpFamily p) ⋆ choice≃Fin _ ⋆ propTrunc≃ (invEquiv (equivΠ {B' = λ x → Y (invEq e x)} e (transpFamily p))) @@ -60,11 +60,11 @@ module _ (X : FinSet ℓ) (Y : X .fst → Type ℓ') where - choice≃ : ((x : X .fst) → ∥ Y x ∥) ≃ ∥ ((x : X .fst) → Y x) ∥ + choice≃ : ((x : X .fst) → ∥ Y x ∥₁) ≃ ∥ ((x : X .fst) → Y x) ∥₁ choice≃ = Prop.rec (isOfHLevel≃ 1 (isPropΠ (λ x → isPropPropTrunc)) isPropPropTrunc) (λ p → choice≃' (X .fst) (_ , p) Y) (X .snd .snd) - choice : ((x : X .fst) → ∥ Y x ∥) → ∥ ((x : X .fst) → Y x) ∥ + choice : ((x : X .fst) → ∥ Y x ∥₁) → ∥ ((x : X .fst) → Y x) ∥₁ choice = choice≃ .fst diff --git a/Cubical/Data/FinSet/Induction.agda b/Cubical/Data/FinSet/Induction.agda index e240fad93c..81cb16ba85 100644 --- a/Cubical/Data/FinSet/Induction.agda +++ b/Cubical/Data/FinSet/Induction.agda @@ -38,7 +38,7 @@ module _ {ℓ : Level} where 𝟘 : FinSet ℓ - 𝟘 = ⊥* , 0 , ∣ uninhabEquiv Empty.rec* Empty.rec ∣ + 𝟘 = ⊥* , 0 , ∣ uninhabEquiv Empty.rec* Empty.rec ∣₁ 𝟙 : FinSet ℓ 𝟙 = Unit* , isContr→isFinSet (isContrUnit*) @@ -92,7 +92,7 @@ module _ -- every finite sets are merely equal to some 𝔽in -∣≡𝔽in∣ : (X : FinSet ℓ) → ∥ Σ[ n ∈ ℕ ] X ≡ 𝔽in n ∥ +∣≡𝔽in∣ : (X : FinSet ℓ) → ∥ Σ[ n ∈ ℕ ] X ≡ 𝔽in n ∥₁ ∣≡𝔽in∣ X = Prop.map (λ (n , p) → n , path X (n , p)) (isFinSet→isFinSet' (X .snd)) where path : (X : FinSet ℓ) → ((n , _) : isFinOrd (X .fst)) → X ≡ 𝔽in n diff --git a/Cubical/Data/FinSet/Properties.agda b/Cubical/Data/FinSet/Properties.agda index 873870480f..666681064e 100644 --- a/Cubical/Data/FinSet/Properties.agda +++ b/Cubical/Data/FinSet/Properties.agda @@ -39,43 +39,43 @@ module _ infixr 30 _⋆̂_ - _⋆̂_ : ∥ A ≃ B ∥ → ∥ B ≃ C ∥ → ∥ A ≃ C ∥ + _⋆̂_ : ∥ A ≃ B ∥₁ → ∥ B ≃ C ∥₁ → ∥ A ≃ C ∥₁ _⋆̂_ = Prop.map2 (_⋆_) module _ {A : Type ℓ}{B : Type ℓ'} where - ∣invEquiv∣ : ∥ A ≃ B ∥ → ∥ B ≃ A ∥ + ∣invEquiv∣ : ∥ A ≃ B ∥₁ → ∥ B ≃ A ∥₁ ∣invEquiv∣ = Prop.map invEquiv -- useful implications EquivPresIsFinSet : A ≃ B → isFinSet A → isFinSet B -EquivPresIsFinSet e (_ , p) = _ , ∣ invEquiv e ∣ ⋆̂ p +EquivPresIsFinSet e (_ , p) = _ , ∣ invEquiv e ∣₁ ⋆̂ p isFinSetFin : {n : ℕ} → isFinSet (Fin n) -isFinSetFin = _ , ∣ idEquiv _ ∣ +isFinSetFin = _ , ∣ idEquiv _ ∣₁ isFinSetUnit : isFinSet Unit -isFinSetUnit = 1 , ∣ invEquiv Fin1≃Unit ∣ +isFinSetUnit = 1 , ∣ invEquiv Fin1≃Unit ∣₁ isFinSetBool : isFinSet Bool -isFinSetBool = 2 , ∣ invEquiv SumFin2≃Bool ∣ +isFinSetBool = 2 , ∣ invEquiv SumFin2≃Bool ∣₁ isFinSet→Discrete : isFinSet A → Discrete A isFinSet→Discrete h = Prop.rec isPropDiscrete (λ p → EquivPresDiscrete (invEquiv p) discreteFin) (h .snd) isContr→isFinSet : isContr A → isFinSet A -isContr→isFinSet h = 1 , ∣ isContr→≃Unit* h ⋆ invEquiv Unit≃Unit* ⋆ invEquiv Fin1≃Unit ∣ +isContr→isFinSet h = 1 , ∣ isContr→≃Unit* h ⋆ invEquiv Unit≃Unit* ⋆ invEquiv Fin1≃Unit ∣₁ isDecProp→isFinSet : isProp A → Dec A → isFinSet A isDecProp→isFinSet h (yes p) = isContr→isFinSet (inhProp→isContr p h) -isDecProp→isFinSet h (no ¬p) = 0 , ∣ uninhabEquiv ¬p (idfun _) ∣ +isDecProp→isFinSet h (no ¬p) = 0 , ∣ uninhabEquiv ¬p (idfun _) ∣₁ -isDec→isFinSet∥∥ : Dec A → isFinSet ∥ A ∥ +isDec→isFinSet∥∥ : Dec A → isFinSet ∥ A ∥₁ isDec→isFinSet∥∥ dec = isDecProp→isFinSet isPropPropTrunc (Dec∥∥ dec) -isFinSet→Dec∥∥ : isFinSet A → Dec ∥ A ∥ +isFinSet→Dec∥∥ : isFinSet A → Dec ∥ A ∥₁ isFinSet→Dec∥∥ h = Prop.rec (isPropDec isPropPropTrunc) (λ p → EquivPresDec (propTrunc≃ (invEquiv p)) (Dec∥Fin∥ _)) (h .snd) @@ -83,11 +83,11 @@ isFinSet→Dec∥∥ h = isFinProp→Dec : isFinSet A → isProp A → Dec A isFinProp→Dec p h = subst Dec (propTruncIdempotent h) (isFinSet→Dec∥∥ p) -PeirceLaw∥∥ : isFinSet A → NonEmpty ∥ A ∥ → ∥ A ∥ +PeirceLaw∥∥ : isFinSet A → NonEmpty ∥ A ∥₁ → ∥ A ∥₁ PeirceLaw∥∥ p = Dec→Stable (isFinSet→Dec∥∥ p) -PeirceLaw : isFinSet A → NonEmpty A → ∥ A ∥ -PeirceLaw p q = PeirceLaw∥∥ p (λ f → q (λ x → f ∣ x ∣)) +PeirceLaw : isFinSet A → NonEmpty A → ∥ A ∥₁ +PeirceLaw p q = PeirceLaw∥∥ p (λ f → q (λ x → f ∣ x ∣₁)) -- transprot family towards Fin diff --git a/Cubical/Data/FinSet/Quotients.agda b/Cubical/Data/FinSet/Quotients.agda index d376366b6d..2f73306e75 100644 --- a/Cubical/Data/FinSet/Quotients.agda +++ b/Cubical/Data/FinSet/Quotients.agda @@ -88,7 +88,7 @@ module _ isFinSetℙEff : isFinSet (ℙEff (X .fst)) isFinSetℙEff = 2 ^ (card X) , Prop.elim (λ _ → isPropPropTrunc {A = _ ≃ Fin _}) - (λ p → ∣ isFinOrdℙEff (X .fst) (_ , p) .snd ∣) + (λ p → ∣ isFinOrdℙEff (X .fst) (_ , p) .snd ∣₁) (X .snd .snd) module _ @@ -97,21 +97,21 @@ module _ (dec : (x x' : X .fst) → isDecProp (R x x')) where isEqClassEff : ℙEff (X .fst) → Type ℓ - isEqClassEff f = ∥ Σ[ x ∈ X .fst ] ((a : X .fst) → f a ≡ dec a x .fst) ∥ + isEqClassEff f = ∥ Σ[ x ∈ X .fst ] ((a : X .fst) → f a ≡ dec a x .fst) ∥₁ isDecPropIsEqClassEff : {f : ℙEff (X .fst)} → isDecProp (isEqClassEff f) isDecPropIsEqClassEff = isDecProp∃ X (λ _ → _ , isDecProp∀ X (λ _ → _ , _ , Bool≡≃ _ _)) isEqClassEff→isEqClass' : (f : ℙEff (X .fst))(x : X .fst) → ((a : X .fst) → f a ≡ dec a x .fst) - → (a : X .fst) → Bool→Type* {ℓ = ℓ'} (f a) ≃ ∥ R a x ∥ + → (a : X .fst) → Bool→Type* {ℓ = ℓ'} (f a) ≃ ∥ R a x ∥₁ isEqClassEff→isEqClass' f x h a = pathToEquiv (cong Bool→Type* (h a)) ⋆ invEquiv (LiftDecProp (dec a x)) ⋆ invEquiv (propTruncIdempotent≃ (isDecProp→isProp (dec a x))) isEqClass→isEqClassEff' : (f : ℙEff (X .fst))(x : X .fst) - → ((a : X .fst) → Bool→Type* {ℓ = ℓ'} (f a) ≃ ∥ R a x ∥) + → ((a : X .fst) → Bool→Type* {ℓ = ℓ'} (f a) ≃ ∥ R a x ∥₁) → (a : X .fst) → f a ≡ dec a x .fst isEqClass→isEqClassEff' f x h a = Bool→TypeInj* _ _ diff --git a/Cubical/Data/FinType/FiniteStructure.agda b/Cubical/Data/FinType/FiniteStructure.agda index b34df802cd..928503ba83 100644 --- a/Cubical/Data/FinType/FiniteStructure.agda +++ b/Cubical/Data/FinType/FiniteStructure.agda @@ -53,9 +53,9 @@ FinSetOfCard≡ _ _ = Σ≡PropEquiv (λ _ → isSetℕ _ _) open Iso -∥FinSetOfCard∥₂≡ : (X Y : FinSetOfCard ℓ n) → ∥ X .fst ≡ Y .fst ∥ → ∣ X ∣₂ ≡ ∣ Y ∣₂ +∥FinSetOfCard∥₂≡ : (X Y : FinSetOfCard ℓ n) → ∥ X .fst ≡ Y .fst ∥₁ → ∣ X ∣₂ ≡ ∣ Y ∣₂ ∥FinSetOfCard∥₂≡ _ _ = - Prop.rec (squash₂ _ _) (λ p → PathIdTrunc₀Iso .inv ∣ FinSetOfCard≡ _ _ .fst p ∣) + Prop.rec (squash₂ _ _) (λ p → PathIdTrunc₀Iso .inv ∣ FinSetOfCard≡ _ _ .fst p ∣₁) isPathConnectedFinSetOfCard : isContr ∥ FinSetOfCard ℓ n ∥₂ isPathConnectedFinSetOfCard {n = n} .fst = ∣ 𝔽in n , card𝔽in n ∣₂ diff --git a/Cubical/Data/Int/Divisibility.agda b/Cubical/Data/Int/Divisibility.agda index 9940e3a286..6a67035014 100644 --- a/Cubical/Data/Int/Divisibility.agda +++ b/Cubical/Data/Int/Divisibility.agda @@ -31,7 +31,6 @@ open import Cubical.Relation.Nullary open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Instances.Int - renaming (ℤ to Ringℤ) open import Cubical.Algebra.CommRingSolver.Reflection private @@ -53,10 +52,10 @@ private helper3 : (n m d r : 𝓡 .fst) → n ≡ d · m + r → n + (- d) · m ≡ r helper3 n m d r p = (λ t → p t + (- d) · m) ∙ helper2 d m r -open Helper Ringℤ +open Helper ℤCommRing -open CommRingStr (Ringℤ .snd) +open CommRingStr (ℤCommRing .snd) -- The Divisibility Relation -- Most definitions are the same as in Cubical.Data.Nat.Divisibility @@ -65,7 +64,7 @@ _∣_ : ℤ → ℤ → Type m ∣ n = ∃[ c ∈ ℤ ] c · m ≡ n isProp∣ : isProp (m ∣ n) -isProp∣ = squash +isProp∣ = squash₁ -- Untruncated divisiblility relation @@ -82,15 +81,15 @@ isProp∣' {m = negsuc m} {n = n} p q = Σ≡Prop (λ _ → isSetℤ _ _) (·rCancel (negsuc m) _ _ (p .snd ∙ sym (q .snd)) (negsucNotpos _ 0)) ∣→∣' : (m n : ℤ) → m ∣ n → m ∣' n -∣→∣' (pos 0) n ∣ c , p ∣ = ·Comm 0 c ∙ p -∣→∣' (pos (suc m)) n ∣ p ∣ = p -∣→∣' (negsuc m) n ∣ p ∣ = p -∣→∣' m n (squash p q i) = isProp∣' (∣→∣' _ _ p) (∣→∣' _ _ q) i +∣→∣' (pos 0) n ∣ c , p ∣₁ = ·Comm 0 c ∙ p +∣→∣' (pos (suc m)) n ∣ p ∣₁ = p +∣→∣' (negsuc m) n ∣ p ∣₁ = p +∣→∣' m n (squash₁ p q i) = isProp∣' (∣→∣' _ _ p) (∣→∣' _ _ q) i ∣'→∣ : (m n : ℤ) → m ∣' n → m ∣ n -∣'→∣ (pos 0) n p = ∣ 0 , p ∣ -∣'→∣ (pos (suc m)) n p = ∣ p ∣ -∣'→∣ (negsuc m) n p = ∣ p ∣ +∣'→∣ (pos 0) n p = ∣ 0 , p ∣₁ +∣'→∣ (pos (suc m)) n p = ∣ p ∣₁ +∣'→∣ (negsuc m) n p = ∣ p ∣₁ ∣≃∣' : (m n : ℤ) → (m ∣ n) ≃ (m ∣' n) ∣≃∣' m n = propBiimpl→Equiv isProp∣ isProp∣' (∣→∣' _ _) (∣'→∣ _ _) @@ -98,19 +97,19 @@ isProp∣' {m = negsuc m} {n = n} p q = -- Properties of divisibility ∣-left : m ∣ (m · k) -∣-left {k = k} = ∣ k , ·Comm k _ ∣ +∣-left {k = k} = ∣ k , ·Comm k _ ∣₁ ∣-right : m ∣ (k · m) -∣-right {k = k} = ∣ k , refl ∣ +∣-right {k = k} = ∣ k , refl ∣₁ ∣-refl : m ≡ n → m ∣ n -∣-refl p = ∣ 1 , p ∣ +∣-refl p = ∣ 1 , p ∣₁ ∣-zeroˡ : 0 ∣ m → 0 ≡ m ∣-zeroˡ = ∣→∣' _ _ ∣-zeroʳ : m ∣ 0 -∣-zeroʳ = ∣ 0 , refl ∣ +∣-zeroʳ = ∣ 0 , refl ∣₁ ∣-+ : k ∣ m → k ∣ n → k ∣ (m + n) ∣-+ = @@ -132,7 +131,7 @@ isProp∣' {m = negsuc m} {n = n} p q = -- Natural numbers back and forth (using abs) ∣→∣ℕ : m ∣ n → abs m ∣ℕ abs n -∣→∣ℕ {m = m} = Prop.rec isProp∣ℕ (λ (c , h) → ∣ abs c , sym (abs· c m) ∙ cong abs h ∣) +∣→∣ℕ {m = m} = Prop.rec isProp∣ℕ (λ (c , h) → ∣ abs c , sym (abs· c m) ∙ cong abs h ∣₁) private ∣ℕ→∣-helper : (m n : ℤ) @@ -169,7 +168,7 @@ private ∙ sym q ∣ℕ→∣ : abs m ∣ℕ abs n → m ∣ n -∣ℕ→∣ = Prop.rec isProp∣ (λ (c , h) → ∣ ∣ℕ→∣-helper _ _ c h (abs→⊎ _ _ refl) (abs→⊎ _ _ refl) ∣) +∣ℕ→∣ = Prop.rec isProp∣ (λ (c , h) → ∣ ∣ℕ→∣-helper _ _ c h (abs→⊎ _ _ refl) (abs→⊎ _ _ refl) ∣₁) ¬∣→¬∣ℕ : ¬ m ∣ n → ¬ abs m ∣ℕ abs n ¬∣→¬∣ℕ p q = p (∣ℕ→∣ q) @@ -294,7 +293,7 @@ module _ (m n : ℤ)(qr : QuotRem m n) where rem≡0→m∣n : qr .rem ≡ 0 → m ∣ n - rem≡0→m∣n p = ∣ qr .div , (λ i → qr .div · m + p (~ i)) ∙ sym (qr .quotEq) ∣ + rem≡0→m∣n p = ∣ qr .div , (λ i → qr .div · m + p (~ i)) ∙ sym (qr .quotEq) ∣₁ m∣n→rem≡0 : m ∣ n → qr .rem ≡ 0 m∣n→rem≡0 p = diff --git a/Cubical/Data/Int/MoreInts/DeltaInt/Properties.agda b/Cubical/Data/Int/MoreInts/DeltaInt/Properties.agda index f3f0119f56..869d8c3052 100644 --- a/Cubical/Data/Int/MoreInts/DeltaInt/Properties.agda +++ b/Cubical/Data/Int/MoreInts/DeltaInt/Properties.agda @@ -30,7 +30,8 @@ isSetDeltaInt : isSet DeltaInt module Cubical.Data.Int.MoreInts.DeltaInt.Properties where -open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism open import Cubical.Data.Nat hiding (zero) open import Cubical.Data.Int hiding (abs; _+_) open import Cubical.Data.Int.MoreInts.DeltaInt.Base @@ -83,4 +84,3 @@ predSucc (cancel a b i) = succPred (cancel a b i) succPredEq : DeltaInt ≡ DeltaInt succPredEq = isoToPath (iso succ pred succPred predSucc) - diff --git a/Cubical/Data/List/Dependent.agda b/Cubical/Data/List/Dependent.agda index 1ee3b7c85d..17feda1d7c 100644 --- a/Cubical/Data/List/Dependent.agda +++ b/Cubical/Data/List/Dependent.agda @@ -1,6 +1,13 @@ {-# OPTIONS --safe #-} -open import Cubical.Foundations.Everything renaming (Iso to _≅_) +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism renaming (Iso to _≅_) +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport + open import Cubical.Data.List open import Cubical.Data.Unit open import Cubical.Data.Prod hiding (map) @@ -73,4 +80,3 @@ mapOverIdfun-∘ : ∀ {ℓA ℓB ℓB' ℓB''} {A : Type ℓA} {B : A → Type → mapOverIdfun (λ a → h a ∘ g a) as ≡ mapOverIdfun h as ∘ mapOverIdfun g as mapOverIdfun-∘ h g [] i [] = [] mapOverIdfun-∘ h g (a ∷ as) i (b ∷ bs) = h a (g a b) ∷ mapOverIdfun-∘ h g as i bs - diff --git a/Cubical/Data/Nat/Coprime.agda b/Cubical/Data/Nat/Coprime.agda index d551ae28c3..58fc21c4f4 100644 --- a/Cubical/Data/Nat/Coprime.agda +++ b/Cubical/Data/Nat/Coprime.agda @@ -42,8 +42,8 @@ module ToCoprime ((m , n) : ℕ × ℕ₊₁) where 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₂ ∣)) + 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, then the constant c (s.t. c · m ≡ n) is also > 0 ∣s-untrunc : m ∣ suc n → Σ[ c ∈ ℕ ] (suc c) · m ≡ suc n diff --git a/Cubical/Data/Queue/1List.agda b/Cubical/Data/Queue/1List.agda index 86b2ff1e9b..e7b61d1fe5 100644 --- a/Cubical/Data/Queue/1List.agda +++ b/Cubical/Data/Queue/1List.agda @@ -1,7 +1,12 @@ {-# OPTIONS --no-exact-split --safe #-} module Cubical.Data.Queue.1List where -open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Structure + open import Cubical.Structures.Queue diff --git a/Cubical/Data/Queue/Finite.agda b/Cubical/Data/Queue/Finite.agda index 3b01a30d87..87f5a5532f 100644 --- a/Cubical/Data/Queue/Finite.agda +++ b/Cubical/Data/Queue/Finite.agda @@ -1,7 +1,8 @@ {-# OPTIONS --no-exact-split --safe #-} module Cubical.Data.Queue.Finite where -open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function open import Cubical.Foundations.SIP open import Cubical.Structures.Queue diff --git a/Cubical/Data/Queue/Truncated2List.agda b/Cubical/Data/Queue/Truncated2List.agda index 1245ba860d..addb596659 100644 --- a/Cubical/Data/Queue/Truncated2List.agda +++ b/Cubical/Data/Queue/Truncated2List.agda @@ -1,7 +1,11 @@ {-# OPTIONS --no-exact-split --safe #-} module Cubical.Data.Queue.Truncated2List where -open import Cubical.Foundations.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.SIP open import Cubical.Structures.Queue diff --git a/Cubical/Data/Queue/Untruncated2List.agda b/Cubical/Data/Queue/Untruncated2List.agda index 749331ae5b..df363ebc3b 100644 --- a/Cubical/Data/Queue/Untruncated2List.agda +++ b/Cubical/Data/Queue/Untruncated2List.agda @@ -1,7 +1,10 @@ {-# OPTIONS --no-exact-split --safe #-} module Cubical.Data.Queue.Untruncated2List where -open import Cubical.Foundations.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.SIP open import Cubical.Structures.Queue diff --git a/Cubical/Data/Queue/Untruncated2ListInvariant.agda b/Cubical/Data/Queue/Untruncated2ListInvariant.agda index d292413e7d..3e7e92636c 100644 --- a/Cubical/Data/Queue/Untruncated2ListInvariant.agda +++ b/Cubical/Data/Queue/Untruncated2ListInvariant.agda @@ -1,7 +1,10 @@ {-# OPTIONS --safe #-} module Cubical.Data.Queue.Untruncated2ListInvariant where -open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels + open import Cubical.Data.Empty as ⊥ open import Cubical.Data.List open import Cubical.Data.Maybe diff --git a/Cubical/Data/Sigma/Base.agda b/Cubical/Data/Sigma/Base.agda index 8197759d20..88f051a356 100644 --- a/Cubical/Data/Sigma/Base.agda +++ b/Cubical/Data/Sigma/Base.agda @@ -29,7 +29,7 @@ infixr 5 _×_ -- Mere existence ∃ : ∀ {ℓ ℓ'} (A : Type ℓ) (B : A → Type ℓ') → Type (ℓ-max ℓ ℓ') -∃ A B = ∥ Σ A B ∥ +∃ A B = ∥ Σ A B ∥₁ infix 2 ∃-syntax diff --git a/Cubical/Data/SumFin/Properties.agda b/Cubical/Data/SumFin/Properties.agda index 36e0bfb4b1..ae453f7166 100644 --- a/Cubical/Data/SumFin/Properties.agda +++ b/Cubical/Data/SumFin/Properties.agda @@ -92,19 +92,19 @@ isNotZero : ℕ → ℕ isNotZero 0 = 0 isNotZero (suc n) = 1 -SumFin∥∥≃ : (n : ℕ) → ∥ Fin n ∥ ≃ Fin (isNotZero n) +SumFin∥∥≃ : (n : ℕ) → ∥ Fin n ∥₁ ≃ Fin (isNotZero n) SumFin∥∥≃ 0 = propTruncIdempotent≃ (isProp⊥) SumFin∥∥≃ (suc n) = - isContr→≃Unit (inhProp→isContr ∣ inl tt ∣ isPropPropTrunc) + isContr→≃Unit (inhProp→isContr ∣ inl tt ∣₁ isPropPropTrunc) ⋆ isContr→≃Unit (isContrUnit) ⋆ invEquiv (⊎-⊥-≃) ℕ→Bool : ℕ → Bool ℕ→Bool 0 = false ℕ→Bool (suc n) = true -SumFin∥∥DecProp : (n : ℕ) → ∥ Fin n ∥ ≃ Bool→Type (ℕ→Bool n) +SumFin∥∥DecProp : (n : ℕ) → ∥ Fin n ∥₁ ≃ Bool→Type (ℕ→Bool n) SumFin∥∥DecProp 0 = uninhabEquiv (Prop.rec isProp⊥ ⊥.rec) ⊥.rec -SumFin∥∥DecProp (suc n) = isContr→≃Unit (inhProp→isContr ∣ inl tt ∣ isPropPropTrunc) +SumFin∥∥DecProp (suc n) = isContr→≃Unit (inhProp→isContr ∣ inl tt ∣₁ isPropPropTrunc) -- negation of SumFin @@ -179,11 +179,11 @@ SumFin∃← (suc n) f = ∘ map-⊎ (invEq (ΣUnit (Bool→Type ∘ f ∘ inl))) (SumFin∃← n (f ∘ inr)) ∘ Bool→Type⊎ _ _ -SumFin∃≃ : (n : ℕ)(f : Fin n → Bool) → ∥ Σ (Fin n) (Bool→Type ∘ f) ∥ ≃ Bool→Type (trueForSome n f) +SumFin∃≃ : (n : ℕ)(f : Fin n → Bool) → ∥ Σ (Fin n) (Bool→Type ∘ f) ∥₁ ≃ Bool→Type (trueForSome n f) SumFin∃≃ n f = propBiimpl→Equiv isPropPropTrunc isPropBool→Type (Prop.rec isPropBool→Type (SumFin∃→ n f)) - (∣_∣ ∘ SumFin∃← n f) + (∣_∣₁ ∘ SumFin∃← n f) SumFin∀≃ : (n : ℕ)(f : Fin n → Bool) → ((x : Fin n) → Bool→Type (f x)) ≃ Bool→Type (trueForAll n f) SumFin∀≃ 0 _ = isContr→≃Unit (isContrΠ⊥) @@ -222,7 +222,7 @@ DecFin (suc n) = yes fzero -- propositional truncation of Fin -Dec∥Fin∥ : (n : ℕ) → Dec ∥ Fin n ∥ +Dec∥Fin∥ : (n : ℕ) → Dec ∥ Fin n ∥₁ Dec∥Fin∥ n = Dec∥∥ (DecFin n) -- some properties about cardinality diff --git a/Cubical/Experiments/Brunerie.agda b/Cubical/Experiments/Brunerie.agda index 4a3ea9138e..18f3258141 100644 --- a/Cubical/Experiments/Brunerie.agda +++ b/Cubical/Experiments/Brunerie.agda @@ -1,10 +1,17 @@ {-# OPTIONS --safe #-} module Cubical.Experiments.Brunerie where -open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence + open import Cubical.Data.Bool open import Cubical.Data.Nat open import Cubical.Data.Int + open import Cubical.HITs.S1 hiding (encode) open import Cubical.HITs.S2 open import Cubical.HITs.S3 @@ -14,6 +21,7 @@ open import Cubical.HITs.GroupoidTruncation as GroupoidTrunc open import Cubical.HITs.2GroupoidTruncation as 2GroupoidTrunc open import Cubical.HITs.Truncation as Trunc open import Cubical.HITs.Susp renaming (toSusp to σ) + open import Cubical.Homotopy.Loopspace open import Cubical.Homotopy.Hopf open S¹Hopf diff --git a/Cubical/Experiments/EscardoSIP.agda b/Cubical/Experiments/EscardoSIP.agda index 56f3acfa6d..217444ea56 100644 --- a/Cubical/Experiments/EscardoSIP.agda +++ b/Cubical/Experiments/EscardoSIP.agda @@ -8,9 +8,15 @@ It would be interesting to compare the proves with the one in Cubical.Foundation {-# OPTIONS --safe #-} module Cubical.Experiments.EscardoSIP where -open import Cubical.Core.Everything -open import Cubical.Foundations.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.Equiv.HalfAdjoint +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Univalence + open import Cubical.Data.Sigma.Properties private diff --git a/Cubical/Experiments/FunExtFromUA.agda b/Cubical/Experiments/FunExtFromUA.agda index 2cf2ae5cf8..3c061efc73 100644 --- a/Cubical/Experiments/FunExtFromUA.agda +++ b/Cubical/Experiments/FunExtFromUA.agda @@ -4,7 +4,11 @@ module Cubical.Experiments.FunExtFromUA where -open import Cubical.Foundations.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.Univalence variable ℓ ℓ' : Level diff --git a/Cubical/Experiments/HoTT-UF.agda b/Cubical/Experiments/HoTT-UF.agda index d5c7d169a8..ecf83854f1 100644 --- a/Cubical/Experiments/HoTT-UF.agda +++ b/Cubical/Experiments/HoTT-UF.agda @@ -48,8 +48,8 @@ open import Cubical.Foundations.Id public ; _≃_ -- The type of equivalences between two given types. ; EquivContr -- A formulation of univalence. - ; ∥_∥ -- Propositional truncation. - ; ∣_∣ -- Map into the propositional truncation. + ; ∥_∥₁ -- Propositional truncation. + ; ∣_∣₁ -- Map into the propositional truncation. ; ∥∥-isProp -- A truncated type is a proposition. ; ∥∥-recursion -- Non-dependent elimination. ; ∥∥-induction -- Dependent elimination. diff --git a/Cubical/Experiments/ZCohomology/Benchmarks.agda b/Cubical/Experiments/ZCohomology/Benchmarks.agda index 6d3a818bb8..cf4a048e9c 100644 --- a/Cubical/Experiments/ZCohomology/Benchmarks.agda +++ b/Cubical/Experiments/ZCohomology/Benchmarks.agda @@ -38,7 +38,7 @@ open import Cubical.Homotopy.Hopf open S¹Hopf open import Cubical.HITs.Truncation open import Cubical.HITs.Susp -open import Cubical.HITs.S1 hiding (rec ; elim ; ind) +open import Cubical.HITs.S1 open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms diff --git a/Cubical/Experiments/ZCohomologyOld/Properties.agda b/Cubical/Experiments/ZCohomologyOld/Properties.agda index 7d1e8dda82..faaadde8c8 100644 --- a/Cubical/Experiments/ZCohomologyOld/Properties.agda +++ b/Cubical/Experiments/ZCohomologyOld/Properties.agda @@ -27,7 +27,7 @@ open import Cubical.Algebra.Group open import Cubical.Algebra.Group.DirProd open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Int open import Cubical.Algebra.Semigroup open import Cubical.Algebra.Monoid diff --git a/Cubical/Foundations/Everything.agda b/Cubical/Foundations/Everything.agda index 8810a1338e..48d7e33aa3 100644 --- a/Cubical/Foundations/Everything.agda +++ b/Cubical/Foundations/Everything.agda @@ -22,8 +22,8 @@ open import Cubical.Foundations.Id ; isEquiv to isEquivId ; equivIsEquiv to equivIsEquivId ; refl to reflId - ; ∥_∥ to propTruncId - ; ∣_∣ to incId + ; ∥_∥₁ to propTruncId + ; ∣_∣₁ to incId ; isPropIsContr to isPropIsContrId ; isPropIsEquiv to isPropIsEquivId ) diff --git a/Cubical/Foundations/Id.agda b/Cubical/Foundations/Id.agda index 505b0f0c81..56c82652a6 100644 --- a/Cubical/Foundations/Id.agda +++ b/Cubical/Foundations/Id.agda @@ -8,7 +8,7 @@ This file contains: - Lemmas for going back and forth between Path and Id -- Function extensionality for Id +- Function extesionality for Id - fiber, isContr, equiv all defined using Id @@ -49,9 +49,7 @@ open import Cubical.Foundations.Univalence renaming ( EquivContr to EquivContrPath ) open import Cubical.Foundations.Isomorphism open import Cubical.HITs.PropositionalTruncation public - renaming ( squash to squashPath - ; rec to recPropTruncPath - ; elim to elimPropTruncPath ) + renaming (rec to recPropTruncPath ; elim to elimPropTruncPath ) open import Cubical.Core.Id public using (Id; ⟨_,_⟩; faceId; pathId; elimId; _≡_) @@ -274,14 +272,14 @@ EquivContr {ℓ = ℓ} A = helper1 f1 f2 f12 (EquivContrPath A) -- Propositional truncation -∥∥-isProp : ∀ (x y : ∥ A ∥) → x ≡ y -∥∥-isProp x y = pathToId (squashPath x y) +∥∥-isProp : ∀ (x y : ∥ A ∥₁) → x ≡ y +∥∥-isProp x y = pathToId (squash₁ x y) -∥∥-recursion : ∀ {A : Type ℓ} {P : Type ℓ} → isProp P → (A → P) → ∥ A ∥ → P +∥∥-recursion : ∀ {A : Type ℓ} {P : Type ℓ} → isProp P → (A → P) → ∥ A ∥₁ → P ∥∥-recursion Pprop f x = recPropTruncPath (isPropToIsPropPath Pprop) f x -∥∥-induction : ∀ {A : Type ℓ} {P : ∥ A ∥ → Type ℓ} → ((a : ∥ A ∥) → isProp (P a)) → - ((x : A) → P ∣ x ∣) → (a : ∥ A ∥) → P 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 diff --git a/Cubical/Foundations/RelationalStructure.agda b/Cubical/Foundations/RelationalStructure.agda index 2b8320fce5..5464dbb2d1 100644 --- a/Cubical/Foundations/RelationalStructure.agda +++ b/Cubical/Foundations/RelationalStructure.agda @@ -140,10 +140,10 @@ strRelQuotientComparison {ρ = ρ} θ α R (eq/ s t r i) = ρs = subst (λ R' → ρ R' s s) (funExt₂ λ x x' → - hPropExt squash (R .fst .snd 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' ∣)) + (λ 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 @@ -248,7 +248,7 @@ structuredQER→structuredEquiv {ρ = ρ} θ (X , s) (Y , t) R r .rel = correction : [R] .fst ≡ graphRel (E.Thm .fst) correction = funExt₂ λ qx qy → - (hPropExt squash (squash/ _ _) + (hPropExt squash₁ (squash/ _ _) (Trunc.rec (squash/ _ _) (λ {(y , qr , py) → Trunc.rec @@ -259,10 +259,10 @@ structuredQER→structuredEquiv {ρ = ρ} θ (X , s) (Y , t) R r .rel = ∙ py}) qr})) (elimProp {P = λ qx → E.Thm .fst qx ≡ qy → [R] .fst qx qy} - (λ _ → isPropΠ λ _ → squash) + (λ _ → isPropΠ λ _ → squash₁) (λ x → elimProp {P = λ qy → E.Thm .fst [ x ] ≡ qy → [R] .fst [ x ] qy} - (λ _ → isPropΠ λ _ → squash) - (λ y p → ∣ _ , ∣ _ , refl , E.fwd≡ToRel p ∣ , refl ∣) + (λ _ → isPropΠ λ _ → squash₁) + (λ y p → ∣ _ , ∣ _ , refl , E.fwd≡ToRel p ∣₁ , refl ∣₁) qy) qx)) diff --git a/Cubical/Functions/Bundle.agda b/Cubical/Functions/Bundle.agda index 98638451f0..2758e7da42 100644 --- a/Cubical/Functions/Bundle.agda +++ b/Cubical/Functions/Bundle.agda @@ -42,9 +42,9 @@ module _ {ℓb ℓf} (B : Type ℓb) (ℓ : Level) (F : Type ℓf) where 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 : (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 ∥ + 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/Logic.agda b/Cubical/Functions/Logic.agda index 3e5439c208..19e9b2e52a 100644 --- a/Cubical/Functions/Logic.agda +++ b/Cubical/Functions/Logic.agda @@ -60,7 +60,7 @@ infix 2 ⇒∶_⇐∶_ infix 2 ⇐∶_⇒∶_ ∥_∥ₚ : Type ℓ → hProp ℓ -∥ A ∥ₚ = ∥ A ∥ , isPropPropTrunc +∥ A ∥ₚ = ∥ A ∥₁ , isPropPropTrunc _≡ₚ_ : (x y : A) → hProp _ x ≡ₚ y = ∥ x ≡ y ∥ₚ @@ -120,16 +120,16 @@ x ≢ₚ y = ¬ x ≡ₚ y -- Disjunction of mere propositions _⊔′_ : Type ℓ → Type ℓ' → Type _ -A ⊔′ B = ∥ A ⊎ B ∥ +A ⊔′ B = ∥ A ⊎ B ∥₁ _⊔_ : hProp ℓ → hProp ℓ' → hProp _ P ⊔ Q = ∥ ⟨ P ⟩ ⊎ ⟨ Q ⟩ ∥ₚ inl : A → A ⊔′ B -inl x = ∣ ⊎.inl x ∣ +inl x = ∣ ⊎.inl x ∣₁ inr : B → A ⊔′ B -inr x = ∣ ⊎.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 ⟩) @@ -194,7 +194,7 @@ Decₚ P = Dec ⟨ P ⟩ , isPropDec (isProp⟨⟩ P) ∥¬A∥≡¬∥A∥ _ = ⇒∶ (λ ¬A A → PropTrunc.elim (λ _ → ⊥.isProp⊥) (PropTrunc.elim (λ _ → isPropΠ λ _ → ⊥.isProp⊥) (λ ¬p p → ¬p p) ¬A) A) - ⇐∶ λ ¬p → ∣ (λ p → ¬p ∣ p ∣) ∣ + ⇐∶ λ ¬p → ∣ (λ p → ¬p ∣ p ∣₁) ∣₁ -------------------------------------------------------------------------------- -- (hProp, ⊔, ⊥) is a bounded ⊔-semilattice @@ -209,11 +209,11 @@ Decₚ P = Dec ⟨ P ⟩ , isPropDec (isProp⟨⟩ P) ⇐∶ assoc2 where assoc2 : (A ⊔′ B) ⊔′ C → A ⊔′ (B ⊔′ C) - 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 + 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 = @@ -267,8 +267,8 @@ Decₚ P = Dec ⟨ P ⟩ , isPropDec (isProp⟨⟩ P) → P ⊓ (Q ⊔ R) ≡ (P ⊓ Q) ⊔ (P ⊓ R) ⊓-⊔-distribˡ P Q R = ⇒∶ (λ { (x , a) → ⊔-elim Q R (λ _ → (P ⊓ Q) ⊔ (P ⊓ R)) - (λ y → ∣ ⊎.inl (x , y) ∣ ) - (λ z → ∣ ⊎.inr (x , z) ∣ ) a }) + (λ y → ∣ ⊎.inl (x , y) ∣₁ ) + (λ z → ∣ ⊎.inr (x , z) ∣₁ ) a }) ⇐∶ ⊔-elim (P ⊓ Q) (P ⊓ R) (λ _ → P ⊓ Q ⊔ R) (λ y → fst y , inl (snd y)) diff --git a/Cubical/Functions/Surjection.agda b/Cubical/Functions/Surjection.agda index 1c833049ad..938d8849ec 100644 --- a/Cubical/Functions/Surjection.agda +++ b/Cubical/Functions/Surjection.agda @@ -20,19 +20,19 @@ private f : A → B isSurjection : (A → B) → Type _ -isSurjection f = ∀ b → ∥ fiber f b ∥ +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 ∣ +section→isSurjection {g = g} s b = ∣ g b , s b ∣₁ isPropIsSurjection : isProp (isSurjection f) -isPropIsSurjection = isPropΠ λ _ → squash +isPropIsSurjection = isPropΠ λ _ → squash₁ isEquiv→isSurjection : isEquiv f → isSurjection f -isEquiv→isSurjection e b = ∣ fst (equiv-proof e b) ∣ +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 @@ -44,7 +44,7 @@ equiv-proof (isEmbedding×isSurjection→isEquiv {f = f} (emb , sur)) b = hpf : hasPropFibers f hpf = isEmbedding→hasPropFibers emb - fib : ∥ fiber f b ∥ + fib : ∥ fiber f b ∥₁ fib = sur b fib' : isProp (fiber f b) @@ -68,12 +68,12 @@ rightCancellable {ℓ} {A} {ℓ'} {B} f = ∀ {C : Type (ℓ-suc (ℓ-max ℓ 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 ∥ + hasPreimage f y = ∥ fiber f y ∥₁ fact₁ : ∀ x → Unit* ≡ hasPreimage f (f x) fact₁ x = hPropExt isPropUnit* isPropPropTrunc - (λ _ → ∣ (x , refl) ∣) + (λ _ → ∣ (x , refl) ∣₁) (λ _ → tt*) fact₂ : ∀ y → Unit* ≡ hasPreimage f y diff --git a/Cubical/HITs/AssocList/Properties.agda b/Cubical/HITs/AssocList/Properties.agda index 2de9f19927..95d8d7337b 100644 --- a/Cubical/HITs/AssocList/Properties.agda +++ b/Cubical/HITs/AssocList/Properties.agda @@ -1,16 +1,26 @@ {-# 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.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.Univalence open import Cubical.Foundations.SIP -open import Cubical.HITs.FiniteMultiset as FMS - hiding (_++_; unitl-++; unitr-++; assoc-++; cons-++; comm-++) -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 +open import Cubical.Structures.MultiSet + +open import Cubical.Data.Nat using (ℕ; zero; suc; _+_; +-assoc; isSetℕ) + +open import Cubical.HITs.AssocList.Base as AL +open import Cubical.HITs.FiniteMultiset as FMS + hiding (_++_; unitl-++; unitr-++; assoc-++; cons-++; comm-++) + private variable ℓ : Level diff --git a/Cubical/HITs/Bouquet/FundamentalGroupProof.agda b/Cubical/HITs/Bouquet/FundamentalGroupProof.agda index b34a4ef97a..d9619247ca 100644 --- a/Cubical/HITs/Bouquet/FundamentalGroupProof.agda +++ b/Cubical/HITs/Bouquet/FundamentalGroupProof.agda @@ -183,17 +183,17 @@ left-homotopy l = decodeEncode base l -- Truncated proofs of right homotopy of winding/looping functions -truncatedPathEquality : (g : FreeGroupoid A) → ∥ cong code (looping g) ≡ ua (equivs g) ∥ +truncatedPathEquality : (g : FreeGroupoid A) → ∥ cong code (looping g) ≡ ua (equivs g) ∥₁ truncatedPathEquality = elimProp Bprop - (λ a → ∣ η-ind a ∣) - (λ g1 g2 → λ ∣ind1∣ ∣ind2∣ → rec2 squash (λ ind1 ind2 → ∣ ·-ind g1 g2 ind1 ind2 ∣) ∣ind1∣ ∣ind2∣) - ∣ ε-ind ∣ - (λ g → λ ∣ind∣ → propRec squash (λ ind → ∣ inv-ind g ind ∣) ∣ind∣) where + (λ a → ∣ η-ind a ∣₁) + (λ g1 g2 → λ ∣ind1∣₁ ∣ind2∣₁ → rec2 squash₁ (λ ind1 ind2 → ∣ ·-ind g1 g2 ind1 ind2 ∣₁) ∣ind1∣₁ ∣ind2∣₁) + ∣ ε-ind ∣₁ + (λ g → λ ∣ind∣₁ → propRec squash₁ (λ ind → ∣ inv-ind g ind ∣₁) ∣ind∣₁) where B : ∀ g → Type _ B g = cong code (looping g) ≡ ua (equivs g) - Bprop : ∀ g → isProp ∥ B g ∥ - Bprop g = squash + Bprop : ∀ g → isProp ∥ B g ∥₁ + Bprop g = squash₁ η-ind : ∀ a → B (η a) η-ind a = refl ·-ind : ∀ g1 g2 → B g1 → B g2 → B (g1 · g2) @@ -218,10 +218,10 @@ truncatedPathEquality = elimProp ≡⟨ sym (invPathsInUNaturality g) ⟩ ua (equivs (inv g)) ∎ -truncatedRight-homotopy : (g : FreeGroupoid A) → ∥ winding (looping g) ≡ g ∥ -truncatedRight-homotopy g = propRec squash recursion (truncatedPathEquality g) where - recursion : cong code (looping g) ≡ ua (equivs g) → ∥ winding (looping g) ≡ g ∥ - recursion hyp = ∣ aux ∣ where +truncatedRight-homotopy : (g : FreeGroupoid A) → ∥ winding (looping g) ≡ g ∥₁ +truncatedRight-homotopy g = propRec squash₁ recursion (truncatedPathEquality g) where + recursion : cong code (looping g) ≡ ua (equivs g) → ∥ winding (looping g) ≡ g ∥₁ + recursion hyp = ∣ aux ∣₁ where aux : winding (looping g) ≡ g aux = winding (looping g) @@ -237,13 +237,13 @@ right-homotopyInTruncatedGroupoid g = Iso.inv PathIdTrunc₀Iso (truncatedRight- -- Truncated encodeDecode over all fibrations -truncatedEncodeDecode : (x : Bouquet A) → (g : code x) → ∥ encode x (decode x g) ≡ g ∥ +truncatedEncodeDecode : (x : Bouquet A) → (g : code x) → ∥ encode x (decode x g) ≡ g ∥₁ truncatedEncodeDecode base = truncatedRight-homotopy truncatedEncodeDecode (loop a i) = isProp→PathP prop truncatedRight-homotopy truncatedRight-homotopy i where - prop : ∀ i → isProp (∀ (g : code (loop a i)) → ∥ encode (loop a i) (decode (loop a i) g) ≡ g ∥) + prop : ∀ i → isProp (∀ (g : code (loop a i)) → ∥ encode (loop a i) (decode (loop a i) g) ≡ g ∥₁) prop i f g = funExt pointwise where - pointwise : (x : code (loop a i)) → PathP (λ _ → ∥ encode (loop a i) (decode (loop a i) x) ≡ x ∥) (f x) (g x) - pointwise x = isProp→PathP (λ i → squash) (f x) (g x) + pointwise : (x : code (loop a i)) → PathP (λ _ → ∥ encode (loop a i) (decode (loop a i) x) ≡ x ∥₁) (f x) (g x) + pointwise x = isProp→PathP (λ i → squash₁) (f x) (g x) encodeDecodeInTruncatedGroupoid : (x : Bouquet A) → (g : code x) → ∣ encode x (decode x g) ∣₂ ≡ ∣ g ∣₂ encodeDecodeInTruncatedGroupoid x g = Iso.inv PathIdTrunc₀Iso (truncatedEncodeDecode x g) diff --git a/Cubical/HITs/Cost/Base.agda b/Cubical/HITs/Cost/Base.agda index 483d1bc7f7..6f6e38f717 100644 --- a/Cubical/HITs/Cost/Base.agda +++ b/Cubical/HITs/Cost/Base.agda @@ -13,11 +13,11 @@ private A B C : Type ℓ Cost : (A : Type ℓ) → Type ℓ -Cost A = A × ∥ ℕ ∥ +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 +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 @@ -28,7 +28,7 @@ _>>=_ : Cost A → (A → Cost B) → Cost B ... | (y , n) = (y , map suc (map2 _+_ m n)) return : A → Cost A -return x = (x , ∣ 0 ∣) +return x = (x , ∣ 0 ∣₁) -- The monad laws all hold by Cost≡ @@ -62,11 +62,11 @@ add (suc x) y = do private -- addBad x y will do x + y calls - _ : addBad 3 5 ≡ (8 , ∣ 8 ∣) + _ : addBad 3 5 ≡ (8 , ∣ 8 ∣₁) _ = refl -- add x y will only do x recursive calls - _ : add 3 5 ≡ (8 , ∣ 3 ∣) + _ : add 3 5 ≡ (8 , ∣ 3 ∣₁) _ = refl -- Despite addBad and add having different cost we can still prove @@ -102,16 +102,16 @@ fibTail (suc x) = fibAux x 1 0 private - _ : fib 10 ≡ (55 , ∣ 176 ∣) + _ : fib 10 ≡ (55 , ∣ 176 ∣₁) _ = refl - _ : fibTail 10 ≡ (55 , ∣ 9 ∣) + _ : fibTail 10 ≡ (55 , ∣ 9 ∣₁) _ = refl - _ : fib 20 ≡ (6765 , ∣ 21890 ∣) + _ : fib 20 ≡ (6765 , ∣ 21890 ∣₁) _ = refl - _ : fibTail 20 ≡ (6765 , ∣ 19 ∣) + _ : fibTail 20 ≡ (6765 , ∣ 19 ∣₁) _ = refl -- Exercise left to the reader: prove that fib and fibTail are equal functions diff --git a/Cubical/HITs/CumulativeHierarchy/Base.agda b/Cubical/HITs/CumulativeHierarchy/Base.agda index 1be68eeeea..f16be78c74 100644 --- a/Cubical/HITs/CumulativeHierarchy/Base.agda +++ b/Cubical/HITs/CumulativeHierarchy/Base.agda @@ -32,8 +32,8 @@ data V (ℓ : Level) : Type (ℓ-suc ℓ) _∈_ : (S T : V ℓ) → hProp (ℓ-suc ℓ) eqImage : {X Y : Type ℓ} (ix : X → V ℓ) (iy : Y → V ℓ) → Type (ℓ-suc ℓ) -eqImage {X = X} {Y = Y} ix iy = (∀ (a : X) → ∥ fiber iy (ix a) ∥) ⊓′ - (∀ (b : Y) → ∥ fiber ix (iy b) ∥) +eqImage {X = X} {Y = Y} ix iy = (∀ (a : X) → ∥ fiber iy (ix a) ∥₁) ⊓′ + (∀ (b : Y) → ∥ fiber ix (iy b) ∥₁) data V ℓ where sett : (X : Type ℓ) → (X → V ℓ) → V ℓ @@ -43,8 +43,8 @@ data V ℓ where A ∈ sett X ix = ∥ Σ[ i ∈ X ] (ix i ≡ A) ∥ₚ A ∈ seteq X Y ix iy (f , g) i = ⇔toPath {P = A ∈ sett X ix} {Q = A ∈ sett Y iy} - (λ ax → do (x , xa) ← ax ; (y , ya) ← f x ; ∣ y , ya ∙ xa ∣) - (λ ay → do (y , ya) ← ay ; (x , xa) ← g y ; ∣ x , xa ∙ ya ∣) i + (λ ax → do (x , xa) ← ax ; (y , ya) ← f x ; ∣ y , ya ∙ xa ∣₁) + (λ ay → do (y , ya) ← ay ; (x , xa) ← g y ; ∣ x , xa ∙ ya ∣₁) i where open PropMonad A ∈ setIsSet a b p q i j = isSetHProp (A ∈ a) (A ∈ b) (λ j → A ∈ p j) (λ j → A ∈ q j) i j @@ -84,16 +84,16 @@ module _ {Z : (s : V ℓ) → Type ℓ'} {isSetZ : ∀ s → isSet (Z s)} (E : E -- using a local definition of Prop.rec satisfies the termination checker rec₁→₂ x₁ = localRec₁ (eq .fst x₁) where localRec₁ : - ∥ fiber ix₂ (ix₁ x₁) ∥ + ∥ fiber ix₂ (ix₁ x₁) ∥₁ → ∃[ (x₂ , p) ∈ fiber ix₂ (ix₁ x₁) ] PathP (λ i → Z (p i)) (elim (ix₂ x₂)) (elim (ix₁ x₁)) - localRec₁ ∣ x₂ , xx ∣ = ∣ (x₂ , xx) , (λ i → elim (xx i)) ∣ - localRec₁ (squash x y i) = squash (localRec₁ x) (localRec₁ y) i + localRec₁ ∣ x₂ , xx ∣₁ = ∣ (x₂ , xx) , (λ i → elim (xx i)) ∣₁ + localRec₁ (squash₁ x y i) = squash₁ (localRec₁ x) (localRec₁ y) i rec₂→₁ x₂ = localRec₂ (eq .snd x₂) where localRec₂ : - ∥ fiber ix₁ (ix₂ x₂) ∥ + ∥ fiber ix₁ (ix₂ x₂) ∥₁ → ∃[ (x₁ , p) ∈ fiber ix₁ (ix₂ x₂) ] PathP (λ i → Z (p i)) (elim (ix₁ x₁)) (elim (ix₂ x₂)) - localRec₂ ∣ x₁ , xx ∣ = ∣ (x₁ , xx) , (λ i → elim (xx i)) ∣ - localRec₂ (squash x y i) = squash (localRec₂ x) (localRec₂ y) i + localRec₂ ∣ x₁ , xx ∣₁ = ∣ (x₁ , xx) , (λ i → elim (xx i)) ∣₁ + localRec₂ (squash₁ x y i) = squash₁ (localRec₂ x) (localRec₂ y) i elim (setIsSet S T x y i j) = isProp→PathP propPathP (cong elim x) (cong elim y) i j where propPathP : (i : I) → isProp (PathP (λ j → Z (setIsSet S T x y i j)) (elim S) (elim T)) propPathP _ = subst isProp (sym (PathP≡Path _ _ _)) (isSetZ _ _ _) @@ -188,8 +188,8 @@ module _ {Z : (s t : V ℓ) → Type ℓ'} {isSetZ : ∀ s t → isSet (Z s t)} ∀ (y₂ : Y₂) → ∃[ (y₁ , p) ∈ fiber iy₁ (iy₂ y₂) ] PathP (λ i → ∀ x → Z (ix x) (p i)) (λ x → rec x (iy₁ y₁)) (λ x → rec x (iy₂ y₂)) - rec₁→₂ y₁ = do (y₂ , yy) ← fst eq y₁ ; ∣ (y₂ , yy) , (λ i x → rec x (yy i)) ∣ - rec₂→₁ y₂ = do (y₁ , yy) ← snd eq y₂ ; ∣ (y₁ , yy) , (λ i x → rec x (yy i)) ∣ + rec₁→₂ y₁ = do (y₂ , yy) ← fst eq y₁ ; ∣ (y₂ , yy) , (λ i x → rec x (yy i)) ∣₁ + rec₂→₁ y₂ = do (y₁ , yy) ← snd eq y₂ ; ∣ (y₁ , yy) , (λ i x → rec x (yy i)) ∣₁ elimImplS : ∀ (X : Type ℓ) (ix : X → V ℓ) @@ -230,8 +230,8 @@ module _ {Z : (s t : V ℓ) → Type ℓ'} {isSetZ : ∀ s t → isSet (Z s t)} ∀ (x₂ : X₂) → ∃[ (x₁ , p) ∈ fiber ix₁ (ix₂ x₂) ] PathP (λ i → ∀ y → Z (p i) (iy y)) (λ y → rec₁ x₁ (iy y)) (λ y → rec₂ x₂ (iy y)) - rec₁→₂Impl x₁ = do ((x₂ , xx) , rx) ← rec₁→₂ x₁ ; ∣ (x₂ , xx) , (λ i y → rx i (iy y)) ∣ - rec₂→₁Impl x₂ = do ((x₁ , xx) , rx) ← rec₂→₁ x₂ ; ∣ (x₁ , xx) , (λ i y → rx i (iy y)) ∣ + rec₁→₂Impl x₁ = do ((x₂ , xx) , rx) ← rec₁→₂ x₁ ; ∣ (x₂ , xx) , (λ i y → rx i (iy y)) ∣₁ + rec₂→₁Impl x₂ = do ((x₁ , xx) , rx) ← rec₂→₁ x₂ ; ∣ (x₁ , xx) , (λ i y → rx i (iy y)) ∣₁ pElim : ElimSet isSetPElim ElimSett pElim = elimImplS diff --git a/Cubical/HITs/CumulativeHierarchy/Constructions.agda b/Cubical/HITs/CumulativeHierarchy/Constructions.agda index 4b3d841426..03336f5c7a 100644 --- a/Cubical/HITs/CumulativeHierarchy/Constructions.agda +++ b/Cubical/HITs/CumulativeHierarchy/Constructions.agda @@ -53,7 +53,7 @@ record SetPackage ℓ ℓ' : Type (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) where field ∈-rep : V ℓ → hProp ℓ' unpack : (x : X) → ⟨ ∈-rep (ix x) ⟩ - repack : {y : V ℓ} → ⟨ ∈-rep y ⟩ → ∥ fiber ix y ∥ + repack : {y : V ℓ} → ⟨ ∈-rep y ⟩ → ∥ fiber ix y ∥₁ open PropMonad classification : ⟨ ∀[ y ] (y ∈ₛ resSet ⇔ ∈-rep y) ⟩ @@ -99,10 +99,10 @@ module UnionSet (S : V ℓ) where UnionPackage : SetPackage _ (ℓ-suc ℓ) structure UnionPackage = UnionStructure ∈-rep UnionPackage y = ∃[ v ] (v ∈ₛ S) ⊓ (y ∈ₛ v) - unpack UnionPackage (vi , yi) = ∣ ⟪ S ⟫↪ vi , ∈ₛ⟪ S ⟫↪ vi , ∈ₛ⟪ ⟪ S ⟫↪ vi ⟫↪ yi ∣ - repack UnionPackage {y = y} = P.rec squash go where - go : Σ[ v ∈ V _ ] ⟨ v ∈ₛ S ⟩ ⊓′ ⟨ y ∈ₛ v ⟩ → ∥ fiber _ y ∥ - go (v , (vi , vS) , xv) = ∣ repFiber≃fiber _ _ .fst ((vi , key .fst) , key .snd) ∣ where + unpack UnionPackage (vi , yi) = ∣ ⟪ S ⟫↪ vi , ∈ₛ⟪ S ⟫↪ vi , ∈ₛ⟪ ⟪ S ⟫↪ vi ⟫↪ yi ∣₁ + repack UnionPackage {y = y} = P.rec squash₁ go where + go : Σ[ v ∈ V _ ] ⟨ v ∈ₛ S ⟩ ⊓′ ⟨ y ∈ₛ v ⟩ → ∥ fiber _ y ∥₁ + go (v , (vi , vS) , xv) = ∣ repFiber≃fiber _ _ .fst ((vi , key .fst) , key .snd) ∣₁ where path : v ≡ ⟪ S ⟫↪ vi path = sym (equivFun identityPrinciple vS) key : Σ[ i ∈ ⟪ ⟪ S ⟫↪ vi ⟫ ] ⟪ ⟪ S ⟫↪ vi ⟫↪ i ≊ y @@ -126,8 +126,8 @@ module PairingSet (a b : V ℓ) where unpack PairingPackage (lift false) = L.inl refl unpack PairingPackage (lift true) = L.inr refl repack PairingPackage {y = y} = _>>= λ where - (_⊎_.inl ya) → ∣ lift false , sym ya ∣ - (_⊎_.inr yb) → ∣ lift true , sym yb ∣ + (_⊎_.inl ya) → ∣ lift false , sym ya ∣₁ + (_⊎_.inr yb) → ∣ lift true , sym yb ∣₁ PAIR : V ℓ PAIR = SetStructure.resSet PairingStructure @@ -147,7 +147,7 @@ module SingletonSet (a : V ℓ) where structure SingletonPackage = SingletonStructure ∈-rep SingletonPackage d = d ≡ₕ a unpack SingletonPackage _ = refl - repack SingletonPackage ya = ∣ lift tt , sym ya ∣ + repack SingletonPackage ya = ∣ lift tt , sym ya ∣₁ SINGL : V ℓ SINGL = SetStructure.resSet SingletonStructure @@ -177,12 +177,12 @@ module InfinitySet {ℓ} where structure ωPackage = ωStructure ∈-rep ωPackage d = (d ≡ₕ ∅) ⊔ (∃[ v ] (d ≡ₕ sucV v) ⊓ (v ∈ₛ ω)) unpack ωPackage (lift zero) = L.inl refl - unpack ωPackage (lift (suc n)) = L.inr ∣ # n , refl , ∈∈ₛ {b = ω} .fst ∣ lift n , refl ∣ ∣ + unpack ωPackage (lift (suc n)) = L.inr ∣ # n , refl , ∈∈ₛ {b = ω} .fst ∣ lift n , refl ∣₁ ∣₁ repack ωPackage {y = y} = ⊔-elim (y ≡ₕ ∅) ∥ _ ∥ₚ (λ _ → ∥ fiber _ y ∥ₚ) - (λ e → ∣ lift zero , sym e ∣) + (λ e → ∣ lift zero , sym e ∣₁) (λ m → do (v , yv , vr) ← m (lift n , eq) ← ∈∈ₛ {b = ω} .snd vr - ∣ lift (suc n) , sym (subst (λ v → y ≡ (v ∪ ⁅ v ⁆s)) (sym eq) yv) ∣ + ∣ lift (suc n) , sym (subst (λ v → y ≡ (v ∪ ⁅ v ⁆s)) (sym eq) yv) ∣₁ ) infinity-ax : ⟨ ∀[ y ] (y ∈ₛ ω ⇔ (y ≡ₕ ∅) ⊔ (∃[ v ] (y ≡ₕ sucV v) ⊓ (v ∈ₛ ω))) ⟩ @@ -193,7 +193,7 @@ module InfinitySet {ℓ} where ω-empty = infinity-ax ∅ .snd (L.inl refl) ω-next : ⟨ ∀[ x ∶ V ℓ ] x ∈ₛ ω ⇒ sucV x ∈ₛ ω ⟩ - ω-next x x∈ω = infinity-ax (sucV x) .snd (L.inr ∣ x , refl , x∈ω ∣) + ω-next x x∈ω = infinity-ax (sucV x) .snd (L.inr ∣ x , refl , x∈ω ∣₁) #-in-ω : ∀ n → ⟨ # n ∈ₛ ω ⟩ #-in-ω zero = ω-empty @@ -213,10 +213,10 @@ module ReplacementSet (r : V ℓ → V ℓ) (a : V ℓ) where ReplacementPackage : SetPackage _ (ℓ-suc ℓ) structure ReplacementPackage = ReplacementStructure ∈-rep ReplacementPackage y = ∃[ z ] (z ∈ₛ a) ⊓ (y ≡ₕ r z) - unpack ReplacementPackage ⟪a⟫ = ∣ ⟪ a ⟫↪ ⟪a⟫ , (∈ₛ⟪ a ⟫↪ ⟪a⟫) , refl ∣ + unpack ReplacementPackage ⟪a⟫ = ∣ ⟪ a ⟫↪ ⟪a⟫ , (∈ₛ⟪ a ⟫↪ ⟪a⟫) , refl ∣₁ repack ReplacementPackage {y = y} m = do (z , (a , za) , yr) ← m - ∣ a , cong r (equivFun identityPrinciple za) ∙ sym yr ∣ + ∣ a , cong r (equivFun identityPrinciple za) ∙ sym yr ∣₁ replacement-ax : ⟨ ∀[ y ] (y ∈ₛ REPLACED ⇔ (∃[ z ] (z ∈ₛ a) ⊓ (y ≡ₕ r z))) ⟩ replacement-ax y = classification y where @@ -234,7 +234,7 @@ module SeparationSet (a : V ℓ) (ϕ : V ℓ → hProp ℓ) where unpack SeparationPackage (⟪a⟫ , phi) = (∈ₛ⟪ a ⟫↪ ⟪a⟫) , phi repack SeparationPackage ((⟪a⟫ , ya) , phi) = ∣ (⟪a⟫ , subst (fst ∘ ϕ) (sym (equivFun identityPrinciple ya)) phi) - , equivFun identityPrinciple ya ∣ + , equivFun identityPrinciple ya ∣₁ SEPAREE : V ℓ SEPAREE = SetStructure.resSet SeparationStructure @@ -250,11 +250,11 @@ module Examples where the1 : V ℓ-zero the1 = # 1 - 1-ok? : ∥ Bool ∥ + 1-ok? : ∥ Bool ∥₁ 1-ok? = do prf ← infinity-ax the1 .fst (#-in-ω 1) - case prf of λ { (⊎.inl _) → ∣ false ∣ ; (⊎.inr _) → ∣ true ∣ } + case prf of λ { (⊎.inl _) → ∣ false ∣₁ ; (⊎.inr _) → ∣ true ∣₁ } where open PropMonad - test : 1-ok? ≡ ∣ true ∣ + test : 1-ok? ≡ ∣ true ∣₁ test = refl diff --git a/Cubical/HITs/CumulativeHierarchy/Properties.agda b/Cubical/HITs/CumulativeHierarchy/Properties.agda index 681500474f..65e3585c36 100644 --- a/Cubical/HITs/CumulativeHierarchy/Properties.agda +++ b/Cubical/HITs/CumulativeHierarchy/Properties.agda @@ -61,11 +61,11 @@ _∼_ = elim2 eliminator where lemma _ rec₁→₂ rec₂→₁ (X₁→Y , Y→X₁) = (λ x₂ → do ((x₁ , c_) , xr₁) ← rec₂→₁ x₂ (y , yr) ← X₁→Y x₁ - ∣ y , subst fst (λ i → xr₁ i y) yr ∣ + ∣ y , subst fst (λ i → xr₁ i y) yr ∣₁ ) , (λ y → do (x₁ , xr₁) ← Y→X₁ y ((x₂ , _) , xr₂) ← rec₁→₂ x₁ - ∣ x₂ , subst fst (λ i → xr₂ (~ i) y) xr₁ ∣ + ∣ x₂ , subst fst (λ i → xr₂ (~ i) y) xr₁ ∣₁ ) open Elim2Set @@ -83,7 +83,7 @@ s ≊ t = ⟨ s ∼ t ⟩ ∼refl : (a : V ℓ) → a ≊ a ∼refl = elimProp (λ a → isProp⟨⟩ (a ∼ a)) - (λ X ix rec → (λ x → ∣ x , rec x ∣) , (λ x → ∣ x , rec x ∣)) + (λ X ix rec → (λ x → ∣ x , rec x ∣₁) , (λ x → ∣ x , rec x ∣₁)) -- keep in mind that the left and right side here live in different universes identityPrinciple : (a ≊ b) ≃ (a ≡ b) @@ -94,8 +94,8 @@ identityPrinciple {a = a} {b = b} = eqImageXY : {X Y : Type ℓ} {ix : X → V ℓ} {iy : Y → V ℓ} → (∀ x y → ⟨ ix x ∼ iy y ⟩ → ix x ≡ iy y) → ⟨ sett X ix ∼ sett Y iy ⟩ → eqImage ix iy - eqImageXY rec rel = (λ x → do (y , y∼x) ← fst rel x ; ∣ y , sym (rec _ _ y∼x) ∣) - , (λ y → do (x , x∼y) ← snd rel y ; ∣ x , rec _ _ x∼y ∣) + eqImageXY rec rel = (λ x → do (y , y∼x) ← fst rel x ; ∣ y , sym (rec _ _ y∼x) ∣₁) + , (λ y → do (x , x∼y) ← snd rel y ; ∣ x , rec _ _ x∼y ∣₁) from : a ≊ b → a ≡ b from = elimProp propB eliminator a b where @@ -137,12 +137,12 @@ isPropMonicPresentation a ((X₁ , ix₁ , isEmb₁) , p) ((X₂ , ix₂ , isEmb fiberwise1 : ∀ b → fiber ix₁ b → fiber ix₂ b fiberwise1 b fbx₁ = proof (_ , isEmbedding→hasPropFibers isEmb₂ b) - by subst (λ A → ⟨ b ∈ A ⟩) (sym p ∙ q) ∣ fbx₁ ∣ + by subst (λ A → ⟨ b ∈ A ⟩) (sym p ∙ q) ∣ fbx₁ ∣₁ fiberwise2 : ∀ b → fiber ix₂ b → fiber ix₁ b fiberwise2 b fbx₂ = proof (_ , isEmbedding→hasPropFibers isEmb₁ b) - by subst (λ A → ⟨ b ∈ A ⟩) (sym q ∙ p) ∣ fbx₂ ∣ + by subst (λ A → ⟨ b ∈ A ⟩) (sym q ∙ p) ∣ fbx₂ ∣₁ sett-repr : (X : Type ℓ) (ix : X → V ℓ) → MonicPresentation (sett X ix) sett-repr {ℓ} X ix = (Rep , ixRep , isEmbIxRep) , seteq X Rep ix ixRep eqImIxRep where @@ -168,7 +168,7 @@ sett-repr {ℓ} X ix = (Rep , ixRep , isEmbIxRep) , seteq X Rep ix ixRep eqImIxR goal = Q.elimProp prop₁ (λ p₁ m → Q.elimProp prop₂ (λ p₂ n → eq/ p₁ p₂ (lemma m n)) p₂ n) p₁ m eqImIxRep : eqImage ix ixRep - eqImIxRep = (λ x → ∣ Q.[ x ] , refl ∣) , Q.elimProp (λ _ → P.squash) (λ b → ∣ b , refl ∣) + eqImIxRep = (λ x → ∣ Q.[ x ] , refl ∣₁) , Q.elimProp (λ _ → P.squash₁) (λ b → ∣ b , refl ∣₁) data DeepMonicPresentation (a : V ℓ) : Type (ℓ-suc ℓ) where dmp : (mp@((_ , ix , _) , _) : MonicPresentation a) @@ -247,7 +247,7 @@ a ∈ₛ b = repFiber ⟪ b ⟫↪ a , isPropRepFiber b a asRep : ⟨ a ∈ sett ⟪ b ⟫ ⟪ b ⟫↪ ⟩ → fiber ⟪ b ⟫↪ a asRep = P.propTruncIdempotent≃ (isEmbedding→hasPropFibers isEmb⟪ b ⟫↪ a) .fst ∈-fromFiber : {a b : V ℓ} → fiber ⟪ b ⟫↪ a → ⟨ a ∈ b ⟩ -∈-fromFiber {a = a} {b = b} = subst (λ br → ⟨ a ∈ br ⟩) (sym ⟪ b ⟫-represents) ∘ ∣_∣ +∈-fromFiber {a = a} {b = b} = subst (λ br → ⟨ a ∈ br ⟩) (sym ⟪ b ⟫-represents) ∘ ∣_∣₁ ∈∈ₛ : {a b : V ℓ} → ⟨ a ∈ b ⇔ a ∈ₛ b ⟩ ∈∈ₛ {a = a} {b = b} = leftToRight , rightToLeft where @@ -260,7 +260,7 @@ a ∈ₛ b = repFiber ⟪ b ⟫↪ a , isPropRepFiber b a ix∈ₛ : {X : Type ℓ} {ix : X → V ℓ} → (x : X) → ⟨ ix x ∈ₛ sett X ix ⟩ -ix∈ₛ {X = X} {ix = ix} x = ∈∈ₛ {a = ix x} {b = sett X ix} .fst ∣ x , refl ∣ +ix∈ₛ {X = X} {ix = ix} x = ∈∈ₛ {a = ix x} {b = sett X ix} .fst ∣ x , refl ∣₁ ∈ₛ⟪_⟫↪_ : (a : V ℓ) (ix : ⟪ a ⟫) → ⟨ ⟪ a ⟫↪ ix ∈ₛ a ⟩ ∈ₛ⟪ a ⟫↪ ix = ix , ∼refl (⟪ a ⟫↪ ix) diff --git a/Cubical/HITs/Cylinder/Base.agda b/Cubical/HITs/Cylinder/Base.agda index 6331d0384b..0b1222e225 100644 --- a/Cubical/HITs/Cylinder/Base.agda +++ b/Cubical/HITs/Cylinder/Base.agda @@ -2,17 +2,19 @@ module Cubical.HITs.Cylinder.Base where -open import Cubical.Core.Everything - -open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv open import Cubical.Data.Sigma open import Cubical.Data.Unit open import Cubical.Data.Sum using (_⊎_; inl; inr) -open import Cubical.HITs.PropositionalTruncation using (∥_∥; ∣_∣; squash) +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. -- -- https://ncatlab.org/nlab/show/cylinder+object @@ -40,12 +42,12 @@ module _ {ℓ} {A : Type ℓ} where -- The above inclusion is surjective includeSurjective : ∀ c → ∃[ s ∈ A ⊎ A ] include s ≡ c - includeSurjective (inl x) = ∣ inl x , refl ∣ - includeSurjective (inr x) = ∣ inr x , refl ∣ + includeSurjective (inl x) = ∣ inl x , refl ∣₁ + includeSurjective (inr x) = ∣ inr x , refl ∣₁ includeSurjective (cross x i) = - squash - ∣ inl x , (λ j → cross x (i ∧ j)) ∣ - ∣ inr x , (λ j → cross x (i ∨ ~ j)) ∣ + squash₁ + ∣ inl x , (λ j → cross x (i ∧ j)) ∣₁ + ∣ inr x , (λ j → cross x (i ∨ ~ j)) ∣₁ i elim diff --git a/Cubical/HITs/EilenbergMacLane1/Properties.agda b/Cubical/HITs/EilenbergMacLane1/Properties.agda index de5d1f51bb..0505d4fd6a 100644 --- a/Cubical/HITs/EilenbergMacLane1/Properties.agda +++ b/Cubical/HITs/EilenbergMacLane1/Properties.agda @@ -29,9 +29,6 @@ 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₂) - private variable ℓG ℓ : Level diff --git a/Cubical/HITs/FreeComMonoids/Properties.agda b/Cubical/HITs/FreeComMonoids/Properties.agda index 588aaaa9a4..59b92e652d 100644 --- a/Cubical/HITs/FreeComMonoids/Properties.agda +++ b/Cubical/HITs/FreeComMonoids/Properties.agda @@ -2,7 +2,13 @@ module Cubical.HITs.FreeComMonoids.Properties where -open import Cubical.Foundations.Everything hiding (assoc; ⟨_⟩) +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.GroupoidLaws hiding (assoc) +open import Cubical.Foundations.Univalence open import Cubical.Data.Nat hiding (_·_ ; _^_) @@ -13,6 +19,8 @@ private variable ℓ : Level A : Type ℓ + + multi-· : A → ℕ → FreeComMonoid A → FreeComMonoid A multi-· x zero xs = xs multi-· x (suc n) xs = ⟦ x ⟧ · multi-· x n xs diff --git a/Cubical/HITs/GroupoidQuotients/Properties.agda b/Cubical/HITs/GroupoidQuotients/Properties.agda index aff8f33131..6ca1bf5db7 100644 --- a/Cubical/HITs/GroupoidQuotients/Properties.agda +++ b/Cubical/HITs/GroupoidQuotients/Properties.agda @@ -23,7 +23,7 @@ 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.PropositionalTruncation as PropTrunc using (∥_∥₁; ∣_∣₁; squash₁) open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂; ∣_∣₂; squash₂) -- Type quotients @@ -73,7 +73,7 @@ elimProp2 Rt Cprop f = elimProp Rt (λ x → isPropΠ (λ y → Cprop x y)) isSurjective[] : (Rt : BinaryRelation.isTrans R) → isSurjection (λ a → [ a ]) -isSurjective[] Rt = elimProp Rt (λ x → squash) (λ a → ∣ a , refl ∣) +isSurjective[] Rt = elimProp Rt (λ x → squash₁) (λ a → ∣ a , refl ∣₁) elim : (Rt : BinaryRelation.isTrans R) → {B : A // Rt → Type ℓ} diff --git a/Cubical/HITs/PropositionalTruncation/Base.agda b/Cubical/HITs/PropositionalTruncation/Base.agda index a453e7a163..226f251f09 100644 --- a/Cubical/HITs/PropositionalTruncation/Base.agda +++ b/Cubical/HITs/PropositionalTruncation/Base.agda @@ -12,6 +12,6 @@ open import Cubical.Core.Primitives -- Propositional truncation as a higher inductive type: -data ∥_∥ {ℓ} (A : Type ℓ) : Type ℓ where - ∣_∣ : A → ∥ A ∥ - squash : ∀ (x y : ∥ A ∥) → x ≡ y +data ∥_∥₁ {ℓ} (A : Type ℓ) : Type ℓ where + ∣_∣₁ : A → ∥ A ∥₁ + squash₁ : ∀ (x y : ∥ A ∥₁) → x ≡ y diff --git a/Cubical/HITs/PropositionalTruncation/MagicTrick.agda b/Cubical/HITs/PropositionalTruncation/MagicTrick.agda index e663da1d42..c43399a0c7 100644 --- a/Cubical/HITs/PropositionalTruncation/MagicTrick.agda +++ b/Cubical/HITs/PropositionalTruncation/MagicTrick.agda @@ -29,37 +29,37 @@ module Recover {ℓ} (A∙ : Pointed ℓ) (h : isHomogeneous A∙) where A = typ A∙ a = pt A∙ - toEquivPtd : ∥ A ∥ → Σ[ B∙ ∈ Pointed ℓ ] (A , a) ≡ B∙ + toEquivPtd : ∥ A ∥₁ → Σ[ B∙ ∈ Pointed ℓ ] (A , a) ≡ B∙ toEquivPtd = rec isPropSingl (λ x → (A , x) , h x) private - B∙ : ∥ A ∥ → Pointed ℓ + B∙ : ∥ A ∥₁ → Pointed ℓ B∙ tx = fst (toEquivPtd tx) - -- the key observation is that B∙ ∣ x ∣ is definitionally equal to (A , x) + -- the key observation is that B∙ ∣ x ∣₁ is definitionally equal to (A , x) private - obvs : ∀ x → B∙ ∣ x ∣ ≡ (A , x) - obvs x = refl -- try it: `C-c C-n B∙ ∣ x ∣` gives `(A , x)` + obvs : ∀ x → B∙ ∣ x ∣₁ ≡ (A , x) + obvs x = refl -- try it: `C-c C-n B∙ ∣ x ∣₁` gives `(A , x)` -- thus any truncated element (of a homogeneous type) can be recovered by agda's normalizer! - recover : ∀ (tx : ∥ A ∥) → typ (B∙ tx) + recover : ∀ (tx : ∥ A ∥₁) → typ (B∙ tx) recover tx = pt (B∙ tx) - recover∣∣ : ∀ (x : A) → recover ∣ x ∣ ≡ x - recover∣∣ x = refl -- try it: `C-c C-n recover ∣ x ∣` gives `x` + recover∣∣ : ∀ (x : A) → recover ∣ x ∣₁ ≡ x + recover∣∣ x = refl -- try it: `C-c C-n recover ∣ x ∣₁` gives `x` private - -- notice that the following typechecks because typ (B∙ ∣ x ∣) is definitionally equal to to A, but - -- `recover : ∥ A ∥ → A` does not because typ (B∙ tx) is not definitionally equal to A (though it is + -- notice that the following typechecks because typ (B∙ ∣ x ∣₁) is definitionally equal to to A, but + -- `recover : ∥ A ∥₁ → A` does not because typ (B∙ tx) is not definitionally equal to A (though it is -- judegmentally equal to A by cong typ (snd (toEquivPtd tx)) : A ≡ typ (B∙ tx)) obvs2 : A → A - obvs2 = recover ∘ ∣_∣ + obvs2 = recover ∘ ∣_∣₁ - -- one might wonder if (cong recover (squash ∣ x ∣ ∣ y ∣)) therefore has type x ≡ y, but thankfully - -- typ (B∙ (squash ∣ x ∣ ∣ y ∣ i)) is *not* A (it's a messy hcomp involving h x and h y) - recover-squash : ∀ x y → -- x ≡ y -- this raises an error - PathP (λ i → typ (B∙ (squash ∣ x ∣ ∣ y ∣ i))) x y - recover-squash x y = cong recover (squash ∣ x ∣ ∣ y ∣) + -- one might wonder if (cong recover (squash₁ ∣ x ∣₁ ∣ y ∣₁)) therefore has type x ≡ y, but thankfully + -- typ (B∙ (squash₁ ∣ x ∣₁ ∣ y ∣₁ i)) is *not* A (it's a messy hcomp involving h x and h y) + recover-squash₁ : ∀ x y → -- x ≡ y -- this raises an error + PathP (λ i → typ (B∙ (squash₁ ∣ x ∣₁ ∣ y ∣₁ i))) x y + recover-squash₁ x y = cong recover (squash₁ ∣ x ∣₁ ∣ y ∣₁) -- Demo, adapted from: @@ -75,8 +75,8 @@ private hidden : ℕ hidden = 17 - ∣hidden∣ : ∥ ℕ ∥ - ∣hidden∣ = ∣ hidden ∣ + ∣hidden∣ : ∥ ℕ ∥₁ + ∣hidden∣ = ∣ hidden ∣₁ -- we can still recover the value, even though agda can no longer see `hidden`! test : recover ∣hidden∣ ≡ 17 @@ -84,5 +84,5 @@ private -- `C-c C-n hidden` gives an error -- Finally, note that the definition of recover is independent of the proof that A is homogeneous. Thus we - -- still can definitionally recover information hidden by ∣_∣ as long as we permit holes. Try replacing + -- still can definitionally recover information hidden by ∣_∣₁ as long as we permit holes. Try replacing -- `isHomogeneousDiscrete discreteℕ` above with a hole (`?`) and notice that everything still works diff --git a/Cubical/HITs/PropositionalTruncation/Monad.agda b/Cubical/HITs/PropositionalTruncation/Monad.agda index 7ecb829a5a..ce757e34c3 100644 --- a/Cubical/HITs/PropositionalTruncation/Monad.agda +++ b/Cubical/HITs/PropositionalTruncation/Monad.agda @@ -17,17 +17,17 @@ private P Q : Type ℓ infix 1 proof_by_ -proof_by_ : (P : hProp ℓ) → ∥ ⟨ P ⟩ ∥ → ⟨ P ⟩ +proof_by_ : (P : hProp ℓ) → ∥ ⟨ P ⟩ ∥₁ → ⟨ P ⟩ proof P by p = rec (isProp⟨⟩ P) (λ p → p) p -return : P → ∥ P ∥ -return p = ∣ p ∣ +return : P → ∥ P ∥₁ +return p = ∣ p ∣₁ -exact_ : ∥ P ∥ → ∥ P ∥ +exact_ : ∥ P ∥₁ → ∥ P ∥₁ exact p = p -_>>=_ : ∥ P ∥ → (P → ∥ Q ∥) → ∥ Q ∥ +_>>=_ : ∥ P ∥₁ → (P → ∥ Q ∥₁) → ∥ Q ∥₁ p >>= f = rec isPropPropTrunc f p -_>>_ : ∥ P ∥ → ∥ Q ∥ → ∥ Q ∥ +_>>_ : ∥ P ∥₁ → ∥ Q ∥₁ → ∥ Q ∥₁ _ >> q = q diff --git a/Cubical/HITs/PropositionalTruncation/Properties.agda b/Cubical/HITs/PropositionalTruncation/Properties.agda index 2e846c31cb..2be49c28b9 100644 --- a/Cubical/HITs/PropositionalTruncation/Properties.agda +++ b/Cubical/HITs/PropositionalTruncation/Properties.agda @@ -28,77 +28,77 @@ private A B C : Type ℓ A′ : Type ℓ' -∥∥-isPropDep : (P : A → Type ℓ) → isOfHLevelDep 1 (λ x → ∥ P x ∥) -∥∥-isPropDep P = isOfHLevel→isOfHLevelDep 1 (λ _ → squash) +∥∥-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 +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 +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) = +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 + (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 : {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 : {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 +isPropPropTrunc : isProp ∥ A ∥₁ +isPropPropTrunc x y = squash₁ x y -propTrunc≃ : A ≃ B → ∥ A ∥ ≃ ∥ B ∥ +propTrunc≃ : A ≃ B → ∥ A ∥₁ ≃ ∥ B ∥₁ propTrunc≃ e = propBiimpl→Equiv isPropPropTrunc isPropPropTrunc - (rec isPropPropTrunc (λ a → ∣ e .fst a ∣)) - (rec isPropPropTrunc (λ b → ∣ invEq e b ∣)) + (rec isPropPropTrunc (λ a → ∣ e .fst a ∣₁)) + (rec isPropPropTrunc (λ b → ∣ invEq e b ∣₁)) -propTruncIdempotent≃ : isProp A → ∥ A ∥ ≃ A +propTruncIdempotent≃ : isProp A → ∥ A ∥₁ ≃ A propTruncIdempotent≃ {A = A} hA = isoToEquiv f where - f : Iso ∥ A ∥ A + f : Iso ∥ A ∥₁ A Iso.fun f = rec hA (idfun A) - Iso.inv f x = ∣ x ∣ + Iso.inv f x = ∣ x ∣₁ Iso.rightInv f _ = refl Iso.leftInv f = elim (λ _ → isProp→isSet isPropPropTrunc _ _) (λ _ → refl) -propTruncIdempotent : isProp A → ∥ A ∥ ≡ A +propTruncIdempotent : isProp A → ∥ A ∥₁ ≡ A propTruncIdempotent hA = ua (propTruncIdempotent≃ hA) -- We could also define the eliminator using the recursor -elim' : {P : ∥ A ∥ → Type ℓ} → ((a : ∥ A ∥) → isProp (P a)) → - ((x : A) → P ∣ x ∣) → (a : ∥ A ∥) → P 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 + 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) +map : (A → B) → (∥ A ∥₁ → ∥ B ∥₁) +map f = rec squash₁ (∣_∣₁ ∘ f) -map2 : (A → B → C) → (∥ A ∥ → ∥ B ∥ → ∥ C ∥) -map2 f = rec (isPropΠ λ _ → squash) (map ∘ 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 @@ -109,21 +109,21 @@ module SetElim (Bset : isSet B) where Bset' : isSet' B Bset' = isSet→isSet' Bset - 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 : 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 - rec→Set f kf ∣ x ∣ = f x - rec→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 + helper f kf ∣ x ∣₁ ∣ y ∣₁ = kf x y + helper f kf (squash₁ t u i) v = Bset' (helper f kf t v) (helper f kf u v) (helper f kf t u) refl i - helper f kf t (squash u v i) + 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 x y = cong f (squash ∣ x ∣ ∣ y ∣) + kcomp : (f : ∥ A ∥₁ → B) → 2-Constant (f ∘ ∣_∣₁) + kcomp f x y = cong f (squash₁ ∣ x ∣₁ ∣ y ∣₁) Fset : isSet (A → B) Fset = isSetΠ (const Bset) @@ -132,14 +132,14 @@ module SetElim (Bset : isSet B) where Kset f = isSetΠ (λ _ → isSetΠ (λ _ → isProp→isSet (Bset _ _))) setRecLemma - : (f : ∥ A ∥ → B) - → rec→Set (f ∘ ∣_∣) (kcomp f) ≡ f + : (f : ∥ A ∥₁ → B) + → rec→Set (f ∘ ∣_∣₁) (kcomp f) ≡ f setRecLemma f i t - = elim {P = λ t → rec→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 + mkKmap : (∥ A ∥₁ → B) → Σ (A → B) 2-Constant + mkKmap f = f ∘ ∣_∣₁ , kcomp f fib : (g : Σ (A → B) 2-Constant) → fiber mkKmap g fib (g , kg) = rec→Set g kg , refl @@ -149,7 +149,7 @@ module SetElim (Bset : isSet B) where Σ≡Prop (λ f → isOfHLevelΣ 2 Fset Kset _ _) (cong (uncurry rec→Set) (sym p) ∙ setRecLemma f) - trunc→Set≃ : (∥ A ∥ → B) ≃ (Σ (A → B) 2-Constant) + trunc→Set≃ : (∥ A ∥₁ → B) ≃ (Σ (A → B) 2-Constant) trunc→Set≃ .fst = mkKmap trunc→Set≃ .snd .equiv-proof g = fib g , eqv g @@ -169,66 +169,66 @@ module SetElim (Bset : isSet B) where e-isEquiv : A → isEquiv (e {A = A}) e-isEquiv a₀ = isoToIsEquiv (iso e (eval a₀) (e-eval a₀) λ _ → refl) - preEquiv₁ : ∥ A ∥ → B ≃ Σ (A → B) 2-Constant + preEquiv₁ : ∥ A ∥₁ → B ≃ Σ (A → B) 2-Constant preEquiv₁ t = e , rec (isPropIsEquiv e) e-isEquiv t - preEquiv₂ : (∥ A ∥ → Σ (A → B) 2-Constant) ≃ Σ (A → B) 2-Constant + preEquiv₂ : (∥ A ∥₁ → Σ (A → B) 2-Constant) ≃ Σ (A → B) 2-Constant preEquiv₂ = isoToEquiv (iso to const (λ _ → refl) retr) where - to : (∥ A ∥ → Σ (A → B) 2-Constant) → Σ (A → B) 2-Constant - to f .fst x = f ∣ x ∣ .fst x - to f .snd x y i = f (squash ∣ x ∣ ∣ y ∣ i) .snd x y i + to : (∥ A ∥₁ → Σ (A → B) 2-Constant) → Σ (A → B) 2-Constant + to f .fst x = f ∣ x ∣₁ .fst x + to f .snd x y i = f (squash₁ ∣ x ∣₁ ∣ y ∣₁ i) .snd x y i retr : retract to const - retr f i t .fst x = f (squash ∣ x ∣ t i) .fst x + retr f i t .fst x = f (squash₁ ∣ x ∣₁ t i) .fst x retr f i t .snd x y = Bset' - (λ j → f (squash ∣ x ∣ ∣ y ∣ j) .snd x y j) + (λ j → f (squash₁ ∣ x ∣₁ ∣ y ∣₁ j) .snd x y j) (f t .snd x y) - (λ j → f (squash ∣ x ∣ t j) .fst x) - (λ j → f (squash ∣ y ∣ t j) .fst y) + (λ j → f (squash₁ ∣ x ∣₁ t j) .fst x) + (λ j → f (squash₁ ∣ y ∣₁ t j) .fst y) i - trunc→Set≃₂ : (∥ A ∥ → B) ≃ Σ (A → B) 2-Constant + trunc→Set≃₂ : (∥ A ∥₁ → B) ≃ Σ (A → B) 2-Constant trunc→Set≃₂ = compEquiv (equivΠCod preEquiv₁) preEquiv₂ open SetElim public using (rec→Set; trunc→Set≃) elim→Set - : {P : ∥ A ∥ → Type ℓ} + : {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 + → (f : (x : A) → P ∣ x ∣₁) + → (kf : ∀ x y → PathP (λ i → P (squash₁ ∣ x ∣₁ ∣ y ∣₁ i)) (f x) (f y)) + → (t : ∥ A ∥₁) → P 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) + g x = transp (λ i → P (squash₁ ∣ x ∣₁ t i)) i0 (f x) gk : 2-Constant g - gk x y i = transp (λ j → P (squash (squash ∣ x ∣ ∣ y ∣ i) t j)) i0 (kf x y i) + gk x y i = transp (λ j → P (squash₁ (squash₁ ∣ x ∣₁ ∣ y ∣₁ i) t j)) i0 (kf x y i) elim2→Set : - {P : ∥ A ∥ → ∥ B ∥ → Type ℓ} + {P : ∥ A ∥₁ → ∥ B ∥₁ → Type ℓ} → (∀ t u → isSet (P t u)) - → (f : (x : A) (y : B) → P ∣ x ∣ ∣ y ∣) - → (kf₁ : ∀ x y v → PathP (λ i → P (squash ∣ x ∣ ∣ y ∣ i) ∣ v ∣) (f x v) (f y v)) - → (kf₂ : ∀ x v w → PathP (λ i → P ∣ x ∣ (squash ∣ v ∣ ∣ w ∣ i)) (f x v) (f x w)) - → (sf : ∀ x y v w → SquareP (λ i j → P (squash ∣ x ∣ ∣ y ∣ i) (squash ∣ v ∣ ∣ w ∣ j)) + → (f : (x : A) (y : B) → P ∣ x ∣₁ ∣ y ∣₁) + → (kf₁ : ∀ x y v → PathP (λ i → P (squash₁ ∣ x ∣₁ ∣ y ∣₁ i) ∣ v ∣₁) (f x v) (f y v)) + → (kf₂ : ∀ x v w → PathP (λ i → P ∣ x ∣₁ (squash₁ ∣ v ∣₁ ∣ w ∣₁ i)) (f x v) (f x w)) + → (sf : ∀ x y v w → SquareP (λ i j → P (squash₁ ∣ x ∣₁ ∣ y ∣₁ i) (squash₁ ∣ v ∣₁ ∣ w ∣₁ j)) (kf₂ x v w) (kf₂ y v w) (kf₁ x y v) (kf₁ x y w)) - → (t : ∥ A ∥) → (u : ∥ B ∥) → P t u + → (t : ∥ A ∥₁) → (u : ∥ B ∥₁) → P t u elim2→Set {A = A} {B = B} {P = P} Pset f kf₁ kf₂ sf = elim→Set (λ _ → isSetΠ (λ _ → Pset _ _)) mapHelper squareHelper where - mapHelper : (x : A) (u : ∥ B ∥) → P ∣ x ∣ u + mapHelper : (x : A) (u : ∥ B ∥₁) → P ∣ x ∣₁ u mapHelper x = elim→Set (λ _ → Pset _ _) (f x) (kf₂ x) squareHelper : (x y : A) - → PathP (λ i → (u : ∥ B ∥) → P (squash ∣ x ∣ ∣ y ∣ i) u) (mapHelper x) (mapHelper y) + → PathP (λ i → (u : ∥ B ∥₁) → P (squash₁ ∣ x ∣₁ ∣ y ∣₁ i) u) (mapHelper x) (mapHelper y) squareHelper x y i = elim→Set (λ _ → Pset _ _) (λ v → kf₁ x y v i) λ v w → sf x y v w i -RecHProp : (P : A → hProp ℓ) (kP : ∀ x y → P x ≡ P y) → ∥ A ∥ → hProp ℓ +RecHProp : (P : A → hProp ℓ) (kP : ∀ x y → P x ≡ P y) → ∥ A ∥₁ → hProp ℓ RecHProp P kP = rec→Set isSetHProp P kP module GpdElim (Bgpd : isGroupoid B) where @@ -238,24 +238,24 @@ module GpdElim (Bgpd : isGroupoid B) where module _ (f : A → B) (3kf : 3-Constant f) where open 3-Constant 3kf - rec→Gpd : ∥ A ∥ → B - pathHelper : (t u : ∥ A ∥) → rec→Gpd t ≡ rec→Gpd u + rec→Gpd : ∥ A ∥₁ → B + pathHelper : (t u : ∥ A ∥₁) → rec→Gpd t ≡ rec→Gpd u triHelper₁ - : (t u v : ∥ A ∥) + : (t u v : ∥ A ∥₁) → Square (pathHelper t u) (pathHelper t v) refl (pathHelper u v) triHelper₂ - : (t u v : ∥ A ∥) + : (t u v : ∥ A ∥₁) → Square (pathHelper t v) (pathHelper u v) (pathHelper t u) refl - rec→Gpd ∣ x ∣ = f x - rec→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 - pathHelper ∣ x ∣ (squash u v j) = triHelper₁ ∣ x ∣ u v j + pathHelper ∣ x ∣₁ ∣ y ∣₁ = link x y + pathHelper (squash₁ t u j) v = triHelper₂ t u v j + pathHelper ∣ x ∣₁ (squash₁ u v j) = triHelper₁ ∣ x ∣₁ u v j - triHelper₁ ∣ x ∣ ∣ y ∣ ∣ z ∣ = coh₁ x y z - triHelper₁ (squash s t i) u v + triHelper₁ ∣ x ∣₁ ∣ y ∣₁ ∣ z ∣₁ = coh₁ x y z + triHelper₁ (squash₁ s t i) u v = Bgpd' (triHelper₁ s u v) (triHelper₁ t u v) @@ -264,27 +264,27 @@ module GpdElim (Bgpd : isGroupoid B) where (λ i → refl) (λ i → pathHelper u v) i - triHelper₁ ∣ x ∣ (squash t u i) v + triHelper₁ ∣ x ∣₁ (squash₁ t u i) v = Bgpd' - (triHelper₁ ∣ x ∣ t v) - (triHelper₁ ∣ x ∣ u v) - (triHelper₁ ∣ x ∣ t u) - (λ i → pathHelper ∣ x ∣ v) + (triHelper₁ ∣ x ∣₁ t v) + (triHelper₁ ∣ x ∣₁ u v) + (triHelper₁ ∣ x ∣₁ t u) + (λ i → pathHelper ∣ x ∣₁ v) (λ i → refl) (triHelper₂ t u v) i - triHelper₁ ∣ x ∣ ∣ y ∣ (squash u v i) + triHelper₁ ∣ x ∣₁ ∣ y ∣₁ (squash₁ u v i) = Bgpd' - (triHelper₁ ∣ x ∣ ∣ y ∣ u) - (triHelper₁ ∣ x ∣ ∣ y ∣ v) + (triHelper₁ ∣ x ∣₁ ∣ y ∣₁ u) + (triHelper₁ ∣ x ∣₁ ∣ y ∣₁ v) (λ i → link x y) - (triHelper₁ ∣ x ∣ u v) + (triHelper₁ ∣ x ∣₁ u v) (λ i → refl) - (triHelper₁ ∣ y ∣ u v) + (triHelper₁ ∣ y ∣₁ u v) i - triHelper₂ ∣ x ∣ ∣ y ∣ ∣ z ∣ = coh₂ x y z - triHelper₂ (squash s t i) u v + triHelper₂ ∣ x ∣₁ ∣ y ∣₁ ∣ z ∣₁ = coh₂ x y z + triHelper₂ (squash₁ s t i) u v = Bgpd' (triHelper₂ s u v) (triHelper₂ t u v) @@ -293,40 +293,40 @@ module GpdElim (Bgpd : isGroupoid B) where (triHelper₂ s t u) (λ i → refl) i - triHelper₂ ∣ x ∣ (squash t u i) v + triHelper₂ ∣ x ∣₁ (squash₁ t u i) v = Bgpd' - (triHelper₂ ∣ x ∣ t v) - (triHelper₂ ∣ x ∣ u v) - (λ i → pathHelper ∣ x ∣ v) + (triHelper₂ ∣ x ∣₁ t v) + (triHelper₂ ∣ x ∣₁ u v) + (λ i → pathHelper ∣ x ∣₁ v) (triHelper₂ t u v) - (triHelper₁ ∣ x ∣ t u) + (triHelper₁ ∣ x ∣₁ t u) (λ i → refl) i - triHelper₂ ∣ x ∣ ∣ y ∣ (squash u v i) + triHelper₂ ∣ x ∣₁ ∣ y ∣₁ (squash₁ u v i) = Bgpd' - (triHelper₂ ∣ x ∣ ∣ y ∣ u) - (triHelper₂ ∣ x ∣ ∣ y ∣ v) - (triHelper₁ ∣ x ∣ u v) - (triHelper₁ ∣ y ∣ u v) + (triHelper₂ ∣ x ∣₁ ∣ y ∣₁ u) + (triHelper₂ ∣ x ∣₁ ∣ y ∣₁ v) + (triHelper₁ ∣ x ∣₁ u v) + (triHelper₁ ∣ y ∣₁ u v) (λ i → link x y) (λ i → refl) i - preEquiv₁ : (∥ A ∥ → Σ (A → B) 3-Constant) ≃ Σ (A → B) 3-Constant + preEquiv₁ : (∥ A ∥₁ → Σ (A → B) 3-Constant) ≃ Σ (A → B) 3-Constant preEquiv₁ = isoToEquiv (iso fn const (λ _ → refl) retr) where open 3-Constant - fn : (∥ A ∥ → Σ (A → B) 3-Constant) → Σ (A → B) 3-Constant - fn f .fst x = f ∣ x ∣ .fst x - fn f .snd .link x y i = f (squash ∣ x ∣ ∣ y ∣ i) .snd .link x y i + fn : (∥ A ∥₁ → Σ (A → B) 3-Constant) → Σ (A → B) 3-Constant + fn f .fst x = f ∣ x ∣₁ .fst x + fn f .snd .link x y i = f (squash₁ ∣ x ∣₁ ∣ y ∣₁ i) .snd .link x y i fn f .snd .coh₁ x y z i j - = f (squash ∣ x ∣ (squash ∣ y ∣ ∣ z ∣ i) j) .snd .coh₁ x y z i j + = f (squash₁ ∣ x ∣₁ (squash₁ ∣ y ∣₁ ∣ z ∣₁ i) j) .snd .coh₁ x y z i j retr : retract fn const - retr f i t .fst x = f (squash ∣ x ∣ t i) .fst x + retr f i t .fst x = f (squash₁ ∣ x ∣₁ t i) .fst x retr f i t .snd .link x y j - = f (squash (squash ∣ x ∣ ∣ y ∣ j) t i) .snd .link x y j + = f (squash₁ (squash₁ ∣ x ∣₁ ∣ y ∣₁ j) t i) .snd .link x y j retr f i t .snd .coh₁ x y z = Bgpd' (λ k j → f (cb k j i0) .snd .coh₁ x y z k j ) @@ -337,8 +337,8 @@ module GpdElim (Bgpd : isGroupoid B) where (λ k j → f (cb j i1 k) .snd .link y z j) i where - cb : I → I → I → ∥ _ ∥ - cb i j k = squash (squash ∣ x ∣ (squash ∣ y ∣ ∣ z ∣ i) j) t k + cb : I → I → I → ∥ _ ∥₁ + cb i j k = squash₁ (squash₁ ∣ x ∣₁ (squash₁ ∣ y ∣₁ ∣ z ∣₁ i) j) t k e : B → Σ (A → B) 3-Constant e b .fst _ = b @@ -368,132 +368,132 @@ module GpdElim (Bgpd : isGroupoid B) where e-isEquiv : A → isEquiv (e {A = A}) e-isEquiv a₀ = isoToIsEquiv (iso e (eval a₀) (e-eval a₀) λ _ → refl) - preEquiv₂ : ∥ A ∥ → B ≃ Σ (A → B) 3-Constant + preEquiv₂ : ∥ A ∥₁ → B ≃ Σ (A → B) 3-Constant preEquiv₂ t = e , rec (isPropIsEquiv e) e-isEquiv t - trunc→Gpd≃ : (∥ A ∥ → B) ≃ Σ (A → B) 3-Constant + trunc→Gpd≃ : (∥ A ∥₁ → B) ≃ Σ (A → B) 3-Constant trunc→Gpd≃ = compEquiv (equivΠCod preEquiv₂) preEquiv₁ open GpdElim using (rec→Gpd; trunc→Gpd≃) public -squashᵗ +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) + → 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 ℓ) + : (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)) + → (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 + → 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) + g x = transp (λ i → P (squash₁ ∣ x ∣₁ t i)) i0 (f x) 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) + = 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) + = 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 : A → TypeOfHLevel ℓ 2) → 3-Constant P → ∥ A ∥₁ → TypeOfHLevel ℓ 2 RecHSet P 3kP = rec→Gpd (isOfHLevelTypeOfHLevel 2) P 3kP -∥∥-IdempotentL-⊎-≃ : ∥ ∥ A ∥ ⊎ A′ ∥ ≃ ∥ 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′ ∥ + 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 ∣ + lem (inr x) = ∣ inr x ∣₁ Iso.inv ∥∥-IdempotentL-⊎-Iso x = map lem x - where lem : A ⊎ A′ → ∥ A ∥ ⊎ A′ - lem (inl x) = inl ∣ 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 + 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-⊎ : ∥ ∥ A ∥₁ ⊎ A′ ∥₁ ≡ ∥ A ⊎ A′ ∥₁ ∥∥-IdempotentL-⊎ = ua ∥∥-IdempotentL-⊎-≃ -∥∥-IdempotentR-⊎-≃ : ∥ A ⊎ ∥ A′ ∥ ∥ ≃ ∥ A ⊎ A′ ∥ +∥∥-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 ∣ + 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′ ∥ + 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 + 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-⊎ : ∥ 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′ ∥ ∎ +∥∥-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-×-≃ : ∥ ∥ 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′ ∣ + 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 + 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-× : ∥ ∥ A ∥₁ × A′ ∥₁ ≡ ∥ A × A′ ∥₁ ∥∥-IdempotentL-× = ua ∥∥-IdempotentL-×-≃ -∥∥-IdempotentR-×-≃ : ∥ A × ∥ A′ ∥ ∥ ≃ ∥ A × A′ ∥ +∥∥-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′ + 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 + 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-× : ∥ 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′ ∥ ∎ +∥∥-Idempotent-× : {A : Type ℓ} {A′ : Type ℓ'} → ∥ ∥ A ∥₁ × ∥ A′ ∥₁ ∥₁ ≡ ∥ A × A′ ∥₁ +∥∥-Idempotent-× {A = A} {A′} = ∥ ∥ A ∥₁ × ∥ A′ ∥₁ ∥₁ ≡⟨ ∥∥-IdempotentR-× ⟩ + ∥ ∥ A ∥₁ × A′ ∥₁ ≡⟨ ∥∥-IdempotentL-× ⟩ + ∥ A × A′ ∥₁ ∎ -∥∥-Idempotent-×-≃ : {A : Type ℓ} {A′ : Type ℓ'} → ∥ ∥ A ∥ × ∥ A′ ∥ ∥ ≃ ∥ A × A′ ∥ +∥∥-Idempotent-×-≃ : {A : Type ℓ} {A′ : Type ℓ'} → ∥ ∥ A ∥₁ × ∥ A′ ∥₁ ∥₁ ≃ ∥ A × A′ ∥₁ ∥∥-Idempotent-×-≃ {A = A} {A′} = compEquiv ∥∥-IdempotentR-×-≃ ∥∥-IdempotentL-×-≃ -∥∥-×-≃ : {A : Type ℓ} {A′ : Type ℓ'} → ∥ A ∥ × ∥ A′ ∥ ≃ ∥ A × A′ ∥ +∥∥-×-≃ : {A : Type ℓ} {A′ : Type ℓ'} → ∥ A ∥₁ × ∥ A′ ∥₁ ≃ ∥ A × A′ ∥₁ ∥∥-×-≃ {A = A} {A′} = compEquiv (invEquiv (propTruncIdempotent≃ (isProp× isPropPropTrunc isPropPropTrunc))) ∥∥-Idempotent-×-≃ -∥∥-× : {A : Type ℓ} {A′ : Type ℓ'} → ∥ A ∥ × ∥ A′ ∥ ≡ ∥ A × A′ ∥ +∥∥-× : {A : Type ℓ} {A′ : Type ℓ'} → ∥ A ∥₁ × ∥ A′ ∥₁ ≡ ∥ A × A′ ∥₁ ∥∥-× = ua ∥∥-×-≃ -- using this we get a convenient recursor/eliminator for binary functions into sets rec2→Set : {A B C : Type ℓ} (Cset : isSet C) → (f : A → B → C) → (∀ (a a' : A) (b b' : B) → f a b ≡ f a' b') - → ∥ A ∥ → ∥ B ∥ → C + → ∥ A ∥₁ → ∥ B ∥₁ → C rec2→Set {A = A} {B = B} {C = C} Cset f fconst = curry (g ∘ ∥∥-×-≃ .fst) where - g : ∥ A × B ∥ → C + g : ∥ A × B ∥₁ → C g = rec→Set Cset (uncurry f) λ x y → fconst (fst x) (fst y) (snd x) (snd y) diff --git a/Cubical/HITs/Pushout/KrausVonRaumer.agda b/Cubical/HITs/Pushout/KrausVonRaumer.agda index 35f48f4557..a3f0847705 100644 --- a/Cubical/HITs/Pushout/KrausVonRaumer.agda +++ b/Cubical/HITs/Pushout/KrausVonRaumer.agda @@ -9,11 +9,20 @@ https://arxiv.org/abs/1901.06022 {-# OPTIONS --safe #-} module Cubical.HITs.Pushout.KrausVonRaumer where -open import Cubical.Foundations.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.Equiv.Properties +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence open import Cubical.Functions.Embedding + open import Cubical.Data.Sigma open import Cubical.Data.Nat open import Cubical.Data.Sum as ⊎ + open import Cubical.HITs.PropositionalTruncation as Trunc open import Cubical.HITs.Pushout.Base as ⊔ open import Cubical.HITs.Pushout.Properties @@ -151,12 +160,12 @@ isEmbeddingInr f g fEmb c₀ c₁ = -- Further Application: Pushouts of emedding-spans of n-Types are n-Types, for n≥0 module _ (f : A → B) (g : A → C) where inlrJointlySurjective : - (z : Pushout f g) → ∥ Σ[ x ∈ (B ⊎ C) ] (⊎.rec inl inr x) ≡ z ∥ + (z : Pushout f g) → ∥ Σ[ x ∈ (B ⊎ C) ] (⊎.rec inl inr x) ≡ z ∥₁ inlrJointlySurjective = elimProp _ (λ _ → isPropPropTrunc) - (λ b → ∣ ⊎.inl b , refl ∣) - (λ c → ∣ ⊎.inr c , refl ∣) + (λ b → ∣ ⊎.inl b , refl ∣₁) + (λ c → ∣ ⊎.inr c , refl ∣₁) preserveHLevelEmbedding : {n : HLevel} diff --git a/Cubical/HITs/RPn/Base.agda b/Cubical/HITs/RPn/Base.agda index f3df232e61..9db399f902 100644 --- a/Cubical/HITs/RPn/Base.agda +++ b/Cubical/HITs/RPn/Base.agda @@ -53,7 +53,7 @@ data RP² : Type₀ where 2-EltPointed₀ = PointedEqvTo ℓ-zero Bool -- Σ[ X ∈ Type₀ ] X × ∥ X ≃ Bool ∥ Bool* : 2-EltType₀ -Bool* = Bool , ∣ idEquiv _ ∣ +Bool* = Bool , ∣ idEquiv _ ∣₁ -- Our first goal is to 'lift' `_⊕_ : Bool → Bool ≃ Bool` to a function `_⊕_ : A → A ≃ Bool` @@ -80,7 +80,7 @@ snd (isContrBoolPointedEquiv x) (e , p) -- 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 ((Bool , false , ∣ idEquiv Bool ∣₁) ≃[ PointedEqvToEquivStr Bool ] X∙) isContr-2-EltPointedEquiv (X , x , ∣e∣) = PropTrunc.rec isPropIsContr (λ e → J (λ X∙ _ → isContr ((Bool , false) ≃[ PointedEquivStr ] X∙)) @@ -110,7 +110,7 @@ module ⊕* (X : 2-EltType₀) where 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 + 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∙)) @@ -133,7 +133,7 @@ module ⊕* (X : 2-EltType₀) where -- 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 ∣) +fst isContr-2-EltPointed = (Bool , false , ∣ idEquiv Bool ∣₁) snd isContr-2-EltPointed A∙ = PointedEqvTo-sip Bool _ A∙ (fst (isContr-2-EltPointedEquiv A∙)) @@ -163,8 +163,8 @@ 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 _ ∣) + ∣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 diff --git a/Cubical/HITs/Rationals/QuoQ/Properties.agda b/Cubical/HITs/Rationals/QuoQ/Properties.agda index cc8932b221..c2c05ffdd5 100644 --- a/Cubical/HITs/Rationals/QuoQ/Properties.agda +++ b/Cubical/HITs/Rationals/QuoQ/Properties.agda @@ -1,7 +1,11 @@ {-# OPTIONS --safe #-} module Cubical.HITs.Rationals.QuoQ.Properties where -open import Cubical.Foundations.Everything hiding (_⁻¹) +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws hiding (_⁻¹) +open import Cubical.Foundations.Univalence 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 _//_) @@ -242,4 +246,3 @@ x - y = x + (- 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/S1/Properties.agda b/Cubical/HITs/S1/Properties.agda index 86f7bd3eec..53b11b6655 100644 --- a/Cubical/HITs/S1/Properties.agda +++ b/Cubical/HITs/S1/Properties.agda @@ -19,14 +19,10 @@ elim : ∀ {ℓ} (C : S¹ → Type ℓ) (b : C base) (l : PathP (λ i → C (loo elim _ b l base = b elim _ b l (loop i) = l i --- Not sure about this name -ind : ∀ {ℓ} (C : S¹ → Type ℓ) (b : C base) (l : subst C loop b ≡ b) → (x : S¹) → C x -ind _ b l = elim _ b (toPathP l) - -isConnectedS¹ : (s : S¹) → ∥ base ≡ s ∥ -isConnectedS¹ base = ∣ refl ∣ +isConnectedS¹ : (s : S¹) → ∥ base ≡ s ∥₁ +isConnectedS¹ base = ∣ refl ∣₁ isConnectedS¹ (loop i) = - squash ∣ (λ j → loop (i ∧ j)) ∣ ∣ (λ j → loop (i ∨ ~ j)) ∣ i + squash₁ ∣ (λ j → loop (i ∧ j)) ∣₁ ∣ (λ j → loop (i ∨ ~ j)) ∣₁ i isGroupoidS¹ : isGroupoid S¹ isGroupoidS¹ s t = diff --git a/Cubical/HITs/SetQuotients/EqClass.agda b/Cubical/HITs/SetQuotients/EqClass.agda index 917cadff1c..3c2398ee1a 100644 --- a/Cubical/HITs/SetQuotients/EqClass.agda +++ b/Cubical/HITs/SetQuotients/EqClass.agda @@ -48,7 +48,7 @@ module _ (R : X → X → Type ℓ'') where isEqClass : ℙ → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') - isEqClass P = ∥ Σ[ x ∈ X ] ((a : X) → P a .fst ≃ ∥ R a x ∥) ∥ + isEqClass P = ∥ Σ[ x ∈ X ] ((a : X) → P a .fst ≃ ∥ R a x ∥₁) ∥₁ isPropIsEqClass : (P : ℙ) → isProp (isEqClass P) isPropIsEqClass P = isPropPropTrunc @@ -89,17 +89,17 @@ module _ (R : X → X → Type ℓ'') (h : isEquivRel R) where - ∥Rx∥Iso : (x x' : X)(r : R x x') → (a : X) → Iso ∥ R a x ∥ ∥ R a x' ∥ - ∥Rx∥Iso x x' r a .fun = Prop.rec isPropPropTrunc (λ r' → ∣ h .transitive _ _ _ r' r ∣) - ∥Rx∥Iso x x' r a .inv = Prop.rec isPropPropTrunc (λ r' → ∣ h .transitive _ _ _ r' (h .symmetric _ _ r) ∣) + ∥Rx∥Iso : (x x' : X)(r : R x x') → (a : X) → Iso ∥ R a x ∥₁ ∥ R a x' ∥₁ + ∥Rx∥Iso x x' r a .fun = Prop.rec isPropPropTrunc (λ r' → ∣ h .transitive _ _ _ r' r ∣₁) + ∥Rx∥Iso x x' r a .inv = Prop.rec isPropPropTrunc (λ r' → ∣ h .transitive _ _ _ r' (h .symmetric _ _ r) ∣₁) ∥Rx∥Iso x x' r a .leftInv _ = isPropPropTrunc _ _ ∥Rx∥Iso x x' r a .rightInv _ = isPropPropTrunc _ _ - isEqClass∥Rx∥ : (x : X) → isEqClass X R (λ a → ∥ R a x ∥ , isPropPropTrunc) - isEqClass∥Rx∥ x = ∣ x , (λ _ → idEquiv _) ∣ + isEqClass∥Rx∥ : (x : X) → isEqClass X R (λ a → ∥ R a x ∥₁ , isPropPropTrunc) + isEqClass∥Rx∥ x = ∣ x , (λ _ → idEquiv _) ∣₁ ∥R∥ : (x : X) → X ∥ R - ∥R∥ x = (λ a → ∥ R a x ∥ , isPropPropTrunc) , isEqClass∥Rx∥ x + ∥R∥ x = (λ a → ∥ R a x ∥₁ , isPropPropTrunc) , isEqClass∥Rx∥ x ∥Rx∥Path : (x x' : X)(r : R x x') → ∥R∥ x ≡ ∥R∥ x' ∥Rx∥Path x x' r i .fst a .fst = ua (isoToEquiv (∥Rx∥Iso x x' r a)) i @@ -114,8 +114,8 @@ module _ /→∥ : X / R → X ∥ R /→∥ = SetQuot.rec (isSet∥ X R) ∥R∥ (λ x x' r → ∥Rx∥Path x x' r) - inj/→∥' : (x x' : X) → ∥R∥ x ≡ ∥R∥ x' → ∥ R x x' ∥ - inj/→∥' x x' p = transport (λ i → p i .fst x .fst) ∣ h .reflexive x ∣ + inj/→∥' : (x x' : X) → ∥R∥ x ≡ ∥R∥ x' → ∥ R x x' ∥₁ + inj/→∥' x x' p = transport (λ i → p i .fst x .fst) ∣ h .reflexive x ∣₁ inj/→∥ : (x y : X / R) → /→∥ x ≡ /→∥ y → x ≡ y inj/→∥ = @@ -126,7 +126,7 @@ module _ isEmbedding/→∥ : isEmbedding /→∥ isEmbedding/→∥ = injEmbedding squash/ (isSet∥ X R) (λ {x} {y} → inj/→∥ x y) - surj/→∥ : (P : X ∥ R) → ((x , _) : Σ[ x ∈ X ] ((a : X) → P .fst a .fst ≃ ∥ R a x ∥)) → ∥R∥ x ≡ P + surj/→∥ : (P : X ∥ R) → ((x , _) : Σ[ x ∈ X ] ((a : X) → P .fst a .fst ≃ ∥ R a x ∥₁)) → ∥R∥ x ≡ P surj/→∥ P (x , p) i .fst a .fst = ua (p a) (~ i) surj/→∥ P (x , p) i .fst a .snd = isProp→PathP {B = λ i → isProp (surj/→∥ P (x , p) i .fst a .fst)} @@ -137,7 +137,7 @@ module _ (isEqClass∥Rx∥ x) (P .snd) i isSurjection/→∥ : isSurjection /→∥ - isSurjection/→∥ P = Prop.rec isPropPropTrunc (λ p → ∣ [ p .fst ] , surj/→∥ P p ∣) (P .snd) + isSurjection/→∥ P = Prop.rec isPropPropTrunc (λ p → ∣ [ p .fst ] , surj/→∥ P p ∣₁) (P .snd) -- both definitions are equivalent equivQuot : X / R ≃ X ∥ R diff --git a/Cubical/HITs/SetQuotients/Properties.agda b/Cubical/HITs/SetQuotients/Properties.agda index ae6d571a50..3d675eed1d 100644 --- a/Cubical/HITs/SetQuotients/Properties.agda +++ b/Cubical/HITs/SetQuotients/Properties.agda @@ -27,9 +27,10 @@ open import Cubical.Relation.Nullary open import Cubical.Relation.Binary.Base open import Cubical.HITs.TypeQuotients as TypeQuot using (_/ₜ_ ; [_] ; eq/) -open import Cubical.HITs.PropositionalTruncation as PropTrunc using (∥_∥ ; ∣_∣ ; squash) renaming (rec to propRec) -open import Cubical.HITs.SetTruncation as SetTrunc using (∥_∥₂ ; ∣_∣₂ ; squash₂ - ; isSetSetTrunc) +open import Cubical.HITs.PropositionalTruncation as PropTrunc + using (∥_∥₁ ; ∣_∣₁ ; squash₁) renaming (rec to propRec) +open import Cubical.HITs.SetTruncation as SetTrunc + using (∥_∥₂ ; ∣_∣₂ ; squash₂ ; isSetSetTrunc) private @@ -92,7 +93,7 @@ elimContr2 contr = -- lemma 6.10.2 in hott book []surjective : (x : A / R) → ∃[ a ∈ A ] [ a ] ≡ x -[]surjective = elimProp (λ x → squash) (λ a → ∣ a , refl ∣) +[]surjective = elimProp (λ x → squash₁) (λ a → ∣ a , refl ∣₁) elim : {P : A / R → Type ℓ} → (∀ x → isSet (P x)) @@ -308,20 +309,20 @@ discreteSetQuotients {A = A} {R = R} Adis Rprop Req Rdec = -- 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 ∣ +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 : 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 : 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/ _ _) @@ -329,8 +330,8 @@ isEquivRel→TruncIso {A = A} {R = R} Req a b = (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 ∣ + ∥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/Fibers.agda b/Cubical/HITs/SetTruncation/Fibers.agda index f183bc1eb2..64a1362b3a 100644 --- a/Cubical/HITs/SetTruncation/Fibers.agda +++ b/Cubical/HITs/SetTruncation/Fibers.agda @@ -68,7 +68,7 @@ module _ Set.elim2 (λ _ _ → isProp→isSet (isPropΠ (λ _ → isSetFiber∥∥₂ _ _))) λ _ _ p → ΣPathP (p , isSet→isSet' squash₂ _ _ _ _) - mereFiber→∥fiber∥₂/R : (x : X) → ∥ f x ≡ y ∥ → ∥ fiber f y ∥₂ / fiberRel + mereFiber→∥fiber∥₂/R : (x : X) → ∥ f x ≡ y ∥₁ → ∥ fiber f y ∥₂ / fiberRel mereFiber→∥fiber∥₂/R x = Prop.rec→Set squash/ (λ p → [ ∣ _ , p ∣₂ ]) (λ _ _ → eq/ _ _ refl) fiber∥∥₂→∥fiber∥₂/R : fiber ∥f∥₂ ∣ y ∣₂ → ∥ fiber f y ∥₂ / fiberRel @@ -132,7 +132,7 @@ module _ ∣transport∣ = Set.rec2 squash₂ (λ s (x , q) → ∣ x , q ∙ s ∣₂) fiberRel2 : (x x' : ∥ fiber f y ∥₂) → Type (ℓ-max ℓ ℓ') - fiberRel2 x x' = ∥ Σ[ s ∈ ∥ y ≡ y ∥₂ ] ∣transport∣ s x ≡ x' ∥ + fiberRel2 x x' = ∥ Σ[ s ∈ ∥ y ≡ y ∥₂ ] ∣transport∣ s x ≡ x' ∥₁ fiberRel2→1 : ∀ x x' → fiberRel2 x x' → fiberRel x x' fiberRel2→1 = @@ -144,13 +144,13 @@ module _ fiberRel1→2 : ∀ x x' → fiberRel x x' → fiberRel2 x x' fiberRel1→2 = - Set.elim2 (λ _ _ → isSetΠ λ _ → isProp→isSet squash) λ a b p → - Prop.rec squash + Set.elim2 (λ _ _ → isSetΠ λ _ → isProp→isSet squash₁) λ a b p → + Prop.rec squash₁ (λ q → let filler = doubleCompPath-filler (sym (a .snd)) (cong f q) (b .snd) in - ∣ ∣ filler i1 ∣₂ , cong ∣_∣₂ (ΣPathP (q , adjustLemma (flipSquare filler))) ∣) + ∣ ∣ filler i1 ∣₂ , cong ∣_∣₂ (ΣPathP (q , adjustLemma (flipSquare filler))) ∣₁) (PathIdTrunc₀Iso .Iso.fun p) where adjustLemma : {x y z w : Y} {p : x ≡ y} {q : x ≡ z} {r : z ≡ w} {s : y ≡ w} @@ -167,4 +167,4 @@ module _ fiberRel1≃2 : ∀ x x' → fiberRel x x' ≃ fiberRel2 x x' fiberRel1≃2 _ _ = - propBiimpl→Equiv (squash₂ _ _) squash (fiberRel1→2 _ _) (fiberRel2→1 _ _) + propBiimpl→Equiv (squash₂ _ _) squash₁ (fiberRel1→2 _ _) (fiberRel2→1 _ _) diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index b6ed287133..31c11082d3 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -91,22 +91,22 @@ module rec→Gpd {A : Type ℓ} {B : Type ℓ'} (Bgpd : isGroupoid B) (f : A → 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 + ε : ∀ (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 → 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 + (ε* 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) = ε* 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) @@ -116,12 +116,12 @@ module rec→Gpd {A : Type ℓ} {B : Type ℓ'} (Bgpd : isGroupoid B) (f : A → 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 + (ε* : ∀ (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) = ε* 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 @@ -131,7 +131,7 @@ module rec→Gpd {A : Type ℓ} {B : Type ℓ'} (Bgpd : isGroupoid B) (f : A → 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 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) @@ -147,7 +147,7 @@ module rec→Gpd {A : Type ℓ} {B : Type ℓ'} (Bgpd : isGroupoid B) (f : A → {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 : 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) @@ -173,11 +173,11 @@ module rec→Gpd {A : Type ℓ} {B : Type ℓ'} (Bgpd : isGroupoid B) (f : A → 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 ∣ + 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 = 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 _ _) @@ -192,10 +192,10 @@ module rec→Gpd {A : Type ℓ} {B : Type ℓ'} (Bgpd : isGroupoid B) (f : A → f₁ : H → B f₁ = Hrec.fun Bgpd f εᶠ λ _ _ _ → refl where - εᶠ : (a b : A) → ∥ a ≡ b ∥ → f a ≡ f b + εᶠ : (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 + -- 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 η @@ -318,11 +318,11 @@ Iso.leftInv IsoSetTruncateSndΣ = elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ _ → refl -PathIdTrunc₀Iso : {a b : A} → Iso (∣ a ∣₂ ≡ ∣ b ∣₂) ∥ a ≡ b ∥ +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 ∣ + (λ a → ∥ a ≡ b ∥₁ , squash₁) (p (~ i)) .fst) + ∣ refl ∣₁ Iso.inv PathIdTrunc₀Iso = pRec (squash₂ _ _) (cong ∣_∣₂) -Iso.rightInv PathIdTrunc₀Iso _ = squash _ _ +Iso.rightInv PathIdTrunc₀Iso _ = squash₁ _ _ Iso.leftInv PathIdTrunc₀Iso _ = squash₂ _ _ _ _ diff --git a/Cubical/HITs/Truncation/FromNegTwo/Properties.agda b/Cubical/HITs/Truncation/FromNegTwo/Properties.agda index cb97ab7df7..e34fb27b53 100644 --- a/Cubical/HITs/Truncation/FromNegTwo/Properties.agda +++ b/Cubical/HITs/Truncation/FromNegTwo/Properties.agda @@ -23,8 +23,7 @@ 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.PropositionalTruncation as PropTrunc using (∥_∥₁; ∣_∣₁; squash₁) 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₄) diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index 8c7bed9f64..2c2b73452a 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -27,8 +27,7 @@ open import Cubical.HITs.S1 hiding (rec ; elim) open import Cubical.HITs.Susp.Base open import Cubical.HITs.Nullification as Null hiding (rec ; elim) -open import Cubical.HITs.PropositionalTruncation as PropTrunc - renaming (∥_∥ to ∥_∥₁; ∣_∣ to ∣_∣₁; squash to squash₁) using () +open import Cubical.HITs.PropositionalTruncation as PropTrunc using (∥_∥₁ ; ∣_∣₁ ; squash₁) 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₄) diff --git a/Cubical/HITs/TypeQuotients/Properties.agda b/Cubical/HITs/TypeQuotients/Properties.agda index 1616700612..4371a7ea2f 100644 --- a/Cubical/HITs/TypeQuotients/Properties.agda +++ b/Cubical/HITs/TypeQuotients/Properties.agda @@ -7,19 +7,13 @@ 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) +open import Cubical.HITs.TypeQuotients.Base private variable diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index b6307b8f09..c157c56c12 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -1,28 +1,36 @@ {-# OPTIONS --safe #-} module Cubical.Homotopy.Connected where -open import Cubical.Core.Everything -open import Cubical.Foundations.Everything +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.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels open import Cubical.Foundations.Path open import Cubical.Foundations.Univalence + open import Cubical.Functions.Fibration + +open import Cubical.Data.Unit +open import Cubical.Data.Bool 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 +open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec) + +open import Cubical.Homotopy.Loopspace + + -- Note that relative to most sources, this notation is off by +2 isConnected : ∀ {ℓ} (n : HLevel) (A : Type ℓ) → Type ℓ diff --git a/Cubical/Homotopy/Freudenthal.agda b/Cubical/Homotopy/Freudenthal.agda index ba05ca7250..d3155658d9 100644 --- a/Cubical/Homotopy/Freudenthal.agda +++ b/Cubical/Homotopy/Freudenthal.agda @@ -6,23 +6,37 @@ Freudenthal suspension theorem {-# OPTIONS --safe #-} module Cubical.Homotopy.Freudenthal where -open import Cubical.Foundations.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.GroupoidLaws +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Path +open import Cubical.Foundations.Univalence + open import Cubical.Data.Nat open import Cubical.Data.Nat.Order open import Cubical.Data.Sigma open import Cubical.Data.Empty renaming (rec to ⊥-rec) -open import Cubical.HITs.Nullification -open import Cubical.HITs.Susp renaming (toSusp to σ) + open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec ; elim to trElim) -open import Cubical.Homotopy.Connected -open import Cubical.Homotopy.WedgeConnectivity -open import Cubical.Homotopy.Loopspace -open import Cubical.HITs.SmashProduct +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.SmashProduct +open import Cubical.HITs.Nullification 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.HITs.Sn + +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.WedgeConnectivity +open import Cubical.Homotopy.Loopspace + + + module _ {ℓ} (n : HLevel) {A : Pointed ℓ} (connA : isConnected (suc (suc n)) (typ A)) where diff --git a/Cubical/Homotopy/Group/LES.agda b/Cubical/Homotopy/Group/LES.agda index 7dcec83017..72363726ec 100644 --- a/Cubical/Homotopy/Group/LES.agda +++ b/Cubical/Homotopy/Group/LES.agda @@ -520,10 +520,10 @@ module setTruncLemmas {ℓ ℓ' ℓ'' : Level} ker⊂im : ((x : typ (Ω ((Ω^ m) B))) → isInKer∙ g x → isInIm∙ f x) → (x : π (suc m) B) → isInKer (_ , e₂) x → isInIm (_ , e₁) x ker⊂im ind = - sElim (λ _ → isSetΠ λ _ → isProp→isSet squash) + sElim (λ _ → isSetΠ λ _ → isProp→isSet squash₁) λ p ker → - pRec squash - (λ ker∙ → ∣ ∣ ind p ker∙ .fst ∣₂ , cong ∣_∣₂ (ind p ker∙ .snd) ∣ ) + pRec squash₁ + (λ ker∙ → ∣ ∣ ind p ker∙ .fst ∣₂ , cong ∣_∣₂ (ind p ker∙ .snd) ∣₁ ) (fun PathIdTrunc₀Iso ker) im⊂ker : ((x : typ (Ω ((Ω^ m) B))) → isInIm∙ f x → isInKer∙ g x) diff --git a/Cubical/Homotopy/Group/Pi3S2.agda b/Cubical/Homotopy/Group/Pi3S2.agda index 17a49c1878..57f3cef1a7 100644 --- a/Cubical/Homotopy/Group/Pi3S2.agda +++ b/Cubical/Homotopy/Group/Pi3S2.agda @@ -30,7 +30,7 @@ open import Cubical.HITs.S1 open import Cubical.Data.Sigma open import Cubical.Data.Nat -open import Cubical.Data.Int hiding (ℤ) +open import Cubical.Data.Int open import Cubical.Algebra.Group open import Cubical.Algebra.Group.ZAction @@ -108,7 +108,7 @@ snd π'₃S²≅π'₃TotalHopf = snd (π'∘∙Hom 2 TotalHopf→∙S²) (IsoSphereJoin 1 1)))) , refl) -π₃S²≅ℤ : GroupEquiv (π'Gr 2 (S₊∙ 2)) ℤ +π₃S²≅ℤ : GroupEquiv (π'Gr 2 (S₊∙ 2)) ℤGroup π₃S²≅ℤ = compGroupEquiv (invGroupEquiv @@ -162,7 +162,7 @@ snd π'₃S²≅π'₃TotalHopf = snd (π'∘∙Hom 2 TotalHopf→∙S²) -- As a consequence, we also get that the Hopf invariant determines -- an iso π₃S²≅ℤ -hopfInvariantEquiv : GroupEquiv (π'Gr 2 (S₊∙ 2)) ℤ +hopfInvariantEquiv : GroupEquiv (π'Gr 2 (S₊∙ 2)) ℤGroup fst (fst hopfInvariantEquiv) = HopfInvariant-π' 0 snd (fst hopfInvariantEquiv) = GroupEquivℤ-isEquiv (invGroupEquiv π₃S²≅ℤ) ∣ HopfMap ∣₂ diff --git a/Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda b/Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda index 0990321eca..051358e8e2 100644 --- a/Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda +++ b/Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda @@ -281,9 +281,9 @@ isSurjective-π₃S³→π₃TotalPushoutPath× = isSurjective-π₃S³→π₃TotalPushoutPath×' : isSurjective π₃S³→π₃TotalPushoutPath× isSurjective-π₃S³→π₃TotalPushoutPath×' = - sElim (λ _ → isProp→isSet squash) - λ p → trRec squash - (λ s → ∣ ∣ fst s ∣₂ , (cong ∣_∣₂ (snd s)) ∣) + sElim (λ _ → isProp→isSet squash₁) + λ p → trRec squash₁ + (λ s → ∣ ∣ fst s ∣₂ , (cong ∣_∣₂ (snd s)) ∣₁) (((isConnectedΩ^→ 3 3 (S³→TotalPushoutPath× , refl) isConnected-toPullback) p) .fst) @@ -306,11 +306,11 @@ S³→S²→Pushout→Unit'' = extendExact4Surjective _ _ _ _ _ _ _ _ _ isSurjective-π₃S³→π₃TotalPushoutPath× (extendExact4Surjective _ _ _ _ _ _ _ _ _ - ((sElim (λ _ → isProp→isSet squash) + ((sElim (λ _ → isProp→isSet squash₁) (λ f → ∣ ∣ (λ x → (tt , fst f x .fst) , sym (fst f x .snd)) , ΣPathP ((ΣPathP (refl , cong fst (snd f))) , λ j i → snd f j .snd (~ i)) ∣₂ - , cong ∣_∣₂ (ΣPathP (refl , sym (rUnit _))) ∣))) + , cong ∣_∣₂ (ΣPathP (refl , sym (rUnit _))) ∣₁))) P→S²→Pushout→P') -- Step 3: We need to show that the map π₃S³ → π₃S² in the above sequence @@ -407,14 +407,14 @@ fold∘W≡Whitehead = where indΠ₃S₂ : ∀ {ℓ} {A : Pointed ℓ} → (f g : A →∙ S₊∙ 2) - → fst f ≡ fst g → ∥ f ≡ g ∥ + → fst f ≡ fst g → ∥ f ≡ g ∥₁ indΠ₃S₂ {A = A} f g p = - trRec squash - (λ r → ∣ ΣPathP (p , r) ∣) + trRec squash₁ + (λ r → ∣ ΣPathP (p , r) ∣₁) (isConnectedPathP 1 {A = (λ i → p i (snd A) ≡ north)} (isConnectedPathSⁿ 1 (fst g (pt A)) north) (snd f) (snd g) .fst ) -BrunerieIsoAbstract : GroupEquiv (π'Gr 3 (S₊∙ 3)) (abstractℤ/ Brunerie) +BrunerieIsoAbstract : GroupEquiv (π'Gr 3 (S₊∙ 3)) (abstractℤGroup/ Brunerie) BrunerieIsoAbstract = compGroupEquiv π₄S³≅π₃coFib-fold∘W∙ (invGroupEquiv @@ -442,6 +442,6 @@ BrunerieIsoAbstract = -- And, finally, we get the actual iso -- (as in Corollary 3.4.5 in Brunerie's thesis) -BrunerieIso : GroupEquiv (π'Gr 3 (S₊∙ 3)) (ℤ/ Brunerie) +BrunerieIso : GroupEquiv (π'Gr 3 (S₊∙ 3)) (ℤGroup/ Brunerie) BrunerieIso = compGroupEquiv BrunerieIsoAbstract (abstractℤ/≅ℤ Brunerie) diff --git a/Cubical/Homotopy/Group/Pi4S3/QuickProof.agda b/Cubical/Homotopy/Group/Pi4S3/QuickProof.agda index 55d2dfbe3e..f1153ba8ee 100644 --- a/Cubical/Homotopy/Group/Pi4S3/QuickProof.agda +++ b/Cubical/Homotopy/Group/Pi4S3/QuickProof.agda @@ -654,7 +654,7 @@ abstract π₃*S³≅π₃*S³-abs = invGroupEquiv π₃S³≅π₃*S³ -- stated in terms of (n : ℕ) to prevent normalisation - π₃'S³≅ℤ-abs : (n : ℕ) → GroupEquiv (π'Gr n (S₊∙ (suc n))) ℤ + π₃'S³≅ℤ-abs : (n : ℕ) → GroupEquiv (π'Gr n (S₊∙ (suc n))) ℤGroup π₃'S³≅ℤ-abs n = GroupIso→GroupEquiv (πₙ'Sⁿ≅ℤ n) η↦η₁-abs : fst (fst π₃S²≅π₃*S²-abs) η ≡ η₁ @@ -686,10 +686,10 @@ abstract IsGroupHom.pres· (π₃'S³≅ℤ-abs 2 .snd) (-π' 2 x) (-π' 2 x) ∙ cong (λ x → x +Z x) - (IsGroupHom.presinv (π₃'S³≅ℤ-abs 2 .snd) x ∙ cong (inv (ℤ .snd)) p) + (IsGroupHom.presinv (π₃'S³≅ℤ-abs 2 .snd) x ∙ cong (inv (ℤGroup .snd)) p) --- Puting it all together, we get our group iso π₃(S²) ≅ ℤ -π₃'S²≅ℤ : GroupEquiv (π'Gr 2 (S₊∙ 2)) ℤ +-- Puting it all together, we get our group iso π₃(S²) ≅ ℤGroup +π₃'S²≅ℤ : GroupEquiv (π'Gr 2 (S₊∙ 2)) ℤGroup π₃'S²≅ℤ = compGroupEquiv π₃S²≅π₃*S²-abs @@ -716,7 +716,7 @@ abstract -- We combine this with the rest of the main conclusions of chapters -- 1-3 in Brunerie's thesis -BrunerieIso : GroupEquiv (π'Gr 3 (S₊∙ 3)) (ℤ/ 2) +BrunerieIso : GroupEquiv (π'Gr 3 (S₊∙ 3)) (ℤGroup/ 2) BrunerieIso = compGroupEquiv (compGroupEquiv π₄S³≅π₃coFib-fold∘W∙ diff --git a/Cubical/Homotopy/Group/Pi4S3/Summary.agda b/Cubical/Homotopy/Group/Pi4S3/Summary.agda index 7cfbd97bd7..f9755bb09a 100644 --- a/Cubical/Homotopy/Group/Pi4S3/Summary.agda +++ b/Cubical/Homotopy/Group/Pi4S3/Summary.agda @@ -56,7 +56,7 @@ open import Cubical.Algebra.Group.ZAction -- The connection to π₄(S³) is then also proved in the BrunerieNumber -- file following Corollary 3.4.5 in Guillaume Brunerie's PhD thesis. -βSpec : GroupEquiv (π 4 𝕊³) (ℤ/ β) +βSpec : GroupEquiv (π 4 𝕊³) (ℤGroup/ β) βSpec = BrunerieIso @@ -69,10 +69,10 @@ open import Cubical.Algebra.Group.ZAction β≡2 = Brunerie≡2 --- This involves a lot of theory, for example that π₃(S²) ≃ ℤ where +-- This involves a lot of theory, for example that π₃(S²) ≃ ℤGroup where -- the underlying map is induced by the Hopf invariant (which involves -- the cup product on cohomology). -_ : GroupEquiv (π 3 𝕊²) ℤ +_ : GroupEquiv (π 3 𝕊²) ℤGroup _ = hopfInvariantEquiv -- Which is a consequence of the fact that π₃(S²) is generated by the @@ -84,15 +84,15 @@ _ = π₂S³-gen-by-HopfMap -- Combining all of this gives us the desired equivalence of groups: -π₄S³≃ℤ/2ℤ : GroupEquiv (π 4 𝕊³) (ℤ/ 2) -π₄S³≃ℤ/2ℤ = subst (GroupEquiv (π 4 𝕊³)) (cong ℤ/_ β≡2) βSpec +π₄S³≃ℤ/2ℤ : GroupEquiv (π 4 𝕊³) (ℤGroup/ 2) +π₄S³≃ℤ/2ℤ = subst (GroupEquiv (π 4 𝕊³)) (cong ℤGroup/_ β≡2) βSpec -- By the SIP this induces an equality of groups: -π₄S³≡ℤ/2ℤ : π 4 𝕊³ ≡ ℤ/ 2 +π₄S³≡ℤ/2ℤ : π 4 𝕊³ ≡ ℤGroup/ 2 π₄S³≡ℤ/2ℤ = GroupPath _ _ .fst π₄S³≃ℤ/2ℤ -- As a sanity check we also establish the equality with Bool: -π₄S³≡Bool : π 4 𝕊³ ≡ Bool -π₄S³≡Bool = π₄S³≡ℤ/2ℤ ∙ GroupPath _ _ .fst (GroupIso→GroupEquiv ℤ/2≅Bool) +π₄S³≡Bool : π 4 𝕊³ ≡ BoolGroup +π₄S³≡Bool = π₄S³≡ℤ/2ℤ ∙ GroupPath _ _ .fst (GroupIso→GroupEquiv ℤGroup/2≅Bool) diff --git a/Cubical/Homotopy/Group/PinSn.agda b/Cubical/Homotopy/Group/PinSn.agda index 91de35793f..90b346cf35 100644 --- a/Cubical/Homotopy/Group/PinSn.agda +++ b/Cubical/Homotopy/Group/PinSn.agda @@ -35,7 +35,7 @@ open import Cubical.Data.Int renaming (_+_ to _+ℤ_) open import Cubical.ZCohomology.Properties open import Cubical.Algebra.Group -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Int open import Cubical.Algebra.Group.ZAction open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties diff --git a/Cubical/Homotopy/Group/SuspensionMap.agda b/Cubical/Homotopy/Group/SuspensionMap.agda index 715ea0a837..16081f0aa6 100644 --- a/Cubical/Homotopy/Group/SuspensionMap.agda +++ b/Cubical/Homotopy/Group/SuspensionMap.agda @@ -591,14 +591,14 @@ isConnectedSuspMap n m = isSurjectiveSuspMap : (n : ℕ) → isSurjective (suspMapπ'Hom {A = S₊∙ (2 + n)} (2 + n)) isSurjectiveSuspMap n = - sElim (λ _ → isProp→isSet squash) + sElim (λ _ → isProp→isSet squash₁) λ f → trRec ((subst (λ x → isOfHLevel x (isInIm (suspMapπ'Hom (2 + n)) ∣ f ∣₂)) (sym (snd (lem n n))) (isProp→isOfHLevelSuc {A = isInIm (suspMapπ'Hom (2 + n)) ∣ f ∣₂} - (fst (lem n n)) squash))) - (λ p → ∣ ∣ fst p ∣₂ , (cong ∣_∣₂ (snd p)) ∣) + (fst (lem n n)) squash₁))) + (λ p → ∣ ∣ fst p ∣₂ , (cong ∣_∣₂ (snd p)) ∣₁) (fst (isConnectedSuspMap (2 + n) (suc n) f)) where lem : (m n : ℕ) → Σ[ x ∈ ℕ ] ((m + suc (suc n) ∸ suc n) ≡ suc x) diff --git a/Cubical/Homotopy/Hopf.agda b/Cubical/Homotopy/Hopf.agda index 0f07a73101..acb0dbb140 100644 --- a/Cubical/Homotopy/Hopf.agda +++ b/Cubical/Homotopy/Hopf.agda @@ -42,7 +42,7 @@ private λ _ → refl module Hopf {ℓ : Level} {A : Pointed ℓ} {e : HSpace A} - (e-ass : AssocHSpace e) (conA : ((x y : typ A) → ∥ x ≡ y ∥)) where + (e-ass : AssocHSpace e) (conA : ((x y : typ A) → ∥ x ≡ y ∥₁)) where isEquiv-μ : (x : typ A) → isEquiv (λ z → (μ e z x)) isEquiv-μ x = pRec (isPropIsEquiv _) (J (λ x _ → isEquiv (λ z → μ e z x)) diff --git a/Cubical/Homotopy/HopfInvariant/Base.agda b/Cubical/Homotopy/HopfInvariant/Base.agda index a0d58cdc3c..3640040682 100644 --- a/Cubical/Homotopy/HopfInvariant/Base.agda +++ b/Cubical/Homotopy/HopfInvariant/Base.agda @@ -38,7 +38,7 @@ open import Cubical.Algebra.Group.ZAction open import Cubical.Algebra.Group.Exact open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Int open import Cubical.Algebra.Group.Instances.Unit open import Cubical.HITs.Pushout @@ -200,9 +200,9 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe Sphere→H : coHom 2+n (S₊ (suc (suc n))) → coHom 2+n H Sphere→H = sMap preSphere→H - conCohom2+n : (x : _) → ∥ x ≡ 0ₖ (suc (suc n)) ∥ + conCohom2+n : (x : _) → ∥ x ≡ 0ₖ (suc (suc n)) ∥₁ conCohom2+n = - coHomK-elim _ (λ _ → isProp→isOfHLevelSuc (suc n) squash) ∣ refl ∣ + coHomK-elim _ (λ _ → isProp→isOfHLevelSuc (suc n) squash₁) ∣ refl ∣₁ HIPSphereCohomIso : Iso (coHom (2 +ℕ n) (HopfInvariantPush n (fst f))) diff --git a/Cubical/Homotopy/HopfInvariant/Brunerie.agda b/Cubical/Homotopy/HopfInvariant/Brunerie.agda index 0868e9b075..355ceec5bd 100644 --- a/Cubical/Homotopy/HopfInvariant/Brunerie.agda +++ b/Cubical/Homotopy/HopfInvariant/Brunerie.agda @@ -55,7 +55,7 @@ open import Cubical.Algebra.Group.ZAction open import Cubical.Algebra.Group.Exact open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Int open import Cubical.Algebra.Group.Instances.Unit open Iso diff --git a/Cubical/Homotopy/HopfInvariant/Homomorphism.agda b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda index c14cd2c779..2af24e0119 100644 --- a/Cubical/Homotopy/HopfInvariant/Homomorphism.agda +++ b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda @@ -37,7 +37,7 @@ open import Cubical.Algebra.Group.ZAction open import Cubical.Algebra.Group.Exact open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Int open import Cubical.Algebra.Group.GroupPath open import Cubical.HITs.Pushout diff --git a/Cubical/Homotopy/HopfInvariant/HopfMap.agda b/Cubical/Homotopy/HopfInvariant/HopfMap.agda index b79995f776..278c064131 100644 --- a/Cubical/Homotopy/HopfInvariant/HopfMap.agda +++ b/Cubical/Homotopy/HopfInvariant/HopfMap.agda @@ -41,7 +41,7 @@ open import Cubical.Algebra.Group.ZAction open import Cubical.Algebra.Group.Exact open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Int open import Cubical.Algebra.Group.GroupPath open import Cubical.HITs.Pushout @@ -63,7 +63,7 @@ snd HopfMap = refl -- We use the Hopf fibration in order to connect it to the Gysin Sequence module hopfS¹ = - Hopf S1-AssocHSpace (sphereElim2 _ (λ _ _ → squash) ∣ refl ∣) + Hopf S1-AssocHSpace (sphereElim2 _ (λ _ _ → squash₁) ∣ refl ∣₁) S¹Hopf = hopfS¹.Hopf E* = hopfS¹.TotalSpacePush²' @@ -184,25 +184,25 @@ H²S²→H²CP² : coHomGr 2 (S₊ 2) .fst → coHomGr 2 CP² .fst H²S²→H²CP² = sMap H²S²→H²CP²-raw cancel-H²S²→H²CP² : (f : CP² → coHomK 2) - → ∥ H²S²→H²CP²-raw (f ∘ inr) ≡ f ∥ + → ∥ H²S²→H²CP²-raw (f ∘ inr) ≡ f ∥₁ cancel-H²S²→H²CP² f = - pRec squash - (λ p → pRec squash - (λ f → ∣ funExt f ∣) + pRec squash₁ + (λ p → pRec squash₁ + (λ f → ∣ funExt f ∣₁) (cancelLem p)) (connLem (f (inl tt))) where - connLem : (x : coHomK 2) → ∥ x ≡ 0ₖ 2 ∥ - connLem = coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣ + connLem : (x : coHomK 2) → ∥ x ≡ 0ₖ 2 ∥₁ + connLem = coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash₁) ∣ refl ∣₁ cancelLem : (p : f (inl tt) ≡ 0ₖ 2) - → ∥ ((x : CP²) → H²S²→H²CP²-raw (f ∘ inr) x ≡ f x) ∥ - cancelLem p = trRec squash (λ pp → + → ∥ ((x : CP²) → H²S²→H²CP²-raw (f ∘ inr) x ≡ f x) ∥₁ + cancelLem p = trRec squash₁ (λ pp → ∣ CP²→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) (sym p) (λ x → (λ i → f (inr x) -ₖ f (push (inl base) (~ i))) ∙∙ (λ i → f (inr x) -ₖ p i) - ∙∙ rUnitₖ 2 (f (inr x))) pp ∣) + ∙∙ rUnitₖ 2 (f (inr x))) pp ∣₁) help where help : @@ -263,7 +263,7 @@ genCP² = ∣ CP²→Groupoid (λ _ → isOfHLevelTrunc 4) ∣_∣ refl ∣₂ -inrInjective : (f g : CP² → coHomK 2) → ∥ f ∘ inr ≡ g ∘ inr ∥ +inrInjective : (f g : CP² → coHomK 2) → ∥ f ∘ inr ≡ g ∘ inr ∥₁ → Path (coHom 2 CP²) ∣ f ∣₂ ∣ g ∣₂ inrInjective f g = pRec (squash₂ _ _) (λ p → pRec (squash₂ _ _) (λ id → trRec (squash₂ _ _) @@ -279,9 +279,9 @@ inrInjective f g = pRec (squash₂ _ _) (conn (f (inl tt)) (g (inl tt)))) where - conn : (x y : coHomK 2) → ∥ x ≡ y ∥ - conn = coHomK-elim _ (λ _ → isSetΠ λ _ → isOfHLevelSuc 1 squash) - (coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣) + conn : (x y : coHomK 2) → ∥ x ≡ y ∥₁ + conn = coHomK-elim _ (λ _ → isSetΠ λ _ → isOfHLevelSuc 1 squash₁) + (coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash₁) ∣ refl ∣₁) conn2 : (x y : coHomK 2) (p q : x ≡ y) → hLevelTrunc 1 (p ≡ q) conn2 x y p q = @@ -316,7 +316,7 @@ private -- as genCP², which is much easier to work with Gysin-e≡genCP² : GysinS².e ≡ genCP² Gysin-e≡genCP² = - inrInjective _ _ ∣ funExt (λ x → funExt⁻ (cong fst (main x)) south) ∣ + inrInjective _ _ ∣ funExt (λ x → funExt⁻ (cong fst (main x)) south) ∣₁ where mainId : (x : Σ (S₊ 2) hopfS¹.Hopf) → Path (hLevelTrunc 4 _) ∣ fst x ∣ ∣ north ∣ diff --git a/Cubical/Homotopy/MayerVietorisCofiber.agda b/Cubical/Homotopy/MayerVietorisCofiber.agda index 894352405b..132df42c28 100644 --- a/Cubical/Homotopy/MayerVietorisCofiber.agda +++ b/Cubical/Homotopy/MayerVietorisCofiber.agda @@ -12,10 +12,14 @@ 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.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv 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 diff --git a/Cubical/Homotopy/WedgeConnectivity.agda b/Cubical/Homotopy/WedgeConnectivity.agda index 27f1795c8d..15009123d4 100644 --- a/Cubical/Homotopy/WedgeConnectivity.agda +++ b/Cubical/Homotopy/WedgeConnectivity.agda @@ -1,13 +1,22 @@ {-# OPTIONS --safe #-} module Cubical.Homotopy.WedgeConnectivity where -open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels + open import Cubical.Data.Nat open import Cubical.Data.Sigma + 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)) diff --git a/Cubical/Homotopy/Whitehead.agda b/Cubical/Homotopy/Whitehead.agda index af7df6635a..2198fdee3b 100644 --- a/Cubical/Homotopy/Whitehead.agda +++ b/Cubical/Homotopy/Whitehead.agda @@ -1,9 +1,12 @@ {-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.Homotopy.Whitehead where -open import Cubical.Foundations.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.Pointed +open import Cubical.Foundations.GroupoidLaws open import Cubical.Data.Nat open import Cubical.Data.Sigma diff --git a/Cubical/Induction/WellFounded.agda b/Cubical/Induction/WellFounded.agda index 561cb3ca90..d82de49855 100644 --- a/Cubical/Induction/WellFounded.agda +++ b/Cubical/Induction/WellFounded.agda @@ -2,7 +2,7 @@ module Cubical.Induction.WellFounded where -open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude Rel : ∀{ℓ} → Type ℓ → ∀ ℓ' → Type _ Rel A ℓ = A → A → Type ℓ diff --git a/Cubical/Modalities/Lex.agda b/Cubical/Modalities/Lex.agda index 6aa0e4770b..2f4579f990 100644 --- a/Cubical/Modalities/Lex.agda +++ b/Cubical/Modalities/Lex.agda @@ -1,9 +1,16 @@ {-# OPTIONS --safe --postfix-projections #-} -open import Cubical.Foundations.Everything renaming (uncurry to λ⟨,⟩_) -open import Cubical.Data.Sigma.Properties +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function renaming (uncurry to λ⟨,⟩_) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport open import Cubical.Foundations.CartesianKanOps +open import Cubical.Data.Sigma.Properties + + module Cubical.Modalities.Lex (◯ : ∀ {ℓ} → Type ℓ → Type ℓ) (η : ∀ {ℓ} {A : Type ℓ} → A → ◯ A) diff --git a/Cubical/Modalities/Modality.agda b/Cubical/Modalities/Modality.agda index fc011e597d..8f5f58c8ca 100644 --- a/Cubical/Modalities/Modality.agda +++ b/Cubical/Modalities/Modality.agda @@ -6,7 +6,12 @@ module Cubical.Modalities.Modality where https://github.com/HoTT/HoTT-Agda/blob/master/core/lib/types/Modality.agda -} -open import Cubical.Foundations.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.Univalence +open import Cubical.Foundations.Structure record Modality ℓ : Type (ℓ-suc ℓ) where field diff --git a/Cubical/Papers/RepresentationIndependence.agda b/Cubical/Papers/RepresentationIndependence.agda index 94c495dcf6..b7d3e49064 100644 --- a/Cubical/Papers/RepresentationIndependence.agda +++ b/Cubical/Papers/RepresentationIndependence.agda @@ -93,14 +93,14 @@ open Equivalences renaming (propBiimpl→Equiv to prop≃) public -- 2.3 Higher Inductive Types -- Propositional Truncation -open PropositionalTruncation using (∥_∥ ; map) public +open PropositionalTruncation using (∥_∥₁ ; map) public open CostMonad using (Cost ; Cost≡ ; _>>=_ ; return ; fib ; fibTail) public -- Computation -_ : fib 20 ≡ (6765 , PropositionalTruncation.∣ 21890 ∣) +_ : fib 20 ≡ (6765 , PropositionalTruncation.∣ 21890 ∣₁) _ = refl -_ : fibTail 20 ≡ (6765 , PropositionalTruncation.∣ 19 ∣) +_ : fibTail 20 ≡ (6765 , PropositionalTruncation.∣ 19 ∣₁) _ = refl -- Set Quotients diff --git a/Cubical/Relation/Binary/Base.agda b/Cubical/Relation/Binary/Base.agda index e476fa5fe2..4ab8757edd 100644 --- a/Cubical/Relation/Binary/Base.agda +++ b/Cubical/Relation/Binary/Base.agda @@ -23,8 +23,8 @@ PropRel : ∀ {ℓ} (A B : Type ℓ) (ℓ' : Level) → Type (ℓ-max ℓ (ℓ-s 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 +idPropRel A .fst a a' = ∥ a ≡ a' ∥₁ +idPropRel A .snd _ _ = squash₁ invPropRel : ∀ {ℓ ℓ'} {A B : Type ℓ} → PropRel A B ℓ' → PropRel B A ℓ' @@ -33,8 +33,8 @@ invPropRel R .snd b a = R .snd a b 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 +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 diff --git a/Cubical/Relation/Nullary/Base.agda b/Cubical/Relation/Nullary/Base.agda index fdd1b3fa6d..d4ec19e12b 100644 --- a/Cubical/Relation/Nullary/Base.agda +++ b/Cubical/Relation/Nullary/Base.agda @@ -32,10 +32,10 @@ Stable A = NonEmpty A → A -- reexport propositional truncation for uniformity open Cubical.HITs.PropositionalTruncation.Base - using (∥_∥) public + using (∥_∥₁) public SplitSupport : Type ℓ → Type ℓ -SplitSupport A = ∥ A ∥ → A +SplitSupport A = ∥ A ∥₁ → A Collapsible : Type ℓ → Type ℓ Collapsible A = Σ[ f ∈ (A → A) ] 2-Constant f diff --git a/Cubical/Relation/Nullary/Properties.agda b/Cubical/Relation/Nullary/Properties.agda index e9209fa533..45ea256988 100644 --- a/Cubical/Relation/Nullary/Properties.agda +++ b/Cubical/Relation/Nullary/Properties.agda @@ -76,12 +76,12 @@ EquivPresDec : ∀ {ℓ ℓ'}{A : Type ℓ} {B : Type ℓ'} → A ≃ B → Dec A → Dec B EquivPresDec p = mapDec (p .fst) (λ f → f ∘ invEq p) -¬→¬∥∥ : ¬ A → ¬ ∥ A ∥ -¬→¬∥∥ ¬p ∣ a ∣ = ¬p a -¬→¬∥∥ ¬p (squash x y i) = isProp⊥ (¬→¬∥∥ ¬p x) (¬→¬∥∥ ¬p y) i +¬→¬∥∥ : ¬ A → ¬ ∥ A ∥₁ +¬→¬∥∥ ¬p ∣ a ∣₁ = ¬p a +¬→¬∥∥ ¬p (squash₁ x y i) = isProp⊥ (¬→¬∥∥ ¬p x) (¬→¬∥∥ ¬p y) i -Dec∥∥ : Dec A → Dec ∥ A ∥ -Dec∥∥ = mapDec ∣_∣ ¬→¬∥∥ +Dec∥∥ : Dec A → Dec ∥ A ∥₁ +Dec∥∥ = mapDec ∣_∣₁ ¬→¬∥∥ -- we have the following implications -- X ── ∣_∣ ─→ ∥ X ∥ @@ -90,13 +90,12 @@ Dec∥∥ = mapDec ∣_∣ ¬→¬∥∥ -- reexport propositional truncation for uniformity open Cubical.HITs.PropositionalTruncation.Base - using (∣_∣) public -populatedBy : ∥ A ∥ → ⟪ A ⟫ +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 + 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 @@ -121,9 +120,9 @@ PStable→SplitSupport pst = pst ∘ populatedBy SplitSupport→Collapsible : SplitSupport A → Collapsible A SplitSupport→Collapsible {A = A} hst = h , hIsConst where h : A → A - h p = hst ∣ p ∣ + h p = hst ∣ p ∣₁ hIsConst : 2-Constant h - hIsConst p q i = hst (squash ∣ p ∣ ∣ q ∣ i) + hIsConst p q i = hst (squash₁ ∣ p ∣₁ ∣ q ∣₁ i) Collapsible→SplitSupport : Collapsible A → SplitSupport A Collapsible→SplitSupport f x = fixpoint (populatedBy x f) @@ -155,9 +154,9 @@ 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 + 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 diff --git a/Cubical/Relation/ZigZag/Applications/MultiSet.agda b/Cubical/Relation/ZigZag/Applications/MultiSet.agda index df50082472..5f5b2bfd48 100644 --- a/Cubical/Relation/ZigZag/Applications/MultiSet.agda +++ b/Cubical/Relation/ZigZag/Applications/MultiSet.agda @@ -181,8 +181,8 @@ module Lists&ALists {A : Type ℓ} (discA : Discrete A) where 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 ∣ + 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) @@ -270,7 +270,7 @@ module Lists&ALists {A : Type ℓ} (discA : Discrete A) where γ : ∀ 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) , (λ 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)) diff --git a/Cubical/Relation/ZigZag/Base.agda b/Cubical/Relation/ZigZag/Base.agda index 01cf6c1191..55256279cb 100644 --- a/Cubical/Relation/ZigZag/Base.agda +++ b/Cubical/Relation/ZigZag/Base.agda @@ -66,7 +66,7 @@ module QER→Equiv {A B : Type ℓ} (R : QuasiEquivRel A B ℓ') where f a = Trunc.rec→Set squash/ ([_] ∘ fst) - (λ {(b , r) (b' , r') → eq/ b b' ∣ a , r , r' ∣}) + (λ {(b , r) (b' , r') → eq/ b b' ∣ a , r , r' ∣₁}) fPath : (a₀ : A) (s₀ : ∃[ b ∈ B ] R .fst .fst a₀ b) @@ -79,7 +79,7 @@ module QER→Equiv {A B : Type ℓ} (R : QuasiEquivRel A B ℓ') where Trunc.elim (λ _ → isPropΠ λ _ → squash/ _ _) (λ {(b₁ , r₁) → Trunc.elim (λ _ → squash/ _ _) - (λ {(b' , r₀' , r₁') → eq/ b₀ b₁ ∣ a₀ , r₀ , sim .zigzag r₀' r₁' r₁ ∣})})}) + (λ {(b' , r₀' , r₁') → eq/ b₀ b₁ ∣ a₀ , r₀ , sim .zigzag r₀' r₁' r₁ ∣₁})})}) φ : A / Rᴸ → B / Rᴿ φ [ a ] = f a (sim .fwd a) @@ -90,7 +90,7 @@ module QER→Equiv {A B : Type ℓ} (R : QuasiEquivRel A B ℓ') where 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 ∣}) + (λ {(b' , r') → eq/ b' b ∣ a , r' , r ∣₁}) (sim .fwd a) fwd≡ToRel : ∀ {a b} → φ [ a ] ≡ [ b ] → R .fst .fst a b @@ -111,7 +111,7 @@ module QER→Equiv {A B : Type ℓ} (R : QuasiEquivRel A B ℓ') where g b = Trunc.rec→Set squash/ ([_] ∘ fst) - (λ {(a , r) (a' , r') → eq/ a a' ∣ b , r , r' ∣}) + (λ {(a , r) (a' , r') → eq/ a a' ∣ b , r , r' ∣₁}) gPath : (b₀ : B) (s₀ : ∃[ a ∈ A ] R .fst .fst a b₀) @@ -124,7 +124,7 @@ module QER→Equiv {A B : Type ℓ} (R : QuasiEquivRel A B ℓ') where Trunc.elim (λ _ → isPropΠ λ _ → squash/ _ _) (λ {(a₁ , r₁) → Trunc.elim (λ _ → squash/ _ _) - (λ {(a' , r₀' , r₁') → eq/ a₀ a₁ ∣ b₀ , r₀ , sim .zigzag r₁ r₁' r₀' ∣})})}) + (λ {(a' , r₀' , r₁') → eq/ a₀ a₁ ∣ b₀ , r₀ , sim .zigzag r₁ r₁' r₀' ∣₁})})}) ψ : B / Rᴿ → A / Rᴸ ψ [ b ] = g b (sim .bwd b) @@ -134,7 +134,7 @@ module QER→Equiv {A B : Type ℓ} (R : QuasiEquivRel A B ℓ') where 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 ∣}) + (λ {(a' , r') → eq/ a' a ∣ b , r' , r ∣₁}) (sim .bwd b) private diff --git a/Cubical/Structures/Relational/Constant.agda b/Cubical/Structures/Relational/Constant.agda index 78394cef36..6c3118fdc5 100644 --- a/Cubical/Structures/Relational/Constant.agda +++ b/Cubical/Structures/Relational/Constant.agda @@ -46,7 +46,7 @@ module _ (A : hSet ℓ') where constantPositiveRel : PositiveStrRel {ℓ = ℓ} constantSuitableRel constantPositiveRel .act = constantRelAction constantPositiveRel .reflexive a = refl - constantPositiveRel .detransitive R R' p = ∣ _ , p , refl ∣ + constantPositiveRel .detransitive R R' p = ∣ _ , p , refl ∣₁ constantPositiveRel .quo R = isoToIsEquiv isom where open Iso diff --git a/Cubical/Structures/Relational/Function.agda b/Cubical/Structures/Relational/Function.agda index e618bd14bb..109002dc13 100644 --- a/Cubical/Structures/Relational/Function.agda +++ b/Cubical/Structures/Relational/Function.agda @@ -40,19 +40,19 @@ private → compPropRel (R .fst) (quotientPropRel (R .fst .fst)) .fst ≡ graphRel [_] composeWith[_] R = funExt₂ λ a t → - hPropExt squash (squash/ _ _) + hPropExt squash₁ (squash/ _ _) (Trunc.rec (squash/ _ _) (λ {(b , r , p) → eq/ a b r ∙ p })) - (λ p → ∣ a , R .snd .reflexive a , 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) + 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 ∣) + (λ r → ∣ _ , eq/ a b r , refl ∣₁) functionSuitableRel : {S : Type ℓ → Type ℓ₁} {T : Type ℓ → Type ℓ₂} {ρ₁ : StrRel S ℓ₁'} {ρ₂ : StrRel T ℓ₂'} @@ -144,7 +144,7 @@ functionSuitableRel {ρ₁ = ρ₁} {ρ₂} θ₁ σ θ₂ .symmetric R h r = θ₂ .symmetric R (h (θ₁ .symmetric (invPropRel R) r)) functionSuitableRel {ρ₁ = ρ₁} {ρ₂} θ₁ σ θ₂ .transitive R R' h h' rr' = Trunc.rec - (θ₂ .prop (λ _ _ → squash) _ _) + (θ₂ .prop (λ _ _ → squash₁) _ _) (λ {(_ , r , r') → θ₂ .transitive R R' (h r) (h' r')}) (σ .detransitive R R' rr') functionSuitableRel {ρ₁ = ρ₁} {ρ₂} θ₁ σ θ₂ .set setX = diff --git a/Cubical/Structures/Relational/Maybe.agda b/Cubical/Structures/Relational/Maybe.agda index aa641c3378..3185bc2787 100644 --- a/Cubical/Structures/Relational/Maybe.agda +++ b/Cubical/Structures/Relational/Maybe.agda @@ -76,7 +76,7 @@ maybePositiveRel : 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' {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 = diff --git a/Cubical/Structures/Relational/Pointed.agda b/Cubical/Structures/Relational/Pointed.agda index 1531eef109..13e83d479e 100644 --- a/Cubical/Structures/Relational/Pointed.agda +++ b/Cubical/Structures/Relational/Pointed.agda @@ -30,7 +30,7 @@ PointedRelStr R = R pointedSuitableRel : SuitableStrRel {ℓ = ℓ} PointedStructure PointedRelStr pointedSuitableRel .quo _ _ _ = isContrSingl _ pointedSuitableRel .symmetric _ r = r -pointedSuitableRel .transitive _ _ r r' = ∣ _ , r , r' ∣ +pointedSuitableRel .transitive _ _ r r' = ∣ _ , r , r' ∣₁ pointedSuitableRel .set setX = setX pointedSuitableRel .prop propR = propR @@ -44,7 +44,7 @@ pointedRelAction .actRel α = α pointedPositiveRel : PositiveStrRel {ℓ = ℓ} pointedSuitableRel pointedPositiveRel .act = pointedRelAction -pointedPositiveRel .reflexive x = ∣ refl ∣ +pointedPositiveRel .reflexive x = ∣ refl ∣₁ pointedPositiveRel .detransitive R R' rr' = rr' pointedPositiveRel .quo R = isoToIsEquiv isom where diff --git a/Cubical/Structures/Relational/Product.agda b/Cubical/Structures/Relational/Product.agda index b2148f36ed..11a127f6e4 100644 --- a/Cubical/Structures/Relational/Product.agda +++ b/Cubical/Structures/Relational/Product.agda @@ -81,10 +81,10 @@ productPositiveRel : productPositiveRel σ₁ σ₂ .act = productRelAction (σ₁ .act) (σ₂ .act) productPositiveRel σ₁ σ₂ .reflexive (s₁ , s₂) = σ₁ .reflexive s₁ , σ₂ .reflexive s₂ productPositiveRel σ₁ σ₂ .detransitive R R' (rr'₁ , rr'₂) = - Trunc.rec squash + Trunc.rec squash₁ (λ {(s₁ , r₁ , r₁') → - Trunc.rec squash - (λ {(s₂ , r₂ , r₂') → ∣ (s₁ , s₂) , (r₁ , r₂) , (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 = diff --git a/Cubical/Structures/TypeEqvTo.agda b/Cubical/Structures/TypeEqvTo.agda index e142bb9135..4ba9f92c39 100644 --- a/Cubical/Structures/TypeEqvTo.agda +++ b/Cubical/Structures/TypeEqvTo.agda @@ -17,22 +17,22 @@ private ℓ ℓ' ℓ'' : Level TypeEqvTo : (ℓ : Level) (X : Type ℓ') → Type (ℓ-max (ℓ-suc ℓ) ℓ') -TypeEqvTo ℓ X = TypeWithStr ℓ (λ Y → ∥ Y ≃ X ∥) +TypeEqvTo ℓ X = TypeWithStr ℓ (λ Y → ∥ Y ≃ X ∥₁) PointedEqvTo : (ℓ : Level) (X : Type ℓ') → Type (ℓ-max (ℓ-suc ℓ) ℓ') -PointedEqvTo ℓ X = TypeWithStr ℓ (λ Y → Y × ∥ Y ≃ X ∥) +PointedEqvTo ℓ X = TypeWithStr ℓ (λ Y → Y × ∥ Y ≃ X ∥₁) module _ (X : Type ℓ') where PointedEqvToStructure : Type ℓ → Type (ℓ-max ℓ ℓ') - PointedEqvToStructure = AxiomsStructure PointedStructure (λ Y _ → ∥ Y ≃ X ∥) + PointedEqvToStructure = AxiomsStructure PointedStructure (λ Y _ → ∥ Y ≃ X ∥₁) PointedEqvToEquivStr : StrEquiv PointedEqvToStructure ℓ'' - PointedEqvToEquivStr = AxiomsEquivStr PointedEquivStr (λ Y _ → ∥ Y ≃ X ∥) + PointedEqvToEquivStr = AxiomsEquivStr PointedEquivStr (λ Y _ → ∥ Y ≃ X ∥₁) pointedEqvToUnivalentStr : UnivalentStr {ℓ} PointedEqvToStructure PointedEqvToEquivStr - pointedEqvToUnivalentStr = axiomsUnivalentStr PointedEquivStr {axioms = λ Y _ → ∥ Y ≃ X ∥} - (λ _ _ → squash) pointedUnivalentStr + pointedEqvToUnivalentStr = axiomsUnivalentStr PointedEquivStr {axioms = λ Y _ → ∥ Y ≃ X ∥₁} + (λ _ _ → squash₁) pointedUnivalentStr PointedEqvToSIP : (A B : PointedEqvTo ℓ X) → A ≃[ PointedEqvToEquivStr ] B ≃ (A ≡ B) PointedEqvToSIP = SIP pointedEqvToUnivalentStr diff --git a/Cubical/ZCohomology/Base.agda b/Cubical/ZCohomology/Base.agda index da01224fad..2ecebf435c 100644 --- a/Cubical/ZCohomology/Base.agda +++ b/Cubical/ZCohomology/Base.agda @@ -1,18 +1,19 @@ {-# OPTIONS --safe #-} module Cubical.ZCohomology.Base where +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed.Base + open import Cubical.Data.Int.Base hiding (_+_) open import Cubical.Data.Nat.Base open import Cubical.Data.Sigma -open import Cubical.Foundations.Prelude -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 + open import Cubical.Homotopy.Loopspace private diff --git a/Cubical/ZCohomology/CohomologyRings/Coproduct.agda b/Cubical/ZCohomology/CohomologyRings/Coproduct.agda index f89658a56b..eacc8e11c5 100644 --- a/Cubical/ZCohomology/CohomologyRings/Coproduct.agda +++ b/Cubical/ZCohomology/CohomologyRings/Coproduct.agda @@ -17,7 +17,7 @@ open import Cubical.Algebra.Group.DirProd open import Cubical.Algebra.AbGroup open import Cubical.Algebra.Ring open import Cubical.Algebra.Ring.DirectProd -open import Cubical.Algebra.Direct-Sum.Base +open import Cubical.Algebra.DirectSum.Base open import Cubical.HITs.SetTruncation as ST diff --git a/Cubical/ZCohomology/CohomologyRings/S0.agda b/Cubical/ZCohomology/CohomologyRings/S0.agda index 2fa4402b15..21a49d3d1e 100644 --- a/Cubical/ZCohomology/CohomologyRings/S0.agda +++ b/Cubical/ZCohomology/CohomologyRings/S0.agda @@ -12,26 +12,26 @@ open import Cubical.Data.Sum open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.AbGroup.Instances.DirectSum +open import Cubical.Algebra.DirectSum.Base open import Cubical.Algebra.Ring open import Cubical.Algebra.Ring.DirectProd open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤ to ℤCR) +open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤCommRing to ℤCR) open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.QuotientRing -open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-notationZ - -open import Cubical.Algebra.Direct-Sum.Base -open import Cubical.Algebra.AbGroup.Instances.Direct-Sum open import Cubical.Algebra.Polynomials.Multivariate.Base renaming (base to baseP) open import Cubical.Algebra.CommRing.Instances.MultivariatePoly open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-Quotient - -open import Cubical.ZCohomology.RingStructure.CohomologyRing +open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-notationZ open import Cubical.HITs.Sn + +open import Cubical.ZCohomology.RingStructure.CohomologyRing open import Cubical.ZCohomology.CohomologyRings.Coproduct open import Cubical.ZCohomology.CohomologyRings.Unit + ----------------------------------------------------------------------------- -- Warning -- H*(S0) is not Z[X]/X² diff --git a/Cubical/ZCohomology/CohomologyRings/S1.agda b/Cubical/ZCohomology/CohomologyRings/S1.agda index 4165e93a20..e6f4db3c63 100644 --- a/Cubical/ZCohomology/CohomologyRings/S1.agda +++ b/Cubical/ZCohomology/CohomologyRings/S1.agda @@ -12,6 +12,7 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Transport open import Cubical.Foundations.HLevels +open import Cubical.Data.Unit open import Cubical.Data.Nat renaming (_+_ to _+n_ ; _·_ to _·n_ ; snotz to nsnotz) open import Cubical.Data.Int open import Cubical.Data.Vec @@ -20,29 +21,27 @@ open import Cubical.Data.FinData open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤG) +open import Cubical.Algebra.Group.Instances.Int renaming (ℤGroup to ℤG) +open import Cubical.Algebra.DirectSum.Base open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤ to ℤCR) open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.QuotientRing open import Cubical.Algebra.Polynomials.Multivariate.Base renaming (base to baseP) +open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤCommRing to ℤCR) open import Cubical.Algebra.CommRing.Instances.MultivariatePoly open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-Quotient open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-notationZ -open import Cubical.Algebra.Direct-Sum.Base open import Cubical.HITs.Truncation open import Cubical.HITs.SetQuotients as SQ renaming (_/_ to _/sq_) -open import Cubical.HITs.PropositionalTruncation as PT renaming (∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Sn open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.GroupStructure open import Cubical.ZCohomology.RingStructure.CupProduct open import Cubical.ZCohomology.RingStructure.CohomologyRing - -open import Cubical.Data.Unit -open import Cubical.HITs.Sn open import Cubical.ZCohomology.Groups.Sn open Iso @@ -303,8 +302,8 @@ module Equiv-S1-Properties where base-neutral-eq (suc (suc n)) = refl base-add-eq : _ - base-add-eq zero a b = (base-Poly+ _ _ _) ∙ (cong (baseP (0 ∷ [])) (sym (IsGroupHom.pres· (snd (Hᵐ-Sⁿ 0 1)) a b))) - base-add-eq one a b = (base-Poly+ _ _ _) ∙ (cong (baseP (1 ∷ [])) (sym (IsGroupHom.pres· (snd (Hᵐ-Sⁿ 1 1)) a b))) + base-add-eq zero a b = (base-poly+ _ _ _) ∙ (cong (baseP (0 ∷ [])) (sym (IsGroupHom.pres· (snd (Hᵐ-Sⁿ 0 1)) a b))) + base-add-eq one a b = (base-poly+ _ _ _) ∙ (cong (baseP (1 ∷ [])) (sym (IsGroupHom.pres· (snd (Hᵐ-Sⁿ 1 1)) a b))) base-add-eq (suc (suc n)) a b = +PℤRid _ H*-S¹→ℤ[x]-pres+ : (x y : H* (S₊ 1)) → H*-S¹→ℤ[x] ( x +H* y) ≡ H*-S¹→ℤ[x] x +Pℤ H*-S¹→ℤ[x] y @@ -347,7 +346,7 @@ module Equiv-S1-Properties where base-case : _ base-case (zero ∷ []) a = cong [_] (cong (baseP (0 ∷ [])) (rightInv (fst (Hᵐ-Sⁿ 0 1)) a)) base-case (one ∷ []) a = cong [_] (cong (baseP (1 ∷ [])) (rightInv (fst (Hᵐ-Sⁿ 1 1)) a)) - base-case (suc (suc n) ∷ []) a = eq/ 0Pℤ (baseP (suc (suc n) ∷ []) a) ∣ ((λ x → baseP (n ∷ []) (-ℤ a)) , helper) ∣₋₁ + base-case (suc (suc n) ∷ []) a = eq/ 0Pℤ (baseP (suc (suc n) ∷ []) a) ∣ ((λ x → baseP (n ∷ []) (-ℤ a)) , helper) ∣₁ where helper : _ helper = (+PℤLid _) ∙ cong₂ baseP (cong (λ X → X ∷ []) (sym (+-comm n 2))) (sym (·ℤRid _)) ∙ (sym (+PℤRid _)) diff --git a/Cubical/ZCohomology/CohomologyRings/Sn.agda b/Cubical/ZCohomology/CohomologyRings/Sn.agda index 82a0375f2b..827d202c93 100644 --- a/Cubical/ZCohomology/CohomologyRings/Sn.agda +++ b/Cubical/ZCohomology/CohomologyRings/Sn.agda @@ -8,6 +8,7 @@ open import Cubical.Foundations.Transport open import Cubical.Foundations.HLevels open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Unit open import Cubical.Data.Nat renaming (_+_ to _+n_ ; +-comm to +n-comm ; _·_ to _·n_ ; snotz to nsnotz) open import Cubical.Data.Nat.Order open import Cubical.Data.Int hiding (_+'_) @@ -18,33 +19,31 @@ open import Cubical.Data.FinData open import Cubical.Relation.Nullary open import Cubical.Algebra.Group -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤG) +open import Cubical.Algebra.Group.Instances.Int renaming (ℤGroup to ℤG) open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.DirectSum.Base open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤ to ℤCR) open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.QuotientRing open import Cubical.Algebra.Polynomials.Multivariate.Base renaming (base to baseP) +open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤCommRing to ℤCR) open import Cubical.Algebra.CommRing.Instances.MultivariatePoly open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-Quotient open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-notationZ -open import Cubical.Algebra.Direct-Sum.Base -open import Cubical.HITs.Truncation open import Cubical.HITs.SetQuotients as SQ renaming (_/_ to _/sq_) +open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.SetTruncation as ST -open import Cubical.HITs.PropositionalTruncation as PT renaming (∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.Truncation +open import Cubical.HITs.Sn 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.RingStructure.CohomologyRing - -open import Cubical.Data.Unit -open import Cubical.HITs.Sn open import Cubical.ZCohomology.Groups.Sn open Iso @@ -407,23 +406,23 @@ module Equiv-Sn-Properties (n : ℕ) where base-add-eq : (k : ℕ) → (a b : coHom k (S₊ (suc n))) → (x : partℕ k) → base-trad-H* k a x +Pℤ base-trad-H* k b x ≡ base-trad-H* k (a +ₕ b) x - base-add-eq k a b (is0 x) = base-Poly+ _ _ _ + base-add-eq k a b (is0 x) = base-poly+ _ _ _ ∙ cong (baseP (0 ∷ [])) (sym (pres· (snd (H⁰-Sⁿ≅ℤ n)) _ _)) ∙ cong (baseP (0 ∷ [])) (cong (fun (fst (H⁰-Sⁿ≅ℤ n))) (sym (subst-+ k a b 0 x))) - base-add-eq k a b (isSn x) = base-Poly+ _ _ _ + base-add-eq k a b (isSn x) = base-poly+ _ _ _ ∙ cong (baseP (1 ∷ [])) (sym (pres· (snd (Hⁿ-Sⁿ≅ℤ n)) _ _)) ∙ cong (baseP (1 ∷ [])) (cong (fun (fst (Hⁿ-Sⁿ≅ℤ n))) (sym (subst-+ k a b (suc n) x))) base-add-eq k a b (else x) = +PℤRid _ - H*-Sⁿ→ℤ[x]-gmorph : (x y : H* (S₊ (suc n))) → H*-Sⁿ→ℤ[x] ( x +H* y) ≡ H*-Sⁿ→ℤ[x] x +Pℤ H*-Sⁿ→ℤ[x] y - H*-Sⁿ→ℤ[x]-gmorph x y = refl + H*-Sⁿ→ℤ[x]-pres+ : (x y : H* (S₊ (suc n))) → H*-Sⁿ→ℤ[x] ( x +H* y) ≡ H*-Sⁿ→ℤ[x] x +Pℤ H*-Sⁿ→ℤ[x] y + H*-Sⁿ→ℤ[x]-pres+ x y = refl H*-Sⁿ→ℤ[x]/x² : H* (S₊ (suc n)) → ℤ[x]/x² H*-Sⁿ→ℤ[x]/x² = [_] ∘ H*-Sⁿ→ℤ[x] - H*-Sⁿ→ℤ[x]/x²-gmorph : (x y : H* (S₊ (suc n))) → H*-Sⁿ→ℤ[x]/x² (x +H* y) ≡ (H*-Sⁿ→ℤ[x]/x² x) +PℤI (H*-Sⁿ→ℤ[x]/x² y) - H*-Sⁿ→ℤ[x]/x²-gmorph x y = refl + H*-Sⁿ→ℤ[x]/x²-pres+ : (x y : H* (S₊ (suc n))) → H*-Sⁿ→ℤ[x]/x² (x +H* y) ≡ (H*-Sⁿ→ℤ[x]/x² x) +PℤI (H*-Sⁿ→ℤ[x]/x² y) + H*-Sⁿ→ℤ[x]/x²-pres+ x y = refl @@ -471,7 +470,7 @@ module Equiv-Sn-Properties (n : ℕ) where ∙ cong [_] (cong (baseP (1 ∷ [])) (cong (fun (fst (Hⁿ-Sⁿ≅ℤ n))) (substReflCoHom (inv (fst (Hⁿ-Sⁿ≅ℤ n)) a)))) ∙ cong [_] (cong (baseP (1 ∷ [])) (rightInv (fst (Hⁿ-Sⁿ≅ℤ n)) a)) - base-case (suc (suc k) ∷ []) a = eq/ 0Pℤ (baseP (suc (suc k) ∷ []) a) ∣ ((λ x → baseP (k ∷ []) (-ℤ a)) , helper) ∣₋₁ + base-case (suc (suc k) ∷ []) a = eq/ 0Pℤ (baseP (suc (suc k) ∷ []) a) ∣ ((λ x → baseP (k ∷ []) (-ℤ a)) , helper) ∣₁ where helper : _ helper = (+PℤLid _) ∙ cong₂ baseP (cong (λ X → X ∷ []) (sym (+n-comm k 2))) (sym (·ℤRid _)) ∙ (sym (+PℤRid _)) diff --git a/Cubical/ZCohomology/CohomologyRings/Unit.agda b/Cubical/ZCohomology/CohomologyRings/Unit.agda index 517c69d9a8..37ed3aef5e 100644 --- a/Cubical/ZCohomology/CohomologyRings/Unit.agda +++ b/Cubical/ZCohomology/CohomologyRings/Unit.agda @@ -13,30 +13,27 @@ open import Cubical.Data.FinData open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤG) +open import Cubical.Algebra.Group.Instances.Int renaming (ℤGroup to ℤG) +open import Cubical.Algebra.DirectSum.Base open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤ to ℤCR) +open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤCommRing to ℤCR) open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.QuotientRing -open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-notationZ - -open import Cubical.Algebra.Direct-Sum.Base open import Cubical.Algebra.Polynomials.Multivariate.Base renaming (base to baseP) -open import Cubical.Algebra.Polynomials.Multivariate.Equiv-A[X]X-A +open import Cubical.Algebra.Polynomials.Multivariate.EquivCarac.A[X]X-A open import Cubical.Algebra.CommRing.Instances.MultivariatePoly open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-Quotient open import Cubical.Algebra.CommRing.Instances.MultivariatePoly-notationZ open import Cubical.HITs.Truncation open import Cubical.HITs.SetQuotients as SQ renaming (_/_ to _/sq_) -open import Cubical.HITs.PropositionalTruncation as PT renaming (∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.GroupStructure open import Cubical.ZCohomology.RingStructure.CupProduct open import Cubical.ZCohomology.RingStructure.CohomologyRing - open import Cubical.ZCohomology.Groups.Unit open Iso @@ -257,7 +254,7 @@ module Equiv-Unit-Properties where base-neutral-eq (suc n) = refl base-add-eq : _ - base-add-eq zero a b = base-Poly+ _ _ _ + base-add-eq zero a b = base-poly+ _ _ _ ∙ cong (baseP (0 ∷ [])) (sym (IsGroupHom.pres· (snd H⁰-Unit≅ℤ) a b)) base-add-eq (suc n) a b = +PℤRid _ @@ -301,15 +298,15 @@ module Equiv-Unit-Properties where where base-case : _ base-case (zero ∷ []) a = refl - base-case (suc n ∷ []) a = eq/ 0Pℤ (baseP (suc n ∷ []) a) ∣ ((λ x → baseP (n ∷ []) (-ℤ a)) , foo) ∣₋₁ + base-case (suc n ∷ []) a = eq/ 0Pℤ (baseP (suc n ∷ []) a) ∣ ((λ x → baseP (n ∷ []) (-ℤ a)) , foo) ∣₁ where - foo : (0P Poly+ baseP (suc n ∷ []) (- a)) ≡ (baseP (n +n 1 ∷ []) (- a · pos 1) Poly+ 0P) - foo = (0P Poly+ baseP (suc n ∷ []) (- a)) ≡⟨ +PℤLid _ ⟩ + foo : (0P poly+ baseP (suc n ∷ []) (- a)) ≡ (baseP (n +n 1 ∷ []) (- a · pos 1) poly+ 0P) + foo = (0P poly+ baseP (suc n ∷ []) (- a)) ≡⟨ +PℤLid _ ⟩ baseP (suc n ∷ []) (- a) ≡⟨ cong₂ baseP (cong (λ X → X ∷ []) (sym ((+-suc n 0) ∙ (cong suc (+-zero n))))) (sym (·ℤRid _)) ⟩ baseP (n +n suc 0 ∷ []) (- a ·ℤ 1ℤ) ≡⟨ refl ⟩ baseP (n +n 1 ∷ []) (- a · pos 1) ≡⟨ sym (+PℤRid _) ⟩ - (baseP (n +n 1 ∷ []) (- a · pos 1) Poly+ 0P) ∎ + (baseP (n +n 1 ∷ []) (- a · pos 1) poly+ 0P) ∎ diff --git a/Cubical/ZCohomology/EilenbergSteenrodZ.agda b/Cubical/ZCohomology/EilenbergSteenrodZ.agda index dba6e4463a..ca99141e34 100644 --- a/Cubical/ZCohomology/EilenbergSteenrodZ.agda +++ b/Cubical/ZCohomology/EilenbergSteenrodZ.agda @@ -4,21 +4,6 @@ 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 @@ -27,26 +12,38 @@ 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.Relation.Nullary + +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Unit +open import Cubical.Data.Bool +open import Cubical.Data.Nat +open import Cubical.Data.Int +open import Cubical.Data.Sigma open import Cubical.Algebra.AbGroup open import Cubical.Algebra.AbGroup.Instances.Unit open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Int open import Cubical.Algebra.Group.Instances.Unit +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Truncation as T +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.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 coHomTheory open Iso open IsGroup @@ -102,7 +99,7 @@ coHomFunctor≡coHomFunctor' = funExt λ {(pos zero) → refl H0-susp : ∀ {ℓ} {A : Pointed ℓ} → isContr (coHomRed 0 (Susp (typ A) , north)) fst H0-susp = 0ₕ∙ _ snd (H0-susp {A = A}) = - sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) (funExt λ {north → sym p @@ -155,11 +152,11 @@ private 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)) + ST.map λ f → suspFunCharacFun {A = A} (suc n) f +inv (suspFunCharac {A = A} n) = ST.map (suspΩFun (suc n)) rightInv (suspFunCharac {A = A} n) = - sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) - λ f → trRec (isProp→isOfHLevelSuc n (isSetSetTrunc _ _)) + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f → T.rec (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ₖ _) @@ -183,20 +180,20 @@ leftInv (suspFunCharac {A = A} n) = -- 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 + ST.map λ 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) +inv suspFunCharac0 = ST.map λ f → suspΩFun 0 (fst f) rightInv (suspFunCharac0 {A = A}) = - sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + ST.elim (λ _ → 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) + ((transport (λ i → helix (T.rec isGroupoidS¹ (λ x → x) (rUnitₖ 1 ∣ intLoop (f x) i ∣ j))) (pos 0)))) ∙ windingℤLoop (f x))))} @@ -209,10 +206,10 @@ 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 + fst (theMorph (pos zero) f) = ST.map λ g → (λ x → fst g (fst f x)) , cong (fst g) (snd f) ∙ snd g snd (theMorph (pos zero) f) = makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (ST.elim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ f g → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl)) theMorph (pos (suc n)) f = coHomMorph _ (fst f) fst (theMorph (negsuc n) f) = idfun _ @@ -229,7 +226,7 @@ private (GroupIso→GroupEquiv ( invIso suspFunCharac0 , makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (ST.elim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ f g → cong ∣_∣₂ (funExt λ { north → refl ; south → refl ; (merid a i) j → helper a (fst f) (fst g) j i})))) @@ -244,7 +241,7 @@ private (GroupIso→GroupEquiv ( invIso (suspFunCharac {A = A} n) , makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (ST.elim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ f g → cong ∣_∣₂ (funExt λ { north → refl ; south → refl ; (merid a i) j → helper a f g j i})))) @@ -262,12 +259,12 @@ private -- naturality of the suspension isomorphism snd (Suspension (isCohomTheoryZ' {ℓ})) (f , p) (pos zero) = - funExt (sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + funExt (ST.elim (λ _ → 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 _ _) + funExt (ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ f → cong ∣_∣₂ (funExt λ {north → refl ; south → refl ; (merid a i) → refl})) @@ -280,23 +277,23 @@ private 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) + uncurry (ST.elim (λ _ → isSetΠ λ _ → isSetΣ isSetSetTrunc λ _ → isProp→isSet isPropPropTrunc) λ {(g , q) inker → ∣ g , q ∣₂ - , pRec isPropPropTrunc + , PT.rec isPropPropTrunc (λ gId → ∣ ∣ (λ { (inl tt) → 0 ; (inr b) → g b ; (push a i) → funExt⁻ (cong fst gId) a (~ i)}) , q ∣₂ - , cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl) ∣) + , cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl) ∣₁) (Iso.fun PathIdTrunc₀Iso inker)}) inv (exactnessIso (pos zero) (f , p)) = - uncurry (sElim (λ _ → isSetΠ λ _ → isSetΣ isSetSetTrunc λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + uncurry (ST.elim (λ _ → isSetΠ λ _ → isSetΣ isSetSetTrunc λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ {(g , q) inim' → ∣ g , q ∣₂ - , pRec (isSetSetTrunc _ _) + , PT.rec (isSetSetTrunc _ _) (uncurry - (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _)) + (ST.elim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _)) (λ pushmap pushId' - → pRec (isSetSetTrunc _ _) + → PT.rec (isSetSetTrunc _ _) (λ pushId → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) (funExt λ x → sym (funExt⁻ (cong fst pushId) (f x)) @@ -306,12 +303,12 @@ private (Iso.fun PathIdTrunc₀Iso pushId')))) inim'}) rightInv (exactnessIso (pos zero) (f , p)) = - uncurry (sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 + uncurry (ST.elim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 (isSetΣ isSetSetTrunc (λ _ → isProp→isSet isPropPropTrunc)) _ _) λ {(p , q) _ → Σ≡Prop (λ _ → isPropPropTrunc) refl}) leftInv (exactnessIso (pos zero) (f , p)) = - uncurry (sElim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 + uncurry (ST.elim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 (isSetΣ isSetSetTrunc (λ _ → isProp→isSet (isSetSetTrunc _ _))) _ _) λ {(p , q) _ → Σ≡Prop (λ _ → isSetSetTrunc _ _) refl}) @@ -322,11 +319,11 @@ private → 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 + λ g gId inker → PT.rec isPropPropTrunc (λ gIdTot → ∣ ∣ (λ { (inl tt) → 0ₖ _ ; (inr b) → g b ; (push a i) → funExt⁻ gIdTot a (~ i)}) ∣₂ - , cong ∣_∣₂ (funExt λ b → refl) ∣) + , 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 @@ -335,15 +332,15 @@ private → isInKer (theMorph (pos (suc n)) {A = A} {B = B} f) x inKer-helper = coHomPointedElim _ (pt B) (λ _ → isPropΠ λ _ → isSetSetTrunc _ _) - λ g gId → pRec (isSetSetTrunc _ _) + λ g gId → PT.rec (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 _ _)) + helper = ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ cg → T.rec (isProp→isOfHLevelSuc n (isSetSetTrunc _ _)) (λ p → (cong ∣_∣₂ (funExt λ x → cong cg (sym (push x)) ∙ p))) (isConnectedPathKn _ (cg (inl tt)) (0ₖ (suc n)) .fst) @@ -353,17 +350,17 @@ private isContr→Iso ((tt* , refl) , λ {(tt* , p) → Σ≡Prop (λ _ → isOfHLevelPath 1 isPropUnit* _ _) refl}) - ((tt* , ∣ tt* , refl ∣) + ((tt* , ∣ tt* , refl ∣₁) , λ {(tt* , p) → Σ≡Prop (λ _ → isPropPropTrunc) refl}) -------------------------- Dimension --------------------------- - Dimension isCohomTheoryZ' (pos zero) p = ⊥-rec (p refl) + 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 _ _)) + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (λ f → T.rec (isProp→isOfHLevelSuc n (isSetSetTrunc _ _)) + (λ f-true → T.rec (isProp→isOfHLevelSuc n (isSetSetTrunc _ _)) (λ f-false → cong ∣_∣₂ (funExt (λ {(lift true) → f-true ; (lift false) → f-false}))) (isConnectedPathKn n (0ₖ _) (f (lift false)) .fst)) diff --git a/Cubical/ZCohomology/GroupStructure.agda b/Cubical/ZCohomology/GroupStructure.agda index 2823ad4bbe..6ca5bc51ed 100644 --- a/Cubical/ZCohomology/GroupStructure.agda +++ b/Cubical/ZCohomology/GroupStructure.agda @@ -16,23 +16,23 @@ open import Cubical.Data.Sigma open import Cubical.Data.Int renaming (_+_ to _+ℤ_ ; -_ to -ℤ_) open import Cubical.Data.Nat renaming (+-assoc to +-assocℕ ; +-comm to +-commℕ) -open import Cubical.HITs.S1 hiding (rec ; elim ; ind) -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 ; isSetSetTrunc to §) -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 open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.DirProd -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Int open import Cubical.Algebra.AbGroup open import Cubical.Algebra.Semigroup open import Cubical.Algebra.Monoid +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation as ST renaming (isSetSetTrunc to §) +open import Cubical.HITs.Truncation as T + +open import Cubical.Homotopy.Loopspace + open Iso renaming (inv to inv') private @@ -57,7 +57,7 @@ infixr 34 _+ₕ∙_ → (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)) _ _) + T.elim2 (λ _ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) (wedgeconFun _ _ (λ _ _ → help _ _) (λ x → lUnit1 ∣ x ∣ ∙ sym (lUnit2 ∣ x ∣)) @@ -118,27 +118,27 @@ wedgeMapS¹ (loop i) (loop j) = _+ₖ_ : {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) +_+ₖ_ {n = suc zero} = T.map2 wedgeMapS¹ +_+ₖ_ {n = suc (suc n)} = T.rec (isOfHLevelΠ (4 + n) λ _ → isOfHLevelTrunc (4 + n)) + λ x → T.rec (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 _)) + T.elim (λ _ → 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) _ _) + help = funExt (T.elim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ _ → refl) isEquiv+ (suc n) = - trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isPropIsEquiv _)) + T.elim (λ _ → 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) + help = funExt (T.elim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) λ _ → refl) Kₙ≃Kₙ : (n : ℕ) (x : coHomK (suc n)) → coHomK (suc n) ≃ coHomK (suc n) @@ -146,8 +146,8 @@ private -ₖ_ : {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 +-ₖ_ {n = suc zero} = T.map λ {base → base ; (loop i) → (loop (~ i))} +-ₖ_ {n = suc (suc n)} = T.map λ {north → north ; south → north ; (merid a i) → ((merid (ptSn (suc n)) ∙ sym (merid a))) i} @@ -172,9 +172,9 @@ syntax -'ₖ-syntax n x y = x -[ n ]ₖ y +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} + T.elim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ { base → refl ; (loop i) → refl} -ₖ^2 {n = suc (suc n)} = - trElim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + T.elim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) λ { north → refl ; south j → ∣ merid (ptSn _) j ∣ₕ ; (merid a i) j @@ -197,14 +197,14 @@ syntax -'ₖ-syntax n x y = x -[ n ]ₖ y commₖ : (n : ℕ) → (x y : coHomK n) → x +[ n ]ₖ y ≡ y +[ n ]ₖ x commₖ zero = +Comm commₖ (suc zero) = - elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + T.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)) _ _) + T.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)) @@ -219,21 +219,21 @@ 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) _ _) + T.elim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ {base → refl ; (loop i) → refl} rUnitₖ (suc (suc n)) = - trElim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + T.elim (λ _ → 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) _ _) + T.elim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ {base → refl ; (loop i) → refl} lUnitₖ (suc (suc n)) = - trElim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + T.elim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) λ x → refl ∙≡+₁ : (p q : typ (Ω (coHomK-ptd 1))) → p ∙ q ≡ cong₂ _+ₖ_ p q @@ -245,13 +245,13 @@ lUnitₖ (suc (suc n)) = 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) _ _) + T.elim (λ _ → 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)) _ _) + T.elim (λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) λ {north → refl ; south → cong ∣_∣ (sym (merid (ptSn (suc n)))) ; (merid a i) → help a i } where s : (a : _) → _ ≡ _ @@ -271,7 +271,7 @@ lCancelₖ (suc (suc n)) = 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) _ _) +rCancelₖ (suc zero) = T.elim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ {base → refl ; (loop i) j → help j i} where help : (λ i → ∣ loop i ∣ₕ +ₖ (-ₖ ∣ loop i ∣ₕ)) ≡ refl @@ -284,7 +284,7 @@ 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) _ _) + T.elim3 (λ _ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ x → wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 3 _ _) (λ y i → rUnitₖ 1 ∣ x ∣ (~ i) +ₖ ∣ y ∣) @@ -297,7 +297,7 @@ assocₖ (suc zero) = (sym (lUnit refl)) assocₖ (suc (suc n)) = - trElim3 (λ _ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + T.elim3 (λ _ _ _ → 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 ∣))) @@ -319,7 +319,7 @@ assocₖ (suc (suc n)) = 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)) _ _) + T.elim3 (λ _ _ _ → 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 ∣))) @@ -375,13 +375,13 @@ isCommΩK (suc (suc n)) p q = ∙≡+₂ n p q ∙∙ cong+ₖ-comm (suc n) p q ∙ +Comm (-ℤ y) (-ℤ x) ∙ sym (cong₂ _+ℤ_ (sym (pos0+ _)) (sym (pos0+ _))) -distrₖ (suc zero) = - elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + T.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)) _ _) + T.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 ∣))) @@ -392,13 +392,13 @@ isCommΩK (suc (suc n)) p q = ∙≡+₂ n p q ∙∙ cong+ₖ-comm (suc n) p q ∙∙ cong (y +ℤ_) (+Comm x (0 - x)) ∙∙ cong (y +ℤ_) (minusPlus x (pos 0)) -cancelRₖ (suc zero) = - elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + T.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)) _ _) + T.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 ∣) @@ -410,13 +410,13 @@ isCommΩK (suc (suc n)) p q = ∙≡+₂ n p q ∙∙ cong+ₖ-comm (suc n) p q -+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) _ _) + T.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)) _ _) + T.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 ∣) @@ -424,13 +424,13 @@ isCommΩK (suc (suc n)) p q = ∙≡+₂ n p q ∙∙ cong+ₖ-comm (suc n) p q ---- 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 = n} = ST.rec2 § λ a b → ∣ (λ x → a x +[ n ]ₖ b x) ∣₂ -ₕ_ : {n : ℕ} → coHom n A → coHom n A --ₕ_ {n = n} = sRec § λ a → ∣ (λ x → -[ n ]ₖ a x) ∣₂ +-ₕ_ {n = n} = ST.rec § λ 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) ∣₂ +_-ₕ_ {n = n} = ST.rec2 § λ a b → ∣ (λ x → a x -[ n ]ₖ b x) ∣₂ +ₕ-syntax : (n : ℕ) → coHom n A → coHom n A → coHom n A +ₕ-syntax n = _+ₕ_ {n = n} @@ -452,57 +452,57 @@ _+'ₕ_ : {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 (§ _ _)) +rUnitₕ n = ST.elim (λ _ → 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 (§ _ _)) +lUnitₕ n = ST.elim (λ _ → 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 (§ _ _)) +rCancelₕ n = ST.elim (λ _ → 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 (§ _ _)) +lCancelₕ n = ST.elim (λ _ → 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 (§ _ _)) +assocₕ n = ST.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 (§ _ _)) +commₕ n = ST.elim2 (λ _ _ → 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 (§ _ _)) +-cancelLₕ n = ST.elim2 (λ _ _ → 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 (§ _ _)) +-cancelRₕ n = ST.elim2 (λ _ _ → 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 (§ _ _)) +-+cancelₕ n = ST.elim2 (λ _ _ → 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) +_+ₕ∙_ {n = zero} = ST.rec2 § λ { (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) +_+ₕ∙_ {n = (suc zero)} = ST.rec2 § λ { (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) + ST.rec2 § λ { (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)) +-ₕ∙_ {n = zero} = ST.rec § λ {(f , p) → ∣ (λ x → -[ 0 ]ₖ (f x)) , cong (λ x → -[ 0 ]ₖ x) p ∣₂} --ₕ∙_ {n = suc zero} = sRec § λ {(f , p) → ∣ (λ x → -ₖ (f x)) +-ₕ∙_ {n = suc zero} = ST.rec § λ {(f , p) → ∣ (λ x → -ₖ (f x)) , cong -ₖ_ p ∣₂} --ₕ∙_ {n = suc (suc n)} = sRec § λ {(f , p) → ∣ (λ x → -ₖ (f x)) +-ₕ∙_ {n = suc (suc n)} = ST.rec § λ {(f , p) → ∣ (λ x → -ₖ (f x)) , cong -ₖ_ p ∣₂} 0ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A @@ -523,16 +523,16 @@ 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 § _ _) + ST.elim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ {(f , p) (g , q) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) λ i x → commₖ 0 (f x) (g x) i)} commₕ∙ (suc zero) = - sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) + ST.elim2 (λ _ _ → 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 § _ _) + ST.elim2 (λ _ _ → 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 @@ -543,24 +543,24 @@ commₕ∙ {A = A} (suc (suc n)) = rUnitₕ∙ : {A : Pointed ℓ} (n : ℕ) (x : coHomRed n A) → x +[ n ]ₕ∙ 0ₕ∙ n ≡ x rUnitₕ∙ zero = - sElim (λ _ → isOfHLevelPath 2 § _ _) + ST.elim (λ _ → isOfHLevelPath 2 § _ _) λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) λ i x → rUnitₖ zero (f x) i)} rUnitₕ∙ (suc zero) = - sElim (λ _ → isOfHLevelPath 2 § _ _) + ST.elim (λ _ → 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 § _ _) + ST.elim (λ _ → 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 § _ _) + ST.elim (λ _ → isOfHLevelPath 2 § _ _) λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) λ i x → lUnitₖ zero (f x) i)} lUnitₕ∙ (suc zero) = - sElim (λ _ → isOfHLevelPath 2 § _ _) + ST.elim (λ _ → 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 § _ _) + ST.elim (λ _ → isOfHLevelPath 2 § _ _) λ {(f , p) → cong ∣_∣₂ (ΣPathP ((λ i x → lUnitₖ (2 + n) (f x) i) , λ i j → lUnitₖ (2 + n) (p j) i))} @@ -578,28 +578,28 @@ private rCancelₕ∙ : {A : Pointed ℓ} (n : ℕ) (x : coHomRed n A) → x +[ n ]ₕ∙ (-[ n ]ₕ∙ x) ≡ 0ₕ∙ n rCancelₕ∙ zero = - sElim (λ _ → isOfHLevelPath 2 § _ _) + ST.elim (λ _ → isOfHLevelPath 2 § _ _) λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) λ i x → rCancelₖ zero (f x) i)} rCancelₕ∙ {A = A} (suc zero) = - sElim (λ _ → isOfHLevelPath 2 § _ _) + ST.elim (λ _ → 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 § _ _) + ST.elim (λ _ → 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 § _ _) + ST.elim (λ _ → isOfHLevelPath 2 § _ _) λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) λ i x → lCancelₖ zero (f x) i)} lCancelₕ∙ {A = A} (suc zero) = - sElim (λ _ → isOfHLevelPath 2 § _ _) + ST.elim (λ _ → 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 § _ _) + ST.elim (λ _ → isOfHLevelPath 2 § _ _) λ {(f , p) → cong ∣_∣₂ (ΣPathP ((λ i x → lCancelₖ (2 + n) (f x) i) , λ i j → lCancelₖ (2 + n) (p j) i))} @@ -607,17 +607,17 @@ lCancelₕ∙ {A = A} (suc (suc n)) = 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 § _ _) + ST.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 § _ _) + ST.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 § _ _) + ST.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))} @@ -670,25 +670,25 @@ 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 ∣₂ +coHomFun n f = ST.rec § λ β → ∣ β ∘ 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 + helper zero = ST.elim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl + helper (suc zero) = ST.elim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl + helper (suc (suc n)) = ST.elim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl coHomIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) → Iso A B → GroupIso (coHomGr n B) (coHomGr n A) fun (fst (coHomIso n is)) = fst (coHomMorph n (fun is)) inv' (fst (coHomIso n is)) = fst (coHomMorph n (inv' is)) rightInv (fst (coHomIso n is)) = - sElim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ x → cong f (leftInv is x)) + ST.elim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ x → cong f (leftInv is x)) leftInv (fst (coHomIso n is)) = - sElim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ x → cong f (rightInv is x)) + ST.elim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ x → cong f (rightInv is x)) snd (coHomIso n is) = snd (coHomMorph n (fun is)) -- Alternative definition of cohomology using ΩKₙ instead. Useful for breaking proofs of group isos @@ -698,20 +698,20 @@ 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) + GroupStr._·_ coHomGrnA = ST.rec2 § λ p q → ∣ (λ x → p x ∙ q x) ∣₂ + inv coHomGrnA = ST.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 § _ _) + (∣ (λ _ → refl) ∣₂) (ST.rec2 § λ p q → ∣ (λ x → p x ∙ q x) ∣₂) (ST.map λ f x → sym (f x)) + helper = makeIsGroup § (ST.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))) + (ST.elim (λ _ → isOfHLevelPath 2 § _ _) λ p → cong ∣_∣₂ (funExt λ x → sym (rUnit (p x)))) + (ST.elim (λ _ → isOfHLevelPath 2 § _ _) λ p → cong ∣_∣₂ (funExt λ x → sym (lUnit (p x)))) + (ST.elim (λ _ → isOfHLevelPath 2 § _ _) λ p → cong ∣_∣₂ (funExt λ x → rCancel (p x))) + (ST.elim (λ _ → 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) @@ -731,17 +731,17 @@ baseChange n x = isoToEquiv is 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) + T.elim (λ _ → isOfHLevelΠ (3 + n) λ _ → isOfHLevelPath (3 + n) (isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) _ _) - (ind n) + (ind-helper n) where - ind : (n : ℕ) (a : S₊ (suc n)) (p : ∣ a ∣ₕ ≡ ∣ a ∣ₕ) → f n ∣ a ∣ₕ (g n ∣ a ∣ₕ p) ≡ p - ind zero = + ind-helper : (n : ℕ) (a : S₊ (suc n)) (p : ∣ a ∣ₕ ≡ ∣ a ∣ₕ) → f n ∣ a ∣ₕ (g n ∣ a ∣ₕ p) ≡ p + ind-helper 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) = + ind-helper (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 @@ -751,16 +751,16 @@ baseChange n x = isoToEquiv is 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) + T.elim (λ _ → isOfHLevelΠ (3 + n) λ _ → isOfHLevelPath (3 + n) (isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) _ _) - (ind n) + (ind-helper 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 = + ind-helper : (n : ℕ) (a : S₊ (suc n)) (p : 0ₖ (suc n) ≡ 0ₖ (suc n)) → g n ∣ a ∣ₕ (f n ∣ a ∣ₕ p) ≡ p + ind-helper 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) = + ind-helper (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)) @@ -862,48 +862,48 @@ module lockedCohom (key : Unit') where -- 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 = ST.rec2 § λ 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)) ∣₂ + -H n = ST.rec § λ 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)) ∣₂ + -Hbin n = ST.rec2 § λ 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 (§ _ _)) + rUnitH n = ST.elim (λ _ → 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 (§ _ _)) + lUnitH n = ST.elim (λ _ → 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 (§ _ _)) + rCancelH n = ST.elim (λ _ → 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 (§ _ _)) + lCancelH n = ST.elim (λ _ → 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 (§ _ _)) + assocH n = ST.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 (§ _ _)) + commH n = ST.elim2 (λ _ _ → 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 (§ _ _)) + -cancelRH n = ST.elim2 (λ _ _ → 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 (§ _ _)) + -cancelLH n = ST.elim2 (λ _ _ → 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 (§ _ _)) + -+cancelH n = ST.elim2 (λ _ _ → 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) diff --git a/Cubical/ZCohomology/Groups/CP2.agda b/Cubical/ZCohomology/Groups/CP2.agda index 1bda474fb0..73d8f0d679 100644 --- a/Cubical/ZCohomology/Groups/CP2.agda +++ b/Cubical/ZCohomology/Groups/CP2.agda @@ -1,15 +1,6 @@ {-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.ZCohomology.Groups.CP2 where -open import Cubical.ZCohomology.Base -open import Cubical.ZCohomology.Groups.Connected -open import Cubical.ZCohomology.GroupStructure -open import Cubical.ZCohomology.Properties -open import Cubical.ZCohomology.MayerVietorisUnreduced -open import Cubical.ZCohomology.Groups.Unit -open import Cubical.ZCohomology.Groups.Sn -open import Cubical.ZCohomology.RingStructure.CupProduct - open import Cubical.Foundations.HLevels open import Cubical.Foundations.Function open import Cubical.Foundations.Univalence @@ -19,15 +10,17 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Equiv -open import Cubical.Data.Empty renaming (rec to ⊥-rec) -open import Cubical.Data.Sigma -open import Cubical.Data.Int +open import Cubical.Relation.Nullary + +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Unit open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) open import Cubical.Data.Nat.Order -open import Cubical.Data.Unit +open import Cubical.Data.Int +open import Cubical.Data.Sigma open import Cubical.Algebra.Group -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Int open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Instances.Unit @@ -37,17 +30,22 @@ open import Cubical.HITs.S1 open import Cubical.HITs.Sn open import Cubical.HITs.Susp open import Cubical.HITs.Join -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 ; elim2 to pElim2 ; ∣_∣ to ∣_∣₁ ; map to pMap) +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.Truncation -open import Cubical.Relation.Nullary - open import Cubical.Homotopy.Hopf -open S¹Hopf +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Groups.Connected +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.RingStructure.CupProduct + +open S¹Hopf open IsGroupHom open Iso @@ -70,8 +68,8 @@ leftInv characFunSpaceCP² _ = H⁰CP²≅ℤ : GroupIso (coHomGr 0 CP²) ℤGroup H⁰CP²≅ℤ = H⁰-connected (inr tt) - (PushoutToProp (λ _ → squash) - (sphereElim _ (λ _ → isOfHLevelSuc 1 squash) + (PushoutToProp (λ _ → squash₁) + (sphereElim _ (λ _ → isOfHLevelSuc 1 squash₁) ∣ sym (push (north , base)) ∣₁) λ _ → ∣ refl ∣₁) @@ -103,7 +101,7 @@ H²CP²≅ℤ = compGroupIso (BijectionIso→GroupIso bij) bij : BijectionIso (coHomGr 2 CP²) (×coHomGr 2 (Susp S¹) Unit) BijectionIso.fun bij = M.i 2 BijectionIso.inj bij x p = - pRec (squash₂ _ _) + PT.rec (squash₂ _ _) (uncurry (λ z q → sym q ∙∙ cong (fst (M.d 1)) (isContr→isProp isContrH¹TotalHopf z (0ₕ _)) @@ -125,7 +123,7 @@ H⁴CP²≅ℤ = compGroupIso (invGroupIso (BijectionIso→GroupIso bij)) bij : BijectionIso (coHomGr 3 TotalHopf) (coHomGr 4 CP²) BijectionIso.fun bij = M.d 3 BijectionIso.inj bij x p = - pRec (squash₂ _ _) + PT.rec (squash₂ _ _) (uncurry (λ z q → sym q ∙∙ cong (M.Δ 3 .fst) @@ -161,22 +159,22 @@ H¹-CP²≅0 = (isOfHLevelRetractFromIso 0 (setTruncIso characFunSpaceCP²) (isOfHLevelRetractFromIso 0 lem₂ lem₃)) where - lem₁ : (f : (Susp S¹ → coHomK 1)) → ∥ (λ _ → 0ₖ _) ≡ f ∥ - lem₁ f = pMap (λ p → p) + lem₁ : (f : (Susp S¹ → coHomK 1)) → ∥ (λ _ → 0ₖ _) ≡ f ∥₁ + lem₁ f = PT.map (λ p → p) (Iso.fun PathIdTrunc₀Iso (isOfHLevelRetractFromIso 1 (fst (Hⁿ-Sᵐ≅0 0 1 (λ p → snotz (sym p)))) isPropUnit (0ₕ _) ∣ f ∣₂)) lem₂ : Iso ∥ (Σ[ x ∈ coHomK 1 ] ( Σ[ f ∈ (Susp S¹ → coHomK 1) ] ((y : TotalHopf) → f (fst y) ≡ x))) ∥₂ ∥ (Σ[ f ∈ (Susp S¹ → coHomK 1) ] ((y : TotalHopf) → f (fst y) ≡ 0ₖ 1)) ∥₂ - fun lem₂ = sMap (uncurry λ x → uncurry λ f p → (λ y → (-ₖ x) +ₖ f y) , λ y → cong ((-ₖ x) +ₖ_) (p y) ∙ lCancelₖ _ x) - inv lem₂ = sMap λ p → 0ₖ _ , p + fun lem₂ = ST.map (uncurry λ x → uncurry λ f p → (λ y → (-ₖ x) +ₖ f y) , λ y → cong ((-ₖ x) +ₖ_) (p y) ∙ lCancelₖ _ x) + inv lem₂ = ST.map λ p → 0ₖ _ , p rightInv lem₂ = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + ST.elim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ {(f , p) → cong ∣_∣₂ (ΣPathP ((funExt (λ x → lUnitₖ _ (f x))) , (funExt (λ y → sym (rUnit (λ i → (-ₖ 0ₖ 1) +ₖ p y i))) ◁ λ j y i → lUnitₖ _ (p y i) j)))} leftInv lem₂ = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + ST.elim (λ _ → isOfHLevelPath 2 squash₂ _ _) (uncurry (coHomK-elim _ (λ _ → isPropΠ (λ _ → squash₂ _ _)) (uncurry λ f p → cong ∣_∣₂ (ΣPathP (refl , (ΣPathP ((funExt (λ x → lUnitₖ _ (f x))) , ((funExt (λ y → sym (rUnit (λ i → (-ₖ 0ₖ 1) +ₖ p y i))) @@ -185,8 +183,8 @@ H¹-CP²≅0 = lem₃ : isContr _ fst lem₃ = ∣ (λ _ → 0ₖ 1) , (λ _ → refl) ∣₂ snd lem₃ = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - (uncurry λ f → pRec (isPropΠ (λ _ → squash₂ _ _)) + ST.elim (λ _ → isOfHLevelPath 2 squash₂ _ _) + (uncurry λ f → PT.rec (isPropΠ (λ _ → squash₂ _ _)) (J (λ f _ → (y : (y₁ : TotalHopf) → f (fst y₁) ≡ 0ₖ 1) → ∣ (λ _ → 0ₖ 1) , (λ _ _ → 0ₖ 1) ∣₂ ≡ ∣ f , y ∣₂) (λ y → cong ∣_∣₂ (ΣPathP ((funExt (λ z → sym (y (north , base)))) , toPathP (s y))))) @@ -210,7 +208,7 @@ Hⁿ-CP²≅0-higher n p = contrGroupIsoUnit ((0ₕ _) , (λ x → sym (main x)) propᵣ = isPropΣ (isOfHLevelRetractFromIso 1 - (fst (Hⁿ-Sᵐ≅0 (2 +ℕ n) 1 λ p → ⊥-rec (snotz (cong predℕ p)))) isPropUnit) + (fst (Hⁿ-Sᵐ≅0 (2 +ℕ n) 1 λ p → ⊥.rec (snotz (cong predℕ p)))) isPropUnit) λ _ → isContr→isProp (isContrHⁿ-Unit _) propₗ : isProp (coHom (2 +ℕ n) TotalHopf) @@ -223,7 +221,7 @@ Hⁿ-CP²≅0-higher n p = contrGroupIsoUnit ((0ₕ _) , (λ x → sym (main x)) main : (x : coHom (3 +ℕ n) CP²) → x ≡ 0ₕ _ main x = - pRec (squash₂ _ _) + PT.rec (squash₂ _ _) (uncurry (λ f p → sym p ∙∙ cong (h .fst) (propₗ f (0ₕ _)) ∙∙ pres1 (snd h))) (inIm x) @@ -231,9 +229,9 @@ Hⁿ-CP²≅0-higher n p = contrGroupIsoUnit ((0ₕ _) , (λ x → sym (main x)) Hⁿ-CP²≅0 : (n : ℕ) → ¬ suc n ≡ 2 → ¬ suc n ≡ 4 → GroupIso (coHomGr (suc n) CP²) UnitGroup₀ Hⁿ-CP²≅0 zero p q = H¹-CP²≅0 -Hⁿ-CP²≅0 (suc zero) p q = ⊥-rec (p refl) +Hⁿ-CP²≅0 (suc zero) p q = ⊥.rec (p refl) Hⁿ-CP²≅0 (suc (suc zero)) p q = Hⁿ-CP²≅0-higher 0 λ p → snotz (sym p) -Hⁿ-CP²≅0 (suc (suc (suc zero))) p q = ⊥-rec (q refl) +Hⁿ-CP²≅0 (suc (suc (suc zero))) p q = ⊥.rec (q refl) Hⁿ-CP²≅0 (suc (suc (suc (suc n)))) p q = Hⁿ-CP²≅0-higher (suc (suc n)) λ p → snotz (cong predℕ p) diff --git a/Cubical/ZCohomology/Groups/Connected.agda b/Cubical/ZCohomology/Groups/Connected.agda index 8d88880773..ce474bb0da 100644 --- a/Cubical/ZCohomology/Groups/Connected.agda +++ b/Cubical/ZCohomology/Groups/Connected.agda @@ -1,48 +1,49 @@ {-# 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.Equiv open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels -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.Data.Nat +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +Comm to +ℤ-comm ; +Assoc to +ℤ-assoc) +open import Cubical.Data.Sigma hiding (_×_) + +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Truncation as T 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.Algebra.Group.Instances.Int open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Homotopy.Connected -open import Cubical.Foundations.Equiv + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Unit 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 + H⁰-connected-type : ∀ {ℓ} {A : Type ℓ} (a : A) → isConnected 2 A → Iso (coHom 0 A) ℤ + Iso.fun (H⁰-connected-type a con) = ST.rec 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)) + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f → cong ∣_∣₂ (funExt λ x → T.rec₊ (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) +H⁰-connected : ∀ {ℓ} {A : Type ℓ} (a : A) → ((x : A) → ∥ a ≡ x ∥₁) → GroupIso (coHomGr 0 A) ℤGroup +fun (fst (H⁰-connected a con)) = ST.rec 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) + ST.elim (λ _ → isProp→isSet (isSetSetTrunc _ _)) + (λ f → cong ∣_∣₂ (funExt λ x → PT.rec (isSetℤ _ _) (cong f) (con x))) +snd (H⁰-connected a con) = makeIsGroupHom (ST.elim2 (λ _ _ → isProp→isSet (isSetℤ _ _)) λ x y → refl) diff --git a/Cubical/ZCohomology/Groups/Coproduct.agda b/Cubical/ZCohomology/Groups/Coproduct.agda index e9911bc3d6..886b7f5b9b 100644 --- a/Cubical/ZCohomology/Groups/Coproduct.agda +++ b/Cubical/ZCohomology/Groups/Coproduct.agda @@ -16,7 +16,6 @@ open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.DirProd open import Cubical.Algebra.Ring open import Cubical.Algebra.Ring.DirectProd -open import Cubical.Algebra.Direct-Sum.Base open import Cubical.HITs.SetTruncation as ST diff --git a/Cubical/ZCohomology/Groups/KleinBottle.agda b/Cubical/ZCohomology/Groups/KleinBottle.agda index 22df515057..14dfaefee0 100644 --- a/Cubical/ZCohomology/Groups/KleinBottle.agda +++ b/Cubical/ZCohomology/Groups/KleinBottle.agda @@ -1,51 +1,47 @@ {-# 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 -open import Cubical.Algebra.Group.DirProd -open import Cubical.Algebra.Group.Morphisms -open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.Bool renaming (Bool to BoolGroup) -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) -open import Cubical.Algebra.Group.Instances.Unit - 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.Path open import Cubical.Foundations.Equiv -open import Cubical.Homotopy.Connected -open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Nat hiding (isEven) +open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Bool open import Cubical.Data.Int renaming (+Comm to +-commℤ ; _+_ to _+ℤ_) +open import Cubical.Data.Sigma +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Truncation as T +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn open import Cubical.HITs.KleinBottle -open import Cubical.Data.Empty -open import Cubical.Foundations.Path +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.DirProd +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Instances.Bool +open import Cubical.Algebra.Group.Instances.Int +open import Cubical.Algebra.Group.Instances.Unit + +open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Loopspace +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn + open IsGroupHom open Iso @@ -99,11 +95,11 @@ private ------ H¹(𝕂²) ≅ 0 -------------- H⁰-𝕂² : GroupIso (coHomGr 0 KleinBottle) ℤGroup -fun (fst H⁰-𝕂²) = sRec isSetℤ λ f → f point +fun (fst H⁰-𝕂²) = ST.rec isSetℤ λ f → f point inv (fst H⁰-𝕂²) x = ∣ (λ _ → x) ∣₂ rightInv (fst H⁰-𝕂²) _ = refl leftInv (fst H⁰-𝕂²) = - sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + ST.elim (λ _ → 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 @@ -118,7 +114,7 @@ leftInv (fst H⁰-𝕂²) = λ i j → f (square i j) helper f = isGroupoid→isGroupoid' (isOfHLevelSuc 2 isSetℤ) _ _ _ _ _ _ snd H⁰-𝕂² = - makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 isSetℤ _ _) λ _ _ → refl) + makeIsGroupHom (ST.elim2 (λ _ _ → isOfHLevelPath 2 isSetℤ _ _) λ _ _ → refl) ------ H¹(𝕂¹) ≅ ℤ ------------ {- @@ -132,7 +128,7 @@ H¹(𝕂²) := ∥ 𝕂² → K₁ ∥₂ nilpotent→≡0 : (x : ℤ) → x +ℤ x ≡ 0 → x ≡ 0 nilpotent→≡0 (pos zero) p = refl nilpotent→≡0 (pos (suc n)) p = - ⊥-rec (negsucNotpos _ _ + ⊥.rec (negsucNotpos _ _ (sym (cong (_- 1) (cong sucℤ (sym (helper2 n)) ∙ p)))) where helper2 : (n : ℕ) → pos (suc n) +pos n ≡ pos (suc (n + n)) @@ -140,7 +136,7 @@ nilpotent→≡0 (pos (suc n)) p = 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)) +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)) @@ -154,7 +150,7 @@ nilpotent→≡0 (negsuc n) p = ⊥-rec (negsucNotpos _ _ (helper2 n p)) nilpotent→≡refl : (x : coHomK 1) (p : x ≡ x) → p ∙ p ≡ refl → p ≡ refl nilpotent→≡refl = - trElim (λ _ → isGroupoidΠ2 λ _ _ → isOfHLevelPlus {n = 1} 2 (isOfHLevelTrunc 3 _ _ _ _)) + T.elim (λ _ → 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) @@ -194,7 +190,7 @@ H¹-𝕂²≅ℤ = compGroupIso theGroupIso (Hⁿ-Sⁿ≅ℤ 0) is-hom : IsGroupHom (coHomGr 1 KleinBottle .snd) (fun theIso) (coHomGr 1 S¹ .snd) is-hom = makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (ST.elim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ f g → cong ∣_∣₂ (funExt λ {base → refl ; (loop i) → refl})) theGroupIso : GroupIso (coHomGr 1 KleinBottle) (coHomGr 1 S¹) @@ -215,22 +211,22 @@ H²(𝕂²) := ∥ 𝕂² → 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) + ST.rec isSetSetTrunc + (uncurry (T.elim (λ _ → 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))) + ST.map λ p → (0ₖ 2) , ((fst p) , (refl , (snd p))) rightInv Iso-H²-𝕂²₁ = - sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ p → refl leftInv Iso-H²-𝕂²₁ = - sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) - (uncurry (trElim (λ _ → is2GroupoidΠ λ _ → isOfHLevelPlus {n = 1} 3 (isSetSetTrunc _ _)) + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry (T.elim (λ _ → is2GroupoidΠ λ _ → isOfHLevelPlus {n = 1} 3 (isSetSetTrunc _ _)) (sphereToPropElim _ (λ _ → isPropΠ λ _ → isSetSetTrunc _ _) λ {(p , (q , sq)) - → trRec (isSetSetTrunc _ _) + → T.rec (isSetSetTrunc _ _) (λ qid → cong ∣_∣₂ (ΣPathP (refl , (ΣPathP (refl , (ΣPathP (sym qid , refl))))))) (fun (PathIdTruncIso _) (isContr→isProp (isConnectedPathKn 1 (0ₖ 2) (0ₖ 2)) ∣ q ∣ ∣ refl ∣))}))) @@ -249,7 +245,7 @@ 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) +ΣKₙNilpot→Bool = uncurry (T.elim (λ _ → 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)}) @@ -297,11 +293,11 @@ which is just ∣ (0 , p) ∣₂ * ∣ (0 , q) ∣₂ ≡ ∣ (0 , p ∙ q) ∣ 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) + _*_ = ST.rec (isSetΠ (λ _ → isSetSetTrunc)) λ a → ST.rec 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) + T.elim2 (λ _ _ → isGroupoidΠ2 λ _ _ → isOfHLevelSuc 2 isSetSetTrunc) (wedgeconFun _ _ (λ _ _ → isSetΠ2 λ _ _ → isSetSetTrunc) (λ x p q → ∣ ∣ x ∣ , cong₂ _+ₖ_ p q ∣₂) @@ -312,7 +308,7 @@ private *=∙ 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 zero p = ⊥.rec (true≢false (sym p)) isEvenNegsuc (suc n) p = p ¬isEvenNegSuc : (n : ℕ) → isEven (pos (suc n)) ≡ false → isEven (negsuc n) ≡ false @@ -324,7 +320,7 @@ evenCharac : (x : ℤ) → isEven x ≡ true ∣ (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 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) ∣))) @@ -333,7 +329,7 @@ evenCharac (pos (suc (suc (suc n)))) isisEven = ∙∙ 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 zero) isisEven = ⊥.rec (true≢false (sym isisEven)) evenCharac (negsuc (suc zero)) isisEven = cong ∣_∣₂ ((λ i → 0ₖ 1 , λ i₁ → hfill (doubleComp-faces (λ i₂ → ∣ base ∣) (λ _ → ∣ base ∣) i₁) @@ -348,7 +344,7 @@ 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 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)) ∣₂ @@ -362,7 +358,7 @@ oddCharac (negsuc zero) isOdd = ; (i = i1) → loop j ; (j = i1) → base}) (loop (j ∨ ~ i)) ∣)) -oddCharac (negsuc (suc zero)) isOdd = ⊥-rec (true≢false isOdd) +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)) @@ -374,13 +370,13 @@ 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 +fun testIso = ST.rec 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 + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry (T.elim (λ _ → isGroupoidΠ λ _ → isOfHLevelPlus {n = 1} 2 (isSetSetTrunc _ _)) (toPropElim (λ _ → isPropΠ (λ _ → isSetSetTrunc _ _)) (λ p → path p (isEven (ΩKn+1→Kn 0 p)) refl)))) @@ -426,7 +422,7 @@ isContrHⁿ-𝕂² n = ∣ x , p , q , P ∣₂ ∣ 0ₖ _ , refl , refl , sym (rUnit refl) ∣₂ helper = - trElim (λ _ → isProp→isOfHLevelSuc (4 + n) (isPropΠ4 λ _ _ _ _ → isPropΠ λ _ → isSetSetTrunc _ _)) + T.elim (λ _ → 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) @@ -437,7 +433,7 @@ isContrHⁿ-𝕂² n = → 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 → T.rec (isProp→isOfHLevelSuc n (isSetSetTrunc _ _)) (λ P≡rUnitrefl i → ∣ 0ₖ (3 + n) , refl , refl , P≡rUnitrefl i ∣₂) (fun (PathIdTruncIso _) (isContr→isProp (isConnectedPath _ (isConnectedPathKn (2 + n) _ _) @@ -447,10 +443,10 @@ isContrHⁿ-𝕂² n = 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 _ _) + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ {(x , p , q , P) - → trRec (isProp→isOfHLevelSuc (suc n) (isSetSetTrunc _ _)) - (λ pId → trRec (isProp→isOfHLevelSuc (suc n) (isSetSetTrunc _ _)) + → T.rec (isProp→isOfHLevelSuc (suc n) (isSetSetTrunc _ _)) + (λ pId → T.rec (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 ∣))) diff --git a/Cubical/ZCohomology/Groups/Prelims.agda b/Cubical/ZCohomology/Groups/Prelims.agda index 8db18ee7f2..76a7ee7ae4 100644 --- a/Cubical/ZCohomology/Groups/Prelims.agda +++ b/Cubical/ZCohomology/Groups/Prelims.agda @@ -1,10 +1,6 @@ {-# 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 @@ -13,16 +9,21 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.Nat +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +Comm to +ℤ-comm ; +Assoc to +ℤ-assoc) +open import Cubical.Data.Sigma + +open import Cubical.HITs.Truncation as T +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.S1 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) + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure infixr 33 _⋄_ @@ -87,10 +88,10 @@ coHomPointedElimSⁿ n m {B = B} isprop ind = 0₄ = 0ₖ 4 S¹map : hLevelTrunc 3 S¹ → S¹ -S¹map = trRec isGroupoidS¹ (idfun _) +S¹map = T.rec 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) _ _) +S¹map-id = T.elim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ a → refl diff --git a/Cubical/ZCohomology/Groups/RP2.agda b/Cubical/ZCohomology/Groups/RP2.agda index 329ffc82d4..fa4f3e4682 100644 --- a/Cubical/ZCohomology/Groups/RP2.agda +++ b/Cubical/ZCohomology/Groups/RP2.agda @@ -11,25 +11,24 @@ open import Cubical.Foundations.Transport open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv -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.HITs.S1 -open import Cubical.HITs.Sn -open import Cubical.HITs.RPn.Base +open import Cubical.Data.Bool +open import Cubical.Data.Int +open import Cubical.Data.Sigma open import Cubical.Algebra.Group open import Cubical.Algebra.Group.DirProd open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.Bool renaming (Bool to BoolGroup) -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Bool +open import Cubical.Algebra.Group.Instances.Int open import Cubical.Algebra.Group.Instances.Unit -open import Cubical.Data.Sigma -open import Cubical.Data.Empty renaming (rec to ⊥-rec) -open import Cubical.Data.Bool -open import Cubical.Data.Int +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Truncation as T +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.RPn.Base open import Cubical.Homotopy.Connected @@ -44,6 +43,8 @@ private ℓ : 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 @@ -63,18 +64,18 @@ private H⁰-RP²≅ℤ : GroupIso (coHomGr 0 RP²) ℤGroup H⁰-RP²≅ℤ = H⁰-connected point connectedRP¹ where - connectedRP¹ : (x : RP²) → ∥ point ≡ x ∥ - connectedRP¹ point = ∣ refl ∣ + connectedRP¹ : (x : RP²) → ∥ point ≡ x ∥₁ + connectedRP¹ point = ∣ refl ∣₁ connectedRP¹ (line i) = - isOfHLevel→isOfHLevelDep 1 {B = λ x → ∥ point ≡ x ∥} - (λ _ → isPropPropTrunc) ∣ refl ∣ ∣ refl ∣ 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)) + 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 _ _ _ _) @@ -82,9 +83,9 @@ H⁰-RP²≅ℤ = H⁰-connected point connectedRP¹ 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 _ _) + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) (uncurry - (trElim (λ _ → isGroupoidΠ λ _ → isOfHLevelPlus {n = 1} 2 (isSetSetTrunc _ _)) + (T.elim (λ _ → isGroupoidΠ λ _ → isOfHLevelPlus {n = 1} 2 (isSetSetTrunc _ _)) (toPropElim (λ _ → isPropΠ (λ _ → isSetSetTrunc _ _)) λ {(p , nilp) → cong ∣_∣₂ (ΣPathP (refl , Σ≡Prop (λ _ → isOfHLevelTrunc 3 _ _ _ _) @@ -107,17 +108,17 @@ H¹-RP²≅0 = 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 + ST.rec isSetSetTrunc (uncurry - (trElim (λ _ → is2GroupoidΠ λ _ → isOfHLevelPlus {n = 2} 2 isSetSetTrunc) + (T.elim (λ _ → 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 _ _) +Iso.inv Iso-H²-RP²₁ = ST.map λ p → (0ₖ 2) , p +Iso.rightInv Iso-H²-RP²₁ = ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ _ → refl Iso.leftInv Iso-H²-RP²₁ = - sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) - (uncurry (trElim (λ _ → is2GroupoidΠ λ _ → isOfHLevelPlus {n = 1} 3 (isSetSetTrunc _ _)) + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry (T.elim (λ _ → is2GroupoidΠ λ _ → isOfHLevelPlus {n = 1} 3 (isSetSetTrunc _ _)) (sphereToPropElim _ (λ _ → isPropΠ (λ _ → isSetSetTrunc _ _)) λ p → refl))) diff --git a/Cubical/ZCohomology/Groups/Sn.agda b/Cubical/ZCohomology/Groups/Sn.agda index 6ec17f1cb9..15ec2a2905 100644 --- a/Cubical/ZCohomology/Groups/Sn.agda +++ b/Cubical/ZCohomology/Groups/Sn.agda @@ -1,13 +1,6 @@ {-# 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 @@ -17,32 +10,38 @@ 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 hiding (rec ; elim ; ind) -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.Empty as ⊥ 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.Data.Int renaming (_+_ to _+ℤ_; +Comm to +ℤ-comm ; +Assoc to +ℤ-assoc) +open import Cubical.Data.Sigma +open import Cubical.Data.Sum open import Cubical.Algebra.Group open import Cubical.Algebra.Group.DirProd open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Instances.Unit -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Int + +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 as ST +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Truncation as T + +open import Cubical.Homotopy.Connected + +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 infixr 31 _□_ _□_ : _ @@ -72,7 +71,7 @@ suspensionAx-Sn n m = (S₊ (suc m) → x ≡ y)) ∥₂ (coHom (suc n) (S₊ (suc m))) Iso.fun helperIso = - sRec isSetSetTrunc + ST.rec isSetSetTrunc (uncurry (coHomK-elim _ (λ _ → isOfHLevelΠ (2 + n) @@ -83,12 +82,12 @@ suspensionAx-Sn n m = λ _ → 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)) + ST.map λ 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 _ _) + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) (uncurry (coHomK-elim _ (λ _ → isProp→isOfHLevelSuc (suc n) (isPropΠ λ _ → isSetSetTrunc _ _)) @@ -130,10 +129,10 @@ 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)) + makeIsGroupHom (ST.elim2 (λ _ _ → isSet→isGroupoid isSetSetTrunc _ _) (λ _ _ → refl)) coHomPushout≅coHomSn (suc n) m = setTruncIso (domIso (invIso PushoutSuspIsoSusp)) , - makeIsGroupHom (sElim2 (λ _ _ → isSet→isGroupoid isSetSetTrunc _ _) (λ _ _ → refl)) + makeIsGroupHom (ST.elim2 (λ _ _ → isSet→isGroupoid isSetSetTrunc _ _) (λ _ _ → refl)) -------------------------- H⁰(S⁰) ----------------------------- S0→ℤ : (a : ℤ × ℤ) → S₊ 0 → ℤ @@ -141,15 +140,15 @@ 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) +fun (fst H⁰-S⁰≅ℤ×ℤ) = ST.rec (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 _ _) + ST.elim (λ _ → isSet→isGroupoid isSetSetTrunc _ _) (λ f → cong ∣_∣₂ (funExt (λ {true → refl ; false → refl}))) snd H⁰-S⁰≅ℤ×ℤ = makeIsGroupHom - (sElim2 (λ _ _ → isSet→isGroupoid (isSet× isSetℤ isSetℤ) _ _) λ a b → refl) + (ST.elim2 (λ _ _ → isSet→isGroupoid (isSet× isSetℤ isSetℤ) _ _) λ a b → refl) ------------------------- Hⁿ(S⁰) ≅ 0 for n ≥ 1 ------------------------------- @@ -171,18 +170,18 @@ private 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) ∣₂ } + snd (isContrHelper zero) = ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ y → T.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) ∣₂ } + , ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ y → T.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) ∣₂ } + , ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ y → T.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) @@ -201,11 +200,11 @@ Hⁿ-S¹≅0 n = contrGroupIsoUnit helper2 : (x : ∥ Σ (hLevelTrunc (4 + n) (S₊ (2 + n))) (λ x → ∥ x ≡ x ∥₂) ∥₂) → ∣ ∣ north ∣ , ∣ refl ∣₂ ∣₂ ≡ x helper2 = - sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) (uncurry - (trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isProp→isOfHLevelSuc (3 + n) (isSetSetTrunc _ _)) + (T.elim (λ _ → isOfHLevelΠ (4 + n) λ _ → isProp→isOfHLevelSuc (3 + n) (isSetSetTrunc _ _)) (suspToPropElim (ptSn (suc n)) (λ _ → isPropΠ λ _ → isSetSetTrunc _ _) - (sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ p → cong ∣_∣₂ (ΣPathP (refl , isContr→isProp helper3 _ _)))))) where @@ -236,13 +235,13 @@ H¹-Sⁿ≅0 (suc n) = contrGroupIsoUnit isContrH¹S³⁺ⁿ isContrH¹S³⁺ⁿ-ish : (f : (S₊ (3 + n) → hLevelTrunc (4 + n) (coHomK 1))) → ∣ (λ _ → ∣ ∣ base ∣ ∣) ∣₂ ≡ ∣ f ∣₂ - isContrH¹S³⁺ⁿ-ish f = ind (f north) refl + isContrH¹S³⁺ⁿ-ish f = ind-helper (f north) refl where - ind : (x : hLevelTrunc (4 + n) (coHomK 1)) + ind-helper : (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 _ _)) + ind-helper = T.elim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelPlus' {n = (3 + n)} 1 (isSetSetTrunc _ _)) + (T.elim (λ _ → 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)) ⟩ @@ -250,7 +249,7 @@ H¹-Sⁿ≅0 (suc n) = contrGroupIsoUnit isContrH¹S³⁺ⁿ isOfHLevelRetractFromIso 0 anIso (∣ (λ _ → ∣ ∣ base ∣ ∣) ∣₂ - , sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) isContrH¹S³⁺ⁿ-ish) + , ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) isContrH¹S³⁺ⁿ-ish) --------- H¹(S¹) ≅ ℤ ------- {- @@ -268,13 +267,13 @@ H¹-S¹≅ℤ = theIso F⁻ = Iso.inv S¹→S¹≡S¹×ℤ theIso : GroupIso (coHomGr 1 (S₊ 1)) ℤGroup - fun (fst theIso) = sRec isSetℤ (λ f → snd (F f)) + fun (fst theIso) = ST.rec 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))) ∣₂) + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ f → cong ((ST.rec isSetSetTrunc ∣_∣₂) + ∘ ST.rec isSetSetTrunc λ x → ∣ F⁻ (x , (snd (F f))) ∣₂) (Iso.inv PathIdTrunc₀Iso (isConnectedS¹ (fst (F f)))) ∙ cong ∣_∣₂ (Iso.leftInv S¹→S¹≡S¹×ℤ f) snd theIso = @@ -298,7 +297,7 @@ 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 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 diff --git a/Cubical/ZCohomology/Groups/SphereProduct.agda b/Cubical/ZCohomology/Groups/SphereProduct.agda index 1acc47055a..7e47b894f5 100644 --- a/Cubical/ZCohomology/Groups/SphereProduct.agda +++ b/Cubical/ZCohomology/Groups/SphereProduct.agda @@ -13,6 +13,8 @@ gₗ, gᵣ are the two generators of H²(S²×S²) and e is the generator of H⁴(S²×S²) -- this of interest since it is used in π₄S³≅ℤ/2. -} + + {-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.ZCohomology.Groups.SphereProduct where @@ -23,31 +25,23 @@ open import Cubical.Foundations.Pointed open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws -open import Cubical.Data.Sigma -open import Cubical.Data.Nat -open import Cubical.Data.Unit - -open import Cubical.HITs.S1 -open import Cubical.HITs.Sn -open import Cubical.HITs.Susp -open import Cubical.HITs.Truncation -open import Cubical.HITs.SetTruncation - renaming (elim to sElim ; elim2 to sElim2 ; map to sMap) -open import Cubical.HITs.PropositionalTruncation - renaming (map to pMap ; rec to pRec) +open import Cubical.Relation.Nullary -open import Cubical.ZCohomology.Base -open import Cubical.ZCohomology.GroupStructure -open import Cubical.ZCohomology.Properties -open import Cubical.ZCohomology.Groups.Sn -open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.Data.Unit +open import Cubical.Data.Nat +open import Cubical.Data.Sigma open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Instances.Int -open import Cubical.Relation.Nullary +open import Cubical.HITs.Truncation as T +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp open import Cubical.Homotopy.Loopspace @@ -65,6 +59,8 @@ private ¬lem n (suc m) p = ¬lem n m (sym (+-suc n m) ∙ cong predℕ p) + + {- Step 1: Hⁿ⁺ᵐ(Sⁿ×Sᵐ)≅H¹⁺ⁿ⁺ᵐ(S¹⁺ⁿ×Sᵐ) -} -- The functions back and forth (suspension on the left component) ↓Sⁿ×Sᵐ→Kₙ₊ₘ : (n m : ℕ) @@ -86,7 +82,7 @@ private ∥HⁿSᵐPath∥ : (n m : ℕ) → (f : S₊ (suc m) → coHomK (suc n)) → ¬ (n ≡ m) - → ∥ f ≡ (λ _ → 0ₖ (suc n)) ∥ + → ∥ f ≡ (λ _ → 0ₖ (suc n)) ∥₁ ∥HⁿSᵐPath∥ n m f p = fun PathIdTrunc₀Iso (isContr→isProp @@ -100,15 +96,15 @@ private (coHomGr (suc (suc ((suc n) + m))) (S₊ (suc (suc n)) × S₊ (suc m))) fun (fst (×leftSuspensionIso n m)) = - sMap (uncurry ∘ ↑Sⁿ×Sᵐ→Kₙ₊ₘ n m ∘ curry) + ST.map (uncurry ∘ ↑Sⁿ×Sᵐ→Kₙ₊ₘ n m ∘ curry) inv (fst (×leftSuspensionIso n m)) = - sMap ((uncurry ∘ ↓Sⁿ×Sᵐ→Kₙ₊ₘ n m ∘ curry)) + ST.map ((uncurry ∘ ↓Sⁿ×Sᵐ→Kₙ₊ₘ n m ∘ curry)) rightInv (fst (×leftSuspensionIso n m)) = - sElim (λ _ → isSetPathImplicit) + ST.elim (λ _ → isSetPathImplicit) λ f → inv PathIdTrunc₀Iso - (pRec squash + (PT.rec squash₁ (uncurry (λ g p - → pMap (λ gid → funExt λ {(x , y) + → PT.map (λ gid → funExt λ {(x , y) → (λ i → uncurry (↑Sⁿ×Sᵐ→Kₙ₊ₘ n m (↓Sⁿ×Sᵐ→Kₙ₊ₘ n m (curry (p (~ i))))) (x , y)) ∙∙ main g gid x y @@ -134,7 +130,7 @@ rightInv (fst (×leftSuspensionIso n m)) = → typ (Ω (coHomK-ptd (suc (suc (suc n + m)))))) ] charac-fun g ≡ f rewrte f = - pMap (λ p + PT.map (λ p → (λ x y → sym (funExt⁻ p y) ∙∙ (λ i → f ((merid x ∙ sym (merid (ptSn (suc n)))) i , y)) ∙∙ funExt⁻ p y) @@ -157,14 +153,14 @@ rightInv (fst (×leftSuspensionIso n m)) = ∥Path∥ : (g : S₊ (suc n) → S₊ (suc m) → typ (Ω (coHomK-ptd (suc (suc (suc n + m)))))) - → ∥ (g (ptSn _)) ≡ (λ _ → refl) ∥ + → ∥ (g (ptSn _)) ≡ (λ _ → refl) ∥₁ ∥Path∥ g = fun PathIdTrunc₀Iso (isContr→isProp (isOfHLevelRetractFromIso 0 ((invIso (fst (coHom≅coHomΩ _ (S₊ (suc m)))))) - (0ₕ _ , sElim (λ _ → isProp→isSet (squash₂ _ _)) - λ f → pRec (squash₂ _ _) (sym ∘ cong ∣_∣₂) + (0ₕ _ , ST.elim (λ _ → isProp→isSet (squash₂ _ _)) + λ f → PT.rec (squash₂ _ _) (sym ∘ cong ∣_∣₂) (∥HⁿSᵐPath∥ _ _ f (¬lem n m)))) ∣ g (ptSn (suc n)) ∣₂ ∣ (λ _ → refl) ∣₂) @@ -190,8 +186,8 @@ rightInv (fst (×leftSuspensionIso n m)) = ∙∙ (cong (g a y ∙_) (cong sym (funExt⁻ gid y)) ∙ sym (rUnit (g a y))) leftInv (fst (×leftSuspensionIso n m)) = - sElim (λ _ → isSetPathImplicit) - λ f → pRec (squash₂ _ _) + ST.elim (λ _ → isSetPathImplicit) + λ f → PT.rec (squash₂ _ _) (λ id → cong ∣_∣₂ (funExt (λ {(x , y) → (λ i → ΩKn+1→Kn _ (sym (rCancel≡refl _ i) @@ -210,7 +206,7 @@ leftInv (fst (×leftSuspensionIso n m)) = (∥HⁿSᵐPath∥ (suc n + m) m (λ x → f (ptSn _ , x)) (¬lem n m)) snd (×leftSuspensionIso n m) = - makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + makeIsGroupHom (ST.elim2 (λ _ _ → isSetPathImplicit) λ f g → cong ∣_∣₂ (funExt λ { (north , y) → refl ; (south , y) → refl @@ -241,12 +237,12 @@ Hⁿ-Sⁿ→Hⁿ-S¹×Sⁿ m f x = Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ : (m : ℕ) → GroupIso (coHomGr (suc m) (S₊ (suc m))) (coHomGr (suc (suc m)) (S₊ (suc zero) × S₊ (suc m))) -fun (fst (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m)) = sMap (uncurry ∘ Hⁿ-S¹×Sⁿ→Hⁿ-Sⁿ m) -inv (fst (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m)) = sMap (Hⁿ-Sⁿ→Hⁿ-S¹×Sⁿ m ∘ curry) +fun (fst (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m)) = ST.map (uncurry ∘ Hⁿ-S¹×Sⁿ→Hⁿ-Sⁿ m) +inv (fst (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m)) = ST.map (Hⁿ-Sⁿ→Hⁿ-S¹×Sⁿ m ∘ curry) rightInv (fst (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m)) = - sElim (λ _ → isSetPathImplicit) + ST.elim (λ _ → isSetPathImplicit) λ f → inv PathIdTrunc₀Iso - (pMap (uncurry (λ g p + (PT.map (uncurry (λ g p → funExt λ {(x , y) → (λ i → uncurry (Hⁿ-S¹×Sⁿ→Hⁿ-Sⁿ m @@ -286,7 +282,7 @@ rightInv (fst (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m)) = rewrte : (f : S₊ 1 × S₊ (suc m) → coHomK (suc (suc m))) → ∃[ g ∈ (S₊ (suc m) → coHomK (suc m)) ] f ≡ characFun g rewrte f = - pMap (λ p → + PT.map (λ p → (λ x → ΩKn+1→Kn _ (sym (funExt⁻ p x) ∙∙ (λ i → f (loop i , x)) ∙∙ funExt⁻ p x)) , funExt λ { (base , y) → funExt⁻ p y @@ -305,7 +301,7 @@ rightInv (fst (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m)) = (funExt⁻ p x) j i)}) (∥HⁿSᵐPath∥ (suc m) m (λ x → f (base , x)) (lem m)) leftInv (fst (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m)) = - sElim (λ _ → isSetPathImplicit) + ST.elim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ x → cong (ΩKn+1→Kn (suc m)) @@ -315,7 +311,7 @@ leftInv (fst (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m)) = ∙ leftInv (Iso-Kn-ΩKn+1 _) (f x)) snd (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m) = makeIsGroupHom - (sElim2 + (ST.elim2 (λ _ _ → isSetPathImplicit) λ f g → cong ∣_∣₂ (funExt λ { (base , y) → refl @@ -326,7 +322,7 @@ snd (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m) = Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ : (n m : ℕ) → GroupIso (coHomGr ((suc n) +' (suc m)) (S₊ (suc n) × S₊ (suc m))) - ℤ + ℤGroup Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ zero m = compGroupIso (invGroupIso (Hⁿ-Sⁿ≅Hⁿ-S¹×Sⁿ m)) (Hⁿ-Sⁿ≅ℤ m) Hⁿ⁺ᵐ-Sⁿ×Sᵐ≅ℤ (suc n) m = diff --git a/Cubical/ZCohomology/Groups/Torus.agda b/Cubical/ZCohomology/Groups/Torus.agda index 2dca3a0b66..4bd077b8dc 100644 --- a/Cubical/ZCohomology/Groups/Torus.agda +++ b/Cubical/ZCohomology/Groups/Torus.agda @@ -1,16 +1,6 @@ {-# 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 @@ -20,39 +10,50 @@ 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.Relation.Nullary + open import Cubical.Data.Unit +open import Cubical.Data.Nat +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +Comm to +ℤ-comm ; +Assoc to +ℤ-assoc) +open import Cubical.Data.Sigma open import Cubical.Algebra.Group open import Cubical.Algebra.Group.DirProd open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.Bool renaming (Bool to BoolGroup) -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Bool +open import Cubical.Algebra.Group.Instances.Int +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Truncation as T +open import Cubical.HITs.Nullification 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.HITs.Wedge + open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Loopspace +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.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. @@ -129,8 +130,8 @@ coHomPointedElimT²' : ∀ {ℓ} (n : ℕ) {B : coHom (2 + n) (S¹ × S¹) → T → (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 _)) + λ p q P → T.rec (isProp→isOfHLevelSuc n (prop _)) + (λ p-refl → T.rec (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) @@ -149,8 +150,8 @@ private H⁰-T²≅ℤ : GroupIso (coHomGr 0 (S₊ 1 × S₊ 1)) ℤGroup H⁰-T²≅ℤ = H⁰-connected (base , base) - λ (a , b) → pRec isPropPropTrunc - (λ id1 → pRec isPropPropTrunc + λ (a , b) → PT.rec isPropPropTrunc + (λ id1 → PT.rec isPropPropTrunc (λ id2 → ∣ ΣPathP (id1 , id2) ∣₁) (Sn-connected 0 b) ) (Sn-connected 0 a) diff --git a/Cubical/ZCohomology/Groups/Unit.agda b/Cubical/ZCohomology/Groups/Unit.agda index 2fe525edf8..35fa2cbfca 100644 --- a/Cubical/ZCohomology/Groups/Unit.agda +++ b/Cubical/ZCohomology/Groups/Unit.agda @@ -5,42 +5,39 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism -open import Cubical.ZCohomology.Base -open import Cubical.ZCohomology.Properties -open import Cubical.ZCohomology.GroupStructure - -open import Cubical.HITs.Sn -open import Cubical.HITs.Susp -open import Cubical.HITs.Truncation -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.Data.Int hiding (ℤ ; _+_ ; +Comm) -open import Cubical.Data.Nat open import Cubical.Data.Unit +open import Cubical.Data.Nat +open import Cubical.Data.Int hiding (ℤ ; _+_ ; +Comm) open import Cubical.Data.Sigma -open import Cubical.Homotopy.Connected - open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Instances.Int open import Cubical.Algebra.Group.Instances.Unit +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Truncation + +open import Cubical.Homotopy.Connected + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.GroupStructure + -- H⁰(Unit) open IsGroupHom open Iso -H⁰-Unit≅ℤ : GroupIso (coHomGr 0 Unit) ℤ -fun (fst H⁰-Unit≅ℤ) = sRec isSetℤ (λ f → f tt) +H⁰-Unit≅ℤ : GroupIso (coHomGr 0 Unit) ℤGroup +fun (fst H⁰-Unit≅ℤ) = ST.rec 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) +leftInv (fst H⁰-Unit≅ℤ) = ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ a → refl +snd H⁰-Unit≅ℤ = makeIsGroupHom (ST.elim2 (λ _ _ → isOfHLevelPath 2 isSetℤ _ _) λ a b → refl) {- Hⁿ(Unit) for n ≥ 1 -} isContrHⁿ-Unit : (n : ℕ) → isContr (coHom (suc n) Unit) @@ -91,7 +88,7 @@ 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 _ _) + ST.elim (λ _ → 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 index ba76d70a0b..12a20b0914 100644 --- a/Cubical/ZCohomology/Groups/Wedge.agda +++ b/Cubical/ZCohomology/Groups/Wedge.agda @@ -9,23 +9,24 @@ open import Cubical.Foundations.GroupoidLaws renaming (assoc to assoc∙) open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv -open import Cubical.HITs.Wedge -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.HITs.Susp -open import Cubical.HITs.S1 -open import Cubical.HITs.Sn -open import Cubical.HITs.Pushout open import Cubical.Data.Nat -open import Cubical.Data.Sigma open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Sigma open import Cubical.Algebra.Group open import Cubical.Algebra.Group.DirProd open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Truncation as T +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Wedge +open import Cubical.HITs.Pushout + open import Cubical.Homotopy.Connected open import Cubical.ZCohomology.Base @@ -34,11 +35,11 @@ open import Cubical.ZCohomology.Properties open import Cubical.ZCohomology.Groups.Unit open import Cubical.ZCohomology.Groups.Sn +open IsGroupHom +open Iso -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) @@ -114,10 +115,10 @@ module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') where Hⁿ-⋁ : (n : ℕ) → GroupIso (coHomGr (suc n) (A ⋁ B)) (×coHomGr (suc n) (typ A) (typ B)) fun (fst (Hⁿ-⋁ zero)) = - sElim (λ _ → isSet× isSetSetTrunc isSetSetTrunc) + ST.elim (λ _ → isSet× isSetSetTrunc isSetSetTrunc) λ f → ∣ (λ x → f (inl x)) ∣₂ , ∣ (λ x → f (inr x)) ∣₂ inv (fst (Hⁿ-⋁ zero)) = - uncurry (sElim2 (λ _ _ → isSetSetTrunc) + uncurry (ST.elim2 (λ _ _ → isSetSetTrunc) λ f g → ∣ wedgeFun⁻ 0 f g ∣₂) rightInv (fst (Hⁿ-⋁ zero)) = uncurry @@ -126,16 +127,16 @@ module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') where λ 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 _ _) + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (λ f → PT.rec (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)) ∥ + → ∥ f ≡ wedgeFun⁻ 0 (λ x → f (inl x)) (λ x → f (inr x)) ∥₁ helper f = - trElim (λ _ → isProp→isOfHLevelSuc 2 (isPropΠ λ _ → isPropPropTrunc)) + T.elim (λ _ → 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) @@ -159,13 +160,13 @@ module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') where λ i _ → (refl ∙ (λ _ → 0ₖ 1)) i snd (Hⁿ-⋁ zero) = makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + (ST.elim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) λ _ _ → refl) fun (fst (Hⁿ-⋁ (suc n))) = - sElim (λ _ → isSet× isSetSetTrunc isSetSetTrunc) + ST.elim (λ _ → isSet× isSetSetTrunc isSetSetTrunc) λ f → ∣ (λ x → f (inl x)) ∣₂ , ∣ (λ x → f (inr x)) ∣₂ inv (fst (Hⁿ-⋁ (suc n))) = - uncurry (sElim2 (λ _ _ → isSetSetTrunc) + uncurry (ST.elim2 (λ _ _ → isSetSetTrunc) λ f g → ∣ wedgeFun⁻ (suc n) f g ∣₂) rightInv (fst (Hⁿ-⋁ (suc n))) = uncurry @@ -174,16 +175,16 @@ module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') where λ 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 _ _) + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (λ f → PT.rec (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)) ∥ + → ∥ f ≡ wedgeFun⁻ (suc n) (λ x → f (inl x)) (λ x → f (inr x)) ∥₁ helper f = - trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isPropΠ λ _ → isPropPropTrunc)) + T.elim (λ _ → 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) @@ -207,28 +208,28 @@ module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') where λ i j → ((λ _ → ∣ north ∣) ∙ refl) i snd (Hⁿ-⋁ (suc n)) = makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + (ST.elim2 (λ _ _ → 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) + ST.rec (isSet× isSetSetTrunc isSetSetTrunc) λ {(f , p) → ∣ (f ∘ inl) , p ∣₂ , ∣ (f ∘ inr) , cong f (sym (push tt)) ∙ p ∣₂} inv (fst H⁰Red-⋁) = - uncurry (sRec2 isSetSetTrunc + uncurry (ST.rec2 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) _ _) + (ST.elim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) λ {(_ , _) (_ , _) → ΣPathP (cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl) , cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) refl))}) leftInv (fst H⁰Red-⋁) = - sElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + ST.elim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) λ {(f , p) → cong ∣_∣₂ (Σ≡Prop (λ _ → isSetℤ _ _) (funExt λ {(inl a) → refl ; (inr b) → refl @@ -239,12 +240,12 @@ module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') where -- Alt. use isOfHLevel→isOfHLevelDep snd H⁰Red-⋁ = makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + (ST.elim2 (λ _ _ → 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 : ((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) + (λ a → PT.rec isPropPropTrunc (λ p → ∣ cong inl p ∣₁) (conA a)) + λ b → PT.rec isPropPropTrunc (λ p → ∣ push tt ∙ cong inr p ∣₁) (conB b) diff --git a/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda index 3ff91ed54d..8c19744493 100644 --- a/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda +++ b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda @@ -6,9 +6,24 @@ 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.Data.Sigma + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.DirProd +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Instances.Bool +open import Cubical.Algebra.Group.Instances.Int +open import Cubical.Algebra.Group.Instances.Unit + +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.ZCohomology.Base open import Cubical.ZCohomology.Properties @@ -19,21 +34,8 @@ 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 -open import Cubical.Algebra.Group.DirProd -open import Cubical.Algebra.Group.Morphisms -open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.Bool renaming (Bool to BoolGroup) -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) -open import Cubical.Algebra.Group.Instances.Unit + S¹⋁S¹ : Type₀ S¹⋁S¹ = S₊∙ 1 ⋁ S₊∙ 1 diff --git a/Cubical/ZCohomology/Gysin.agda b/Cubical/ZCohomology/Gysin.agda index 42b596b7fe..5f848e8580 100644 --- a/Cubical/ZCohomology/Gysin.agda +++ b/Cubical/ZCohomology/Gysin.agda @@ -18,34 +18,30 @@ open import Cubical.Functions.Embedding open import Cubical.Relation.Nullary -open import Cubical.Data.Sum -open import Cubical.Data.Fin -open import Cubical.Data.Empty renaming (rec to ⊥-rec) -open import Cubical.Data.Sigma -open import Cubical.Data.Int hiding (_+'_) -open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) -open import Cubical.Data.Nat.Order open import Cubical.Data.Unit open import Cubical.Data.Bool +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Int hiding (_+'_) +open import Cubical.Data.Sigma +open import Cubical.Data.Sum +open import Cubical.Data.Fin open import Cubical.Algebra.AbGroup open import Cubical.Algebra.Group open import Cubical.Algebra.Group.ZAction open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Int +open import Cubical.HITs.Truncation as T +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.Pushout.Flattening open import Cubical.HITs.Pushout open import Cubical.HITs.Sn open import Cubical.HITs.Susp open import Cubical.HITs.S1 renaming (_·_ to _*_) -open import Cubical.HITs.Truncation - renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) -open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap ; elim3 to sElim3) -open import Cubical.HITs.PropositionalTruncation - renaming (rec to pRec ; elim to pElim) open import Cubical.HITs.Join open import Cubical.Homotopy.Connected @@ -65,6 +61,8 @@ open import Cubical.ZCohomology.RingStructure.CupProduct open import Cubical.ZCohomology.RingStructure.RingLaws open import Cubical.ZCohomology.RingStructure.GradedCommutativity + + -- There seems to be some problems with the termination checker. -- Spelling out integer induction with 3 base cases like this -- solves the issue. @@ -147,15 +145,15 @@ K : (n : ℕ) → GroupIso (coHomRedGrDir n (S₊∙ n)) (πS n) fst (K n) = setTruncIdempotentIso (isSetπS n) snd (K zero) = makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 0) _ _) + (ST.elim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 0) _ _) λ f g → →∙Homogeneous≡ (isHomogeneousKn 0) refl) snd (K (suc zero)) = makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 1) _ _) + (ST.elim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 1) _ _) λ f g → →∙Homogeneous≡ (isHomogeneousKn 1) refl) snd (K (suc (suc n))) = makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS _) _ _) + (ST.elim2 (λ _ _ → isOfHLevelPath 2 (isSetπS _) _ _) λ f g → →∙Homogeneous≡ (isHomogeneousKn _) refl) -- πS has the following generator @@ -522,7 +520,7 @@ module g-base where -- We now generealise the equivalence g to also apply to arbitrary fibrations (Q : B → Type) -- satisfying (Q * ≃∙ Sⁿ) module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) - (conB : (x y : typ B) → ∥ x ≡ y ∥) + (conB : (x y : typ B) → ∥ x ≡ y ∥₁) (n : ℕ) (Q-is : Iso (typ (Q (pt B))) (S₊ n)) (Q-is-ptd : Iso.fun Q-is (pt (Q (pt B))) ≡ snd (S₊∙ n)) (c : (b : typ B) → (Q b →∙ coHomK-ptd n)) @@ -560,7 +558,7 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) g-equiv : (b : typ B) (i : ℕ) → isEquiv (g b i) g-equiv b i = - pRec (isPropIsEquiv _) + PT.rec (isPropIsEquiv _) (J (λ b _ → isEquiv (g b i)) (g-base i)) (conB (pt B) b) @@ -572,7 +570,7 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) is-setQ→K : (b : typ B) → isSet (Q b →∙ coHomK-ptd n) is-setQ→K b = - sRec (isProp→isOfHLevelSuc 1 isPropIsSet) + ST.rec (isProp→isOfHLevelSuc 1 isPropIsSet) (J (λ b _ → isSet (Q b →∙ coHomK-ptd n)) (subst isSet (cong (_→∙ coHomK-ptd n) (ua∙ (isoToEquiv (invIso Q-is)) @@ -584,8 +582,8 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) isConnB : isConnected 3 (typ B) fst isConnB = ∣ pt B ∣ snd isConnB = - trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) - λ a → sRec (isOfHLevelTrunc 3 _ _) (cong ∣_∣ₕ) (conB (pt B) a) + T.elim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → ST.rec (isOfHLevelTrunc 3 _ _) (cong ∣_∣ₕ) (conB (pt B) a) isPropPath : isProp (∥ pt B ≡ pt B ∥₂) isPropPath = @@ -597,13 +595,13 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) c* : Σ[ c ∈ ((b : typ B) → (Q b →∙ coHomK-ptd n)) ] (c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) fst c* b = - sRec (is-setQ→K b) + ST.rec (is-setQ→K b) (J (λ b _ → Q b →∙ coHomK-ptd n) ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) , cong (genFunSpace n .fst) Q-is-ptd ∙ genFunSpace n .snd)) (conB (pt B) b) snd c* = - funExt λ x → (λ i → sRec (is-setQ→K (pt B)) + funExt λ x → (λ i → ST.rec (is-setQ→K (pt B)) (J (λ b _ → Q b →∙ coHomK-ptd n) ((λ x₁ → genFunSpace n .fst (Iso.fun Q-is x₁)) , (λ i → genFunSpace n .fst (Q-is-ptd i)) @@ -621,21 +619,21 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) -- This form of c* will make things somewhat easier to work with later on. c≡ : (b : fst B) (p : ∥ pt B ≡ b ∥₂) → (c* .fst b) - ≡ sRec (is-setQ→K b) + ≡ ST.rec (is-setQ→K b) (λ pp → (λ qb → genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) , cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) ∙ ((λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) p c≡ b = - sElim (λ _ → isOfHLevelPath 2 (is-setQ→K b) _ _) + ST.elim (λ _ → isOfHLevelPath 2 (is-setQ→K b) _ _) (J (λ b a → c* .fst b ≡ - sRec (is-setQ→K b) (λ pp → + ST.rec (is-setQ→K b) (λ pp → (λ qb → genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) , cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) ∙ (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd) ∣ a ∣₂) - ((λ i → sRec (is-setQ→K (pt B)) + ((λ i → ST.rec (is-setQ→K (pt B)) (J (λ b₁ _ → Q b₁ →∙ coHomK-ptd n) ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) , (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) @@ -786,7 +784,7 @@ module preThom {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where -- the nᵗʰ cohomology of B and the (n+i)ᵗʰ reduced cohomology of Ẽ, -- as defined above module Thom {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) - (conB : (x y : typ B) → ∥ x ≡ y ∥) + (conB : (x y : typ B) → ∥ x ≡ y ∥₁) (n : ℕ) (Q-is : Iso (typ (preThom.Q B P (pt B))) (S₊ n)) (Q-is-ptd : Iso.fun Q-is (pt (preThom.Q B P (pt B))) ≡ snd (S₊∙ n)) (c : (b : typ B) → (preThom.Q B P b →∙ coHomK-ptd n)) @@ -807,7 +805,7 @@ module Thom {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) (preThom.ι B P (i +' n)))) snd (ϕ i) = makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + (ST.elim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ F G → cong ∣_∣₂ (cong (Iso.fun (preThom.ι B P (i +' n))) (funExt (λ a → @@ -823,8 +821,8 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) (Q-is-ptd : Iso.fun Q-is (pt (preThom.Q B P (pt B))) ≡ snd (S₊∙ n)) where - 0-connB : (x y : typ B) → ∥ x ≡ y ∥ - 0-connB x y = sRec (isProp→isSet squash) (∥_∥.∣_∣) (conB x y) + 0-connB : (x y : typ B) → ∥ x ≡ y ∥₁ + 0-connB x y = ST.rec (isProp→isSet squash₁) (∥_∥₁.∣_∣₁) (conB x y) c = (c* B (preThom.Q B P) conB n Q-is Q-is-ptd .fst) c-ptd = (c* B (preThom.Q B P) conB n Q-is Q-is-ptd .snd) @@ -859,11 +857,11 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) E-susp : (i : ℕ) → GroupHom (coHomGr i E') (coHomRedGrDir (suc i) (E'̃ , inl tt)) fst (E-susp i) = - sMap λ f → (λ { (inl x) → 0ₖ _ + ST.map λ f → (λ { (inl x) → 0ₖ _ ; (inr x) → 0ₖ _ ; (push a j) → Kn→ΩKn+1 _ (f a) j}) , refl snd (E-susp zero) = - makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + makeIsGroupHom (ST.elim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn 1) (funExt λ { (inl x) → refl @@ -872,7 +870,7 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) (Kn→ΩKn+1-hom zero (f a) (g a) ∙ ∙≡+₁ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a))) k j}))) snd (E-susp (suc i)) = - makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + makeIsGroupHom (ST.elim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) (funExt λ { (inl x) → refl @@ -884,42 +882,42 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) module cofibSeq where j* : (i : ℕ) → GroupHom (coHomRedGrDir i (E'̃ , (inl tt))) (coHomGr i (typ B)) - fst (j* i) = sMap λ f → λ x → fst f (inr x) + fst (j* i) = ST.map λ f → λ x → fst f (inr x) snd (j* zero) = makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) + (ST.elim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) snd (j* (suc zero)) = makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) + (ST.elim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) snd (j* (suc (suc i₁))) = makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) + (ST.elim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) Im-j⊂Ker-p : (i : ℕ) (x : _) → isInIm (j* i) x → isInKer (p-hom i) x Im-j⊂Ker-p i = - sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) - λ f → pRec (squash₂ _ _) - (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + ST.elim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ f → PT.rec (squash₂ _ _) + (uncurry (ST.elim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) λ g P → subst (isInKer (p-hom i)) P (cong ∣_∣₂ (funExt λ x → cong (g .fst) (sym (push x)) ∙ g .snd)))) Ker-p⊂Im-j : (i : ℕ) (x : _) → isInKer (p-hom i) x → isInIm (j* i) x Ker-p⊂Im-j i = - sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) - λ f ker → pRec squash + ST.elim (λ _ → isSetΠ (λ _ → isProp→isSet squash₁)) + λ f ker → PT.rec squash₁ (λ id → ∣ ∣ (λ { (inl x) → 0ₖ _ ; (inr x) → f x ; (push a i₁) → funExt⁻ id a (~ i₁)}) , refl ∣₂ - , refl ∣) + , refl ∣₁) (Iso.fun PathIdTrunc₀Iso ker) Im-p⊂Ker-Susp : (i : ℕ) (x : _) → isInIm (p-hom i) x → isInKer (E-susp i) x Im-p⊂Ker-Susp i = - sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) - λ f → pRec (squash₂ _ _) - (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + ST.elim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ f → PT.rec (squash₂ _ _) + (uncurry (ST.elim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) λ g y → subst (isInKer (E-susp i)) y (cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) (funExt λ { (inl x) → refl @@ -929,8 +927,8 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) Ker-Susp⊂Im-p : (i : ℕ) (x : _) → isInKer (E-susp i) x → isInIm (p-hom i) x Ker-Susp⊂Im-p i = - sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) - λ f ker → pRec squash + ST.elim (λ _ → isSetΠ (λ _ → isProp→isSet squash₁)) + λ f ker → PT.rec squash₁ (λ id → ∣ ∣ (λ x → ΩKn+1→Kn i (sym (funExt⁻ (cong fst id) (inr x)))) ∣₂ , cong ∣_∣₂ (funExt (λ { (a , b) → cong (ΩKn+1→Kn i) @@ -944,16 +942,16 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) ∙ (sym (assoc _ _ _) ∙∙ cong (Kn→ΩKn+1 i (f (a , b)) ∙_) (rCancel _) ∙∙ sym (rUnit _))) - ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f (a , b))})) ∣) + ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f (a , b))})) ∣₁) (Iso.fun PathIdTrunc₀Iso ker) Im-Susp⊂Ker-j : (i : ℕ) (x : _) → isInIm (E-susp i) x → isInKer (cofibSeq.j* (suc i)) x Im-Susp⊂Ker-j i = - sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) - λ g → pRec (squash₂ _ _) - (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) - λ f id → pRec (squash₂ _ _) + ST.elim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ g → PT.rec (squash₂ _ _) + (uncurry (ST.elim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ f id → PT.rec (squash₂ _ _) (λ P → subst (isInKer (cofibSeq.j* (suc i))) (cong ∣_∣₂ P) (cong ∣_∣₂ refl)) (Iso.fun PathIdTrunc₀Iso id))) @@ -961,16 +959,16 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) Ker-j⊂Im-Susp : (i : ℕ) (x : _) → isInKer (cofibSeq.j* (suc i)) x → isInIm (E-susp i) x Ker-j⊂Im-Susp i = - sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) + ST.elim (λ _ → isSetΠ (λ _ → isProp→isSet squash₁)) λ f ker - → pRec squash + → PT.rec squash₁ (λ p → ∣ ∣ (λ x → ΩKn+1→Kn _ (sym (snd f) ∙∙ cong (fst f) (push x) ∙∙ funExt⁻ p (fst x))) ∣₂ , cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) (funExt (λ { (inl x) → sym (snd f) ; (inr x) → sym (funExt⁻ p x) - ; (push a j) k → h3 f p a k j}))) ∣) + ; (push a j) k → h3 f p a k j}))) ∣₁) (Iso.fun PathIdTrunc₀Iso ker) where h3 : (f : (E'̃ , inl tt) →∙ coHomK-ptd (suc i)) @@ -1040,14 +1038,14 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) Im-ϕ∘j⊂Ker-p : (i : ℕ) (x : _) → isInIm (ϕ∘j i) x → isInKer (p-hom _) x Im-ϕ∘j⊂Ker-p i x p = cofibSeq.Im-j⊂Ker-p _ x - (pRec squash (uncurry (λ f p → ∣ fst (fst (ϕ _)) f , p ∣)) p) + (PT.rec squash₁ (uncurry (λ f p → ∣ fst (fst (ϕ _)) f , p ∣₁)) p) Ker-p⊂Im-ϕ∘j : (i : ℕ) (x : _) → isInKer (p-hom _) x → isInIm (ϕ∘j i) x Ker-p⊂Im-ϕ∘j i x p = - pRec squash + PT.rec squash₁ (uncurry (λ f p → ∣ (invEq (fst (ϕ _)) f) - , (cong (fst (cofibSeq.j* _)) (secEq (fst (ϕ _)) f) ∙ p) ∣)) + , (cong (fst (cofibSeq.j* _)) (secEq (fst (ϕ _)) f) ∙ p) ∣₁)) (cofibSeq.Ker-p⊂Im-j _ x p) Im-p⊂KerSusp∘ϕ : (i : ℕ) (x : _) @@ -1065,7 +1063,7 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) Im-Susp∘ϕ⊂Ker-ϕ∘j : (i : ℕ) → (x : _) → isInIm (susp∘ϕ i) x → isInKer (ϕ∘j (suc i)) x Im-Susp∘ϕ⊂Ker-ϕ∘j i x = - pRec (squash₂ _ _) + PT.rec (squash₂ _ _) (uncurry λ f → J (λ x p → isInKer (ϕ∘j (suc i)) x) ((λ i → fst (cofibSeq.j* _) (fst (fst (ϕ _)) (fst (ϕ' _) (fst (E-susp _) f)))) ∙∙ cong (fst (cofibSeq.j* _)) @@ -1073,7 +1071,7 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) (invGroupEquiv (ϕ (suc i))) (+'-suc i n)) (fst (E-susp _) f)) ∙∙ (natTranspLem _ (λ n → fst (cofibSeq.j* n)) (sym (+'-suc i n)) ∙ cong (subst (λ z → coHomGr z (typ B) .fst) (sym (+'-suc i n))) - (Im-Susp⊂Ker-j _ (fst (E-susp (i +' n)) f) ∣ f , refl ∣) + (Im-Susp⊂Ker-j _ (fst (E-susp (i +' n)) f) ∣ f , refl ∣₁) ∙ tLem i n))) where tLem : (i n : ℕ) → subst (λ z → coHomGr z (typ B) .fst) (sym (+'-suc i n)) @@ -1086,9 +1084,9 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) Ker-ϕ∘j⊂Im-Susp∘ϕ : (i : ℕ) (x : _) → isInKer (ϕ∘j (suc i)) x → isInIm (susp∘ϕ i) x Ker-ϕ∘j⊂Im-Susp∘ϕ i x p = - pRec squash + PT.rec squash₁ (uncurry (λ f p → ∣ f , cong (fst (fst (thomIso' i))) p - ∙ secEq (fst (thomIso' _)) x ∣)) + ∙ secEq (fst (thomIso' _)) x ∣₁)) (Ker-j⊂Im-Susp _ (invEq (fst (thomIso' _)) x) ((cong (cofibSeq.j* (suc (i +' n)) .fst ) lem₁ ∙ natTranspLem _ (λ n → cofibSeq.j* n .fst) (+'-suc i n)) @@ -1115,7 +1113,7 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) ϕ∘j≡ : (i : ℕ) → ϕ∘j i ≡ ⌣-hom i ϕ∘j≡ i = Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + (funExt (ST.elim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl)) -- We can now restate the previous resluts for (λ x → x ⌣ e) diff --git a/Cubical/ZCohomology/MayerVietorisUnreduced.agda b/Cubical/ZCohomology/MayerVietorisUnreduced.agda index 031cd8c68b..9cf0def897 100644 --- a/Cubical/ZCohomology/MayerVietorisUnreduced.agda +++ b/Cubical/ZCohomology/MayerVietorisUnreduced.agda @@ -8,24 +8,25 @@ open import Cubical.Foundations.Path 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.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) + open import Cubical.Data.Nat +open import Cubical.Data.Sigma open import Cubical.Algebra.Group open import Cubical.Algebra.Group.DirProd open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.Group.Instances.Bool renaming (Bool to BoolGroup) -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.Instances.Bool +open import Cubical.Algebra.Group.Instances.Int open import Cubical.Algebra.Group.Instances.Unit +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.S1 +open import Cubical.HITs.Susp + open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.Properties open import Cubical.ZCohomology.GroupStructure @@ -38,11 +39,11 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : 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)) ∣₂ + i* n = ST.rec (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) + makeIsGroupHom (ST.elim2 (λ _ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) λ _ _ → refl) i : (n : ℕ) → GroupHom (coHomGr n (Pushout f g)) (×coHomGr n A B) fst (i n) = i* n @@ -91,24 +92,24 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : ; (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 : ℕ) → IsGroupHom (coHomGr n C .snd) (ST.rec isSetSetTrunc λ a → ∣ d-pre n a ∣₂) (coHomGr (suc n) (Pushout f g) .snd) dIsHom n = makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + (ST.elim2 (λ _ _ → 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 ∣₂ + fst (d n) = ST.rec 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) _ _) + Im-d⊂Ker-i n = ST.elim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) + λ a → PT.rec (isOfHLevelPath' 1 (isSet× isSetSetTrunc isSetSetTrunc) _ _) (sigmaElim (λ _ → isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) - λ δ b i → sRec (isSet× isSetSetTrunc isSetSetTrunc) + λ δ b i → ST.rec (isSet× isSetSetTrunc isSetSetTrunc) (λ δ → ∣ (λ x → δ (inl x)) ∣₂ , ∣ (λ x → δ (inr x)) ∣₂ ) (b (~ i))) @@ -116,10 +117,10 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : → 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)} + ST.elim (λ _ → isSetΠ λ _ → isProp→isSet isPropPropTrunc) + λ a p → PT.rec {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) + (λ p1 → PT.rec 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 δ)) ∣₁) @@ -150,11 +151,11 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : → isInIm (i n) x → isInKer (Δ n) x Im-i⊂Ker-Δ n (Fa , Fb) = - sElim {B = λ Fa → (Fb : _) → isInIm (i n) (Fa , Fb) + ST.elim {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 _ _) + (λ Fa → ST.elim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ Fb → PT.rec (isSetSetTrunc _ _) (sigmaElim (λ x → isProp→isSet (isSetSetTrunc _ _)) λ Fd p → helper n Fa Fb Fd p)) Fa @@ -171,7 +172,7 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : → isInKer (Δ n) a → isInIm (i n) a Ker-Δ⊂Im-i n = prodElim (λ _ → isSetΠ (λ _ → isProp→isSet isPropPropTrunc)) - (λ Fa Fb p → pRec isPropPropTrunc + (λ Fa Fb p → PT.rec isPropPropTrunc (λ q → ∣ ∣ helpFun Fa Fb q ∣₂ , refl ∣₁) (helper Fa Fb p)) where @@ -210,8 +211,8 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : → 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)) ∣₂ , + ST.elim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 isPropPropTrunc) + λ Fc p → PT.rec 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) @@ -230,10 +231,10 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : → isInIm (Δ n) a → isInKer (d n) a Im-Δ⊂Ker-d n = - sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) - λ Fc → pRec (isOfHLevelPath' 1 isSetSetTrunc _ _) + ST.elim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) + λ Fc → PT.rec (isOfHLevelPath' 1 isSetSetTrunc _ _) (sigmaProdElim (λ _ → isOfHLevelPath 2 isSetSetTrunc _ _) - λ Fa Fb p → pRec (isOfHLevelPath' 1 isSetSetTrunc _ _) + λ Fa Fb p → PT.rec (isOfHLevelPath' 1 isSetSetTrunc _ _) (λ q → ((λ i → fst (d n) ∣ (q (~ i)) ∣₂) ∙ dΔ-Id n Fa Fb)) (Iso.fun (PathIdTrunc₀Iso) p)) diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 3ee13d2a64..16caa740e0 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -29,26 +29,24 @@ open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Functions.Morphism -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.HITs.S1 hiding (encode ; decode ; _·_) -open import Cubical.HITs.Sn -open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; map2 to trMap2; rec to trRec ; elim3 to trElim3) - -open import Cubical.Data.Sum.Base hiding (map) -open import Cubical.Data.Int renaming (_+_ to _ℤ+_) hiding (-_) open import Cubical.Data.Nat +open import Cubical.Data.Int renaming (_+_ to _ℤ+_) hiding (-_) open import Cubical.Data.Sigma - -open import Cubical.Homotopy.Loopspace -open import Cubical.Homotopy.Connected +open import Cubical.Data.Sum.Base open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.AbGroup -open import Cubical.Algebra.Semigroup -open import Cubical.Algebra.Monoid + +open import Cubical.HITs.Truncation as T +open import Cubical.HITs.SetTruncation as ST renaming (isSetSetTrunc to §) +open import Cubical.HITs.S1 hiding (encode ; decode ; _·_) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp + +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.GroupStructure @@ -62,12 +60,12 @@ private ------------------- Connectedness --------------------- is2ConnectedKn : (n : ℕ) → isConnected 2 (coHomK (suc n)) is2ConnectedKn zero = ∣ ∣ base ∣ ∣ - , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) - (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isOfHLevelTrunc 2)) _ _) + , T.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (T.elim (λ _ → 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 _ _)) + , T.elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (T.elim (λ _ → isProp→isOfHLevelSuc (3 + n) (isOfHLevelTrunc 2 _ _)) (suspToPropElim (ptSn (suc n)) (λ _ → isOfHLevelTrunc 2 _ _) refl)) isConnectedKn : (n : ℕ) → isConnected (2 + n) (coHomK (suc n)) @@ -76,19 +74,19 @@ isConnectedKn n = isOfHLevelRetractFromIso 0 (invIso (truncOfTruncIso (2 + n) 1) -- 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)) + T.elim (λ _ → 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)) _ _) + , T.elim (λ _ → 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)))) + T.elim (λ _ → 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 ∣ @@ -105,7 +103,7 @@ coHomPointedElim : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → Typ → ((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 _)) + ST.elim (λ _ → isOfHLevelSuc 1 (isprop _)) λ f → helper n isprop indp f (f a) refl where helper : (n : ℕ) {B : coHom (suc n) A → Type ℓ'} @@ -116,16 +114,16 @@ coHomPointedElim {ℓ' = ℓ'} {A = A} n a isprop indp = → 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 _)) + T.elim (λ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ λ _ → isprop _)) (toPropElim (λ _ → isPropΠ λ _ → isprop _) (ind f)) helper (suc zero) isprop ind f = - trElim (λ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ λ _ → isprop _)) + T.elim (λ _ → 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 _)) + T.elim (λ _ → 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 _)) + T.elim (λ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ λ _ → isprop _)) (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) @@ -133,7 +131,7 @@ coHomPointedElim2 : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → co → ((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 _ _)) +coHomPointedElim2 {ℓ' = ℓ'} {A = A} n a isprop indp = ST.elim2 (λ _ _ → 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 ℓ'} @@ -144,16 +142,16 @@ coHomPointedElim2 {ℓ' = ℓ'} {A = A} n a isprop indp = sElim2 (λ _ _ → isO → f a ≡ x → g a ≡ y → B ∣ f ∣₂ ∣ g ∣₂ helper zero a isprop indp f g = - elim2 (λ _ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ2 λ _ _ → isprop _ _)) + T.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 _ _)) + T.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 _ _)) + T.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 _ _)) + T.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 ℓ} @@ -161,7 +159,7 @@ coHomK-elim : ∀ {ℓ} (n : ℕ) {B : coHomK (suc n) → Type ℓ} → B (0ₖ (suc n)) → (x : _) → B x coHomK-elim n {B = B } hlev b = - trElim (λ _ → isOfHLevelPlus {n = (suc n)} 2 (hlev _)) + T.elim (λ _ → isOfHLevelPlus {n = (suc n)} 2 (hlev _)) (sphereElim _ (hlev ∘ ∣_∣) b) {- Equivalence between cohomology of A and reduced cohomology of (A + 1) -} @@ -171,40 +169,40 @@ coHomRed+1Equiv : (n : ℕ) → 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 + helpLemma {C = C} = isoToPath (iso map-helper1 + map-helper2 (λ b → linvPf b) (λ _ → refl)) where - map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →∙ C)) - map1 f = map1' , refl + map-helper1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →∙ C)) + map-helper1 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) + map-helper2 : ((((A ⊎ Unit) , inr (tt)) →∙ C)) → (A → typ C) + map-helper2 (g , pf) x = g (inl x) - linvPf : (b :((((A ⊎ Unit) , inr (tt)) →∙ C))) → map1 (map2 b) ≡ b + linvPf : (b :((((A ⊎ Unit) , inr (tt)) →∙ C))) → map-helper1 (map-helper2 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') (map-helper2 (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ₖ _ _ +fun (Iso-coHom-coHomRed {A = A , a} n) = ST.map fst +inv' (Iso-coHom-coHomRed {A = A , a} n) = ST.map λ f → (λ x → f x -ₖ f a) , rCancelₖ _ _ rightInv (Iso-coHom-coHomRed {A = A , a} n) = - sElim (λ _ → isOfHLevelPath 2 § _ _) - λ f → trRec (isProp→isOfHLevelSuc _ (§ _ _)) + ST.elim (λ _ → isOfHLevelPath 2 § _ _) + λ f → T.rec (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 § _ _) + ST.elim (λ _ → isOfHLevelPath 2 § _ _) λ {(f , p) → cong ∣_∣₂ (ΣPathP (((funExt λ x → (cong (λ y → f x -ₖ y) p ∙∙ cong (λ y → f x +ₖ y) -0ₖ ∙∙ rUnitₖ _ (f x)) ∙ refl)) @@ -227,8 +225,8 @@ leftInv (Iso-coHom-coHomRed {A = A , a} n) = → 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 ++∙≡+ zero = ST.elim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl ++∙≡+ (suc n) = ST.elim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl private homhelp : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) (x y : coHom (suc n) (typ A)) @@ -267,7 +265,7 @@ abstract private module _ (n : ℕ) where σ : {n : ℕ} → coHomK (suc n) → Path (coHomK (2 + n)) ∣ north ∣ ∣ north ∣ - σ {n = n} = trRec (isOfHLevelTrunc (4 + n) _ _) + σ {n = n} = T.rec (isOfHLevelTrunc (4 + n) _ _) λ a → cong ∣_∣ (merid a ∙ sym (merid (ptSn (suc n)))) σ-hom-helper : ∀ {ℓ} {A : Type ℓ} {a : A} (p : a ≡ a) (r : refl ≡ p) @@ -276,7 +274,7 @@ private σ-hom : {n : ℕ} (x y : coHomK (suc n)) → σ (x +ₖ y) ≡ σ x ∙ σ y σ-hom {n = zero} = - elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 4 _ _) _ _) + T.elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 4 _ _) _ _) (wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) (λ x → lUnit _ @@ -286,7 +284,7 @@ private ∙∙ 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) _ _) _ _) + T.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))))) @@ -310,7 +308,7 @@ private -- we define the code using addIso Code : (n : ℕ) → coHomK (2 + n) → Type₀ - Code n x = (trRec {B = TypeOfHLevel ℓ-zero (3 + n)} (isOfHLevelTypeOfHLevel (3 + n)) + Code n x = (T.rec {B = TypeOfHLevel ℓ-zero (3 + n)} (isOfHLevelTypeOfHLevel (3 + n)) λ a → Code' a , hLevCode' a) x .fst where Code' : (S₊ (2 + n)) → Type₀ @@ -323,11 +321,11 @@ private 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)) _ _) + symMeridLem n x = T.elim (λ _ → 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 {n = n} = T.elim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) decode-elim where north≡merid : (a : S₊ (suc n)) @@ -337,7 +335,7 @@ private 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) _ _) + decode-elim south = T.rec (isOfHLevelTrunc (4 + n) _ _) λ a → cong ∣_∣ (merid a) decode-elim (merid a i) = hcomp (λ k → λ { (i = i0) → σ @@ -346,8 +344,8 @@ private 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) _ _) _ _) + ≡ T.rec (isOfHLevelTrunc (4 + n) _ _) λ a → cong ∣_∣ (merid a) + mainPath a = funExt (T.elim (λ _ → 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))) @@ -377,7 +375,7 @@ private -- 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))) + T.elim (λ _ → isProp→isOfHLevelSuc (3 + n) (isPropIsOfHLevel (3 + n))) (sphereToPropElim _ (λ _ → (isPropIsOfHLevel (3 + n))) (isOfHLevelTrunc (3 + n))) @@ -401,14 +399,14 @@ private → 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))) _ _) + (T.elim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelΠ (3 + n) (λ _ → isOfHLevelTrunc (3 + n))) _ _) λ x → funExt - (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) + (T.elim (λ _ → 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)) + T.elim (λ 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) @@ -423,7 +421,7 @@ 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)) _ _) + T.elim (λ _ → 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) ∣) @@ -476,8 +474,8 @@ 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) + cong winding (congFunct (T.rec isGroupoidS¹ (λ x → x)) p q) + ∙ winding-hom (cong (T.rec isGroupoidS¹ (λ x → x)) p) (cong (T.rec 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) @@ -501,17 +499,17 @@ isHomogeneousKn n = 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) +fun (fst (coHom≅coHomΩ n A)) = ST.map λ f a → Kn→ΩKn+1 n (f a) +inv' (fst (coHom≅coHomΩ n A)) = ST.map λ f a → ΩKn+1→Kn n (f a) rightInv (fst (coHom≅coHomΩ n A)) = - sElim (λ _ → isOfHLevelPath 2 § _ _) + ST.elim (λ _ → isOfHLevelPath 2 § _ _) λ f → cong ∣_∣₂ (funExt λ x → rightInv (Iso-Kn-ΩKn+1 n) (f x)) leftInv (fst (coHom≅coHomΩ n A)) = - sElim (λ _ → isOfHLevelPath 2 § _ _) + ST.elim (λ _ → isOfHLevelPath 2 § _ _) λ f → cong ∣_∣₂ (funExt λ x → leftInv (Iso-Kn-ΩKn+1 n) (f x)) snd (coHom≅coHomΩ n A) = makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) + (ST.elim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ f g → cong ∣_∣₂ (funExt λ x → Kn→ΩKn+1-hom n (f x) (g x))) module lockedKnIso (key : Unit') where @@ -547,11 +545,11 @@ 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ℤ) _ _) + (funExt (T.elim (λ _ → 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} + T.elim {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 @@ -612,7 +610,7 @@ snd (contr∙-lem n m) (f , p) = →∙Homogeneous≡ (isHomogeneous→∙ (isHo where help : (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) help = - trElim (λ _ → isOfHLevelPath (3 + n) + T.elim (λ _ → 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) @@ -657,7 +655,7 @@ isoType→isoCohom : {A : Type ℓ} {B : Type ℓ'} (n : ℕ) → 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₂ _ _) + makeIsGroupHom (ST.elim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) (λ _ _ → refl)) -- Explicit index swapping for cohomology groups diff --git a/Cubical/ZCohomology/RingStructure/CohomologyRing.agda b/Cubical/ZCohomology/RingStructure/CohomologyRing.agda index da0365cec6..6970c1dd2f 100644 --- a/Cubical/ZCohomology/RingStructure/CohomologyRing.agda +++ b/Cubical/ZCohomology/RingStructure/CohomologyRing.agda @@ -14,8 +14,8 @@ open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.AbGroup open import Cubical.Algebra.Ring -open import Cubical.Algebra.Direct-Sum.Base -open import Cubical.Algebra.AbGroup.Instances.Direct-Sum +open import Cubical.Algebra.DirectSum.Base +open import Cubical.Algebra.AbGroup.Instances.DirectSum open import Cubical.HITs.SetTruncation as ST diff --git a/Cubical/ZCohomology/RingStructure/CupProduct.agda b/Cubical/ZCohomology/RingStructure/CupProduct.agda index b2a0f26976..900e102271 100644 --- a/Cubical/ZCohomology/RingStructure/CupProduct.agda +++ b/Cubical/ZCohomology/RingStructure/CupProduct.agda @@ -1,20 +1,22 @@ {-# 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 +open import Cubical.Data.Int hiding (_+'_ ; +'≡+ ; _+_) + +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.Truncation as T +open import Cubical.HITs.S1 hiding (_·_) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties infixl 30 _·₀_ infixr 35 _⌣ₖ_ @@ -53,7 +55,7 @@ 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) +⌣ₖ∙ (suc n) (suc m) = T.rec (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ₖ _ @@ -87,4 +89,4 @@ 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) ∣₂ +_⌣_ = ST.rec2 squash₂ λ f g → ∣ (λ x → f x ⌣ₖ g x) ∣₂ diff --git a/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda index 445d07429d..e33024424d 100644 --- a/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda +++ b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda @@ -1,22 +1,5 @@ {-# 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 @@ -27,19 +10,35 @@ 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.Empty as ⊥ open import Cubical.Data.Nat +open import Cubical.Data.Int + renaming (_+_ to _ℤ+_ ; _·_ to _ℤ∙_ ; +Comm to +ℤ-comm ; ·Comm to ∙-comm ; +Assoc to ℤ+-assoc ; -_ to -ℤ_) + hiding (_+'_ ; +'≡+) open import Cubical.Data.Sigma -open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.Data.Sum +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.Truncation as T +open import Cubical.HITs.S1 hiding (_·_) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp + +open import Cubical.Homotopy.Loopspace + +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 + private variable ℓ : Level + + 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) @@ -90,7 +89,7 @@ private (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) +-ₖ'-gen {k = suc k} n m p q = T.map (-ₖ'-helper {k = suc k} n m p q) -- -ₖ'ⁿ̇*ᵐ -ₖ'^_·_ : {k : ℕ} (n m : ℕ) → coHomK k → coHomK k @@ -98,18 +97,18 @@ private -- cohomology version -ₕ'^_·_ : {k : ℕ} {A : Type ℓ} (n m : ℕ) → coHom k A → coHom k A --ₕ'^_·_ n m = sMap λ f x → (-ₖ'^ n · m) (f x) +-ₕ'^_·_ n m = ST.map λ 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)) + T.elim ((λ _ → isOfHLevelTruncPath)) λ { base → refl ; (loop i) → refl} -ₖ'-gen-inr≡-ₖ' {k = suc (suc k)} n m p q = - trElim ((λ _ → isOfHLevelTruncPath)) + T.elim ((λ _ → isOfHLevelTruncPath)) λ { north → refl ; south → refl ; (merid a i) k → ∣ symDistr (merid (ptSn _)) (sym (merid a)) (~ k) (~ i) ∣ₕ} @@ -119,10 +118,10 @@ private → -ₖ'-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) + T.elim (λ _ → isOfHLevelTruncPath) λ { base → refl ; (loop i) → refl} -ₖ'-gen-inl-left {k = suc (suc k)} n m p q = - trElim (λ _ → isOfHLevelPath (4 + k) (isOfHLevelTrunc (4 + k)) _ _) + T.elim (λ _ → 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 ∣ₕ} @@ -133,18 +132,18 @@ private -ₖ'-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) + T.elim (λ _ → isOfHLevelTruncPath) λ { base → refl ; (loop i) → refl} -ₖ'-gen-inl-right {k = suc zero} n m (inr x₁) q = - trElim (λ _ → isOfHLevelTruncPath) + T.elim (λ _ → isOfHLevelTruncPath) λ { base → refl ; (loop i) → refl} -ₖ'-gen-inl-right {k = suc (suc k)} n m (inl x₁) q = - trElim (λ _ → isOfHLevelTruncPath) + T.elim (λ _ → 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) + T.elim (λ _ → isOfHLevelTruncPath) λ { north → refl ; south → cong ∣_∣ₕ (merid (ptSn _)) ; (merid a i) k → ∣ compPath-filler (merid a) (sym (merid (ptSn _))) (~ k) i ∣ₕ} @@ -186,7 +185,7 @@ cong-ₖ'-gen-inr : {k : ℕ} (n m : ℕ) (p : _) (q : _) (P : Path (coHomK (2 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) + code = T.elim (λ _ → 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} @@ -874,9 +873,9 @@ gradedComm'-elimCase k zero (suc m) term (inr tt) q 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)) + ⊥.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)) + ⊥.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 = @@ -1048,7 +1047,7 @@ gradedComm'-⌣ₖ∙ : (n m : ℕ) (p : _) (q : _) (a : _) (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)) _ _) + T.elim (λ _ → 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) @@ -1067,7 +1066,7 @@ gradedComm'-⌣ₖ∙ n m p q = f₁≡f₂ : (b : _) → f₁ b ≡ f₂ b f₁≡f₂ = - trElim (λ _ → isOfHLevelPath (3 + m) + T.elim (λ _ → isOfHLevelPath (3 + m) (subst (isOfHLevel (3 + m)) (λ i → S₊∙ (suc n) →∙ coHomK-ptd (+'-comm (suc n) (suc m) (~ i))) (isOfHLevel↑∙' (suc m) n)) _ _) @@ -1095,7 +1094,7 @@ gradedComm'-⌣ₖ (suc n) (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₂ _ _) + ST.elim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → cong ∣_∣₂ (funExt (λ x → gradedComm'-⌣ₖ n m (f x) (g x) @@ -1120,10 +1119,10 @@ gradedComm'-⌣ n m = -ₕ^-gen-eq : ∀ {ℓ} {k : ℕ} {A : Type ℓ} (n m : ℕ) → (p : isEvenT n ⊎ isOddT n) (q : isEvenT m ⊎ isOddT m) → (x : coHom k A) - → -ₕ^-gen n m p q x ≡ (sMap λ f x → (-ₖ'-gen n m p q) (f x)) x --ₕ^-gen-eq {k = k} n m (inl p) q = sElim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ x → sym (-ₖ'-gen-inl-left n m p q (f x))) --ₕ^-gen-eq {k = k} n m (inr p) (inl q) = sElim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ z → sym (-ₖ'-gen-inl-right n m (inr p) q (f z))) --ₕ^-gen-eq {k = k} n m (inr p) (inr q) = sElim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ z → sym (-ₖ'-gen-inr≡-ₖ' n m p q (f z))) + → -ₕ^-gen n m p q x ≡ (ST.map λ f x → (-ₖ'-gen n m p q) (f x)) x +-ₕ^-gen-eq {k = k} n m (inl p) q = ST.elim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ x → sym (-ₖ'-gen-inl-left n m p q (f x))) +-ₕ^-gen-eq {k = k} n m (inr p) (inl q) = ST.elim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ z → sym (-ₖ'-gen-inl-right n m (inr p) q (f z))) +-ₕ^-gen-eq {k = k} n m (inr p) (inr q) = ST.elim (λ _ → isSetPathImplicit) λ f → cong ∣_∣₂ (funExt λ z → sym (-ₖ'-gen-inr≡-ₖ' n m p q (f z))) -ₕ^-eq : ∀ {ℓ} {k : ℕ} {A : Type ℓ} (n m : ℕ) → (a : coHom k A) → (-ₕ^ n · m) a ≡ (-ₕ'^ n · m) a diff --git a/Cubical/ZCohomology/RingStructure/RingLaws.agda b/Cubical/ZCohomology/RingStructure/RingLaws.agda index 6d85e407bf..c7436bbf15 100644 --- a/Cubical/ZCohomology/RingStructure/RingLaws.agda +++ b/Cubical/ZCohomology/RingStructure/RingLaws.agda @@ -1,19 +1,6 @@ {-# 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.Path @@ -21,17 +8,32 @@ 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.Int + renaming (_+_ to _ℤ+_ ; _·_ to _ℤ∙_ ; +Comm to +ℤ-comm ; ·Comm to ∙-comm ; +Assoc to ℤ+-assoc ; -_ to -ℤ_) + hiding (_+'_ ; +'≡+) open import Cubical.Data.Sigma +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.Truncation as T +open import Cubical.HITs.S1 hiding (_·_) +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp + +open import Cubical.Homotopy.Loopspace + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.RingStructure.CupProduct + private variable ℓ : Level + + -- Some boring lemmas ·₀≡·ℤ : (x y : ℤ) → _·₀_ {n = zero} x y ≡ x ℤ∙ y ·₀≡·ℤ (pos zero) y = refl @@ -105,7 +107,7 @@ private 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 _ _)) + T.elim2 (λ _ _ → isOfHLevelSuc (2 + n) (hLevHelp n m _ _)) main where hLevHelp : (n m : ℕ) (x y : _) @@ -235,7 +237,7 @@ private 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) _ _) + T.elim2 (λ _ _ → isOfHLevelPath (3 + n) (isOfHLevel↑∙ (suc n) m) _ _) main where fst-left : (n : ℕ) (y : S₊ (suc n)) → @@ -360,7 +362,7 @@ private -- 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) _ _) + T.elim (λ _ → isOfHLevelPath (3 + n) (hLev-assocer-helpFun n m) _ _) λ a → →∙Homogeneous≡ (subst isHomogeneous Kn≃ΩKn+1∙ (isHomogeneousKn _)) (funExt (main n a)) where @@ -393,7 +395,7 @@ 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) + T.elim (λ _ → 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))) _ _) @@ -593,33 +595,33 @@ assoc-⌣ₖ (suc n) (suc m) (suc k) x y z = -- Ring laws for ⌣ module _ {A : Type ℓ} (n m : ℕ) where ⌣-0ₕ : (f : coHom n A) → (f ⌣ 0ₕ m) ≡ 0ₕ _ - ⌣-0ₕ = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + ⌣-0ₕ = ST.elim (λ _ → 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₂ _ _) + 0ₕ-⌣ = ST.elim (λ _ → 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₂ _ _) + ST.elim (λ _ → isSetΠ2 λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → ST.elim2 (λ _ _ → 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₂ _ _) + ST.elim2 (λ _ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ g h → ST.elim (λ _ → 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 → + ST.elim (λ _ → isSetΠ2 λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → ST.elim (λ _ → isSetΠ λ _ → isOfHLevelPath 2 squash₂ _ _) + λ g → ST.elim (λ _ → 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))) @@ -628,12 +630,12 @@ module _ {A : Type ℓ} (n m : ℕ) where -- Additive unit(s) 0⌣ : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) (x : coHom n A) → x ⌣ (0ₕ m) ≡ 0ₕ _ -0⌣ n m = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +0⌣ n m = ST.elim (λ _ → 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₂ _ _) +⌣0 n m = ST.elim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ f → cong ∣_∣₂ (funExt λ x → 0ₖ-⌣ₖ m n (f x)) -- Multiplicative unit @@ -651,31 +653,31 @@ lUnit⌣ₖ (suc n) x = rUnitₖ _ x lUnit⌣ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) (x : coHom n A) → x ⌣ 1⌣ ≡ subst (λ n → coHom n A) (sym (n+'0 n)) x -lUnit⌣ zero = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +lUnit⌣ zero = ST.elim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ f → cong ∣_∣₂ (funExt (λ x → comm-·₀ (f x) (pos 1))) ∙ sym (transportRefl ∣ f ∣₂) lUnit⌣ (suc n) = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + ST.elim (λ _ → 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₂ _ _) +rUnit⌣ zero = ST.elim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ f → refl rUnit⌣ (suc n) = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + ST.elim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ f → cong ∣_∣₂ (funExt λ x → rUnitₖ _ (f x)) -ₕDistᵣ : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) (x : coHom n A) (y : coHom m A) → (-ₕ (x ⌣ y)) ≡ x ⌣ (-ₕ y) -ₕDistᵣ n m = - sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + ST.elim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → cong ∣_∣₂ (funExt λ x → -Distᵣ n m (f x) (g x)) -ₕDistₗ : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) (x : coHom n A) (y : coHom m A) → (-ₕ (x ⌣ y)) ≡ (-ₕ x) ⌣ y -ₕDistₗ n m = - sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + ST.elim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → cong ∣_∣₂ (funExt λ x → -Distₗ n m (f x) (g x)) -ₕDistₗᵣ : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) @@ -683,7 +685,7 @@ rUnit⌣ (suc n) = -ₕDistₗᵣ n m x y = sym (-ₕDistₗ n m x (-ₕ y)) ∙∙ cong -ₕ_ (sym (-ₕDistᵣ n m x y)) - ∙∙ sElim2 {C = λ x y → (-ₕ (-ₕ (x ⌣ y))) ≡ x ⌣ y} + ∙∙ ST.elim2 {C = λ x y → (-ₕ (-ₕ (x ⌣ y))) ≡ x ⌣ y} (λ _ _ → isSetPathImplicit) (λ f g → cong ∣_∣₂ (funExt λ _ → -ₖ^2 _)) x y diff --git a/NAMING.md b/NAMING.md index 542592c50a..2fc666c4e6 100644 --- a/NAMING.md +++ b/NAMING.md @@ -44,3 +44,15 @@ For naming conventions specific to the Algebra subfolder, see should appear in the order they appear in the type (like `isContrUnit`). For functions things can either be separated by `→` (like `isProp→isSet`) or `To` (like `isoToEquiv`). + +* The `elim` and `rec` should be use as much as possible without + renaming but by importing and renaming the module. + For instance use "open import Cubical.Data.Empty as ⊥` + then use `⊥.rec` or `⊥.elim' rather than doing + `renaming (rec to rec-⊥)` and using `rec-⊥`. + + Some convetional naming : + - Empty -> ⊥ + - PropositionalTruncation -> PT + - SetTruncation -> ST + - SetQuotient -> SQ