Skip to content

Commit 1777421

Browse files
committed
Add subgroups and submodules
1 parent dd64a3a commit 1777421

File tree

3 files changed

+94
-0
lines changed

3 files changed

+94
-0
lines changed

CHANGELOG.md

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

82+
* `Algebra.Construct.Sub.Group` for the definition of subgroups.
83+
84+
* `Algebra.Module.Construct.Sub.Module` for the definition of submodules.
85+
8286
* `Algebra.Properties.BooleanRing`.
8387

8488
* `Algebra.Properties.BooleanSemiring`.
Lines changed: 40 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,40 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Definition of subgroups
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --safe --cubical-compatible #-}
8+
9+
open import Algebra.Bundles using (Group; RawGroup)
10+
11+
module Algebra.Construct.Sub.Group {c ℓ} (G : Group c ℓ) where
12+
13+
private
14+
module G = Group G
15+
16+
open import Algebra.Structures using (IsGroup)
17+
open import Algebra.Morphism.Structures using (IsGroupMonomorphism)
18+
import Algebra.Morphism.GroupMonomorphism as GroupMonomorphism
19+
open import Level using (suc; _⊔_)
20+
21+
record Subgroup c′ ℓ′ : Set (c ⊔ ℓ ⊔ suc (c′ ⊔ ℓ′)) where
22+
field
23+
Sub : RawGroup c′ ℓ′
24+
25+
private
26+
module Sub = RawGroup Sub
27+
28+
field
29+
ι : Sub.Carrier G.Carrier
30+
ι-monomorphism : IsGroupMonomorphism Sub G.rawGroup ι
31+
32+
module ι = IsGroupMonomorphism ι-monomorphism
33+
34+
isGroup : IsGroup Sub._≈_ Sub._∙_ Sub.ε Sub._⁻¹
35+
isGroup = GroupMonomorphism.isGroup ι-monomorphism G.isGroup
36+
37+
group : Group _ _
38+
group = record { isGroup = isGroup }
39+
40+
open Group group public hiding (isGroup)
Lines changed: 50 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,50 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Definition of submodules
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
open import Algebra.Bundles using (CommutativeRing)
10+
open import Algebra.Module.Bundles using (Module; RawModule)
11+
12+
module Algebra.Module.Construct.Sub.Module {c ℓ cm ℓm} {R : CommutativeRing c ℓ} (M : Module R cm ℓm) where
13+
14+
private
15+
module R = CommutativeRing R
16+
module M = Module M
17+
18+
open import Algebra.Construct.Sub.Group M.+ᴹ-group
19+
open import Algebra.Module.Structures using (IsModule)
20+
open import Algebra.Module.Morphism.Structures using (IsModuleMonomorphism)
21+
import Algebra.Module.Morphism.ModuleMonomorphism as ModuleMonomorphism
22+
open import Level using (suc; _⊔_)
23+
24+
record Submodule cm′ ℓm′ : Set (c ⊔ cm ⊔ ℓm ⊔ suc (cm′ ⊔ ℓm′)) where
25+
field
26+
Sub : RawModule R.Carrier cm′ ℓm′
27+
28+
private
29+
module Sub = RawModule Sub
30+
31+
field
32+
ι : Sub.Carrierᴹ M.Carrierᴹ
33+
ι-monomorphism : IsModuleMonomorphism Sub M.rawModule ι
34+
35+
module ι = IsModuleMonomorphism ι-monomorphism
36+
37+
isModule : IsModule R Sub._≈ᴹ_ Sub._+ᴹ_ Sub.0ᴹ Sub.-ᴹ_ Sub._*ₗ_ Sub._*ᵣ_
38+
isModule = ModuleMonomorphism.isModule ι-monomorphism R.isCommutativeRing M.isModule
39+
40+
⟨module⟩ : Module R _ _
41+
⟨module⟩ = record { isModule = isModule }
42+
43+
open Module ⟨module⟩ public hiding (isModule)
44+
45+
subgroup : Subgroup cm′ ℓm′
46+
subgroup = record
47+
{ N = Sub.+ᴹ-rawGroup
48+
; ι = ι
49+
; ι-monomorphism = ι.+ᴹ-isGroupMonomorphism
50+
}

0 commit comments

Comments
 (0)