Skip to content
4 changes: 4 additions & 0 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -79,10 +79,14 @@ Deprecated names
New modules
-----------

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

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

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

* `Algebra.NormalSubgroup` for the definition of normal subgroups.

* `Algebra.Properties.BooleanRing`.

* `Algebra.Properties.BooleanSemiring`.
Expand Down
124 changes: 124 additions & 0 deletions src/Algebra/Construct/Quotient/Group.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,124 @@
------------------------------------------------------------------------
-- The Agda standard library
--
-- Quotient groups
------------------------------------------------------------------------

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

open import Algebra.Bundles using (Group)
open import Algebra.NormalSubgroup using (NormalSubgroup)

module Algebra.Construct.Quotient.Group {c ℓ} (G : Group c ℓ) {c′ ℓ′} (N : NormalSubgroup G c′ ℓ′) where

open Group G

import Algebra.Definitions as AlgDefs
open import Algebra.Morphism.Structures using (IsGroupHomomorphism)
open import Algebra.Properties.Monoid monoid
open import Algebra.Properties.Group G using (⁻¹-anti-homo-∙)
open import Algebra.Structures using (IsGroup)
Copy link
Contributor

Choose a reason for hiding this comment

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

Do we need this import? Cf. #2391 and see below.

open import Data.Product.Base using (_,_)
open import Function.Base using (_∘_)
open import Function.Definitions using (Surjective)
open import Level using (_⊔_)
open import Relation.Binary.Core using (_⇒_)
open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive)
open import Relation.Binary.Structures using (IsEquivalence)
open import Relation.Binary.Reasoning.Setoid setoid

private
module N = NormalSubgroup N
open NormalSubgroup N using (ι; module ι; conjugate; normal)
Comment on lines +30 to +32
Copy link
Contributor

Choose a reason for hiding this comment

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

Can simplify this to:

Suggested change
private
module N = NormalSubgroup N
open NormalSubgroup N using (ι; module ι; conjugate; normal)
private
open module N = NormalSubgroup N using (ι; module ι; conjugate; normal)


infix 0 _by_

data _≋_ (x y : Carrier) : Set (c ⊔ ℓ ⊔ c′) where
_by_ : ∀ g → ι g ∙ x ≈ y → x ≋ y
Copy link
Member Author

Choose a reason for hiding this comment

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

If this is instead _by_ : ∀ g → x ∙ ι g ≈ y → x ≋ y some things line up nicer when we get to integers (in particular it matches Data.Integer.DivMod.a≡a%n+[a/n]*n)

Copy link
Contributor

Choose a reason for hiding this comment

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

Well, so be it, but be careful of having a tail wag the dog. It wouldn't be the first time that existing stdlib modules expose argument order/definitional inconsistency when exposed to new additions/refactoriings (the saga of Algebra.Properties.*.Divisibility and Data.Nat.DivMod.* being a case in point)...

Copy link
Member Author

Choose a reason for hiding this comment

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

The alternative here would be putting more lemmas in Data.Integer.DivMod, which I want to do anyway

Copy link
Contributor

@jamesmckinna jamesmckinna Nov 22, 2025

Choose a reason for hiding this comment

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

Well, after my recent flurry of activity around Central/Centre.* #2872 #2873 #2885 I've been turning over-and-over this question of what is the mathematically/formalistically 'best' account of the equivalence relation on the Carrier such that id is the function underlying the quotient map, and such that the identification of the Kernel of the quotient homomorphism with the Normal subgroup which gives rise to it.

Envisaging Kernels #2871 #1966 being designed/defined on top of the recent addition Relation.Binary.Morphism.Construct.On in #2872 , the conclusion I come to is that the symmetric form of the by constructor would be more streamlined...

... and avoid having to commit you to one or other 'handed' choice...

  • the carrier of (the kernel of a hom h onto) the quotient becomes triples (m,x,n) st h(m)x = xh(n),
  • this is obviously Normal... the conjugate swaps m and n (up to h...)

conversely, given a Normal subgroup, we obtain an element of the kernel as (m, x, conjugate(m))...

etc. details to work out, but I think overall the symmetric choice for the by constructor would work more smoothly.


≈⇒≋ : _≈_ ⇒ _≋_
≈⇒≋ {x} {y} x≈y = N.ε by trans (∙-cong ι.ε-homo x≈y) (identityˡ y)

≋-refl : Reflexive _≋_
≋-refl = ≈⇒≋ refl

≋-sym : Symmetric _≋_
≋-sym {x} {y} (g by ιg∙x≈y) = g N.⁻¹ by begin
ι (g N.⁻¹) ∙ y ≈⟨ ∙-cong (ι.⁻¹-homo g) (sym ιg∙x≈y) ⟩
ι g ⁻¹ ∙ (ι g ∙ x) ≈⟨ cancelˡ (inverseˡ (ι g)) x ⟩
x ∎

≋-trans : Transitive _≋_
≋-trans {x} {y} {z} (g by ιg∙x≈y) (h by ιh∙y≈z) = h N.∙ g by begin
ι (h N.∙ g) ∙ x ≈⟨ ∙-congʳ (ι.∙-homo h g) ⟩
(ι h ∙ ι g) ∙ x ≈⟨ uv≈w⇒xu∙v≈xw ιg∙x≈y (ι h) ⟩
ι h ∙ y ≈⟨ ιh∙y≈z ⟩
z ∎

