From d4648741cb5a07716f717368dde6081c0a3fceb6 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 29 Jul 2024 17:58:16 +0200 Subject: [PATCH 01/83] move univalence code --- Cubical/Algebra/CommRing/Base.agda | 101 ------------------ Cubical/Algebra/CommRing/Univalence.agda | 129 +++++++++++++++++++++++ 2 files changed, 129 insertions(+), 101 deletions(-) create mode 100644 Cubical/Algebra/CommRing/Univalence.agda diff --git a/Cubical/Algebra/CommRing/Base.agda b/Cubical/Algebra/CommRing/Base.agda index 36f44a8001..bedb3fd91b 100644 --- a/Cubical/Algebra/CommRing/Base.agda +++ b/Cubical/Algebra/CommRing/Base.agda @@ -130,104 +130,3 @@ isPropIsCommRing 0r 1r _+_ _·_ -_ = (λ ring → isPropΠ2 (λ _ _ → is-set ring _ _))) where open IsRing - -𝒮ᴰ-CommRing : DUARel (𝒮-Univ ℓ) CommRingStr ℓ -𝒮ᴰ-CommRing = - 𝒮ᴰ-Record (𝒮-Univ _) IsCommRingEquiv - (fields: - data[ 0r ∣ null ∣ pres0 ] - data[ 1r ∣ null ∣ pres1 ] - data[ _+_ ∣ bin ∣ pres+ ] - data[ _·_ ∣ bin ∣ pres· ] - data[ -_ ∣ autoDUARel _ _ ∣ pres- ] - prop[ isCommRing ∣ (λ _ _ → isPropIsCommRing _ _ _ _ _) ]) - where - open CommRingStr - open IsRingHom - - -- faster with some sharing - null = autoDUARel (𝒮-Univ _) (λ A → A) - bin = autoDUARel (𝒮-Univ _) (λ A → A → A → A) - -CommRingPath : (R S : CommRing ℓ) → CommRingEquiv R S ≃ (R ≡ S) -CommRingPath = ∫ 𝒮ᴰ-CommRing .UARel.ua - -uaCommRing : {A B : CommRing ℓ} → CommRingEquiv A B → A ≡ B -uaCommRing {A = A} {B = B} = equivFun (CommRingPath A B) - -CommRingIso : (R : CommRing ℓ) (S : CommRing ℓ') → Type (ℓ-max ℓ ℓ') -CommRingIso R S = Σ[ e ∈ Iso (R .fst) (S .fst) ] - IsRingHom (CommRingStr→RingStr (R .snd)) (e .fun) (CommRingStr→RingStr (S .snd)) - -CommRingEquivIsoCommRingIso : (R : CommRing ℓ) (S : CommRing ℓ') → Iso (CommRingEquiv R S) (CommRingIso R S) -fst (fun (CommRingEquivIsoCommRingIso R S) e) = equivToIso (e .fst) -snd (fun (CommRingEquivIsoCommRingIso R S) e) = e .snd -fst (inv (CommRingEquivIsoCommRingIso R S) e) = isoToEquiv (e .fst) -snd (inv (CommRingEquivIsoCommRingIso R S) e) = e .snd -rightInv (CommRingEquivIsoCommRingIso R S) (e , he) = - Σ≡Prop (λ e → isPropIsRingHom (snd (CommRing→Ring R)) (e .fun) (snd (CommRing→Ring S))) - rem - where - rem : equivToIso (isoToEquiv e) ≡ e - fun (rem i) x = fun e x - inv (rem i) x = inv e x - rightInv (rem i) b j = CommRingStr.is-set (snd S) (fun e (inv e b)) b (rightInv e b) (rightInv e b) i j - leftInv (rem i) a j = CommRingStr.is-set (snd R) (inv e (fun e a)) a (retEq (isoToEquiv e) a) (leftInv e a) i j -leftInv (CommRingEquivIsoCommRingIso R S) e = - Σ≡Prop (λ e → isPropIsRingHom (snd (CommRing→Ring R)) (e .fst) (snd (CommRing→Ring S))) - (equivEq refl) - -isGroupoidCommRing : isGroupoid (CommRing ℓ) -isGroupoidCommRing _ _ = isOfHLevelRespectEquiv 2 (CommRingPath _ _) (isSetRingEquiv _ _) - -open CommRingStr -open IsRingHom - --- TODO: Induced structure results are temporarily inconvenient while we transition between algebra --- representations -module _ (R : CommRing ℓ) {A : Type ℓ} - (0a 1a : A) - (add mul : A → A → A) - (inv : A → A) - (e : ⟨ R ⟩ ≃ A) - (p0 : e .fst (R .snd .0r) ≡ 0a) - (p1 : e .fst (R .snd .1r) ≡ 1a) - (p+ : ∀ x y → e .fst (R .snd ._+_ x y) ≡ add (e .fst x) (e .fst y)) - (p· : ∀ x y → e .fst (R .snd ._·_ x y) ≡ mul (e .fst x) (e .fst y)) - (pinv : ∀ x → e .fst (R .snd .-_ x) ≡ inv (e .fst x)) - where - - private - module R = CommRingStr (R .snd) - - BaseΣ : Type (ℓ-suc ℓ) - BaseΣ = Σ[ B ∈ Type ℓ ] B × B × (B → B → B) × (B → B → B) × (B → B) - - FamilyΣ : BaseΣ → Type ℓ - FamilyΣ (B , u0 , u1 , a , m , i) = IsCommRing u0 u1 a m i - - inducedΣ : FamilyΣ (A , 0a , 1a , add , mul , inv) - inducedΣ = - subst FamilyΣ - (UARel.≅→≡ (autoUARel BaseΣ) (e , p0 , p1 , p+ , p· , pinv)) - R.isCommRing - - InducedCommRing : CommRing ℓ - InducedCommRing .fst = A - 0r (InducedCommRing .snd) = 0a - 1r (InducedCommRing .snd) = 1a - _+_ (InducedCommRing .snd) = add - _·_ (InducedCommRing .snd) = mul - - InducedCommRing .snd = inv - isCommRing (InducedCommRing .snd) = inducedΣ - - InducedCommRingEquiv : CommRingEquiv R InducedCommRing - fst InducedCommRingEquiv = e - pres0 (snd InducedCommRingEquiv) = p0 - pres1 (snd InducedCommRingEquiv) = p1 - pres+ (snd InducedCommRingEquiv) = p+ - pres· (snd InducedCommRingEquiv) = p· - pres- (snd InducedCommRingEquiv) = pinv - - InducedCommRingPath : R ≡ InducedCommRing - InducedCommRingPath = CommRingPath _ _ .fst InducedCommRingEquiv diff --git a/Cubical/Algebra/CommRing/Univalence.agda b/Cubical/Algebra/CommRing/Univalence.agda new file mode 100644 index 0000000000..a5cc63b970 --- /dev/null +++ b/Cubical/Algebra/CommRing/Univalence.agda @@ -0,0 +1,129 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommRing.Univalence where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.CommRing.Base +open import Cubical.Algebra.Ring.Base + +open import Cubical.Reflection.RecordEquiv + + +private + variable + ℓ ℓ' : Level + +open Iso + +𝒮ᴰ-CommRing : DUARel (𝒮-Univ ℓ) CommRingStr ℓ +𝒮ᴰ-CommRing = + 𝒮ᴰ-Record (𝒮-Univ _) IsCommRingEquiv + (fields: + data[ 0r ∣ null ∣ pres0 ] + data[ 1r ∣ null ∣ pres1 ] + data[ _+_ ∣ bin ∣ pres+ ] + data[ _·_ ∣ bin ∣ pres· ] + data[ -_ ∣ autoDUARel _ _ ∣ pres- ] + prop[ isCommRing ∣ (λ _ _ → isPropIsCommRing _ _ _ _ _) ]) + where + open CommRingStr + open IsRingHom + + -- faster with some sharing + null = autoDUARel (𝒮-Univ _) (λ A → A) + bin = autoDUARel (𝒮-Univ _) (λ A → A → A → A) + +CommRingPath : (R S : CommRing ℓ) → CommRingEquiv R S ≃ (R ≡ S) +CommRingPath = ∫ 𝒮ᴰ-CommRing .UARel.ua + +uaCommRing : {A B : CommRing ℓ} → CommRingEquiv A B → A ≡ B +uaCommRing {A = A} {B = B} = equivFun (CommRingPath A B) + +CommRingIso : (R : CommRing ℓ) (S : CommRing ℓ') → Type (ℓ-max ℓ ℓ') +CommRingIso R S = Σ[ e ∈ Iso (R .fst) (S .fst) ] + IsRingHom (CommRingStr→RingStr (R .snd)) (e .fun) (CommRingStr→RingStr (S .snd)) + +CommRingEquivIsoCommRingIso : (R : CommRing ℓ) (S : CommRing ℓ') → Iso (CommRingEquiv R S) (CommRingIso R S) +fst (fun (CommRingEquivIsoCommRingIso R S) e) = equivToIso (e .fst) +snd (fun (CommRingEquivIsoCommRingIso R S) e) = e .snd +fst (inv (CommRingEquivIsoCommRingIso R S) e) = isoToEquiv (e .fst) +snd (inv (CommRingEquivIsoCommRingIso R S) e) = e .snd +rightInv (CommRingEquivIsoCommRingIso R S) (e , he) = + Σ≡Prop (λ e → isPropIsRingHom (snd (CommRing→Ring R)) (e .fun) (snd (CommRing→Ring S))) + rem + where + rem : equivToIso (isoToEquiv e) ≡ e + fun (rem i) x = fun e x + inv (rem i) x = inv e x + rightInv (rem i) b j = CommRingStr.is-set (snd S) (fun e (inv e b)) b (rightInv e b) (rightInv e b) i j + leftInv (rem i) a j = CommRingStr.is-set (snd R) (inv e (fun e a)) a (retEq (isoToEquiv e) a) (leftInv e a) i j +leftInv (CommRingEquivIsoCommRingIso R S) e = + Σ≡Prop (λ e → isPropIsRingHom (snd (CommRing→Ring R)) (e .fst) (snd (CommRing→Ring S))) + (equivEq refl) + +isGroupoidCommRing : isGroupoid (CommRing ℓ) +isGroupoidCommRing _ _ = isOfHLevelRespectEquiv 2 (CommRingPath _ _) (isSetRingEquiv _ _) + + +open CommRingStr +open IsRingHom +-- TODO: Induced structure results are temporarily inconvenient while we transition between algebra +-- representations +module _ (R : CommRing ℓ) {A : Type ℓ} + (0a 1a : A) + (add mul : A → A → A) + (inv : A → A) + (e : ⟨ R ⟩ ≃ A) + (p0 : e .fst (R .snd .0r) ≡ 0a) + (p1 : e .fst (R .snd .1r) ≡ 1a) + (p+ : ∀ x y → e .fst (R .snd ._+_ x y) ≡ add (e .fst x) (e .fst y)) + (p· : ∀ x y → e .fst (R .snd ._·_ x y) ≡ mul (e .fst x) (e .fst y)) + (pinv : ∀ x → e .fst (R .snd .-_ x) ≡ inv (e .fst x)) + where + + private + module R = CommRingStr (R .snd) + + BaseΣ : Type (ℓ-suc ℓ) + BaseΣ = Σ[ B ∈ Type ℓ ] B × B × (B → B → B) × (B → B → B) × (B → B) + + FamilyΣ : BaseΣ → Type ℓ + FamilyΣ (B , u0 , u1 , a , m , i) = IsCommRing u0 u1 a m i + + inducedΣ : FamilyΣ (A , 0a , 1a , add , mul , inv) + inducedΣ = + subst FamilyΣ + (UARel.≅→≡ (autoUARel BaseΣ) (e , p0 , p1 , p+ , p· , pinv)) + R.isCommRing + + InducedCommRing : CommRing ℓ + InducedCommRing .fst = A + 0r (InducedCommRing .snd) = 0a + 1r (InducedCommRing .snd) = 1a + _+_ (InducedCommRing .snd) = add + _·_ (InducedCommRing .snd) = mul + - InducedCommRing .snd = inv + isCommRing (InducedCommRing .snd) = inducedΣ + + InducedCommRingEquiv : CommRingEquiv R InducedCommRing + fst InducedCommRingEquiv = e + pres0 (snd InducedCommRingEquiv) = p0 + pres1 (snd InducedCommRingEquiv) = p1 + pres+ (snd InducedCommRingEquiv) = p+ + pres· (snd InducedCommRingEquiv) = p· + pres- (snd InducedCommRingEquiv) = pinv + + InducedCommRingPath : R ≡ InducedCommRing + InducedCommRingPath = CommRingPath _ _ .fst InducedCommRingEquiv From 905410353d9b97166019950c846fd6d952e99001 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 29 Jul 2024 17:59:09 +0200 Subject: [PATCH 02/83] start... --- Cubical/Algebra/CommRing/Base.agda | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/Cubical/Algebra/CommRing/Base.agda b/Cubical/Algebra/CommRing/Base.agda index bedb3fd91b..b8b6a4e495 100644 --- a/Cubical/Algebra/CommRing/Base.agda +++ b/Cubical/Algebra/CommRing/Base.agda @@ -109,6 +109,22 @@ CommRingStr.- snd (Ring→CommRing R p) = RingStr.-_ (snd R) IsCommRing.isRing (CommRingStr.isCommRing (snd (Ring→CommRing R p))) = RingStr.isRing (snd R) IsCommRing.·Comm (CommRingStr.isCommRing (snd (Ring→CommRing R p))) = p +record IsCommRingHom {A : Type ℓ} {B : Type ℓ'} (R : CommRingStr A) (f : A → B) (S : CommRingStr B) + : Type (ℓ-max ℓ ℓ') + where + + -- Shorter qualified names + private + module R = CommRingStr R + module S = CommRingStr S + + field + pres0 : f R.0r ≡ S.0r + pres1 : f R.1r ≡ S.1r + pres+ : (x y : A) → f (x R.+ y) ≡ f x S.+ f y + pres· : (x y : A) → f (x R.· y) ≡ f x S.· f y + pres- : (x : A) → f (R.- x) ≡ S.- (f x) + CommRingHom : (R : CommRing ℓ) (S : CommRing ℓ') → Type (ℓ-max ℓ ℓ') CommRingHom R S = RingHom (CommRing→Ring R) (CommRing→Ring S) From 6bd61760bcc2413dbbf4bce11410f5a0ef1400a0 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Tue, 30 Jul 2024 11:18:26 +0200 Subject: [PATCH 03/83] boilerplate --- Cubical/Algebra/CommRing/Base.agda | 86 +++++++++++++++++++++++- Cubical/Algebra/CommRing/Kernel.agda | 4 +- Cubical/Algebra/CommRing/Properties.agda | 41 +++++------ Cubical/Algebra/CommRing/Univalence.agda | 13 ++-- 4 files changed, 112 insertions(+), 32 deletions(-) diff --git a/Cubical/Algebra/CommRing/Base.agda b/Cubical/Algebra/CommRing/Base.agda index b8b6a4e495..5de6caca0e 100644 --- a/Cubical/Algebra/CommRing/Base.agda +++ b/Cubical/Algebra/CommRing/Base.agda @@ -113,7 +113,6 @@ record IsCommRingHom {A : Type ℓ} {B : Type ℓ'} (R : CommRingStr A) (f : A : Type (ℓ-max ℓ ℓ') where - -- Shorter qualified names private module R = CommRingStr R module S = CommRingStr S @@ -125,12 +124,14 @@ record IsCommRingHom {A : Type ℓ} {B : Type ℓ'} (R : CommRingStr A) (f : A pres· : (x y : A) → f (x R.· y) ≡ f x S.· f y pres- : (x : A) → f (R.- x) ≡ S.- (f x) +unquoteDecl IsCommRingHomIsoΣ = declareRecordIsoΣ IsCommRingHomIsoΣ (quote IsCommRingHom) + CommRingHom : (R : CommRing ℓ) (S : CommRing ℓ') → Type (ℓ-max ℓ ℓ') -CommRingHom R S = RingHom (CommRing→Ring R) (CommRing→Ring S) +CommRingHom R S = Σ[ f ∈ (⟨ R ⟩ → ⟨ S ⟩) ] IsCommRingHom (R .snd) f (S .snd) IsCommRingEquiv : {A : Type ℓ} {B : Type ℓ'} (R : CommRingStr A) (e : A ≃ B) (S : CommRingStr B) → Type (ℓ-max ℓ ℓ') -IsCommRingEquiv R e S = IsRingHom (CommRingStr→RingStr R) (e .fst) (CommRingStr→RingStr S) +IsCommRingEquiv R e S = IsCommRingHom R (e .fst) S CommRingEquiv : (R : CommRing ℓ) (S : CommRing ℓ') → Type (ℓ-max ℓ ℓ') CommRingEquiv R S = Σ[ e ∈ (R .fst ≃ S .fst) ] IsCommRingEquiv (R .snd) e (S .snd) @@ -146,3 +147,82 @@ isPropIsCommRing 0r 1r _+_ _·_ -_ = (λ ring → isPropΠ2 (λ _ _ → is-set ring _ _))) where open IsRing + +isPropIsCommRingHom : {A : Type ℓ} {B : Type ℓ'} (R : CommRingStr A) (f : A → B) (S : CommRingStr B) + → isProp (IsCommRingHom R f S) +isPropIsCommRingHom R f S = isOfHLevelRetractFromIso 1 IsCommRingHomIsoΣ + (isProp×4 (is-set _ _) + (is-set _ _) + (isPropΠ2 λ _ _ → is-set _ _) + (isPropΠ2 λ _ _ → is-set _ _) + (isPropΠ λ _ → is-set _ _)) + where + open CommRingStr S using (is-set) + +isSetCommRingHom : (R : CommRing ℓ) (S : CommRing ℓ') → isSet (CommRingHom R S) +isSetCommRingHom R S = isSetΣSndProp (isSetΠ λ _ → is-set) (λ f → isPropIsCommRingHom (snd R) f (snd S)) + where + open CommRingStr (str S) using (is-set) + +isSetCommRingEquiv : (A : CommRing ℓ) (B : CommRing ℓ') → isSet (CommRingEquiv A B) +isSetCommRingEquiv A B = isSetΣSndProp (isOfHLevel≃ 2 A.is-set B.is-set) + (λ e → isPropIsCommRingHom (snd A) (fst e) (snd B)) + where + module A = CommRingStr (str A) + module B = CommRingStr (str B) + +RingHom→CommRingHom : {R : CommRing ℓ} {S : CommRing ℓ'} + → RingHom (CommRing→Ring R) (CommRing→Ring S) + → CommRingHom R S +RingHom→CommRingHom f .fst = f .fst +RingHom→CommRingHom {R = R} {S = S} f .snd = copy + where open IsCommRingHom + copy : IsCommRingHom (R .snd) (f .fst) (S .snd) + copy .pres0 = f .snd .IsRingHom.pres0 + copy .pres1 = f .snd .IsRingHom.pres1 + copy .pres+ = f .snd .IsRingHom.pres+ + copy .pres· = f .snd .IsRingHom.pres· + copy .pres- = f .snd .IsRingHom.pres- + +IsRingHom→IsCommRingHom : (R : CommRing ℓ) (S : CommRing ℓ') + → (f : ⟨ R ⟩ → ⟨ S ⟩) + → IsRingHom ((CommRing→Ring R) .snd) f ((CommRing→Ring S) .snd) + → IsCommRingHom (R .snd) f (S .snd) +IsRingHom→IsCommRingHom R S f p = RingHom→CommRingHom (f , p) .snd + +CommRingHom→RingHom : {R : CommRing ℓ} {S : CommRing ℓ'} + → CommRingHom R S + → RingHom (CommRing→Ring R) (CommRing→Ring S) +CommRingHom→RingHom f .fst = f .fst +CommRingHom→RingHom {R = R} {S = S} f .snd = copy + where open IsRingHom + copy : IsRingHom ((CommRing→Ring R) .snd) (f .fst) ((CommRing→Ring S) .snd) + copy .pres0 = f .snd .IsCommRingHom.pres0 + copy .pres1 = f .snd .IsCommRingHom.pres1 + copy .pres+ = f .snd .IsCommRingHom.pres+ + copy .pres· = f .snd .IsCommRingHom.pres· + copy .pres- = f .snd .IsCommRingHom.pres- + +IsCommRingHom→IsRingHom : (R : CommRing ℓ) (S : CommRing ℓ') + → (f : ⟨ R ⟩ → ⟨ S ⟩) + → IsCommRingHom (R .snd) f (S .snd) + → IsRingHom ((CommRing→Ring R) .snd) f ((CommRing→Ring S) .snd) +IsCommRingHom→IsRingHom R S f p = CommRingHom→RingHom (f , p) .snd + +module _ {R : CommRing ℓ} {S : CommRing ℓ'} {f : ⟨ R ⟩ → ⟨ S ⟩} where + + private + module R = CommRingStr (R .snd) + module S = CommRingStr (S .snd) + + module _ + (p1 : f R.1r ≡ S.1r) + (p+ : (x y : ⟨ R ⟩) → f (x R.+ y) ≡ f x S.+ f y) + (p· : (x y : ⟨ R ⟩) → f (x R.· y) ≡ f x S.· f y) + where + + makeIsCommRingHom : IsCommRingHom (R .snd) f (S .snd) + makeIsCommRingHom = IsRingHom→IsCommRingHom _ _ _ (makeIsRingHom p1 p+ p·) + +_$cr_ : {R : CommRing ℓ} {S : CommRing ℓ'} → (φ : CommRingHom R S) → (x : ⟨ R ⟩) → ⟨ S ⟩ +φ $cr x = φ .fst x diff --git a/Cubical/Algebra/CommRing/Kernel.agda b/Cubical/Algebra/CommRing/Kernel.agda index c9ae8d7b3e..f585c91875 100644 --- a/Cubical/Algebra/CommRing/Kernel.agda +++ b/Cubical/Algebra/CommRing/Kernel.agda @@ -19,7 +19,7 @@ module _ (R S : CommRing ℓ) (f : CommRingHom R S) where -- If R and S were implicit, their ·Comm component could (almost?) never be inferred. kernelIdeal : IdealsIn R - kernelIdeal = Ideal→CommIdeal (ringKernelIdeal f) + kernelIdeal = Ideal→CommIdeal (ringKernelIdeal (CommRingHom→RingHom f)) kernelFiber : (x y : ⟨ R ⟩) → fst f x ≡ fst f y → (x - y) ∈ fst kernelIdeal - kernelFiber x y p = ringKernelFiber f x y p + kernelFiber x y p = ringKernelFiber (CommRingHom→RingHom f) x y p diff --git a/Cubical/Algebra/CommRing/Properties.agda b/Cubical/Algebra/CommRing/Properties.agda index 14900652e1..9244d3d9b0 100644 --- a/Cubical/Algebra/CommRing/Properties.agda +++ b/Cubical/Algebra/CommRing/Properties.agda @@ -26,6 +26,7 @@ open import Cubical.Algebra.Monoid open import Cubical.Algebra.AbGroup open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing.Base +open import Cubical.Algebra.CommRing.Univalence open import Cubical.HITs.PropositionalTruncation @@ -153,11 +154,13 @@ module _ where open RingHoms idCommRingHom : (R : CommRing ℓ) → CommRingHom R R - idCommRingHom R = idRingHom (CommRing→Ring R) + idCommRingHom R = RingHom→CommRingHom (idRingHom (CommRing→Ring R)) compCommRingHom : (R : CommRing ℓ) (S : CommRing ℓ') (T : CommRing ℓ'') → CommRingHom R S → CommRingHom S T → CommRingHom R T - compCommRingHom R S T = compRingHom {R = CommRing→Ring R} {CommRing→Ring S} {CommRing→Ring T} + compCommRingHom R S T f g = + RingHom→CommRingHom + (compRingHom (CommRingHom→RingHom f) (CommRingHom→RingHom g)) _∘cr_ : {R : CommRing ℓ} {S : CommRing ℓ'} {T : CommRing ℓ''} → CommRingHom S T → CommRingHom R S → CommRingHom R T @@ -165,17 +168,17 @@ module _ where compIdCommRingHom : {R S : CommRing ℓ} (f : CommRingHom R S) → compCommRingHom _ _ _ (idCommRingHom R) f ≡ f - compIdCommRingHom = compIdRingHom + compIdCommRingHom f = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) refl idCompCommRingHom : {R S : CommRing ℓ} (f : CommRingHom R S) → compCommRingHom _ _ _ f (idCommRingHom S) ≡ f - idCompCommRingHom = idCompRingHom + idCompCommRingHom f = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) refl compAssocCommRingHom : {R S T U : CommRing ℓ} (f : CommRingHom R S) (g : CommRingHom S T) (h : CommRingHom T U) → compCommRingHom _ _ _ (compCommRingHom _ _ _ f g) h ≡ compCommRingHom _ _ _ f (compCommRingHom _ _ _ g h) - compAssocCommRingHom = compAssocRingHom + compAssocCommRingHom f g h = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) refl open Iso @@ -184,22 +187,21 @@ module _ where injCommRingIso f x y h = sym (f .fst .leftInv x) ∙∙ cong (f .fst .inv) h ∙∙ f .fst .leftInv y module CommRingEquivs where - open RingEquivs + open RingEquivs - compCommRingEquiv : {A : CommRing ℓ} {B : CommRing ℓ'} {C : CommRing ℓ''} - → CommRingEquiv A B → CommRingEquiv B C → CommRingEquiv A C - compCommRingEquiv {A = A} {B = B} {C = C} = compRingEquiv {A = CommRing→Ring A} - {B = CommRing→Ring B} - {C = CommRing→Ring C} + compCommRingEquiv : {A : CommRing ℓ} {B : CommRing ℓ'} {C : CommRing ℓ''} + → CommRingEquiv A B → CommRingEquiv B C → CommRingEquiv A C + compCommRingEquiv f g .fst = compEquiv (f .fst) (g .fst) + compCommRingEquiv f g .snd = compCommRingHom _ _ _ (f .fst .fst , f .snd) (g .fst .fst , g .snd) .snd - invCommRingEquiv : (A : CommRing ℓ) → (B : CommRing ℓ') → CommRingEquiv A B → CommRingEquiv B A - fst (invCommRingEquiv A B e) = invEquiv (fst e) - snd (invCommRingEquiv A B e) = isRingHomInv e - - idCommRingEquiv : (A : CommRing ℓ) → CommRingEquiv A A - fst (idCommRingEquiv A) = idEquiv (fst A) - snd (idCommRingEquiv A) = makeIsRingHom refl (λ _ _ → refl) (λ _ _ → refl) + invCommRingEquiv : (A : CommRing ℓ) → (B : CommRing ℓ') → CommRingEquiv A B → CommRingEquiv B A + fst (invCommRingEquiv A B e) = invEquiv (fst e) + snd (invCommRingEquiv A B e) = + IsRingHom→IsCommRingHom _ _ _ $ isRingHomInv (e .fst , CommRingHom→RingHom (e .fst .fst , e .snd) .snd) + idCommRingEquiv : (A : CommRing ℓ) → CommRingEquiv A A + fst (idCommRingEquiv A) = idEquiv (fst A) + snd (idCommRingEquiv A) = makeIsCommRingHom refl (λ _ _ → refl) (λ _ _ → refl) module Exponentiation (R' : CommRing ℓ) where open CommRingStr (snd R') @@ -247,7 +249,6 @@ module Exponentiation (R' : CommRing ℓ) where ^-presUnit f zero f∈Rˣ = RˣContainsOne ^-presUnit f (suc n) f∈Rˣ = RˣMultClosed f (f ^ n) ⦃ f∈Rˣ ⦄ ⦃ ^-presUnit f n f∈Rˣ ⦄ - module CommRingHomTheory {A' B' : CommRing ℓ} (φ : CommRingHom A' B') where open Units A' renaming (Rˣ to Aˣ ; _⁻¹ to _⁻¹ᵃ ; ·-rinv to ·A-rinv ; ·-linv to ·A-linv) open Units B' renaming (Rˣ to Bˣ ; _⁻¹ to _⁻¹ᵇ ; ·-rinv to ·B-rinv) @@ -260,7 +261,7 @@ module CommRingHomTheory {A' B' : CommRing ℓ} (φ : CommRingHom A' B') where instance _ = A' .snd _ = B' .snd - open IsRingHom (φ .snd) + open IsCommRingHom (φ .snd) RingHomRespInv : (r : A) ⦃ r∈Aˣ : r ∈ Aˣ ⦄ → f r ∈ Bˣ RingHomRespInv r = f (r ⁻¹ᵃ) , (sym (pres· r (r ⁻¹ᵃ)) ∙∙ cong (f) (·A-rinv r) ∙∙ pres1) diff --git a/Cubical/Algebra/CommRing/Univalence.agda b/Cubical/Algebra/CommRing/Univalence.agda index a5cc63b970..5fa0972963 100644 --- a/Cubical/Algebra/CommRing/Univalence.agda +++ b/Cubical/Algebra/CommRing/Univalence.agda @@ -16,7 +16,6 @@ open import Cubical.Displayed.Universe open import Cubical.Algebra.AbGroup open import Cubical.Algebra.CommRing.Base -open import Cubical.Algebra.Ring.Base open import Cubical.Reflection.RecordEquiv @@ -39,7 +38,7 @@ open Iso prop[ isCommRing ∣ (λ _ _ → isPropIsCommRing _ _ _ _ _) ]) where open CommRingStr - open IsRingHom + open IsCommRingHom -- faster with some sharing null = autoDUARel (𝒮-Univ _) (λ A → A) @@ -53,7 +52,7 @@ uaCommRing {A = A} {B = B} = equivFun (CommRingPath A B) CommRingIso : (R : CommRing ℓ) (S : CommRing ℓ') → Type (ℓ-max ℓ ℓ') CommRingIso R S = Σ[ e ∈ Iso (R .fst) (S .fst) ] - IsRingHom (CommRingStr→RingStr (R .snd)) (e .fun) (CommRingStr→RingStr (S .snd)) + IsCommRingHom (R .snd) (e .fun) (S .snd) CommRingEquivIsoCommRingIso : (R : CommRing ℓ) (S : CommRing ℓ') → Iso (CommRingEquiv R S) (CommRingIso R S) fst (fun (CommRingEquivIsoCommRingIso R S) e) = equivToIso (e .fst) @@ -61,7 +60,7 @@ snd (fun (CommRingEquivIsoCommRingIso R S) e) = e .snd fst (inv (CommRingEquivIsoCommRingIso R S) e) = isoToEquiv (e .fst) snd (inv (CommRingEquivIsoCommRingIso R S) e) = e .snd rightInv (CommRingEquivIsoCommRingIso R S) (e , he) = - Σ≡Prop (λ e → isPropIsRingHom (snd (CommRing→Ring R)) (e .fun) (snd (CommRing→Ring S))) + Σ≡Prop (λ e → isPropIsCommRingHom (snd R) (e .fun) (snd S)) rem where rem : equivToIso (isoToEquiv e) ≡ e @@ -70,15 +69,15 @@ rightInv (CommRingEquivIsoCommRingIso R S) (e , he) = rightInv (rem i) b j = CommRingStr.is-set (snd S) (fun e (inv e b)) b (rightInv e b) (rightInv e b) i j leftInv (rem i) a j = CommRingStr.is-set (snd R) (inv e (fun e a)) a (retEq (isoToEquiv e) a) (leftInv e a) i j leftInv (CommRingEquivIsoCommRingIso R S) e = - Σ≡Prop (λ e → isPropIsRingHom (snd (CommRing→Ring R)) (e .fst) (snd (CommRing→Ring S))) + Σ≡Prop (λ e → isPropIsCommRingHom (snd R) (e .fst) (snd S)) (equivEq refl) isGroupoidCommRing : isGroupoid (CommRing ℓ) -isGroupoidCommRing _ _ = isOfHLevelRespectEquiv 2 (CommRingPath _ _) (isSetRingEquiv _ _) +isGroupoidCommRing _ _ = isOfHLevelRespectEquiv 2 (CommRingPath _ _) (isSetCommRingEquiv _ _) open CommRingStr -open IsRingHom +open IsCommRingHom -- TODO: Induced structure results are temporarily inconvenient while we transition between algebra -- representations module _ (R : CommRing ℓ) {A : Type ℓ} From acde70a2870d0286d93dc326332df26e76755c35 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Tue, 14 May 2024 22:53:00 +0200 Subject: [PATCH 04/83] fix auto-univalence --- Cubical/Algebra/CommAlgebra/Base.agda | 49 ++++++++++++++++++++------- 1 file changed, 37 insertions(+), 12 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/Base.agda index 4022db273f..e7638c9f6d 100644 --- a/Cubical/Algebra/CommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/Base.agda @@ -29,8 +29,7 @@ record IsCommAlgebra (R : CommRing ℓ) {A : Type ℓ'} (0a : A) (1a : A) (_+_ : A → A → A) (_·_ : A → A → A) (-_ : A → A) (_⋆_ : ⟨ R ⟩ → A → A) : Type (ℓ-max ℓ ℓ') where - - constructor iscommalgebra + no-eta-equality field isAlgebra : IsAlgebra (CommRing→Ring R) 0a 1a _+_ _·_ -_ _⋆_ @@ -41,8 +40,7 @@ record IsCommAlgebra (R : CommRing ℓ) {A : Type ℓ'} unquoteDecl IsCommAlgebraIsoΣ = declareRecordIsoΣ IsCommAlgebraIsoΣ (quote IsCommAlgebra) record CommAlgebraStr (R : CommRing ℓ) (A : Type ℓ') : Type (ℓ-max ℓ ℓ') where - - constructor commalgebrastr + no-eta-equality field 0a : A @@ -67,15 +65,34 @@ module _ {R : CommRing ℓ} where open CommRingStr (snd R) using (1r) renaming (_+_ to _+r_; _·_ to _·s_) CommAlgebraStr→AlgebraStr : {A : Type ℓ'} → CommAlgebraStr R A → AlgebraStr (CommRing→Ring R) A - CommAlgebraStr→AlgebraStr (commalgebrastr _ _ _ _ _ _ (iscommalgebra isAlgebra ·-comm)) = - algebrastr _ _ _ _ _ _ isAlgebra + CommAlgebraStr→AlgebraStr {A = A} cstr = x + where open AlgebraStr + x : AlgebraStr (CommRing→Ring R) A + 0a x = _ + 1a x = _ + _+_ x = _ + _·_ x = _ + - x = _ + _⋆_ x = _ + isAlgebra x = IsCommAlgebra.isAlgebra (CommAlgebraStr.isCommAlgebra cstr) CommAlgebra→Algebra : (A : CommAlgebra R ℓ') → Algebra (CommRing→Ring R) ℓ' - CommAlgebra→Algebra (_ , str) = (_ , CommAlgebraStr→AlgebraStr str) + fst (CommAlgebra→Algebra A) = fst A + snd (CommAlgebra→Algebra A) = CommAlgebraStr→AlgebraStr (snd A) CommAlgebra→CommRing : (A : CommAlgebra R ℓ') → CommRing ℓ' - CommAlgebra→CommRing (_ , commalgebrastr _ _ _ _ _ _ (iscommalgebra isAlgebra ·-comm)) = - _ , commringstr _ _ _ _ _ (iscommring (IsAlgebra.isRing isAlgebra) ·-comm) + CommAlgebra→CommRing (A , str) = x + where open CommRingStr + open CommAlgebraStr + x : CommRing _ + fst x = A + 0r (snd x) = _ + 1r (snd x) = _ + _+_ (snd x) = _ + _·_ (snd x) = _ + - snd x = _ + IsCommRing.isRing (isCommRing (snd x)) = RingStr.isRing (Algebra→Ring (_ , CommAlgebraStr→AlgebraStr str) .snd) + IsCommRing.·Comm (isCommRing (snd x)) = CommAlgebraStr.·Comm str module _ {A : Type ℓ'} {0a 1a : A} @@ -253,7 +270,7 @@ isPropIsCommAlgebra R _ _ _ _ _ _ = (λ alg → isPropΠ2 λ _ _ → alg .IsAlgebra.is-set _ _)) 𝒮ᴰ-CommAlgebra : (R : CommRing ℓ) → DUARel (𝒮-Univ ℓ') (CommAlgebraStr R) (ℓ-max ℓ ℓ') -𝒮ᴰ-CommAlgebra R = +𝒮ᴰ-CommAlgebra {ℓ' = ℓ'} R = 𝒮ᴰ-Record (𝒮-Univ _) (IsCommAlgebraEquiv {R = R}) (fields: data[ 0a ∣ nul ∣ pres0 ] @@ -262,7 +279,16 @@ isPropIsCommAlgebra R _ _ _ _ _ _ = data[ _·_ ∣ bin ∣ pres· ] data[ -_ ∣ autoDUARel _ _ ∣ pres- ] data[ _⋆_ ∣ autoDUARel _ _ ∣ pres⋆ ] - prop[ isCommAlgebra ∣ (λ _ _ → isPropIsCommAlgebra _ _ _ _ _ _ _) ]) + prop[ isCommAlgebra ∣ (λ A ΣA + → isPropIsCommAlgebra + {ℓ' = ℓ'} + R + (snd (fst (fst (fst (fst (fst ΣA)))))) + (snd (fst (fst (fst (fst ΣA))))) + (snd (fst (fst (fst ΣA)))) + (snd (fst (fst ΣA))) + (snd (fst ΣA)) + (snd ΣA)) ]) where open CommAlgebraStr open IsAlgebraHom @@ -279,4 +305,3 @@ uaCommAlgebra {R = R} {A = A} {B = B} = equivFun (CommAlgebraPath R A B) isGroupoidCommAlgebra : {R : CommRing ℓ} → isGroupoid (CommAlgebra R ℓ') isGroupoidCommAlgebra A B = isOfHLevelRespectEquiv 2 (CommAlgebraPath _ _ _) (isSetAlgebraEquiv _ _) --- -} From 2d9b18f139c713a87431e74290c163d44a644496 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Wed, 29 May 2024 09:49:44 +0200 Subject: [PATCH 05/83] fix FreeCommAlg --- Cubical/Algebra/CommAlgebra/Base.agda | 16 +++++++++++++ .../CommAlgebra/FreeCommAlgebra/Base.agda | 6 ++--- Cubical/Algebra/CommAlgebra/Properties.agda | 24 +++++++++++++------ 3 files changed, 36 insertions(+), 10 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/Base.agda index e7638c9f6d..8a26c5937a 100644 --- a/Cubical/Algebra/CommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/Base.agda @@ -141,6 +141,22 @@ module _ {R : CommRing ℓ} where x · (r ⋆ y) ∎ makeIsCommAlgebra .IsCommAlgebra.·Comm = ·Comm + makeCommAlgebraStr : + (A : Type ℓ') (0a 1a : A) + (_+_ _·_ : A → A → A) ( -_ : A → A) (_⋆_ : ⟨ R ⟩ → A → A) + (isCommAlg : IsCommAlgebra R 0a 1a _+_ _·_ -_ _⋆_) + → CommAlgebraStr R A + makeCommAlgebraStr A 0a 1a _+_ _·_ -_ _⋆_ isCommAlg = + record + { 0a = 0a + ; 1a = 1a + ; _+_ = _+_ + ; _·_ = _·_ + ; -_ = -_ + ; _⋆_ = _⋆_ + ; isCommAlgebra = isCommAlg + } + module _ (S : CommRing ℓ') where open CommRingStr (snd S) renaming (1r to 1S) open CommRingStr (snd R) using () renaming (_·_ to _·R_; _+_ to _+R_; 1r to 1R) diff --git a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Base.agda index b66cd0433f..b5291c432a 100644 --- a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Base.agda @@ -91,6 +91,6 @@ module Construction (R : CommRing ℓ) where ⋆-assoc ⋆-rdist-+ ⋆-ldist-+ ·-lid ⋆-assoc-· _[_] : (R : CommRing ℓ) (I : Type ℓ') → CommAlgebra R (ℓ-max ℓ ℓ') -(R [ I ]) = R[ I ] , commalgebrastr 0a 1a _+_ _·_ -_ _⋆_ isCommAlgebra - where - open Construction R +fst (R [ I ]) = Construction.R[_] R I +snd (R [ I ]) = makeCommAlgebraStr _ _ _ _ _ _ _ isCommAlgebra + where open Construction R diff --git a/Cubical/Algebra/CommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebra/Properties.agda index fd75d7d0cd..42a7890c51 100644 --- a/Cubical/Algebra/CommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/Properties.agda @@ -234,7 +234,7 @@ module CommAlgebraEquivs {R : CommRing ℓ} where module CommAlgebraUAFunctoriality {R : CommRing ℓ} where open CommAlgebraStr open CommAlgebraEquivs - +{- CommAlgebra≡ : (A B : CommAlgebra R ℓ') → ( Σ[ p ∈ ⟨ A ⟩ ≡ ⟨ B ⟩ ] Σ[ q0 ∈ PathP (λ i → p i) (0a (snd A)) (0a (snd B)) ] @@ -247,16 +247,25 @@ module CommAlgebraUAFunctoriality {R : CommRing ℓ} where (isCommAlgebra (snd B))) ≃ (A ≡ B) CommAlgebra≡ A B = isoToEquiv theIso - where + where -- commalgebrastr (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i) (t i) open Iso theIso : Iso _ _ - fun theIso (p , q0 , q1 , r+ , r· , s- , s⋆ , t) i = p i - , commalgebrastr (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i) (t i) + fst (fun theIso (p , q0 , q1 , r+ , r· , s- , s⋆ , t) i) = p i + snd (fun theIso (p , q0 , q1 , r+ , r· , s- , s⋆ , t) i) = Astr + where + Astr : CommAlgebraStr R (p i) + 0a Astr = ? + 1a Astr = ? + _+_ Astr = ? + _·_ Astr = ? + - Astr = ? + _⋆_ Astr = ? + isCommAlgebra Astr = ? inv theIso x = cong ⟨_⟩ x , cong (0a ∘ snd) x , cong (1a ∘ snd) x , cong (_+_ ∘ snd) x , cong (_·_ ∘ snd) x , cong (-_ ∘ snd) x , cong (_⋆_ ∘ snd) x , cong (isCommAlgebra ∘ snd) x - rightInv theIso _ = refl - leftInv theIso _ = refl + rightInv theIso _ = ? + leftInv theIso _ = ? caracCommAlgebra≡ : {A B : CommAlgebra R ℓ'} (p q : A ≡ B) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q caracCommAlgebra≡ {A = A} {B = B} p q P = @@ -285,7 +294,6 @@ module CommAlgebraUAFunctoriality {R : CommRing ℓ} where ≡⟨ sym (cong-∙ ⟨_⟩ (uaCommAlgebra f) (uaCommAlgebra g)) ⟩ cong ⟨_⟩ (uaCommAlgebra f ∙ uaCommAlgebra g) ∎) - open CommAlgebraHoms open CommAlgebraEquivs open CommAlgebraUAFunctoriality @@ -298,7 +306,9 @@ recPT→CommAlgebra 𝓕 σ compCoh = GpdElim.rec→Gpd isGroupoidCommAlgebra (3-ConstantCompChar 𝓕 (λ x y → uaCommAlgebra (σ x y)) λ x y z → sym ( cong uaCommAlgebra (compCoh x y z) ∙ uaCompCommAlgebraEquiv (σ x y) (σ y z))) +-} +open CommAlgebraHoms contrCommAlgebraHom→contrCommAlgebraEquiv : {R : CommRing ℓ} {A : Type ℓ'} (σ : A → CommAlgebra R ℓ'') From 16ba911e67584b013850631dec05019df51917be Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 30 May 2024 16:51:42 +0200 Subject: [PATCH 06/83] fix subalgebra --- Cubical/Algebra/CommAlgebra/Subalgebra.agda | 18 +++++++++--------- 1 file changed, 9 insertions(+), 9 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/Subalgebra.agda b/Cubical/Algebra/CommAlgebra/Subalgebra.agda index 107a54f296..c4ca6f28ba 100644 --- a/Cubical/Algebra/CommAlgebra/Subalgebra.agda +++ b/Cubical/Algebra/CommAlgebra/Subalgebra.agda @@ -20,14 +20,15 @@ module _ (S : Subalgebra) where Subalgebra→CommAlgebra≡ = Subalgebra→Algebra≡ S Subalgebra→CommAlgebra : CommAlgebra R ℓ' - Subalgebra→CommAlgebra = - Subalgebra→Algebra S .fst - , record - { AlgebraStr (Subalgebra→Algebra S .snd) - ; isCommAlgebra = iscommalgebra - (Subalgebra→Algebra S .snd .AlgebraStr.isAlgebra) - (λ x y → Subalgebra→CommAlgebra≡ - (CommAlgebraStr.·Comm (snd A) (fst x) (fst y)))} + fst Subalgebra→CommAlgebra = Subalgebra→Algebra S .fst + snd Subalgebra→CommAlgebra = record + { AlgebraStr (Subalgebra→Algebra S .snd) + ; isCommAlgebra = record { + isAlgebra = + Subalgebra→Algebra S .snd .AlgebraStr.isAlgebra ; + ·Comm = λ x y → Subalgebra→CommAlgebra≡ + (CommAlgebraStr.·Comm (snd A) (fst x) (fst y)) } + } Subalgebra→CommAlgebraHom : CommAlgebraHom Subalgebra→CommAlgebra A Subalgebra→CommAlgebraHom = Subalgebra→AlgebraHom S @@ -41,4 +42,3 @@ module _ (S : Subalgebra) where (λ x y → Subalgebra→CommAlgebra≡ (pres+ x y)) (λ x y → Subalgebra→CommAlgebra≡ (pres· x y)) (λ x y → Subalgebra→CommAlgebra≡ (pres⋆ x y)) - From b1bac5cceb1e52a35976a8ec16f7985c5c327ef1 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 30 May 2024 16:57:24 +0200 Subject: [PATCH 07/83] fix FPAlgebra --- Cubical/Algebra/CommAlgebra/FPAlgebra/Base.agda | 2 +- Cubical/Algebra/CommAlgebra/FPAlgebra/Instances.agda | 4 ++-- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/FPAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/FPAlgebra/Base.agda index 229c8cb20a..322e06a5f9 100644 --- a/Cubical/Algebra/CommAlgebra/FPAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/FPAlgebra/Base.agda @@ -51,7 +51,7 @@ module _ {R : CommRing ℓ} where evPolyHomomorphic A B f P values = (fst f) (evPoly A P values) ≡⟨ refl ⟩ (fst f) (fst (freeInducedHom A values) P) ≡⟨ refl ⟩ - fst (f ∘a freeInducedHom A values) P ≡⟨ cong (λ u → fst u P) (natIndHomR f values) ⟩ + fst (f ∘a freeInducedHom A values) P ≡⟨ cong (λ u → fst u P) (natIndHomR {A = A} {B = B} f values) ⟩ fst (freeInducedHom B (fst f ∘ values)) P ≡⟨ refl ⟩ evPoly B P (fst f ∘ values) ∎ where open AlgebraHoms diff --git a/Cubical/Algebra/CommAlgebra/FPAlgebra/Instances.agda b/Cubical/Algebra/CommAlgebra/FPAlgebra/Instances.agda index 26e21cad66..24319606d8 100644 --- a/Cubical/Algebra/CommAlgebra/FPAlgebra/Instances.agda +++ b/Cubical/Algebra/CommAlgebra/FPAlgebra/Instances.agda @@ -78,13 +78,13 @@ module _ (R : CommRing ℓ) where inverse1 : fromA ∘a toA ≡ idAlgebraHom _ inverse1 = fromA ∘a toA - ≡⟨ sym (unique _ _ _ _ _ _ (λ i → cong (fst fromA) ( + ≡⟨ sym (unique _ _ B _ _ _ (λ i → cong (fst fromA) ( fst toA (generator n emptyGen i) ≡⟨ inducedHomOnGenerators _ _ _ _ _ _ ⟩ Construction.var i ∎))) ⟩ inducedHom n emptyGen B (generator _ _) (relationsHold _ _) - ≡⟨ unique _ _ _ _ _ _ (λ i → refl) ⟩ + ≡⟨ unique _ _ B _ _ _ (λ i → refl) ⟩ idAlgebraHom _ ∎ inverse2 : toA ∘a fromA ≡ idAlgebraHom _ From 487fe8da2d17e60bd349a0920417753cd5161b44 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Wed, 19 Jun 2024 00:12:30 +0200 Subject: [PATCH 08/83] copy Evans approach for Groups --- Cubical/Algebra/CommAlgebra/Base.agda | 2 + Cubical/Algebra/CommAlgebra/Properties.agda | 83 +++++++-------------- 2 files changed, 28 insertions(+), 57 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/Base.agda index 8a26c5937a..12a890574e 100644 --- a/Cubical/Algebra/CommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/Base.agda @@ -58,6 +58,8 @@ record CommAlgebraStr (R : CommRing ℓ) (A : Type ℓ') : Type (ℓ-max ℓ ℓ infixl 7 _⋆_ infixl 6 _+_ +unquoteDecl CommAlgebraStrIsoΣ = declareRecordIsoΣ CommAlgebraStrIsoΣ (quote CommAlgebraStr) + CommAlgebra : (R : CommRing ℓ) → ∀ ℓ' → Type (ℓ-max ℓ (ℓ-suc ℓ')) CommAlgebra R ℓ' = Σ[ A ∈ Type ℓ' ] CommAlgebraStr R A diff --git a/Cubical/Algebra/CommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebra/Properties.agda index 42a7890c51..38123c1107 100644 --- a/Cubical/Algebra/CommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/Properties.agda @@ -14,6 +14,8 @@ open import Cubical.Foundations.SIP open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Path +open import Cubical.Functions.Embedding + open import Cubical.Data.Sigma open import Cubical.Reflection.StrictEquiv @@ -221,69 +223,37 @@ module CommAlgebraHoms {R : CommRing ℓ} where compAssocCommAlgebraHom = compAssocAlgebraHom module CommAlgebraEquivs {R : CommRing ℓ} where - open AlgebraEquivs - - compCommAlgebraEquiv : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} - → CommAlgebraEquiv A B → CommAlgebraEquiv B C → CommAlgebraEquiv A C - compCommAlgebraEquiv {A = A} {B = B} {C = C} = compAlgebraEquiv {A = CommAlgebra→Algebra A} - {B = CommAlgebra→Algebra B} - {C = CommAlgebra→Algebra C} - + open AlgebraEquivs + + compCommAlgebraEquiv : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} + → CommAlgebraEquiv A B → CommAlgebraEquiv B C → CommAlgebraEquiv A C + compCommAlgebraEquiv {A = A} {B = B} {C = C} = compAlgebraEquiv {A = CommAlgebra→Algebra A} + {B = CommAlgebra→Algebra B} + {C = CommAlgebra→Algebra C} + + + isSetCommAlgStr : (A : Type ℓ') → isSet (CommAlgebraStr R A) + isSetCommAlgStr A = + let open CommAlgebraStr + in isOfHLevelSucIfInhabited→isOfHLevelSuc 1 (λ str → + isOfHLevelRetractFromIso 2 CommAlgebraStrIsoΣ $ + isSetΣ (str .is-set) λ _ → + isSetΣ (str .is-set) (λ _ → + isSetΣ (isSet→ (isSet→ (str .is-set))) λ _ → + isSetΣ (isSet→ (isSet→ (str .is-set))) (λ _ → + isSetΣ (isSet→ (str .is-set)) λ _ → + isSetΣSndProp (isSet→ (isSet→ (str .is-set))) (λ _ → + isPropIsCommAlgebra R _ _ _ _ _ _)))) -- the CommAlgebra version of uaCompEquiv module CommAlgebraUAFunctoriality {R : CommRing ℓ} where open CommAlgebraStr open CommAlgebraEquivs -{- - CommAlgebra≡ : (A B : CommAlgebra R ℓ') → ( - Σ[ p ∈ ⟨ A ⟩ ≡ ⟨ B ⟩ ] - Σ[ q0 ∈ PathP (λ i → p i) (0a (snd A)) (0a (snd B)) ] - Σ[ q1 ∈ PathP (λ i → p i) (1a (snd A)) (1a (snd B)) ] - Σ[ r+ ∈ PathP (λ i → p i → p i → p i) (_+_ (snd A)) (_+_ (snd B)) ] - Σ[ r· ∈ PathP (λ i → p i → p i → p i) (_·_ (snd A)) (_·_ (snd B)) ] - Σ[ s- ∈ PathP (λ i → p i → p i) (-_ (snd A)) (-_ (snd B)) ] - Σ[ s⋆ ∈ PathP (λ i → ⟨ R ⟩ → p i → p i) (_⋆_ (snd A)) (_⋆_ (snd B)) ] - PathP (λ i → IsCommAlgebra R (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i)) (isCommAlgebra (snd A)) - (isCommAlgebra (snd B))) - ≃ (A ≡ B) - CommAlgebra≡ A B = isoToEquiv theIso - where -- commalgebrastr (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i) (t i) - open Iso - theIso : Iso _ _ - fst (fun theIso (p , q0 , q1 , r+ , r· , s- , s⋆ , t) i) = p i - snd (fun theIso (p , q0 , q1 , r+ , r· , s- , s⋆ , t) i) = Astr - where - Astr : CommAlgebraStr R (p i) - 0a Astr = ? - 1a Astr = ? - _+_ Astr = ? - _·_ Astr = ? - - Astr = ? - _⋆_ Astr = ? - isCommAlgebra Astr = ? - inv theIso x = cong ⟨_⟩ x , cong (0a ∘ snd) x , cong (1a ∘ snd) x - , cong (_+_ ∘ snd) x , cong (_·_ ∘ snd) x , cong (-_ ∘ snd) x , cong (_⋆_ ∘ snd) x - , cong (isCommAlgebra ∘ snd) x - rightInv theIso _ = ? - leftInv theIso _ = ? caracCommAlgebra≡ : {A B : CommAlgebra R ℓ'} (p q : A ≡ B) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q - caracCommAlgebra≡ {A = A} {B = B} p q P = - sym (transportTransport⁻ (ua (CommAlgebra≡ A B)) p) - ∙∙ cong (transport (ua (CommAlgebra≡ A B))) helper - ∙∙ transportTransport⁻ (ua (CommAlgebra≡ A B)) q - where - helper : transport (sym (ua (CommAlgebra≡ A B))) p ≡ transport (sym (ua (CommAlgebra≡ A B))) q - helper = Σ≡Prop - (λ _ → isPropΣ - (isOfHLevelPathP' 1 (is-set (snd B)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (is-set (snd B)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ λ _ → is-set (snd B)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) - λ _ → isOfHLevelPathP 1 (isPropIsCommAlgebra _ _ _ _ _ _ _) _ _) - (transportRefl (cong ⟨_⟩ p) ∙ P ∙ sym (transportRefl (cong ⟨_⟩ q))) + caracCommAlgebra≡ _ _ α = + isEmbedding→Inj (iso→isEmbedding (invIso ΣPathIsoPathΣ)) _ _ $ + Σ≡Prop (λ _ → isOfHLevelPathP' 1 (isSetCommAlgStr _) _ _) α uaCompCommAlgebraEquiv : {A B C : CommAlgebra R ℓ'} (f : CommAlgebraEquiv A B) (g : CommAlgebraEquiv B C) → uaCommAlgebra (compCommAlgebraEquiv f g) ≡ uaCommAlgebra f ∙ uaCommAlgebra g @@ -306,7 +276,6 @@ recPT→CommAlgebra 𝓕 σ compCoh = GpdElim.rec→Gpd isGroupoidCommAlgebra (3-ConstantCompChar 𝓕 (λ x y → uaCommAlgebra (σ x y)) λ x y z → sym ( cong uaCommAlgebra (compCoh x y z) ∙ uaCompCommAlgebraEquiv (σ x y) (σ y z))) --} open CommAlgebraHoms From a33924c5d4f7cda65781fee6707ce32c64d7f21a Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Wed, 19 Jun 2024 14:44:06 +0200 Subject: [PATCH 09/83] refactor --- Cubical/Algebra/Algebra/Base.agda | 4 +- Cubical/Algebra/Algebra/Properties.agda | 69 +++++-------------------- Cubical/Algebra/Algebra/Univalence.agda | 45 ++++++++++++++++ 3 files changed, 60 insertions(+), 58 deletions(-) create mode 100644 Cubical/Algebra/Algebra/Univalence.agda diff --git a/Cubical/Algebra/Algebra/Base.agda b/Cubical/Algebra/Algebra/Base.agda index 97afcced0f..8ba29d9f5a 100644 --- a/Cubical/Algebra/Algebra/Base.agda +++ b/Cubical/Algebra/Algebra/Base.agda @@ -36,8 +36,6 @@ record IsAlgebra (R : Ring ℓ) {A : Type ℓ'} (0a 1a : A) (_+_ _·_ : A → A → A) (-_ : A → A) (_⋆_ : ⟨ R ⟩ → A → A) : Type (ℓ-max ℓ ℓ') where - constructor isalgebra - open RingStr (snd R) using (1r) renaming (_+_ to _+r_; _·_ to _·r_) field @@ -72,6 +70,8 @@ record AlgebraStr (R : Ring ℓ) (A : Type ℓ') : Type (ℓ-max ℓ ℓ') where open IsAlgebra isAlgebra public +unquoteDecl AlgebraStrIsoΣ = declareRecordIsoΣ AlgebraStrIsoΣ (quote AlgebraStr) + Algebra : (R : Ring ℓ) → ∀ ℓ' → Type (ℓ-max ℓ (ℓ-suc ℓ')) Algebra R ℓ' = Σ[ A ∈ Type ℓ' ] AlgebraStr R A diff --git a/Cubical/Algebra/Algebra/Properties.agda b/Cubical/Algebra/Algebra/Properties.agda index 0bffb60665..a384174696 100644 --- a/Cubical/Algebra/Algebra/Properties.agda +++ b/Cubical/Algebra/Algebra/Properties.agda @@ -10,8 +10,7 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Path open import Cubical.Foundations.Transport -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.SIP +open import Cubical.Foundations.SIP using (⟨_⟩) open import Cubical.Data.Sigma @@ -185,57 +184,15 @@ module AlgebraEquivs where (λ g → isPropIsAlgebraHom _ (B .snd) g (C .snd)) (isoOnTypes .leftInv (g .fst)) --- the Algebra version of uaCompEquiv -module AlgebraUAFunctoriality where - open AlgebraStr - open AlgebraEquivs - - Algebra≡ : (A B : Algebra R ℓ') → ( - Σ[ p ∈ ⟨ A ⟩ ≡ ⟨ B ⟩ ] - Σ[ q0 ∈ PathP (λ i → p i) (0a (snd A)) (0a (snd B)) ] - Σ[ q1 ∈ PathP (λ i → p i) (1a (snd A)) (1a (snd B)) ] - Σ[ r+ ∈ PathP (λ i → p i → p i → p i) (_+_ (snd A)) (_+_ (snd B)) ] - Σ[ r· ∈ PathP (λ i → p i → p i → p i) (_·_ (snd A)) (_·_ (snd B)) ] - Σ[ s- ∈ PathP (λ i → p i → p i) (-_ (snd A)) (-_ (snd B)) ] - Σ[ s⋆ ∈ PathP (λ i → ⟨ R ⟩ → p i → p i) (_⋆_ (snd A)) (_⋆_ (snd B)) ] - PathP (λ i → IsAlgebra R (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i)) (isAlgebra (snd A)) - (isAlgebra (snd B))) - ≃ (A ≡ B) - Algebra≡ A B = isoToEquiv theIso - where - open Iso - theIso : Iso _ _ - fun theIso (p , q0 , q1 , r+ , r· , s- , s⋆ , t) i = p i - , algebrastr (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i) (t i) - inv theIso x = cong ⟨_⟩ x , cong (0a ∘ snd) x , cong (1a ∘ snd) x - , cong (_+_ ∘ snd) x , cong (_·_ ∘ snd) x , cong (-_ ∘ snd) x , cong (_⋆_ ∘ snd) x - , cong (isAlgebra ∘ snd) x - rightInv theIso _ = refl - leftInv theIso _ = refl - - caracAlgebra≡ : (p q : A ≡ B) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q - caracAlgebra≡ {A = A} {B = B} p q P = - sym (transportTransport⁻ (ua (Algebra≡ A B)) p) - ∙∙ cong (transport (ua (Algebra≡ A B))) helper - ∙∙ transportTransport⁻ (ua (Algebra≡ A B)) q - where - helper : transport (sym (ua (Algebra≡ A B))) p ≡ transport (sym (ua (Algebra≡ A B))) q - helper = Σ≡Prop - (λ _ → isPropΣ - (isOfHLevelPathP' 1 (is-set (snd B)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (is-set (snd B)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ λ _ → is-set (snd B)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) - λ _ → isOfHLevelPathP 1 (isPropIsAlgebra _ _ _ _ _ _ _) _ _) - (transportRefl (cong ⟨_⟩ p) ∙ P ∙ sym (transportRefl (cong ⟨_⟩ q))) - - uaCompAlgebraEquiv : (f : AlgebraEquiv A B) (g : AlgebraEquiv B C) - → uaAlgebra (compAlgebraEquiv f g) ≡ uaAlgebra f ∙ uaAlgebra g - uaCompAlgebraEquiv f g = caracAlgebra≡ _ _ ( - cong ⟨_⟩ (uaAlgebra (compAlgebraEquiv f g)) - ≡⟨ uaCompEquiv _ _ ⟩ - cong ⟨_⟩ (uaAlgebra f) ∙ cong ⟨_⟩ (uaAlgebra g) - ≡⟨ sym (cong-∙ ⟨_⟩ (uaAlgebra f) (uaAlgebra g)) ⟩ - cong ⟨_⟩ (uaAlgebra f ∙ uaAlgebra g) ∎) +isSetAlgebraStr : (A : Type ℓ') → isSet (AlgebraStr R A) +isSetAlgebraStr A = + let open AlgebraStr + in isOfHLevelSucIfInhabited→isOfHLevelSuc 1 λ str → + isOfHLevelRetractFromIso 2 AlgebraStrIsoΣ $ + isSetΣ (str .is-set) λ _ → + isSetΣ (str .is-set) λ _ → + isSetΣ (isSet→ (isSet→ (str .is-set))) λ _ → + isSetΣ (isSet→ (isSet→ (str .is-set))) λ _ → + isSetΣ (isSet→ (str .is-set)) (λ _ → + isSetΣSndProp (isSet→ (isSet→ (str .is-set))) λ _ → + isPropIsAlgebra _ _ _ _ _ _ _) diff --git a/Cubical/Algebra/Algebra/Univalence.agda b/Cubical/Algebra/Algebra/Univalence.agda new file mode 100644 index 0000000000..feac7a461f --- /dev/null +++ b/Cubical/Algebra/Algebra/Univalence.agda @@ -0,0 +1,45 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Algebra.Univalence where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.SIP + +open import Cubical.Functions.Embedding + +open import Cubical.Data.Sigma + +open import Cubical.Algebra.Ring.Base +open import Cubical.Algebra.Algebra.Base +open import Cubical.Algebra.Algebra.Properties + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + R : Ring ℓ + A B C D : Algebra R ℓ + + +-- the Algebra version of uaCompEquiv +module AlgebraUAFunctoriality where + open AlgebraStr + open AlgebraEquivs + + caracAlgebra≡ : (p q : A ≡ B) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q + caracAlgebra≡ _ _ α = + isEmbedding→Inj (iso→isEmbedding (invIso ΣPathIsoPathΣ)) _ _ $ + Σ≡Prop (λ _ → isOfHLevelPathP' 1 (isSetAlgebraStr _) _ _) α + + uaCompAlgebraEquiv : (f : AlgebraEquiv A B) (g : AlgebraEquiv B C) + → uaAlgebra (compAlgebraEquiv f g) ≡ uaAlgebra f ∙ uaAlgebra g + uaCompAlgebraEquiv f g = caracAlgebra≡ _ _ ( + cong ⟨_⟩ (uaAlgebra (compAlgebraEquiv f g)) + ≡⟨ uaCompEquiv _ _ ⟩ + cong ⟨_⟩ (uaAlgebra f) ∙ cong ⟨_⟩ (uaAlgebra g) + ≡⟨ sym (cong-∙ ⟨_⟩ (uaAlgebra f) (uaAlgebra g)) ⟩ + cong ⟨_⟩ (uaAlgebra f ∙ uaAlgebra g) ∎) From ef58e50a4fd58a6ea5485e301a823a18bdef6cec Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Wed, 19 Jun 2024 15:07:39 +0200 Subject: [PATCH 10/83] refactor --- Cubical/Algebra/Algebra/Base.agda | 50 ++++++------------------- Cubical/Algebra/Algebra/Univalence.agda | 33 ++++++++++++++++ 2 files changed, 44 insertions(+), 39 deletions(-) diff --git a/Cubical/Algebra/Algebra/Base.agda b/Cubical/Algebra/Algebra/Base.agda index 8ba29d9f5a..047ed059aa 100644 --- a/Cubical/Algebra/Algebra/Base.agda +++ b/Cubical/Algebra/Algebra/Base.agda @@ -10,11 +10,6 @@ open import Cubical.Foundations.SIP open import Cubical.Data.Sigma -open import Cubical.Displayed.Base -open import Cubical.Displayed.Auto -open import Cubical.Displayed.Record -open import Cubical.Displayed.Universe - open import Cubical.Reflection.RecordEquiv open import Cubical.Algebra.Monoid @@ -36,6 +31,8 @@ record IsAlgebra (R : Ring ℓ) {A : Type ℓ'} (0a 1a : A) (_+_ _·_ : A → A → A) (-_ : A → A) (_⋆_ : ⟨ R ⟩ → A → A) : Type (ℓ-max ℓ ℓ') where + no-eta-equality + open RingStr (snd R) using (1r) renaming (_+_ to _+r_; _·_ to _·r_) field @@ -208,13 +205,15 @@ isPropIsAlgebra R _ _ _ _ _ _ = let open IsLeftModule in isPropIsAlgebraHom : (R : Ring ℓ) {A : Type ℓ'} {B : Type ℓ''} (AS : AlgebraStr R A) (f : A → B) (BS : AlgebraStr R B) → isProp (IsAlgebraHom AS f BS) -isPropIsAlgebraHom R AS f BS = isOfHLevelRetractFromIso 1 IsAlgebraHomIsoΣ - (isProp×5 (is-set _ _) - (is-set _ _) - (isPropΠ2 λ _ _ → is-set _ _) - (isPropΠ2 λ _ _ → is-set _ _) - (isPropΠ λ _ → is-set _ _) - (isPropΠ2 λ _ _ → is-set _ _)) +isPropIsAlgebraHom R AS f BS = + isOfHLevelRetractFromIso 1 + IsAlgebraHomIsoΣ + (isProp×5 (is-set _ _) + (is-set _ _) + (isPropΠ2 λ _ _ → is-set _ _) + (isPropΠ2 λ _ _ → is-set _ _) + (isPropΠ λ _ → is-set _ _) + (isPropΠ2 λ _ _ → is-set _ _)) where open AlgebraStr BS @@ -237,33 +236,6 @@ isSetAlgebraEquiv M N = isSetΣ (isOfHLevel≃ 2 M.is-set N.is-set) AlgebraHom≡ : {φ ψ : AlgebraHom A B} → fst φ ≡ fst ψ → φ ≡ ψ AlgebraHom≡ = Σ≡Prop λ f → isPropIsAlgebraHom _ _ f _ -𝒮ᴰ-Algebra : (R : Ring ℓ) → DUARel (𝒮-Univ ℓ') (AlgebraStr R) (ℓ-max ℓ ℓ') -𝒮ᴰ-Algebra R = - 𝒮ᴰ-Record (𝒮-Univ _) (IsAlgebraEquiv {R = R}) - (fields: - data[ 0a ∣ nul ∣ pres0 ] - data[ 1a ∣ nul ∣ pres1 ] - data[ _+_ ∣ bin ∣ pres+ ] - data[ _·_ ∣ bin ∣ pres· ] - data[ -_ ∣ autoDUARel _ _ ∣ pres- ] - data[ _⋆_ ∣ autoDUARel _ _ ∣ pres⋆ ] - prop[ isAlgebra ∣ (λ _ _ → isPropIsAlgebra _ _ _ _ _ _ _) ]) - where - open AlgebraStr - - -- faster with some sharing - nul = autoDUARel (𝒮-Univ _) (λ A → A) - bin = autoDUARel (𝒮-Univ _) (λ A → A → A → A) - -AlgebraPath : (A B : Algebra R ℓ') → (AlgebraEquiv A B) ≃ (A ≡ B) -AlgebraPath {R = R} = ∫ (𝒮ᴰ-Algebra R) .UARel.ua - -uaAlgebra : AlgebraEquiv A B → A ≡ B -uaAlgebra {A = A} {B = B} = equivFun (AlgebraPath A B) - -isGroupoidAlgebra : isGroupoid (Algebra R ℓ') -isGroupoidAlgebra _ _ = isOfHLevelRespectEquiv 2 (AlgebraPath _ _) (isSetAlgebraEquiv _ _) - -- Smart constructor for algebra homomorphisms -- that infers the other equations from pres1, pres+, pres·, and pres⋆ diff --git a/Cubical/Algebra/Algebra/Univalence.agda b/Cubical/Algebra/Algebra/Univalence.agda index feac7a461f..de22d30c39 100644 --- a/Cubical/Algebra/Algebra/Univalence.agda +++ b/Cubical/Algebra/Algebra/Univalence.agda @@ -14,6 +14,11 @@ open import Cubical.Functions.Embedding open import Cubical.Data.Sigma +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + open import Cubical.Algebra.Ring.Base open import Cubical.Algebra.Algebra.Base open import Cubical.Algebra.Algebra.Properties @@ -24,6 +29,34 @@ private R : Ring ℓ A B C D : Algebra R ℓ +open IsAlgebraHom + +𝒮ᴰ-Algebra : (R : Ring ℓ) → DUARel (𝒮-Univ ℓ') (AlgebraStr R) (ℓ-max ℓ ℓ') +𝒮ᴰ-Algebra R = + 𝒮ᴰ-Record (𝒮-Univ _) (IsAlgebraEquiv {R = R}) + (fields: + data[ 0a ∣ nul ∣ pres0 ] + data[ 1a ∣ nul ∣ pres1 ] + data[ _+_ ∣ bin ∣ pres+ ] + data[ _·_ ∣ bin ∣ pres· ] + data[ -_ ∣ autoDUARel _ _ ∣ pres- ] + data[ _⋆_ ∣ autoDUARel _ _ ∣ pres⋆ ] + prop[ isAlgebra ∣ (λ _ _ → isPropIsAlgebra _ _ _ _ _ _ _) ]) + where + open AlgebraStr + + -- faster with some sharing + nul = autoDUARel (𝒮-Univ _) (λ A → A) + bin = autoDUARel (𝒮-Univ _) (λ A → A → A → A) + +AlgebraPath : (A B : Algebra R ℓ') → (AlgebraEquiv A B) ≃ (A ≡ B) +AlgebraPath {R = R} = ∫ (𝒮ᴰ-Algebra R) .UARel.ua + +uaAlgebra : AlgebraEquiv A B → A ≡ B +uaAlgebra {A = A} {B = B} = equivFun (AlgebraPath A B) + +isGroupoidAlgebra : isGroupoid (Algebra R ℓ') +isGroupoidAlgebra _ _ = isOfHLevelRespectEquiv 2 (AlgebraPath _ _) (isSetAlgebraEquiv _ _) -- the Algebra version of uaCompEquiv module AlgebraUAFunctoriality where From 90bbfaccc4954dab6c88bc26df67052d653fd863 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Wed, 19 Jun 2024 15:30:58 +0200 Subject: [PATCH 11/83] no-eta for algebra structures --- Cubical/Algebra/Algebra/Base.agda | 2 +- Cubical/Algebra/Algebra/Univalence.agda | 10 +++++++++- Cubical/Algebra/CommAlgebra/Subalgebra.agda | 3 ++- 3 files changed, 12 insertions(+), 3 deletions(-) diff --git a/Cubical/Algebra/Algebra/Base.agda b/Cubical/Algebra/Algebra/Base.agda index 047ed059aa..21c74bce72 100644 --- a/Cubical/Algebra/Algebra/Base.agda +++ b/Cubical/Algebra/Algebra/Base.agda @@ -54,7 +54,7 @@ unquoteDecl IsAlgebraIsoΣ = declareRecordIsoΣ IsAlgebraIsoΣ (quote IsAlgebra) record AlgebraStr (R : Ring ℓ) (A : Type ℓ') : Type (ℓ-max ℓ ℓ') where - constructor algebrastr + no-eta-equality field 0a : A diff --git a/Cubical/Algebra/Algebra/Univalence.agda b/Cubical/Algebra/Algebra/Univalence.agda index de22d30c39..3394ba74f9 100644 --- a/Cubical/Algebra/Algebra/Univalence.agda +++ b/Cubical/Algebra/Algebra/Univalence.agda @@ -41,7 +41,15 @@ open IsAlgebraHom data[ _·_ ∣ bin ∣ pres· ] data[ -_ ∣ autoDUARel _ _ ∣ pres- ] data[ _⋆_ ∣ autoDUARel _ _ ∣ pres⋆ ] - prop[ isAlgebra ∣ (λ _ _ → isPropIsAlgebra _ _ _ _ _ _ _) ]) + prop[ isAlgebra ∣ (λ A ΣA → + isPropIsAlgebra + R + (snd (fst (fst (fst (fst (fst ΣA)))))) + (snd (fst (fst (fst (fst ΣA))))) + (snd (fst (fst (fst ΣA)))) + (snd (fst (fst ΣA))) + (snd (fst ΣA)) + (snd ΣA)) ]) where open AlgebraStr diff --git a/Cubical/Algebra/CommAlgebra/Subalgebra.agda b/Cubical/Algebra/CommAlgebra/Subalgebra.agda index c4ca6f28ba..fd052280c8 100644 --- a/Cubical/Algebra/CommAlgebra/Subalgebra.agda +++ b/Cubical/Algebra/CommAlgebra/Subalgebra.agda @@ -31,7 +31,8 @@ module _ (S : Subalgebra) where } Subalgebra→CommAlgebraHom : CommAlgebraHom Subalgebra→CommAlgebra A - Subalgebra→CommAlgebraHom = Subalgebra→AlgebraHom S + fst Subalgebra→CommAlgebraHom = Subalgebra→AlgebraHom S .fst + snd Subalgebra→CommAlgebraHom = record { IsAlgebraHom (Subalgebra→AlgebraHom S .snd) } SubalgebraHom : (B : CommAlgebra R ℓ') (f : CommAlgebraHom B A) → ((b : ⟨ B ⟩) → fst f b ∈ fst S) From 83e9919a08896e9afc72c8b466b3d07bdfe65563 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 26 Jul 2024 23:18:54 +0200 Subject: [PATCH 12/83] opaque quotients, fix unicity --- Cubical/Algebra/Ring/Quotient.agda | 148 +++++++++++++++++------------ 1 file changed, 89 insertions(+), 59 deletions(-) diff --git a/Cubical/Algebra/Ring/Quotient.agda b/Cubical/Algebra/Ring/Quotient.agda index a8ce89ad88..432193b185 100644 --- a/Cubical/Algebra/Ring/Quotient.agda +++ b/Cubical/Algebra/Ring/Quotient.agda @@ -12,8 +12,8 @@ open import Cubical.Data.Sigma using (Σ≡Prop) open import Cubical.Relation.Binary -open import Cubical.HITs.SetQuotients.Base renaming (_/_ to _/ₛ_) -open import Cubical.HITs.SetQuotients.Properties +open import Cubical.HITs.SetQuotients as SQ renaming (_/_ to _/ₛ_) +open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.Algebra.Ring open import Cubical.Algebra.Ring.Ideal @@ -55,7 +55,7 @@ module _ (R' : Ring ℓ) (I : ⟨ R' ⟩ → hProp ℓ) (I-isIdeal : isIdeal R' ((x + a) - (y + a)) ∎ pre-+/I : R → R/I → R/I - pre-+/I x = elim + pre-+/I x = SQ.elim (λ _ → squash/) (λ y → [ x + y ]) λ y y' diffrenceInIdeal @@ -71,7 +71,7 @@ module _ (R' : Ring ℓ) (I : ⟨ R' ⟩ → hProp ℓ) (I-isIdeal : isIdeal R' (λ a → lemma x y a x-y∈I) _+/I_ : R/I → R/I → R/I - x +/I y = (elim R/I→R/I-isSet pre-+/I pre-+/I-DescendsToQuotient x) y + x +/I y = (SQ.elim R/I→R/I-isSet pre-+/I pre-+/I-DescendsToQuotient x) y where R/I→R/I-isSet : R/I → isSet (R/I → R/I) R/I→R/I-isSet _ = isSetΠ (λ _ → squash/) @@ -98,7 +98,7 @@ module _ (R' : Ring ℓ) (I : ⟨ R' ⟩ → hProp ℓ) (I-isIdeal : isIdeal R' 1/I = [ 1r ] -/I : R/I → R/I - -/I = elim (λ _ → squash/) (λ x' → [ - x' ]) eq + -/I = SQ.elim (λ _ → squash/) (λ x' → [ - x' ]) eq where eq : (x y : R) → (x - y ∈ I) → [ - x ] ≡ [ - y ] eq x y x-y∈I = eq/ (- x) (- y) (subst (λ u → u ∈ I) eq' (isIdeal.-closed I-isIdeal x-y∈I)) @@ -121,7 +121,7 @@ module _ (R' : Ring ℓ) (I : ⟨ R' ⟩ → hProp ℓ) (I-isIdeal : isIdeal R' _·/I_ : R/I → R/I → R/I _·/I_ = - elim (λ _ → isSetΠ (λ _ → squash/)) + SQ.elim (λ _ → isSetΠ (λ _ → squash/)) (λ x → left· x) eq' where @@ -134,7 +134,7 @@ module _ (R' : Ring ℓ) (I : ⟨ R' ⟩ → hProp ℓ) (I-isIdeal : isIdeal R' (x · y) - (x · y') ∎) (isIdeal.·-closedLeft I-isIdeal x y-y'∈I)) left· : (x : R) → R/I → R/I - left· x = elim (λ y → squash/) + left· x = SQ.elim (λ y → squash/) (λ y → [ x · y ]) (eq x) eq' : (x x' : R) → (x - x' ∈ I) → left· x ≡ left· x' @@ -189,24 +189,45 @@ module _ (R' : Ring ℓ) (I : ⟨ R' ⟩ → hProp ℓ) (I-isIdeal : isIdeal R' +/I-assoc +/I-rid +/I-rinv +/I-comm ·/I-assoc ·/I-rid ·/I-lid /I-rdist /I-ldist -_/_ : (R : Ring ℓ) → (I : IdealsIn R) → Ring ℓ -R / (I , IisIdeal) = asRing R I IisIdeal - -[_]/I : {R : Ring ℓ} {I : IdealsIn R} → (a : ⟨ R ⟩) → ⟨ R / I ⟩ -[ a ]/I = [ a ] - -quotientHom : (R : Ring ℓ) → (I : IdealsIn R) → RingHom R (R / I) -fst (quotientHom R I) = [_] -IsRingHom.pres0 (snd (quotientHom R I)) = refl -IsRingHom.pres1 (snd (quotientHom R I)) = refl -IsRingHom.pres+ (snd (quotientHom R I)) _ _ = refl -IsRingHom.pres· (snd (quotientHom R I)) _ _ = refl -IsRingHom.pres- (snd (quotientHom R I)) _ = refl - -quotientHomSurjective : (R : Ring ℓ) → (I : IdealsIn R) - → isSurjection (fst (quotientHom R I)) -quotientHomSurjective R I = []surjective - +opaque + _/_ : (R : Ring ℓ) → (I : IdealsIn R) → Ring ℓ + R / (I , IisIdeal) = asRing R I IisIdeal + +opaque + unfolding _/_ + [_]/I : {R : Ring ℓ} {I : IdealsIn R} → (a : ⟨ R ⟩) → ⟨ R / I ⟩ + [ a ]/I = [ a ] + + quotientHom : (R : Ring ℓ) → (I : IdealsIn R) → RingHom R (R / I) + fst (quotientHom R I) = [_] + IsRingHom.pres0 (snd (quotientHom R I)) = refl + IsRingHom.pres1 (snd (quotientHom R I)) = refl + IsRingHom.pres+ (snd (quotientHom R I)) _ _ = refl + IsRingHom.pres· (snd (quotientHom R I)) _ _ = refl + IsRingHom.pres- (snd (quotientHom R I)) _ = refl + +module _ (R : Ring ℓ) (I : IdealsIn R) where + opaque + unfolding _/_ quotientHom + quotientHomSurjective : isSurjection (fst (quotientHom R I)) + quotientHomSurjective = []surjective + + open RingHoms + quotientHomEpi : (S : Ring ℓ') + → (f g : RingHom (R / I) S) + → f ∘r quotientHom R I ≡ g ∘r quotientHom R I + → f ≡ g + quotientHomEpi S f g p = + Σ≡Prop (λ _ → isPropIsRingHom _ _ _) + (funExt λ x + → PT.rec + (is-set _ _) + (λ {(x' , [x']≡x) → f .fst x ≡⟨ cong (λ y → f .fst y) (sym [x']≡x) ⟩ + fst (f ∘r quotientHom R I) x' ≡⟨ cong (λ h → fst h x') p ⟩ + fst (g ∘r quotientHom R I) x' ≡⟨ cong (λ y → g .fst y) [x']≡x ⟩ + g .fst x ∎ }) + (quotientHomSurjective x )) + where open RingStr (S .snd) using (is-set) module UniversalProperty (R : Ring ℓ) (I : IdealsIn R) where open RingStr ⦃...⦄ @@ -230,33 +251,38 @@ module UniversalProperty (R : Ring ℓ) (I : IdealsIn R) where if S is from a different universe. Instead, the condition, that Iₛ is contained in the kernel of φ is rephrased explicitly. -} - inducedHom : ((x : ⟨ R ⟩) → x ∈ Iₛ → φ $r x ≡ 0r) → RingHom (R / I) S - fst (inducedHom Iₛ⊆kernel) = - elim - (λ _ → is-set) - f - λ r₁ r₂ r₁-r₂∈I → equalByDifference (f r₁) (f r₂) - (f r₁ - f r₂ ≡⟨ cong (λ u → f r₁ + u) (sym (φ.pres- _)) ⟩ - f r₁ + f (- r₂) ≡⟨ sym (φ.pres+ _ _) ⟩ - f (r₁ - r₂) ≡⟨ Iₛ⊆kernel (r₁ - r₂) r₁-r₂∈I ⟩ - 0r ∎) - pres0 (snd (inducedHom Iₛ⊆kernel)) = φ.pres0 - pres1 (snd (inducedHom Iₛ⊆kernel)) = φ.pres1 - pres+ (snd (inducedHom Iₛ⊆kernel)) = - elimProp2 (λ _ _ → is-set _ _) φ.pres+ - pres· (snd (inducedHom Iₛ⊆kernel)) = - elimProp2 (λ _ _ → is-set _ _) φ.pres· - pres- (snd (inducedHom Iₛ⊆kernel)) = - elimProp (λ _ → is-set _ _) φ.pres- - - solution : (p : ((x : ⟨ R ⟩) → x ∈ Iₛ → φ $r x ≡ 0r)) - → (x : ⟨ R ⟩) → inducedHom p $r [ x ] ≡ φ $r x - solution p x = refl - - unique : (p : ((x : ⟨ R ⟩) → x ∈ Iₛ → φ $r x ≡ 0r)) - → (ψ : RingHom (R / I) S) → (ψIsSolution : (x : ⟨ R ⟩) → ψ $r [ x ] ≡ φ $r x) - → (x : ⟨ R ⟩) → ψ $r [ x ] ≡ inducedHom p $r [ x ] - unique p ψ ψIsSolution x = ψIsSolution x + opaque + unfolding _/_ + inducedHom : ((x : ⟨ R ⟩) → x ∈ Iₛ → φ $r x ≡ 0r) → RingHom (R / I) S + fst (inducedHom Iₛ⊆kernel) = + SQ.elim + (λ _ → is-set) + f + λ r₁ r₂ r₁-r₂∈I → equalByDifference (f r₁) (f r₂) + (f r₁ - f r₂ ≡⟨ cong (λ u → f r₁ + u) (sym (φ.pres- _)) ⟩ + f r₁ + f (- r₂) ≡⟨ sym (φ.pres+ _ _) ⟩ + f (r₁ - r₂) ≡⟨ Iₛ⊆kernel (r₁ - r₂) r₁-r₂∈I ⟩ + 0r ∎) + pres0 (snd (inducedHom Iₛ⊆kernel)) = φ.pres0 + pres1 (snd (inducedHom Iₛ⊆kernel)) = φ.pres1 + pres+ (snd (inducedHom Iₛ⊆kernel)) = + elimProp2 (λ _ _ → is-set _ _) φ.pres+ + pres· (snd (inducedHom Iₛ⊆kernel)) = + elimProp2 (λ _ _ → is-set _ _) φ.pres· + pres- (snd (inducedHom Iₛ⊆kernel)) = + elimProp (λ _ → is-set _ _) φ.pres- + + opaque + unfolding _/_ [_]/I inducedHom + open RingHoms + solution : (p : ((x : ⟨ R ⟩) → x ∈ Iₛ → φ $r x ≡ 0r)) + → inducedHom p ∘r quotientHom R I ≡ φ + solution p = Σ≡Prop (λ _ → isPropIsRingHom _ _ _) refl + + unique : (p : ((x : ⟨ R ⟩) → x ∈ Iₛ → φ $r x ≡ 0r)) + → (ψ : RingHom (R / I) S) → (ψIsSolution : ψ ∘r quotientHom R I ≡ φ) + → ψ ≡ inducedHom p + unique p ψ ψIsSolution = quotientHomEpi R I S ψ (inducedHom p) (ψIsSolution ∙ sym (solution p)) {- Show that the kernel of the quotient map @@ -276,8 +302,10 @@ module idealIsKernel {R : Ring ℓ} (I : IdealsIn R) where x + 0r ≡⟨ +IdR x ⟩ x ∎ - I⊆ker : fst I ⊆ kernel π - I⊆ker x x∈I = eq/ _ _ (subst (_∈ fst I) (sym (x-0≡x x)) x∈I) + opaque + unfolding _/_ [_]/I + I⊆ker : fst I ⊆ kernel π + I⊆ker x x∈I = eq/ _ _ (subst (_∈ fst I) (sym (x-0≡x x)) x∈I) private _~_ : Rel ⟨ R ⟩ ⟨ R ⟩ ℓ @@ -310,11 +338,13 @@ module idealIsKernel {R : Ring ℓ} (I : IdealsIn R) where symmetric ~IsEquivRel x y x~y = subst (_∈ fst I) -[x-y]≡y-x (-closed x~y) transitive ~IsEquivRel x y z x~y y~z = subst (_∈ fst I) x-y+y-z≡x-z (+-closed x~y y~z) - ker⊆I : kernel π ⊆ fst I - ker⊆I x x∈ker = subst (_∈ fst I) (x-0≡x x) x-0∈I - where - x-0∈I : x - 0r ∈ fst I - x-0∈I = effective ~IsPropValued ~IsEquivRel x 0r x∈ker + opaque + unfolding [_]/I + ker⊆I : kernel π ⊆ fst I + ker⊆I x x∈ker = subst (_∈ fst I) (x-0≡x x) x-0∈I + where + x-0∈I : x - 0r ∈ fst I + x-0∈I = effective ~IsPropValued ~IsEquivRel x 0r x∈ker kernel≡I : {R : Ring ℓ} (I : IdealsIn R) → kernelIdeal (quotientHom R I) ≡ I From bbda88b2048c548197d002c7604a91e74237cbfd Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 29 Jul 2024 11:28:28 +0200 Subject: [PATCH 13/83] fast tc for quotient comm rings --- Cubical/Algebra/CommRing/Quotient/Base.agda | 131 ++++++++++++++------ Cubical/Algebra/Ring/Quotient.agda | 129 ++++++++++--------- 2 files changed, 162 insertions(+), 98 deletions(-) diff --git a/Cubical/Algebra/CommRing/Quotient/Base.agda b/Cubical/Algebra/CommRing/Quotient/Base.agda index 960717c018..404a538d1e 100644 --- a/Cubical/Algebra/CommRing/Quotient/Base.agda +++ b/Cubical/Algebra/CommRing/Quotient/Base.agda @@ -8,8 +8,9 @@ open import Cubical.Functions.Surjection open import Cubical.Data.Nat open import Cubical.Data.FinData +open import Cubical.Data.Sigma using (Σ≡Prop) -open import Cubical.HITs.SetQuotients using ([_]; squash/; elimProp2) +open import Cubical.HITs.SetQuotients as SQ renaming (_/_ to _/ₛ_) open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.Algebra.CommRing @@ -23,57 +24,104 @@ private variable ℓ ℓ' : Level -_/_ : (R : CommRing ℓ) → (I : IdealsIn R) → CommRing ℓ -R / I = - fst asRing , commringstr _ _ _ _ _ - (iscommring (RingStr.isRing (snd asRing)) - (elimProp2 (λ _ _ → squash/ _ _) - commEq)) - where - asRing = (CommRing→Ring R) Ring./ (CommIdeal→Ideal I) - _·/_ : fst asRing → fst asRing → fst asRing - _·/_ = RingStr._·_ (snd asRing) - commEq : (x y : fst R) → ([ x ] ·/ [ y ]) ≡ ([ y ] ·/ [ x ]) - commEq x y i = [ CommRingStr.·Comm (snd R) x y i ] +module _ (R : CommRing ℓ) (I : IdealsIn R) where + open CommRingStr (snd R) + R/I = ⟨ R ⟩ /ₛ (λ x y → x - y ∈ (fst I)) + + opaque + unfolding Ring.quotientRingStr + quotientCommRingStr : CommRingStr R/I + quotientCommRingStr = snd + (Ring→CommRing + ((CommRing→Ring R) Ring./ (CommIdeal→Ideal I)) + (elimProp2 (λ _ _ → squash/ _ _) + λ x y i → [ CommRingStr.·Comm (snd R) x y i ])) + + +_/_ : (R : CommRing ℓ) (I : IdealsIn R) → CommRing ℓ +fst (R / I) = R/I R I +snd (R / I) = quotientCommRingStr R I [_]/ : {R : CommRing ℓ} {I : IdealsIn R} → (a : fst R) → fst (R / I) -[ a ]/ = [ a ] +[ a ]/ = SQ.[ a ] + + + +module Coherence (R : CommRing ℓ) (I : IdealsIn R) where + opaque + unfolding quotientCommRingStr Ring.quotientRingStr + isRingHomCoh : IsRingHom (snd (CommRing→Ring (R / I))) + (λ x → x) + (snd ((CommRing→Ring R) Ring./ (CommIdeal→Ideal I))) + IsRingHom.pres0 isRingHomCoh = refl + IsRingHom.pres1 isRingHomCoh = refl + IsRingHom.pres+ isRingHomCoh = λ _ _ → refl + IsRingHom.pres· isRingHomCoh = λ _ _ → refl + IsRingHom.pres- isRingHomCoh = λ _ → refl + isRingHomCohInv : IsRingHom (snd ((CommRing→Ring R) Ring./ (CommIdeal→Ideal I))) + (λ x → x) + (snd (CommRing→Ring (R / I))) + IsRingHom.pres0 isRingHomCohInv = refl + IsRingHom.pres1 isRingHomCohInv = refl + IsRingHom.pres+ isRingHomCohInv = λ _ _ → refl + IsRingHom.pres· isRingHomCohInv = λ _ _ → refl + IsRingHom.pres- isRingHomCohInv = λ _ → refl + + 0r≡ : RingStr.0r (CommRing→Ring (R / I) .snd) ≡ RingStr.0r ((CommRing→Ring R Ring./ CommIdeal→Ideal I) .snd) + 0r≡ = refl + + ringStr : RingHom (CommRing→Ring (R / I)) + ((CommRing→Ring R) Ring./ (CommIdeal→Ideal I)) + fst ringStr x = x + (snd ringStr) = isRingHomCoh + + ringStrInv : RingHom ((CommRing→Ring R) Ring./ (CommIdeal→Ideal I)) + (CommRing→Ring (R / I)) + + fst ringStrInv x = x + (snd ringStrInv) = isRingHomCohInv + + +open RingHoms module Quotient-FGideal-CommRing-Ring - (A : CommRing ℓ) - (B : Ring ℓ') - (g : RingHom (CommRing→Ring A) B) + (R : CommRing ℓ) + (S : Ring ℓ') + (f : RingHom (CommRing→Ring R) S) {n : ℕ} - (v : FinVec ⟨ A ⟩ n) - (gnull : (k : Fin n) → g $r v k ≡ RingStr.0r (snd B)) + (v : FinVec ⟨ R ⟩ n) + (fnull : (k : Fin n) → f $r v k ≡ RingStr.0r (snd S)) where - open RingStr (snd B) using (0r; is-set) + open RingStr (snd S) using (0r; is-set) - zeroOnGeneratedIdeal : (x : ⟨ A ⟩) → x ∈ fst (generatedIdeal A v) → g $r x ≡ 0r + Iv = generatedIdeal R v + + zeroOnGeneratedIdeal : (x : ⟨ R ⟩) → x ∈ fst Iv → f $r x ≡ 0r zeroOnGeneratedIdeal x x∈FGIdeal = PT.elim - (λ _ → is-set (g $r x) 0r) - (λ {(α , isLC) → subst _ (sym isLC) (cancelLinearCombination A B g _ α v gnull)}) + (λ _ → is-set (f $r x) 0r) + (λ {(α , isLC) → subst _ (sym isLC) (cancelLinearCombination R S f _ α v fnull)}) x∈FGIdeal - inducedHom : RingHom (CommRing→Ring (A / (generatedIdeal _ v))) B - inducedHom = Ring.UniversalProperty.inducedHom (CommRing→Ring A) (CommIdeal→Ideal ideal) g zeroOnGeneratedIdeal - where ideal = generatedIdeal A v + inducedHom : RingHom (CommRing→Ring (R / (generatedIdeal _ v))) S + inducedHom = Ring.UniversalProperty.inducedHom + (CommRing→Ring R) (CommIdeal→Ideal Iv) f zeroOnGeneratedIdeal + ∘r (Coherence.ringStr R Iv) module Quotient-FGideal-CommRing-CommRing - (A : CommRing ℓ) - (B : CommRing ℓ') - (g : CommRingHom A B) + (R : CommRing ℓ) + (S : CommRing ℓ') + (f : CommRingHom R S) {n : ℕ} - (v : FinVec ⟨ A ⟩ n) - (gnull : (k : Fin n) → g $r v k ≡ CommRingStr.0r (snd B)) + (v : FinVec ⟨ R ⟩ n) + (fnull : (k : Fin n) → f $r v k ≡ CommRingStr.0r (snd S)) where - inducedHom : CommRingHom (A / (generatedIdeal _ v)) B - inducedHom = Quotient-FGideal-CommRing-Ring.inducedHom A (CommRing→Ring B) g v gnull + inducedHom : CommRingHom (R / (generatedIdeal _ v)) S + inducedHom = Quotient-FGideal-CommRing-Ring.inducedHom R (CommRing→Ring S) f v fnull module UniversalProperty (R S : CommRing ℓ) @@ -84,15 +132,16 @@ module UniversalProperty inducedHom : CommRingHom (R / I) S inducedHom = Ring.UniversalProperty.inducedHom (CommRing→Ring R) (CommIdeal→Ideal I) f I⊆ker - + ∘r Coherence.ringStr R I quotientHom : (R : CommRing ℓ) → (I : IdealsIn R) → CommRingHom R (R / I) -quotientHom R I = Ring.quotientHom (CommRing→Ring R) (CommIdeal→Ideal I) +quotientHom R I = Coherence.ringStrInv R I ∘r Ring.quotientHom (CommRing→Ring R) (CommIdeal→Ideal I) quotientHomSurjective : (R : CommRing ℓ) → (I : IdealsIn R) → isSurjection (fst (quotientHom R I)) quotientHomSurjective R I = Ring.quotientHomSurjective (CommRing→Ring R) (CommIdeal→Ideal I) + module _ {R : CommRing ℓ} (I : IdealsIn R) where open CommRingStr ⦃...⦄ private @@ -100,8 +149,16 @@ module _ {R : CommRing ℓ} (I : IdealsIn R) where instance _ = snd R _ = snd (R / I) - kernel≡I : kernelIdeal R (R / I) π ≡ I - kernel≡I = cong Ideal→CommIdeal (Ring.kernel≡I (CommIdeal→Ideal I)) + opaque + kernel≡I : kernelIdeal R (R / I) π ≡ I + kernel≡I = + Σ≡Prop (CommIdeal.isPropIsCommIdeal _) + (funExt + λ x → Σ≡Prop (λ _ → isPropIsProp) + let reason = cong (λ y → π .fst x ≡ y) (Coherence.0r≡ R I) + in (π .fst x ≡ RingStr.0r (CommRing→Ring (R / I) .snd) ≡⟨ reason ⟩ + π .fst x ≡ RingStr.0r ((CommRing→Ring R Ring./ CommIdeal→Ideal I) .snd) ∎)) + ∙ cong Ideal→CommIdeal (Ring.kernel≡I (CommIdeal→Ideal I)) zeroOnIdeal : (x : ⟨ R ⟩) → x ∈ fst I → fst π x ≡ 0r zeroOnIdeal x x∈I = subst (λ P → fst ((fst P) x)) (sym kernel≡I) x∈I diff --git a/Cubical/Algebra/Ring/Quotient.agda b/Cubical/Algebra/Ring/Quotient.agda index 432193b185..8b4b879019 100644 --- a/Cubical/Algebra/Ring/Quotient.agda +++ b/Cubical/Algebra/Ring/Quotient.agda @@ -184,50 +184,54 @@ module _ (R' : Ring ℓ) (I : ⟨ R' ⟩ → hProp ℓ) (I-isIdeal : isIdeal R' eq : (x y z : R) → [ x ] ·/I ([ y ] +/I [ z ]) ≡ ([ x ] ·/I [ y ]) +/I ([ x ] ·/I [ z ]) eq x y z i = [ ·DistR+ x y z i ] - asRing : Ring ℓ - asRing = makeRing 0/I 1/I _+/I_ _·/I_ -/I isSetR/I - +/I-assoc +/I-rid +/I-rinv +/I-comm - ·/I-assoc ·/I-rid ·/I-lid /I-rdist /I-ldist + opaque + quotientRingStr : RingStr R/I + quotientRingStr = snd (makeRing 0/I 1/I _+/I_ _·/I_ -/I isSetR/I + +/I-assoc +/I-rid +/I-rinv +/I-comm + ·/I-assoc ·/I-rid ·/I-lid /I-rdist /I-ldist) -opaque - _/_ : (R : Ring ℓ) → (I : IdealsIn R) → Ring ℓ - R / (I , IisIdeal) = asRing R I IisIdeal +_/_ : (R : Ring ℓ) → (I : IdealsIn R) → Ring ℓ +fst (R / I) = R/I R (fst I) (snd I) +snd (R / I) = quotientRingStr R (fst I) (snd I) + +[_]/I : {R : Ring ℓ} {I : IdealsIn R} → (a : ⟨ R ⟩) → ⟨ R / I ⟩ +[ a ]/I = [ a ] opaque - unfolding _/_ - [_]/I : {R : Ring ℓ} {I : IdealsIn R} → (a : ⟨ R ⟩) → ⟨ R / I ⟩ - [ a ]/I = [ a ] - - quotientHom : (R : Ring ℓ) → (I : IdealsIn R) → RingHom R (R / I) - fst (quotientHom R I) = [_] - IsRingHom.pres0 (snd (quotientHom R I)) = refl - IsRingHom.pres1 (snd (quotientHom R I)) = refl - IsRingHom.pres+ (snd (quotientHom R I)) _ _ = refl - IsRingHom.pres· (snd (quotientHom R I)) _ _ = refl - IsRingHom.pres- (snd (quotientHom R I)) _ = refl + unfolding quotientRingStr + isRingHomQuotientHom : (R : Ring ℓ) (I : IdealsIn R) + → IsRingHom (R .snd) [_] ((R / I) .snd) + IsRingHom.pres0 (isRingHomQuotientHom R I) = refl + IsRingHom.pres1 (isRingHomQuotientHom R I) = refl + IsRingHom.pres+ (isRingHomQuotientHom R I) _ _ = refl + IsRingHom.pres· (isRingHomQuotientHom R I) _ _ = refl + IsRingHom.pres- (isRingHomQuotientHom R I) _ = refl + +quotientHom : (R : Ring ℓ) → (I : IdealsIn R) → RingHom R (R / I) +fst (quotientHom R I) = [_] +snd (quotientHom R I) = isRingHomQuotientHom R I module _ (R : Ring ℓ) (I : IdealsIn R) where - opaque - unfolding _/_ quotientHom - quotientHomSurjective : isSurjection (fst (quotientHom R I)) - quotientHomSurjective = []surjective + quotientHomSurjective : isSurjection (fst (quotientHom R I)) + quotientHomSurjective = []surjective + + open RingHoms + quotientHomEpi : (S : Ring ℓ') + → (f g : RingHom (R / I) S) + → f ∘r quotientHom R I ≡ g ∘r quotientHom R I + → f ≡ g + quotientHomEpi S f g p = + Σ≡Prop (λ _ → isPropIsRingHom _ _ _) + (funExt λ x + → PT.rec + (is-set _ _) + (λ {(x' , [x']≡x) → f .fst x ≡⟨ cong (λ y → f .fst y) (sym [x']≡x) ⟩ + fst (f ∘r quotientHom R I) x' ≡⟨ cong (λ h → fst h x') p ⟩ + fst (g ∘r quotientHom R I) x' ≡⟨ cong (λ y → g .fst y) [x']≡x ⟩ + g .fst x ∎ }) + (quotientHomSurjective x )) + where open RingStr (S .snd) using (is-set) - open RingHoms - quotientHomEpi : (S : Ring ℓ') - → (f g : RingHom (R / I) S) - → f ∘r quotientHom R I ≡ g ∘r quotientHom R I - → f ≡ g - quotientHomEpi S f g p = - Σ≡Prop (λ _ → isPropIsRingHom _ _ _) - (funExt λ x - → PT.rec - (is-set _ _) - (λ {(x' , [x']≡x) → f .fst x ≡⟨ cong (λ y → f .fst y) (sym [x']≡x) ⟩ - fst (f ∘r quotientHom R I) x' ≡⟨ cong (λ h → fst h x') p ⟩ - fst (g ∘r quotientHom R I) x' ≡⟨ cong (λ y → g .fst y) [x']≡x ⟩ - g .fst x ∎ }) - (quotientHomSurjective x )) - where open RingStr (S .snd) using (is-set) module UniversalProperty (R : Ring ℓ) (I : IdealsIn R) where open RingStr ⦃...⦄ @@ -251,30 +255,33 @@ module UniversalProperty (R : Ring ℓ) (I : IdealsIn R) where if S is from a different universe. Instead, the condition, that Iₛ is contained in the kernel of φ is rephrased explicitly. -} + inducedHomMap : ((x : ⟨ R ⟩) → x ∈ Iₛ → φ $r x ≡ 0r) → ⟨ R / I ⟩ → ⟨ S ⟩ + inducedHomMap Iₛ⊆kernel = + SQ.elim + (λ _ → is-set) + f + λ r₁ r₂ r₁-r₂∈I → equalByDifference (f r₁) (f r₂) + (f r₁ - f r₂ ≡⟨ cong (λ u → f r₁ + u) (sym (φ.pres- _)) ⟩ + f r₁ + f (- r₂) ≡⟨ sym (φ.pres+ _ _) ⟩ + f (r₁ - r₂) ≡⟨ Iₛ⊆kernel (r₁ - r₂) r₁-r₂∈I ⟩ + 0r ∎) + opaque - unfolding _/_ - inducedHom : ((x : ⟨ R ⟩) → x ∈ Iₛ → φ $r x ≡ 0r) → RingHom (R / I) S - fst (inducedHom Iₛ⊆kernel) = - SQ.elim - (λ _ → is-set) - f - λ r₁ r₂ r₁-r₂∈I → equalByDifference (f r₁) (f r₂) - (f r₁ - f r₂ ≡⟨ cong (λ u → f r₁ + u) (sym (φ.pres- _)) ⟩ - f r₁ + f (- r₂) ≡⟨ sym (φ.pres+ _ _) ⟩ - f (r₁ - r₂) ≡⟨ Iₛ⊆kernel (r₁ - r₂) r₁-r₂∈I ⟩ - 0r ∎) - pres0 (snd (inducedHom Iₛ⊆kernel)) = φ.pres0 - pres1 (snd (inducedHom Iₛ⊆kernel)) = φ.pres1 - pres+ (snd (inducedHom Iₛ⊆kernel)) = - elimProp2 (λ _ _ → is-set _ _) φ.pres+ - pres· (snd (inducedHom Iₛ⊆kernel)) = - elimProp2 (λ _ _ → is-set _ _) φ.pres· - pres- (snd (inducedHom Iₛ⊆kernel)) = - elimProp (λ _ → is-set _ _) φ.pres- + unfolding quotientRingStr + inducedMapIsRingHom : (Iₛ⊆kernel : (x : ⟨ R ⟩) → x ∈ Iₛ → φ $r x ≡ 0r) + → IsRingHom ((R / I) .snd) (inducedHomMap Iₛ⊆kernel) (S .snd) + pres0 (inducedMapIsRingHom Iₛ⊆kernel) = φ.pres0 + pres1 (inducedMapIsRingHom Iₛ⊆kernel) = φ.pres1 + pres+ (inducedMapIsRingHom Iₛ⊆kernel) = elimProp2 (λ _ _ → is-set _ _) φ.pres+ + pres· (inducedMapIsRingHom Iₛ⊆kernel) = elimProp2 (λ _ _ → is-set _ _) φ.pres· + pres- (inducedMapIsRingHom Iₛ⊆kernel) = elimProp (λ _ → is-set _ _) φ.pres- + + inducedHom : ((x : ⟨ R ⟩) → x ∈ Iₛ → φ $r x ≡ 0r) → RingHom (R / I) S + fst (inducedHom Iₛ⊆kernel) = inducedHomMap Iₛ⊆kernel + snd (inducedHom Iₛ⊆kernel) = inducedMapIsRingHom Iₛ⊆kernel + open RingHoms opaque - unfolding _/_ [_]/I inducedHom - open RingHoms solution : (p : ((x : ⟨ R ⟩) → x ∈ Iₛ → φ $r x ≡ 0r)) → inducedHom p ∘r quotientHom R I ≡ φ solution p = Σ≡Prop (λ _ → isPropIsRingHom _ _ _) refl @@ -303,7 +310,7 @@ module idealIsKernel {R : Ring ℓ} (I : IdealsIn R) where x ∎ opaque - unfolding _/_ [_]/I + unfolding isRingHomQuotientHom I⊆ker : fst I ⊆ kernel π I⊆ker x x∈I = eq/ _ _ (subst (_∈ fst I) (sym (x-0≡x x)) x∈I) @@ -339,7 +346,7 @@ module idealIsKernel {R : Ring ℓ} (I : IdealsIn R) where transitive ~IsEquivRel x y z x~y y~z = subst (_∈ fst I) x-y+y-z≡x-z (+-closed x~y y~z) opaque - unfolding [_]/I + unfolding isRingHomQuotientHom ker⊆I : kernel π ⊆ fst I ker⊆I x x∈ker = subst (_∈ fst I) (x-0≡x x) x-0∈I where From 35274b50798157b3178e542ef541a86d094868e4 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 29 Jul 2024 15:37:13 +0200 Subject: [PATCH 14/83] uniform indentation --- Cubical/Algebra/CommRing/Ideal/Base.agda | 138 +++++++++++------------ 1 file changed, 69 insertions(+), 69 deletions(-) diff --git a/Cubical/Algebra/CommRing/Ideal/Base.agda b/Cubical/Algebra/CommRing/Ideal/Base.agda index 413d51ff84..9a29cfab2d 100644 --- a/Cubical/Algebra/CommRing/Ideal/Base.agda +++ b/Cubical/Algebra/CommRing/Ideal/Base.agda @@ -35,75 +35,75 @@ private ℓ : Level module CommIdeal (R' : CommRing ℓ) where - private R = fst R' - open CommRingStr (snd R') - open Sum (CommRing→Ring R') - open CommRingTheory R' - open RingTheory (CommRing→Ring R') - - record isCommIdeal (I : ℙ R) : Type ℓ where - constructor - makeIsCommIdeal - field - +Closed : ∀ {x y : R} → x ∈p I → y ∈p I → (x + y) ∈p I - contains0 : 0r ∈p I - ·Closed : ∀ {x : R} (r : R) → x ∈p I → r · x ∈p I - - ·RClosed : ∀ {x : R} (r : R) → x ∈p I → x · r ∈p I - ·RClosed r x∈pI = subst-∈p I (·Comm _ _) (·Closed r x∈pI) - - open isCommIdeal - isPropIsCommIdeal : (I : ℙ R) → isProp (isCommIdeal I) - +Closed (isPropIsCommIdeal I ici₁ ici₂ i) x∈pI y∈pI = - I _ .snd (ici₁ .+Closed x∈pI y∈pI) (ici₂ .+Closed x∈pI y∈pI) i - contains0 (isPropIsCommIdeal I ici₁ ici₂ i) = I 0r .snd (ici₁ .contains0) (ici₂ .contains0) i - ·Closed (isPropIsCommIdeal I ici₁ ici₂ i) r x∈pI = - I _ .snd (ici₁ .·Closed r x∈pI) (ici₂ .·Closed r x∈pI) i - - CommIdeal : Type (ℓ-suc ℓ) - CommIdeal = Σ[ I ∈ ℙ R ] isCommIdeal I - - isSetCommIdeal : isSet CommIdeal - isSetCommIdeal = isSetΣSndProp isSetℙ isPropIsCommIdeal - - --inclusion and containment of ideals - _⊆_ : CommIdeal → CommIdeal → Type ℓ - I ⊆ J = I .fst ⊆p J .fst - - infix 5 _∈_ - _∈_ : R → CommIdeal → Type ℓ - x ∈ I = x ∈p I .fst - - subst-∈ : (I : CommIdeal) {x y : R} → x ≡ y → x ∈ I → y ∈ I - subst-∈ I = subst-∈p (I .fst) - - CommIdeal≡Char : {I J : CommIdeal} → I ⊆ J → J ⊆ I → I ≡ J - CommIdeal≡Char I⊆J J⊆I = Σ≡Prop isPropIsCommIdeal (⊆-extensionality _ _ (I⊆J , J⊆I)) - - -Closed : (I : CommIdeal) (x : R) - → x ∈ I → (- x) ∈ I - -Closed I x x∈I = subst (_∈ I) (solve! R') (·Closed (snd I) (- 1r) x∈I) - - ∑Closed : (I : CommIdeal) {n : ℕ} (V : FinVec R n) - → (∀ i → V i ∈ I) → ∑ V ∈ I - ∑Closed I {n = zero} _ _ = I .snd .contains0 - ∑Closed I {n = suc n} V h = I .snd .+Closed (h zero) (∑Closed I (V ∘ suc) (h ∘ suc)) - - 0Ideal : CommIdeal - fst 0Ideal x = (x ≡ 0r) , is-set _ _ - +Closed (snd 0Ideal) x≡0 y≡0 = cong₂ (_+_) x≡0 y≡0 ∙ +IdR _ - contains0 (snd 0Ideal) = refl - ·Closed (snd 0Ideal) r x≡0 = cong (r ·_) x≡0 ∙ 0RightAnnihilates _ - - 1Ideal : CommIdeal - fst 1Ideal x = ⊤ - +Closed (snd 1Ideal) _ _ = lift tt - contains0 (snd 1Ideal) = lift tt - ·Closed (snd 1Ideal) _ _ = lift tt - - contains1Is1 : (I : CommIdeal) → 1r ∈ I → I ≡ 1Ideal - contains1Is1 I 1∈I = CommIdeal≡Char (λ _ _ → lift tt) - λ x _ → subst-∈ I (·IdR _) (I .snd .·Closed x 1∈I) -- x≡x·1 ∈ I + private R = fst R' + open CommRingStr (snd R') + open Sum (CommRing→Ring R') + open CommRingTheory R' + open RingTheory (CommRing→Ring R') + + record isCommIdeal (I : ℙ R) : Type ℓ where + constructor + makeIsCommIdeal + field + +Closed : ∀ {x y : R} → x ∈p I → y ∈p I → (x + y) ∈p I + contains0 : 0r ∈p I + ·Closed : ∀ {x : R} (r : R) → x ∈p I → r · x ∈p I + + ·RClosed : ∀ {x : R} (r : R) → x ∈p I → x · r ∈p I + ·RClosed r x∈pI = subst-∈p I (·Comm _ _) (·Closed r x∈pI) + + open isCommIdeal + isPropIsCommIdeal : (I : ℙ R) → isProp (isCommIdeal I) + +Closed (isPropIsCommIdeal I ici₁ ici₂ i) x∈pI y∈pI = + I _ .snd (ici₁ .+Closed x∈pI y∈pI) (ici₂ .+Closed x∈pI y∈pI) i + contains0 (isPropIsCommIdeal I ici₁ ici₂ i) = I 0r .snd (ici₁ .contains0) (ici₂ .contains0) i + ·Closed (isPropIsCommIdeal I ici₁ ici₂ i) r x∈pI = + I _ .snd (ici₁ .·Closed r x∈pI) (ici₂ .·Closed r x∈pI) i + + CommIdeal : Type (ℓ-suc ℓ) + CommIdeal = Σ[ I ∈ ℙ R ] isCommIdeal I + + isSetCommIdeal : isSet CommIdeal + isSetCommIdeal = isSetΣSndProp isSetℙ isPropIsCommIdeal + + --inclusion and containment of ideals + _⊆_ : CommIdeal → CommIdeal → Type ℓ + I ⊆ J = I .fst ⊆p J .fst + + infix 5 _∈_ + _∈_ : R → CommIdeal → Type ℓ + x ∈ I = x ∈p I .fst + + subst-∈ : (I : CommIdeal) {x y : R} → x ≡ y → x ∈ I → y ∈ I + subst-∈ I = subst-∈p (I .fst) + + CommIdeal≡Char : {I J : CommIdeal} → I ⊆ J → J ⊆ I → I ≡ J + CommIdeal≡Char I⊆J J⊆I = Σ≡Prop isPropIsCommIdeal (⊆-extensionality _ _ (I⊆J , J⊆I)) + + -Closed : (I : CommIdeal) (x : R) + → x ∈ I → (- x) ∈ I + -Closed I x x∈I = subst (_∈ I) (solve! R') (·Closed (snd I) (- 1r) x∈I) + + ∑Closed : (I : CommIdeal) {n : ℕ} (V : FinVec R n) + → (∀ i → V i ∈ I) → ∑ V ∈ I + ∑Closed I {n = zero} _ _ = I .snd .contains0 + ∑Closed I {n = suc n} V h = I .snd .+Closed (h zero) (∑Closed I (V ∘ suc) (h ∘ suc)) + + 0Ideal : CommIdeal + fst 0Ideal x = (x ≡ 0r) , is-set _ _ + +Closed (snd 0Ideal) x≡0 y≡0 = cong₂ (_+_) x≡0 y≡0 ∙ +IdR _ + contains0 (snd 0Ideal) = refl + ·Closed (snd 0Ideal) r x≡0 = cong (r ·_) x≡0 ∙ 0RightAnnihilates _ + + 1Ideal : CommIdeal + fst 1Ideal x = ⊤ + +Closed (snd 1Ideal) _ _ = lift tt + contains0 (snd 1Ideal) = lift tt + ·Closed (snd 1Ideal) _ _ = lift tt + + contains1Is1 : (I : CommIdeal) → 1r ∈ I → I ≡ 1Ideal + contains1Is1 I 1∈I = CommIdeal≡Char (λ _ _ → lift tt) + λ x _ → subst-∈ I (·IdR _) (I .snd .·Closed x 1∈I) -- x≡x·1 ∈ I IdealsIn : (R : CommRing ℓ) → Type _ IdealsIn R = CommIdeal.CommIdeal R From 41bc2a35164708b1466d443f2a15cc98a67302ea Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 29 Jul 2024 15:42:36 +0200 Subject: [PATCH 15/83] cleanup /remove flag --- Cubical/Algebra/CommRing/Quotient/Base.agda | 8 +------- 1 file changed, 1 insertion(+), 7 deletions(-) diff --git a/Cubical/Algebra/CommRing/Quotient/Base.agda b/Cubical/Algebra/CommRing/Quotient/Base.agda index 404a538d1e..66eb50e4d6 100644 --- a/Cubical/Algebra/CommRing/Quotient/Base.agda +++ b/Cubical/Algebra/CommRing/Quotient/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --safe --lossy-unification #-} +{-# OPTIONS --safe #-} module Cubical.Algebra.CommRing.Quotient.Base where open import Cubical.Foundations.Prelude @@ -37,7 +37,6 @@ module _ (R : CommRing ℓ) (I : IdealsIn R) where (elimProp2 (λ _ _ → squash/ _ _) λ x y i → [ CommRingStr.·Comm (snd R) x y i ])) - _/_ : (R : CommRing ℓ) (I : IdealsIn R) → CommRing ℓ fst (R / I) = R/I R I snd (R / I) = quotientCommRingStr R I @@ -45,8 +44,6 @@ snd (R / I) = quotientCommRingStr R I [_]/ : {R : CommRing ℓ} {I : IdealsIn R} → (a : fst R) → fst (R / I) [ a ]/ = SQ.[ a ] - - module Coherence (R : CommRing ℓ) (I : IdealsIn R) where opaque unfolding quotientCommRingStr Ring.quotientRingStr @@ -82,10 +79,8 @@ module Coherence (R : CommRing ℓ) (I : IdealsIn R) where fst ringStrInv x = x (snd ringStrInv) = isRingHomCohInv - open RingHoms - module Quotient-FGideal-CommRing-Ring (R : CommRing ℓ) (S : Ring ℓ') @@ -141,7 +136,6 @@ quotientHomSurjective : (R : CommRing ℓ) → (I : IdealsIn R) → isSurjection (fst (quotientHom R I)) quotientHomSurjective R I = Ring.quotientHomSurjective (CommRing→Ring R) (CommIdeal→Ideal I) - module _ {R : CommRing ℓ} (I : IdealsIn R) where open CommRingStr ⦃...⦄ private From a4980391c1e9f7417a17e90f99d817a35e53d368 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 29 Jul 2024 17:16:34 +0200 Subject: [PATCH 16/83] improve names --- Cubical/Algebra/DirectSum/DirectSumHIT/Base.agda | 14 +++++++------- 1 file changed, 7 insertions(+), 7 deletions(-) diff --git a/Cubical/Algebra/DirectSum/DirectSumHIT/Base.agda b/Cubical/Algebra/DirectSum/DirectSumHIT/Base.agda index 987f196b30..7884b5235b 100644 --- a/Cubical/Algebra/DirectSum/DirectSumHIT/Base.agda +++ b/Cubical/Algebra/DirectSum/DirectSumHIT/Base.agda @@ -90,7 +90,7 @@ module _ (Idx : Type ℓ) (P : Idx → Type ℓ') (AGP : (r : Idx) → AbGroupSt module DS-Ind-Prop (Q : (x : ⊕HIT Idx P AGP) → Type ℓ'') - (ispd : (x : ⊕HIT Idx P AGP) → isProp (Q x)) + (isPropQ : (x : ⊕HIT Idx P AGP) → isProp (Q x)) -- elements (neutral* : Q neutral) (base* : (r : Idx) → (a : P r) → Q (base r a)) @@ -98,12 +98,12 @@ module _ (Idx : Type ℓ) (P : Idx → Type ℓ') (AGP : (r : Idx) → AbGroupSt where f : (x : ⊕HIT Idx P AGP) → Q x - f x = DS-Ind-Set.f Q (λ x → isProp→isSet (ispd x)) neutral* base* _add*_ - (λ {x} {y} {z} xs ys zs → toPathP (ispd _ (transport (λ i → Q (addAssoc x y z i)) (xs add* (ys add* zs))) ((xs add* ys) add* zs))) - (λ {x} xs → toPathP (ispd _ (transport (λ i → Q (addRid x i)) (xs add* neutral*)) xs)) - (λ {x} {y} xs ys → toPathP (ispd _ (transport (λ i → Q (addComm x y i)) (xs add* ys)) (ys add* xs))) - (λ r → toPathP (ispd _ (transport (λ i → Q (base-neutral r i)) (base* r (AbGroupStr.0g (AGP r)))) neutral*)) - (λ r a b → toPathP (ispd _ (transport (λ i → Q (base-add r a b i)) ((base* r a) add* (base* r b))) (base* r (AbGroupStr._+_ (AGP r) a b) ))) + f x = DS-Ind-Set.f Q (λ x → isProp→isSet (isPropQ x)) neutral* base* _add*_ + (λ {x} {y} {z} xs ys zs → toPathP (isPropQ _ (transport (λ i → Q (addAssoc x y z i)) (xs add* (ys add* zs))) ((xs add* ys) add* zs))) + (λ {x} xs → toPathP (isPropQ _ (transport (λ i → Q (addRid x i)) (xs add* neutral*)) xs)) + (λ {x} {y} xs ys → toPathP (isPropQ _ (transport (λ i → Q (addComm x y i)) (xs add* ys)) (ys add* xs))) + (λ r → toPathP (isPropQ _ (transport (λ i → Q (base-neutral r i)) (base* r (AbGroupStr.0g (AGP r)))) neutral*)) + (λ r a b → toPathP (isPropQ _ (transport (λ i → Q (base-add r a b i)) ((base* r a) add* (base* r b))) (base* r (AbGroupStr._+_ (AGP r) a b) ))) x From 3e57fdd827d4219aae9f5a74483487ceb06bda7d Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 29 Jul 2024 17:16:43 +0200 Subject: [PATCH 17/83] clean up --- Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda index 0419f097ab..3199f0e951 100644 --- a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda @@ -47,6 +47,7 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn A) where _ = snd A opaque + unfolding CommRing.quotientCommRingStr _/_ : CommAlgebra R ℓ _/_ = commAlgebraFromCommRing A/IAsCommRing @@ -178,7 +179,6 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn A) where _ : CommAlgebraStr R ⟨ B ⟩ _ = str B - {- trivial quotient -} module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) where open CommAlgebraStr (snd A) @@ -211,8 +211,6 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) where → (a : fst A) → fst (A / I) [_]/ = fst (quotientHom _ _) - - module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn A) where open CommIdeal using (isPropIsCommIdeal) @@ -231,7 +229,6 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn A) where _ ≡⟨ CommRing.kernel≡I {R = CommAlgebra→CommRing A} I ⟩ I ∎ - module _ {R : CommRing ℓ} {A : CommAlgebra R ℓ} From 79f932cd1215eb7df92166672ccb85da90c3a5c4 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 29 Jul 2024 17:17:00 +0200 Subject: [PATCH 18/83] WIP: deactivate problematic code --- .../Multivariate/EquivCarac/A[X]X-A.agda | 3 +- .../Multivariate/EquivCarac/An[X]X-A.agda | 62 +- .../EilenbergMacLane/Rings/KleinBottle.agda | 600 +++++++++--------- .../EilenbergMacLane/Rings/RP2.agda | 118 ++-- .../EilenbergMacLane/Rings/RP2wedgeS1.agda | 34 +- .../Cohomology/EilenbergMacLane/Rings/Sn.agda | 218 +++---- Cubical/ZCohomology/CohomologyRings/CP2.agda | 3 +- .../CohomologyRings/KleinBottle.agda | 3 +- Cubical/ZCohomology/CohomologyRings/RP2.agda | 3 +- .../CohomologyRings/RP2wedgeS1.agda | 3 +- Cubical/ZCohomology/CohomologyRings/S0.agda | 3 +- Cubical/ZCohomology/CohomologyRings/S1.agda | 3 +- .../CohomologyRings/S2wedgeS4.agda | 3 +- Cubical/ZCohomology/CohomologyRings/Sn.agda | 52 +- Cubical/ZCohomology/CohomologyRings/Unit.agda | 3 +- 15 files changed, 572 insertions(+), 539 deletions(-) diff --git a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/A[X]X-A.agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/A[X]X-A.agda index dc41f35a12..71f5b38c06 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/A[X]X-A.agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/A[X]X-A.agda @@ -190,7 +190,7 @@ module Properties-Equiv-QuotientXn-A A→A[x]/x : A → A[x]/x A→A[x]/x = [_] ∘ A→A[x] - +{- A→A[x]/x-pres+ : (a a' : A) → A→A[x]/x (a +A a') ≡ A→A[x]/x a +PAI A→A[x]/x a' A→A[x]/x-pres+ a a' = cong [_] (A→A[x]-pres+ a a') @@ -247,3 +247,4 @@ module _ Equiv-ℤ[X]/X-ℤ : RingEquiv (CommRing→Ring ℤ[X]/X) (CommRing→Ring ℤCR) Equiv-ℤ[X]/X-ℤ = Equiv-A[X]/X-A ℤCR +-- -} diff --git a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[X]X-A.agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[X]X-A.agda index 2d2f46f80f..f77fa7c596 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[X]X-A.agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[X]X-A.agda @@ -19,6 +19,7 @@ open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.Quotient +open import Cubical.Algebra.Monoid.Instances.NatVec open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤCommRing to ℤCR) @@ -216,37 +217,39 @@ module Properties-Equiv-QuotientXn-A ... | no ¬p = ⊥.rec (¬p refl) - +{- ----------------------------------------------------------------------------- -- Retraction open RingStr open IsRing - e-retr : (x : A[x1,···,xn]/ Ar n) → A→PAI (PAI→A x) ≡ x - e-retr = SQ.elimProp (λ _ → isSetPAI _ _) - (DS-Ind-Prop.f _ _ _ _ (λ _ → isSetPAI _ _) - base0-eq - base-eq - λ {U V} ind-U ind-V → cong [_] (A→PA-pres+ _ _) ∙ cong₂ _+PAI_ ind-U ind-V) - where - base0-eq : A→PAI (PAI→A [ 0PA ]) ≡ [ 0PA ] - base0-eq = cong [_] (base-neutral (replicate 0)) - - base-eq : (v : Vec ℕ n) → (a : A ) → [ A→PA (PA→A (base v a)) ] ≡ [ base v a ] - base-eq v a with (discreteVecℕn v (replicate 0)) - ... | 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) ∣₁ - where - helper : _ - helper = cong (λ X → X +PA base v (-A a)) (base-neutral (replicate 0)) - ∙ +PAIdL (base v (-A a)) - ∙ sym ( - genδ-FinVec-ℕLinearCombi ((A[X1,···,Xn] Ar n)) n k infkn (base v' (-A a)) ( Ar n) - ∙ cong₂ base (cong (λ X → v' +n-vec δℕ-Vec n X) (toFromId' n k infkn)) (·AIdR _) - ∙ cong (λ X → base X (-A a)) (sym eqvv')) + opaque + unfolding quotientCommRingStr + e-retr : (x : A[x1,···,xn]/ Ar n) → A→PAI (PAI→A x) ≡ x + e-retr = SQ.elimProp (λ _ → isSetPAI _ _) + (DS-Ind-Prop.f (NatVecMonoid n .fst) (λ _ → A) _ _ (λ _ → isSetPAI _ _) + base0-eq + base-eq + λ {U V} ind-U ind-V → cong [_] (A→PA-pres+ _ _) ∙ cong₂ _+PAI_ ind-U ind-V) + where + base0-eq : A→PAI (PAI→A [ 0PA ]) ≡ [ 0PA ] + base0-eq = cong [_] (base-neutral (replicate 0)) + + base-eq : (v : Vec ℕ n) → (a : A ) → [ A→PA (PA→A (base v a)) ] ≡ [ base v a ] + base-eq v a with (discreteVecℕn v (replicate 0)) + ... | 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) ∣₁ + where + helper : {!!} ≡ _ + helper = cong (λ X → X +PA base v (-A a)) (base-neutral (replicate 0)) + ∙ +PAIdL (base v (-A a)) + ∙ sym ( + genδ-FinVec-ℕLinearCombi ((A[X1,···,Xn] Ar n)) n k infkn (base v' (-A a)) ( Ar n) + ∙ cong₂ base (cong (λ X → v' +n-vec δℕ-Vec n X) (toFromId' n k infkn)) (·AIdR _) + ∙ cong (λ X → base X (-A a)) (sym eqvv')) @@ -254,17 +257,17 @@ module Properties-Equiv-QuotientXn-A -- Equiv module _ - (Ar@(A , Astr) : CommRing ℓ) + (R : CommRing ℓ) (n : ℕ) where open Iso - open Properties-Equiv-QuotientXn-A Ar n + open Properties-Equiv-QuotientXn-A R n - Equiv-QuotientX-A : CommRingEquiv (A[X1,···,Xn]/ Ar n) Ar + Equiv-QuotientX-A : CommRingEquiv (A[X1,···,Xn]/ R n) R fst Equiv-QuotientX-A = isoToEquiv is where - is : Iso (A[x1,···,xn]/ Ar n) A + is : Iso (A[x1,···,xn]/ R n) (R .fst) fun is = PAI→A inv is = A→PAI rightInv is = e-sect @@ -273,3 +276,4 @@ module _ -- Warning this doesn't prove Z[X]/X ≅ ℤ because you get two definition, -- see notation Multivariate-Quotient-notationZ for more details +-} diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/KleinBottle.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/KleinBottle.agda index 06ff30e2b5..d7bcb5b82e 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/KleinBottle.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/KleinBottle.agda @@ -853,312 +853,318 @@ isHomℤ/2[X,Y]→H*Klein = makeIsRingHom refl (λ _ _ → refl) ∙∙ H²K²→ℤ/2→H²K² (_⌣_ {G'' = ℤ/2Ring} K²gen.α K²gen.α) -- Map H*(K²) → ℤ/2[X,Y]/I -H*Klein→ℤ/2[X,Y]/I : - fst (H*R ℤ/2Ring KleinBottle) - → fst (CommRing→Ring ℤ/2[X,Y]/) -H*Klein→ℤ/2[X,Y]/I = - DS-Rec-Set.f _ _ _ _ squash/ [ neutral ] - HⁿKlein→ℤ/2[X,Y]/I _ - (+Assoc (snd (CommRing→Ring ℤ/2[X,Y]/))) - (+IdR (snd (CommRing→Ring ℤ/2[X,Y]/))) - (+Comm (snd (CommRing→Ring ℤ/2[X,Y]/))) - (λ { zero → cong [_] (base-neutral _) - ; one → cong [_] (cong₂ _add_ (base-neutral _) (base-neutral _) - ∙ addRid neutral) - ; two → cong [_] (cong (base (2 ∷ 0 ∷ [])) - (IsGroupHom.pres1 (snd (H²[K²,ℤ/2]≅ℤ/2*))) - ∙ base-neutral _) - ; (suc (suc (suc r))) → refl}) - λ { zero a b → cong [_] (base-add _ _ _ ∙ cong (base (0 ∷ 0 ∷ [])) - (sym (IsGroupHom.pres· (snd (H⁰[K²,ℤ/2]≅ℤ/2)) a b))) - ; one a b → cong [_] (move4 _ _ _ _ _add_ addAssoc addComm - ∙ cong₂ _add_ (base-add _ _ _ ∙ cong (base (1 ∷ 0 ∷ [])) - (cong fst (sym - (IsGroupHom.pres· (snd (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2)) a b)))) - (base-add _ _ _ ∙ cong (base (0 ∷ 1 ∷ [])) - (cong snd - (sym (IsGroupHom.pres· (snd (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2)) a b))))) - ; two a b → cong [_] (base-add _ _ _ ∙ cong (base (2 ∷ 0 ∷ [])) - (sym (IsGroupHom.pres· (snd (H²[K²,ℤ/2]≅ℤ/2*)) a b))) - ; (suc (suc (suc n))) → λ a b → cong [_] (addRid neutral)} +opaque + unfolding quotientCommRingStr + H*Klein→ℤ/2[X,Y]/I : + fst (H*R ℤ/2Ring KleinBottle) + → fst (CommRing→Ring ℤ/2[X,Y]/) + H*Klein→ℤ/2[X,Y]/I = + DS-Rec-Set.f _ _ _ _ squash/ [ neutral ] + HⁿKlein→ℤ/2[X,Y]/I _ + (+Assoc (snd (CommRing→Ring ℤ/2[X,Y]/))) + (+IdR (snd (CommRing→Ring ℤ/2[X,Y]/))) + (+Comm (snd (CommRing→Ring ℤ/2[X,Y]/))) + (λ { zero → cong [_] (base-neutral _) + ; one → cong [_] (cong₂ _add_ (base-neutral _) (base-neutral _) + ∙ addRid neutral) + ; two → cong [_] (cong (base (2 ∷ 0 ∷ [])) + (IsGroupHom.pres1 (snd (H²[K²,ℤ/2]≅ℤ/2*))) + ∙ base-neutral _) + ; (suc (suc (suc r))) → refl}) + λ { zero a b → cong [_] (base-add _ _ _ ∙ cong (base (0 ∷ 0 ∷ [])) + (sym (IsGroupHom.pres· (snd (H⁰[K²,ℤ/2]≅ℤ/2)) a b))) + ; one a b → cong [_] (move4 _ _ _ _ _add_ addAssoc addComm + ∙ cong₂ _add_ (base-add _ _ _ ∙ cong (base (1 ∷ 0 ∷ [])) + (cong fst (sym + (IsGroupHom.pres· (snd (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2)) a b)))) + (base-add _ _ _ ∙ cong (base (0 ∷ 1 ∷ [])) + (cong snd + (sym (IsGroupHom.pres· (snd (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2)) a b))))) + ; two a b → cong [_] (base-add _ _ _ ∙ cong (base (2 ∷ 0 ∷ [])) + (sym (IsGroupHom.pres· (snd (H²[K²,ℤ/2]≅ℤ/2*)) a b))) + ; (suc (suc (suc n))) → λ a b → cong [_] (addRid neutral)} -- The equivalence -ℤ/2[X,Y]/≅H*KleinBottle : - RingEquiv (CommRing→Ring ℤ/2[X,Y]/) - (H*R ℤ/2Ring KleinBottle) -fst ℤ/2[X,Y]/≅H*KleinBottle = isoToEquiv is - where - is : Iso _ _ - fun is = ℤ/2[X,Y]/I→H*Klein .fst - inv is = H*Klein→ℤ/2[X,Y]/I - rightInv is = DS-Ind-Prop.f _ _ _ _ - (λ _ → trunc _ _) - refl - (λ { zero a → lem₀ a _ refl - ; one a → lem₁ a _ _ refl - ; two a → lem₂ a _ refl - ; (suc (suc (suc r))) a → - sym (base-neutral _) - ∙ cong (base (3 + r)) - (isContr→isProp (isContr-HⁿKleinBottle r ℤ/2) - (0ₕ (3 + r)) a)}) - λ {x} {y} ind1 ind2 - → IsRingHom.pres+ (ℤ/2[X,Y]/I→H*Klein .snd) - (H*Klein→ℤ/2[X,Y]/I x) (H*Klein→ℤ/2[X,Y]/I y) - ∙ cong₂ _add_ ind1 ind2 +{- +opaque + unfolding quotientCommRingStr H*Klein→ℤ/2[X,Y]/I + ℤ/2[X,Y]/≅H*KleinBottle : + RingEquiv (CommRing→Ring ℤ/2[X,Y]/) + (H*R ℤ/2Ring KleinBottle) + fst ℤ/2[X,Y]/≅H*KleinBottle = isoToEquiv is where - lem₀ : (a : _) (x : _) - → H⁰[K²,ℤ/2]≅ℤ/2 .fst .fst a ≡ x - → ℤ/2[X,Y]/I→H*Klein .fst (H*Klein→ℤ/2[X,Y]/I (base zero a)) - ≡ base zero a - lem₀ a = - ℤ/2-elim - (λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ H*Klein→ℤ/2[X,Y]/I) (help id) - ∙ sym (help id)) - λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst) - (cong [_] (cong (base (0 ∷ 0 ∷ [])) id)) - ∙ cong (base zero) - (sym (cong (invEq (H⁰[K²,ℤ/2]≅ℤ/2 .fst)) id) - ∙ retEq (fst H⁰[K²,ℤ/2]≅ℤ/2) a) + is : Iso _ _ + fun is = ℤ/2[X,Y]/I→H*Klein .fst + inv is = H*Klein→ℤ/2[X,Y]/I + rightInv is = DS-Ind-Prop.f _ _ _ _ + (λ _ → trunc _ _) + refl + (λ { zero a → lem₀ a _ refl + ; one a → lem₁ a _ _ refl + ; two a → lem₂ a _ refl + ; (suc (suc (suc r))) a → + sym (base-neutral _) + ∙ cong (base (3 + r)) + (isContr→isProp (isContr-HⁿKleinBottle r ℤ/2) + (0ₕ (3 + r)) a)}) + λ {x} {y} ind1 ind2 + → IsRingHom.pres+ (ℤ/2[X,Y]/I→H*Klein .snd) + (H*Klein→ℤ/2[X,Y]/I x) (H*Klein→ℤ/2[X,Y]/I y) + ∙ cong₂ _add_ ind1 ind2 where - help : H⁰[K²,ℤ/2]≅ℤ/2 .fst .fst a ≡ 0 - → Path (fst (H*R ℤ/2Ring KleinBottle)) (base zero a) neutral - help id' = - sym (cong (base zero) - (sym (cong (invEq (H⁰[K²,ℤ/2]≅ℤ/2 .fst)) id' - ∙ IsGroupHom.pres1 (isGroupHomInv (H⁰[K²,ℤ/2]≅ℤ/2))) - ∙ retEq (fst H⁰[K²,ℤ/2]≅ℤ/2) a)) - ∙ base-neutral zero - - lem₁ : (a : _) → (x y : _) - → H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst .fst a ≡ (x , y) - → ℤ/2[X,Y]/I→H*Klein .fst (H*Klein→ℤ/2[X,Y]/I (base one a)) - ≡ base one a - lem₁ a = - ℤ/2-elim - (ℤ/2-elim - (λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) - (cong₂ _add_ (cong (base (1 ∷ 0 ∷ [])) - (cong fst id)) - (cong (base (0 ∷ 1 ∷ [])) - (cong snd id))) - ∙ addRid neutral - ∙ sym (help a id)) - λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) - (cong₂ _add_ - (cong (base (1 ∷ 0 ∷ [])) (cong fst id) - ∙ base-neutral _) - (cong (base (0 ∷ 1 ∷ [])) (cong snd id)) - ∙ addComm _ _ ∙ addRid _) - ∙∙ cong (base 1) (1ₕ-⌣ 1 K²gen.β) - ∙∙ cong (base 1) (sym (retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) K²gen.β) - ∙∙ cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) (sym id) - ∙∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a)) - (ℤ/2-elim - (λ id → (cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) - (cong₂ _add_ - (cong (base (1 ∷ 0 ∷ [])) (cong fst id)) - (cong (base (0 ∷ 1 ∷ [])) (cong snd id) ∙ base-neutral _) - ∙ addRid _) - ∙ cong (base 1) - ( (⌣-1ₕ 1 K²gen.α ∙ transportRefl K²gen.α) - ∙ (sym (retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) K²gen.α) - ∙∙ cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) (α↦1 ∙ sym id) - ∙∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a)))) - λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) - (cong₂ _add_ - (cong (base (1 ∷ 0 ∷ [])) (cong fst id)) - (cong (base (0 ∷ 1 ∷ [])) (cong snd id))) - ∙ IsRingHom.pres+ (snd ℤ/2[X,Y]/I→H*Klein) - [ base (1 ∷ 0 ∷ []) 1 ] [ base (0 ∷ 1 ∷ []) 1 ] - ∙ cong₂ _add_ - (cong (base one) (⌣-1ₕ 1 (incL 1) ∙ transportRefl K²gen.α)) - (cong (base one) (1ₕ-⌣ 1 (incR 1))) - ∙ base-add 1 K²gen.α K²gen.β - ∙ cong (base one) - (sym (retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) (K²gen.α +ₕ K²gen.β)) - ∙∙ (cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) (sym id)) - ∙∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a)) - where - help : (a : _) → H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst .fst a ≡ (0 , 0) - → Path (fst (H*R ℤ/2Ring KleinBottle)) (base one a) neutral - help a p = - (sym (cong (base one) - (sym (cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) p - ∙ IsGroupHom.pres1 (isGroupHomInv (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2))) - ∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a))) - ∙ base-neutral one - - lem₂ : (a : _) (x : _) → H²[K²,ℤ/2]≅ℤ/2* .fst .fst a ≡ x - → ℤ/2[X,Y]/I→H*Klein .fst (H*Klein→ℤ/2[X,Y]/I (base two a)) - ≡ base two a - lem₂ a = - ℤ/2-elim - (λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ H*Klein→ℤ/2[X,Y]/I) - (cong (base 2) (help1 id) ∙ base-neutral _) - ∙∙ sym (base-neutral _) - ∙∙ cong (base 2) (sym (help1 id))) + lem₀ : (a : _) (x : _) + → H⁰[K²,ℤ/2]≅ℤ/2 .fst .fst a ≡ x + → ℤ/2[X,Y]/I→H*Klein .fst (H*Klein→ℤ/2[X,Y]/I (base zero a)) + ≡ base zero a + lem₀ a = + ℤ/2-elim + (λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ H*Klein→ℤ/2[X,Y]/I) (help id) + ∙ sym (help id)) λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst) - (cong [_] (cong (base (2 ∷ 0 ∷ [])) - (cong (H²[K²,ℤ/2]≅ℤ/2* .fst .fst) (help2 id) - ∙ α²↦1') )) - ∙∙ cong (base 2) (⌣-1ₕ 2 (incL 2) ∙ transportRefl (incL 2)) - ∙∙ cong (base two) (sym (help2 id)) - where - help1 : H²[K²,ℤ/2]≅ℤ/2* .fst .fst a ≡ 0 → a ≡ 0ₕ 2 - help1 p = sym (retEq (H²[K²,ℤ/2]≅ℤ/2* .fst) a) - ∙ cong (invEq (H²[K²,ℤ/2]≅ℤ/2* .fst)) p - ∙ IsGroupHom.pres1 (isGroupHomInv H²[K²,ℤ/2]≅ℤ/2*) - - help2 : H²[K²,ℤ/2]≅ℤ/2* .fst .fst a ≡ 1 → a ≡ α⌣α - help2 p = sym (retEq (H²[K²,ℤ/2]≅ℤ/2* .fst) a) - ∙∙ cong (invEq (H²[K²,ℤ/2]≅ℤ/2* .fst)) (p ∙ sym α²↦1') - ∙∙ retEq (H²[K²,ℤ/2]≅ℤ/2* .fst) α⌣α - - leftInv is = - SQ.elimProp - (λ _ → squash/ _ _) - (DS-Ind-Prop.f _ _ _ _ + (cong [_] (cong (base (0 ∷ 0 ∷ [])) id)) + ∙ cong (base zero) + (sym (cong (invEq (H⁰[K²,ℤ/2]≅ℤ/2 .fst)) id) + ∙ retEq (fst H⁰[K²,ℤ/2]≅ℤ/2) a) + where + help : H⁰[K²,ℤ/2]≅ℤ/2 .fst .fst a ≡ 0 + → Path (fst (H*R ℤ/2Ring KleinBottle)) (base zero a) neutral + help id' = + sym (cong (base zero) + (sym (cong (invEq (H⁰[K²,ℤ/2]≅ℤ/2 .fst)) id' + ∙ IsGroupHom.pres1 (isGroupHomInv (H⁰[K²,ℤ/2]≅ℤ/2))) + ∙ retEq (fst H⁰[K²,ℤ/2]≅ℤ/2) a)) + ∙ base-neutral zero + + lem₁ : (a : _) → (x y : _) + → H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst .fst a ≡ (x , y) + → ℤ/2[X,Y]/I→H*Klein .fst (H*Klein→ℤ/2[X,Y]/I (base one a)) + ≡ base one a + lem₁ a = + ℤ/2-elim + (ℤ/2-elim + (λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) + (cong₂ _add_ (cong (base (1 ∷ 0 ∷ [])) + (cong fst id)) + (cong (base (0 ∷ 1 ∷ [])) + (cong snd id))) + ∙ addRid neutral + ∙ sym (help a id)) + λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) + (cong₂ _add_ + (cong (base (1 ∷ 0 ∷ [])) (cong fst id) + ∙ base-neutral _) + (cong (base (0 ∷ 1 ∷ [])) (cong snd id)) + ∙ addComm _ _ ∙ addRid _) + ∙∙ cong (base 1) (1ₕ-⌣ 1 K²gen.β) + ∙∙ cong (base 1) (sym (retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) K²gen.β) + ∙∙ cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) (sym id) + ∙∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a)) + (ℤ/2-elim + (λ id → (cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) + (cong₂ _add_ + (cong (base (1 ∷ 0 ∷ [])) (cong fst id)) + (cong (base (0 ∷ 1 ∷ [])) (cong snd id) ∙ base-neutral _) + ∙ addRid _) + ∙ cong (base 1) + ( (⌣-1ₕ 1 K²gen.α ∙ transportRefl K²gen.α) + ∙ (sym (retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) K²gen.α) + ∙∙ cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) (α↦1 ∙ sym id) + ∙∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a)))) + λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) + (cong₂ _add_ + (cong (base (1 ∷ 0 ∷ [])) (cong fst id)) + (cong (base (0 ∷ 1 ∷ [])) (cong snd id))) + ∙ IsRingHom.pres+ (snd ℤ/2[X,Y]/I→H*Klein) + [ base (1 ∷ 0 ∷ []) 1 ] [ base (0 ∷ 1 ∷ []) 1 ] + ∙ cong₂ _add_ + (cong (base one) (⌣-1ₕ 1 (incL 1) ∙ transportRefl K²gen.α)) + (cong (base one) (1ₕ-⌣ 1 (incR 1))) + ∙ base-add 1 K²gen.α K²gen.β + ∙ cong (base one) + (sym (retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) (K²gen.α +ₕ K²gen.β)) + ∙∙ (cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) (sym id)) + ∙∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a)) + where + help : (a : _) → H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst .fst a ≡ (0 , 0) + → Path (fst (H*R ℤ/2Ring KleinBottle)) (base one a) neutral + help a p = + (sym (cong (base one) + (sym (cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) p + ∙ IsGroupHom.pres1 (isGroupHomInv (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2))) + ∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a))) + ∙ base-neutral one + + lem₂ : (a : _) (x : _) → H²[K²,ℤ/2]≅ℤ/2* .fst .fst a ≡ x + → ℤ/2[X,Y]/I→H*Klein .fst (H*Klein→ℤ/2[X,Y]/I (base two a)) + ≡ base two a + lem₂ a = + ℤ/2-elim + (λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ H*Klein→ℤ/2[X,Y]/I) + (cong (base 2) (help1 id) ∙ base-neutral _) + ∙∙ sym (base-neutral _) + ∙∙ cong (base 2) (sym (help1 id))) + λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst) + (cong [_] (cong (base (2 ∷ 0 ∷ [])) + (cong (H²[K²,ℤ/2]≅ℤ/2* .fst .fst) (help2 id) + ∙ α²↦1') )) + ∙∙ cong (base 2) (⌣-1ₕ 2 (incL 2) ∙ transportRefl (incL 2)) + ∙∙ cong (base two) (sym (help2 id)) + where + help1 : H²[K²,ℤ/2]≅ℤ/2* .fst .fst a ≡ 0 → a ≡ 0ₕ 2 + help1 p = sym (retEq (H²[K²,ℤ/2]≅ℤ/2* .fst) a) + ∙ cong (invEq (H²[K²,ℤ/2]≅ℤ/2* .fst)) p + ∙ IsGroupHom.pres1 (isGroupHomInv H²[K²,ℤ/2]≅ℤ/2*) + + help2 : H²[K²,ℤ/2]≅ℤ/2* .fst .fst a ≡ 1 → a ≡ α⌣α + help2 p = sym (retEq (H²[K²,ℤ/2]≅ℤ/2* .fst) a) + ∙∙ cong (invEq (H²[K²,ℤ/2]≅ℤ/2* .fst)) (p ∙ sym α²↦1') + ∙∙ retEq (H²[K²,ℤ/2]≅ℤ/2* .fst) α⌣α + + leftInv is = + SQ.elimProp (λ _ → squash/ _ _) - refl - (λ r a → main a r) - λ {x} {y} ind1 ind2 - → cong₂ (_+r_ (snd (CommRing→Ring ℤ/2[X,Y]/))) - ind1 ind2) - where - clem : (x y : ℕ) - → H*Klein→ℤ/2[X,Y]/I - (ℤ/2[X,Y]/I→H*Klein .fst - [ base (suc (suc (suc x)) ∷ y ∷ []) 1 ]) - ≡ [ neutral ] - clem x zero = refl - clem x (suc n) = refl - - help : (y : ℕ) → - Path (fst (CommRing→Ring ℤ/2[X,Y]/)) - [ base (one ∷ suc (suc y) ∷ []) 1 ] - [ neutral ] - help y = eq/ _ _ - ∣ (λ { zero → neutral - ; one → base (1 ∷ y ∷ []) 1 - ; two → neutral}) - , (sym (addRid _) - ∙ addComm (base (1 ∷ suc (suc y) ∷ []) 1 add neutral) neutral) - ∙ (λ i → neutral add (base (1 ∷ (+-comm 2 y i) ∷ []) 1 - add (addRid neutral (~ i)))) ∣₁ - - help2 : (y : ℕ) - → ℤ/2[X,Y]/I→H*Klein .fst [ base (zero ∷ suc (suc y) ∷ []) 1 ] - ≡ neutral - help2 zero = cong (base 2) (1ₕ-⌣ 2 (incR two)) - ∙ base-neutral _ - help2 (suc y) = base-neutral _ - - help3 : (x y : ℕ) - → ℤ/2[X,Y]/I→H*Klein .fst [ base (x ∷ suc (suc y) ∷ []) 1 ] - ≡ neutral - help3 zero y = help2 y - help3 (suc x) y = - (λ i → base (suc (suc (+-suc x y i))) - (transp (λ j → coHom (suc (suc (+-suc x y (i ∧ j)))) - ℤ/2 KleinBottle) (~ i) - (_⌣_ {G'' = ℤ/2Ring} (incL (suc x)) (incR (suc (suc y)))))) - ∙ cong (base (suc (suc (suc (x + y))))) - (isContr→isProp (isContr-HⁿKleinBottle (x + y) ℤ/2) _ _) - ∙ base-neutral _ - - main-1 : (r : _) - → H*Klein→ℤ/2[X,Y]/I (ℤ/2[X,Y]/I→H*Klein .fst [ base r 1 ]) - ≡ [ base r 1 ] - main-1 (zero ∷ zero ∷ []) = refl - main-1 (zero ∷ one ∷ []) = - cong (H*Klein→ℤ/2[X,Y]/I) - (cong (base 1) (1ₕ-⌣ 1 (incR 1))) - ∙ cong [_] (cong₂ _add_ (base-neutral _) - (λ _ → base (0 ∷ 1 ∷ []) 1) - ∙ addComm _ _ ∙ addRid _) - main-1 (zero ∷ suc (suc y) ∷ []) = - cong H*Klein→ℤ/2[X,Y]/I (help2 y) - ∙ eq/ _ _ - ∣ (λ {zero → neutral - ; one → base (0 ∷ (y ∷ [])) 1 - ; two → neutral}) - , cong (neutral add_) - (((λ i → base (0 ∷ (+-comm 2 y i) ∷ []) 1) - ∙ sym (addRid (base (0 ∷ (y + 2) ∷ []) 1))) - ∙ cong (base (0 ∷ (y + 2) ∷ []) 1 add_) (sym (addRid _))) ∣₁ - main-1 (one ∷ zero ∷ []) = - cong H*Klein→ℤ/2[X,Y]/I - (cong (base 1) (⌣-1ₕ 1 (incL one) ∙ transportRefl _)) - ∙ cong [_] (cong₂ _add_ (cong (base (1 ∷ 0 ∷ [])) - (cong fst α↦1)) - (base-neutral _) - ∙ addRid _) - main-1 (one ∷ one ∷ []) = - cong [_] (cong (base (2 ∷ 0 ∷ [])) αβ↦1') - ∙ eq/ _ _ ∣ (λ {zero → neutral - ; one → neutral - ; two → base (0 ∷ 0 ∷ []) 1}) - , ((addComm _ _ - ∙ sym (addRid _) - ∙ addComm (base (1 ∷ 1 ∷ []) 1 - add (base (2 ∷ 0 ∷ []) 1)) - neutral - ∙ sym (addRid _) - ∙ addComm (neutral add (base (1 ∷ 1 ∷ []) 1 - add (base (2 ∷ 0 ∷ []) 1))) neutral) - ∙ λ i → neutral add - (neutral add (addRid - (base (1 ∷ 1 ∷ []) 1 - add (base (2 ∷ 0 ∷ []) 1)) (~ i)))) ∣₁ - main-1 (one ∷ suc (suc y) ∷ []) = - cong H*Klein→ℤ/2[X,Y]/I (help3 one y) ∙ sym (help y) - main-1 (two ∷ zero ∷ []) = - cong [_] (cong (base (2 ∷ 0 ∷ [])) - (cong (H²[K²,ℤ/2]≅ℤ/2* .fst .fst) - (⌣-1ₕ 2 (incL 2) ∙ transportRefl _) - ∙ α²↦1')) - main-1 (two ∷ suc y ∷ []) = - eq/ neutral _ - ∣ (λ {zero → base (0 ∷ y ∷ []) 1 - ; one → neutral - ; two → base (1 ∷ y ∷ []) 1}) - , ((addComm _ _ ∙ addRid _ - ∙ ((((λ i → base (2 ∷ +-comm 1 y i ∷ []) 1) - ∙ sym (addRid _)) - ∙ cong (base (2 ∷ y + 1 ∷ []) (fsuc fzero) add_) - (sym (base-neutral _) - ∙ sym (base-add (3 ∷ y + 0 ∷ []) 1 1))) - ∙ addComm _ _) - ∙ sym (addAssoc _ _ _)) - ∙∙ cong (base (3 ∷ y + 0 ∷ []) (fsuc fzero) add_) - (addComm _ _ - ∙ sym (addComm _ _ ∙ addRid _)) - ∙∙ (λ i → base (3 ∷ (y + 0) ∷ []) 1 - add (neutral - add addRid (base (2 ∷ (y + 1) ∷ []) 1 - add base (3 ∷ (y + 0) ∷ []) 1 ) (~ i)))) ∣₁ - main-1 (suc (suc (suc x)) ∷ y ∷ []) = - clem x y - ∙ eq/ neutral _ - ∣ (λ {zero → base (x ∷ y ∷ []) 1 - ; one → neutral - ; two → neutral}) - , ((addComm neutral (base ((3 + x) ∷ y ∷ []) 1) - ∙ cong (base ((3 + x) ∷ y ∷ []) 1 add_) (sym (addRid neutral))) - ∙ λ i → base ((+-comm 3 x i) ∷ (+-comm 0 y i) ∷ []) 1 - add (neutral add (addRid neutral (~ i)))) ∣₁ - - main : (a : ℤ/2 .fst) (r : _) - → H*Klein→ℤ/2[X,Y]/I (ℤ/2[X,Y]/I→H*Klein .fst [ base r a ]) ≡ [ base r a ] - main = ℤ/2-elim (λ r → cong (H*Klein→ℤ/2[X,Y]/I ∘ ℤ/2[X,Y]/I→H*Klein .fst) - (cong [_] (base-neutral r)) - ∙ cong [_] (sym (base-neutral r))) - main-1 -snd ℤ/2[X,Y]/≅H*KleinBottle = ℤ/2[X,Y]/I→H*Klein .snd + (DS-Ind-Prop.f _ _ _ _ + (λ _ → squash/ _ _) + refl + (λ r a → main a r) + λ {x} {y} ind1 ind2 + → cong₂ (_+r_ (snd (CommRing→Ring ℤ/2[X,Y]/))) + ind1 ind2) + where + clem : (x y : ℕ) + → H*Klein→ℤ/2[X,Y]/I + (ℤ/2[X,Y]/I→H*Klein .fst + [ base (suc (suc (suc x)) ∷ y ∷ []) 1 ]) + ≡ [ neutral ] + clem x zero = refl + clem x (suc n) = refl + + help : (y : ℕ) → + Path (fst (CommRing→Ring ℤ/2[X,Y]/)) + [ base (one ∷ suc (suc y) ∷ []) 1 ] + [ neutral ] + help y = eq/ _ _ + ∣ (λ { zero → neutral + ; one → base (1 ∷ y ∷ []) 1 + ; two → neutral}) + , (sym (addRid _) + ∙ addComm (base (1 ∷ suc (suc y) ∷ []) 1 add neutral) neutral) + ∙ (λ i → neutral add (base (1 ∷ (+-comm 2 y i) ∷ []) 1 + add (addRid neutral (~ i)))) ∣₁ + + help2 : (y : ℕ) + → ℤ/2[X,Y]/I→H*Klein .fst [ base (zero ∷ suc (suc y) ∷ []) 1 ] + ≡ neutral + help2 zero = cong (base 2) (1ₕ-⌣ 2 (incR two)) + ∙ base-neutral _ + help2 (suc y) = base-neutral _ + + help3 : (x y : ℕ) + → ℤ/2[X,Y]/I→H*Klein .fst [ base (x ∷ suc (suc y) ∷ []) 1 ] + ≡ neutral + help3 zero y = help2 y + help3 (suc x) y = + (λ i → base (suc (suc (+-suc x y i))) + (transp (λ j → coHom (suc (suc (+-suc x y (i ∧ j)))) + ℤ/2 KleinBottle) (~ i) + (_⌣_ {G'' = ℤ/2Ring} (incL (suc x)) (incR (suc (suc y)))))) + ∙ cong (base (suc (suc (suc (x + y))))) + (isContr→isProp (isContr-HⁿKleinBottle (x + y) ℤ/2) _ _) + ∙ base-neutral _ + + main-1 : (r : _) + → H*Klein→ℤ/2[X,Y]/I (ℤ/2[X,Y]/I→H*Klein .fst [ base r 1 ]) + ≡ [ base r 1 ] + main-1 (zero ∷ zero ∷ []) = refl + main-1 (zero ∷ one ∷ []) = + cong (H*Klein→ℤ/2[X,Y]/I) + (cong (base 1) (1ₕ-⌣ 1 (incR 1))) + ∙ cong [_] (cong₂ _add_ (base-neutral _) + (λ _ → base (0 ∷ 1 ∷ []) 1) + ∙ addComm _ _ ∙ addRid _) + main-1 (zero ∷ suc (suc y) ∷ []) = + cong H*Klein→ℤ/2[X,Y]/I (help2 y) + ∙ eq/ _ _ + ∣ (λ {zero → neutral + ; one → base (0 ∷ (y ∷ [])) 1 + ; two → neutral}) + , cong (neutral add_) + (((λ i → base (0 ∷ (+-comm 2 y i) ∷ []) 1) + ∙ sym (addRid (base (0 ∷ (y + 2) ∷ []) 1))) + ∙ cong (base (0 ∷ (y + 2) ∷ []) 1 add_) (sym (addRid _))) ∣₁ + main-1 (one ∷ zero ∷ []) = + cong H*Klein→ℤ/2[X,Y]/I + (cong (base 1) (⌣-1ₕ 1 (incL one) ∙ transportRefl _)) + ∙ cong [_] (cong₂ _add_ (cong (base (1 ∷ 0 ∷ [])) + (cong fst α↦1)) + (base-neutral _) + ∙ addRid _) + main-1 (one ∷ one ∷ []) = + cong [_] (cong (base (2 ∷ 0 ∷ [])) αβ↦1') + ∙ eq/ _ _ ∣ (λ {zero → neutral + ; one → neutral + ; two → base (0 ∷ 0 ∷ []) 1}) + , ((addComm _ _ + ∙ sym (addRid _) + ∙ addComm (base (1 ∷ 1 ∷ []) 1 + add (base (2 ∷ 0 ∷ []) 1)) + neutral + ∙ sym (addRid _) + ∙ addComm (neutral add (base (1 ∷ 1 ∷ []) 1 + add (base (2 ∷ 0 ∷ []) 1))) neutral) + ∙ λ i → neutral add + (neutral add (addRid + (base (1 ∷ 1 ∷ []) 1 + add (base (2 ∷ 0 ∷ []) 1)) (~ i)))) ∣₁ + main-1 (one ∷ suc (suc y) ∷ []) = + cong H*Klein→ℤ/2[X,Y]/I (help3 one y) ∙ sym (help y) + main-1 (two ∷ zero ∷ []) = + cong [_] (cong (base (2 ∷ 0 ∷ [])) + (cong (H²[K²,ℤ/2]≅ℤ/2* .fst .fst) + (⌣-1ₕ 2 (incL 2) ∙ transportRefl _) + ∙ α²↦1')) + main-1 (two ∷ suc y ∷ []) = + eq/ neutral _ + ∣ (λ {zero → base (0 ∷ y ∷ []) 1 + ; one → neutral + ; two → base (1 ∷ y ∷ []) 1}) + , ((addComm _ _ ∙ addRid _ + ∙ ((((λ i → base (2 ∷ +-comm 1 y i ∷ []) 1) + ∙ sym (addRid _)) + ∙ cong (base (2 ∷ y + 1 ∷ []) (fsuc fzero) add_) + (sym (base-neutral _) + ∙ sym (base-add (3 ∷ y + 0 ∷ []) 1 1))) + ∙ addComm _ _) + ∙ sym (addAssoc _ _ _)) + ∙∙ cong (base (3 ∷ y + 0 ∷ []) (fsuc fzero) add_) + (addComm _ _ + ∙ sym (addComm _ _ ∙ addRid _)) + ∙∙ (λ i → base (3 ∷ (y + 0) ∷ []) 1 + add (neutral + add addRid (base (2 ∷ (y + 1) ∷ []) 1 + add base (3 ∷ (y + 0) ∷ []) 1 ) (~ i)))) ∣₁ + main-1 (suc (suc (suc x)) ∷ y ∷ []) = + clem x y + ∙ eq/ neutral _ + ∣ (λ {zero → base (x ∷ y ∷ []) 1 + ; one → neutral + ; two → neutral}) + , ((addComm neutral (base ((3 + x) ∷ y ∷ []) 1) + ∙ cong (base ((3 + x) ∷ y ∷ []) 1 add_) (sym (addRid neutral))) + ∙ λ i → base ((+-comm 3 x i) ∷ (+-comm 0 y i) ∷ []) 1 + add (neutral add (addRid neutral (~ i)))) ∣₁ + + main : (a : ℤ/2 .fst) (r : _) + → H*Klein→ℤ/2[X,Y]/I (ℤ/2[X,Y]/I→H*Klein .fst [ base r a ]) ≡ [ base r a ] + main = ℤ/2-elim (λ r → cong (H*Klein→ℤ/2[X,Y]/I ∘ ℤ/2[X,Y]/I→H*Klein .fst) + (cong [_] (base-neutral r)) + ∙ cong [_] (sym (base-neutral r))) + main-1 + snd ℤ/2[X,Y]/≅H*KleinBottle = ℤ/2[X,Y]/I→H*Klein .snd H*KleinBottle≅ℤ/2[X,Y]/ : RingEquiv (H*R ℤ/2Ring KleinBottle) (CommRing→Ring ℤ/2[X,Y]/) H*KleinBottle≅ℤ/2[X,Y]/ = RingEquivs.invRingEquiv ℤ/2[X,Y]/≅H*KleinBottle +-- -} diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/RP2.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/RP2.agda index db570dcc10..2ca66de1ff 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/RP2.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/RP2.agda @@ -379,64 +379,66 @@ module _ {ℓ : Level} (n : ℕ) where H*[RP²,ℤ₂]→ℤ₂[X]/-fun : (r : ℕ) → coHom r ℤ/2 RP² → ℤ₂[X]/ .fst H*[RP²,ℤ₂]→ℤ₂[X]/-fun r x = [ base (r ∷ []) (H*[RP²,ℤ₂]→ℤ₂[X]/-fun' r x) ] - H*[RP²,ℤ₂]→ℤ₂[X]/ : H*R ℤ/2Ring RP² .fst → ℤ₂[X]/ .fst - H*[RP²,ℤ₂]→ℤ₂[X]/ = DS-Rec-Set.f _ _ _ _ squash/ - [ neutral ] - H*[RP²,ℤ₂]→ℤ₂[X]/-fun - (CommRingStr._+_ (snd ℤ₂[X]/)) - (CommRingStr.+Assoc (snd ℤ₂[X]/)) - (CommRingStr.+IdR (snd ℤ₂[X]/)) - (CommRingStr.+Comm (snd ℤ₂[X]/)) - (λ { zero → cong [_] (base-neutral (0 ∷ [])) - ; one → cong [_] (cong (base (1 ∷ [])) (IsGroupHom.pres1 (snd (H¹[RP²,ℤ/2]≅ℤ/2))) - ∙ base-neutral (1 ∷ [])) - ; two → cong [_] (cong (base (2 ∷ [])) (IsGroupHom.pres1 (snd (H²[RP²,ℤ/2]≅ℤ/2))) - ∙ base-neutral (2 ∷ [])) - ; (suc (suc (suc r))) → cong [_] (base-neutral _)}) - (λ { zero a b → cong [_] (base-add (0 ∷ []) _ _ - ∙ cong (base (0 ∷ [])) - (sym (IsGroupHom.pres· (snd (H⁰[RP²,ℤ/2]≅ℤ/2)) a b))) - ; one a b → cong [_] (base-add (1 ∷ []) _ _ - ∙ cong (base (1 ∷ [])) - (sym (IsGroupHom.pres· (snd (H¹[RP²,ℤ/2]≅ℤ/2)) a b))) - ; two a b → cong [_] (base-add (2 ∷ []) _ _ - ∙ cong (base (2 ∷ [])) - (sym (IsGroupHom.pres· (snd (H²[RP²,ℤ/2]≅ℤ/2)) a b))) - ; (suc (suc (suc r))) a b - → cong [_] (cong₂ _add_ refl (base-neutral (3 +ℕ r ∷ [])) - ∙ addRid _)}) - - ℤ₂[X]/→H*[RP²,ℤ₂]→ℤ₂[X]/ : (x : _) - → H*[RP²,ℤ₂]→ℤ₂[X]/ (ℤ₂[X]/→H*[RP²,ℤ₂] .fst x) ≡ x - ℤ₂[X]/→H*[RP²,ℤ₂]→ℤ₂[X]/ = SQ.elimProp (λ _ → squash/ _ _) - (DS-Ind-Prop.f _ _ _ _ (λ _ → squash/ _ _) - refl - (λ { (zero ∷ []) a → cong [_] (cong (base (zero ∷ [])) - (secEq (H⁰[RP²,ℤ/2]≅ℤ/2 .fst) a)) - ; (one ∷ []) a → cong [_] (cong (base (one ∷ [])) - (secEq (H¹[RP²,ℤ/2]≅ℤ/2 .fst) a)) - ; (two ∷ []) a → cong [_] (cong (base (two ∷ [])) - (secEq (H²[RP²,ℤ/2]≅ℤ/2 .fst) a)) - ; (suc (suc (suc r)) ∷ []) → ℤ/2-elim refl - (sym (eq/ _ _ ∣ (λ {zero → base (r ∷ []) fone}) - , ((cong₂ _add_ refl (base-neutral _) - ∙ addRid _ - ∙ λ i → base (+-comm 3 r i ∷ []) fone)) - ∙ sym (addRid _) ∣₁))}) - (λ p q → cong₂ (CommRingStr._+_ (snd ℤ₂[X]/)) p q)) - - H*[RP²,ℤ₂]→ℤ₂[X]/→H*[RP²,ℤ₂] : (x : H*R ℤ/2Ring RP² .fst) - → ℤ₂[X]/→H*[RP²,ℤ₂] .fst (H*[RP²,ℤ₂]→ℤ₂[X]/ x) ≡ x - H*[RP²,ℤ₂]→ℤ₂[X]/→H*[RP²,ℤ₂] = DS-Ind-Prop.f _ _ _ _ - (λ _ → trunc _ _) - refl - (λ { zero x → cong (base zero) (retEq (H⁰[RP²,ℤ/2]≅ℤ/2 .fst) x) - ; one x → cong (base one) (retEq (H¹[RP²,ℤ/2]≅ℤ/2 .fst) x) - ; two x → cong (base two) (retEq (H²[RP²,ℤ/2]≅ℤ/2 .fst) x) - ; (suc (suc (suc r))) x → cong (base (3 +ℕ r)) - (isContr→isProp (isContrH³⁺ⁿ[RP²,G] r ℤ/2) _ _)}) - λ {x} {y} p q - → (IsRingHom.pres+ (ℤ₂[X]/→H*[RP²,ℤ₂] .snd) _ _ ∙ cong₂ _add_ p q) + opaque + unfolding H*Klein→ℤ/2[X,Y]/I quotientCommRingStr + H*[RP²,ℤ₂]→ℤ₂[X]/ : H*R ℤ/2Ring RP² .fst → ℤ₂[X]/ .fst + H*[RP²,ℤ₂]→ℤ₂[X]/ = DS-Rec-Set.f _ _ _ _ squash/ + [ neutral ] + H*[RP²,ℤ₂]→ℤ₂[X]/-fun + (CommRingStr._+_ (snd ℤ₂[X]/)) + (CommRingStr.+Assoc (snd ℤ₂[X]/)) + (CommRingStr.+IdR (snd ℤ₂[X]/)) + (CommRingStr.+Comm (snd ℤ₂[X]/)) + (λ { zero → cong [_] (base-neutral (0 ∷ [])) + ; one → cong [_] (cong (base (1 ∷ [])) (IsGroupHom.pres1 (snd (H¹[RP²,ℤ/2]≅ℤ/2))) + ∙ base-neutral (1 ∷ [])) + ; two → cong [_] (cong (base (2 ∷ [])) (IsGroupHom.pres1 (snd (H²[RP²,ℤ/2]≅ℤ/2))) + ∙ base-neutral (2 ∷ [])) + ; (suc (suc (suc r))) → cong [_] (base-neutral _)}) + (λ { zero a b → cong [_] (base-add (0 ∷ []) _ _ + ∙ cong (base (0 ∷ [])) + (sym (IsGroupHom.pres· (snd (H⁰[RP²,ℤ/2]≅ℤ/2)) a b))) + ; one a b → cong [_] (base-add (1 ∷ []) _ _ + ∙ cong (base (1 ∷ [])) + (sym (IsGroupHom.pres· (snd (H¹[RP²,ℤ/2]≅ℤ/2)) a b))) + ; two a b → cong [_] (base-add (2 ∷ []) _ _ + ∙ cong (base (2 ∷ [])) + (sym (IsGroupHom.pres· (snd (H²[RP²,ℤ/2]≅ℤ/2)) a b))) + ; (suc (suc (suc r))) a b + → cong [_] (cong₂ _add_ refl (base-neutral (3 +ℕ r ∷ [])) + ∙ addRid _)}) + + ℤ₂[X]/→H*[RP²,ℤ₂]→ℤ₂[X]/ : (x : _) + → H*[RP²,ℤ₂]→ℤ₂[X]/ (ℤ₂[X]/→H*[RP²,ℤ₂] .fst x) ≡ x + ℤ₂[X]/→H*[RP²,ℤ₂]→ℤ₂[X]/ = SQ.elimProp (λ _ → squash/ _ _) + (DS-Ind-Prop.f _ _ _ _ (λ _ → squash/ _ _) + refl + (λ { (zero ∷ []) a → cong [_] (cong (base (zero ∷ [])) + (secEq (H⁰[RP²,ℤ/2]≅ℤ/2 .fst) a)) + ; (one ∷ []) a → cong [_] (cong (base (one ∷ [])) + (secEq (H¹[RP²,ℤ/2]≅ℤ/2 .fst) a)) + ; (two ∷ []) a → cong [_] (cong (base (two ∷ [])) + (secEq (H²[RP²,ℤ/2]≅ℤ/2 .fst) a)) + ; (suc (suc (suc r)) ∷ []) → ℤ/2-elim refl + (sym (eq/ _ _ ∣ (λ {zero → base (r ∷ []) fone}) + , ((cong₂ _add_ refl (base-neutral _) + ∙ addRid _ + ∙ λ i → base (+-comm 3 r i ∷ []) fone)) + ∙ sym (addRid _) ∣₁))}) + (λ p q → cong₂ (CommRingStr._+_ (snd ℤ₂[X]/)) p q)) + + H*[RP²,ℤ₂]→ℤ₂[X]/→H*[RP²,ℤ₂] : (x : H*R ℤ/2Ring RP² .fst) + → ℤ₂[X]/→H*[RP²,ℤ₂] .fst (H*[RP²,ℤ₂]→ℤ₂[X]/ x) ≡ x + H*[RP²,ℤ₂]→ℤ₂[X]/→H*[RP²,ℤ₂] = DS-Ind-Prop.f _ _ _ _ + (λ _ → trunc _ _) + refl + (λ { zero x → cong (base zero) (retEq (H⁰[RP²,ℤ/2]≅ℤ/2 .fst) x) + ; one x → cong (base one) (retEq (H¹[RP²,ℤ/2]≅ℤ/2 .fst) x) + ; two x → cong (base two) (retEq (H²[RP²,ℤ/2]≅ℤ/2 .fst) x) + ; (suc (suc (suc r))) x → cong (base (3 +ℕ r)) + (isContr→isProp (isContrH³⁺ⁿ[RP²,G] r ℤ/2) _ _)}) + λ {x} {y} p q + → (IsRingHom.pres+ (ℤ₂[X]/→H*[RP²,ℤ₂] .snd) _ _ ∙ cong₂ _add_ p q) ℤ₂[X]/≅H*[RP²,ℤ₂] : RingEquiv (CommRing→Ring ℤ₂[X]/) (H*R ℤ/2Ring RP²) fst ℤ₂[X]/≅H*[RP²,ℤ₂] = diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda index d328f4afcf..0896082589 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda @@ -618,8 +618,10 @@ module Equiv-RP²∨S¹-Properties ℤ/2[x,y]/→H*-RP²∨S¹ : ℤ/2[x,y]/ → H*ℤ/2 RP²∨S¹ ℤ/2[x,y]/→H*-RP²∨S¹ = fst ℤ/2[X,Y]/→H*R-RP²∨S¹ - ℤ/2[x,y]/→H*-RP²∨S¹-pres0 : ℤ/2[x,y]/→H*-RP²∨S¹ 0Pℤ/2I ≡ 0H* - ℤ/2[x,y]/→H*-RP²∨S¹-pres0 = refl + opaque + unfolding H*Klein→ℤ/2[X,Y]/I + ℤ/2[x,y]/→H*-RP²∨S¹-pres0 : ℤ/2[x,y]/→H*-RP²∨S¹ 0Pℤ/2I ≡ 0H* + ℤ/2[x,y]/→H*-RP²∨S¹-pres0 = refl ℤ/2[x,y]/→H*-RP²∨S¹-pres+ : (x y : ℤ/2[x,y]/) → ℤ/2[x,y]/→H*-RP²∨S¹ ( x +Pℤ/2I y) @@ -670,13 +672,15 @@ module Equiv-RP²∨S¹-Properties H*-RP²∨S¹→ℤ/2[x,y]/ : H*ℤ/2 RP²∨S¹ → ℤ/2[x,y]/ H*-RP²∨S¹→ℤ/2[x,y]/ = [_] ∘ H*-RP²∨S¹→ℤ/2[x,y] - H*-RP²∨S¹→ℤ/2[x,y]/-pres0 : H*-RP²∨S¹→ℤ/2[x,y]/ 0H* ≡ 0Pℤ/2I - H*-RP²∨S¹→ℤ/2[x,y]/-pres0 = refl + opaque + unfolding H*Klein→ℤ/2[X,Y]/I + H*-RP²∨S¹→ℤ/2[x,y]/-pres0 : H*-RP²∨S¹→ℤ/2[x,y]/ 0H* ≡ 0Pℤ/2I + H*-RP²∨S¹→ℤ/2[x,y]/-pres0 = refl - H*-RP²∨S¹→ℤ/2[x,y]/-pres+ : (x y : H*ℤ/2 RP²∨S¹) → - H*-RP²∨S¹→ℤ/2[x,y]/ (x +H* y) - ≡ (H*-RP²∨S¹→ℤ/2[x,y]/ x) +Pℤ/2I (H*-RP²∨S¹→ℤ/2[x,y]/ y) - H*-RP²∨S¹→ℤ/2[x,y]/-pres+ x y = refl + H*-RP²∨S¹→ℤ/2[x,y]/-pres+ : (x y : H*ℤ/2 RP²∨S¹) → + H*-RP²∨S¹→ℤ/2[x,y]/ (x +H* y) + ≡ (H*-RP²∨S¹→ℤ/2[x,y]/ x) +Pℤ/2I (H*-RP²∨S¹→ℤ/2[x,y]/ y) + H*-RP²∨S¹→ℤ/2[x,y]/-pres+ x y = refl ----------------------------------------------------------------------------- @@ -740,12 +744,14 @@ module Equiv-RP²∨S¹-Properties ∙ cong₂ baseP (cong₂ (λ X Y → X ∷ Y ∷ []) (+-comm _ _) (+-comm _ _)) (·ℤ/2IdR _)) - e-retr : (x : ℤ/2[x,y]/) → H*-RP²∨S¹→ℤ/2[x,y]/ (ℤ/2[x,y]/→H*-RP²∨S¹ x) ≡ x - e-retr = SQ.elimProp (λ _ → isSetPℤ/2I _ _) - (DS-Ind-Prop.f _ _ _ _ (λ _ → isSetPℤ/2I _ _) - refl - e-retr-base - λ {U V} ind-U ind-V → cong₂ _+Pℤ/2I_ ind-U ind-V) + opaque + unfolding H*Klein→ℤ/2[X,Y]/I + e-retr : (x : ℤ/2[x,y]/) → H*-RP²∨S¹→ℤ/2[x,y]/ (ℤ/2[x,y]/→H*-RP²∨S¹ x) ≡ x + e-retr = SQ.elimProp (λ _ → isSetPℤ/2I _ _) + (DS-Ind-Prop.f _ _ _ _ (λ _ → isSetPℤ/2I _ _) + refl + e-retr-base + λ {U V} ind-U ind-V → cong₂ _+Pℤ/2I_ ind-U ind-V) -- Computation of the Cohomology Ring diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda index 7f494a47c9..62f4c0bb4d 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda @@ -245,116 +245,118 @@ module _ {ℓ : Level} (G : CommRing ℓ) (n : ℕ) where ∙ base-neutral (1 ∷ [])) help (gt x) = refl - H*[Sⁿ,G]→G[X]/-fun-coh₂ : (r : ℕ) (x y : coHom r GAb (S₊ (suc n))) - → CommRingStr._+_ (snd (G[X]/)) - (H*[Sⁿ,G]→G[X]/-fun r x) (H*[Sⁿ,G]→G[X]/-fun r y) - ≡ H*[Sⁿ,G]→G[X]/-fun r (x +ₕ y) - H*[Sⁿ,G]→G[X]/-fun-coh₂ zero x y = - cong [_] (base-add (0 ∷ []) _ _ - ∙ cong (base (0 ∷ [])) (sym (IsGroupHom.pres· (H⁰[Sⁿ,G]≅G GAb n .snd) _ _))) - H*[Sⁿ,G]→G[X]/-fun-coh₂ (suc r) x y = help (n ≟ r) x y - where - help : (p : _) (x y : coHom (suc r) GAb (S₊ (suc n))) + opaque + unfolding quotientCommRingStr + H*[Sⁿ,G]→G[X]/-fun-coh₂ : (r : ℕ) (x y : coHom r GAb (S₊ (suc n))) → CommRingStr._+_ (snd (G[X]/)) - (H*[Sⁿ,G]→G[X]/-fun^ r p x) - (H*[Sⁿ,G]→G[X]/-fun^ r p y) - ≡ H*[Sⁿ,G]→G[X]/-fun^ r p (x +ₕ y) - help (lt p) x y = cong [_] (addRid neutral) - help (eq p) = help' r p + (H*[Sⁿ,G]→G[X]/-fun r x) (H*[Sⁿ,G]→G[X]/-fun r y) + ≡ H*[Sⁿ,G]→G[X]/-fun r (x +ₕ y) + H*[Sⁿ,G]→G[X]/-fun-coh₂ zero x y = + cong [_] (base-add (0 ∷ []) _ _ + ∙ cong (base (0 ∷ [])) (sym (IsGroupHom.pres· (H⁰[Sⁿ,G]≅G GAb n .snd) _ _))) + H*[Sⁿ,G]→G[X]/-fun-coh₂ (suc r) x y = help (n ≟ r) x y where - help' : (r : ℕ) (p : n ≡ r) (x y : _) - → CommRingStr._+_ (snd G[X]/) - (H*[Sⁿ,G]→G[X]/-fun^ r (eq p) x) - (H*[Sⁿ,G]→G[X]/-fun^ r (eq p) y) - ≡ H*[Sⁿ,G]→G[X]/-fun^ r (eq p) (x +ₕ y) - help' = J> λ a b → cong [_] - (cong₂ (λ x y → (base (1 ∷ []) x) add (base (1 ∷ []) y)) - (cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) (transportRefl a)) - (cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) (transportRefl b)) - ∙ base-add _ _ _ - ∙ cong (base (1 ∷ [])) - (sym (IsGroupHom.pres· (Hⁿ[Sⁿ,G]≅G GAb n .snd) _ _) - ∙ cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) (sym (transportRefl (a +ₕ b))))) - help (gt p) x y = cong [_] (addRid neutral) - - H*[Sⁿ,G]→G[X]/ : fst (H*R GRing (S₊ (suc n))) → G[X]/ .fst - H*[Sⁿ,G]→G[X]/ = DS-Rec-Set.f _ _ _ _ squash/ - [ neutral ] - H*[Sⁿ,G]→G[X]/-fun - (CommRingStr._+_ (snd (G[X]/))) - (λ _ _ _ → CommRingStr.+Assoc (snd (G[X]/)) _ _ _) - (λ _ → CommRingStr.+IdR (snd (G[X]/)) _) - (λ _ _ → CommRingStr.+Comm (snd (G[X]/)) _ _) - H*[Sⁿ,G]→G[X]/-fun-coh₁ - H*[Sⁿ,G]→G[X]/-fun-coh₂ - - H*[Sⁿ,G]→G[X]/→H*[Sⁿ,G] : (x : fst (H*R GRing (S₊ (suc n)))) - → G[X]/→H*[Sⁿ,G] .fst (H*[Sⁿ,G]→G[X]/ x) ≡ x - H*[Sⁿ,G]→G[X]/→H*[Sⁿ,G] = - DS-Ind-Prop.f _ _ _ _ (λ _ → trunc _ _) - refl - (λ { zero a → cong (base zero) (retEq (H⁰[Sⁿ,G]≅G GAb n .fst) a) - ; (suc r) → help r (n ≟ r)}) - λ p q → IsRingHom.pres+ (G[X]/→H*[Sⁿ,G] .snd) _ _ - ∙ cong₂ _add_ p q - where - help : (r : ℕ) (p : _) (a : _) - → G[X]/→H*[Sⁿ,G] .fst (H*[Sⁿ,G]→G[X]/-fun^ r p a) - ≡ base (suc r) a - help r (lt x) a = - sym (base-neutral (suc r)) - ∙ cong (base (suc r)) - (isContr→isProp - (subst (λ k → isContr (coHom k GAb (S₊ (suc n)))) - (cong suc (snd x)) - (isContr-Hᵐ⁺ⁿ[Sⁿ,G] GAb n (fst x))) _ _) - help r (eq x) a = - J (λ r x → (a : coHom (suc r) GAb (S₊ (suc n))) - → G[X]/→H*[Sⁿ,G] .fst (H*[Sⁿ,G]→G[X]/-fun^ r (eq x) a) - ≡ base (suc r) a) - (λ a → cong (base (suc n)) - (retEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) _ ∙ transportRefl a)) x a - help r (gt x) a = - sym (base-neutral (suc r)) - ∙ cong (base (suc r)) - (isContr→isProp - (subst (λ k → isContr (coHom (suc r) GAb (S₊ k))) - (cong suc (snd x)) - (isOfHLevelRetractFromIso 0 - (equivToIso (fst (Hⁿ[Sᵐ⁺ⁿ,G]≅0 GAb r (fst x)))) - isContrUnit*)) _ _) - - G[X]/→H*[Sⁿ,G]→H*[Sⁿ,G] : (x : G[X]/ .fst) - → H*[Sⁿ,G]→G[X]/ (G[X]/→H*[Sⁿ,G] .fst x) ≡ x - G[X]/→H*[Sⁿ,G]→H*[Sⁿ,G] = SQ.elimProp (λ _ → squash/ _ _) - (DS-Ind-Prop.f _ _ _ _ (λ _ → squash/ _ _) - refl - (λ { (zero ∷ []) g → cong [_] (cong (base (0 ∷ [])) - (secEq (H⁰[Sⁿ,G]≅G GAb n .fst) g)) - ; (one ∷ []) g → h2 g _ - ; (suc (suc x) ∷ []) g - → sym (eq/ _ neutral ∣ hh x g - , (λ i → base (+-comm 2 x i ∷ []) (CommRingStr.·IdR (snd G) g (~ i)) - add neutral) ∣₁)}) - λ p q → cong₂ (CommRingStr._+_ (snd (G[X]/))) p q) - where - hh : (x : ℕ) (g : fst G) → FinVec (fst (PolyCommRing G 1)) 1 - hh x g zero = base (x ∷ []) g - - h2 : (g : fst G) (p : _) - → H*[Sⁿ,G]→G[X]/-fun^ n p - (invEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) g) - ≡ [ base (one ∷ []) g ] - h2 g (lt x) = ⊥.rec (¬m)) + (H*[Sⁿ,G]→G[X]/-fun^ r p x) + (H*[Sⁿ,G]→G[X]/-fun^ r p y) + ≡ H*[Sⁿ,G]→G[X]/-fun^ r p (x +ₕ y) + help (lt p) x y = cong [_] (addRid neutral) + help (eq p) = help' r p + where + help' : (r : ℕ) (p : n ≡ r) (x y : _) + → CommRingStr._+_ (snd G[X]/) + (H*[Sⁿ,G]→G[X]/-fun^ r (eq p) x) + (H*[Sⁿ,G]→G[X]/-fun^ r (eq p) y) + ≡ H*[Sⁿ,G]→G[X]/-fun^ r (eq p) (x +ₕ y) + help' = J> λ a b → cong [_] + (cong₂ (λ x y → (base (1 ∷ []) x) add (base (1 ∷ []) y)) + (cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) (transportRefl a)) + (cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) (transportRefl b)) + ∙ base-add _ _ _ + ∙ cong (base (1 ∷ [])) + (sym (IsGroupHom.pres· (Hⁿ[Sⁿ,G]≅G GAb n .snd) _ _) + ∙ cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) (sym (transportRefl (a +ₕ b))))) + help (gt p) x y = cong [_] (addRid neutral) + + H*[Sⁿ,G]→G[X]/ : fst (H*R GRing (S₊ (suc n))) → G[X]/ .fst + H*[Sⁿ,G]→G[X]/ = DS-Rec-Set.f _ _ _ _ squash/ + [ neutral ] + H*[Sⁿ,G]→G[X]/-fun + (CommRingStr._+_ (snd (G[X]/))) + (λ _ _ _ → CommRingStr.+Assoc (snd (G[X]/)) _ _ _) + (λ _ → CommRingStr.+IdR (snd (G[X]/)) _) + (λ _ _ → CommRingStr.+Comm (snd (G[X]/)) _ _) + H*[Sⁿ,G]→G[X]/-fun-coh₁ + H*[Sⁿ,G]→G[X]/-fun-coh₂ + + H*[Sⁿ,G]→G[X]/→H*[Sⁿ,G] : (x : fst (H*R GRing (S₊ (suc n)))) + → G[X]/→H*[Sⁿ,G] .fst (H*[Sⁿ,G]→G[X]/ x) ≡ x + H*[Sⁿ,G]→G[X]/→H*[Sⁿ,G] = + DS-Ind-Prop.f _ _ _ _ (λ _ → trunc _ _) + refl + (λ { zero a → cong (base zero) (retEq (H⁰[Sⁿ,G]≅G GAb n .fst) a) + ; (suc r) → help r (n ≟ r)}) + λ p q → IsRingHom.pres+ (G[X]/→H*[Sⁿ,G] .snd) _ _ + ∙ cong₂ _add_ p q + where + help : (r : ℕ) (p : _) (a : _) + → G[X]/→H*[Sⁿ,G] .fst (H*[Sⁿ,G]→G[X]/-fun^ r p a) + ≡ base (suc r) a + help r (lt x) a = + sym (base-neutral (suc r)) + ∙ cong (base (suc r)) + (isContr→isProp + (subst (λ k → isContr (coHom k GAb (S₊ (suc n)))) + (cong suc (snd x)) + (isContr-Hᵐ⁺ⁿ[Sⁿ,G] GAb n (fst x))) _ _) + help r (eq x) a = + J (λ r x → (a : coHom (suc r) GAb (S₊ (suc n))) + → G[X]/→H*[Sⁿ,G] .fst (H*[Sⁿ,G]→G[X]/-fun^ r (eq x) a) + ≡ base (suc r) a) + (λ a → cong (base (suc n)) + (retEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) _ ∙ transportRefl a)) x a + help r (gt x) a = + sym (base-neutral (suc r)) + ∙ cong (base (suc r)) + (isContr→isProp + (subst (λ k → isContr (coHom (suc r) GAb (S₊ k))) + (cong suc (snd x)) + (isOfHLevelRetractFromIso 0 + (equivToIso (fst (Hⁿ[Sᵐ⁺ⁿ,G]≅0 GAb r (fst x)))) + isContrUnit*)) _ _) + + G[X]/→H*[Sⁿ,G]→H*[Sⁿ,G] : (x : G[X]/ .fst) + → H*[Sⁿ,G]→G[X]/ (G[X]/→H*[Sⁿ,G] .fst x) ≡ x + G[X]/→H*[Sⁿ,G]→H*[Sⁿ,G] = SQ.elimProp (λ _ → squash/ _ _) + (DS-Ind-Prop.f _ _ _ _ (λ _ → squash/ _ _) + refl + (λ { (zero ∷ []) g → cong [_] (cong (base (0 ∷ [])) + (secEq (H⁰[Sⁿ,G]≅G GAb n .fst) g)) + ; (one ∷ []) g → h2 g _ + ; (suc (suc x) ∷ []) g + → sym (eq/ _ neutral ∣ hh x g + , (λ i → base (+-comm 2 x i ∷ []) (CommRingStr.·IdR (snd G) g (~ i)) + add neutral) ∣₁)}) + λ p q → cong₂ (CommRingStr._+_ (snd (G[X]/))) p q) + where + hh : (x : ℕ) (g : fst G) → FinVec (fst (PolyCommRing G 1)) 1 + hh x g zero = base (x ∷ []) g + + h2 : (g : fst G) (p : _) + → H*[Sⁿ,G]→G[X]/-fun^ n p + (invEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) g) + ≡ [ base (one ∷ []) g ] + h2 g (lt x) = ⊥.rec (¬m≅H*[Sⁿ,G] : RingEquiv (CommRing→Ring G[X]/) (H*R GRing (S₊ (suc n))) diff --git a/Cubical/ZCohomology/CohomologyRings/CP2.agda b/Cubical/ZCohomology/CohomologyRings/CP2.agda index 1b5ebe1b1f..6cfd837ffb 100644 --- a/Cubical/ZCohomology/CohomologyRings/CP2.agda +++ b/Cubical/ZCohomology/CohomologyRings/CP2.agda @@ -362,7 +362,7 @@ module ComputeCP²Notation H*-CP²→ℤ[x]/x³ : H* CP² → ℤ[x]/x³ H*-CP²→ℤ[x]/x³ = [_] ∘ H*-CP²→ℤ[x] - +{- H*-CP²→ℤ[x]/x³-gmorph : (x y : H* CP²) → H*-CP²→ℤ[x]/x³ (x +H* y) ≡ (H*-CP²→ℤ[x]/x³ x) +PℤI (H*-CP²→ℤ[x]/x³ y) H*-CP²→ℤ[x]/x³-gmorph x y = refl @@ -454,3 +454,4 @@ snd CP²-CohomologyRing = snd ℤ[X]/X³→H*R-CP² CohomologyRing-CP² : RingEquiv (H*R CP²) (CommRing→Ring ℤ[X]/X³) CohomologyRing-CP² = RingEquivs.invRingEquiv CP²-CohomologyRing +-- -} diff --git a/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda b/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda index 0454d41bbd..b21c3cf935 100644 --- a/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda +++ b/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda @@ -370,7 +370,7 @@ module Equiv-𝕂²-Properties ℤ[x,y]/<2y,y²,xy,x²>→H*-𝕂² : ℤ[x,y]/<2y,y²,xy,x²> → H* KleinBottle ℤ[x,y]/<2y,y²,xy,x²>→H*-𝕂² = fst ℤ[X,Y]/<2Y,Y²,XY,X²>→H*R-𝕂² - +{- ℤ[x,y]/<2y,y²,xy,x²>→H*-𝕂²-pres0 : ℤ[x,y]/<2y,y²,xy,x²>→H*-𝕂² 0PℤI ≡ 0H* ℤ[x,y]/<2y,y²,xy,x²>→H*-𝕂²-pres0 = refl @@ -570,3 +570,4 @@ module _ where CohomologyRing-𝕂² : RingEquiv (H*R KleinBottle) (CommRing→Ring ℤ[X,Y]/<2Y,Y²,XY,X²>) CohomologyRing-𝕂² = RingEquivs.invRingEquiv 𝕂²-CohomologyRing +-- -} diff --git a/Cubical/ZCohomology/CohomologyRings/RP2.agda b/Cubical/ZCohomology/CohomologyRings/RP2.agda index b17c6f6ab4..668e5db41d 100644 --- a/Cubical/ZCohomology/CohomologyRings/RP2.agda +++ b/Cubical/ZCohomology/CohomologyRings/RP2.agda @@ -320,7 +320,7 @@ module Equiv-RP2-Properties where ℤ[x]/<2x,x²>→H*-RP² : ℤ[x]/<2x,x²> → H* RP² ℤ[x]/<2x,x²>→H*-RP² = fst ℤ[X]/<2X,X²>→H*R-RP² - +{- ℤ[x]/<2x,x²>→H*-RP²-pres0 : ℤ[x]/<2x,x²>→H*-RP² 0PℤI ≡ 0H* ℤ[x]/<2x,x²>→H*-RP²-pres0 = refl @@ -512,3 +512,4 @@ module _ where CohomologyRing-RP² : RingEquiv (H*R RP²) (CommRing→Ring ℤ[X]/<2X,X²>) CohomologyRing-RP² = RingEquivs.invRingEquiv RP²-CohomologyRing +-- -} diff --git a/Cubical/ZCohomology/CohomologyRings/RP2wedgeS1.agda b/Cubical/ZCohomology/CohomologyRings/RP2wedgeS1.agda index 4c358d2960..66fee63268 100644 --- a/Cubical/ZCohomology/CohomologyRings/RP2wedgeS1.agda +++ b/Cubical/ZCohomology/CohomologyRings/RP2wedgeS1.agda @@ -371,7 +371,7 @@ module Equiv-RP²⋁S¹-Properties ℤ[x,y]/<2y,y²,xy,x²>→H*-RP²⋁S¹ : ℤ[x,y]/<2y,y²,xy,x²> → H* RP²⋁S¹ ℤ[x,y]/<2y,y²,xy,x²>→H*-RP²⋁S¹ = fst ℤ[X,Y]/<2Y,Y²,XY,X²>→H*R-RP²⋁S¹ - +{- ℤ[x,y]/<2y,y²,xy,x²>→H*-RP²⋁S¹-pres0 : ℤ[x,y]/<2y,y²,xy,x²>→H*-RP²⋁S¹ 0PℤI ≡ 0H* ℤ[x,y]/<2y,y²,xy,x²>→H*-RP²⋁S¹-pres0 = refl @@ -571,3 +571,4 @@ module _ where CohomologyRing-RP²⋁S¹ : RingEquiv (H*R RP²⋁S¹) (CommRing→Ring ℤ[X,Y]/<2Y,Y²,XY,X²>) CohomologyRing-RP²⋁S¹ = RingEquivs.invRingEquiv RP²⋁S¹-CohomologyRing +-- -} diff --git a/Cubical/ZCohomology/CohomologyRings/S0.agda b/Cubical/ZCohomology/CohomologyRings/S0.agda index 1ce3385eb1..6ac6cfaf53 100644 --- a/Cubical/ZCohomology/CohomologyRings/S0.agda +++ b/Cubical/ZCohomology/CohomologyRings/S0.agda @@ -34,7 +34,7 @@ open import Cubical.ZCohomology.CohomologyRings.Unit -- Computation of the cohomology ring open RingEquivs - +{- Cohomology-Ring-S⁰P : RingEquiv (H*R (S₊ 0)) (DirectProd-Ring (CommRing→Ring ℤ[X]/X) (CommRing→Ring ℤ[X]/X)) Cohomology-Ring-S⁰P = compRingEquiv (CohomologyRing-Equiv (invIso Iso-⊤⊎⊤-Bool)) (compRingEquiv (CohomologyRing-Coproduct Unit Unit) @@ -44,3 +44,4 @@ Cohomology-Ring-S⁰ℤ : RingEquiv (H*R (S₊ 0)) (DirectProd-Ring (CommRing→ Cohomology-Ring-S⁰ℤ = compRingEquiv (CohomologyRing-Equiv (invIso Iso-⊤⊎⊤-Bool)) (compRingEquiv (CohomologyRing-Coproduct Unit Unit) (Coproduct-Equiv.Coproduct-Equiv-12 CohomologyRing-Unitℤ CohomologyRing-Unitℤ)) +-} diff --git a/Cubical/ZCohomology/CohomologyRings/S1.agda b/Cubical/ZCohomology/CohomologyRings/S1.agda index ac7b26e0f2..f8ba962782 100644 --- a/Cubical/ZCohomology/CohomologyRings/S1.agda +++ b/Cubical/ZCohomology/CohomologyRings/S1.agda @@ -276,7 +276,7 @@ module Equiv-S1-Properties where H*-S¹→ℤ[x]/x² : H* (S₊ 1) → ℤ[x]/x² H*-S¹→ℤ[x]/x² = [_] ∘ H*-S¹→ℤ[x] - +{- H*-S¹→ℤ[x]/x²-pres+ : (x y : H* (S₊ 1)) → 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 @@ -337,3 +337,4 @@ module _ where CohomologyRing-S¹ : RingEquiv (H*R (S₊ 1)) (CommRing→Ring ℤ[X]/X²) CohomologyRing-S¹ = RingEquivs.invRingEquiv S¹-CohomologyRing +-- -} diff --git a/Cubical/ZCohomology/CohomologyRings/S2wedgeS4.agda b/Cubical/ZCohomology/CohomologyRings/S2wedgeS4.agda index 410ee7b9de..6e4b8e1a38 100644 --- a/Cubical/ZCohomology/CohomologyRings/S2wedgeS4.agda +++ b/Cubical/ZCohomology/CohomologyRings/S2wedgeS4.agda @@ -376,7 +376,7 @@ module Equiv-RP2-Properties where ℤ[x,y]/→H*-S²⋁S⁴ : ℤ[x,y]/ → H* S²⋁S⁴ ℤ[x,y]/→H*-S²⋁S⁴ = fst ℤ[X,Y]/→H*R-S²⋁S⁴ - +{- ℤ[x,y]/→H*-S²⋁S⁴-pres0 : ℤ[x,y]/→H*-S²⋁S⁴ 0PℤI ≡ 0H* ℤ[x,y]/→H*-S²⋁S⁴-pres0 = refl @@ -522,3 +522,4 @@ module _ where CohomologyRing-S²⋁S⁴ : RingEquiv (H*R S²⋁S⁴) (CommRing→Ring ℤ[X,Y]/) CohomologyRing-S²⋁S⁴ = RingEquivs.invRingEquiv S²⋁S⁴-CohomologyRing +-- -} diff --git a/Cubical/ZCohomology/CohomologyRings/Sn.agda b/Cubical/ZCohomology/CohomologyRings/Sn.agda index 1c0c287859..47b0f283b7 100644 --- a/Cubical/ZCohomology/CohomologyRings/Sn.agda +++ b/Cubical/ZCohomology/CohomologyRings/Sn.agda @@ -334,8 +334,10 @@ module Equiv-Sn-Properties (n : ℕ) where H*-Sⁿ→ℤ[x]/x² : H* (S₊ (suc n)) → ℤ[x]/x² H*-Sⁿ→ℤ[x]/x² = [_] ∘ H*-Sⁿ→ℤ[x] - 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 + opaque + unfolding quotientCommRingStr + 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 @@ -366,29 +368,30 @@ module Equiv-Sn-Properties (n : ℕ) where ----------------------------------------------------------------------------- -- Retraction - - e-retr : (x : ℤ[x]/x²) → H*-Sⁿ→ℤ[x]/x² (ℤ[x]/x²→H*-Sⁿ x) ≡ x - e-retr = SQ.elimProp (λ _ → isSetPℤI _ _) - (DS-Ind-Prop.f _ _ _ _ (λ _ → isSetPℤI _ _) - refl - base-case - λ {U V} ind-U ind-V → cong₂ _+PℤI_ ind-U ind-V) - where - base-case : _ - base-case (zero ∷ []) a = cong [_] (cong (base-trad-H* 0 (ϕ₀ a)) part0) - ∙ cong [_] (cong (base (0 ∷ [])) (cong ϕ₀⁻¹ (transportRefl (ϕ₀ a)))) - ∙ cong [_] (cong (base (0 ∷ [])) (ϕ₀-retr a)) - base-case (one ∷ []) a = cong [_] (cong (base-trad-H* (suc n) (ϕₙ a)) (partSn (part (suc n)))) - ∙ cong [_] (cong (base (1 ∷ [])) (cong ϕₙ⁻¹ (transportRefl (ϕₙ a)))) - ∙ cong [_] (cong (base (1 ∷ [])) (ϕₙ-retr a)) - base-case (suc (suc k) ∷ []) a = eq/ _ _ ∣ ((λ x → base (k ∷ []) (-ℤ a)) , helper) ∣₁ +{- + opaque + unfolding quotientCommRingStr + e-retr : (x : ℤ[x]/x²) → H*-Sⁿ→ℤ[x]/x² (ℤ[x]/x²→H*-Sⁿ x) ≡ x + e-retr = SQ.elimProp (λ _ → isSetPℤI _ _) + (DS-Ind-Prop.f _ _ _ _ (λ _ → isSetPℤI _ _) + refl + base-case + λ {U V} ind-U ind-V → cong₂ _+PℤI_ ind-U ind-V) where - helper : _ - helper = (+PℤIdL _) - ∙ cong₂ base (cong (λ X → X ∷ []) - (sym (+n-comm k 2))) (sym (·ℤIdR _)) - ∙ (sym (+PℤIdR _)) - + base-case : _ + base-case (zero ∷ []) a = cong [_] (cong (base-trad-H* 0 (ϕ₀ a)) part0) + ∙ cong [_] (cong (base (0 ∷ [])) (cong ϕ₀⁻¹ (transportRefl (ϕ₀ a)))) + ∙ cong [_] (cong (base (0 ∷ [])) (ϕ₀-retr a)) + base-case (one ∷ []) a = cong [_] (cong (base-trad-H* (suc n) (ϕₙ a)) (partSn (part (suc n)))) + ∙ cong [_] (cong (base (1 ∷ [])) (cong ϕₙ⁻¹ (transportRefl (ϕₙ a)))) + ∙ cong [_] (cong (base (1 ∷ [])) (ϕₙ-retr a)) + base-case (suc (suc k) ∷ []) a = eq/ _ _ ∣ ((λ x → base (k ∷ []) (-ℤ a)) , helper) ∣₁ + where + helper : _ + helper = (+PℤIdL _) + ∙ cong₂ base (cong (λ X → X ∷ []) + (sym (+n-comm k 2))) (sym (·ℤIdR _)) + ∙ (sym (+PℤIdR _)) ----------------------------------------------------------------------------- @@ -410,3 +413,4 @@ module _ (n : ℕ) where CohomologyRing-Sⁿ : RingEquiv (H*R (S₊ (suc n))) (CommRing→Ring ℤ[X]/X²) CohomologyRing-Sⁿ = RingEquivs.invRingEquiv Sⁿ-CohomologyRing +-- -} diff --git a/Cubical/ZCohomology/CohomologyRings/Unit.agda b/Cubical/ZCohomology/CohomologyRings/Unit.agda index 600ac9858a..405b90e40e 100644 --- a/Cubical/ZCohomology/CohomologyRings/Unit.agda +++ b/Cubical/ZCohomology/CohomologyRings/Unit.agda @@ -241,7 +241,7 @@ module Equiv-Unit-Properties where H*-Unit→ℤ[x]/x : H* Unit → ℤ[x]/x H*-Unit→ℤ[x]/x = [_] ∘ H*-Unit→ℤ[x] - +{- H*-Unit→ℤ[x]/x-pres+ : (x y : H* Unit) → H*-Unit→ℤ[x]/x (x +H* y) ≡ (H*-Unit→ℤ[x]/x x) +PℤI (H*-Unit→ℤ[x]/x y) H*-Unit→ℤ[x]/x-pres+ x y = cong [_] (H*-Unit→ℤ[x]-pres+ x y) @@ -314,3 +314,4 @@ module _ where CohomologyRing-Unitℤ : RingEquiv (H*R Unit) (CommRing→Ring ℤCR) CohomologyRing-Unitℤ = compRingEquiv CohomologyRing-UnitP Equiv-ℤ[X]/X-ℤ +-- -} From 34fc5c9d0613db1ed7b9409555dbc7215971fa40 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 20 Jun 2024 19:26:57 +0200 Subject: [PATCH 19/83] refactor/simplify --- Cubical/Algebra/Ring/Base.agda | 46 +++++++++++++------------ Cubical/Algebra/Ring/Properties.agda | 51 +++++++++------------------- 2 files changed, 41 insertions(+), 56 deletions(-) diff --git a/Cubical/Algebra/Ring/Base.agda b/Cubical/Algebra/Ring/Base.agda index 91ce9d7099..3d4b2a0143 100644 --- a/Cubical/Algebra/Ring/Base.agda +++ b/Cubical/Algebra/Ring/Base.agda @@ -78,6 +78,9 @@ record RingStr (A : Type ℓ) : Type (ℓ-suc ℓ) where open IsRing isRing public + +unquoteDecl RingStrIsoΣ = declareRecordIsoΣ RingStrIsoΣ (quote RingStr) + Ring : ∀ ℓ → Type (ℓ-suc ℓ) Ring ℓ = TypeWithStr ℓ RingStr @@ -234,6 +237,28 @@ uaRing {A = A} {B = B} = equivFun (RingPath A B) isGroupoidRing : isGroupoid (Ring ℓ) isGroupoidRing _ _ = isOfHLevelRespectEquiv 2 (RingPath _ _) (isSetRingEquiv _ _) + +-- Rings have an abelian group and a monoid + +module _ (A : Ring ℓ) where + open RingStr (A .snd) + -- (A , (ringstr 0r 1r _+_ _·_ -_ R)) + Ring→AbGroup : AbGroup ℓ + Ring→AbGroup .fst = ⟨ A ⟩ + Ring→AbGroup .snd .AbGroupStr.0g = 0r + Ring→AbGroup .snd .AbGroupStr._+_ = _+_ + Ring→AbGroup .snd .AbGroupStr.-_ = -_ + Ring→AbGroup .snd .AbGroupStr.isAbGroup = IsRing.+IsAbGroup isRing + + Ring→MultMonoid : Monoid ℓ + Ring→MultMonoid = monoid ⟨ A ⟩ 1r _·_ (IsRing.·IsMonoid isRing) + +Ring→Group : Ring ℓ → Group ℓ +Ring→Group = AbGroup→Group ∘ Ring→AbGroup + +Ring→AddMonoid : Ring ℓ → Monoid ℓ +Ring→AddMonoid = Group→Monoid ∘ Ring→Group + open RingStr open IsRingHom @@ -287,27 +312,6 @@ module _ (R : Ring ℓ) {A : Type ℓ} InducedRingPath = RingPath _ _ .fst InducedRingEquiv - - --- Rings have an abelian group and a monoid - -module _ ((A , (ringstr 0r 1r _+_ _·_ -_ R)) : Ring ℓ) where - Ring→AbGroup : AbGroup ℓ - Ring→AbGroup .fst = A - Ring→AbGroup .snd .AbGroupStr.0g = 0r - Ring→AbGroup .snd .AbGroupStr._+_ = _+_ - Ring→AbGroup .snd .AbGroupStr.-_ = -_ - Ring→AbGroup .snd .AbGroupStr.isAbGroup = IsRing.+IsAbGroup R - - Ring→MultMonoid : Monoid ℓ - Ring→MultMonoid = monoid A 1r _·_ (IsRing.·IsMonoid R) - -Ring→Group : Ring ℓ → Group ℓ -Ring→Group = AbGroup→Group ∘ Ring→AbGroup - -Ring→AddMonoid : Ring ℓ → Monoid ℓ -Ring→AddMonoid = Group→Monoid ∘ Ring→Group - -- Smart constructor for ring homomorphisms -- that infers the other equations from pres1, pres+, and pres· diff --git a/Cubical/Algebra/Ring/Properties.agda b/Cubical/Algebra/Ring/Properties.agda index afe940f880..9896e62723 100644 --- a/Cubical/Algebra/Ring/Properties.agda +++ b/Cubical/Algebra/Ring/Properties.agda @@ -12,6 +12,8 @@ open import Cubical.Foundations.Structure open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Path +open import Cubical.Functions.Embedding + open import Cubical.Data.Sigma open import Cubical.Algebra.Monoid @@ -263,47 +265,26 @@ module RingHomTheory {R : Ring ℓ} {S : Ring ℓ'} (φ : RingHom R S) where 0r ∎ +isSetRingStr : (R : Type ℓ) → isSet (RingStr R) +isSetRingStr R = + let open RingStr + in isOfHLevelSucIfInhabited→isOfHLevelSuc 1 λ str → + isOfHLevelRetractFromIso 2 RingStrIsoΣ $ + isSetΣ (str .is-set) (λ _ → + isSetΣ (str .is-set) (λ _ → + isSetΣ (isSet→ (isSet→ (str .is-set))) λ _ → + isSetΣ (isSet→ (isSet→ (str .is-set))) (λ _ → + isSetΣSndProp (isSet→ (str .is-set)) (λ _ → isPropIsRing _ _ _ _ _)))) + -- the Ring version of uaCompEquiv module RingUAFunctoriality where open RingStr open RingEquivs - Ring≡ : (A B : Ring ℓ) → ( - Σ[ p ∈ ⟨ A ⟩ ≡ ⟨ B ⟩ ] - Σ[ q0 ∈ PathP (λ i → p i) (0r (snd A)) (0r (snd B)) ] - Σ[ q1 ∈ PathP (λ i → p i) (1r (snd A)) (1r (snd B)) ] - Σ[ r+ ∈ PathP (λ i → p i → p i → p i) (_+_ (snd A)) (_+_ (snd B)) ] - Σ[ r· ∈ PathP (λ i → p i → p i → p i) (_·_ (snd A)) (_·_ (snd B)) ] - Σ[ s ∈ PathP (λ i → p i → p i) (-_ (snd A)) (-_ (snd B)) ] - PathP (λ i → IsRing (q0 i) (q1 i) (r+ i) (r· i) (s i)) (isRing (snd A)) (isRing (snd B))) - ≃ (A ≡ B) - Ring≡ A B = isoToEquiv theIso - where - open Iso - theIso : Iso _ _ - fun theIso (p , q0 , q1 , r+ , r· , s , t) i = p i - , ringstr (q0 i) (q1 i) (r+ i) (r· i) (s i) (t i) - inv theIso x = cong ⟨_⟩ x , cong (0r ∘ snd) x , cong (1r ∘ snd) x - , cong (_+_ ∘ snd) x , cong (_·_ ∘ snd) x , cong (-_ ∘ snd) x , cong (isRing ∘ snd) x - rightInv theIso _ = refl - leftInv theIso _ = refl - caracRing≡ : {A B : Ring ℓ} (p q : A ≡ B) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q - caracRing≡ {A = A} {B = B} p q P = - sym (transportTransport⁻ (ua (Ring≡ A B)) p) - ∙∙ cong (transport (ua (Ring≡ A B))) helper - ∙∙ transportTransport⁻ (ua (Ring≡ A B)) q - where - helper : transport (sym (ua (Ring≡ A B))) p ≡ transport (sym (ua (Ring≡ A B))) q - helper = Σ≡Prop - (λ _ → isPropΣ - (isOfHLevelPathP' 1 (is-set (snd B)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (is-set (snd B)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ → is-set (snd B)) _ _) - λ _ → isPropΣ (isOfHLevelPathP' 1 (isSetΠ λ _ → is-set (snd B)) _ _) - λ _ → isOfHLevelPathP 1 (isPropIsRing _ _ _ _ _) _ _) - (transportRefl (cong ⟨_⟩ p) ∙ P ∙ sym (transportRefl (cong ⟨_⟩ q))) + caracRing≡ _ _ α = + isEmbedding→Inj (iso→isEmbedding (invIso ΣPathIsoPathΣ)) _ _ $ + Σ≡Prop (λ _ → isOfHLevelPathP' 1 (isSetRingStr _) _ _) α uaCompRingEquiv : {A B C : Ring ℓ} (f : RingEquiv A B) (g : RingEquiv B C) → uaRing (compRingEquiv f g) ≡ uaRing f ∙ uaRing g From 76f8c222fd7e91ce3ec025487bbd03ae6cf3f91c Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 26 Jul 2024 19:33:11 +0200 Subject: [PATCH 20/83] WIP --- Cubical/Algebra/CommRing/Base.agda | 15 +++++++++++++-- Cubical/Algebra/Ring/Base.agda | 13 +++++++++---- 2 files changed, 22 insertions(+), 6 deletions(-) diff --git a/Cubical/Algebra/CommRing/Base.agda b/Cubical/Algebra/CommRing/Base.agda index 5de6caca0e..1a0701f16f 100644 --- a/Cubical/Algebra/CommRing/Base.agda +++ b/Cubical/Algebra/CommRing/Base.agda @@ -91,10 +91,21 @@ makeCommRing 0r 1r _+_ _·_ -_ is-setR +Assoc +IdR +InvR +Comm ·Assoc ·IdR ·D _ , commringstr _ _ _ _ _ (makeIsCommRing is-setR +Assoc +IdR +InvR +Comm ·Assoc ·IdR ·DistR+ ·Comm) CommRingStr→RingStr : {A : Type ℓ} → CommRingStr A → RingStr A -CommRingStr→RingStr (commringstr _ _ _ _ _ H) = ringstr _ _ _ _ _ (IsCommRing.isRing H) +CommRingStr→RingStr cringStr = + record + { 0r = _ ; 1r = _ ; _+_ = _ ; _·_ = _ ; -_ = _ ; + isRing = IsCommRing.isRing (CommRingStr.isCommRing cringStr) } CommRing→Ring : CommRing ℓ → Ring ℓ -CommRing→Ring (_ , commringstr _ _ _ _ _ H) = _ , ringstr _ _ _ _ _ (IsCommRing.isRing H) +fst (CommRing→Ring CRing) = fst CRing +snd (CommRing→Ring CRing) = record + { 0r = _ + ; 1r = _ + ; _+_ = _ + ; _·_ = _ + ; -_ = _ + ; isRing = IsCommRing.isRing (CommRingStr.isCommRing (snd CRing)) + } CommRing→AbGroup : CommRing ℓ → AbGroup ℓ CommRing→AbGroup R = Ring→AbGroup (CommRing→Ring R) diff --git a/Cubical/Algebra/Ring/Base.agda b/Cubical/Algebra/Ring/Base.agda index 3d4b2a0143..402277dfa5 100644 --- a/Cubical/Algebra/Ring/Base.agda +++ b/Cubical/Algebra/Ring/Base.agda @@ -33,7 +33,7 @@ private record IsRing {R : Type ℓ} (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) : Type ℓ where - + no-eta-equality constructor isring field @@ -61,8 +61,7 @@ unquoteDecl IsRingIsoΣ = declareRecordIsoΣ IsRingIsoΣ (quote IsRing) record RingStr (A : Type ℓ) : Type (ℓ-suc ℓ) where - - constructor ringstr + no-eta-equality field 0r : A @@ -218,7 +217,13 @@ RingHom≡ = Σ≡Prop λ f → isPropIsRingHom _ f _ data[ _+_ ∣ bin ∣ pres+ ] data[ _·_ ∣ bin ∣ pres· ] data[ -_ ∣ un ∣ pres- ] - prop[ isRing ∣ (λ _ _ → isPropIsRing _ _ _ _ _) ]) + prop[ isRing ∣ (λ R ΣR → + isPropIsRing + (snd (fst (fst (fst (fst ΣR))))) + (snd (fst (fst (fst ΣR)))) + (snd (fst (fst ΣR))) + (snd (fst ΣR)) + (snd ΣR)) ]) where open RingStr open IsRingHom From 8a0c069bc98d2a4d7595547274ba0a47182ed38a Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 26 Jul 2024 23:18:22 +0200 Subject: [PATCH 21/83] start with stuff on the new comm alg --- Cubical/Algebra/CommAlgebraAlt/Base.agda | 76 +++++ Cubical/Algebra/CommAlgebraAlt/Ideal.agda | 36 +++ .../CommAlgebraAlt/QuotientAlgebra.agda | 260 ++++++++++++++++++ 3 files changed, 372 insertions(+) create mode 100644 Cubical/Algebra/CommAlgebraAlt/Base.agda create mode 100644 Cubical/Algebra/CommAlgebraAlt/Ideal.agda create mode 100644 Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda diff --git a/Cubical/Algebra/CommAlgebraAlt/Base.agda b/Cubical/Algebra/CommAlgebraAlt/Base.agda new file mode 100644 index 0000000000..29941b54ee --- /dev/null +++ b/Cubical/Algebra/CommAlgebraAlt/Base.agda @@ -0,0 +1,76 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebraAlt.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Structure using (⟨_⟩) +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Sigma + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Ring + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + +CommAlgebra : (R : CommRing ℓ) (ℓ' : Level) → Type _ +CommAlgebra R ℓ' = Σ[ A ∈ CommRing ℓ' ] CommRingHom R A + +module _ {R : CommRing ℓ} where + open RingHoms + + CommAlgebra→CommRing : (A : CommAlgebra R ℓ') → CommRing ℓ' + CommAlgebra→CommRing = fst + + CommAlgebraHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') → Type _ + CommAlgebraHom A B = Σ[ f ∈ CommRingHom (fst A) (fst B) ] f ∘r snd A ≡ snd B + + idCAlgHom : (A : CommAlgebra R ℓ) → CommAlgebraHom A A + idCAlgHom A = (idCommRingHom (fst A)) , Σ≡Prop (λ _ → isPropIsCommRingHom (R .snd) _ (A .fst .snd)) refl + + compCommAlgebraHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} + → (g : CommAlgebraHom B C) → (f : CommAlgebraHom A B) + → CommAlgebraHom A C + compCommAlgebraHom {A = A} {B = B} {C = C} g f = + ((fst g) ∘r (fst f)) , + Σ≡Prop (λ _ → isPropIsCommRingHom (R .snd) _ (C .fst .snd)) + ( + (fst g ∘r (fst f ∘r snd A)) .fst ≡⟨ cong (λ h → (fst g ∘r h) .fst) (f .snd) ⟩ + (fst g ∘r snd B) .fst ≡⟨ cong fst (g .snd) ⟩ + (C .snd .fst) ∎) + + ⟨_⟩ₐ : (A : CommAlgebra R ℓ') → Type ℓ' + ⟨ A ⟩ₐ = A .fst .fst + + module CommAlgebraStr (A : CommAlgebra R ℓ') where + open CommRingStr {{...}} + instance + _ = A + _ : CommRingStr ⟨ A ⟩ₐ + _ = (A .fst .snd) + _ : CommRingStr (R .fst) + _ = (R .snd) + + infixl 7 _⋆_ + _⋆_ : ⟨ R ⟩ → ⟨ A ⟩ₐ → ⟨ A ⟩ₐ + _⋆_ r x = (A .snd .fst r) · x + + ⋆Assoc : (r s : ⟨ R ⟩) (x : ⟨ A ⟩ₐ) → (r · s) ⋆ x ≡ r ⋆ (s ⋆ x) + ⋆Assoc r s x = cong (_· x) (pres· r s) ∙ sym (·Assoc _ _ _) + where open IsRingHom (A .snd .snd) + + ⋆DistR+ : (r : ⟨ R ⟩) (x y : ⟨ A ⟩ₐ) → r ⋆ (x + y) ≡ (r ⋆ x) + (r ⋆ y) + ⋆DistR+ r x y = ·DistR+ _ _ _ + + ⋆DistL+ : (r s : ⟨ R ⟩) (x : ⟨ A ⟩ₐ) → (r + s) ⋆ x ≡ (r ⋆ x) + (s ⋆ x) + ⋆DistL+ r s x = cong (_· x) (pres+ r s) ∙ ·DistL+ _ _ _ + where open IsRingHom (A .snd .snd) + + ⋆IdL : (x : ⟨ A ⟩ₐ) → 1r ⋆ x ≡ x + ⋆IdL x = cong (_· x) pres1 ∙ ·IdL x + where open IsRingHom (A .snd .snd) + + ⋆AssocL : (r : ⟨ R ⟩) (x y : ⟨ A ⟩ₐ) → (r ⋆ x) · y ≡ r ⋆ (x · y) + ⋆AssocL r x y = sym (·Assoc (A .snd .fst r) x y) diff --git a/Cubical/Algebra/CommAlgebraAlt/Ideal.agda b/Cubical/Algebra/CommAlgebraAlt/Ideal.agda new file mode 100644 index 0000000000..775e11684c --- /dev/null +++ b/Cubical/Algebra/CommAlgebraAlt/Ideal.agda @@ -0,0 +1,36 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebraAlt.Ideal where +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Powerset + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Ideal renaming (IdealsIn to IdealsInCommRing; + makeIdeal to makeIdealCommRing) +open import Cubical.Algebra.CommAlgebraAlt.Base +-- open import Cubical.Algebra.Ring + +open import Cubical.Data.Unit + +private + variable + ℓ : Level + +module _ (R : CommRing ℓ) (A : CommAlgebra R ℓ) where + IdealsIn : Type _ + IdealsIn = IdealsInCommRing (fst A) + + open CommRingStr (A .fst .snd) + + makeIdeal : (I : A .fst .fst → hProp ℓ) + → (+-closed : {x y : A .fst .fst} → x ∈ I → y ∈ I → (x + y) ∈ I) + → (0-closed : 0r ∈ I) + → (·-closedLeft : {x : A .fst .fst} → (r : A .fst .fst) → x ∈ I → r · x ∈ I) + → IdealsIn + makeIdeal = makeIdealCommRing {R = fst A} + + 0Ideal : IdealsIn + 0Ideal = CommIdeal.0Ideal (fst A) + + 1Ideal : IdealsIn + 1Ideal = CommIdeal.1Ideal (fst A) diff --git a/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda new file mode 100644 index 0000000000..8ca755bd0c --- /dev/null +++ b/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda @@ -0,0 +1,260 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebraAlt.QuotientAlgebra where +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Powerset using (_∈_; _⊆_) +open import Cubical.Foundations.Structure + +open import Cubical.HITs.SetQuotients hiding (_/_) +open import Cubical.Data.Unit +open import Cubical.Data.Sigma.Properties using (Σ≡Prop) + +open import Cubical.Algebra.CommRing +import Cubical.Algebra.CommRing.Quotient as CommRing +import Cubical.Algebra.Ring.Quotient as Ring +open import Cubical.Algebra.CommRing.Ideal hiding (IdealsIn) +open import Cubical.Algebra.CommAlgebraAlt.Base +open import Cubical.Algebra.CommAlgebraAlt.Ideal +-- open import Cubical.Algebra.CommAlgebra.Kernel +-- open import Cubical.Algebra.CommAlgebra.Instances.Unit +-- open import Cubical.Algebra.Algebra.Base using (IsAlgebraHom; isPropIsAlgebraHom) +open import Cubical.Algebra.Ring +-- open import Cubical.Algebra.Ring.Ideal using (isIdeal) +open import Cubical.Tactics.CommRingSolver +-- open import Cubical.Algebra.Algebra.Properties +-- open AlgebraHoms using (_∘a_) + +private + variable + ℓ : Level + +{- + The definition of the quotient algebra (_/_ below) is marked opaque to avoid + long type checking times. All other definitions that need to "look into" this + opaque definition must be in an opaque block that unfolds the definition of _/_. +-} + +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where + open CommRingStr {{...}} -- hiding (_-_; -_; ·IdL ; ·DistR+) renaming (_·_ to _·R_; _+_ to _+R_) + open CommAlgebraStr {{...}} + open RingTheory (CommRing→Ring (fst A)) using (-DistR·) + open RingHoms + instance + _ : CommRingStr ⟨ R ⟩ + _ = snd R + _ : CommRingStr ⟨ A .fst ⟩ + _ = A .fst .snd + _ = A + + _/_ : CommAlgebra R ℓ + _/_ = ((fst A) CommRing./ I) , + (CommRing.quotientHom (fst A) I ∘r A .snd) + + quotientHom : CommAlgebraHom A (_/_) + quotientHom = (CommRing.quotientHom (fst A) I) , refl + +{- + commAlgebraFromCommRing + A/IAsCommRing + (λ r → elim (λ _ → squash/) (λ x → [ r ⋆ x ]) (eq r)) + (λ r s → elimProp (λ _ → squash/ _ _) + λ x i → [ ((r ·R s) ⋆ x ≡⟨ ⋆Assoc r s x ⟩ + r ⋆ (s ⋆ x) ∎) i ]) + (λ r → elimProp2 (λ _ _ → squash/ _ _) + λ x y i → [ (r ⋆ (x + y) ≡⟨ ⋆DistR+ r x y ⟩ + r ⋆ x + r ⋆ y ∎) i ]) + (λ r s → elimProp (λ _ → squash/ _ _) + λ x i → [ ((r +R s) ⋆ x ≡⟨ ⋆DistL+ r s x ⟩ + r ⋆ x + s ⋆ x ∎) i ]) + (elimProp (λ _ → squash/ _ _) + (λ x i → [ (1r ⋆ x ≡⟨ ⋆IdL x ⟩ x ∎) i ])) + λ r → elimProp2 (λ _ _ → squash/ _ _) + λ x y i → [ ((r ⋆ x) · y ≡⟨ ⋆AssocL r x y ⟩ + r ⋆ (x · y) ∎) i ] + + where + A/IAsCommRing : CommRing ℓ + A/IAsCommRing = (CommAlgebra→CommRing A) CommRing./ I + [_]/ : ⟨ A ⟩ → ⟨ A/IAsCommRing ⟩ + [_]/ = CommRing.[_]/ {R = CommAlgebra→CommRing A} {I = I} + open CommIdeal using (isCommIdeal) + eq : (r : fst R) (x y : A .fst .fst) → x - y ∈ (fst I) → [ r ⋆ x ]/ ≡ [ r ⋆ y ]/ + eq r x y x-y∈I = eq/ _ _ + (subst (λ u → u ∈ fst I) + ((r ⋆ 1r) · (x - y) ≡⟨ ·DistR+ (r ⋆ 1r) x (- y) ⟩ + (r ⋆ 1r) · x + (r ⋆ 1r) · (- y) ≡[ i ]⟨ (r ⋆ 1r) · x + -DistR· (r ⋆ 1r) y i ⟩ + (r ⋆ 1r) · x - (r ⋆ 1r) · y ≡[ i ]⟨ ⋆AssocL r 1r x i + - ⋆AssocL r 1r y i ⟩ + r ⋆ (1r · x) - r ⋆ (1r · y) ≡[ i ]⟨ r ⋆ (·IdL x i) - r ⋆ (·IdL y i) ⟩ + r ⋆ x - r ⋆ y ∎ ) + (isCommIdeal.·Closed (snd I) _ x-y∈I)) + +-} + +{- + + opaque + unfolding _/_ + + quotientHom : CommAlgebraHom A (_/_) + fst quotientHom x = [ x ] + IsAlgebraHom.pres0 (snd quotientHom) = refl + IsAlgebraHom.pres1 (snd quotientHom) = refl + IsAlgebraHom.pres+ (snd quotientHom) _ _ = refl + IsAlgebraHom.pres· (snd quotientHom) _ _ = refl + IsAlgebraHom.pres- (snd quotientHom) _ = refl + IsAlgebraHom.pres⋆ (snd quotientHom) _ _ = refl + +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn A) where + open CommRingStr {{...}} + hiding (_-_; -_; ·IdL; ·DistR+; is-set) + renaming (_·_ to _·R_; _+_ to _+R_) + open CommAlgebraStr ⦃...⦄ + + instance + _ : CommRingStr ⟨ R ⟩ + _ = snd R + _ : CommAlgebraStr R ⟨ A ⟩ + _ = snd A + + opaque + unfolding _/_ + + -- sanity check / maybe a helper function some day + -- (These two rings are not definitionally equal, but only because of proofs, not data.) + CommForget/ : RingEquiv (CommAlgebra→Ring (A / I)) ((CommAlgebra→Ring A) Ring./ (CommIdeal→Ideal I)) + fst CommForget/ = idEquiv _ + IsRingHom.pres0 (snd CommForget/) = refl + IsRingHom.pres1 (snd CommForget/) = refl + IsRingHom.pres+ (snd CommForget/) = λ _ _ → refl + IsRingHom.pres· (snd CommForget/) = λ _ _ → refl + IsRingHom.pres- (snd CommForget/) = λ _ → refl + + module _ + (B : CommAlgebra R ℓ) + (ϕ : CommAlgebraHom A B) + (I⊆kernel : (fst I) ⊆ (fst (kernel A B ϕ))) + where + + open IsAlgebraHom + open RingTheory (CommRing→Ring (CommAlgebra→CommRing B)) + + private + instance + _ : CommAlgebraStr R ⟨ B ⟩ + _ = snd B + _ : CommRingStr ⟨ B ⟩ + _ = snd (CommAlgebra→CommRing B) + + opaque + unfolding _/_ + + inducedHom : CommAlgebraHom (A / I) B + fst inducedHom = + rec is-set (λ x → fst ϕ x) + λ a b a-b∈I → + equalByDifference + (fst ϕ a) (fst ϕ b) + ((fst ϕ a) - (fst ϕ b) ≡⟨ cong (λ u → (fst ϕ a) + u) (sym (IsAlgebraHom.pres- (snd ϕ) _)) ⟩ + (fst ϕ a) + (fst ϕ (- b)) ≡⟨ sym (IsAlgebraHom.pres+ (snd ϕ) _ _) ⟩ + fst ϕ (a - b) ≡⟨ I⊆kernel (a - b) a-b∈I ⟩ + 0r ∎) + pres0 (snd inducedHom) = pres0 (snd ϕ) + pres1 (snd inducedHom) = pres1 (snd ϕ) + pres+ (snd inducedHom) = elimProp2 (λ _ _ → is-set _ _) (pres+ (snd ϕ)) + pres· (snd inducedHom) = elimProp2 (λ _ _ → is-set _ _) (pres· (snd ϕ)) + pres- (snd inducedHom) = elimProp (λ _ → is-set _ _) (pres- (snd ϕ)) + pres⋆ (snd inducedHom) = λ r → elimProp (λ _ → is-set _ _) (pres⋆ (snd ϕ) r) + + opaque + unfolding inducedHom quotientHom + + inducedHom∘quotientHom : inducedHom ∘a quotientHom A I ≡ ϕ + inducedHom∘quotientHom = Σ≡Prop (isPropIsCommAlgebraHom {M = A} {N = B}) (funExt (λ a → refl)) + + opaque + unfolding quotientHom + + injectivePrecomp : (B : CommAlgebra R ℓ) (f g : CommAlgebraHom (A / I) B) + → f ∘a (quotientHom A I) ≡ g ∘a (quotientHom A I) + → f ≡ g + injectivePrecomp B f g p = + Σ≡Prop + (λ h → isPropIsCommAlgebraHom {M = A / I} {N = B} h) + (descendMapPath (fst f) (fst g) is-set + λ x → λ i → fst (p i) x) + where + instance + _ : CommAlgebraStr R ⟨ B ⟩ + _ = str B + + +{- trivial quotient -} +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) where + open CommAlgebraStr (snd A) + + opaque + unfolding _/_ + + oneIdealQuotient : CommAlgebraEquiv (A / (1Ideal A)) (UnitCommAlgebra R {ℓ' = ℓ}) + fst oneIdealQuotient = + isoToEquiv (iso (fst (terminalMap R (A / (1Ideal A)))) + (λ _ → [ 0a ]) + (λ _ → isPropUnit* _ _) + (elimProp (λ _ → squash/ _ _) + λ a → eq/ 0a a tt*)) + snd oneIdealQuotient = snd (terminalMap R (A / (1Ideal A))) + + opaque + unfolding quotientHom + + zeroIdealQuotient : CommAlgebraEquiv A (A / (0Ideal A)) + fst zeroIdealQuotient = + let open RingTheory (CommRing→Ring (CommAlgebra→CommRing A)) + in isoToEquiv (iso (fst (quotientHom A (0Ideal A))) + (rec is-set (λ x → x) λ x y x-y≡0 → equalByDifference x y x-y≡0) + (elimProp (λ _ → squash/ _ _) λ _ → refl) + λ _ → refl) + snd zeroIdealQuotient = snd (quotientHom A (0Ideal A)) + +[_]/ : {R : CommRing ℓ} {A : CommAlgebra R ℓ} {I : IdealsIn A} + → (a : fst A) → fst (A / I) +[_]/ = fst (quotientHom _ _) + + + +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn A) where + open CommIdeal using (isPropIsCommIdeal) + + private + π : CommAlgebraHom A (A / I) + π = quotientHom A I + + opaque + unfolding quotientHom + + kernel≡I : kernel A (A / I) π ≡ I + kernel≡I = + kernel A (A / I) π ≡⟨ Σ≡Prop + (isPropIsCommIdeal (CommAlgebra→CommRing A)) + refl ⟩ + _ ≡⟨ CommRing.kernel≡I {R = CommAlgebra→CommRing A} I ⟩ + I ∎ + + +module _ + {R : CommRing ℓ} + {A : CommAlgebra R ℓ} + {I : IdealsIn A} + where + + opaque + unfolding quotientHom + + isZeroFromIdeal : (x : ⟨ A ⟩) → x ∈ (fst I) → fst (quotientHom A I) x ≡ CommAlgebraStr.0a (snd (A / I)) + isZeroFromIdeal x x∈I = eq/ x 0a (subst (_∈ fst I) (solve! (CommAlgebra→CommRing A)) x∈I ) + where + open CommAlgebraStr (snd A) +-} From 5801fd35765faebd241197c83f20f4d50ea0dc0c Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 29 Jul 2024 17:35:40 +0200 Subject: [PATCH 22/83] Revert "Auxiliary commit to revert individual files from c5e74aa95d26512451a1a2bd2edae0df18561e13" This reverts commit cea8a34f6ab8876640fad0a558743f1499fa15c5. --- Cubical/Algebra/Ring/Base.agda | 13 ++++--------- 1 file changed, 4 insertions(+), 9 deletions(-) diff --git a/Cubical/Algebra/Ring/Base.agda b/Cubical/Algebra/Ring/Base.agda index 402277dfa5..3d4b2a0143 100644 --- a/Cubical/Algebra/Ring/Base.agda +++ b/Cubical/Algebra/Ring/Base.agda @@ -33,7 +33,7 @@ private record IsRing {R : Type ℓ} (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) : Type ℓ where - no-eta-equality + constructor isring field @@ -61,7 +61,8 @@ unquoteDecl IsRingIsoΣ = declareRecordIsoΣ IsRingIsoΣ (quote IsRing) record RingStr (A : Type ℓ) : Type (ℓ-suc ℓ) where - no-eta-equality + + constructor ringstr field 0r : A @@ -217,13 +218,7 @@ RingHom≡ = Σ≡Prop λ f → isPropIsRingHom _ f _ data[ _+_ ∣ bin ∣ pres+ ] data[ _·_ ∣ bin ∣ pres· ] data[ -_ ∣ un ∣ pres- ] - prop[ isRing ∣ (λ R ΣR → - isPropIsRing - (snd (fst (fst (fst (fst ΣR))))) - (snd (fst (fst (fst ΣR)))) - (snd (fst (fst ΣR))) - (snd (fst ΣR)) - (snd ΣR)) ]) + prop[ isRing ∣ (λ _ _ → isPropIsRing _ _ _ _ _) ]) where open RingStr open IsRingHom From 2725661f393f350f8c1f2c002f7d8a15e37bf377 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 29 Jul 2024 17:49:20 +0200 Subject: [PATCH 23/83] Quotients --- .../CommAlgebraAlt/QuotientAlgebra.agda | 68 +++---------------- Cubical/Algebra/CommRing/Quotient/Base.agda | 12 ++-- 2 files changed, 17 insertions(+), 63 deletions(-) diff --git a/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda index 8ca755bd0c..f5fe36945e 100644 --- a/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda @@ -52,72 +52,22 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where _/_ = ((fst A) CommRing./ I) , (CommRing.quotientHom (fst A) I ∘r A .snd) - quotientHom : CommAlgebraHom A (_/_) + quotientHom : CommAlgebraHom {R = R} A (_/_) quotientHom = (CommRing.quotientHom (fst A) I) , refl -{- - commAlgebraFromCommRing - A/IAsCommRing - (λ r → elim (λ _ → squash/) (λ x → [ r ⋆ x ]) (eq r)) - (λ r s → elimProp (λ _ → squash/ _ _) - λ x i → [ ((r ·R s) ⋆ x ≡⟨ ⋆Assoc r s x ⟩ - r ⋆ (s ⋆ x) ∎) i ]) - (λ r → elimProp2 (λ _ _ → squash/ _ _) - λ x y i → [ (r ⋆ (x + y) ≡⟨ ⋆DistR+ r x y ⟩ - r ⋆ x + r ⋆ y ∎) i ]) - (λ r s → elimProp (λ _ → squash/ _ _) - λ x i → [ ((r +R s) ⋆ x ≡⟨ ⋆DistL+ r s x ⟩ - r ⋆ x + s ⋆ x ∎) i ]) - (elimProp (λ _ → squash/ _ _) - (λ x i → [ (1r ⋆ x ≡⟨ ⋆IdL x ⟩ x ∎) i ])) - λ r → elimProp2 (λ _ _ → squash/ _ _) - λ x y i → [ ((r ⋆ x) · y ≡⟨ ⋆AssocL r x y ⟩ - r ⋆ (x · y) ∎) i ] - - where - A/IAsCommRing : CommRing ℓ - A/IAsCommRing = (CommAlgebra→CommRing A) CommRing./ I - [_]/ : ⟨ A ⟩ → ⟨ A/IAsCommRing ⟩ - [_]/ = CommRing.[_]/ {R = CommAlgebra→CommRing A} {I = I} - open CommIdeal using (isCommIdeal) - eq : (r : fst R) (x y : A .fst .fst) → x - y ∈ (fst I) → [ r ⋆ x ]/ ≡ [ r ⋆ y ]/ - eq r x y x-y∈I = eq/ _ _ - (subst (λ u → u ∈ fst I) - ((r ⋆ 1r) · (x - y) ≡⟨ ·DistR+ (r ⋆ 1r) x (- y) ⟩ - (r ⋆ 1r) · x + (r ⋆ 1r) · (- y) ≡[ i ]⟨ (r ⋆ 1r) · x + -DistR· (r ⋆ 1r) y i ⟩ - (r ⋆ 1r) · x - (r ⋆ 1r) · y ≡[ i ]⟨ ⋆AssocL r 1r x i - - ⋆AssocL r 1r y i ⟩ - r ⋆ (1r · x) - r ⋆ (1r · y) ≡[ i ]⟨ r ⋆ (·IdL x i) - r ⋆ (·IdL y i) ⟩ - r ⋆ x - r ⋆ y ∎ ) - (isCommIdeal.·Closed (snd I) _ x-y∈I)) - --} - -{- - - opaque - unfolding _/_ - - quotientHom : CommAlgebraHom A (_/_) - fst quotientHom x = [ x ] - IsAlgebraHom.pres0 (snd quotientHom) = refl - IsAlgebraHom.pres1 (snd quotientHom) = refl - IsAlgebraHom.pres+ (snd quotientHom) _ _ = refl - IsAlgebraHom.pres· (snd quotientHom) _ _ = refl - IsAlgebraHom.pres- (snd quotientHom) _ = refl - IsAlgebraHom.pres⋆ (snd quotientHom) _ _ = refl - -module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn A) where - open CommRingStr {{...}} - hiding (_-_; -_; ·IdL; ·DistR+; is-set) - renaming (_·_ to _·R_; _+_ to _+R_) +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where + open CommRingStr ⦃...⦄ open CommAlgebraStr ⦃...⦄ + instance _ : CommRingStr ⟨ R ⟩ _ = snd R - _ : CommAlgebraStr R ⟨ A ⟩ - _ = snd A + _ : CommRingStr ⟨ A ⟩ₐ + _ = A .fst .snd + _ = A + +{- opaque unfolding _/_ diff --git a/Cubical/Algebra/CommRing/Quotient/Base.agda b/Cubical/Algebra/CommRing/Quotient/Base.agda index 66eb50e4d6..d89fd6deff 100644 --- a/Cubical/Algebra/CommRing/Quotient/Base.agda +++ b/Cubical/Algebra/CommRing/Quotient/Base.agda @@ -3,6 +3,7 @@ module Cubical.Algebra.CommRing.Quotient.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function open import Cubical.Foundations.Powerset open import Cubical.Functions.Surjection @@ -112,11 +113,11 @@ module Quotient-FGideal-CommRing-CommRing (f : CommRingHom R S) {n : ℕ} (v : FinVec ⟨ R ⟩ n) - (fnull : (k : Fin n) → f $r v k ≡ CommRingStr.0r (snd S)) + (fnull : (k : Fin n) → f $cr v k ≡ CommRingStr.0r (snd S)) where inducedHom : CommRingHom (R / (generatedIdeal _ v)) S - inducedHom = Quotient-FGideal-CommRing-Ring.inducedHom R (CommRing→Ring S) f v fnull + inducedHom = RingHom→CommRingHom $ Quotient-FGideal-CommRing-Ring.inducedHom R (CommRing→Ring S) (CommRingHom→RingHom f) v fnull module UniversalProperty (R S : CommRing ℓ) @@ -126,11 +127,14 @@ module UniversalProperty where inducedHom : CommRingHom (R / I) S - inducedHom = Ring.UniversalProperty.inducedHom (CommRing→Ring R) (CommIdeal→Ideal I) f I⊆ker + inducedHom = RingHom→CommRingHom $ + Ring.UniversalProperty.inducedHom (CommRing→Ring R) (CommIdeal→Ideal I) (CommRingHom→RingHom f) I⊆ker ∘r Coherence.ringStr R I quotientHom : (R : CommRing ℓ) → (I : IdealsIn R) → CommRingHom R (R / I) -quotientHom R I = Coherence.ringStrInv R I ∘r Ring.quotientHom (CommRing→Ring R) (CommIdeal→Ideal I) +quotientHom R I = RingHom→CommRingHom $ + Coherence.ringStrInv R I + ∘r Ring.quotientHom (CommRing→Ring R) (CommIdeal→Ideal I) quotientHomSurjective : (R : CommRing ℓ) → (I : IdealsIn R) → isSurjection (fst (quotientHom R I)) From 9b9e0ace1bb175288975048088c471bec893299f Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Tue, 30 Jul 2024 12:06:01 +0200 Subject: [PATCH 24/83] update comm ring hom interface --- Cubical/Algebra/CommRing/Base.agda | 3 ++ Cubical/Algebra/CommRing/FiberedProduct.agda | 20 +++++------ .../CommRing/Ideal/SurjectiveImage.agda | 2 +- .../Localisation/InvertingElements.agda | 9 ++--- .../Algebra/CommRing/Localisation/Limit.agda | 22 ++++++------ .../Localisation/UniversalProperty.agda | 13 +++---- .../Algebra/CommRing/Quotient/IdealSum.agda | 6 ++-- .../Functorial/ZFunctors/Base.agda | 18 +++++----- .../Functorial/ZFunctors/CompactOpen.agda | 2 +- .../ZariskiLattice/UniversalProperty.agda | 2 +- Cubical/Categories/Instances/CommRings.agda | 32 ++++++++--------- .../Site/Instances/ZariskiCommRing.agda | 36 +++++++++---------- 12 files changed, 85 insertions(+), 80 deletions(-) diff --git a/Cubical/Algebra/CommRing/Base.agda b/Cubical/Algebra/CommRing/Base.agda index 1a0701f16f..c2786d580d 100644 --- a/Cubical/Algebra/CommRing/Base.agda +++ b/Cubical/Algebra/CommRing/Base.agda @@ -237,3 +237,6 @@ module _ {R : CommRing ℓ} {S : CommRing ℓ'} {f : ⟨ R ⟩ → ⟨ S ⟩} wh _$cr_ : {R : CommRing ℓ} {S : CommRing ℓ'} → (φ : CommRingHom R S) → (x : ⟨ R ⟩) → ⟨ S ⟩ φ $cr x = φ .fst x + +CommRingHom≡ : {R : CommRing ℓ} {S : CommRing ℓ'} {φ ψ : CommRingHom R S} → fst φ ≡ fst ψ → φ ≡ ψ +CommRingHom≡ = Σ≡Prop λ f → isPropIsCommRingHom _ f _ diff --git a/Cubical/Algebra/CommRing/FiberedProduct.agda b/Cubical/Algebra/CommRing/FiberedProduct.agda index ba1b1cb97f..a64c67a600 100644 --- a/Cubical/Algebra/CommRing/FiberedProduct.agda +++ b/Cubical/Algebra/CommRing/FiberedProduct.agda @@ -20,11 +20,11 @@ module _ (A B C : CommRing ℓ) (α : CommRingHom A C) (β : CommRingHom B C) wh module A = CommRingStr (snd A) module B = CommRingStr (snd B) module C = CommRingStr (snd C) - module α = IsRingHom (snd α) - module β = IsRingHom (snd β) + module α = IsCommRingHom (snd α) + module β = IsCommRingHom (snd β) open CommRingStr - open IsRingHom + open IsCommRingHom fbT : Type ℓ fbT = Σ[ ab ∈ fst A × fst B ] (fst α (fst ab) ≡ fst β (snd ab)) @@ -69,15 +69,15 @@ module _ (A B C : CommRing ℓ) (α : CommRingHom A C) (β : CommRingHom B C) wh fiberedProductPr₁ : CommRingHom fiberedProduct A fst fiberedProductPr₁ = fst ∘ fst - snd fiberedProductPr₁ = makeIsRingHom refl (λ _ _ → refl) (λ _ _ → refl) + snd fiberedProductPr₁ = makeIsCommRingHom refl (λ _ _ → refl) (λ _ _ → refl) fiberedProductPr₂ : CommRingHom fiberedProduct B fst fiberedProductPr₂ = snd ∘ fst - snd fiberedProductPr₂ = makeIsRingHom refl (λ _ _ → refl) (λ _ _ → refl) + snd fiberedProductPr₂ = makeIsCommRingHom refl (λ _ _ → refl) (λ _ _ → refl) fiberedProductPr₁₂Commutes : compCommRingHom fiberedProduct A C fiberedProductPr₁ α ≡ compCommRingHom fiberedProduct B C fiberedProductPr₂ β - fiberedProductPr₁₂Commutes = RingHom≡ (funExt (λ x → x .snd)) + fiberedProductPr₁₂Commutes = CommRingHom≡ (funExt (λ x → x .snd)) fiberedProductUnivProp : (D : CommRing ℓ) (h : CommRingHom D A) (k : CommRingHom D B) → @@ -86,14 +86,14 @@ module _ (A B C : CommRing ℓ) (α : CommRingHom A C) (β : CommRingHom B C) wh (h ≡ compCommRingHom D fiberedProduct A l fiberedProductPr₁) × (k ≡ compCommRingHom D fiberedProduct B l fiberedProductPr₂) fiberedProductUnivProp D (h , hh) (k , hk) H = - uniqueExists f (RingHom≡ refl , RingHom≡ refl) - (λ _ → isProp× (isSetRingHom _ _ _ _) (isSetRingHom _ _ _ _)) + uniqueExists f (CommRingHom≡ refl , CommRingHom≡ refl) + (λ _ → isProp× (isSetCommRingHom _ _ _ _) (isSetCommRingHom _ _ _ _)) (λ { (g , _) (Hh , Hk) → - RingHom≡ (funExt (λ d → fbT≡ (funExt⁻ (cong fst Hh) d) + CommRingHom≡ (funExt (λ d → fbT≡ (funExt⁻ (cong fst Hh) d) (funExt⁻ (cong fst Hk) d))) }) where f : CommRingHom D fiberedProduct fst f d = (h d , k d) , funExt⁻ (cong fst H) d - snd f = makeIsRingHom (fbT≡ (hh .pres1) (hk .pres1)) + snd f = makeIsCommRingHom (fbT≡ (hh .pres1) (hk .pres1)) (λ _ _ → fbT≡ (hh .pres+ _ _) (hk .pres+ _ _)) (λ _ _ → fbT≡ (hh .pres· _ _) (hk .pres· _ _)) diff --git a/Cubical/Algebra/CommRing/Ideal/SurjectiveImage.agda b/Cubical/Algebra/CommRing/Ideal/SurjectiveImage.agda index 71126412d9..09bfacff80 100644 --- a/Cubical/Algebra/CommRing/Ideal/SurjectiveImage.agda +++ b/Cubical/Algebra/CommRing/Ideal/SurjectiveImage.agda @@ -23,4 +23,4 @@ imageIdeal : {R S : CommRing ℓ} (f : CommRingHom R S) (f-epi : isSurjection (fst f)) (I : IdealsIn R) → IdealsIn S -imageIdeal f f-epi I = Ideal→CommIdeal (Ring.imageIdeal f f-epi (CommIdeal→Ideal I)) +imageIdeal f f-epi I = Ideal→CommIdeal (Ring.imageIdeal (CommRingHom→RingHom f) f-epi (CommIdeal→Ideal I)) diff --git a/Cubical/Algebra/CommRing/Localisation/InvertingElements.agda b/Cubical/Algebra/CommRing/Localisation/InvertingElements.agda index 6441cd684e..2aea295444 100644 --- a/Cubical/Algebra/CommRing/Localisation/InvertingElements.agda +++ b/Cubical/Algebra/CommRing/Localisation/InvertingElements.agda @@ -44,6 +44,7 @@ open import Cubical.Algebra.CommRing.Localisation.UniversalProperty open import Cubical.Algebra.CommRing.Ideal open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.RadicalIdeal +open import Cubical.Algebra.CommRing.Univalence open import Cubical.Tactics.CommRingSolver @@ -92,7 +93,7 @@ module InvertingElementsBase (R' : CommRing ℓ) where R[1/0]≡0 : R[1/ 0r ]AsCommRing ≡ UnitCommRing R[1/0]≡0 = uaCommRing (e , eIsRHom) where - open IsRingHom + open IsCommRingHom e : R[1/ 0r ]AsCommRing .fst ≃ UnitCommRing .fst e = isContr→Equiv isContrR[1/0] isContrUnit* @@ -280,11 +281,11 @@ module _ {A B : CommRing ℓ} (φ : CommRingHom A B) (f : fst A) where module A = InvertingElementsBase A module B = InvertingElementsBase B module AU = A.UniversalProp f - module BU = B.UniversalProp (φ $r f) + module BU = B.UniversalProp (φ $cr f) φ/1 = BU./1AsCommRingHom ∘cr φ - uniqInvElemHom : ∃![ χ ∈ CommRingHom A.R[1/ f ]AsCommRing B.R[1/ φ $r f ]AsCommRing ] + uniqInvElemHom : ∃![ χ ∈ CommRingHom A.R[1/ f ]AsCommRing B.R[1/ φ $cr f ]AsCommRing ] (fst χ) ∘ (fst AU./1AsCommRingHom) ≡ (fst φ/1) uniqInvElemHom = AU.invElemUniversalProp _ φ/1 (BU.S/1⊆S⁻¹Rˣ _ ∣ 1 , sym (·IdR _) ∣₁) @@ -381,7 +382,7 @@ module DoubleLoc (R' : CommRing ℓ) (f g : fst R') where /1/1AsCommRingHom : CommRingHom R' R[1/f][1/g]AsCommRing fst /1/1AsCommRingHom = _/1/1 - snd /1/1AsCommRingHom = makeIsRingHom refl lem+ lem· + snd /1/1AsCommRingHom = makeIsCommRingHom refl lem+ lem· where lem+ : _ lem+ r r' = diff --git a/Cubical/Algebra/CommRing/Localisation/Limit.agda b/Cubical/Algebra/CommRing/Localisation/Limit.agda index 6f65e200ac..079ac5d113 100644 --- a/Cubical/Algebra/CommRing/Localisation/Limit.agda +++ b/Cubical/Algebra/CommRing/Localisation/Limit.agda @@ -512,7 +512,7 @@ module _ (R' : CommRing ℓ) {n : ℕ} (f : FinVec (fst R') n) where 0at i ∎ where instance _ = L.S⁻¹RAsCommRing i .snd - open IsRingHom + open IsCommRingHom @@ -539,8 +539,8 @@ module _ (R' : CommRing ℓ) {n : ℕ} (f : FinVec (fst R') n) where coneOut locCone (sing i) = U./1AsCommRingHom i coneOut locCone (pair i j _) = UP./1AsCommRingHom i j coneOutCommutes locCone idAr = ⋆IdR _ - coneOutCommutes locCone singPairL = RingHom≡ (χˡUnique _ _ .fst .snd) - coneOutCommutes locCone singPairR = RingHom≡ (χʳUnique _ _ .fst .snd) + coneOutCommutes locCone singPairL = CommRingHom≡ (χˡUnique _ _ .fst .snd) + coneOutCommutes locCone singPairR = CommRingHom≡ (χʳUnique _ _ .fst .snd) isLimConeLocCone : 1r ∈ ⟨f₀,⋯,fₙ⟩ → isLimCone _ _ locCone isLimConeLocCone 1∈⟨f₀,⋯,fₙ⟩ A' cᴬ = (ψ , isConeMorψ) , ψUniqueness @@ -557,19 +557,19 @@ module _ (R' : CommRing ℓ) {n : ℕ} (f : FinVec (fst R') n) where where χφSquare< : ∀ i j → i < j → χˡ i j .fst (φ i .fst a) ≡ χʳ i j .fst (φ j .fst a) χφSquare< i j i Date: Tue, 30 Jul 2024 14:11:13 +0200 Subject: [PATCH 25/83] more boilerplate --- Cubical/Algebra/CommAlgebraAlt/Base.agda | 30 +++++++++++++++++------- 1 file changed, 21 insertions(+), 9 deletions(-) diff --git a/Cubical/Algebra/CommAlgebraAlt/Base.agda b/Cubical/Algebra/CommAlgebraAlt/Base.agda index 29941b54ee..e73fe1ff49 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Base.agda @@ -3,6 +3,7 @@ module Cubical.Algebra.CommAlgebraAlt.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function open import Cubical.Foundations.Structure using (⟨_⟩) open import Cubical.Foundations.HLevels @@ -19,13 +20,24 @@ CommAlgebra : (R : CommRing ℓ) (ℓ' : Level) → Type _ CommAlgebra R ℓ' = Σ[ A ∈ CommRing ℓ' ] CommRingHom R A module _ {R : CommRing ℓ} where - open RingHoms - CommAlgebra→CommRing : (A : CommAlgebra R ℓ') → CommRing ℓ' CommAlgebra→CommRing = fst + CommAlgebra→Ring : (A : CommAlgebra R ℓ') → Ring ℓ' + CommAlgebra→Ring A = CommRing→Ring (fst A) + CommAlgebraHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') → Type _ - CommAlgebraHom A B = Σ[ f ∈ CommRingHom (fst A) (fst B) ] f ∘r snd A ≡ snd B + CommAlgebraHom A B = Σ[ f ∈ CommRingHom (fst A) (fst B) ] f ∘cr snd A ≡ snd B + + CommAlgebraHom→CommRingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → CommAlgebraHom A B + → CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) + CommAlgebraHom→CommRingHom = fst + + CommAlgebraHom→RingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → CommAlgebraHom A B + → RingHom (CommAlgebra→Ring A) (CommAlgebra→Ring B) + CommAlgebraHom→RingHom = CommRingHom→RingHom ∘ CommAlgebraHom→CommRingHom idCAlgHom : (A : CommAlgebra R ℓ) → CommAlgebraHom A A idCAlgHom A = (idCommRingHom (fst A)) , Σ≡Prop (λ _ → isPropIsCommRingHom (R .snd) _ (A .fst .snd)) refl @@ -34,11 +46,11 @@ module _ {R : CommRing ℓ} where → (g : CommAlgebraHom B C) → (f : CommAlgebraHom A B) → CommAlgebraHom A C compCommAlgebraHom {A = A} {B = B} {C = C} g f = - ((fst g) ∘r (fst f)) , + ((fst g) ∘cr (fst f)) , Σ≡Prop (λ _ → isPropIsCommRingHom (R .snd) _ (C .fst .snd)) ( - (fst g ∘r (fst f ∘r snd A)) .fst ≡⟨ cong (λ h → (fst g ∘r h) .fst) (f .snd) ⟩ - (fst g ∘r snd B) .fst ≡⟨ cong fst (g .snd) ⟩ + (fst g ∘cr (fst f ∘cr snd A)) .fst ≡⟨ cong (λ h → (fst g ∘cr h) .fst) (f .snd) ⟩ + (fst g ∘cr snd B) .fst ≡⟨ cong fst (g .snd) ⟩ (C .snd .fst) ∎) ⟨_⟩ₐ : (A : CommAlgebra R ℓ') → Type ℓ' @@ -59,18 +71,18 @@ module _ {R : CommRing ℓ} where ⋆Assoc : (r s : ⟨ R ⟩) (x : ⟨ A ⟩ₐ) → (r · s) ⋆ x ≡ r ⋆ (s ⋆ x) ⋆Assoc r s x = cong (_· x) (pres· r s) ∙ sym (·Assoc _ _ _) - where open IsRingHom (A .snd .snd) + where open IsCommRingHom (A .snd .snd) ⋆DistR+ : (r : ⟨ R ⟩) (x y : ⟨ A ⟩ₐ) → r ⋆ (x + y) ≡ (r ⋆ x) + (r ⋆ y) ⋆DistR+ r x y = ·DistR+ _ _ _ ⋆DistL+ : (r s : ⟨ R ⟩) (x : ⟨ A ⟩ₐ) → (r + s) ⋆ x ≡ (r ⋆ x) + (s ⋆ x) ⋆DistL+ r s x = cong (_· x) (pres+ r s) ∙ ·DistL+ _ _ _ - where open IsRingHom (A .snd .snd) + where open IsCommRingHom (A .snd .snd) ⋆IdL : (x : ⟨ A ⟩ₐ) → 1r ⋆ x ≡ x ⋆IdL x = cong (_· x) pres1 ∙ ·IdL x - where open IsRingHom (A .snd .snd) + where open IsCommRingHom (A .snd .snd) ⋆AssocL : (r : ⟨ R ⟩) (x y : ⟨ A ⟩ₐ) → (r ⋆ x) · y ≡ r ⋆ (x · y) ⋆AssocL r x y = sym (·Assoc (A .snd .fst r) x y) From a07d7d7bae517d78e429ea9d2a9fc95934ff5a1a Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Tue, 30 Jul 2024 14:11:18 +0200 Subject: [PATCH 26/83] kernels --- Cubical/Algebra/CommAlgebraAlt/Base.agda | 45 +++++++++++-------- Cubical/Algebra/CommAlgebraAlt/Kernel.agda | 21 +++++++++ .../CommAlgebraAlt/QuotientAlgebra.agda | 10 ++--- 3 files changed, 52 insertions(+), 24 deletions(-) create mode 100644 Cubical/Algebra/CommAlgebraAlt/Kernel.agda diff --git a/Cubical/Algebra/CommAlgebraAlt/Base.agda b/Cubical/Algebra/CommAlgebraAlt/Base.agda index e73fe1ff49..d4437bf5ac 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Base.agda @@ -23,22 +23,9 @@ module _ {R : CommRing ℓ} where CommAlgebra→CommRing : (A : CommAlgebra R ℓ') → CommRing ℓ' CommAlgebra→CommRing = fst - CommAlgebra→Ring : (A : CommAlgebra R ℓ') → Ring ℓ' - CommAlgebra→Ring A = CommRing→Ring (fst A) - CommAlgebraHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') → Type _ CommAlgebraHom A B = Σ[ f ∈ CommRingHom (fst A) (fst B) ] f ∘cr snd A ≡ snd B - CommAlgebraHom→CommRingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} - → CommAlgebraHom A B - → CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) - CommAlgebraHom→CommRingHom = fst - - CommAlgebraHom→RingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} - → CommAlgebraHom A B - → RingHom (CommAlgebra→Ring A) (CommAlgebra→Ring B) - CommAlgebraHom→RingHom = CommRingHom→RingHom ∘ CommAlgebraHom→CommRingHom - idCAlgHom : (A : CommAlgebra R ℓ) → CommAlgebraHom A A idCAlgHom A = (idCommRingHom (fst A)) , Σ≡Prop (λ _ → isPropIsCommRingHom (R .snd) _ (A .fst .snd)) refl @@ -46,12 +33,16 @@ module _ {R : CommRing ℓ} where → (g : CommAlgebraHom B C) → (f : CommAlgebraHom A B) → CommAlgebraHom A C compCommAlgebraHom {A = A} {B = B} {C = C} g f = - ((fst g) ∘cr (fst f)) , - Σ≡Prop (λ _ → isPropIsCommRingHom (R .snd) _ (C .fst .snd)) - ( - (fst g ∘cr (fst f ∘cr snd A)) .fst ≡⟨ cong (λ h → (fst g ∘cr h) .fst) (f .snd) ⟩ - (fst g ∘cr snd B) .fst ≡⟨ cong fst (g .snd) ⟩ - (C .snd .fst) ∎) + ((fst g) ∘cr (fst f)) , pasting + where + opaque + pasting : ((fst g) ∘cr (fst f)) ∘cr snd A ≡ snd C + pasting = + Σ≡Prop (λ _ → isPropIsCommRingHom (R .snd) _ (C .fst .snd)) + ( + (fst g ∘cr (fst f ∘cr snd A)) .fst ≡⟨ cong (λ h → (fst g ∘cr h) .fst) (f .snd) ⟩ + (fst g ∘cr snd B) .fst ≡⟨ cong fst (g .snd) ⟩ + (C .snd .fst) ∎) ⟨_⟩ₐ : (A : CommAlgebra R ℓ') → Type ℓ' ⟨ A ⟩ₐ = A .fst .fst @@ -86,3 +77,19 @@ module _ {R : CommRing ℓ} where ⋆AssocL : (r : ⟨ R ⟩) (x y : ⟨ A ⟩ₐ) → (r ⋆ x) · y ≡ r ⋆ (x · y) ⋆AssocL r x y = sym (·Assoc (A .snd .fst r) x y) + + +{- Convenience forgetful functions -} +module _ {R : CommRing ℓ} where + CommAlgebra→Ring : (A : CommAlgebra R ℓ') → Ring ℓ' + CommAlgebra→Ring A = CommRing→Ring (fst A) + + CommAlgebraHom→CommRingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → CommAlgebraHom A B + → CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) + CommAlgebraHom→CommRingHom = fst + + CommAlgebraHom→RingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → CommAlgebraHom A B + → RingHom (CommAlgebra→Ring A) (CommAlgebra→Ring B) + CommAlgebraHom→RingHom = CommRingHom→RingHom ∘ CommAlgebraHom→CommRingHom diff --git a/Cubical/Algebra/CommAlgebraAlt/Kernel.agda b/Cubical/Algebra/CommAlgebraAlt/Kernel.agda new file mode 100644 index 0000000000..20a8f48c56 --- /dev/null +++ b/Cubical/Algebra/CommAlgebraAlt/Kernel.agda @@ -0,0 +1,21 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebraAlt.Kernel where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.Algebra.CommRing.Base +open import Cubical.Algebra.CommRing.Ideal using (Ideal→CommIdeal) +open import Cubical.Algebra.Ring.Kernel using () renaming (kernelIdeal to ringKernel) +open import Cubical.Algebra.CommAlgebraAlt.Base +--open import Cubical.Algebra.CommAlgebra.Properties +open import Cubical.Algebra.CommAlgebraAlt.Ideal + +private + variable + ℓ : Level + +module _ {R : CommRing ℓ} (A B : CommAlgebra R ℓ) (ϕ : CommAlgebraHom A B) where + + kernel : IdealsIn R A + kernel = Ideal→CommIdeal (ringKernel (CommAlgebraHom→RingHom {A = A} {B = B} ϕ)) diff --git a/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda index f5fe36945e..f0c5bcbda3 100644 --- a/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda @@ -17,7 +17,7 @@ import Cubical.Algebra.Ring.Quotient as Ring open import Cubical.Algebra.CommRing.Ideal hiding (IdealsIn) open import Cubical.Algebra.CommAlgebraAlt.Base open import Cubical.Algebra.CommAlgebraAlt.Ideal --- open import Cubical.Algebra.CommAlgebra.Kernel +open import Cubical.Algebra.CommAlgebraAlt.Kernel -- open import Cubical.Algebra.CommAlgebra.Instances.Unit -- open import Cubical.Algebra.Algebra.Base using (IsAlgebraHom; isPropIsAlgebraHom) open import Cubical.Algebra.Ring @@ -40,7 +40,6 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where open CommRingStr {{...}} -- hiding (_-_; -_; ·IdL ; ·DistR+) renaming (_·_ to _·R_; _+_ to _+R_) open CommAlgebraStr {{...}} open RingTheory (CommRing→Ring (fst A)) using (-DistR·) - open RingHoms instance _ : CommRingStr ⟨ R ⟩ _ = snd R @@ -50,16 +49,15 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where _/_ : CommAlgebra R ℓ _/_ = ((fst A) CommRing./ I) , - (CommRing.quotientHom (fst A) I ∘r A .snd) + (CommRing.quotientHom (fst A) I ∘cr A .snd) - quotientHom : CommAlgebraHom {R = R} A (_/_) + quotientHom : CommAlgebraHom A (_/_) quotientHom = (CommRing.quotientHom (fst A) I) , refl module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where open CommRingStr ⦃...⦄ open CommAlgebraStr ⦃...⦄ - instance _ : CommRingStr ⟨ R ⟩ _ = snd R @@ -81,6 +79,7 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where IsRingHom.pres+ (snd CommForget/) = λ _ _ → refl IsRingHom.pres· (snd CommForget/) = λ _ _ → refl IsRingHom.pres- (snd CommForget/) = λ _ → refl +-} module _ (B : CommAlgebra R ℓ) @@ -88,6 +87,7 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where (I⊆kernel : (fst I) ⊆ (fst (kernel A B ϕ))) where +{- open IsAlgebraHom open RingTheory (CommRing→Ring (CommAlgebra→CommRing B)) From bd528216067ff3d197d0fd1e1988cea5a67a4aad Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Tue, 30 Jul 2024 17:50:35 +0200 Subject: [PATCH 27/83] syntax for new algebras --- Cubical/Algebra/CommAlgebraAlt/Base.agda | 21 +++++++++++++++++++++ 1 file changed, 21 insertions(+) diff --git a/Cubical/Algebra/CommAlgebraAlt/Base.agda b/Cubical/Algebra/CommAlgebraAlt/Base.agda index d4437bf5ac..fd487e0c92 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Base.agda @@ -44,9 +44,30 @@ module _ {R : CommRing ℓ} where (fst g ∘cr snd B) .fst ≡⟨ cong fst (g .snd) ⟩ (C .snd .fst) ∎) + ⟨_⟩ₐ : (A : CommAlgebra R ℓ') → Type ℓ' ⟨ A ⟩ₐ = A .fst .fst + _$ca_ : {A : CommAlgebra R ℓ} {B : CommAlgebra R ℓ'} → (φ : CommAlgebraHom A B) → (x : ⟨ A ⟩ₐ) → ⟨ B ⟩ₐ + φ $ca x = φ .fst .fst x + + _∘ca_ : {A : CommAlgebra R ℓ} {B : CommAlgebra R ℓ'} {C : CommAlgebra R ℓ''} + → CommAlgebraHom B C → CommAlgebraHom A B → CommAlgebraHom A C + g ∘ca f = compCommAlgebraHom g f + + opaque + CommAlgebraHom≡ : + {A : CommAlgebra R ℓ} {B : CommAlgebra R ℓ'} {f g : CommAlgebraHom A B} + → f .fst .fst ≡ g .fst .fst + → f ≡ g + CommAlgebraHom≡ {B = B} p = + Σ≡Prop (λ _ → isSetΣSndProp (isSetΠ (λ _ → is-set)) + (λ _ → isPropIsCommRingHom _ _ _) + _ _) + (Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) + p) + where open CommRingStr (B .fst .snd) using (is-set) + module CommAlgebraStr (A : CommAlgebra R ℓ') where open CommRingStr {{...}} instance From 3a9a310a4376c3dd0d8af30e8c2e23fa92ddf826 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Tue, 30 Jul 2024 17:50:54 +0200 Subject: [PATCH 28/83] more on comm ring quotients --- Cubical/Algebra/CommRing/Quotient/Base.agda | 21 +++++++++++++-------- 1 file changed, 13 insertions(+), 8 deletions(-) diff --git a/Cubical/Algebra/CommRing/Quotient/Base.agda b/Cubical/Algebra/CommRing/Quotient/Base.agda index d89fd6deff..0f63b3eff3 100644 --- a/Cubical/Algebra/CommRing/Quotient/Base.agda +++ b/Cubical/Algebra/CommRing/Quotient/Base.agda @@ -80,8 +80,18 @@ module Coherence (R : CommRing ℓ) (I : IdealsIn R) where fst ringStrInv x = x (snd ringStrInv) = isRingHomCohInv + open RingHoms +quotientHom : (R : CommRing ℓ) → (I : IdealsIn R) → CommRingHom R (R / I) +quotientHom R I = RingHom→CommRingHom $ + Coherence.ringStrInv R I + ∘r Ring.quotientHom (CommRing→Ring R) (CommIdeal→Ideal I) + +quotientHomSurjective : (R : CommRing ℓ) → (I : IdealsIn R) + → isSurjection (fst (quotientHom R I)) +quotientHomSurjective R I = Ring.quotientHomSurjective (CommRing→Ring R) (CommIdeal→Ideal I) + module Quotient-FGideal-CommRing-Ring (R : CommRing ℓ) (S : Ring ℓ') @@ -131,14 +141,9 @@ module UniversalProperty Ring.UniversalProperty.inducedHom (CommRing→Ring R) (CommIdeal→Ideal I) (CommRingHom→RingHom f) I⊆ker ∘r Coherence.ringStr R I -quotientHom : (R : CommRing ℓ) → (I : IdealsIn R) → CommRingHom R (R / I) -quotientHom R I = RingHom→CommRingHom $ - Coherence.ringStrInv R I - ∘r Ring.quotientHom (CommRing→Ring R) (CommIdeal→Ideal I) - -quotientHomSurjective : (R : CommRing ℓ) → (I : IdealsIn R) - → isSurjection (fst (quotientHom R I)) -quotientHomSurjective R I = Ring.quotientHomSurjective (CommRing→Ring R) (CommIdeal→Ideal I) + isSolution : inducedHom ∘cr quotientHom R I ≡ f + isSolution = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) + (cong fst (Ring.UniversalProperty.solution (CommRing→Ring R) (CommIdeal→Ideal I) (CommRingHom→RingHom f) I⊆ker)) module _ {R : CommRing ℓ} (I : IdealsIn R) where open CommRingStr ⦃...⦄ From 4a9842feaa8882dbe331409fc2c5ed2ee26df94a Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Tue, 30 Jul 2024 17:51:02 +0200 Subject: [PATCH 29/83] fix --- Cubical/Algebra/CommAlgebra/Base.agda | 10 ++++---- Cubical/Algebra/CommAlgebra/Properties.agda | 28 ++++++++++----------- 2 files changed, 19 insertions(+), 19 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/Base.agda index 12a890574e..fe04691d36 100644 --- a/Cubical/Algebra/CommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/Base.agda @@ -224,11 +224,11 @@ module _ {R : CommRing ℓ} where → CommAlgebraHom A B → CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) fst (CommAlgebraHom→CommRingHom A B f) = fst f - IsRingHom.pres0 (snd (CommAlgebraHom→CommRingHom A B f)) = IsAlgebraHom.pres0 (snd f) - IsRingHom.pres1 (snd (CommAlgebraHom→CommRingHom A B f)) = IsAlgebraHom.pres1 (snd f) - IsRingHom.pres+ (snd (CommAlgebraHom→CommRingHom A B f)) = IsAlgebraHom.pres+ (snd f) - IsRingHom.pres· (snd (CommAlgebraHom→CommRingHom A B f)) = IsAlgebraHom.pres· (snd f) - IsRingHom.pres- (snd (CommAlgebraHom→CommRingHom A B f)) = IsAlgebraHom.pres- (snd f) + IsCommRingHom.pres0 (snd (CommAlgebraHom→CommRingHom A B f)) = IsAlgebraHom.pres0 (snd f) + IsCommRingHom.pres1 (snd (CommAlgebraHom→CommRingHom A B f)) = IsAlgebraHom.pres1 (snd f) + IsCommRingHom.pres+ (snd (CommAlgebraHom→CommRingHom A B f)) = IsAlgebraHom.pres+ (snd f) + IsCommRingHom.pres· (snd (CommAlgebraHom→CommRingHom A B f)) = IsAlgebraHom.pres· (snd f) + IsCommRingHom.pres- (snd (CommAlgebraHom→CommRingHom A B f)) = IsAlgebraHom.pres- (snd f) module _ {M : CommAlgebra R ℓ'} {N : CommAlgebra R ℓ''} where open CommAlgebraStr {{...}} diff --git a/Cubical/Algebra/CommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebra/Properties.agda index 38123c1107..8d10d81366 100644 --- a/Cubical/Algebra/CommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/Properties.agda @@ -78,7 +78,7 @@ module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where λ r x y → sym (·Assoc _ _ _) where open CommRingStr (snd A) - open IsRingHom φIsHom + open IsCommRingHom φIsHom fromCommAlg : CommAlgebra R ℓ' → CommRingWithHom fromCommAlg A = (CommAlgebra→CommRing A) , φ , φIsHom @@ -88,8 +88,8 @@ module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where open AlgebraTheory (CommRing→Ring R) (CommAlgebra→Algebra A) φ : ⟨ R ⟩ → ⟨ A ⟩ φ r = r ⋆ 1a - φIsHom : IsRingHom (CommRing→Ring R .snd) φ (CommRing→Ring (CommAlgebra→CommRing A) .snd) - φIsHom = makeIsRingHom (⋆IdL _) (λ _ _ → ⋆DistL+ _ _ _) + φIsHom : IsCommRingHom (R .snd) φ ((CommAlgebra→CommRing A) .snd) + φIsHom = makeIsCommRingHom (⋆IdL _) (λ _ _ → ⋆DistL+ _ _ _) λ x y → cong (λ a → (x ·r y) ⋆ a) (sym (·IdL _)) ∙ ⋆Dist· _ _ _ _ -- helpful for localisations @@ -107,7 +107,7 @@ module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where -- note that the proofs of the axioms might differ! isCommRing (snd (CommAlgebra→CommRing≡ i)) = isProp→PathP (λ i → isPropIsCommRing _ _ _ _ _ ) (isCommRing (snd (CommAlgebra→CommRing (toCommAlg Aφ)))) (isCommRing (snd A)) i - +{- CommRingWithHomRoundTrip : (Aφ : CommRingWithHom) → fromCommAlg (toCommAlg Aφ) ≡ Aφ CommRingWithHomRoundTrip (A , φ) = ΣPathP (CommAlgebra→CommRing≡ (A , φ) , φPathP) where @@ -116,7 +116,7 @@ module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where φPathP : PathP (λ i → CommRingHom R (CommAlgebra→CommRing≡ (A , φ) i)) (snd (fromCommAlg (toCommAlg (A , φ)))) φ φPathP = RingHomPathP _ _ _ _ _ _ λ i x → ·IdR (snd A) (fst φ x) i - +-} CommAlgRoundTrip : (A : CommAlgebra R ℓ') → toCommAlg (fromCommAlg A) ≡ A CommAlgRoundTrip A = ΣPathP (refl , AlgStrPathP) where @@ -133,21 +133,20 @@ module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where CommAlgebraStr.isCommAlgebra (AlgStrPathP i) = isProp→PathP (λ i → isPropIsCommAlgebra _ _ _ _ _ _ (CommAlgebraStr._⋆_ (AlgStrPathP i))) (CommAlgebraStr.isCommAlgebra (snd (toCommAlg (fromCommAlg A)))) isCommAlgebra i - +{- CommAlgIso : Iso (CommAlgebra R ℓ') CommRingWithHom fun CommAlgIso = fromCommAlg inv CommAlgIso = toCommAlg rightInv CommAlgIso = CommRingWithHomRoundTrip leftInv CommAlgIso = CommAlgRoundTrip - - open RingHoms - open IsRingHom +-} + open IsCommRingHom isCommRingWithHomHom : (A B : CommRingWithHom) → CommRingHom (fst A) (fst B) → Type (ℓ-max ℓ ℓ') - isCommRingWithHomHom (_ , f) (_ , g) h = h ∘r f ≡ g + isCommRingWithHomHom (_ , f) (_ , g) h = h ∘cr f ≡ g CommRingWithHomHom : CommRingWithHom → CommRingWithHom → Type (ℓ-max ℓ ℓ') - CommRingWithHomHom (A , f) (B , g) = Σ[ h ∈ CommRingHom A B ] h ∘r f ≡ g + CommRingWithHomHom (A , f) (B , g) = Σ[ h ∈ CommRingHom A B ] h ∘cr f ≡ g toCommAlgebraHom : (A B : CommRingWithHom) (h : CommRingHom (fst A) (fst B)) → isCommRingWithHomHom A B h @@ -173,13 +172,13 @@ module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where pres· (snd (fst (fromCommAlgebraHom A B f))) = IsAlgebraHom.pres· (snd f) pres- (snd (fst (fromCommAlgebraHom A B f))) = IsAlgebraHom.pres- (snd f) snd (fromCommAlgebraHom A B f) = - RingHom≡ (funExt (λ x → IsAlgebraHom.pres⋆ (snd f) x 1a ∙ cong (x ⋆_) (IsAlgebraHom.pres1 (snd f)))) + CommRingHom≡ (funExt (λ x → IsAlgebraHom.pres⋆ (snd f) x 1a ∙ cong (x ⋆_) (IsAlgebraHom.pres1 (snd f)))) where open CommAlgebraStr (snd A) using (1a) open CommAlgebraStr (snd B) using (_⋆_) isCommRingWithHomEquiv : (A B : CommRingWithHom) → CommRingEquiv (fst A) (fst B) → Type (ℓ-max ℓ ℓ') - isCommRingWithHomEquiv A B e = isCommRingWithHomHom A B (RingEquiv→RingHom e) + isCommRingWithHomEquiv A B e = isCommRingWithHomHom A B (CommRingEquiv→CommRingHom e) CommRingWithHomEquiv : CommRingWithHom → CommRingWithHom → Type (ℓ-max ℓ ℓ') CommRingWithHomEquiv A B = Σ[ e ∈ CommRingEquiv (fst A) (fst B) ] isCommRingWithHomEquiv A B e @@ -337,7 +336,7 @@ module _ {R : CommRing ℓ} {A B : CommAlgebra R ℓ'} where pres- (snd (CommAlgebraHomFromRingHom ϕ pres*)) = IsRingHom.pres- (snd ϕ) pres⋆ (snd (CommAlgebraHomFromRingHom ϕ pres*)) = pres* - +{- module _ {R S : CommRing ℓ} (f : CommRingHom S R) where baseChange : CommAlgebra R ℓ' → CommAlgebra S ℓ' baseChange A = @@ -359,3 +358,4 @@ module _ {R S : CommRing ℓ} (f : CommRingHom S R) where pbSliceHom : Σ[ k ∈ CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) ] k ∘r ((snd homA) ∘r f) ≡ ((snd homB) ∘r f) pbSliceHom = fst asSliceHom , Σ≡Prop (λ _ → isPropIsRingHom _ _ _) λ i x → fst ((snd asSliceHom) i) (fst f x) +-} From f1433a8cf785714f333fec9d218dd8806dc75223 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Tue, 30 Jul 2024 18:08:29 +0200 Subject: [PATCH 30/83] algebra quotients --- Cubical/Algebra/CommAlgebraAlt/Base.agda | 3 + Cubical/Algebra/CommAlgebraAlt/FGIdeal.agda | 37 +++++++++ .../CommAlgebraAlt/QuotientAlgebra.agda | 80 +++++++++---------- Cubical/Algebra/CommRing/Quotient/Base.agda | 28 ++++++- Cubical/Algebra/Ring/Quotient.agda | 6 ++ 5 files changed, 110 insertions(+), 44 deletions(-) create mode 100644 Cubical/Algebra/CommAlgebraAlt/FGIdeal.agda diff --git a/Cubical/Algebra/CommAlgebraAlt/Base.agda b/Cubical/Algebra/CommAlgebraAlt/Base.agda index fd487e0c92..b50375ccef 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Base.agda @@ -48,6 +48,9 @@ module _ {R : CommRing ℓ} where ⟨_⟩ₐ : (A : CommAlgebra R ℓ') → Type ℓ' ⟨ A ⟩ₐ = A .fst .fst + ⟨_⟩ₐ→ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} (f : CommAlgebraHom A B) → ⟨ A ⟩ₐ → ⟨ B ⟩ₐ + ⟨ f ⟩ₐ→ = f .fst .fst + _$ca_ : {A : CommAlgebra R ℓ} {B : CommAlgebra R ℓ'} → (φ : CommAlgebraHom A B) → (x : ⟨ A ⟩ₐ) → ⟨ B ⟩ₐ φ $ca x = φ .fst .fst x diff --git a/Cubical/Algebra/CommAlgebraAlt/FGIdeal.agda b/Cubical/Algebra/CommAlgebraAlt/FGIdeal.agda new file mode 100644 index 0000000000..6de95db747 --- /dev/null +++ b/Cubical/Algebra/CommAlgebraAlt/FGIdeal.agda @@ -0,0 +1,37 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebraAlt.FGIdeal where +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Powerset + +open import Cubical.Data.FinData +open import Cubical.Data.Nat +open import Cubical.Data.Vec + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.FGIdeal using () + renaming (generatedIdeal to generatedIdealCommRing; + indInIdeal to ringIncInIdeal; + 0FGIdeal to 0FGIdealCommRing) +open import Cubical.Algebra.CommAlgebraAlt.Base +open import Cubical.Algebra.CommAlgebraAlt.Ideal + +private + variable + ℓ : Level + R : CommRing ℓ + +generatedIdeal : {n : ℕ} (A : CommAlgebra R ℓ) → FinVec ⟨ A ⟩ₐ n → IdealsIn R A +generatedIdeal A = generatedIdealCommRing (CommAlgebra→CommRing A) + +incInIdeal : {n : ℕ} (A : CommAlgebra R ℓ) + (U : FinVec ⟨ A ⟩ₐ n) (i : Fin n) → U i ∈ fst (generatedIdeal A U) +incInIdeal A = ringIncInIdeal (CommAlgebra→CommRing A) + +syntax generatedIdeal A V = ⟨ V ⟩[ A ] + +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) where + open CommRingStr (A .fst .snd) + + 0FGIdeal : {n : ℕ} → ⟨ replicateFinVec n 0r ⟩[ A ] ≡ (0Ideal R A) + 0FGIdeal = 0FGIdealCommRing (CommAlgebra→CommRing A) diff --git a/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda index f0c5bcbda3..135d7b469e 100644 --- a/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda @@ -3,6 +3,7 @@ module Cubical.Algebra.CommAlgebraAlt.QuotientAlgebra where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Powerset using (_∈_; _⊆_) open import Cubical.Foundations.Structure @@ -13,7 +14,6 @@ open import Cubical.Data.Sigma.Properties using (Σ≡Prop) open import Cubical.Algebra.CommRing import Cubical.Algebra.CommRing.Quotient as CommRing -import Cubical.Algebra.Ring.Quotient as Ring open import Cubical.Algebra.CommRing.Ideal hiding (IdealsIn) open import Cubical.Algebra.CommAlgebraAlt.Base open import Cubical.Algebra.CommAlgebraAlt.Ideal @@ -65,64 +65,60 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where _ = A .fst .snd _ = A -{- - - opaque - unfolding _/_ - - -- sanity check / maybe a helper function some day - -- (These two rings are not definitionally equal, but only because of proofs, not data.) - CommForget/ : RingEquiv (CommAlgebra→Ring (A / I)) ((CommAlgebra→Ring A) Ring./ (CommIdeal→Ideal I)) - fst CommForget/ = idEquiv _ - IsRingHom.pres0 (snd CommForget/) = refl - IsRingHom.pres1 (snd CommForget/) = refl - IsRingHom.pres+ (snd CommForget/) = λ _ _ → refl - IsRingHom.pres· (snd CommForget/) = λ _ _ → refl - IsRingHom.pres- (snd CommForget/) = λ _ → refl --} - module _ (B : CommAlgebra R ℓ) (ϕ : CommAlgebraHom A B) (I⊆kernel : (fst I) ⊆ (fst (kernel A B ϕ))) where -{- - open IsAlgebraHom open RingTheory (CommRing→Ring (CommAlgebra→CommRing B)) + open CommAlgebraStr ⦃...⦄ private instance - _ : CommAlgebraStr R ⟨ B ⟩ - _ = snd B - _ : CommRingStr ⟨ B ⟩ + _ = B + _ : CommRingStr ⟨ B ⟩ₐ _ = snd (CommAlgebra→CommRing B) + inducedHom : CommAlgebraHom (A / I) B + inducedHom .fst = CommRing.UniversalProperty.inducedHom + (CommAlgebra→CommRing A) + (CommAlgebra→CommRing B) + I + (CommAlgebraHom→CommRingHom ϕ) + I⊆kernel + inducedHom .snd = p + where + step1 = cong (inducedHom .fst ∘cr_) (sym (quotientHom A I .snd)) + step2 = compAssocCommRingHom (A .snd) (CommRing.quotientHom (A .fst) I) (inducedHom .fst) + step3 = cong (_∘cr A .snd) (CommRing.UniversalProperty.isSolution + (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) I (CommAlgebraHom→CommRingHom ϕ) I⊆kernel) + opaque + p : (inducedHom .fst ∘cr snd (A / I)) ≡ snd B + p = (inducedHom .fst ∘cr snd (A / I)) ≡⟨ step1 ⟩ + (inducedHom .fst ∘cr (CommRing.quotientHom (A .fst) I ∘cr A .snd)) ≡⟨ step2 ⟩ + (inducedHom .fst ∘cr CommRing.quotientHom (A .fst) I) ∘cr A .snd ≡⟨ step3 ⟩ + (ϕ .fst) ∘cr A .snd ≡⟨ ϕ .snd ⟩ + snd B ∎ + opaque - unfolding _/_ - - inducedHom : CommAlgebraHom (A / I) B - fst inducedHom = - rec is-set (λ x → fst ϕ x) - λ a b a-b∈I → - equalByDifference - (fst ϕ a) (fst ϕ b) - ((fst ϕ a) - (fst ϕ b) ≡⟨ cong (λ u → (fst ϕ a) + u) (sym (IsAlgebraHom.pres- (snd ϕ) _)) ⟩ - (fst ϕ a) + (fst ϕ (- b)) ≡⟨ sym (IsAlgebraHom.pres+ (snd ϕ) _ _) ⟩ - fst ϕ (a - b) ≡⟨ I⊆kernel (a - b) a-b∈I ⟩ - 0r ∎) - pres0 (snd inducedHom) = pres0 (snd ϕ) - pres1 (snd inducedHom) = pres1 (snd ϕ) - pres+ (snd inducedHom) = elimProp2 (λ _ _ → is-set _ _) (pres+ (snd ϕ)) - pres· (snd inducedHom) = elimProp2 (λ _ _ → is-set _ _) (pres· (snd ϕ)) - pres- (snd inducedHom) = elimProp (λ _ → is-set _ _) (pres- (snd ϕ)) - pres⋆ (snd inducedHom) = λ r → elimProp (λ _ → is-set _ _) (pres⋆ (snd ϕ) r) + inducedHom∘quotientHom : inducedHom ∘ca quotientHom A I ≡ ϕ + inducedHom∘quotientHom = CommAlgebraHom≡ (funExt (λ _ → refl)) + + opaque + isUnique : (ψ : CommAlgebraHom (A / I) B) (ψIsSolution : ⟨ ψ ⟩ₐ→ ∘ ⟨ quotientHom A I ⟩ₐ→ ≡ ⟨ ϕ ⟩ₐ→) + → ψ ≡ inducedHom + isUnique ψ ψIsSolution = + CommAlgebraHom≡ + (cong fst $ + CommRing.UniversalProperty.isUnique + (A .fst) (B .fst) I (ϕ .fst) I⊆kernel (ψ .fst) ψIsSolution) +{- + opaque unfolding inducedHom quotientHom - inducedHom∘quotientHom : inducedHom ∘a quotientHom A I ≡ ϕ - inducedHom∘quotientHom = Σ≡Prop (isPropIsCommAlgebraHom {M = A} {N = B}) (funExt (λ a → refl)) opaque unfolding quotientHom diff --git a/Cubical/Algebra/CommRing/Quotient/Base.agda b/Cubical/Algebra/CommRing/Quotient/Base.agda index 0f63b3eff3..506bae8386 100644 --- a/Cubical/Algebra/CommRing/Quotient/Base.agda +++ b/Cubical/Algebra/CommRing/Quotient/Base.agda @@ -138,12 +138,36 @@ module UniversalProperty inducedHom : CommRingHom (R / I) S inducedHom = RingHom→CommRingHom $ - Ring.UniversalProperty.inducedHom (CommRing→Ring R) (CommIdeal→Ideal I) (CommRingHom→RingHom f) I⊆ker + Ring.UniversalProperty.inducedHom + (CommRing→Ring R) + (CommIdeal→Ideal I) + (CommRingHom→RingHom f) + I⊆ker ∘r Coherence.ringStr R I isSolution : inducedHom ∘cr quotientHom R I ≡ f isSolution = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) - (cong fst (Ring.UniversalProperty.solution (CommRing→Ring R) (CommIdeal→Ideal I) (CommRingHom→RingHom f) I⊆ker)) + (cong fst (Ring.UniversalProperty.solution + (CommRing→Ring R) + (CommIdeal→Ideal I) + (CommRingHom→RingHom f) + I⊆ker)) + + opaque + unfolding quotientCommRingStr + isUnique : (ψ : CommRingHom (R / I) S) → (ψIsSolution : ψ .fst ∘ quotientHom R I .fst ≡ f .fst) + → ψ ≡ inducedHom + isUnique ψ ψIsSolution = + Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) + (cong fst + (Ring.UniversalProperty.unique' + (CommRing→Ring R) + (CommIdeal→Ideal I) + (CommRingHom→RingHom f) + I⊆ker + (CommRingHom→RingHom ψ) + ψIsSolution)) + module _ {R : CommRing ℓ} (I : IdealsIn R) where open CommRingStr ⦃...⦄ diff --git a/Cubical/Algebra/Ring/Quotient.agda b/Cubical/Algebra/Ring/Quotient.agda index 8b4b879019..3ca0a47743 100644 --- a/Cubical/Algebra/Ring/Quotient.agda +++ b/Cubical/Algebra/Ring/Quotient.agda @@ -3,6 +3,7 @@ module Cubical.Algebra.Ring.Quotient where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels open import Cubical.Foundations.Structure open import Cubical.Foundations.Powerset using (_∈_; _⊆_; ⊆-extensionality) -- \in, \sub= @@ -291,6 +292,11 @@ module UniversalProperty (R : Ring ℓ) (I : IdealsIn R) where → ψ ≡ inducedHom p unique p ψ ψIsSolution = quotientHomEpi R I S ψ (inducedHom p) (ψIsSolution ∙ sym (solution p)) + unique' : (p : ((x : ⟨ R ⟩) → x ∈ Iₛ → φ $r x ≡ 0r)) + → (ψ : RingHom (R / I) S) → (ψIsSolution : ψ .fst ∘ quotientHom R I .fst ≡ φ .fst) + → ψ ≡ inducedHom p + unique' p ψ ψIsSolution = unique p ψ (RingHom≡ ψIsSolution) + {- Show that the kernel of the quotient map π : R ─→ R/I From d8b07b48ff6a5b73972758dbb6950f6e961d74ae Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Wed, 31 Jul 2024 10:41:10 +0200 Subject: [PATCH 31/83] more opaque definitions, more use of set-level equalities --- .../CommAlgebraAlt/QuotientAlgebra.agda | 31 +++++------- Cubical/Algebra/CommRing/Quotient/Base.agda | 47 ++++++++++++------- Cubical/Algebra/Ring/Quotient.agda | 30 ++++++------ Cubical/Foundations/Structure.agda | 14 ++++++ 4 files changed, 72 insertions(+), 50 deletions(-) diff --git a/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda index 135d7b469e..d241aedeee 100644 --- a/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda @@ -30,14 +30,8 @@ private variable ℓ : Level -{- - The definition of the quotient algebra (_/_ below) is marked opaque to avoid - long type checking times. All other definitions that need to "look into" this - opaque definition must be in an opaque block that unfolds the definition of _/_. --} - module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where - open CommRingStr {{...}} -- hiding (_-_; -_; ·IdL ; ·DistR+) renaming (_·_ to _·R_; _+_ to _+R_) + open CommRingStr {{...}} open CommAlgebraStr {{...}} open RingTheory (CommRing→Ring (fst A)) using (-DistR·) instance @@ -113,28 +107,25 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where (cong fst $ CommRing.UniversalProperty.isUnique (A .fst) (B .fst) I (ϕ .fst) I⊆kernel (ψ .fst) ψIsSolution) -{- - opaque - unfolding inducedHom quotientHom +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where + open CommRingStr {{...}} + open CommAlgebraStr {{...}} opaque - unfolding quotientHom +-- unfolding quotientHom injectivePrecomp : (B : CommAlgebra R ℓ) (f g : CommAlgebraHom (A / I) B) - → f ∘a (quotientHom A I) ≡ g ∘a (quotientHom A I) + → f ∘ca (quotientHom A I) ≡ g ∘ca (quotientHom A I) → f ≡ g injectivePrecomp B f g p = - Σ≡Prop - (λ h → isPropIsCommAlgebraHom {M = A / I} {N = B} h) - (descendMapPath (fst f) (fst g) is-set - λ x → λ i → fst (p i) x) - where - instance - _ : CommAlgebraStr R ⟨ B ⟩ - _ = str B + CommAlgebraHom≡ ? +-- (descendMapPath ⟨ f ⟩ₐ→ ⟨ g ⟩ₐ→ is-set +-- λ x → λ i → fst (p i) x) +-- where +{- {- trivial quotient -} diff --git a/Cubical/Algebra/CommRing/Quotient/Base.agda b/Cubical/Algebra/CommRing/Quotient/Base.agda index 506bae8386..a377b436f8 100644 --- a/Cubical/Algebra/CommRing/Quotient/Base.agda +++ b/Cubical/Algebra/CommRing/Quotient/Base.agda @@ -1,8 +1,9 @@ -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.Algebra.CommRing.Quotient.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure +open import Cubical.Foundations.HLevels open import Cubical.Foundations.Function open import Cubical.Foundations.Powerset open import Cubical.Functions.Surjection @@ -23,7 +24,7 @@ import Cubical.Algebra.Ring.Quotient as Ring private variable - ℓ ℓ' : Level + ℓ ℓ' ℓ'' : Level module _ (R : CommRing ℓ) (I : IdealsIn R) where open CommRingStr (snd R) @@ -83,14 +84,26 @@ module Coherence (R : CommRing ℓ) (I : IdealsIn R) where open RingHoms -quotientHom : (R : CommRing ℓ) → (I : IdealsIn R) → CommRingHom R (R / I) -quotientHom R I = RingHom→CommRingHom $ - Coherence.ringStrInv R I - ∘r Ring.quotientHom (CommRing→Ring R) (CommIdeal→Ideal I) +module _ (R : CommRing ℓ) (I : IdealsIn R) where -quotientHomSurjective : (R : CommRing ℓ) → (I : IdealsIn R) - → isSurjection (fst (quotientHom R I)) -quotientHomSurjective R I = Ring.quotientHomSurjective (CommRing→Ring R) (CommIdeal→Ideal I) + quotientHom : CommRingHom R (R / I) + quotientHom = + withOpaqueStr $ + RingHom→CommRingHom $ + Coherence.ringStrInv R I + ∘r Ring.quotientHom (CommRing→Ring R) (CommIdeal→Ideal I) + + quotientHomSurjective : isSurjection (quotientHom .fst) + quotientHomSurjective = Ring.quotientHomSurjective (CommRing→Ring R) (CommIdeal→Ideal I) + + quotientHomEpi : (S : hSet ℓ') + → (f g : ⟨ R / I ⟩ → ⟨ S ⟩) + → f ∘ quotientHom .fst ≡ g ∘ quotientHom .fst + → f ≡ g + quotientHomEpi S f g p = + (Ring.quotientHomEpi + (CommRing→Ring R) (CommIdeal→Ideal I) S + f g p) module Quotient-FGideal-CommRing-Ring (R : CommRing ℓ) @@ -137,13 +150,15 @@ module UniversalProperty where inducedHom : CommRingHom (R / I) S - inducedHom = RingHom→CommRingHom $ - Ring.UniversalProperty.inducedHom - (CommRing→Ring R) - (CommIdeal→Ideal I) - (CommRingHom→RingHom f) - I⊆ker - ∘r Coherence.ringStr R I + inducedHom = + withOpaqueStr $ + RingHom→CommRingHom $ + Ring.UniversalProperty.inducedHom + (CommRing→Ring R) + (CommIdeal→Ideal I) + (CommRingHom→RingHom f) + I⊆ker + ∘r Coherence.ringStr R I isSolution : inducedHom ∘cr quotientHom R I ≡ f isSolution = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) diff --git a/Cubical/Algebra/Ring/Quotient.agda b/Cubical/Algebra/Ring/Quotient.agda index 3ca0a47743..326cfb6036 100644 --- a/Cubical/Algebra/Ring/Quotient.agda +++ b/Cubical/Algebra/Ring/Quotient.agda @@ -217,21 +217,19 @@ module _ (R : Ring ℓ) (I : IdealsIn R) where quotientHomSurjective = []surjective open RingHoms - quotientHomEpi : (S : Ring ℓ') - → (f g : RingHom (R / I) S) - → f ∘r quotientHom R I ≡ g ∘r quotientHom R I + quotientHomEpi : (S : hSet ℓ') + → (f g : ⟨ R / I ⟩ → ⟨ S ⟩) + → f ∘ fst (quotientHom R I) ≡ g ∘ fst (quotientHom R I) → f ≡ g quotientHomEpi S f g p = - Σ≡Prop (λ _ → isPropIsRingHom _ _ _) - (funExt λ x - → PT.rec - (is-set _ _) - (λ {(x' , [x']≡x) → f .fst x ≡⟨ cong (λ y → f .fst y) (sym [x']≡x) ⟩ - fst (f ∘r quotientHom R I) x' ≡⟨ cong (λ h → fst h x') p ⟩ - fst (g ∘r quotientHom R I) x' ≡⟨ cong (λ y → g .fst y) [x']≡x ⟩ - g .fst x ∎ }) - (quotientHomSurjective x )) - where open RingStr (S .snd) using (is-set) + (funExt λ x + → PT.rec + (S .snd _ _) + (λ {(x' , [x']≡x) → f x ≡⟨ cong (λ y → f y) (sym [x']≡x) ⟩ + (f ∘ fst (quotientHom R I)) x' ≡⟨ cong (λ h → h x') p ⟩ + (g ∘ fst (quotientHom R I)) x' ≡⟨ cong (λ y → g y) [x']≡x ⟩ + g x ∎ }) + (quotientHomSurjective x )) module UniversalProperty (R : Ring ℓ) (I : IdealsIn R) where @@ -290,7 +288,11 @@ module UniversalProperty (R : Ring ℓ) (I : IdealsIn R) where unique : (p : ((x : ⟨ R ⟩) → x ∈ Iₛ → φ $r x ≡ 0r)) → (ψ : RingHom (R / I) S) → (ψIsSolution : ψ ∘r quotientHom R I ≡ φ) → ψ ≡ inducedHom p - unique p ψ ψIsSolution = quotientHomEpi R I S ψ (inducedHom p) (ψIsSolution ∙ sym (solution p)) + unique p ψ ψIsSolution = + RingHom≡ + (quotientHomEpi R I (S .fst , isSetS) (ψ .fst) (inducedHom p .fst) + (cong fst $ ψIsSolution ∙ sym (solution p))) + where open RingStr (S .snd) using () renaming (is-set to isSetS) unique' : (p : ((x : ⟨ R ⟩) → x ∈ Iₛ → φ $r x ≡ 0r)) → (ψ : RingHom (R / I) S) → (ψIsSolution : ψ .fst ∘ quotientHom R I .fst ≡ φ .fst) diff --git a/Cubical/Foundations/Structure.agda b/Cubical/Foundations/Structure.agda index eba47e0b78..8d143f0b73 100644 --- a/Cubical/Foundations/Structure.agda +++ b/Cubical/Foundations/Structure.agda @@ -15,6 +15,20 @@ private TypeWithStr : (ℓ : Level) (S : Type ℓ → Type ℓ') → Type (ℓ-max (ℓ-suc ℓ) ℓ') TypeWithStr ℓ S = Σ[ X ∈ Type ℓ ] S X +{- + A helper to make the second component opaque + it is a good idea to use this for set-level structures, + where the second component consists only of proofs of identities + or other propositions +-} +withOpaqueStr : {A : Type ℓ'} {B : A → Type ℓ''} + → Σ[ x ∈ A ] B x + → Σ[ x ∈ A ] B x +withOpaqueStr {B = B} (x , str) = x , str' + where opaque + str' : B x + str' = str + typ : TypeWithStr ℓ S → Type ℓ typ = fst From 55d58c677dedecb5ae2c798adc70eded9fc9a1ae Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 1 Aug 2024 10:12:06 +0200 Subject: [PATCH 32/83] commRingHom fixes --- Cubical/Algebra/CommAlgebra/Localisation.agda | 24 +++++++++---------- .../CommAlgebra/LocalisationAlgebra.agda | 10 ++++---- Cubical/Algebra/CommAlgebra/Properties.agda | 4 ++++ .../Polynomials/UnivariatePolyList.agda | 2 +- 4 files changed, 22 insertions(+), 18 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/Localisation.agda b/Cubical/Algebra/CommAlgebra/Localisation.agda index 38856fbbbc..17745f411f 100644 --- a/Cubical/Algebra/CommAlgebra/Localisation.agda +++ b/Cubical/Algebra/CommAlgebra/Localisation.agda @@ -91,15 +91,15 @@ module AlgLoc (R' : CommRing ℓ) χᴬ : CommAlgebraHom S⁻¹RAsCommAlg B' fst χᴬ = fst χ - pres0 (snd χᴬ) = IsRingHom.pres0 (snd χ) - pres1 (snd χᴬ) = IsRingHom.pres1 (snd χ) - pres+ (snd χᴬ) = IsRingHom.pres+ (snd χ) - pres· (snd χᴬ) = IsRingHom.pres· (snd χ) - pres- (snd χᴬ) = IsRingHom.pres- (snd χ) + pres0 (snd χᴬ) = IsCommRingHom.pres0 (snd χ) + pres1 (snd χᴬ) = IsCommRingHom.pres1 (snd χ) + pres+ (snd χᴬ) = IsCommRingHom.pres+ (snd χ) + pres· (snd χᴬ) = IsCommRingHom.pres· (snd χ) + pres- (snd χᴬ) = IsCommRingHom.pres- (snd χ) pres⋆ (snd χᴬ) r x = path where path : fst χ ((r /1) ·ₗ x) ≡ _⋆_ (snd B') r (fst χ x) - path = fst χ ((r /1) ·ₗ x) ≡⟨ IsRingHom.pres· (snd χ) _ _ ⟩ + path = fst χ ((r /1) ·ₗ x) ≡⟨ IsCommRingHom.pres· (snd χ) _ _ ⟩ fst χ (r /1) ·b fst χ x ≡⟨ cong (_·b fst χ x) (χcomp r) ⟩ fst φ r ·b fst χ x ≡⟨ refl ⟩ _⋆_ (snd B') r 1b ·b fst χ x ≡⟨ ⋆AssocL (snd B') _ _ _ ⟩ @@ -115,11 +115,11 @@ module AlgLoc (R' : CommRing ℓ) ψ' : CommRingHom S⁻¹RAsCommRing B fst ψ' = fst ψ - IsRingHom.pres0 (snd ψ') = pres0 (snd ψ) - IsRingHom.pres1 (snd ψ') = pres1 (snd ψ) - IsRingHom.pres+ (snd ψ') = pres+ (snd ψ) - IsRingHom.pres· (snd ψ') = pres· (snd ψ) - IsRingHom.pres- (snd ψ') = pres- (snd ψ) + IsCommRingHom.pres0 (snd ψ') = pres0 (snd ψ) + IsCommRingHom.pres1 (snd ψ') = pres1 (snd ψ) + IsCommRingHom.pres+ (snd ψ') = pres+ (snd ψ) + IsCommRingHom.pres· (snd ψ') = pres· (snd ψ) + IsCommRingHom.pres- (snd ψ') = pres- (snd ψ) ψ'r/1≡φr : ∀ r → fst ψ (r /1) ≡ fst φ r ψ'r/1≡φr r = @@ -138,7 +138,7 @@ module AlgLoc (R' : CommRing ℓ) → CommAlgebraEquiv S⁻¹RAsCommAlg (toCommAlg (A' , φ)) S⁻¹RAlgCharEquiv A' φ cond = toCommAlgebraEquiv (S⁻¹RAsCommRing , /1AsCommRingHom) (A' , φ) (S⁻¹RCharEquiv R' S' SMultClosedSubset A' φ cond) - (RingHom≡ (S⁻¹RHasUniversalProp A' φ (cond .φS⊆Aˣ) .fst .snd)) + (CommRingHom≡ (S⁻¹RHasUniversalProp A' φ (cond .φS⊆Aˣ) .fst .snd)) where open PathToS⁻¹R diff --git a/Cubical/Algebra/CommAlgebra/LocalisationAlgebra.agda b/Cubical/Algebra/CommAlgebra/LocalisationAlgebra.agda index 703104ce86..1fbf864dfc 100644 --- a/Cubical/Algebra/CommAlgebra/LocalisationAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra/LocalisationAlgebra.agda @@ -151,7 +151,7 @@ module _ /1AsCommAlgebraHom = RUniv./1AsCommRingHom .fst , record - { IsRingHom (RUniv./1AsCommRingHom .snd) + { IsCommRingHom (RUniv./1AsCommRingHom .snd) ; pres⋆ = λ r x → refl} -- /1AsCommAlgebraHom and /1AsCommRingHom are equal over equality of the @@ -160,7 +160,7 @@ module _ (λ i → CommRingHom Aᵣ (S⁻¹AAsCommAlgebra→CommRing i)) (CommAlgebraHom→CommRingHom A S⁻¹AAsCommAlgebra /1AsCommAlgebraHom) RUniv./1AsCommRingHom - /1AsCommAlgebraHom→CommRingHom = ΣPathPProp (λ f → isPropIsRingHom _ f _) + /1AsCommAlgebraHom→CommRingHom = ΣPathPProp (λ f → isPropIsCommRingHom _ f _) (λ i → RUniv._/1) S⁻¹AHasUniversalProp : hasLocUniversalProp S⁻¹AAsCommAlgebra @@ -186,7 +186,7 @@ module _ original-univ : type-univ RLoc.S⁻¹RAsCommRing RUniv./1AsCommRingHom original-univ = RUniv.S⁻¹RHasUniversalProp (CommAlgebra→CommRing B) - (CommAlgebraHom→RingHom {A = A} {B = B} ψ) + (CommAlgebraHom→CommRingHom A B ψ) ψS⊂Bˣ univ : type-univ (CommAlgebra→CommRing S⁻¹AAsCommAlgebra) @@ -237,7 +237,7 @@ module _ χₐ : CommAlgebraHom S⁻¹AAsCommAlgebra B χₐ = univ .fst .fst .fst , record - { IsRingHom (univ .fst .fst .snd) + { IsCommRingHom (univ .fst .fst .snd) ; pres⋆ = pres⋆ } -- Commutativity is the same as the one for rings, since it only cares @@ -252,7 +252,7 @@ module _ χₐunique (φ' , φ'comm) = Σ≡Prop ((λ _ → isSetΠ (λ _ → is-set) _ _)) $ AlgebraHom≡ $ cong (fst ∘ fst) -- Get underlying bare function. - (univ .snd (CommAlgebraHom→RingHom {A = S⁻¹AAsCommAlgebra} {B = B} + (univ .snd (CommAlgebraHom→CommRingHom S⁻¹AAsCommAlgebra B φ' , φ'comm)) -- The above universal property leads to a generic induction principle for diff --git a/Cubical/Algebra/CommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebra/Properties.agda index 8d10d81366..ededbddde3 100644 --- a/Cubical/Algebra/CommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/Properties.agda @@ -107,6 +107,7 @@ module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where -- note that the proofs of the axioms might differ! isCommRing (snd (CommAlgebra→CommRing≡ i)) = isProp→PathP (λ i → isPropIsCommRing _ _ _ _ _ ) (isCommRing (snd (CommAlgebra→CommRing (toCommAlg Aφ)))) (isCommRing (snd A)) i + {- CommRingWithHomRoundTrip : (Aφ : CommRingWithHom) → fromCommAlg (toCommAlg Aφ) ≡ Aφ CommRingWithHomRoundTrip (A , φ) = ΣPathP (CommAlgebra→CommRing≡ (A , φ) , φPathP) @@ -117,6 +118,7 @@ module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where (snd (fromCommAlg (toCommAlg (A , φ)))) φ φPathP = RingHomPathP _ _ _ _ _ _ λ i x → ·IdR (snd A) (fst φ x) i -} + CommAlgRoundTrip : (A : CommAlgebra R ℓ') → toCommAlg (fromCommAlg A) ≡ A CommAlgRoundTrip A = ΣPathP (refl , AlgStrPathP) where @@ -133,6 +135,7 @@ module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where CommAlgebraStr.isCommAlgebra (AlgStrPathP i) = isProp→PathP (λ i → isPropIsCommAlgebra _ _ _ _ _ _ (CommAlgebraStr._⋆_ (AlgStrPathP i))) (CommAlgebraStr.isCommAlgebra (snd (toCommAlg (fromCommAlg A)))) isCommAlgebra i + {- CommAlgIso : Iso (CommAlgebra R ℓ') CommRingWithHom fun CommAlgIso = fromCommAlg @@ -140,6 +143,7 @@ module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where rightInv CommAlgIso = CommRingWithHomRoundTrip leftInv CommAlgIso = CommAlgRoundTrip -} + open IsCommRingHom isCommRingWithHomHom : (A B : CommRingWithHom) → CommRingHom (fst A) (fst B) → Type (ℓ-max ℓ ℓ') diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/UnivariatePolyList.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/UnivariatePolyList.agda index 22191cbf98..62f6db0e4d 100644 --- a/Cubical/Algebra/CommRing/Instances/Polynomials/UnivariatePolyList.agda +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/UnivariatePolyList.agda @@ -35,7 +35,7 @@ module _ (R : CommRing ℓ) where constantPolynomialHom : CommRingHom R UnivariatePolyList constantPolynomialHom = [_] , - makeIsRingHom + makeIsCommRingHom refl (λ r s → refl) λ r s → sym (MultHom[-] r s) From 4bacbaadd27cece0d54b9ec02af26e4c311c90fd Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Wed, 31 Jul 2024 16:49:22 +0200 Subject: [PATCH 33/83] unit comm algebra --- Cubical/Algebra/CommAlgebraAlt/Base.agda | 3 ++ Cubical/Algebra/CommAlgebraAlt/Ideal.agda | 6 +-- .../CommAlgebraAlt/Instances/Unit.agda | 52 +++++++++++++++++++ Cubical/Algebra/CommAlgebraAlt/Kernel.agda | 4 +- Cubical/Algebra/CommRing/Instances/Unit.agda | 15 +++++- Cubical/Algebra/CommRing/Properties.agda | 4 +- Cubical/Algebra/CommRing/Quotient/Base.agda | 18 ++++--- 7 files changed, 85 insertions(+), 17 deletions(-) create mode 100644 Cubical/Algebra/CommAlgebraAlt/Instances/Unit.agda diff --git a/Cubical/Algebra/CommAlgebraAlt/Base.agda b/Cubical/Algebra/CommAlgebraAlt/Base.agda index b50375ccef..e3417122ff 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Base.agda @@ -71,6 +71,9 @@ module _ {R : CommRing ℓ} where p) where open CommRingStr (B .fst .snd) using (is-set) + CommAlgebraEquiv : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') → Type _ + CommAlgebraEquiv A B = Σ[ f ∈ CommRingEquiv (A .fst) (B .fst) ] (f .fst .fst , f .snd) ∘cr A .snd ≡ B .snd + module CommAlgebraStr (A : CommAlgebra R ℓ') where open CommRingStr {{...}} instance diff --git a/Cubical/Algebra/CommAlgebraAlt/Ideal.agda b/Cubical/Algebra/CommAlgebraAlt/Ideal.agda index 775e11684c..47eeebaabc 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Ideal.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Ideal.agda @@ -14,15 +14,15 @@ open import Cubical.Data.Unit private variable - ℓ : Level + ℓ ℓ' : Level -module _ (R : CommRing ℓ) (A : CommAlgebra R ℓ) where +module _ (R : CommRing ℓ) (A : CommAlgebra R ℓ') where IdealsIn : Type _ IdealsIn = IdealsInCommRing (fst A) open CommRingStr (A .fst .snd) - makeIdeal : (I : A .fst .fst → hProp ℓ) + makeIdeal : (I : A .fst .fst → hProp ℓ') → (+-closed : {x y : A .fst .fst} → x ∈ I → y ∈ I → (x + y) ∈ I) → (0-closed : 0r ∈ I) → (·-closedLeft : {x : A .fst .fst} → (r : A .fst .fst) → x ∈ I → r · x ∈ I) diff --git a/Cubical/Algebra/CommAlgebraAlt/Instances/Unit.agda b/Cubical/Algebra/CommAlgebraAlt/Instances/Unit.agda new file mode 100644 index 0000000000..dc68c46351 --- /dev/null +++ b/Cubical/Algebra/CommAlgebraAlt/Instances/Unit.agda @@ -0,0 +1,52 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebraAlt.Instances.Unit where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Structure + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma.Properties using (Σ≡Prop) + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommAlgebraAlt.Base +open import Cubical.Algebra.CommRing.Instances.Unit +open import Cubical.Tactics.CommRingSolver + +private + variable + ℓ ℓ' : Level + +module _ (R : CommRing ℓ) where + UnitCommAlgebra : CommAlgebra R ℓ' + UnitCommAlgebra = UnitCommRing , mapToUnitCommRing R + + module _ (A : CommAlgebra R ℓ') where + terminalMap : CommAlgebraHom A (UnitCommAlgebra {ℓ' = ℓ}) + terminalMap .fst = mapToUnitCommRing (A .fst) + terminalMap .snd = isPropMapToUnitCommRing _ _ _ + +{- + module _ (A : CommAlgebra R ℓ) where + + terminalityContr : isContr (CommAlgebraHom A UnitCommAlgebra) + terminalityContr = terminalMap , path + where path : (ϕ : CommAlgebraHom A UnitCommAlgebra) → terminalMap ≡ ϕ + path ϕ = Σ≡Prop (isPropIsCommAlgebraHom {M = A} {N = UnitCommAlgebra}) + λ i _ → isPropUnit* _ _ i + + open CommAlgebraStr (snd A) + module _ (1≡0 : 1a ≡ 0a) where + + 1≡0→isContr : isContr ⟨ A ⟩ + 1≡0→isContr = 0a , λ a → + 0a ≡⟨ solve! S ⟩ + a · 0a ≡⟨ cong (λ b → a · b) (sym 1≡0) ⟩ + a · 1a ≡⟨ solve! S ⟩ + a ∎ + where S = CommAlgebra→CommRing A + + equivFrom1≡0 : CommAlgebraEquiv A UnitCommAlgebra + equivFrom1≡0 = isContr→Equiv 1≡0→isContr isContrUnit* , + snd terminalMap +-} diff --git a/Cubical/Algebra/CommAlgebraAlt/Kernel.agda b/Cubical/Algebra/CommAlgebraAlt/Kernel.agda index 20a8f48c56..484d0d7bc8 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Kernel.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Kernel.agda @@ -13,9 +13,9 @@ open import Cubical.Algebra.CommAlgebraAlt.Ideal private variable - ℓ : Level + ℓ ℓ' : Level -module _ {R : CommRing ℓ} (A B : CommAlgebra R ℓ) (ϕ : CommAlgebraHom A B) where +module _ {R : CommRing ℓ} (A B : CommAlgebra R ℓ') (ϕ : CommAlgebraHom A B) where kernel : IdealsIn R A kernel = Ideal→CommIdeal (ringKernel (CommAlgebraHom→RingHom {A = A} {B = B} ϕ)) diff --git a/Cubical/Algebra/CommRing/Instances/Unit.agda b/Cubical/Algebra/CommRing/Instances/Unit.agda index 3a9e6a3685..4d63b49c7c 100644 --- a/Cubical/Algebra/CommRing/Instances/Unit.agda +++ b/Cubical/Algebra/CommRing/Instances/Unit.agda @@ -10,11 +10,11 @@ open import Cubical.Algebra.CommRing private variable - ℓ : Level + ℓ ℓ' : Level open CommRingStr -UnitCommRing : ∀ {ℓ} → CommRing ℓ +UnitCommRing : CommRing ℓ fst UnitCommRing = Unit* 0r (snd UnitCommRing) = tt* 1r (snd UnitCommRing) = tt* @@ -25,3 +25,14 @@ isCommRing (snd UnitCommRing) = makeIsCommRing isSetUnit* (λ _ _ _ → refl) (λ { tt* → refl }) (λ _ → refl) (λ _ _ → refl) (λ _ _ _ → refl) (λ { tt* → refl }) (λ _ _ _ → refl) (λ _ _ → refl) + +mapToUnitCommRing : {ℓ : Level} (R : CommRing ℓ') → CommRingHom R (UnitCommRing {ℓ = ℓ}) +mapToUnitCommRing R .fst = λ _ → tt* +mapToUnitCommRing R .snd .IsCommRingHom.pres0 = refl +mapToUnitCommRing R .snd .IsCommRingHom.pres1 = refl +mapToUnitCommRing R .snd .IsCommRingHom.pres+ = λ _ _ → refl +mapToUnitCommRing R .snd .IsCommRingHom.pres· = λ _ _ → refl +mapToUnitCommRing R .snd .IsCommRingHom.pres- = λ _ → refl + +isPropMapToUnitCommRing : {ℓ : Level} (R : CommRing ℓ') → isProp (CommRingHom R (UnitCommRing {ℓ = ℓ})) +isPropMapToUnitCommRing R f g = CommRingHom≡ (funExt λ _ → refl) diff --git a/Cubical/Algebra/CommRing/Properties.agda b/Cubical/Algebra/CommRing/Properties.agda index 9244d3d9b0..3e6e4f41a5 100644 --- a/Cubical/Algebra/CommRing/Properties.agda +++ b/Cubical/Algebra/CommRing/Properties.agda @@ -32,7 +32,7 @@ open import Cubical.HITs.PropositionalTruncation private variable - ℓ ℓ' ℓ'' : Level + ℓ ℓ' ℓ'' ℓ''' : Level module Units (R' : CommRing ℓ) where open CommRingStr (snd R') @@ -174,7 +174,7 @@ module _ where → compCommRingHom _ _ _ f (idCommRingHom S) ≡ f idCompCommRingHom f = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) refl - compAssocCommRingHom : {R S T U : CommRing ℓ} + compAssocCommRingHom : {R : CommRing ℓ} {S : CommRing ℓ'} {T : CommRing ℓ''} {U : CommRing ℓ'''} (f : CommRingHom R S) (g : CommRingHom S T) (h : CommRingHom T U) → compCommRingHom _ _ _ (compCommRingHom _ _ _ f g) h ≡ compCommRingHom _ _ _ f (compCommRingHom _ _ _ g h) diff --git a/Cubical/Algebra/CommRing/Quotient/Base.agda b/Cubical/Algebra/CommRing/Quotient/Base.agda index a377b436f8..c1ac1a1541 100644 --- a/Cubical/Algebra/CommRing/Quotient/Base.agda +++ b/Cubical/Algebra/CommRing/Quotient/Base.agda @@ -140,7 +140,8 @@ module Quotient-FGideal-CommRing-CommRing where inducedHom : CommRingHom (R / (generatedIdeal _ v)) S - inducedHom = RingHom→CommRingHom $ Quotient-FGideal-CommRing-Ring.inducedHom R (CommRing→Ring S) (CommRingHom→RingHom f) v fnull + inducedHom = RingHom→CommRingHom $ + Quotient-FGideal-CommRing-Ring.inducedHom R (CommRing→Ring S) (CommRingHom→RingHom f) v fnull module UniversalProperty (R S : CommRing ℓ) @@ -160,13 +161,14 @@ module UniversalProperty I⊆ker ∘r Coherence.ringStr R I - isSolution : inducedHom ∘cr quotientHom R I ≡ f - isSolution = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) - (cong fst (Ring.UniversalProperty.solution - (CommRing→Ring R) - (CommIdeal→Ideal I) - (CommRingHom→RingHom f) - I⊆ker)) + opaque + isSolution : inducedHom ∘cr quotientHom R I ≡ f + isSolution = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) + (cong fst (Ring.UniversalProperty.solution + (CommRing→Ring R) + (CommIdeal→Ideal I) + (CommRingHom→RingHom f) + I⊆ker)) opaque unfolding quotientCommRingStr From 194e4560d786092cb73f683c7c85b52a4fb45480 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Wed, 31 Jul 2024 16:59:11 +0200 Subject: [PATCH 34/83] quotients --- .../CommAlgebraAlt/QuotientAlgebra.agda | 100 ++++++++---------- 1 file changed, 42 insertions(+), 58 deletions(-) diff --git a/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda index d241aedeee..ee89da6450 100644 --- a/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.Algebra.CommAlgebraAlt.QuotientAlgebra where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels @@ -18,19 +18,15 @@ open import Cubical.Algebra.CommRing.Ideal hiding (IdealsIn) open import Cubical.Algebra.CommAlgebraAlt.Base open import Cubical.Algebra.CommAlgebraAlt.Ideal open import Cubical.Algebra.CommAlgebraAlt.Kernel --- open import Cubical.Algebra.CommAlgebra.Instances.Unit --- open import Cubical.Algebra.Algebra.Base using (IsAlgebraHom; isPropIsAlgebraHom) +open import Cubical.Algebra.CommAlgebraAlt.Instances.Unit open import Cubical.Algebra.Ring --- open import Cubical.Algebra.Ring.Ideal using (isIdeal) open import Cubical.Tactics.CommRingSolver --- open import Cubical.Algebra.Algebra.Properties --- open AlgebraHoms using (_∘a_) private variable - ℓ : Level + ℓ ℓ' ℓ'' : Level -module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ') (I : IdealsIn R A) where open CommRingStr {{...}} open CommAlgebraStr {{...}} open RingTheory (CommRing→Ring (fst A)) using (-DistR·) @@ -41,12 +37,12 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where _ = A .fst .snd _ = A - _/_ : CommAlgebra R ℓ - _/_ = ((fst A) CommRing./ I) , - (CommRing.quotientHom (fst A) I ∘cr A .snd) + _/_ : CommAlgebra R ℓ' + _/_ = ((fst A) CommRing./ I) , + (CommRing.quotientHom (fst A) I ∘cr A .snd) quotientHom : CommAlgebraHom A (_/_) - quotientHom = (CommRing.quotientHom (fst A) I) , refl + quotientHom = withOpaqueStr $ (CommRing.quotientHom (fst A) I) , refl module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where open CommRingStr ⦃...⦄ @@ -110,59 +106,50 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where -module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ') (I : IdealsIn R A) where open CommRingStr {{...}} open CommAlgebraStr {{...}} opaque --- unfolding quotientHom + quotientHomEpi : (S : hSet ℓ'') + → (f g : ⟨ A / I ⟩ₐ → ⟨ S ⟩) + → f ∘ ⟨ quotientHom A I ⟩ₐ→ ≡ g ∘ ⟨ quotientHom A I ⟩ₐ→ + → f ≡ g + quotientHomEpi = CommRing.quotientHomEpi (fst A) I - injectivePrecomp : (B : CommAlgebra R ℓ) (f g : CommAlgebraHom (A / I) B) - → f ∘ca (quotientHom A I) ≡ g ∘ca (quotientHom A I) - → f ≡ g - injectivePrecomp B f g p = - CommAlgebraHom≡ ? --- (descendMapPath ⟨ f ⟩ₐ→ ⟨ g ⟩ₐ→ is-set --- λ x → λ i → fst (p i) x) --- where -{- {- trivial quotient -} module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) where - open CommAlgebraStr (snd A) - - opaque - unfolding _/_ + open CommRingStr (A .fst .snd) - oneIdealQuotient : CommAlgebraEquiv (A / (1Ideal A)) (UnitCommAlgebra R {ℓ' = ℓ}) - fst oneIdealQuotient = - isoToEquiv (iso (fst (terminalMap R (A / (1Ideal A)))) - (λ _ → [ 0a ]) + oneIdealQuotient : CommAlgebraEquiv (A / (1Ideal R A)) (UnitCommAlgebra R {ℓ' = ℓ}) + oneIdealQuotient .fst .fst = + withOpaqueStr $ + isoToEquiv (iso ⟨ terminalMap R (A / 1Ideal R A) ⟩ₐ→ + (λ _ → [ 0r ]) (λ _ → isPropUnit* _ _) - (elimProp (λ _ → squash/ _ _) - λ a → eq/ 0a a tt*)) - snd oneIdealQuotient = snd (terminalMap R (A / (1Ideal A))) + (elimProp (λ _ → squash/ _ _) (λ a → eq/ 0r a tt*))) + oneIdealQuotient .fst .snd = terminalMap R (A / 1Ideal R A) .fst .snd + oneIdealQuotient .snd = terminalMap R (A / (1Ideal R A)) .snd - opaque - unfolding quotientHom - zeroIdealQuotient : CommAlgebraEquiv A (A / (0Ideal A)) - fst zeroIdealQuotient = - let open RingTheory (CommRing→Ring (CommAlgebra→CommRing A)) - in isoToEquiv (iso (fst (quotientHom A (0Ideal A))) + zeroIdealQuotient : CommAlgebraEquiv A (A / (0Ideal R A)) + zeroIdealQuotient .fst .fst = + withOpaqueStr $ + let open RingTheory (CommRing→Ring (CommAlgebra→CommRing A)) + in isoToEquiv (iso ⟨ (quotientHom A (0Ideal R A)) ⟩ₐ→ (rec is-set (λ x → x) λ x y x-y≡0 → equalByDifference x y x-y≡0) (elimProp (λ _ → squash/ _ _) λ _ → refl) λ _ → refl) - snd zeroIdealQuotient = snd (quotientHom A (0Ideal A)) - -[_]/ : {R : CommRing ℓ} {A : CommAlgebra R ℓ} {I : IdealsIn A} - → (a : fst A) → fst (A / I) -[_]/ = fst (quotientHom _ _) - + zeroIdealQuotient .fst .snd = quotientHom A (0Ideal R A) .fst .snd + zeroIdealQuotient .snd = quotientHom A (0Ideal R A) .snd +[_]/ : {R : CommRing ℓ} {A : CommAlgebra R ℓ} {I : IdealsIn R A} + → (a : ⟨ A ⟩ₐ) → ⟨ A / I ⟩ₐ +[_]/ = ⟨ quotientHom _ _ ⟩ₐ→ -module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn A) where +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where open CommIdeal using (isPropIsCommIdeal) private @@ -170,8 +157,6 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn A) where π = quotientHom A I opaque - unfolding quotientHom - kernel≡I : kernel A (A / I) π ≡ I kernel≡I = kernel A (A / I) π ≡⟨ Σ≡Prop @@ -180,18 +165,17 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn A) where _ ≡⟨ CommRing.kernel≡I {R = CommAlgebra→CommRing A} I ⟩ I ∎ - module _ {R : CommRing ℓ} {A : CommAlgebra R ℓ} - {I : IdealsIn A} + {I : IdealsIn R A} where + open CommRingStr ⦃...⦄ + private instance + _ = A .fst .snd + _ = (A / I).fst .snd opaque - unfolding quotientHom - - isZeroFromIdeal : (x : ⟨ A ⟩) → x ∈ (fst I) → fst (quotientHom A I) x ≡ CommAlgebraStr.0a (snd (A / I)) - isZeroFromIdeal x x∈I = eq/ x 0a (subst (_∈ fst I) (solve! (CommAlgebra→CommRing A)) x∈I ) - where - open CommAlgebraStr (snd A) --} + unfolding CommRing.quotientCommRingStr + isZeroFromIdeal : (x : ⟨ A ⟩ₐ) → x ∈ (fst I) → ⟨ quotientHom A I ⟩ₐ→ x ≡ 0r + isZeroFromIdeal x x∈I = eq/ x 0r (subst (_∈ fst I) (solve! (CommAlgebra→CommRing A)) x∈I ) From 62d9af2f15531971f8c1b5ce5247706e5749fb2f Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 1 Aug 2024 10:17:18 +0200 Subject: [PATCH 35/83] no-eta for the equations of a ring --- Cubical/Algebra/Ring/Base.agda | 1 + 1 file changed, 1 insertion(+) diff --git a/Cubical/Algebra/Ring/Base.agda b/Cubical/Algebra/Ring/Base.agda index 3d4b2a0143..bd22398f72 100644 --- a/Cubical/Algebra/Ring/Base.agda +++ b/Cubical/Algebra/Ring/Base.agda @@ -34,6 +34,7 @@ private record IsRing {R : Type ℓ} (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) : Type ℓ where + no-eta-equality constructor isring field From b4475a4c8fee4d854586b5dd6bdf4fdffa4fc7ce Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 1 Aug 2024 11:18:48 +0200 Subject: [PATCH 36/83] WIP: Univalence --- Cubical/Algebra/CommAlgebraAlt/Base.agda | 10 ++++-- .../Algebra/CommAlgebraAlt/Univalence.agda | 35 +++++++++++++++++++ 2 files changed, 43 insertions(+), 2 deletions(-) create mode 100644 Cubical/Algebra/CommAlgebraAlt/Univalence.agda diff --git a/Cubical/Algebra/CommAlgebraAlt/Base.agda b/Cubical/Algebra/CommAlgebraAlt/Base.agda index e3417122ff..d51f74fe71 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Base.agda @@ -6,6 +6,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Function open import Cubical.Foundations.Structure using (⟨_⟩) open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence open import Cubical.Data.Sigma @@ -44,7 +45,6 @@ module _ {R : CommRing ℓ} where (fst g ∘cr snd B) .fst ≡⟨ cong fst (g .snd) ⟩ (C .snd .fst) ∎) - ⟨_⟩ₐ : (A : CommAlgebra R ℓ') → Type ℓ' ⟨ A ⟩ₐ = A .fst .fst @@ -71,6 +71,13 @@ module _ {R : CommRing ℓ} where p) where open CommRingStr (B .fst .snd) using (is-set) + CommAlgebra≡ : + {A B : CommAlgebra R ℓ'} + → (p : (A .fst) ≡ (B .fst)) + → (pathToEquiv $ cong fst p) .fst ∘ (A .snd) .fst ≡ (B .snd) .fst + → A ≡ B + CommAlgebra≡ p q = {!!} + CommAlgebraEquiv : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') → Type _ CommAlgebraEquiv A B = Σ[ f ∈ CommRingEquiv (A .fst) (B .fst) ] (f .fst .fst , f .snd) ∘cr A .snd ≡ B .snd @@ -105,7 +112,6 @@ module _ {R : CommRing ℓ} where ⋆AssocL : (r : ⟨ R ⟩) (x y : ⟨ A ⟩ₐ) → (r ⋆ x) · y ≡ r ⋆ (x · y) ⋆AssocL r x y = sym (·Assoc (A .snd .fst r) x y) - {- Convenience forgetful functions -} module _ {R : CommRing ℓ} where CommAlgebra→Ring : (A : CommAlgebra R ℓ') → Ring ℓ' diff --git a/Cubical/Algebra/CommAlgebraAlt/Univalence.agda b/Cubical/Algebra/CommAlgebraAlt/Univalence.agda new file mode 100644 index 0000000000..e47411001f --- /dev/null +++ b/Cubical/Algebra/CommAlgebraAlt/Univalence.agda @@ -0,0 +1,35 @@ +module Cubical.Algebra.CommAlgebraAlt.Univalence where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Structure using (⟨_⟩) +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Sigma + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Univalence +open import Cubical.Algebra.CommAlgebraAlt.Base + +private + variable + ℓ ℓ' ℓ'' : Level + +CommAlgebraPath : (R : CommRing ℓ) → (A B : CommAlgebra R ℓ') → (CommAlgebraEquiv A B) ≃ (A ≡ B) +CommAlgebraPath R A B = isoToEquiv (iso to {!!} {!!} {!!}) + where + to : CommAlgebraEquiv A B → A ≡ B + to e = Σ≡Prop {!!} (CommRingPath (A .fst) (B .fst) .fst (e .fst)) + +{- +CommAlgebraPath : (R : CommRing ℓ) → (A B : CommAlgebra R ℓ') → (CommAlgebraEquiv A B) ≃ (A ≡ B) +CommAlgebraPath R = ∫ (𝒮ᴰ-CommAlgebra R) .UARel.ua + +uaCommAlgebra : {R : CommRing ℓ} {A B : CommAlgebra R ℓ'} → CommAlgebraEquiv A B → A ≡ B +uaCommAlgebra {R = R} {A = A} {B = B} = equivFun (CommAlgebraPath R A B) + +isGroupoidCommAlgebra : {R : CommRing ℓ} → isGroupoid (CommAlgebra R ℓ') +isGroupoidCommAlgebra A B = isOfHLevelRespectEquiv 2 (CommAlgebraPath _ _ _) (isSetAlgebraEquiv _ _) +-} From 24582d20051ccf4b964e2d3dcc4cfb00f90fb8f5 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 1 Aug 2024 15:32:12 +0200 Subject: [PATCH 37/83] equality characterization with help from Evan --- Cubical/Algebra/CommAlgebraAlt/Base.agda | 21 ++++++++++++++++++++- 1 file changed, 20 insertions(+), 1 deletion(-) diff --git a/Cubical/Algebra/CommAlgebraAlt/Base.agda b/Cubical/Algebra/CommAlgebraAlt/Base.agda index d51f74fe71..60a6f49dbd 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Base.agda @@ -7,6 +7,7 @@ open import Cubical.Foundations.Function open import Cubical.Foundations.Structure using (⟨_⟩) open import Cubical.Foundations.HLevels open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport open import Cubical.Data.Sigma @@ -76,7 +77,20 @@ module _ {R : CommRing ℓ} where → (p : (A .fst) ≡ (B .fst)) → (pathToEquiv $ cong fst p) .fst ∘ (A .snd) .fst ≡ (B .snd) .fst → A ≡ B - CommAlgebra≡ p q = {!!} + CommAlgebra≡ {A = A} {B = B} p q = ΣPathTransport→PathΣ A B (p , pComm) + where + pComm : subst (CommRingHom R) p (A .snd) ≡ B .snd + pComm = + CommRingHom≡ + (fst (subst (CommRingHom R) p (A .snd)) + ≡⟨ sym (substCommSlice (CommRingHom R) (λ X → ⟨ R ⟩ → ⟨ X ⟩) (λ _ → fst) p (A .snd)) ⟩ + subst (λ X → ⟨ R ⟩ → ⟨ X ⟩) p (A .snd .fst) + ≡⟨ fromPathP (funTypeTransp (λ _ → ⟨ R ⟩) ⟨_⟩ p (A .snd .fst)) ⟩ + subst ⟨_⟩ p ∘ A .snd .fst ∘ subst (λ _ → ⟨ R ⟩) (sym p) + ≡⟨ cong ((subst ⟨_⟩ p ∘ A .snd .fst) ∘_) (funExt (λ _ → transportRefl _)) ⟩ + (pathToEquiv $ cong (λ r → fst r) p) .fst ∘ A .snd .fst + ≡⟨ q ⟩ + fst (B .snd) ∎) CommAlgebraEquiv : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') → Type _ CommAlgebraEquiv A B = Σ[ f ∈ CommRingEquiv (A .fst) (B .fst) ] (f .fst .fst , f .snd) ∘cr A .snd ≡ B .snd @@ -126,3 +140,8 @@ module _ {R : CommRing ℓ} where → CommAlgebraHom A B → RingHom (CommAlgebra→Ring A) (CommAlgebra→Ring B) CommAlgebraHom→RingHom = CommRingHom→RingHom ∘ CommAlgebraHom→CommRingHom + + CommAlgebraEquiv→CommRingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → CommAlgebraEquiv A B + → CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) + CommAlgebraEquiv→CommRingHom e = CommRingEquiv→CommRingHom (e .fst) From 58d2e8954956bca01063e787e41f592e72eb226a Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 1 Aug 2024 17:37:02 +0200 Subject: [PATCH 38/83] WIP --- Cubical/Algebra/CommAlgebraAlt/Base.agda | 20 ++++++++++-- .../Algebra/CommAlgebraAlt/Univalence.agda | 31 ++++++++++++++++--- Cubical/Algebra/CommRing/Properties.agda | 6 ++++ 3 files changed, 49 insertions(+), 8 deletions(-) diff --git a/Cubical/Algebra/CommAlgebraAlt/Base.agda b/Cubical/Algebra/CommAlgebraAlt/Base.agda index 60a6f49dbd..a3e3c600d0 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Base.agda @@ -4,6 +4,7 @@ module Cubical.Algebra.CommAlgebraAlt.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Structure using (⟨_⟩) open import Cubical.Foundations.HLevels open import Cubical.Foundations.Univalence @@ -28,8 +29,8 @@ module _ {R : CommRing ℓ} where CommAlgebraHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') → Type _ CommAlgebraHom A B = Σ[ f ∈ CommRingHom (fst A) (fst B) ] f ∘cr snd A ≡ snd B - idCAlgHom : (A : CommAlgebra R ℓ) → CommAlgebraHom A A - idCAlgHom A = (idCommRingHom (fst A)) , Σ≡Prop (λ _ → isPropIsCommRingHom (R .snd) _ (A .fst .snd)) refl + idCAlgHom : (A : CommAlgebra R ℓ') → CommAlgebraHom A A + idCAlgHom A = (idCommRingHom (fst A)) , CommRingHom≡ refl compCommAlgebraHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} → (g : CommAlgebraHom B C) → (f : CommAlgebraHom A B) @@ -95,6 +96,19 @@ module _ {R : CommRing ℓ} where CommAlgebraEquiv : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') → Type _ CommAlgebraEquiv A B = Σ[ f ∈ CommRingEquiv (A .fst) (B .fst) ] (f .fst .fst , f .snd) ∘cr A .snd ≡ B .snd + idCAlgEquiv : (A : CommAlgebra R ℓ') → CommAlgebraEquiv A A + idCAlgEquiv A = (CommRingEquivs.idCommRingEquiv (fst A)) , + CommRingHom≡ refl + + CommAlgebraEquiv≡ : + {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {f g : CommAlgebraEquiv A B} + → f .fst .fst .fst ≡ g .fst .fst .fst + → f ≡ g + CommAlgebraEquiv≡ {B = B} p = + Σ≡Prop (λ _ → isSetΣSndProp (isSetΠ (λ _ → isSetB)) (λ _ → isPropIsCommRingHom _ _ _) _ _) + (CommRingEquivs.CommRingEquiv≡ p) + where open CommRingStr (B .fst .snd) using () renaming (is-set to isSetB) + module CommAlgebraStr (A : CommAlgebra R ℓ') where open CommRingStr {{...}} instance @@ -126,7 +140,7 @@ module _ {R : CommRing ℓ} where ⋆AssocL : (r : ⟨ R ⟩) (x y : ⟨ A ⟩ₐ) → (r ⋆ x) · y ≡ r ⋆ (x · y) ⋆AssocL r x y = sym (·Assoc (A .snd .fst r) x y) -{- Convenience forgetful functions -} +{- Convenient forgetful functions -} module _ {R : CommRing ℓ} where CommAlgebra→Ring : (A : CommAlgebra R ℓ') → Ring ℓ' CommAlgebra→Ring A = CommRing→Ring (fst A) diff --git a/Cubical/Algebra/CommAlgebraAlt/Univalence.agda b/Cubical/Algebra/CommAlgebraAlt/Univalence.agda index e47411001f..0e595efabb 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Univalence.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Univalence.agda @@ -6,6 +6,7 @@ open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Structure using (⟨_⟩) open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence open import Cubical.Data.Sigma @@ -17,11 +18,31 @@ private variable ℓ ℓ' ℓ'' : Level -CommAlgebraPath : (R : CommRing ℓ) → (A B : CommAlgebra R ℓ') → (CommAlgebraEquiv A B) ≃ (A ≡ B) -CommAlgebraPath R A B = isoToEquiv (iso to {!!} {!!} {!!}) - where - to : CommAlgebraEquiv A B → A ≡ B - to e = Σ≡Prop {!!} (CommRingPath (A .fst) (B .fst) .fst (e .fst)) + +module _ (R : CommRing ℓ) where + private + to : (A B : CommAlgebra R ℓ') → CommAlgebraEquiv A B → A ≡ B + to A B e = CommAlgebra≡ + r≡ + ((pathToEquiv $ cong fst r≡) .fst ∘ A .snd .fst + ≡⟨ cong (_∘ A .snd .fst) (cong fst $ pathToEquiv-ua (e .fst .fst) ) ⟩ + CommAlgebraEquiv→CommRingHom e .fst ∘ A .snd .fst + ≡⟨ cong fst (e .snd) ⟩ + B .snd .fst ∎) + where r≡ : (A .fst) ≡ (B .fst) + r≡ = CommRingPath (A .fst) (B .fst) .fst (e .fst) + + test : (A : CommAlgebra R ℓ') → to A A (idCAlgEquiv A) ≡ refl + test A = {!to A A (idCAlgEquiv A)!} ≡⟨ {!!} ⟩ ? ≡⟨ {!!} ⟩ refl ∎ + + toIsEquiv : (A : CommAlgebra R ℓ') {B' : CommAlgebra R ℓ'} → (p : A ≡ B') → isContr (fiber (to A B') p) + toIsEquiv A = + J (λ B' A≡B' → isContr (fiber (to A B') A≡B')) + ((idCAlgEquiv A , {!test A!}) , + λ {(e , toe≡refl) → {!!}}) + + CommAlgebraPath : (A B : CommAlgebra R ℓ') → (CommAlgebraEquiv A B) ≃ (A ≡ B) + CommAlgebraPath {ℓ' = ℓ'} A B = to A B , {!!} {- CommAlgebraPath : (R : CommRing ℓ) → (A B : CommAlgebra R ℓ') → (CommAlgebraEquiv A B) ≃ (A ≡ B) diff --git a/Cubical/Algebra/CommRing/Properties.agda b/Cubical/Algebra/CommRing/Properties.agda index 3e6e4f41a5..f21d6be99a 100644 --- a/Cubical/Algebra/CommRing/Properties.agda +++ b/Cubical/Algebra/CommRing/Properties.agda @@ -203,6 +203,12 @@ module CommRingEquivs where fst (idCommRingEquiv A) = idEquiv (fst A) snd (idCommRingEquiv A) = makeIsCommRingHom refl (λ _ _ → refl) (λ _ _ → refl) + CommRingEquiv≡ : {A : CommRing ℓ} {B : CommRing ℓ'} {f g : CommRingEquiv A B} + → f .fst .fst ≡ g .fst .fst + → f ≡ g + CommRingEquiv≡ p = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) + (Σ≡Prop isPropIsEquiv p) + module Exponentiation (R' : CommRing ℓ) where open CommRingStr (snd R') private R = fst R' From 78a6beb6ce01e2418c9b94c6eb1931a8e17215df Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 2 Aug 2024 12:01:09 +0200 Subject: [PATCH 39/83] typevariate polynomials as comm ring --- .../Instances/Polynomials/Typevariate.agda | 86 +++++++++++++++++++ 1 file changed, 86 insertions(+) create mode 100644 Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate.agda diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate.agda new file mode 100644 index 0000000000..d5c2fee719 --- /dev/null +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate.agda @@ -0,0 +1,86 @@ +{-# OPTIONS --safe #-} + +module Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate where +{- + The ring of polynomials with coefficients in a given ring, + or in other words the free commutative algebra over a commutative ring. + Note that this is a constructive definition, which entails that polynomials + cannot be represented by lists of coefficients, where the last one is non-zero. + For rings with decidable equality, that is still possible. + + I (Felix Cherubini) learned about this (and other) definition(s) from David Jaz Myers. + You can watch him talk about these things here: + https://www.youtube.com/watch?v=VNp-f_9MnVk + + This file contains + * the definition of the free commutative algebra on a type I over a commutative ring R as a HIT + (let us call that R[I]) + * a prove that the construction is an commutative R-algebra + For more, see the Properties file. +-} +open import Cubical.Foundations.Prelude + +open import Cubical.Algebra.CommRing.Base + +private + variable + ℓ ℓ' : Level + +module Construction (R : CommRing ℓ) where + open CommRingStr (snd R) using (1r; 0r) renaming (_+_ to _+r_; _·_ to _·r_) + + data R[_] (I : Type ℓ') : Type (ℓ-max ℓ ℓ') where + var : I → R[ I ] + const : fst R → R[ I ] + _+_ : R[ I ] → R[ I ] → R[ I ] + -_ : R[ I ] → R[ I ] + _·_ : R[ I ] → R[ I ] → R[ I ] -- \cdot + + +-assoc : (x y z : R[ I ]) → x + (y + z) ≡ (x + y) + z + +-rid : (x : R[ I ]) → x + (const 0r) ≡ x + +-rinv : (x : R[ I ]) → x + (- x) ≡ (const 0r) + +-comm : (x y : R[ I ]) → x + y ≡ y + x + + ·-assoc : (x y z : R[ I ]) → x · (y · z) ≡ (x · y) · z + ·-lid : (x : R[ I ]) → (const 1r) · x ≡ x + ·-comm : (x y : R[ I ]) → x · y ≡ y · x + + ldist : (x y z : R[ I ]) → (x + y) · z ≡ (x · z) + (y · z) + + +HomConst : (s t : fst R) → const (s +r t) ≡ const s + const t + ·HomConst : (s t : fst R) → const (s ·r t) ≡ (const s) · (const t) + + 0-trunc : (x y : R[ I ]) (p q : x ≡ y) → p ≡ q + + + opaque + isCommRing : (I : Type ℓ') → IsCommRing (const {I = I} 0r) (const 1r) _+_ _·_ -_ + isCommRing I = + makeIsCommRing 0-trunc + +-assoc +-rid +-rinv +-comm + ·-assoc (λ x → ·-comm _ _ ∙ ·-lid x) + (λ x y z → (x · (y + z)) ≡⟨ ·-comm _ _ ⟩ + ((y + z) · x) ≡⟨ ldist _ _ _ ⟩ + ((y · x) + (z · x)) ≡[ i ]⟨ (·-comm y x i + ·-comm z x i) ⟩ + ((x · y) + (x · z)) ∎) + ·-comm + + module _ (I : Type ℓ') where + open CommRingStr hiding (isCommRing) + commRingStr : CommRingStr (R[_] I) + commRingStr .0r = _ + commRingStr .1r = _ + commRingStr .CommRingStr._+_ = _ + commRingStr .CommRingStr._·_ = _ + CommRingStr.- commRingStr = _ + commRingStr .CommRingStr.isCommRing = isCommRing I + + + opaque + open IsCommRingHom + constIsCommRingHom : (I : Type ℓ') → IsCommRingHom (R .snd) (const {I = I}) (commRingStr I) + constIsCommRingHom I = makeIsCommRingHom refl +HomConst ·HomConst + +_[_]ᵣ : (R : CommRing ℓ) (I : Type ℓ') → CommRing (ℓ-max ℓ ℓ') +fst (R [ I ]ᵣ) = Construction.R[_] R I +snd (R [ I ]ᵣ) = Construction.commRingStr R I From bb224c51c729462b90026369f6e290c86607aaf4 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 2 Aug 2024 13:41:01 +0200 Subject: [PATCH 40/83] polynomials as comm algebra --- .../CommAlgebraAlt/Instances/Polynomials.agda | 15 +++++++++++++++ .../Instances/Polynomials/Typevariate.agda | 6 +++++- 2 files changed, 20 insertions(+), 1 deletion(-) create mode 100644 Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda diff --git a/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda b/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda new file mode 100644 index 0000000000..614c807471 --- /dev/null +++ b/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda @@ -0,0 +1,15 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebraAlt.Instances.Polynomials where + +open import Cubical.Foundations.Prelude + +open import Cubical.Algebra.CommRing.Base +open import Cubical.Algebra.CommAlgebraAlt.Base +open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate + +private + variable + ℓ ℓ' : Level + +_[_] : (R : CommRing ℓ) (I : Type ℓ') → CommAlgebra R _ +R [ I ] = (R [ I ]ᵣ) , {!!} diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate.agda index d5c2fee719..0b83c987b3 100644 --- a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate.agda +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate.agda @@ -52,7 +52,6 @@ module Construction (R : CommRing ℓ) where 0-trunc : (x y : R[ I ]) (p q : x ≡ y) → p ≡ q - opaque isCommRing : (I : Type ℓ') → IsCommRing (const {I = I} 0r) (const 1r) _+_ _·_ -_ isCommRing I = @@ -84,3 +83,8 @@ module Construction (R : CommRing ℓ) where _[_]ᵣ : (R : CommRing ℓ) (I : Type ℓ') → CommRing (ℓ-max ℓ ℓ') fst (R [ I ]ᵣ) = Construction.R[_] R I snd (R [ I ]ᵣ) = Construction.commRingStr R I + +const : (R : CommRing ℓ) (I : Type ℓ') → CommRingHom R (R [ I ]ᵣ) +const R I .fst = let open Construction R + in R[_].const +const R I .snd = Construction.constIsCommRingHom R I From 6667098e9625543a4b5cf556189d0a6e678ca303 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 2 Aug 2024 14:02:01 +0200 Subject: [PATCH 41/83] remove opaque structure again the idea is, that it might be not worth the effort and the isRing part is still 'opaque' due to the no-eta-equality of isRing. --- Cubical/Algebra/CommRing/Quotient/Base.agda | 16 ++++++---------- Cubical/Algebra/Ring/Quotient.agda | 11 ++++------- 2 files changed, 10 insertions(+), 17 deletions(-) diff --git a/Cubical/Algebra/CommRing/Quotient/Base.agda b/Cubical/Algebra/CommRing/Quotient/Base.agda index c1ac1a1541..8dde0f6e8c 100644 --- a/Cubical/Algebra/CommRing/Quotient/Base.agda +++ b/Cubical/Algebra/CommRing/Quotient/Base.agda @@ -30,14 +30,12 @@ module _ (R : CommRing ℓ) (I : IdealsIn R) where open CommRingStr (snd R) R/I = ⟨ R ⟩ /ₛ (λ x y → x - y ∈ (fst I)) - opaque - unfolding Ring.quotientRingStr - quotientCommRingStr : CommRingStr R/I - quotientCommRingStr = snd - (Ring→CommRing - ((CommRing→Ring R) Ring./ (CommIdeal→Ideal I)) - (elimProp2 (λ _ _ → squash/ _ _) - λ x y i → [ CommRingStr.·Comm (snd R) x y i ])) + quotientCommRingStr : CommRingStr R/I + quotientCommRingStr = snd + (Ring→CommRing + ((CommRing→Ring R) Ring./ (CommIdeal→Ideal I)) + (elimProp2 (λ _ _ → squash/ _ _) + λ x y i → [ CommRingStr.·Comm (snd R) x y i ])) _/_ : (R : CommRing ℓ) (I : IdealsIn R) → CommRing ℓ fst (R / I) = R/I R I @@ -48,7 +46,6 @@ snd (R / I) = quotientCommRingStr R I module Coherence (R : CommRing ℓ) (I : IdealsIn R) where opaque - unfolding quotientCommRingStr Ring.quotientRingStr isRingHomCoh : IsRingHom (snd (CommRing→Ring (R / I))) (λ x → x) (snd ((CommRing→Ring R) Ring./ (CommIdeal→Ideal I))) @@ -171,7 +168,6 @@ module UniversalProperty I⊆ker)) opaque - unfolding quotientCommRingStr isUnique : (ψ : CommRingHom (R / I) S) → (ψIsSolution : ψ .fst ∘ quotientHom R I .fst ≡ f .fst) → ψ ≡ inducedHom isUnique ψ ψIsSolution = diff --git a/Cubical/Algebra/Ring/Quotient.agda b/Cubical/Algebra/Ring/Quotient.agda index 326cfb6036..188b3a96b1 100644 --- a/Cubical/Algebra/Ring/Quotient.agda +++ b/Cubical/Algebra/Ring/Quotient.agda @@ -185,11 +185,10 @@ module _ (R' : Ring ℓ) (I : ⟨ R' ⟩ → hProp ℓ) (I-isIdeal : isIdeal R' eq : (x y z : R) → [ x ] ·/I ([ y ] +/I [ z ]) ≡ ([ x ] ·/I [ y ]) +/I ([ x ] ·/I [ z ]) eq x y z i = [ ·DistR+ x y z i ] - opaque - quotientRingStr : RingStr R/I - quotientRingStr = snd (makeRing 0/I 1/I _+/I_ _·/I_ -/I isSetR/I - +/I-assoc +/I-rid +/I-rinv +/I-comm - ·/I-assoc ·/I-rid ·/I-lid /I-rdist /I-ldist) + quotientRingStr : RingStr R/I + quotientRingStr = snd (makeRing 0/I 1/I _+/I_ _·/I_ -/I isSetR/I + +/I-assoc +/I-rid +/I-rinv +/I-comm + ·/I-assoc ·/I-rid ·/I-lid /I-rdist /I-ldist) _/_ : (R : Ring ℓ) → (I : IdealsIn R) → Ring ℓ fst (R / I) = R/I R (fst I) (snd I) @@ -199,7 +198,6 @@ snd (R / I) = quotientRingStr R (fst I) (snd I) [ a ]/I = [ a ] opaque - unfolding quotientRingStr isRingHomQuotientHom : (R : Ring ℓ) (I : IdealsIn R) → IsRingHom (R .snd) [_] ((R / I) .snd) IsRingHom.pres0 (isRingHomQuotientHom R I) = refl @@ -266,7 +264,6 @@ module UniversalProperty (R : Ring ℓ) (I : IdealsIn R) where 0r ∎) opaque - unfolding quotientRingStr inducedMapIsRingHom : (Iₛ⊆kernel : (x : ⟨ R ⟩) → x ∈ Iₛ → φ $r x ≡ 0r) → IsRingHom ((R / I) .snd) (inducedHomMap Iₛ⊆kernel) (S .snd) pres0 (inducedMapIsRingHom Iₛ⊆kernel) = φ.pres0 From f13c4c1676911e2b0581cc85eaf52acaf6782c5d Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 2 Aug 2024 14:02:16 +0200 Subject: [PATCH 42/83] Revert "WIP: deactivate problematic code" This reverts commit d7192ee08aba2d4bcaa6248fa49e953f22826dad. --- .../Multivariate/EquivCarac/A[X]X-A.agda | 3 +- .../Multivariate/EquivCarac/An[X]X-A.agda | 62 +- .../EilenbergMacLane/Rings/KleinBottle.agda | 600 +++++++++--------- .../EilenbergMacLane/Rings/RP2.agda | 118 ++-- .../EilenbergMacLane/Rings/RP2wedgeS1.agda | 34 +- .../Cohomology/EilenbergMacLane/Rings/Sn.agda | 218 ++++--- Cubical/ZCohomology/CohomologyRings/CP2.agda | 3 +- .../CohomologyRings/KleinBottle.agda | 3 +- Cubical/ZCohomology/CohomologyRings/RP2.agda | 3 +- .../CohomologyRings/RP2wedgeS1.agda | 3 +- Cubical/ZCohomology/CohomologyRings/S0.agda | 3 +- Cubical/ZCohomology/CohomologyRings/S1.agda | 3 +- .../CohomologyRings/S2wedgeS4.agda | 3 +- Cubical/ZCohomology/CohomologyRings/Sn.agda | 52 +- Cubical/ZCohomology/CohomologyRings/Unit.agda | 3 +- 15 files changed, 539 insertions(+), 572 deletions(-) diff --git a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/A[X]X-A.agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/A[X]X-A.agda index 71f5b38c06..dc41f35a12 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/A[X]X-A.agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/A[X]X-A.agda @@ -190,7 +190,7 @@ module Properties-Equiv-QuotientXn-A A→A[x]/x : A → A[x]/x A→A[x]/x = [_] ∘ A→A[x] -{- + A→A[x]/x-pres+ : (a a' : A) → A→A[x]/x (a +A a') ≡ A→A[x]/x a +PAI A→A[x]/x a' A→A[x]/x-pres+ a a' = cong [_] (A→A[x]-pres+ a a') @@ -247,4 +247,3 @@ module _ Equiv-ℤ[X]/X-ℤ : RingEquiv (CommRing→Ring ℤ[X]/X) (CommRing→Ring ℤCR) Equiv-ℤ[X]/X-ℤ = Equiv-A[X]/X-A ℤCR --- -} diff --git a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[X]X-A.agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[X]X-A.agda index f77fa7c596..2d2f46f80f 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[X]X-A.agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[X]X-A.agda @@ -19,7 +19,6 @@ open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.Quotient -open import Cubical.Algebra.Monoid.Instances.NatVec open import Cubical.Algebra.CommRing.Instances.Int renaming (ℤCommRing to ℤCR) @@ -217,39 +216,37 @@ module Properties-Equiv-QuotientXn-A ... | no ¬p = ⊥.rec (¬p refl) -{- + ----------------------------------------------------------------------------- -- Retraction open RingStr open IsRing - opaque - unfolding quotientCommRingStr - e-retr : (x : A[x1,···,xn]/ Ar n) → A→PAI (PAI→A x) ≡ x - e-retr = SQ.elimProp (λ _ → isSetPAI _ _) - (DS-Ind-Prop.f (NatVecMonoid n .fst) (λ _ → A) _ _ (λ _ → isSetPAI _ _) - base0-eq - base-eq - λ {U V} ind-U ind-V → cong [_] (A→PA-pres+ _ _) ∙ cong₂ _+PAI_ ind-U ind-V) - where - base0-eq : A→PAI (PAI→A [ 0PA ]) ≡ [ 0PA ] - base0-eq = cong [_] (base-neutral (replicate 0)) - - base-eq : (v : Vec ℕ n) → (a : A ) → [ A→PA (PA→A (base v a)) ] ≡ [ base v a ] - base-eq v a with (discreteVecℕn v (replicate 0)) - ... | 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) ∣₁ - where - helper : {!!} ≡ _ - helper = cong (λ X → X +PA base v (-A a)) (base-neutral (replicate 0)) - ∙ +PAIdL (base v (-A a)) - ∙ sym ( - genδ-FinVec-ℕLinearCombi ((A[X1,···,Xn] Ar n)) n k infkn (base v' (-A a)) ( Ar n) - ∙ cong₂ base (cong (λ X → v' +n-vec δℕ-Vec n X) (toFromId' n k infkn)) (·AIdR _) - ∙ cong (λ X → base X (-A a)) (sym eqvv')) + e-retr : (x : A[x1,···,xn]/ Ar n) → A→PAI (PAI→A x) ≡ x + e-retr = SQ.elimProp (λ _ → isSetPAI _ _) + (DS-Ind-Prop.f _ _ _ _ (λ _ → isSetPAI _ _) + base0-eq + base-eq + λ {U V} ind-U ind-V → cong [_] (A→PA-pres+ _ _) ∙ cong₂ _+PAI_ ind-U ind-V) + where + base0-eq : A→PAI (PAI→A [ 0PA ]) ≡ [ 0PA ] + base0-eq = cong [_] (base-neutral (replicate 0)) + + base-eq : (v : Vec ℕ n) → (a : A ) → [ A→PA (PA→A (base v a)) ] ≡ [ base v a ] + base-eq v a with (discreteVecℕn v (replicate 0)) + ... | 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) ∣₁ + where + helper : _ + helper = cong (λ X → X +PA base v (-A a)) (base-neutral (replicate 0)) + ∙ +PAIdL (base v (-A a)) + ∙ sym ( + genδ-FinVec-ℕLinearCombi ((A[X1,···,Xn] Ar n)) n k infkn (base v' (-A a)) ( Ar n) + ∙ cong₂ base (cong (λ X → v' +n-vec δℕ-Vec n X) (toFromId' n k infkn)) (·AIdR _) + ∙ cong (λ X → base X (-A a)) (sym eqvv')) @@ -257,17 +254,17 @@ module Properties-Equiv-QuotientXn-A -- Equiv module _ - (R : CommRing ℓ) + (Ar@(A , Astr) : CommRing ℓ) (n : ℕ) where open Iso - open Properties-Equiv-QuotientXn-A R n + open Properties-Equiv-QuotientXn-A Ar n - Equiv-QuotientX-A : CommRingEquiv (A[X1,···,Xn]/ R n) R + Equiv-QuotientX-A : CommRingEquiv (A[X1,···,Xn]/ Ar n) Ar fst Equiv-QuotientX-A = isoToEquiv is where - is : Iso (A[x1,···,xn]/ R n) (R .fst) + is : Iso (A[x1,···,xn]/ Ar n) A fun is = PAI→A inv is = A→PAI rightInv is = e-sect @@ -276,4 +273,3 @@ module _ -- Warning this doesn't prove Z[X]/X ≅ ℤ because you get two definition, -- see notation Multivariate-Quotient-notationZ for more details --} diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/KleinBottle.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/KleinBottle.agda index d7bcb5b82e..06ff30e2b5 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/KleinBottle.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/KleinBottle.agda @@ -853,318 +853,312 @@ isHomℤ/2[X,Y]→H*Klein = makeIsRingHom refl (λ _ _ → refl) ∙∙ H²K²→ℤ/2→H²K² (_⌣_ {G'' = ℤ/2Ring} K²gen.α K²gen.α) -- Map H*(K²) → ℤ/2[X,Y]/I -opaque - unfolding quotientCommRingStr - H*Klein→ℤ/2[X,Y]/I : - fst (H*R ℤ/2Ring KleinBottle) - → fst (CommRing→Ring ℤ/2[X,Y]/) - H*Klein→ℤ/2[X,Y]/I = - DS-Rec-Set.f _ _ _ _ squash/ [ neutral ] - HⁿKlein→ℤ/2[X,Y]/I _ - (+Assoc (snd (CommRing→Ring ℤ/2[X,Y]/))) - (+IdR (snd (CommRing→Ring ℤ/2[X,Y]/))) - (+Comm (snd (CommRing→Ring ℤ/2[X,Y]/))) - (λ { zero → cong [_] (base-neutral _) - ; one → cong [_] (cong₂ _add_ (base-neutral _) (base-neutral _) - ∙ addRid neutral) - ; two → cong [_] (cong (base (2 ∷ 0 ∷ [])) - (IsGroupHom.pres1 (snd (H²[K²,ℤ/2]≅ℤ/2*))) - ∙ base-neutral _) - ; (suc (suc (suc r))) → refl}) - λ { zero a b → cong [_] (base-add _ _ _ ∙ cong (base (0 ∷ 0 ∷ [])) - (sym (IsGroupHom.pres· (snd (H⁰[K²,ℤ/2]≅ℤ/2)) a b))) - ; one a b → cong [_] (move4 _ _ _ _ _add_ addAssoc addComm - ∙ cong₂ _add_ (base-add _ _ _ ∙ cong (base (1 ∷ 0 ∷ [])) - (cong fst (sym - (IsGroupHom.pres· (snd (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2)) a b)))) - (base-add _ _ _ ∙ cong (base (0 ∷ 1 ∷ [])) - (cong snd - (sym (IsGroupHom.pres· (snd (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2)) a b))))) - ; two a b → cong [_] (base-add _ _ _ ∙ cong (base (2 ∷ 0 ∷ [])) - (sym (IsGroupHom.pres· (snd (H²[K²,ℤ/2]≅ℤ/2*)) a b))) - ; (suc (suc (suc n))) → λ a b → cong [_] (addRid neutral)} +H*Klein→ℤ/2[X,Y]/I : + fst (H*R ℤ/2Ring KleinBottle) + → fst (CommRing→Ring ℤ/2[X,Y]/) +H*Klein→ℤ/2[X,Y]/I = + DS-Rec-Set.f _ _ _ _ squash/ [ neutral ] + HⁿKlein→ℤ/2[X,Y]/I _ + (+Assoc (snd (CommRing→Ring ℤ/2[X,Y]/))) + (+IdR (snd (CommRing→Ring ℤ/2[X,Y]/))) + (+Comm (snd (CommRing→Ring ℤ/2[X,Y]/))) + (λ { zero → cong [_] (base-neutral _) + ; one → cong [_] (cong₂ _add_ (base-neutral _) (base-neutral _) + ∙ addRid neutral) + ; two → cong [_] (cong (base (2 ∷ 0 ∷ [])) + (IsGroupHom.pres1 (snd (H²[K²,ℤ/2]≅ℤ/2*))) + ∙ base-neutral _) + ; (suc (suc (suc r))) → refl}) + λ { zero a b → cong [_] (base-add _ _ _ ∙ cong (base (0 ∷ 0 ∷ [])) + (sym (IsGroupHom.pres· (snd (H⁰[K²,ℤ/2]≅ℤ/2)) a b))) + ; one a b → cong [_] (move4 _ _ _ _ _add_ addAssoc addComm + ∙ cong₂ _add_ (base-add _ _ _ ∙ cong (base (1 ∷ 0 ∷ [])) + (cong fst (sym + (IsGroupHom.pres· (snd (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2)) a b)))) + (base-add _ _ _ ∙ cong (base (0 ∷ 1 ∷ [])) + (cong snd + (sym (IsGroupHom.pres· (snd (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2)) a b))))) + ; two a b → cong [_] (base-add _ _ _ ∙ cong (base (2 ∷ 0 ∷ [])) + (sym (IsGroupHom.pres· (snd (H²[K²,ℤ/2]≅ℤ/2*)) a b))) + ; (suc (suc (suc n))) → λ a b → cong [_] (addRid neutral)} -- The equivalence -{- -opaque - unfolding quotientCommRingStr H*Klein→ℤ/2[X,Y]/I - ℤ/2[X,Y]/≅H*KleinBottle : - RingEquiv (CommRing→Ring ℤ/2[X,Y]/) - (H*R ℤ/2Ring KleinBottle) - fst ℤ/2[X,Y]/≅H*KleinBottle = isoToEquiv is +ℤ/2[X,Y]/≅H*KleinBottle : + RingEquiv (CommRing→Ring ℤ/2[X,Y]/) + (H*R ℤ/2Ring KleinBottle) +fst ℤ/2[X,Y]/≅H*KleinBottle = isoToEquiv is + where + is : Iso _ _ + fun is = ℤ/2[X,Y]/I→H*Klein .fst + inv is = H*Klein→ℤ/2[X,Y]/I + rightInv is = DS-Ind-Prop.f _ _ _ _ + (λ _ → trunc _ _) + refl + (λ { zero a → lem₀ a _ refl + ; one a → lem₁ a _ _ refl + ; two a → lem₂ a _ refl + ; (suc (suc (suc r))) a → + sym (base-neutral _) + ∙ cong (base (3 + r)) + (isContr→isProp (isContr-HⁿKleinBottle r ℤ/2) + (0ₕ (3 + r)) a)}) + λ {x} {y} ind1 ind2 + → IsRingHom.pres+ (ℤ/2[X,Y]/I→H*Klein .snd) + (H*Klein→ℤ/2[X,Y]/I x) (H*Klein→ℤ/2[X,Y]/I y) + ∙ cong₂ _add_ ind1 ind2 where - is : Iso _ _ - fun is = ℤ/2[X,Y]/I→H*Klein .fst - inv is = H*Klein→ℤ/2[X,Y]/I - rightInv is = DS-Ind-Prop.f _ _ _ _ - (λ _ → trunc _ _) - refl - (λ { zero a → lem₀ a _ refl - ; one a → lem₁ a _ _ refl - ; two a → lem₂ a _ refl - ; (suc (suc (suc r))) a → - sym (base-neutral _) - ∙ cong (base (3 + r)) - (isContr→isProp (isContr-HⁿKleinBottle r ℤ/2) - (0ₕ (3 + r)) a)}) - λ {x} {y} ind1 ind2 - → IsRingHom.pres+ (ℤ/2[X,Y]/I→H*Klein .snd) - (H*Klein→ℤ/2[X,Y]/I x) (H*Klein→ℤ/2[X,Y]/I y) - ∙ cong₂ _add_ ind1 ind2 - where - lem₀ : (a : _) (x : _) - → H⁰[K²,ℤ/2]≅ℤ/2 .fst .fst a ≡ x - → ℤ/2[X,Y]/I→H*Klein .fst (H*Klein→ℤ/2[X,Y]/I (base zero a)) - ≡ base zero a - lem₀ a = - ℤ/2-elim - (λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ H*Klein→ℤ/2[X,Y]/I) (help id) - ∙ sym (help id)) - λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst) - (cong [_] (cong (base (0 ∷ 0 ∷ [])) id)) - ∙ cong (base zero) - (sym (cong (invEq (H⁰[K²,ℤ/2]≅ℤ/2 .fst)) id) - ∙ retEq (fst H⁰[K²,ℤ/2]≅ℤ/2) a) - where - help : H⁰[K²,ℤ/2]≅ℤ/2 .fst .fst a ≡ 0 - → Path (fst (H*R ℤ/2Ring KleinBottle)) (base zero a) neutral - help id' = - sym (cong (base zero) - (sym (cong (invEq (H⁰[K²,ℤ/2]≅ℤ/2 .fst)) id' - ∙ IsGroupHom.pres1 (isGroupHomInv (H⁰[K²,ℤ/2]≅ℤ/2))) - ∙ retEq (fst H⁰[K²,ℤ/2]≅ℤ/2) a)) - ∙ base-neutral zero - - lem₁ : (a : _) → (x y : _) - → H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst .fst a ≡ (x , y) - → ℤ/2[X,Y]/I→H*Klein .fst (H*Klein→ℤ/2[X,Y]/I (base one a)) - ≡ base one a - lem₁ a = - ℤ/2-elim - (ℤ/2-elim - (λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) - (cong₂ _add_ (cong (base (1 ∷ 0 ∷ [])) - (cong fst id)) - (cong (base (0 ∷ 1 ∷ [])) - (cong snd id))) - ∙ addRid neutral - ∙ sym (help a id)) - λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) - (cong₂ _add_ - (cong (base (1 ∷ 0 ∷ [])) (cong fst id) - ∙ base-neutral _) - (cong (base (0 ∷ 1 ∷ [])) (cong snd id)) - ∙ addComm _ _ ∙ addRid _) - ∙∙ cong (base 1) (1ₕ-⌣ 1 K²gen.β) - ∙∙ cong (base 1) (sym (retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) K²gen.β) - ∙∙ cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) (sym id) - ∙∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a)) - (ℤ/2-elim - (λ id → (cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) - (cong₂ _add_ - (cong (base (1 ∷ 0 ∷ [])) (cong fst id)) - (cong (base (0 ∷ 1 ∷ [])) (cong snd id) ∙ base-neutral _) - ∙ addRid _) - ∙ cong (base 1) - ( (⌣-1ₕ 1 K²gen.α ∙ transportRefl K²gen.α) - ∙ (sym (retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) K²gen.α) - ∙∙ cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) (α↦1 ∙ sym id) - ∙∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a)))) - λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) - (cong₂ _add_ - (cong (base (1 ∷ 0 ∷ [])) (cong fst id)) - (cong (base (0 ∷ 1 ∷ [])) (cong snd id))) - ∙ IsRingHom.pres+ (snd ℤ/2[X,Y]/I→H*Klein) - [ base (1 ∷ 0 ∷ []) 1 ] [ base (0 ∷ 1 ∷ []) 1 ] - ∙ cong₂ _add_ - (cong (base one) (⌣-1ₕ 1 (incL 1) ∙ transportRefl K²gen.α)) - (cong (base one) (1ₕ-⌣ 1 (incR 1))) - ∙ base-add 1 K²gen.α K²gen.β - ∙ cong (base one) - (sym (retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) (K²gen.α +ₕ K²gen.β)) - ∙∙ (cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) (sym id)) - ∙∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a)) - where - help : (a : _) → H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst .fst a ≡ (0 , 0) - → Path (fst (H*R ℤ/2Ring KleinBottle)) (base one a) neutral - help a p = - (sym (cong (base one) - (sym (cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) p - ∙ IsGroupHom.pres1 (isGroupHomInv (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2))) - ∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a))) - ∙ base-neutral one - - lem₂ : (a : _) (x : _) → H²[K²,ℤ/2]≅ℤ/2* .fst .fst a ≡ x - → ℤ/2[X,Y]/I→H*Klein .fst (H*Klein→ℤ/2[X,Y]/I (base two a)) - ≡ base two a - lem₂ a = - ℤ/2-elim - (λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ H*Klein→ℤ/2[X,Y]/I) - (cong (base 2) (help1 id) ∙ base-neutral _) - ∙∙ sym (base-neutral _) - ∙∙ cong (base 2) (sym (help1 id))) - λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst) - (cong [_] (cong (base (2 ∷ 0 ∷ [])) - (cong (H²[K²,ℤ/2]≅ℤ/2* .fst .fst) (help2 id) - ∙ α²↦1') )) - ∙∙ cong (base 2) (⌣-1ₕ 2 (incL 2) ∙ transportRefl (incL 2)) - ∙∙ cong (base two) (sym (help2 id)) - where - help1 : H²[K²,ℤ/2]≅ℤ/2* .fst .fst a ≡ 0 → a ≡ 0ₕ 2 - help1 p = sym (retEq (H²[K²,ℤ/2]≅ℤ/2* .fst) a) - ∙ cong (invEq (H²[K²,ℤ/2]≅ℤ/2* .fst)) p - ∙ IsGroupHom.pres1 (isGroupHomInv H²[K²,ℤ/2]≅ℤ/2*) - - help2 : H²[K²,ℤ/2]≅ℤ/2* .fst .fst a ≡ 1 → a ≡ α⌣α - help2 p = sym (retEq (H²[K²,ℤ/2]≅ℤ/2* .fst) a) - ∙∙ cong (invEq (H²[K²,ℤ/2]≅ℤ/2* .fst)) (p ∙ sym α²↦1') - ∙∙ retEq (H²[K²,ℤ/2]≅ℤ/2* .fst) α⌣α - - leftInv is = - SQ.elimProp - (λ _ → squash/ _ _) - (DS-Ind-Prop.f _ _ _ _ - (λ _ → squash/ _ _) - refl - (λ r a → main a r) - λ {x} {y} ind1 ind2 - → cong₂ (_+r_ (snd (CommRing→Ring ℤ/2[X,Y]/))) - ind1 ind2) + lem₀ : (a : _) (x : _) + → H⁰[K²,ℤ/2]≅ℤ/2 .fst .fst a ≡ x + → ℤ/2[X,Y]/I→H*Klein .fst (H*Klein→ℤ/2[X,Y]/I (base zero a)) + ≡ base zero a + lem₀ a = + ℤ/2-elim + (λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ H*Klein→ℤ/2[X,Y]/I) (help id) + ∙ sym (help id)) + λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst) + (cong [_] (cong (base (0 ∷ 0 ∷ [])) id)) + ∙ cong (base zero) + (sym (cong (invEq (H⁰[K²,ℤ/2]≅ℤ/2 .fst)) id) + ∙ retEq (fst H⁰[K²,ℤ/2]≅ℤ/2) a) where - clem : (x y : ℕ) - → H*Klein→ℤ/2[X,Y]/I - (ℤ/2[X,Y]/I→H*Klein .fst - [ base (suc (suc (suc x)) ∷ y ∷ []) 1 ]) - ≡ [ neutral ] - clem x zero = refl - clem x (suc n) = refl - - help : (y : ℕ) → - Path (fst (CommRing→Ring ℤ/2[X,Y]/)) - [ base (one ∷ suc (suc y) ∷ []) 1 ] - [ neutral ] - help y = eq/ _ _ - ∣ (λ { zero → neutral - ; one → base (1 ∷ y ∷ []) 1 - ; two → neutral}) - , (sym (addRid _) - ∙ addComm (base (1 ∷ suc (suc y) ∷ []) 1 add neutral) neutral) - ∙ (λ i → neutral add (base (1 ∷ (+-comm 2 y i) ∷ []) 1 - add (addRid neutral (~ i)))) ∣₁ - - help2 : (y : ℕ) - → ℤ/2[X,Y]/I→H*Klein .fst [ base (zero ∷ suc (suc y) ∷ []) 1 ] - ≡ neutral - help2 zero = cong (base 2) (1ₕ-⌣ 2 (incR two)) - ∙ base-neutral _ - help2 (suc y) = base-neutral _ - - help3 : (x y : ℕ) - → ℤ/2[X,Y]/I→H*Klein .fst [ base (x ∷ suc (suc y) ∷ []) 1 ] - ≡ neutral - help3 zero y = help2 y - help3 (suc x) y = - (λ i → base (suc (suc (+-suc x y i))) - (transp (λ j → coHom (suc (suc (+-suc x y (i ∧ j)))) - ℤ/2 KleinBottle) (~ i) - (_⌣_ {G'' = ℤ/2Ring} (incL (suc x)) (incR (suc (suc y)))))) - ∙ cong (base (suc (suc (suc (x + y))))) - (isContr→isProp (isContr-HⁿKleinBottle (x + y) ℤ/2) _ _) - ∙ base-neutral _ - - main-1 : (r : _) - → H*Klein→ℤ/2[X,Y]/I (ℤ/2[X,Y]/I→H*Klein .fst [ base r 1 ]) - ≡ [ base r 1 ] - main-1 (zero ∷ zero ∷ []) = refl - main-1 (zero ∷ one ∷ []) = - cong (H*Klein→ℤ/2[X,Y]/I) - (cong (base 1) (1ₕ-⌣ 1 (incR 1))) - ∙ cong [_] (cong₂ _add_ (base-neutral _) - (λ _ → base (0 ∷ 1 ∷ []) 1) + help : H⁰[K²,ℤ/2]≅ℤ/2 .fst .fst a ≡ 0 + → Path (fst (H*R ℤ/2Ring KleinBottle)) (base zero a) neutral + help id' = + sym (cong (base zero) + (sym (cong (invEq (H⁰[K²,ℤ/2]≅ℤ/2 .fst)) id' + ∙ IsGroupHom.pres1 (isGroupHomInv (H⁰[K²,ℤ/2]≅ℤ/2))) + ∙ retEq (fst H⁰[K²,ℤ/2]≅ℤ/2) a)) + ∙ base-neutral zero + + lem₁ : (a : _) → (x y : _) + → H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst .fst a ≡ (x , y) + → ℤ/2[X,Y]/I→H*Klein .fst (H*Klein→ℤ/2[X,Y]/I (base one a)) + ≡ base one a + lem₁ a = + ℤ/2-elim + (ℤ/2-elim + (λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) + (cong₂ _add_ (cong (base (1 ∷ 0 ∷ [])) + (cong fst id)) + (cong (base (0 ∷ 1 ∷ [])) + (cong snd id))) + ∙ addRid neutral + ∙ sym (help a id)) + λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) + (cong₂ _add_ + (cong (base (1 ∷ 0 ∷ [])) (cong fst id) + ∙ base-neutral _) + (cong (base (0 ∷ 1 ∷ [])) (cong snd id)) ∙ addComm _ _ ∙ addRid _) - main-1 (zero ∷ suc (suc y) ∷ []) = - cong H*Klein→ℤ/2[X,Y]/I (help2 y) - ∙ eq/ _ _ - ∣ (λ {zero → neutral - ; one → base (0 ∷ (y ∷ [])) 1 - ; two → neutral}) - , cong (neutral add_) - (((λ i → base (0 ∷ (+-comm 2 y i) ∷ []) 1) - ∙ sym (addRid (base (0 ∷ (y + 2) ∷ []) 1))) - ∙ cong (base (0 ∷ (y + 2) ∷ []) 1 add_) (sym (addRid _))) ∣₁ - main-1 (one ∷ zero ∷ []) = - cong H*Klein→ℤ/2[X,Y]/I - (cong (base 1) (⌣-1ₕ 1 (incL one) ∙ transportRefl _)) - ∙ cong [_] (cong₂ _add_ (cong (base (1 ∷ 0 ∷ [])) - (cong fst α↦1)) - (base-neutral _) + ∙∙ cong (base 1) (1ₕ-⌣ 1 K²gen.β) + ∙∙ cong (base 1) (sym (retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) K²gen.β) + ∙∙ cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) (sym id) + ∙∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a)) + (ℤ/2-elim + (λ id → (cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) + (cong₂ _add_ + (cong (base (1 ∷ 0 ∷ [])) (cong fst id)) + (cong (base (0 ∷ 1 ∷ [])) (cong snd id) ∙ base-neutral _) ∙ addRid _) - main-1 (one ∷ one ∷ []) = - cong [_] (cong (base (2 ∷ 0 ∷ [])) αβ↦1') - ∙ eq/ _ _ ∣ (λ {zero → neutral - ; one → neutral - ; two → base (0 ∷ 0 ∷ []) 1}) - , ((addComm _ _ - ∙ sym (addRid _) - ∙ addComm (base (1 ∷ 1 ∷ []) 1 - add (base (2 ∷ 0 ∷ []) 1)) - neutral - ∙ sym (addRid _) - ∙ addComm (neutral add (base (1 ∷ 1 ∷ []) 1 - add (base (2 ∷ 0 ∷ []) 1))) neutral) - ∙ λ i → neutral add - (neutral add (addRid - (base (1 ∷ 1 ∷ []) 1 - add (base (2 ∷ 0 ∷ []) 1)) (~ i)))) ∣₁ - main-1 (one ∷ suc (suc y) ∷ []) = - cong H*Klein→ℤ/2[X,Y]/I (help3 one y) ∙ sym (help y) - main-1 (two ∷ zero ∷ []) = - cong [_] (cong (base (2 ∷ 0 ∷ [])) - (cong (H²[K²,ℤ/2]≅ℤ/2* .fst .fst) - (⌣-1ₕ 2 (incL 2) ∙ transportRefl _) - ∙ α²↦1')) - main-1 (two ∷ suc y ∷ []) = - eq/ neutral _ - ∣ (λ {zero → base (0 ∷ y ∷ []) 1 - ; one → neutral - ; two → base (1 ∷ y ∷ []) 1}) - , ((addComm _ _ ∙ addRid _ - ∙ ((((λ i → base (2 ∷ +-comm 1 y i ∷ []) 1) - ∙ sym (addRid _)) - ∙ cong (base (2 ∷ y + 1 ∷ []) (fsuc fzero) add_) - (sym (base-neutral _) - ∙ sym (base-add (3 ∷ y + 0 ∷ []) 1 1))) - ∙ addComm _ _) - ∙ sym (addAssoc _ _ _)) - ∙∙ cong (base (3 ∷ y + 0 ∷ []) (fsuc fzero) add_) - (addComm _ _ - ∙ sym (addComm _ _ ∙ addRid _)) - ∙∙ (λ i → base (3 ∷ (y + 0) ∷ []) 1 - add (neutral - add addRid (base (2 ∷ (y + 1) ∷ []) 1 - add base (3 ∷ (y + 0) ∷ []) 1 ) (~ i)))) ∣₁ - main-1 (suc (suc (suc x)) ∷ y ∷ []) = - clem x y - ∙ eq/ neutral _ - ∣ (λ {zero → base (x ∷ y ∷ []) 1 - ; one → neutral - ; two → neutral}) - , ((addComm neutral (base ((3 + x) ∷ y ∷ []) 1) - ∙ cong (base ((3 + x) ∷ y ∷ []) 1 add_) (sym (addRid neutral))) - ∙ λ i → base ((+-comm 3 x i) ∷ (+-comm 0 y i) ∷ []) 1 - add (neutral add (addRid neutral (~ i)))) ∣₁ - - main : (a : ℤ/2 .fst) (r : _) - → H*Klein→ℤ/2[X,Y]/I (ℤ/2[X,Y]/I→H*Klein .fst [ base r a ]) ≡ [ base r a ] - main = ℤ/2-elim (λ r → cong (H*Klein→ℤ/2[X,Y]/I ∘ ℤ/2[X,Y]/I→H*Klein .fst) - (cong [_] (base-neutral r)) - ∙ cong [_] (sym (base-neutral r))) - main-1 - snd ℤ/2[X,Y]/≅H*KleinBottle = ℤ/2[X,Y]/I→H*Klein .snd + ∙ cong (base 1) + ( (⌣-1ₕ 1 K²gen.α ∙ transportRefl K²gen.α) + ∙ (sym (retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) K²gen.α) + ∙∙ cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) (α↦1 ∙ sym id) + ∙∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a)))) + λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ [_]) + (cong₂ _add_ + (cong (base (1 ∷ 0 ∷ [])) (cong fst id)) + (cong (base (0 ∷ 1 ∷ [])) (cong snd id))) + ∙ IsRingHom.pres+ (snd ℤ/2[X,Y]/I→H*Klein) + [ base (1 ∷ 0 ∷ []) 1 ] [ base (0 ∷ 1 ∷ []) 1 ] + ∙ cong₂ _add_ + (cong (base one) (⌣-1ₕ 1 (incL 1) ∙ transportRefl K²gen.α)) + (cong (base one) (1ₕ-⌣ 1 (incR 1))) + ∙ base-add 1 K²gen.α K²gen.β + ∙ cong (base one) + (sym (retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) (K²gen.α +ₕ K²gen.β)) + ∙∙ (cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) (sym id)) + ∙∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a)) + where + help : (a : _) → H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst .fst a ≡ (0 , 0) + → Path (fst (H*R ℤ/2Ring KleinBottle)) (base one a) neutral + help a p = + (sym (cong (base one) + (sym (cong (invEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst)) p + ∙ IsGroupHom.pres1 (isGroupHomInv (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2))) + ∙ retEq (H¹[K²,ℤ/2]≅ℤ/2×ℤ/2 .fst) a))) + ∙ base-neutral one + + lem₂ : (a : _) (x : _) → H²[K²,ℤ/2]≅ℤ/2* .fst .fst a ≡ x + → ℤ/2[X,Y]/I→H*Klein .fst (H*Klein→ℤ/2[X,Y]/I (base two a)) + ≡ base two a + lem₂ a = + ℤ/2-elim + (λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst ∘ H*Klein→ℤ/2[X,Y]/I) + (cong (base 2) (help1 id) ∙ base-neutral _) + ∙∙ sym (base-neutral _) + ∙∙ cong (base 2) (sym (help1 id))) + λ id → cong (ℤ/2[X,Y]/I→H*Klein .fst) + (cong [_] (cong (base (2 ∷ 0 ∷ [])) + (cong (H²[K²,ℤ/2]≅ℤ/2* .fst .fst) (help2 id) + ∙ α²↦1') )) + ∙∙ cong (base 2) (⌣-1ₕ 2 (incL 2) ∙ transportRefl (incL 2)) + ∙∙ cong (base two) (sym (help2 id)) + where + help1 : H²[K²,ℤ/2]≅ℤ/2* .fst .fst a ≡ 0 → a ≡ 0ₕ 2 + help1 p = sym (retEq (H²[K²,ℤ/2]≅ℤ/2* .fst) a) + ∙ cong (invEq (H²[K²,ℤ/2]≅ℤ/2* .fst)) p + ∙ IsGroupHom.pres1 (isGroupHomInv H²[K²,ℤ/2]≅ℤ/2*) + + help2 : H²[K²,ℤ/2]≅ℤ/2* .fst .fst a ≡ 1 → a ≡ α⌣α + help2 p = sym (retEq (H²[K²,ℤ/2]≅ℤ/2* .fst) a) + ∙∙ cong (invEq (H²[K²,ℤ/2]≅ℤ/2* .fst)) (p ∙ sym α²↦1') + ∙∙ retEq (H²[K²,ℤ/2]≅ℤ/2* .fst) α⌣α + + leftInv is = + SQ.elimProp + (λ _ → squash/ _ _) + (DS-Ind-Prop.f _ _ _ _ + (λ _ → squash/ _ _) + refl + (λ r a → main a r) + λ {x} {y} ind1 ind2 + → cong₂ (_+r_ (snd (CommRing→Ring ℤ/2[X,Y]/))) + ind1 ind2) + where + clem : (x y : ℕ) + → H*Klein→ℤ/2[X,Y]/I + (ℤ/2[X,Y]/I→H*Klein .fst + [ base (suc (suc (suc x)) ∷ y ∷ []) 1 ]) + ≡ [ neutral ] + clem x zero = refl + clem x (suc n) = refl + + help : (y : ℕ) → + Path (fst (CommRing→Ring ℤ/2[X,Y]/)) + [ base (one ∷ suc (suc y) ∷ []) 1 ] + [ neutral ] + help y = eq/ _ _ + ∣ (λ { zero → neutral + ; one → base (1 ∷ y ∷ []) 1 + ; two → neutral}) + , (sym (addRid _) + ∙ addComm (base (1 ∷ suc (suc y) ∷ []) 1 add neutral) neutral) + ∙ (λ i → neutral add (base (1 ∷ (+-comm 2 y i) ∷ []) 1 + add (addRid neutral (~ i)))) ∣₁ + + help2 : (y : ℕ) + → ℤ/2[X,Y]/I→H*Klein .fst [ base (zero ∷ suc (suc y) ∷ []) 1 ] + ≡ neutral + help2 zero = cong (base 2) (1ₕ-⌣ 2 (incR two)) + ∙ base-neutral _ + help2 (suc y) = base-neutral _ + + help3 : (x y : ℕ) + → ℤ/2[X,Y]/I→H*Klein .fst [ base (x ∷ suc (suc y) ∷ []) 1 ] + ≡ neutral + help3 zero y = help2 y + help3 (suc x) y = + (λ i → base (suc (suc (+-suc x y i))) + (transp (λ j → coHom (suc (suc (+-suc x y (i ∧ j)))) + ℤ/2 KleinBottle) (~ i) + (_⌣_ {G'' = ℤ/2Ring} (incL (suc x)) (incR (suc (suc y)))))) + ∙ cong (base (suc (suc (suc (x + y))))) + (isContr→isProp (isContr-HⁿKleinBottle (x + y) ℤ/2) _ _) + ∙ base-neutral _ + + main-1 : (r : _) + → H*Klein→ℤ/2[X,Y]/I (ℤ/2[X,Y]/I→H*Klein .fst [ base r 1 ]) + ≡ [ base r 1 ] + main-1 (zero ∷ zero ∷ []) = refl + main-1 (zero ∷ one ∷ []) = + cong (H*Klein→ℤ/2[X,Y]/I) + (cong (base 1) (1ₕ-⌣ 1 (incR 1))) + ∙ cong [_] (cong₂ _add_ (base-neutral _) + (λ _ → base (0 ∷ 1 ∷ []) 1) + ∙ addComm _ _ ∙ addRid _) + main-1 (zero ∷ suc (suc y) ∷ []) = + cong H*Klein→ℤ/2[X,Y]/I (help2 y) + ∙ eq/ _ _ + ∣ (λ {zero → neutral + ; one → base (0 ∷ (y ∷ [])) 1 + ; two → neutral}) + , cong (neutral add_) + (((λ i → base (0 ∷ (+-comm 2 y i) ∷ []) 1) + ∙ sym (addRid (base (0 ∷ (y + 2) ∷ []) 1))) + ∙ cong (base (0 ∷ (y + 2) ∷ []) 1 add_) (sym (addRid _))) ∣₁ + main-1 (one ∷ zero ∷ []) = + cong H*Klein→ℤ/2[X,Y]/I + (cong (base 1) (⌣-1ₕ 1 (incL one) ∙ transportRefl _)) + ∙ cong [_] (cong₂ _add_ (cong (base (1 ∷ 0 ∷ [])) + (cong fst α↦1)) + (base-neutral _) + ∙ addRid _) + main-1 (one ∷ one ∷ []) = + cong [_] (cong (base (2 ∷ 0 ∷ [])) αβ↦1') + ∙ eq/ _ _ ∣ (λ {zero → neutral + ; one → neutral + ; two → base (0 ∷ 0 ∷ []) 1}) + , ((addComm _ _ + ∙ sym (addRid _) + ∙ addComm (base (1 ∷ 1 ∷ []) 1 + add (base (2 ∷ 0 ∷ []) 1)) + neutral + ∙ sym (addRid _) + ∙ addComm (neutral add (base (1 ∷ 1 ∷ []) 1 + add (base (2 ∷ 0 ∷ []) 1))) neutral) + ∙ λ i → neutral add + (neutral add (addRid + (base (1 ∷ 1 ∷ []) 1 + add (base (2 ∷ 0 ∷ []) 1)) (~ i)))) ∣₁ + main-1 (one ∷ suc (suc y) ∷ []) = + cong H*Klein→ℤ/2[X,Y]/I (help3 one y) ∙ sym (help y) + main-1 (two ∷ zero ∷ []) = + cong [_] (cong (base (2 ∷ 0 ∷ [])) + (cong (H²[K²,ℤ/2]≅ℤ/2* .fst .fst) + (⌣-1ₕ 2 (incL 2) ∙ transportRefl _) + ∙ α²↦1')) + main-1 (two ∷ suc y ∷ []) = + eq/ neutral _ + ∣ (λ {zero → base (0 ∷ y ∷ []) 1 + ; one → neutral + ; two → base (1 ∷ y ∷ []) 1}) + , ((addComm _ _ ∙ addRid _ + ∙ ((((λ i → base (2 ∷ +-comm 1 y i ∷ []) 1) + ∙ sym (addRid _)) + ∙ cong (base (2 ∷ y + 1 ∷ []) (fsuc fzero) add_) + (sym (base-neutral _) + ∙ sym (base-add (3 ∷ y + 0 ∷ []) 1 1))) + ∙ addComm _ _) + ∙ sym (addAssoc _ _ _)) + ∙∙ cong (base (3 ∷ y + 0 ∷ []) (fsuc fzero) add_) + (addComm _ _ + ∙ sym (addComm _ _ ∙ addRid _)) + ∙∙ (λ i → base (3 ∷ (y + 0) ∷ []) 1 + add (neutral + add addRid (base (2 ∷ (y + 1) ∷ []) 1 + add base (3 ∷ (y + 0) ∷ []) 1 ) (~ i)))) ∣₁ + main-1 (suc (suc (suc x)) ∷ y ∷ []) = + clem x y + ∙ eq/ neutral _ + ∣ (λ {zero → base (x ∷ y ∷ []) 1 + ; one → neutral + ; two → neutral}) + , ((addComm neutral (base ((3 + x) ∷ y ∷ []) 1) + ∙ cong (base ((3 + x) ∷ y ∷ []) 1 add_) (sym (addRid neutral))) + ∙ λ i → base ((+-comm 3 x i) ∷ (+-comm 0 y i) ∷ []) 1 + add (neutral add (addRid neutral (~ i)))) ∣₁ + + main : (a : ℤ/2 .fst) (r : _) + → H*Klein→ℤ/2[X,Y]/I (ℤ/2[X,Y]/I→H*Klein .fst [ base r a ]) ≡ [ base r a ] + main = ℤ/2-elim (λ r → cong (H*Klein→ℤ/2[X,Y]/I ∘ ℤ/2[X,Y]/I→H*Klein .fst) + (cong [_] (base-neutral r)) + ∙ cong [_] (sym (base-neutral r))) + main-1 +snd ℤ/2[X,Y]/≅H*KleinBottle = ℤ/2[X,Y]/I→H*Klein .snd H*KleinBottle≅ℤ/2[X,Y]/ : RingEquiv (H*R ℤ/2Ring KleinBottle) (CommRing→Ring ℤ/2[X,Y]/) H*KleinBottle≅ℤ/2[X,Y]/ = RingEquivs.invRingEquiv ℤ/2[X,Y]/≅H*KleinBottle --- -} diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/RP2.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/RP2.agda index 2ca66de1ff..db570dcc10 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/RP2.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/RP2.agda @@ -379,66 +379,64 @@ module _ {ℓ : Level} (n : ℕ) where H*[RP²,ℤ₂]→ℤ₂[X]/-fun : (r : ℕ) → coHom r ℤ/2 RP² → ℤ₂[X]/ .fst H*[RP²,ℤ₂]→ℤ₂[X]/-fun r x = [ base (r ∷ []) (H*[RP²,ℤ₂]→ℤ₂[X]/-fun' r x) ] - opaque - unfolding H*Klein→ℤ/2[X,Y]/I quotientCommRingStr - H*[RP²,ℤ₂]→ℤ₂[X]/ : H*R ℤ/2Ring RP² .fst → ℤ₂[X]/ .fst - H*[RP²,ℤ₂]→ℤ₂[X]/ = DS-Rec-Set.f _ _ _ _ squash/ - [ neutral ] - H*[RP²,ℤ₂]→ℤ₂[X]/-fun - (CommRingStr._+_ (snd ℤ₂[X]/)) - (CommRingStr.+Assoc (snd ℤ₂[X]/)) - (CommRingStr.+IdR (snd ℤ₂[X]/)) - (CommRingStr.+Comm (snd ℤ₂[X]/)) - (λ { zero → cong [_] (base-neutral (0 ∷ [])) - ; one → cong [_] (cong (base (1 ∷ [])) (IsGroupHom.pres1 (snd (H¹[RP²,ℤ/2]≅ℤ/2))) - ∙ base-neutral (1 ∷ [])) - ; two → cong [_] (cong (base (2 ∷ [])) (IsGroupHom.pres1 (snd (H²[RP²,ℤ/2]≅ℤ/2))) - ∙ base-neutral (2 ∷ [])) - ; (suc (suc (suc r))) → cong [_] (base-neutral _)}) - (λ { zero a b → cong [_] (base-add (0 ∷ []) _ _ - ∙ cong (base (0 ∷ [])) - (sym (IsGroupHom.pres· (snd (H⁰[RP²,ℤ/2]≅ℤ/2)) a b))) - ; one a b → cong [_] (base-add (1 ∷ []) _ _ - ∙ cong (base (1 ∷ [])) - (sym (IsGroupHom.pres· (snd (H¹[RP²,ℤ/2]≅ℤ/2)) a b))) - ; two a b → cong [_] (base-add (2 ∷ []) _ _ - ∙ cong (base (2 ∷ [])) - (sym (IsGroupHom.pres· (snd (H²[RP²,ℤ/2]≅ℤ/2)) a b))) - ; (suc (suc (suc r))) a b - → cong [_] (cong₂ _add_ refl (base-neutral (3 +ℕ r ∷ [])) - ∙ addRid _)}) - - ℤ₂[X]/→H*[RP²,ℤ₂]→ℤ₂[X]/ : (x : _) - → H*[RP²,ℤ₂]→ℤ₂[X]/ (ℤ₂[X]/→H*[RP²,ℤ₂] .fst x) ≡ x - ℤ₂[X]/→H*[RP²,ℤ₂]→ℤ₂[X]/ = SQ.elimProp (λ _ → squash/ _ _) - (DS-Ind-Prop.f _ _ _ _ (λ _ → squash/ _ _) - refl - (λ { (zero ∷ []) a → cong [_] (cong (base (zero ∷ [])) - (secEq (H⁰[RP²,ℤ/2]≅ℤ/2 .fst) a)) - ; (one ∷ []) a → cong [_] (cong (base (one ∷ [])) - (secEq (H¹[RP²,ℤ/2]≅ℤ/2 .fst) a)) - ; (two ∷ []) a → cong [_] (cong (base (two ∷ [])) - (secEq (H²[RP²,ℤ/2]≅ℤ/2 .fst) a)) - ; (suc (suc (suc r)) ∷ []) → ℤ/2-elim refl - (sym (eq/ _ _ ∣ (λ {zero → base (r ∷ []) fone}) - , ((cong₂ _add_ refl (base-neutral _) - ∙ addRid _ - ∙ λ i → base (+-comm 3 r i ∷ []) fone)) - ∙ sym (addRid _) ∣₁))}) - (λ p q → cong₂ (CommRingStr._+_ (snd ℤ₂[X]/)) p q)) - - H*[RP²,ℤ₂]→ℤ₂[X]/→H*[RP²,ℤ₂] : (x : H*R ℤ/2Ring RP² .fst) - → ℤ₂[X]/→H*[RP²,ℤ₂] .fst (H*[RP²,ℤ₂]→ℤ₂[X]/ x) ≡ x - H*[RP²,ℤ₂]→ℤ₂[X]/→H*[RP²,ℤ₂] = DS-Ind-Prop.f _ _ _ _ - (λ _ → trunc _ _) - refl - (λ { zero x → cong (base zero) (retEq (H⁰[RP²,ℤ/2]≅ℤ/2 .fst) x) - ; one x → cong (base one) (retEq (H¹[RP²,ℤ/2]≅ℤ/2 .fst) x) - ; two x → cong (base two) (retEq (H²[RP²,ℤ/2]≅ℤ/2 .fst) x) - ; (suc (suc (suc r))) x → cong (base (3 +ℕ r)) - (isContr→isProp (isContrH³⁺ⁿ[RP²,G] r ℤ/2) _ _)}) - λ {x} {y} p q - → (IsRingHom.pres+ (ℤ₂[X]/→H*[RP²,ℤ₂] .snd) _ _ ∙ cong₂ _add_ p q) + H*[RP²,ℤ₂]→ℤ₂[X]/ : H*R ℤ/2Ring RP² .fst → ℤ₂[X]/ .fst + H*[RP²,ℤ₂]→ℤ₂[X]/ = DS-Rec-Set.f _ _ _ _ squash/ + [ neutral ] + H*[RP²,ℤ₂]→ℤ₂[X]/-fun + (CommRingStr._+_ (snd ℤ₂[X]/)) + (CommRingStr.+Assoc (snd ℤ₂[X]/)) + (CommRingStr.+IdR (snd ℤ₂[X]/)) + (CommRingStr.+Comm (snd ℤ₂[X]/)) + (λ { zero → cong [_] (base-neutral (0 ∷ [])) + ; one → cong [_] (cong (base (1 ∷ [])) (IsGroupHom.pres1 (snd (H¹[RP²,ℤ/2]≅ℤ/2))) + ∙ base-neutral (1 ∷ [])) + ; two → cong [_] (cong (base (2 ∷ [])) (IsGroupHom.pres1 (snd (H²[RP²,ℤ/2]≅ℤ/2))) + ∙ base-neutral (2 ∷ [])) + ; (suc (suc (suc r))) → cong [_] (base-neutral _)}) + (λ { zero a b → cong [_] (base-add (0 ∷ []) _ _ + ∙ cong (base (0 ∷ [])) + (sym (IsGroupHom.pres· (snd (H⁰[RP²,ℤ/2]≅ℤ/2)) a b))) + ; one a b → cong [_] (base-add (1 ∷ []) _ _ + ∙ cong (base (1 ∷ [])) + (sym (IsGroupHom.pres· (snd (H¹[RP²,ℤ/2]≅ℤ/2)) a b))) + ; two a b → cong [_] (base-add (2 ∷ []) _ _ + ∙ cong (base (2 ∷ [])) + (sym (IsGroupHom.pres· (snd (H²[RP²,ℤ/2]≅ℤ/2)) a b))) + ; (suc (suc (suc r))) a b + → cong [_] (cong₂ _add_ refl (base-neutral (3 +ℕ r ∷ [])) + ∙ addRid _)}) + + ℤ₂[X]/→H*[RP²,ℤ₂]→ℤ₂[X]/ : (x : _) + → H*[RP²,ℤ₂]→ℤ₂[X]/ (ℤ₂[X]/→H*[RP²,ℤ₂] .fst x) ≡ x + ℤ₂[X]/→H*[RP²,ℤ₂]→ℤ₂[X]/ = SQ.elimProp (λ _ → squash/ _ _) + (DS-Ind-Prop.f _ _ _ _ (λ _ → squash/ _ _) + refl + (λ { (zero ∷ []) a → cong [_] (cong (base (zero ∷ [])) + (secEq (H⁰[RP²,ℤ/2]≅ℤ/2 .fst) a)) + ; (one ∷ []) a → cong [_] (cong (base (one ∷ [])) + (secEq (H¹[RP²,ℤ/2]≅ℤ/2 .fst) a)) + ; (two ∷ []) a → cong [_] (cong (base (two ∷ [])) + (secEq (H²[RP²,ℤ/2]≅ℤ/2 .fst) a)) + ; (suc (suc (suc r)) ∷ []) → ℤ/2-elim refl + (sym (eq/ _ _ ∣ (λ {zero → base (r ∷ []) fone}) + , ((cong₂ _add_ refl (base-neutral _) + ∙ addRid _ + ∙ λ i → base (+-comm 3 r i ∷ []) fone)) + ∙ sym (addRid _) ∣₁))}) + (λ p q → cong₂ (CommRingStr._+_ (snd ℤ₂[X]/)) p q)) + + H*[RP²,ℤ₂]→ℤ₂[X]/→H*[RP²,ℤ₂] : (x : H*R ℤ/2Ring RP² .fst) + → ℤ₂[X]/→H*[RP²,ℤ₂] .fst (H*[RP²,ℤ₂]→ℤ₂[X]/ x) ≡ x + H*[RP²,ℤ₂]→ℤ₂[X]/→H*[RP²,ℤ₂] = DS-Ind-Prop.f _ _ _ _ + (λ _ → trunc _ _) + refl + (λ { zero x → cong (base zero) (retEq (H⁰[RP²,ℤ/2]≅ℤ/2 .fst) x) + ; one x → cong (base one) (retEq (H¹[RP²,ℤ/2]≅ℤ/2 .fst) x) + ; two x → cong (base two) (retEq (H²[RP²,ℤ/2]≅ℤ/2 .fst) x) + ; (suc (suc (suc r))) x → cong (base (3 +ℕ r)) + (isContr→isProp (isContrH³⁺ⁿ[RP²,G] r ℤ/2) _ _)}) + λ {x} {y} p q + → (IsRingHom.pres+ (ℤ₂[X]/→H*[RP²,ℤ₂] .snd) _ _ ∙ cong₂ _add_ p q) ℤ₂[X]/≅H*[RP²,ℤ₂] : RingEquiv (CommRing→Ring ℤ₂[X]/) (H*R ℤ/2Ring RP²) fst ℤ₂[X]/≅H*[RP²,ℤ₂] = diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda index 0896082589..d328f4afcf 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/RP2wedgeS1.agda @@ -618,10 +618,8 @@ module Equiv-RP²∨S¹-Properties ℤ/2[x,y]/→H*-RP²∨S¹ : ℤ/2[x,y]/ → H*ℤ/2 RP²∨S¹ ℤ/2[x,y]/→H*-RP²∨S¹ = fst ℤ/2[X,Y]/→H*R-RP²∨S¹ - opaque - unfolding H*Klein→ℤ/2[X,Y]/I - ℤ/2[x,y]/→H*-RP²∨S¹-pres0 : ℤ/2[x,y]/→H*-RP²∨S¹ 0Pℤ/2I ≡ 0H* - ℤ/2[x,y]/→H*-RP²∨S¹-pres0 = refl + ℤ/2[x,y]/→H*-RP²∨S¹-pres0 : ℤ/2[x,y]/→H*-RP²∨S¹ 0Pℤ/2I ≡ 0H* + ℤ/2[x,y]/→H*-RP²∨S¹-pres0 = refl ℤ/2[x,y]/→H*-RP²∨S¹-pres+ : (x y : ℤ/2[x,y]/) → ℤ/2[x,y]/→H*-RP²∨S¹ ( x +Pℤ/2I y) @@ -672,15 +670,13 @@ module Equiv-RP²∨S¹-Properties H*-RP²∨S¹→ℤ/2[x,y]/ : H*ℤ/2 RP²∨S¹ → ℤ/2[x,y]/ H*-RP²∨S¹→ℤ/2[x,y]/ = [_] ∘ H*-RP²∨S¹→ℤ/2[x,y] - opaque - unfolding H*Klein→ℤ/2[X,Y]/I - H*-RP²∨S¹→ℤ/2[x,y]/-pres0 : H*-RP²∨S¹→ℤ/2[x,y]/ 0H* ≡ 0Pℤ/2I - H*-RP²∨S¹→ℤ/2[x,y]/-pres0 = refl + H*-RP²∨S¹→ℤ/2[x,y]/-pres0 : H*-RP²∨S¹→ℤ/2[x,y]/ 0H* ≡ 0Pℤ/2I + H*-RP²∨S¹→ℤ/2[x,y]/-pres0 = refl - H*-RP²∨S¹→ℤ/2[x,y]/-pres+ : (x y : H*ℤ/2 RP²∨S¹) → - H*-RP²∨S¹→ℤ/2[x,y]/ (x +H* y) - ≡ (H*-RP²∨S¹→ℤ/2[x,y]/ x) +Pℤ/2I (H*-RP²∨S¹→ℤ/2[x,y]/ y) - H*-RP²∨S¹→ℤ/2[x,y]/-pres+ x y = refl + H*-RP²∨S¹→ℤ/2[x,y]/-pres+ : (x y : H*ℤ/2 RP²∨S¹) → + H*-RP²∨S¹→ℤ/2[x,y]/ (x +H* y) + ≡ (H*-RP²∨S¹→ℤ/2[x,y]/ x) +Pℤ/2I (H*-RP²∨S¹→ℤ/2[x,y]/ y) + H*-RP²∨S¹→ℤ/2[x,y]/-pres+ x y = refl ----------------------------------------------------------------------------- @@ -744,14 +740,12 @@ module Equiv-RP²∨S¹-Properties ∙ cong₂ baseP (cong₂ (λ X Y → X ∷ Y ∷ []) (+-comm _ _) (+-comm _ _)) (·ℤ/2IdR _)) - opaque - unfolding H*Klein→ℤ/2[X,Y]/I - e-retr : (x : ℤ/2[x,y]/) → H*-RP²∨S¹→ℤ/2[x,y]/ (ℤ/2[x,y]/→H*-RP²∨S¹ x) ≡ x - e-retr = SQ.elimProp (λ _ → isSetPℤ/2I _ _) - (DS-Ind-Prop.f _ _ _ _ (λ _ → isSetPℤ/2I _ _) - refl - e-retr-base - λ {U V} ind-U ind-V → cong₂ _+Pℤ/2I_ ind-U ind-V) + e-retr : (x : ℤ/2[x,y]/) → H*-RP²∨S¹→ℤ/2[x,y]/ (ℤ/2[x,y]/→H*-RP²∨S¹ x) ≡ x + e-retr = SQ.elimProp (λ _ → isSetPℤ/2I _ _) + (DS-Ind-Prop.f _ _ _ _ (λ _ → isSetPℤ/2I _ _) + refl + e-retr-base + λ {U V} ind-U ind-V → cong₂ _+Pℤ/2I_ ind-U ind-V) -- Computation of the Cohomology Ring diff --git a/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda b/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda index 62f4c0bb4d..7f494a47c9 100644 --- a/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda +++ b/Cubical/Cohomology/EilenbergMacLane/Rings/Sn.agda @@ -245,118 +245,116 @@ module _ {ℓ : Level} (G : CommRing ℓ) (n : ℕ) where ∙ base-neutral (1 ∷ [])) help (gt x) = refl - opaque - unfolding quotientCommRingStr - H*[Sⁿ,G]→G[X]/-fun-coh₂ : (r : ℕ) (x y : coHom r GAb (S₊ (suc n))) + H*[Sⁿ,G]→G[X]/-fun-coh₂ : (r : ℕ) (x y : coHom r GAb (S₊ (suc n))) + → CommRingStr._+_ (snd (G[X]/)) + (H*[Sⁿ,G]→G[X]/-fun r x) (H*[Sⁿ,G]→G[X]/-fun r y) + ≡ H*[Sⁿ,G]→G[X]/-fun r (x +ₕ y) + H*[Sⁿ,G]→G[X]/-fun-coh₂ zero x y = + cong [_] (base-add (0 ∷ []) _ _ + ∙ cong (base (0 ∷ [])) (sym (IsGroupHom.pres· (H⁰[Sⁿ,G]≅G GAb n .snd) _ _))) + H*[Sⁿ,G]→G[X]/-fun-coh₂ (suc r) x y = help (n ≟ r) x y + where + help : (p : _) (x y : coHom (suc r) GAb (S₊ (suc n))) → CommRingStr._+_ (snd (G[X]/)) - (H*[Sⁿ,G]→G[X]/-fun r x) (H*[Sⁿ,G]→G[X]/-fun r y) - ≡ H*[Sⁿ,G]→G[X]/-fun r (x +ₕ y) - H*[Sⁿ,G]→G[X]/-fun-coh₂ zero x y = - cong [_] (base-add (0 ∷ []) _ _ - ∙ cong (base (0 ∷ [])) (sym (IsGroupHom.pres· (H⁰[Sⁿ,G]≅G GAb n .snd) _ _))) - H*[Sⁿ,G]→G[X]/-fun-coh₂ (suc r) x y = help (n ≟ r) x y - where - help : (p : _) (x y : coHom (suc r) GAb (S₊ (suc n))) - → CommRingStr._+_ (snd (G[X]/)) - (H*[Sⁿ,G]→G[X]/-fun^ r p x) - (H*[Sⁿ,G]→G[X]/-fun^ r p y) - ≡ H*[Sⁿ,G]→G[X]/-fun^ r p (x +ₕ y) - help (lt p) x y = cong [_] (addRid neutral) - help (eq p) = help' r p - where - help' : (r : ℕ) (p : n ≡ r) (x y : _) - → CommRingStr._+_ (snd G[X]/) - (H*[Sⁿ,G]→G[X]/-fun^ r (eq p) x) - (H*[Sⁿ,G]→G[X]/-fun^ r (eq p) y) - ≡ H*[Sⁿ,G]→G[X]/-fun^ r (eq p) (x +ₕ y) - help' = J> λ a b → cong [_] - (cong₂ (λ x y → (base (1 ∷ []) x) add (base (1 ∷ []) y)) - (cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) (transportRefl a)) - (cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) (transportRefl b)) - ∙ base-add _ _ _ - ∙ cong (base (1 ∷ [])) - (sym (IsGroupHom.pres· (Hⁿ[Sⁿ,G]≅G GAb n .snd) _ _) - ∙ cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) (sym (transportRefl (a +ₕ b))))) - help (gt p) x y = cong [_] (addRid neutral) - - H*[Sⁿ,G]→G[X]/ : fst (H*R GRing (S₊ (suc n))) → G[X]/ .fst - H*[Sⁿ,G]→G[X]/ = DS-Rec-Set.f _ _ _ _ squash/ - [ neutral ] - H*[Sⁿ,G]→G[X]/-fun - (CommRingStr._+_ (snd (G[X]/))) - (λ _ _ _ → CommRingStr.+Assoc (snd (G[X]/)) _ _ _) - (λ _ → CommRingStr.+IdR (snd (G[X]/)) _) - (λ _ _ → CommRingStr.+Comm (snd (G[X]/)) _ _) - H*[Sⁿ,G]→G[X]/-fun-coh₁ - H*[Sⁿ,G]→G[X]/-fun-coh₂ - - H*[Sⁿ,G]→G[X]/→H*[Sⁿ,G] : (x : fst (H*R GRing (S₊ (suc n)))) - → G[X]/→H*[Sⁿ,G] .fst (H*[Sⁿ,G]→G[X]/ x) ≡ x - H*[Sⁿ,G]→G[X]/→H*[Sⁿ,G] = - DS-Ind-Prop.f _ _ _ _ (λ _ → trunc _ _) - refl - (λ { zero a → cong (base zero) (retEq (H⁰[Sⁿ,G]≅G GAb n .fst) a) - ; (suc r) → help r (n ≟ r)}) - λ p q → IsRingHom.pres+ (G[X]/→H*[Sⁿ,G] .snd) _ _ - ∙ cong₂ _add_ p q + (H*[Sⁿ,G]→G[X]/-fun^ r p x) + (H*[Sⁿ,G]→G[X]/-fun^ r p y) + ≡ H*[Sⁿ,G]→G[X]/-fun^ r p (x +ₕ y) + help (lt p) x y = cong [_] (addRid neutral) + help (eq p) = help' r p where - help : (r : ℕ) (p : _) (a : _) - → G[X]/→H*[Sⁿ,G] .fst (H*[Sⁿ,G]→G[X]/-fun^ r p a) - ≡ base (suc r) a - help r (lt x) a = - sym (base-neutral (suc r)) - ∙ cong (base (suc r)) - (isContr→isProp - (subst (λ k → isContr (coHom k GAb (S₊ (suc n)))) - (cong suc (snd x)) - (isContr-Hᵐ⁺ⁿ[Sⁿ,G] GAb n (fst x))) _ _) - help r (eq x) a = - J (λ r x → (a : coHom (suc r) GAb (S₊ (suc n))) - → G[X]/→H*[Sⁿ,G] .fst (H*[Sⁿ,G]→G[X]/-fun^ r (eq x) a) - ≡ base (suc r) a) - (λ a → cong (base (suc n)) - (retEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) _ ∙ transportRefl a)) x a - help r (gt x) a = - sym (base-neutral (suc r)) - ∙ cong (base (suc r)) - (isContr→isProp - (subst (λ k → isContr (coHom (suc r) GAb (S₊ k))) - (cong suc (snd x)) - (isOfHLevelRetractFromIso 0 - (equivToIso (fst (Hⁿ[Sᵐ⁺ⁿ,G]≅0 GAb r (fst x)))) - isContrUnit*)) _ _) - - G[X]/→H*[Sⁿ,G]→H*[Sⁿ,G] : (x : G[X]/ .fst) - → H*[Sⁿ,G]→G[X]/ (G[X]/→H*[Sⁿ,G] .fst x) ≡ x - G[X]/→H*[Sⁿ,G]→H*[Sⁿ,G] = SQ.elimProp (λ _ → squash/ _ _) - (DS-Ind-Prop.f _ _ _ _ (λ _ → squash/ _ _) - refl - (λ { (zero ∷ []) g → cong [_] (cong (base (0 ∷ [])) - (secEq (H⁰[Sⁿ,G]≅G GAb n .fst) g)) - ; (one ∷ []) g → h2 g _ - ; (suc (suc x) ∷ []) g - → sym (eq/ _ neutral ∣ hh x g - , (λ i → base (+-comm 2 x i ∷ []) (CommRingStr.·IdR (snd G) g (~ i)) - add neutral) ∣₁)}) - λ p q → cong₂ (CommRingStr._+_ (snd (G[X]/))) p q) - where - hh : (x : ℕ) (g : fst G) → FinVec (fst (PolyCommRing G 1)) 1 - hh x g zero = base (x ∷ []) g - - h2 : (g : fst G) (p : _) - → H*[Sⁿ,G]→G[X]/-fun^ n p - (invEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) g) - ≡ [ base (one ∷ []) g ] - h2 g (lt x) = ⊥.rec (¬m) + (H*[Sⁿ,G]→G[X]/-fun^ r (eq p) x) + (H*[Sⁿ,G]→G[X]/-fun^ r (eq p) y) + ≡ H*[Sⁿ,G]→G[X]/-fun^ r (eq p) (x +ₕ y) + help' = J> λ a b → cong [_] + (cong₂ (λ x y → (base (1 ∷ []) x) add (base (1 ∷ []) y)) + (cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) (transportRefl a)) + (cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) (transportRefl b)) + ∙ base-add _ _ _ + ∙ cong (base (1 ∷ [])) + (sym (IsGroupHom.pres· (Hⁿ[Sⁿ,G]≅G GAb n .snd) _ _) + ∙ cong (Hⁿ[Sⁿ,G]≅G GAb n .fst .fst) (sym (transportRefl (a +ₕ b))))) + help (gt p) x y = cong [_] (addRid neutral) + + H*[Sⁿ,G]→G[X]/ : fst (H*R GRing (S₊ (suc n))) → G[X]/ .fst + H*[Sⁿ,G]→G[X]/ = DS-Rec-Set.f _ _ _ _ squash/ + [ neutral ] + H*[Sⁿ,G]→G[X]/-fun + (CommRingStr._+_ (snd (G[X]/))) + (λ _ _ _ → CommRingStr.+Assoc (snd (G[X]/)) _ _ _) + (λ _ → CommRingStr.+IdR (snd (G[X]/)) _) + (λ _ _ → CommRingStr.+Comm (snd (G[X]/)) _ _) + H*[Sⁿ,G]→G[X]/-fun-coh₁ + H*[Sⁿ,G]→G[X]/-fun-coh₂ + + H*[Sⁿ,G]→G[X]/→H*[Sⁿ,G] : (x : fst (H*R GRing (S₊ (suc n)))) + → G[X]/→H*[Sⁿ,G] .fst (H*[Sⁿ,G]→G[X]/ x) ≡ x + H*[Sⁿ,G]→G[X]/→H*[Sⁿ,G] = + DS-Ind-Prop.f _ _ _ _ (λ _ → trunc _ _) + refl + (λ { zero a → cong (base zero) (retEq (H⁰[Sⁿ,G]≅G GAb n .fst) a) + ; (suc r) → help r (n ≟ r)}) + λ p q → IsRingHom.pres+ (G[X]/→H*[Sⁿ,G] .snd) _ _ + ∙ cong₂ _add_ p q + where + help : (r : ℕ) (p : _) (a : _) + → G[X]/→H*[Sⁿ,G] .fst (H*[Sⁿ,G]→G[X]/-fun^ r p a) + ≡ base (suc r) a + help r (lt x) a = + sym (base-neutral (suc r)) + ∙ cong (base (suc r)) + (isContr→isProp + (subst (λ k → isContr (coHom k GAb (S₊ (suc n)))) + (cong suc (snd x)) + (isContr-Hᵐ⁺ⁿ[Sⁿ,G] GAb n (fst x))) _ _) + help r (eq x) a = + J (λ r x → (a : coHom (suc r) GAb (S₊ (suc n))) + → G[X]/→H*[Sⁿ,G] .fst (H*[Sⁿ,G]→G[X]/-fun^ r (eq x) a) + ≡ base (suc r) a) + (λ a → cong (base (suc n)) + (retEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) _ ∙ transportRefl a)) x a + help r (gt x) a = + sym (base-neutral (suc r)) + ∙ cong (base (suc r)) + (isContr→isProp + (subst (λ k → isContr (coHom (suc r) GAb (S₊ k))) + (cong suc (snd x)) + (isOfHLevelRetractFromIso 0 + (equivToIso (fst (Hⁿ[Sᵐ⁺ⁿ,G]≅0 GAb r (fst x)))) + isContrUnit*)) _ _) + + G[X]/→H*[Sⁿ,G]→H*[Sⁿ,G] : (x : G[X]/ .fst) + → H*[Sⁿ,G]→G[X]/ (G[X]/→H*[Sⁿ,G] .fst x) ≡ x + G[X]/→H*[Sⁿ,G]→H*[Sⁿ,G] = SQ.elimProp (λ _ → squash/ _ _) + (DS-Ind-Prop.f _ _ _ _ (λ _ → squash/ _ _) + refl + (λ { (zero ∷ []) g → cong [_] (cong (base (0 ∷ [])) + (secEq (H⁰[Sⁿ,G]≅G GAb n .fst) g)) + ; (one ∷ []) g → h2 g _ + ; (suc (suc x) ∷ []) g + → sym (eq/ _ neutral ∣ hh x g + , (λ i → base (+-comm 2 x i ∷ []) (CommRingStr.·IdR (snd G) g (~ i)) + add neutral) ∣₁)}) + λ p q → cong₂ (CommRingStr._+_ (snd (G[X]/))) p q) + where + hh : (x : ℕ) (g : fst G) → FinVec (fst (PolyCommRing G 1)) 1 + hh x g zero = base (x ∷ []) g + + h2 : (g : fst G) (p : _) + → H*[Sⁿ,G]→G[X]/-fun^ n p + (invEq (Hⁿ[Sⁿ,G]≅G GAb n .fst) g) + ≡ [ base (one ∷ []) g ] + h2 g (lt x) = ⊥.rec (¬m≅H*[Sⁿ,G] : RingEquiv (CommRing→Ring G[X]/) (H*R GRing (S₊ (suc n))) diff --git a/Cubical/ZCohomology/CohomologyRings/CP2.agda b/Cubical/ZCohomology/CohomologyRings/CP2.agda index 6cfd837ffb..1b5ebe1b1f 100644 --- a/Cubical/ZCohomology/CohomologyRings/CP2.agda +++ b/Cubical/ZCohomology/CohomologyRings/CP2.agda @@ -362,7 +362,7 @@ module ComputeCP²Notation H*-CP²→ℤ[x]/x³ : H* CP² → ℤ[x]/x³ H*-CP²→ℤ[x]/x³ = [_] ∘ H*-CP²→ℤ[x] -{- + H*-CP²→ℤ[x]/x³-gmorph : (x y : H* CP²) → H*-CP²→ℤ[x]/x³ (x +H* y) ≡ (H*-CP²→ℤ[x]/x³ x) +PℤI (H*-CP²→ℤ[x]/x³ y) H*-CP²→ℤ[x]/x³-gmorph x y = refl @@ -454,4 +454,3 @@ snd CP²-CohomologyRing = snd ℤ[X]/X³→H*R-CP² CohomologyRing-CP² : RingEquiv (H*R CP²) (CommRing→Ring ℤ[X]/X³) CohomologyRing-CP² = RingEquivs.invRingEquiv CP²-CohomologyRing --- -} diff --git a/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda b/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda index b21c3cf935..0454d41bbd 100644 --- a/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda +++ b/Cubical/ZCohomology/CohomologyRings/KleinBottle.agda @@ -370,7 +370,7 @@ module Equiv-𝕂²-Properties ℤ[x,y]/<2y,y²,xy,x²>→H*-𝕂² : ℤ[x,y]/<2y,y²,xy,x²> → H* KleinBottle ℤ[x,y]/<2y,y²,xy,x²>→H*-𝕂² = fst ℤ[X,Y]/<2Y,Y²,XY,X²>→H*R-𝕂² -{- + ℤ[x,y]/<2y,y²,xy,x²>→H*-𝕂²-pres0 : ℤ[x,y]/<2y,y²,xy,x²>→H*-𝕂² 0PℤI ≡ 0H* ℤ[x,y]/<2y,y²,xy,x²>→H*-𝕂²-pres0 = refl @@ -570,4 +570,3 @@ module _ where CohomologyRing-𝕂² : RingEquiv (H*R KleinBottle) (CommRing→Ring ℤ[X,Y]/<2Y,Y²,XY,X²>) CohomologyRing-𝕂² = RingEquivs.invRingEquiv 𝕂²-CohomologyRing --- -} diff --git a/Cubical/ZCohomology/CohomologyRings/RP2.agda b/Cubical/ZCohomology/CohomologyRings/RP2.agda index 668e5db41d..b17c6f6ab4 100644 --- a/Cubical/ZCohomology/CohomologyRings/RP2.agda +++ b/Cubical/ZCohomology/CohomologyRings/RP2.agda @@ -320,7 +320,7 @@ module Equiv-RP2-Properties where ℤ[x]/<2x,x²>→H*-RP² : ℤ[x]/<2x,x²> → H* RP² ℤ[x]/<2x,x²>→H*-RP² = fst ℤ[X]/<2X,X²>→H*R-RP² -{- + ℤ[x]/<2x,x²>→H*-RP²-pres0 : ℤ[x]/<2x,x²>→H*-RP² 0PℤI ≡ 0H* ℤ[x]/<2x,x²>→H*-RP²-pres0 = refl @@ -512,4 +512,3 @@ module _ where CohomologyRing-RP² : RingEquiv (H*R RP²) (CommRing→Ring ℤ[X]/<2X,X²>) CohomologyRing-RP² = RingEquivs.invRingEquiv RP²-CohomologyRing --- -} diff --git a/Cubical/ZCohomology/CohomologyRings/RP2wedgeS1.agda b/Cubical/ZCohomology/CohomologyRings/RP2wedgeS1.agda index 66fee63268..4c358d2960 100644 --- a/Cubical/ZCohomology/CohomologyRings/RP2wedgeS1.agda +++ b/Cubical/ZCohomology/CohomologyRings/RP2wedgeS1.agda @@ -371,7 +371,7 @@ module Equiv-RP²⋁S¹-Properties ℤ[x,y]/<2y,y²,xy,x²>→H*-RP²⋁S¹ : ℤ[x,y]/<2y,y²,xy,x²> → H* RP²⋁S¹ ℤ[x,y]/<2y,y²,xy,x²>→H*-RP²⋁S¹ = fst ℤ[X,Y]/<2Y,Y²,XY,X²>→H*R-RP²⋁S¹ -{- + ℤ[x,y]/<2y,y²,xy,x²>→H*-RP²⋁S¹-pres0 : ℤ[x,y]/<2y,y²,xy,x²>→H*-RP²⋁S¹ 0PℤI ≡ 0H* ℤ[x,y]/<2y,y²,xy,x²>→H*-RP²⋁S¹-pres0 = refl @@ -571,4 +571,3 @@ module _ where CohomologyRing-RP²⋁S¹ : RingEquiv (H*R RP²⋁S¹) (CommRing→Ring ℤ[X,Y]/<2Y,Y²,XY,X²>) CohomologyRing-RP²⋁S¹ = RingEquivs.invRingEquiv RP²⋁S¹-CohomologyRing --- -} diff --git a/Cubical/ZCohomology/CohomologyRings/S0.agda b/Cubical/ZCohomology/CohomologyRings/S0.agda index 6ac6cfaf53..1ce3385eb1 100644 --- a/Cubical/ZCohomology/CohomologyRings/S0.agda +++ b/Cubical/ZCohomology/CohomologyRings/S0.agda @@ -34,7 +34,7 @@ open import Cubical.ZCohomology.CohomologyRings.Unit -- Computation of the cohomology ring open RingEquivs -{- + Cohomology-Ring-S⁰P : RingEquiv (H*R (S₊ 0)) (DirectProd-Ring (CommRing→Ring ℤ[X]/X) (CommRing→Ring ℤ[X]/X)) Cohomology-Ring-S⁰P = compRingEquiv (CohomologyRing-Equiv (invIso Iso-⊤⊎⊤-Bool)) (compRingEquiv (CohomologyRing-Coproduct Unit Unit) @@ -44,4 +44,3 @@ Cohomology-Ring-S⁰ℤ : RingEquiv (H*R (S₊ 0)) (DirectProd-Ring (CommRing→ Cohomology-Ring-S⁰ℤ = compRingEquiv (CohomologyRing-Equiv (invIso Iso-⊤⊎⊤-Bool)) (compRingEquiv (CohomologyRing-Coproduct Unit Unit) (Coproduct-Equiv.Coproduct-Equiv-12 CohomologyRing-Unitℤ CohomologyRing-Unitℤ)) --} diff --git a/Cubical/ZCohomology/CohomologyRings/S1.agda b/Cubical/ZCohomology/CohomologyRings/S1.agda index f8ba962782..ac7b26e0f2 100644 --- a/Cubical/ZCohomology/CohomologyRings/S1.agda +++ b/Cubical/ZCohomology/CohomologyRings/S1.agda @@ -276,7 +276,7 @@ module Equiv-S1-Properties where H*-S¹→ℤ[x]/x² : H* (S₊ 1) → ℤ[x]/x² H*-S¹→ℤ[x]/x² = [_] ∘ H*-S¹→ℤ[x] -{- + H*-S¹→ℤ[x]/x²-pres+ : (x y : H* (S₊ 1)) → 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 @@ -337,4 +337,3 @@ module _ where CohomologyRing-S¹ : RingEquiv (H*R (S₊ 1)) (CommRing→Ring ℤ[X]/X²) CohomologyRing-S¹ = RingEquivs.invRingEquiv S¹-CohomologyRing --- -} diff --git a/Cubical/ZCohomology/CohomologyRings/S2wedgeS4.agda b/Cubical/ZCohomology/CohomologyRings/S2wedgeS4.agda index 6e4b8e1a38..410ee7b9de 100644 --- a/Cubical/ZCohomology/CohomologyRings/S2wedgeS4.agda +++ b/Cubical/ZCohomology/CohomologyRings/S2wedgeS4.agda @@ -376,7 +376,7 @@ module Equiv-RP2-Properties where ℤ[x,y]/→H*-S²⋁S⁴ : ℤ[x,y]/ → H* S²⋁S⁴ ℤ[x,y]/→H*-S²⋁S⁴ = fst ℤ[X,Y]/→H*R-S²⋁S⁴ -{- + ℤ[x,y]/→H*-S²⋁S⁴-pres0 : ℤ[x,y]/→H*-S²⋁S⁴ 0PℤI ≡ 0H* ℤ[x,y]/→H*-S²⋁S⁴-pres0 = refl @@ -522,4 +522,3 @@ module _ where CohomologyRing-S²⋁S⁴ : RingEquiv (H*R S²⋁S⁴) (CommRing→Ring ℤ[X,Y]/) CohomologyRing-S²⋁S⁴ = RingEquivs.invRingEquiv S²⋁S⁴-CohomologyRing --- -} diff --git a/Cubical/ZCohomology/CohomologyRings/Sn.agda b/Cubical/ZCohomology/CohomologyRings/Sn.agda index 47b0f283b7..1c0c287859 100644 --- a/Cubical/ZCohomology/CohomologyRings/Sn.agda +++ b/Cubical/ZCohomology/CohomologyRings/Sn.agda @@ -334,10 +334,8 @@ module Equiv-Sn-Properties (n : ℕ) where H*-Sⁿ→ℤ[x]/x² : H* (S₊ (suc n)) → ℤ[x]/x² H*-Sⁿ→ℤ[x]/x² = [_] ∘ H*-Sⁿ→ℤ[x] - opaque - unfolding quotientCommRingStr - 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 + 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 @@ -368,30 +366,29 @@ module Equiv-Sn-Properties (n : ℕ) where ----------------------------------------------------------------------------- -- Retraction -{- - opaque - unfolding quotientCommRingStr - e-retr : (x : ℤ[x]/x²) → H*-Sⁿ→ℤ[x]/x² (ℤ[x]/x²→H*-Sⁿ x) ≡ x - e-retr = SQ.elimProp (λ _ → isSetPℤI _ _) - (DS-Ind-Prop.f _ _ _ _ (λ _ → isSetPℤI _ _) - refl - base-case - λ {U V} ind-U ind-V → cong₂ _+PℤI_ ind-U ind-V) + + e-retr : (x : ℤ[x]/x²) → H*-Sⁿ→ℤ[x]/x² (ℤ[x]/x²→H*-Sⁿ x) ≡ x + e-retr = SQ.elimProp (λ _ → isSetPℤI _ _) + (DS-Ind-Prop.f _ _ _ _ (λ _ → isSetPℤI _ _) + refl + base-case + λ {U V} ind-U ind-V → cong₂ _+PℤI_ ind-U ind-V) + where + base-case : _ + base-case (zero ∷ []) a = cong [_] (cong (base-trad-H* 0 (ϕ₀ a)) part0) + ∙ cong [_] (cong (base (0 ∷ [])) (cong ϕ₀⁻¹ (transportRefl (ϕ₀ a)))) + ∙ cong [_] (cong (base (0 ∷ [])) (ϕ₀-retr a)) + base-case (one ∷ []) a = cong [_] (cong (base-trad-H* (suc n) (ϕₙ a)) (partSn (part (suc n)))) + ∙ cong [_] (cong (base (1 ∷ [])) (cong ϕₙ⁻¹ (transportRefl (ϕₙ a)))) + ∙ cong [_] (cong (base (1 ∷ [])) (ϕₙ-retr a)) + base-case (suc (suc k) ∷ []) a = eq/ _ _ ∣ ((λ x → base (k ∷ []) (-ℤ a)) , helper) ∣₁ where - base-case : _ - base-case (zero ∷ []) a = cong [_] (cong (base-trad-H* 0 (ϕ₀ a)) part0) - ∙ cong [_] (cong (base (0 ∷ [])) (cong ϕ₀⁻¹ (transportRefl (ϕ₀ a)))) - ∙ cong [_] (cong (base (0 ∷ [])) (ϕ₀-retr a)) - base-case (one ∷ []) a = cong [_] (cong (base-trad-H* (suc n) (ϕₙ a)) (partSn (part (suc n)))) - ∙ cong [_] (cong (base (1 ∷ [])) (cong ϕₙ⁻¹ (transportRefl (ϕₙ a)))) - ∙ cong [_] (cong (base (1 ∷ [])) (ϕₙ-retr a)) - base-case (suc (suc k) ∷ []) a = eq/ _ _ ∣ ((λ x → base (k ∷ []) (-ℤ a)) , helper) ∣₁ - where - helper : _ - helper = (+PℤIdL _) - ∙ cong₂ base (cong (λ X → X ∷ []) - (sym (+n-comm k 2))) (sym (·ℤIdR _)) - ∙ (sym (+PℤIdR _)) + helper : _ + helper = (+PℤIdL _) + ∙ cong₂ base (cong (λ X → X ∷ []) + (sym (+n-comm k 2))) (sym (·ℤIdR _)) + ∙ (sym (+PℤIdR _)) + ----------------------------------------------------------------------------- @@ -413,4 +410,3 @@ module _ (n : ℕ) where CohomologyRing-Sⁿ : RingEquiv (H*R (S₊ (suc n))) (CommRing→Ring ℤ[X]/X²) CohomologyRing-Sⁿ = RingEquivs.invRingEquiv Sⁿ-CohomologyRing --- -} diff --git a/Cubical/ZCohomology/CohomologyRings/Unit.agda b/Cubical/ZCohomology/CohomologyRings/Unit.agda index 405b90e40e..600ac9858a 100644 --- a/Cubical/ZCohomology/CohomologyRings/Unit.agda +++ b/Cubical/ZCohomology/CohomologyRings/Unit.agda @@ -241,7 +241,7 @@ module Equiv-Unit-Properties where H*-Unit→ℤ[x]/x : H* Unit → ℤ[x]/x H*-Unit→ℤ[x]/x = [_] ∘ H*-Unit→ℤ[x] -{- + H*-Unit→ℤ[x]/x-pres+ : (x y : H* Unit) → H*-Unit→ℤ[x]/x (x +H* y) ≡ (H*-Unit→ℤ[x]/x x) +PℤI (H*-Unit→ℤ[x]/x y) H*-Unit→ℤ[x]/x-pres+ x y = cong [_] (H*-Unit→ℤ[x]-pres+ x y) @@ -314,4 +314,3 @@ module _ where CohomologyRing-Unitℤ : RingEquiv (H*R Unit) (CommRing→Ring ℤCR) CohomologyRing-Unitℤ = compRingEquiv CohomologyRing-UnitP Equiv-ℤ[X]/X-ℤ --- -} From 6052d37e0b3802f39bb4de3c0557af0f662e391e Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 2 Aug 2024 18:11:33 +0200 Subject: [PATCH 43/83] make one example more meaningful --- Cubical/Tactics/CommRingSolver/Examples.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Cubical/Tactics/CommRingSolver/Examples.agda b/Cubical/Tactics/CommRingSolver/Examples.agda index 6d38ab13d5..531af32e9a 100644 --- a/Cubical/Tactics/CommRingSolver/Examples.agda +++ b/Cubical/Tactics/CommRingSolver/Examples.agda @@ -65,7 +65,7 @@ module Test (R : CommRing ℓ) (x y z : fst R) where ex1 : x ≡ x ex1 = solve! R - ex2 : x ≡ x + ex2 : (0r - 1r) · x ≡ 0r - x ex2 = solve! R ex3 : x + y ≡ y + x From 3079c8d32e6d8d9543fa190abac399c6b99e1478 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 2 Aug 2024 18:04:47 +0200 Subject: [PATCH 44/83] univalence for comm algebras with Evans help --- Cubical/Algebra/CommAlgebraAlt/Base.agda | 30 ++++------ .../Algebra/CommAlgebraAlt/Univalence.agda | 58 +++++++++---------- 2 files changed, 38 insertions(+), 50 deletions(-) diff --git a/Cubical/Algebra/CommAlgebraAlt/Base.agda b/Cubical/Algebra/CommAlgebraAlt/Base.agda index a3e3c600d0..1b8143839f 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Base.agda @@ -9,10 +9,12 @@ open import Cubical.Foundations.Structure using (⟨_⟩) open import Cubical.Foundations.HLevels open import Cubical.Foundations.Univalence open import Cubical.Foundations.Transport +open import Cubical.Foundations.Path open import Cubical.Data.Sigma open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Univalence open import Cubical.Algebra.Ring private @@ -73,26 +75,6 @@ module _ {R : CommRing ℓ} where p) where open CommRingStr (B .fst .snd) using (is-set) - CommAlgebra≡ : - {A B : CommAlgebra R ℓ'} - → (p : (A .fst) ≡ (B .fst)) - → (pathToEquiv $ cong fst p) .fst ∘ (A .snd) .fst ≡ (B .snd) .fst - → A ≡ B - CommAlgebra≡ {A = A} {B = B} p q = ΣPathTransport→PathΣ A B (p , pComm) - where - pComm : subst (CommRingHom R) p (A .snd) ≡ B .snd - pComm = - CommRingHom≡ - (fst (subst (CommRingHom R) p (A .snd)) - ≡⟨ sym (substCommSlice (CommRingHom R) (λ X → ⟨ R ⟩ → ⟨ X ⟩) (λ _ → fst) p (A .snd)) ⟩ - subst (λ X → ⟨ R ⟩ → ⟨ X ⟩) p (A .snd .fst) - ≡⟨ fromPathP (funTypeTransp (λ _ → ⟨ R ⟩) ⟨_⟩ p (A .snd .fst)) ⟩ - subst ⟨_⟩ p ∘ A .snd .fst ∘ subst (λ _ → ⟨ R ⟩) (sym p) - ≡⟨ cong ((subst ⟨_⟩ p ∘ A .snd .fst) ∘_) (funExt (λ _ → transportRefl _)) ⟩ - (pathToEquiv $ cong (λ r → fst r) p) .fst ∘ A .snd .fst - ≡⟨ q ⟩ - fst (B .snd) ∎) - CommAlgebraEquiv : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') → Type _ CommAlgebraEquiv A B = Σ[ f ∈ CommRingEquiv (A .fst) (B .fst) ] (f .fst .fst , f .snd) ∘cr A .snd ≡ B .snd @@ -109,6 +91,14 @@ module _ {R : CommRing ℓ} where (CommRingEquivs.CommRingEquiv≡ p) where open CommRingStr (B .fst .snd) using () renaming (is-set to isSetB) + isSetCommAlgebraEquiv : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') + → isSet (CommAlgebraEquiv A B) + isSetCommAlgebraEquiv A B = + isSetΣSndProp + (isSetCommRingEquiv _ _) + (λ e → isSetΣSndProp (isSet→ isSetB) (λ _ → isPropIsCommRingHom _ _ _) _ _) + where open CommRingStr (B .fst .snd) using () renaming (is-set to isSetB) + module CommAlgebraStr (A : CommAlgebra R ℓ') where open CommRingStr {{...}} instance diff --git a/Cubical/Algebra/CommAlgebraAlt/Univalence.agda b/Cubical/Algebra/CommAlgebraAlt/Univalence.agda index 0e595efabb..7095418b0b 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Univalence.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Univalence.agda @@ -7,6 +7,8 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Structure using (⟨_⟩) open import Cubical.Foundations.HLevels open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Path open import Cubical.Data.Sigma @@ -19,38 +21,34 @@ private ℓ ℓ' ℓ'' : Level -module _ (R : CommRing ℓ) where - private - to : (A B : CommAlgebra R ℓ') → CommAlgebraEquiv A B → A ≡ B - to A B e = CommAlgebra≡ - r≡ - ((pathToEquiv $ cong fst r≡) .fst ∘ A .snd .fst - ≡⟨ cong (_∘ A .snd .fst) (cong fst $ pathToEquiv-ua (e .fst .fst) ) ⟩ - CommAlgebraEquiv→CommRingHom e .fst ∘ A .snd .fst - ≡⟨ cong fst (e .snd) ⟩ - B .snd .fst ∎) - where r≡ : (A .fst) ≡ (B .fst) - r≡ = CommRingPath (A .fst) (B .fst) .fst (e .fst) - - test : (A : CommAlgebra R ℓ') → to A A (idCAlgEquiv A) ≡ refl - test A = {!to A A (idCAlgEquiv A)!} ≡⟨ {!!} ⟩ ? ≡⟨ {!!} ⟩ refl ∎ - - toIsEquiv : (A : CommAlgebra R ℓ') {B' : CommAlgebra R ℓ'} → (p : A ≡ B') → isContr (fiber (to A B') p) - toIsEquiv A = - J (λ B' A≡B' → isContr (fiber (to A B') A≡B')) - ((idCAlgEquiv A , {!test A!}) , - λ {(e , toe≡refl) → {!!}}) - - CommAlgebraPath : (A B : CommAlgebra R ℓ') → (CommAlgebraEquiv A B) ≃ (A ≡ B) - CommAlgebraPath {ℓ' = ℓ'} A B = to A B , {!!} - -{- -CommAlgebraPath : (R : CommRing ℓ) → (A B : CommAlgebra R ℓ') → (CommAlgebraEquiv A B) ≃ (A ≡ B) -CommAlgebraPath R = ∫ (𝒮ᴰ-CommAlgebra R) .UARel.ua +CommAlgebraPath : + (R : CommRing ℓ) + (A B : CommAlgebra R ℓ') + → (Σ[ f ∈ CommRingEquiv (A .fst) (B .fst) ] (f .fst .fst , f .snd) ∘cr A .snd ≡ B .snd) + ≃ (A ≡ B) +CommAlgebraPath R A B = + (Σ-cong-equiv + (CommRingPath _ _) + (λ e → compPathlEquiv (computeSubst e))) + ∙ₑ ΣPathTransport≃PathΣ A B + where computeSubst : + (e : CommRingEquiv (fst A) (fst B)) + → subst (CommRingHom R) (uaCommRing e) (A .snd) + ≡ (CommRingEquiv→CommRingHom e) ∘cr A .snd + computeSubst e = + CommRingHom≡ $ + (subst (CommRingHom R) (uaCommRing e) (A .snd)) .fst + ≡⟨ sym (substCommSlice (CommRingHom R) (λ X → ⟨ R ⟩ → ⟨ X ⟩) (λ _ → fst) (uaCommRing e) (A .snd)) ⟩ + subst (λ X → ⟨ R ⟩ → ⟨ X ⟩) (uaCommRing e) (A .snd .fst) + ≡⟨ fromPathP (funTypeTransp (λ _ → ⟨ R ⟩) ⟨_⟩ (uaCommRing e) (A .snd .fst)) ⟩ + subst ⟨_⟩ (uaCommRing e) ∘ A .snd .fst ∘ subst (λ _ → ⟨ R ⟩) (sym (uaCommRing e)) + ≡⟨ cong ((subst ⟨_⟩ (uaCommRing e) ∘ A .snd .fst) ∘_) (funExt (λ _ → transportRefl _)) ⟩ + (subst ⟨_⟩ (uaCommRing e) ∘ (A .snd .fst)) + ≡⟨ funExt (λ _ → uaβ (e .fst) _) ⟩ + (CommRingEquiv→CommRingHom e ∘cr A .snd) .fst ∎ uaCommAlgebra : {R : CommRing ℓ} {A B : CommAlgebra R ℓ'} → CommAlgebraEquiv A B → A ≡ B uaCommAlgebra {R = R} {A = A} {B = B} = equivFun (CommAlgebraPath R A B) isGroupoidCommAlgebra : {R : CommRing ℓ} → isGroupoid (CommAlgebra R ℓ') -isGroupoidCommAlgebra A B = isOfHLevelRespectEquiv 2 (CommAlgebraPath _ _ _) (isSetAlgebraEquiv _ _) --} +isGroupoidCommAlgebra A B = isOfHLevelRespectEquiv 2 (CommAlgebraPath _ _ _) (isSetCommAlgebraEquiv _ _) From 527861fcdaa23fbc2b68fcbf1369b9343eb31bda Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 2 Aug 2024 18:22:45 +0200 Subject: [PATCH 45/83] various fixes, continue with typevariate polynomials --- .../CommAlgebraAlt/Instances/Polynomials.agda | 2 +- Cubical/Algebra/CommRing/FiberedProduct.agda | 10 +- .../Instances/Polynomials/Typevariate.agda | 88 +--- .../Polynomials/Typevariate/Base.agda | 90 ++++ .../Polynomials/Typevariate/Properties.agda | 416 ++++++++++++++++++ Cubical/Algebra/CommRing/Properties.agda | 16 +- Cubical/Algebra/GradedRing/DirectSumFun.agda | 1 + .../ZariskiLattice/StructureSheaf.agda | 14 +- .../Categories/Instances/CommAlgebras.agda | 157 ++++--- Cubical/Categories/Instances/CommRings.agda | 2 +- 10 files changed, 609 insertions(+), 187 deletions(-) create mode 100644 Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda create mode 100644 Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda diff --git a/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda b/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda index 614c807471..c6626e88b0 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda @@ -12,4 +12,4 @@ private ℓ ℓ' : Level _[_] : (R : CommRing ℓ) (I : Type ℓ') → CommAlgebra R _ -R [ I ] = (R [ I ]ᵣ) , {!!} +R [ I ] = (R [ I ]ᵣ) , const R I diff --git a/Cubical/Algebra/CommRing/FiberedProduct.agda b/Cubical/Algebra/CommRing/FiberedProduct.agda index a64c67a600..1f1eb0519b 100644 --- a/Cubical/Algebra/CommRing/FiberedProduct.agda +++ b/Cubical/Algebra/CommRing/FiberedProduct.agda @@ -75,16 +75,16 @@ module _ (A B C : CommRing ℓ) (α : CommRingHom A C) (β : CommRingHom B C) wh fst fiberedProductPr₂ = snd ∘ fst snd fiberedProductPr₂ = makeIsCommRingHom refl (λ _ _ → refl) (λ _ _ → refl) - fiberedProductPr₁₂Commutes : compCommRingHom fiberedProduct A C fiberedProductPr₁ α - ≡ compCommRingHom fiberedProduct B C fiberedProductPr₂ β + fiberedProductPr₁₂Commutes : compCommRingHom fiberedProductPr₁ α + ≡ compCommRingHom fiberedProductPr₂ β fiberedProductPr₁₂Commutes = CommRingHom≡ (funExt (λ x → x .snd)) fiberedProductUnivProp : (D : CommRing ℓ) (h : CommRingHom D A) (k : CommRingHom D B) → - compCommRingHom D A C h α ≡ compCommRingHom D B C k β → + compCommRingHom h α ≡ compCommRingHom k β → ∃![ l ∈ CommRingHom D fiberedProduct ] - (h ≡ compCommRingHom D fiberedProduct A l fiberedProductPr₁) - × (k ≡ compCommRingHom D fiberedProduct B l fiberedProductPr₂) + (h ≡ compCommRingHom l fiberedProductPr₁) + × (k ≡ compCommRingHom l fiberedProductPr₂) fiberedProductUnivProp D (h , hh) (k , hk) H = uniqueExists f (CommRingHom≡ refl , CommRingHom≡ refl) (λ _ → isProp× (isSetCommRingHom _ _ _ _) (isSetCommRingHom _ _ _ _)) diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate.agda index 0b83c987b3..949fa66a93 100644 --- a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate.agda +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate.agda @@ -1,90 +1,6 @@ {-# OPTIONS --safe #-} module Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate where -{- - The ring of polynomials with coefficients in a given ring, - or in other words the free commutative algebra over a commutative ring. - Note that this is a constructive definition, which entails that polynomials - cannot be represented by lists of coefficients, where the last one is non-zero. - For rings with decidable equality, that is still possible. - I (Felix Cherubini) learned about this (and other) definition(s) from David Jaz Myers. - You can watch him talk about these things here: - https://www.youtube.com/watch?v=VNp-f_9MnVk - - This file contains - * the definition of the free commutative algebra on a type I over a commutative ring R as a HIT - (let us call that R[I]) - * a prove that the construction is an commutative R-algebra - For more, see the Properties file. --} -open import Cubical.Foundations.Prelude - -open import Cubical.Algebra.CommRing.Base - -private - variable - ℓ ℓ' : Level - -module Construction (R : CommRing ℓ) where - open CommRingStr (snd R) using (1r; 0r) renaming (_+_ to _+r_; _·_ to _·r_) - - data R[_] (I : Type ℓ') : Type (ℓ-max ℓ ℓ') where - var : I → R[ I ] - const : fst R → R[ I ] - _+_ : R[ I ] → R[ I ] → R[ I ] - -_ : R[ I ] → R[ I ] - _·_ : R[ I ] → R[ I ] → R[ I ] -- \cdot - - +-assoc : (x y z : R[ I ]) → x + (y + z) ≡ (x + y) + z - +-rid : (x : R[ I ]) → x + (const 0r) ≡ x - +-rinv : (x : R[ I ]) → x + (- x) ≡ (const 0r) - +-comm : (x y : R[ I ]) → x + y ≡ y + x - - ·-assoc : (x y z : R[ I ]) → x · (y · z) ≡ (x · y) · z - ·-lid : (x : R[ I ]) → (const 1r) · x ≡ x - ·-comm : (x y : R[ I ]) → x · y ≡ y · x - - ldist : (x y z : R[ I ]) → (x + y) · z ≡ (x · z) + (y · z) - - +HomConst : (s t : fst R) → const (s +r t) ≡ const s + const t - ·HomConst : (s t : fst R) → const (s ·r t) ≡ (const s) · (const t) - - 0-trunc : (x y : R[ I ]) (p q : x ≡ y) → p ≡ q - - opaque - isCommRing : (I : Type ℓ') → IsCommRing (const {I = I} 0r) (const 1r) _+_ _·_ -_ - isCommRing I = - makeIsCommRing 0-trunc - +-assoc +-rid +-rinv +-comm - ·-assoc (λ x → ·-comm _ _ ∙ ·-lid x) - (λ x y z → (x · (y + z)) ≡⟨ ·-comm _ _ ⟩ - ((y + z) · x) ≡⟨ ldist _ _ _ ⟩ - ((y · x) + (z · x)) ≡[ i ]⟨ (·-comm y x i + ·-comm z x i) ⟩ - ((x · y) + (x · z)) ∎) - ·-comm - - module _ (I : Type ℓ') where - open CommRingStr hiding (isCommRing) - commRingStr : CommRingStr (R[_] I) - commRingStr .0r = _ - commRingStr .1r = _ - commRingStr .CommRingStr._+_ = _ - commRingStr .CommRingStr._·_ = _ - CommRingStr.- commRingStr = _ - commRingStr .CommRingStr.isCommRing = isCommRing I - - - opaque - open IsCommRingHom - constIsCommRingHom : (I : Type ℓ') → IsCommRingHom (R .snd) (const {I = I}) (commRingStr I) - constIsCommRingHom I = makeIsCommRingHom refl +HomConst ·HomConst - -_[_]ᵣ : (R : CommRing ℓ) (I : Type ℓ') → CommRing (ℓ-max ℓ ℓ') -fst (R [ I ]ᵣ) = Construction.R[_] R I -snd (R [ I ]ᵣ) = Construction.commRingStr R I - -const : (R : CommRing ℓ) (I : Type ℓ') → CommRingHom R (R [ I ]ᵣ) -const R I .fst = let open Construction R - in R[_].const -const R I .snd = Construction.constIsCommRingHom R I +open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.Base public +open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.Properties public diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda new file mode 100644 index 0000000000..8ac50123d4 --- /dev/null +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda @@ -0,0 +1,90 @@ +{-# OPTIONS --safe #-} + +module Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.Base where +{- + The ring of polynomials with coefficients in a given ring, + or in other words the free commutative algebra over a commutative ring. + Note that this is a constructive definition, which entails that polynomials + cannot be represented by lists of coefficients, where the last one is non-zero. + For rings with decidable equality, that is still possible. + + I (Felix Cherubini) learned about this (and other) definition(s) from David Jaz Myers. + You can watch him talk about these things here: + https://www.youtube.com/watch?v=VNp-f_9MnVk + + This file contains + * the definition of the free commutative algebra on a type I over a commutative ring R as a HIT + (let us call that R[I]) + * a prove that the construction is an commutative R-algebra + For more, see the Properties file. +-} +open import Cubical.Foundations.Prelude + +open import Cubical.Algebra.CommRing.Base + +private + variable + ℓ ℓ' : Level + +module Construction (R : CommRing ℓ) where + open CommRingStr (snd R) using (1r; 0r) renaming (_+_ to _+r_; _·_ to _·r_) + + data R[_] (I : Type ℓ') : Type (ℓ-max ℓ ℓ') where + var : I → R[ I ] + const : fst R → R[ I ] + _+_ : R[ I ] → R[ I ] → R[ I ] + -_ : R[ I ] → R[ I ] + _·_ : R[ I ] → R[ I ] → R[ I ] -- \cdot + + +-assoc : (x y z : R[ I ]) → x + (y + z) ≡ (x + y) + z + +-rid : (x : R[ I ]) → x + (const 0r) ≡ x + +-rinv : (x : R[ I ]) → x + (- x) ≡ (const 0r) + +-comm : (x y : R[ I ]) → x + y ≡ y + x + + ·-assoc : (x y z : R[ I ]) → x · (y · z) ≡ (x · y) · z + ·-lid : (x : R[ I ]) → (const 1r) · x ≡ x + ·-comm : (x y : R[ I ]) → x · y ≡ y · x + + ldist : (x y z : R[ I ]) → (x + y) · z ≡ (x · z) + (y · z) + + +HomConst : (s t : fst R) → const (s +r t) ≡ const s + const t + ·HomConst : (s t : fst R) → const (s ·r t) ≡ (const s) · (const t) + + 0-trunc : (x y : R[ I ]) (p q : x ≡ y) → p ≡ q + + opaque + isCommRing : (I : Type ℓ') → IsCommRing (const {I = I} 0r) (const 1r) _+_ _·_ -_ + isCommRing I = + makeIsCommRing 0-trunc + +-assoc +-rid +-rinv +-comm + ·-assoc (λ x → ·-comm _ _ ∙ ·-lid x) + (λ x y z → (x · (y + z)) ≡⟨ ·-comm _ _ ⟩ + ((y + z) · x) ≡⟨ ldist _ _ _ ⟩ + ((y · x) + (z · x)) ≡[ i ]⟨ (·-comm y x i + ·-comm z x i) ⟩ + ((x · y) + (x · z)) ∎) + ·-comm + + module _ (I : Type ℓ') where + open CommRingStr hiding (isCommRing) + commRingStr : CommRingStr (R[_] I) + commRingStr .0r = _ + commRingStr .1r = _ + commRingStr .CommRingStr._+_ = _ + commRingStr .CommRingStr._·_ = _ + CommRingStr.- commRingStr = _ + commRingStr .CommRingStr.isCommRing = isCommRing I + + + opaque + open IsCommRingHom + constIsCommRingHom : (I : Type ℓ') → IsCommRingHom (R .snd) (const {I = I}) (commRingStr I) + constIsCommRingHom I = makeIsCommRingHom refl +HomConst ·HomConst + +_[_]ᵣ : (R : CommRing ℓ) (I : Type ℓ') → CommRing (ℓ-max ℓ ℓ') +fst (R [ I ]ᵣ) = Construction.R[_] R I +snd (R [ I ]ᵣ) = Construction.commRingStr R I + +constPolynomial : (R : CommRing ℓ) (I : Type ℓ') → CommRingHom R (R [ I ]ᵣ) +constPolynomial R I .fst = let open Construction R + in R[_].const +constPolynomial R I .snd = Construction.constIsCommRingHom R I diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda new file mode 100644 index 0000000000..92d08e922f --- /dev/null +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda @@ -0,0 +1,416 @@ +{-# OPTIONS --safe #-} + +module Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.Properties where +{- + This file contains + * an elimination principle for proving some proposition for all elements of R[I] + ('elimProp') + * definitions of the induced maps appearing in the universal property of R[I], + that is: * for any map I → A, where A is a commutative R-algebra, + the induced algebra homomorphism R[I] → A + ('inducedHom') + * for any hom R[I] → A, the 'restricttion to variables' I → A + ('evaluateAt') + * a proof that the two constructions are inverse to each other + ('homMapIso') + * proofs that the constructions are natural in various ways + * a proof that the Polynomials on zero generators are the initial R-Algebra + ('freeOn⊥') +-} + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function hiding (const) +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Sigma.Properties using (Σ≡Prop) +open import Cubical.HITs.SetTruncation + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.Base +open import Cubical.Algebra.Ring.Properties +{- +open import Cubical.Algebra.CommAlgebra.FreeCommAlgebra.Base +open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebra.Instances.Initial +open import Cubical.Algebra.Algebra +open import Cubical.Algebra.Module using (module ModuleTheory) +-} +open import Cubical.Data.Empty +open import Cubical.Data.Sigma + +open import Cubical.Tactics.CommRingSolver + +private + variable + ℓ ℓ' ℓ'' : Level + +module _ {R : CommRing ℓ} {I : Type ℓ'} where + open CommRingStr ⦃...⦄ -- (snd R) + private instance + _ = snd R + _ = snd (R [ I ]ᵣ) + + module C = Construction + open C using (var; const) + + {- + Construction of the 'elimProp' eliminator. + -} + module _ + {P : ⟨ R [ I ]ᵣ ⟩ → Type ℓ''} + (isPropP : {x : _} → isProp (P x)) + (onVar : {x : I} → P (var x)) + (onConst : {r : ⟨ R ⟩} → P (const r)) + (on+ : {x y : ⟨ R [ I ]ᵣ ⟩} → P x → P y → P (x + y)) + (on· : {x y : ⟨ R [ I ]ᵣ ⟩} → P x → P y → P (x · y)) + where + + private + on- : {x : ⟨ R [ I ]ᵣ ⟩} → P x → P (- x) + on- {x} Px = subst P (minusByMult x) (on· onConst Px) + where minusByMult : (x : _) → (const (- 1r)) · x ≡ - x + minusByMult x = + (const (- 1r)) · x ≡⟨ cong (_· x) (pres- 1r) ⟩ + (- const (1r)) · x ≡⟨ cong (λ u → (- u) · x) pres1 ⟩ + (- 1r) · x ≡⟨ sym (-IsMult-1 x) ⟩ + - x ∎ + where open RingTheory (CommRing→Ring (R [ I ]ᵣ)) using (-IsMult-1) + open IsCommRingHom (constPolynomial R I .snd) + + -- A helper to deal with the path constructor cases. + mkPathP : + {x0 x1 : ⟨ R [ I ]ᵣ ⟩} → + (p : x0 ≡ x1) → + (Px0 : P (x0)) → + (Px1 : P (x1)) → + PathP (λ i → P (p i)) Px0 Px1 + mkPathP _ = isProp→PathP λ i → isPropP + + elimProp : ((x : _) → P x) + + elimProp (var _) = onVar + elimProp (const _) = onConst + elimProp (x C.+ y) = on+ (elimProp x) (elimProp y) + elimProp (C.- x) = on- (elimProp x) + elimProp (x C.· y) = on· (elimProp x) (elimProp y) + + elimProp (C.+-assoc x y z i) = + mkPathP (C.+-assoc x y z) + (on+ (elimProp x) (on+ (elimProp y) (elimProp z))) + (on+ (on+ (elimProp x) (elimProp y)) (elimProp z)) + i + elimProp (C.+-rid x i) = + mkPathP (C.+-rid x) + (on+ (elimProp x) onConst) + (elimProp x) + i + elimProp (C.+-rinv x i) = + mkPathP (C.+-rinv x) + (on+ (elimProp x) (on- (elimProp x))) + onConst + i + elimProp (C.+-comm x y i) = + mkPathP (C.+-comm x y) + (on+ (elimProp x) (elimProp y)) + (on+ (elimProp y) (elimProp x)) + i + elimProp (C.·-assoc x y z i) = + mkPathP (C.·-assoc x y z) + (on· (elimProp x) (on· (elimProp y) (elimProp z))) + (on· (on· (elimProp x) (elimProp y)) (elimProp z)) + i + elimProp (C.·-lid x i) = + mkPathP (C.·-lid x) + (on· onConst (elimProp x)) + (elimProp x) + i + elimProp (C.·-comm x y i) = + mkPathP (C.·-comm x y) + (on· (elimProp x) (elimProp y)) + (on· (elimProp y) (elimProp x)) + i + elimProp (C.ldist x y z i) = + mkPathP (C.ldist x y z) + (on· (on+ (elimProp x) (elimProp y)) (elimProp z)) + (on+ (on· (elimProp x) (elimProp z)) (on· (elimProp y) (elimProp z))) + i + elimProp (C.+HomConst s t i) = + mkPathP (C.+HomConst s t) + onConst + (on+ onConst onConst) + i + elimProp (C.·HomConst s t i) = + mkPathP (C.·HomConst s t) + onConst + (on· onConst onConst) + i + + elimProp (C.0-trunc x y p q i j) = + isOfHLevel→isOfHLevelDep 2 (λ _ → isProp→isSet isPropP) + (elimProp x) + (elimProp y) + (cong elimProp p) + (cong elimProp q) + (C.0-trunc x y p q) + i + j + + {- + Construction of the induced map. + -} + module _ (S : CommRing ℓ'') (f : CommRingHom R S) (φ : I → ⟨ S ⟩) where + open CommRingStr ⦃...⦄ + private instance + _ = R .snd + _ = S .snd +-- open AlgebraTheory (CommRing→Ring R) (CommAlgebra→Algebra A) + + imageOf0Works : (f .fst 0r) · 1r ≡ 0r + imageOf0Works = ? -- ⋆AnnihilL 1a + + imageOf1Works : 1r · 1r ≡ 1r + imageOf1Works = ? -- ⋆IdL 1a + + + inducedMap : ⟨ R [ I ]ᵣ ⟩ → ⟨ S ⟩ + inducedMap (var x) = φ x + inducedMap (const r) = (f .fst r) · 1r + inducedMap (P C.+ Q) = (inducedMap P) + (inducedMap Q) + inducedMap (C.- P) = - inducedMap P + inducedMap (C.+-assoc P Q S i) = +Assoc (inducedMap P) (inducedMap Q) (inducedMap S) i + inducedMap (C.+-rid P i) = + let + eq : (inducedMap P) + (inducedMap (const 0r)) ≡ (inducedMap P) + eq = (inducedMap P) + (inducedMap (const 0r)) ≡⟨ refl ⟩ + (inducedMap P) + ((f .fst 0r) · 1r) ≡⟨ cong + (λ u → (inducedMap P) + u) + (imageOf0Works) ⟩ + (inducedMap P) + 0r ≡⟨ +IdR _ ⟩ + (inducedMap P) ∎ + in eq i + inducedMap (C.+-rinv P i) = + let eq : (inducedMap P - inducedMap P) ≡ (inducedMap (const 0r)) + eq = (inducedMap P - inducedMap P) ≡⟨ +InvR _ ⟩ + 0a ≡⟨ sym imageOf0Works ⟩ + (inducedMap (const 0r))∎ + in eq i + inducedMap (C.+-comm P Q i) = +Comm (inducedMap P) (inducedMap Q) i + inducedMap (P C.· Q) = inducedMap P · inducedMap Q + inducedMap (C.·-assoc P Q S i) = ·Assoc (inducedMap P) (inducedMap Q) (inducedMap S) i + inducedMap (C.·-lid P i) = + let eq = inducedMap (const 1r) · inducedMap P ≡⟨ cong (λ u → u · inducedMap P) imageOf1Works ⟩ + 1a · inducedMap P ≡⟨ ·IdL (inducedMap P) ⟩ + inducedMap P ∎ + in eq i + inducedMap (C.·-comm P Q i) = ·Comm (inducedMap P) (inducedMap Q) i + inducedMap (C.ldist P Q S i) = ·DistL+ (inducedMap P) (inducedMap Q) (inducedMap S) i + inducedMap (C.+HomConst s t i) = ⋆DistL+ s t 1a i + inducedMap (C.·HomConst s t i) = + let eq = (s ·r t) ⋆ 1a ≡⟨ cong (λ u → u ⋆ 1a) (·r-comm _ _) ⟩ + (t ·r s) ⋆ 1a ≡⟨ ⋆Assoc t s 1a ⟩ + t ⋆ (s ⋆ 1a) ≡⟨ cong (λ u → t ⋆ u) (sym (·IdR _)) ⟩ + t ⋆ ((s ⋆ 1a) · 1a) ≡⟨ ⋆AssocR t (s ⋆ 1a) 1a ⟩ + (s ⋆ 1a) · (t ⋆ 1a) ∎ + in eq i + inducedMap (C.0-trunc P Q p q i j) = + is-set (inducedMap P) (inducedMap Q) (cong _ p) (cong _ q) i j + +{- + + module _ where + open IsAlgebraHom + + inducedHom : CommAlgebraHom (R [ I ]ᵣ) A + inducedHom .fst = inducedMap + inducedHom .snd .pres0 = ⋆AnnihilL _ + inducedHom .snd .pres1 = imageOf1Works + inducedHom .snd .pres+ x y = refl + inducedHom .snd .pres· x y = refl + inducedHom .snd .pres- x = refl + inducedHom .snd .pres⋆ r x = + (r ⋆ 1a) · inducedMap x ≡⟨ ⋆AssocL r 1a (inducedMap x) ⟩ + r ⋆ (1a · inducedMap x) ≡⟨ cong (λ u → r ⋆ u) (·IdL (inducedMap x)) ⟩ + r ⋆ inducedMap x ∎ + +-} + +{- + + module _ (A : CommAlgebra R ℓ'') where + open CommAlgebraStr (A .snd) + open AlgebraTheory (CommRing→Ring R) (CommAlgebra→Algebra A) + + Hom = CommAlgebraHom (R [ I ]) A + open IsAlgebraHom + + evaluateAt : Hom → I → ⟨ A ⟩ + evaluateAt φ x = φ .fst (var x) + + mapRetrievable : ∀ (φ : I → ⟨ A ⟩) + → evaluateAt (inducedHom A φ) ≡ φ + mapRetrievable φ = refl + + homRetrievable : ∀ (f : Hom) + → inducedMap A (evaluateAt f) ≡ fst f + homRetrievable f = funExt ( + elimProp + {P = λ x → ι x ≡ f $a x} + (is-set _ _) + (λ {x} → refl) + (λ {r} → + r ⋆ 1a ≡⟨ cong (λ u → r ⋆ u) (sym f.pres1) ⟩ + r ⋆ (f $a (const 1r)) ≡⟨ sym (f.pres⋆ r _) ⟩ + f $a (const r C.· const 1r) ≡⟨ cong (λ u → f $a u) (sym (C.·HomConst r 1r)) ⟩ + f $a (const (r ·r 1r)) ≡⟨ cong (λ u → f $a (const u)) (·r-rid r) ⟩ + f $a (const r) ∎) + + (λ {x} {y} eq-x eq-y → + ι (x C.+ y) ≡⟨ refl ⟩ + (ι x + ι y) ≡⟨ cong (λ u → u + ι y) eq-x ⟩ + ((f $a x) + ι y) ≡⟨ cong (λ u → (f $a x) + u) eq-y ⟩ + ((f $a x) + (f $a y)) ≡⟨ sym (f.pres+ _ _) ⟩ + (f $a (x C.+ y)) ∎) + + (λ {x} {y} eq-x eq-y → + ι (x C.· y) ≡⟨ refl ⟩ + ι x · ι y ≡⟨ cong (λ u → u · ι y) eq-x ⟩ + (f $a x) · (ι y) ≡⟨ cong (λ u → (f $a x) · u) eq-y ⟩ + (f $a x) · (f $a y) ≡⟨ sym (f.pres· _ _) ⟩ + f $a (x C.· y) ∎) + ) + where + ι = inducedMap A (evaluateAt f) + module f = IsAlgebraHom (f .snd) + + +evaluateAt : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') + (f : CommAlgebraHom (R [ I ]) A) + → (I → fst A) +evaluateAt A f x = f $a (Construction.var x) + +inducedHom : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') + (φ : I → fst A ) + → CommAlgebraHom (R [ I ]) A +inducedHom A φ = Theory.inducedHom A φ + + +homMapIso : {R : CommRing ℓ} {I : Type ℓ''} (A : CommAlgebra R ℓ') + → Iso (CommAlgebraHom (R [ I ]) A) (I → (fst A)) +Iso.fun (homMapIso A) = evaluateAt A +Iso.inv (homMapIso A) = inducedHom A +Iso.rightInv (homMapIso A) = λ ϕ → Theory.mapRetrievable A ϕ +Iso.leftInv (homMapIso {R = R} {I = I} A) = + λ f → Σ≡Prop (λ f → isPropIsCommAlgebraHom {M = R [ I ]} {N = A} f) + (Theory.homRetrievable A f) + +inducedHomUnique : + {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') (φ : I → fst A ) + → (f : CommAlgebraHom (R [ I ]) A) → ((i : I) → fst f (Construction.var i) ≡ φ i) + → f ≡ inducedHom A φ +inducedHomUnique {I = I} A φ f p = + isoFunInjective (homMapIso A) f (inducedHom A φ) λ j i → p i j + +homMapPath : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R (ℓ-max ℓ ℓ')) + → CommAlgebraHom (R [ I ]) A ≡ (I → fst A) +homMapPath A = isoToPath (homMapIso A) + +{- Corollary: Two homomorphisms with the same values on generators are equal -} +equalByUMP : {R : CommRing ℓ} {I : Type ℓ'} + → (A : CommAlgebra R ℓ'') + → (f g : CommAlgebraHom (R [ I ]) A) + → ((i : I) → fst f (Construction.var i) ≡ fst g (Construction.var i)) + → (x : ⟨ R [ I ] ⟩) → fst f x ≡ fst g x +equalByUMP {R = R} {I = I} A f g = funExt⁻ ∘ cong fst ∘ isoFunInjective (homMapIso A) f g ∘ funExt + +{- A corollary, which is useful for constructing isomorphisms to + algebras with the same universal property -} +isIdByUMP : {R : CommRing ℓ} {I : Type ℓ'} + → (f : CommAlgebraHom (R [ I ]) (R [ I ])) + → ((i : I) → fst f (Construction.var i) ≡ Construction.var i) + → (x : ⟨ R [ I ] ⟩) → fst f x ≡ x +isIdByUMP {R = R} {I = I} f p = equalByUMP (R [ I ]) f (idCAlgHom (R [ I ])) p + +-- The homomorphism induced by the variables is the identity. +inducedHomVar : (R : CommRing ℓ) (I : Type ℓ') + → inducedHom (R [ I ]) Construction.var ≡ idCAlgHom (R [ I ]) +inducedHomVar R I = isoFunInjective (homMapIso (R [ I ])) _ _ refl + +module _ {R : CommRing ℓ} {A B : CommAlgebra R ℓ''} where + open AlgebraHoms + A′ = CommAlgebra→Algebra A + B′ = CommAlgebra→Algebra B + R′ = (CommRing→Ring R) + ν : AlgebraHom A′ B′ → (⟨ A ⟩ → ⟨ B ⟩) + ν φ = φ .fst + + {- + Hom(R[I],A) → (I → A) + ↓ ↓ψ + Hom(R[I],B) → (I → B) + -} + naturalEvR : {I : Type ℓ'} (ψ : CommAlgebraHom A B) + (f : CommAlgebraHom (R [ I ]) A) + → (fst ψ) ∘ evaluateAt A f ≡ evaluateAt B (ψ ∘a f) + naturalEvR ψ f = refl + + {- + Hom(R[I],A) ← (I → A) + ↓ ↓ψ + Hom(R[I],B) ← (I → B) + -} + natIndHomR : {I : Type ℓ'} (ψ : CommAlgebraHom A B) + (ϕ : I → ⟨ A ⟩) + → ψ ∘a inducedHom A ϕ ≡ inducedHom B (fst ψ ∘ ϕ) + natIndHomR ψ ϕ = isoFunInjective (homMapIso B) _ _ + (evaluateAt B (ψ ∘a (inducedHom A ϕ)) ≡⟨ refl ⟩ + fst ψ ∘ evaluateAt A (inducedHom A ϕ) ≡⟨ refl ⟩ + fst ψ ∘ ϕ ≡⟨ Iso.rightInv (homMapIso B) _ ⟩ + evaluateAt B (inducedHom B (fst ψ ∘ ϕ)) ∎) + + {- + Hom(R[I],A) → (I → A) + ↓ ↓ + Hom(R[J],A) → (J → A) + -} + naturalEvL : {I J : Type ℓ'} (φ : J → I) + (f : CommAlgebraHom (R [ I ]) A) + → (evaluateAt A f) ∘ φ + ≡ evaluateAt A (f ∘a (inducedHom (R [ I ]) (λ x → Construction.var (φ x)))) + naturalEvL φ f = refl + +module _ {R : CommRing ℓ} where + {- + Prove that the FreeCommAlgebra over R on zero generators is + isomorphic to the initial R-Algebra - R itsself. + -} + freeOn⊥ : CommAlgebraEquiv (R [ ⊥ ]) (initialCAlg R) + freeOn⊥ = + equivByInitiality + R (R [ ⊥ ]) + {- Show that R[⊥] has the universal property of the + initial R-Algbera and conclude that those are isomorphic -} + λ B → let to : CommAlgebraHom (R [ ⊥ ]) B → (⊥ → fst B) + to = evaluateAt B + + from : (⊥ → fst B) → CommAlgebraHom (R [ ⊥ ]) B + from = inducedHom B + + from-to : (x : _) → from (to x) ≡ x + from-to x = + Σ≡Prop (λ f → isPropIsCommAlgebraHom {M = R [ ⊥ ]} {N = B} f) + (Theory.homRetrievable B x) + + equiv : CommAlgebraHom (R [ ⊥ ]) B ≃ (⊥ → fst B) + equiv = + isoToEquiv + (iso to from (λ x → isContr→isOfHLevel 1 isContr⊥→A _ _) from-to) + in isOfHLevelRespectEquiv 0 (invEquiv equiv) isContr⊥→A + +module _ {R : CommRing ℓ} {I : Type ℓ} where + baseRingHom : CommRingHom R (CommAlgebra→CommRing (R [ I ])) + baseRingHom = snd (Iso.fun (CommAlgChar.CommAlgIso R) (R [ I ])) +-} diff --git a/Cubical/Algebra/CommRing/Properties.agda b/Cubical/Algebra/CommRing/Properties.agda index f21d6be99a..ffda3e340f 100644 --- a/Cubical/Algebra/CommRing/Properties.agda +++ b/Cubical/Algebra/CommRing/Properties.agda @@ -156,28 +156,28 @@ module _ where idCommRingHom : (R : CommRing ℓ) → CommRingHom R R idCommRingHom R = RingHom→CommRingHom (idRingHom (CommRing→Ring R)) - compCommRingHom : (R : CommRing ℓ) (S : CommRing ℓ') (T : CommRing ℓ'') + compCommRingHom : {R : CommRing ℓ} {S : CommRing ℓ'} {T : CommRing ℓ''} → CommRingHom R S → CommRingHom S T → CommRingHom R T - compCommRingHom R S T f g = + compCommRingHom f g = RingHom→CommRingHom (compRingHom (CommRingHom→RingHom f) (CommRingHom→RingHom g)) _∘cr_ : {R : CommRing ℓ} {S : CommRing ℓ'} {T : CommRing ℓ''} → CommRingHom S T → CommRingHom R S → CommRingHom R T - g ∘cr f = compCommRingHom _ _ _ f g + g ∘cr f = compCommRingHom f g compIdCommRingHom : {R S : CommRing ℓ} (f : CommRingHom R S) - → compCommRingHom _ _ _ (idCommRingHom R) f ≡ f + → compCommRingHom (idCommRingHom R) f ≡ f compIdCommRingHom f = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) refl idCompCommRingHom : {R S : CommRing ℓ} (f : CommRingHom R S) - → compCommRingHom _ _ _ f (idCommRingHom S) ≡ f + → compCommRingHom f (idCommRingHom S) ≡ f idCompCommRingHom f = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) refl compAssocCommRingHom : {R : CommRing ℓ} {S : CommRing ℓ'} {T : CommRing ℓ''} {U : CommRing ℓ'''} (f : CommRingHom R S) (g : CommRingHom S T) (h : CommRingHom T U) - → compCommRingHom _ _ _ (compCommRingHom _ _ _ f g) h - ≡ compCommRingHom _ _ _ f (compCommRingHom _ _ _ g h) + → compCommRingHom (compCommRingHom f g) h + ≡ compCommRingHom f (compCommRingHom g h) compAssocCommRingHom f g h = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) refl open Iso @@ -192,7 +192,7 @@ module CommRingEquivs where compCommRingEquiv : {A : CommRing ℓ} {B : CommRing ℓ'} {C : CommRing ℓ''} → CommRingEquiv A B → CommRingEquiv B C → CommRingEquiv A C compCommRingEquiv f g .fst = compEquiv (f .fst) (g .fst) - compCommRingEquiv f g .snd = compCommRingHom _ _ _ (f .fst .fst , f .snd) (g .fst .fst , g .snd) .snd + compCommRingEquiv f g .snd = compCommRingHom (f .fst .fst , f .snd) (g .fst .fst , g .snd) .snd invCommRingEquiv : (A : CommRing ℓ) → (B : CommRing ℓ') → CommRingEquiv A B → CommRingEquiv B A fst (invCommRingEquiv A B e) = invEquiv (fst e) diff --git a/Cubical/Algebra/GradedRing/DirectSumFun.agda b/Cubical/Algebra/GradedRing/DirectSumFun.agda index c88e643ac2..961eb8db97 100644 --- a/Cubical/Algebra/GradedRing/DirectSumFun.agda +++ b/Cubical/Algebra/GradedRing/DirectSumFun.agda @@ -39,6 +39,7 @@ open import Cubical.Algebra.Ring open import Cubical.Algebra.GradedRing.Base open import Cubical.Algebra.GradedRing.DirectSumHIT open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Univalence open import Cubical.HITs.PropositionalTruncation as PT diff --git a/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda index 0ca7597380..a51e4e950e 100644 --- a/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda +++ b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda @@ -280,8 +280,8 @@ module _ {ℓ : Level} (R' : CommRing ℓ) where where open Units R[1/ h ]AsCommRing open Sum (CommRing→Ring R[1/ h ]AsCommRing) - open IsRingHom (snd /1AsCommRingHom) - open SumMap _ _ /1AsCommRingHom + open IsCommRingHom (snd /1AsCommRingHom) + open SumMap _ _ (CommRingHom→RingHom /1AsCommRingHom) instance h⁻ᵐ : (h ^ m) /1 ∈ₚ (R[1/ h ]AsCommRing ˣ) h⁻ᵐ = [ 1r , h ^ m , ∣ m , refl ∣₁ ] @@ -326,7 +326,7 @@ module _ {ℓ : Level} (R' : CommRing ℓ) where -- R[1/h][(fᵢfⱼ)/1/1] =/= R[1/h][(fᵢ/1 · fⱼ/1)/1] -- definitionally open Cone - open IsRingHom + open IsCommRingHom module D i = DoubleLoc R' h (f i) @@ -347,11 +347,11 @@ module _ {ℓ : Level} (R' : CommRing ℓ) where (Σ≡Prop (λ _ → isPropPropTrunc) (sym (·IdR 1r)))) pres- (snd (coneOut /1/1Cone (pair i j i Date: Fri, 2 Aug 2024 18:53:17 +0200 Subject: [PATCH 46/83] induce a com ring hom --- .../Polynomials/Typevariate/Properties.agda | 64 +++++-------------- 1 file changed, 17 insertions(+), 47 deletions(-) diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda index 92d08e922f..21db0b7cc4 100644 --- a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda @@ -32,13 +32,7 @@ open import Cubical.HITs.SetTruncation open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.Base open import Cubical.Algebra.Ring.Properties -{- -open import Cubical.Algebra.CommAlgebra.FreeCommAlgebra.Base -open import Cubical.Algebra.CommAlgebra -open import Cubical.Algebra.CommAlgebra.Instances.Initial -open import Cubical.Algebra.Algebra -open import Cubical.Algebra.Module using (module ModuleTheory) --} + open import Cubical.Data.Empty open import Cubical.Data.Sigma @@ -163,80 +157,56 @@ module _ {R : CommRing ℓ} {I : Type ℓ'} where Construction of the induced map. -} module _ (S : CommRing ℓ'') (f : CommRingHom R S) (φ : I → ⟨ S ⟩) where - open CommRingStr ⦃...⦄ private instance - _ = R .snd _ = S .snd --- open AlgebraTheory (CommRing→Ring R) (CommAlgebra→Algebra A) - - imageOf0Works : (f .fst 0r) · 1r ≡ 0r - imageOf0Works = ? -- ⋆AnnihilL 1a - - imageOf1Works : 1r · 1r ≡ 1r - imageOf1Works = ? -- ⋆IdL 1a + open IsCommRingHom inducedMap : ⟨ R [ I ]ᵣ ⟩ → ⟨ S ⟩ inducedMap (var x) = φ x - inducedMap (const r) = (f .fst r) · 1r + inducedMap (const r) = (f .fst r) inducedMap (P C.+ Q) = (inducedMap P) + (inducedMap Q) inducedMap (C.- P) = - inducedMap P inducedMap (C.+-assoc P Q S i) = +Assoc (inducedMap P) (inducedMap Q) (inducedMap S) i inducedMap (C.+-rid P i) = let eq : (inducedMap P) + (inducedMap (const 0r)) ≡ (inducedMap P) - eq = (inducedMap P) + (inducedMap (const 0r)) ≡⟨ refl ⟩ - (inducedMap P) + ((f .fst 0r) · 1r) ≡⟨ cong - (λ u → (inducedMap P) + u) - (imageOf0Works) ⟩ + eq = (inducedMap P) + (inducedMap (const 0r)) ≡⟨⟩ + (inducedMap P) + ((f .fst 0r)) ≡⟨ cong ((inducedMap P) +_) (pres0 (f .snd)) ⟩ (inducedMap P) + 0r ≡⟨ +IdR _ ⟩ (inducedMap P) ∎ in eq i inducedMap (C.+-rinv P i) = let eq : (inducedMap P - inducedMap P) ≡ (inducedMap (const 0r)) eq = (inducedMap P - inducedMap P) ≡⟨ +InvR _ ⟩ - 0a ≡⟨ sym imageOf0Works ⟩ + 0r ≡⟨ sym (pres0 (f .snd)) ⟩ (inducedMap (const 0r))∎ in eq i inducedMap (C.+-comm P Q i) = +Comm (inducedMap P) (inducedMap Q) i inducedMap (P C.· Q) = inducedMap P · inducedMap Q inducedMap (C.·-assoc P Q S i) = ·Assoc (inducedMap P) (inducedMap Q) (inducedMap S) i inducedMap (C.·-lid P i) = - let eq = inducedMap (const 1r) · inducedMap P ≡⟨ cong (λ u → u · inducedMap P) imageOf1Works ⟩ - 1a · inducedMap P ≡⟨ ·IdL (inducedMap P) ⟩ + let eq = inducedMap (const 1r) · inducedMap P ≡⟨ cong (λ u → u · inducedMap P) (pres1 (f .snd)) ⟩ + 1r · inducedMap P ≡⟨ ·IdL (inducedMap P) ⟩ inducedMap P ∎ in eq i inducedMap (C.·-comm P Q i) = ·Comm (inducedMap P) (inducedMap Q) i inducedMap (C.ldist P Q S i) = ·DistL+ (inducedMap P) (inducedMap Q) (inducedMap S) i - inducedMap (C.+HomConst s t i) = ⋆DistL+ s t 1a i - inducedMap (C.·HomConst s t i) = - let eq = (s ·r t) ⋆ 1a ≡⟨ cong (λ u → u ⋆ 1a) (·r-comm _ _) ⟩ - (t ·r s) ⋆ 1a ≡⟨ ⋆Assoc t s 1a ⟩ - t ⋆ (s ⋆ 1a) ≡⟨ cong (λ u → t ⋆ u) (sym (·IdR _)) ⟩ - t ⋆ ((s ⋆ 1a) · 1a) ≡⟨ ⋆AssocR t (s ⋆ 1a) 1a ⟩ - (s ⋆ 1a) · (t ⋆ 1a) ∎ - in eq i + inducedMap (C.+HomConst s t i) = pres+ (f .snd) s t i + inducedMap (C.·HomConst s t i) = pres· (f .snd) s t i inducedMap (C.0-trunc P Q p q i j) = is-set (inducedMap P) (inducedMap Q) (cong _ p) (cong _ q) i j -{- - module _ where - open IsAlgebraHom + open IsCommRingHom - inducedHom : CommAlgebraHom (R [ I ]ᵣ) A + inducedHom : CommRingHom (R [ I ]ᵣ) S inducedHom .fst = inducedMap - inducedHom .snd .pres0 = ⋆AnnihilL _ - inducedHom .snd .pres1 = imageOf1Works - inducedHom .snd .pres+ x y = refl - inducedHom .snd .pres· x y = refl - inducedHom .snd .pres- x = refl - inducedHom .snd .pres⋆ r x = - (r ⋆ 1a) · inducedMap x ≡⟨ ⋆AssocL r 1a (inducedMap x) ⟩ - r ⋆ (1a · inducedMap x) ≡⟨ cong (λ u → r ⋆ u) (·IdL (inducedMap x)) ⟩ - r ⋆ inducedMap x ∎ - --} + inducedHom .snd .pres0 = pres0 (f .snd) + inducedHom .snd .pres1 = pres1 (f. snd) + inducedHom .snd .pres+ = λ _ _ → refl + inducedHom .snd .pres· = λ _ _ → refl + inducedHom .snd .pres- = λ _ → refl {- From 2f43b22b6bbf1c4e76741e2ea9739daf17f875b9 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 2 Aug 2024 21:47:16 +0200 Subject: [PATCH 47/83] Comm ring of polynomials on a general variable type and its universal property --- .../Polynomials/Typevariate/Base.agda | 5 + .../Polynomials/Typevariate/Properties.agda | 251 +++++------------- 2 files changed, 68 insertions(+), 188 deletions(-) diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda index 8ac50123d4..058e2b2c4c 100644 --- a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda @@ -19,6 +19,7 @@ module Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.Base where For more, see the Properties file. -} open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure using (⟨_⟩) open import Cubical.Algebra.CommRing.Base @@ -88,3 +89,7 @@ constPolynomial : (R : CommRing ℓ) (I : Type ℓ') → CommRingHom R (R [ I ] constPolynomial R I .fst = let open Construction R in R[_].const constPolynomial R I .snd = Construction.constIsCommRingHom R I + +opaque + var : {R : CommRing ℓ} {I : Type ℓ'} → I → ⟨ R [ I ]ᵣ ⟩ + var = Construction.var diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda index 21db0b7cc4..e721203840 100644 --- a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda @@ -1,21 +1,10 @@ {-# OPTIONS --safe #-} - module Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.Properties where {- This file contains - * an elimination principle for proving some proposition for all elements of R[I] + * an elimination principle for proving some proposition for all elements of R[I]ᵣ ('elimProp') * definitions of the induced maps appearing in the universal property of R[I], - that is: * for any map I → A, where A is a commutative R-algebra, - the induced algebra homomorphism R[I] → A - ('inducedHom') - * for any hom R[I] → A, the 'restricttion to variables' I → A - ('evaluateAt') - * a proof that the two constructions are inverse to each other - ('homMapIso') - * proofs that the constructions are natural in various ways - * a proof that the Polynomials on zero generators are the initial R-Algebra - ('freeOn⊥') -} open import Cubical.Foundations.Prelude @@ -43,13 +32,13 @@ private ℓ ℓ' ℓ'' : Level module _ {R : CommRing ℓ} {I : Type ℓ'} where - open CommRingStr ⦃...⦄ -- (snd R) + open CommRingStr ⦃...⦄ private instance _ = snd R _ = snd (R [ I ]ᵣ) module C = Construction - open C using (var; const) + open C using (const) {- Construction of the 'elimProp' eliminator. @@ -57,7 +46,7 @@ module _ {R : CommRing ℓ} {I : Type ℓ'} where module _ {P : ⟨ R [ I ]ᵣ ⟩ → Type ℓ''} (isPropP : {x : _} → isProp (P x)) - (onVar : {x : I} → P (var x)) + (onVar : {x : I} → P (C.var x)) (onConst : {r : ⟨ R ⟩} → P (const r)) (on+ : {x y : ⟨ R [ I ]ᵣ ⟩} → P x → P y → P (x + y)) (on· : {x y : ⟨ R [ I ]ᵣ ⟩} → P x → P y → P (x · y)) @@ -86,7 +75,7 @@ module _ {R : CommRing ℓ} {I : Type ℓ'} where elimProp : ((x : _) → P x) - elimProp (var _) = onVar + elimProp (C.var _) = onVar elimProp (const _) = onConst elimProp (x C.+ y) = on+ (elimProp x) (elimProp y) elimProp (C.- x) = on- (elimProp x) @@ -155,6 +144,15 @@ module _ {R : CommRing ℓ} {I : Type ℓ'} where {- Construction of the induced map. + In this module and the module below, we will show the universal property + of the polynomial ring as a commutative ring. + + R ──→ R[I] + \ ∣ + f ∃! for a given φ : I → S + ↘ ↙ + S + -} module _ (S : CommRing ℓ'') (f : CommRingHom R S) (φ : I → ⟨ S ⟩) where private instance @@ -163,7 +161,7 @@ module _ {R : CommRing ℓ} {I : Type ℓ'} where open IsCommRingHom inducedMap : ⟨ R [ I ]ᵣ ⟩ → ⟨ S ⟩ - inducedMap (var x) = φ x + inducedMap (C.var x) = φ x inducedMap (const r) = (f .fst r) inducedMap (P C.+ Q) = (inducedMap P) + (inducedMap Q) inducedMap (C.- P) = - inducedMap P @@ -208,179 +206,56 @@ module _ {R : CommRing ℓ} {I : Type ℓ'} where inducedHom .snd .pres· = λ _ _ → refl inducedHom .snd .pres- = λ _ → refl -{- - - module _ (A : CommAlgebra R ℓ'') where - open CommAlgebraStr (A .snd) - open AlgebraTheory (CommRing→Ring R) (CommAlgebra→Algebra A) + opaque + inducedHomComm : inducedHom ∘cr constPolynomial R I ≡ f + inducedHomComm = CommRingHom≡ $ funExt (λ r → refl) - Hom = CommAlgebraHom (R [ I ]) A - open IsAlgebraHom - - evaluateAt : Hom → I → ⟨ A ⟩ - evaluateAt φ x = φ .fst (var x) - - mapRetrievable : ∀ (φ : I → ⟨ A ⟩) - → evaluateAt (inducedHom A φ) ≡ φ - mapRetrievable φ = refl - - homRetrievable : ∀ (f : Hom) - → inducedMap A (evaluateAt f) ≡ fst f - homRetrievable f = funExt ( +module _ {R : CommRing ℓ} {I : Type ℓ'} (S : CommRing ℓ'') (f : CommRingHom R S) where + open CommRingStr ⦃...⦄ + private instance + _ = S .snd + _ = (R [ I ]ᵣ) .snd + open IsCommRingHom + + evalVar : CommRingHom (R [ I ]ᵣ) S → I → ⟨ S ⟩ + evalVar f = f .fst ∘ var + + opaque + unfolding var + evalInduce : ∀ (φ : I → ⟨ S ⟩) + → evalVar (inducedHom S f φ) ≡ φ + evalInduce φ = refl + + opaque + unfolding var + induceEval : (g : CommRingHom (R [ I ]ᵣ) S) + (p : g .fst ∘ constPolynomial R I .fst ≡ f .fst) + → (inducedHom S f (evalVar g)) ≡ g + induceEval g p = + let theMap : ⟨ R [ I ]ᵣ ⟩ → ⟨ S ⟩ + theMap = inducedMap S f (evalVar g) + in + CommRingHom≡ $ + funExt $ elimProp - {P = λ x → ι x ≡ f $a x} (is-set _ _) (λ {x} → refl) - (λ {r} → - r ⋆ 1a ≡⟨ cong (λ u → r ⋆ u) (sym f.pres1) ⟩ - r ⋆ (f $a (const 1r)) ≡⟨ sym (f.pres⋆ r _) ⟩ - f $a (const r C.· const 1r) ≡⟨ cong (λ u → f $a u) (sym (C.·HomConst r 1r)) ⟩ - f $a (const (r ·r 1r)) ≡⟨ cong (λ u → f $a (const u)) (·r-rid r) ⟩ - f $a (const r) ∎) - + (λ {r} → sym (cong (λ f → f r) p)) (λ {x} {y} eq-x eq-y → - ι (x C.+ y) ≡⟨ refl ⟩ - (ι x + ι y) ≡⟨ cong (λ u → u + ι y) eq-x ⟩ - ((f $a x) + ι y) ≡⟨ cong (λ u → (f $a x) + u) eq-y ⟩ - ((f $a x) + (f $a y)) ≡⟨ sym (f.pres+ _ _) ⟩ - (f $a (x C.+ y)) ∎) - - (λ {x} {y} eq-x eq-y → - ι (x C.· y) ≡⟨ refl ⟩ - ι x · ι y ≡⟨ cong (λ u → u · ι y) eq-x ⟩ - (f $a x) · (ι y) ≡⟨ cong (λ u → (f $a x) · u) eq-y ⟩ - (f $a x) · (f $a y) ≡⟨ sym (f.pres· _ _) ⟩ - f $a (x C.· y) ∎) - ) - where - ι = inducedMap A (evaluateAt f) - module f = IsAlgebraHom (f .snd) - - -evaluateAt : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') - (f : CommAlgebraHom (R [ I ]) A) - → (I → fst A) -evaluateAt A f x = f $a (Construction.var x) - -inducedHom : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') - (φ : I → fst A ) - → CommAlgebraHom (R [ I ]) A -inducedHom A φ = Theory.inducedHom A φ - - -homMapIso : {R : CommRing ℓ} {I : Type ℓ''} (A : CommAlgebra R ℓ') - → Iso (CommAlgebraHom (R [ I ]) A) (I → (fst A)) -Iso.fun (homMapIso A) = evaluateAt A -Iso.inv (homMapIso A) = inducedHom A -Iso.rightInv (homMapIso A) = λ ϕ → Theory.mapRetrievable A ϕ -Iso.leftInv (homMapIso {R = R} {I = I} A) = - λ f → Σ≡Prop (λ f → isPropIsCommAlgebraHom {M = R [ I ]} {N = A} f) - (Theory.homRetrievable A f) - -inducedHomUnique : - {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') (φ : I → fst A ) - → (f : CommAlgebraHom (R [ I ]) A) → ((i : I) → fst f (Construction.var i) ≡ φ i) - → f ≡ inducedHom A φ -inducedHomUnique {I = I} A φ f p = - isoFunInjective (homMapIso A) f (inducedHom A φ) λ j i → p i j - -homMapPath : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R (ℓ-max ℓ ℓ')) - → CommAlgebraHom (R [ I ]) A ≡ (I → fst A) -homMapPath A = isoToPath (homMapIso A) - -{- Corollary: Two homomorphisms with the same values on generators are equal -} -equalByUMP : {R : CommRing ℓ} {I : Type ℓ'} - → (A : CommAlgebra R ℓ'') - → (f g : CommAlgebraHom (R [ I ]) A) - → ((i : I) → fst f (Construction.var i) ≡ fst g (Construction.var i)) - → (x : ⟨ R [ I ] ⟩) → fst f x ≡ fst g x -equalByUMP {R = R} {I = I} A f g = funExt⁻ ∘ cong fst ∘ isoFunInjective (homMapIso A) f g ∘ funExt - -{- A corollary, which is useful for constructing isomorphisms to - algebras with the same universal property -} -isIdByUMP : {R : CommRing ℓ} {I : Type ℓ'} - → (f : CommAlgebraHom (R [ I ]) (R [ I ])) - → ((i : I) → fst f (Construction.var i) ≡ Construction.var i) - → (x : ⟨ R [ I ] ⟩) → fst f x ≡ x -isIdByUMP {R = R} {I = I} f p = equalByUMP (R [ I ]) f (idCAlgHom (R [ I ])) p - --- The homomorphism induced by the variables is the identity. -inducedHomVar : (R : CommRing ℓ) (I : Type ℓ') - → inducedHom (R [ I ]) Construction.var ≡ idCAlgHom (R [ I ]) -inducedHomVar R I = isoFunInjective (homMapIso (R [ I ])) _ _ refl - -module _ {R : CommRing ℓ} {A B : CommAlgebra R ℓ''} where - open AlgebraHoms - A′ = CommAlgebra→Algebra A - B′ = CommAlgebra→Algebra B - R′ = (CommRing→Ring R) - ν : AlgebraHom A′ B′ → (⟨ A ⟩ → ⟨ B ⟩) - ν φ = φ .fst - - {- - Hom(R[I],A) → (I → A) - ↓ ↓ψ - Hom(R[I],B) → (I → B) - -} - naturalEvR : {I : Type ℓ'} (ψ : CommAlgebraHom A B) - (f : CommAlgebraHom (R [ I ]) A) - → (fst ψ) ∘ evaluateAt A f ≡ evaluateAt B (ψ ∘a f) - naturalEvR ψ f = refl - - {- - Hom(R[I],A) ← (I → A) - ↓ ↓ψ - Hom(R[I],B) ← (I → B) - -} - natIndHomR : {I : Type ℓ'} (ψ : CommAlgebraHom A B) - (ϕ : I → ⟨ A ⟩) - → ψ ∘a inducedHom A ϕ ≡ inducedHom B (fst ψ ∘ ϕ) - natIndHomR ψ ϕ = isoFunInjective (homMapIso B) _ _ - (evaluateAt B (ψ ∘a (inducedHom A ϕ)) ≡⟨ refl ⟩ - fst ψ ∘ evaluateAt A (inducedHom A ϕ) ≡⟨ refl ⟩ - fst ψ ∘ ϕ ≡⟨ Iso.rightInv (homMapIso B) _ ⟩ - evaluateAt B (inducedHom B (fst ψ ∘ ϕ)) ∎) - - {- - Hom(R[I],A) → (I → A) - ↓ ↓ - Hom(R[J],A) → (J → A) - -} - naturalEvL : {I J : Type ℓ'} (φ : J → I) - (f : CommAlgebraHom (R [ I ]) A) - → (evaluateAt A f) ∘ φ - ≡ evaluateAt A (f ∘a (inducedHom (R [ I ]) (λ x → Construction.var (φ x)))) - naturalEvL φ f = refl - -module _ {R : CommRing ℓ} where - {- - Prove that the FreeCommAlgebra over R on zero generators is - isomorphic to the initial R-Algebra - R itsself. - -} - freeOn⊥ : CommAlgebraEquiv (R [ ⊥ ]) (initialCAlg R) - freeOn⊥ = - equivByInitiality - R (R [ ⊥ ]) - {- Show that R[⊥] has the universal property of the - initial R-Algbera and conclude that those are isomorphic -} - λ B → let to : CommAlgebraHom (R [ ⊥ ]) B → (⊥ → fst B) - to = evaluateAt B - - from : (⊥ → fst B) → CommAlgebraHom (R [ ⊥ ]) B - from = inducedHom B - - from-to : (x : _) → from (to x) ≡ x - from-to x = - Σ≡Prop (λ f → isPropIsCommAlgebraHom {M = R [ ⊥ ]} {N = B} f) - (Theory.homRetrievable B x) - - equiv : CommAlgebraHom (R [ ⊥ ]) B ≃ (⊥ → fst B) - equiv = - isoToEquiv - (iso to from (λ x → isContr→isOfHLevel 1 isContr⊥→A _ _) from-to) - in isOfHLevelRespectEquiv 0 (invEquiv equiv) isContr⊥→A - -module _ {R : CommRing ℓ} {I : Type ℓ} where - baseRingHom : CommRingHom R (CommAlgebra→CommRing (R [ I ])) - baseRingHom = snd (Iso.fun (CommAlgChar.CommAlgIso R) (R [ I ])) --} + theMap (x + y) ≡⟨ pres+ (inducedHom S f (evalVar g) .snd) x y ⟩ + theMap x + theMap y ≡[ i ]⟨ (eq-x i + eq-y i) ⟩ + (g $cr x + g $cr y) ≡⟨ sym (pres+ (g .snd) _ _) ⟩ + (g $cr (x + y)) ∎) + λ {x} {y} eq-x eq-y → + theMap (x · y) ≡⟨ pres· (inducedHom S f (evalVar g) .snd) x y ⟩ + theMap x · theMap y ≡[ i ]⟨ (eq-x i · eq-y i) ⟩ + (g $cr x · g $cr y) ≡⟨ sym (pres· (g .snd) _ _) ⟩ + (g $cr (x · y)) ∎ + + opaque + inducedHomUnique : (φ : I → ⟨ S ⟩) + (g : CommRingHom (R [ I ]ᵣ) S) + (p : g .fst ∘ constPolynomial R I .fst ≡ f .fst) + (q : evalVar g ≡ φ) + → inducedHom S f φ ≡ g + inducedHomUnique φ g p q = cong (inducedHom S f) (sym q) ∙ induceEval g p From 57c7b168bd83de41390990a497ae741ea090e877 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Sat, 3 Aug 2024 14:42:51 +0200 Subject: [PATCH 48/83] improve interface --- Cubical/Algebra/CommAlgebraAlt/Base.agda | 92 +++++++++++++++++------- 1 file changed, 68 insertions(+), 24 deletions(-) diff --git a/Cubical/Algebra/CommAlgebraAlt/Base.agda b/Cubical/Algebra/CommAlgebraAlt/Base.agda index 1b8143839f..89eeaac07c 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Base.agda @@ -55,7 +55,7 @@ module _ {R : CommRing ℓ} where ⟨_⟩ₐ→ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} (f : CommAlgebraHom A B) → ⟨ A ⟩ₐ → ⟨ B ⟩ₐ ⟨ f ⟩ₐ→ = f .fst .fst - _$ca_ : {A : CommAlgebra R ℓ} {B : CommAlgebra R ℓ'} → (φ : CommAlgebraHom A B) → (x : ⟨ A ⟩ₐ) → ⟨ B ⟩ₐ + _$ca_ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} → (φ : CommAlgebraHom A B) → (x : ⟨ A ⟩ₐ) → ⟨ B ⟩ₐ φ $ca x = φ .fst .fst x _∘ca_ : {A : CommAlgebra R ℓ} {B : CommAlgebra R ℓ'} {C : CommAlgebra R ℓ''} @@ -99,36 +99,80 @@ module _ {R : CommRing ℓ} where (λ e → isSetΣSndProp (isSet→ isSetB) (λ _ → isPropIsCommRingHom _ _ _) _ _) where open CommRingStr (B .fst .snd) using () renaming (is-set to isSetB) - module CommAlgebraStr (A : CommAlgebra R ℓ') where + {- + Contrary to most algebraic structures, this one only contains + law and structure of a CommAlgebra, which it is *in addition* + to its CommRing structure. + -} + record CommAlgebraStr (A : Type ℓ') : Type (ℓ-suc (ℓ-max ℓ ℓ')) where open CommRingStr {{...}} instance - _ = A - _ : CommRingStr ⟨ A ⟩ₐ - _ = (A .fst .snd) _ : CommRingStr (R .fst) _ = (R .snd) + field + crStr : CommRingStr A + _⋆_ : ⟨ R ⟩ → A → A + ⋆Assoc : (r s : ⟨ R ⟩) (x : A) → (r · s) ⋆ x ≡ r ⋆ (s ⋆ x) + ⋆DistR+ : (r : ⟨ R ⟩) (x y : A) → r ⋆ (CommRingStr._+_ crStr x y) ≡ CommRingStr._+_ crStr(r ⋆ x) (r ⋆ y) + ⋆DistL+ : (r s : ⟨ R ⟩) (x : A) → (r + s) ⋆ x ≡ CommRingStr._+_ crStr (r ⋆ x) (s ⋆ x) + ⋆IdL : (x : A) → 1r ⋆ x ≡ x + ⋆AssocL : (r : ⟨ R ⟩) (x y : A) → CommRingStr._·_ crStr (r ⋆ x) y ≡ r ⋆ (CommRingStr._·_ crStr x y) infixl 7 _⋆_ - _⋆_ : ⟨ R ⟩ → ⟨ A ⟩ₐ → ⟨ A ⟩ₐ - _⋆_ r x = (A .snd .fst r) · x - ⋆Assoc : (r s : ⟨ R ⟩) (x : ⟨ A ⟩ₐ) → (r · s) ⋆ x ≡ r ⋆ (s ⋆ x) - ⋆Assoc r s x = cong (_· x) (pres· r s) ∙ sym (·Assoc _ _ _) - where open IsCommRingHom (A .snd .snd) - - ⋆DistR+ : (r : ⟨ R ⟩) (x y : ⟨ A ⟩ₐ) → r ⋆ (x + y) ≡ (r ⋆ x) + (r ⋆ y) - ⋆DistR+ r x y = ·DistR+ _ _ _ - - ⋆DistL+ : (r s : ⟨ R ⟩) (x : ⟨ A ⟩ₐ) → (r + s) ⋆ x ≡ (r ⋆ x) + (s ⋆ x) - ⋆DistL+ r s x = cong (_· x) (pres+ r s) ∙ ·DistL+ _ _ _ - where open IsCommRingHom (A .snd .snd) - - ⋆IdL : (x : ⟨ A ⟩ₐ) → 1r ⋆ x ≡ x - ⋆IdL x = cong (_· x) pres1 ∙ ·IdL x - where open IsCommRingHom (A .snd .snd) - - ⋆AssocL : (r : ⟨ R ⟩) (x y : ⟨ A ⟩ₐ) → (r ⋆ x) · y ≡ r ⋆ (x · y) - ⋆AssocL r x y = sym (·Assoc (A .snd .fst r) x y) + CommAlgebra→CommAlgebraStr : (A : CommAlgebra R ℓ') → CommAlgebraStr ⟨ A ⟩ₐ + CommAlgebra→CommAlgebraStr A = + let open CommRingStr ⦃...⦄ + instance + _ : CommRingStr (R .fst) + _ = R .snd + _ = A .fst .snd + in record + { crStr = A .fst .snd + ; _⋆_ = λ r x → (A .snd .fst r) · x + ; ⋆Assoc = λ r s x → cong (_· x) (IsCommRingHom.pres· (A .snd .snd) r s) ∙ sym (·Assoc _ _ _) + ; ⋆DistR+ = λ r x y → ·DistR+ _ _ _ + ; ⋆DistL+ = λ r s x → cong (_· x) (IsCommRingHom.pres+ (A .snd .snd) r s) ∙ ·DistL+ _ _ _ + ; ⋆IdL = λ x → cong (_· x) (IsCommRingHom.pres1 (A .snd .snd)) ∙ ·IdL x + ; ⋆AssocL = λ r x y → sym (·Assoc (A .snd .fst r) x y) + } + + module IsCommAlgebraHom {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} (f : CommAlgebraHom A B) where + open IsCommRingHom (f .fst .snd) + open CommRingStr ⦃...⦄ + open CommAlgebraStr ⦃...⦄ + private instance + _ = CommAlgebra→CommAlgebraStr A + _ = CommAlgebra→CommAlgebraStr B + _ = B .fst .snd + + opaque + pres⋆ : (r : ⟨ R ⟩) (x : ⟨ A ⟩ₐ) → f $ca (r ⋆ x) ≡ r ⋆ f $ca x + pres⋆ r x = f $ca (r ⋆ x) ≡⟨ pres· (A .snd .fst r) x ⟩ + (f $ca (A .snd .fst r)) · (f $ca x) ≡⟨ cong (_· (f $ca x)) (cong (λ g → g .fst r) (f .snd)) ⟩ + r ⋆ f $ca x ∎ + + CommAlgebraHomFromCommRingHom : + {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → (f : CommRingHom (A .fst) (B .fst)) + → ((r : ⟨ R ⟩) (x : ⟨ A ⟩ₐ) → f .fst (CommAlgebraStr._⋆_ (CommAlgebra→CommAlgebraStr A) r x) + ≡ CommAlgebraStr._⋆_ (CommAlgebra→CommAlgebraStr B) r (f .fst x)) + → CommAlgebraHom A B + CommAlgebraHomFromCommRingHom f pres⋆ .fst = f + CommAlgebraHomFromCommRingHom {A = A} {B = B} f pres⋆ .snd = + let open CommRingStr ⦃...⦄ + open CommAlgebraStr ⦃...⦄ + instance + _ = A .fst .snd + _ = B .fst .snd + _ = CommAlgebra→CommAlgebraStr B + in CommRingHom≡ + (funExt (λ (r : ⟨ R ⟩) → + f .fst (A .snd .fst r) ≡⟨ cong (λ u → f .fst u) (sym (·IdR _)) ⟩ + f .fst ((A .snd .fst r) · 1r) ≡⟨ pres⋆ r (CommRingStr.1r (A .fst .snd)) ⟩ + r ⋆ (f .fst 1r) ≡⟨ cong (r ⋆_) (IsCommRingHom.pres1 (f .snd)) ⟩ + r ⋆ 1r ≡⟨ ·IdR _ ⟩ + B .snd .fst r ∎)) {- Convenient forgetful functions -} module _ {R : CommRing ℓ} where From b576f32ef897f9e41c2ed86b3797681070409d7a Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Sat, 3 Aug 2024 15:04:27 +0200 Subject: [PATCH 49/83] more interface, more opaque --- Cubical/Algebra/CommAlgebraAlt/Base.agda | 16 +++++++++++----- 1 file changed, 11 insertions(+), 5 deletions(-) diff --git a/Cubical/Algebra/CommAlgebraAlt/Base.agda b/Cubical/Algebra/CommAlgebraAlt/Base.agda index 89eeaac07c..3cacda5f4e 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Base.agda @@ -5,7 +5,7 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Structure using (⟨_⟩) +open import Cubical.Foundations.Structure using (⟨_⟩; withOpaqueStr) open import Cubical.Foundations.HLevels open import Cubical.Foundations.Univalence open import Cubical.Foundations.Transport @@ -32,13 +32,13 @@ module _ {R : CommRing ℓ} where CommAlgebraHom A B = Σ[ f ∈ CommRingHom (fst A) (fst B) ] f ∘cr snd A ≡ snd B idCAlgHom : (A : CommAlgebra R ℓ') → CommAlgebraHom A A - idCAlgHom A = (idCommRingHom (fst A)) , CommRingHom≡ refl + idCAlgHom A = withOpaqueStr $ (idCommRingHom (fst A)) , CommRingHom≡ refl compCommAlgebraHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} → (g : CommAlgebraHom B C) → (f : CommAlgebraHom A B) → CommAlgebraHom A C compCommAlgebraHom {A = A} {B = B} {C = C} g f = - ((fst g) ∘cr (fst f)) , pasting + withOpaqueStr $ ((fst g) ∘cr (fst f)) , pasting where opaque pasting : ((fst g) ∘cr (fst f)) ∘cr snd A ≡ snd C @@ -120,6 +120,7 @@ module _ {R : CommRing ℓ} where ⋆AssocL : (r : ⟨ R ⟩) (x y : A) → CommRingStr._·_ crStr (r ⋆ x) y ≡ r ⋆ (CommRingStr._·_ crStr x y) infixl 7 _⋆_ + {- TODO: make laws opaque -} CommAlgebra→CommAlgebraStr : (A : CommAlgebra R ℓ') → CommAlgebraStr ⟨ A ⟩ₐ CommAlgebra→CommAlgebraStr A = let open CommRingStr ⦃...⦄ @@ -158,8 +159,7 @@ module _ {R : CommRing ℓ} where → ((r : ⟨ R ⟩) (x : ⟨ A ⟩ₐ) → f .fst (CommAlgebraStr._⋆_ (CommAlgebra→CommAlgebraStr A) r x) ≡ CommAlgebraStr._⋆_ (CommAlgebra→CommAlgebraStr B) r (f .fst x)) → CommAlgebraHom A B - CommAlgebraHomFromCommRingHom f pres⋆ .fst = f - CommAlgebraHomFromCommRingHom {A = A} {B = B} f pres⋆ .snd = + CommAlgebraHomFromCommRingHom {A = A} {B = B} f pres⋆ = withOpaqueStr $ f , let open CommRingStr ⦃...⦄ open CommAlgebraStr ⦃...⦄ instance @@ -193,3 +193,9 @@ module _ {R : CommRing ℓ} where → CommAlgebraEquiv A B → CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) CommAlgebraEquiv→CommRingHom e = CommRingEquiv→CommRingHom (e .fst) + + CommAlgebraEquiv→CommAlgebraHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → CommAlgebraEquiv A B + → CommAlgebraHom A B + CommAlgebraEquiv→CommAlgebraHom f = + withOpaqueStr $ (CommRingEquiv→CommRingHom (f .fst)) , CommRingHom≡ (cong (fst) (f .snd)) From a19fbd5e811dc29b61aec8c540ca48820233e309 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Sat, 3 Aug 2024 15:04:43 +0200 Subject: [PATCH 50/83] fix --- Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda | 1 - 1 file changed, 1 deletion(-) diff --git a/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda index ee89da6450..1ee161b93f 100644 --- a/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda @@ -176,6 +176,5 @@ module _ _ = (A / I).fst .snd opaque - unfolding CommRing.quotientCommRingStr isZeroFromIdeal : (x : ⟨ A ⟩ₐ) → x ∈ (fst I) → ⟨ quotientHom A I ⟩ₐ→ x ≡ 0r isZeroFromIdeal x x∈I = eq/ x 0r (subst (_∈ fst I) (solve! (CommAlgebra→CommRing A)) x∈I ) From 5089c7f4eab2dca3809a5c27f31a08531fbb32d0 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Sat, 3 Aug 2024 19:26:26 +0200 Subject: [PATCH 51/83] start with localization, fixes, more opaque things --- .../Algebra/CommAlgebraAlt/Localisation.agda | 313 ++++++++++++++++++ .../Algebra/CommAlgebraAlt/Univalence.agda | 1 + .../Algebra/CommRing/Localisation/Base.agda | 1 + Cubical/Foundations/Structure.agda | 17 +- 4 files changed, 326 insertions(+), 6 deletions(-) create mode 100644 Cubical/Algebra/CommAlgebraAlt/Localisation.agda diff --git a/Cubical/Algebra/CommAlgebraAlt/Localisation.agda b/Cubical/Algebra/CommAlgebraAlt/Localisation.agda new file mode 100644 index 0000000000..f366f42c1c --- /dev/null +++ b/Cubical/Algebra/CommAlgebraAlt/Localisation.agda @@ -0,0 +1,313 @@ +{-# OPTIONS --safe --lossy-unification #-} +module Cubical.Algebra.CommAlgebraAlt.Localisation where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Structure using (⟨_⟩; withOpaqueStr; makeOpaque) +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Powerset + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat renaming ( _+_ to _+ℕ_ ; _·_ to _·ℕ_ ; _^_ to _^ℕ_ + ; +-comm to +ℕ-comm ; +-assoc to +ℕ-assoc + ; ·-assoc to ·ℕ-assoc ; ·-comm to ·ℕ-comm) +open import Cubical.Data.FinData + +open import Cubical.Reflection.StrictEquiv + +open import Cubical.Structures.Axioms +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.CommRing.Base +open import Cubical.Algebra.CommRing.Properties +open import Cubical.Algebra.CommRing.Ideal +open import Cubical.Algebra.CommRing.FGIdeal +open import Cubical.Algebra.CommRing.RadicalIdeal +open import Cubical.Algebra.CommRing.Localisation +open import Cubical.Algebra.Ring +open import Cubical.Algebra.CommAlgebraAlt.Base +open import Cubical.Algebra.CommAlgebraAlt.Univalence + +open import Cubical.Tactics.CommRingSolver + +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.HITs.PropositionalTruncation as PT + +private + variable + ℓ ℓ′ : Level + +module AlgLoc (R : CommRing ℓ) + (S : ℙ ⟨ R ⟩) (SMultClosedSubset : isMultClosedSubset R S) where + open isMultClosedSubset + open CommRingStr ⦃...⦄ + open CommAlgebraStr ⦃...⦄ + open RingTheory (CommRing→Ring R) + open CommRingTheory R + open Loc R S SMultClosedSubset hiding (S) + open S⁻¹RUniversalProp R S SMultClosedSubset + open IsCommRingHom + + + S⁻¹RAsCommAlg : CommAlgebra R ℓ + S⁻¹RAsCommAlg = (S⁻¹RAsCommRing , /1AsCommRingHom) + private instance + _ = CommAlgebra→CommAlgebraStr S⁻¹RAsCommAlg + _ = S⁻¹RAsCommAlg .fst .snd + _ = R .snd + + module _ (A : CommAlgebra R ℓ) where + private instance + _ = CommAlgebra→CommAlgebraStr A + _ = A .fst .snd + + hasLocAlgUniversalProp : (∀ s → s ∈ S → s ⋆ 1r ∈ A .fst ˣ) + → Type (ℓ-suc ℓ) + hasLocAlgUniversalProp _ = (B : CommAlgebra R ℓ) + → (let instance + _ = (CommAlgebra→CommAlgebraStr B) + _ = B .fst .snd + in ∀ s → s ∈ S → s ⋆ 1r ∈ B .fst ˣ) + → isContr (CommAlgebraHom A B) + + S⋆1⊆S⁻¹Rˣ : ∀ s → s ∈ S → s ⋆ 1r ∈ S⁻¹Rˣ + S⋆1⊆S⁻¹Rˣ s s∈S' = subst-∈ S⁻¹Rˣ + (cong [_] (≡-× (sym (·IdR s)) (Σ≡Prop (λ x → S x .snd) (sym (·IdR _))))) + (S/1⊆S⁻¹Rˣ s s∈S') + S⁻¹RHasAlgUniversalProp : hasLocAlgUniversalProp S⁻¹RAsCommAlg S⋆1⊆S⁻¹Rˣ + S⁻¹RHasAlgUniversalProp B S⋆1⊆Bˣ' = χᴬ , χᴬuniqueness + where + φ = B .snd + instance + _ = CommAlgebra→CommAlgebraStr B + _ = B .fst .snd + + S⋆1⊆Bˣ : (s : fst R) → s ∈ S → fst φ s ∈ (B .fst ˣ) + S⋆1⊆Bˣ s s∈S = subst (_∈ (B .fst ˣ)) (·IdR _) (S⋆1⊆Bˣ' s s∈S) + + χ : CommRingHom S⁻¹RAsCommRing (B .fst) + χ = S⁻¹RHasUniversalProp (B .fst) φ S⋆1⊆Bˣ .fst .fst + + χcomp : ∀ r → fst χ (r /1) ≡ fst φ r + χcomp = funExt⁻ (S⁻¹RHasUniversalProp (B .fst) φ S⋆1⊆Bˣ .fst .snd) + + χᴬ : CommAlgebraHom S⁻¹RAsCommAlg B + χᴬ = CommAlgebraHomFromCommRingHom χ path + where + opaque + path : ∀ r x → fst χ ((r /1) ·ₗ x) ≡ r ⋆ (fst χ x) + path r x = fst χ ((r /1) ·ₗ x) ≡⟨ IsCommRingHom.pres· (snd χ) _ _ ⟩ + fst χ (r /1) · fst χ x ≡⟨ cong (_· fst χ x) (χcomp r) ⟩ + (B .snd .fst r) · (fst χ x) ≡⟨ cong (_· fst χ x) (sym (·IdR _)) ⟩ + (r ⋆ 1r) · fst χ x ≡⟨ ⋆AssocL _ _ _ ⟩ + r ⋆ (1r · fst χ x) ≡⟨ cong (r ⋆_) (·IdL _) ⟩ + r ⋆ (fst χ x) ∎ + + χᴬuniqueness : (ψ : CommAlgebraHom S⁻¹RAsCommAlg B) → χᴬ ≡ ψ + χᴬuniqueness ψ = + CommAlgebraHom≡ {f = χᴬ} + (χᴬ .fst .fst ≡⟨ cong (fst ∘ fst) (χuniqueness (ψ .fst , funExt ψ'r/1≡φr)) ⟩ ψ .fst .fst ∎) + + where + χuniqueness = S⁻¹RHasUniversalProp (B .fst) φ S⋆1⊆Bˣ .snd + + ψ'r/1≡φr : ∀ r → ψ .fst .fst (r /1) ≡ fst φ r + ψ'r/1≡φr r = + ψ .fst .fst (r /1) ≡⟨ cong (ψ .fst .fst) (sym (·ₗ-rid _)) ⟩ + ψ .fst .fst (r ⋆ 1r) ≡⟨ IsCommAlgebraHom.pres⋆ ψ _ _ ⟩ + r ⋆ (ψ .fst .fst 1r) ≡⟨ cong (λ u → r ⋆ u) (pres1 (ψ .fst .snd)) ⟩ + r ⋆ 1r ≡⟨ solve! (B .fst) ⟩ + fst φ r ∎ + + + -- an immediate corollary: + isContrHomS⁻¹RS⁻¹R : isContr (CommAlgebraHom S⁻¹RAsCommAlg S⁻¹RAsCommAlg) + isContrHomS⁻¹RS⁻¹R = S⁻¹RHasAlgUniversalProp S⁻¹RAsCommAlg S⋆1⊆S⁻¹Rˣ + + S⁻¹RAlgCharEquiv : (A : CommRing ℓ) (φ : CommRingHom R A) + → PathToS⁻¹R R S SMultClosedSubset A φ + → CommAlgebraEquiv S⁻¹RAsCommAlg (A , φ) + S⁻¹RAlgCharEquiv A φ cond .fst = S⁻¹RCharEquiv R S SMultClosedSubset A φ cond + S⁻¹RAlgCharEquiv A φ cond .snd = CommRingHom≡ (S⁻¹RHasUniversalProp A φ (cond .φS⊆Aˣ) .fst .snd) + where open PathToS⁻¹R + +-- the special case of localizing at a single element +R[1/_]AsCommAlgebra : {R : CommRing ℓ} → fst R → CommAlgebra R ℓ +R[1/_]AsCommAlgebra {R = R} f = S⁻¹RAsCommAlg [ f ⁿ|n≥0] (powersFormMultClosedSubset f) + where + open AlgLoc R + open InvertingElementsBase R + +module _ {R : CommRing ℓ} (f : fst R) where + open InvertingElementsBase R + open AlgLoc R [ f ⁿ|n≥0] (powersFormMultClosedSubset f) + + invElCommAlgebra→CommRingPath : CommAlgebra→CommRing R[1/ f ]AsCommAlgebra ≡ R[1/ f ]AsCommRing + invElCommAlgebra→CommRingPath = refl + +module AlgLocTwoSubsets (R' : CommRing ℓ) + (S₁ : ℙ (fst R')) (S₁MultClosedSubset : isMultClosedSubset R' S₁) + (S₂ : ℙ (fst R')) (S₂MultClosedSubset : isMultClosedSubset R' S₂) where + open isMultClosedSubset + open RingTheory (CommRing→Ring R') + open Loc R' S₁ S₁MultClosedSubset renaming (S⁻¹R to S₁⁻¹R ; + S⁻¹RAsCommRing to S₁⁻¹RAsCommRing) + open Loc R' S₂ S₂MultClosedSubset renaming (S⁻¹R to S₂⁻¹R ; + S⁻¹RAsCommRing to S₂⁻¹RAsCommRing) + open AlgLoc R' S₁ S₁MultClosedSubset renaming ( S⁻¹RAsCommAlg to S₁⁻¹RAsCommAlg + ; S⋆1⊆S⁻¹Rˣ to S₁⋆1⊆S₁⁻¹Rˣ + ; S⁻¹RHasAlgUniversalProp to S₁⁻¹RHasAlgUniversalProp + ; isContrHomS⁻¹RS⁻¹R to isContrHomS₁⁻¹RS₁⁻¹R) + open AlgLoc R' S₂ S₂MultClosedSubset renaming ( S⁻¹RAsCommAlg to S₂⁻¹RAsCommAlg + ; S⋆1⊆S⁻¹Rˣ to S₂⋆1⊆S₂⁻¹Rˣ + ; S⁻¹RHasAlgUniversalProp to S₂⁻¹RHasAlgUniversalProp + ; isContrHomS⁻¹RS⁻¹R to isContrHomS₂⁻¹RS₂⁻¹R) + + open CommAlgebraStr ⦃...⦄ + open CommRingStr ⦃...⦄ + + private + S₁⁻¹Rˣ = S₁⁻¹RAsCommRing ˣ + S₂⁻¹Rˣ = S₂⁻¹RAsCommRing ˣ + instance + _ = CommAlgebra→CommAlgebraStr S₁⁻¹RAsCommAlg + _ = CommAlgebra→CommAlgebraStr S₂⁻¹RAsCommAlg + _ = S₂⁻¹RAsCommAlg .fst .snd + _ = S₁⁻¹RAsCommAlg .fst .snd + _ = R' .snd + + + isContrS₁⁻¹R≡S₂⁻¹R : (∀ (s₁ : ⟨ R' ⟩) → s₁ ∈ S₁ → s₁ ⋆ 1r ∈ S₂⁻¹Rˣ) + → (∀ s₂ → s₂ ∈ S₂ → s₂ ⋆ 1r ∈ S₁⁻¹Rˣ) + → isContr (S₁⁻¹RAsCommAlg ≡ S₂⁻¹RAsCommAlg) + isContrS₁⁻¹R≡S₂⁻¹R S₁⊆S₂⁻¹Rˣ S₂⊆S₁⁻¹Rˣ = isOfHLevelRetractFromIso 0 + (equivToIso (invEquiv (CommAlgebraPath _ _ _))) + isContrS₁⁻¹R≅S₂⁻¹R + where + χ₁ : CommAlgebraHom S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg + χ₁ = S₁⁻¹RHasAlgUniversalProp S₂⁻¹RAsCommAlg S₁⊆S₂⁻¹Rˣ .fst + + χ₂ : CommAlgebraHom S₂⁻¹RAsCommAlg S₁⁻¹RAsCommAlg + χ₂ = S₂⁻¹RHasAlgUniversalProp S₁⁻¹RAsCommAlg S₂⊆S₁⁻¹Rˣ .fst + + χ₁∘χ₂≡id : χ₁ ∘ca χ₂ ≡ idCAlgHom _ + χ₁∘χ₂≡id = isContr→isProp isContrHomS₂⁻¹RS₂⁻¹R _ _ + + χ₂∘χ₁≡id : χ₂ ∘ca χ₁ ≡ idCAlgHom _ + χ₂∘χ₁≡id = isContr→isProp isContrHomS₁⁻¹RS₁⁻¹R _ _ + + IsoS₁⁻¹RS₂⁻¹R : Iso S₁⁻¹R S₂⁻¹R + Iso.fun IsoS₁⁻¹RS₂⁻¹R = ⟨ χ₁ ⟩ₐ→ + Iso.inv IsoS₁⁻¹RS₂⁻¹R = ⟨ χ₂ ⟩ₐ→ + Iso.rightInv IsoS₁⁻¹RS₂⁻¹R = funExt⁻ (cong ⟨_⟩ₐ→ χ₁∘χ₂≡id) + Iso.leftInv IsoS₁⁻¹RS₂⁻¹R = funExt⁻ (cong ⟨_⟩ₐ→ χ₂∘χ₁≡id) + + isContrS₁⁻¹R≅S₂⁻¹R : isContr (CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg) + isContrS₁⁻¹R≅S₂⁻¹R = withOpaqueStr $ center , uniqueness + where + X₁asCRHom = fst χ₁ + center : CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg + fst center = withOpaqueStr $ (isoToEquiv IsoS₁⁻¹RS₂⁻¹R) , snd X₁asCRHom + snd center = makeOpaque $ CommRingHom≡ (cong fst (χ₁ .snd)) + + uniqueness : (φ : CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg) → center ≡ φ + uniqueness φ = + CommAlgebraEquiv≡ (cong ⟨_⟩ₐ→ $ S₁⁻¹RHasAlgUniversalProp S₂⁻¹RAsCommAlg S₁⊆S₂⁻¹Rˣ .snd + (CommAlgebraEquiv→CommAlgebraHom φ)) + + + opaque + isPropS₁⁻¹R≡S₂⁻¹R : isProp (S₁⁻¹RAsCommAlg ≡ S₂⁻¹RAsCommAlg) + isPropS₁⁻¹R≡S₂⁻¹R S₁⁻¹R≡S₂⁻¹R = + isContr→isProp (isContrS₁⁻¹R≡S₂⁻¹R S₁⊆S₂⁻¹Rˣ S₂⊆S₁⁻¹Rˣ) S₁⁻¹R≡S₂⁻¹R + where + + S₁⊆S₂⁻¹Rˣ : ∀ s₁ → s₁ ∈ S₁ → s₁ ⋆ 1r ∈ S₂⁻¹Rˣ + S₁⊆S₂⁻¹Rˣ s₁ s₁∈S₁ = + transport (λ i → let instance + _ = (S₁⁻¹R≡S₂⁻¹R i) .fst .snd + _ = CommAlgebra→CommAlgebraStr (S₁⁻¹R≡S₂⁻¹R i) + in s₁ ⋆ 1r ∈ (CommAlgebra→CommRing (S₁⁻¹R≡S₂⁻¹R i)) ˣ) (S₁⋆1⊆S₁⁻¹Rˣ s₁ s₁∈S₁) + + S₂⊆S₁⁻¹Rˣ : ∀ s₂ → s₂ ∈ S₂ → s₂ ⋆ 1r ∈ S₁⁻¹Rˣ + S₂⊆S₁⁻¹Rˣ s₂ s₂∈S₂ = + transport (λ i → let instance + _ = ((sym S₁⁻¹R≡S₂⁻¹R) i) .fst .snd + _ = CommAlgebra→CommAlgebraStr ((sym S₁⁻¹R≡S₂⁻¹R) i) + in s₂ ⋆ 1r ∈ (CommAlgebra→CommRing ((sym S₁⁻¹R≡S₂⁻¹R) i)) ˣ) (S₂⋆1⊆S₂⁻¹Rˣ s₂ s₂∈S₂) + + +-- A crucial result for the construction of the structure sheaf +module DoubleAlgLoc (R : CommRing ℓ) (f g : (fst R)) where + open Exponentiation R + open InvertingElementsBase + open CommRingStr (snd R) hiding (·IdR) + open isMultClosedSubset + open DoubleLoc R f g hiding (R[1/fg]≡R[1/f][1/g]) + open AlgLoc R ([_ⁿ|n≥0] R (f · g)) (powersFormMultClosedSubset R (f · g)) + renaming (S⁻¹RAlgCharEquiv to R[1/fg]AlgCharEquiv) + open CommIdeal R hiding (subst-∈) renaming (_∈_ to _∈ᵢ_) + open isCommIdeal + open RadicalIdeal R + +{- + + private + ⟨_⟩ : (fst R) → CommIdeal + ⟨ f ⟩ = ⟨ replicateFinVec 1 f ⟩[ R ] + + R[1/fg]AsCommAlgebra = R[1/_]AsCommAlgebra {R = R} (f · g) + R[1/fg]ˣ = R[1/_]AsCommRing R (f · g) ˣ + + R[1/g]AsCommAlgebra = R[1/_]AsCommAlgebra {R = R} g + R[1/g]ˣ = R[1/_]AsCommRing R g ˣ + + R[1/f][1/g]AsCommRing = R[1/_]AsCommRing (R[1/_]AsCommRing R f) + [ g , 1r , powersFormMultClosedSubset R f .containsOne ] + R[1/f][1/g]AsCommAlgebra = toCommAlg (R[1/f][1/g]AsCommRing , /1/1AsCommRingHom) + + R[1/fg]≡R[1/f][1/g] : R[1/fg]AsCommAlgebra ≡ R[1/f][1/g]AsCommAlgebra + R[1/fg]≡R[1/f][1/g] = uaCommAlgebra (R[1/fg]AlgCharEquiv _ _ pathtoR[1/fg]) + + doubleLocCancel : g ∈ᵢ √ ⟨ f ⟩ → R[1/f][1/g]AsCommAlgebra ≡ R[1/g]AsCommAlgebra + doubleLocCancel g∈√⟨f⟩ = sym R[1/fg]≡R[1/f][1/g] ∙ isContrR[1/fg]≡R[1/g] toUnit1 toUnit2 .fst + where + open S⁻¹RUniversalProp R ([_ⁿ|n≥0] R g) (powersFormMultClosedSubset R g) + renaming (_/1 to _/1ᵍ) + open S⁻¹RUniversalProp R ([_ⁿ|n≥0] R (f · g)) (powersFormMultClosedSubset R (f · g)) + renaming (_/1 to _/1ᶠᵍ) + open AlgLocTwoSubsets R ([_ⁿ|n≥0] R (f · g)) (powersFormMultClosedSubset R (f · g)) + ([_ⁿ|n≥0] R g) (powersFormMultClosedSubset R g) + renaming (isContrS₁⁻¹R≡S₂⁻¹R to isContrR[1/fg]≡R[1/g]) + open CommAlgebraStr ⦃...⦄ hiding (_·_ ; _+_) + instance + _ = snd R[1/fg]AsCommAlgebra + _ = snd R[1/g]AsCommAlgebra + + toUnit1 : ∀ s → s ∈ [_ⁿ|n≥0] R (f · g) → s ⋆ 1a ∈ R[1/g]ˣ + toUnit1 s s∈[fgⁿ|n≥0] = subst-∈ R[1/g]ˣ (sym (·IdR (s /1ᵍ))) + (RadicalLemma.toUnit R g (f · g) (radHelper _ _ g∈√⟨f⟩) s s∈[fgⁿ|n≥0]) + where + radHelper : ∀ x y → x ∈ᵢ √ ⟨ y ⟩ → x ∈ᵢ √ ⟨ y · x ⟩ + radHelper x y = PT.rec ((√ ⟨ y · x ⟩) .fst x .snd) (uncurry helper1) + where + helper1 : (n : ℕ) → x ^ n ∈ᵢ ⟨ y ⟩ → x ∈ᵢ √ ⟨ y · x ⟩ + helper1 n = PT.rec ((√ ⟨ y · x ⟩) .fst x .snd) (uncurry helper2) + where + helper2 : (α : FinVec (fst R) 1) + → x ^ n ≡ linearCombination R α (replicateFinVec 1 y) + → x ∈ᵢ √ ⟨ y · x ⟩ + helper2 α p = ∣ (suc n) , ∣ α , cong (x ·_) p ∙ solve! R ∣₁ ∣₁ + + toUnit2 : ∀ s → s ∈ [_ⁿ|n≥0] R g → s ⋆ 1a ∈ R[1/fg]ˣ + toUnit2 s s∈[gⁿ|n≥0] = subst-∈ R[1/fg]ˣ (sym (·IdR (s /1ᶠᵍ))) + (RadicalLemma.toUnit R (f · g) g radHelper s s∈[gⁿ|n≥0]) + where + radHelper : (f · g) ∈ᵢ √ ⟨ g ⟩ + radHelper = ·Closed (snd (√ ⟨ g ⟩)) f (∈→∈√ _ _ (indInIdeal R _ zero)) +-} +-- -} +-- -} diff --git a/Cubical/Algebra/CommAlgebraAlt/Univalence.agda b/Cubical/Algebra/CommAlgebraAlt/Univalence.agda index 7095418b0b..6bbde392f1 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Univalence.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Univalence.agda @@ -1,3 +1,4 @@ +{-# OPTIONS --safe #-} module Cubical.Algebra.CommAlgebraAlt.Univalence where open import Cubical.Foundations.Prelude diff --git a/Cubical/Algebra/CommRing/Localisation/Base.agda b/Cubical/Algebra/CommRing/Localisation/Base.agda index a6123ebdae..e74d3fd47c 100644 --- a/Cubical/Algebra/CommRing/Localisation/Base.agda +++ b/Cubical/Algebra/CommRing/Localisation/Base.agda @@ -44,6 +44,7 @@ private -- A multiplicatively closed subset is assumed to contain 1 record isMultClosedSubset (R' : CommRing ℓ) (S' : ℙ (fst R')) : Type ℓ where + no-eta-equality constructor multclosedsubset field diff --git a/Cubical/Foundations/Structure.agda b/Cubical/Foundations/Structure.agda index 8d143f0b73..9374d3190b 100644 --- a/Cubical/Foundations/Structure.agda +++ b/Cubical/Foundations/Structure.agda @@ -16,18 +16,23 @@ TypeWithStr : (ℓ : Level) (S : Type ℓ → Type ℓ') → Type (ℓ-max (ℓ- TypeWithStr ℓ S = Σ[ X ∈ Type ℓ ] S X {- - A helper to make the second component opaque - it is a good idea to use this for set-level structures, + A helper to make a definition or the second component of a type with structure + opaque. It is a good idea to use this for set-level structures, where the second component consists only of proofs of identities or other propositions -} +makeOpaque : {A : Type ℓ'} → A → A +makeOpaque {A = A} a = a' + where + opaque + a' : A + a' = a + withOpaqueStr : {A : Type ℓ'} {B : A → Type ℓ''} → Σ[ x ∈ A ] B x → Σ[ x ∈ A ] B x -withOpaqueStr {B = B} (x , str) = x , str' - where opaque - str' : B x - str' = str +withOpaqueStr {B = B} (x , str) = x , makeOpaque str + typ : TypeWithStr ℓ S → Type ℓ typ = fst From e9f061690ccaa313cb141a760c1cd4eacc62280a Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Sat, 3 Aug 2024 20:12:09 +0200 Subject: [PATCH 52/83] make ring quotients' isRing opaque instead of no-eta --- Cubical/Algebra/Ring/Base.agda | 1 - Cubical/Algebra/Ring/Quotient.agda | 25 ++++++++++++++++++------- 2 files changed, 18 insertions(+), 8 deletions(-) diff --git a/Cubical/Algebra/Ring/Base.agda b/Cubical/Algebra/Ring/Base.agda index bd22398f72..3d4b2a0143 100644 --- a/Cubical/Algebra/Ring/Base.agda +++ b/Cubical/Algebra/Ring/Base.agda @@ -34,7 +34,6 @@ private record IsRing {R : Type ℓ} (0r 1r : R) (_+_ _·_ : R → R → R) (-_ : R → R) : Type ℓ where - no-eta-equality constructor isring field diff --git a/Cubical/Algebra/Ring/Quotient.agda b/Cubical/Algebra/Ring/Quotient.agda index 188b3a96b1..5558b59fa6 100644 --- a/Cubical/Algebra/Ring/Quotient.agda +++ b/Cubical/Algebra/Ring/Quotient.agda @@ -185,10 +185,20 @@ module _ (R' : Ring ℓ) (I : ⟨ R' ⟩ → hProp ℓ) (I-isIdeal : isIdeal R' eq : (x y z : R) → [ x ] ·/I ([ y ] +/I [ z ]) ≡ ([ x ] ·/I [ y ]) +/I ([ x ] ·/I [ z ]) eq x y z i = [ ·DistR+ x y z i ] + quotientRingStr' : RingStr R/I + quotientRingStr' = snd (makeRing 0/I 1/I _+/I_ _·/I_ -/I isSetR/I + +/I-assoc +/I-rid +/I-rinv +/I-comm + ·/I-assoc ·/I-rid ·/I-lid /I-rdist /I-ldist) + + opaque + isRingQuotient : IsRing 0/I 1/I _+/I_ _·/I_ -/I + isRingQuotient = RingStr.isRing quotientRingStr' + quotientRingStr : RingStr R/I - quotientRingStr = snd (makeRing 0/I 1/I _+/I_ _·/I_ -/I isSetR/I - +/I-assoc +/I-rid +/I-rinv +/I-comm - ·/I-assoc ·/I-rid ·/I-lid /I-rdist /I-ldist) + quotientRingStr = + record + { RingStr quotientRingStr'; isRing = isRingQuotient } + _/_ : (R : Ring ℓ) → (I : IdealsIn R) → Ring ℓ fst (R / I) = R/I R (fst I) (snd I) @@ -358,7 +368,8 @@ module idealIsKernel {R : Ring ℓ} (I : IdealsIn R) where x-0∈I : x - 0r ∈ fst I x-0∈I = effective ~IsPropValued ~IsEquivRel x 0r x∈ker -kernel≡I : {R : Ring ℓ} (I : IdealsIn R) - → kernelIdeal (quotientHom R I) ≡ I -kernel≡I {R = R} I = Σ≡Prop (isPropIsIdeal R) (⊆-extensionality _ _ (ker⊆I , I⊆ker)) - where open idealIsKernel I +opaque + kernel≡I : {R : Ring ℓ} (I : IdealsIn R) + → kernelIdeal (quotientHom R I) ≡ I + kernel≡I {R = R} I = Σ≡Prop (isPropIsIdeal R) (⊆-extensionality _ _ (ker⊆I , I⊆ker)) + where open idealIsKernel I From 2de9d4f8be86e8ad100249fc31eab2d252b50ca7 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Sun, 4 Aug 2024 11:37:48 +0200 Subject: [PATCH 53/83] update to new interface for comm ring homs --- .../FreeCommAlgebra/OnCoproduct.agda | 10 ++++---- .../FreeCommAlgebra/Properties.agda | 2 +- Cubical/Algebra/CommAlgebra/Properties.agda | 20 +++++++--------- .../Algebra/CommAlgebra/QuotientAlgebra.agda | 1 - Cubical/Algebra/CommRing/Base.agda | 16 +++++++++++-- .../CommRing/Localisation/PullbackSquare.agda | 24 +++++++++---------- Cubical/Algebra/CommRing/Properties.agda | 16 +++++++++---- .../EquivCarac/AB-An[X]Bn[X].agda | 10 ++++---- .../Multivariate/EquivCarac/A[X]X-A.agda | 4 ++-- .../EquivCarac/An[Am[X]]-Anm[X].agda | 2 +- .../Multivariate/EquivCarac/An[X]X-A.agda | 2 +- .../Multivariate/EquivCarac/Poly0-A.agda | 2 +- .../UnivariateHIT/Polyn-nPoly.agda | 2 +- .../UnivariateList/Poly1-1Poly.agda | 2 +- Cubical/Algebra/Ring/Quotient.agda | 2 +- 15 files changed, 64 insertions(+), 51 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/OnCoproduct.agda b/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/OnCoproduct.agda index a20195fca7..61fcb8c920 100644 --- a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/OnCoproduct.agda +++ b/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/OnCoproduct.agda @@ -76,11 +76,11 @@ module CalculateFreeCommAlgebraOnCoproduct (R : CommRing ℓ) (I J : Type ℓ) w asHomOverR[I] = Iso.fun isoR[I] R[I⊎J]overR[I] asHomOverR = Iso.fun isoR (R [ I ⊎ J ]) - ≡RingHoms : snd asHomOverR[I] ∘r baseRingHom ≡ baseRingHom + ≡RingHoms : snd asHomOverR[I] ∘cr baseRingHom ≡ baseRingHom ≡RingHoms = - RingHom≡ + CommRingHom≡ (funExt λ x → - fst (snd asHomOverR[I] ∘r baseRingHom) x ≡⟨⟩ + fst (snd asHomOverR[I] ∘cr baseRingHom) x ≡⟨⟩ fst (snd asHomOverR[I]) (const x · 1a) ≡⟨⟩ (const x · 1a) ⋆ 1a ≡⟨ cong (_⋆ 1a) (·IdR (const x)) ⟩ const x ⋆ 1a ≡⟨⟩ @@ -90,12 +90,12 @@ module CalculateFreeCommAlgebraOnCoproduct (R : CommRing ℓ) (I J : Type ℓ) w ≡R[I⊎J] = baseChange baseRingHom R[I⊎J]overR[I] ≡⟨⟩ - Iso.inv isoR ((CommAlgebra→CommRing R[I⊎J]overR[I]) , (snd asHomOverR[I]) ∘r baseRingHom) ≡⟨ step1 ⟩ + Iso.inv isoR ((CommAlgebra→CommRing R[I⊎J]overR[I]) , (snd asHomOverR[I]) ∘cr baseRingHom) ≡⟨ step1 ⟩ Iso.inv isoR (CommAlgebra→CommRing (R [ I ⊎ J ]) , baseRingHom) ≡⟨⟩ Iso.inv isoR asHomOverR ≡⟨ step2 ⟩ R [ I ⊎ J ] ∎ where - step1 : Iso.inv isoR ((CommAlgebra→CommRing R[I⊎J]overR[I]) , (snd asHomOverR[I]) ∘r baseRingHom) + step1 : Iso.inv isoR ((CommAlgebra→CommRing R[I⊎J]overR[I]) , (snd asHomOverR[I]) ∘cr baseRingHom) ≡ Iso.inv isoR (CommAlgebra→CommRing (R [ I ⊎ J ]) , baseRingHom) step1 i = Iso.inv isoR ((CommAlgebra→CommRing R[I⊎J]overR[I]) , ≡RingHoms i) diff --git a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Properties.agda index a7dea81c7c..76c462ee48 100644 --- a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Properties.agda @@ -30,8 +30,8 @@ open import Cubical.Data.Sigma.Properties using (Σ≡Prop) open import Cubical.HITs.SetTruncation open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommAlgebra.FreeCommAlgebra.Base open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebra.FreeCommAlgebra.Base open import Cubical.Algebra.CommAlgebra.Instances.Initial open import Cubical.Algebra.Algebra open import Cubical.Algebra.Module using (module ModuleTheory) diff --git a/Cubical/Algebra/CommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebra/Properties.agda index ededbddde3..feb73d2da4 100644 --- a/Cubical/Algebra/CommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/Properties.agda @@ -108,7 +108,7 @@ module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where isCommRing (snd (CommAlgebra→CommRing≡ i)) = isProp→PathP (λ i → isPropIsCommRing _ _ _ _ _ ) (isCommRing (snd (CommAlgebra→CommRing (toCommAlg Aφ)))) (isCommRing (snd A)) i -{- + CommRingWithHomRoundTrip : (Aφ : CommRingWithHom) → fromCommAlg (toCommAlg Aφ) ≡ Aφ CommRingWithHomRoundTrip (A , φ) = ΣPathP (CommAlgebra→CommRing≡ (A , φ) , φPathP) where @@ -116,8 +116,8 @@ module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where -- this only works because fst (CommAlgebra→CommRing≡ (A , φ) i) = fst A definitionally! φPathP : PathP (λ i → CommRingHom R (CommAlgebra→CommRing≡ (A , φ) i)) (snd (fromCommAlg (toCommAlg (A , φ)))) φ - φPathP = RingHomPathP _ _ _ _ _ _ λ i x → ·IdR (snd A) (fst φ x) i --} + φPathP = CommRingHomPathP _ _ _ _ _ _ λ i x → ·IdR (snd A) (fst φ x) i + CommAlgRoundTrip : (A : CommAlgebra R ℓ') → toCommAlg (fromCommAlg A) ≡ A CommAlgRoundTrip A = ΣPathP (refl , AlgStrPathP) @@ -136,13 +136,11 @@ module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where (λ i → isPropIsCommAlgebra _ _ _ _ _ _ (CommAlgebraStr._⋆_ (AlgStrPathP i))) (CommAlgebraStr.isCommAlgebra (snd (toCommAlg (fromCommAlg A)))) isCommAlgebra i -{- CommAlgIso : Iso (CommAlgebra R ℓ') CommRingWithHom fun CommAlgIso = fromCommAlg inv CommAlgIso = toCommAlg rightInv CommAlgIso = CommRingWithHomRoundTrip leftInv CommAlgIso = CommAlgRoundTrip --} open IsCommRingHom @@ -340,26 +338,24 @@ module _ {R : CommRing ℓ} {A B : CommAlgebra R ℓ'} where pres- (snd (CommAlgebraHomFromRingHom ϕ pres*)) = IsRingHom.pres- (snd ϕ) pres⋆ (snd (CommAlgebraHomFromRingHom ϕ pres*)) = pres* -{- module _ {R S : CommRing ℓ} (f : CommRingHom S R) where baseChange : CommAlgebra R ℓ' → CommAlgebra S ℓ' baseChange A = - Iso.inv (CommAlgChar.CommAlgIso S) (fst asRingHom , compCommRingHom _ _ _ f (snd asRingHom)) + Iso.inv (CommAlgChar.CommAlgIso S) (fst asRingHom , compCommRingHom f (snd asRingHom)) where asRingHom : CommAlgChar.CommRingWithHom R asRingHom = Iso.fun (CommAlgChar.CommAlgIso R) A baseChangeHom : (A B : CommAlgebra R ℓ') → CommAlgebraHom A B → CommAlgebraHom (baseChange A) (baseChange B) baseChangeHom A B ϕ = - CommAlgChar.toCommAlgebraHom S (fst homA , snd homA ∘r f) (fst homB , snd homB ∘r f) (fst pbSliceHom) (snd pbSliceHom) + CommAlgChar.toCommAlgebraHom S (fst homA , snd homA ∘cr f) (fst homB , snd homB ∘cr f) (fst pbSliceHom) (snd pbSliceHom) where open RingHoms homA = Iso.fun (CommAlgChar.CommAlgIso R) A homB = Iso.fun (CommAlgChar.CommAlgIso R) B - asSliceHom : Σ[ h ∈ CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) ] h ∘r (snd homA) ≡ snd homB + asSliceHom : Σ[ h ∈ CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) ] h ∘cr (snd homA) ≡ snd homB asSliceHom = CommAlgChar.fromCommAlgebraHom R A B ϕ pbSliceHom : Σ[ k ∈ CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) ] - k ∘r ((snd homA) ∘r f) ≡ ((snd homB) ∘r f) - pbSliceHom = fst asSliceHom , Σ≡Prop (λ _ → isPropIsRingHom _ _ _) λ i x → fst ((snd asSliceHom) i) (fst f x) --} + k ∘cr ((snd homA) ∘cr f) ≡ ((snd homB) ∘cr f) + pbSliceHom = fst asSliceHom , CommRingHom≡ λ i x → fst ((snd asSliceHom) i) (fst f x) diff --git a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda index 3199f0e951..3390614d92 100644 --- a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda @@ -47,7 +47,6 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn A) where _ = snd A opaque - unfolding CommRing.quotientCommRingStr _/_ : CommAlgebra R ℓ _/_ = commAlgebraFromCommRing A/IAsCommRing diff --git a/Cubical/Algebra/CommRing/Base.agda b/Cubical/Algebra/CommRing/Base.agda index c2786d580d..b4441b63de 100644 --- a/Cubical/Algebra/CommRing/Base.agda +++ b/Cubical/Algebra/CommRing/Base.agda @@ -3,6 +3,7 @@ module Cubical.Algebra.CommRing.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function using (_$_) open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels open import Cubical.Foundations.SIP @@ -220,6 +221,11 @@ IsCommRingHom→IsRingHom : (R : CommRing ℓ) (S : CommRing ℓ') → IsRingHom ((CommRing→Ring R) .snd) f ((CommRing→Ring S) .snd) IsCommRingHom→IsRingHom R S f p = CommRingHom→RingHom (f , p) .snd +CommRingEquiv→RingEquiv : {A : CommRing ℓ} {B : CommRing ℓ'} + → CommRingEquiv A B → RingEquiv (CommRing→Ring A) (CommRing→Ring B) +CommRingEquiv→RingEquiv e .fst = e .fst +CommRingEquiv→RingEquiv e .snd = IsCommRingHom→IsRingHom _ _ (e .fst .fst) (e .snd) + module _ {R : CommRing ℓ} {S : CommRing ℓ'} {f : ⟨ R ⟩ → ⟨ S ⟩} where private @@ -238,5 +244,11 @@ module _ {R : CommRing ℓ} {S : CommRing ℓ'} {f : ⟨ R ⟩ → ⟨ S ⟩} wh _$cr_ : {R : CommRing ℓ} {S : CommRing ℓ'} → (φ : CommRingHom R S) → (x : ⟨ R ⟩) → ⟨ S ⟩ φ $cr x = φ .fst x -CommRingHom≡ : {R : CommRing ℓ} {S : CommRing ℓ'} {φ ψ : CommRingHom R S} → fst φ ≡ fst ψ → φ ≡ ψ -CommRingHom≡ = Σ≡Prop λ f → isPropIsCommRingHom _ f _ +opaque + CommRingHom≡ : {R : CommRing ℓ} {S : CommRing ℓ'} {φ ψ : CommRingHom R S} → fst φ ≡ fst ψ → φ ≡ ψ + CommRingHom≡ = Σ≡Prop λ f → isPropIsCommRingHom _ f _ + + CommRingHomPathP : (R : CommRing ℓ) (S T : CommRing ℓ') (p : S ≡ T) (φ : CommRingHom R S) (ψ : CommRingHom R T) + → PathP (λ i → R .fst → p i .fst) (φ .fst) (ψ .fst) + → PathP (λ i → CommRingHom R (p i)) φ ψ + CommRingHomPathP R S T p φ ψ q = ΣPathP (q , isProp→PathP (λ _ → isPropIsCommRingHom _ _ _) _ _) diff --git a/Cubical/Algebra/CommRing/Localisation/PullbackSquare.agda b/Cubical/Algebra/CommRing/Localisation/PullbackSquare.agda index 231223071b..c608e3d896 100644 --- a/Cubical/Algebra/CommRing/Localisation/PullbackSquare.agda +++ b/Cubical/Algebra/CommRing/Localisation/PullbackSquare.agda @@ -73,7 +73,7 @@ private module _ (R' : CommRing ℓ) (f g : (fst R')) where - open IsRingHom + open IsCommRingHom open isMultClosedSubset open CommRingTheory R' open RingTheory (CommRing→Ring R') @@ -490,7 +490,7 @@ module _ (R' : CommRing ℓ) (f g : (fst R')) where /1χComm x = eq/ _ _ ((1r , powersFormMultClosedSubset (f · g) .containsOne) , refl) /1χHomComm : /1ᵍAsCommRingHom ⋆ χ₂ ≡ /1ᶠAsCommRingHom ⋆ χ₁ - /1χHomComm = RingHom≡ (funExt /1χComm) + /1χHomComm = CommRingHom≡ (funExt /1χComm) fgSquare : 1r ∈ ⟨f,g⟩ → isPullback _ fgCospan /1ᵍAsCommRingHom /1ᶠAsCommRingHom /1χHomComm @@ -501,16 +501,16 @@ module _ (R' : CommRing ℓ) (f g : (fst R')) where applyEqualizerLemma : ∀ a → ∃![ χa ∈ R ] (χa /1ᶠ ≡ fst φ a) × (χa /1ᵍ ≡ fst ψ a) applyEqualizerLemma a = - equalizerLemma 1∈⟨f,g⟩ (fst φ a) (fst ψ a) (cong (_$r a) (sym ψχ₂≡φχ₁)) + equalizerLemma 1∈⟨f,g⟩ (fst φ a) (fst ψ a) (cong (_$cr a) (sym ψχ₂≡φχ₁)) χ : CommRingHom A R' fst χ a = applyEqualizerLemma a .fst .fst - snd χ = makeIsRingHom + snd χ = makeIsCommRingHom (cong fst (applyEqualizerLemma 1r .snd (1r , 1Coh))) (λ x y → cong fst (applyEqualizerLemma (x + y) .snd (_ , +Coh x y))) (λ x y → cong fst (applyEqualizerLemma (x · y) .snd (_ , ·Coh x y))) where - open IsRingHom + open IsCommRingHom 1Coh : (1r /1ᶠ ≡ fst φ 1r) × (1r /1ᵍ ≡ fst ψ 1r) 1Coh = (sym (φ .snd .pres1)) , sym (ψ .snd .pres1) @@ -538,19 +538,19 @@ module _ (R' : CommRing ℓ) (f g : (fst R')) where χCoh : (ψ ≡ χ ⋆ /1ᵍAsCommRingHom) × (φ ≡ χ ⋆ /1ᶠAsCommRingHom) - fst χCoh = RingHom≡ (funExt (λ a → sym (applyEqualizerLemma a .fst .snd .snd))) - snd χCoh = RingHom≡ (funExt (λ a → sym (applyEqualizerLemma a .fst .snd .fst))) + fst χCoh = CommRingHom≡ (funExt (λ a → sym (applyEqualizerLemma a .fst .snd .snd))) + snd χCoh = CommRingHom≡ (funExt (λ a → sym (applyEqualizerLemma a .fst .snd .fst))) χUniqueness : (y : Σ[ θ ∈ CommRingHom A R' ] (ψ ≡ θ ⋆ /1ᵍAsCommRingHom) × (φ ≡ θ ⋆ /1ᶠAsCommRingHom)) → (χ , χCoh) ≡ y - χUniqueness (θ , θCoh) = Σ≡Prop (λ _ → isProp× (isSetRingHom _ _ _ _) - (isSetRingHom _ _ _ _)) - (RingHom≡ (funExt (λ a → cong fst (applyEqualizerLemma a .snd (θtriple a))))) + χUniqueness (θ , θCoh) = Σ≡Prop (λ _ → isProp× (isSetCommRingHom _ _ _ _) + (isSetCommRingHom _ _ _ _)) + (CommRingHom≡ (funExt (λ a → cong fst (applyEqualizerLemma a .snd (θtriple a))))) where θtriple : ∀ a → Σ[ x ∈ R ] (x /1ᶠ ≡ fst φ a) × (x /1ᵍ ≡ fst ψ a) - θtriple a = fst θ a , sym (cong (_$r a) (θCoh .snd)) - , sym (cong (_$r a) (θCoh .fst)) + θtriple a = fst θ a , sym (cong (_$cr a) (θCoh .snd)) + , sym (cong (_$cr a) (θCoh .fst)) -- packaging it all up diff --git a/Cubical/Algebra/CommRing/Properties.agda b/Cubical/Algebra/CommRing/Properties.agda index ffda3e340f..ea3424ab15 100644 --- a/Cubical/Algebra/CommRing/Properties.agda +++ b/Cubical/Algebra/CommRing/Properties.agda @@ -168,17 +168,17 @@ module _ where compIdCommRingHom : {R S : CommRing ℓ} (f : CommRingHom R S) → compCommRingHom (idCommRingHom R) f ≡ f - compIdCommRingHom f = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) refl + compIdCommRingHom f = CommRingHom≡ refl idCompCommRingHom : {R S : CommRing ℓ} (f : CommRingHom R S) → compCommRingHom f (idCommRingHom S) ≡ f - idCompCommRingHom f = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) refl + idCompCommRingHom f = CommRingHom≡ refl compAssocCommRingHom : {R : CommRing ℓ} {S : CommRing ℓ'} {T : CommRing ℓ''} {U : CommRing ℓ'''} (f : CommRingHom R S) (g : CommRingHom S T) (h : CommRingHom T U) → compCommRingHom (compCommRingHom f g) h ≡ compCommRingHom f (compCommRingHom g h) - compAssocCommRingHom f g h = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) refl + compAssocCommRingHom f g h = CommRingHom≡ refl open Iso @@ -194,10 +194,16 @@ module CommRingEquivs where compCommRingEquiv f g .fst = compEquiv (f .fst) (g .fst) compCommRingEquiv f g .snd = compCommRingHom (f .fst .fst , f .snd) (g .fst .fst , g .snd) .snd + opaque + isCommRingHomInv : {A : CommRing ℓ} {B : CommRing ℓ'} + → (e : CommRingEquiv A B) → IsCommRingHom (snd B) (invEq (fst e)) (snd A) + isCommRingHomInv e = + IsRingHom→IsCommRingHom _ _ _ $ + isRingHomInv (e .fst , CommRingHom→RingHom (e .fst .fst , e .snd) .snd) + invCommRingEquiv : (A : CommRing ℓ) → (B : CommRing ℓ') → CommRingEquiv A B → CommRingEquiv B A fst (invCommRingEquiv A B e) = invEquiv (fst e) - snd (invCommRingEquiv A B e) = - IsRingHom→IsCommRingHom _ _ _ $ isRingHomInv (e .fst , CommRingHom→RingHom (e .fst .fst , e .snd) .snd) + snd (invCommRingEquiv A B e) = isCommRingHomInv e idCommRingEquiv : (A : CommRing ℓ) → CommRingEquiv A A fst (idCommRingEquiv A) = idEquiv (fst A) diff --git a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/AB-An[X]Bn[X].agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/AB-An[X]Bn[X].agda index c008d4d45f..bc540c77bf 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/AB-An[X]Bn[X].agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/AB-An[X]Bn[X].agda @@ -23,7 +23,7 @@ private variable ----------------------------------------------------------------------------- -- Lift -open IsRingHom +open IsCommRingHom open CommRingStr module _ @@ -46,7 +46,7 @@ module _ (+Comm (snd PB)) (λ v → (cong (base v) (pres0 fstr)) ∙ (base-neutral v)) (λ v a b → (base-add v (f a) (f b)) ∙ (cong (base v) (sym (pres+ fstr a b)))) - snd (makeCommRingHomPoly (f , fstr)) = makeIsRingHom + snd (makeCommRingHomPoly (f , fstr)) = makeIsCommRingHom (cong (base (replicate zero)) (pres1 fstr)) (λ _ _ → refl) (DS-Ind-Prop.f _ _ _ _ (λ _ → isPropΠ λ _ → trunc _ _) @@ -72,15 +72,15 @@ module _ PB = PolyCommRing B' n open Iso - open RingEquivs + open CommRingEquivs lift-equiv-poly : (e : CommRingEquiv A' B') → CommRingEquiv (PolyCommRing A' n) (PolyCommRing B' n) fst (lift-equiv-poly (e , fstr)) = isoToEquiv is where f = fst e g = invEq e - gstr : IsRingHom (snd (CommRing→Ring B')) g (snd (CommRing→Ring A')) - gstr = isRingHomInv (e , fstr) + gstr : IsCommRingHom (B' .snd) g (A' .snd) + gstr = isCommRingHomInv (e , fstr) is : Iso (fst PA) (fst PB) fun is = fst (makeCommRingHomPoly A' B' n (f , fstr)) diff --git a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/A[X]X-A.agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/A[X]X-A.agda index dc41f35a12..067a81fe05 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/A[X]X-A.agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/A[X]X-A.agda @@ -166,7 +166,7 @@ module Properties-Equiv-QuotientXn-A A[X]→A : CommRingHom A[X] Ar fst A[X]→A = A[x]→A - snd A[X]→A = makeIsRingHom A[x]→A-pres1 A[x]→A-pres+ A[x]→A-pres· + snd A[X]→A = makeIsCommRingHom A[x]→A-pres1 A[x]→A-pres+ A[x]→A-pres· A[x]→A-cancel : (k : Fin 1) → A[x]→A ( Ar 1 0 1 k) ≡ 0A A[x]→A-cancel zero = refl @@ -246,4 +246,4 @@ module _ snd Equiv-A[X]/X-A = snd A[X]/X→A Equiv-ℤ[X]/X-ℤ : RingEquiv (CommRing→Ring ℤ[X]/X) (CommRing→Ring ℤCR) -Equiv-ℤ[X]/X-ℤ = Equiv-A[X]/X-A ℤCR +Equiv-ℤ[X]/X-ℤ = CommRingEquiv→RingEquiv $ Equiv-A[X]/X-A ℤCR diff --git a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[Am[X]]-Anm[X].agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[Am[X]]-Anm[X].agda index 830cd13dd7..542cf7a241 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[Am[X]]-Anm[X].agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[Am[X]]-Anm[X].agda @@ -143,4 +143,4 @@ module _ (A' : CommRing ℓ) (n m : ℕ) where Iso.rightInv is = e-sect Iso.leftInv is = e-retr - snd CRE-PolyN∘M-PolyN+M = makeIsRingHom PAmn→PAn+m-pres1 PAmn→PAn+m-pres+ PAmn→PAn+m-pres· + snd CRE-PolyN∘M-PolyN+M = makeIsCommRingHom PAmn→PAn+m-pres1 PAmn→PAn+m-pres+ PAmn→PAn+m-pres· diff --git a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[X]X-A.agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[X]X-A.agda index 2d2f46f80f..1b28c2b7e1 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[X]X-A.agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/An[X]X-A.agda @@ -177,7 +177,7 @@ module Properties-Equiv-QuotientXn-A PAr→Ar : CommRingHom (A[X1,···,Xn] Ar n) Ar fst PAr→Ar = PA→A - snd PAr→Ar = makeIsRingHom PA→A-pres1 PA→A-pres+ PA→A-pres· + snd PAr→Ar = makeIsCommRingHom PA→A-pres1 PA→A-pres+ PA→A-pres· PAIr→Ar : CommRingHom (A[X1,···,Xn]/ Ar n) Ar PAIr→Ar = Quotient-FGideal-CommRing-CommRing.inducedHom (A[X1,···,Xn] Ar n) Ar PAr→Ar ( Ar n) PA→A-cancel diff --git a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/Poly0-A.agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/Poly0-A.agda index a11e066fb2..21588143bd 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/Poly0-A.agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/Poly0-A.agda @@ -88,4 +88,4 @@ module _ (A' : CommRing ℓ) where Iso.inv is = A→Poly0 Iso.rightInv is = e-sect Iso.leftInv is = e-retr - snd CRE-Poly0-A = makeIsRingHom Poly0→A-pres1 Poly0→A-pres+ Poly0→A-pres· + snd CRE-Poly0-A = makeIsCommRingHom Poly0→A-pres1 Poly0→A-pres+ Poly0→A-pres· diff --git a/Cubical/Algebra/Polynomials/UnivariateHIT/Polyn-nPoly.agda b/Cubical/Algebra/Polynomials/UnivariateHIT/Polyn-nPoly.agda index fc79661799..961d93510b 100644 --- a/Cubical/Algebra/Polynomials/UnivariateHIT/Polyn-nPoly.agda +++ b/Cubical/Algebra/Polynomials/UnivariateHIT/Polyn-nPoly.agda @@ -98,7 +98,7 @@ module equiv1 (A'@(A , Astr) : CommRing ℓ) inv is = convSense rightInv is = sect leftInv is = retr - snd equivR = makeIsRingHom refl (λ _ _ → refl) converseSense-pres· + snd equivR = makeIsCommRingHom refl (λ _ _ → refl) converseSense-pres· open equiv1 diff --git a/Cubical/Algebra/Polynomials/UnivariateList/Poly1-1Poly.agda b/Cubical/Algebra/Polynomials/UnivariateList/Poly1-1Poly.agda index 49570647b8..b37a071a18 100644 --- a/Cubical/Algebra/Polynomials/UnivariateList/Poly1-1Poly.agda +++ b/Cubical/Algebra/Polynomials/UnivariateList/Poly1-1Poly.agda @@ -174,7 +174,7 @@ module _ (Acr : CommRing ℓ) where Iso.inv is = Poly:→Poly1 Iso.rightInv is = e-sect Iso.leftInv is = e-retr - snd CRE-Poly1-Poly: = makeIsRingHom + snd CRE-Poly1-Poly: = makeIsCommRingHom Poly1→Poly:-pres1 Poly1→Poly:-pres+ Poly1→Poly:-pres· diff --git a/Cubical/Algebra/Ring/Quotient.agda b/Cubical/Algebra/Ring/Quotient.agda index 5558b59fa6..f98e86455f 100644 --- a/Cubical/Algebra/Ring/Quotient.agda +++ b/Cubical/Algebra/Ring/Quotient.agda @@ -197,7 +197,7 @@ module _ (R' : Ring ℓ) (I : ⟨ R' ⟩ → hProp ℓ) (I-isIdeal : isIdeal R' quotientRingStr : RingStr R/I quotientRingStr = record - { RingStr quotientRingStr'; isRing = isRingQuotient } + { isRing = isRingQuotient; RingStr quotientRingStr' } _/_ : (R : Ring ℓ) → (I : IdealsIn R) → Ring ℓ From 94770d3da1a36a1ab9fb394d380740e073283d3a Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Sun, 4 Aug 2024 12:07:05 +0200 Subject: [PATCH 54/83] update to new interface for comm ring homs --- .../Functorial/ZFunctors/OpenSubscheme.agda | 17 +++++++-------- .../StructureSheafPullback.agda | 21 +++++++++---------- 2 files changed, 17 insertions(+), 21 deletions(-) diff --git a/Cubical/AlgebraicGeometry/Functorial/ZFunctors/OpenSubscheme.agda b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/OpenSubscheme.agda index 31be9b51e5..f4909bf66e 100644 --- a/Cubical/AlgebraicGeometry/Functorial/ZFunctors/OpenSubscheme.agda +++ b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/OpenSubscheme.agda @@ -27,7 +27,6 @@ open import Cubical.Data.Nat using (ℕ) open import Cubical.Data.FinData -open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Localisation open import Cubical.Algebra.Semilattice @@ -66,8 +65,7 @@ module StandardOpens {ℓ : Level} (R : CommRing ℓ) (f : R .fst) where open isIsoC open DistLatticeStr ⦃...⦄ open CommRingStr ⦃...⦄ - open IsRingHom - open RingHoms + open IsCommRingHom open IsLatticeHom open ZarLat @@ -84,7 +82,7 @@ module StandardOpens {ℓ : Level} (R : CommRing ℓ) (f : R .fst) where D = yonedaᴾ ZarLatFun R .inv (ZL.D R f) SpR[1/f]≅⟦Df⟧ : NatIso (Sp .F-ob R[1/ f ]AsCommRing) ⟦ D ⟧ᶜᵒ - N-ob (trans SpR[1/f]≅⟦Df⟧) B φ = (φ ∘r /1AsCommRingHom) , ∨lRid _ ∙ path + N-ob (trans SpR[1/f]≅⟦Df⟧) B φ = (φ ∘cr /1AsCommRingHom) , ∨lRid _ ∙ path where open CommRingHomTheory φ open IsSupport (ZL.isSupportD B) @@ -98,18 +96,18 @@ module StandardOpens {ℓ : Level} (R : CommRing ℓ) (f : R .fst) where path : ZL.D B (φ .fst (f /1)) ≡ 1l path = supportUnit _ isUnitφ[f/1] - N-hom (trans SpR[1/f]≅⟦Df⟧) _ = funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) (RingHom≡ refl) + N-hom (trans SpR[1/f]≅⟦Df⟧) _ = funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) (CommRingHom≡ refl) inv (nIso SpR[1/f]≅⟦Df⟧ B) (φ , Dφf≡D1) = invElemUniversalProp B φ isUnitφf .fst .fst where instance _ = ZariskiLattice B .snd isUnitφf : φ .fst f ∈ B ˣ - isUnitφf = unitLemmaZarLat B (φ $r f) (sym (∨lRid _) ∙ Dφf≡D1) + isUnitφf = unitLemmaZarLat B (φ $cr f) (sym (∨lRid _) ∙ Dφf≡D1) sec (nIso SpR[1/f]≅⟦Df⟧ B) = - funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) (RingHom≡ (invElemUniversalProp _ _ _ .fst .snd)) + funExt λ _ → Σ≡Prop (λ _ → squash/ _ _) (CommRingHom≡ (invElemUniversalProp _ _ _ .fst .snd)) ret (nIso SpR[1/f]≅⟦Df⟧ B) = - funExt λ φ → cong fst (invElemUniversalProp B (φ ∘r /1AsCommRingHom) _ .snd (φ , refl)) + funExt λ φ → cong fst (invElemUniversalProp B (φ ∘cr /1AsCommRingHom) _ .snd (φ , refl)) isAffineD : isAffineCompactOpen D isAffineD = ∣ R[1/ f ]AsCommRing , SpR[1/f]≅⟦Df⟧ ∣₁ @@ -128,8 +126,7 @@ module _ {ℓ : Level} (R : CommRing ℓ) (W : CompactOpen (Sp ⟅ R ⟆)) where open DistLatticeStr ⦃...⦄ open CommRingStr ⦃...⦄ open PosetStr ⦃...⦄ - open IsRingHom - open RingHoms + open IsCommRingHom open IsLatticeHom open ZarLat diff --git a/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheafPullback.agda b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheafPullback.agda index 41695c786a..80ac11afc6 100644 --- a/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheafPullback.agda +++ b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheafPullback.agda @@ -36,7 +36,6 @@ open import Cubical.Relation.Nullary open import Cubical.Relation.Binary open import Cubical.Relation.Binary.Order.Poset -open import Cubical.Algebra.Ring open import Cubical.Algebra.Ring.Properties open import Cubical.Algebra.Ring.BigOps open import Cubical.Algebra.Algebra @@ -358,7 +357,7 @@ module _ (R' : CommRing ℓ) where R[1/h][1/fg]AsCommRing = InvertingElementsBase.R[1/_]AsCommRing R[1/ h ]AsCommRing ((f /1) · (g /1)) - open IsRingHom + open IsCommRingHom /1/1AsCommRingHomFG : CommRingHom R' R[1/h][1/fg]AsCommRing fst /1/1AsCommRingHomFG r = [ [ r , 1r , ∣ 0 , refl ∣₁ ] , 1r , ∣ 0 , refl ∣₁ ] pres0 (snd /1/1AsCommRingHomFG) = refl @@ -375,20 +374,20 @@ module _ (R' : CommRing ℓ) where open Cospan open Pullback open RingHoms - isRHomR[1/h]→R[1/h][1/f] : theRingPullback .pbPr₂ ∘r /1AsCommRingHom ≡ /1/1AsCommRingHom f - isRHomR[1/h]→R[1/h][1/f] = RingHom≡ (funExt (λ x → refl)) + isRHomR[1/h]→R[1/h][1/f] : theRingPullback .pbPr₂ ∘cr /1AsCommRingHom ≡ /1/1AsCommRingHom f + isRHomR[1/h]→R[1/h][1/f] = CommRingHom≡ (funExt (λ x → refl)) - isRHomR[1/h]→R[1/h][1/g] : theRingPullback .pbPr₁ ∘r /1AsCommRingHom ≡ /1/1AsCommRingHom g - isRHomR[1/h]→R[1/h][1/g] = RingHom≡ (funExt (λ x → refl)) + isRHomR[1/h]→R[1/h][1/g] : theRingPullback .pbPr₁ ∘cr /1AsCommRingHom ≡ /1/1AsCommRingHom g + isRHomR[1/h]→R[1/h][1/g] = CommRingHom≡ (funExt (λ x → refl)) - isRHomR[1/h][1/f]→R[1/h][1/fg] : theRingCospan .s₂ ∘r /1/1AsCommRingHom f ≡ /1/1AsCommRingHomFG - isRHomR[1/h][1/f]→R[1/h][1/fg] = RingHom≡ (funExt + isRHomR[1/h][1/f]→R[1/h][1/fg] : theRingCospan .s₂ ∘cr /1/1AsCommRingHom f ≡ /1/1AsCommRingHomFG + isRHomR[1/h][1/f]→R[1/h][1/fg] = CommRingHom≡ (funExt (λ x → cong [_] (≡-× (cong [_] (≡-× (cong (x ·_) (transportRefl 1r) ∙ ·IdR x) (Σ≡Prop (λ _ → isPropPropTrunc) (cong (1r ·_) (transportRefl 1r) ∙ ·IdR 1r)))) (Σ≡Prop (λ _ → isPropPropTrunc) (cong (1r ·_) (transportRefl 1r) ∙ ·IdR 1r))))) - isRHomR[1/h][1/g]→R[1/h][1/fg] : theRingCospan .s₁ ∘r /1/1AsCommRingHom g ≡ /1/1AsCommRingHomFG - isRHomR[1/h][1/g]→R[1/h][1/fg] = RingHom≡ (funExt + isRHomR[1/h][1/g]→R[1/h][1/fg] : theRingCospan .s₁ ∘cr /1/1AsCommRingHom g ≡ /1/1AsCommRingHomFG + isRHomR[1/h][1/g]→R[1/h][1/fg] = CommRingHom≡ (funExt (λ x → cong [_] (≡-× (cong [_] (≡-× (cong (x ·_) (transportRefl 1r) ∙ ·IdR x) (Σ≡Prop (λ _ → isPropPropTrunc) (cong (1r ·_) (transportRefl 1r) ∙ ·IdR 1r)))) (Σ≡Prop (λ _ → isPropPropTrunc) (cong (1r ·_) (transportRefl 1r) ∙ ·IdR 1r))))) @@ -435,7 +434,7 @@ module _ (R' : CommRing ℓ) where p i = InvertingElementsBase.R[1/_]AsCommRing R[1/ h ]AsCommRing (eqInR[1/h] i) q : PathP (λ i → CommRingHom R' (p i)) /1/1AsCommRingHomFG (/1/1AsCommRingHom (f · g)) - q = toPathP (RingHom≡ (funExt ( + q = toPathP (CommRingHom≡ (funExt ( λ x → cong [_] (≡-× (cong [_] (≡-× (transportRefl _ ∙ transportRefl x) (Σ≡Prop (λ _ → isPropPropTrunc) (transportRefl 1r)))) (Σ≡Prop (λ _ → isPropPropTrunc) (transportRefl 1r)))))) From 74d51ab541b9b294ea2e5c56524d8f667bf9a4ba Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Sun, 4 Aug 2024 12:34:39 +0200 Subject: [PATCH 55/83] stub for polynomials, fix --- .../CommAlgebraAlt/Instances/Polynomials.agda | 128 +++++++++++++++++- 1 file changed, 127 insertions(+), 1 deletion(-) diff --git a/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda b/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda index c6626e88b0..a34e78391f 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda @@ -12,4 +12,130 @@ private ℓ ℓ' : Level _[_] : (R : CommRing ℓ) (I : Type ℓ') → CommAlgebra R _ -R [ I ] = (R [ I ]ᵣ) , const R I +R [ I ] = (R [ I ]ᵣ) , constPolynomial R I + + +{- +evaluateAt : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') + (f : CommAlgebraHom (R [ I ]) A) + → (I → fst A) +evaluateAt A f x = f $a (Construction.var x) + +inducedHom : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') + (φ : I → fst A ) + → CommAlgebraHom (R [ I ]) A +inducedHom A φ = Theory.inducedHom A φ + + +homMapIso : {R : CommRing ℓ} {I : Type ℓ''} (A : CommAlgebra R ℓ') + → Iso (CommAlgebraHom (R [ I ]) A) (I → (fst A)) +Iso.fun (homMapIso A) = evaluateAt A +Iso.inv (homMapIso A) = inducedHom A +Iso.rightInv (homMapIso A) = λ ϕ → Theory.mapRetrievable A ϕ +Iso.leftInv (homMapIso {R = R} {I = I} A) = + λ f → Σ≡Prop (λ f → isPropIsCommAlgebraHom {M = R [ I ]} {N = A} f) + (Theory.homRetrievable A f) + +inducedHomUnique : + {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') (φ : I → fst A ) + → (f : CommAlgebraHom (R [ I ]) A) → ((i : I) → fst f (Construction.var i) ≡ φ i) + → f ≡ inducedHom A φ +inducedHomUnique {I = I} A φ f p = + isoFunInjective (homMapIso A) f (inducedHom A φ) λ j i → p i j + +homMapPath : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R (ℓ-max ℓ ℓ')) + → CommAlgebraHom (R [ I ]) A ≡ (I → fst A) +homMapPath A = isoToPath (homMapIso A) + +{- Corollary: Two homomorphisms with the same values on generators are equal -} +equalByUMP : {R : CommRing ℓ} {I : Type ℓ'} + → (A : CommAlgebra R ℓ'') + → (f g : CommAlgebraHom (R [ I ]) A) + → ((i : I) → fst f (Construction.var i) ≡ fst g (Construction.var i)) + → (x : ⟨ R [ I ] ⟩) → fst f x ≡ fst g x +equalByUMP {R = R} {I = I} A f g = funExt⁻ ∘ cong fst ∘ isoFunInjective (homMapIso A) f g ∘ funExt + +{- A corollary, which is useful for constructing isomorphisms to + algebras with the same universal property -} +isIdByUMP : {R : CommRing ℓ} {I : Type ℓ'} + → (f : CommAlgebraHom (R [ I ]) (R [ I ])) + → ((i : I) → fst f (Construction.var i) ≡ Construction.var i) + → (x : ⟨ R [ I ] ⟩) → fst f x ≡ x +isIdByUMP {R = R} {I = I} f p = equalByUMP (R [ I ]) f (idCAlgHom (R [ I ])) p + +-- The homomorphism induced by the variables is the identity. +inducedHomVar : (R : CommRing ℓ) (I : Type ℓ') + → inducedHom (R [ I ]) Construction.var ≡ idCAlgHom (R [ I ]) +inducedHomVar R I = isoFunInjective (homMapIso (R [ I ])) _ _ refl + +module _ {R : CommRing ℓ} {A B : CommAlgebra R ℓ''} where + open AlgebraHoms + A′ = CommAlgebra→Algebra A + B′ = CommAlgebra→Algebra B + R′ = (CommRing→Ring R) + ν : AlgebraHom A′ B′ → (⟨ A ⟩ → ⟨ B ⟩) + ν φ = φ .fst + + {- + Hom(R[I],A) → (I → A) + ↓ ↓ψ + Hom(R[I],B) → (I → B) + -} + naturalEvR : {I : Type ℓ'} (ψ : CommAlgebraHom A B) + (f : CommAlgebraHom (R [ I ]) A) + → (fst ψ) ∘ evaluateAt A f ≡ evaluateAt B (ψ ∘a f) + naturalEvR ψ f = refl + + {- + Hom(R[I],A) ← (I → A) + ↓ ↓ψ + Hom(R[I],B) ← (I → B) + -} + natIndHomR : {I : Type ℓ'} (ψ : CommAlgebraHom A B) + (ϕ : I → ⟨ A ⟩) + → ψ ∘a inducedHom A ϕ ≡ inducedHom B (fst ψ ∘ ϕ) + natIndHomR ψ ϕ = isoFunInjective (homMapIso B) _ _ + (evaluateAt B (ψ ∘a (inducedHom A ϕ)) ≡⟨ refl ⟩ + fst ψ ∘ evaluateAt A (inducedHom A ϕ) ≡⟨ refl ⟩ + fst ψ ∘ ϕ ≡⟨ Iso.rightInv (homMapIso B) _ ⟩ + evaluateAt B (inducedHom B (fst ψ ∘ ϕ)) ∎) + + {- + Hom(R[I],A) → (I → A) + ↓ ↓ + Hom(R[J],A) → (J → A) + -} + naturalEvL : {I J : Type ℓ'} (φ : J → I) + (f : CommAlgebraHom (R [ I ]) A) + → (evaluateAt A f) ∘ φ + ≡ evaluateAt A (f ∘a (inducedHom (R [ I ]) (λ x → Construction.var (φ x)))) + naturalEvL φ f = refl + +module _ {R : CommRing ℓ} where + {- + Prove that the FreeCommAlgebra over R on zero generators is + isomorphic to the initial R-Algebra - R itsself. + -} + freeOn⊥ : CommAlgebraEquiv (R [ ⊥ ]) (initialCAlg R) + freeOn⊥ = + equivByInitiality + R (R [ ⊥ ]) + {- Show that R[⊥] has the universal property of the + initial R-Algbera and conclude that those are isomorphic -} + λ B → let to : CommAlgebraHom (R [ ⊥ ]) B → (⊥ → fst B) + to = evaluateAt B + + from : (⊥ → fst B) → CommAlgebraHom (R [ ⊥ ]) B + from = inducedHom B + + from-to : (x : _) → from (to x) ≡ x + from-to x = + Σ≡Prop (λ f → isPropIsCommAlgebraHom {M = R [ ⊥ ]} {N = B} f) + (Theory.homRetrievable B x) + + equiv : CommAlgebraHom (R [ ⊥ ]) B ≃ (⊥ → fst B) + equiv = + isoToEquiv + (iso to from (λ x → isContr→isOfHLevel 1 isContr⊥→A _ _) from-to) + in isOfHLevelRespectEquiv 0 (invEquiv equiv) isContr⊥→A +-} From 3573a4c41712798169fee9ca41bb7c631b28a574 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Sun, 4 Aug 2024 15:40:40 +0200 Subject: [PATCH 56/83] finish localization --- .../Algebra/CommAlgebraAlt/Localisation.agda | 44 ++++++++++--------- 1 file changed, 23 insertions(+), 21 deletions(-) diff --git a/Cubical/Algebra/CommAlgebraAlt/Localisation.agda b/Cubical/Algebra/CommAlgebraAlt/Localisation.agda index f366f42c1c..91e7e28922 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Localisation.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Localisation.agda @@ -242,9 +242,11 @@ module AlgLocTwoSubsets (R' : CommRing ℓ) -- A crucial result for the construction of the structure sheaf module DoubleAlgLoc (R : CommRing ℓ) (f g : (fst R)) where + open CommRingStr ⦃...⦄ + private instance + _ = R .snd open Exponentiation R open InvertingElementsBase - open CommRingStr (snd R) hiding (·IdR) open isMultClosedSubset open DoubleLoc R f g hiding (R[1/fg]≡R[1/f][1/g]) open AlgLoc R ([_ⁿ|n≥0] R (f · g)) (powersFormMultClosedSubset R (f · g)) @@ -253,11 +255,10 @@ module DoubleAlgLoc (R : CommRing ℓ) (f g : (fst R)) where open isCommIdeal open RadicalIdeal R -{- private - ⟨_⟩ : (fst R) → CommIdeal - ⟨ f ⟩ = ⟨ replicateFinVec 1 f ⟩[ R ] + ⟨_⟩ᵢ : (fst R) → CommIdeal + ⟨ f ⟩ᵢ = ⟨ replicateFinVec 1 f ⟩[ R ] R[1/fg]AsCommAlgebra = R[1/_]AsCommAlgebra {R = R} (f · g) R[1/fg]ˣ = R[1/_]AsCommRing R (f · g) ˣ @@ -267,12 +268,14 @@ module DoubleAlgLoc (R : CommRing ℓ) (f g : (fst R)) where R[1/f][1/g]AsCommRing = R[1/_]AsCommRing (R[1/_]AsCommRing R f) [ g , 1r , powersFormMultClosedSubset R f .containsOne ] - R[1/f][1/g]AsCommAlgebra = toCommAlg (R[1/f][1/g]AsCommRing , /1/1AsCommRingHom) + + R[1/f][1/g]AsCommAlgebra : CommAlgebra R _ + R[1/f][1/g]AsCommAlgebra = R[1/f][1/g]AsCommRing , /1/1AsCommRingHom R[1/fg]≡R[1/f][1/g] : R[1/fg]AsCommAlgebra ≡ R[1/f][1/g]AsCommAlgebra R[1/fg]≡R[1/f][1/g] = uaCommAlgebra (R[1/fg]AlgCharEquiv _ _ pathtoR[1/fg]) - doubleLocCancel : g ∈ᵢ √ ⟨ f ⟩ → R[1/f][1/g]AsCommAlgebra ≡ R[1/g]AsCommAlgebra + doubleLocCancel : g ∈ᵢ √ ⟨ f ⟩ᵢ → R[1/f][1/g]AsCommAlgebra ≡ R[1/g]AsCommAlgebra doubleLocCancel g∈√⟨f⟩ = sym R[1/fg]≡R[1/f][1/g] ∙ isContrR[1/fg]≡R[1/g] toUnit1 toUnit2 .fst where open S⁻¹RUniversalProp R ([_ⁿ|n≥0] R g) (powersFormMultClosedSubset R g) @@ -282,32 +285,31 @@ module DoubleAlgLoc (R : CommRing ℓ) (f g : (fst R)) where open AlgLocTwoSubsets R ([_ⁿ|n≥0] R (f · g)) (powersFormMultClosedSubset R (f · g)) ([_ⁿ|n≥0] R g) (powersFormMultClosedSubset R g) renaming (isContrS₁⁻¹R≡S₂⁻¹R to isContrR[1/fg]≡R[1/g]) - open CommAlgebraStr ⦃...⦄ hiding (_·_ ; _+_) + open CommAlgebraStr ⦃...⦄ instance - _ = snd R[1/fg]AsCommAlgebra - _ = snd R[1/g]AsCommAlgebra + _ = CommAlgebra→CommAlgebraStr R[1/fg]AsCommAlgebra + _ = CommAlgebra→CommAlgebraStr R[1/g]AsCommAlgebra + _ = R[1/fg]AsCommAlgebra .fst .snd + _ = R[1/g]AsCommAlgebra .fst .snd - toUnit1 : ∀ s → s ∈ [_ⁿ|n≥0] R (f · g) → s ⋆ 1a ∈ R[1/g]ˣ + toUnit1 : ∀ s → s ∈ [_ⁿ|n≥0] R (f · g) → s ⋆ 1r ∈ R[1/g]ˣ toUnit1 s s∈[fgⁿ|n≥0] = subst-∈ R[1/g]ˣ (sym (·IdR (s /1ᵍ))) (RadicalLemma.toUnit R g (f · g) (radHelper _ _ g∈√⟨f⟩) s s∈[fgⁿ|n≥0]) where - radHelper : ∀ x y → x ∈ᵢ √ ⟨ y ⟩ → x ∈ᵢ √ ⟨ y · x ⟩ - radHelper x y = PT.rec ((√ ⟨ y · x ⟩) .fst x .snd) (uncurry helper1) + radHelper : ∀ x y → x ∈ᵢ √ ⟨ y ⟩ᵢ → x ∈ᵢ √ ⟨ y · x ⟩ᵢ + radHelper x y = PT.rec ((√ ⟨ y · x ⟩ᵢ) .fst x .snd) (uncurry helper1) where - helper1 : (n : ℕ) → x ^ n ∈ᵢ ⟨ y ⟩ → x ∈ᵢ √ ⟨ y · x ⟩ - helper1 n = PT.rec ((√ ⟨ y · x ⟩) .fst x .snd) (uncurry helper2) + helper1 : (n : ℕ) → x ^ n ∈ᵢ ⟨ y ⟩ᵢ → x ∈ᵢ √ ⟨ y · x ⟩ᵢ + helper1 n = PT.rec ((√ ⟨ y · x ⟩ᵢ) .fst x .snd) (uncurry helper2) where helper2 : (α : FinVec (fst R) 1) → x ^ n ≡ linearCombination R α (replicateFinVec 1 y) - → x ∈ᵢ √ ⟨ y · x ⟩ + → x ∈ᵢ √ ⟨ y · x ⟩ᵢ helper2 α p = ∣ (suc n) , ∣ α , cong (x ·_) p ∙ solve! R ∣₁ ∣₁ - toUnit2 : ∀ s → s ∈ [_ⁿ|n≥0] R g → s ⋆ 1a ∈ R[1/fg]ˣ + toUnit2 : ∀ s → s ∈ [_ⁿ|n≥0] R g → s ⋆ 1r ∈ R[1/fg]ˣ toUnit2 s s∈[gⁿ|n≥0] = subst-∈ R[1/fg]ˣ (sym (·IdR (s /1ᶠᵍ))) (RadicalLemma.toUnit R (f · g) g radHelper s s∈[gⁿ|n≥0]) where - radHelper : (f · g) ∈ᵢ √ ⟨ g ⟩ - radHelper = ·Closed (snd (√ ⟨ g ⟩)) f (∈→∈√ _ _ (indInIdeal R _ zero)) --} --- -} --- -} + radHelper : (f · g) ∈ᵢ √ ⟨ g ⟩ᵢ + radHelper = ·Closed (snd (√ ⟨ g ⟩ᵢ)) f (∈→∈√ _ _ (indInIdeal R _ zero)) From 414b2f28e7ffc31df9f4a1ede621d777d2034787 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Sun, 4 Aug 2024 17:32:41 +0200 Subject: [PATCH 57/83] start with OnCoproduct, change default notation for R[I] --- .../CommAlgebraAlt/Instances/Polynomials.agda | 7 +- .../Polynomials/Typevariate/Base.agda | 10 +-- .../Polynomials/Typevariate/OnCoproduct.agda | 81 +++++++++++++++++++ .../Polynomials/Typevariate/Properties.agda | 48 +++++++---- 4 files changed, 122 insertions(+), 24 deletions(-) create mode 100644 Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/OnCoproduct.agda diff --git a/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda b/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda index a34e78391f..481aed53bd 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda @@ -9,11 +9,10 @@ open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate private variable - ℓ ℓ' : Level - -_[_] : (R : CommRing ℓ) (I : Type ℓ') → CommAlgebra R _ -R [ I ] = (R [ I ]ᵣ) , constPolynomial R I + ℓ ℓ' ℓ'' : Level +_[_]ₐ : (R : CommRing ℓ) (I : Type ℓ') → CommAlgebra R _ +R [ I ]ₐ = (R [ I ]) , constPolynomial R I {- evaluateAt : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda index 058e2b2c4c..93c792d110 100644 --- a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda @@ -81,15 +81,15 @@ module Construction (R : CommRing ℓ) where constIsCommRingHom : (I : Type ℓ') → IsCommRingHom (R .snd) (const {I = I}) (commRingStr I) constIsCommRingHom I = makeIsCommRingHom refl +HomConst ·HomConst -_[_]ᵣ : (R : CommRing ℓ) (I : Type ℓ') → CommRing (ℓ-max ℓ ℓ') -fst (R [ I ]ᵣ) = Construction.R[_] R I -snd (R [ I ]ᵣ) = Construction.commRingStr R I +_[_] : (R : CommRing ℓ) (I : Type ℓ') → CommRing (ℓ-max ℓ ℓ') +fst (R [ I ]) = Construction.R[_] R I +snd (R [ I ]) = Construction.commRingStr R I -constPolynomial : (R : CommRing ℓ) (I : Type ℓ') → CommRingHom R (R [ I ]ᵣ) +constPolynomial : (R : CommRing ℓ) (I : Type ℓ') → CommRingHom R (R [ I ]) constPolynomial R I .fst = let open Construction R in R[_].const constPolynomial R I .snd = Construction.constIsCommRingHom R I opaque - var : {R : CommRing ℓ} {I : Type ℓ'} → I → ⟨ R [ I ]ᵣ ⟩ + var : {R : CommRing ℓ} {I : Type ℓ'} → I → ⟨ R [ I ] ⟩ var = Construction.var diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/OnCoproduct.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/OnCoproduct.agda new file mode 100644 index 0000000000..eea5dda7a5 --- /dev/null +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/OnCoproduct.agda @@ -0,0 +1,81 @@ +{-# OPTIONS --safe #-} +{- + The goal of this module is to show that for two types I,J, there is an + isomorphism of algebras + + R[I][J] ≃ R[ I ⊎ J ] + + where '⊎' is the disjoint sum. +-} +module Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.OnCoproduct where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function using (_∘_) +open import Cubical.Foundations.Structure using (⟨_⟩) + +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Sigma + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate + +private + variable + ℓ ℓ' : Level + +module CalculatePolynomialsOnCoproduct (R : CommRing ℓ) (I J : Type ℓ) where + private + I→I+J : CommRingHom (R [ I ]) (R [ I ⊎ J ]) + I→I+J = inducedHom (R [ I ⊎ J ]) (constPolynomial R (I ⊎ J)) (var ∘ inl) + + to : CommRingHom ((R [ I ]) [ J ]) (R [ I ⊎ J ]) + to = inducedHom (R [ I ⊎ J ]) I→I+J (var ∘ inr) + + constPolynomialIJ : CommRingHom R ((R [ I ]) [ J ]) + constPolynomialIJ = constPolynomial _ _ ∘cr constPolynomial _ _ + + evalVarTo : to .fst ∘ var ≡ var ∘ inr + evalVarTo = evalInduce (R [ I ⊎ J ]) I→I+J (var ∘ inr) + + commConstTo : to ∘cr constPolynomialIJ ≡ constPolynomial _ _ + commConstTo = CommRingHom≡ refl + + mapVars : I ⊎ J → ⟨ (R [ I ]) [ J ] ⟩ + mapVars (inl i) = constPolynomial _ _ $cr var i + mapVars (inr j) = var j + + to∘MapVars : to .fst ∘ mapVars ≡ var + to∘MapVars = funExt λ {(inl i) → to .fst (constPolynomial _ _ $cr var i) + ≡⟨ cong (λ z → z i) (evalInduce (R [ I ⊎ J ]) (constPolynomial R (I ⊎ J)) (var ∘ inl)) ⟩ + var (inl i) ∎; + (inr j) → (to .fst (var j) ≡⟨ cong (λ z → z j) evalVarTo ⟩ var (inr j) ∎)} + + from : CommRingHom (R [ I ⊎ J ]) ((R [ I ]) [ J ]) + from = inducedHom + ((R [ I ]) [ J ]) + (constPolynomial (R [ I ]) J ∘cr constPolynomial R I) + mapVars + + evalVarFrom : from .fst ∘ var ≡ mapVars + evalVarFrom = evalInduce ((R [ I ]) [ J ]) (constPolynomial (R [ I ]) J ∘cr constPolynomial R I) mapVars + + toFrom : to ∘cr from ≡ (idCommRingHom _) + toFrom = + idByIdOnVars + (to ∘cr from) + (to .fst ∘ from .fst ∘ constPolynomial R (I ⊎ J) .fst ≡⟨⟩ + constPolynomial R (I ⊎ J) .fst ∎) + (to .fst ∘ from .fst ∘ var ≡⟨ cong (to .fst ∘_) evalVarFrom ⟩ + to .fst ∘ mapVars ≡⟨ to∘MapVars ⟩ + var ∎) +{- + fromTo : from ∘cr to ≡ (idCommRingHom _) + fromTo = + idByIdOnVars + (from ∘cr to) + (from .fst ∘ to .fst ∘ constPolynomial (R [ I ]) J .fst ≡⟨⟩ + from .fst ∘ I→I+J .fst + ≡⟨ cong fst (hom≡ByValuesOnVars ((R [ I ]) [ J ]) {!from ∘cr I→I+J!} {!I→I+J!} {!!} {!!} {!!} {!!}) ⟩ + constPolynomial (R [ I ]) J .fst ∎) + (from .fst ∘ to .fst ∘ var ≡⟨ {!!} ⟩ var ∎) +-} diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda index e721203840..65675589b5 100644 --- a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda @@ -35,7 +35,7 @@ module _ {R : CommRing ℓ} {I : Type ℓ'} where open CommRingStr ⦃...⦄ private instance _ = snd R - _ = snd (R [ I ]ᵣ) + _ = snd (R [ I ]) module C = Construction open C using (const) @@ -44,16 +44,16 @@ module _ {R : CommRing ℓ} {I : Type ℓ'} where Construction of the 'elimProp' eliminator. -} module _ - {P : ⟨ R [ I ]ᵣ ⟩ → Type ℓ''} + {P : ⟨ R [ I ] ⟩ → Type ℓ''} (isPropP : {x : _} → isProp (P x)) (onVar : {x : I} → P (C.var x)) (onConst : {r : ⟨ R ⟩} → P (const r)) - (on+ : {x y : ⟨ R [ I ]ᵣ ⟩} → P x → P y → P (x + y)) - (on· : {x y : ⟨ R [ I ]ᵣ ⟩} → P x → P y → P (x · y)) + (on+ : {x y : ⟨ R [ I ] ⟩} → P x → P y → P (x + y)) + (on· : {x y : ⟨ R [ I ] ⟩} → P x → P y → P (x · y)) where private - on- : {x : ⟨ R [ I ]ᵣ ⟩} → P x → P (- x) + on- : {x : ⟨ R [ I ] ⟩} → P x → P (- x) on- {x} Px = subst P (minusByMult x) (on· onConst Px) where minusByMult : (x : _) → (const (- 1r)) · x ≡ - x minusByMult x = @@ -61,12 +61,12 @@ module _ {R : CommRing ℓ} {I : Type ℓ'} where (- const (1r)) · x ≡⟨ cong (λ u → (- u) · x) pres1 ⟩ (- 1r) · x ≡⟨ sym (-IsMult-1 x) ⟩ - x ∎ - where open RingTheory (CommRing→Ring (R [ I ]ᵣ)) using (-IsMult-1) + where open RingTheory (CommRing→Ring (R [ I ])) using (-IsMult-1) open IsCommRingHom (constPolynomial R I .snd) -- A helper to deal with the path constructor cases. mkPathP : - {x0 x1 : ⟨ R [ I ]ᵣ ⟩} → + {x0 x1 : ⟨ R [ I ] ⟩} → (p : x0 ≡ x1) → (Px0 : P (x0)) → (Px1 : P (x1)) → @@ -160,7 +160,7 @@ module _ {R : CommRing ℓ} {I : Type ℓ'} where open IsCommRingHom - inducedMap : ⟨ R [ I ]ᵣ ⟩ → ⟨ S ⟩ + inducedMap : ⟨ R [ I ] ⟩ → ⟨ S ⟩ inducedMap (C.var x) = φ x inducedMap (const r) = (f .fst r) inducedMap (P C.+ Q) = (inducedMap P) + (inducedMap Q) @@ -198,7 +198,7 @@ module _ {R : CommRing ℓ} {I : Type ℓ'} where module _ where open IsCommRingHom - inducedHom : CommRingHom (R [ I ]ᵣ) S + inducedHom : CommRingHom (R [ I ]) S inducedHom .fst = inducedMap inducedHom .snd .pres0 = pres0 (f .snd) inducedHom .snd .pres1 = pres1 (f. snd) @@ -214,11 +214,11 @@ module _ {R : CommRing ℓ} {I : Type ℓ'} (S : CommRing ℓ'') (f : CommRingH open CommRingStr ⦃...⦄ private instance _ = S .snd - _ = (R [ I ]ᵣ) .snd + _ = (R [ I ]) .snd open IsCommRingHom - evalVar : CommRingHom (R [ I ]ᵣ) S → I → ⟨ S ⟩ - evalVar f = f .fst ∘ var + evalVar : CommRingHom (R [ I ]) S → I → ⟨ S ⟩ + evalVar h = h .fst ∘ var opaque unfolding var @@ -228,11 +228,11 @@ module _ {R : CommRing ℓ} {I : Type ℓ'} (S : CommRing ℓ'') (f : CommRingH opaque unfolding var - induceEval : (g : CommRingHom (R [ I ]ᵣ) S) + induceEval : (g : CommRingHom (R [ I ]) S) (p : g .fst ∘ constPolynomial R I .fst ≡ f .fst) → (inducedHom S f (evalVar g)) ≡ g induceEval g p = - let theMap : ⟨ R [ I ]ᵣ ⟩ → ⟨ S ⟩ + let theMap : ⟨ R [ I ] ⟩ → ⟨ S ⟩ theMap = inducedMap S f (evalVar g) in CommRingHom≡ $ @@ -254,8 +254,26 @@ module _ {R : CommRing ℓ} {I : Type ℓ'} (S : CommRing ℓ'') (f : CommRingH opaque inducedHomUnique : (φ : I → ⟨ S ⟩) - (g : CommRingHom (R [ I ]ᵣ) S) + (g : CommRingHom (R [ I ]) S) (p : g .fst ∘ constPolynomial R I .fst ≡ f .fst) (q : evalVar g ≡ φ) → inducedHom S f φ ≡ g inducedHomUnique φ g p q = cong (inducedHom S f) (sym q) ∙ induceEval g p + + opaque + hom≡ByValuesOnVars : (g h : CommRingHom (R [ I ]) S) + (p : g .fst ∘ constPolynomial _ _ .fst ≡ f .fst) (q : h .fst ∘ constPolynomial _ _ .fst ≡ f .fst) + → (evalVar g ≡ evalVar h) + → g ≡ h + hom≡ByValuesOnVars g h p q ≡OnVars = + sym (inducedHomUnique ϕ g p refl) ∙ inducedHomUnique ϕ h q (sym ≡OnVars) + where ϕ : I → ⟨ S ⟩ + ϕ = evalVar g + +opaque + idByIdOnVars : {R : CommRing ℓ} {I : Type ℓ'} + (g : CommRingHom (R [ I ]) (R [ I ])) + (p : g .fst ∘ constPolynomial _ _ .fst ≡ constPolynomial _ _ .fst) + → (g .fst ∘ var ≡ idfun _ ∘ var) + → g ≡ idCommRingHom (R [ I ]) + idByIdOnVars g p idOnVars = hom≡ByValuesOnVars _ (constPolynomial _ _) g (idCommRingHom _) p refl idOnVars From 0b76ab78babaf0203c6375bb50cd4b8a21cfcc29 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Wed, 7 Aug 2024 15:08:42 +0200 Subject: [PATCH 58/83] universal property of polynomials as rings --- .../Polynomials/Typevariate/OnCoproduct.agda | 48 ++++++++++++++----- .../Polynomials/Typevariate/Properties.agda | 8 +++- Cubical/Algebra/CommRing/Properties.agda | 1 + 3 files changed, 44 insertions(+), 13 deletions(-) diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/OnCoproduct.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/OnCoproduct.agda index eea5dda7a5..3290a2276f 100644 --- a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/OnCoproduct.agda +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/OnCoproduct.agda @@ -10,7 +10,7 @@ module Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.OnCoproduct where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Function using (_∘_) +open import Cubical.Foundations.Function using (_∘_; _$_) open import Cubical.Foundations.Structure using (⟨_⟩) open import Cubical.Data.Sum as ⊎ @@ -45,10 +45,13 @@ module CalculatePolynomialsOnCoproduct (R : CommRing ℓ) (I J : Type ℓ) where mapVars (inr j) = var j to∘MapVars : to .fst ∘ mapVars ≡ var - to∘MapVars = funExt λ {(inl i) → to .fst (constPolynomial _ _ $cr var i) - ≡⟨ cong (λ z → z i) (evalInduce (R [ I ⊎ J ]) (constPolynomial R (I ⊎ J)) (var ∘ inl)) ⟩ - var (inl i) ∎; - (inr j) → (to .fst (var j) ≡⟨ cong (λ z → z j) evalVarTo ⟩ var (inr j) ∎)} + to∘MapVars = + funExt λ {(inl i) → to .fst (constPolynomial _ _ $cr var i) + ≡⟨ cong (λ z → z i) + (evalInduce (R [ I ⊎ J ]) + (constPolynomial R (I ⊎ J)) (var ∘ inl)) ⟩ + var (inl i) ∎; + (inr j) → (to .fst (var j) ≡⟨ cong (λ z → z j) evalVarTo ⟩ var (inr j) ∎)} from : CommRingHom (R [ I ⊎ J ]) ((R [ I ]) [ J ]) from = inducedHom @@ -59,7 +62,7 @@ module CalculatePolynomialsOnCoproduct (R : CommRing ℓ) (I J : Type ℓ) where evalVarFrom : from .fst ∘ var ≡ mapVars evalVarFrom = evalInduce ((R [ I ]) [ J ]) (constPolynomial (R [ I ]) J ∘cr constPolynomial R I) mapVars - toFrom : to ∘cr from ≡ (idCommRingHom _) + toFrom : to ∘cr from ≡ (idCommRingHom (R [ I ⊎ J ])) toFrom = idByIdOnVars (to ∘cr from) @@ -68,14 +71,35 @@ module CalculatePolynomialsOnCoproduct (R : CommRing ℓ) (I J : Type ℓ) where (to .fst ∘ from .fst ∘ var ≡⟨ cong (to .fst ∘_) evalVarFrom ⟩ to .fst ∘ mapVars ≡⟨ to∘MapVars ⟩ var ∎) -{- - fromTo : from ∘cr to ≡ (idCommRingHom _) + + {- from and to are R[I]-algebra homomorphisms: + this is true definitionally for to. + -} + fromAlgebraHom : from ∘cr I→I+J ≡ constPolynomial (R [ I ]) J + fromAlgebraHom = + hom≡ByValuesOnVars + ((R [ I ]) [ J ]) (constPolynomial (R [ I ]) J ∘cr constPolynomial R I) + (from ∘cr I→I+J) (constPolynomial (R [ I ]) J) + refl refl + (funExt + (λ i → (from ∘cr I→I+J) .fst (var i) + ≡⟨ cong (from .fst) + (evalOnVars (R [ I ⊎ J ]) + (constPolynomial R (I ⊎ J)) + (var ∘ inl) i) ⟩ + from .fst (var (inl i)) + ≡⟨ evalOnVars ((R [ I ]) [ J ]) + (constPolynomial (R [ I ]) J ∘cr constPolynomial R I) + mapVars (inl i) ⟩ + constPolynomial (R [ I ]) J $cr var i ∎)) + + fromTo : from ∘cr to ≡ (idCommRingHom $ (R [ I ]) [ J ]) fromTo = idByIdOnVars (from ∘cr to) (from .fst ∘ to .fst ∘ constPolynomial (R [ I ]) J .fst ≡⟨⟩ - from .fst ∘ I→I+J .fst - ≡⟨ cong fst (hom≡ByValuesOnVars ((R [ I ]) [ J ]) {!from ∘cr I→I+J!} {!I→I+J!} {!!} {!!} {!!} {!!}) ⟩ + from .fst ∘ I→I+J .fst ≡⟨ cong fst fromAlgebraHom ⟩ constPolynomial (R [ I ]) J .fst ∎) - (from .fst ∘ to .fst ∘ var ≡⟨ {!!} ⟩ var ∎) --} + (from .fst ∘ to .fst ∘ var ≡⟨ cong (from .fst ∘_) (funExt $ evalOnVars (R [ I ⊎ J ]) I→I+J (var ∘ inr) ) ⟩ + from .fst ∘ var ∘ inr ≡⟨ (funExt $ λ j → evalOnVars _ _ _ (inr j)) ⟩ + var ∎) diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda index 65675589b5..0af77c592c 100644 --- a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda @@ -223,9 +223,15 @@ module _ {R : CommRing ℓ} {I : Type ℓ'} (S : CommRing ℓ'') (f : CommRingH opaque unfolding var evalInduce : ∀ (φ : I → ⟨ S ⟩) - → evalVar (inducedHom S f φ) ≡ φ + → evalVar (inducedHom S f φ) ≡ φ evalInduce φ = refl + opaque + unfolding var + evalOnVars : ∀ (φ : I → ⟨ S ⟩) + → (i : I) → inducedHom S f φ .fst (var i) ≡ φ i + evalOnVars φ i = refl + opaque unfolding var induceEval : (g : CommRingHom (R [ I ]) S) diff --git a/Cubical/Algebra/CommRing/Properties.agda b/Cubical/Algebra/CommRing/Properties.agda index ea3424ab15..da9ac7a4d7 100644 --- a/Cubical/Algebra/CommRing/Properties.agda +++ b/Cubical/Algebra/CommRing/Properties.agda @@ -162,6 +162,7 @@ module _ where RingHom→CommRingHom (compRingHom (CommRingHom→RingHom f) (CommRingHom→RingHom g)) + infixr 9 _∘cr_ -- same as functions _∘cr_ : {R : CommRing ℓ} {S : CommRing ℓ'} {T : CommRing ℓ''} → CommRingHom S T → CommRingHom R S → CommRingHom R T g ∘cr f = compCommRingHom f g From 1c71ebde4939c706fa4b4dde2b5d05bcc9c43da8 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Wed, 7 Aug 2024 17:34:52 +0200 Subject: [PATCH 59/83] move old CommAlgebra --- Cubical/Algebra/CommAlgebra.agda | 4 ++-- .../Algebra/CommAlgebra/FreeCommAlgebra.agda | 6 ------ Cubical/Algebra/CommAlgebraAlt.agda | 5 +++++ .../AsModule}/Base.agda | 5 ++++- .../AsModule}/FGIdeal.agda | 4 ++-- .../AsModule}/FPAlgebra.agda | 4 ++-- .../AsModule}/FPAlgebra/Base.agda | 12 ++++++------ .../AsModule}/FPAlgebra/Instances.agda | 18 +++++++++--------- .../AsModule/FreeCommAlgebra.agda | 6 ++++++ .../AsModule}/FreeCommAlgebra/Base.agda | 4 ++-- .../AsModule}/FreeCommAlgebra/OnCoproduct.agda | 0 .../AsModule}/FreeCommAlgebra/Properties.agda | 6 +++--- .../AsModule}/Ideal.agda | 2 +- .../AsModule}/Instances/Initial.agda | 2 +- .../AsModule}/Instances/Pointwise.agda | 4 ++-- .../AsModule}/Instances/Unit.agda | 4 ++-- .../AsModule}/Kernel.agda | 8 ++++---- .../AsModule}/Localisation.agda | 6 +++--- .../AsModule}/LocalisationAlgebra.agda | 4 ++-- .../AsModule}/Properties.agda | 4 ++-- .../AsModule}/QuotientAlgebra.agda | 8 ++++---- .../AsModule}/Subalgebra.agda | 2 +- .../AsModule}/UnivariatePolyList.agda | 2 +- .../ZariskiLattice/StructureSheaf.agda | 6 +++--- Cubical/Categories/Instances/CommAlgebras.agda | 2 +- Cubical/Papers/AffineSchemes.agda | 2 +- 26 files changed, 69 insertions(+), 61 deletions(-) delete mode 100644 Cubical/Algebra/CommAlgebra/FreeCommAlgebra.agda create mode 100644 Cubical/Algebra/CommAlgebraAlt.agda rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/Base.agda (99%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/FGIdeal.agda (91%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/FPAlgebra.agda (67%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/FPAlgebra/Base.agda (96%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/FPAlgebra/Instances.agda (93%) create mode 100644 Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra.agda rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/FreeCommAlgebra/Base.agda (96%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/FreeCommAlgebra/OnCoproduct.agda (100%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/FreeCommAlgebra/Properties.agda (98%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/Ideal.agda (95%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/Instances/Initial.agda (98%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/Instances/Pointwise.agda (90%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/Instances/Unit.agda (94%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/Kernel.agda (69%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/Localisation.agda (98%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/LocalisationAlgebra.agda (99%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/Properties.agda (99%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/QuotientAlgebra.agda (97%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/Subalgebra.agda (97%) rename Cubical/Algebra/{CommAlgebra => CommAlgebraAlt/AsModule}/UnivariatePolyList.agda (99%) diff --git a/Cubical/Algebra/CommAlgebra.agda b/Cubical/Algebra/CommAlgebra.agda index d150261beb..8600ad3ada 100644 --- a/Cubical/Algebra/CommAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} module Cubical.Algebra.CommAlgebra where -open import Cubical.Algebra.CommAlgebra.Base public -open import Cubical.Algebra.CommAlgebra.Properties public +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base public +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Properties public diff --git a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra.agda b/Cubical/Algebra/CommAlgebra/FreeCommAlgebra.agda deleted file mode 100644 index ac7f45ffea..0000000000 --- a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra.agda +++ /dev/null @@ -1,6 +0,0 @@ -{-# OPTIONS --safe #-} - -module Cubical.Algebra.CommAlgebra.FreeCommAlgebra where - -open import Cubical.Algebra.CommAlgebra.FreeCommAlgebra.Base public -open import Cubical.Algebra.CommAlgebra.FreeCommAlgebra.Properties public diff --git a/Cubical/Algebra/CommAlgebraAlt.agda b/Cubical/Algebra/CommAlgebraAlt.agda new file mode 100644 index 0000000000..f7f8ed410e --- /dev/null +++ b/Cubical/Algebra/CommAlgebraAlt.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebraAlt where + +open import Cubical.Algebra.CommAlgebraAlt.Base public +open import Cubical.Algebra.CommAlgebraAlt.Properties public diff --git a/Cubical/Algebra/CommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/Base.agda similarity index 99% rename from Cubical/Algebra/CommAlgebra/Base.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/Base.agda index fe04691d36..36addd0c39 100644 --- a/Cubical/Algebra/CommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/Base.agda @@ -1,5 +1,8 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.Base where +{- + This used to be the default definition of CommAlgebra. +-} +module Cubical.Algebra.CommAlgebraAlt.AsModule.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv diff --git a/Cubical/Algebra/CommAlgebra/FGIdeal.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/FGIdeal.agda similarity index 91% rename from Cubical/Algebra/CommAlgebra/FGIdeal.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/FGIdeal.agda index e06ed81b11..e9464a4059 100644 --- a/Cubical/Algebra/CommAlgebra/FGIdeal.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/FGIdeal.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.FGIdeal where +module Cubical.Algebra.CommAlgebraAlt.AsModule.FGIdeal where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure open import Cubical.Foundations.Powerset @@ -14,7 +14,7 @@ open import Cubical.Algebra.CommRing.FGIdeal using () indInIdeal to ringIncInIdeal; 0FGIdeal to 0FGIdealCommRing) open import Cubical.Algebra.CommAlgebra -open import Cubical.Algebra.CommAlgebra.Ideal +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Ideal private variable diff --git a/Cubical/Algebra/CommAlgebra/FPAlgebra.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra.agda similarity index 67% rename from Cubical/Algebra/CommAlgebra/FPAlgebra.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra.agda index d454e4f0a1..23d84f5ba9 100644 --- a/Cubical/Algebra/CommAlgebra/FPAlgebra.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra.agda @@ -9,6 +9,6 @@ Our definition is more explicit. -} {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.FPAlgebra where +module Cubical.Algebra.CommAlgebraAlt.AsModule.FPAlgebra where -open import Cubical.Algebra.CommAlgebra.FPAlgebra.Base public +open import Cubical.Algebra.CommAlgebraAlt.AsModule.FPAlgebra.Base public diff --git a/Cubical/Algebra/CommAlgebra/FPAlgebra/Base.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Base.agda similarity index 96% rename from Cubical/Algebra/CommAlgebra/FPAlgebra/Base.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Base.agda index 322e06a5f9..77615ad8c9 100644 --- a/Cubical/Algebra/CommAlgebra/FPAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Base.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.FPAlgebra.Base where +module Cubical.Algebra.CommAlgebraAlt.AsModule.FPAlgebra.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -19,13 +19,13 @@ open import Cubical.HITs.PropositionalTruncation open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.FGIdeal using (inclOfFGIdeal) open import Cubical.Algebra.CommAlgebra -open import Cubical.Algebra.CommAlgebra.FreeCommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra renaming (inducedHom to freeInducedHom) -open import Cubical.Algebra.CommAlgebra.QuotientAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule.QuotientAlgebra renaming (inducedHom to quotientInducedHom) -open import Cubical.Algebra.CommAlgebra.Ideal -open import Cubical.Algebra.CommAlgebra.FGIdeal -open import Cubical.Algebra.CommAlgebra.Kernel +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Ideal +open import Cubical.Algebra.CommAlgebraAlt.AsModule.FGIdeal +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Kernel open import Cubical.Algebra.Algebra.Properties open import Cubical.Algebra.Algebra diff --git a/Cubical/Algebra/CommAlgebra/FPAlgebra/Instances.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Instances.agda similarity index 93% rename from Cubical/Algebra/CommAlgebra/FPAlgebra/Instances.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Instances.agda index 24319606d8..1d16ad4262 100644 --- a/Cubical/Algebra/CommAlgebra/FPAlgebra/Instances.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Instances.agda @@ -8,7 +8,7 @@ * R/⟨x⟩ (as special case of the above) -} {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.FPAlgebra.Instances where +module Cubical.Algebra.CommAlgebraAlt.AsModule.FPAlgebra.Instances where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -28,19 +28,19 @@ open import Cubical.HITs.PropositionalTruncation open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.FGIdeal using (inclOfFGIdeal) open import Cubical.Algebra.CommAlgebra -open import Cubical.Algebra.CommAlgebra.FreeCommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra renaming (inducedHom to freeInducedHom) -open import Cubical.Algebra.CommAlgebra.QuotientAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule.QuotientAlgebra renaming (inducedHom to quotientInducedHom) -open import Cubical.Algebra.CommAlgebra.Ideal using (IdealsIn) -open import Cubical.Algebra.CommAlgebra.FGIdeal -open import Cubical.Algebra.CommAlgebra.Instances.Initial -open import Cubical.Algebra.CommAlgebra.Instances.Unit +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Ideal using (IdealsIn) +open import Cubical.Algebra.CommAlgebraAlt.AsModule.FGIdeal +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Initial +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Unit renaming (UnitCommAlgebra to TerminalCAlg) -open import Cubical.Algebra.CommAlgebra.Kernel +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Kernel open import Cubical.Algebra.Algebra -open import Cubical.Algebra.CommAlgebra.FPAlgebra.Base +open import Cubical.Algebra.CommAlgebraAlt.AsModule.FPAlgebra.Base private variable diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra.agda new file mode 100644 index 0000000000..9a98574f8a --- /dev/null +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} + +module Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra where + +open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra.Base public +open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra.Properties public diff --git a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Base.agda similarity index 96% rename from Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Base.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Base.agda index b5291c432a..f45047d1ae 100644 --- a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Base.agda @@ -1,6 +1,6 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.FreeCommAlgebra.Base where +module Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra.Base where {- The free commutative algebra over a commutative ring, or in other words the ring of polynomials with coefficients in a given ring. @@ -21,7 +21,7 @@ module Cubical.Algebra.CommAlgebra.FreeCommAlgebra.Base where open import Cubical.Foundations.Prelude open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base private variable diff --git a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/OnCoproduct.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/OnCoproduct.agda similarity index 100% rename from Cubical/Algebra/CommAlgebra/FreeCommAlgebra/OnCoproduct.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/OnCoproduct.agda diff --git a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Properties.agda similarity index 98% rename from Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Properties.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Properties.agda index 76c462ee48..626731008a 100644 --- a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Properties.agda @@ -1,6 +1,6 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.FreeCommAlgebra.Properties where +module Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra.Properties where {- This file contains * an elimination principle for proving some proposition for all elements of R[I] @@ -31,8 +31,8 @@ open import Cubical.HITs.SetTruncation open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommAlgebra -open import Cubical.Algebra.CommAlgebra.FreeCommAlgebra.Base -open import Cubical.Algebra.CommAlgebra.Instances.Initial +open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra.Base +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Initial open import Cubical.Algebra.Algebra open import Cubical.Algebra.Module using (module ModuleTheory) diff --git a/Cubical/Algebra/CommAlgebra/Ideal.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/Ideal.agda similarity index 95% rename from Cubical/Algebra/CommAlgebra/Ideal.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/Ideal.agda index b3f8b0d126..3ac7be5f81 100644 --- a/Cubical/Algebra/CommAlgebra/Ideal.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/Ideal.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.Ideal where +module Cubical.Algebra.CommAlgebraAlt.AsModule.Ideal where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Powerset diff --git a/Cubical/Algebra/CommAlgebra/Instances/Initial.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Initial.agda similarity index 98% rename from Cubical/Algebra/CommAlgebra/Instances/Initial.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Initial.agda index 26c8b6bdd6..d9b4708eb5 100644 --- a/Cubical/Algebra/CommAlgebra/Instances/Initial.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Initial.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.Instances.Initial where +module Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Initial where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels diff --git a/Cubical/Algebra/CommAlgebra/Instances/Pointwise.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Pointwise.agda similarity index 90% rename from Cubical/Algebra/CommAlgebra/Instances/Pointwise.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Pointwise.agda index 9ce845e935..fe3f048572 100644 --- a/Cubical/Algebra/CommAlgebra/Instances/Pointwise.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Pointwise.agda @@ -1,12 +1,12 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.Instances.Pointwise where +module Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Pointwise where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Algebra.CommRing.Base open import Cubical.Algebra.CommRing.Instances.Pointwise -open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base open import Cubical.Algebra.Algebra using (IsAlgebraHom) private diff --git a/Cubical/Algebra/CommAlgebra/Instances/Unit.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Unit.agda similarity index 94% rename from Cubical/Algebra/CommAlgebra/Instances/Unit.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Unit.agda index 7154fc6b84..710ec2069a 100644 --- a/Cubical/Algebra/CommAlgebra/Instances/Unit.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Unit.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.Instances.Unit where +module Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Unit where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -9,7 +9,7 @@ open import Cubical.Data.Unit open import Cubical.Data.Sigma.Properties using (Σ≡Prop) open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base open import Cubical.Algebra.CommRing.Instances.Unit open import Cubical.Algebra.Algebra.Base using (IsAlgebraHom) open import Cubical.Tactics.CommRingSolver diff --git a/Cubical/Algebra/CommAlgebra/Kernel.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/Kernel.agda similarity index 69% rename from Cubical/Algebra/CommAlgebra/Kernel.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/Kernel.agda index 41d8e9af96..fa8ab8b4b9 100644 --- a/Cubical/Algebra/CommAlgebra/Kernel.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/Kernel.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.Kernel where +module Cubical.Algebra.CommAlgebraAlt.AsModule.Kernel where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -7,9 +7,9 @@ open import Cubical.Foundations.Equiv open import Cubical.Algebra.CommRing.Base open import Cubical.Algebra.CommRing.Ideal using (Ideal→CommIdeal) open import Cubical.Algebra.Ring.Kernel using () renaming (kernelIdeal to ringKernel) -open import Cubical.Algebra.CommAlgebra.Base -open import Cubical.Algebra.CommAlgebra.Properties -open import Cubical.Algebra.CommAlgebra.Ideal +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Properties +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Ideal private variable diff --git a/Cubical/Algebra/CommAlgebra/Localisation.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/Localisation.agda similarity index 98% rename from Cubical/Algebra/CommAlgebra/Localisation.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/Localisation.agda index 17745f411f..a3045e0ee7 100644 --- a/Cubical/Algebra/CommAlgebra/Localisation.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/Localisation.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --lossy-unification #-} -module Cubical.Algebra.CommAlgebra.Localisation where +module Cubical.Algebra.CommAlgebraAlt.AsModule.Localisation where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -28,8 +28,8 @@ open import Cubical.Algebra.CommRing.RadicalIdeal open import Cubical.Algebra.CommRing.Localisation open import Cubical.Algebra.Ring open import Cubical.Algebra.Algebra -open import Cubical.Algebra.CommAlgebra.Base -open import Cubical.Algebra.CommAlgebra.Properties +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Properties open import Cubical.Tactics.CommRingSolver diff --git a/Cubical/Algebra/CommAlgebra/LocalisationAlgebra.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/LocalisationAlgebra.agda similarity index 99% rename from Cubical/Algebra/CommAlgebra/LocalisationAlgebra.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/LocalisationAlgebra.agda index 1fbf864dfc..88179e9626 100644 --- a/Cubical/Algebra/CommAlgebra/LocalisationAlgebra.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/LocalisationAlgebra.agda @@ -23,12 +23,12 @@ import Cubical.HITs.PropositionalTruncation as PropTrunc open import Cubical.Algebra.Algebra open import Cubical.Algebra.CommAlgebra -open import Cubical.Algebra.CommAlgebra.Subalgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Subalgebra open import Cubical.Algebra.CommRing as CommRing hiding (_ˣ;module Units) open import Cubical.Algebra.CommRing.Localisation using (isMultClosedSubset) open import Cubical.Algebra.Ring -module Cubical.Algebra.CommAlgebra.LocalisationAlgebra +module Cubical.Algebra.CommAlgebraAlt.AsModule.LocalisationAlgebra {ℓR : Level} (R : CommRing ℓR) where diff --git a/Cubical/Algebra/CommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/Properties.agda similarity index 99% rename from Cubical/Algebra/CommAlgebra/Properties.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/Properties.agda index feb73d2da4..db2b937858 100644 --- a/Cubical/Algebra/CommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/Properties.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --lossy-unification #-} -module Cubical.Algebra.CommAlgebra.Properties where +module Cubical.Algebra.CommAlgebraAlt.AsModule.Properties where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function @@ -26,7 +26,7 @@ open import Cubical.Algebra.Monoid open import Cubical.Algebra.CommRing open import Cubical.Algebra.Ring open import Cubical.Algebra.Algebra -open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base open import Cubical.Algebra.CommRing using (CommRing→Ring) diff --git a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/QuotientAlgebra.agda similarity index 97% rename from Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/QuotientAlgebra.agda index 3390614d92..ccb19527ec 100644 --- a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/QuotientAlgebra.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.QuotientAlgebra where +module Cubical.Algebra.CommAlgebraAlt.AsModule.QuotientAlgebra where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv @@ -16,9 +16,9 @@ import Cubical.Algebra.CommRing.Quotient as CommRing import Cubical.Algebra.Ring.Quotient as Ring open import Cubical.Algebra.CommRing.Ideal hiding (IdealsIn) open import Cubical.Algebra.CommAlgebra -open import Cubical.Algebra.CommAlgebra.Ideal -open import Cubical.Algebra.CommAlgebra.Kernel -open import Cubical.Algebra.CommAlgebra.Instances.Unit +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Ideal +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Kernel +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Unit open import Cubical.Algebra.Algebra.Base using (IsAlgebraHom; isPropIsAlgebraHom) open import Cubical.Algebra.Ring open import Cubical.Algebra.Ring.Ideal using (isIdeal) diff --git a/Cubical/Algebra/CommAlgebra/Subalgebra.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/Subalgebra.agda similarity index 97% rename from Cubical/Algebra/CommAlgebra/Subalgebra.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/Subalgebra.agda index fd052280c8..a386292275 100644 --- a/Cubical/Algebra/CommAlgebra/Subalgebra.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/Subalgebra.agda @@ -7,7 +7,7 @@ open import Cubical.Algebra.Algebra open import Cubical.Algebra.CommAlgebra open import Cubical.Algebra.CommRing -module Cubical.Algebra.CommAlgebra.Subalgebra +module Cubical.Algebra.CommAlgebraAlt.AsModule.Subalgebra {ℓ ℓ' : Level} (R : CommRing ℓ) (A : CommAlgebra R ℓ') where diff --git a/Cubical/Algebra/CommAlgebra/UnivariatePolyList.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/UnivariatePolyList.agda similarity index 99% rename from Cubical/Algebra/CommAlgebra/UnivariatePolyList.agda rename to Cubical/Algebra/CommAlgebraAlt/AsModule/UnivariatePolyList.agda index 982b966c1a..ddd41b50db 100644 --- a/Cubical/Algebra/CommAlgebra/UnivariatePolyList.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/UnivariatePolyList.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.UnivariatePolyList where +module Cubical.Algebra.CommAlgebraAlt.AsModule.UnivariatePolyList where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure diff --git a/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda index a51e4e950e..007875d8fe 100644 --- a/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda +++ b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda @@ -47,9 +47,9 @@ open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.RadicalIdeal open import Cubical.Algebra.CommRing.Localisation open import Cubical.Algebra.CommRing.Instances.Unit -open import Cubical.Algebra.CommAlgebra.Base -open import Cubical.Algebra.CommAlgebra.Properties -open import Cubical.Algebra.CommAlgebra.Localisation +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Properties +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Localisation open import Cubical.Tactics.CommRingSolver open import Cubical.Algebra.Semilattice open import Cubical.Algebra.Lattice diff --git a/Cubical/Categories/Instances/CommAlgebras.agda b/Cubical/Categories/Instances/CommAlgebras.agda index d81f5486e7..3ac204bb19 100644 --- a/Cubical/Categories/Instances/CommAlgebras.agda +++ b/Cubical/Categories/Instances/CommAlgebras.agda @@ -15,7 +15,7 @@ open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing open import Cubical.Algebra.Algebra open import Cubical.Algebra.CommAlgebra -open import Cubical.Algebra.CommAlgebra.Instances.Unit +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Unit open import Cubical.Categories.Category open import Cubical.Categories.Functor diff --git a/Cubical/Papers/AffineSchemes.agda b/Cubical/Papers/AffineSchemes.agda index 5d679863bd..5ff4d80c9f 100644 --- a/Cubical/Papers/AffineSchemes.agda +++ b/Cubical/Papers/AffineSchemes.agda @@ -41,7 +41,7 @@ module Localization = L.Loc import Cubical.Algebra.CommRing.Localisation.UniversalProperty as LocalizationUnivProp import Cubical.Algebra.CommRing.Localisation.InvertingElements as LocalizationInvEl import Cubical.Algebra.CommAlgebra as R-Algs -import Cubical.Algebra.CommAlgebra.Localisation as LocalizationR-Alg +import Cubical.Algebra.CommAlgebraAlt.AsModule.Localisation as LocalizationR-Alg -- 3.2: The Zariski Lattice From 33c060b618eedd6af51784ad492a28f41f419f03 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Wed, 7 Aug 2024 18:36:59 +0200 Subject: [PATCH 60/83] category structure --- Cubical/Algebra/CommAlgebraAlt/Base.agda | 7 ++-- .../Algebra/CommAlgebraAlt/Properties.agda | 33 +++++++++++++++++++ 2 files changed, 38 insertions(+), 2 deletions(-) create mode 100644 Cubical/Algebra/CommAlgebraAlt/Properties.agda diff --git a/Cubical/Algebra/CommAlgebraAlt/Base.agda b/Cubical/Algebra/CommAlgebraAlt/Base.agda index 3cacda5f4e..f10fa406e6 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/Base.agda @@ -34,6 +34,8 @@ module _ {R : CommRing ℓ} where idCAlgHom : (A : CommAlgebra R ℓ') → CommAlgebraHom A A idCAlgHom A = withOpaqueStr $ (idCommRingHom (fst A)) , CommRingHom≡ refl + idCommAlgebraHom = idCAlgHom + compCommAlgebraHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} → (g : CommAlgebraHom B C) → (f : CommAlgebraHom A B) → CommAlgebraHom A C @@ -58,13 +60,14 @@ module _ {R : CommRing ℓ} where _$ca_ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} → (φ : CommAlgebraHom A B) → (x : ⟨ A ⟩ₐ) → ⟨ B ⟩ₐ φ $ca x = φ .fst .fst x - _∘ca_ : {A : CommAlgebra R ℓ} {B : CommAlgebra R ℓ'} {C : CommAlgebra R ℓ''} + infixr 9 _∘ca_ -- same as functions + _∘ca_ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} → CommAlgebraHom B C → CommAlgebraHom A B → CommAlgebraHom A C g ∘ca f = compCommAlgebraHom g f opaque CommAlgebraHom≡ : - {A : CommAlgebra R ℓ} {B : CommAlgebra R ℓ'} {f g : CommAlgebraHom A B} + {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {f g : CommAlgebraHom A B} → f .fst .fst ≡ g .fst .fst → f ≡ g CommAlgebraHom≡ {B = B} p = diff --git a/Cubical/Algebra/CommAlgebraAlt/Properties.agda b/Cubical/Algebra/CommAlgebraAlt/Properties.agda new file mode 100644 index 0000000000..893b6b386f --- /dev/null +++ b/Cubical/Algebra/CommAlgebraAlt/Properties.agda @@ -0,0 +1,33 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebraAlt.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Structure using (⟨_⟩; withOpaqueStr) + +open import Cubical.Data.Sigma + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommAlgebraAlt.Base + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + +module _ {R : CommRing ℓ} where + + opaque + compIdCommAlgebraHom : {A B : CommAlgebra R ℓ'} (f : CommAlgebraHom A B) + → (f ∘ca idCAlgHom A) ≡ f + compIdCommAlgebraHom f = CommAlgebraHom≡ refl + + idCompCommAlgebraHom : {A B : CommAlgebra R ℓ'} (f : CommAlgebraHom A B) + → (idCommAlgebraHom B) ∘ca f ≡ f + idCompCommAlgebraHom f = CommAlgebraHom≡ refl + + compAssocCommAlgebraHom : {A B C D : CommAlgebra R ℓ'} + (f : CommAlgebraHom A B) (g : CommAlgebraHom B C) (h : CommAlgebraHom C D) + → h ∘ca (g ∘ca f) ≡ (h ∘ca g) ∘ca f + compAssocCommAlgebraHom f g h = CommAlgebraHom≡ refl From 5ec0ba45d66cb9a88ffe3db963f5b5e252b9701b Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 8 Aug 2024 12:04:04 +0200 Subject: [PATCH 61/83] fix ref --- Cubical/Experiments/ZariskiLatticeBasicOpens.agda | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/Cubical/Experiments/ZariskiLatticeBasicOpens.agda b/Cubical/Experiments/ZariskiLatticeBasicOpens.agda index 693831b702..c642600eeb 100644 --- a/Cubical/Experiments/ZariskiLatticeBasicOpens.agda +++ b/Cubical/Experiments/ZariskiLatticeBasicOpens.agda @@ -31,9 +31,8 @@ open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Localisation.Base open import Cubical.Algebra.CommRing.Localisation.UniversalProperty open import Cubical.Algebra.CommRing.Localisation.InvertingElements -open import Cubical.Algebra.CommAlgebra.Base -open import Cubical.Algebra.CommAlgebra.Properties -open import Cubical.Algebra.CommAlgebra.Localisation +open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Localisation open import Cubical.Tactics.CommRingSolver open import Cubical.Algebra.Semilattice From acfd7470a3a4a4252917dca76ad0d504b4c8bcc9 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 8 Aug 2024 12:13:01 +0200 Subject: [PATCH 62/83] fix/move algebra as module --- Cubical/Algebra/CommAlgebraAlt/AsModule.agda | 5 +++++ .../CommAlgebraAlt/AsModule/FreeCommAlgebra/OnCoproduct.agda | 4 ++-- 2 files changed, 7 insertions(+), 2 deletions(-) create mode 100644 Cubical/Algebra/CommAlgebraAlt/AsModule.agda diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule.agda new file mode 100644 index 0000000000..406d5ff81e --- /dev/null +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebraAlt.AsModule where + +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base public +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Properties public diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/OnCoproduct.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/OnCoproduct.agda index 61fcb8c920..82dc9d441d 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/OnCoproduct.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/OnCoproduct.agda @@ -7,7 +7,7 @@ where '⊎' is the disjoint sum. -} -module Cubical.Algebra.CommAlgebra.FreeCommAlgebra.OnCoproduct where +module Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra.OnCoproduct where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -24,7 +24,7 @@ open import Cubical.Algebra.Ring open import Cubical.Algebra.Algebra open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommAlgebra -open import Cubical.Algebra.CommAlgebra.FreeCommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra private variable ℓ ℓ' : Level From 0f5053e24a3349b96bed875f93deedf368359aad Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 8 Aug 2024 12:28:48 +0200 Subject: [PATCH 63/83] refs --- Cubical/Algebra/Polynomials/TypevariateHIT/Base.agda | 2 +- .../Polynomials/UnivariateList/UniversalProperty.agda | 2 +- .../ZariskiLattice/StructureSheafPullback.agda | 8 ++++---- 3 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Cubical/Algebra/Polynomials/TypevariateHIT/Base.agda b/Cubical/Algebra/Polynomials/TypevariateHIT/Base.agda index 44a2232770..c19fb0b5db 100644 --- a/Cubical/Algebra/Polynomials/TypevariateHIT/Base.agda +++ b/Cubical/Algebra/Polynomials/TypevariateHIT/Base.agda @@ -13,4 +13,4 @@ The typevariate polynomials are constructed as a free commutative algebra on the They are justified by a proof of the universal property of a free commutative algebra. -} -open import Cubical.Algebra.CommAlgebra.FreeCommAlgebra public +open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra public diff --git a/Cubical/Algebra/Polynomials/UnivariateList/UniversalProperty.agda b/Cubical/Algebra/Polynomials/UnivariateList/UniversalProperty.agda index a4167cfcfd..9847c6292e 100644 --- a/Cubical/Algebra/Polynomials/UnivariateList/UniversalProperty.agda +++ b/Cubical/Algebra/Polynomials/UnivariateList/UniversalProperty.agda @@ -5,4 +5,4 @@ module Cubical.Algebra.Polynomials.UnivariateList.UniversalProperty where Export an CommAlgebra-Instance for the UnivariateList-Polynomials. Also export the universal property with respect to not-necessarily commutative algebras. -} -open import Cubical.Algebra.CommAlgebra.UnivariatePolyList public +open import Cubical.Algebra.CommAlgebraAlt.AsModule.UnivariatePolyList public diff --git a/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheafPullback.agda b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheafPullback.agda index 80ac11afc6..e067b98047 100644 --- a/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheafPullback.agda +++ b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheafPullback.agda @@ -48,10 +48,10 @@ open import Cubical.Algebra.CommRing.Localisation.Base open import Cubical.Algebra.CommRing.Localisation.UniversalProperty open import Cubical.Algebra.CommRing.Localisation.InvertingElements open import Cubical.Algebra.CommRing.Localisation.PullbackSquare -open import Cubical.Algebra.CommAlgebra.Base -open import Cubical.Algebra.CommAlgebra.Properties -open import Cubical.Algebra.CommAlgebra.Localisation -open import Cubical.Algebra.CommAlgebra.Instances.Unit +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Properties +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Localisation +open import Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Unit open import Cubical.Tactics.CommRingSolver open import Cubical.Algebra.Semilattice open import Cubical.Algebra.Lattice From 367e8bd900b3be3a6e420fff1e88aeb99aeadf93 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 8 Aug 2024 12:35:31 +0200 Subject: [PATCH 64/83] remove old CommAlgebra.agda, move all refs --- Cubical/Algebra/CommAlgebra.agda | 5 ----- Cubical/Algebra/CommAlgebraAlt/AsModule/FGIdeal.agda | 2 +- Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Base.agda | 2 +- .../Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Instances.agda | 2 +- .../CommAlgebraAlt/AsModule/FreeCommAlgebra/OnCoproduct.agda | 2 +- .../CommAlgebraAlt/AsModule/FreeCommAlgebra/Properties.agda | 2 +- Cubical/Algebra/CommAlgebraAlt/AsModule/Ideal.agda | 2 +- .../Algebra/CommAlgebraAlt/AsModule/Instances/Initial.agda | 2 +- .../Algebra/CommAlgebraAlt/AsModule/LocalisationAlgebra.agda | 2 +- Cubical/Algebra/CommAlgebraAlt/AsModule/QuotientAlgebra.agda | 2 +- Cubical/Algebra/CommAlgebraAlt/AsModule/Subalgebra.agda | 2 +- .../Algebra/CommAlgebraAlt/AsModule/UnivariatePolyList.agda | 2 +- .../Polynomials/TypevariateHIT/EquivUnivariateListPoly.agda | 2 +- Cubical/Categories/Instances/CommAlgebras.agda | 2 +- Cubical/Papers/AffineSchemes.agda | 2 +- Cubical/Tactics/CommRingSolver/Examples.agda | 2 +- 16 files changed, 15 insertions(+), 20 deletions(-) delete mode 100644 Cubical/Algebra/CommAlgebra.agda diff --git a/Cubical/Algebra/CommAlgebra.agda b/Cubical/Algebra/CommAlgebra.agda deleted file mode 100644 index 8600ad3ada..0000000000 --- a/Cubical/Algebra/CommAlgebra.agda +++ /dev/null @@ -1,5 +0,0 @@ -{-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra where - -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base public -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Properties public diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/FGIdeal.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/FGIdeal.agda index e9464a4059..a7c5d2bd45 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/FGIdeal.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/FGIdeal.agda @@ -13,7 +13,7 @@ open import Cubical.Algebra.CommRing.FGIdeal using () renaming (generatedIdeal to generatedIdealCommRing; indInIdeal to ringIncInIdeal; 0FGIdeal to 0FGIdealCommRing) -open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule open import Cubical.Algebra.CommAlgebraAlt.AsModule.Ideal private diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Base.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Base.agda index 77615ad8c9..89f031326f 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Base.agda @@ -18,7 +18,7 @@ open import Cubical.HITs.PropositionalTruncation open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.FGIdeal using (inclOfFGIdeal) -open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra renaming (inducedHom to freeInducedHom) open import Cubical.Algebra.CommAlgebraAlt.AsModule.QuotientAlgebra diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Instances.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Instances.agda index 1d16ad4262..08bf00c5b8 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Instances.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Instances.agda @@ -27,7 +27,7 @@ open import Cubical.HITs.PropositionalTruncation open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.FGIdeal using (inclOfFGIdeal) -open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra renaming (inducedHom to freeInducedHom) open import Cubical.Algebra.CommAlgebraAlt.AsModule.QuotientAlgebra diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/OnCoproduct.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/OnCoproduct.agda index 82dc9d441d..293f5fcf45 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/OnCoproduct.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/OnCoproduct.agda @@ -23,7 +23,7 @@ open import Cubical.Data.Sigma open import Cubical.Algebra.Ring open import Cubical.Algebra.Algebra open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra private variable diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Properties.agda index 626731008a..aa0aa542b0 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Properties.agda @@ -30,7 +30,7 @@ open import Cubical.Data.Sigma.Properties using (Σ≡Prop) open import Cubical.HITs.SetTruncation open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra.Base open import Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Initial open import Cubical.Algebra.Algebra diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/Ideal.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/Ideal.agda index 3ac7be5f81..3fc1c16b04 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/Ideal.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/Ideal.agda @@ -7,7 +7,7 @@ open import Cubical.Foundations.Powerset open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Ideal renaming (IdealsIn to IdealsInCommRing; makeIdeal to makeIdealCommRing) -open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule open import Cubical.Algebra.Ring open import Cubical.Data.Unit diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Initial.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Initial.agda index d9b4708eb5..9c8bdb8871 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Initial.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Initial.agda @@ -12,7 +12,7 @@ open import Cubical.Algebra.CommRing open import Cubical.Algebra.Ring open import Cubical.Algebra.Algebra.Base using (IsAlgebraHom) open import Cubical.Algebra.Algebra.Properties -open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule import Cubical.Algebra.Algebra.Properties open AlgebraHoms diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/LocalisationAlgebra.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/LocalisationAlgebra.agda index 88179e9626..b7919a193e 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/LocalisationAlgebra.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/LocalisationAlgebra.agda @@ -22,7 +22,7 @@ import Cubical.HITs.SetQuotients as SQ import Cubical.HITs.PropositionalTruncation as PropTrunc open import Cubical.Algebra.Algebra -open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule open import Cubical.Algebra.CommAlgebraAlt.AsModule.Subalgebra open import Cubical.Algebra.CommRing as CommRing hiding (_ˣ;module Units) open import Cubical.Algebra.CommRing.Localisation using (isMultClosedSubset) diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/QuotientAlgebra.agda index ccb19527ec..22b849c462 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/QuotientAlgebra.agda @@ -15,7 +15,7 @@ open import Cubical.Algebra.CommRing import Cubical.Algebra.CommRing.Quotient as CommRing import Cubical.Algebra.Ring.Quotient as Ring open import Cubical.Algebra.CommRing.Ideal hiding (IdealsIn) -open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule open import Cubical.Algebra.CommAlgebraAlt.AsModule.Ideal open import Cubical.Algebra.CommAlgebraAlt.AsModule.Kernel open import Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Unit diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/Subalgebra.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/Subalgebra.agda index a386292275..ce33d50538 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/Subalgebra.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/Subalgebra.agda @@ -4,7 +4,7 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure open import Cubical.Algebra.Algebra -open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule open import Cubical.Algebra.CommRing module Cubical.Algebra.CommAlgebraAlt.AsModule.Subalgebra diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/UnivariatePolyList.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/UnivariatePolyList.agda index ddd41b50db..ba16bd1081 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/UnivariatePolyList.agda +++ b/Cubical/Algebra/CommAlgebraAlt/AsModule/UnivariatePolyList.agda @@ -11,7 +11,7 @@ open import Cubical.Data.Sigma open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing open import Cubical.Algebra.AbGroup -open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule open import Cubical.Algebra.Algebra open import Cubical.Algebra.CommRing.Instances.Polynomials.UnivariatePolyList open import Cubical.Algebra.Polynomials.UnivariateList.Base diff --git a/Cubical/Algebra/Polynomials/TypevariateHIT/EquivUnivariateListPoly.agda b/Cubical/Algebra/Polynomials/TypevariateHIT/EquivUnivariateListPoly.agda index ac9f570c94..f8a8a11dc5 100644 --- a/Cubical/Algebra/Polynomials/TypevariateHIT/EquivUnivariateListPoly.agda +++ b/Cubical/Algebra/Polynomials/TypevariateHIT/EquivUnivariateListPoly.agda @@ -10,7 +10,7 @@ open import Cubical.Data.Unit open import Cubical.Algebra.CommRing open import Cubical.Algebra.Algebra -open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule open import Cubical.Algebra.Polynomials.TypevariateHIT renaming (inducedHomUnique to inducedHomUniqueHIT; isIdByUMP to isIdByUMP-HIT) diff --git a/Cubical/Categories/Instances/CommAlgebras.agda b/Cubical/Categories/Instances/CommAlgebras.agda index 3ac204bb19..6b8df7cedb 100644 --- a/Cubical/Categories/Instances/CommAlgebras.agda +++ b/Cubical/Categories/Instances/CommAlgebras.agda @@ -14,7 +14,7 @@ open import Cubical.Data.Sigma open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing open import Cubical.Algebra.Algebra -open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule open import Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Unit open import Cubical.Categories.Category diff --git a/Cubical/Papers/AffineSchemes.agda b/Cubical/Papers/AffineSchemes.agda index 5ff4d80c9f..cffebdbc66 100644 --- a/Cubical/Papers/AffineSchemes.agda +++ b/Cubical/Papers/AffineSchemes.agda @@ -40,7 +40,7 @@ module Localization = L.Loc import Cubical.Algebra.CommRing.Localisation.UniversalProperty as LocalizationUnivProp import Cubical.Algebra.CommRing.Localisation.InvertingElements as LocalizationInvEl -import Cubical.Algebra.CommAlgebra as R-Algs +import Cubical.Algebra.CommAlgebraAlt.AsModule as R-Algs import Cubical.Algebra.CommAlgebraAlt.AsModule.Localisation as LocalizationR-Alg diff --git a/Cubical/Tactics/CommRingSolver/Examples.agda b/Cubical/Tactics/CommRingSolver/Examples.agda index 531af32e9a..a22f3f4079 100644 --- a/Cubical/Tactics/CommRingSolver/Examples.agda +++ b/Cubical/Tactics/CommRingSolver/Examples.agda @@ -10,7 +10,7 @@ open import Cubical.Data.Nat using (ℕ; suc; zero) open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Instances.Int -open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule open import Cubical.Tactics.CommRingSolver open import Cubical.Tactics.CommRingSolver.RawAlgebra using (scalar) From 71d44b995b19675a8450162615af7e970dfacede Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 8 Aug 2024 13:44:00 +0200 Subject: [PATCH 65/83] fix --- Cubical/Experiments/ZariskiLatticeBasicOpens.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Cubical/Experiments/ZariskiLatticeBasicOpens.agda b/Cubical/Experiments/ZariskiLatticeBasicOpens.agda index c642600eeb..4da00892d9 100644 --- a/Cubical/Experiments/ZariskiLatticeBasicOpens.agda +++ b/Cubical/Experiments/ZariskiLatticeBasicOpens.agda @@ -31,7 +31,7 @@ open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Localisation.Base open import Cubical.Algebra.CommRing.Localisation.UniversalProperty open import Cubical.Algebra.CommRing.Localisation.InvertingElements -open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebraAlt.AsModule open import Cubical.Algebra.CommAlgebraAlt.AsModule.Localisation open import Cubical.Tactics.CommRingSolver open import Cubical.Algebra.Semilattice From 7cc50a655e02198fafb0fc5ac40a2e80afb7d526 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 8 Aug 2024 14:01:00 +0200 Subject: [PATCH 66/83] make the new CommAlgebra the default --- Cubical/Algebra/CommAlgebra.agda | 5 +++++ Cubical/Algebra/CommAlgebra/AsModule.agda | 5 +++++ .../AsModule/Base.agda | 2 +- .../AsModule/FGIdeal.agda | 6 +++--- .../AsModule/FPAlgebra.agda | 4 ++-- .../AsModule/FPAlgebra/Base.agda | 14 ++++++------- .../AsModule/FPAlgebra/Instances.agda | 20 +++++++++---------- .../CommAlgebra/AsModule/FreeCommAlgebra.agda | 6 ++++++ .../AsModule/FreeCommAlgebra/Base.agda | 4 ++-- .../AsModule/FreeCommAlgebra/OnCoproduct.agda | 6 +++--- .../AsModule/FreeCommAlgebra/Properties.agda | 8 ++++---- .../AsModule/Ideal.agda | 4 ++-- .../AsModule/Instances/Initial.agda | 4 ++-- .../AsModule/Instances/Pointwise.agda | 4 ++-- .../AsModule/Instances/Unit.agda | 4 ++-- .../AsModule/Kernel.agda | 8 ++++---- .../AsModule/Localisation.agda | 6 +++--- .../AsModule/LocalisationAlgebra.agda | 6 +++--- .../AsModule/Properties.agda | 4 ++-- .../AsModule/QuotientAlgebra.agda | 10 +++++----- .../AsModule/Subalgebra.agda | 4 ++-- .../AsModule/UnivariatePolyList.agda | 4 ++-- .../{CommAlgebraAlt => CommAlgebra}/Base.agda | 2 +- .../FGIdeal.agda | 6 +++--- .../Ideal.agda | 4 ++-- .../Instances/Polynomials.agda | 4 ++-- .../Instances/Unit.agda | 4 ++-- .../Kernel.agda | 6 +++--- .../Localisation.agda | 6 +++--- .../Properties.agda | 4 ++-- .../QuotientAlgebra.agda | 10 +++++----- .../Univalence.agda | 4 ++-- Cubical/Algebra/CommAlgebraAlt.agda | 5 ----- Cubical/Algebra/CommAlgebraAlt/AsModule.agda | 5 ----- .../AsModule/FreeCommAlgebra.agda | 6 ------ .../Polynomials/TypevariateHIT/Base.agda | 2 +- .../EquivUnivariateListPoly.agda | 2 +- .../UnivariateList/UniversalProperty.agda | 2 +- .../ZariskiLattice/StructureSheaf.agda | 6 +++--- .../StructureSheafPullback.agda | 8 ++++---- .../Categories/Instances/CommAlgebras.agda | 4 ++-- .../Experiments/ZariskiLatticeBasicOpens.agda | 4 ++-- Cubical/Papers/AffineSchemes.agda | 4 ++-- Cubical/Tactics/CommRingSolver/Examples.agda | 2 +- 44 files changed, 119 insertions(+), 119 deletions(-) create mode 100644 Cubical/Algebra/CommAlgebra.agda create mode 100644 Cubical/Algebra/CommAlgebra/AsModule.agda rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/Base.agda (99%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/FGIdeal.agda (87%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/FPAlgebra.agda (67%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/FPAlgebra/Base.agda (96%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/FPAlgebra/Instances.agda (92%) create mode 100644 Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra.agda rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/FreeCommAlgebra/Base.agda (96%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/FreeCommAlgebra/OnCoproduct.agda (98%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/FreeCommAlgebra/Properties.agda (98%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/Ideal.agda (90%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/Instances/Initial.agda (97%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/Instances/Pointwise.agda (90%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/Instances/Unit.agda (94%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/Kernel.agda (69%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/Localisation.agda (98%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/LocalisationAlgebra.agda (98%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/Properties.agda (99%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/QuotientAlgebra.agda (96%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/Subalgebra.agda (94%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/AsModule/UnivariatePolyList.agda (98%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/Base.agda (99%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/FGIdeal.agda (88%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/Ideal.agda (91%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/Instances/Polynomials.agda (98%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/Instances/Unit.agda (93%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/Kernel.agda (80%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/Localisation.agda (98%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/Properties.agda (91%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/QuotientAlgebra.agda (96%) rename Cubical/Algebra/{CommAlgebraAlt => CommAlgebra}/Univalence.agda (95%) delete mode 100644 Cubical/Algebra/CommAlgebraAlt.agda delete mode 100644 Cubical/Algebra/CommAlgebraAlt/AsModule.agda delete mode 100644 Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra.agda diff --git a/Cubical/Algebra/CommAlgebra.agda b/Cubical/Algebra/CommAlgebra.agda new file mode 100644 index 0000000000..d150261beb --- /dev/null +++ b/Cubical/Algebra/CommAlgebra.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra where + +open import Cubical.Algebra.CommAlgebra.Base public +open import Cubical.Algebra.CommAlgebra.Properties public diff --git a/Cubical/Algebra/CommAlgebra/AsModule.agda b/Cubical/Algebra/CommAlgebra/AsModule.agda new file mode 100644 index 0000000000..215f11888d --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/AsModule.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.AsModule where + +open import Cubical.Algebra.CommAlgebra.AsModule.Base public +open import Cubical.Algebra.CommAlgebra.AsModule.Properties public diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/Base.agda b/Cubical/Algebra/CommAlgebra/AsModule/Base.agda similarity index 99% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/Base.agda rename to Cubical/Algebra/CommAlgebra/AsModule/Base.agda index 36addd0c39..ae6602f03b 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/Base.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/Base.agda @@ -2,7 +2,7 @@ {- This used to be the default definition of CommAlgebra. -} -module Cubical.Algebra.CommAlgebraAlt.AsModule.Base where +module Cubical.Algebra.CommAlgebra.AsModule.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/FGIdeal.agda b/Cubical/Algebra/CommAlgebra/AsModule/FGIdeal.agda similarity index 87% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/FGIdeal.agda rename to Cubical/Algebra/CommAlgebra/AsModule/FGIdeal.agda index a7c5d2bd45..8bfbf30385 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/FGIdeal.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/FGIdeal.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule.FGIdeal where +module Cubical.Algebra.CommAlgebra.AsModule.FGIdeal where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure open import Cubical.Foundations.Powerset @@ -13,8 +13,8 @@ open import Cubical.Algebra.CommRing.FGIdeal using () renaming (generatedIdeal to generatedIdealCommRing; indInIdeal to ringIncInIdeal; 0FGIdeal to 0FGIdealCommRing) -open import Cubical.Algebra.CommAlgebraAlt.AsModule -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Ideal +open import Cubical.Algebra.CommAlgebra.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule.Ideal private variable diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra.agda b/Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra.agda similarity index 67% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra.agda rename to Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra.agda index 23d84f5ba9..36605f2a27 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra.agda @@ -9,6 +9,6 @@ Our definition is more explicit. -} {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule.FPAlgebra where +module Cubical.Algebra.CommAlgebra.AsModule.FPAlgebra where -open import Cubical.Algebra.CommAlgebraAlt.AsModule.FPAlgebra.Base public +open import Cubical.Algebra.CommAlgebra.AsModule.FPAlgebra.Base public diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra/Base.agda similarity index 96% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Base.agda rename to Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra/Base.agda index 89f031326f..5b2c16d99f 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra/Base.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule.FPAlgebra.Base where +module Cubical.Algebra.CommAlgebra.AsModule.FPAlgebra.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -18,14 +18,14 @@ open import Cubical.HITs.PropositionalTruncation open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.FGIdeal using (inclOfFGIdeal) -open import Cubical.Algebra.CommAlgebraAlt.AsModule -open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra +open import Cubical.Algebra.CommAlgebra.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule.FreeCommAlgebra renaming (inducedHom to freeInducedHom) -open import Cubical.Algebra.CommAlgebraAlt.AsModule.QuotientAlgebra +open import Cubical.Algebra.CommAlgebra.AsModule.QuotientAlgebra renaming (inducedHom to quotientInducedHom) -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Ideal -open import Cubical.Algebra.CommAlgebraAlt.AsModule.FGIdeal -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Kernel +open import Cubical.Algebra.CommAlgebra.AsModule.Ideal +open import Cubical.Algebra.CommAlgebra.AsModule.FGIdeal +open import Cubical.Algebra.CommAlgebra.AsModule.Kernel open import Cubical.Algebra.Algebra.Properties open import Cubical.Algebra.Algebra diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Instances.agda b/Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra/Instances.agda similarity index 92% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Instances.agda rename to Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra/Instances.agda index 08bf00c5b8..e206f60c9c 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/FPAlgebra/Instances.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra/Instances.agda @@ -8,7 +8,7 @@ * R/⟨x⟩ (as special case of the above) -} {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule.FPAlgebra.Instances where +module Cubical.Algebra.CommAlgebra.AsModule.FPAlgebra.Instances where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -27,20 +27,20 @@ open import Cubical.HITs.PropositionalTruncation open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.FGIdeal using (inclOfFGIdeal) -open import Cubical.Algebra.CommAlgebraAlt.AsModule -open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra +open import Cubical.Algebra.CommAlgebra.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule.FreeCommAlgebra renaming (inducedHom to freeInducedHom) -open import Cubical.Algebra.CommAlgebraAlt.AsModule.QuotientAlgebra +open import Cubical.Algebra.CommAlgebra.AsModule.QuotientAlgebra renaming (inducedHom to quotientInducedHom) -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Ideal using (IdealsIn) -open import Cubical.Algebra.CommAlgebraAlt.AsModule.FGIdeal -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Initial -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Unit +open import Cubical.Algebra.CommAlgebra.AsModule.Ideal using (IdealsIn) +open import Cubical.Algebra.CommAlgebra.AsModule.FGIdeal +open import Cubical.Algebra.CommAlgebra.AsModule.Instances.Initial +open import Cubical.Algebra.CommAlgebra.AsModule.Instances.Unit renaming (UnitCommAlgebra to TerminalCAlg) -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Kernel +open import Cubical.Algebra.CommAlgebra.AsModule.Kernel open import Cubical.Algebra.Algebra -open import Cubical.Algebra.CommAlgebraAlt.AsModule.FPAlgebra.Base +open import Cubical.Algebra.CommAlgebra.AsModule.FPAlgebra.Base private variable diff --git a/Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra.agda b/Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra.agda new file mode 100644 index 0000000000..1a9b94af54 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} + +module Cubical.Algebra.CommAlgebra.AsModule.FreeCommAlgebra where + +open import Cubical.Algebra.CommAlgebra.AsModule.FreeCommAlgebra.Base public +open import Cubical.Algebra.CommAlgebra.AsModule.FreeCommAlgebra.Properties public diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/Base.agda similarity index 96% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Base.agda rename to Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/Base.agda index f45047d1ae..8ce217225c 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/Base.agda @@ -1,6 +1,6 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra.Base where +module Cubical.Algebra.CommAlgebra.AsModule.FreeCommAlgebra.Base where {- The free commutative algebra over a commutative ring, or in other words the ring of polynomials with coefficients in a given ring. @@ -21,7 +21,7 @@ module Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra.Base where open import Cubical.Foundations.Prelude open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base +open import Cubical.Algebra.CommAlgebra.AsModule.Base private variable diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/OnCoproduct.agda b/Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/OnCoproduct.agda similarity index 98% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/OnCoproduct.agda rename to Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/OnCoproduct.agda index 293f5fcf45..03eb8d10a4 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/OnCoproduct.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/OnCoproduct.agda @@ -7,7 +7,7 @@ where '⊎' is the disjoint sum. -} -module Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra.OnCoproduct where +module Cubical.Algebra.CommAlgebra.AsModule.FreeCommAlgebra.OnCoproduct where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -23,8 +23,8 @@ open import Cubical.Data.Sigma open import Cubical.Algebra.Ring open import Cubical.Algebra.Algebra open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommAlgebraAlt.AsModule -open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra +open import Cubical.Algebra.CommAlgebra.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule.FreeCommAlgebra private variable ℓ ℓ' : Level diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/Properties.agda similarity index 98% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Properties.agda rename to Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/Properties.agda index aa0aa542b0..791a4deddd 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/Properties.agda @@ -1,6 +1,6 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra.Properties where +module Cubical.Algebra.CommAlgebra.AsModule.FreeCommAlgebra.Properties where {- This file contains * an elimination principle for proving some proposition for all elements of R[I] @@ -30,9 +30,9 @@ open import Cubical.Data.Sigma.Properties using (Σ≡Prop) open import Cubical.HITs.SetTruncation open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommAlgebraAlt.AsModule -open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra.Base -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Initial +open import Cubical.Algebra.CommAlgebra.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule.FreeCommAlgebra.Base +open import Cubical.Algebra.CommAlgebra.AsModule.Instances.Initial open import Cubical.Algebra.Algebra open import Cubical.Algebra.Module using (module ModuleTheory) diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/Ideal.agda b/Cubical/Algebra/CommAlgebra/AsModule/Ideal.agda similarity index 90% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/Ideal.agda rename to Cubical/Algebra/CommAlgebra/AsModule/Ideal.agda index 3fc1c16b04..623d038c26 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/Ideal.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/Ideal.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule.Ideal where +module Cubical.Algebra.CommAlgebra.AsModule.Ideal where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Powerset @@ -7,7 +7,7 @@ open import Cubical.Foundations.Powerset open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Ideal renaming (IdealsIn to IdealsInCommRing; makeIdeal to makeIdealCommRing) -open import Cubical.Algebra.CommAlgebraAlt.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule open import Cubical.Algebra.Ring open import Cubical.Data.Unit diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Initial.agda b/Cubical/Algebra/CommAlgebra/AsModule/Instances/Initial.agda similarity index 97% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Initial.agda rename to Cubical/Algebra/CommAlgebra/AsModule/Instances/Initial.agda index 9c8bdb8871..34678b461c 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Initial.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/Instances/Initial.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Initial where +module Cubical.Algebra.CommAlgebra.AsModule.Instances.Initial where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels @@ -12,7 +12,7 @@ open import Cubical.Algebra.CommRing open import Cubical.Algebra.Ring open import Cubical.Algebra.Algebra.Base using (IsAlgebraHom) open import Cubical.Algebra.Algebra.Properties -open import Cubical.Algebra.CommAlgebraAlt.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule import Cubical.Algebra.Algebra.Properties open AlgebraHoms diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Pointwise.agda b/Cubical/Algebra/CommAlgebra/AsModule/Instances/Pointwise.agda similarity index 90% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Pointwise.agda rename to Cubical/Algebra/CommAlgebra/AsModule/Instances/Pointwise.agda index fe3f048572..edbe6beef0 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Pointwise.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/Instances/Pointwise.agda @@ -1,12 +1,12 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Pointwise where +module Cubical.Algebra.CommAlgebra.AsModule.Instances.Pointwise where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Algebra.CommRing.Base open import Cubical.Algebra.CommRing.Instances.Pointwise -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base +open import Cubical.Algebra.CommAlgebra.AsModule.Base open import Cubical.Algebra.Algebra using (IsAlgebraHom) private diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Unit.agda b/Cubical/Algebra/CommAlgebra/AsModule/Instances/Unit.agda similarity index 94% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Unit.agda rename to Cubical/Algebra/CommAlgebra/AsModule/Instances/Unit.agda index 710ec2069a..e73e3c8d1f 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/Instances/Unit.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/Instances/Unit.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Unit where +module Cubical.Algebra.CommAlgebra.AsModule.Instances.Unit where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -9,7 +9,7 @@ open import Cubical.Data.Unit open import Cubical.Data.Sigma.Properties using (Σ≡Prop) open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base +open import Cubical.Algebra.CommAlgebra.AsModule.Base open import Cubical.Algebra.CommRing.Instances.Unit open import Cubical.Algebra.Algebra.Base using (IsAlgebraHom) open import Cubical.Tactics.CommRingSolver diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/Kernel.agda b/Cubical/Algebra/CommAlgebra/AsModule/Kernel.agda similarity index 69% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/Kernel.agda rename to Cubical/Algebra/CommAlgebra/AsModule/Kernel.agda index fa8ab8b4b9..6c98c8a8f0 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/Kernel.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/Kernel.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule.Kernel where +module Cubical.Algebra.CommAlgebra.AsModule.Kernel where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -7,9 +7,9 @@ open import Cubical.Foundations.Equiv open import Cubical.Algebra.CommRing.Base open import Cubical.Algebra.CommRing.Ideal using (Ideal→CommIdeal) open import Cubical.Algebra.Ring.Kernel using () renaming (kernelIdeal to ringKernel) -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Properties -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Ideal +open import Cubical.Algebra.CommAlgebra.AsModule.Base +open import Cubical.Algebra.CommAlgebra.AsModule.Properties +open import Cubical.Algebra.CommAlgebra.AsModule.Ideal private variable diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/Localisation.agda b/Cubical/Algebra/CommAlgebra/AsModule/Localisation.agda similarity index 98% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/Localisation.agda rename to Cubical/Algebra/CommAlgebra/AsModule/Localisation.agda index a3045e0ee7..b17ed15f7a 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/Localisation.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/Localisation.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --lossy-unification #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule.Localisation where +module Cubical.Algebra.CommAlgebra.AsModule.Localisation where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -28,8 +28,8 @@ open import Cubical.Algebra.CommRing.RadicalIdeal open import Cubical.Algebra.CommRing.Localisation open import Cubical.Algebra.Ring open import Cubical.Algebra.Algebra -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Properties +open import Cubical.Algebra.CommAlgebra.AsModule.Base +open import Cubical.Algebra.CommAlgebra.AsModule.Properties open import Cubical.Tactics.CommRingSolver diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/LocalisationAlgebra.agda b/Cubical/Algebra/CommAlgebra/AsModule/LocalisationAlgebra.agda similarity index 98% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/LocalisationAlgebra.agda rename to Cubical/Algebra/CommAlgebra/AsModule/LocalisationAlgebra.agda index b7919a193e..d42248a94d 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/LocalisationAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/LocalisationAlgebra.agda @@ -22,13 +22,13 @@ import Cubical.HITs.SetQuotients as SQ import Cubical.HITs.PropositionalTruncation as PropTrunc open import Cubical.Algebra.Algebra -open import Cubical.Algebra.CommAlgebraAlt.AsModule -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Subalgebra +open import Cubical.Algebra.CommAlgebra.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule.Subalgebra open import Cubical.Algebra.CommRing as CommRing hiding (_ˣ;module Units) open import Cubical.Algebra.CommRing.Localisation using (isMultClosedSubset) open import Cubical.Algebra.Ring -module Cubical.Algebra.CommAlgebraAlt.AsModule.LocalisationAlgebra +module Cubical.Algebra.CommAlgebra.AsModule.LocalisationAlgebra {ℓR : Level} (R : CommRing ℓR) where diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/Properties.agda b/Cubical/Algebra/CommAlgebra/AsModule/Properties.agda similarity index 99% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/Properties.agda rename to Cubical/Algebra/CommAlgebra/AsModule/Properties.agda index db2b937858..040c18d643 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/Properties.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --lossy-unification #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule.Properties where +module Cubical.Algebra.CommAlgebra.AsModule.Properties where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function @@ -26,7 +26,7 @@ open import Cubical.Algebra.Monoid open import Cubical.Algebra.CommRing open import Cubical.Algebra.Ring open import Cubical.Algebra.Algebra -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base +open import Cubical.Algebra.CommAlgebra.AsModule.Base open import Cubical.Algebra.CommRing using (CommRing→Ring) diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebra/AsModule/QuotientAlgebra.agda similarity index 96% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/QuotientAlgebra.agda rename to Cubical/Algebra/CommAlgebra/AsModule/QuotientAlgebra.agda index 22b849c462..6bd6d62ae4 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/QuotientAlgebra.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule.QuotientAlgebra where +module Cubical.Algebra.CommAlgebra.AsModule.QuotientAlgebra where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv @@ -15,10 +15,10 @@ open import Cubical.Algebra.CommRing import Cubical.Algebra.CommRing.Quotient as CommRing import Cubical.Algebra.Ring.Quotient as Ring open import Cubical.Algebra.CommRing.Ideal hiding (IdealsIn) -open import Cubical.Algebra.CommAlgebraAlt.AsModule -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Ideal -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Kernel -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Unit +open import Cubical.Algebra.CommAlgebra.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule.Ideal +open import Cubical.Algebra.CommAlgebra.AsModule.Kernel +open import Cubical.Algebra.CommAlgebra.AsModule.Instances.Unit open import Cubical.Algebra.Algebra.Base using (IsAlgebraHom; isPropIsAlgebraHom) open import Cubical.Algebra.Ring open import Cubical.Algebra.Ring.Ideal using (isIdeal) diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/Subalgebra.agda b/Cubical/Algebra/CommAlgebra/AsModule/Subalgebra.agda similarity index 94% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/Subalgebra.agda rename to Cubical/Algebra/CommAlgebra/AsModule/Subalgebra.agda index ce33d50538..8deb268d56 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/Subalgebra.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/Subalgebra.agda @@ -4,10 +4,10 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure open import Cubical.Algebra.Algebra -open import Cubical.Algebra.CommAlgebraAlt.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule open import Cubical.Algebra.CommRing -module Cubical.Algebra.CommAlgebraAlt.AsModule.Subalgebra +module Cubical.Algebra.CommAlgebra.AsModule.Subalgebra {ℓ ℓ' : Level} (R : CommRing ℓ) (A : CommAlgebra R ℓ') where diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/UnivariatePolyList.agda b/Cubical/Algebra/CommAlgebra/AsModule/UnivariatePolyList.agda similarity index 98% rename from Cubical/Algebra/CommAlgebraAlt/AsModule/UnivariatePolyList.agda rename to Cubical/Algebra/CommAlgebra/AsModule/UnivariatePolyList.agda index ba16bd1081..839e7ade9e 100644 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/UnivariatePolyList.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/UnivariatePolyList.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule.UnivariatePolyList where +module Cubical.Algebra.CommAlgebra.AsModule.UnivariatePolyList where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure @@ -11,7 +11,7 @@ open import Cubical.Data.Sigma open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing open import Cubical.Algebra.AbGroup -open import Cubical.Algebra.CommAlgebraAlt.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule open import Cubical.Algebra.Algebra open import Cubical.Algebra.CommRing.Instances.Polynomials.UnivariatePolyList open import Cubical.Algebra.Polynomials.UnivariateList.Base diff --git a/Cubical/Algebra/CommAlgebraAlt/Base.agda b/Cubical/Algebra/CommAlgebra/Base.agda similarity index 99% rename from Cubical/Algebra/CommAlgebraAlt/Base.agda rename to Cubical/Algebra/CommAlgebra/Base.agda index f10fa406e6..9c2629f6e2 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Base.agda +++ b/Cubical/Algebra/CommAlgebra/Base.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.Base where +module Cubical.Algebra.CommAlgebra.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv diff --git a/Cubical/Algebra/CommAlgebraAlt/FGIdeal.agda b/Cubical/Algebra/CommAlgebra/FGIdeal.agda similarity index 88% rename from Cubical/Algebra/CommAlgebraAlt/FGIdeal.agda rename to Cubical/Algebra/CommAlgebra/FGIdeal.agda index 6de95db747..1ef76185f2 100644 --- a/Cubical/Algebra/CommAlgebraAlt/FGIdeal.agda +++ b/Cubical/Algebra/CommAlgebra/FGIdeal.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.FGIdeal where +module Cubical.Algebra.CommAlgebra.FGIdeal where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure open import Cubical.Foundations.Powerset @@ -13,8 +13,8 @@ open import Cubical.Algebra.CommRing.FGIdeal using () renaming (generatedIdeal to generatedIdealCommRing; indInIdeal to ringIncInIdeal; 0FGIdeal to 0FGIdealCommRing) -open import Cubical.Algebra.CommAlgebraAlt.Base -open import Cubical.Algebra.CommAlgebraAlt.Ideal +open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebra.Ideal private variable diff --git a/Cubical/Algebra/CommAlgebraAlt/Ideal.agda b/Cubical/Algebra/CommAlgebra/Ideal.agda similarity index 91% rename from Cubical/Algebra/CommAlgebraAlt/Ideal.agda rename to Cubical/Algebra/CommAlgebra/Ideal.agda index 47eeebaabc..f80a3b70b3 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Ideal.agda +++ b/Cubical/Algebra/CommAlgebra/Ideal.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.Ideal where +module Cubical.Algebra.CommAlgebra.Ideal where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Powerset @@ -7,7 +7,7 @@ open import Cubical.Foundations.Powerset open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Ideal renaming (IdealsIn to IdealsInCommRing; makeIdeal to makeIdealCommRing) -open import Cubical.Algebra.CommAlgebraAlt.Base +open import Cubical.Algebra.CommAlgebra.Base -- open import Cubical.Algebra.Ring open import Cubical.Data.Unit diff --git a/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda b/Cubical/Algebra/CommAlgebra/Instances/Polynomials.agda similarity index 98% rename from Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda rename to Cubical/Algebra/CommAlgebra/Instances/Polynomials.agda index 481aed53bd..dc97a52b41 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Instances/Polynomials.agda +++ b/Cubical/Algebra/CommAlgebra/Instances/Polynomials.agda @@ -1,10 +1,10 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.Instances.Polynomials where +module Cubical.Algebra.CommAlgebra.Instances.Polynomials where open import Cubical.Foundations.Prelude open import Cubical.Algebra.CommRing.Base -open import Cubical.Algebra.CommAlgebraAlt.Base +open import Cubical.Algebra.CommAlgebra.Base open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate private diff --git a/Cubical/Algebra/CommAlgebraAlt/Instances/Unit.agda b/Cubical/Algebra/CommAlgebra/Instances/Unit.agda similarity index 93% rename from Cubical/Algebra/CommAlgebraAlt/Instances/Unit.agda rename to Cubical/Algebra/CommAlgebra/Instances/Unit.agda index dc68c46351..c84177967a 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Instances/Unit.agda +++ b/Cubical/Algebra/CommAlgebra/Instances/Unit.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.Instances.Unit where +module Cubical.Algebra.CommAlgebra.Instances.Unit where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -9,7 +9,7 @@ open import Cubical.Data.Unit open import Cubical.Data.Sigma.Properties using (Σ≡Prop) open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommAlgebraAlt.Base +open import Cubical.Algebra.CommAlgebra.Base open import Cubical.Algebra.CommRing.Instances.Unit open import Cubical.Tactics.CommRingSolver diff --git a/Cubical/Algebra/CommAlgebraAlt/Kernel.agda b/Cubical/Algebra/CommAlgebra/Kernel.agda similarity index 80% rename from Cubical/Algebra/CommAlgebraAlt/Kernel.agda rename to Cubical/Algebra/CommAlgebra/Kernel.agda index 484d0d7bc8..352b42aa68 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Kernel.agda +++ b/Cubical/Algebra/CommAlgebra/Kernel.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.Kernel where +module Cubical.Algebra.CommAlgebra.Kernel where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -7,9 +7,9 @@ open import Cubical.Foundations.Equiv open import Cubical.Algebra.CommRing.Base open import Cubical.Algebra.CommRing.Ideal using (Ideal→CommIdeal) open import Cubical.Algebra.Ring.Kernel using () renaming (kernelIdeal to ringKernel) -open import Cubical.Algebra.CommAlgebraAlt.Base +open import Cubical.Algebra.CommAlgebra.Base --open import Cubical.Algebra.CommAlgebra.Properties -open import Cubical.Algebra.CommAlgebraAlt.Ideal +open import Cubical.Algebra.CommAlgebra.Ideal private variable diff --git a/Cubical/Algebra/CommAlgebraAlt/Localisation.agda b/Cubical/Algebra/CommAlgebra/Localisation.agda similarity index 98% rename from Cubical/Algebra/CommAlgebraAlt/Localisation.agda rename to Cubical/Algebra/CommAlgebra/Localisation.agda index 91e7e28922..f13fab0c21 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Localisation.agda +++ b/Cubical/Algebra/CommAlgebra/Localisation.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --lossy-unification #-} -module Cubical.Algebra.CommAlgebraAlt.Localisation where +module Cubical.Algebra.CommAlgebra.Localisation where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -28,8 +28,8 @@ open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.RadicalIdeal open import Cubical.Algebra.CommRing.Localisation open import Cubical.Algebra.Ring -open import Cubical.Algebra.CommAlgebraAlt.Base -open import Cubical.Algebra.CommAlgebraAlt.Univalence +open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebra.Univalence open import Cubical.Tactics.CommRingSolver diff --git a/Cubical/Algebra/CommAlgebraAlt/Properties.agda b/Cubical/Algebra/CommAlgebra/Properties.agda similarity index 91% rename from Cubical/Algebra/CommAlgebraAlt/Properties.agda rename to Cubical/Algebra/CommAlgebra/Properties.agda index 893b6b386f..cde7f76a1e 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/Properties.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.Properties where +module Cubical.Algebra.CommAlgebra.Properties where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -10,7 +10,7 @@ open import Cubical.Foundations.Structure using (⟨_⟩; withOpaqueStr) open import Cubical.Data.Sigma open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommAlgebraAlt.Base +open import Cubical.Algebra.CommAlgebra.Base private variable diff --git a/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda similarity index 96% rename from Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda rename to Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda index 1ee161b93f..9bee35adba 100644 --- a/Cubical/Algebra/CommAlgebraAlt/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe --lossy-unification #-} -module Cubical.Algebra.CommAlgebraAlt.QuotientAlgebra where +module Cubical.Algebra.CommAlgebra.QuotientAlgebra where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv @@ -15,10 +15,10 @@ open import Cubical.Data.Sigma.Properties using (Σ≡Prop) open import Cubical.Algebra.CommRing import Cubical.Algebra.CommRing.Quotient as CommRing open import Cubical.Algebra.CommRing.Ideal hiding (IdealsIn) -open import Cubical.Algebra.CommAlgebraAlt.Base -open import Cubical.Algebra.CommAlgebraAlt.Ideal -open import Cubical.Algebra.CommAlgebraAlt.Kernel -open import Cubical.Algebra.CommAlgebraAlt.Instances.Unit +open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebra.Ideal +open import Cubical.Algebra.CommAlgebra.Kernel +open import Cubical.Algebra.CommAlgebra.Instances.Unit open import Cubical.Algebra.Ring open import Cubical.Tactics.CommRingSolver diff --git a/Cubical/Algebra/CommAlgebraAlt/Univalence.agda b/Cubical/Algebra/CommAlgebra/Univalence.agda similarity index 95% rename from Cubical/Algebra/CommAlgebraAlt/Univalence.agda rename to Cubical/Algebra/CommAlgebra/Univalence.agda index 6bbde392f1..27755deb9e 100644 --- a/Cubical/Algebra/CommAlgebraAlt/Univalence.agda +++ b/Cubical/Algebra/CommAlgebra/Univalence.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.Univalence where +module Cubical.Algebra.CommAlgebra.Univalence where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -15,7 +15,7 @@ open import Cubical.Data.Sigma open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Univalence -open import Cubical.Algebra.CommAlgebraAlt.Base +open import Cubical.Algebra.CommAlgebra.Base private variable diff --git a/Cubical/Algebra/CommAlgebraAlt.agda b/Cubical/Algebra/CommAlgebraAlt.agda deleted file mode 100644 index f7f8ed410e..0000000000 --- a/Cubical/Algebra/CommAlgebraAlt.agda +++ /dev/null @@ -1,5 +0,0 @@ -{-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt where - -open import Cubical.Algebra.CommAlgebraAlt.Base public -open import Cubical.Algebra.CommAlgebraAlt.Properties public diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule.agda deleted file mode 100644 index 406d5ff81e..0000000000 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule.agda +++ /dev/null @@ -1,5 +0,0 @@ -{-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebraAlt.AsModule where - -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base public -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Properties public diff --git a/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra.agda b/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra.agda deleted file mode 100644 index 9a98574f8a..0000000000 --- a/Cubical/Algebra/CommAlgebraAlt/AsModule/FreeCommAlgebra.agda +++ /dev/null @@ -1,6 +0,0 @@ -{-# OPTIONS --safe #-} - -module Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra where - -open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra.Base public -open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra.Properties public diff --git a/Cubical/Algebra/Polynomials/TypevariateHIT/Base.agda b/Cubical/Algebra/Polynomials/TypevariateHIT/Base.agda index c19fb0b5db..6ad5a7a4c8 100644 --- a/Cubical/Algebra/Polynomials/TypevariateHIT/Base.agda +++ b/Cubical/Algebra/Polynomials/TypevariateHIT/Base.agda @@ -13,4 +13,4 @@ The typevariate polynomials are constructed as a free commutative algebra on the They are justified by a proof of the universal property of a free commutative algebra. -} -open import Cubical.Algebra.CommAlgebraAlt.AsModule.FreeCommAlgebra public +open import Cubical.Algebra.CommAlgebra.AsModule.FreeCommAlgebra public diff --git a/Cubical/Algebra/Polynomials/TypevariateHIT/EquivUnivariateListPoly.agda b/Cubical/Algebra/Polynomials/TypevariateHIT/EquivUnivariateListPoly.agda index f8a8a11dc5..2450de79b0 100644 --- a/Cubical/Algebra/Polynomials/TypevariateHIT/EquivUnivariateListPoly.agda +++ b/Cubical/Algebra/Polynomials/TypevariateHIT/EquivUnivariateListPoly.agda @@ -10,7 +10,7 @@ open import Cubical.Data.Unit open import Cubical.Algebra.CommRing open import Cubical.Algebra.Algebra -open import Cubical.Algebra.CommAlgebraAlt.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule open import Cubical.Algebra.Polynomials.TypevariateHIT renaming (inducedHomUnique to inducedHomUniqueHIT; isIdByUMP to isIdByUMP-HIT) diff --git a/Cubical/Algebra/Polynomials/UnivariateList/UniversalProperty.agda b/Cubical/Algebra/Polynomials/UnivariateList/UniversalProperty.agda index 9847c6292e..0adb140f97 100644 --- a/Cubical/Algebra/Polynomials/UnivariateList/UniversalProperty.agda +++ b/Cubical/Algebra/Polynomials/UnivariateList/UniversalProperty.agda @@ -5,4 +5,4 @@ module Cubical.Algebra.Polynomials.UnivariateList.UniversalProperty where Export an CommAlgebra-Instance for the UnivariateList-Polynomials. Also export the universal property with respect to not-necessarily commutative algebras. -} -open import Cubical.Algebra.CommAlgebraAlt.AsModule.UnivariatePolyList public +open import Cubical.Algebra.CommAlgebra.AsModule.UnivariatePolyList public diff --git a/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda index 007875d8fe..621e87171b 100644 --- a/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda +++ b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda @@ -47,9 +47,9 @@ open import Cubical.Algebra.CommRing.FGIdeal open import Cubical.Algebra.CommRing.RadicalIdeal open import Cubical.Algebra.CommRing.Localisation open import Cubical.Algebra.CommRing.Instances.Unit -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Properties -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Localisation +open import Cubical.Algebra.CommAlgebra.AsModule.Base +open import Cubical.Algebra.CommAlgebra.AsModule.Properties +open import Cubical.Algebra.CommAlgebra.AsModule.Localisation open import Cubical.Tactics.CommRingSolver open import Cubical.Algebra.Semilattice open import Cubical.Algebra.Lattice diff --git a/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheafPullback.agda b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheafPullback.agda index e067b98047..412f3bc966 100644 --- a/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheafPullback.agda +++ b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheafPullback.agda @@ -48,10 +48,10 @@ open import Cubical.Algebra.CommRing.Localisation.Base open import Cubical.Algebra.CommRing.Localisation.UniversalProperty open import Cubical.Algebra.CommRing.Localisation.InvertingElements open import Cubical.Algebra.CommRing.Localisation.PullbackSquare -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Base -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Properties -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Localisation -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Unit +open import Cubical.Algebra.CommAlgebra.AsModule.Base +open import Cubical.Algebra.CommAlgebra.AsModule.Properties +open import Cubical.Algebra.CommAlgebra.AsModule.Localisation +open import Cubical.Algebra.CommAlgebra.AsModule.Instances.Unit open import Cubical.Tactics.CommRingSolver open import Cubical.Algebra.Semilattice open import Cubical.Algebra.Lattice diff --git a/Cubical/Categories/Instances/CommAlgebras.agda b/Cubical/Categories/Instances/CommAlgebras.agda index 6b8df7cedb..38b0140c17 100644 --- a/Cubical/Categories/Instances/CommAlgebras.agda +++ b/Cubical/Categories/Instances/CommAlgebras.agda @@ -14,8 +14,8 @@ open import Cubical.Data.Sigma open import Cubical.Algebra.Ring open import Cubical.Algebra.CommRing open import Cubical.Algebra.Algebra -open import Cubical.Algebra.CommAlgebraAlt.AsModule -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Instances.Unit +open import Cubical.Algebra.CommAlgebra.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule.Instances.Unit open import Cubical.Categories.Category open import Cubical.Categories.Functor diff --git a/Cubical/Experiments/ZariskiLatticeBasicOpens.agda b/Cubical/Experiments/ZariskiLatticeBasicOpens.agda index 4da00892d9..b5820781d9 100644 --- a/Cubical/Experiments/ZariskiLatticeBasicOpens.agda +++ b/Cubical/Experiments/ZariskiLatticeBasicOpens.agda @@ -31,8 +31,8 @@ open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Localisation.Base open import Cubical.Algebra.CommRing.Localisation.UniversalProperty open import Cubical.Algebra.CommRing.Localisation.InvertingElements -open import Cubical.Algebra.CommAlgebraAlt.AsModule -open import Cubical.Algebra.CommAlgebraAlt.AsModule.Localisation +open import Cubical.Algebra.CommAlgebra.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule.Localisation open import Cubical.Tactics.CommRingSolver open import Cubical.Algebra.Semilattice diff --git a/Cubical/Papers/AffineSchemes.agda b/Cubical/Papers/AffineSchemes.agda index cffebdbc66..6b0b6927fc 100644 --- a/Cubical/Papers/AffineSchemes.agda +++ b/Cubical/Papers/AffineSchemes.agda @@ -40,8 +40,8 @@ module Localization = L.Loc import Cubical.Algebra.CommRing.Localisation.UniversalProperty as LocalizationUnivProp import Cubical.Algebra.CommRing.Localisation.InvertingElements as LocalizationInvEl -import Cubical.Algebra.CommAlgebraAlt.AsModule as R-Algs -import Cubical.Algebra.CommAlgebraAlt.AsModule.Localisation as LocalizationR-Alg +import Cubical.Algebra.CommAlgebra.AsModule as R-Algs +import Cubical.Algebra.CommAlgebra.AsModule.Localisation as LocalizationR-Alg -- 3.2: The Zariski Lattice diff --git a/Cubical/Tactics/CommRingSolver/Examples.agda b/Cubical/Tactics/CommRingSolver/Examples.agda index a22f3f4079..c931620558 100644 --- a/Cubical/Tactics/CommRingSolver/Examples.agda +++ b/Cubical/Tactics/CommRingSolver/Examples.agda @@ -10,7 +10,7 @@ open import Cubical.Data.Nat using (ℕ; suc; zero) open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Instances.Int -open import Cubical.Algebra.CommAlgebraAlt.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule open import Cubical.Tactics.CommRingSolver open import Cubical.Tactics.CommRingSolver.RawAlgebra using (scalar) From f21b3a1cf6dfcdda9333e17d27f382a91b60b77b Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 8 Aug 2024 14:02:28 +0200 Subject: [PATCH 67/83] remove no-eta for structure --- Cubical/Algebra/Algebra/Base.agda | 2 -- 1 file changed, 2 deletions(-) diff --git a/Cubical/Algebra/Algebra/Base.agda b/Cubical/Algebra/Algebra/Base.agda index 21c74bce72..8cae5bdab6 100644 --- a/Cubical/Algebra/Algebra/Base.agda +++ b/Cubical/Algebra/Algebra/Base.agda @@ -54,8 +54,6 @@ unquoteDecl IsAlgebraIsoΣ = declareRecordIsoΣ IsAlgebraIsoΣ (quote IsAlgebra) record AlgebraStr (R : Ring ℓ) (A : Type ℓ') : Type (ℓ-max ℓ ℓ') where - no-eta-equality - field 0a : A 1a : A From 703e71711c7dea9c2db7ef7840e3f7f050d4ff18 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 8 Aug 2024 14:51:38 +0200 Subject: [PATCH 68/83] start with fp algebras --- Cubical/Algebra/CommAlgebra/FP.agda | 14 +++++++ Cubical/Algebra/CommAlgebra/FP/Base.agda | 52 ++++++++++++++++++++++++ 2 files changed, 66 insertions(+) create mode 100644 Cubical/Algebra/CommAlgebra/FP.agda create mode 100644 Cubical/Algebra/CommAlgebra/FP/Base.agda diff --git a/Cubical/Algebra/CommAlgebra/FP.agda b/Cubical/Algebra/CommAlgebra/FP.agda new file mode 100644 index 0000000000..41f07ffa70 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/FP.agda @@ -0,0 +1,14 @@ +{- + Finitely presented algebras. + An R-algebra A is finitely presented, if there merely is an exact sequence + of R-modules: + + (f₁,⋯,fₘ) → R[X₁,⋯,Xₙ] → A → 0 + + (where f₁,⋯,fₘ ∈ R[X₁,⋯,Xₙ]) + Our definition is more explicit. +-} +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.FP where + +open import Cubical.Algebra.CommAlgebra.FP.Base public diff --git a/Cubical/Algebra/CommAlgebra/FP/Base.agda b/Cubical/Algebra/CommAlgebra/FP/Base.agda new file mode 100644 index 0000000000..28f95bd918 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/FP/Base.agda @@ -0,0 +1,52 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.FP.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure + +open import Cubical.Data.FinData +open import Cubical.Data.Nat +open import Cubical.Data.Vec +open import Cubical.Data.Sigma + +open import Cubical.HITs.PropositionalTruncation + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebra.Instances.Polynomials +open import Cubical.Algebra.CommAlgebra.QuotientAlgebra +open import Cubical.Algebra.CommAlgebra.Ideal +open import Cubical.Algebra.CommAlgebra.FGIdeal +open import Cubical.Algebra.CommAlgebra.Kernel + + +private + variable + ℓ ℓ' : Level + +module _ (R : CommRing ℓ) where + Polynomials : (n : ℕ) → CommAlgebra R ℓ + Polynomials n = R [ Fin n ]ₐ + + FPCAlgebra : {m : ℕ} (n : ℕ) (relations : FinVec ⟨ Polynomials n ⟩ₐ m) → CommAlgebra R ℓ + FPCAlgebra n relations = Polynomials n / ⟨ relations ⟩[ Polynomials n ] + + record FinitePresentation (A : CommAlgebra R ℓ') : Type (ℓ-max ℓ ℓ') where + no-eta-equality + field + n : ℕ + m : ℕ + relations : FinVec ⟨ Polynomials n ⟩ₐ m + equiv : CommAlgebraEquiv (FPCAlgebra n relations) A + + abstract + isFP : (A : CommAlgebra R ℓ') → Type (ℓ-max ℓ ℓ') + isFP A = ∥ FinitePresentation A ∥₁ + + isFPIsProp : {A : CommAlgebra R ℓ} → isProp (isFP A) + isFPIsProp = isPropPropTrunc From 538486223ea26fc5224ed335b7df0771e6d10994 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 8 Aug 2024 17:07:45 +0200 Subject: [PATCH 69/83] small additions --- Cubical/Algebra/CommAlgebra/Properties.agda | 15 +++++++++++++++ Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda | 6 +++--- Cubical/Algebra/CommRing/Properties.agda | 4 +--- Cubical/Data/FinData/Base.agda | 3 +++ 4 files changed, 22 insertions(+), 6 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebra/Properties.agda index cde7f76a1e..6100bb9027 100644 --- a/Cubical/Algebra/CommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/Properties.agda @@ -31,3 +31,18 @@ module _ {R : CommRing ℓ} where (f : CommAlgebraHom A B) (g : CommAlgebraHom B C) (h : CommAlgebraHom C D) → h ∘ca (g ∘ca f) ≡ (h ∘ca g) ∘ca f compAssocCommAlgebraHom f g h = CommAlgebraHom≡ refl + + invCommAlgebraEquiv : {A B : CommAlgebra R ℓ'} + → CommAlgebraEquiv A B + → CommAlgebraEquiv B A + invCommAlgebraEquiv {A = A} {B = B} f = f⁻¹ , commutes + where f⁻¹ = invCommRingEquiv (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) (f .fst) + f⁻¹∘f≡Id : f⁻¹ .fst .fst ∘ f .fst .fst .fst ≡ idfun _ + f⁻¹∘f≡Id = funExt (secIsEq (equivIsEquiv (f⁻¹ .fst))) + abstract + commutes : (CommRingEquiv→CommRingHom f⁻¹) ∘cr B .snd ≡ A .snd + commutes = + CommRingHom≡ $ + f⁻¹ .fst .fst ∘ B .snd .fst ≡⟨ sym ( cong ((f⁻¹ .fst .fst) ∘_) (cong fst (f .snd))) ⟩ + f⁻¹ .fst .fst ∘ f .fst .fst .fst ∘ A .snd .fst ≡⟨ cong (_∘ A .snd .fst) f⁻¹∘f≡Id ⟩ + A .snd .fst ∎ diff --git a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda index 9bee35adba..7b15f0264d 100644 --- a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda @@ -38,11 +38,11 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ') (I : IdealsIn R A) where _ = A _/_ : CommAlgebra R ℓ' - _/_ = ((fst A) CommRing./ I) , - (CommRing.quotientHom (fst A) I ∘cr A .snd) + _/_ = ((fst A) CommRing./ I) , + (withOpaqueStr $ (CommRing.quotientHom (fst A) I ∘cr A .snd)) quotientHom : CommAlgebraHom A (_/_) - quotientHom = withOpaqueStr $ (CommRing.quotientHom (fst A) I) , refl + quotientHom = withOpaqueStr $ (CommRing.quotientHom (fst A) I) , CommRingHom≡ refl module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where open CommRingStr ⦃...⦄ diff --git a/Cubical/Algebra/CommRing/Properties.agda b/Cubical/Algebra/CommRing/Properties.agda index da9ac7a4d7..e08765fe21 100644 --- a/Cubical/Algebra/CommRing/Properties.agda +++ b/Cubical/Algebra/CommRing/Properties.agda @@ -187,7 +187,7 @@ module _ where → (x y : R .fst) → f .fst .fun x ≡ f .fst .fun y → x ≡ y injCommRingIso f x y h = sym (f .fst .leftInv x) ∙∙ cong (f .fst .inv) h ∙∙ f .fst .leftInv y -module CommRingEquivs where +module _ where open RingEquivs compCommRingEquiv : {A : CommRing ℓ} {B : CommRing ℓ'} {C : CommRing ℓ''} @@ -319,7 +319,6 @@ module CommRingTheory (R' : CommRing ℓ) where -- the CommRing version of uaCompEquiv module CommRingUAFunctoriality where open CommRingStr - open CommRingEquivs CommRing≡ : (A B : CommRing ℓ) → ( Σ[ p ∈ ⟨ A ⟩ ≡ ⟨ B ⟩ ] @@ -369,7 +368,6 @@ module CommRingUAFunctoriality where -open CommRingEquivs open CommRingUAFunctoriality recPT→CommRing : {A : Type ℓ'} (𝓕 : A → CommRing ℓ) → (σ : ∀ x y → CommRingEquiv (𝓕 x) (𝓕 y)) diff --git a/Cubical/Data/FinData/Base.agda b/Cubical/Data/FinData/Base.agda index f1d6ce2286..d8e7f6daf5 100644 --- a/Cubical/Data/FinData/Base.agda +++ b/Cubical/Data/FinData/Base.agda @@ -76,6 +76,9 @@ rec a0 aS (suc x) = aS FinVec : (A : Type ℓ) (n : ℕ) → Type ℓ FinVec A n = Fin n → A +emptyFinVec : (A : Type ℓ) → FinVec A 0 +emptyFinVec A () + replicateFinVec : (n : ℕ) → A → FinVec A n replicateFinVec _ a _ = a From 6e12f6c8fa951f0fda2ae83457914e9685525fea Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 8 Aug 2024 18:08:51 +0200 Subject: [PATCH 70/83] WIP: start with FP algebra examples --- Cubical/Algebra/CommAlgebra/FP/Instances.agda | 245 ++++++++++++++++++ 1 file changed, 245 insertions(+) create mode 100644 Cubical/Algebra/CommAlgebra/FP/Instances.agda diff --git a/Cubical/Algebra/CommAlgebra/FP/Instances.agda b/Cubical/Algebra/CommAlgebra/FP/Instances.agda new file mode 100644 index 0000000000..76dd856a38 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/FP/Instances.agda @@ -0,0 +1,245 @@ +{- + This module shows, that a couple of R-algebras are finitely presented. + So far, in all cases a finite presentation is constructed. + Here is a list of the fp-algebras in this file, with their presentations: + * R[X₁,...,Xₙ] = R[X₁,...,Xₙ]/⟨⟩ (ideal generated by zero elements) + * R = R[⊥]/⟨⟩ + * R/⟨x₁,...,xₙ⟩ = R[⊥]/⟨x₁,...,xₙ⟩ + * R/⟨x⟩ (as special case of the above) +-} +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.FP.Instances where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure + +open import Cubical.Data.FinData +open import Cubical.Data.Nat +open import Cubical.Data.Vec +open import Cubical.Data.Sigma +open import Cubical.Data.Empty + +open import Cubical.HITs.PropositionalTruncation + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.FGIdeal using (inclOfFGIdeal) +open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebra.Instances.Polynomials +open import Cubical.Algebra.CommAlgebra.QuotientAlgebra +open import Cubical.Algebra.CommAlgebra.Ideal using (IdealsIn) +open import Cubical.Algebra.CommAlgebra.FGIdeal +-- open import Cubical.Algebra.CommAlgebra.Instances.Initial +open import Cubical.Algebra.CommAlgebra.Instances.Unit +open import Cubical.Algebra.CommAlgebra.AsModule.Kernel + +open import Cubical.Algebra.CommAlgebra.FP.Base + +private + variable + ℓ ℓ' : Level + + +module _ (R : CommRing ℓ) where + open FinitePresentation + + module _ (n : ℕ) where + polynomialsFP : FinitePresentation R (Polynomials R n) + polynomialsFP = + record { + n = n ; + m = 0 ; + relations = emptyFinVec ⟨ Polynomials R n ⟩ₐ ; + equiv = invCommAlgebraEquiv {!zeroIdealQuotient!} + } + +{- + {- Every (multivariate) polynomial algebra is finitely presented -} + module _ (n : ℕ) where + private + A : CommAlgebra R ℓ + A = Polynomials n + + emptyGen : FinVec (fst A) 0 + emptyGen = λ () + + B : CommAlgebra R ℓ + B = FPAlgebra n emptyGen + + polynomialAlgFP : FinitePresentation A + FinitePresentation.n polynomialAlgFP = n + m polynomialAlgFP = 0 + relations polynomialAlgFP = emptyGen + equiv polynomialAlgFP = + -- Idea: A and B enjoy the same universal property. + toAAsEquiv , snd toA + where + toA : CommAlgebraHom B A + toA = inducedHom n emptyGen A Construction.var (λ ()) + fromA : CommAlgebraHom A B + fromA = freeInducedHom B (generator _ _) + open AlgebraHoms + inverse1 : fromA ∘a toA ≡ idAlgebraHom _ + inverse1 = + fromA ∘a toA + ≡⟨ sym (unique _ _ B _ _ _ (λ i → cong (fst fromA) ( + fst toA (generator n emptyGen i) + ≡⟨ inducedHomOnGenerators _ _ _ _ _ _ ⟩ + Construction.var i + ∎))) ⟩ + inducedHom n emptyGen B (generator _ _) (relationsHold _ _) + ≡⟨ unique _ _ B _ _ _ (λ i → refl) ⟩ + idAlgebraHom _ + ∎ + inverse2 : toA ∘a fromA ≡ idAlgebraHom _ + inverse2 = isoFunInjective (homMapIso A) _ _ ( + evaluateAt A (toA ∘a fromA) ≡⟨ sym (naturalEvR {A = B} {B = A} toA fromA) ⟩ + fst toA ∘ evaluateAt B fromA ≡⟨ refl ⟩ + fst toA ∘ generator _ _ ≡⟨ funExt (inducedHomOnGenerators _ _ _ _ _)⟩ + Construction.var ∎) + toAAsEquiv : ⟨ B ⟩ ≃ ⟨ A ⟩ + toAAsEquiv = isoToEquiv (iso (fst toA) + (fst fromA) + (λ a i → fst (inverse2 i) a) + (λ b i → fst (inverse1 i) b)) + + {- The initial R-algebra is finitely presented -} + private + R[⊥] : CommAlgebra R ℓ + R[⊥] = Polynomials 0 + + emptyGen : FinVec (fst R[⊥]) 0 + emptyGen = λ () + + R[⊥]/⟨0⟩ : CommAlgebra R ℓ + R[⊥]/⟨0⟩ = FPAlgebra 0 emptyGen + + R[⊥]/⟨0⟩IsInitial : (B : CommAlgebra R ℓ) + → isContr (CommAlgebraHom R[⊥]/⟨0⟩ B) + R[⊥]/⟨0⟩IsInitial B = iHom , uniqueness + where + iHom : CommAlgebraHom R[⊥]/⟨0⟩ B + iHom = inducedHom 0 emptyGen B (λ ()) (λ ()) + uniqueness : (f : CommAlgebraHom R[⊥]/⟨0⟩ B) → + iHom ≡ f + uniqueness f = unique 0 emptyGen B (λ ()) (λ ()) f (λ ()) + + initialCAlgFP : FinitePresentation (initialCAlg R) + n initialCAlgFP = 0 + m initialCAlgFP = 0 + relations initialCAlgFP = emptyGen + equiv initialCAlgFP = + equivByInitiality R R[⊥]/⟨0⟩ R[⊥]/⟨0⟩IsInitial + + {- The terminal R-algebra is finitely presented -} + private + unitGen : FinVec (fst R[⊥]) 1 + unitGen zero = 1a + where open CommAlgebraStr (snd R[⊥]) + + R[⊥]/⟨1⟩ : CommAlgebra R ℓ + R[⊥]/⟨1⟩ = FPAlgebra 0 unitGen + + terminalCAlgFP : FinitePresentation (TerminalCAlg R) + n terminalCAlgFP = 0 + m terminalCAlgFP = 1 + relations terminalCAlgFP = unitGen + equiv terminalCAlgFP = equivFrom1≡0 R R[⊥]/⟨1⟩ (sym (⋆IdL 1a) ∙ relationsHold 0 unitGen zero) + where open CommAlgebraStr (snd R[⊥]/⟨1⟩) + + + {- + Quotients of the base ring by finitely generated ideals are finitely presented. + -} + module _ {m : ℕ} (xs : FinVec ⟨ R ⟩ m) where + ⟨xs⟩ : IdealsIn (initialCAlg R) + ⟨xs⟩ = generatedIdeal (initialCAlg R) xs + + R/⟨xs⟩ = (initialCAlg R) / ⟨xs⟩ + + open CommAlgebraStr ⦃...⦄ + private + rels : FinVec ⟨ Polynomials {R = R} 0 ⟩ m + rels = Construction.const ∘ xs + + B = FPAlgebra 0 rels + + π = quotientHom (initialCAlg R) ⟨xs⟩ + instance + _ = snd R/⟨xs⟩ + _ = snd (initialCAlg R) + _ = snd B + + πxs≡0 : (i : Fin m) → π $a xs i ≡ 0a + πxs≡0 i = isZeroFromIdeal {A = initialCAlg R} {I = ⟨xs⟩} (xs i) + (incInIdeal (initialCAlg R) xs i) + + + + R/⟨xs⟩FP : FinitePresentation R/⟨xs⟩ + n R/⟨xs⟩FP = 0 + FinitePresentation.m R/⟨xs⟩FP = m + relations R/⟨xs⟩FP = rels + equiv R/⟨xs⟩FP = (isoToEquiv (iso (fst toA) (fst fromA) + (λ a i → toFrom i $a a) + λ a i → fromTo i $a a)) + , (snd toA) + where + toA : CommAlgebraHom B R/⟨xs⟩ + toA = inducedHom 0 rels R/⟨xs⟩ (λ ()) relation-holds + where + vals : FinVec ⟨ R/⟨xs⟩ ⟩ 0 + vals () + vals' : FinVec ⟨ initialCAlg R ⟩ 0 + vals' () + relation-holds = λ i → + evPoly R/⟨xs⟩ (rels i) vals ≡⟨ sym + (evPolyHomomorphic + (initialCAlg R) + R/⟨xs⟩ + π + (rels i) + vals') ⟩ + π $a (evPoly (initialCAlg R) + (rels i) + vals') ≡⟨ cong (π $a_) (·IdR (xs i)) ⟩ + π $a xs i ≡⟨ πxs≡0 i ⟩ + 0a ∎ + {- + R ─→ R/⟨xs⟩ + id↓ ↓ ∃! + R ─→ R[⊥]/⟨rels⟩ + -} + + fromA : CommAlgebraHom R/⟨xs⟩ B + fromA = + quotientInducedHom + (initialCAlg R) + ⟨xs⟩ + B + (initialMap R B) + (inclOfFGIdeal + (CommAlgebra→CommRing (initialCAlg R)) + xs + (kernel (initialCAlg R) B (initialMap R B)) + (relationsHold 0 rels)) + + open AlgebraHoms + + fromTo : fromA ∘a toA ≡ idCAlgHom B + fromTo = cong fst + (isContr→isProp (universal 0 rels B (λ ()) (relationsHold 0 rels)) + (fromA ∘a toA , (λ ())) + (idCAlgHom B , (λ ()))) + + toFrom : toA ∘a fromA ≡ idCAlgHom R/⟨xs⟩ + toFrom = injectivePrecomp (initialCAlg R) ⟨xs⟩ R/⟨xs⟩ (toA ∘a fromA) (idCAlgHom R/⟨xs⟩) + (isContr→isProp (initialityContr R R/⟨xs⟩) _ _) + + module _ {m : ℕ} (x : ⟨ R ⟩) where + R/⟨x⟩FP : FinitePresentation (initialCAlg R / generatedIdeal (initialCAlg R) (replicateFinVec 1 x)) + R/⟨x⟩FP = R/⟨xs⟩FP (replicateFinVec 1 x) +-- -} From 1f44a505cccc10a844ad7f856757a1be0cbe6428 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 8 Aug 2024 18:09:15 +0200 Subject: [PATCH 71/83] change def, fix thinghs --- Cubical/Algebra/CommAlgebra/Base.agda | 94 ++++++++++++------- Cubical/Algebra/CommAlgebra/FP/Base.agda | 8 +- .../CommAlgebra/Instances/Polynomials.agda | 5 +- .../Algebra/CommAlgebra/Instances/Unit.agda | 4 +- .../Algebra/CommAlgebra/QuotientAlgebra.agda | 50 +++++----- Cubical/Algebra/CommRing/Base.agda | 2 +- 6 files changed, 97 insertions(+), 66 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/Base.agda index 9c2629f6e2..01ff56c891 100644 --- a/Cubical/Algebra/CommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/Base.agda @@ -17,6 +17,8 @@ open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Univalence open import Cubical.Algebra.Ring +open import Cubical.Reflection.RecordEquiv + private variable ℓ ℓ' ℓ'' ℓ''' : Level @@ -28,11 +30,52 @@ module _ {R : CommRing ℓ} where CommAlgebra→CommRing : (A : CommAlgebra R ℓ') → CommRing ℓ' CommAlgebra→CommRing = fst + ⟨_⟩ₐ : (A : CommAlgebra R ℓ') → Type ℓ' + ⟨ A ⟩ₐ = A .fst .fst + + record IsCommAlgebraHom (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') (f : ⟨ A ⟩ₐ → ⟨ B ⟩ₐ) : Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) where + no-eta-equality + field + isCommRingHom : IsCommRingHom (A .fst .snd) f (B .fst .snd) + commutes : (f , isCommRingHom) ∘cr snd A ≡ snd B + + unquoteDecl IsCommAlgebraHomIsoΣ = declareRecordIsoΣ IsCommAlgebraHomIsoΣ (quote IsCommAlgebraHom) + + isPropIsCommAlgebraHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') (f : ⟨ A ⟩ₐ → ⟨ B ⟩ₐ) + → isProp (IsCommAlgebraHom A B f) + isPropIsCommAlgebraHom A B f = + isOfHLevelRetractFromIso 1 IsCommAlgebraHomIsoΣ $ + isPropΣ (isPropIsCommRingHom (A .fst .snd) f (B .fst .snd)) + λ _ → isSetΣSndProp (isSet→ is-set) (λ _ → isPropIsCommRingHom _ _ _) _ _ + where + open CommRingStr (B .fst .snd) using (is-set) + CommAlgebraHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') → Type _ - CommAlgebraHom A B = Σ[ f ∈ CommRingHom (fst A) (fst B) ] f ∘cr snd A ≡ snd B + CommAlgebraHom A B = Σ[ f ∈ (⟨ A ⟩ₐ → ⟨ B ⟩ₐ) ] IsCommAlgebraHom A B f + + CommRingHom→CommAlgebraHom : + {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → (f : CommRingHom (A .fst) (B .fst)) + → f ∘cr snd A ≡ snd B + → CommAlgebraHom A B + CommRingHom→CommAlgebraHom f commutes .fst = f .fst + CommRingHom→CommAlgebraHom f commutes .snd = + record { isCommRingHom = f .snd ; commutes = commutes } + + CommAlgebraHom→CommRingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → CommAlgebraHom A B + → CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) + CommAlgebraHom→CommRingHom f = f .fst , IsCommAlgebraHom.isCommRingHom (f .snd) + + ⟨_⟩ᵣ→ = CommAlgebraHom→CommRingHom + + CommAlgebraHom→Triangle : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → (f : CommAlgebraHom A B) + → ⟨ f ⟩ᵣ→ ∘cr A .snd ≡ B .snd + CommAlgebraHom→Triangle f = IsCommAlgebraHom.commutes (f .snd) idCAlgHom : (A : CommAlgebra R ℓ') → CommAlgebraHom A A - idCAlgHom A = withOpaqueStr $ (idCommRingHom (fst A)) , CommRingHom≡ refl + idCAlgHom A = CommRingHom→CommAlgebraHom (idCommRingHom (fst A)) (CommRingHom≡ refl) idCommAlgebraHom = idCAlgHom @@ -40,25 +83,21 @@ module _ {R : CommRing ℓ} where → (g : CommAlgebraHom B C) → (f : CommAlgebraHom A B) → CommAlgebraHom A C compCommAlgebraHom {A = A} {B = B} {C = C} g f = - withOpaqueStr $ ((fst g) ∘cr (fst f)) , pasting + CommRingHom→CommAlgebraHom (⟨ g ⟩ᵣ→ ∘cr ⟨ f ⟩ᵣ→) pasting where opaque - pasting : ((fst g) ∘cr (fst f)) ∘cr snd A ≡ snd C + pasting : (⟨ g ⟩ᵣ→ ∘cr ⟨ f ⟩ᵣ→) ∘cr snd A ≡ snd C pasting = - Σ≡Prop (λ _ → isPropIsCommRingHom (R .snd) _ (C .fst .snd)) - ( - (fst g ∘cr (fst f ∘cr snd A)) .fst ≡⟨ cong (λ h → (fst g ∘cr h) .fst) (f .snd) ⟩ - (fst g ∘cr snd B) .fst ≡⟨ cong fst (g .snd) ⟩ - (C .snd .fst) ∎) - - ⟨_⟩ₐ : (A : CommAlgebra R ℓ') → Type ℓ' - ⟨ A ⟩ₐ = A .fst .fst + CommRingHom≡ $ + (⟨ g ⟩ᵣ→ ∘cr (⟨ f ⟩ᵣ→ ∘cr snd A)) .fst ≡⟨ cong (λ h → (⟨ g ⟩ᵣ→ ∘cr h) .fst) (CommAlgebraHom→Triangle f) ⟩ + (⟨ g ⟩ᵣ→ ∘cr snd B) .fst ≡⟨ cong fst (CommAlgebraHom→Triangle g) ⟩ + (C .snd .fst) ∎ ⟨_⟩ₐ→ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} (f : CommAlgebraHom A B) → ⟨ A ⟩ₐ → ⟨ B ⟩ₐ - ⟨ f ⟩ₐ→ = f .fst .fst + ⟨ f ⟩ₐ→ = f .fst _$ca_ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} → (φ : CommAlgebraHom A B) → (x : ⟨ A ⟩ₐ) → ⟨ B ⟩ₐ - φ $ca x = φ .fst .fst x + _$ca_ φ x = φ .fst x infixr 9 _∘ca_ -- same as functions _∘ca_ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} @@ -68,21 +107,16 @@ module _ {R : CommRing ℓ} where opaque CommAlgebraHom≡ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {f g : CommAlgebraHom A B} - → f .fst .fst ≡ g .fst .fst + → f .fst ≡ g .fst → f ≡ g - CommAlgebraHom≡ {B = B} p = - Σ≡Prop (λ _ → isSetΣSndProp (isSetΠ (λ _ → is-set)) - (λ _ → isPropIsCommRingHom _ _ _) - _ _) - (Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) - p) + CommAlgebraHom≡ {B = B} p = Σ≡Prop (λ _ → isPropIsCommAlgebraHom _ _ _) p where open CommRingStr (B .fst .snd) using (is-set) CommAlgebraEquiv : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') → Type _ CommAlgebraEquiv A B = Σ[ f ∈ CommRingEquiv (A .fst) (B .fst) ] (f .fst .fst , f .snd) ∘cr A .snd ≡ B .snd idCAlgEquiv : (A : CommAlgebra R ℓ') → CommAlgebraEquiv A A - idCAlgEquiv A = (CommRingEquivs.idCommRingEquiv (fst A)) , + idCAlgEquiv A = (idCommRingEquiv (fst A)) , CommRingHom≡ refl CommAlgebraEquiv≡ : @@ -91,7 +125,7 @@ module _ {R : CommRing ℓ} where → f ≡ g CommAlgebraEquiv≡ {B = B} p = Σ≡Prop (λ _ → isSetΣSndProp (isSetΠ (λ _ → isSetB)) (λ _ → isPropIsCommRingHom _ _ _) _ _) - (CommRingEquivs.CommRingEquiv≡ p) + (CommRingEquiv≡ p) where open CommRingStr (B .fst .snd) using () renaming (is-set to isSetB) isSetCommAlgebraEquiv : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') @@ -141,8 +175,8 @@ module _ {R : CommRing ℓ} where ; ⋆AssocL = λ r x y → sym (·Assoc (A .snd .fst r) x y) } - module IsCommAlgebraHom {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} (f : CommAlgebraHom A B) where - open IsCommRingHom (f .fst .snd) + module CommAlgebraHom {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} (f : CommAlgebraHom A B) where + open IsCommRingHom (IsCommAlgebraHom.isCommRingHom (f .snd)) open CommRingStr ⦃...⦄ open CommAlgebraStr ⦃...⦄ private instance @@ -153,7 +187,7 @@ module _ {R : CommRing ℓ} where opaque pres⋆ : (r : ⟨ R ⟩) (x : ⟨ A ⟩ₐ) → f $ca (r ⋆ x) ≡ r ⋆ f $ca x pres⋆ r x = f $ca (r ⋆ x) ≡⟨ pres· (A .snd .fst r) x ⟩ - (f $ca (A .snd .fst r)) · (f $ca x) ≡⟨ cong (_· (f $ca x)) (cong (λ g → g .fst r) (f .snd)) ⟩ + (f $ca (A .snd .fst r)) · (f $ca x) ≡⟨ cong (_· (f $ca x)) (cong (λ g → g .fst r) (CommAlgebraHom→Triangle f)) ⟩ r ⋆ f $ca x ∎ CommAlgebraHomFromCommRingHom : @@ -162,7 +196,7 @@ module _ {R : CommRing ℓ} where → ((r : ⟨ R ⟩) (x : ⟨ A ⟩ₐ) → f .fst (CommAlgebraStr._⋆_ (CommAlgebra→CommAlgebraStr A) r x) ≡ CommAlgebraStr._⋆_ (CommAlgebra→CommAlgebraStr B) r (f .fst x)) → CommAlgebraHom A B - CommAlgebraHomFromCommRingHom {A = A} {B = B} f pres⋆ = withOpaqueStr $ f , + CommAlgebraHomFromCommRingHom {A = A} {B = B} f pres⋆ = CommRingHom→CommAlgebraHom f let open CommRingStr ⦃...⦄ open CommAlgebraStr ⦃...⦄ instance @@ -182,10 +216,6 @@ module _ {R : CommRing ℓ} where CommAlgebra→Ring : (A : CommAlgebra R ℓ') → Ring ℓ' CommAlgebra→Ring A = CommRing→Ring (fst A) - CommAlgebraHom→CommRingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} - → CommAlgebraHom A B - → CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) - CommAlgebraHom→CommRingHom = fst CommAlgebraHom→RingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} → CommAlgebraHom A B @@ -201,4 +231,4 @@ module _ {R : CommRing ℓ} where → CommAlgebraEquiv A B → CommAlgebraHom A B CommAlgebraEquiv→CommAlgebraHom f = - withOpaqueStr $ (CommRingEquiv→CommRingHom (f .fst)) , CommRingHom≡ (cong (fst) (f .snd)) + CommRingHom→CommAlgebraHom (CommRingEquiv→CommRingHom (f .fst)) $ CommRingHom≡ (cong (fst) (f .snd)) diff --git a/Cubical/Algebra/CommAlgebra/FP/Base.agda b/Cubical/Algebra/CommAlgebra/FP/Base.agda index 28f95bd918..85abeaf3bd 100644 --- a/Cubical/Algebra/CommAlgebra/FP/Base.agda +++ b/Cubical/Algebra/CommAlgebra/FP/Base.agda @@ -4,7 +4,6 @@ module Cubical.Algebra.CommAlgebra.FP.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Powerset open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels open import Cubical.Foundations.Structure @@ -12,7 +11,6 @@ open import Cubical.Foundations.Structure open import Cubical.Data.FinData open import Cubical.Data.Nat open import Cubical.Data.Vec -open import Cubical.Data.Sigma open import Cubical.HITs.PropositionalTruncation @@ -22,7 +20,6 @@ open import Cubical.Algebra.CommAlgebra.Instances.Polynomials open import Cubical.Algebra.CommAlgebra.QuotientAlgebra open import Cubical.Algebra.CommAlgebra.Ideal open import Cubical.Algebra.CommAlgebra.FGIdeal -open import Cubical.Algebra.CommAlgebra.Kernel private @@ -30,8 +27,9 @@ private ℓ ℓ' : Level module _ (R : CommRing ℓ) where - Polynomials : (n : ℕ) → CommAlgebra R ℓ - Polynomials n = R [ Fin n ]ₐ + abstract + Polynomials : (n : ℕ) → CommAlgebra R ℓ + Polynomials n = R [ Fin n ]ₐ FPCAlgebra : {m : ℕ} (n : ℕ) (relations : FinVec ⟨ Polynomials n ⟩ₐ m) → CommAlgebra R ℓ FPCAlgebra n relations = Polynomials n / ⟨ relations ⟩[ Polynomials n ] diff --git a/Cubical/Algebra/CommAlgebra/Instances/Polynomials.agda b/Cubical/Algebra/CommAlgebra/Instances/Polynomials.agda index dc97a52b41..02e268daab 100644 --- a/Cubical/Algebra/CommAlgebra/Instances/Polynomials.agda +++ b/Cubical/Algebra/CommAlgebra/Instances/Polynomials.agda @@ -11,8 +11,9 @@ private variable ℓ ℓ' ℓ'' : Level -_[_]ₐ : (R : CommRing ℓ) (I : Type ℓ') → CommAlgebra R _ -R [ I ]ₐ = (R [ I ]) , constPolynomial R I +opaque + _[_]ₐ : (R : CommRing ℓ) (I : Type ℓ') → CommAlgebra R (ℓ-max ℓ ℓ') + R [ I ]ₐ = (R [ I ]) , constPolynomial R I {- evaluateAt : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') diff --git a/Cubical/Algebra/CommAlgebra/Instances/Unit.agda b/Cubical/Algebra/CommAlgebra/Instances/Unit.agda index c84177967a..4dd3c760c1 100644 --- a/Cubical/Algebra/CommAlgebra/Instances/Unit.agda +++ b/Cubical/Algebra/CommAlgebra/Instances/Unit.agda @@ -2,6 +2,7 @@ module Cubical.Algebra.CommAlgebra.Instances.Unit where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv open import Cubical.Foundations.Structure @@ -23,8 +24,7 @@ module _ (R : CommRing ℓ) where module _ (A : CommAlgebra R ℓ') where terminalMap : CommAlgebraHom A (UnitCommAlgebra {ℓ' = ℓ}) - terminalMap .fst = mapToUnitCommRing (A .fst) - terminalMap .snd = isPropMapToUnitCommRing _ _ _ + terminalMap = CommRingHom→CommAlgebraHom (mapToUnitCommRing (A .fst)) $ isPropMapToUnitCommRing _ _ _ {- module _ (A : CommAlgebra R ℓ) where diff --git a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda index 7b15f0264d..bdab5a6ddb 100644 --- a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda @@ -27,8 +27,8 @@ private ℓ ℓ' ℓ'' : Level module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ') (I : IdealsIn R A) where - open CommRingStr {{...}} - open CommAlgebraStr {{...}} + open CommRingStr ⦃...⦄ + open CommAlgebraStr ⦃...⦄ open RingTheory (CommRing→Ring (fst A)) using (-DistR·) instance _ : CommRingStr ⟨ R ⟩ @@ -42,7 +42,7 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ') (I : IdealsIn R A) where (withOpaqueStr $ (CommRing.quotientHom (fst A) I ∘cr A .snd)) quotientHom : CommAlgebraHom A (_/_) - quotientHom = withOpaqueStr $ (CommRing.quotientHom (fst A) I) , CommRingHom≡ refl + quotientHom = CommRingHom→CommAlgebraHom (CommRing.quotientHom (fst A) I) $ CommRingHom≡ refl module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where open CommRingStr ⦃...⦄ @@ -71,25 +71,26 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where _ = snd (CommAlgebra→CommRing B) inducedHom : CommAlgebraHom (A / I) B - inducedHom .fst = CommRing.UniversalProperty.inducedHom - (CommAlgebra→CommRing A) - (CommAlgebra→CommRing B) - I - (CommAlgebraHom→CommRingHom ϕ) - I⊆kernel - inducedHom .snd = p - where - step1 = cong (inducedHom .fst ∘cr_) (sym (quotientHom A I .snd)) - step2 = compAssocCommRingHom (A .snd) (CommRing.quotientHom (A .fst) I) (inducedHom .fst) - step3 = cong (_∘cr A .snd) (CommRing.UniversalProperty.isSolution - (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) I (CommAlgebraHom→CommRingHom ϕ) I⊆kernel) - opaque - p : (inducedHom .fst ∘cr snd (A / I)) ≡ snd B - p = (inducedHom .fst ∘cr snd (A / I)) ≡⟨ step1 ⟩ - (inducedHom .fst ∘cr (CommRing.quotientHom (A .fst) I ∘cr A .snd)) ≡⟨ step2 ⟩ - (inducedHom .fst ∘cr CommRing.quotientHom (A .fst) I) ∘cr A .snd ≡⟨ step3 ⟩ - (ϕ .fst) ∘cr A .snd ≡⟨ ϕ .snd ⟩ - snd B ∎ + inducedHom = + CommRingHom→CommAlgebraHom ind (CommRingHom≡ p) + where + ind = (CommRing.UniversalProperty.inducedHom + (CommAlgebra→CommRing A) + (CommAlgebra→CommRing B) + I + (CommAlgebraHom→CommRingHom ϕ) + I⊆kernel) + + step1 = cong (ind .fst ∘_) (sym (cong fst (CommAlgebraHom→Triangle $ quotientHom A I))) + step2 = cong fst (compAssocCommRingHom (A .snd) (CommRing.quotientHom (A .fst) I) ind) + step3 = cong (_∘ A .snd .fst) (cong fst (CommRing.UniversalProperty.isSolution (A .fst) (B .fst) I (CommAlgebraHom→CommRingHom ϕ) I⊆kernel)) + opaque + p : (ind .fst ∘ snd (A / I) .fst) ≡ B .snd .fst + p = ind .fst ∘ snd (A / I) .fst ≡⟨ step1 ⟩ + ind .fst ∘ quotientHom A I .fst ∘ A .snd .fst ≡⟨ step2 ⟩ + ind .fst ∘ quotientHom A I .fst ∘ A .snd .fst ≡⟨ step3 ⟩ + ϕ .fst ∘ A .snd .fst ≡⟨ cong fst (CommAlgebraHom→Triangle ϕ) ⟩ + B .snd .fst ∎ opaque inducedHom∘quotientHom : inducedHom ∘ca quotientHom A I ≡ ϕ @@ -102,7 +103,7 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where CommAlgebraHom≡ (cong fst $ CommRing.UniversalProperty.isUnique - (A .fst) (B .fst) I (ϕ .fst) I⊆kernel (ψ .fst) ψIsSolution) + (A .fst) (B .fst) I ⟨ ϕ ⟩ᵣ→ I⊆kernel ⟨ ψ ⟩ᵣ→ ψIsSolution) @@ -118,7 +119,6 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ') (I : IdealsIn R A) where quotientHomEpi = CommRing.quotientHomEpi (fst A) I - {- trivial quotient -} module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) where open CommRingStr (A .fst .snd) @@ -133,6 +133,7 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) where oneIdealQuotient .fst .snd = terminalMap R (A / 1Ideal R A) .fst .snd oneIdealQuotient .snd = terminalMap R (A / (1Ideal R A)) .snd +{- zeroIdealQuotient : CommAlgebraEquiv A (A / (0Ideal R A)) zeroIdealQuotient .fst .fst = @@ -178,3 +179,4 @@ module _ opaque isZeroFromIdeal : (x : ⟨ A ⟩ₐ) → x ∈ (fst I) → ⟨ quotientHom A I ⟩ₐ→ x ≡ 0r isZeroFromIdeal x x∈I = eq/ x 0r (subst (_∈ fst I) (solve! (CommAlgebra→CommRing A)) x∈I ) +-} diff --git a/Cubical/Algebra/CommRing/Base.agda b/Cubical/Algebra/CommRing/Base.agda index b4441b63de..77249feec7 100644 --- a/Cubical/Algebra/CommRing/Base.agda +++ b/Cubical/Algebra/CommRing/Base.agda @@ -124,7 +124,7 @@ IsCommRing.·Comm (CommRingStr.isCommRing (snd (Ring→CommRing R p))) = p record IsCommRingHom {A : Type ℓ} {B : Type ℓ'} (R : CommRingStr A) (f : A → B) (S : CommRingStr B) : Type (ℓ-max ℓ ℓ') where - + no-eta-equality private module R = CommRingStr R module S = CommRingStr S From a95079aca6f258d52a75ff3d9d43a7442fdb34ed Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 9 Aug 2024 16:20:43 +0200 Subject: [PATCH 72/83] fix localization, adapt univalence, rearrange base --- Cubical/Algebra/CommAlgebra/Base.agda | 151 +++++++++--------- Cubical/Algebra/CommAlgebra/Localisation.agda | 34 ++-- .../Algebra/CommAlgebra/QuotientAlgebra.agda | 7 +- Cubical/Algebra/CommAlgebra/Univalence.agda | 107 +++++++++---- 4 files changed, 178 insertions(+), 121 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/Base.agda index 01ff56c891..dda937fb97 100644 --- a/Cubical/Algebra/CommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/Base.agda @@ -33,12 +33,70 @@ module _ {R : CommRing ℓ} where ⟨_⟩ₐ : (A : CommAlgebra R ℓ') → Type ℓ' ⟨ A ⟩ₐ = A .fst .fst + {- + Contrary to most algebraic structures, this one only contains + law and structure of a CommAlgebra, which it is *in addition* + to its CommRing structure. + -} + record CommAlgebraStr (A : Type ℓ') : Type (ℓ-suc (ℓ-max ℓ ℓ')) where + open CommRingStr {{...}} + instance + _ : CommRingStr (R .fst) + _ = (R .snd) + + field + crStr : CommRingStr A + _⋆_ : ⟨ R ⟩ → A → A + ⋆Assoc : (r s : ⟨ R ⟩) (x : A) → (r · s) ⋆ x ≡ r ⋆ (s ⋆ x) + ⋆DistR+ : (r : ⟨ R ⟩) (x y : A) → r ⋆ (CommRingStr._+_ crStr x y) ≡ CommRingStr._+_ crStr(r ⋆ x) (r ⋆ y) + ⋆DistL+ : (r s : ⟨ R ⟩) (x : A) → (r + s) ⋆ x ≡ CommRingStr._+_ crStr (r ⋆ x) (s ⋆ x) + ⋆IdL : (x : A) → 1r ⋆ x ≡ x + ⋆AssocL : (r : ⟨ R ⟩) (x y : A) → CommRingStr._·_ crStr (r ⋆ x) y ≡ r ⋆ (CommRingStr._·_ crStr x y) + infixl 7 _⋆_ + + {- TODO: make laws opaque -} + CommAlgebra→CommAlgebraStr : (A : CommAlgebra R ℓ') → CommAlgebraStr ⟨ A ⟩ₐ + CommAlgebra→CommAlgebraStr A = + let open CommRingStr ⦃...⦄ + instance + _ : CommRingStr (R .fst) + _ = R .snd + _ = A .fst .snd + in record + { crStr = A .fst .snd + ; _⋆_ = λ r x → (A .snd .fst r) · x + ; ⋆Assoc = λ r s x → cong (_· x) (IsCommRingHom.pres· (A .snd .snd) r s) ∙ sym (·Assoc _ _ _) + ; ⋆DistR+ = λ r x y → ·DistR+ _ _ _ + ; ⋆DistL+ = λ r s x → cong (_· x) (IsCommRingHom.pres+ (A .snd .snd) r s) ∙ ·DistL+ _ _ _ + ; ⋆IdL = λ x → cong (_· x) (IsCommRingHom.pres1 (A .snd .snd)) ∙ ·IdL x + ; ⋆AssocL = λ r x y → sym (·Assoc (A .snd .fst r) x y) + } + + +{- + Homomorphisms of CommAlgebras + ----------------------------- + This is defined as a record instead of a Σ-type to make type inference work + better. +-} record IsCommAlgebraHom (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') (f : ⟨ A ⟩ₐ → ⟨ B ⟩ₐ) : Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) where no-eta-equality field isCommRingHom : IsCommRingHom (A .fst .snd) f (B .fst .snd) commutes : (f , isCommRingHom) ∘cr snd A ≡ snd B + open IsCommRingHom isCommRingHom public + open CommRingStr ⦃...⦄ + open CommAlgebraStr ⦃...⦄ + private instance + _ = CommAlgebra→CommAlgebraStr A + _ = CommAlgebra→CommAlgebraStr B + _ = B .fst .snd + pres⋆ : (r : ⟨ R ⟩) (x : ⟨ A ⟩ₐ) → f (r ⋆ x) ≡ r ⋆ f x + pres⋆ r x = f (r ⋆ x) ≡⟨ pres· (A .snd .fst r) x ⟩ + (f (A .snd .fst r)) · (f x) ≡⟨ cong (_· (f x)) (cong (λ g → g .fst r) commutes) ⟩ + r ⋆ f x ∎ + unquoteDecl IsCommAlgebraHomIsoΣ = declareRecordIsoΣ IsCommAlgebraHomIsoΣ (quote IsCommAlgebraHom) isPropIsCommAlgebraHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') (f : ⟨ A ⟩ₐ → ⟨ B ⟩ₐ) @@ -112,83 +170,32 @@ module _ {R : CommRing ℓ} where CommAlgebraHom≡ {B = B} p = Σ≡Prop (λ _ → isPropIsCommAlgebraHom _ _ _) p where open CommRingStr (B .fst .snd) using (is-set) + +{- Equivalences of CommAlgebras -} CommAlgebraEquiv : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') → Type _ - CommAlgebraEquiv A B = Σ[ f ∈ CommRingEquiv (A .fst) (B .fst) ] (f .fst .fst , f .snd) ∘cr A .snd ≡ B .snd + CommAlgebraEquiv A B = Σ[ f ∈ (⟨ A ⟩ₐ ≃ ⟨ B ⟩ₐ) ] IsCommAlgebraHom A B (f .fst) idCAlgEquiv : (A : CommAlgebra R ℓ') → CommAlgebraEquiv A A - idCAlgEquiv A = (idCommRingEquiv (fst A)) , - CommRingHom≡ refl + idCAlgEquiv A = (idEquiv ⟨ A ⟩ₐ) , isHom + where + opaque + isHom : IsCommAlgebraHom A A (idfun ⟨ A ⟩ₐ) + isHom = record { isCommRingHom = idCommRingHom (A .fst) .snd ; + commutes = CommRingHom≡ refl } CommAlgebraEquiv≡ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {f g : CommAlgebraEquiv A B} - → f .fst .fst .fst ≡ g .fst .fst .fst + → f .fst .fst ≡ g .fst .fst → f ≡ g CommAlgebraEquiv≡ {B = B} p = - Σ≡Prop (λ _ → isSetΣSndProp (isSetΠ (λ _ → isSetB)) (λ _ → isPropIsCommRingHom _ _ _) _ _) - (CommRingEquiv≡ p) - where open CommRingStr (B .fst .snd) using () renaming (is-set to isSetB) + Σ≡Prop (isPropIsCommAlgebraHom _ _ ∘ fst) (Σ≡Prop isPropIsEquiv p) isSetCommAlgebraEquiv : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') → isSet (CommAlgebraEquiv A B) isSetCommAlgebraEquiv A B = - isSetΣSndProp - (isSetCommRingEquiv _ _) - (λ e → isSetΣSndProp (isSet→ isSetB) (λ _ → isPropIsCommRingHom _ _ _) _ _) - where open CommRingStr (B .fst .snd) using () renaming (is-set to isSetB) - - {- - Contrary to most algebraic structures, this one only contains - law and structure of a CommAlgebra, which it is *in addition* - to its CommRing structure. - -} - record CommAlgebraStr (A : Type ℓ') : Type (ℓ-suc (ℓ-max ℓ ℓ')) where - open CommRingStr {{...}} - instance - _ : CommRingStr (R .fst) - _ = (R .snd) - - field - crStr : CommRingStr A - _⋆_ : ⟨ R ⟩ → A → A - ⋆Assoc : (r s : ⟨ R ⟩) (x : A) → (r · s) ⋆ x ≡ r ⋆ (s ⋆ x) - ⋆DistR+ : (r : ⟨ R ⟩) (x y : A) → r ⋆ (CommRingStr._+_ crStr x y) ≡ CommRingStr._+_ crStr(r ⋆ x) (r ⋆ y) - ⋆DistL+ : (r s : ⟨ R ⟩) (x : A) → (r + s) ⋆ x ≡ CommRingStr._+_ crStr (r ⋆ x) (s ⋆ x) - ⋆IdL : (x : A) → 1r ⋆ x ≡ x - ⋆AssocL : (r : ⟨ R ⟩) (x y : A) → CommRingStr._·_ crStr (r ⋆ x) y ≡ r ⋆ (CommRingStr._·_ crStr x y) - infixl 7 _⋆_ - - {- TODO: make laws opaque -} - CommAlgebra→CommAlgebraStr : (A : CommAlgebra R ℓ') → CommAlgebraStr ⟨ A ⟩ₐ - CommAlgebra→CommAlgebraStr A = - let open CommRingStr ⦃...⦄ - instance - _ : CommRingStr (R .fst) - _ = R .snd - _ = A .fst .snd - in record - { crStr = A .fst .snd - ; _⋆_ = λ r x → (A .snd .fst r) · x - ; ⋆Assoc = λ r s x → cong (_· x) (IsCommRingHom.pres· (A .snd .snd) r s) ∙ sym (·Assoc _ _ _) - ; ⋆DistR+ = λ r x y → ·DistR+ _ _ _ - ; ⋆DistL+ = λ r s x → cong (_· x) (IsCommRingHom.pres+ (A .snd .snd) r s) ∙ ·DistL+ _ _ _ - ; ⋆IdL = λ x → cong (_· x) (IsCommRingHom.pres1 (A .snd .snd)) ∙ ·IdL x - ; ⋆AssocL = λ r x y → sym (·Assoc (A .snd .fst r) x y) - } - - module CommAlgebraHom {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} (f : CommAlgebraHom A B) where - open IsCommRingHom (IsCommAlgebraHom.isCommRingHom (f .snd)) - open CommRingStr ⦃...⦄ - open CommAlgebraStr ⦃...⦄ - private instance - _ = CommAlgebra→CommAlgebraStr A - _ = CommAlgebra→CommAlgebraStr B - _ = B .fst .snd - - opaque - pres⋆ : (r : ⟨ R ⟩) (x : ⟨ A ⟩ₐ) → f $ca (r ⋆ x) ≡ r ⋆ f $ca x - pres⋆ r x = f $ca (r ⋆ x) ≡⟨ pres· (A .snd .fst r) x ⟩ - (f $ca (A .snd .fst r)) · (f $ca x) ≡⟨ cong (_· (f $ca x)) (cong (λ g → g .fst r) (CommAlgebraHom→Triangle f)) ⟩ - r ⋆ f $ca x ∎ + isSetΣSndProp (isSetΣSndProp (isSet→ isSetB) isPropIsEquiv) + (isPropIsCommAlgebraHom _ _ ∘ fst) + where open CommRingStr (B .fst .snd) using () renaming (is-set to isSetB) CommAlgebraHomFromCommRingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} @@ -211,6 +218,7 @@ module _ {R : CommRing ℓ} where r ⋆ 1r ≡⟨ ·IdR _ ⟩ B .snd .fst r ∎)) + {- Convenient forgetful functions -} module _ {R : CommRing ℓ} where CommAlgebra→Ring : (A : CommAlgebra R ℓ') → Ring ℓ' @@ -222,13 +230,12 @@ module _ {R : CommRing ℓ} where → RingHom (CommAlgebra→Ring A) (CommAlgebra→Ring B) CommAlgebraHom→RingHom = CommRingHom→RingHom ∘ CommAlgebraHom→CommRingHom - CommAlgebraEquiv→CommRingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} - → CommAlgebraEquiv A B - → CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) - CommAlgebraEquiv→CommRingHom e = CommRingEquiv→CommRingHom (e .fst) - CommAlgebraEquiv→CommAlgebraHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} → CommAlgebraEquiv A B → CommAlgebraHom A B - CommAlgebraEquiv→CommAlgebraHom f = - CommRingHom→CommAlgebraHom (CommRingEquiv→CommRingHom (f .fst)) $ CommRingHom≡ (cong (fst) (f .snd)) + CommAlgebraEquiv→CommAlgebraHom f = (f .fst .fst , f .snd) + + CommAlgebraEquiv→CommRingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → CommAlgebraEquiv A B + → CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) + CommAlgebraEquiv→CommRingHom = CommAlgebraHom→CommRingHom ∘ CommAlgebraEquiv→CommAlgebraHom diff --git a/Cubical/Algebra/CommAlgebra/Localisation.agda b/Cubical/Algebra/CommAlgebra/Localisation.agda index f13fab0c21..1e5f57f47d 100644 --- a/Cubical/Algebra/CommAlgebra/Localisation.agda +++ b/Cubical/Algebra/CommAlgebra/Localisation.agda @@ -109,17 +109,17 @@ module AlgLoc (R : CommRing ℓ) χᴬuniqueness : (ψ : CommAlgebraHom S⁻¹RAsCommAlg B) → χᴬ ≡ ψ χᴬuniqueness ψ = CommAlgebraHom≡ {f = χᴬ} - (χᴬ .fst .fst ≡⟨ cong (fst ∘ fst) (χuniqueness (ψ .fst , funExt ψ'r/1≡φr)) ⟩ ψ .fst .fst ∎) + (χᴬ .fst ≡⟨ cong (fst ∘ fst) (χuniqueness (CommAlgebraHom→CommRingHom ψ , funExt ψ'r/1≡φr)) ⟩ ψ .fst ∎) where χuniqueness = S⁻¹RHasUniversalProp (B .fst) φ S⋆1⊆Bˣ .snd - ψ'r/1≡φr : ∀ r → ψ .fst .fst (r /1) ≡ fst φ r + ψ'r/1≡φr : ∀ r → ψ .fst (r /1) ≡ fst φ r ψ'r/1≡φr r = - ψ .fst .fst (r /1) ≡⟨ cong (ψ .fst .fst) (sym (·ₗ-rid _)) ⟩ - ψ .fst .fst (r ⋆ 1r) ≡⟨ IsCommAlgebraHom.pres⋆ ψ _ _ ⟩ - r ⋆ (ψ .fst .fst 1r) ≡⟨ cong (λ u → r ⋆ u) (pres1 (ψ .fst .snd)) ⟩ - r ⋆ 1r ≡⟨ solve! (B .fst) ⟩ + ψ .fst (r /1) ≡⟨ cong (ψ .fst) (sym (·ₗ-rid _)) ⟩ + ψ .fst (r ⋆ 1r) ≡⟨ IsCommAlgebraHom.pres⋆ (ψ .snd) _ _ ⟩ + r ⋆ (ψ .fst 1r) ≡⟨ cong (λ u → r ⋆ u) (IsCommAlgebraHom.pres1 (ψ .snd)) ⟩ + r ⋆ 1r ≡⟨ solve! (B .fst) ⟩ fst φ r ∎ @@ -130,9 +130,11 @@ module AlgLoc (R : CommRing ℓ) S⁻¹RAlgCharEquiv : (A : CommRing ℓ) (φ : CommRingHom R A) → PathToS⁻¹R R S SMultClosedSubset A φ → CommAlgebraEquiv S⁻¹RAsCommAlg (A , φ) - S⁻¹RAlgCharEquiv A φ cond .fst = S⁻¹RCharEquiv R S SMultClosedSubset A φ cond - S⁻¹RAlgCharEquiv A φ cond .snd = CommRingHom≡ (S⁻¹RHasUniversalProp A φ (cond .φS⊆Aˣ) .fst .snd) - where open PathToS⁻¹R + S⁻¹RAlgCharEquiv A φ cond = + CommRingEquiv→CommAlgebraEquiv + (S⁻¹RCharEquiv R S SMultClosedSubset A φ cond) + (S⁻¹RHasUniversalProp A φ (cond .φS⊆Aˣ) .fst .snd) + where open PathToS⁻¹R -- the special case of localizing at a single element R[1/_]AsCommAlgebra : {R : CommRing ℓ} → fst R → CommAlgebra R ℓ @@ -184,8 +186,8 @@ module AlgLocTwoSubsets (R' : CommRing ℓ) → (∀ s₂ → s₂ ∈ S₂ → s₂ ⋆ 1r ∈ S₁⁻¹Rˣ) → isContr (S₁⁻¹RAsCommAlg ≡ S₂⁻¹RAsCommAlg) isContrS₁⁻¹R≡S₂⁻¹R S₁⊆S₂⁻¹Rˣ S₂⊆S₁⁻¹Rˣ = isOfHLevelRetractFromIso 0 - (equivToIso (invEquiv (CommAlgebraPath _ _ _))) - isContrS₁⁻¹R≅S₂⁻¹R + (equivToIso (invEquiv (CommAlgebraPath _ _))) + isContrS₁⁻¹R≅S₂⁻¹R where χ₁ : CommAlgebraHom S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg χ₁ = S₁⁻¹RHasAlgUniversalProp S₂⁻¹RAsCommAlg S₁⊆S₂⁻¹Rˣ .fst @@ -206,12 +208,14 @@ module AlgLocTwoSubsets (R' : CommRing ℓ) Iso.leftInv IsoS₁⁻¹RS₂⁻¹R = funExt⁻ (cong ⟨_⟩ₐ→ χ₂∘χ₁≡id) isContrS₁⁻¹R≅S₂⁻¹R : isContr (CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg) - isContrS₁⁻¹R≅S₂⁻¹R = withOpaqueStr $ center , uniqueness + isContrS₁⁻¹R≅S₂⁻¹R = center , uniqueness where - X₁asCRHom = fst χ₁ + X₁asCRHom = CommAlgebraHom→CommRingHom χ₁ center : CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg - fst center = withOpaqueStr $ (isoToEquiv IsoS₁⁻¹RS₂⁻¹R) , snd X₁asCRHom - snd center = makeOpaque $ CommRingHom≡ (cong fst (χ₁ .snd)) + center = + CommRingEquiv→CommAlgebraEquiv + ((isoToEquiv IsoS₁⁻¹RS₂⁻¹R) , snd X₁asCRHom) + (cong fst (CommAlgebraHom→Triangle χ₁)) uniqueness : (φ : CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg) → center ≡ φ uniqueness φ = diff --git a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda index bdab5a6ddb..3b7417f854 100644 --- a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda @@ -124,7 +124,10 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) where open CommRingStr (A .fst .snd) oneIdealQuotient : CommAlgebraEquiv (A / (1Ideal R A)) (UnitCommAlgebra R {ℓ' = ℓ}) - oneIdealQuotient .fst .fst = + oneIdealQuotient = ? + +{- +.fst .fst = withOpaqueStr $ isoToEquiv (iso ⟨ terminalMap R (A / 1Ideal R A) ⟩ₐ→ (λ _ → [ 0r ]) @@ -133,8 +136,6 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) where oneIdealQuotient .fst .snd = terminalMap R (A / 1Ideal R A) .fst .snd oneIdealQuotient .snd = terminalMap R (A / (1Ideal R A)) .snd -{- - zeroIdealQuotient : CommAlgebraEquiv A (A / (0Ideal R A)) zeroIdealQuotient .fst .fst = withOpaqueStr $ diff --git a/Cubical/Algebra/CommAlgebra/Univalence.agda b/Cubical/Algebra/CommAlgebra/Univalence.agda index 27755deb9e..4af65b17f9 100644 --- a/Cubical/Algebra/CommAlgebra/Univalence.agda +++ b/Cubical/Algebra/CommAlgebra/Univalence.agda @@ -22,34 +22,79 @@ private ℓ ℓ' ℓ'' : Level -CommAlgebraPath : - (R : CommRing ℓ) - (A B : CommAlgebra R ℓ') - → (Σ[ f ∈ CommRingEquiv (A .fst) (B .fst) ] (f .fst .fst , f .snd) ∘cr A .snd ≡ B .snd) - ≃ (A ≡ B) -CommAlgebraPath R A B = - (Σ-cong-equiv - (CommRingPath _ _) - (λ e → compPathlEquiv (computeSubst e))) - ∙ₑ ΣPathTransport≃PathΣ A B - where computeSubst : - (e : CommRingEquiv (fst A) (fst B)) - → subst (CommRingHom R) (uaCommRing e) (A .snd) - ≡ (CommRingEquiv→CommRingHom e) ∘cr A .snd - computeSubst e = - CommRingHom≡ $ - (subst (CommRingHom R) (uaCommRing e) (A .snd)) .fst - ≡⟨ sym (substCommSlice (CommRingHom R) (λ X → ⟨ R ⟩ → ⟨ X ⟩) (λ _ → fst) (uaCommRing e) (A .snd)) ⟩ - subst (λ X → ⟨ R ⟩ → ⟨ X ⟩) (uaCommRing e) (A .snd .fst) - ≡⟨ fromPathP (funTypeTransp (λ _ → ⟨ R ⟩) ⟨_⟩ (uaCommRing e) (A .snd .fst)) ⟩ - subst ⟨_⟩ (uaCommRing e) ∘ A .snd .fst ∘ subst (λ _ → ⟨ R ⟩) (sym (uaCommRing e)) - ≡⟨ cong ((subst ⟨_⟩ (uaCommRing e) ∘ A .snd .fst) ∘_) (funExt (λ _ → transportRefl _)) ⟩ - (subst ⟨_⟩ (uaCommRing e) ∘ (A .snd .fst)) - ≡⟨ funExt (λ _ → uaβ (e .fst) _) ⟩ - (CommRingEquiv→CommRingHom e ∘cr A .snd) .fst ∎ - -uaCommAlgebra : {R : CommRing ℓ} {A B : CommAlgebra R ℓ'} → CommAlgebraEquiv A B → A ≡ B -uaCommAlgebra {R = R} {A = A} {B = B} = equivFun (CommAlgebraPath R A B) - -isGroupoidCommAlgebra : {R : CommRing ℓ} → isGroupoid (CommAlgebra R ℓ') -isGroupoidCommAlgebra A B = isOfHLevelRespectEquiv 2 (CommAlgebraPath _ _ _) (isSetCommAlgebraEquiv _ _) +module _ {R : CommRing ℓ} where + + CommAlgebraPath' : + (A B : CommAlgebra R ℓ') + → (Σ[ f ∈ CommRingEquiv (A .fst) (B .fst) ] (f .fst .fst , f .snd) ∘cr A .snd ≡ B .snd) + ≃ (A ≡ B) + CommAlgebraPath' A B = + (Σ-cong-equiv + (CommRingPath _ _) + (λ e → compPathlEquiv (computeSubst e))) + ∙ₑ ΣPathTransport≃PathΣ A B + where computeSubst : + (e : CommRingEquiv (fst A) (fst B)) + → subst (CommRingHom R) (uaCommRing e) (A .snd) + ≡ (CommRingEquiv→CommRingHom e) ∘cr A .snd + computeSubst e = + CommRingHom≡ $ + (subst (CommRingHom R) (uaCommRing e) (A .snd)) .fst + ≡⟨ sym (substCommSlice (CommRingHom R) (λ X → ⟨ R ⟩ → ⟨ X ⟩) (λ _ → fst) (uaCommRing e) (A .snd)) ⟩ + subst (λ X → ⟨ R ⟩ → ⟨ X ⟩) (uaCommRing e) (A .snd .fst) + ≡⟨ fromPathP (funTypeTransp (λ _ → ⟨ R ⟩) ⟨_⟩ (uaCommRing e) (A .snd .fst)) ⟩ + subst ⟨_⟩ (uaCommRing e) ∘ A .snd .fst ∘ subst (λ _ → ⟨ R ⟩) (sym (uaCommRing e)) + ≡⟨ cong ((subst ⟨_⟩ (uaCommRing e) ∘ A .snd .fst) ∘_) (funExt (λ _ → transportRefl _)) ⟩ + (subst ⟨_⟩ (uaCommRing e) ∘ (A .snd .fst)) + ≡⟨ funExt (λ _ → uaβ (e .fst) _) ⟩ + (CommRingEquiv→CommRingHom e ∘cr A .snd) .fst ∎ + + + CommRingEquiv→CommAlgebraEquiv : + {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → (e : CommRingEquiv (A .fst) (B .fst)) + → e .fst .fst ∘ A .snd .fst ≡ B .snd .fst + → CommAlgebraEquiv A B + CommRingEquiv→CommAlgebraEquiv {A = A} {B = B} e p = e .fst , isHom + where + opaque + isHom : IsCommAlgebraHom A B (e .fst .fst) + isHom = record { isCommRingHom = e .snd ; commutes = CommRingHom≡ p } + + CommAlgebraEquiv→CREquivΣ : + {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → (CommAlgebraEquiv A B) + ≃ (Σ[ f ∈ CommRingEquiv (A .fst) (B .fst) ] (f .fst .fst , f .snd) ∘cr A .snd ≡ B .snd) + CommAlgebraEquiv→CREquivΣ = + isoToEquiv $ + iso to from + (λ _ → Σ≡Prop (λ _ → isSetCommRingHom _ _ _ _) (Σ≡Prop (λ f → isPropIsCommRingHom _ (f .fst) _) refl)) + (λ _ → Σ≡Prop (isPropIsCommAlgebraHom _ _ ∘ fst) refl) + where to : CommAlgebraEquiv _ _ → _ + to e = (e .fst , IsCommAlgebraHom.isCommRingHom (e .snd)) , IsCommAlgebraHom.commutes (e .snd) + + from : _ → CommAlgebraEquiv _ _ + from (e , commutes) = CommRingEquiv→CommAlgebraEquiv e (cong fst commutes) + + CommAlgebraPath : + (A B : CommAlgebra R ℓ') + → CommAlgebraEquiv A B + ≃ (A ≡ B) + CommAlgebraPath A B = + CommAlgebraEquiv A B + ≃⟨ CommAlgebraEquiv→CREquivΣ ⟩ + Σ[ f ∈ CommRingEquiv (A .fst) (B .fst) ] (f .fst .fst , f .snd) ∘cr A .snd ≡ B .snd + ≃⟨ CommAlgebraPath' A B ⟩ + A ≡ B ■ + + uaCommAlgebra : {A B : CommAlgebra R ℓ'} → CommAlgebraEquiv A B → A ≡ B + uaCommAlgebra {A = A} {B = B} = + equivFun $ CommAlgebraPath A B + + + isGroupoidCommAlgebra : isGroupoid (CommAlgebra R ℓ') + isGroupoidCommAlgebra A B = + isOfHLevelRespectEquiv + 2 + (CommAlgebraPath' _ _) + (isSetΣSndProp (isSetCommRingEquiv _ _) λ _ → isSetCommRingHom _ _ _ _) From 4ba4d01519bc90e4b7f7ec0cf7b0f6bb078f31c7 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 9 Aug 2024 17:09:57 +0200 Subject: [PATCH 73/83] fix quotients --- Cubical/Algebra/CommAlgebra/Base.agda | 6 ++++++ Cubical/Algebra/CommAlgebra/Properties.agda | 14 ++++++------- .../Algebra/CommAlgebra/QuotientAlgebra.agda | 21 +++++++------------ 3 files changed, 21 insertions(+), 20 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/Base.agda index dda937fb97..3be2bfd9a2 100644 --- a/Cubical/Algebra/CommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/Base.agda @@ -239,3 +239,9 @@ module _ {R : CommRing ℓ} where → CommAlgebraEquiv A B → CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) CommAlgebraEquiv→CommRingHom = CommAlgebraHom→CommRingHom ∘ CommAlgebraEquiv→CommAlgebraHom + + CommAlgebraEquiv→CommRingEquiv : + {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → CommAlgebraEquiv A B + → CommRingEquiv (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) + CommAlgebraEquiv→CommRingEquiv e = (e .fst) , (IsCommAlgebraHom.isCommRingHom (e .snd)) diff --git a/Cubical/Algebra/CommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebra/Properties.agda index 6100bb9027..0cf21ba7ac 100644 --- a/Cubical/Algebra/CommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/Properties.agda @@ -11,6 +11,7 @@ open import Cubical.Data.Sigma open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebra.Univalence private variable @@ -35,14 +36,13 @@ module _ {R : CommRing ℓ} where invCommAlgebraEquiv : {A B : CommAlgebra R ℓ'} → CommAlgebraEquiv A B → CommAlgebraEquiv B A - invCommAlgebraEquiv {A = A} {B = B} f = f⁻¹ , commutes - where f⁻¹ = invCommRingEquiv (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) (f .fst) - f⁻¹∘f≡Id : f⁻¹ .fst .fst ∘ f .fst .fst .fst ≡ idfun _ + invCommAlgebraEquiv {A = A} {B = B} f = CommRingEquiv→CommAlgebraEquiv f⁻¹ commutes + where f⁻¹ = invCommRingEquiv (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) (CommAlgebraEquiv→CommRingEquiv f) + f⁻¹∘f≡Id : f⁻¹ .fst .fst ∘ f .fst .fst ≡ idfun _ f⁻¹∘f≡Id = funExt (secIsEq (equivIsEquiv (f⁻¹ .fst))) abstract - commutes : (CommRingEquiv→CommRingHom f⁻¹) ∘cr B .snd ≡ A .snd + commutes : (f⁻¹ .fst .fst) ∘ B .snd .fst ≡ A .snd .fst commutes = - CommRingHom≡ $ - f⁻¹ .fst .fst ∘ B .snd .fst ≡⟨ sym ( cong ((f⁻¹ .fst .fst) ∘_) (cong fst (f .snd))) ⟩ - f⁻¹ .fst .fst ∘ f .fst .fst .fst ∘ A .snd .fst ≡⟨ cong (_∘ A .snd .fst) f⁻¹∘f≡Id ⟩ + f⁻¹ .fst .fst ∘ B .snd .fst ≡⟨ sym ( cong ((f⁻¹ .fst .fst) ∘_) (cong fst (IsCommAlgebraHom.commutes (f .snd)))) ⟩ + f⁻¹ .fst .fst ∘ f .fst .fst ∘ A .snd .fst ≡⟨ cong (_∘ A .snd .fst) f⁻¹∘f≡Id ⟩ A .snd .fst ∎ diff --git a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda index 3b7417f854..70cff09a32 100644 --- a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda @@ -124,28 +124,25 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) where open CommRingStr (A .fst .snd) oneIdealQuotient : CommAlgebraEquiv (A / (1Ideal R A)) (UnitCommAlgebra R {ℓ' = ℓ}) - oneIdealQuotient = ? - -{- -.fst .fst = - withOpaqueStr $ + oneIdealQuotient .fst = + (withOpaqueStr $ isoToEquiv (iso ⟨ terminalMap R (A / 1Ideal R A) ⟩ₐ→ (λ _ → [ 0r ]) (λ _ → isPropUnit* _ _) - (elimProp (λ _ → squash/ _ _) (λ a → eq/ 0r a tt*))) - oneIdealQuotient .fst .snd = terminalMap R (A / 1Ideal R A) .fst .snd - oneIdealQuotient .snd = terminalMap R (A / (1Ideal R A)) .snd + (elimProp (λ _ → squash/ _ _) (λ a → eq/ 0r a tt*)))) + oneIdealQuotient .snd = makeOpaque $ terminalMap R (A / 1Ideal R A) .snd zeroIdealQuotient : CommAlgebraEquiv A (A / (0Ideal R A)) - zeroIdealQuotient .fst .fst = + zeroIdealQuotient .fst = withOpaqueStr $ let open RingTheory (CommRing→Ring (CommAlgebra→CommRing A)) in isoToEquiv (iso ⟨ (quotientHom A (0Ideal R A)) ⟩ₐ→ (rec is-set (λ x → x) λ x y x-y≡0 → equalByDifference x y x-y≡0) (elimProp (λ _ → squash/ _ _) λ _ → refl) λ _ → refl) - zeroIdealQuotient .fst .snd = quotientHom A (0Ideal R A) .fst .snd - zeroIdealQuotient .snd = quotientHom A (0Ideal R A) .snd + zeroIdealQuotient .snd = + makeOpaque $ quotientHom A (0Ideal R A) .snd + [_]/ : {R : CommRing ℓ} {A : CommAlgebra R ℓ} {I : IdealsIn R A} → (a : ⟨ A ⟩ₐ) → ⟨ A / I ⟩ₐ @@ -166,7 +163,6 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn R A) where refl ⟩ _ ≡⟨ CommRing.kernel≡I {R = CommAlgebra→CommRing A} I ⟩ I ∎ - module _ {R : CommRing ℓ} {A : CommAlgebra R ℓ} @@ -180,4 +176,3 @@ module _ opaque isZeroFromIdeal : (x : ⟨ A ⟩ₐ) → x ∈ (fst I) → ⟨ quotientHom A I ⟩ₐ→ x ≡ 0r isZeroFromIdeal x x∈I = eq/ x 0r (subst (_∈ fst I) (solve! (CommAlgebra→CommRing A)) x∈I ) --} From b2f707292fb3cd1635d27961db4ae2b37d5aaf5f Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 9 Aug 2024 17:22:48 +0200 Subject: [PATCH 74/83] fix --- Cubical/Algebra/CommAlgebra/FP/Instances.agda | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/FP/Instances.agda b/Cubical/Algebra/CommAlgebra/FP/Instances.agda index 76dd856a38..237c8b0a32 100644 --- a/Cubical/Algebra/CommAlgebra/FP/Instances.agda +++ b/Cubical/Algebra/CommAlgebra/FP/Instances.agda @@ -34,7 +34,6 @@ open import Cubical.Algebra.CommAlgebra.Ideal using (IdealsIn) open import Cubical.Algebra.CommAlgebra.FGIdeal -- open import Cubical.Algebra.CommAlgebra.Instances.Initial open import Cubical.Algebra.CommAlgebra.Instances.Unit -open import Cubical.Algebra.CommAlgebra.AsModule.Kernel open import Cubical.Algebra.CommAlgebra.FP.Base @@ -45,6 +44,7 @@ private module _ (R : CommRing ℓ) where open FinitePresentation +{- module _ (n : ℕ) where polynomialsFP : FinitePresentation R (Polynomials R n) @@ -56,7 +56,6 @@ module _ (R : CommRing ℓ) where equiv = invCommAlgebraEquiv {!zeroIdealQuotient!} } -{- {- Every (multivariate) polynomial algebra is finitely presented -} module _ (n : ℕ) where private From 89ecaa2594c5b45c49215598ac44d3c1969d89a7 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 23 Aug 2024 11:59:44 +0200 Subject: [PATCH 75/83] fix --- Cubical/Algebra/CommRing/Properties.agda | 5 + .../EquivCarac/AB-An[X]Bn[X].agda | 1 - .../UnivariateFun/Polyn-nPoly.agda | 39 ------- .../UnivariateHIT/Polyn-nPoly.agda | 110 ------------------ .../UnivariateList/Polyn-nPoly.agda | 35 ------ 5 files changed, 5 insertions(+), 185 deletions(-) delete mode 100644 Cubical/Algebra/Polynomials/UnivariateFun/Polyn-nPoly.agda delete mode 100644 Cubical/Algebra/Polynomials/UnivariateHIT/Polyn-nPoly.agda delete mode 100644 Cubical/Algebra/Polynomials/UnivariateList/Polyn-nPoly.agda diff --git a/Cubical/Algebra/CommRing/Properties.agda b/Cubical/Algebra/CommRing/Properties.agda index e08765fe21..cafdd6e6fa 100644 --- a/Cubical/Algebra/CommRing/Properties.agda +++ b/Cubical/Algebra/CommRing/Properties.agda @@ -216,6 +216,11 @@ module _ where CommRingEquiv≡ p = Σ≡Prop (λ _ → isPropIsCommRingHom _ _ _) (Σ≡Prop isPropIsEquiv p) + infixr 9 _∘cre_ -- same as functions + _∘cre_ : {R : CommRing ℓ} {S : CommRing ℓ'} {T : CommRing ℓ''} + → CommRingEquiv S T → CommRingEquiv R S → CommRingEquiv R T + g ∘cre f = compCommRingEquiv f g + module Exponentiation (R' : CommRing ℓ) where open CommRingStr (snd R') private R = fst R' diff --git a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/AB-An[X]Bn[X].agda b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/AB-An[X]Bn[X].agda index bc540c77bf..8f875ed23d 100644 --- a/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/AB-An[X]Bn[X].agda +++ b/Cubical/Algebra/Polynomials/Multivariate/EquivCarac/AB-An[X]Bn[X].agda @@ -72,7 +72,6 @@ module _ PB = PolyCommRing B' n open Iso - open CommRingEquivs lift-equiv-poly : (e : CommRingEquiv A' B') → CommRingEquiv (PolyCommRing A' n) (PolyCommRing B' n) fst (lift-equiv-poly (e , fstr)) = isoToEquiv is diff --git a/Cubical/Algebra/Polynomials/UnivariateFun/Polyn-nPoly.agda b/Cubical/Algebra/Polynomials/UnivariateFun/Polyn-nPoly.agda deleted file mode 100644 index 6c6d445f0b..0000000000 --- a/Cubical/Algebra/Polynomials/UnivariateFun/Polyn-nPoly.agda +++ /dev/null @@ -1,39 +0,0 @@ -{-# OPTIONS --safe --lossy-unification #-} -module Cubical.Algebra.Polynomials.UnivariateFun.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.Sigma - -open import Cubical.Algebra.Ring -open import Cubical.Algebra.GradedRing.DirectSumFun -open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommRing.Instances.Polynomials.UnivariatePolyHIT -open import Cubical.Algebra.CommRing.Instances.Polynomials.UnivariatePolyFun -open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly - -open import Cubical.Algebra.Polynomials.Multivariate.EquivCarac.Poly0-A -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 import Cubical.Algebra.Polynomials.UnivariateHIT.Polyn-nPoly - -open CommRingEquivs renaming (compCommRingEquiv to _∘-ecr_ ; invCommRingEquiv to inv-ecr) - -private variable - ℓ : Level - - ------------------------------------------------------------------------------ --- Definition - -Equiv1 : (A' : CommRing ℓ) → CommRingEquiv (nUnivariatePolyHIT A' 1) (nUnivariatePolyFun A' 1) -Equiv1 A' = CommRingEquiv-DirectSumGradedRing _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ - -Equiv-Polyn-nPolyFun : (A' : CommRing ℓ) → (n : ℕ) → CommRingEquiv (PolyCommRing A' n) (nUnivariatePolyFun A' n) -Equiv-Polyn-nPolyFun A' zero = CRE-Poly0-A A' -Equiv-Polyn-nPolyFun A' (suc n) = inv-ecr _ _ (CRE-PolyN∘M-PolyN+M A' 1 n) - ∘-ecr (lift-equiv-poly _ _ 1 (Equiv-Polyn-nPolyFun A' n) - ∘-ecr (Equiv-Polyn-nPolyHIT (nUnivariatePolyFun A' n) 1 - ∘-ecr Equiv1 (nUnivariatePolyFun A' n))) diff --git a/Cubical/Algebra/Polynomials/UnivariateHIT/Polyn-nPoly.agda b/Cubical/Algebra/Polynomials/UnivariateHIT/Polyn-nPoly.agda deleted file mode 100644 index 961d93510b..0000000000 --- a/Cubical/Algebra/Polynomials/UnivariateHIT/Polyn-nPoly.agda +++ /dev/null @@ -1,110 +0,0 @@ -{-# OPTIONS --safe --lossy-unification #-} -module Cubical.Algebra.Polynomials.UnivariateHIT.Polyn-nPoly where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.HLevels - -open import Cubical.Data.Nat renaming (_+_ to _+n_; _·_ to _·n_) -open import Cubical.Data.Vec -open import Cubical.Data.Sigma - -open import Cubical.Algebra.DirectSum.DirectSumHIT.Base -open import Cubical.Algebra.Ring -open import Cubical.Algebra.GradedRing.DirectSumFun -open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommRing.Instances.Polynomials.UnivariatePolyHIT -open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly - -open import Cubical.Algebra.Polynomials.Multivariate.EquivCarac.Poly0-A -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 - -module equiv1 (A'@(A , Astr) : CommRing ℓ) - where - - private - PA = PolyCommRing A' 1 - PAstr = snd PA - PA' = nUnivariatePolyHIT A' 1 - PA'str = snd PA' - - open CommRingStr - - directSense : fst (PolyCommRing A' 1) → fst (nUnivariatePolyHIT A' 1) - directSense = DS-Rec-Set.f _ _ _ _ - (is-set PA'str) - (0r PA'str) - (λ { (n ∷ []) a → base n a}) - (_+_ PA'str) - (+Assoc PA'str) - (+IdR PA'str) - (+Comm PA'str) - (λ { (n ∷ []) → base-neutral _ }) - λ { (n ∷ []) a b → base-add _ _ _} - - convSense : fst (nUnivariatePolyHIT A' 1) → fst (PolyCommRing A' 1) - convSense = DS-Rec-Set.f _ _ _ _ - (is-set PAstr) - (0r PAstr) - (λ n a → base (n ∷ []) a) - (_+_ PAstr) - (+Assoc PAstr) - (+IdR PAstr) - (+Comm PAstr) - (λ n → base-neutral _) - λ _ _ _ → base-add _ _ _ - - retr : (x : fst (PolyCommRing A' 1)) → convSense (directSense x) ≡ x - retr = DS-Ind-Prop.f _ _ _ _ (λ _ → is-set PAstr _ _) - refl - (λ { (n ∷ []) a → refl}) - (λ {U V} ind-U ind-V → cong₂ (_+_ PAstr) ind-U ind-V) - - sect : (x : fst (nUnivariatePolyHIT A' 1)) → (directSense (convSense x) ≡ x) - sect = DS-Ind-Prop.f _ _ _ _ (λ _ → is-set PA'str _ _) - refl - (λ n a → refl) - (λ {U V} ind-U ind-V → cong₂ (_+_ PA'str) ind-U ind-V) - - converseSense-pres· : (x y : fst (PolyCommRing A' 1)) → - directSense (_·_ PAstr x y) ≡ _·_ PA'str (directSense x) (directSense y) - converseSense-pres· = DS-Ind-Prop.f _ _ _ _ - (λ _ → isPropΠ λ _ → is-set PA'str _ _) - (λ _ → refl) - (λ { (n ∷ []) a → DS-Ind-Prop.f _ _ _ _ (λ _ → is-set PA'str _ _) - refl - (λ { (m ∷ []) b → refl}) - λ {U V} ind-U ind-V → cong₂ (_+_ PA'str) ind-U ind-V}) - λ {U V} ind-U ind-V y → cong₂ (_+_ PA'str) (ind-U y) (ind-V y) - - open Iso - - equivR : CommRingEquiv PA PA' - fst equivR = isoToEquiv is - where - is : Iso (PA .fst) (PA' .fst) - fun is = directSense - inv is = convSense - rightInv is = sect - leftInv is = retr - snd equivR = makeIsCommRingHom refl (λ _ _ → refl) converseSense-pres· - - -open equiv1 - -Equiv-Polyn-nPolyHIT : (A' : CommRing ℓ) → (n : ℕ) → CommRingEquiv (PolyCommRing A' n) (nUnivariatePolyHIT A' n) -Equiv-Polyn-nPolyHIT A' zero = CRE-Poly0-A A' -Equiv-Polyn-nPolyHIT A' (suc n) = inv-ecr _ _ (CRE-PolyN∘M-PolyN+M A' 1 n) - ∘-ecr (lift-equiv-poly _ _ 1 (Equiv-Polyn-nPolyHIT A' n) - ∘-ecr equivR (nUnivariatePolyHIT A' n)) diff --git a/Cubical/Algebra/Polynomials/UnivariateList/Polyn-nPoly.agda b/Cubical/Algebra/Polynomials/UnivariateList/Polyn-nPoly.agda deleted file mode 100644 index 72053b9405..0000000000 --- a/Cubical/Algebra/Polynomials/UnivariateList/Polyn-nPoly.agda +++ /dev/null @@ -1,35 +0,0 @@ -{-# OPTIONS --safe --lossy-unification #-} -module Cubical.Algebra.Polynomials.UnivariateList.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.Polynomials.UnivariatePolyList -open import Cubical.Algebra.CommRing.Instances.Polynomials.MultivariatePoly - -open import Cubical.Algebra.Polynomials.Multivariate.EquivCarac.Poly0-A -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 import Cubical.Algebra.Polynomials.UnivariateList.Poly1-1Poly - -open CommRingEquivs renaming (compCommRingEquiv to _∘-ecr_ ; invCommRingEquiv to inv-ecr) - -private variable - ℓ : Level - - ------------------------------------------------------------------------------ --- Definition - -Equiv-Polyn-nPoly : (A' : CommRing ℓ) → (n : ℕ) → CommRingEquiv (PolyCommRing A' n) (nUnivariatePolyList 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 _ _ 1 (Equiv-Polyn-nPoly A' n) - ∘-ecr CRE-Poly1-Poly: (nUnivariatePolyList A' n)) From 7293a0db74e84f33f08d4a207fba998aeb9565b2 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 26 Aug 2024 15:03:21 +0200 Subject: [PATCH 76/83] solver examples with new comm algebras --- Cubical/Tactics/CommRingSolver/Examples.agda | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/Cubical/Tactics/CommRingSolver/Examples.agda b/Cubical/Tactics/CommRingSolver/Examples.agda index c931620558..71242c6a90 100644 --- a/Cubical/Tactics/CommRingSolver/Examples.agda +++ b/Cubical/Tactics/CommRingSolver/Examples.agda @@ -10,7 +10,7 @@ open import Cubical.Data.Nat using (ℕ; suc; zero) open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Instances.Int -open import Cubical.Algebra.CommAlgebra.AsModule +open import Cubical.Algebra.CommAlgebra open import Cubical.Tactics.CommRingSolver open import Cubical.Tactics.CommRingSolver.RawAlgebra using (scalar) @@ -114,15 +114,16 @@ module Test (R : CommRing ℓ) (x y z : fst R) where ex11 x = solve! R module _ (R : CommRing ℓ) (A : CommAlgebra R ℓ') where - open CommAlgebraStr {{...}} + open CommRingStr ⦃...⦄ private instance - _ = (snd A) + _ : CommRingStr ⟨ A ⟩ₐ + _ = (A .fst .snd) {- The ring solver should also be able to deal with more complicated arguments and operations with that are not given as the exact names in CommRingStr. -} - ex12 : (x y : ⟨ A ⟩) → x + y ≡ y + x + ex12 : (x y : ⟨ A ⟩ₐ) → x + y ≡ y + x ex12 x y = solve! (CommAlgebra→CommRing A) module TestInPlaceSolving (R : CommRing ℓ) where From bbd32fd5bc19a9a25ea3dfaccaf686f97b491aff Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 26 Aug 2024 15:38:08 +0200 Subject: [PATCH 77/83] fix --- Cubical/Algebra/CommAlgebra/Base.agda | 2 + .../Experiments/ZariskiLatticeBasicOpens.agda | 84 ++++++++++--------- 2 files changed, 48 insertions(+), 38 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/Base.agda index 3be2bfd9a2..2cd0555673 100644 --- a/Cubical/Algebra/CommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/Base.agda @@ -224,6 +224,8 @@ module _ {R : CommRing ℓ} where CommAlgebra→Ring : (A : CommAlgebra R ℓ') → Ring ℓ' CommAlgebra→Ring A = CommRing→Ring (fst A) + CommAlgebra→CommRingStr : (A : CommAlgebra R ℓ') → CommRingStr ⟨ A ⟩ₐ + CommAlgebra→CommRingStr A = A .fst .snd CommAlgebraHom→RingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} → CommAlgebraHom A B diff --git a/Cubical/Experiments/ZariskiLatticeBasicOpens.agda b/Cubical/Experiments/ZariskiLatticeBasicOpens.agda index b5820781d9..09814e7790 100644 --- a/Cubical/Experiments/ZariskiLatticeBasicOpens.agda +++ b/Cubical/Experiments/ZariskiLatticeBasicOpens.agda @@ -23,7 +23,6 @@ open import Cubical.Data.Sigma.Properties open import Cubical.Data.FinData open import Cubical.Relation.Nullary open import Cubical.Relation.Binary --- open import Cubical.Relation.Binary.Poset open import Cubical.Algebra.Ring open import Cubical.Algebra.Algebra @@ -31,8 +30,9 @@ open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Localisation.Base open import Cubical.Algebra.CommRing.Localisation.UniversalProperty open import Cubical.Algebra.CommRing.Localisation.InvertingElements -open import Cubical.Algebra.CommAlgebra.AsModule -open import Cubical.Algebra.CommAlgebra.AsModule.Localisation +open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebra.Univalence +open import Cubical.Algebra.CommAlgebra.Localisation open import Cubical.Tactics.CommRingSolver open import Cubical.Algebra.Semilattice @@ -49,8 +49,12 @@ private module Presheaf (A' : CommRing ℓ) where - open CommRingStr (snd A') renaming (_·_ to _·r_ ; ·Comm to ·r-comm ; ·Assoc to ·rAssoc - ; ·IdL to ·rLid ; ·IdR to ·rRid) + open CommRingStr ⦃...⦄ -- renaming (_·_ to _·r_ ; ·Comm to ·r-comm ; ·Assoc to ·rAssoc + -- ; ·IdL to ·rLid ; ·IdR to ·rRid) + + private + instance + _ = A' .snd open Exponentiation A' open CommRingTheory A' open InvertingElementsBase A' @@ -62,32 +66,32 @@ module Presheaf (A' : CommRing ℓ) where A[1/_] : A → CommAlgebra A' ℓ A[1/ x ] = AlgLoc.S⁻¹RAsCommAlg A' [ x ⁿ|n≥0] (powersFormMultClosedSubset _) - A[1/_]ˣ : (x : A) → ℙ (fst A[1/ x ]) + A[1/_]ˣ : (x : A) → ℙ ⟨ A[1/ x ] ⟩ₐ A[1/ x ]ˣ = (CommAlgebra→CommRing A[1/ x ]) ˣ _≼_ : A → A → Type ℓ - x ≼ y = ∃[ n ∈ ℕ ] Σ[ z ∈ A ] x ^ n ≡ z ·r y -- rad(x) ⊆ rad(y) + x ≼ y = ∃[ n ∈ ℕ ] Σ[ z ∈ A ] x ^ n ≡ z · y -- rad(x) ⊆ rad(y) -- ≼ is a pre-order: Refl≼ : isRefl _≼_ - Refl≼ x = PT.∣ 1 , 1r , ·r-comm _ _ ∣₁ + Refl≼ x = PT.∣ 1 , 1r , ·Comm _ _ ∣₁ Trans≼ : isTrans _≼_ Trans≼ x y z = map2 Trans≼Σ where - Trans≼Σ : Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a ·r y - → Σ[ n ∈ ℕ ] Σ[ a ∈ A ] y ^ n ≡ a ·r z - → Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a ·r z - Trans≼Σ (n , a , p) (m , b , q) = n ·ℕ m , (a ^ m ·r b) , path + Trans≼Σ : Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a · y + → Σ[ n ∈ ℕ ] Σ[ a ∈ A ] y ^ n ≡ a · z + → Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a · z + Trans≼Σ (n , a , p) (m , b , q) = n ·ℕ m , (a ^ m · b) , path where - path : x ^ (n ·ℕ m) ≡ a ^ m ·r b ·r z + path : x ^ (n ·ℕ m) ≡ a ^ m · b · z path = x ^ (n ·ℕ m) ≡⟨ ^-rdist-·ℕ x n m ⟩ (x ^ n) ^ m ≡⟨ cong (_^ m) p ⟩ - (a ·r y) ^ m ≡⟨ ^-ldist-· a y m ⟩ - a ^ m ·r y ^ m ≡⟨ cong (a ^ m ·r_) q ⟩ - a ^ m ·r (b ·r z) ≡⟨ ·rAssoc _ _ _ ⟩ - a ^ m ·r b ·r z ∎ + (a · y) ^ m ≡⟨ ^-ldist-· a y m ⟩ + a ^ m · y ^ m ≡⟨ cong (a ^ m ·_) q ⟩ + a ^ m · (b · z) ≡⟨ ·Assoc _ _ _ ⟩ + a ^ m · b · z ∎ R : A → A → Type ℓ @@ -105,23 +109,24 @@ module Presheaf (A' : CommRing ℓ) where powerIs≽ : (x a : A) → x ∈ [ a ⁿ|n≥0] → a ≼ x powerIs≽ x a = map powerIs≽Σ where - powerIs≽Σ : Σ[ n ∈ ℕ ] (x ≡ a ^ n) → Σ[ n ∈ ℕ ] Σ[ z ∈ A ] (a ^ n ≡ z ·r x) - powerIs≽Σ (n , p) = n , 1r , sym p ∙ sym (·rLid _) + powerIs≽Σ : Σ[ n ∈ ℕ ] (x ≡ a ^ n) → Σ[ n ∈ ℕ ] Σ[ z ∈ A ] (a ^ n ≡ z · x) + powerIs≽Σ (n , p) = n , 1r , sym p ∙ sym (·IdL _) module ≼ToLoc (x y : A) where private instance - _ = snd A[1/ x ] + _ = CommAlgebra→CommRingStr A[1/ x ] + _ = CommAlgebra→CommAlgebraStr A[1/ x ] - lemma : x ≼ y → y ⋆ 1a ∈ A[1/ x ]ˣ -- y/1 ∈ A[1/x]ˣ - lemma = PT.rec (A[1/ x ]ˣ (y ⋆ 1a) .snd) lemmaΣ + lemma : x ≼ y → y ⋆ 1r ∈ A[1/ x ]ˣ -- y/1 ∈ A[1/x]ˣ + lemma = PT.rec (A[1/ x ]ˣ (y ⋆ 1r) .snd) lemmaΣ where - path1 : (y z : A) → 1r ·r (y ·r 1r ·r z) ·r 1r ≡ z ·r y + path1 : (y z : A) → 1r · (y · 1r · z) · 1r ≡ z · y path1 _ _ = solve! A' - path2 : (xn : A) → xn ≡ 1r ·r 1r ·r (1r ·r 1r ·r xn) + path2 : (xn : A) → xn ≡ 1r · 1r · (1r · 1r · xn) path2 _ = solve! A' - lemmaΣ : Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a ·r y → y ⋆ 1a ∈ A[1/ x ]ˣ + lemmaΣ : Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a · y → y ⋆ 1r ∈ A[1/ x ]ˣ lemmaΣ (n , z , p) = [ z , (x ^ n) , PT.∣ n , refl ∣₁ ] -- xⁿ≡zy → y⁻¹ ≡ z/xⁿ , eq/ _ _ ((1r , powersFormMultClosedSubset _ .containsOne) , (path1 _ _ ∙∙ sym p ∙∙ path2 _)) @@ -130,8 +135,10 @@ module Presheaf (A' : CommRing ℓ) where private [yⁿ|n≥0] = [ y ⁿ|n≥0] instance - _ = snd A[1/ x ] - lemma : ∀ (s : A) → s ∈ [yⁿ|n≥0] → s ⋆ 1a ∈ A[1/ x ]ˣ + _ = CommAlgebra→CommRingStr A[1/ x ] + _ = CommAlgebra→CommAlgebraStr A[1/ x ] + + lemma : ∀ (s : A) → s ∈ [yⁿ|n≥0] → s ⋆ 1r ∈ A[1/ x ]ˣ lemma _ s∈[yⁿ|n≥0] = ≼ToLoc.lemma _ _ (Trans≼ _ y _ x≼y (powerIs≽ _ _ s∈[yⁿ|n≥0])) @@ -157,23 +164,24 @@ module Presheaf (A' : CommRing ℓ) where -- Multiplication lifts to the quotient and corresponds to intersection -- of basic opens, i.e. we get a meet-semilattice with: _∧/_ : A / R → A / R → A / R - _∧/_ = setQuotSymmBinOp (RequivRel .reflexive) (RequivRel .transitive) _·r_ - (λ a b → subst (λ x → R (a ·r b) x) (·r-comm a b) (RequivRel .reflexive (a ·r b))) ·r-lcoh + _∧/_ = setQuotSymmBinOp (RequivRel .reflexive) (RequivRel .transitive) _·_ + (λ a b → subst (λ x → R (a · b) x) (·Comm a b) (RequivRel .reflexive (a · b))) ·-lcoh where - ·r-lcoh-≼ : (x y z : A) → x ≼ y → (x ·r z) ≼ (y ·r z) - ·r-lcoh-≼ x y z = map ·r-lcoh-≼Σ + ·-lcoh-≼ : (x y z : A) → x ≼ y → (x · z) ≼ (y · z) + ·-lcoh-≼ x y z = map ·-lcoh-≼Σ where - path : (x z a y zn : A) → x ·r z ·r (a ·r y ·r zn) ≡ x ·r zn ·r a ·r (y ·r z) + path : (x z a y zn : A) → x · z · (a · y · zn) ≡ x · zn · a · (y · z) path _ _ _ _ _ = solve! A' - ·r-lcoh-≼Σ : Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a ·r y - → Σ[ n ∈ ℕ ] Σ[ a ∈ A ] (x ·r z) ^ n ≡ a ·r (y ·r z) - ·r-lcoh-≼Σ (n , a , p) = suc n , (x ·r z ^ n ·r a) , (cong (x ·r z ·r_) (^-ldist-· _ _ _) - ∙∙ cong (λ v → x ·r z ·r (v ·r z ^ n)) p + ·-lcoh-≼Σ : Σ[ n ∈ ℕ ] Σ[ a ∈ A ] x ^ n ≡ a · y + → Σ[ n ∈ ℕ ] Σ[ a ∈ A ] (x · z) ^ n ≡ a · (y · z) + ·-lcoh-≼Σ (n , a , p) = suc n , (x · z ^ n · a) , (cong (x · z ·_) (^-ldist-· _ _ _) + ∙∙ cong (λ v → x · z · (v · z ^ n)) p ∙∙ path _ _ _ _ _) - ·r-lcoh : (x y z : A) → R x y → R (x ·r z) (y ·r z) - ·r-lcoh x y z Rxy = ·r-lcoh-≼ x y z (Rxy .fst) , ·r-lcoh-≼ y x z (Rxy .snd) + ·-lcoh : (x y z : A) → R x y → R (x · z) (y · z) + ·-lcoh x y z Rxy = ·-lcoh-≼ x y z (Rxy .fst) , ·-lcoh-≼ y x z (Rxy .snd) + {- BasicOpens : Semilattice ℓ BasicOpens = makeSemilattice [ 1r ] _∧/_ squash/ From 9286b261ad8a18c88419cbeecd07e09ff2f9acdd Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 26 Aug 2024 16:39:08 +0200 Subject: [PATCH 78/83] more boilterplate, morphism contructor --- Cubical/Algebra/CommAlgebra/Base.agda | 15 +++++++-- Cubical/Algebra/CommAlgebra/Properties.agda | 34 +++++++++++++++++++++ 2 files changed, 46 insertions(+), 3 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/Base.agda index 2cd0555673..10668f649f 100644 --- a/Cubical/Algebra/CommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/Base.agda @@ -138,9 +138,9 @@ module _ {R : CommRing ℓ} where idCommAlgebraHom = idCAlgHom compCommAlgebraHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} - → (g : CommAlgebraHom B C) → (f : CommAlgebraHom A B) + → (f : CommAlgebraHom A B) → (g : CommAlgebraHom B C) → CommAlgebraHom A C - compCommAlgebraHom {A = A} {B = B} {C = C} g f = + compCommAlgebraHom {A = A} {B = B} {C = C} f g = CommRingHom→CommAlgebraHom (⟨ g ⟩ᵣ→ ∘cr ⟨ f ⟩ᵣ→) pasting where opaque @@ -160,7 +160,7 @@ module _ {R : CommRing ℓ} where infixr 9 _∘ca_ -- same as functions _∘ca_ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} → CommAlgebraHom B C → CommAlgebraHom A B → CommAlgebraHom A C - g ∘ca f = compCommAlgebraHom g f + g ∘ca f = compCommAlgebraHom f g opaque CommAlgebraHom≡ : @@ -227,6 +227,15 @@ module _ {R : CommRing ℓ} where CommAlgebra→CommRingStr : (A : CommAlgebra R ℓ') → CommRingStr ⟨ A ⟩ₐ CommAlgebra→CommRingStr A = A .fst .snd + CommAlgebra→Set : (A : CommAlgebra R ℓ') → hSet ℓ' + CommAlgebra→Set A = ⟨ A ⟩ₐ , is-set + where open CommRingStr (CommAlgebra→CommRingStr A) using (is-set) + + isSetCommAlgebraHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') + → isSet (CommAlgebraHom A B) + isSetCommAlgebraHom A B = isSetΣSndProp (isSet→ is-set) (λ _ → isPropIsCommAlgebraHom _ _ _) + where open CommRingStr (CommAlgebra→CommRingStr B) using (is-set) + CommAlgebraHom→RingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} → CommAlgebraHom A B → RingHom (CommAlgebra→Ring A) (CommAlgebra→Ring B) diff --git a/Cubical/Algebra/CommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebra/Properties.agda index 0cf21ba7ac..068fb46522 100644 --- a/Cubical/Algebra/CommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/Properties.agda @@ -46,3 +46,37 @@ module _ {R : CommRing ℓ} where f⁻¹ .fst .fst ∘ B .snd .fst ≡⟨ sym ( cong ((f⁻¹ .fst .fst) ∘_) (cong fst (IsCommAlgebraHom.commutes (f .snd)))) ⟩ f⁻¹ .fst .fst ∘ f .fst .fst ∘ A .snd .fst ≡⟨ cong (_∘ A .snd .fst) f⁻¹∘f≡Id ⟩ A .snd .fst ∎ + +module _ + -- Variable generalization would fail below without the module parameters A and B. + {R : CommRing ℓ} + {A : CommAlgebra R ℓ'} + {B : CommAlgebra R ℓ''} + {f : ⟨ A ⟩ₐ → ⟨ B ⟩ₐ} + where + + open CommAlgebraStr ⦃...⦄ + open CommRingStr ⦃...⦄ + private instance + _ = CommAlgebra→CommRingStr A + _ = CommAlgebra→CommRingStr B + _ = CommAlgebra→CommAlgebraStr A + _ = CommAlgebra→CommAlgebraStr B + + module _ + (p1 : f 1r ≡ 1r) + (p+ : (x y : ⟨ A ⟩ₐ) → f (x + y) ≡ f x + f y) + (p· : (x y : ⟨ A ⟩ₐ) → f (x · y) ≡ f x · f y) + (p⋆ : (r : ⟨ R ⟩) (x : ⟨ A ⟩ₐ) → f (r ⋆ x) ≡ r ⋆ f x) + where + + makeIsCommAlgebraHom : IsCommAlgebraHom A B f + makeIsCommAlgebraHom .IsCommAlgebraHom.isCommRingHom = makeIsCommRingHom p1 p+ p· + makeIsCommAlgebraHom .IsCommAlgebraHom.commutes = + CommRingHom≡ $ funExt λ r → f (A .snd .fst r) ≡⟨ cong f (sym (·IdR _)) ⟩ + f ((A .snd .fst r) · 1r) ≡⟨⟩ + f (r ⋆ 1r) ≡⟨ p⋆ _ _ ⟩ + r ⋆ (f 1r) ≡⟨⟩ + (B .snd .fst r) · f 1r ≡⟨ cong ((B .snd .fst r) ·_) p1 ⟩ + (B .snd .fst r) · 1r ≡⟨ ·IdR _ ⟩ + B .snd .fst r ∎ From 14c6d4cfbb6c3d44c9aee1de33119bd79f9dbba7 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 30 Aug 2024 16:27:56 +0200 Subject: [PATCH 79/83] polynomial ring is fp --- Cubical/Algebra/CommAlgebra/FGIdeal.agda | 4 ++-- Cubical/Algebra/CommAlgebra/FP/Base.agda | 5 ++--- Cubical/Algebra/CommAlgebra/FP/Instances.agda | 21 +++++++++++++++---- 3 files changed, 21 insertions(+), 9 deletions(-) diff --git a/Cubical/Algebra/CommAlgebra/FGIdeal.agda b/Cubical/Algebra/CommAlgebra/FGIdeal.agda index 1ef76185f2..d9b1a25ee0 100644 --- a/Cubical/Algebra/CommAlgebra/FGIdeal.agda +++ b/Cubical/Algebra/CommAlgebra/FGIdeal.agda @@ -33,5 +33,5 @@ syntax generatedIdeal A V = ⟨ V ⟩[ A ] module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) where open CommRingStr (A .fst .snd) - 0FGIdeal : {n : ℕ} → ⟨ replicateFinVec n 0r ⟩[ A ] ≡ (0Ideal R A) - 0FGIdeal = 0FGIdealCommRing (CommAlgebra→CommRing A) + 0FGIdeal≡0Ideal : {n : ℕ} → ⟨ replicateFinVec n 0r ⟩[ A ] ≡ (0Ideal R A) + 0FGIdeal≡0Ideal = 0FGIdealCommRing (CommAlgebra→CommRing A) diff --git a/Cubical/Algebra/CommAlgebra/FP/Base.agda b/Cubical/Algebra/CommAlgebra/FP/Base.agda index 85abeaf3bd..2ea0b39dcb 100644 --- a/Cubical/Algebra/CommAlgebra/FP/Base.agda +++ b/Cubical/Algebra/CommAlgebra/FP/Base.agda @@ -27,9 +27,8 @@ private ℓ ℓ' : Level module _ (R : CommRing ℓ) where - abstract - Polynomials : (n : ℕ) → CommAlgebra R ℓ - Polynomials n = R [ Fin n ]ₐ + Polynomials : (n : ℕ) → CommAlgebra R ℓ + Polynomials n = R [ Fin n ]ₐ FPCAlgebra : {m : ℕ} (n : ℕ) (relations : FinVec ⟨ Polynomials n ⟩ₐ m) → CommAlgebra R ℓ FPCAlgebra n relations = Polynomials n / ⟨ relations ⟩[ Polynomials n ] diff --git a/Cubical/Algebra/CommAlgebra/FP/Instances.agda b/Cubical/Algebra/CommAlgebra/FP/Instances.agda index 237c8b0a32..09738084d1 100644 --- a/Cubical/Algebra/CommAlgebra/FP/Instances.agda +++ b/Cubical/Algebra/CommAlgebra/FP/Instances.agda @@ -7,7 +7,7 @@ * R/⟨x₁,...,xₙ⟩ = R[⊥]/⟨x₁,...,xₙ⟩ * R/⟨x⟩ (as special case of the above) -} -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.Algebra.CommAlgebra.FP.Instances where open import Cubical.Foundations.Prelude @@ -30,7 +30,7 @@ open import Cubical.Algebra.CommRing.FGIdeal using (inclOfFGIdeal) open import Cubical.Algebra.CommAlgebra open import Cubical.Algebra.CommAlgebra.Instances.Polynomials open import Cubical.Algebra.CommAlgebra.QuotientAlgebra -open import Cubical.Algebra.CommAlgebra.Ideal using (IdealsIn) +open import Cubical.Algebra.CommAlgebra.Ideal using (IdealsIn; 0Ideal) open import Cubical.Algebra.CommAlgebra.FGIdeal -- open import Cubical.Algebra.CommAlgebra.Instances.Initial open import Cubical.Algebra.CommAlgebra.Instances.Unit @@ -44,18 +44,31 @@ private module _ (R : CommRing ℓ) where open FinitePresentation -{- module _ (n : ℕ) where + private + abstract + p : 0Ideal R (Polynomials R n) ≡ ⟨ emptyFinVec ⟨ Polynomials R n ⟩ₐ ⟩[ _ ] + p = sym $ 0FGIdeal≡0Ideal (Polynomials R n) + + compute : + CommAlgebraEquiv (Polynomials R n) $ (Polynomials R n) / ⟨ emptyFinVec ⟨ Polynomials R n ⟩ₐ ⟩[ _ ] + compute = + transport (λ i → CommAlgebraEquiv (Polynomials R n) ((Polynomials R n) / (p i))) $ + zeroIdealQuotient (Polynomials R n) + polynomialsFP : FinitePresentation R (Polynomials R n) polynomialsFP = record { n = n ; m = 0 ; relations = emptyFinVec ⟨ Polynomials R n ⟩ₐ ; - equiv = invCommAlgebraEquiv {!zeroIdealQuotient!} + equiv = invCommAlgebraEquiv compute } +{- + + {- Every (multivariate) polynomial algebra is finitely presented -} module _ (n : ℕ) where private From ed48931b83444c3acd4de33f6f42a6e85ed435cc Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Fri, 30 Aug 2024 17:15:01 +0200 Subject: [PATCH 80/83] add more calcualtions --- Cubical/Algebra/CommAlgebra/FP/Instances.agda | 52 +--- .../Polynomials/Typevariate/OnCoproduct.agda | 3 +- .../Polynomials/Typevariate/Properties.agda | 275 ++--------------- .../Typevariate/UniversalProperty.agda | 285 ++++++++++++++++++ 4 files changed, 312 insertions(+), 303 deletions(-) create mode 100644 Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/UniversalProperty.agda diff --git a/Cubical/Algebra/CommAlgebra/FP/Instances.agda b/Cubical/Algebra/CommAlgebra/FP/Instances.agda index 09738084d1..6e3a423104 100644 --- a/Cubical/Algebra/CommAlgebra/FP/Instances.agda +++ b/Cubical/Algebra/CommAlgebra/FP/Instances.agda @@ -45,6 +45,7 @@ private module _ (R : CommRing ℓ) where open FinitePresentation + {- Every (multivariate) polynomial algebra is finitely presented -} module _ (n : ℕ) where private abstract @@ -52,7 +53,8 @@ module _ (R : CommRing ℓ) where p = sym $ 0FGIdeal≡0Ideal (Polynomials R n) compute : - CommAlgebraEquiv (Polynomials R n) $ (Polynomials R n) / ⟨ emptyFinVec ⟨ Polynomials R n ⟩ₐ ⟩[ _ ] + CommAlgebraEquiv (Polynomials R n) + ((Polynomials R n) / ⟨ emptyFinVec ⟨ Polynomials R n ⟩ₐ ⟩[ _ ]) compute = transport (λ i → CommAlgebraEquiv (Polynomials R n) ((Polynomials R n) / (p i))) $ zeroIdealQuotient (Polynomials R n) @@ -69,54 +71,6 @@ module _ (R : CommRing ℓ) where {- - {- Every (multivariate) polynomial algebra is finitely presented -} - module _ (n : ℕ) where - private - A : CommAlgebra R ℓ - A = Polynomials n - - emptyGen : FinVec (fst A) 0 - emptyGen = λ () - - B : CommAlgebra R ℓ - B = FPAlgebra n emptyGen - - polynomialAlgFP : FinitePresentation A - FinitePresentation.n polynomialAlgFP = n - m polynomialAlgFP = 0 - relations polynomialAlgFP = emptyGen - equiv polynomialAlgFP = - -- Idea: A and B enjoy the same universal property. - toAAsEquiv , snd toA - where - toA : CommAlgebraHom B A - toA = inducedHom n emptyGen A Construction.var (λ ()) - fromA : CommAlgebraHom A B - fromA = freeInducedHom B (generator _ _) - open AlgebraHoms - inverse1 : fromA ∘a toA ≡ idAlgebraHom _ - inverse1 = - fromA ∘a toA - ≡⟨ sym (unique _ _ B _ _ _ (λ i → cong (fst fromA) ( - fst toA (generator n emptyGen i) - ≡⟨ inducedHomOnGenerators _ _ _ _ _ _ ⟩ - Construction.var i - ∎))) ⟩ - inducedHom n emptyGen B (generator _ _) (relationsHold _ _) - ≡⟨ unique _ _ B _ _ _ (λ i → refl) ⟩ - idAlgebraHom _ - ∎ - inverse2 : toA ∘a fromA ≡ idAlgebraHom _ - inverse2 = isoFunInjective (homMapIso A) _ _ ( - evaluateAt A (toA ∘a fromA) ≡⟨ sym (naturalEvR {A = B} {B = A} toA fromA) ⟩ - fst toA ∘ evaluateAt B fromA ≡⟨ refl ⟩ - fst toA ∘ generator _ _ ≡⟨ funExt (inducedHomOnGenerators _ _ _ _ _)⟩ - Construction.var ∎) - toAAsEquiv : ⟨ B ⟩ ≃ ⟨ A ⟩ - toAAsEquiv = isoToEquiv (iso (fst toA) - (fst fromA) - (λ a i → fst (inverse2 i) a) - (λ b i → fst (inverse1 i) b)) {- The initial R-algebra is finitely presented -} private diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/OnCoproduct.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/OnCoproduct.agda index 3290a2276f..7ce31c7d22 100644 --- a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/OnCoproduct.agda +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/OnCoproduct.agda @@ -17,7 +17,8 @@ open import Cubical.Data.Sum as ⊎ open import Cubical.Data.Sigma open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate +open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.Base +open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.UniversalProperty private variable diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda index 0af77c592c..bde4d50eae 100644 --- a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda @@ -1,18 +1,12 @@ {-# OPTIONS --safe #-} module Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.Properties where -{- - This file contains - * an elimination principle for proving some proposition for all elements of R[I]ᵣ - ('elimProp') - * definitions of the induced maps appearing in the universal property of R[I], --} open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.HLevels open import Cubical.Foundations.Structure -open import Cubical.Foundations.Function hiding (const) +open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism open import Cubical.Data.Sigma.Properties using (Σ≡Prop) @@ -20,9 +14,9 @@ open import Cubical.HITs.SetTruncation open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.Base -open import Cubical.Algebra.Ring.Properties +open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.UniversalProperty public -open import Cubical.Data.Empty +open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Sigma open import Cubical.Tactics.CommRingSolver @@ -31,255 +25,30 @@ private variable ℓ ℓ' ℓ'' : Level -module _ {R : CommRing ℓ} {I : Type ℓ'} where +module _ {R : CommRing ℓ} where open CommRingStr ⦃...⦄ private instance _ = snd R - _ = snd (R [ I ]) - module C = Construction - open C using (const) + polynomialsOn⊥ : CommRingEquiv (R [ ⊥ ]) R + polynomialsOn⊥ = + isoToEquiv + (iso (to .fst) (from .fst) + (λ x → cong (λ f → f .fst x) to∘from) + (λ x → cong (λ f → f .fst x) from∘to)) , + isHom + where to : CommRingHom (R [ ⊥ ]) R + to = inducedHom R (idCommRingHom R) ⊥.rec - {- - Construction of the 'elimProp' eliminator. - -} - module _ - {P : ⟨ R [ I ] ⟩ → Type ℓ''} - (isPropP : {x : _} → isProp (P x)) - (onVar : {x : I} → P (C.var x)) - (onConst : {r : ⟨ R ⟩} → P (const r)) - (on+ : {x y : ⟨ R [ I ] ⟩} → P x → P y → P (x + y)) - (on· : {x y : ⟨ R [ I ] ⟩} → P x → P y → P (x · y)) - where + from : CommRingHom R (R [ ⊥ ]) + from = constPolynomial R ⊥ - private - on- : {x : ⟨ R [ I ] ⟩} → P x → P (- x) - on- {x} Px = subst P (minusByMult x) (on· onConst Px) - where minusByMult : (x : _) → (const (- 1r)) · x ≡ - x - minusByMult x = - (const (- 1r)) · x ≡⟨ cong (_· x) (pres- 1r) ⟩ - (- const (1r)) · x ≡⟨ cong (λ u → (- u) · x) pres1 ⟩ - (- 1r) · x ≡⟨ sym (-IsMult-1 x) ⟩ - - x ∎ - where open RingTheory (CommRing→Ring (R [ I ])) using (-IsMult-1) - open IsCommRingHom (constPolynomial R I .snd) + to∘from : to ∘cr from ≡ idCommRingHom R + to∘from = CommRingHom≡ refl - -- A helper to deal with the path constructor cases. - mkPathP : - {x0 x1 : ⟨ R [ I ] ⟩} → - (p : x0 ≡ x1) → - (Px0 : P (x0)) → - (Px1 : P (x1)) → - PathP (λ i → P (p i)) Px0 Px1 - mkPathP _ = isProp→PathP λ i → isPropP + from∘to : from ∘cr to ≡ idCommRingHom (R [ ⊥ ]) + from∘to = idByIdOnVars (from ∘cr to) refl (funExt ⊥.elim) - elimProp : ((x : _) → P x) - - elimProp (C.var _) = onVar - elimProp (const _) = onConst - elimProp (x C.+ y) = on+ (elimProp x) (elimProp y) - elimProp (C.- x) = on- (elimProp x) - elimProp (x C.· y) = on· (elimProp x) (elimProp y) - - elimProp (C.+-assoc x y z i) = - mkPathP (C.+-assoc x y z) - (on+ (elimProp x) (on+ (elimProp y) (elimProp z))) - (on+ (on+ (elimProp x) (elimProp y)) (elimProp z)) - i - elimProp (C.+-rid x i) = - mkPathP (C.+-rid x) - (on+ (elimProp x) onConst) - (elimProp x) - i - elimProp (C.+-rinv x i) = - mkPathP (C.+-rinv x) - (on+ (elimProp x) (on- (elimProp x))) - onConst - i - elimProp (C.+-comm x y i) = - mkPathP (C.+-comm x y) - (on+ (elimProp x) (elimProp y)) - (on+ (elimProp y) (elimProp x)) - i - elimProp (C.·-assoc x y z i) = - mkPathP (C.·-assoc x y z) - (on· (elimProp x) (on· (elimProp y) (elimProp z))) - (on· (on· (elimProp x) (elimProp y)) (elimProp z)) - i - elimProp (C.·-lid x i) = - mkPathP (C.·-lid x) - (on· onConst (elimProp x)) - (elimProp x) - i - elimProp (C.·-comm x y i) = - mkPathP (C.·-comm x y) - (on· (elimProp x) (elimProp y)) - (on· (elimProp y) (elimProp x)) - i - elimProp (C.ldist x y z i) = - mkPathP (C.ldist x y z) - (on· (on+ (elimProp x) (elimProp y)) (elimProp z)) - (on+ (on· (elimProp x) (elimProp z)) (on· (elimProp y) (elimProp z))) - i - elimProp (C.+HomConst s t i) = - mkPathP (C.+HomConst s t) - onConst - (on+ onConst onConst) - i - elimProp (C.·HomConst s t i) = - mkPathP (C.·HomConst s t) - onConst - (on· onConst onConst) - i - - elimProp (C.0-trunc x y p q i j) = - isOfHLevel→isOfHLevelDep 2 (λ _ → isProp→isSet isPropP) - (elimProp x) - (elimProp y) - (cong elimProp p) - (cong elimProp q) - (C.0-trunc x y p q) - i - j - - {- - Construction of the induced map. - In this module and the module below, we will show the universal property - of the polynomial ring as a commutative ring. - - R ──→ R[I] - \ ∣ - f ∃! for a given φ : I → S - ↘ ↙ - S - - -} - module _ (S : CommRing ℓ'') (f : CommRingHom R S) (φ : I → ⟨ S ⟩) where - private instance - _ = S .snd - - open IsCommRingHom - - inducedMap : ⟨ R [ I ] ⟩ → ⟨ S ⟩ - inducedMap (C.var x) = φ x - inducedMap (const r) = (f .fst r) - inducedMap (P C.+ Q) = (inducedMap P) + (inducedMap Q) - inducedMap (C.- P) = - inducedMap P - inducedMap (C.+-assoc P Q S i) = +Assoc (inducedMap P) (inducedMap Q) (inducedMap S) i - inducedMap (C.+-rid P i) = - let - eq : (inducedMap P) + (inducedMap (const 0r)) ≡ (inducedMap P) - eq = (inducedMap P) + (inducedMap (const 0r)) ≡⟨⟩ - (inducedMap P) + ((f .fst 0r)) ≡⟨ cong ((inducedMap P) +_) (pres0 (f .snd)) ⟩ - (inducedMap P) + 0r ≡⟨ +IdR _ ⟩ - (inducedMap P) ∎ - in eq i - inducedMap (C.+-rinv P i) = - let eq : (inducedMap P - inducedMap P) ≡ (inducedMap (const 0r)) - eq = (inducedMap P - inducedMap P) ≡⟨ +InvR _ ⟩ - 0r ≡⟨ sym (pres0 (f .snd)) ⟩ - (inducedMap (const 0r))∎ - in eq i - inducedMap (C.+-comm P Q i) = +Comm (inducedMap P) (inducedMap Q) i - inducedMap (P C.· Q) = inducedMap P · inducedMap Q - inducedMap (C.·-assoc P Q S i) = ·Assoc (inducedMap P) (inducedMap Q) (inducedMap S) i - inducedMap (C.·-lid P i) = - let eq = inducedMap (const 1r) · inducedMap P ≡⟨ cong (λ u → u · inducedMap P) (pres1 (f .snd)) ⟩ - 1r · inducedMap P ≡⟨ ·IdL (inducedMap P) ⟩ - inducedMap P ∎ - in eq i - inducedMap (C.·-comm P Q i) = ·Comm (inducedMap P) (inducedMap Q) i - inducedMap (C.ldist P Q S i) = ·DistL+ (inducedMap P) (inducedMap Q) (inducedMap S) i - inducedMap (C.+HomConst s t i) = pres+ (f .snd) s t i - inducedMap (C.·HomConst s t i) = pres· (f .snd) s t i - inducedMap (C.0-trunc P Q p q i j) = - is-set (inducedMap P) (inducedMap Q) (cong _ p) (cong _ q) i j - - module _ where - open IsCommRingHom - - inducedHom : CommRingHom (R [ I ]) S - inducedHom .fst = inducedMap - inducedHom .snd .pres0 = pres0 (f .snd) - inducedHom .snd .pres1 = pres1 (f. snd) - inducedHom .snd .pres+ = λ _ _ → refl - inducedHom .snd .pres· = λ _ _ → refl - inducedHom .snd .pres- = λ _ → refl - - opaque - inducedHomComm : inducedHom ∘cr constPolynomial R I ≡ f - inducedHomComm = CommRingHom≡ $ funExt (λ r → refl) - -module _ {R : CommRing ℓ} {I : Type ℓ'} (S : CommRing ℓ'') (f : CommRingHom R S) where - open CommRingStr ⦃...⦄ - private instance - _ = S .snd - _ = (R [ I ]) .snd - open IsCommRingHom - - evalVar : CommRingHom (R [ I ]) S → I → ⟨ S ⟩ - evalVar h = h .fst ∘ var - - opaque - unfolding var - evalInduce : ∀ (φ : I → ⟨ S ⟩) - → evalVar (inducedHom S f φ) ≡ φ - evalInduce φ = refl - - opaque - unfolding var - evalOnVars : ∀ (φ : I → ⟨ S ⟩) - → (i : I) → inducedHom S f φ .fst (var i) ≡ φ i - evalOnVars φ i = refl - - opaque - unfolding var - induceEval : (g : CommRingHom (R [ I ]) S) - (p : g .fst ∘ constPolynomial R I .fst ≡ f .fst) - → (inducedHom S f (evalVar g)) ≡ g - induceEval g p = - let theMap : ⟨ R [ I ] ⟩ → ⟨ S ⟩ - theMap = inducedMap S f (evalVar g) - in - CommRingHom≡ $ - funExt $ - elimProp - (is-set _ _) - (λ {x} → refl) - (λ {r} → sym (cong (λ f → f r) p)) - (λ {x} {y} eq-x eq-y → - theMap (x + y) ≡⟨ pres+ (inducedHom S f (evalVar g) .snd) x y ⟩ - theMap x + theMap y ≡[ i ]⟨ (eq-x i + eq-y i) ⟩ - (g $cr x + g $cr y) ≡⟨ sym (pres+ (g .snd) _ _) ⟩ - (g $cr (x + y)) ∎) - λ {x} {y} eq-x eq-y → - theMap (x · y) ≡⟨ pres· (inducedHom S f (evalVar g) .snd) x y ⟩ - theMap x · theMap y ≡[ i ]⟨ (eq-x i · eq-y i) ⟩ - (g $cr x · g $cr y) ≡⟨ sym (pres· (g .snd) _ _) ⟩ - (g $cr (x · y)) ∎ - - opaque - inducedHomUnique : (φ : I → ⟨ S ⟩) - (g : CommRingHom (R [ I ]) S) - (p : g .fst ∘ constPolynomial R I .fst ≡ f .fst) - (q : evalVar g ≡ φ) - → inducedHom S f φ ≡ g - inducedHomUnique φ g p q = cong (inducedHom S f) (sym q) ∙ induceEval g p - - opaque - hom≡ByValuesOnVars : (g h : CommRingHom (R [ I ]) S) - (p : g .fst ∘ constPolynomial _ _ .fst ≡ f .fst) (q : h .fst ∘ constPolynomial _ _ .fst ≡ f .fst) - → (evalVar g ≡ evalVar h) - → g ≡ h - hom≡ByValuesOnVars g h p q ≡OnVars = - sym (inducedHomUnique ϕ g p refl) ∙ inducedHomUnique ϕ h q (sym ≡OnVars) - where ϕ : I → ⟨ S ⟩ - ϕ = evalVar g - -opaque - idByIdOnVars : {R : CommRing ℓ} {I : Type ℓ'} - (g : CommRingHom (R [ I ]) (R [ I ])) - (p : g .fst ∘ constPolynomial _ _ .fst ≡ constPolynomial _ _ .fst) - → (g .fst ∘ var ≡ idfun _ ∘ var) - → g ≡ idCommRingHom (R [ I ]) - idByIdOnVars g p idOnVars = hom≡ByValuesOnVars _ (constPolynomial _ _) g (idCommRingHom _) p refl idOnVars + abstract + isHom : IsCommRingHom ((R [ ⊥ ]) .snd) (to .fst) (R .snd) + isHom = to .snd diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/UniversalProperty.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/UniversalProperty.agda new file mode 100644 index 0000000000..50e6aaf814 --- /dev/null +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/UniversalProperty.agda @@ -0,0 +1,285 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.UniversalProperty where +{- + This file contains + * an elimination principle for proving some proposition for all elements of R[I]ᵣ + ('elimProp') + * definitions of the induced maps appearing in the universal property of R[I], +-} + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function hiding (const) +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Sigma.Properties using (Σ≡Prop) +open import Cubical.HITs.SetTruncation + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.Base +open import Cubical.Algebra.Ring.Properties + +open import Cubical.Data.Empty +open import Cubical.Data.Sigma + +open import Cubical.Tactics.CommRingSolver + +private + variable + ℓ ℓ' ℓ'' : Level + +module _ {R : CommRing ℓ} {I : Type ℓ'} where + open CommRingStr ⦃...⦄ + private instance + _ = snd R + _ = snd (R [ I ]) + + module C = Construction + open C using (const) + + {- + Construction of the 'elimProp' eliminator. + -} + module _ + {P : ⟨ R [ I ] ⟩ → Type ℓ''} + (isPropP : {x : _} → isProp (P x)) + (onVar : {x : I} → P (C.var x)) + (onConst : {r : ⟨ R ⟩} → P (const r)) + (on+ : {x y : ⟨ R [ I ] ⟩} → P x → P y → P (x + y)) + (on· : {x y : ⟨ R [ I ] ⟩} → P x → P y → P (x · y)) + where + + private + on- : {x : ⟨ R [ I ] ⟩} → P x → P (- x) + on- {x} Px = subst P (minusByMult x) (on· onConst Px) + where minusByMult : (x : _) → (const (- 1r)) · x ≡ - x + minusByMult x = + (const (- 1r)) · x ≡⟨ cong (_· x) (pres- 1r) ⟩ + (- const (1r)) · x ≡⟨ cong (λ u → (- u) · x) pres1 ⟩ + (- 1r) · x ≡⟨ sym (-IsMult-1 x) ⟩ + - x ∎ + where open RingTheory (CommRing→Ring (R [ I ])) using (-IsMult-1) + open IsCommRingHom (constPolynomial R I .snd) + + -- A helper to deal with the path constructor cases. + mkPathP : + {x0 x1 : ⟨ R [ I ] ⟩} → + (p : x0 ≡ x1) → + (Px0 : P (x0)) → + (Px1 : P (x1)) → + PathP (λ i → P (p i)) Px0 Px1 + mkPathP _ = isProp→PathP λ i → isPropP + + elimProp : ((x : _) → P x) + + elimProp (C.var _) = onVar + elimProp (const _) = onConst + elimProp (x C.+ y) = on+ (elimProp x) (elimProp y) + elimProp (C.- x) = on- (elimProp x) + elimProp (x C.· y) = on· (elimProp x) (elimProp y) + + elimProp (C.+-assoc x y z i) = + mkPathP (C.+-assoc x y z) + (on+ (elimProp x) (on+ (elimProp y) (elimProp z))) + (on+ (on+ (elimProp x) (elimProp y)) (elimProp z)) + i + elimProp (C.+-rid x i) = + mkPathP (C.+-rid x) + (on+ (elimProp x) onConst) + (elimProp x) + i + elimProp (C.+-rinv x i) = + mkPathP (C.+-rinv x) + (on+ (elimProp x) (on- (elimProp x))) + onConst + i + elimProp (C.+-comm x y i) = + mkPathP (C.+-comm x y) + (on+ (elimProp x) (elimProp y)) + (on+ (elimProp y) (elimProp x)) + i + elimProp (C.·-assoc x y z i) = + mkPathP (C.·-assoc x y z) + (on· (elimProp x) (on· (elimProp y) (elimProp z))) + (on· (on· (elimProp x) (elimProp y)) (elimProp z)) + i + elimProp (C.·-lid x i) = + mkPathP (C.·-lid x) + (on· onConst (elimProp x)) + (elimProp x) + i + elimProp (C.·-comm x y i) = + mkPathP (C.·-comm x y) + (on· (elimProp x) (elimProp y)) + (on· (elimProp y) (elimProp x)) + i + elimProp (C.ldist x y z i) = + mkPathP (C.ldist x y z) + (on· (on+ (elimProp x) (elimProp y)) (elimProp z)) + (on+ (on· (elimProp x) (elimProp z)) (on· (elimProp y) (elimProp z))) + i + elimProp (C.+HomConst s t i) = + mkPathP (C.+HomConst s t) + onConst + (on+ onConst onConst) + i + elimProp (C.·HomConst s t i) = + mkPathP (C.·HomConst s t) + onConst + (on· onConst onConst) + i + + elimProp (C.0-trunc x y p q i j) = + isOfHLevel→isOfHLevelDep 2 (λ _ → isProp→isSet isPropP) + (elimProp x) + (elimProp y) + (cong elimProp p) + (cong elimProp q) + (C.0-trunc x y p q) + i + j + + {- + Construction of the induced map. + In this module and the module below, we will show the universal property + of the polynomial ring as a commutative ring. + + R ──→ R[I] + \ ∣ + f ∃! for a given φ : I → S + ↘ ↙ + S + + -} + module _ (S : CommRing ℓ'') (f : CommRingHom R S) (φ : I → ⟨ S ⟩) where + private instance + _ = S .snd + + open IsCommRingHom + + inducedMap : ⟨ R [ I ] ⟩ → ⟨ S ⟩ + inducedMap (C.var x) = φ x + inducedMap (const r) = (f .fst r) + inducedMap (P C.+ Q) = (inducedMap P) + (inducedMap Q) + inducedMap (C.- P) = - inducedMap P + inducedMap (C.+-assoc P Q S i) = +Assoc (inducedMap P) (inducedMap Q) (inducedMap S) i + inducedMap (C.+-rid P i) = + let + eq : (inducedMap P) + (inducedMap (const 0r)) ≡ (inducedMap P) + eq = (inducedMap P) + (inducedMap (const 0r)) ≡⟨⟩ + (inducedMap P) + ((f .fst 0r)) ≡⟨ cong ((inducedMap P) +_) (pres0 (f .snd)) ⟩ + (inducedMap P) + 0r ≡⟨ +IdR _ ⟩ + (inducedMap P) ∎ + in eq i + inducedMap (C.+-rinv P i) = + let eq : (inducedMap P - inducedMap P) ≡ (inducedMap (const 0r)) + eq = (inducedMap P - inducedMap P) ≡⟨ +InvR _ ⟩ + 0r ≡⟨ sym (pres0 (f .snd)) ⟩ + (inducedMap (const 0r))∎ + in eq i + inducedMap (C.+-comm P Q i) = +Comm (inducedMap P) (inducedMap Q) i + inducedMap (P C.· Q) = inducedMap P · inducedMap Q + inducedMap (C.·-assoc P Q S i) = ·Assoc (inducedMap P) (inducedMap Q) (inducedMap S) i + inducedMap (C.·-lid P i) = + let eq = inducedMap (const 1r) · inducedMap P ≡⟨ cong (λ u → u · inducedMap P) (pres1 (f .snd)) ⟩ + 1r · inducedMap P ≡⟨ ·IdL (inducedMap P) ⟩ + inducedMap P ∎ + in eq i + inducedMap (C.·-comm P Q i) = ·Comm (inducedMap P) (inducedMap Q) i + inducedMap (C.ldist P Q S i) = ·DistL+ (inducedMap P) (inducedMap Q) (inducedMap S) i + inducedMap (C.+HomConst s t i) = pres+ (f .snd) s t i + inducedMap (C.·HomConst s t i) = pres· (f .snd) s t i + inducedMap (C.0-trunc P Q p q i j) = + is-set (inducedMap P) (inducedMap Q) (cong _ p) (cong _ q) i j + + module _ where + open IsCommRingHom + + inducedHom : CommRingHom (R [ I ]) S + inducedHom .fst = inducedMap + inducedHom .snd .pres0 = pres0 (f .snd) + inducedHom .snd .pres1 = pres1 (f. snd) + inducedHom .snd .pres+ = λ _ _ → refl + inducedHom .snd .pres· = λ _ _ → refl + inducedHom .snd .pres- = λ _ → refl + + opaque + inducedHomComm : inducedHom ∘cr constPolynomial R I ≡ f + inducedHomComm = CommRingHom≡ $ funExt (λ r → refl) + +module _ {R : CommRing ℓ} {I : Type ℓ'} (S : CommRing ℓ'') (f : CommRingHom R S) where + open CommRingStr ⦃...⦄ + private instance + _ = S .snd + _ = (R [ I ]) .snd + open IsCommRingHom + + evalVar : CommRingHom (R [ I ]) S → I → ⟨ S ⟩ + evalVar h = h .fst ∘ var + + opaque + unfolding var + evalInduce : ∀ (φ : I → ⟨ S ⟩) + → evalVar (inducedHom S f φ) ≡ φ + evalInduce φ = refl + + opaque + unfolding var + evalOnVars : ∀ (φ : I → ⟨ S ⟩) + → (i : I) → inducedHom S f φ .fst (var i) ≡ φ i + evalOnVars φ i = refl + + opaque + unfolding var + induceEval : (g : CommRingHom (R [ I ]) S) + (p : g .fst ∘ constPolynomial R I .fst ≡ f .fst) + → (inducedHom S f (evalVar g)) ≡ g + induceEval g p = + let theMap : ⟨ R [ I ] ⟩ → ⟨ S ⟩ + theMap = inducedMap S f (evalVar g) + in + CommRingHom≡ $ + funExt $ + elimProp + (is-set _ _) + (λ {x} → refl) + (λ {r} → sym (cong (λ f → f r) p)) + (λ {x} {y} eq-x eq-y → + theMap (x + y) ≡⟨ pres+ (inducedHom S f (evalVar g) .snd) x y ⟩ + theMap x + theMap y ≡[ i ]⟨ (eq-x i + eq-y i) ⟩ + (g $cr x + g $cr y) ≡⟨ sym (pres+ (g .snd) _ _) ⟩ + (g $cr (x + y)) ∎) + λ {x} {y} eq-x eq-y → + theMap (x · y) ≡⟨ pres· (inducedHom S f (evalVar g) .snd) x y ⟩ + theMap x · theMap y ≡[ i ]⟨ (eq-x i · eq-y i) ⟩ + (g $cr x · g $cr y) ≡⟨ sym (pres· (g .snd) _ _) ⟩ + (g $cr (x · y)) ∎ + + opaque + inducedHomUnique : (φ : I → ⟨ S ⟩) + (g : CommRingHom (R [ I ]) S) + (p : g .fst ∘ constPolynomial R I .fst ≡ f .fst) + (q : evalVar g ≡ φ) + → inducedHom S f φ ≡ g + inducedHomUnique φ g p q = cong (inducedHom S f) (sym q) ∙ induceEval g p + + opaque + hom≡ByValuesOnVars : (g h : CommRingHom (R [ I ]) S) + (p : g .fst ∘ constPolynomial _ _ .fst ≡ f .fst) (q : h .fst ∘ constPolynomial _ _ .fst ≡ f .fst) + → (evalVar g ≡ evalVar h) + → g ≡ h + hom≡ByValuesOnVars g h p q ≡OnVars = + sym (inducedHomUnique ϕ g p refl) ∙ inducedHomUnique ϕ h q (sym ≡OnVars) + where ϕ : I → ⟨ S ⟩ + ϕ = evalVar g + +opaque + idByIdOnVars : {R : CommRing ℓ} {I : Type ℓ'} + (g : CommRingHom (R [ I ]) (R [ I ])) + (p : g .fst ∘ constPolynomial _ _ .fst ≡ constPolynomial _ _ .fst) + → (g .fst ∘ var ≡ idfun _ ∘ var) + → g ≡ idCommRingHom (R [ I ]) + idByIdOnVars g p idOnVars = hom≡ByValuesOnVars _ (constPolynomial _ _) g (idCommRingHom _) p refl idOnVars From c6f62d40f172ea11eefa3cc12182bfe674b1251c Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 30 Sep 2024 18:02:43 +0200 Subject: [PATCH 81/83] port initial algebra --- .../CommAlgebra/Instances/Initial.agda | 94 +++++++++++++++++++ 1 file changed, 94 insertions(+) create mode 100644 Cubical/Algebra/CommAlgebra/Instances/Initial.agda diff --git a/Cubical/Algebra/CommAlgebra/Instances/Initial.agda b/Cubical/Algebra/CommAlgebra/Instances/Initial.agda new file mode 100644 index 0000000000..899d6d049c --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/Instances/Initial.agda @@ -0,0 +1,94 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.Instances.Initial where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function using (_$_) + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma.Properties using (Σ≡Prop) + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Algebra.Base using (IsAlgebraHom) +open import Cubical.Algebra.Algebra.Properties +open import Cubical.Algebra.CommAlgebra +import Cubical.Algebra.Algebra.Properties + +open AlgebraHoms + +private + variable + ℓ : Level + +module _ (R : CommRing ℓ) where + module _ where + open CommRingStr (snd R) + + initialCAlg : CommAlgebra R ℓ + initialCAlg .fst = R + initialCAlg .snd = idCommRingHom R + + module _ (A : CommAlgebra R ℓ) where +-- open CommAlgebraStr ⦃... ⦄ + + initialMap : CommAlgebraHom initialCAlg A + initialMap = CommRingHom→CommAlgebraHom (A .snd) (CommRingHom≡ refl) + + initialMapEq : (f : CommAlgebraHom initialCAlg A) + → f ≡ initialMap + initialMapEq f = CommAlgebraHom≡ $ + funExt $ + λ r → f $ca r ≡⟨ (cong (λ h → h .fst r) $ f .snd .IsCommAlgebraHom.commutes) ⟩ + initialMap $ca r ∎ + + initialMapProp : (f g : CommAlgebraHom initialCAlg A) + → f ≡ g + initialMapProp f g = initialMapEq f ∙ sym (initialMapEq g) + + + initialityIso : Iso (CommAlgebraHom initialCAlg A) (Unit* {ℓ = ℓ}) + initialityIso = iso (λ _ → tt*) + (λ _ → initialMap) + (λ {tt*x → refl}) + λ f → sym (initialMapEq f) + + initialityPath : CommAlgebraHom initialCAlg A ≡ Unit* + initialityPath = isoToPath initialityIso + + initialityContr : isContr (CommAlgebraHom initialCAlg A) + initialityContr = initialMap , λ ϕ → sym (initialMapEq ϕ) + + + {- + Show that any R-Algebra with the same universal property + as the initial R-Algebra, is isomorphic to the initial + R-Algebra. + -} + module _ (A : CommAlgebra R ℓ) where + equivByInitiality : + (isInitial : (B : CommAlgebra R ℓ) → isContr (CommAlgebraHom A B)) + → CommAlgebraEquiv A (initialCAlg) + equivByInitiality isInitial = isoToEquiv asIso , snd to + where + -- open CommAlgebraHoms + to : CommAlgebraHom A initialCAlg + to = fst (isInitial initialCAlg) + + from : CommAlgebraHom initialCAlg A + from = initialMap A + + asIso : Iso ⟨ A ⟩ₐ ⟨ initialCAlg ⟩ₐ + Iso.fun asIso = fst to + Iso.inv asIso = fst from + Iso.rightInv asIso = + λ x i → cong + fst + (isContr→isProp (initialityContr initialCAlg) (to ∘ca from) (idCAlgHom initialCAlg)) + i x + Iso.leftInv asIso = + λ x i → cong + fst + (isContr→isProp (isInitial A) (from ∘ca to) (idCAlgHom A)) + i x From bc7a25255a4892efe84faa5aac3395edf1468667 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Mon, 30 Sep 2024 18:30:09 +0200 Subject: [PATCH 82/83] fix comment --- .../CommRing/Instances/Polynomials/Typevariate/Base.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda index 93c792d110..3f06b0e2c8 100644 --- a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda @@ -15,7 +15,7 @@ module Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.Base where This file contains * the definition of the free commutative algebra on a type I over a commutative ring R as a HIT (let us call that R[I]) - * a prove that the construction is an commutative R-algebra + * the homomorphism R -> R[I], which turns this CommRing into a CommAlgebra For more, see the Properties file. -} open import Cubical.Foundations.Prelude From 7917b850b3a09fa4dfe9fbe4e8a44fde64365133 Mon Sep 17 00:00:00 2001 From: Felix Cherubini Date: Thu, 10 Oct 2024 19:32:00 +0200 Subject: [PATCH 83/83] WIP --- Cubical/Algebra/CommAlgebra/FP/Base.agda | 31 ++++++++----- Cubical/Algebra/CommAlgebra/FP/Instances.agda | 35 ++++++++++----- .../Algebra/CommAlgebra/FP/Properties.agda | 33 ++++++++++++++ .../CommAlgebra/Instances/Polynomials.agda | 45 +++++++++++++++---- .../Polynomials/Typevariate/Properties.agda | 2 +- 5 files changed, 116 insertions(+), 30 deletions(-) create mode 100644 Cubical/Algebra/CommAlgebra/FP/Properties.agda diff --git a/Cubical/Algebra/CommAlgebra/FP/Base.agda b/Cubical/Algebra/CommAlgebra/FP/Base.agda index 2ea0b39dcb..ec78eea357 100644 --- a/Cubical/Algebra/CommAlgebra/FP/Base.agda +++ b/Cubical/Algebra/CommAlgebra/FP/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.Algebra.CommAlgebra.FP.Base where open import Cubical.Foundations.Prelude @@ -11,6 +11,7 @@ open import Cubical.Foundations.Structure open import Cubical.Data.FinData open import Cubical.Data.Nat open import Cubical.Data.Vec +open import Cubical.Data.Sigma open import Cubical.HITs.PropositionalTruncation @@ -30,20 +31,30 @@ module _ (R : CommRing ℓ) where Polynomials : (n : ℕ) → CommAlgebra R ℓ Polynomials n = R [ Fin n ]ₐ - FPCAlgebra : {m : ℕ} (n : ℕ) (relations : FinVec ⟨ Polynomials n ⟩ₐ m) → CommAlgebra R ℓ - FPCAlgebra n relations = Polynomials n / ⟨ relations ⟩[ Polynomials n ] - - record FinitePresentation (A : CommAlgebra R ℓ') : Type (ℓ-max ℓ ℓ') where + record FinitePresentation : Type ℓ where no-eta-equality field n : ℕ m : ℕ relations : FinVec ⟨ Polynomials n ⟩ₐ m - equiv : CommAlgebraEquiv (FPCAlgebra n relations) A - abstract - isFP : (A : CommAlgebra R ℓ') → Type (ℓ-max ℓ ℓ') - isFP A = ∥ FinitePresentation A ∥₁ + opaque + ideal : IdealsIn R (Polynomials n) + ideal = ⟨ relations ⟩[ Polynomials n ] + + opaque + algebra : CommAlgebra R ℓ + algebra = Polynomials n / ideal + + π : CommAlgebraHom (Polynomials n) algebra + π = quotientHom (Polynomials n) ideal + + generator : (i : Fin n) → ⟨ algebra ⟩ₐ + generator = ⟨ π ⟩ₐ→ ∘ var + + isFP : (A : CommAlgebra R ℓ') → Type (ℓ-max ℓ ℓ') + isFP A = ∃[ p ∈ FinitePresentation ] CommAlgebraEquiv (FinitePresentation.algebra p) A - isFPIsProp : {A : CommAlgebra R ℓ} → isProp (isFP A) + opaque + isFPIsProp : {A : CommAlgebra R ℓ'} → isProp (isFP A) isFPIsProp = isPropPropTrunc diff --git a/Cubical/Algebra/CommAlgebra/FP/Instances.agda b/Cubical/Algebra/CommAlgebra/FP/Instances.agda index 6e3a423104..9535984546 100644 --- a/Cubical/Algebra/CommAlgebra/FP/Instances.agda +++ b/Cubical/Algebra/CommAlgebra/FP/Instances.agda @@ -32,7 +32,7 @@ open import Cubical.Algebra.CommAlgebra.Instances.Polynomials open import Cubical.Algebra.CommAlgebra.QuotientAlgebra open import Cubical.Algebra.CommAlgebra.Ideal using (IdealsIn; 0Ideal) open import Cubical.Algebra.CommAlgebra.FGIdeal --- open import Cubical.Algebra.CommAlgebra.Instances.Initial +open import Cubical.Algebra.CommAlgebra.Instances.Initial open import Cubical.Algebra.CommAlgebra.Instances.Unit open import Cubical.Algebra.CommAlgebra.FP.Base @@ -48,7 +48,7 @@ module _ (R : CommRing ℓ) where {- Every (multivariate) polynomial algebra is finitely presented -} module _ (n : ℕ) where private - abstract + opaque p : 0Ideal R (Polynomials R n) ≡ ⟨ emptyFinVec ⟨ Polynomials R n ⟩ₐ ⟩[ _ ] p = sym $ 0FGIdeal≡0Ideal (Polynomials R n) @@ -59,29 +59,44 @@ module _ (R : CommRing ℓ) where transport (λ i → CommAlgebraEquiv (Polynomials R n) ((Polynomials R n) / (p i))) $ zeroIdealQuotient (Polynomials R n) - polynomialsFP : FinitePresentation R (Polynomials R n) + polynomialsFP : FinitePresentation R polynomialsFP = record { n = n ; m = 0 ; - relations = emptyFinVec ⟨ Polynomials R n ⟩ₐ ; - equiv = invCommAlgebraEquiv compute + relations = emptyFinVec ⟨ Polynomials R n ⟩ₐ } -{- - + opaque + unfolding algebra ideal + isFPPolynomials : isFP R (Polynomials R n) + isFPPolynomials = ∣ polynomialsFP , invCommAlgebraEquiv compute ∣₁ {- The initial R-algebra is finitely presented -} private R[⊥] : CommAlgebra R ℓ - R[⊥] = Polynomials 0 + R[⊥] = Polynomials R 0 - emptyGen : FinVec (fst R[⊥]) 0 + emptyGen : FinVec ⟨ R[⊥] ⟩ₐ 0 emptyGen = λ () + initialAlgFP : FinitePresentation R + initialAlgFP = record { n = 0 ; m = 0 ; relations = emptyGen } + R[⊥]/⟨0⟩ : CommAlgebra R ℓ - R[⊥]/⟨0⟩ = FPAlgebra 0 emptyGen + R[⊥]/⟨0⟩ = algebra initialAlgFP +{- + + R[⊥]/⟨0⟩IsInitial : (B : CommAlgebra R ℓ) + → isContr (CommAlgebraHom R[⊥]/⟨0⟩ B) + R[⊥]/⟨0⟩IsInitial B = {!!} , {!!} + where + iHom : CommAlgebraHom R[⊥]/⟨0⟩ B + iHom = inducedHom 0 emptyGen B (λ ()) (λ ()) + uniqueness : (f : CommAlgebraHom R[⊥]/⟨0⟩ B) → + iHom ≡ f + uniqueness f = unique 0 emptyGen B (λ ()) (λ ()) f (λ ()) R[⊥]/⟨0⟩IsInitial : (B : CommAlgebra R ℓ) → isContr (CommAlgebraHom R[⊥]/⟨0⟩ B) diff --git a/Cubical/Algebra/CommAlgebra/FP/Properties.agda b/Cubical/Algebra/CommAlgebra/FP/Properties.agda new file mode 100644 index 0000000000..8192d6cfdf --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/FP/Properties.agda @@ -0,0 +1,33 @@ +{-# OPTIONS --safe #-} +{- + Universal property of finitely presented algebras +-} +module Cubical.Algebra.CommAlgebra.FP.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure + +open import Cubical.Data.FinData +open import Cubical.Data.Nat +open import Cubical.Data.Vec + +open import Cubical.HITs.PropositionalTruncation + +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommAlgebra +open import Cubical.Algebra.CommAlgebra.Instances.Polynomials +open import Cubical.Algebra.CommAlgebra.QuotientAlgebra +open import Cubical.Algebra.CommAlgebra.Ideal +open import Cubical.Algebra.CommAlgebra.FGIdeal +open import Cubical.Algebra.CommAlgebra.FP.Base + + +private + variable + ℓ ℓ' : Level + +module _ (R : CommRing ℓ) (fp : FinitePresentation R) where diff --git a/Cubical/Algebra/CommAlgebra/Instances/Polynomials.agda b/Cubical/Algebra/CommAlgebra/Instances/Polynomials.agda index 02e268daab..907ae3f3b4 100644 --- a/Cubical/Algebra/CommAlgebra/Instances/Polynomials.agda +++ b/Cubical/Algebra/CommAlgebra/Instances/Polynomials.agda @@ -2,18 +2,50 @@ module Cubical.Algebra.CommAlgebra.Instances.Polynomials where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function using (_$_; _∘_) +open import Cubical.Foundations.Structure using (withOpaqueStr) +open import Cubical.Foundations.Isomorphism using (isoFunInjective) + +open import Cubical.Data.Nat +open import Cubical.Data.FinData open import Cubical.Algebra.CommRing.Base open import Cubical.Algebra.CommAlgebra.Base -open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate +open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate as Poly hiding (var) +import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.UniversalProperty + as Poly private variable ℓ ℓ' ℓ'' : Level -opaque - _[_]ₐ : (R : CommRing ℓ) (I : Type ℓ') → CommAlgebra R (ℓ-max ℓ ℓ') - R [ I ]ₐ = (R [ I ]) , constPolynomial R I +_[_]ₐ : (R : CommRing ℓ) (I : Type ℓ') → CommAlgebra R (ℓ-max ℓ ℓ') +R [ I ]ₐ = (R [ I ]) , constPolynomial R I + +module _ {R : CommRing ℓ} where + evPolyIn : {n : ℕ} (A : CommAlgebra R ℓ') + → ⟨ R [ Fin n ]ₐ ⟩ₐ → FinVec ⟨ A ⟩ₐ n → ⟨ A ⟩ₐ + evPolyIn {n = n} A P v = Poly.inducedHom (CommAlgebra→CommRing A) (A .snd) v $cr P + +module _ {R : CommRing ℓ} {I : Type ℓ'} where + var : I → ⟨ R [ I ]ₐ ⟩ₐ + var = Poly.var + + inducedHom : (A : CommAlgebra R ℓ'') (φ : I → ⟨ A ⟩ₐ ) + → CommAlgebraHom (R [ I ]ₐ) A + inducedHom A ϕ = + CommAlgebraHomFromCommRingHom + f + (λ _ _ → refl) + where f : CommRingHom _ _ + f = Poly.inducedHom (CommAlgebra→CommRing A) (A .snd) ϕ + + inducedHomUnique : + (A : CommAlgebra R ℓ'') (φ : I → ⟨ A ⟩ₐ ) + → (f : CommAlgebraHom (R [ I ]ₐ) A) → ((i : I) → ⟨ f ⟩ₐ→ ∘ var ≡ φ) + → f ≡ inducedHom A φ + inducedHomUnique A φ f p = {!isoFunInjective (homMapIso A) f (inducedHom A φ) λ j i → p i j!} + -- {- evaluateAt : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') @@ -21,11 +53,6 @@ evaluateAt : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') → (I → fst A) evaluateAt A f x = f $a (Construction.var x) -inducedHom : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'') - (φ : I → fst A ) - → CommAlgebraHom (R [ I ]) A -inducedHom A φ = Theory.inducedHom A φ - homMapIso : {R : CommRing ℓ} {I : Type ℓ''} (A : CommAlgebra R ℓ') → Iso (CommAlgebraHom (R [ I ]) A) (I → (fst A)) diff --git a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda index bde4d50eae..99a33504ec 100644 --- a/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda @@ -14,7 +14,7 @@ open import Cubical.HITs.SetTruncation open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.Base -open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.UniversalProperty public +open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.UniversalProperty open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Sigma