Skip to content

Commit 08f24a6

Browse files
Port two Endomorphism submodules over to new Function hierarchy (#2342)
* port over two modules * and add to CHANGELOG * fix whitespace * fix warning: it was pointing to a record that did not exist. * fix things as per Matthew's review - though this remains a breaking change. * take care of comments from James. * adjust CHANGELOG for what will be implemented shortly * Revert "take care of comments from James." This reverts commit 93e9e0f. * Revert "fix things as per Matthew's review - though this remains a breaking change." This reverts commit d1cae72. * Revert "fix whitespace" This reverts commit 81230ec. * Revert "port over two modules" This reverts commit 6619f11. * rename these * fix tiny merge issue * get deprecations right (remove where not needed, make more global where needed) * style guide - missing blank lines * fix a bad merge * fixed deprecations * fix #2394 * minor tweaks --------- Co-authored-by: James McKinna <[email protected]>
1 parent c0fafe9 commit 08f24a6

File tree

6 files changed

+266
-11
lines changed

6 files changed

+266
-11
lines changed

CHANGELOG.md

+16-2
Original file line numberDiff line numberDiff line change
@@ -51,6 +51,11 @@ Deprecated modules
5151
* `Data.List.Relation.Binary.Sublist.Propositional.Disjoint` deprecated in favour of
5252
`Data.List.Relation.Binary.Sublist.Propositional.Slice`.
5353

54+
* The modules `Function.Endomorphism.Propositional` and
55+
`Function.Endomorphism.Setoid` that use the old `Function`
56+
hierarchy. Use `Function.Endo.Propositional` and
57+
`Function.Endo.Setoid` instead.
58+
5459
Deprecated names
5560
----------------
5661

@@ -167,6 +172,11 @@ New modules
167172
indexedSetoid : {A : Set a} → IndexedSetoid ℕ a _
168173
```
169174

175+
* The modules `Function.Endo.Propositional` and
176+
`Function.Endo.Setoid` are new but are actually proper ports of
177+
`Function.Endomorphism.Propositional` and
178+
`Function.Endomorphism.Setoid`.
179+
170180
* `Function.Relation.Binary.Equality`
171181
```agda
172182
setoid : Setoid a₁ a₂ → Setoid b₁ b₂ → Setoid _ _
@@ -292,13 +302,17 @@ Additions to existing modules
292302
rawModule : RawModule R c ℓ
293303
```
294304

295-
* In `Algebra.Morphism.Structures`
305+
* In `Algebra.Morphism.Structures`:
296306
```agda
297307
module SuccessorSetMorphisms (N₁ : RawSuccessorSet a ℓ₁) (N₂ : RawSuccessorSet b ℓ₂) where
298308
record IsSuccessorSetHomomorphism (⟦_⟧ : N₁.Carrier → N₂.Carrier) : Set _
299309
record IsSuccessorSetMonomorphism (⟦_⟧ : N₁.Carrier → N₂.Carrier) : Set _
300-
record IsSuccessorSetIsomorphism (⟦_⟧ : N₁.Carrier → N₂.Carrier) : Set _
310+
record IsSuccessorSetIsomorphism (⟦_⟧ : N₁.Carrier → N₂.Carrier) : Set _
301311
312+
IsSemigroupHomomorphism : (A → B) → Set _
313+
IsSemigroupMonomorphism : (A → B) → Set _
314+
IsSemigroupIsomorphism : (A → B) → Set _
315+
```
302316
* In `Algebra.Properties.AbelianGroup`:
303317
```
304318
⁻¹-anti-homo‿- : (x - y) ⁻¹ ≈ y - x

src/Function/Endo/Propositional.agda

+121
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,121 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Endomorphisms on a Set
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
module Function.Endo.Propositional {a} (A : Set a) where
10+
11+
open import Algebra using (Semigroup; Magma; RawMagma; Monoid; RawMonoid)
12+
open import Algebra.Core
13+
import Algebra.Definitions.RawMonoid as RawMonoidDefinitions
14+
import Algebra.Properties.Monoid.Mult as MonoidMultProperties
15+
open import Algebra.Structures using (IsMagma; IsSemigroup; IsMonoid)
16+
open import Algebra.Morphism
17+
using (module Definitions; IsMagmaHomomorphism; IsMonoidHomomorphism)
18+
open Definitions using (Homomorphic₂)
19+
20+
open import Data.Nat.Base using (ℕ; zero; suc; _+_; +-rawMagma; +-0-rawMonoid)
21+
open import Data.Nat.Properties using (+-0-monoid; +-semigroup)
22+
open import Data.Product.Base using (_,_)
23+
24+
open import Function.Base using (id; _∘′_; _∋_; flip)
25+
open import Function.Bundles using (Func; _⟶ₛ_; _⟨$⟩_)
26+
open import Relation.Binary.Core using (_Preserves_⟶_)
27+
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong; cong₂)
28+
import Relation.Binary.PropositionalEquality.Properties as ≡
29+
30+
import Function.Endo.Setoid (≡.setoid A) as Setoid
31+
32+
------------------------------------------------------------------------
33+
-- Basic type and raw bundles
34+
35+
Endo : Set a
36+
Endo = A A
37+
38+
private
39+
40+
_∘_ : Op₂ Endo
41+
_∘_ = _∘′_
42+
43+
∘-id-rawMonoid : RawMonoid a a
44+
∘-id-rawMonoid = record { Carrier = Endo; _≈_ = _≡_ ; _∙_ = _∘_ ; ε = id }
45+
46+
open RawMonoid ∘-id-rawMonoid
47+
using ()
48+
renaming (rawMagma to ∘-rawMagma)
49+
50+
51+
------------------------------------------------------------------------
52+
-- Conversion back and forth with the Setoid-based notion of Endomorphism
53+
54+
fromSetoidEndo : Setoid.Endo Endo
55+
fromSetoidEndo = _⟨$⟩_
56+
57+
toSetoidEndo : Endo Setoid.Endo
58+
toSetoidEndo f = record
59+
{ to = f
60+
; cong = cong f
61+
}
62+
63+
------------------------------------------------------------------------
64+
-- Structures
65+
66+
∘-isMagma : IsMagma _≡_ _∘_
67+
∘-isMagma = record
68+
{ isEquivalence = ≡.isEquivalence
69+
; ∙-cong = cong₂ _∘_
70+
}
71+
72+
∘-magma : Magma _ _
73+
∘-magma = record { isMagma = ∘-isMagma }
74+
75+
∘-isSemigroup : IsSemigroup _≡_ _∘_
76+
∘-isSemigroup = record
77+
{ isMagma = ∘-isMagma
78+
; assoc = λ _ _ _ refl
79+
}
80+
81+
∘-semigroup : Semigroup _ _
82+
∘-semigroup = record { isSemigroup = ∘-isSemigroup }
83+
84+
∘-id-isMonoid : IsMonoid _≡_ _∘_ id
85+
∘-id-isMonoid = record
86+
{ isSemigroup = ∘-isSemigroup
87+
; identity = (λ _ refl) , (λ _ refl)
88+
}
89+
90+
∘-id-monoid : Monoid _ _
91+
∘-id-monoid = record { isMonoid = ∘-id-isMonoid }
92+
93+
------------------------------------------------------------------------
94+
-- n-th iterated composition
95+
96+
infixr 8 _^_
97+
98+
_^_ : Endo Endo
99+
_^_ = flip _×_ where open RawMonoidDefinitions ∘-id-rawMonoid using (_×_)
100+
101+
------------------------------------------------------------------------
102+
-- Homomorphism
103+
104+
module _ (f : Endo) where
105+
106+
open MonoidMultProperties ∘-id-monoid using (×-homo-+)
107+
108+
^-homo : Homomorphic₂ ℕ Endo _≡_ (f ^_) _+_ _∘_
109+
^-homo = ×-homo-+ f
110+
111+
^-isMagmaHomomorphism : IsMagmaHomomorphism +-rawMagma ∘-rawMagma (f ^_)
112+
^-isMagmaHomomorphism = record
113+
{ isRelHomomorphism = record { cong = cong (f ^_) }
114+
; homo = ^-homo
115+
}
116+
117+
^-isMonoidHomomorphism : IsMonoidHomomorphism +-0-rawMonoid ∘-id-rawMonoid (f ^_)
118+
^-isMonoidHomomorphism = record
119+
{ isMagmaHomomorphism = ^-isMagmaHomomorphism
120+
; ε-homo = refl
121+
}

