Skip to content

Commit 672d0e3

Browse files
Add proofs for Ring and commutative semigroup (#1875)
1 parent 7036cfd commit 672d0e3

File tree

3 files changed

+25
-1
lines changed

3 files changed

+25
-1
lines changed

CHANGELOG.md

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1617,6 +1617,7 @@ Other minor changes
16171617
leftSemimedial : LeftSemimedial _∙_
16181618
rightSemimedial : RightSemimedial _∙_
16191619
middleSemimedial : ∀ x y z → (x ∙ y) ∙ (z ∙ x) ≈ (x ∙ z) ∙ (y ∙ x)
1620+
semimedial : Semimedial _∙_
16201621
```
16211622

16221623
* Added new proofs to `Algebra.Properties.Semigroup`:
@@ -1630,6 +1631,9 @@ Other minor changes
16301631
* Added new proofs to `Algebra.Properties.Ring`:
16311632
```agda
16321633
-1*x≈-x : ∀ x → - 1# * x ≈ - x
1634+
x+x≈x⇒x≈0 : ∀ x → x + x ≈ x → x ≈ 0#
1635+
x[y-z]≈xy-xz : ∀ x y z → x * (y - z) ≈ x * y - x * z
1636+
[y-z]x≈yx-zx : ∀ x y z → (y - z) * x ≈ (y * x) - (z * x)
16331637
```
16341638

16351639
* Added new definitions to `Algebra.Structures`:

src/Algebra/Properties/CommutativeSemigroup.agda

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -16,6 +16,7 @@ open CommutativeSemigroup CS
1616

1717
open import Algebra.Definitions _≈_
1818
open import Relation.Binary.Reasoning.Setoid setoid
19+
open import Data.Product
1920

2021
------------------------------------------------------------------------------
2122
-- Re-export the contents of semigroup
@@ -168,3 +169,6 @@ middleSemimedial x y z = begin
168169
x ∙ ((z ∙ y) ∙ x) ≈⟨ ∙-congˡ ( assoc z y x) ⟩
169170
x ∙ (z ∙ (y ∙ x)) ≈⟨ sym (assoc x z ((y ∙ x))) ⟩
170171
(x ∙ z) ∙ (y ∙ x) ∎
172+
173+
semimedial : Semimedial _∙_
174+
semimedial = semimedialˡ , semimedialʳ

src/Algebra/Properties/Ring.agda

Lines changed: 17 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -6,7 +6,7 @@
66

77
{-# OPTIONS --without-K --safe #-}
88

9-
open import Algebra
9+
open import Algebra using (Ring)
1010

1111
module Algebra.Properties.Ring {r₁ r₂} (R : Ring r₁ r₂) where
1212

@@ -15,6 +15,8 @@ open Ring R
1515
import Algebra.Properties.AbelianGroup as AbelianGroupProperties
1616
open import Function.Base using (_$_)
1717
open import Relation.Binary.Reasoning.Setoid setoid
18+
open import Algebra.Definitions _≈_
19+
open import Data.Product
1820

1921
------------------------------------------------------------------------
2022
-- Export properties of abelian groups
@@ -67,3 +69,17 @@ open AbelianGroupProperties +-abelianGroup public
6769
- (1# * x) ≈⟨ -‿cong ( *-identityˡ x ) ⟩
6870
- x ∎
6971

72+
x+x≈x⇒x≈0 : x x + x ≈ x x ≈ 0#
73+
x+x≈x⇒x≈0 x eq = +-identityˡ-unique x x eq
74+
75+
x[y-z]≈xy-xz : x y z x * (y - z) ≈ x * y - x * z
76+
x[y-z]≈xy-xz x y z = begin
77+
x * (y - z) ≈⟨ distribˡ x y (- z) ⟩
78+
x * y + x * - z ≈⟨ +-congˡ (sym (-‿distribʳ-* x z)) ⟩
79+
x * y - x * z ∎
80+
81+
[y-z]x≈yx-zx : x y z (y - z) * x ≈ (y * x) - (z * x)
82+
[y-z]x≈yx-zx x y z = begin
83+
(y - z) * x ≈⟨ distribʳ x y (- z) ⟩
84+
y * x + - z * x ≈⟨ +-congˡ (sym (-‿distribˡ-* z x)) ⟩
85+
y * x - z * x ∎

0 commit comments

Comments
 (0)