Skip to content

Commit 4529e73

Browse files
jamesmckinnaMatthewDaggitt
authored andcommitted
Add unique morphisms from/to Initial and Terminal algebras (#2296)
* added unique morphisms * refactored for uniformity's sake * exploit the uniformity * add missing instances * finish up, for now * `CHANGELOG` * `CHANGELOG` * `TheMorphism` * comment * comment * comment * `The` to `Unique` * lifted out istantiated `import` * blank line * note on instantiated `import`s * parametrise on the `Raw` bundle * parametrise on the `Raw` bundle * Rerranged to get rid of lots of boilerplate --------- Co-authored-by: MatthewDaggitt <[email protected]>
1 parent 2f5d88d commit 4529e73

File tree

7 files changed

+179
-7
lines changed

7 files changed

+179
-7
lines changed

CHANGELOG.md

+12
Original file line numberDiff line numberDiff line change
@@ -50,6 +50,12 @@ New modules
5050
-----------
5151
* `Algebra.Module.Bundles.Raw`: raw bundles for module-like algebraic structures
5252

53+
* The unique morphism from the initial, resp. terminal, algebra:
54+
```agda
55+
Algebra.Morphism.Construct.Initial
56+
Algebra.Morphism.Construct.Terminal
57+
```
58+
5359
* Nagata's construction of the "idealization of a module":
5460
```agda
5561
Algebra.Module.Construct.Idealization
@@ -117,6 +123,12 @@ Additions to existing modules
117123
rawModule : RawModule R c ℓ
118124
```
119125

126+
* In `Algebra.Construct.Terminal`:
127+
```agda
128+
rawNearSemiring : RawNearSemiring c ℓ
129+
nearSemiring : NearSemiring c ℓ
130+
```
131+
120132
* In `Algebra.Properties.Monoid.Mult`:
121133
```agda
122134
×-homo-0 : ∀ x → 0 × x ≈ 0#

doc/style-guide.md

+4
Original file line numberDiff line numberDiff line change
@@ -131,6 +131,10 @@ automate most of this.
131131
open SetoidEquality S
132132
```
133133

134+
* If importing a parametrised module, qualified or otherwise, with its
135+
parameters instantiated, then such 'instantiated imports' should be placed
136+
*after* the main block of `import`s, and *before* any `variable` declarations.
137+
134138
* Naming conventions for qualified `import`s: if importing a module under
135139
a root of the form `Data.X` (e.g. the `Base` module for basic operations,
136140
or `Properties` for lemmas about them etc.) then conventionally, the

src/Algebra/Construct/Initial.agda

+2-2
Original file line numberDiff line numberDiff line change
@@ -4,8 +4,8 @@
44
-- Instances of algebraic structures where the carrier is ⊥.
55
-- In mathematics, this is usually called 0.
66
--
7-
-- From monoids up, these are zero-objects – i.e, the initial
8-
-- object is the terminal object in the relevant category.
7+
-- From monoids up, these are zero-objects – i.e, the terminal
8+
-- object is *also* the initial object in the relevant category.
99
------------------------------------------------------------------------
1010

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

src/Algebra/Construct/Terminal.agda

+9-3
Original file line numberDiff line numberDiff line change
@@ -1,8 +1,8 @@
11
------------------------------------------------------------------------
22
-- The Agda standard library
33
--
4-
-- Instances of algebraic structures where the carrier is ⊤.
5-
-- In mathematics, this is usually called 0 (1 in the case of Group).
4+
-- Instances of algebraic structures where the carrier is ⊤. In
5+
-- mathematics, this is usually called 0 (1 in the case of Monoid, Group).
66
--
77
-- From monoids up, these are zero-objects – i.e, both the initial
88
-- and the terminal object in the relevant category.
@@ -27,7 +27,7 @@ module 𝕆ne where
2727
Carrier : Set c
2828
Carrier =
2929

30-
_≈_ : Rel Carrier ℓ
30+
_≈_ : Rel Carrier ℓ
3131
_ ≈ _ =
3232

3333
------------------------------------------------------------------------
@@ -42,6 +42,9 @@ rawMonoid = record { 𝕆ne }
4242
rawGroup : RawGroup c ℓ
4343
rawGroup = record { 𝕆ne }
4444

45+
rawNearSemiring : RawNearSemiring c ℓ
46+
rawNearSemiring = record { 𝕆ne }
47+
4548
rawSemiring : RawSemiring c ℓ
4649
rawSemiring = record { 𝕆ne }
4750

@@ -78,6 +81,9 @@ group = record { 𝕆ne }
7881
abelianGroup : AbelianGroup c ℓ
7982
abelianGroup = record { 𝕆ne }
8083

84+
nearSemiring : NearSemiring c ℓ
85+
nearSemiring = record { 𝕆ne }
86+
8187
semiring : Semiring c ℓ
8288
semiring = record { 𝕆ne }
8389

src/Algebra/Construct/Zero.agda

+2-2
Original file line numberDiff line numberDiff line change
@@ -1,8 +1,8 @@
11
------------------------------------------------------------------------
22
-- The Agda standard library
33
--
4-
-- Instances of algebraic structures where the carrier is ⊤.
5-
-- In mathematics, this is usually called 0 (1 in the case of Group).
4+
-- Instances of algebraic structures where the carrier is ⊤. In
5+
-- mathematics, this is usually called 0 (1 in the case of Monoid, Group).
66
--
77
-- From monoids up, these are are zero-objects – i.e, both the initial
88
-- and the terminal object in the relevant category.
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,62 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- The unique morphism from the initial object,
5+
-- for each of the relevant categories. Since
6+
-- `Semigroup` and `Band` are simply `Magma`s
7+
-- satisfying additional properties, it suffices to
8+
-- define the morphism on the underlying `RawMagma`.
9+
------------------------------------------------------------------------
10+
11+
{-# OPTIONS --cubical-compatible --safe #-}
12+
13+
open import Level using (Level)
14+
15+
module Algebra.Morphism.Construct.Initial {c ℓ : Level} where
16+
17+
open import Algebra.Bundles.Raw using (RawMagma)
18+
open import Algebra.Morphism.Structures
19+
open import Function.Definitions using (Injective)
20+
import Relation.Binary.Morphism.Definitions as Rel
21+
open import Relation.Binary.Morphism.Structures
22+
open import Relation.Binary.Core using (Rel)
23+
24+
open import Algebra.Construct.Initial {c} {ℓ}
25+
26+
private
27+
variable
28+
a m ℓm : Level
29+
A : Set a
30+
: Rel A ℓm
31+
32+
------------------------------------------------------------------------
33+
-- The unique morphism
34+
35+
zero : ℤero.Carrier A
36+
zero ()
37+
38+
------------------------------------------------------------------------
39+
-- Basic properties
40+
41+
cong : (≈ : Rel A ℓm) Rel.Homomorphic₂ ℤero.Carrier A ℤero._≈_ ≈ zero
42+
cong _ {x = ()}
43+
44+
injective : (≈ : Rel A ℓm) Injective ℤero._≈_ ≈ zero
45+
injective _ {x = ()}
46+
47+
------------------------------------------------------------------------
48+
-- Morphism structures
49+
50+
isMagmaHomomorphism : (M : RawMagma m ℓm)
51+
IsMagmaHomomorphism rawMagma M zero
52+
isMagmaHomomorphism M = record
53+
{ isRelHomomorphism = record { cong = cong (RawMagma._≈_ M) }
54+
; homo = λ()
55+
}
56+
57+
isMagmaMonomorphism : (M : RawMagma m ℓm)
58+
IsMagmaMonomorphism rawMagma M zero
59+
isMagmaMonomorphism M = record
60+
{ isMagmaHomomorphism = isMagmaHomomorphism M
61+
; injective = injective (RawMagma._≈_ M)
62+
}
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,88 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- The unique morphism to the terminal object,
5+
-- for each of the relevant categories. Since
6+
-- each terminal algebra builds on `Monoid`,
7+
-- possibly with additional (trivial) operations,
8+
-- satisfying additional properties, it suffices to
9+
-- define the morphism on the underlying `RawMonoid`
10+
------------------------------------------------------------------------
11+
12+
{-# OPTIONS --cubical-compatible --safe #-}
13+
14+
open import Level using (Level)
15+
16+
module Algebra.Morphism.Construct.Terminal {c ℓ : Level} where
17+
18+
open import Algebra.Bundles.Raw
19+
using (RawMagma; RawMonoid; RawGroup; RawNearSemiring; RawSemiring; RawRing)
20+
open import Algebra.Morphism.Structures
21+
22+
open import Data.Product.Base using (_,_)
23+
open import Function.Definitions using (StrictlySurjective)
24+
import Relation.Binary.Morphism.Definitions as Rel
25+
open import Relation.Binary.Morphism.Structures
26+
27+
open import Algebra.Construct.Terminal {c} {ℓ}
28+
29+
private
30+
variable
31+
a ℓa : Level
32+
A : Set a
33+
34+
------------------------------------------------------------------------
35+
-- The unique morphism
36+
37+
one : A 𝕆ne.Carrier
38+
one _ = _
39+
40+
------------------------------------------------------------------------
41+
-- Basic properties
42+
43+
strictlySurjective : A StrictlySurjective 𝕆ne._≈_ one
44+
strictlySurjective x _ = x , _
45+
46+
------------------------------------------------------------------------
47+
-- Homomorphisms
48+
49+
isMagmaHomomorphism : (M : RawMagma a ℓa)
50+
IsMagmaHomomorphism M rawMagma one
51+
isMagmaHomomorphism M = record
52+
{ isRelHomomorphism = record { cong = _ }
53+
; homo = _
54+
}
55+
56+
isMonoidHomomorphism : (M : RawMonoid a ℓa)
57+
IsMonoidHomomorphism M rawMonoid one
58+
isMonoidHomomorphism M = record
59+
{ isMagmaHomomorphism = isMagmaHomomorphism (RawMonoid.rawMagma M)
60+
; ε-homo = _
61+
}
62+
63+
isGroupHomomorphism : (G : RawGroup a ℓa)
64+
IsGroupHomomorphism G rawGroup one
65+
isGroupHomomorphism G = record
66+
{ isMonoidHomomorphism = isMonoidHomomorphism (RawGroup.rawMonoid G)
67+
; ⁻¹-homo = λ _ _
68+
}
69+
70+
isNearSemiringHomomorphism : (N : RawNearSemiring a ℓa)
71+
IsNearSemiringHomomorphism N rawNearSemiring one
72+
isNearSemiringHomomorphism N = record
73+
{ +-isMonoidHomomorphism = isMonoidHomomorphism (RawNearSemiring.+-rawMonoid N)
74+
; *-homo = λ _ _ _
75+
}
76+
77+
isSemiringHomomorphism : (S : RawSemiring a ℓa)
78+
IsSemiringHomomorphism S rawSemiring one
79+
isSemiringHomomorphism S = record
80+
{ isNearSemiringHomomorphism = isNearSemiringHomomorphism (RawSemiring.rawNearSemiring S)
81+
; 1#-homo = _
82+
}
83+
84+
isRingHomomorphism : (R : RawRing a ℓa) IsRingHomomorphism R rawRing one
85+
isRingHomomorphism R = record
86+
{ isSemiringHomomorphism = isSemiringHomomorphism (RawRing.rawSemiring R)
87+
; -‿homo = λ _ _
88+
}

0 commit comments

Comments
 (0)