src/Function/Endo/Setoid.agda

+117
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,117 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Endomorphisms on a Setoid
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
open import Relation.Binary.Bundles using (Setoid)
10+
11+
module Function.Endo.Setoid {c e} (S : Setoid c e) where
12+
13+
open import Agda.Builtin.Equality using (_≡_)
14+
15+
open import Algebra using (Semigroup; Magma; RawMagma; Monoid; RawMonoid)
16+
import Algebra.Definitions.RawMonoid as RawMonoidDefinitions
17+
import Algebra.Properties.Monoid.Mult as MonoidMultProperties
18+
open import Algebra.Structures using (IsMagma; IsSemigroup; IsMonoid)
19+
open import Algebra.Morphism
20+
using (module Definitions; IsMagmaHomomorphism; IsMonoidHomomorphism)
21+
open Definitions using (Homomorphic₂)
22+
open import Data.Nat.Base using (ℕ; zero; suc; _+_; +-rawMagma; +-0-rawMonoid)
23+
open import Data.Nat.Properties using (+-semigroup; +-identityʳ)
24+
open import Data.Product.Base using (_,_)
25+
open import Function.Bundles using (Func; _⟶ₛ_; _⟨$⟩_)
26+
open import Function.Construct.Identity using () renaming (function to identity)
27+
open import Function.Construct.Composition using () renaming (function to _∘_)
28+
open import Function.Relation.Binary.Setoid.Equality as Eq using (_⇨_)
29+
open import Level using (Level; _⊔_)
30+
open import Relation.Binary.Core using (_Preserves_⟶_)
31+
32+
private
33+
open module E = Setoid (S ⇨ S) hiding (refl)
34+
module S = Setoid S
35+
open Func using (cong)
36+
37+
38+
------------------------------------------------------------------------
39+
-- Basic type and raw bundles
40+
41+
Endo : Set _
42+
Endo = S ⟶ₛ S
43+
44+
private
45+
id : Endo
46+
id = identity S
47+
48+
∘-id-rawMonoid : RawMonoid (c ⊔ e) (c ⊔ e)
49+
∘-id-rawMonoid = record { Carrier = Endo; _≈_ = _≈_ ; _∙_ = _∘_ ; ε = id }
50+
51+
open RawMonoid ∘-id-rawMonoid
52+
using ()
53+
renaming (rawMagma to ∘-rawMagma)
54+
55+
--------------------------------------------------------------
56+
-- Structures
57+
58+
∘-isMagma : IsMagma _≈_ _∘_
59+
∘-isMagma = record
60+
{ isEquivalence = isEquivalence
61+
; ∙-cong = λ {_} {_} {_} {v} x≈y u≈v S.trans u≈v (cong v x≈y)
62+
}
63+
64+
∘-magma : Magma (c ⊔ e) (c ⊔ e)
65+
∘-magma = record { isMagma = ∘-isMagma }
66+
67+
∘-isSemigroup : IsSemigroup _≈_ _∘_
68+
∘-isSemigroup = record
69+
{ isMagma = ∘-isMagma
70+
; assoc = λ _ _ _ S.refl
71+
}
72+
73+
∘-semigroup : Semigroup (c ⊔ e) (c ⊔ e)
74+
∘-semigroup = record { isSemigroup = ∘-isSemigroup }
75+
76+
∘-id-isMonoid : IsMonoid _≈_ _∘_ id
77+
∘-id-isMonoid = record
78+
{ isSemigroup = ∘-isSemigroup
79+
; identity = (λ _ S.refl) , (λ _ S.refl)
80+
}
81+
82+
∘-id-monoid : Monoid (c ⊔ e) (c ⊔ e)
83+
∘-id-monoid = record { isMonoid = ∘-id-isMonoid }
84+
85+
------------------------------------------------------------------------
86+
-- -- n-th iterated composition
87+
88+
infixr 8 _^_
89+
90+
_^_ : Endo Endo
91+
f ^ n = n × f where open RawMonoidDefinitions ∘-id-rawMonoid using (_×_)
92+
93+
------------------------------------------------------------------------
94+
-- Homomorphism
95+
96+
module _ (f : Endo) where
97+
98+
open MonoidMultProperties ∘-id-monoid using (×-congˡ; ×-homo-+)
99+
100+
^-cong₂ : (f ^_) Preserves _≡_ ⟶ _≈_
101+
^-cong₂ = ×-congˡ {f}
102+
103+
^-homo : Homomorphic₂ ℕ Endo _≈_ (f ^_) _+_ _∘_
104+
^-homo = ×-homo-+ f
105+
106+
^-isMagmaHomomorphism : IsMagmaHomomorphism +-rawMagma ∘-rawMagma (f ^_)
107+
^-isMagmaHomomorphism = record
108+
{ isRelHomomorphism = record { cong = ^-cong₂ }
109+
; homo = ^-homo
110+
}
111+
112+
^-isMonoidHomomorphism : IsMonoidHomomorphism +-0-rawMonoid ∘-id-rawMonoid (f ^_)
113+
^-isMonoidHomomorphism = record
114+
{ isMagmaHomomorphism = ^-isMagmaHomomorphism
115+
; ε-homo = S.refl
116+
}
117+

