|
| 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 | + |
0 commit comments