Skip to content
Draft
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
4 changes: 4 additions & 0 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -81,8 +81,12 @@ New modules

* `Algebra.Construct.Quotient.Group` for the definition of quotient groups.

* `Algebra.Construct.Quotient.Ring` for the definition of quotient rings.

* `Algebra.Construct.Sub.Group` for the definition of subgroups.

* `Algebra.Ideal` for the definition of (two sided) ideals of rings.

* `Algebra.Module.Construct.Sub.Bimodule` for the definition of subbimodules.

* `Algebra.NormalSubgroup` for the definition of normal subgroups.
Expand Down
89 changes: 89 additions & 0 deletions src/Algebra/Construct/Quotient/Ring.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,89 @@
------------------------------------------------------------------------
-- The Agda standard library
--
-- Quotient rings
------------------------------------------------------------------------

{-# OPTIONS --safe --cubical-compatible #-}

open import Algebra.Bundles using (Ring; RawRing)
open import Algebra.Ideal using (Ideal)

module Algebra.Construct.Quotient.Ring {c ℓ} (R : Ring c ℓ) {c′ ℓ′} (I : Ideal R c′ ℓ′) where

open import Algebra.Morphism.Structures using (IsRingHomomorphism)
open import Algebra.Properties.Ring R
open import Algebra.Structures
open import Level

open Ring R
private module I = Ideal I
open I using (ι; normalSubgroup)
Comment on lines +20 to +21
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

You can, somewhat counterintuitively, simplify this to

Suggested change
private module I = Ideal I
open I using (ι; normalSubgroup)
private open module I = Ideal I using (ι; normalSubgroup)


open import Algebra.Construct.Quotient.Group +-group normalSubgroup public
using (_≋_; _by_; ≋-refl; ≋-sym; ≋-trans; ≋-isEquivalence; ≈⇒≋; quotientIsGroup; quotientGroup; π; π-surjective)
renaming (≋-∙-cong to ≋-+-cong; ≋-⁻¹-cong to ≋‿-‿cong)
open import Algebra.Definitions _≋_
open import Algebra.Properties.Semiring semiring
open import Relation.Binary.Reasoning.Setoid setoid
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Ditto. should this be part of the API exported by Algebra.Properties.X? Do we ever invoke algebraic properties not in a context where we have access to equational reasoning wrt that algebra?


≋-*-cong : Congruent₂ _*_
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I would be tempted to use a local module (I forget the syntax for doing just inside the proof part) where I is opened so that all these qualifiers can go away. It's an idiom used a lot in 1lab that I need to integrate into my own Agda.

≋-*-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
ι (ι 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)) ⟩
ι (ι j I.*ₗ k I.+ᴹ j I.*ᵣ u) + ι (x I.*ₗ k) + x * u ≈⟨ +-congʳ (+-congʳ (ι.+ᴹ-homo (ι j I.*ₗ k) (j I.*ᵣ u))) ⟩
ι (ι j I.*ₗ k) + ι (j I.*ᵣ u) + ι (x I.*ₗ k) + x * u ≈⟨ +-congʳ (+-cong (+-cong (ι.*ₗ-homo (ι j) k) (ι.*ᵣ-homo u j)) (ι.*ₗ-homo x k)) ⟩
ι j * ι k + ι j * u + x * ι k + x * u ≈⟨ binomial-expansion (ι j) x (ι k) u ⟨
(ι j + x) * (ι k + u) ≈⟨ *-cong ιj+x≈y ιk+u≈v ⟩
y * v ∎

quotientRawRing : RawRing c (c ⊔ ℓ ⊔ c′)
quotientRawRing = record
{ Carrier = Carrier
; _≈_ = _≋_
; _+_ = _+_
; _*_ = _*_
; -_ = -_
; 0# = 0#
; 1# = 1#
}
Comment on lines +41 to +48
Copy link
Contributor

@jamesmckinna jamesmckinna Nov 5, 2025

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Suggested change
{ Carrier = Carrier
; _≈_ = _≋_
; _+_ = _+_
; _*_ = _*_
; -_ = -_
; 0# = 0#
; 1# = 1#
}
{ RawRing R.rawRing hiding (_≈_)
; _≈_ = _≋_
}

avoids redundant DRY ...?


quotientIsRing : IsRing _≋_ _+_ _*_ (-_) 0# 1#
quotientIsRing = record
{ +-isAbelianGroup = record
{ isGroup = quotientIsGroup
; comm = λ x y → ≈⇒≋ (+-comm x y)
}
Comment on lines +53 to +55
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Don't need to create, nor import, quotientIsGroup here; can be recreated on the fly:

Suggested change
{ isGroup = quotientIsGroup
; comm = λ x y ≈⇒≋ (+-comm x y)
}
{ isGroup = isGroup
; comm = λ x y ≈⇒≋ (+-comm x y)
} where open Group quotientGroup using (isGroup)

; *-cong = ≋-*-cong
; *-assoc = λ x y z → ≈⇒≋ (*-assoc x y z)
; *-identity = record
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

as this is 'just' a pair, use _,_ instead of a record? And maybe also eta contract?

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I personally find lambdas on the left hand side of an operator (such as _,_) really ugly, so I use the record syntax to avoid it where I can...

{ fst = λ x → ≈⇒≋ (*-identityˡ x)
; snd = λ x → ≈⇒≋ (*-identityʳ x)
}
; distrib = record
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

_,_ would work here too (but not eta)

{ fst = λ x y z → ≈⇒≋ (distribˡ x y z)
; snd = λ x y z → ≈⇒≋ (distribʳ x y z)
}
}

quotientRing : Ring c (c ⊔ ℓ ⊔ c′)
quotientRing = record { isRing = quotientIsRing }
Comment on lines +68 to +69
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Similarly, can we inline the definition of quotientIsRing here?


π-isHomomorphism : IsRingHomomorphism rawRing quotientRawRing π
π-isHomomorphism = record
{ isSemiringHomomorphism = record
{ isNearSemiringHomomorphism = record
{ +-isMonoidHomomorphism = record
{ isMagmaHomomorphism = record
{ isRelHomomorphism = record
{ cong = ≈⇒≋
}
; homo = λ _ _ → ≋-refl
}
; ε-homo = ≋-refl
}
; *-homo = λ _ _ → ≋-refl
}
; 1#-homo = ≋-refl
}
; -‿homo = λ _ → ≋-refl
}
33 changes: 33 additions & 0 deletions src/Algebra/Ideal.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,33 @@
------------------------------------------------------------------------
-- The Agda standard library
--
-- Ideals of a ring
------------------------------------------------------------------------

{-# OPTIONS --safe --cubical-compatible #-}

open import Algebra.Bundles using (Ring)

module Algebra.Ideal {c ℓ} (R : Ring c ℓ) where

open Ring R using (+-group; +-comm)

open import Algebra.Module.Construct.Sub.Bimodule using (Subbimodule)
import Algebra.Module.Construct.TensorUnit as TU
open import Algebra.NormalSubgroup (+-group)
open import Level

------------------------------------------------------------------------
-- Definition

-- a (two sided) ideal is exactly a subbimodule of R considered as a bimodule over itself
record Ideal c′ ℓ′ : Set (c ⊔ ℓ ⊔ suc (c′ ⊔ ℓ′)) where
field
subbimodule : Subbimodule {R = R} TU.bimodule c′ ℓ′

open Subbimodule subbimodule public

normalSubgroup : NormalSubgroup c′ ℓ′
normalSubgroup = record
{ isNormal = abelian⇒subgroup-normal +-comm subgroup
}