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

* `Algebra.Construct.Quotient.{Abelian}Group` for the definition of quotient (Abelian) groups.

* `Algebra.Construct.Sub.{Abelian}Group` for the definition of sub-(Abelian)groups.

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

* `Algebra.Construct.Sub.Ring.Ideal` for the definition of ideals of a ring.

* `Algebra.Module.Construct.Sub.Bimodule` for the definition of sub-bimodules.

* `Algebra.Properties.BooleanRing`.

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

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

open import Algebra.Bundles using (Group; AbelianGroup)
import Algebra.Construct.Sub.AbelianGroup as AbelianSubgroup

module Algebra.Construct.Quotient.AbelianGroup
{c ℓ} (G : AbelianGroup c ℓ)
(open AbelianSubgroup G using (Subgroup; normalSubgroup))
{c′ ℓ′} (N : Subgroup c′ ℓ′)
where

private
module G = AbelianGroup G

-- Re-export the quotient group

open import Algebra.Construct.Quotient.Group G.group (normalSubgroup N) public

-- With its additional bundle

quotientAbelianGroup : AbelianGroup c _
quotientAbelianGroup = record
{ isAbelianGroup = record
{ isGroup = isGroup
; comm = λ g h → ≈⇒≋ (G.comm g h)
}
} where open Group quotientGroup

-- Public re-exports, as needed?
126 changes: 126 additions & 0 deletions src/Algebra/Construct/Quotient/Group.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,126 @@
------------------------------------------------------------------------
-- The Agda standard library
--
-- Quotient groups
------------------------------------------------------------------------

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

open import Algebra.Bundles using (Group)
open import Algebra.Construct.Sub.Group.Normal using (NormalSubgroup)

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

open import Algebra.Definitions using (Congruent₁; Congruent₂)
open import Algebra.Morphism.Structures
using (IsMagmaHomomorphism; IsMonoidHomomorphism; IsGroupHomomorphism)
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)

private
open module G = Group G

open import Algebra.Properties.Group G using (⁻¹-anti-homo-∙)
open import Algebra.Properties.Monoid monoid
open import Relation.Binary.Reasoning.Setoid setoid

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

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

≋-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 ∎

≋-∙-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 ⁻¹)

quotientGroup : Group c (c ⊔ ℓ ⊔ c′)
quotientGroup = record
{ isGroup = record
{ isMonoid = record
{ isSemigroup = record
{ isMagma = record
{ isEquivalence = record
Copy link
Contributor

Choose a reason for hiding this comment

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

I know this is just an experiment, but I'd certainly with that this IsEquivalence record was pulled out.

Copy link
Contributor Author

Choose a reason for hiding this comment

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

One thing that impressed me about @Taneb 's code was the extent to which, in various places, it manages to produces huge nested record blocks which nevertheless 'do the right thing'.

I've flip-flopped over the years between this style, and a much more incremental, define every sub-record at top-level, put the pieces together step-by-step style, which the library has to date favoured.

One reason (perhaps two or more!) I am now turning towards the 'one big record' style arises from #2391 :

  • a single large object, suitably re-opened as public, (re-)introduces all the 'correct' substructure, moreover with the 'rectified'/'official' names, rather than trying to invent proxy names for the top-level substructure in the other style
  • when any client imports the new constructions, they aren't faced with a (DRY) choice between cherry-picking the ad-hoc names (which may be the 'correct' ones), or facing a clash/choicepoint between such names and those names introduced by open of substructures

In this instance (as perhaps elsewhere...) the defintion of distinct isEquivalence : IsEquivalence ... fields at top-level is starting to emerge as an anti-pattern for me: once an Algebra.Structures.IsX object is in scope, or has been defined, then isEquivalance is available as a canonical projection from that, and indeed that should be the preferred mode-of-access (via open if necessary) for a substructure which exists solely

  • to export canonical names refl, sym, trans for the properties, and their derived forms
  • to construct Setoids from which Relation.Binary.Reasoning.Setoid syntax may then be brought into scope
  • ...

... much of which functionality/behaviour can be achieved at a higher-level of the nesting hierarchy by suitable re-organisation of the Algebra.Properties.X hierarchy #2804 #2858 etc.

{ refl = ≋-refl
; sym = ≋-sym
; trans = ≋-trans
}
; ∙-cong = ≋-∙-cong
}
; assoc = λ x y z → ≈⇒≋ (assoc x y z)
}
; identity = ≈⇒≋ ∘ identityˡ , ≈⇒≋ ∘ identityʳ
}
; inverse = ≈⇒≋ ∘ inverseˡ , ≈⇒≋ ∘ inverseʳ
; ⁻¹-cong = ≋-⁻¹-cong
}
}

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

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

