open import 1Lab.Prelude open import Algebra.Semigroup module Algebra.Monoid where
Monoids🔗
A monoid is a semigroup equipped with a two-sided identity element: An element such that . For any particular choice of binary operator , if a two-sided identity exists, then it is unique; In this sense, “being a monoid” could be considered an “axiom” that semigroups may satisfy.
However, since semigroup homomorphisms do not automatically preserve the identity element1, it is part of the type signature for is-monoid, being considered structure that a semigroup may be equipped with.
record is-monoid (id : A) (_⋆_ : A → A → A) : Type (level-of A) where field has-is-semigroup : is-semigroup _⋆_ open is-semigroup has-is-semigroup public field idl : {x : A} → id ⋆ x ≡ x idr : {x : A} → x ⋆ id ≡ x open is-monoid public
The condition of defining a monoid is a proposition, so that we may safely decompose monoids as the structure , which has to satisfy the property of being a monoid.
private unquoteDecl eqv = declare-record-iso eqv (quote is-monoid) instance H-Level-is-monoid : ∀ {id : A} {_⋆_ : A → A → A} {n} → H-Level (is-monoid id _⋆_) (suc n) H-Level-is-monoid = prop-instance λ x → let open is-monoid x in is-hlevel≃ 1 (Iso→Equiv eqv e⁻¹) (hlevel 1) x
A monoid structure on a type is given by the choice of identity element, the choice of binary operation, and the witness that these choices form a monoid. A Monoid, then, is a type with a monoid structure.
record Monoid-on (A : Type ℓ) : Type ℓ where field identity : A _⋆_ : A → A → A has-is-monoid : is-monoid identity _⋆_ open is-monoid has-is-monoid public Monoid : (ℓ : Level) → Type (lsuc ℓ) Monoid ℓ = Σ (Monoid-on {ℓ = ℓ}) open Monoid-on
There is also a predicate which witnesses when an equivalence between monoids is a monoid homomorphism. It has to preserve the identity, and commute with the multiplication:
record Monoid-hom (A B : Σ (Monoid-on {ℓ = ℓ})) (e : A .fst → B .fst) : Type ℓ where private module A = Monoid-on (A .snd) module B = Monoid-on (B .snd) field pres-id : e A.identity ≡ B.identity pres-⋆ : (x y : A .fst) → e (x A.⋆ y) ≡ e x B.⋆ e y open Monoid-hom Monoid≃ : (A B : Σ (Monoid-on {ℓ = ℓ})) (e : A .fst ≃ B .fst) → Type _ Monoid≃ A B (e , _) = Monoid-hom A B e
We automatically derive a proof that Monoid-on is univalent for the structure induced by Monoid≃:
Monoid-univalent : is-univalent {ℓ = ℓ} (HomT→Str Monoid≃) Monoid-univalent {ℓ = ℓ} = Derive-univalent-record (record-desc (Monoid-on {ℓ = ℓ}) Monoid≃ (record: field[ identity by pres-id ] field[ _⋆_ by pres-⋆ ] axiom[ has-is-monoid by (λ _ → hlevel 1) ]))
From this, we automatically get a specialisation of the SIP for Monoids, which says that identification of monoids is monoid isomorphism.
Monoid≡ : {A B : Monoid ℓ} → (A ≃[ HomT→Str Monoid≃ ] B) ≃ (A ≡ B) Monoid≡ = SIP Monoid-univalent
Relationships to Unital Magmas🔗
open import Algebra.Magma.Unital
By definition, every monoid is exactly a unital magma that is also a semigroup. However, adopting this as a definition yields several issues especially when it comes to metaprogramming, which is why this is instead expressed by explicitly proving the implications between the properties.
First, we show that every monoid is a unital magma:
module _ {id : A} {_⋆_ : A → A → A} where is-monoid→is-unital-magma : is-monoid id _⋆_ → is-unital-magma id _⋆_ is-monoid→is-unital-magma mon .has-is-magma = mon .has-is-semigroup .has-is-magma is-monoid→is-unital-magma mon .idl = mon .idl is-monoid→is-unital-magma mon .idr = mon .idr
“Reshuffling” the record fields also allows us to show the reverse direction, namely, that every unital semigroup is a monoid.
is-unital-magma→is-semigroup→is-monoid : is-unital-magma id _⋆_ → is-semigroup _⋆_ → is-monoid id _⋆_ is-unital-magma→is-semigroup→is-monoid uni sem .has-is-semigroup = sem is-unital-magma→is-semigroup→is-monoid uni sem .idl = uni .idl is-unital-magma→is-semigroup→is-monoid uni sem .idr = uni .idr
Inverses🔗
A useful application of the monoid laws is in showing that having an inverse is a proprety of a specific element, not structure on that element. To make this precise: Let be an element of a monoid, say ; If there are and such that and , then .
monoid-inverse-unique : ∀ {1M : A} {_⋆_ : A → A → A} → (m : is-monoid 1M _⋆_) → (e x y : A) → (x ⋆ e ≡ 1M) → (e ⋆ y ≡ 1M) → x ≡ y
This is a happy theorem where stating the assumptions takes as many lines as the proof, which is a rather boring calculation. Since is an identity for , we can freely multiply by to “sneak in” a :
monoid-inverse-unique {1M = 1M} {_⋆_} m e x y li1 ri2 = x ≡⟨ sym (m .idr) ⟩≡ x ⋆ 1M ≡⟨ ap₂ _⋆_ refl (sym ri2) ⟩≡ x ⋆ (e ⋆ y) ≡⟨ m .associative ⟩≡ (x ⋆ e) ⋆ y ≡⟨ ap₂ _⋆_ li1 refl ⟩≡ 1M ⋆ y ≡⟨ m .idl ⟩≡ y ∎
Counterexample: The map which sends everything to zero is a semigroup homomorphism, but does not preserve the unit of .↩︎