Skip to content

Commit 0147df3

Browse files
committed
Add principal ideal construction
1 parent e18c588 commit 0147df3

File tree

2 files changed

+56
-0
lines changed

2 files changed

+56
-0
lines changed

CHANGELOG.md

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -87,6 +87,8 @@ New modules
8787

8888
* `Algebra.Ideal` for the definition of (two sided) ideals of rings.
8989

90+
* `Algebra.Ideal.Construct.Principal` for ideals generated by a single element of a commutative ring.
91+
9092
* `Algebra.Module.Construct.Sub.Bimodule` for the definition of subbimodules.
9193

9294
* `Algebra.NormalSubgroup` for the definition of normal subgroups.
Lines changed: 54 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,54 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Ideals generated by a single element of a commutative ring
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --safe --cubical-compatible #-}
8+
9+
open import Algebra.Bundles using (CommutativeRing)
10+
11+
module Algebra.Ideal.Construct.Principal {c ℓ} (R : CommutativeRing c ℓ) where
12+
13+
open import Function.Base using (id; _on_)
14+
15+
open CommutativeRing R
16+
17+
open import Algebra.Ideal ring
18+
open import Algebra.Properties.CommutativeSemigroup *-commutativeSemigroup
19+
open import Algebra.Properties.Ring ring
20+
21+
⟨_⟩ : Carrier Ideal c ℓ
22+
⟨ a ⟩ = record
23+
{ subbimodule = record
24+
{ domain = record
25+
{ Carrierᴹ = Carrier
26+
; _≈ᴹ_ = _≈_ on a *_
27+
; _+ᴹ_ = _+_
28+
; _*ₗ_ = _*_
29+
; _*ᵣ_ = _*_
30+
; 0ᴹ = 0#
31+
; -ᴹ_ = -_
32+
}
33+
; ι = a *_
34+
; ι-monomorphism = record
35+
{ isBimoduleHomomorphism = record
36+
{ +ᴹ-isGroupHomomorphism = record
37+
{ isMonoidHomomorphism = record
38+
{ isMagmaHomomorphism = record
39+
{ isRelHomomorphism = record
40+
{ cong = id
41+
}
42+
; homo = distribˡ a
43+
}
44+
; ε-homo = zeroʳ a
45+
}
46+
; ⁻¹-homo = λ x sym (-‿distribʳ-* a x)
47+
}
48+
; *ₗ-homo = x∙yz≈y∙xz a
49+
; *ᵣ-homo = λ r x sym (*-assoc a x r)
50+
}
51+
; injective = id
52+
}
53+
}
54+
}

0 commit comments

Comments
 (0)