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 @@ -79,10 +79,14 @@ Deprecated names
New modules
-----------

* `Algebra.Cosntruct.Quotient.Group` for the definition of quotient groups.
Copy link
Contributor

Choose a reason for hiding this comment

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

Typo

Suggested change
* `Algebra.Cosntruct.Quotient.Group` for the definition of quotient groups.
* `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
129 changes: 129 additions & 0 deletions src/Algebra/Construct/Quotient/Group.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,129 @@
------------------------------------------------------------------------
-- 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.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 +29 to +31
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

≋-refl : Reflexive _≋_
≋-refl {x} = N.ε by trans (∙-congʳ ι.ε-homo) (identityˡ x)
Comment on lines +38 to +39
Copy link
Contributor

Choose a reason for hiding this comment

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

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

Unsurprisingly, these proofs are definitionally equal to the previous versions, but put ≈⇒≋ ('abstract' reflexivity...) higher up the food chain.


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

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

Comment on lines +61 to +63
Copy link
Contributor

Choose a reason for hiding this comment

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

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

In view of the above refactoring of reflexivity...

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 = record
{ fst = λ x ≈⇒≋ (identityˡ x)
; snd = λ x ≈⇒≋ (identityʳ x)
}
}
; inverse = record
{ fst = λ x ≈⇒≋ (inverseˡ x)
; snd = λ x ≈⇒≋ (inverseʳ x)
}
; ⁻¹-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

project : Group.Carrier G Group.Carrier quotientGroup
Copy link
Contributor

Choose a reason for hiding this comment

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

My personal preference is for π (for projection), echoing ι (for inclusion), but I won't fight for it...
... maybe I should?

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've got no strong feeling here

project x = x -- because we do all the work in the relation

project-isHomomorphism : IsGroupHomomorphism rawGroup (Group.rawGroup quotientGroup) project
project-isHomomorphism = record
Comment on lines +114 to +115
Copy link
Contributor

Choose a reason for hiding this comment

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

Naming?

Suggested change
project-isHomomorphism : IsGroupHomomorphism rawGroup (Group.rawGroup quotientGroup) project
project-isHomomorphism = record
project-isGroupHomomorphism : IsGroupHomomorphism rawGroup (Group.rawGroup quotientGroup) project
project-isGroupHomomorphism = record

Twofold reasons:

  • Subsequently, for Ring R / Ideal I, you also define project-isHomomorphism : IsRingHomomorphism, and maybe that's enough (let the X path context of the algebra determine what kind of IsXHomomorphism is defined
  • BUT, for Ring, what you actually need is to separate out the isMonoidHomomorphism field, so why not also isGroupHomomorphism?

{ isMonoidHomomorphism = record
{ isMagmaHomomorphism = record
{ isRelHomomorphism = record
{ cong = ≈⇒≋
}
; homo = λ _ _ ≋-refl
}
; ε-homo = ≋-refl
}
Comment on lines +116 to +124
Copy link
Contributor

Choose a reason for hiding this comment

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

This should be lifted out as a separate (prior) definition

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

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

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

; ⁻¹-homo = λ _ ≋-refl
}

project-surjective : Surjective _≈_ _≋_ project
project-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
Copy link
Contributor

Choose a reason for hiding this comment

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

Unused!

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
Comment on lines +35 to +37
Copy link
Contributor

Choose a reason for hiding this comment

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

I'd maybe be tempted to lift this out to Algebra.Construct.Quotient.AbelianGroup to bundle up the quotient-out-of-any-subgroup construction, and simplify the downstream imports in #2855 ?