|
6 | 6 |
|
7 | 7 | {-# OPTIONS --cubical-compatible --safe #-}
|
8 | 8 |
|
9 |
| --- Disabled to prevent warnings from deprecated names |
10 |
| -{-# OPTIONS --warn=noUserWarning #-} |
11 |
| - |
12 |
| -open import Relation.Binary.Core using (_Preserves_⟶_) |
13 | 9 | open import Relation.Binary.Bundles using (Setoid)
|
14 | 10 |
|
15 | 11 | module Function.Endomorphism.Setoid {c e} (S : Setoid c e) where
|
16 | 12 |
|
17 | 13 | open import Agda.Builtin.Equality
|
18 | 14 |
|
19 |
| -open import Algebra |
20 |
| -open import Algebra.Structures |
21 |
| -open import Algebra.Morphism; open Definitions |
22 |
| -open import Function.Equality using (setoid; _⟶_; id; _∘_; cong) |
23 |
| -open import Data.Nat.Base using (ℕ; _+_); open ℕ |
24 |
| -open import Data.Nat.Properties |
| 15 | +open import Algebra using (Semigroup; Magma; RawMagma; Monoid; RawMonoid) |
| 16 | +open import Algebra.Structures using (IsMagma; IsSemigroup; IsMonoid) |
| 17 | +open import Algebra.Morphism |
| 18 | + using (module Definitions; IsSemigroupHomomorphism; IsMonoidHomomorphism) |
| 19 | +open Definitions using (Homomorphic₂) |
| 20 | +open import Data.Nat.Base using (ℕ; _+_; +-rawMagma; +-0-rawMonoid) |
| 21 | +open ℕ |
| 22 | +open import Data.Nat.Properties using (+-semigroup; +-identityʳ) |
25 | 23 | open import Data.Product.Base using (_,_)
|
26 |
| - |
27 |
| -import Relation.Binary.Indexed.Heterogeneous.Construct.Trivial as Trivial |
| 24 | +open import Function.Bundles using (Func; _⟶ₛ_; _⟨$⟩_) |
| 25 | +open import Function.Construct.Identity using () renaming (function to identity) |
| 26 | +open import Function.Construct.Composition using () renaming (function to _∘_) |
| 27 | +open import Function.Relation.Binary.Setoid.Equality as Eq using (_⇨_) |
| 28 | +open import Level using (Level; _⊔_) |
| 29 | +open import Relation.Binary.Core using (_Preserves_⟶_) |
28 | 30 |
|
29 | 31 | private
|
30 |
| - module E = Setoid (setoid S (Trivial.indexedSetoid S)) |
| 32 | + module E = Setoid (S ⇨ S) |
31 | 33 | open E hiding (refl)
|
32 |
| - |
| 34 | + open Func using (cong) |
33 | 35 | ------------------------------------------------------------------------
|
34 | 36 | -- Basic type and functions
|
35 | 37 |
|
36 | 38 | Endo : Set _
|
37 |
| -Endo = S ⟶ S |
| 39 | +Endo = S ⟶ₛ S |
38 | 40 |
|
39 | 41 | infixr 8 _^_
|
40 | 42 |
|
| 43 | +private |
| 44 | + id : Endo |
| 45 | + id = identity S |
| 46 | + |
41 | 47 | _^_ : Endo → ℕ → Endo
|
42 | 48 | f ^ zero = id
|
43 | 49 | f ^ suc n = f ∘ (f ^ n)
|
44 | 50 |
|
45 | 51 | ^-cong₂ : ∀ f → (f ^_) Preserves _≡_ ⟶ _≈_
|
46 |
| -^-cong₂ f {n} refl = cong (f ^ n) |
| 52 | +^-cong₂ f {n} refl = cong (f ^ n) (Setoid.refl S) |
47 | 53 |
|
48 | 54 | ^-homo : ∀ f → Homomorphic₂ ℕ Endo _≈_ (f ^_) _+_ _∘_
|
49 |
| -^-homo f zero n x≈y = cong (f ^ n) x≈y |
50 |
| -^-homo f (suc m) n x≈y = cong f (^-homo f m n x≈y) |
| 55 | +^-homo f zero n = Setoid.refl S |
| 56 | +^-homo f (suc m) zero = ^-cong₂ f (+-identityʳ m) |
| 57 | +^-homo f (suc m) (suc n) = ^-homo f m (suc n) |
51 | 58 |
|
52 | 59 | ------------------------------------------------------------------------
|
53 | 60 | -- Structures
|
54 | 61 |
|
55 | 62 | ∘-isMagma : IsMagma _≈_ _∘_
|
56 | 63 | ∘-isMagma = record
|
57 | 64 | { isEquivalence = isEquivalence
|
58 |
| - ; ∙-cong = λ g f x → g (f x) |
| 65 | + ; ∙-cong = λ {_} {_} {_} {v} x≈y u≈v → S.trans u≈v (cong v x≈y) |
59 | 66 | }
|
| 67 | + where |
| 68 | + module S = Setoid S |
60 | 69 |
|
61 |
| -∘-magma : Magma _ _ |
| 70 | +∘-magma : Magma (c ⊔ e) (c ⊔ e) |
62 | 71 | ∘-magma = record { isMagma = ∘-isMagma }
|
63 | 72 |
|
64 | 73 | ∘-isSemigroup : IsSemigroup _≈_ _∘_
|
65 | 74 | ∘-isSemigroup = record
|
66 | 75 | { isMagma = ∘-isMagma
|
67 |
| - ; assoc = λ h g f x≈y → cong h (cong g (cong f x≈y)) |
| 76 | + ; assoc = λ h g f {s} → Setoid.refl S |
68 | 77 | }
|
69 | 78 |
|
70 |
| -∘-semigroup : Semigroup _ _ |
| 79 | +∘-semigroup : Semigroup (c ⊔ e) (c ⊔ e) |
71 | 80 | ∘-semigroup = record { isSemigroup = ∘-isSemigroup }
|
72 | 81 |
|
73 | 82 | ∘-id-isMonoid : IsMonoid _≈_ _∘_ id
|
74 | 83 | ∘-id-isMonoid = record
|
75 | 84 | { isSemigroup = ∘-isSemigroup
|
76 |
| - ; identity = cong , cong |
| 85 | + ; identity = (λ f → Setoid.refl S) , (λ _ → Setoid.refl S) |
77 | 86 | }
|
78 | 87 |
|
79 |
| -∘-id-monoid : Monoid _ _ |
| 88 | +∘-id-monoid : Monoid (c ⊔ e) (c ⊔ e) |
80 | 89 | ∘-id-monoid = record { isMonoid = ∘-id-isMonoid }
|
81 | 90 |
|
| 91 | +private |
| 92 | + ∘-rawMagma : RawMagma (c ⊔ e) (c ⊔ e) |
| 93 | + ∘-rawMagma = Semigroup.rawMagma ∘-semigroup |
| 94 | + |
| 95 | + ∘-id-rawMonoid : RawMonoid (c ⊔ e) (c ⊔ e) |
| 96 | + ∘-id-rawMonoid = Monoid.rawMonoid ∘-id-monoid |
| 97 | + |
82 | 98 | ------------------------------------------------------------------------
|
83 | 99 | -- Homomorphism
|
84 | 100 |
|
85 |
| -^-isSemigroupMorphism : ∀ f → IsSemigroupMorphism +-semigroup ∘-semigroup (f ^_) |
86 |
| -^-isSemigroupMorphism f = record |
87 |
| - { ⟦⟧-cong = ^-cong₂ f |
88 |
| - ; ∙-homo = ^-homo f |
| 101 | +^-isSemigroupHomomorphism : ∀ f → IsSemigroupHomomorphism +-rawMagma ∘-rawMagma (f ^_) |
| 102 | +^-isSemigroupHomomorphism f = record |
| 103 | + { isRelHomomorphism = record { cong = ^-cong₂ f } |
| 104 | + ; homo = ^-homo f |
89 | 105 | }
|
90 | 106 |
|
91 |
| -^-isMonoidMorphism : ∀ f → IsMonoidMorphism +-0-monoid ∘-id-monoid (f ^_) |
92 |
| -^-isMonoidMorphism f = record |
93 |
| - { sm-homo = ^-isSemigroupMorphism f |
94 |
| - ; ε-homo = λ x≈y → x≈y |
| 107 | +^-isMonoidHomoorphism : ∀ f → IsMonoidHomomorphism +-0-rawMonoid ∘-id-rawMonoid (f ^_) |
| 108 | +^-isMonoidHomoorphism f = record |
| 109 | + { isMagmaHomomorphism = record |
| 110 | + { isRelHomomorphism = record { cong = ^-cong₂ f } |
| 111 | + ; homo = ^-homo f |
| 112 | + } |
| 113 | + ; ε-homo = Setoid.refl S |
95 | 114 | }
|
| 115 | + |
0 commit comments