|
| 1 | +------------------------------------------------------------------------ |
| 2 | +-- The Agda standard library |
| 3 | +-- |
| 4 | +-- Quotient rings |
| 5 | +------------------------------------------------------------------------ |
| 6 | + |
| 7 | +{-# OPTIONS --safe --cubical-compatible #-} |
| 8 | + |
| 9 | +open import Algebra.Bundles using (Ring; RawRing) |
| 10 | +open import Algebra.Ideal using (Ideal) |
| 11 | + |
| 12 | +module Algebra.Construct.Quotient.Ring {c ℓ} (R : Ring c ℓ) {c′ ℓ′} (I : Ideal R c′ ℓ′) where |
| 13 | + |
| 14 | +open Ring R |
| 15 | +private module I = Ideal I |
| 16 | +open I using (ι; normalSubgroup) |
| 17 | + |
| 18 | +open import Algebra.Construct.Quotient.Group +-group normalSubgroup public |
| 19 | + using (_≋_; _by_; ≋-refl; ≋-sym; ≋-trans; ≋-isEquivalence; ≈⇒≋; quotientIsGroup; quotientGroup; project; project-surjective) |
| 20 | + renaming (≋-∙-cong to ≋-+-cong; ≋-⁻¹-cong to ≋‿-‿cong) |
| 21 | + |
| 22 | +open import Algebra.Definitions _≋_ |
| 23 | +open import Algebra.Morphism.Structures using (IsRingHomomorphism) |
| 24 | +open import Algebra.Properties.Semiring semiring |
| 25 | +open import Algebra.Properties.Ring R |
| 26 | +open import Algebra.Structures |
| 27 | +open import Function.Definitions using (Surjective) |
| 28 | +open import Level |
| 29 | +open import Relation.Binary.Reasoning.Setoid setoid |
| 30 | + |
| 31 | +≋-*-cong : Congruent₂ _*_ |
| 32 | +≋-*-cong {x} {y} {u} {v} (j by ιj+x≈y) (k by ιk+u≈v) = ι j I.*ₗ k I.+ᴹ j I.*ᵣ u I.+ᴹ x I.*ₗ k by begin |
| 33 | + ι (ι j I.*ₗ k I.+ᴹ j I.*ᵣ u I.+ᴹ x I.*ₗ k) + x * u ≈⟨ +-congʳ (ι.+ᴹ-homo (ι j I.*ₗ k I.+ᴹ j I.*ᵣ u) (x I.*ₗ k)) ⟩ |
| 34 | + ι (ι j I.*ₗ k I.+ᴹ j I.*ᵣ u) + ι (x I.*ₗ k) + x * u ≈⟨ +-congʳ (+-congʳ (ι.+ᴹ-homo (ι j I.*ₗ k) (j I.*ᵣ u))) ⟩ |
| 35 | + ι (ι j I.*ₗ k) + ι (j I.*ᵣ u) + ι (x I.*ₗ k) + x * u ≈⟨ +-congʳ (+-cong (+-cong (ι.*ₗ-homo (ι j) k) (ι.*ᵣ-homo u j)) (ι.*ₗ-homo x k)) ⟩ |
| 36 | + ι j * ι k + ι j * u + x * ι k + x * u ≈⟨ binomial-expansion (ι j) x (ι k) u ⟨ |
| 37 | + (ι j + x) * (ι k + u) ≈⟨ *-cong ιj+x≈y ιk+u≈v ⟩ |
| 38 | + y * v ∎ |
| 39 | + |
| 40 | +quotientRawRing : RawRing c (c ⊔ ℓ ⊔ c′) |
| 41 | +quotientRawRing = record |
| 42 | + { Carrier = Carrier |
| 43 | + ; _≈_ = _≋_ |
| 44 | + ; _+_ = _+_ |
| 45 | + ; _*_ = _*_ |
| 46 | + ; -_ = -_ |
| 47 | + ; 0# = 0# |
| 48 | + ; 1# = 1# |
| 49 | + } |
| 50 | + |
| 51 | +quotientIsRing : IsRing _≋_ _+_ _*_ (-_) 0# 1# |
| 52 | +quotientIsRing = record |
| 53 | + { +-isAbelianGroup = record |
| 54 | + { isGroup = quotientIsGroup |
| 55 | + ; comm = λ x y → ≈⇒≋ (+-comm x y) |
| 56 | + } |
| 57 | + ; *-cong = ≋-*-cong |
| 58 | + ; *-assoc = λ x y z → ≈⇒≋ (*-assoc x y z) |
| 59 | + ; *-identity = record |
| 60 | + { fst = λ x → ≈⇒≋ (*-identityˡ x) |
| 61 | + ; snd = λ x → ≈⇒≋ (*-identityʳ x) |
| 62 | + } |
| 63 | + ; distrib = record |
| 64 | + { fst = λ x y z → ≈⇒≋ (distribˡ x y z) |
| 65 | + ; snd = λ x y z → ≈⇒≋ (distribʳ x y z) |
| 66 | + } |
| 67 | + } |
| 68 | + |
| 69 | +quotientRing : Ring c (c ⊔ ℓ ⊔ c′) |
| 70 | +quotientRing = record { isRing = quotientIsRing } |
| 71 | + |
| 72 | +project-isHomomorphism : IsRingHomomorphism rawRing quotientRawRing project |
| 73 | +project-isHomomorphism = record |
| 74 | + { isSemiringHomomorphism = record |
| 75 | + { isNearSemiringHomomorphism = record |
| 76 | + { +-isMonoidHomomorphism = record |
| 77 | + { isMagmaHomomorphism = record |
| 78 | + { isRelHomomorphism = record |
| 79 | + { cong = ≈⇒≋ |
| 80 | + } |
| 81 | + ; homo = λ _ _ → ≋-refl |
| 82 | + } |
| 83 | + ; ε-homo = ≋-refl |
| 84 | + } |
| 85 | + ; *-homo = λ _ _ → ≋-refl |
| 86 | + } |
| 87 | + ; 1#-homo = ≋-refl |
| 88 | + } |
| 89 | + ; -‿homo = λ _ → ≋-refl |
| 90 | + } |
0 commit comments