≋-isEquivalence : IsEquivalence _≋_
≋-isEquivalence = record
{ refl = ≋-refl
; sym = ≋-sym
; trans = ≋-trans
}

open AlgDefs _≋_

≋-∙-cong : Congruent₂ _∙_
≋-∙-cong {x} {y} {u} {v} (g by ιg∙x≈y) (h by ιh∙u≈v) = g N.∙ h′ by begin
ι (g N.∙ h′) ∙ (x ∙ u) ≈⟨ ∙-congʳ (ι.∙-homo g h′) ⟩
(ι g ∙ ι h′) ∙ (x ∙ u) ≈⟨ uv≈wx⇒yu∙vz≈yw∙xz (normal h x) (ι g) u ⟩
(ι g ∙ x) ∙ (ι h ∙ u) ≈⟨ ∙-cong ιg∙x≈y ιh∙u≈v ⟩
y ∙ v ∎
where h′ = conjugate h x

≋-⁻¹-cong : Congruent₁ _⁻¹
≋-⁻¹-cong {x} {y} (g by ιg∙x≈y) = h by begin
ι h ∙ x ⁻¹ ≈⟨ normal (g N.⁻¹) (x ⁻¹) ⟩
x ⁻¹ ∙ ι (g N.⁻¹) ≈⟨ ∙-congˡ (ι.⁻¹-homo g) ⟩
x ⁻¹ ∙ ι g ⁻¹ ≈⟨ ⁻¹-anti-homo-∙ (ι g) x ⟨
(ι g ∙ x) ⁻¹ ≈⟨ ⁻¹-cong ιg∙x≈y ⟩
y ⁻¹ ∎
where h = conjugate (g N.⁻¹) (x ⁻¹)

quotientIsGroup : IsGroup _≋_ _∙_ ε _⁻¹
quotientIsGroup = record
{ isMonoid = record
{ isSemigroup = record
{ isMagma = record
{ isEquivalence = ≋-isEquivalence
; ∙-cong = ≋-∙-cong
}
; assoc = λ x y z → ≈⇒≋ (assoc x y z)
}
; identity = ≈⇒≋ ∘ identityˡ , ≈⇒≋ ∘ identityʳ
}
; inverse = ≈⇒≋ ∘ inverseˡ , ≈⇒≋ ∘ inverseʳ
; ⁻¹-cong = ≋-⁻¹-cong
}

quotientGroup : Group c (c ⊔ ℓ ⊔ c′)
quotientGroup = record { isGroup = quotientIsGroup }
Copy link
Contributor

Choose a reason for hiding this comment

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

And now inline the definition above?

Suggested change
quotientGroup = record { isGroup = quotientIsGroup }
quotientGroup = record
{ isGroup = ...
}


_/_ : Group c (c ⊔ ℓ ⊔ c′)
_/_ = quotientGroup

π : Group.Carrier G → Group.Carrier quotientGroup
π x = x -- because we do all the work in the relation

π-isHomomorphism : IsGroupHomomorphism rawGroup (Group.rawGroup quotientGroup) π
π-isHomomorphism = record
{ isMonoidHomomorphism = record
{ isMagmaHomomorphism = record
{ isRelHomomorphism = record
{ cong = ≈⇒≋
}
; homo = λ _ _ → ≋-refl
}
; ε-homo = ≋-refl
}
; ⁻¹-homo = λ _ → ≋-refl
}

π-surjective : Surjective _≈_ _≋_ π
π-surjective g = g , ≈⇒≋
37 changes: 37 additions & 0 deletions src/Algebra/NormalSubgroup.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,37 @@
------------------------------------------------------------------------
-- The Agda standard library
--
-- Definition of normal subgroups
------------------------------------------------------------------------

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

open import Algebra.Bundles using (Group)

module Algebra.NormalSubgroup {c ℓ} (G : Group c ℓ) where

open import Algebra.Definitions using (Commutative)
open import Algebra.Construct.Sub.Group G using (Subgroup)
open import Level using (suc; _⊔_)

private
module G = Group G

-- every element of the subgroup commutes in G
record IsNormal {c′ ℓ′} (subgroup : Subgroup c′ ℓ′) : Set (c ⊔ ℓ ⊔ c′) where
open Subgroup subgroup
field
conjugate : ∀ n g → Carrier
normal : ∀ n g → ι (conjugate n g) G.∙ g G.≈ g G.∙ ι n

record NormalSubgroup c′ ℓ′ : Set (c ⊔ ℓ ⊔ suc (c′ ⊔ ℓ′)) where
field
subgroup : Subgroup c′ ℓ′
isNormal : IsNormal subgroup

open Subgroup subgroup public
open IsNormal isNormal public

abelian⇒subgroup-normal : ∀ {c′ ℓ′} → Commutative G._≈_ G._∙_ → (subgroup : Subgroup c′ ℓ′) → IsNormal subgroup
abelian⇒subgroup-normal ∙-comm subgroup = record { normal = λ n g → ∙-comm (ι n) g }
where open Subgroup subgroup