Skip to content

Commit ad622b2

Browse files
Fix agda#2396 by removing redundant zero in IsNonAssociativeRing
The zero field in the IsNonAssociativeRing was redundant, and could be replaced with a proof based on the other properties.
1 parent d5bb6cf commit ad622b2

File tree

1 file changed

+9
-7
lines changed

1 file changed

+9
-7
lines changed

src/Algebra/Structures.agda

+9-7
Original file line numberDiff line numberDiff line change
@@ -805,7 +805,6 @@ record IsNonAssociativeRing (+ * : Op₂ A) (-_ : Op₁ A) (0# 1# : A) : Set (a
805805
*-cong : Congruent₂ *
806806
*-identity : Identity 1# *
807807
distrib : * DistributesOver +
808-
zero : Zero 0# *
809808

810809
open IsAbelianGroup +-isAbelianGroup public
811810
renaming
@@ -833,18 +832,21 @@ record IsNonAssociativeRing (+ * : Op₂ A) (-_ : Op₁ A) (0# 1# : A) : Set (a
833832
; isGroup to +-isGroup
834833
)
835834

836-
zeroˡ : LeftZero 0# *
837-
zeroˡ = proj₁ zero
838-
839-
zeroʳ : RightZero 0# *
840-
zeroʳ = proj₂ zero
841-
842835
distribˡ : * DistributesOverˡ +
843836
distribˡ = proj₁ distrib
844837

845838
distribʳ : * DistributesOverʳ +
846839
distribʳ = proj₂ distrib
847840

841+
zeroˡ : LeftZero 0# *
842+
zeroˡ = Consequences.assoc∧distribʳ∧idʳ∧invʳ⇒zeˡ setoid +-cong *-cong +-assoc distribʳ +-identityʳ -‿inverseʳ
843+
844+
zeroʳ : RightZero 0# *
845+
zeroʳ = Consequences.assoc∧distribˡ∧idʳ∧invʳ⇒zeʳ setoid +-cong *-cong +-assoc distribˡ +-identityʳ -‿inverseʳ
846+
847+
zero : Zero 0# *
848+
zero = zeroˡ , zeroʳ
849+
848850
*-isMagma : IsMagma *
849851
*-isMagma = record
850852
{ isEquivalence = isEquivalence

0 commit comments

Comments
 (0)