Skip to content

Commit b6ac7e0

Browse files
MatthewDaggittandreasabel
authored andcommitted
Fix #2195 by removing redundant zero from IsRing (#2209)
* Fix #2195 by removing redundant zero from IsRing * Moved proofs eariler to IsSemiringWithoutOne * Updated CHANGELOG * Fix bug * Refactored Properties.Ring * Fix renaming
1 parent 67afe7c commit b6ac7e0

File tree

12 files changed

+98
-151
lines changed

12 files changed

+98
-151
lines changed

CHANGELOG.md

+6
Original file line numberDiff line numberDiff line change
@@ -33,6 +33,12 @@ Highlights
3333
Bug-fixes
3434
---------
3535

36+
* In `Algebra.Structures` the records `IsRing` and `IsRingWithoutOne` contained an unnecessary field
37+
`zero : RightZero 0# *`, which could be derived from the other ring axioms.
38+
Consequently this field has been removed from the record, and the record
39+
`IsRingWithoutAnnihilatingZero` in `Algebra.Structures.Biased` has been
40+
deprecated as it is now identical to is `IsRing`.
41+
3642
* In `Algebra.Definitions.RawSemiring` the record `Prime` did not
3743
enforce that the number was not divisible by `1#`. To fix this
3844
`p∤1 : p ∤ 1#` has been added as a field.

src/Algebra/Bundles.agda

+3
Original file line numberDiff line numberDiff line change
@@ -944,6 +944,9 @@ record Ring c ℓ : Set (suc (c ⊔ ℓ)) where
944944
+-abelianGroup : AbelianGroup _ _
945945
+-abelianGroup = record { isAbelianGroup = +-isAbelianGroup }
946946

947+
ringWithoutOne : RingWithoutOne _ _
948+
ringWithoutOne = record { isRingWithoutOne = isRingWithoutOne }
949+
947950
semiring : Semiring _ _
948951
semiring = record { isSemiring = isSemiring }
949952

src/Algebra/Construct/DirectProduct.agda

-3
Original file line numberDiff line numberDiff line change
@@ -335,8 +335,6 @@ ringWithoutOne R S = record
335335
; *-assoc = Semigroup.assoc (semigroup R.*-semigroup S.*-semigroup)
336336
; distrib = (λ x y z (R.distribˡ , S.distribˡ) <*> x <*> y <*> z)
337337
, (λ x y z (R.distribʳ , S.distribʳ) <*> x <*> y <*> z)
338-
; zero = uncurry (λ x y R.zeroˡ x , S.zeroˡ y)
339-
, uncurry (λ x y R.zeroʳ x , S.zeroʳ y)
340338
}
341339

342340
} where module R = RingWithoutOne R; module S = RingWithoutOne S
@@ -389,7 +387,6 @@ ring R S = record
389387
; *-assoc = Semiring.*-assoc Semi
390388
; *-identity = Semiring.*-identity Semi
391389
; distrib = Semiring.distrib Semi
392-
; zero = Semiring.zero Semi
393390
}
394391
}
395392
where

src/Algebra/Lattice/Properties/BooleanAlgebra.agda

-1
Original file line numberDiff line numberDiff line change
@@ -516,7 +516,6 @@ module XorRing
516516
; *-assoc = ∧-assoc
517517
; *-identity = ∧-identity
518518
; distrib = ∧-distrib-⊕
519-
; zero = ∧-zero
520519
}
521520

522521
⊕-∧-isCommutativeRing : IsCommutativeRing _⊕_ _∧_ id ⊥ ⊤

src/Algebra/Morphism/RingMonomorphism.agda

-1
Original file line numberDiff line numberDiff line change
@@ -150,7 +150,6 @@ isRing isRing = record
150150
; *-assoc = *-assoc R.*-isMagma R.*-assoc
151151
; *-identity = *-identity R.*-isMagma R.*-identity
152152
; distrib = distrib R.+-isGroup R.*-isMagma R.distrib
153-
; zero = zero R.+-isGroup R.*-isMagma R.zero
154153
} where module R = IsRing isRing
155154

156155
isCommutativeRing : IsCommutativeRing _≈₂_ _⊕_ _⊛_ ⊝_ 0#₂ 1#₂

src/Algebra/Properties/Ring.agda