src/Function/Endomorphism/Propositional.agda

+6-4
Original file line numberDiff line numberDiff line change
@@ -1,16 +1,18 @@
11
------------------------------------------------------------------------
22
-- The Agda standard library
33
--
4-
-- Endomorphisms on a Set
4+
-- This module is DEPRECATED.
55
------------------------------------------------------------------------
66

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

9-
-- Disabled to prevent warnings from deprecated names
10-
{-# OPTIONS --warn=noUserWarning #-}
11-
129
module Function.Endomorphism.Propositional {a} (A : Set a) where
1310

11+
{-# WARNING_ON_IMPORT
12+
"Function.Endomorphism.Propositional was deprecated in v2.1.
13+
Use Function.Endo.Propositional instead."
14+
#-}
15+
1416
open import Algebra
1517
open import Algebra.Morphism; open Definitions
1618

src/Function/Endomorphism/Setoid.agda

+6-4
Original file line numberDiff line numberDiff line change
@@ -1,19 +1,21 @@
11
------------------------------------------------------------------------
22
-- The Agda standard library
33
--
4-
-- Endomorphisms on a Setoid
4+
-- This module is DEPRECATED.
55
------------------------------------------------------------------------
66

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

9-
-- Disabled to prevent warnings from deprecated names
10-
{-# OPTIONS --warn=noUserWarning #-}
11-
129
open import Relation.Binary.Core using (_Preserves_⟶_)
1310
open import Relation.Binary.Bundles using (Setoid)
1411

1512
module Function.Endomorphism.Setoid {c e} (S : Setoid c e) where
1613

14+
{-# WARNING_ON_IMPORT
15+
"Function.Endomorphism.Setoid was deprecated in v2.1.
16+
Use Function.Endo.Setoid instead."
17+
#-}
18+
1719
open import Agda.Builtin.Equality
1820

1921
open import Algebra

src/Function/Equality.agda

-1
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,6 @@
55
------------------------------------------------------------------------
66

77
{-# OPTIONS --cubical-compatible --safe #-}
8-
{-# OPTIONS --warn=noUserWarning #-}
98

109
module Function.Equality where
1110

0 commit comments

Comments
 (0)