Skip to content

Commit 276ca7e

Browse files
authored
[ add ] rule of signs for RingWithoutOne (#2761)
* add: rule of signs for `RingWithoutOne` * reset: `CHANGELOG` for v2.4
1 parent 604b1f4 commit 276ca7e

File tree

2 files changed

+20
-6
lines changed

2 files changed

+20
-6
lines changed

CHANGELOG.md

Lines changed: 5 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -47,8 +47,12 @@ New modules
4747
Additions to existing modules
4848
-----------------------------
4949

50+
* In `Algebra.Properties.RingWithoutOne`:
51+
```agda
52+
[-x][-y]≈xy : ∀ x y → - x * - y ≈ x * y
53+
```
54+
5055
* In `Data.Nat.Properties`:
5156
```agda
5257
∸-suc : m ≤ n → suc n ∸ m ≡ suc (n ∸ m)
5358
```
54-

src/Algebra/Properties/RingWithoutOne.agda

Lines changed: 15 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -17,7 +17,7 @@ open import Function.Base using (_$_)
1717
open import Relation.Binary.Reasoning.Setoid setoid
1818

1919
------------------------------------------------------------------------
20-
-- Export properties of abelian groups
20+
-- Re-export abelian group properties for addition
2121

2222
open AbelianGroupProperties +-abelianGroup public
2323
renaming
@@ -36,6 +36,12 @@ open AbelianGroupProperties +-abelianGroup public
3636
; ⁻¹-∙-comm to -‿+-comm
3737
)
3838

39+
x+x≈x⇒x≈0 : x x + x ≈ x x ≈ 0#
40+
x+x≈x⇒x≈0 x eq = +-identityˡ-unique x x eq
41+
42+
------------------------------------------------------------------------
43+
-- Consequences of distributivity
44+
3945
-‿distribˡ-* : x y - (x * y) ≈ - x * y
4046
-‿distribˡ-* x y = sym $ begin
4147
- x * y ≈⟨ +-identityʳ (- x * y) ⟨
@@ -58,17 +64,21 @@ open AbelianGroupProperties +-abelianGroup public
5864
- (x * y) + 0# ≈⟨ +-identityʳ (- (x * y)) ⟩
5965
- (x * y) ∎
6066

61-
x+x≈x⇒x≈0 : x x + x ≈ x x ≈ 0#
62-
x+x≈x⇒x≈0 x eq = +-identityˡ-unique x x eq
67+
[-x][-y]≈xy : x y - x * - y ≈ x * y
68+
[-x][-y]≈xy x y = begin
69+
- x * - y ≈⟨ -‿distribˡ-* x (- y) ⟨
70+
- (x * - y) ≈⟨ -‿cong (-‿distribʳ-* x y) ⟨
71+
- (- (x * y)) ≈⟨ -‿involutive (x * y) ⟩
72+
x * y ∎
6373

6474
x[y-z]≈xy-xz : x y z x * (y - z) ≈ x * y - x * z
6575
x[y-z]≈xy-xz x y z = begin
6676
x * (y - z) ≈⟨ distribˡ x y (- z) ⟩
67-
x * y + x * - z ≈⟨ +-congˡ (sym (-‿distribʳ-* x z)) ⟩
77+
x * y + x * - z ≈⟨ +-congˡ (-‿distribʳ-* x z)
6878
x * y - x * z ∎
6979

7080
[y-z]x≈yx-zx : x y z (y - z) * x ≈ (y * x) - (z * x)
7181
[y-z]x≈yx-zx x y z = begin
7282
(y - z) * x ≈⟨ distribʳ x y (- z) ⟩
73-
y * x + - z * x ≈⟨ +-congˡ (sym (-‿distribˡ-* z x)) ⟩
83+
y * x + - z * x ≈⟨ +-congˡ (-‿distribˡ-* z x)
7484
y * x - z * x ∎

0 commit comments

Comments
 (0)