Skip to content

Commit 7367bae

Browse files
jmougeotjamesmckinnaJacquesCarette
authored
[Add] Consequences of identity for monoids (#2692)
* Adds reasonig combinator for semigroup * Adds reasonig combinator for semigroup * Adds reasonig combinator for semigroup * Adds reasonig combinator for semigroup * Add some more missing reasoning combinators * add module Extends * rename SemiGroup to Semigroup * fix-whitespace * Update CHANGELOG.md Co-authored-by: jamesmckinna <[email protected]> * New names * improve syntx * fix-whitespace * Name changes * Names change * Space * Proof of assoc with PUshes and Pulles * Proof of assoc with PUshes and Pulles * white space * Update src/Algebra/Properties/Semigroup/Reasoning.agda Co-authored-by: jamesmckinna <[email protected]> * Reasoning to Semigroup and explicit variables * Update CHANGELOG.md Co-authored-by: jamesmckinna <[email protected]> * Add reasonining combinators for monoids * chore: move imports around in Algebra.Definitions merge * Syntax modification * rebase * update semigroup * One line proof * move importation * fix-whitespace * Add reasonining combinators for monoids * chore: move imports around in Algebra.Definitions merge * Syntax modification * rebase merge * move importation * fix-whitespace * explicit varaibles * used of semigroup propreties * update * Update CHANGELOG.md Co-authored-by: jamesmckinna <[email protected]> * merge and proof refactoring * Changelog update * Update src/Algebra/Properties/Monoid.agda Co-authored-by: jamesmckinna <[email protected]> * Update src/Algebra/Properties/Monoid.agda Co-authored-by: jamesmckinna <[email protected]> * Update src/Algebra/Properties/Monoid.agda * rename id to ε --------- Co-authored-by: jamesmckinna <[email protected]> Co-authored-by: Jacques Carette <[email protected]>
1 parent 5c64117 commit 7367bae

File tree

2 files changed

+71
-0
lines changed

2 files changed

+71
-0
lines changed

CHANGELOG.md

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -186,6 +186,10 @@ New modules
186186
Additions to existing modules
187187
-----------------------------
188188

189+
* In `Algebra.Properties.Monoid` adding consequences for identity for monoids
190+
191+
* In `Algebra.Properties.Semigroup` adding consequences for associativity for semigroups
192+
189193
* In `Algebra.Consequences.Base`:
190194
```agda
191195
module Congruence (_≈_ : Rel A ℓ) (cong : Congruent₂ _≈_ _∙_) (refl : Reflexive _≈_)

src/Algebra/Properties/Monoid.agda

Lines changed: 67 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,67 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Equational reasoning for monoids
5+
-- (Utilities for identity and cancellation reasoning, extending semigroup reasoning)
6+
------------------------------------------------------------------------
7+
8+
{-# OPTIONS --cubical-compatible --safe #-}
9+
10+
open import Algebra.Bundles using (Monoid)
11+
open import Function using (_∘_)
12+
13+
module Algebra.Properties.Monoid {o ℓ} (M : Monoid o ℓ) where
14+
15+
open Monoid M
16+
using (Carrier; _∙_; _≈_; setoid; isMagma; semigroup; ε; sym; identityˡ
17+
; identityʳ ; ∙-cong; refl; assoc; ∙-congˡ; ∙-congʳ; trans)
18+
open import Relation.Binary.Reasoning.Setoid setoid
19+
20+
open import Algebra.Properties.Semigroup semigroup public
21+
22+
private
23+
variable
24+
a b c d : Carrier
25+
26+
ε-unique : a ( b b ∙ a ≈ b) a ≈ ε
27+
ε-unique a b∙a≈b = trans (sym (identityˡ a)) (b∙a≈b ε)
28+
29+
ε-comm : a a ∙ ε ≈ ε ∙ a
30+
ε-comm a = trans (identityʳ a) (sym (identityˡ a))
31+
32+
module _ (a≈ε : a ≈ ε) where
33+
elimʳ : b b ∙ a ≈ b
34+
elimʳ = trans (∙-congˡ a≈ε) ∘ identityʳ
35+
36+
elimˡ : b a ∙ b ≈ b
37+
elimˡ = trans (∙-congʳ a≈ε) ∘ identityˡ
38+
39+
introʳ : b b ≈ b ∙ a
40+
introʳ = sym ∘ elimʳ
41+
42+
introˡ : b b ≈ a ∙ b
43+
introˡ = sym ∘ elimˡ
44+
45+
introᶜ : c b ∙ c ≈ b ∙ (a ∙ c)
46+
introᶜ c = trans (∙-congˡ (sym (identityˡ c))) (∙-congˡ (∙-congʳ (sym a≈ε)))
47+
48+
module _ (inv : a ∙ c ≈ ε) where
49+
50+
cancelʳ : b (b ∙ a) ∙ c ≈ b
51+
cancelʳ b = trans (uv≈w⇒xu∙v≈xw inv b) (identityʳ b)
52+
53+
cancelˡ : b a ∙ (c ∙ b) ≈ b
54+
cancelˡ b = trans (uv≈w⇒u∙vx≈wx inv b) (identityˡ b)
55+
56+
insertˡ : b b ≈ a ∙ (c ∙ b)
57+
insertˡ = sym ∘ cancelˡ
58+
59+
insertʳ : b b ≈ (b ∙ a) ∙ c
60+
insertʳ = sym ∘ cancelʳ
61+
62+
cancelᶜ : b d (b ∙ a) ∙ (c ∙ d) ≈ b ∙ d
63+
cancelᶜ b d = trans (uv≈w⇒xu∙vy≈x∙wy inv b d) (∙-congˡ (identityˡ d))
64+
65+
insertᶜ : b d b ∙ d ≈ (b ∙ a) ∙ (c ∙ d)
66+
insertᶜ = λ b d sym (cancelᶜ b d)
67+

0 commit comments

Comments
 (0)