Skip to content

Commit 4b50bf3

Browse files
[ add ] Algebra.Definitions.Central plus Consequences for Identity and Zero (#2873)
* add: `Algebra.Definitions.Central` * add: `Identity` and `Zero` elements are `Central` * Update src/Algebra/Definitions.agda Co-authored-by: Jacques Carette <carette@mcmaster.ca> * fix: line break --------- Co-authored-by: Jacques Carette <carette@mcmaster.ca>
1 parent 38f22ec commit 4b50bf3

File tree

3 files changed

+23
-0
lines changed

3 files changed

+23
-0
lines changed

CHANGELOG.md

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -118,12 +118,21 @@ Additions to existing modules
118118
```agda
119119
binomial-expansion : Associative _∙_ → _◦_ DistributesOver _∙_ →
120120
∀ w x y z → ((w ∙ x) ◦ (y ∙ z)) ≡ ((((w ◦ y) ∙ (w ◦ z)) ∙ (x ◦ y)) ∙ (x ◦ z))
121+
identity⇒central : Identity e _∙_ → Central _∙_ e
122+
zero⇒central : Zero e _∙_ → Central _∙_ e
121123
```
122124

123125
* In `Algebra.Consequences.Setoid`:
124126
```agda
125127
binomial-expansion : Congruent₂ _∙_ → Associative _∙_ → _◦_ DistributesOver _∙_ →
126128
∀ w x y z → ((w ∙ x) ◦ (y ∙ z)) ≈ ((((w ◦ y) ∙ (w ◦ z)) ∙ (x ◦ y)) ∙ (x ◦ z))
129+
identity⇒central : Identity e _∙_ → Central _∙_ e
130+
zero⇒central : Zero e _∙_ → Central _∙_ e
131+
```
132+
133+
* In `Algebra.Definitions`:
134+
```agda
135+
Central : Op₂ A → A → Set _
127136
```
128137

129138
* In `Algebra.Lattice.Properties.BooleanAlgebra.XorRing`:

src/Algebra/Consequences/Setoid.agda

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -184,6 +184,15 @@ module _ {_∙_ : Op₂ A} (comm : Commutative _∙_) {e : A} where
184184
x ∙ z ≈⟨ comm x z ⟩
185185
z ∙ x ∎
186186

187+
module _ {_∙_ : Op₂ A} {e : A} where
188+
189+
identity⇒central : Identity e _∙_ Central _∙_ e
190+
identity⇒central (identityˡ , identityʳ) x = trans (identityˡ x) (sym (identityʳ x))
191+
192+
zero⇒central : Zero e _∙_ Central _∙_ e
193+
zero⇒central (zeroˡ , zeroʳ) x = trans (zeroˡ x) (sym (zeroʳ x))
194+
195+
187196
------------------------------------------------------------------------
188197
-- Group-like structures
189198

src/Algebra/Definitions.agda

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -50,6 +50,11 @@ Associative _∙_ = ∀ x y z → ((x ∙ y) ∙ z) ≈ (x ∙ (y ∙ z))
5050
Commutative : Op₂ A Set _
5151
Commutative _∙_ = x y (x ∙ y) ≈ (y ∙ x)
5252

53+
-- An element is called `Central` for a binary operation
54+
-- if it commutes with all other elements.
55+
Central : Op₂ A A Set _
56+
Central _∙_ x = y (x ∙ y) ≈ (y ∙ x)
57+
5358
LeftIdentity : A Op₂ A Set _
5459
LeftIdentity e _∙_ = x (e ∙ x) ≈ x
5560

0 commit comments

Comments
 (0)