|
| 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