+5-57
Original file line numberDiff line numberDiff line change
@@ -12,73 +12,21 @@ module Algebra.Properties.Ring {r₁ r₂} (R : Ring r₁ r₂) where
1212

1313
open Ring R
1414

15-
import Algebra.Properties.AbelianGroup as AbelianGroupProperties
15+
import Algebra.Properties.RingWithoutOne as RingWithoutOneProperties
1616
open import Function.Base using (_$_)
1717
open import Relation.Binary.Reasoning.Setoid setoid
1818
open import Algebra.Definitions _≈_
1919

2020
------------------------------------------------------------------------
21-
-- Export properties of abelian groups
21+
-- Export properties of rings without a 1#.
2222

23-
open AbelianGroupProperties +-abelianGroup public
24-
renaming
25-
( ε⁻¹≈ε to -0#≈0#
26-
; ∙-cancelˡ to +-cancelˡ
27-
; ∙-cancelʳ to +-cancelʳ
28-
; ∙-cancel to +-cancel
29-
; ⁻¹-involutive to -‿involutive
30-
; ⁻¹-injective to -‿injective
31-
; ⁻¹-anti-homo-∙ to -‿anti-homo-+
32-
; identityˡ-unique to +-identityˡ-unique
33-
; identityʳ-unique to +-identityʳ-unique
34-
; identity-unique to +-identity-unique
35-
; inverseˡ-unique to +-inverseˡ-unique
36-
; inverseʳ-unique to +-inverseʳ-unique
37-
; ⁻¹-∙-comm to -‿+-comm
38-
)
23+
open RingWithoutOneProperties ringWithoutOne public
3924

4025
------------------------------------------------------------------------
41-
-- Properties of -_
42-
43-
-‿distribˡ-* : x y - (x * y) ≈ - x * y
44-
-‿distribˡ-* x y = sym $ begin
45-
- x * y ≈⟨ sym $ +-identityʳ _ ⟩
46-
- x * y + 0# ≈⟨ +-congˡ $ sym (-‿inverseʳ _) ⟩
47-
- x * y + (x * y + - (x * y)) ≈⟨ sym $ +-assoc _ _ _ ⟩
48-
- x * y + x * y + - (x * y) ≈⟨ +-congʳ $ sym (distribʳ _ _ _) ⟩
49-
(- x + x) * y + - (x * y) ≈⟨ +-congʳ $ *-congʳ $ -‿inverseˡ _ ⟩
50-
0# * y + - (x * y) ≈⟨ +-congʳ $ zeroˡ _ ⟩
51-
0# + - (x * y) ≈⟨ +-identityˡ _ ⟩
52-
- (x * y) ∎
53-
54-
-‿distribʳ-* : x y - (x * y) ≈ x * - y
55-
-‿distribʳ-* x y = sym $ begin
56-
x * - y ≈⟨ sym $ +-identityˡ _ ⟩
57-
0# + x * - y ≈⟨ +-congʳ $ sym (-‿inverseˡ _) ⟩
58-
- (x * y) + x * y + x * - y ≈⟨ +-assoc _ _ _ ⟩
59-
- (x * y) + (x * y + x * - y) ≈⟨ +-congˡ $ sym (distribˡ _ _ _) ⟩
60-
- (x * y) + x * (y + - y) ≈⟨ +-congˡ $ *-congˡ $ -‿inverseʳ _ ⟩
61-
- (x * y) + x * 0# ≈⟨ +-congˡ $ zeroʳ _ ⟩
62-
- (x * y) + 0# ≈⟨ +-identityʳ _ ⟩
63-
- (x * y) ∎
26+
-- Extra properties of 1#
6427

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

src/Algebra/Properties/RingWithoutOne.agda

+23-8
Original file line numberDiff line numberDiff line change
@@ -38,22 +38,37 @@ open AbelianGroupProperties +-abelianGroup public
3838

