Skip to content

Commit 21b7243

Browse files
[fixes #2273] Add SuccessorSet and associated boilerplate (#2277)
* added `RawNNO` * [fix #2273] Add `NNO`-related modules * start again, fail better... * rectify names * tighten `import`s * rectify names: `CHANGELOG` --------- Co-authored-by: MatthewDaggitt <[email protected]>
1 parent 65c296a commit 21b7243

File tree

5 files changed

+126
-11
lines changed

5 files changed

+126
-11
lines changed

CHANGELOG.md

+27-6
Original file line numberDiff line numberDiff line change
@@ -128,13 +128,29 @@ New modules
128128
Additions to existing modules
129129
-----------------------------
130130

131+
* In `Algebra.Bundles`
132+
```agda
133+
record SuccessorSet c ℓ : Set (suc (c ⊔ ℓ))
134+
```
135+
136+
* In `Algebra.Bundles.Raw`
137+
```agda
138+
record RawSuccessorSet c ℓ : Set (suc (c ⊔ ℓ))
139+
```
140+
131141
* Exporting more `Raw` substructures from `Algebra.Bundles.Ring`:
132142
```agda
133143
rawNearSemiring : RawNearSemiring _ _
134144
rawRingWithoutOne : RawRingWithoutOne _ _
135145
+-rawGroup : RawGroup _ _
136146
```
137147

148+
* In `Algebra.Construct.Terminal`:
149+
```agda
150+
rawNearSemiring : RawNearSemiring c ℓ
151+
nearSemiring : NearSemiring c ℓ
152+
```
153+
138154
* In `Algebra.Module.Bundles`, raw bundles are re-exported and the bundles expose their raw counterparts.
139155

140156
* In `Algebra.Module.Construct.DirectProduct`:
@@ -173,6 +189,13 @@ Additions to existing modules
173189
rawModule : RawModule R c ℓ
174190
```
175191

192+
* In `Algebra.Morphism.Structures`
193+
```agda
194+
module SuccessorSetMorphisms (N₁ : RawSuccessorSet a ℓ₁) (N₂ : RawSuccessorSet b ℓ₂) where
195+
record IsSuccessorSetHomomorphism (⟦_⟧ : N₁.Carrier → N₂.Carrier) : Set _
196+
record IsSuccessorSetMonomorphism (⟦_⟧ : N₁.Carrier → N₂.Carrier) : Set _
197+
record IsSuccessorSetIsomorphism (⟦_⟧ : N₁.Carrier → N₂.Carrier) : Set _
198+
176199
* In `Algebra.Properties.Group`:
177200
```agda
178201
isQuasigroup : IsQuasigroup _∙_ _\\_ _//_
@@ -199,12 +222,6 @@ Additions to existing modules
199222
identity-unique : Identity x _∙_ → x ≈ ε
200223
```
201224

202-
* In `Algebra.Construct.Terminal`:
203-
```agda
204-
rawNearSemiring : RawNearSemiring c ℓ
205-
nearSemiring : NearSemiring c ℓ
206-
```
207-
208225
* In `Algebra.Properties.Monoid.Mult`:
209226
```agda
210227
×-homo-0 : ∀ x → 0 × x ≈ 0#
@@ -218,6 +235,10 @@ Additions to existing modules
218235
idem-×-homo-* : (_*_ IdempotentOn x) → (m × x) * (n × x) ≈ (m ℕ.* n) × x
219236
```
220237

238+
* In `Algebra.Structures`
239+
```agda
240+
record IsSuccessorSet (suc# : Op₁ A) (zero# : A) : Set _
241+
221242
* In `Algebra.Structures.IsGroup`:
222243
```agda
223244
infixl 6 _//_

src/Algebra/Bundles.agda

+21-3
Original file line numberDiff line numberDiff line change
@@ -15,19 +15,37 @@ import Algebra.Bundles.Raw as Raw
1515
open import Algebra.Core
1616
open import Algebra.Structures
1717
open import Relation.Binary.Core using (Rel)
18-
open import Function.Base
19-
import Relation.Nullary as N
2018
open import Level
2119

2220
------------------------------------------------------------------------
2321
-- Re-export definitions of 'raw' bundles
2422

2523
open Raw public
26-
using (RawMagma; RawMonoid; RawGroup
24+
using ( RawSuccessorSet; RawMagma; RawMonoid; RawGroup
2725
; RawNearSemiring; RawSemiring
2826
; RawRingWithoutOne; RawRing
2927
; RawQuasigroup; RawLoop; RawKleeneAlgebra)
3028

29+
------------------------------------------------------------------------
30+
-- Bundles with 1 unary operation & 1 element
31+
------------------------------------------------------------------------
32+
33+
record SuccessorSet c ℓ : Set (suc (c ⊔ ℓ)) where
34+
infix 4 _≈_
35+
field
36+
Carrier : Set c
37+
_≈_ : Rel Carrier ℓ
38+
suc# : Op₁ Carrier
39+
zero# : Carrier
40+
isSuccessorSet : IsSuccessorSet _≈_ suc# zero#
41+
42+
open IsSuccessorSet isSuccessorSet public
43+
44+
rawSuccessorSet : RawSuccessorSet _ _
45+
rawSuccessorSet = record { _≈_ = _≈_; suc# = suc#; zero# = zero# }
46+
47+
open RawSuccessorSet rawSuccessorSet public
48+
3149
------------------------------------------------------------------------
3250
-- Bundles with 1 binary operation
3351
------------------------------------------------------------------------

src/Algebra/Bundles/Raw.agda

+14
Original file line numberDiff line numberDiff line change
@@ -13,6 +13,20 @@ open import Relation.Binary.Core using (Rel)
1313
open import Level using (suc; _⊔_)
1414
open import Relation.Nullary.Negation.Core using (¬_)
1515

16+
------------------------------------------------------------------------
17+
-- Raw bundles with 1 unary operation & 1 element
18+
------------------------------------------------------------------------
19+
20+
-- A raw SuccessorSet is a SuccessorSet without any laws.
21+
22+
record RawSuccessorSet c ℓ : Set (suc (c ⊔ ℓ)) where
23+
infix 4 _≈_
24+
field
25+
Carrier : Set c
26+
_≈_ : Rel Carrier ℓ
27+
suc# : Op₁ Carrier
28+
zero# : Carrier
29+
1630
------------------------------------------------------------------------
1731
-- Raw bundles with 1 binary operation
1832
------------------------------------------------------------------------

src/Algebra/Morphism/Structures.agda

+50-2
Original file line numberDiff line numberDiff line change
@@ -6,21 +6,69 @@
66

77
{-# OPTIONS --cubical-compatible --safe #-}
88

9-
open import Relation.Binary.Core
10-
119
module Algebra.Morphism.Structures where
1210

1311
open import Algebra.Core
1412
open import Algebra.Bundles
1513
import Algebra.Morphism.Definitions as MorphismDefinitions
1614
open import Level using (Level; _⊔_)
1715
open import Function.Definitions
16+
open import Relation.Binary.Core
1817
open import Relation.Binary.Morphism.Structures
1918

2019
private
2120
variable
2221
a b ℓ₁ ℓ₂ : Level
2322

23+
------------------------------------------------------------------------
24+
-- Morphisms over SuccessorSet-like structures
25+
------------------------------------------------------------------------
26+
27+
module SuccessorSetMorphisms
28+
(N₁ : RawSuccessorSet a ℓ₁) (N₂ : RawSuccessorSet b ℓ₂)
29+
where
30+
31+
open RawSuccessorSet N₁
32+
renaming (Carrier to A; _≈_ to _≈₁_; suc# to suc#₁; zero# to zero#₁)
33+
open RawSuccessorSet N₂
34+
renaming (Carrier to B; _≈_ to _≈₂_; suc# to suc#₂; zero# to zero#₂)
35+
open MorphismDefinitions A B _≈₂_
36+
37+
38+
record IsSuccessorSetHomomorphism (⟦_⟧ : A B) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where
39+
field
40+
isRelHomomorphism : IsRelHomomorphism _≈₁_ _≈₂_ ⟦_⟧
41+
suc#-homo : Homomorphic₁ ⟦_⟧ suc#₁ suc#₂
42+
zero#-homo : Homomorphic₀ ⟦_⟧ zero#₁ zero#₂
43+
44+
record IsSuccessorSetMonomorphism (⟦_⟧ : A B) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where
45+
field
46+
isSuccessorSetHomomorphism : IsSuccessorSetHomomorphism ⟦_⟧
47+
injective : Injective _≈₁_ _≈₂_ ⟦_⟧
48+
49+
open IsSuccessorSetHomomorphism isSuccessorSetHomomorphism public
50+
51+
isRelMonomorphism : IsRelMonomorphism _≈₁_ _≈₂_ ⟦_⟧
52+
isRelMonomorphism = record
53+
{ isHomomorphism = isRelHomomorphism
54+
; injective = injective
55+
}
56+
57+
58+
record IsSuccessorSetIsomorphism (⟦_⟧ : A B) : Set (a ⊔ b ⊔ ℓ₁ ⊔ ℓ₂) where
59+
field
60+
isSuccessorSetMonomorphism : IsSuccessorSetMonomorphism ⟦_⟧
61+
surjective : Surjective _≈₁_ _≈₂_ ⟦_⟧
62+
63+
open IsSuccessorSetMonomorphism isSuccessorSetMonomorphism public
64+
65+
isRelIsomorphism : IsRelIsomorphism _≈₁_ _≈₂_ ⟦_⟧
66+
isRelIsomorphism = record
67+
{ isMonomorphism = isRelMonomorphism
68+
; surjective = surjective
69+
}
70+
71+
2472
------------------------------------------------------------------------
2573
-- Morphisms over magma-like structures
2674
------------------------------------------------------------------------

src/Algebra/Structures.agda

+14
Original file line numberDiff line numberDiff line change
@@ -27,6 +27,20 @@ import Algebra.Consequences.Setoid as Consequences
2727
open import Data.Product.Base using (_,_; proj₁; proj₂)
2828
open import Level using (_⊔_)
2929

30+
------------------------------------------------------------------------
31+
-- Structures with 1 unary operation & 1 element
32+
------------------------------------------------------------------------
33+
34+
record IsSuccessorSet (suc# : Op₁ A) (zero# : A) : Set (a ⊔ ℓ) where
35+
field
36+
isEquivalence : IsEquivalence _≈_
37+
suc#-cong : Congruent₁ suc#
38+
39+
open IsEquivalence isEquivalence public
40+
41+
setoid : Setoid a ℓ
42+
setoid = record { isEquivalence = isEquivalence }
43+
3044
------------------------------------------------------------------------
3145
-- Structures with 1 binary operation
3246
------------------------------------------------------------------------

0 commit comments

Comments
 (0)