diff --git a/Cubical/Algebra/Algebra/Base.agda b/Cubical/Algebra/Algebra/Base.agda index 97afcced0f..8cae5bdab6 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,7 +31,7 @@ record IsAlgebra (R : Ring ℓ) {A : Type ℓ'} (0a 1a : A) (_+_ _·_ : A → A → A) (-_ : A → A) (_⋆_ : ⟨ R ⟩ → A → A) : Type (ℓ-max ℓ ℓ') where - constructor isalgebra + no-eta-equality open RingStr (snd R) using (1r) renaming (_+_ to _+r_; _·_ to _·r_) @@ -59,8 +54,6 @@ unquoteDecl IsAlgebraIsoΣ = declareRecordIsoΣ IsAlgebraIsoΣ (quote IsAlgebra) record AlgebraStr (R : Ring ℓ) (A : Type ℓ') : Type (ℓ-max ℓ ℓ') where - constructor algebrastr - field 0a : A 1a : A @@ -72,6 +65,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 @@ -208,13 +203,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 +234,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/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..3394ba74f9 --- /dev/null +++ b/Cubical/Algebra/Algebra/Univalence.agda @@ -0,0 +1,86 @@ +{-# 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.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 + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + 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 ∣ (λ 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 + + -- 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 + 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) ∎) 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/CommAlgebra/AsModule/Base.agda b/Cubical/Algebra/CommAlgebra/AsModule/Base.agda new file mode 100644 index 0000000000..ae6602f03b --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/AsModule/Base.agda @@ -0,0 +1,328 @@ +{-# OPTIONS --safe #-} +{- + This used to be the default definition of CommAlgebra. +-} +module Cubical.Algebra.CommAlgebra.AsModule.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.SIP + +open import Cubical.Data.Sigma + +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Algebra + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open import Cubical.Reflection.RecordEquiv + +private + variable + ℓ ℓ' ℓ'' : Level + +record IsCommAlgebra (R : CommRing ℓ) {A : Type ℓ'} + (0a : A) (1a : A) + (_+_ : A → A → A) (_·_ : A → A → A) (-_ : A → A) + (_⋆_ : ⟨ R ⟩ → A → A) : Type (ℓ-max ℓ ℓ') where + no-eta-equality + + field + isAlgebra : IsAlgebra (CommRing→Ring R) 0a 1a _+_ _·_ -_ _⋆_ + ·Comm : (x y : A) → x · y ≡ y · x + + open IsAlgebra isAlgebra public + +unquoteDecl IsCommAlgebraIsoΣ = declareRecordIsoΣ IsCommAlgebraIsoΣ (quote IsCommAlgebra) + +record CommAlgebraStr (R : CommRing ℓ) (A : Type ℓ') : Type (ℓ-max ℓ ℓ') where + no-eta-equality + + field + 0a : A + 1a : A + _+_ : A → A → A + _·_ : A → A → A + -_ : A → A + _⋆_ : ⟨ R ⟩ → A → A + isCommAlgebra : IsCommAlgebra R 0a 1a _+_ _·_ -_ _⋆_ + + open IsCommAlgebra isCommAlgebra public + + infix 8 -_ + infixl 7 _·_ + infixl 7 _⋆_ + infixl 6 _+_ + +unquoteDecl CommAlgebraStrIsoΣ = declareRecordIsoΣ CommAlgebraStrIsoΣ (quote CommAlgebraStr) + +CommAlgebra : (R : CommRing ℓ) → ∀ ℓ' → Type (ℓ-max ℓ (ℓ-suc ℓ')) +CommAlgebra R ℓ' = Σ[ A ∈ Type ℓ' ] CommAlgebraStr R A + +module _ {R : CommRing ℓ} where + open CommRingStr (snd R) using (1r) renaming (_+_ to _+r_; _·_ to _·s_) + + CommAlgebraStr→AlgebraStr : {A : Type ℓ'} → CommAlgebraStr R A → AlgebraStr (CommRing→Ring R) A + CommAlgebraStr→AlgebraStr {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) ℓ' + fst (CommAlgebra→Algebra A) = fst A + snd (CommAlgebra→Algebra A) = CommAlgebraStr→AlgebraStr (snd A) + + CommAlgebra→CommRing : (A : CommAlgebra R ℓ') → CommRing ℓ' + 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} + {_+_ _·_ : A → A → A} { -_ : A → A} {_⋆_ : ⟨ R ⟩ → A → A} + (isSet-A : isSet A) + (+Assoc : (x y z : A) → x + (y + z) ≡ (x + y) + z) + (+IdR : (x : A) → x + 0a ≡ x) + (+InvR : (x : A) → x + (- x) ≡ 0a) + (+Comm : (x y : A) → x + y ≡ y + x) + (·Assoc : (x y z : A) → x · (y · z) ≡ (x · y) · z) + (·IdL : (x : A) → 1a · x ≡ x) + (·DistL+ : (x y z : A) → (x + y) · z ≡ (x · z) + (y · z)) + (·Comm : (x y : A) → x · y ≡ y · x) + (⋆Assoc : (r s : ⟨ R ⟩) (x : A) → (r ·s s) ⋆ x ≡ r ⋆ (s ⋆ x)) + (⋆DistR+ : (r : ⟨ R ⟩) (x y : A) → r ⋆ (x + y) ≡ (r ⋆ x) + (r ⋆ y)) + (⋆DistL+ : (r s : ⟨ R ⟩) (x : A) → (r +r s) ⋆ x ≡ (r ⋆ x) + (s ⋆ x)) + (⋆IdL : (x : A) → 1r ⋆ x ≡ x) + (⋆AssocL : (r : ⟨ R ⟩) (x y : A) → (r ⋆ x) · y ≡ r ⋆ (x · y)) + where + + + makeIsCommAlgebra : IsCommAlgebra R 0a 1a _+_ _·_ -_ _⋆_ + makeIsCommAlgebra .IsCommAlgebra.isAlgebra = makeIsAlgebra + isSet-A + +Assoc +IdR +InvR +Comm + ·Assoc ·IdR ·IdL ·DistR+ ·DistL+ + ⋆Assoc + ⋆DistR+ + ⋆DistL+ + ⋆IdL + ⋆AssocR + ⋆AssocL + where + ·IdR : _ + ·IdR x = x · 1a ≡⟨ ·Comm _ _ ⟩ 1a · x ≡⟨ ·IdL _ ⟩ x ∎ + ·DistR+ : _ + ·DistR+ x y z = x · (y + z) ≡⟨ ·Comm _ _ ⟩ + (y + z) · x ≡⟨ ·DistL+ _ _ _ ⟩ + (y · x) + (z · x) ≡⟨ cong (λ u → (y · x) + u) (·Comm _ _) ⟩ + (y · x) + (x · z) ≡⟨ cong (λ u → u + (x · z)) (·Comm _ _) ⟩ + (x · y) + (x · z) ∎ + ⋆AssocR : _ + ⋆AssocR r x y = r ⋆ (x · y) ≡⟨ cong (λ u → r ⋆ u) (·Comm _ _) ⟩ + r ⋆ (y · x) ≡⟨ sym (⋆AssocL _ _ _) ⟩ + (r ⋆ y) · x ≡⟨ ·Comm _ _ ⟩ + 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) + + module _ + (_⋆_ : fst R → fst S → fst S) + (·Assoc⋆ : (r s : fst R) (x : fst S) → (r ·R s) ⋆ x ≡ r ⋆ (s ⋆ x)) + (⋆DistR+ : (r : fst R) (x y : fst S) → r ⋆ (x + y) ≡ (r ⋆ x) + (r ⋆ y)) + (⋆DistL+ : (r s : fst R) (x : fst S) → (r +R s) ⋆ x ≡ (r ⋆ x) + (s ⋆ x)) + (⋆IdL : (x : fst S) → 1R ⋆ x ≡ x) + (⋆AssocL : (r : fst R) (x y : fst S) → (r ⋆ x) · y ≡ r ⋆ (x · y)) + where + + commAlgebraFromCommRing : CommAlgebra R ℓ' + commAlgebraFromCommRing .fst = fst S + commAlgebraFromCommRing .snd .CommAlgebraStr.0a = 0r + commAlgebraFromCommRing .snd .CommAlgebraStr.1a = 1S + commAlgebraFromCommRing .snd .CommAlgebraStr._+_ = _+_ + commAlgebraFromCommRing .snd .CommAlgebraStr._·_ = _·_ + commAlgebraFromCommRing .snd .CommAlgebraStr.-_ = -_ + commAlgebraFromCommRing .snd .CommAlgebraStr._⋆_ = _⋆_ + commAlgebraFromCommRing .snd .CommAlgebraStr.isCommAlgebra = + makeIsCommAlgebra is-set +Assoc +IdR +InvR +Comm ·Assoc ·IdL ·DistL+ ·Comm + ·Assoc⋆ ⋆DistR+ ⋆DistL+ ⋆IdL ⋆AssocL + + commAlgebraFromCommRing→CommRing : CommAlgebra→CommRing commAlgebraFromCommRing ≡ S + -- Note that this is not definitional: the proofs of the axioms might differ. + commAlgebraFromCommRing→CommRing i .fst = ⟨ S ⟩ + commAlgebraFromCommRing→CommRing i .snd .CommRingStr.0r = 0r + commAlgebraFromCommRing→CommRing i .snd .CommRingStr.1r = 1S + commAlgebraFromCommRing→CommRing i .snd .CommRingStr._+_ = _+_ + commAlgebraFromCommRing→CommRing i .snd .CommRingStr._·_ = _·_ + commAlgebraFromCommRing→CommRing i .snd .CommRingStr.-_ = -_ + commAlgebraFromCommRing→CommRing i .snd .CommRingStr.isCommRing = + isProp→PathP (λ i → isPropIsCommRing _ _ _ _ _) + (CommRingStr.isCommRing (snd (CommAlgebra→CommRing commAlgebraFromCommRing))) + isCommRing + i + + IsCommAlgebraEquiv : {A : Type ℓ'} {B : Type ℓ''} + (M : CommAlgebraStr R A) (e : A ≃ B) (N : CommAlgebraStr R B) + → Type _ + IsCommAlgebraEquiv M e N = + IsAlgebraHom (CommAlgebraStr→AlgebraStr M) (e .fst) (CommAlgebraStr→AlgebraStr N) + + CommAlgebraEquiv : (M : CommAlgebra R ℓ') (N : CommAlgebra R ℓ'') → Type _ + CommAlgebraEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsCommAlgebraEquiv (M .snd) e (N .snd) + + IsCommAlgebraHom : {A : Type ℓ'} {B : Type ℓ''} + (M : CommAlgebraStr R A) (f : A → B) (N : CommAlgebraStr R B) + → Type _ + IsCommAlgebraHom M f N = + IsAlgebraHom (CommAlgebraStr→AlgebraStr M) f (CommAlgebraStr→AlgebraStr N) + + CommAlgebraHom : (M : CommAlgebra R ℓ') (N : CommAlgebra R ℓ'') → Type _ + CommAlgebraHom M N = Σ[ f ∈ (⟨ M ⟩ → ⟨ N ⟩) ] IsCommAlgebraHom (M .snd) f (N .snd) + + CommAlgebraEquiv→CommAlgebraHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → CommAlgebraEquiv A B → CommAlgebraHom A B + CommAlgebraEquiv→CommAlgebraHom (e , eIsHom) = e .fst , eIsHom + + CommAlgebraHom→CommRingHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') + → CommAlgebraHom A B + → CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) + fst (CommAlgebraHom→CommRingHom A B f) = fst 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 {{...}} + open IsAlgebraHom + private + instance + _ = snd M + _ = snd N + + makeCommAlgebraHom : (f : fst M → fst N) + → (fPres1 : f 1a ≡ 1a) + → (fPres+ : (x y : fst M) → f (x + y) ≡ f x + f y) + → (fPres· : (x y : fst M) → f (x · y) ≡ f x · f y) + → (fPres⋆ : (r : fst R) (x : fst M) → f (r ⋆ x) ≡ r ⋆ f x) + → CommAlgebraHom M N + makeCommAlgebraHom f fPres1 fPres+ fPres· fPres⋆ = f , isHom + where fPres0 = + f 0a ≡⟨ sym (+IdR _) ⟩ + f 0a + 0a ≡⟨ cong (λ u → f 0a + u) (sym (+InvR (f 0a))) ⟩ + f 0a + (f 0a - f 0a) ≡⟨ +Assoc (f 0a) (f 0a) (- f 0a) ⟩ + (f 0a + f 0a) - f 0a ≡⟨ cong (λ u → u - f 0a) (sym (fPres+ 0a 0a)) ⟩ + f (0a + 0a) - f 0a ≡⟨ cong (λ u → f u - f 0a) (+IdL 0a) ⟩ + f 0a - f 0a ≡⟨ +InvR (f 0a) ⟩ + 0a ∎ + + isHom : IsCommAlgebraHom (snd M) f (snd N) + pres0 isHom = fPres0 + pres1 isHom = fPres1 + pres+ isHom = fPres+ + pres· isHom = fPres· + pres- isHom = (λ x → + f (- x) ≡⟨ sym (+IdR _) ⟩ + (f (- x) + 0a) ≡⟨ cong (λ u → f (- x) + u) (sym (+InvR (f x))) ⟩ + (f (- x) + (f x - f x)) ≡⟨ +Assoc _ _ _ ⟩ + ((f (- x) + f x) - f x) ≡⟨ cong (λ u → u - f x) (sym (fPres+ _ _)) ⟩ + (f ((- x) + x) - f x) ≡⟨ cong (λ u → f u - f x) (+InvL x) ⟩ + (f 0a - f x) ≡⟨ cong (λ u → u - f x) fPres0 ⟩ + (0a - f x) ≡⟨ +IdL _ ⟩ (- f x) ∎) + pres⋆ isHom = fPres⋆ + + isPropIsCommAlgebraHom : (f : fst M → fst N) → isProp (IsCommAlgebraHom (snd M) f (snd N)) + isPropIsCommAlgebraHom f = isPropIsAlgebraHom + (CommRing→Ring R) + (snd (CommAlgebra→Algebra M)) + f + (snd (CommAlgebra→Algebra N)) + +isPropIsCommAlgebra : (R : CommRing ℓ) {A : Type ℓ'} + (0a 1a : A) + (_+_ _·_ : A → A → A) + (-_ : A → A) + (_⋆_ : ⟨ R ⟩ → A → A) + → isProp (IsCommAlgebra R 0a 1a _+_ _·_ -_ _⋆_) +isPropIsCommAlgebra R _ _ _ _ _ _ = + isOfHLevelRetractFromIso 1 IsCommAlgebraIsoΣ + (isPropΣ (isPropIsAlgebra _ _ _ _ _ _ _) + (λ alg → isPropΠ2 λ _ _ → alg .IsAlgebra.is-set _ _)) + +𝒮ᴰ-CommAlgebra : (R : CommRing ℓ) → DUARel (𝒮-Univ ℓ') (CommAlgebraStr R) (ℓ-max ℓ ℓ') +𝒮ᴰ-CommAlgebra {ℓ' = ℓ'} R = + 𝒮ᴰ-Record (𝒮-Univ _) (IsCommAlgebraEquiv {R = R}) + (fields: + data[ 0a ∣ nul ∣ pres0 ] + data[ 1a ∣ nul ∣ pres1 ] + data[ _+_ ∣ bin ∣ pres+ ] + data[ _·_ ∣ bin ∣ pres· ] + data[ -_ ∣ autoDUARel _ _ ∣ pres- ] + data[ _⋆_ ∣ autoDUARel _ _ ∣ pres⋆ ] + prop[ isCommAlgebra ∣ (λ 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 + + -- faster with some sharing + nul = autoDUARel (𝒮-Univ _) (λ A → A) + bin = autoDUARel (𝒮-Univ _) (λ A → A → A → A) + +CommAlgebraPath : (R : CommRing ℓ) → (A B : CommAlgebra R ℓ') → (CommAlgebraEquiv A B) ≃ (A ≡ B) +CommAlgebraPath R = ∫ (𝒮ᴰ-CommAlgebra R) .UARel.ua + +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 _ _) diff --git a/Cubical/Algebra/CommAlgebra/AsModule/FGIdeal.agda b/Cubical/Algebra/CommAlgebra/AsModule/FGIdeal.agda new file mode 100644 index 0000000000..8bfbf30385 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/AsModule/FGIdeal.agda @@ -0,0 +1,37 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.AsModule.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.CommAlgebra.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule.Ideal + +private + variable + ℓ : Level + R : CommRing ℓ + +generatedIdeal : {n : ℕ} (A : CommAlgebra R ℓ) → FinVec (fst A) n → IdealsIn 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 CommAlgebraStr (snd A) + + 0FGIdeal : {n : ℕ} → ⟨ replicateFinVec n 0a ⟩[ A ] ≡ (0Ideal A) + 0FGIdeal = 0FGIdealCommRing (CommAlgebra→CommRing A) diff --git a/Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra.agda b/Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra.agda new file mode 100644 index 0000000000..36605f2a27 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra.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.AsModule.FPAlgebra where + +open import Cubical.Algebra.CommAlgebra.AsModule.FPAlgebra.Base public diff --git a/Cubical/Algebra/CommAlgebra/FPAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra/Base.agda similarity index 95% rename from Cubical/Algebra/CommAlgebra/FPAlgebra/Base.agda rename to Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra/Base.agda index 229c8cb20a..5b2c16d99f 100644 --- a/Cubical/Algebra/CommAlgebra/FPAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra/Base.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.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.CommAlgebra -open import Cubical.Algebra.CommAlgebra.FreeCommAlgebra +open import Cubical.Algebra.CommAlgebra.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule.FreeCommAlgebra renaming (inducedHom to freeInducedHom) -open import Cubical.Algebra.CommAlgebra.QuotientAlgebra +open import Cubical.Algebra.CommAlgebra.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.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 @@ -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/AsModule/FPAlgebra/Instances.agda similarity index 91% rename from Cubical/Algebra/CommAlgebra/FPAlgebra/Instances.agda rename to Cubical/Algebra/CommAlgebra/AsModule/FPAlgebra/Instances.agda index 26e21cad66..e206f60c9c 100644 --- a/Cubical/Algebra/CommAlgebra/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.CommAlgebra.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.CommAlgebra -open import Cubical.Algebra.CommAlgebra.FreeCommAlgebra +open import Cubical.Algebra.CommAlgebra.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule.FreeCommAlgebra renaming (inducedHom to freeInducedHom) -open import Cubical.Algebra.CommAlgebra.QuotientAlgebra +open import Cubical.Algebra.CommAlgebra.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.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.CommAlgebra.Kernel +open import Cubical.Algebra.CommAlgebra.AsModule.Kernel open import Cubical.Algebra.Algebra -open import Cubical.Algebra.CommAlgebra.FPAlgebra.Base +open import Cubical.Algebra.CommAlgebra.AsModule.FPAlgebra.Base private variable @@ -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 _ 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/CommAlgebra/FreeCommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/Base.agda similarity index 94% rename from Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Base.agda rename to Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/Base.agda index b66cd0433f..8ce217225c 100644 --- a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/Base.agda @@ -1,6 +1,6 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.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.CommAlgebra.FreeCommAlgebra.Base where open import Cubical.Foundations.Prelude open import Cubical.Algebra.CommRing -open import Cubical.Algebra.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebra.AsModule.Base private variable @@ -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/FreeCommAlgebra/OnCoproduct.agda b/Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/OnCoproduct.agda similarity index 95% rename from Cubical/Algebra/CommAlgebra/FreeCommAlgebra/OnCoproduct.agda rename to Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/OnCoproduct.agda index a20195fca7..03eb8d10a4 100644 --- a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/OnCoproduct.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/OnCoproduct.agda @@ -7,7 +7,7 @@ where '⊎' is the disjoint sum. -} -module Cubical.Algebra.CommAlgebra.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.CommAlgebra -open import Cubical.Algebra.CommAlgebra.FreeCommAlgebra +open import Cubical.Algebra.CommAlgebra.AsModule +open import Cubical.Algebra.CommAlgebra.AsModule.FreeCommAlgebra private variable ℓ ℓ' : Level @@ -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/AsModule/FreeCommAlgebra/Properties.agda similarity index 98% rename from Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Properties.agda rename to Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/Properties.agda index a7dea81c7c..791a4deddd 100644 --- a/Cubical/Algebra/CommAlgebra/FreeCommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/FreeCommAlgebra/Properties.agda @@ -1,6 +1,6 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.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.CommAlgebra.FreeCommAlgebra.Base -open import Cubical.Algebra.CommAlgebra -open import Cubical.Algebra.CommAlgebra.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/CommAlgebra/AsModule/Ideal.agda b/Cubical/Algebra/CommAlgebra/AsModule/Ideal.agda new file mode 100644 index 0000000000..623d038c26 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/AsModule/Ideal.agda @@ -0,0 +1,36 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.AsModule.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.CommAlgebra.AsModule +open import Cubical.Algebra.Ring + +open import Cubical.Data.Unit + +private + variable + ℓ : Level + +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) where + IdealsIn : Type _ + IdealsIn = IdealsInCommRing (CommAlgebra→CommRing A) + + open CommAlgebraStr (snd A) + + makeIdeal : (I : fst A → hProp ℓ) + → (+-closed : {x y : fst A} → x ∈ I → y ∈ I → (x + y) ∈ I) + → (0-closed : 0a ∈ I) + → (·-closedLeft : {x : fst A} → (r : fst A) → x ∈ I → r · x ∈ I) + → IdealsIn + makeIdeal = makeIdealCommRing {R = CommAlgebra→CommRing A} + + 0Ideal : IdealsIn + 0Ideal = CommIdeal.0Ideal (CommAlgebra→CommRing A) + + 1Ideal : IdealsIn + 1Ideal = CommIdeal.1Ideal (CommAlgebra→CommRing A) diff --git a/Cubical/Algebra/CommAlgebra/AsModule/Instances/Initial.agda b/Cubical/Algebra/CommAlgebra/AsModule/Instances/Initial.agda new file mode 100644 index 0000000000..34678b461c --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/AsModule/Instances/Initial.agda @@ -0,0 +1,130 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.AsModule.Instances.Initial where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism + +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.AsModule +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 = fst R + initialCAlg .snd .CommAlgebraStr.0a = _ + initialCAlg .snd .CommAlgebraStr.1a = _ + initialCAlg .snd .CommAlgebraStr._+_ = _ + initialCAlg .snd .CommAlgebraStr._·_ = _ + initialCAlg .snd .CommAlgebraStr.-_ = _ + initialCAlg .snd .CommAlgebraStr._⋆_ r x = r · x + initialCAlg .snd .CommAlgebraStr.isCommAlgebra = + makeIsCommAlgebra is-set + +Assoc +IdR +InvR +Comm + ·Assoc ·IdL + ·DistL+ ·Comm + (λ x y z → sym (·Assoc x y z)) ·DistR+ ·DistL+ ·IdL + λ x y z → sym (·Assoc x y z) + + module _ (A : CommAlgebra R ℓ) where + open CommAlgebraStr ⦃... ⦄ + private + instance + _ : CommAlgebraStr R (fst A) + _ = snd A + _ : CommAlgebraStr R (fst R) + _ = snd initialCAlg + + _*_ : fst R → (fst A) → (fst A) + r * a = CommAlgebraStr._⋆_ (snd A) r a + + initialMap : CommAlgebraHom initialCAlg A + initialMap = + makeCommAlgebraHom {M = initialCAlg} {N = A} + (λ r → r * 1a) + (⋆IdL _) + (λ x y → ⋆DistL+ x y 1a) + (λ x y → (x · y) * 1a ≡⟨ ⋆Assoc _ _ _ ⟩ + x * (y * 1a) ≡[ i ]⟨ x * (·IdL (y * 1a) (~ i)) ⟩ + x * (1a · (y * 1a)) ≡⟨ sym (⋆AssocL _ _ _) ⟩ + (x * 1a) · (y * 1a) ∎) + (λ r x → (r · x) * 1a ≡⟨ ⋆Assoc _ _ _ ⟩ + (r * (x * 1a)) ∎) + + initialMapEq : (f : CommAlgebraHom initialCAlg A) + → f ≡ initialMap + initialMapEq f = + let open IsAlgebraHom (snd f) + in Σ≡Prop + (isPropIsCommAlgebraHom {M = initialCAlg} {N = A}) + λ i x → + ((fst f) x ≡⟨ cong (fst f) (sym (·IdR _)) ⟩ + fst f (x · 1a) ≡⟨ pres⋆ x 1a ⟩ + CommAlgebraStr._⋆_ (snd A) x (fst f 1a) ≡⟨ cong + (λ u → (snd A CommAlgebraStr.⋆ x) u) + pres1 ⟩ + (CommAlgebraStr._⋆_ (snd A) x 1a) ∎) i + + 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 (fst A) (fst initialCAlg) + Iso.fun asIso = fst to + Iso.inv asIso = fst from + Iso.rightInv asIso = + λ x i → cong + fst + (isContr→isProp (initialityContr initialCAlg) (to ∘a from) (idCAlgHom initialCAlg)) + i x + Iso.leftInv asIso = + λ x i → cong + fst + (isContr→isProp (isInitial A) (from ∘a to) (idCAlgHom A)) + i x diff --git a/Cubical/Algebra/CommAlgebra/Instances/Pointwise.agda b/Cubical/Algebra/CommAlgebra/AsModule/Instances/Pointwise.agda similarity index 91% rename from Cubical/Algebra/CommAlgebra/Instances/Pointwise.agda rename to Cubical/Algebra/CommAlgebra/AsModule/Instances/Pointwise.agda index 9ce845e935..edbe6beef0 100644 --- a/Cubical/Algebra/CommAlgebra/Instances/Pointwise.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/Instances/Pointwise.agda @@ -1,12 +1,12 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.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.CommAlgebra.Base +open import Cubical.Algebra.CommAlgebra.AsModule.Base open import Cubical.Algebra.Algebra using (IsAlgebraHom) private diff --git a/Cubical/Algebra/CommAlgebra/AsModule/Instances/Unit.agda b/Cubical/Algebra/CommAlgebra/AsModule/Instances/Unit.agda new file mode 100644 index 0000000000..e73e3c8d1f --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/AsModule/Instances/Unit.agda @@ -0,0 +1,60 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.AsModule.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.CommAlgebra.AsModule.Base +open import Cubical.Algebra.CommRing.Instances.Unit +open import Cubical.Algebra.Algebra.Base using (IsAlgebraHom) +open import Cubical.Tactics.CommRingSolver + +private + variable + ℓ ℓ' : Level + +module _ (R : CommRing ℓ) where + UnitCommAlgebra : CommAlgebra R ℓ' + UnitCommAlgebra = + commAlgebraFromCommRing + UnitCommRing + (λ _ _ → tt*) (λ _ _ _ → refl) (λ _ _ _ → refl) + (λ _ _ _ → refl) (λ _ → refl) (λ _ _ _ → refl) + + module _ (A : CommAlgebra R ℓ) where + terminalMap : CommAlgebraHom A (UnitCommAlgebra {ℓ' = ℓ}) + terminalMap = (λ _ → tt*) , isHom + where open IsAlgebraHom + isHom : IsCommAlgebraHom (snd A) _ (snd UnitCommAlgebra) + pres0 isHom = isPropUnit* _ _ + pres1 isHom = isPropUnit* _ _ + pres+ isHom = λ _ _ → isPropUnit* _ _ + pres· isHom = λ _ _ → isPropUnit* _ _ + pres- isHom = λ _ → isPropUnit* _ _ + pres⋆ isHom = λ _ _ → isPropUnit* _ _ + + 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/CommAlgebra/AsModule/Kernel.agda b/Cubical/Algebra/CommAlgebra/AsModule/Kernel.agda new file mode 100644 index 0000000000..6c98c8a8f0 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/AsModule/Kernel.agda @@ -0,0 +1,21 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.AsModule.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.CommAlgebra.AsModule.Base +open import Cubical.Algebra.CommAlgebra.AsModule.Properties +open import Cubical.Algebra.CommAlgebra.AsModule.Ideal + +private + variable + ℓ : Level + +module _ {R : CommRing ℓ} (A B : CommAlgebra R ℓ) (ϕ : CommAlgebraHom A B) where + + kernel : IdealsIn A + kernel = Ideal→CommIdeal (ringKernel (CommAlgebraHom→RingHom {A = A} {B = B} ϕ)) diff --git a/Cubical/Algebra/CommAlgebra/AsModule/Localisation.agda b/Cubical/Algebra/CommAlgebra/AsModule/Localisation.agda new file mode 100644 index 0000000000..b17ed15f7a --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/AsModule/Localisation.agda @@ -0,0 +1,318 @@ +{-# OPTIONS --safe --lossy-unification #-} +module Cubical.Algebra.CommAlgebra.AsModule.Localisation where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.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.Algebra +open import Cubical.Algebra.CommAlgebra.AsModule.Base +open import Cubical.Algebra.CommAlgebra.AsModule.Properties + +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' : ℙ (fst R')) (SMultClosedSubset : isMultClosedSubset R' S') where + open isMultClosedSubset + private R = fst R' + open CommAlgebraStr + open IsAlgebraHom + open CommRingStr (snd R') renaming (_+_ to _+r_ ; _·_ to _·r_ ; ·IdR to ·rRid) + open RingTheory (CommRing→Ring R') + open CommRingTheory R' + open Loc R' S' SMultClosedSubset + open S⁻¹RUniversalProp R' S' SMultClosedSubset + open CommAlgChar R' + + + S⁻¹RAsCommAlg : CommAlgebra R' ℓ + S⁻¹RAsCommAlg = toCommAlg (S⁻¹RAsCommRing , /1AsCommRingHom) + + LocCommAlg→CommRingPath : CommAlgebra→CommRing S⁻¹RAsCommAlg ≡ S⁻¹RAsCommRing + LocCommAlg→CommRingPath = CommAlgebra→CommRing≡ (S⁻¹RAsCommRing , /1AsCommRingHom) + + hasLocAlgUniversalProp : (A : CommAlgebra R' ℓ) + → (∀ s → s ∈ S' → _⋆_ (snd A) s (1a (snd A)) ∈ (CommAlgebra→CommRing A) ˣ) + → Type (ℓ-suc ℓ) + hasLocAlgUniversalProp A _ = (B : CommAlgebra R' ℓ) + → (∀ s → s ∈ S' → _⋆_ (snd B) s (1a (snd B)) ∈ (CommAlgebra→CommRing B) ˣ) + → isContr (CommAlgebraHom A B) + + S⋆1⊆S⁻¹Rˣ : ∀ s → s ∈ S' → _⋆_ (snd S⁻¹RAsCommAlg) s (1a (snd S⁻¹RAsCommAlg)) ∈ S⁻¹Rˣ + S⋆1⊆S⁻¹Rˣ s s∈S' = subst-∈ S⁻¹Rˣ + (cong [_] (≡-× (sym (·rRid s)) (Σ≡Prop (λ x → S' x .snd) (sym (·rRid _))))) + (S/1⊆S⁻¹Rˣ s s∈S') + + + S⁻¹RHasAlgUniversalProp : hasLocAlgUniversalProp S⁻¹RAsCommAlg S⋆1⊆S⁻¹Rˣ + S⁻¹RHasAlgUniversalProp B' S⋆1⊆Bˣ = χᴬ , χᴬuniqueness + where + B = fromCommAlg B' .fst + φ = fromCommAlg B' .snd + open CommRingStr (snd B) renaming (_·_ to _·b_ ; 1r to 1b ; ·IdL to ·bLid) + + χ : CommRingHom S⁻¹RAsCommRing B + χ = S⁻¹RHasUniversalProp B φ S⋆1⊆Bˣ .fst .fst + + χcomp : ∀ r → fst χ (r /1) ≡ fst φ r + χcomp = funExt⁻ (S⁻¹RHasUniversalProp B φ S⋆1⊆Bˣ .fst .snd) + + χᴬ : CommAlgebraHom S⁻¹RAsCommAlg B' + fst χᴬ = fst χ + pres0 (snd χᴬ) = 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) ≡⟨ 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') _ _ _ ⟩ + _⋆_ (snd B') r (1b ·b fst χ x) ≡⟨ cong (_⋆_ (snd B') r) (·bLid _) ⟩ + _⋆_ (snd B') r (fst χ x) ∎ + + + χᴬuniqueness : (ψ : CommAlgebraHom S⁻¹RAsCommAlg B') → χᴬ ≡ ψ + χᴬuniqueness ψ = Σ≡Prop (λ _ → isPropIsAlgebraHom _ _ _ _) + (cong (fst ∘ fst) (χuniqueness (ψ' , funExt ψ'r/1≡φr))) + where + χuniqueness = S⁻¹RHasUniversalProp B φ S⋆1⊆Bˣ .snd + + ψ' : CommRingHom S⁻¹RAsCommRing B + fst ψ' = fst ψ + 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 = + fst ψ (r /1) ≡⟨ cong (fst ψ) (sym (·ₗ-rid _)) ⟩ + fst ψ (_⋆_ (snd S⁻¹RAsCommAlg) r (1a (snd S⁻¹RAsCommAlg))) ≡⟨ pres⋆ (snd ψ) _ _ ⟩ + _⋆_ (snd B') r (fst ψ (1a (snd S⁻¹RAsCommAlg))) ≡⟨ cong (_⋆_ (snd B') r) (pres1 (snd ψ)) ⟩ + _⋆_ (snd B') r 1b ∎ + + + -- an immediate 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 (toCommAlg (A' , φ)) + S⁻¹RAlgCharEquiv A' φ cond = toCommAlgebraEquiv (S⁻¹RAsCommRing , /1AsCommRingHom) (A' , φ) + (S⁻¹RCharEquiv R' S' SMultClosedSubset A' φ cond) + (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 = LocCommAlg→CommRingPath + +module AlgLocTwoSubsets (R' : CommRing ℓ) + (S₁ : ℙ (fst R')) (S₁MultClosedSubset : isMultClosedSubset R' S₁) + (S₂ : ℙ (fst R')) (S₂MultClosedSubset : isMultClosedSubset R' S₂) where + open isMultClosedSubset + open CommRingStr (snd R') hiding (is-set) + open RingTheory (CommRing→Ring R') + open Loc R' S₁ S₁MultClosedSubset renaming (S⁻¹R to S₁⁻¹R ; + S⁻¹RAsCommRing to S₁⁻¹RAsCommRing) + open Loc R' S₂ S₂MultClosedSubset renaming (S⁻¹R to S₂⁻¹R ; + S⁻¹RAsCommRing to S₂⁻¹RAsCommRing) + open AlgLoc R' S₁ S₁MultClosedSubset renaming ( S⁻¹RAsCommAlg to S₁⁻¹RAsCommAlg + ; S⋆1⊆S⁻¹Rˣ to S₁⋆1⊆S₁⁻¹Rˣ + ; S⁻¹RHasAlgUniversalProp to S₁⁻¹RHasAlgUniversalProp + ; isContrHomS⁻¹RS⁻¹R to isContrHomS₁⁻¹RS₁⁻¹R) + open AlgLoc R' S₂ S₂MultClosedSubset renaming ( S⁻¹RAsCommAlg to S₂⁻¹RAsCommAlg + ; S⋆1⊆S⁻¹Rˣ to S₂⋆1⊆S₂⁻¹Rˣ + ; S⁻¹RHasAlgUniversalProp to S₂⁻¹RHasAlgUniversalProp + ; isContrHomS⁻¹RS⁻¹R to isContrHomS₂⁻¹RS₂⁻¹R) + + open IsAlgebraHom + open AlgebraHoms + open CommAlgebraHoms + open CommAlgebraStr ⦃...⦄ + + private + R = fst R' + S₁⁻¹Rˣ = S₁⁻¹RAsCommRing ˣ + S₂⁻¹Rˣ = S₂⁻¹RAsCommRing ˣ + instance + _ = snd S₁⁻¹RAsCommAlg + _ = snd S₂⁻¹RAsCommAlg + + + isContrS₁⁻¹R≡S₂⁻¹R : (∀ s₁ → s₁ ∈ S₁ → s₁ ⋆ 1a ∈ S₂⁻¹Rˣ) + → (∀ s₂ → s₂ ∈ S₂ → s₂ ⋆ 1a ∈ S₁⁻¹Rˣ) + → isContr (S₁⁻¹RAsCommAlg ≡ S₂⁻¹RAsCommAlg) + isContrS₁⁻¹R≡S₂⁻¹R S₁⊆S₂⁻¹Rˣ S₂⊆S₁⁻¹Rˣ = isOfHLevelRetractFromIso 0 + (equivToIso (invEquiv (CommAlgebraPath _ _ _))) + isContrS₁⁻¹R≅S₂⁻¹R + where + χ₁ : CommAlgebraHom S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg + χ₁ = S₁⁻¹RHasAlgUniversalProp S₂⁻¹RAsCommAlg S₁⊆S₂⁻¹Rˣ .fst + + χ₂ : CommAlgebraHom S₂⁻¹RAsCommAlg S₁⁻¹RAsCommAlg + χ₂ = S₂⁻¹RHasAlgUniversalProp S₁⁻¹RAsCommAlg S₂⊆S₁⁻¹Rˣ .fst + + χ₁∘χ₂≡id : χ₁ ∘a χ₂ ≡ idCommAlgebraHom _ + χ₁∘χ₂≡id = isContr→isProp isContrHomS₂⁻¹RS₂⁻¹R _ _ + + χ₂∘χ₁≡id : χ₂ ∘a χ₁ ≡ idCommAlgebraHom _ + χ₂∘χ₁≡id = isContr→isProp isContrHomS₁⁻¹RS₁⁻¹R _ _ + + IsoS₁⁻¹RS₂⁻¹R : Iso S₁⁻¹R S₂⁻¹R + Iso.fun IsoS₁⁻¹RS₂⁻¹R = fst χ₁ + Iso.inv IsoS₁⁻¹RS₂⁻¹R = fst χ₂ + Iso.rightInv IsoS₁⁻¹RS₂⁻¹R = funExt⁻ (cong fst χ₁∘χ₂≡id) + Iso.leftInv IsoS₁⁻¹RS₂⁻¹R = funExt⁻ (cong fst χ₂∘χ₁≡id) + + isContrS₁⁻¹R≅S₂⁻¹R : isContr (CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg) + isContrS₁⁻¹R≅S₂⁻¹R = center , uniqueness + where + center : CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg + fst center = isoToEquiv IsoS₁⁻¹RS₂⁻¹R + pres0 (snd center) = pres0 (snd χ₁) + pres1 (snd center) = pres1 (snd χ₁) + pres+ (snd center) = pres+ (snd χ₁) + pres· (snd center) = pres· (snd χ₁) + pres- (snd center) = pres- (snd χ₁) + pres⋆ (snd center) = pres⋆ (snd χ₁) + + uniqueness : (φ : CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg) → center ≡ φ + uniqueness φ = Σ≡Prop (λ _ → isPropIsAlgebraHom _ _ _ _) + (equivEq (cong fst + (S₁⁻¹RHasAlgUniversalProp S₂⁻¹RAsCommAlg S₁⊆S₂⁻¹Rˣ .snd + (AlgebraEquiv→AlgebraHom φ)))) + + + isPropS₁⁻¹R≡S₂⁻¹R : isProp (S₁⁻¹RAsCommAlg ≡ S₂⁻¹RAsCommAlg) + isPropS₁⁻¹R≡S₂⁻¹R S₁⁻¹R≡S₂⁻¹R = + isContr→isProp (isContrS₁⁻¹R≡S₂⁻¹R S₁⊆S₂⁻¹Rˣ S₂⊆S₁⁻¹Rˣ) S₁⁻¹R≡S₂⁻¹R + where + S₁⊆S₂⁻¹Rˣ : ∀ s₁ → s₁ ∈ S₁ → s₁ ⋆ 1a ∈ S₂⁻¹Rˣ + S₁⊆S₂⁻¹Rˣ s₁ s₁∈S₁ = + transport (λ i → _⋆_ ⦃ S₁⁻¹R≡S₂⁻¹R i .snd ⦄ s₁ (1a ⦃ S₁⁻¹R≡S₂⁻¹R i .snd ⦄) + ∈ (CommAlgebra→CommRing (S₁⁻¹R≡S₂⁻¹R i)) ˣ) (S₁⋆1⊆S₁⁻¹Rˣ s₁ s₁∈S₁) + + S₂⊆S₁⁻¹Rˣ : ∀ s₂ → s₂ ∈ S₂ → s₂ ⋆ 1a ∈ S₁⁻¹Rˣ + S₂⊆S₁⁻¹Rˣ s₂ s₂∈S₂ = + transport (λ i → _⋆_ ⦃ (sym S₁⁻¹R≡S₂⁻¹R) i .snd ⦄ s₂ (1a ⦃ (sym S₁⁻¹R≡S₂⁻¹R) i .snd ⦄) + ∈ (CommAlgebra→CommRing ((sym S₁⁻¹R≡S₂⁻¹R) i)) ˣ) (S₂⋆1⊆S₂⁻¹Rˣ s₂ s₂∈S₂) + + + +-- 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 CommAlgChar R + 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/CommAlgebra/LocalisationAlgebra.agda b/Cubical/Algebra/CommAlgebra/AsModule/LocalisationAlgebra.agda similarity index 97% rename from Cubical/Algebra/CommAlgebra/LocalisationAlgebra.agda rename to Cubical/Algebra/CommAlgebra/AsModule/LocalisationAlgebra.agda index 703104ce86..d42248a94d 100644 --- a/Cubical/Algebra/CommAlgebra/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.CommAlgebra -open import Cubical.Algebra.CommAlgebra.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.CommAlgebra.LocalisationAlgebra +module Cubical.Algebra.CommAlgebra.AsModule.LocalisationAlgebra {ℓR : Level} (R : CommRing ℓR) where @@ -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/AsModule/Properties.agda b/Cubical/Algebra/CommAlgebra/AsModule/Properties.agda new file mode 100644 index 0000000000..040c18d643 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/AsModule/Properties.agda @@ -0,0 +1,361 @@ +{-# OPTIONS --safe --lossy-unification #-} +module Cubical.Algebra.CommAlgebra.AsModule.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.SIP +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path + +open import Cubical.Functions.Embedding + +open import Cubical.Data.Sigma + +open import Cubical.Reflection.StrictEquiv + +open import Cubical.Structures.Axioms +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid +open import Cubical.Algebra.CommRing +open import Cubical.Algebra.Ring +open import Cubical.Algebra.Algebra +open import Cubical.Algebra.CommAlgebra.AsModule.Base + +open import Cubical.Algebra.CommRing using (CommRing→Ring) + +open import Cubical.HITs.PropositionalTruncation + +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + R : CommRing ℓ + +open AlgebraHoms + +idCAlgHom : (A : CommAlgebra R ℓ) → _ +idCAlgHom A = idAlgebraHom (CommAlgebra→Algebra A) + +idCAlgEquiv : (A : CommAlgebra R ℓ) → CommAlgebraEquiv A A +fst (idCAlgEquiv A) = idEquiv (fst A) +snd (idCAlgEquiv A) = snd (idCAlgHom A) + +infix 3 _≃CAlg∎ +infixr 2 _≃CAlg⟨_⟩_ + +_≃CAlg∎ : (A : CommAlgebra R ℓ) → CommAlgebraEquiv A A +A ≃CAlg∎ = idCAlgEquiv A + +_≃CAlg⟨_⟩_ : {B C : CommAlgebra R ℓ} + (A : CommAlgebra R ℓ) (f : CommAlgebraEquiv A B) (g : CommAlgebraEquiv B C) + → CommAlgebraEquiv A C +A ≃CAlg⟨ f ⟩ g = g ∘≃a f + +-- An R-algebra is the same as a CommRing A with a CommRingHom φ : R → A +module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where + open Iso + open CommRingTheory + + + CommRingWithHom : Type (ℓ-max ℓ (ℓ-suc ℓ')) + CommRingWithHom = Σ[ A ∈ CommRing ℓ' ] CommRingHom R A + + toCommAlg : CommRingWithHom → CommAlgebra R ℓ' + toCommAlg (A , φ , φIsHom) = + commAlgebraFromCommRing + A + (λ r a → (φ r) · a) + (λ r s x → cong (_· x) (pres· r s) ∙ sym (·Assoc _ _ _)) + (λ r x y → ·DistR+ _ _ _) + (λ r s x → cong (_· x) (pres+ r s) ∙ ·DistL+ _ _ _) + (λ x → cong (_· x) pres1 ∙ ·IdL x) + λ r x y → sym (·Assoc _ _ _) + where + open CommRingStr (snd A) + open IsCommRingHom φIsHom + + fromCommAlg : CommAlgebra R ℓ' → CommRingWithHom + fromCommAlg A = (CommAlgebra→CommRing A) , φ , φIsHom + where + open CommRingStr (snd R) renaming (_·_ to _·r_) hiding (·IdL) + open CommAlgebraStr (snd A) + open AlgebraTheory (CommRing→Ring R) (CommAlgebra→Algebra A) + φ : ⟨ R ⟩ → ⟨ A ⟩ + φ r = r ⋆ 1a + φ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 + module _ (Aφ : CommRingWithHom) where + open CommRingStr + private + A = fst Aφ + CommAlgebra→CommRing≡ : CommAlgebra→CommRing (toCommAlg Aφ) ≡ A + fst (CommAlgebra→CommRing≡ i) = fst A + 0r (snd (CommAlgebra→CommRing≡ i)) = 0r (snd A) + 1r (snd (CommAlgebra→CommRing≡ i)) = 1r (snd A) + _+_ (snd (CommAlgebra→CommRing≡ i)) = _+_ (snd A) + _·_ (snd (CommAlgebra→CommRing≡ i)) = _·_ (snd A) + -_ (snd (CommAlgebra→CommRing≡ i)) = -_ (snd A) + -- 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 + open CommRingStr + -- 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 = CommRingHomPathP _ _ _ _ _ _ λ i x → ·IdR (snd A) (fst φ x) i + + + CommAlgRoundTrip : (A : CommAlgebra R ℓ') → toCommAlg (fromCommAlg A) ≡ A + CommAlgRoundTrip A = ΣPathP (refl , AlgStrPathP) + where + open CommAlgebraStr ⦃...⦄ + instance + _ = snd A + AlgStrPathP : PathP (λ i → CommAlgebraStr R ⟨ A ⟩) (snd (toCommAlg (fromCommAlg A))) (snd A) + CommAlgebraStr.0a (AlgStrPathP i) = 0a + CommAlgebraStr.1a (AlgStrPathP i) = 1a + CommAlgebraStr._+_ (AlgStrPathP i) = _+_ + CommAlgebraStr._·_ (AlgStrPathP i) = _·_ + CommAlgebraStr.-_ (AlgStrPathP i) = -_ + CommAlgebraStr._⋆_ (AlgStrPathP i) r x = (⋆AssocL r 1a x ∙ cong (r ⋆_) (·IdL x)) i + CommAlgebraStr.isCommAlgebra (AlgStrPathP i) = isProp→PathP + (λ i → isPropIsCommAlgebra _ _ _ _ _ _ (CommAlgebraStr._⋆_ (AlgStrPathP i))) + (CommAlgebraStr.isCommAlgebra (snd (toCommAlg (fromCommAlg A)))) isCommAlgebra i + + CommAlgIso : Iso (CommAlgebra R ℓ') CommRingWithHom + fun CommAlgIso = fromCommAlg + inv CommAlgIso = toCommAlg + rightInv CommAlgIso = CommRingWithHomRoundTrip + leftInv CommAlgIso = CommAlgRoundTrip + + open IsCommRingHom + + isCommRingWithHomHom : (A B : CommRingWithHom) → CommRingHom (fst A) (fst B) → Type (ℓ-max ℓ ℓ') + isCommRingWithHomHom (_ , f) (_ , g) h = h ∘cr f ≡ g + + CommRingWithHomHom : CommRingWithHom → CommRingWithHom → Type (ℓ-max ℓ ℓ') + 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 + → CommAlgebraHom (toCommAlg A) (toCommAlg B) + toCommAlgebraHom (A , f) (B , g) h commDiag = + makeCommAlgebraHom (fst h) (pres1 (snd h)) (pres+ (snd h)) (pres· (snd h)) pres⋆h + where + open CommRingStr ⦃...⦄ + instance + _ = snd A + _ = snd B + pres⋆h : ∀ r x → fst h (fst f r · x) ≡ fst g r · fst h x + pres⋆h r x = fst h (fst f r · x) ≡⟨ pres· (snd h) _ _ ⟩ + fst h (fst f r) · fst h x ≡⟨ cong (λ φ → fst φ r · fst h x) commDiag ⟩ + fst g r · fst h x ∎ + + fromCommAlgebraHom : (A B : CommAlgebra R ℓ') → CommAlgebraHom A B + → CommRingWithHomHom (fromCommAlg A) (fromCommAlg B) + fst (fst (fromCommAlgebraHom A B f)) = fst f + pres0 (snd (fst (fromCommAlgebraHom A B f))) = IsAlgebraHom.pres0 (snd f) + pres1 (snd (fst (fromCommAlgebraHom A B f))) = IsAlgebraHom.pres1 (snd f) + pres+ (snd (fst (fromCommAlgebraHom A B f))) = IsAlgebraHom.pres+ (snd f) + 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) = + 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 (CommRingEquiv→CommRingHom e) + + CommRingWithHomEquiv : CommRingWithHom → CommRingWithHom → Type (ℓ-max ℓ ℓ') + CommRingWithHomEquiv A B = Σ[ e ∈ CommRingEquiv (fst A) (fst B) ] isCommRingWithHomEquiv A B e + + toCommAlgebraEquiv : (A B : CommRingWithHom) (e : CommRingEquiv (fst A) (fst B)) + → isCommRingWithHomEquiv A B e + → CommAlgebraEquiv (toCommAlg A) (toCommAlg B) + fst (toCommAlgebraEquiv A B e eCommDiag) = e .fst + snd (toCommAlgebraEquiv A B e eCommDiag) = toCommAlgebraHom A B _ eCommDiag .snd + + + +module CommAlgebraHoms {R : CommRing ℓ} where + open AlgebraHoms + + idCommAlgebraHom : (A : CommAlgebra R ℓ') → CommAlgebraHom A A + idCommAlgebraHom A = idAlgebraHom (CommAlgebra→Algebra A) + + compCommAlgebraHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') (C : CommAlgebra R ℓ''') + → CommAlgebraHom A B → CommAlgebraHom B C → CommAlgebraHom A C + compCommAlgebraHom A B C = compAlgebraHom {A = CommAlgebra→Algebra A} + {B = CommAlgebra→Algebra B} + {C = CommAlgebra→Algebra C} + + _∘ca_ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} + → CommAlgebraHom B C → CommAlgebraHom A B → CommAlgebraHom A C + g ∘ca f = compCommAlgebraHom _ _ _ f g + + compIdCommAlgebraHom : {A B : CommAlgebra R ℓ'} (f : CommAlgebraHom A B) + → compCommAlgebraHom _ _ _ (idCommAlgebraHom A) f ≡ f + compIdCommAlgebraHom = compIdAlgebraHom + + idCompCommAlgebraHom : {A B : CommAlgebra R ℓ'} (f : CommAlgebraHom A B) + → compCommAlgebraHom _ _ _ f (idCommAlgebraHom B) ≡ f + idCompCommAlgebraHom = idCompAlgebraHom + + compAssocCommAlgebraHom : {A B C D : CommAlgebra R ℓ'} + (f : CommAlgebraHom A B) (g : CommAlgebraHom B C) (h : CommAlgebraHom C D) + → compCommAlgebraHom _ _ _ (compCommAlgebraHom _ _ _ f g) h + ≡ compCommAlgebraHom _ _ _ f (compCommAlgebraHom _ _ _ g h) + 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} + + + 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 + + caracCommAlgebra≡ : {A B : CommAlgebra R ℓ'} (p q : A ≡ B) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ 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 + uaCompCommAlgebraEquiv f g = caracCommAlgebra≡ _ _ ( + cong ⟨_⟩ (uaCommAlgebra (compCommAlgebraEquiv f g)) + ≡⟨ uaCompEquiv _ _ ⟩ + cong ⟨_⟩ (uaCommAlgebra f) ∙ cong ⟨_⟩ (uaCommAlgebra g) + ≡⟨ sym (cong-∙ ⟨_⟩ (uaCommAlgebra f) (uaCommAlgebra g)) ⟩ + cong ⟨_⟩ (uaCommAlgebra f ∙ uaCommAlgebra g) ∎) + +open CommAlgebraHoms +open CommAlgebraEquivs +open CommAlgebraUAFunctoriality +recPT→CommAlgebra : {R : CommRing ℓ} {A : Type ℓ'} (𝓕 : A → CommAlgebra R ℓ'') + → (σ : ∀ x y → CommAlgebraEquiv (𝓕 x) (𝓕 y)) + → (∀ x y z → σ x z ≡ compCommAlgebraEquiv (σ x y) (σ y z)) + ------------------------------------------------------ + → ∥ A ∥₁ → CommAlgebra R ℓ'' +recPT→CommAlgebra 𝓕 σ compCoh = GpdElim.rec→Gpd isGroupoidCommAlgebra 𝓕 + (3-ConstantCompChar 𝓕 (λ x y → uaCommAlgebra (σ x y)) + λ x y z → sym ( cong uaCommAlgebra (compCoh x y z) + ∙ uaCompCommAlgebraEquiv (σ x y) (σ y z))) + +open CommAlgebraHoms + +contrCommAlgebraHom→contrCommAlgebraEquiv : {R : CommRing ℓ} {A : Type ℓ'} + (σ : A → CommAlgebra R ℓ'') + → (∀ x y → isContr (CommAlgebraHom (σ x) (σ y))) + ---------------------------------------------------------------------------- + → ∀ x y → isContr (CommAlgebraEquiv (σ x) (σ y)) +contrCommAlgebraHom→contrCommAlgebraEquiv σ contrHom x y = σEquiv , + λ e → Σ≡Prop (λ _ → isPropIsAlgebraHom _ _ _ _) + (Σ≡Prop isPropIsEquiv (cong fst (contrHom _ _ .snd (CommAlgebraEquiv→CommAlgebraHom e)))) + where + open Iso + χ₁ = contrHom x y .fst + χ₂ = contrHom y x .fst + χ₁∘χ₂≡id : χ₁ ∘ca χ₂ ≡ idCommAlgebraHom _ + χ₁∘χ₂≡id = isContr→isProp (contrHom _ _) _ _ + + χ₂∘χ₁≡id : χ₂ ∘ca χ₁ ≡ idCommAlgebraHom _ + χ₂∘χ₁≡id = isContr→isProp (contrHom _ _) _ _ + + σIso : Iso ⟨ σ x ⟩ ⟨ σ y ⟩ + fun σIso = fst χ₁ + inv σIso = fst χ₂ + rightInv σIso = funExt⁻ (cong fst χ₁∘χ₂≡id) + leftInv σIso = funExt⁻ (cong fst χ₂∘χ₁≡id) + + σEquiv : CommAlgebraEquiv (σ x) (σ y) + fst σEquiv = isoToEquiv σIso + snd σEquiv = snd χ₁ + +CommAlgebra→Ring : {R : CommRing ℓ} → CommAlgebra R ℓ' → Ring ℓ' +CommAlgebra→Ring = CommRing→Ring ∘ CommAlgebra→CommRing + +module _ {R : CommRing ℓ} {A B : CommAlgebra R ℓ'} where + open CommAlgebraStr ⦃...⦄ + instance + _ = snd A + _ = snd B + open IsAlgebraHom + + CommAlgebraHom→RingHom : CommAlgebraHom A B → RingHom (CommAlgebra→Ring A) (CommAlgebra→Ring B) + fst (CommAlgebraHom→RingHom ϕ) = fst ϕ + IsRingHom.pres0 (snd (CommAlgebraHom→RingHom ϕ)) = pres0 (snd ϕ) + IsRingHom.pres1 (snd (CommAlgebraHom→RingHom ϕ)) = pres1 (snd ϕ) + IsRingHom.pres+ (snd (CommAlgebraHom→RingHom ϕ)) = pres+ (snd ϕ) + IsRingHom.pres· (snd (CommAlgebraHom→RingHom ϕ)) = pres· (snd ϕ) + IsRingHom.pres- (snd (CommAlgebraHom→RingHom ϕ)) = pres- (snd ϕ) + + CommAlgebraHomFromRingHom : + (ϕ : RingHom (CommAlgebra→Ring A) (CommAlgebra→Ring B)) + → ((r : fst R) (x : fst A) → (fst ϕ) (r ⋆ x) ≡ r ⋆ (fst ϕ x)) + → CommAlgebraHom A B + fst (CommAlgebraHomFromRingHom ϕ pres*) = fst ϕ + pres0 (snd (CommAlgebraHomFromRingHom ϕ pres*)) = IsRingHom.pres0 (snd ϕ) + pres1 (snd (CommAlgebraHomFromRingHom ϕ pres*)) = IsRingHom.pres1 (snd ϕ) + pres+ (snd (CommAlgebraHomFromRingHom ϕ pres*)) = IsRingHom.pres+ (snd ϕ) + pres· (snd (CommAlgebraHomFromRingHom ϕ pres*)) = IsRingHom.pres· (snd ϕ) + 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)) + 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 ∘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 ∘cr (snd homA) ≡ snd homB + asSliceHom = CommAlgChar.fromCommAlgebraHom R A B ϕ + + pbSliceHom : Σ[ k ∈ CommRingHom (CommAlgebra→CommRing A) (CommAlgebra→CommRing B) ] + 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/AsModule/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebra/AsModule/QuotientAlgebra.agda new file mode 100644 index 0000000000..6bd6d62ae4 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/AsModule/QuotientAlgebra.agda @@ -0,0 +1,243 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.AsModule.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.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) +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 A) where + open CommRingStr {{...}} hiding (_-_; -_; ·IdL ; ·DistR+) renaming (_·_ to _·R_; _+_ to _+R_) + open CommAlgebraStr {{...}} + open RingTheory (CommRing→Ring (CommAlgebra→CommRing A)) using (-DistR·) + instance + _ : CommRingStr ⟨ R ⟩ + _ = snd R + _ : CommAlgebraStr R ⟨ A ⟩ + _ = snd A + + opaque + _/_ : CommAlgebra R ℓ + _/_ = 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 : fst A) → x - y ∈ (fst I) → [ r ⋆ x ]/ ≡ [ r ⋆ y ]/ + eq r x y x-y∈I = eq/ _ _ + (subst (λ u → u ∈ fst I) + ((r ⋆ 1a) · (x - y) ≡⟨ ·DistR+ (r ⋆ 1a) x (- y) ⟩ + (r ⋆ 1a) · x + (r ⋆ 1a) · (- y) ≡[ i ]⟨ (r ⋆ 1a) · x + -DistR· (r ⋆ 1a) y i ⟩ + (r ⋆ 1a) · x - (r ⋆ 1a) · y ≡[ i ]⟨ ⋆AssocL r 1a x i + - ⋆AssocL r 1a y i ⟩ + r ⋆ (1a · x) - r ⋆ (1a · 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) diff --git a/Cubical/Algebra/CommAlgebra/Subalgebra.agda b/Cubical/Algebra/CommAlgebra/AsModule/Subalgebra.agda similarity index 59% rename from Cubical/Algebra/CommAlgebra/Subalgebra.agda rename to Cubical/Algebra/CommAlgebra/AsModule/Subalgebra.agda index 107a54f296..8deb268d56 100644 --- a/Cubical/Algebra/CommAlgebra/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.CommAlgebra +open import Cubical.Algebra.CommAlgebra.AsModule open import Cubical.Algebra.CommRing -module Cubical.Algebra.CommAlgebra.Subalgebra +module Cubical.Algebra.CommAlgebra.AsModule.Subalgebra {ℓ ℓ' : Level} (R : CommRing ℓ) (A : CommAlgebra R ℓ') where @@ -20,17 +20,19 @@ 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 + 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) @@ -41,4 +43,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)) - diff --git a/Cubical/Algebra/CommAlgebra/UnivariatePolyList.agda b/Cubical/Algebra/CommAlgebra/AsModule/UnivariatePolyList.agda similarity index 98% rename from Cubical/Algebra/CommAlgebra/UnivariatePolyList.agda rename to Cubical/Algebra/CommAlgebra/AsModule/UnivariatePolyList.agda index 982b966c1a..839e7ade9e 100644 --- a/Cubical/Algebra/CommAlgebra/UnivariatePolyList.agda +++ b/Cubical/Algebra/CommAlgebra/AsModule/UnivariatePolyList.agda @@ -1,5 +1,5 @@ {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.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.CommAlgebra +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/CommAlgebra/Base.agda b/Cubical/Algebra/CommAlgebra/Base.agda index 4022db273f..10668f649f 100644 --- a/Cubical/Algebra/CommAlgebra/Base.agda +++ b/Cubical/Algebra/CommAlgebra/Base.agda @@ -3,280 +3,256 @@ module Cubical.Algebra.CommAlgebra.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 (⟨_⟩; withOpaqueStr) open import Cubical.Foundations.HLevels -open import Cubical.Foundations.SIP +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Path open import Cubical.Data.Sigma -open import Cubical.Algebra.Semigroup -open import Cubical.Algebra.Monoid open import Cubical.Algebra.CommRing +open import Cubical.Algebra.CommRing.Univalence open import Cubical.Algebra.Ring -open import Cubical.Algebra.Algebra - -open import Cubical.Displayed.Base -open import Cubical.Displayed.Auto -open import Cubical.Displayed.Record -open import Cubical.Displayed.Universe open import Cubical.Reflection.RecordEquiv private variable - ℓ ℓ' ℓ'' : Level - -record IsCommAlgebra (R : CommRing ℓ) {A : Type ℓ'} - (0a : A) (1a : A) - (_+_ : A → A → A) (_·_ : A → A → A) (-_ : A → A) - (_⋆_ : ⟨ R ⟩ → A → A) : Type (ℓ-max ℓ ℓ') where - - constructor iscommalgebra - - field - isAlgebra : IsAlgebra (CommRing→Ring R) 0a 1a _+_ _·_ -_ _⋆_ - ·Comm : (x y : A) → x · y ≡ y · x - - open IsAlgebra isAlgebra public - -unquoteDecl IsCommAlgebraIsoΣ = declareRecordIsoΣ IsCommAlgebraIsoΣ (quote IsCommAlgebra) - -record CommAlgebraStr (R : CommRing ℓ) (A : Type ℓ') : Type (ℓ-max ℓ ℓ') where - - constructor commalgebrastr - - field - 0a : A - 1a : A - _+_ : A → A → A - _·_ : A → A → A - -_ : A → A - _⋆_ : ⟨ R ⟩ → A → A - isCommAlgebra : IsCommAlgebra R 0a 1a _+_ _·_ -_ _⋆_ - - open IsCommAlgebra isCommAlgebra public + ℓ ℓ' ℓ'' ℓ''' : Level - infix 8 -_ - infixl 7 _·_ - infixl 7 _⋆_ - infixl 6 _+_ - -CommAlgebra : (R : CommRing ℓ) → ∀ ℓ' → Type (ℓ-max ℓ (ℓ-suc ℓ')) -CommAlgebra R ℓ' = Σ[ A ∈ Type ℓ' ] CommAlgebraStr R A +CommAlgebra : (R : CommRing ℓ) (ℓ' : Level) → Type _ +CommAlgebra R ℓ' = Σ[ A ∈ CommRing ℓ' ] CommRingHom R A module _ {R : CommRing ℓ} where - open CommRingStr (snd R) using (1r) renaming (_+_ to _+r_; _·_ to _·s_) - - CommAlgebraStr→AlgebraStr : {A : Type ℓ'} → CommAlgebraStr R A → AlgebraStr (CommRing→Ring R) A - CommAlgebraStr→AlgebraStr (commalgebrastr _ _ _ _ _ _ (iscommalgebra isAlgebra ·-comm)) = - algebrastr _ _ _ _ _ _ isAlgebra - - CommAlgebra→Algebra : (A : CommAlgebra R ℓ') → Algebra (CommRing→Ring R) ℓ' - CommAlgebra→Algebra (_ , str) = (_ , CommAlgebraStr→AlgebraStr str) - CommAlgebra→CommRing : (A : CommAlgebra R ℓ') → CommRing ℓ' - CommAlgebra→CommRing (_ , commalgebrastr _ _ _ _ _ _ (iscommalgebra isAlgebra ·-comm)) = - _ , commringstr _ _ _ _ _ (iscommring (IsAlgebra.isRing isAlgebra) ·-comm) - - module _ - {A : Type ℓ'} {0a 1a : A} - {_+_ _·_ : A → A → A} { -_ : A → A} {_⋆_ : ⟨ R ⟩ → A → A} - (isSet-A : isSet A) - (+Assoc : (x y z : A) → x + (y + z) ≡ (x + y) + z) - (+IdR : (x : A) → x + 0a ≡ x) - (+InvR : (x : A) → x + (- x) ≡ 0a) - (+Comm : (x y : A) → x + y ≡ y + x) - (·Assoc : (x y z : A) → x · (y · z) ≡ (x · y) · z) - (·IdL : (x : A) → 1a · x ≡ x) - (·DistL+ : (x y z : A) → (x + y) · z ≡ (x · z) + (y · z)) - (·Comm : (x y : A) → x · y ≡ y · x) - (⋆Assoc : (r s : ⟨ R ⟩) (x : A) → (r ·s s) ⋆ x ≡ r ⋆ (s ⋆ x)) - (⋆DistR+ : (r : ⟨ R ⟩) (x y : A) → r ⋆ (x + y) ≡ (r ⋆ x) + (r ⋆ y)) - (⋆DistL+ : (r s : ⟨ R ⟩) (x : A) → (r +r s) ⋆ x ≡ (r ⋆ x) + (s ⋆ x)) - (⋆IdL : (x : A) → 1r ⋆ x ≡ x) - (⋆AssocL : (r : ⟨ R ⟩) (x y : A) → (r ⋆ x) · y ≡ r ⋆ (x · y)) + CommAlgebra→CommRing = fst + + ⟨_⟩ₐ : (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 ⟩ₐ) + → 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 ∈ (⟨ 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 = CommRingHom→CommAlgebraHom (idCommRingHom (fst A)) (CommRingHom≡ refl) + + idCommAlgebraHom = idCAlgHom + + compCommAlgebraHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} + → (f : CommAlgebraHom A B) → (g : CommAlgebraHom B C) + → CommAlgebraHom A C + compCommAlgebraHom {A = A} {B = B} {C = C} f g = + CommRingHom→CommAlgebraHom (⟨ g ⟩ᵣ→ ∘cr ⟨ f ⟩ᵣ→) pasting where + opaque + pasting : (⟨ g ⟩ᵣ→ ∘cr ⟨ f ⟩ᵣ→) ∘cr snd A ≡ snd C + pasting = + 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 + + _$ca_ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} → (φ : CommAlgebraHom A B) → (x : ⟨ A ⟩ₐ) → ⟨ B ⟩ₐ + _$ca_ φ x = φ .fst x + + 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 f g + + opaque + CommAlgebraHom≡ : + {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {f g : CommAlgebraHom A B} + → f .fst ≡ g .fst + → f ≡ g + 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 ∈ (⟨ A ⟩ₐ ≃ ⟨ B ⟩ₐ) ] IsCommAlgebraHom A B (f .fst) + + idCAlgEquiv : (A : CommAlgebra R ℓ') → CommAlgebraEquiv A A + 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 ≡ g .fst .fst + → f ≡ g + CommAlgebraEquiv≡ {B = B} p = + Σ≡Prop (isPropIsCommAlgebraHom _ _ ∘ fst) (Σ≡Prop isPropIsEquiv p) + + isSetCommAlgebraEquiv : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') + → isSet (CommAlgebraEquiv A B) + isSetCommAlgebraEquiv A B = + 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 ℓ''} + → (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 {A = A} {B = B} f pres⋆ = CommRingHom→CommAlgebraHom f + 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 + 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 + + 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) - makeIsCommAlgebra : IsCommAlgebra R 0a 1a _+_ _·_ -_ _⋆_ - makeIsCommAlgebra .IsCommAlgebra.isAlgebra = makeIsAlgebra - isSet-A - +Assoc +IdR +InvR +Comm - ·Assoc ·IdR ·IdL ·DistR+ ·DistL+ - ⋆Assoc - ⋆DistR+ - ⋆DistL+ - ⋆IdL - ⋆AssocR - ⋆AssocL - where - ·IdR : _ - ·IdR x = x · 1a ≡⟨ ·Comm _ _ ⟩ 1a · x ≡⟨ ·IdL _ ⟩ x ∎ - ·DistR+ : _ - ·DistR+ x y z = x · (y + z) ≡⟨ ·Comm _ _ ⟩ - (y + z) · x ≡⟨ ·DistL+ _ _ _ ⟩ - (y · x) + (z · x) ≡⟨ cong (λ u → (y · x) + u) (·Comm _ _) ⟩ - (y · x) + (x · z) ≡⟨ cong (λ u → u + (x · z)) (·Comm _ _) ⟩ - (x · y) + (x · z) ∎ - ⋆AssocR : _ - ⋆AssocR r x y = r ⋆ (x · y) ≡⟨ cong (λ u → r ⋆ u) (·Comm _ _) ⟩ - r ⋆ (y · x) ≡⟨ sym (⋆AssocL _ _ _) ⟩ - (r ⋆ y) · x ≡⟨ ·Comm _ _ ⟩ - x · (r ⋆ y) ∎ - makeIsCommAlgebra .IsCommAlgebra.·Comm = ·Comm - - 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) - - module _ - (_⋆_ : fst R → fst S → fst S) - (·Assoc⋆ : (r s : fst R) (x : fst S) → (r ·R s) ⋆ x ≡ r ⋆ (s ⋆ x)) - (⋆DistR+ : (r : fst R) (x y : fst S) → r ⋆ (x + y) ≡ (r ⋆ x) + (r ⋆ y)) - (⋆DistL+ : (r s : fst R) (x : fst S) → (r +R s) ⋆ x ≡ (r ⋆ x) + (s ⋆ x)) - (⋆IdL : (x : fst S) → 1R ⋆ x ≡ x) - (⋆AssocL : (r : fst R) (x y : fst S) → (r ⋆ x) · y ≡ r ⋆ (x · y)) - where - - commAlgebraFromCommRing : CommAlgebra R ℓ' - commAlgebraFromCommRing .fst = fst S - commAlgebraFromCommRing .snd .CommAlgebraStr.0a = 0r - commAlgebraFromCommRing .snd .CommAlgebraStr.1a = 1S - commAlgebraFromCommRing .snd .CommAlgebraStr._+_ = _+_ - commAlgebraFromCommRing .snd .CommAlgebraStr._·_ = _·_ - commAlgebraFromCommRing .snd .CommAlgebraStr.-_ = -_ - commAlgebraFromCommRing .snd .CommAlgebraStr._⋆_ = _⋆_ - commAlgebraFromCommRing .snd .CommAlgebraStr.isCommAlgebra = - makeIsCommAlgebra is-set +Assoc +IdR +InvR +Comm ·Assoc ·IdL ·DistL+ ·Comm - ·Assoc⋆ ⋆DistR+ ⋆DistL+ ⋆IdL ⋆AssocL - - commAlgebraFromCommRing→CommRing : CommAlgebra→CommRing commAlgebraFromCommRing ≡ S - -- Note that this is not definitional: the proofs of the axioms might differ. - commAlgebraFromCommRing→CommRing i .fst = ⟨ S ⟩ - commAlgebraFromCommRing→CommRing i .snd .CommRingStr.0r = 0r - commAlgebraFromCommRing→CommRing i .snd .CommRingStr.1r = 1S - commAlgebraFromCommRing→CommRing i .snd .CommRingStr._+_ = _+_ - commAlgebraFromCommRing→CommRing i .snd .CommRingStr._·_ = _·_ - commAlgebraFromCommRing→CommRing i .snd .CommRingStr.-_ = -_ - commAlgebraFromCommRing→CommRing i .snd .CommRingStr.isCommRing = - isProp→PathP (λ i → isPropIsCommRing _ _ _ _ _) - (CommRingStr.isCommRing (snd (CommAlgebra→CommRing commAlgebraFromCommRing))) - isCommRing - i - - IsCommAlgebraEquiv : {A : Type ℓ'} {B : Type ℓ''} - (M : CommAlgebraStr R A) (e : A ≃ B) (N : CommAlgebraStr R B) - → Type _ - IsCommAlgebraEquiv M e N = - IsAlgebraHom (CommAlgebraStr→AlgebraStr M) (e .fst) (CommAlgebraStr→AlgebraStr N) - - CommAlgebraEquiv : (M : CommAlgebra R ℓ') (N : CommAlgebra R ℓ'') → Type _ - CommAlgebraEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsCommAlgebraEquiv (M .snd) e (N .snd) - - IsCommAlgebraHom : {A : Type ℓ'} {B : Type ℓ''} - (M : CommAlgebraStr R A) (f : A → B) (N : CommAlgebraStr R B) - → Type _ - IsCommAlgebraHom M f N = - IsAlgebraHom (CommAlgebraStr→AlgebraStr M) f (CommAlgebraStr→AlgebraStr N) - - CommAlgebraHom : (M : CommAlgebra R ℓ') (N : CommAlgebra R ℓ'') → Type _ - CommAlgebraHom M N = Σ[ f ∈ (⟨ M ⟩ → ⟨ N ⟩) ] IsCommAlgebraHom (M .snd) f (N .snd) + CommAlgebraHom→RingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} + → CommAlgebraHom A B + → RingHom (CommAlgebra→Ring A) (CommAlgebra→Ring B) + CommAlgebraHom→RingHom = CommRingHom→RingHom ∘ CommAlgebraHom→CommRingHom CommAlgebraEquiv→CommAlgebraHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} - → CommAlgebraEquiv A B → CommAlgebraHom A B - CommAlgebraEquiv→CommAlgebraHom (e , eIsHom) = e .fst , eIsHom - - CommAlgebraHom→CommRingHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') - → 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) - - module _ {M : CommAlgebra R ℓ'} {N : CommAlgebra R ℓ''} where - open CommAlgebraStr {{...}} - open IsAlgebraHom - private - instance - _ = snd M - _ = snd N - - makeCommAlgebraHom : (f : fst M → fst N) - → (fPres1 : f 1a ≡ 1a) - → (fPres+ : (x y : fst M) → f (x + y) ≡ f x + f y) - → (fPres· : (x y : fst M) → f (x · y) ≡ f x · f y) - → (fPres⋆ : (r : fst R) (x : fst M) → f (r ⋆ x) ≡ r ⋆ f x) - → CommAlgebraHom M N - makeCommAlgebraHom f fPres1 fPres+ fPres· fPres⋆ = f , isHom - where fPres0 = - f 0a ≡⟨ sym (+IdR _) ⟩ - f 0a + 0a ≡⟨ cong (λ u → f 0a + u) (sym (+InvR (f 0a))) ⟩ - f 0a + (f 0a - f 0a) ≡⟨ +Assoc (f 0a) (f 0a) (- f 0a) ⟩ - (f 0a + f 0a) - f 0a ≡⟨ cong (λ u → u - f 0a) (sym (fPres+ 0a 0a)) ⟩ - f (0a + 0a) - f 0a ≡⟨ cong (λ u → f u - f 0a) (+IdL 0a) ⟩ - f 0a - f 0a ≡⟨ +InvR (f 0a) ⟩ - 0a ∎ - - isHom : IsCommAlgebraHom (snd M) f (snd N) - pres0 isHom = fPres0 - pres1 isHom = fPres1 - pres+ isHom = fPres+ - pres· isHom = fPres· - pres- isHom = (λ x → - f (- x) ≡⟨ sym (+IdR _) ⟩ - (f (- x) + 0a) ≡⟨ cong (λ u → f (- x) + u) (sym (+InvR (f x))) ⟩ - (f (- x) + (f x - f x)) ≡⟨ +Assoc _ _ _ ⟩ - ((f (- x) + f x) - f x) ≡⟨ cong (λ u → u - f x) (sym (fPres+ _ _)) ⟩ - (f ((- x) + x) - f x) ≡⟨ cong (λ u → f u - f x) (+InvL x) ⟩ - (f 0a - f x) ≡⟨ cong (λ u → u - f x) fPres0 ⟩ - (0a - f x) ≡⟨ +IdL _ ⟩ (- f x) ∎) - pres⋆ isHom = fPres⋆ - - isPropIsCommAlgebraHom : (f : fst M → fst N) → isProp (IsCommAlgebraHom (snd M) f (snd N)) - isPropIsCommAlgebraHom f = isPropIsAlgebraHom - (CommRing→Ring R) - (snd (CommAlgebra→Algebra M)) - f - (snd (CommAlgebra→Algebra N)) - -isPropIsCommAlgebra : (R : CommRing ℓ) {A : Type ℓ'} - (0a 1a : A) - (_+_ _·_ : A → A → A) - (-_ : A → A) - (_⋆_ : ⟨ R ⟩ → A → A) - → isProp (IsCommAlgebra R 0a 1a _+_ _·_ -_ _⋆_) -isPropIsCommAlgebra R _ _ _ _ _ _ = - isOfHLevelRetractFromIso 1 IsCommAlgebraIsoΣ - (isPropΣ (isPropIsAlgebra _ _ _ _ _ _ _) - (λ alg → isPropΠ2 λ _ _ → alg .IsAlgebra.is-set _ _)) - -𝒮ᴰ-CommAlgebra : (R : CommRing ℓ) → DUARel (𝒮-Univ ℓ') (CommAlgebraStr R) (ℓ-max ℓ ℓ') -𝒮ᴰ-CommAlgebra R = - 𝒮ᴰ-Record (𝒮-Univ _) (IsCommAlgebraEquiv {R = R}) - (fields: - data[ 0a ∣ nul ∣ pres0 ] - data[ 1a ∣ nul ∣ pres1 ] - data[ _+_ ∣ bin ∣ pres+ ] - data[ _·_ ∣ bin ∣ pres· ] - data[ -_ ∣ autoDUARel _ _ ∣ pres- ] - data[ _⋆_ ∣ autoDUARel _ _ ∣ pres⋆ ] - prop[ isCommAlgebra ∣ (λ _ _ → isPropIsCommAlgebra _ _ _ _ _ _ _) ]) - where - open CommAlgebraStr - open IsAlgebraHom - - -- faster with some sharing - nul = autoDUARel (𝒮-Univ _) (λ A → A) - bin = autoDUARel (𝒮-Univ _) (λ A → A → A → A) - -CommAlgebraPath : (R : CommRing ℓ) → (A B : CommAlgebra R ℓ') → (CommAlgebraEquiv A B) ≃ (A ≡ B) -CommAlgebraPath R = ∫ (𝒮ᴰ-CommAlgebra R) .UARel.ua - -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 _ _) --- -} + → CommAlgebraEquiv A B + → CommAlgebraHom A B + 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 + + 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/FGIdeal.agda b/Cubical/Algebra/CommAlgebra/FGIdeal.agda index e06ed81b11..d9b1a25ee0 100644 --- a/Cubical/Algebra/CommAlgebra/FGIdeal.agda +++ b/Cubical/Algebra/CommAlgebra/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.CommAlgebra.Base open import Cubical.Algebra.CommAlgebra.Ideal private @@ -21,17 +21,17 @@ private ℓ : Level R : CommRing ℓ -generatedIdeal : {n : ℕ} (A : CommAlgebra R ℓ) → FinVec (fst A) n → IdealsIn A +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) + (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 CommAlgebraStr (snd A) + open CommRingStr (A .fst .snd) - 0FGIdeal : {n : ℕ} → ⟨ replicateFinVec n 0a ⟩[ A ] ≡ (0Ideal 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/FPAlgebra.agda b/Cubical/Algebra/CommAlgebra/FP.agda similarity index 71% rename from Cubical/Algebra/CommAlgebra/FPAlgebra.agda rename to Cubical/Algebra/CommAlgebra/FP.agda index d454e4f0a1..41f07ffa70 100644 --- a/Cubical/Algebra/CommAlgebra/FPAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra/FP.agda @@ -9,6 +9,6 @@ Our definition is more explicit. -} {-# OPTIONS --safe #-} -module Cubical.Algebra.CommAlgebra.FPAlgebra where +module Cubical.Algebra.CommAlgebra.FP where -open import Cubical.Algebra.CommAlgebra.FPAlgebra.Base public +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..ec78eea357 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/FP/Base.agda @@ -0,0 +1,60 @@ +{-# OPTIONS --safe --lossy-unification #-} +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.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 + + +private + variable + ℓ ℓ' : Level + +module _ (R : CommRing ℓ) where + Polynomials : (n : ℕ) → CommAlgebra R ℓ + Polynomials n = R [ Fin n ]ₐ + + record FinitePresentation : Type ℓ where + no-eta-equality + field + n : ℕ + m : ℕ + relations : FinVec ⟨ Polynomials n ⟩ₐ m + + 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 + + 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 new file mode 100644 index 0000000000..9535984546 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/FP/Instances.agda @@ -0,0 +1,226 @@ +{- + 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 --lossy-unification #-} +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; 0Ideal) +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.FP.Base + +private + variable + ℓ ℓ' : Level + + +module _ (R : CommRing ℓ) where + open FinitePresentation + + {- Every (multivariate) polynomial algebra is finitely presented -} + module _ (n : ℕ) where + private + opaque + 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 + polynomialsFP = + record { + n = n ; + m = 0 ; + 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 R 0 + + emptyGen : FinVec ⟨ R[⊥] ⟩ₐ 0 + emptyGen = λ () + + initialAlgFP : FinitePresentation R + initialAlgFP = record { n = 0 ; m = 0 ; relations = emptyGen } + + R[⊥]/⟨0⟩ : CommAlgebra R ℓ + 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) + 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) +-- -} 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/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/CommAlgebra/Ideal.agda b/Cubical/Algebra/CommAlgebra/Ideal.agda index b3f8b0d126..f80a3b70b3 100644 --- a/Cubical/Algebra/CommAlgebra/Ideal.agda +++ b/Cubical/Algebra/CommAlgebra/Ideal.agda @@ -7,30 +7,30 @@ 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.Ring +open import Cubical.Algebra.CommAlgebra.Base +-- open import Cubical.Algebra.Ring 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 (CommAlgebra→CommRing A) + IdealsIn = IdealsInCommRing (fst A) - open CommAlgebraStr (snd A) + open CommRingStr (A .fst .snd) - makeIdeal : (I : fst A → hProp ℓ) - → (+-closed : {x y : fst A} → x ∈ I → y ∈ I → (x + y) ∈ I) - → (0-closed : 0a ∈ I) - → (·-closedLeft : {x : fst A} → (r : fst A) → x ∈ I → r · x ∈ I) + 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 = CommAlgebra→CommRing A} + makeIdeal = makeIdealCommRing {R = fst A} 0Ideal : IdealsIn - 0Ideal = CommIdeal.0Ideal (CommAlgebra→CommRing A) + 0Ideal = CommIdeal.0Ideal (fst A) 1Ideal : IdealsIn - 1Ideal = CommIdeal.1Ideal (CommAlgebra→CommRing A) + 1Ideal = CommIdeal.1Ideal (fst A) diff --git a/Cubical/Algebra/CommAlgebra/Instances/Initial.agda b/Cubical/Algebra/CommAlgebra/Instances/Initial.agda index 26c8b6bdd6..899d6d049c 100644 --- a/Cubical/Algebra/CommAlgebra/Instances/Initial.agda +++ b/Cubical/Algebra/CommAlgebra/Instances/Initial.agda @@ -4,6 +4,7 @@ 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) @@ -22,69 +23,31 @@ private ℓ : Level module _ (R : CommRing ℓ) where - module _ where open CommRingStr (snd R) initialCAlg : CommAlgebra R ℓ - initialCAlg .fst = fst R - initialCAlg .snd .CommAlgebraStr.0a = _ - initialCAlg .snd .CommAlgebraStr.1a = _ - initialCAlg .snd .CommAlgebraStr._+_ = _ - initialCAlg .snd .CommAlgebraStr._·_ = _ - initialCAlg .snd .CommAlgebraStr.-_ = _ - initialCAlg .snd .CommAlgebraStr._⋆_ r x = r · x - initialCAlg .snd .CommAlgebraStr.isCommAlgebra = - makeIsCommAlgebra is-set - +Assoc +IdR +InvR +Comm - ·Assoc ·IdL - ·DistL+ ·Comm - (λ x y z → sym (·Assoc x y z)) ·DistR+ ·DistL+ ·IdL - λ x y z → sym (·Assoc x y z) + initialCAlg .fst = R + initialCAlg .snd = idCommRingHom R module _ (A : CommAlgebra R ℓ) where - open CommAlgebraStr ⦃... ⦄ - private - instance - _ : CommAlgebraStr R (fst A) - _ = snd A - _ : CommAlgebraStr R (fst R) - _ = snd initialCAlg - - _*_ : fst R → (fst A) → (fst A) - r * a = CommAlgebraStr._⋆_ (snd A) r a +-- open CommAlgebraStr ⦃... ⦄ initialMap : CommAlgebraHom initialCAlg A - initialMap = - makeCommAlgebraHom {M = initialCAlg} {N = A} - (λ r → r * 1a) - (⋆IdL _) - (λ x y → ⋆DistL+ x y 1a) - (λ x y → (x · y) * 1a ≡⟨ ⋆Assoc _ _ _ ⟩ - x * (y * 1a) ≡[ i ]⟨ x * (·IdL (y * 1a) (~ i)) ⟩ - x * (1a · (y * 1a)) ≡⟨ sym (⋆AssocL _ _ _) ⟩ - (x * 1a) · (y * 1a) ∎) - (λ r x → (r · x) * 1a ≡⟨ ⋆Assoc _ _ _ ⟩ - (r * (x * 1a)) ∎) + initialMap = CommRingHom→CommAlgebraHom (A .snd) (CommRingHom≡ refl) initialMapEq : (f : CommAlgebraHom initialCAlg A) → f ≡ initialMap - initialMapEq f = - let open IsAlgebraHom (snd f) - in Σ≡Prop - (isPropIsCommAlgebraHom {M = initialCAlg} {N = A}) - λ i x → - ((fst f) x ≡⟨ cong (fst f) (sym (·IdR _)) ⟩ - fst f (x · 1a) ≡⟨ pres⋆ x 1a ⟩ - CommAlgebraStr._⋆_ (snd A) x (fst f 1a) ≡⟨ cong - (λ u → (snd A CommAlgebraStr.⋆ x) u) - pres1 ⟩ - (CommAlgebraStr._⋆_ (snd A) x 1a) ∎) i + 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) @@ -97,6 +60,7 @@ module _ (R : CommRing ℓ) where 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 @@ -108,23 +72,23 @@ module _ (R : CommRing ℓ) where → CommAlgebraEquiv A (initialCAlg) equivByInitiality isInitial = isoToEquiv asIso , snd to where - open CommAlgebraHoms + -- open CommAlgebraHoms to : CommAlgebraHom A initialCAlg to = fst (isInitial initialCAlg) from : CommAlgebraHom initialCAlg A from = initialMap A - asIso : Iso (fst A) (fst initialCAlg) + 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 ∘a from) (idCAlgHom initialCAlg)) + (isContr→isProp (initialityContr initialCAlg) (to ∘ca from) (idCAlgHom initialCAlg)) i x Iso.leftInv asIso = λ x i → cong fst - (isContr→isProp (isInitial A) (from ∘a to) (idCAlgHom A)) + (isContr→isProp (isInitial A) (from ∘ca to) (idCAlgHom A)) i x diff --git a/Cubical/Algebra/CommAlgebra/Instances/Polynomials.agda b/Cubical/Algebra/CommAlgebra/Instances/Polynomials.agda new file mode 100644 index 0000000000..907ae3f3b4 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/Instances/Polynomials.agda @@ -0,0 +1,168 @@ +{-# OPTIONS --safe #-} +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 as Poly hiding (var) +import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.UniversalProperty + as Poly + +private + variable + ℓ ℓ' ℓ'' : Level + +_[_]ₐ : (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 ℓ'') + (f : CommAlgebraHom (R [ I ]) A) + → (I → fst A) +evaluateAt A f x = f $a (Construction.var x) + + +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 +-} diff --git a/Cubical/Algebra/CommAlgebra/Instances/Unit.agda b/Cubical/Algebra/CommAlgebra/Instances/Unit.agda index 7154fc6b84..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 @@ -11,7 +12,6 @@ open import Cubical.Data.Sigma.Properties using (Σ≡Prop) open import Cubical.Algebra.CommRing open import Cubical.Algebra.CommAlgebra.Base open import Cubical.Algebra.CommRing.Instances.Unit -open import Cubical.Algebra.Algebra.Base using (IsAlgebraHom) open import Cubical.Tactics.CommRingSolver private @@ -20,23 +20,14 @@ private module _ (R : CommRing ℓ) where UnitCommAlgebra : CommAlgebra R ℓ' - UnitCommAlgebra = - commAlgebraFromCommRing - UnitCommRing - (λ _ _ → tt*) (λ _ _ _ → refl) (λ _ _ _ → refl) - (λ _ _ _ → refl) (λ _ → refl) (λ _ _ _ → refl) + UnitCommAlgebra = UnitCommRing , mapToUnitCommRing R - module _ (A : CommAlgebra R ℓ) where + module _ (A : CommAlgebra R ℓ') where terminalMap : CommAlgebraHom A (UnitCommAlgebra {ℓ' = ℓ}) - terminalMap = (λ _ → tt*) , isHom - where open IsAlgebraHom - isHom : IsCommAlgebraHom (snd A) _ (snd UnitCommAlgebra) - pres0 isHom = isPropUnit* _ _ - pres1 isHom = isPropUnit* _ _ - pres+ isHom = λ _ _ → isPropUnit* _ _ - pres· isHom = λ _ _ → isPropUnit* _ _ - pres- isHom = λ _ → isPropUnit* _ _ - pres⋆ isHom = λ _ _ → isPropUnit* _ _ + terminalMap = CommRingHom→CommAlgebraHom (mapToUnitCommRing (A .fst)) $ isPropMapToUnitCommRing _ _ _ + +{- + module _ (A : CommAlgebra R ℓ) where terminalityContr : isContr (CommAlgebraHom A UnitCommAlgebra) terminalityContr = terminalMap , path @@ -58,3 +49,4 @@ module _ (R : CommRing ℓ) where equivFrom1≡0 : CommAlgebraEquiv A UnitCommAlgebra equivFrom1≡0 = isContr→Equiv 1≡0→isContr isContrUnit* , snd terminalMap +-} diff --git a/Cubical/Algebra/CommAlgebra/Kernel.agda b/Cubical/Algebra/CommAlgebra/Kernel.agda index 41d8e9af96..352b42aa68 100644 --- a/Cubical/Algebra/CommAlgebra/Kernel.agda +++ b/Cubical/Algebra/CommAlgebra/Kernel.agda @@ -8,14 +8,14 @@ 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.Properties open import Cubical.Algebra.CommAlgebra.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 A + kernel : IdealsIn R A kernel = Ideal→CommIdeal (ringKernel (CommAlgebraHom→RingHom {A = A} {B = B} ϕ)) diff --git a/Cubical/Algebra/CommAlgebra/Localisation.agda b/Cubical/Algebra/CommAlgebra/Localisation.agda index 38856fbbbc..1e5f57f47d 100644 --- a/Cubical/Algebra/CommAlgebra/Localisation.agda +++ b/Cubical/Algebra/CommAlgebra/Localisation.agda @@ -6,6 +6,7 @@ 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 @@ -27,120 +28,113 @@ 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.Algebra open import Cubical.Algebra.CommAlgebra.Base -open import Cubical.Algebra.CommAlgebra.Properties +open import Cubical.Algebra.CommAlgebra.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' : ℙ (fst R')) (SMultClosedSubset : isMultClosedSubset R' S') where +module AlgLoc (R : CommRing ℓ) + (S : ℙ ⟨ R ⟩) (SMultClosedSubset : isMultClosedSubset R S) where open isMultClosedSubset - private R = fst R' - open CommAlgebraStr - open IsAlgebraHom - open CommRingStr (snd R') renaming (_+_ to _+r_ ; _·_ to _·r_ ; ·IdR to ·rRid) - open RingTheory (CommRing→Ring R') - open CommRingTheory R' - open Loc R' S' SMultClosedSubset - open S⁻¹RUniversalProp R' S' SMultClosedSubset - open CommAlgChar R' - - - S⁻¹RAsCommAlg : CommAlgebra R' ℓ - S⁻¹RAsCommAlg = toCommAlg (S⁻¹RAsCommRing , /1AsCommRingHom) - - LocCommAlg→CommRingPath : CommAlgebra→CommRing S⁻¹RAsCommAlg ≡ S⁻¹RAsCommRing - LocCommAlg→CommRingPath = CommAlgebra→CommRing≡ (S⁻¹RAsCommRing , /1AsCommRingHom) - - hasLocAlgUniversalProp : (A : CommAlgebra R' ℓ) - → (∀ s → s ∈ S' → _⋆_ (snd A) s (1a (snd A)) ∈ (CommAlgebra→CommRing A) ˣ) - → Type (ℓ-suc ℓ) - hasLocAlgUniversalProp A _ = (B : CommAlgebra R' ℓ) - → (∀ s → s ∈ S' → _⋆_ (snd B) s (1a (snd B)) ∈ (CommAlgebra→CommRing B) ˣ) - → isContr (CommAlgebraHom A B) - - S⋆1⊆S⁻¹Rˣ : ∀ s → s ∈ S' → _⋆_ (snd S⁻¹RAsCommAlg) s (1a (snd S⁻¹RAsCommAlg)) ∈ S⁻¹Rˣ + 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 (·rRid s)) (Σ≡Prop (λ x → S' x .snd) (sym (·rRid _))))) + (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 + S⁻¹RHasAlgUniversalProp B S⋆1⊆Bˣ' = χᴬ , χᴬuniqueness where - B = fromCommAlg B' .fst - φ = fromCommAlg B' .snd - open CommRingStr (snd B) renaming (_·_ to _·b_ ; 1r to 1b ; ·IdL to ·bLid) + φ = 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 - χ = S⁻¹RHasUniversalProp B φ S⋆1⊆Bˣ .fst .fst + χ : 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 φ S⋆1⊆Bˣ .fst .snd) - - χᴬ : CommAlgebraHom S⁻¹RAsCommAlg B' - fst χᴬ = fst χ - pres0 (snd χᴬ) = IsRingHom.pres0 (snd χ) - pres1 (snd χᴬ) = IsRingHom.pres1 (snd χ) - pres+ (snd χᴬ) = IsRingHom.pres+ (snd χ) - pres· (snd χᴬ) = IsRingHom.pres· (snd χ) - pres- (snd χᴬ) = IsRingHom.pres- (snd χ) - pres⋆ (snd χᴬ) r x = path - where - path : fst χ ((r /1) ·ₗ x) ≡ _⋆_ (snd B') r (fst χ x) - path = fst χ ((r /1) ·ₗ x) ≡⟨ IsRingHom.pres· (snd χ) _ _ ⟩ - fst χ (r /1) ·b fst χ x ≡⟨ cong (_·b fst χ x) (χcomp r) ⟩ - fst φ r ·b fst χ x ≡⟨ refl ⟩ - _⋆_ (snd B') r 1b ·b fst χ x ≡⟨ ⋆AssocL (snd B') _ _ _ ⟩ - _⋆_ (snd B') r (1b ·b fst χ x) ≡⟨ cong (_⋆_ (snd B') r) (·bLid _) ⟩ - _⋆_ (snd B') r (fst χ x) ∎ - - - χᴬuniqueness : (ψ : CommAlgebraHom S⁻¹RAsCommAlg B') → χᴬ ≡ ψ - χᴬuniqueness ψ = Σ≡Prop (λ _ → isPropIsAlgebraHom _ _ _ _) - (cong (fst ∘ fst) (χuniqueness (ψ' , funExt ψ'r/1≡φr))) + χcomp = funExt⁻ (S⁻¹RHasUniversalProp (B .fst) φ S⋆1⊆Bˣ .fst .snd) + + χᴬ : CommAlgebraHom S⁻¹RAsCommAlg B + χᴬ = CommAlgebraHomFromCommRingHom χ path where - χuniqueness = S⁻¹RHasUniversalProp B φ S⋆1⊆Bˣ .snd + 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 ≡⟨ cong (fst ∘ fst) (χuniqueness (CommAlgebraHom→CommRingHom ψ , funExt ψ'r/1≡φr)) ⟩ ψ .fst ∎) - ψ' : 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 ψ) + where + χuniqueness = S⁻¹RHasUniversalProp (B .fst) φ S⋆1⊆Bˣ .snd - ψ'r/1≡φr : ∀ r → fst ψ (r /1) ≡ fst φ r + ψ'r/1≡φr : ∀ r → ψ .fst (r /1) ≡ fst φ r ψ'r/1≡φr r = - fst ψ (r /1) ≡⟨ cong (fst ψ) (sym (·ₗ-rid _)) ⟩ - fst ψ (_⋆_ (snd S⁻¹RAsCommAlg) r (1a (snd S⁻¹RAsCommAlg))) ≡⟨ pres⋆ (snd ψ) _ _ ⟩ - _⋆_ (snd B') r (fst ψ (1a (snd S⁻¹RAsCommAlg))) ≡⟨ cong (_⋆_ (snd B') r) (pres1 (snd ψ)) ⟩ - _⋆_ (snd B') r 1b ∎ + ψ .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 ∎ -- 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 (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)) - where open PathToS⁻¹R - + S⁻¹RAlgCharEquiv : (A : CommRing ℓ) (φ : CommRingHom R A) + → PathToS⁻¹R R S SMultClosedSubset A φ + → CommAlgebraEquiv S⁻¹RAsCommAlg (A , φ) + 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 ℓ @@ -154,13 +148,12 @@ module _ {R : CommRing ℓ} (f : fst R) where open AlgLoc R [ f ⁿ|n≥0] (powersFormMultClosedSubset f) invElCommAlgebra→CommRingPath : CommAlgebra→CommRing R[1/ f ]AsCommAlgebra ≡ R[1/ f ]AsCommRing - invElCommAlgebra→CommRingPath = LocCommAlg→CommRingPath + 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 CommRingStr (snd R') hiding (is-set) open RingTheory (CommRing→Ring R') open Loc R' S₁ S₁MultClosedSubset renaming (S⁻¹R to S₁⁻¹R ; S⁻¹RAsCommRing to S₁⁻¹RAsCommRing) @@ -175,26 +168,26 @@ module AlgLocTwoSubsets (R' : CommRing ℓ) ; S⁻¹RHasAlgUniversalProp to S₂⁻¹RHasAlgUniversalProp ; isContrHomS⁻¹RS⁻¹R to isContrHomS₂⁻¹RS₂⁻¹R) - open IsAlgebraHom - open AlgebraHoms - open CommAlgebraHoms open CommAlgebraStr ⦃...⦄ + open CommRingStr ⦃...⦄ private - R = fst R' S₁⁻¹Rˣ = S₁⁻¹RAsCommRing ˣ S₂⁻¹Rˣ = S₂⁻¹RAsCommRing ˣ instance - _ = snd S₁⁻¹RAsCommAlg - _ = snd S₂⁻¹RAsCommAlg + _ = CommAlgebra→CommAlgebraStr S₁⁻¹RAsCommAlg + _ = CommAlgebra→CommAlgebraStr S₂⁻¹RAsCommAlg + _ = S₂⁻¹RAsCommAlg .fst .snd + _ = S₁⁻¹RAsCommAlg .fst .snd + _ = R' .snd - isContrS₁⁻¹R≡S₂⁻¹R : (∀ s₁ → s₁ ∈ S₁ → s₁ ⋆ 1a ∈ S₂⁻¹Rˣ) - → (∀ s₂ → s₂ ∈ S₂ → s₂ ⋆ 1a ∈ S₁⁻¹Rˣ) + 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 + (equivToIso (invEquiv (CommAlgebraPath _ _))) + isContrS₁⁻¹R≅S₂⁻¹R where χ₁ : CommAlgebraHom S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg χ₁ = S₁⁻¹RHasAlgUniversalProp S₂⁻¹RAsCommAlg S₁⊆S₂⁻¹Rˣ .fst @@ -202,70 +195,74 @@ module AlgLocTwoSubsets (R' : CommRing ℓ) χ₂ : CommAlgebraHom S₂⁻¹RAsCommAlg S₁⁻¹RAsCommAlg χ₂ = S₂⁻¹RHasAlgUniversalProp S₁⁻¹RAsCommAlg S₂⊆S₁⁻¹Rˣ .fst - χ₁∘χ₂≡id : χ₁ ∘a χ₂ ≡ idCommAlgebraHom _ + χ₁∘χ₂≡id : χ₁ ∘ca χ₂ ≡ idCAlgHom _ χ₁∘χ₂≡id = isContr→isProp isContrHomS₂⁻¹RS₂⁻¹R _ _ - χ₂∘χ₁≡id : χ₂ ∘a χ₁ ≡ idCommAlgebraHom _ + χ₂∘χ₁≡id : χ₂ ∘ca χ₁ ≡ idCAlgHom _ χ₂∘χ₁≡id = isContr→isProp isContrHomS₁⁻¹RS₁⁻¹R _ _ IsoS₁⁻¹RS₂⁻¹R : Iso S₁⁻¹R S₂⁻¹R - Iso.fun IsoS₁⁻¹RS₂⁻¹R = fst χ₁ - Iso.inv IsoS₁⁻¹RS₂⁻¹R = fst χ₂ - Iso.rightInv IsoS₁⁻¹RS₂⁻¹R = funExt⁻ (cong fst χ₁∘χ₂≡id) - Iso.leftInv IsoS₁⁻¹RS₂⁻¹R = funExt⁻ (cong fst χ₂∘χ₁≡id) + 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 = center , uniqueness where + X₁asCRHom = CommAlgebraHom→CommRingHom χ₁ center : CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg - fst center = isoToEquiv IsoS₁⁻¹RS₂⁻¹R - pres0 (snd center) = pres0 (snd χ₁) - pres1 (snd center) = pres1 (snd χ₁) - pres+ (snd center) = pres+ (snd χ₁) - pres· (snd center) = pres· (snd χ₁) - pres- (snd center) = pres- (snd χ₁) - pres⋆ (snd center) = pres⋆ (snd χ₁) + center = + CommRingEquiv→CommAlgebraEquiv + ((isoToEquiv IsoS₁⁻¹RS₂⁻¹R) , snd X₁asCRHom) + (cong fst (CommAlgebraHom→Triangle χ₁)) uniqueness : (φ : CommAlgebraEquiv S₁⁻¹RAsCommAlg S₂⁻¹RAsCommAlg) → center ≡ φ - uniqueness φ = Σ≡Prop (λ _ → isPropIsAlgebraHom _ _ _ _) - (equivEq (cong fst - (S₁⁻¹RHasAlgUniversalProp S₂⁻¹RAsCommAlg S₁⊆S₂⁻¹Rˣ .snd - (AlgebraEquiv→AlgebraHom φ)))) + uniqueness φ = + CommAlgebraEquiv≡ (cong ⟨_⟩ₐ→ $ S₁⁻¹RHasAlgUniversalProp S₂⁻¹RAsCommAlg S₁⊆S₂⁻¹Rˣ .snd + (CommAlgebraEquiv→CommAlgebraHom φ)) - isPropS₁⁻¹R≡S₂⁻¹R : isProp (S₁⁻¹RAsCommAlg ≡ S₂⁻¹RAsCommAlg) - isPropS₁⁻¹R≡S₂⁻¹R S₁⁻¹R≡S₂⁻¹R = - isContr→isProp (isContrS₁⁻¹R≡S₂⁻¹R S₁⊆S₂⁻¹Rˣ S₂⊆S₁⁻¹Rˣ) S₁⁻¹R≡S₂⁻¹R - where - S₁⊆S₂⁻¹Rˣ : ∀ s₁ → s₁ ∈ S₁ → s₁ ⋆ 1a ∈ S₂⁻¹Rˣ - S₁⊆S₂⁻¹Rˣ s₁ s₁∈S₁ = - transport (λ i → _⋆_ ⦃ S₁⁻¹R≡S₂⁻¹R i .snd ⦄ s₁ (1a ⦃ S₁⁻¹R≡S₂⁻¹R i .snd ⦄) - ∈ (CommAlgebra→CommRing (S₁⁻¹R≡S₂⁻¹R i)) ˣ) (S₁⋆1⊆S₁⁻¹Rˣ s₁ s₁∈S₁) + 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₂ ⋆ 1a ∈ S₁⁻¹Rˣ - S₂⊆S₁⁻¹Rˣ s₂ s₂∈S₂ = - transport (λ i → _⋆_ ⦃ (sym S₁⁻¹R≡S₂⁻¹R) i .snd ⦄ s₂ (1a ⦃ (sym S₁⁻¹R≡S₂⁻¹R) i .snd ⦄) - ∈ (CommAlgebra→CommRing ((sym S₁⁻¹R≡S₂⁻¹R) i)) ˣ) (S₂⋆1⊆S₂⁻¹Rˣ s₂ s₂∈S₂) + 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 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 CommAlgChar R 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 ] + ⟨_⟩ᵢ : (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) ˣ @@ -275,12 +272,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) @@ -290,29 +289,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)) diff --git a/Cubical/Algebra/CommAlgebra/Properties.agda b/Cubical/Algebra/CommAlgebra/Properties.agda index fd75d7d0cd..068fb46522 100644 --- a/Cubical/Algebra/CommAlgebra/Properties.agda +++ b/Cubical/Algebra/CommAlgebra/Properties.agda @@ -1,382 +1,82 @@ -{-# OPTIONS --safe --lossy-unification #-} +{-# OPTIONS --safe #-} module Cubical.Algebra.CommAlgebra.Properties where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv open import Cubical.Foundations.Function -open import Cubical.Foundations.Equiv.HalfAdjoint -open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Transport -open import Cubical.Foundations.SIP -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Path +open import Cubical.Foundations.Structure using (⟨_⟩; withOpaqueStr) open import Cubical.Data.Sigma -open import Cubical.Reflection.StrictEquiv - -open import Cubical.Structures.Axioms -open import Cubical.Algebra.Semigroup -open import Cubical.Algebra.Monoid open import Cubical.Algebra.CommRing -open import Cubical.Algebra.Ring -open import Cubical.Algebra.Algebra open import Cubical.Algebra.CommAlgebra.Base - -open import Cubical.Algebra.CommRing using (CommRing→Ring) - -open import Cubical.HITs.PropositionalTruncation +open import Cubical.Algebra.CommAlgebra.Univalence private variable ℓ ℓ' ℓ'' ℓ''' : Level - R : CommRing ℓ - -open AlgebraHoms - -idCAlgHom : (A : CommAlgebra R ℓ) → _ -idCAlgHom A = idAlgebraHom (CommAlgebra→Algebra A) - -idCAlgEquiv : (A : CommAlgebra R ℓ) → CommAlgebraEquiv A A -fst (idCAlgEquiv A) = idEquiv (fst A) -snd (idCAlgEquiv A) = snd (idCAlgHom A) - -infix 3 _≃CAlg∎ -infixr 2 _≃CAlg⟨_⟩_ - -_≃CAlg∎ : (A : CommAlgebra R ℓ) → CommAlgebraEquiv A A -A ≃CAlg∎ = idCAlgEquiv A - -_≃CAlg⟨_⟩_ : {B C : CommAlgebra R ℓ} - (A : CommAlgebra R ℓ) (f : CommAlgebraEquiv A B) (g : CommAlgebraEquiv B C) - → CommAlgebraEquiv A C -A ≃CAlg⟨ f ⟩ g = g ∘≃a f - --- An R-algebra is the same as a CommRing A with a CommRingHom φ : R → A -module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where - open Iso - open CommRingTheory - - - CommRingWithHom : Type (ℓ-max ℓ (ℓ-suc ℓ')) - CommRingWithHom = Σ[ A ∈ CommRing ℓ' ] CommRingHom R A - - toCommAlg : CommRingWithHom → CommAlgebra R ℓ' - toCommAlg (A , φ , φIsHom) = - commAlgebraFromCommRing - A - (λ r a → (φ r) · a) - (λ r s x → cong (_· x) (pres· r s) ∙ sym (·Assoc _ _ _)) - (λ r x y → ·DistR+ _ _ _) - (λ r s x → cong (_· x) (pres+ r s) ∙ ·DistL+ _ _ _) - (λ x → cong (_· x) pres1 ∙ ·IdL x) - λ r x y → sym (·Assoc _ _ _) - where - open CommRingStr (snd A) - open IsRingHom φIsHom - - fromCommAlg : CommAlgebra R ℓ' → CommRingWithHom - fromCommAlg A = (CommAlgebra→CommRing A) , φ , φIsHom - where - open CommRingStr (snd R) renaming (_·_ to _·r_) hiding (·IdL) - open CommAlgebraStr (snd A) - open AlgebraTheory (CommRing→Ring R) (CommAlgebra→Algebra A) - φ : ⟨ R ⟩ → ⟨ A ⟩ - φ r = r ⋆ 1a - φIsHom : IsRingHom (CommRing→Ring R .snd) φ (CommRing→Ring (CommAlgebra→CommRing A) .snd) - φIsHom = makeIsRingHom (⋆IdL _) (λ _ _ → ⋆DistL+ _ _ _) - λ x y → cong (λ a → (x ·r y) ⋆ a) (sym (·IdL _)) ∙ ⋆Dist· _ _ _ _ - - -- helpful for localisations - module _ (Aφ : CommRingWithHom) where - open CommRingStr - private - A = fst Aφ - CommAlgebra→CommRing≡ : CommAlgebra→CommRing (toCommAlg Aφ) ≡ A - fst (CommAlgebra→CommRing≡ i) = fst A - 0r (snd (CommAlgebra→CommRing≡ i)) = 0r (snd A) - 1r (snd (CommAlgebra→CommRing≡ i)) = 1r (snd A) - _+_ (snd (CommAlgebra→CommRing≡ i)) = _+_ (snd A) - _·_ (snd (CommAlgebra→CommRing≡ i)) = _·_ (snd A) - -_ (snd (CommAlgebra→CommRing≡ i)) = -_ (snd A) - -- 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 - open CommRingStr - -- 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 - - CommAlgRoundTrip : (A : CommAlgebra R ℓ') → toCommAlg (fromCommAlg A) ≡ A - CommAlgRoundTrip A = ΣPathP (refl , AlgStrPathP) - where - open CommAlgebraStr ⦃...⦄ - instance - _ = snd A - AlgStrPathP : PathP (λ i → CommAlgebraStr R ⟨ A ⟩) (snd (toCommAlg (fromCommAlg A))) (snd A) - CommAlgebraStr.0a (AlgStrPathP i) = 0a - CommAlgebraStr.1a (AlgStrPathP i) = 1a - CommAlgebraStr._+_ (AlgStrPathP i) = _+_ - CommAlgebraStr._·_ (AlgStrPathP i) = _·_ - CommAlgebraStr.-_ (AlgStrPathP i) = -_ - CommAlgebraStr._⋆_ (AlgStrPathP i) r x = (⋆AssocL r 1a x ∙ cong (r ⋆_) (·IdL x)) i - CommAlgebraStr.isCommAlgebra (AlgStrPathP i) = isProp→PathP - (λ i → isPropIsCommAlgebra _ _ _ _ _ _ (CommAlgebraStr._⋆_ (AlgStrPathP i))) - (CommAlgebraStr.isCommAlgebra (snd (toCommAlg (fromCommAlg A)))) isCommAlgebra i - - CommAlgIso : Iso (CommAlgebra R ℓ') CommRingWithHom - fun CommAlgIso = fromCommAlg - inv CommAlgIso = toCommAlg - rightInv CommAlgIso = CommRingWithHomRoundTrip - leftInv CommAlgIso = CommAlgRoundTrip - - open RingHoms - open IsRingHom - - isCommRingWithHomHom : (A B : CommRingWithHom) → CommRingHom (fst A) (fst B) → Type (ℓ-max ℓ ℓ') - isCommRingWithHomHom (_ , f) (_ , g) h = h ∘r f ≡ g - - CommRingWithHomHom : CommRingWithHom → CommRingWithHom → Type (ℓ-max ℓ ℓ') - CommRingWithHomHom (A , f) (B , g) = Σ[ h ∈ CommRingHom A B ] h ∘r f ≡ g - - toCommAlgebraHom : (A B : CommRingWithHom) (h : CommRingHom (fst A) (fst B)) - → isCommRingWithHomHom A B h - → CommAlgebraHom (toCommAlg A) (toCommAlg B) - toCommAlgebraHom (A , f) (B , g) h commDiag = - makeCommAlgebraHom (fst h) (pres1 (snd h)) (pres+ (snd h)) (pres· (snd h)) pres⋆h - where - open CommRingStr ⦃...⦄ - instance - _ = snd A - _ = snd B - pres⋆h : ∀ r x → fst h (fst f r · x) ≡ fst g r · fst h x - pres⋆h r x = fst h (fst f r · x) ≡⟨ pres· (snd h) _ _ ⟩ - fst h (fst f r) · fst h x ≡⟨ cong (λ φ → fst φ r · fst h x) commDiag ⟩ - fst g r · fst h x ∎ - - fromCommAlgebraHom : (A B : CommAlgebra R ℓ') → CommAlgebraHom A B - → CommRingWithHomHom (fromCommAlg A) (fromCommAlg B) - fst (fst (fromCommAlgebraHom A B f)) = fst f - pres0 (snd (fst (fromCommAlgebraHom A B f))) = IsAlgebraHom.pres0 (snd f) - pres1 (snd (fst (fromCommAlgebraHom A B f))) = IsAlgebraHom.pres1 (snd f) - pres+ (snd (fst (fromCommAlgebraHom A B f))) = IsAlgebraHom.pres+ (snd f) - 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)))) - 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) - CommRingWithHomEquiv : CommRingWithHom → CommRingWithHom → Type (ℓ-max ℓ ℓ') - CommRingWithHomEquiv A B = Σ[ e ∈ CommRingEquiv (fst A) (fst B) ] isCommRingWithHomEquiv A B e - - toCommAlgebraEquiv : (A B : CommRingWithHom) (e : CommRingEquiv (fst A) (fst B)) - → isCommRingWithHomEquiv A B e - → CommAlgebraEquiv (toCommAlg A) (toCommAlg B) - fst (toCommAlgebraEquiv A B e eCommDiag) = e .fst - snd (toCommAlgebraEquiv A B e eCommDiag) = toCommAlgebraHom A B _ eCommDiag .snd - - - -module CommAlgebraHoms {R : CommRing ℓ} where - open AlgebraHoms - - idCommAlgebraHom : (A : CommAlgebra R ℓ') → CommAlgebraHom A A - idCommAlgebraHom A = idAlgebraHom (CommAlgebra→Algebra A) - - compCommAlgebraHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'') (C : CommAlgebra R ℓ''') - → CommAlgebraHom A B → CommAlgebraHom B C → CommAlgebraHom A C - compCommAlgebraHom A B C = compAlgebraHom {A = CommAlgebra→Algebra A} - {B = CommAlgebra→Algebra B} - {C = CommAlgebra→Algebra C} - - _∘ca_ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''} - → CommAlgebraHom B C → CommAlgebraHom A B → CommAlgebraHom A C - g ∘ca f = compCommAlgebraHom _ _ _ f g - - compIdCommAlgebraHom : {A B : CommAlgebra R ℓ'} (f : CommAlgebraHom A B) - → compCommAlgebraHom _ _ _ (idCommAlgebraHom A) f ≡ f - compIdCommAlgebraHom = compIdAlgebraHom - - idCompCommAlgebraHom : {A B : CommAlgebra R ℓ'} (f : CommAlgebraHom A B) - → compCommAlgebraHom _ _ _ f (idCommAlgebraHom B) ≡ f - idCompCommAlgebraHom = idCompAlgebraHom - - compAssocCommAlgebraHom : {A B C D : CommAlgebra R ℓ'} - (f : CommAlgebraHom A B) (g : CommAlgebraHom B C) (h : CommAlgebraHom C D) - → compCommAlgebraHom _ _ _ (compCommAlgebraHom _ _ _ f g) h - ≡ compCommAlgebraHom _ _ _ f (compCommAlgebraHom _ _ _ g h) - 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} - - --- 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 - 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) - 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 - - 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))) - - uaCompCommAlgebraEquiv : {A B C : CommAlgebra R ℓ'} (f : CommAlgebraEquiv A B) (g : CommAlgebraEquiv B C) - → uaCommAlgebra (compCommAlgebraEquiv f g) ≡ uaCommAlgebra f ∙ uaCommAlgebra g - uaCompCommAlgebraEquiv f g = caracCommAlgebra≡ _ _ ( - cong ⟨_⟩ (uaCommAlgebra (compCommAlgebraEquiv f g)) - ≡⟨ uaCompEquiv _ _ ⟩ - cong ⟨_⟩ (uaCommAlgebra f) ∙ cong ⟨_⟩ (uaCommAlgebra g) - ≡⟨ sym (cong-∙ ⟨_⟩ (uaCommAlgebra f) (uaCommAlgebra g)) ⟩ - cong ⟨_⟩ (uaCommAlgebra f ∙ uaCommAlgebra g) ∎) - - -open CommAlgebraHoms -open CommAlgebraEquivs -open CommAlgebraUAFunctoriality -recPT→CommAlgebra : {R : CommRing ℓ} {A : Type ℓ'} (𝓕 : A → CommAlgebra R ℓ'') - → (σ : ∀ x y → CommAlgebraEquiv (𝓕 x) (𝓕 y)) - → (∀ x y z → σ x z ≡ compCommAlgebraEquiv (σ x y) (σ y z)) - ------------------------------------------------------ - → ∥ A ∥₁ → CommAlgebra R ℓ'' -recPT→CommAlgebra 𝓕 σ compCoh = GpdElim.rec→Gpd isGroupoidCommAlgebra 𝓕 - (3-ConstantCompChar 𝓕 (λ x y → uaCommAlgebra (σ x y)) - λ x y z → sym ( cong uaCommAlgebra (compCoh x y z) - ∙ uaCompCommAlgebraEquiv (σ x y) (σ y z))) - - -contrCommAlgebraHom→contrCommAlgebraEquiv : {R : CommRing ℓ} {A : Type ℓ'} - (σ : A → CommAlgebra R ℓ'') - → (∀ x y → isContr (CommAlgebraHom (σ x) (σ y))) - ---------------------------------------------------------------------------- - → ∀ x y → isContr (CommAlgebraEquiv (σ x) (σ y)) -contrCommAlgebraHom→contrCommAlgebraEquiv σ contrHom x y = σEquiv , - λ e → Σ≡Prop (λ _ → isPropIsAlgebraHom _ _ _ _) - (Σ≡Prop isPropIsEquiv (cong fst (contrHom _ _ .snd (CommAlgebraEquiv→CommAlgebraHom e)))) +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 + + invCommAlgebraEquiv : {A B : CommAlgebra R ℓ'} + → CommAlgebraEquiv A B + → CommAlgebraEquiv B A + 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 : (f⁻¹ .fst .fst) ∘ B .snd .fst ≡ A .snd .fst + commutes = + 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 Iso - χ₁ = contrHom x y .fst - χ₂ = contrHom y x .fst - χ₁∘χ₂≡id : χ₁ ∘ca χ₂ ≡ idCommAlgebraHom _ - χ₁∘χ₂≡id = isContr→isProp (contrHom _ _) _ _ - - χ₂∘χ₁≡id : χ₂ ∘ca χ₁ ≡ idCommAlgebraHom _ - χ₂∘χ₁≡id = isContr→isProp (contrHom _ _) _ _ - - σIso : Iso ⟨ σ x ⟩ ⟨ σ y ⟩ - fun σIso = fst χ₁ - inv σIso = fst χ₂ - rightInv σIso = funExt⁻ (cong fst χ₁∘χ₂≡id) - leftInv σIso = funExt⁻ (cong fst χ₂∘χ₁≡id) - - σEquiv : CommAlgebraEquiv (σ x) (σ y) - fst σEquiv = isoToEquiv σIso - snd σEquiv = snd χ₁ -CommAlgebra→Ring : {R : CommRing ℓ} → CommAlgebra R ℓ' → Ring ℓ' -CommAlgebra→Ring = CommRing→Ring ∘ CommAlgebra→CommRing - -module _ {R : CommRing ℓ} {A B : CommAlgebra R ℓ'} where open CommAlgebraStr ⦃...⦄ - instance - _ = snd A - _ = snd B - open IsAlgebraHom - - CommAlgebraHom→RingHom : CommAlgebraHom A B → RingHom (CommAlgebra→Ring A) (CommAlgebra→Ring B) - fst (CommAlgebraHom→RingHom ϕ) = fst ϕ - IsRingHom.pres0 (snd (CommAlgebraHom→RingHom ϕ)) = pres0 (snd ϕ) - IsRingHom.pres1 (snd (CommAlgebraHom→RingHom ϕ)) = pres1 (snd ϕ) - IsRingHom.pres+ (snd (CommAlgebraHom→RingHom ϕ)) = pres+ (snd ϕ) - IsRingHom.pres· (snd (CommAlgebraHom→RingHom ϕ)) = pres· (snd ϕ) - IsRingHom.pres- (snd (CommAlgebraHom→RingHom ϕ)) = pres- (snd ϕ) - - CommAlgebraHomFromRingHom : - (ϕ : RingHom (CommAlgebra→Ring A) (CommAlgebra→Ring B)) - → ((r : fst R) (x : fst A) → (fst ϕ) (r ⋆ x) ≡ r ⋆ (fst ϕ x)) - → CommAlgebraHom A B - fst (CommAlgebraHomFromRingHom ϕ pres*) = fst ϕ - pres0 (snd (CommAlgebraHomFromRingHom ϕ pres*)) = IsRingHom.pres0 (snd ϕ) - pres1 (snd (CommAlgebraHomFromRingHom ϕ pres*)) = IsRingHom.pres1 (snd ϕ) - pres+ (snd (CommAlgebraHomFromRingHom ϕ pres*)) = IsRingHom.pres+ (snd ϕ) - pres· (snd (CommAlgebraHomFromRingHom ϕ pres*)) = IsRingHom.pres· (snd ϕ) - 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)) + 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 - 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) - 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 = 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) + 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 ∎ diff --git a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda index 0419f097ab..70cff09a32 100644 --- a/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda +++ b/Cubical/Algebra/CommAlgebra/QuotientAlgebra.agda @@ -1,8 +1,9 @@ -{-# OPTIONS --safe #-} +{-# OPTIONS --safe --lossy-unification #-} module Cubical.Algebra.CommAlgebra.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,112 +14,46 @@ 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.CommAlgebra +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.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 + ℓ ℓ' ℓ'' : 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 A) where - open CommRingStr {{...}} hiding (_-_; -_; ·IdL ; ·DistR+) renaming (_·_ to _·R_; _+_ to _+R_) - open CommAlgebraStr {{...}} - open RingTheory (CommRing→Ring (CommAlgebra→CommRing A)) using (-DistR·) +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ') (I : IdealsIn R A) where + open CommRingStr ⦃...⦄ + open CommAlgebraStr ⦃...⦄ + open RingTheory (CommRing→Ring (fst A)) using (-DistR·) instance _ : CommRingStr ⟨ R ⟩ _ = snd R - _ : CommAlgebraStr R ⟨ A ⟩ - _ = snd A + _ : CommRingStr ⟨ A .fst ⟩ + _ = A .fst .snd + _ = A - opaque - _/_ : CommAlgebra R ℓ - _/_ = 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 : fst A) → x - y ∈ (fst I) → [ r ⋆ x ]/ ≡ [ r ⋆ y ]/ - eq r x y x-y∈I = eq/ _ _ - (subst (λ u → u ∈ fst I) - ((r ⋆ 1a) · (x - y) ≡⟨ ·DistR+ (r ⋆ 1a) x (- y) ⟩ - (r ⋆ 1a) · x + (r ⋆ 1a) · (- y) ≡[ i ]⟨ (r ⋆ 1a) · x + -DistR· (r ⋆ 1a) y i ⟩ - (r ⋆ 1a) · x - (r ⋆ 1a) · y ≡[ i ]⟨ ⋆AssocL r 1a x i - - ⋆AssocL r 1a y i ⟩ - r ⋆ (1a · x) - r ⋆ (1a · y) ≡[ i ]⟨ r ⋆ (·IdL x i) - r ⋆ (·IdL y i) ⟩ - r ⋆ x - r ⋆ y ∎ ) - (isCommIdeal.·Closed (snd I) _ x-y∈I)) + _/_ : CommAlgebra R ℓ' + _/_ = ((fst A) CommRing./ I) , + (withOpaqueStr $ (CommRing.quotientHom (fst A) I ∘cr A .snd)) - 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_) + quotientHom : CommAlgebraHom A (_/_) + quotientHom = CommRingHom→CommAlgebraHom (CommRing.quotientHom (fst A) I) $ CommRingHom≡ refl + +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 - - 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 + _ : CommRingStr ⟨ A ⟩ₐ + _ = A .fst .snd + _ = A module _ (B : CommAlgebra R ℓ) @@ -126,94 +61,94 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn A) where (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 = + 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 - 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 - unfolding inducedHom quotientHom + isUnique : (ψ : CommAlgebraHom (A / I) B) (ψIsSolution : ⟨ ψ ⟩ₐ→ ∘ ⟨ quotientHom A I ⟩ₐ→ ≡ ⟨ ϕ ⟩ₐ→) + → ψ ≡ inducedHom + isUnique ψ ψIsSolution = + CommAlgebraHom≡ + (cong fst $ + CommRing.UniversalProperty.isUnique + (A .fst) (B .fst) I ⟨ ϕ ⟩ᵣ→ I⊆kernel ⟨ ψ ⟩ᵣ→ ψIsSolution) - inducedHom∘quotientHom : inducedHom ∘a quotientHom A I ≡ ϕ - inducedHom∘quotientHom = Σ≡Prop (isPropIsCommAlgebraHom {M = A} {N = B}) (funExt (λ a → refl)) + + +module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ') (I : IdealsIn R A) where + open CommRingStr {{...}} + open CommAlgebraStr {{...}} 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 + quotientHomEpi : (S : hSet ℓ'') + → (f g : ⟨ A / I ⟩ₐ → ⟨ S ⟩) + → f ∘ ⟨ quotientHom A I ⟩ₐ→ ≡ g ∘ ⟨ quotientHom A I ⟩ₐ→ + → f ≡ g + quotientHomEpi = CommRing.quotientHomEpi (fst A) I {- 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 = + (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))) - - 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))) + (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 = + 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 .snd = + makeOpaque $ 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 @@ -221,8 +156,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 @@ -230,18 +163,16 @@ module _ {R : CommRing ℓ} (A : CommAlgebra R ℓ) (I : IdealsIn A) where refl ⟩ _ ≡⟨ 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) + 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/CommAlgebra/Univalence.agda b/Cubical/Algebra/CommAlgebra/Univalence.agda new file mode 100644 index 0000000000..4af65b17f9 --- /dev/null +++ b/Cubical/Algebra/CommAlgebra/Univalence.agda @@ -0,0 +1,100 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommAlgebra.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.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.CommAlgebra.Base + +private + variable + ℓ ℓ' ℓ'' : Level + + +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 _ _ _ _) diff --git a/Cubical/Algebra/CommRing/Base.agda b/Cubical/Algebra/CommRing/Base.agda index 36f44a8001..77249feec7 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 @@ -91,10 +92,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) @@ -109,12 +121,29 @@ 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 + no-eta-equality + 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) + +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) @@ -131,103 +160,95 @@ isPropIsCommRing 0r 1r _+_ _·_ -_ = 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 +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 - 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)) + 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 + +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 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 + +opaque + CommRingHom≡ : {R : CommRing ℓ} {S : CommRing ℓ'} {φ ψ : CommRingHom R S} → fst φ ≡ fst ψ → φ ≡ ψ + CommRingHom≡ = Σ≡Prop λ f → isPropIsCommRingHom _ f _ - 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 + 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/FiberedProduct.agda b/Cubical/Algebra/CommRing/FiberedProduct.agda index ba1b1cb97f..1f1eb0519b 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,31 +69,31 @@ 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 : 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 (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/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 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/Instances/Polynomials/Typevariate.agda b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate.agda new file mode 100644 index 0000000000..949fa66a93 --- /dev/null +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate.agda @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} + +module Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate where + +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..3f06b0e2c8 --- /dev/null +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Base.agda @@ -0,0 +1,95 @@ +{-# 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]) + * the homomorphism R -> R[I], which turns this CommRing into a CommAlgebra + For more, see the Properties file. +-} +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure using (⟨_⟩) + +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 + +opaque + 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..7ce31c7d22 --- /dev/null +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/OnCoproduct.agda @@ -0,0 +1,106 @@ +{-# 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.Base +open import Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.UniversalProperty + +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 (R [ I ⊎ J ])) + 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 ∎) + + {- 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 fromAlgebraHom ⟩ + constPolynomial (R [ I ]) J .fst ∎) + (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 new file mode 100644 index 0000000000..99a33504ec --- /dev/null +++ b/Cubical/Algebra/CommRing/Instances/Polynomials/Typevariate/Properties.agda @@ -0,0 +1,54 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.CommRing.Instances.Polynomials.Typevariate.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +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.CommRing.Instances.Polynomials.Typevariate.UniversalProperty + +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Sigma + +open import Cubical.Tactics.CommRingSolver + +private + variable + ℓ ℓ' ℓ'' : Level + +module _ {R : CommRing ℓ} where + open CommRingStr ⦃...⦄ + private instance + _ = snd R + + 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 + + from : CommRingHom R (R [ ⊥ ]) + from = constPolynomial R ⊥ + + to∘from : to ∘cr from ≡ idCommRingHom R + to∘from = CommRingHom≡ refl + + from∘to : from ∘cr to ≡ idCommRingHom (R [ ⊥ ]) + from∘to = idByIdOnVars (from ∘cr to) refl (funExt ⊥.elim) + + 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 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) 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/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/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/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 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/TypevariateHIT/Base.agda b/Cubical/Algebra/Polynomials/TypevariateHIT/Base.agda index 44a2232770..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.CommAlgebra.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 ac9f570c94..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.CommAlgebra +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/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 fc79661799..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 = makeIsRingHom 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/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/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)) diff --git a/Cubical/Algebra/Polynomials/UnivariateList/UniversalProperty.agda b/Cubical/Algebra/Polynomials/UnivariateList/UniversalProperty.agda index a4167cfcfd..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.CommAlgebra.UnivariatePolyList public +open import Cubical.Algebra.CommAlgebra.AsModule.UnivariatePolyList public 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 diff --git a/Cubical/Algebra/Ring/Quotient.agda b/Cubical/Algebra/Ring/Quotient.agda index a8ce89ad88..f98e86455f 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= @@ -12,8 +13,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 +56,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 +72,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 +99,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 +122,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 +135,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' @@ -184,28 +185,59 @@ 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 + 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 = + record + { isRing = isRingQuotient; RingStr quotientRingStr' } + _/_ : (R : Ring ℓ) → (I : IdealsIn R) → Ring ℓ -R / (I , IisIdeal) = asRing R I IisIdeal +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 + 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) = [_] -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 +snd (quotientHom R I) = isRingHomQuotientHom R I + +module _ (R : Ring ℓ) (I : IdealsIn R) where + quotientHomSurjective : isSurjection (fst (quotientHom R I)) + quotientHomSurjective = []surjective + + open RingHoms + 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 = + (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 @@ -230,9 +262,9 @@ 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 + 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₂) @@ -240,23 +272,39 @@ module UniversalProperty (R : Ring ℓ) (I : IdealsIn R) where 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 + 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 + 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 = + 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) + → ψ ≡ inducedHom p + unique' p ψ ψIsSolution = unique p ψ (RingHom≡ ψIsSolution) {- Show that the kernel of the quotient map @@ -276,8 +324,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 isRingHomQuotientHom + 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,13 +360,16 @@ 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 - -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 + unfolding isRingHomQuotientHom + 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 + 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 diff --git a/Cubical/AlgebraicGeometry/Functorial/ZFunctors/Base.agda b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/Base.agda index 8524743648..24a60c7adc 100644 --- a/Cubical/AlgebraicGeometry/Functorial/ZFunctors/Base.agda +++ b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/Base.agda @@ -58,7 +58,7 @@ module _ {ℓ : Level} where open Functor open NatTrans open CommRingStr ⦃...⦄ - open IsRingHom + open IsCommRingHom -- using the naming conventions of Demazure & Gabriel @@ -159,8 +159,8 @@ module _ {ℓ : Level} where pres- (snd (F-hom 𝓞 α)) _ = makeNatTransPath (funExt₂ λ _ _ → refl) -- functoriality of 𝓞 - F-id 𝓞 = RingHom≡ (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) - F-seq 𝓞 _ _ = RingHom≡ (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) + F-id 𝓞 = CommRingHom≡ (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) + F-seq 𝓞 _ _ = CommRingHom≡ (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) @@ -174,7 +174,7 @@ module AdjBij {ℓ : Level} where open Functor open NatTrans open Iso - open IsRingHom + open IsCommRingHom private module _ {A : CommRing ℓ} {X : ℤFunctor {ℓ}} where _♭ : CommRingHom A (𝓞 .F-ob X) → X ⇒ Sp .F-ob A @@ -186,7 +186,7 @@ module AdjBij {ℓ : Level} where pres· (snd (N-ob (φ ♭) B x)) _ _ = cong (λ y → y .N-ob B x) (φ .snd .pres· _ _) pres- (snd (N-ob (φ ♭) B x)) _ = cong (λ y → y .N-ob B x) (φ .snd .pres- _) - N-hom (φ ♭) ψ = funExt (λ x → RingHom≡ (funExt λ a → funExt⁻ (φ .fst a .N-hom ψ) x)) + N-hom (φ ♭) ψ = funExt (λ x → CommRingHom≡ (funExt λ a → funExt⁻ (φ .fst a .N-hom ψ) x)) -- the other direction is just precomposition modulo Yoneda @@ -202,10 +202,10 @@ module AdjBij {ℓ : Level} where -- the two maps are inverse to each other ♭♯Id : ∀ (α : X ⇒ Sp .F-ob A) → ((α ♯) ♭) ≡ α - ♭♯Id _ = makeNatTransPath (funExt₂ λ _ _ → RingHom≡ (funExt (λ _ → refl))) + ♭♯Id _ = makeNatTransPath (funExt₂ λ _ _ → CommRingHom≡ (funExt (λ _ → refl))) ♯♭Id : ∀ (φ : CommRingHom A (𝓞 .F-ob X)) → ((φ ♭) ♯) ≡ φ - ♯♭Id _ = RingHom≡ (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) + ♯♭Id _ = CommRingHom≡ (funExt λ _ → makeNatTransPath (funExt₂ λ _ _ → refl)) -- we get a relative adjunction 𝓞 ⊣ᵢ Sp @@ -219,12 +219,12 @@ module AdjBij {ℓ : Level} where 𝓞⊣SpNatℤFunctor : {A : CommRing ℓ} {X Y : ℤFunctor {ℓ}} (α : X ⇒ Sp .F-ob A) (β : Y ⇒ X) → (β ●ᵛ α) ♯ ≡ (𝓞 .F-hom β) ∘cr (α ♯) - 𝓞⊣SpNatℤFunctor _ _ = RingHom≡ (funExt (λ _ → makeNatTransPath (funExt₂ (λ _ _ → refl)))) + 𝓞⊣SpNatℤFunctor _ _ = CommRingHom≡ (funExt (λ _ → makeNatTransPath (funExt₂ (λ _ _ → refl)))) 𝓞⊣SpNatCommRing : {X : ℤFunctor {ℓ}} {A B : CommRing ℓ} (φ : CommRingHom A (𝓞 .F-ob X)) (ψ : CommRingHom B A) → (φ ∘cr ψ) ♭ ≡ (φ ♭) ●ᵛ Sp .F-hom ψ - 𝓞⊣SpNatCommRing _ _ = makeNatTransPath (funExt₂ λ _ _ → RingHom≡ (funExt (λ _ → refl))) + 𝓞⊣SpNatCommRing _ _ = makeNatTransPath (funExt₂ λ _ _ → CommRingHom≡ (funExt (λ _ → refl))) -- the counit is an equivalence private diff --git a/Cubical/AlgebraicGeometry/Functorial/ZFunctors/CompactOpen.agda b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/CompactOpen.agda index d24abaf8ba..a8af5d7436 100644 --- a/Cubical/AlgebraicGeometry/Functorial/ZFunctors/CompactOpen.agda +++ b/Cubical/AlgebraicGeometry/Functorial/ZFunctors/CompactOpen.agda @@ -74,7 +74,7 @@ module _ {ℓ : Level} where open NatIso open DistLatticeStr ⦃...⦄ open CommRingStr ⦃...⦄ - open IsRingHom + open IsCommRingHom open IsLatticeHom open ZarLat open ZarLatUniversalProp 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/StructureSheaf.agda b/Cubical/AlgebraicGeometry/ZariskiLattice/StructureSheaf.agda index 0ca7597380..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.CommAlgebra.Base -open import Cubical.Algebra.CommAlgebra.Properties -open import Cubical.Algebra.CommAlgebra.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 @@ -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