Skip to content

Commit 1518194

Browse files
committed
Add normal subgroups and quotient groups
1 parent 07ccc28 commit 1518194

File tree

3 files changed

+159
-0
lines changed

3 files changed

+159
-0
lines changed

CHANGELOG.md

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -79,10 +79,14 @@ Deprecated names
7979
New modules
8080
-----------
8181

82+
* `Algebra.Cosntruct.Quotient.Group` for the definition of quotient groups.
83+
8284
* `Algebra.Construct.Sub.Group` for the definition of subgroups.
8385

8486
* `Algebra.Module.Construct.Sub.Module` for the definition of submodules.
8587

88+
* `Algebra.NormalSubgroup` for the definition of normal subgroups.
89+
8690
* `Algebra.Properties.BooleanRing`.
8791

8892
* `Algebra.Properties.BooleanSemiring`.
Lines changed: 127 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,127 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Quotient groups
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --safe --cubical-compatible #-}
8+
9+
open import Algebra.Bundles using (Group)
10+
open import Algebra.NormalSubgroup using (NormalSubgroup)
11+
12+
module Algebra.Construct.Quotient.Group {c ℓ} (G : Group c ℓ) {c′ ℓ′} (N : NormalSubgroup G c′ ℓ′) where
13+
14+
open Group G
15+
16+
import Algebra.Definitions as AlgDefs
17+
open import Algebra.Morphism.Structures using (IsGroupHomomorphism)
18+
open import Algebra.Properties.Monoid monoid
19+
open import Algebra.Properties.Group G using (⁻¹-anti-homo-∙)
20+
open import Algebra.Structures using (IsGroup)
21+
open import Data.Product.Base using (_,_; proj₁; proj₂)
22+
open import Function.Definitions using (Surjective)
23+
open import Level using (_⊔_)
24+
open import Relation.Binary.Core using (_⇒_)
25+
open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive)
26+
open import Relation.Binary.Structures using (IsEquivalence)
27+
open import Relation.Binary.Reasoning.Setoid setoid
28+
29+
private
30+
module N = NormalSubgroup N
31+
open NormalSubgroup N using (ι; module ι; normal)
32+
33+
infix 0 _by_
34+
35+
data _≋_ (x y : Carrier) : Set (c ⊔ ℓ ⊔ c′) where
36+
_by_ : g ι g ∙ x ≈ y x ≋ y
37+
38+
≋-refl : Reflexive _≋_
39+
≋-refl {x} = N.ε by trans (∙-congʳ ι.ε-homo) (identityˡ x)
40+
41+
≋-sym : Symmetric _≋_
42+
≋-sym {x} {y} (g by ιg∙x≈y) = g N.⁻¹ by begin
43+
ι (g N.⁻¹) ∙ y ≈⟨ ∙-cong (ι.⁻¹-homo g) (sym ιg∙x≈y) ⟩
44+
ι g ⁻¹ ∙ (ι g ∙ x) ≈⟨ cancelˡ (inverseˡ (ι g)) x ⟩
45+
x ∎
46+
47+
≋-trans : Transitive _≋_
48+
≋-trans {x} {y} {z} (g by ιg∙x≈y) (h by ιh∙y≈z) = h N.∙ g by begin
49+
ι (h N.∙ g) ∙ x ≈⟨ ∙-congʳ (ι.∙-homo h g) ⟩
50+
(ι h ∙ ι g) ∙ x ≈⟨ uv≈w⇒xu∙v≈xw ιg∙x≈y (ι h) ⟩
51+
ι h ∙ y ≈⟨ ιh∙y≈z ⟩
52+
z ∎
53+
54+
≋-isEquivalence : IsEquivalence _≋_
55+
≋-isEquivalence = record
56+
{ refl = ≋-refl
57+
; sym = ≋-sym
58+
; trans = ≋-trans
59+
}
60+
61+
≈⇒≋ : _≈_ ⇒ _≋_
62+
≈⇒≋ {x} {y} x≈y = N.ε by trans (∙-cong ι.ε-homo x≈y) (identityˡ y)
63+
64+
open AlgDefs _≋_
65+
66+
≋-∙-cong : Congruent₂ _∙_
67+
≋-∙-cong {x} {y} {u} {v} (g by ιg∙x≈y) (h by ιh∙u≈v) = g N.∙ h′ by begin
68+
ι (g N.∙ h′) ∙ (x ∙ u) ≈⟨ ∙-congʳ (ι.∙-homo g h′) ⟩
69+
(ι g ∙ ι h′) ∙ (x ∙ u) ≈⟨ uv∙wx≈u[vw∙x] (ι g) (ι h′) x u ⟩
70+
ι g ∙ ((ι h′ ∙ x) ∙ u) ≈⟨ uv≈w⇒xu∙v≈xw (uv≈w⇒u∙vx≈wx (normal h x .proj₂) u) (ι g) ⟨
71+
(ι g ∙ x) ∙ (ι h ∙ u) ≈⟨ ∙-cong ιg∙x≈y ιh∙u≈v ⟩
72+
y ∙ v ∎
73+
where h′ = normal h x .proj₁
74+
75+
≋-⁻¹-cong : Congruent₁ _⁻¹
76+
≋-⁻¹-cong {x} {y} (g by ιg∙x≈y) = h by begin
77+
ι h ∙ x ⁻¹ ≈⟨ normal (g N.⁻¹) (x ⁻¹) .proj₂ ⟨
78+
x ⁻¹ ∙ ι (g N.⁻¹) ≈⟨ ∙-congˡ (ι.⁻¹-homo g) ⟩
79+
x ⁻¹ ∙ ι g ⁻¹ ≈⟨ ⁻¹-anti-homo-∙ (ι g) x ⟨
80+
(ι g ∙ x) ⁻¹ ≈⟨ ⁻¹-cong ιg∙x≈y ⟩
81+
y ⁻¹ ∎
82+
where h = normal (g N.⁻¹) (x ⁻¹) .proj₁
83+
84+
quotientIsGroup : IsGroup _≋_ _∙_ ε _⁻¹
85+
quotientIsGroup = record
86+
{ isMonoid = record
87+
{ isSemigroup = record
88+
{ isMagma = record
89+
{ isEquivalence = ≋-isEquivalence
90+
; ∙-cong = ≋-∙-cong
91+
}
92+
; assoc = λ x y z ≈⇒≋ (assoc x y z)
93+
}
94+
; identity = record
95+
{ fst = λ x ≈⇒≋ (identityˡ x)
96+
; snd = λ x ≈⇒≋ (identityʳ x)
97+
}
98+
}
99+
; inverse = record
100+
{ fst = λ x ≈⇒≋ (inverseˡ x)
101+
; snd = λ x ≈⇒≋ (inverseʳ x)
102+
}
103+
; ⁻¹-cong = ≋-⁻¹-cong
104+
}
105+
106+
quotientGroup : Group c (c ⊔ ℓ ⊔ c′)
107+
quotientGroup = record { isGroup = quotientIsGroup }
108+
109+
project : Group.Carrier G Group.Carrier quotientGroup
110+
project x = x -- because we do all the work in the relation
111+
112+
project-isHomomorphism : IsGroupHomomorphism rawGroup (Group.rawGroup quotientGroup) project
113+
project-isHomomorphism = record
114+
{ isMonoidHomomorphism = record
115+
{ isMagmaHomomorphism = record
116+
{ isRelHomomorphism = record
117+
{ cong = ≈⇒≋
118+
}
119+
; homo = λ _ _ ≋-refl
120+
}
121+
; ε-homo = ≋-refl
122+
}
123+
; ⁻¹-homo = λ _ ≋-refl
124+
}
125+
126+
project-surjective : Surjective _≈_ _≋_ project
127+
project-surjective g = g , ≈⇒≋

src/Algebra/NormalSubgroup.agda

Lines changed: 28 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,28 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Definition of normal subgroups
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --safe --cubical-compatible #-}
8+
9+
open import Algebra.Bundles using (Group)
10+
11+
module Algebra.NormalSubgroup {c ℓ} (G : Group c ℓ) where
12+
13+
open import Algebra.Construct.Sub.Group G using (Subgroup)
14+
open import Data.Product.Base using (∃-syntax)
15+
open import Level using (suc; _⊔_)
16+
17+
private
18+
module G = Group G
19+
20+
record NormalSubgroup c′ ℓ′ : Set (c ⊔ ℓ ⊔ suc (c′ ⊔ ℓ′)) where
21+
field
22+
subgroup : Subgroup c′ ℓ′
23+
24+
open Subgroup subgroup public
25+
26+
field
27+
-- every element of N commutes in G
28+
normal : n g ∃[ n′ ] g G.∙ ι n G.≈ ι n′ G.∙ g

0 commit comments

Comments
 (0)