3939
-‿distribˡ-* : x y - (x * y) ≈ - x * y
4040
-‿distribˡ-* x y = sym $ begin
41-
- x * y ≈⟨ sym $ +-identityʳ (- x * y)
42-
- x * y + 0# ≈⟨ +-congˡ $ sym ( -‿inverseʳ (x * y) ) ⟩
43-
- x * y + (x * y + - (x * y)) ≈⟨ sym $ +-assoc (- x * y) (x * y) (- (x * y))
44-
- x * y + x * y + - (x * y) ≈⟨ +-congʳ $ sym ( distribʳ y (- x) x ) ⟩
41+
- x * y ≈⟨ +-identityʳ (- x * y)
42+
- x * y + 0# ≈⟨ +-congˡ $ -‿inverseʳ (x * y)
43+
- x * y + (x * y + - (x * y)) ≈⟨ +-assoc (- x * y) (x * y) (- (x * y))
44+
- x * y + x * y + - (x * y) ≈⟨ +-congʳ $ distribʳ y (- x) x
4545
(- x + x) * y + - (x * y) ≈⟨ +-congʳ $ *-congʳ $ -‿inverseˡ x ⟩
4646
0# * y + - (x * y) ≈⟨ +-congʳ $ zeroˡ y ⟩
4747
0# + - (x * y) ≈⟨ +-identityˡ (- (x * y)) ⟩
4848
- (x * y) ∎
4949

5050
-‿distribʳ-* : x y - (x * y) ≈ x * - y
5151
-‿distribʳ-* x y = sym $ begin
52-
x * - y ≈⟨ sym $ +-identityˡ (x * (- y)) ⟩
53-
0# + x * - y ≈⟨ +-congʳ $ sym ( -‿inverseˡ (x * y) ) ⟩
54-
- (x * y) + x * y + x * - y ≈⟨ +-assoc (- (x * y)) (x * y) (x * (- y)) ⟩
55-
- (x * y) + (x * y + x * - y) ≈⟨ +-congˡ $ sym ( distribˡ x y ( - y) ) ⟩
52+
x * - y ≈⟨ +-identityˡ (x * - y)
53+
0# + x * - y ≈⟨ +-congʳ $ -‿inverseˡ (x * y)
54+
- (x * y) + x * y + x * - y ≈⟨ +-assoc (- (x * y)) (x * y) (x * - y) ⟩
55+
- (x * y) + (x * y + x * - y) ≈⟨ +-congˡ $ distribˡ x y (- y)
5656
- (x * y) + x * (y + - y) ≈⟨ +-congˡ $ *-congˡ $ -‿inverseʳ y ⟩
5757
- (x * y) + x * 0# ≈⟨ +-congˡ $ zeroʳ x ⟩
5858
- (x * y) + 0# ≈⟨ +-identityʳ (- (x * y)) ⟩
5959
- (x * y) ∎
60+
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
63+
64+
x[y-z]≈xy-xz : x y z x * (y - z) ≈ x * y - x * z
65+
x[y-z]≈xy-xz x y z = begin
66+
x * (y - z) ≈⟨ distribˡ x y (- z) ⟩
67+
x * y + x * - z ≈⟨ +-congˡ (sym (-‿distribʳ-* x z)) ⟩
68+
x * y - x * z ∎
69+
70+
[y-z]x≈yx-zx : x y z (y - z) * x ≈ (y * x) - (z * x)
71+
[y-z]x≈yx-zx x y z = begin
72+
(y - z) * x ≈⟨ distribʳ x y (- z) ⟩
73+
y * x + - z * x ≈⟨ +-congˡ (sym (-‿distribˡ-* z x)) ⟩
74+
y * x - z * x ∎

src/Algebra/Structures.agda