π-isMagmaHomomorphism : IsMagmaHomomorphism rawMagma (Group.rawMagma quotientGroup) π
π-isMagmaHomomorphism = record
{ isRelHomomorphism = record
{ cong = ≈⇒≋
}
; homo = λ _ _ → ≋-refl
}

π-isMonoidHomomorphism : IsMonoidHomomorphism rawMonoid (Group.rawMonoid quotientGroup) π
π-isMonoidHomomorphism = record
{ isMagmaHomomorphism = π-isMagmaHomomorphism
; ε-homo = ≋-refl
}

π-isGroupHomomorphism : IsGroupHomomorphism rawGroup (Group.rawGroup quotientGroup) π
π-isGroupHomomorphism = record
{ isMonoidHomomorphism = π-isMonoidHomomorphism
; ⁻¹-homo = λ _ → ≋-refl
}

π-surjective : Surjective _≈_ _≋_ π
π-surjective g = g , ≈⇒≋
76 changes: 76 additions & 0 deletions src/Algebra/Construct/Quotient/Ring.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,76 @@
------------------------------------------------------------------------
-- The Agda standard library
--
-- Quotient rings
------------------------------------------------------------------------

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

open import Algebra.Bundles using (AbelianGroup; Ring; RawRing)
open import Algebra.Construct.Sub.Ring.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 Data.Product.Base using (_,_)
open import Function.Base using (_∘_)
open import Level using (_⊔_)

import Algebra.Construct.Quotient.AbelianGroup as Quotient

private
module R = Ring R
module I = Ideal I
module R/I = Quotient R.+-abelianGroup I.subgroup


open R/I public
using (_≋_; _by_; ≋-refl; ≈⇒≋
; quotientAbelianGroup
; π; π-isMonoidHomomorphism; π-surjective
)

open import Algebra.Definitions _≋_ using (Congruent₂)