+26-58
Original file line numberDiff line numberDiff line change
@@ -651,7 +651,6 @@ record IsRingWithoutOne (+ * : Op₂ A) (-_ : Op₁ A) (0# : A) : Set (a ⊔ ℓ
651651
*-cong : Congruent₂ *
652652
*-assoc : Associative *
653653
distrib : * DistributesOver +
654-
zero : Zero 0# *
655654

656655
open IsAbelianGroup +-isAbelianGroup public
657656
renaming
@@ -679,23 +678,28 @@ record IsRingWithoutOne (+ * : Op₂ A) (-_ : Op₁ A) (0# : A) : Set (a ⊔ ℓ
679678
; isGroup to +-isGroup
680679
)
681680

682-
*-isMagma : IsMagma *
683-
*-isMagma = record
684-
{ isEquivalence = isEquivalence
685-
; ∙-cong = *-cong
686-
}
681+
distribˡ : * DistributesOverˡ +
682+
distribˡ = proj₁ distrib
683+
684+
distribʳ : * DistributesOverʳ +
685+
distribʳ = proj₂ distrib
687686

688687
zeroˡ : LeftZero 0# *
689-
zeroˡ = proj₁ zero
688+
zeroˡ = Consequences.assoc∧distribʳ∧idʳ∧invʳ⇒zeˡ setoid
689+
+-cong *-cong +-assoc distribʳ +-identityʳ -‿inverseʳ
690690

691691
zeroʳ : RightZero 0# *
692-
zeroʳ = proj₂ zero
692+
zeroʳ = Consequences.assoc∧distribˡ∧idʳ∧invʳ⇒zeʳ setoid
693+
+-cong *-cong +-assoc distribˡ +-identityʳ -‿inverseʳ
693694

694-
distribˡ : * DistributesOverˡ +
695-
distribˡ = proj₁ distrib
695+
zero : Zero 0# *
696+
zero = zeroˡ , zeroʳ
696697

697-
distribʳ : * DistributesOverʳ +
698-
distribʳ = proj₂ distrib
698+
*-isMagma : IsMagma *
699+
*-isMagma = record
700+
{ isEquivalence = isEquivalence
701+
; ∙-cong = *-cong
702+
}
699703

700704
*-isSemigroup : IsSemigroup *
701705
*-isSemigroup = record
@@ -806,45 +810,17 @@ record IsRing (+ * : Op₂ A) (-_ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where
806810
*-assoc : Associative *
807811
*-identity : Identity 1# *
808812
distrib : * DistributesOver +
809-
zero : Zero 0# *
810813

811-
open IsAbelianGroup +-isAbelianGroup public
812-
renaming
813-
( assoc to +-assoc
814-
; ∙-cong to +-cong
815-
; ∙-congˡ to +-congˡ
816-
; ∙-congʳ to +-congʳ
817-
; identity to +-identity
818-
; identityˡ to +-identityˡ
819-
; identityʳ to +-identityʳ
820-
; inverse to -‿inverse
821-
; inverseˡ to -‿inverseˡ
822-
; inverseʳ to -‿inverseʳ
823-
; ⁻¹-cong to -‿cong
824-
; comm to +-comm
825-
; isMagma to +-isMagma
826-
; isSemigroup to +-isSemigroup
827-
; isMonoid to +-isMonoid
828-
; isUnitalMagma to +-isUnitalMagma
829-
; isCommutativeMagma to +-isCommutativeMagma
830-
; isCommutativeMonoid to +-isCommutativeMonoid
831-
; isCommutativeSemigroup to +-isCommutativeSemigroup
832-
; isInvertibleMagma to +-isInvertibleMagma
833-
; isInvertibleUnitalMagma to +-isInvertibleUnitalMagma
834-
; isGroup to +-isGroup
835-
)
836-
837-
*-isMagma : IsMagma *
838-
*-isMagma = record
839-
{ isEquivalence = isEquivalence
840-
; ∙-cong = *-cong
814+
isRingWithoutOne : IsRingWithoutOne + * -_ 0#
815+
isRingWithoutOne = record
816+
{ +-isAbelianGroup = +-isAbelianGroup
817+
; *-cong = *-cong
818+
; *-assoc = *-assoc
819+
; distrib = distrib
841820
}
842821

843-
*-isSemigroup : IsSemigroup *
844-
*-isSemigroup = record
845-
{ isMagma = *-isMagma
846-
; assoc = *-assoc
847-
}
822+
open IsRingWithoutOne isRingWithoutOne public
823+
hiding (+-isAbelianGroup; *-cong; *-assoc; distrib)
848824

849825
*-isMonoid : IsMonoid * 1#
850826
*-isMonoid = record
@@ -855,18 +831,10 @@ record IsRing (+ * : Op₂ A) (-_ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where
855831
open IsMonoid *-isMonoid public
856832
using ()
857833
renaming
858-
( ∙-congˡ to *-congˡ
859-
; ∙-congʳ to *-congʳ
860-
; identityˡ to *-identityˡ
834+
( identityˡ to *-identityˡ
861835
; identityʳ to *-identityʳ
862836
)
863837

864-
zeroˡ : LeftZero 0# *
865-
zeroˡ = proj₁ zero
866-
867-
zeroʳ : RightZero 0# *
868-
zeroʳ = proj₂ zero
869-
870838
isSemiringWithoutAnnihilatingZero
871839
: IsSemiringWithoutAnnihilatingZero + * 0# 1#
872840
isSemiringWithoutAnnihilatingZero = record
@@ -885,7 +853,7 @@ record IsRing (+ * : Op₂ A) (-_ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where
885853
}
886854

887855
open IsSemiring isSemiring public
888-
using (distribˡ; distribʳ; isNearSemiring; isSemiringWithoutOne)
856+
using (isNearSemiring; isSemiringWithoutOne)
889857

890858

891859
record IsCommutativeRing

src/Algebra/Structures/Biased.agda

+35-20
Original file line numberDiff line numberDiff line change
@@ -192,6 +192,33 @@ open IsCommutativeSemiringʳ public
192192
------------------------------------------------------------------------
193193
-- IsRing
194194

195+
record IsRing* (+ * : Op₂ A) (-_ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where
196+
field
197+
+-isAbelianGroup : IsAbelianGroup + 0# -_
198+
*-isMonoid : IsMonoid * 1#
199+
distrib : * DistributesOver +
200+
zero : Zero 0# *
201+
202+
isRing : IsRing + * -_ 0# 1#
203+
isRing = record
204+
{ +-isAbelianGroup = +-isAbelianGroup
205+
; *-cong = ∙-cong
206+
; *-assoc = assoc
207+
; *-identity = identity
208+
; distrib = distrib
209+
} where open IsMonoid *-isMonoid
210+
211+
open IsRing* public
212+
using () renaming (isRing to isRing*)
213+
214+
215+
216+
------------------------------------------------------------------------
217+
-- Deprecated
218+
------------------------------------------------------------------------
219+
220+
-- Version 2.0
221+
195222
-- We can recover a ring without proving that 0# annihilates *.
196223
record IsRingWithoutAnnihilatingZero (+ * : Op₂ A) (-_ : Op₁ A) (0# 1# : A)
197224
: Set (a ⊔ ℓ) where
@@ -224,28 +251,16 @@ record IsRingWithoutAnnihilatingZero (+ * : Op₂ A) (-_ : Op₁ A) (0# 1# : A)
224251
; *-assoc = *.assoc
225252
; *-identity = *.identity
226253
; distrib = distrib
227-
; zero = zero
228254
}
229255

230256
open IsRingWithoutAnnihilatingZero public
231257
using () renaming (isRing to isRingWithoutAnnihilatingZero)
232258

233-
record IsRing* (+ * : Op₂ A) (-_ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where
234-
field
235-
+-isAbelianGroup : IsAbelianGroup + 0# -_
236-
*-isMonoid : IsMonoid * 1#
237-
distrib : * DistributesOver +
238-
zero : Zero 0# *
239-
240-
isRing : IsRing + * -_ 0# 1#
241-
isRing = record
242-
{ +-isAbelianGroup = +-isAbelianGroup
243-
; *-cong = ∙-cong
244-
; *-assoc = assoc
245-
; *-identity = identity
246-
; distrib = distrib
247-
; zero = zero
248-
} where open IsMonoid *-isMonoid
249-
250-
open IsRing* public
251-
using () renaming (isRing to isRing*)
259+
{-# WARNING_ON_USAGE IsRingWithoutAnnihilatingZero
260+
"Warning: IsRingWithoutAnnihilatingZero was deprecated in v2.0.
261+
Please use the standard `IsRing` instead."
262+
#-}
263+
{-# WARNING_ON_USAGE isRingWithoutAnnihilatingZero
264+
"Warning: isRingWithoutAnnihilatingZero was deprecated in v2.0.
265+
Please use the standard `IsRing` instead."
266+
#-}

src/Data/Integer/Properties.agda

-1
Original file line numberDiff line numberDiff line change
@@ -1532,7 +1532,6 @@ private
15321532
; *-assoc = *-assoc
15331533
; *-identity = *-identity
15341534
; distrib = *-distrib-+
1535-
; zero = *-zero
15361535
}
15371536

15381537
+-*-isCommutativeRing : IsCommutativeRing _+_ _*_ -_ 0ℤ 1ℤ

src/Data/Parity/Properties.agda

-1
Original file line numberDiff line numberDiff line change
@@ -371,7 +371,6 @@ p+p≡0ℙ 1ℙ = refl
371371
; *-assoc = *-assoc
372372
; *-identity = *-identity
373373
; distrib = *-distrib-+
374-
; zero = *-zero
375374
}
376375

377376
+-*-ring : Ring 0ℓ 0ℓ

0 commit comments

Comments
 (0)