≋-*-cong : Congruent₂ R._*_
≋-*-cong {x} {y} {u} {v} (j by ιj+x≈y) (k by ιk+u≈v) = ι j *ₗ k +ᴹ j *ᵣ u +ᴹ x *ₗ k by begin
ι (ι j *ₗ k +ᴹ j *ᵣ u +ᴹ x *ₗ k) + x * u ≈⟨ +-congʳ (ι.+ᴹ-homo (ι j *ₗ k +ᴹ j *ᵣ u) (x *ₗ k)) ⟩
ι (ι j *ₗ k +ᴹ j *ᵣ u) + ι (x *ₗ k) + x * u ≈⟨ +-congʳ (+-congʳ (ι.+ᴹ-homo (ι j *ₗ k) (j *ᵣ u))) ⟩
ι (ι j *ₗ k) + ι (j *ᵣ u) + ι (x *ₗ 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 ∎
where
open R using (_+_; _*_; +-congʳ ;+-cong; *-cong)
open import Algebra.Properties.Semiring R.semiring using (binomial-expansion)
open import Relation.Binary.Reasoning.Setoid R.setoid
open I using (ι; _*ₗ_; _*ᵣ_; _+ᴹ_)

quotientRawRing : RawRing c (c ⊔ ℓ ⊔ c′)
quotientRawRing = record { RawRing R.rawRing hiding (_≈_) ; _≈_ = _≋_ }

quotientRing : Ring c (c ⊔ ℓ ⊔ c′)
quotientRing = record
{ isRing = record
{ +-isAbelianGroup = isAbelianGroup
; *-cong = ≋-*-cong
; *-assoc = λ x y z → ≈⇒≋ (R.*-assoc x y z)
; *-identity = ≈⇒≋ ∘ R.*-identityˡ , ≈⇒≋ ∘ R.*-identityʳ
; distrib = (λ x y z → ≈⇒≋ (R.distribˡ x y z)) , (λ x y z → ≈⇒≋ (R.distribʳ x y z))
}
} where open AbelianGroup quotientAbelianGroup using (isAbelianGroup)

π-isRingHomomorphism : IsRingHomomorphism R.rawRing quotientRawRing π
π-isRingHomomorphism = record
{ isSemiringHomomorphism = record
{ isNearSemiringHomomorphism = record
{ +-isMonoidHomomorphism = π-isMonoidHomomorphism
; *-homo = λ _ _ → ≋-refl
}
; 1#-homo = ≋-refl
}
; -‿homo = λ _ → ≋-refl
}

78 changes: 78 additions & 0 deletions src/Algebra/Construct/Sub/AbelianGroup.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,78 @@
------------------------------------------------------------------------
-- The Agda standard library
--
-- Subgroups of Abelian groups: necessarily Normal
--
-- This is a strict addition/extension to Nathan van Doorn (Taneb)'s PR
-- https://github.com/agda/agda-stdlib/pull/2852
-- and avoids the lemma `Algebra.NormalSubgroup.abelian⇒subgroup-normal`
-- introduced in PR https://github.com/agda/agda-stdlib/pull/2854
-- in favour of the direct definition of the field `normal` below.
------------------------------------------------------------------------

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

open import Algebra.Bundles using (AbelianGroup)

module Algebra.Construct.Sub.AbelianGroup {c ℓ} (G : AbelianGroup c ℓ) where

-- As with the corresponding appeal to `IsGroup` in the module
-- `Algebra.Construct.Sub.AbelianGroup`, this import drives the export
-- of the `X` bundle corresponding to the `subX` structure/bundle
-- being defined here.

open import Algebra.Morphism.GroupMonomorphism using (isAbelianGroup)

private
module G = AbelianGroup G

-- Here, we could chose simply to expose the `NormalSubgroup` definition
-- from `Algebra.Construct.Sub.Group.Normal`, but as this module will use
-- type inference to allow the creation of `NormalSubgroup`s based on their
-- `isNormal` fields alone, it makes sense also to export the type `IsNormal`.

open import Algebra.Construct.Sub.Group.Normal G.group
using (IsNormal; NormalSubgroup)

-- Re-export the notion of subgroup of the underlying Group

open import Algebra.Construct.Sub.Group G.group public
using (Subgroup)

-- Then, for any such Subgroup:
-- * it is, in fact, Normal
-- * and defines an AbelianGroup, not just a Group

module _ {c′ ℓ′} (subgroup : Subgroup c′ ℓ′) where

-- Here, we do need both the underlying function `ι` of the mono, and
-- the proof `ι-monomorphism`, in order to be able to construct an
-- `IsAbelianGroup` structure on the way to the `AbelianGroup` bundle.

open Subgroup subgroup public
using (ι; ι-monomorphism)

-- Here, we eta-contract the use of `G.comm` in defining the `normal` field.

isNormal : IsNormal subgroup
isNormal = record { normal = λ n → G.comm (ι n) }
Copy link
Contributor

Choose a reason for hiding this comment

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

eta contract?

Copy link
Contributor Author

Choose a reason for hiding this comment

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

Yes! In that, it is itself an eta-contraction of the original

isNormal = record { normal = λ n g  G.comm (ι n) g }

Is it harmful to have done this?

Or are you asking whether the definition could be contracted further? Perhaps, but only via one of those odd Function.Base composition operators which compose only in one argument position of an n+2-ary function, which I find obscure rather than illuminate (my point-free fu is not always as strong as it might be).


-- And the use of `record`s throughout permits this 'boilerplate' construction.

normalSubgroup : NormalSubgroup c′ ℓ′
normalSubgroup = record { isNormal = isNormal }

-- And the `public` re-export of its substructure.

open NormalSubgroup normalSubgroup public
using (conjugate; normal)

-- As with `Algebra.Construct.Sub.Group`, there seems no need to create
-- an intermediate manifest field `isAbelianGroup`, when this may be, and
-- is, obtained by opening the bundle for `public` export.

abelianGroup : AbelianGroup c′ ℓ′
abelianGroup = record
{ isAbelianGroup = isAbelianGroup ι-monomorphism G.isAbelianGroup }

open AbelianGroup abelianGroup public
Loading