@@ -26,6 +26,7 @@ open import Algebra.Definitions _≈_
26
26
import Algebra.Consequences.Setoid as Consequences
27
27
open import Data.Product.Base using (_,_; proj₁; proj₂)
28
28
open import Level using (_⊔_)
29
+ open import Relation.Nullary.Negation.Core using (¬_)
29
30
30
31
------------------------------------------------------------------------
31
32
-- Structures with 1 unary operation & 1 element
@@ -649,6 +650,20 @@ record IsIntegralSemiring (+ * : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ) where
649
650
open IsSemiring isSemiring public
650
651
651
652
653
+ record IsIntegralCommutativeSemiring (+ * : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ) where
654
+ field
655
+ isCommutativeSemiring : IsCommutativeSemiring + * 0# 1#
656
+ integral : Integral 1# 0# *
657
+
658
+ open IsCommutativeSemiring isCommutativeSemiring public
659
+
660
+ isIntegralSemiring : IsIntegralSemiring + * 0# 1#
661
+ isIntegralSemiring = record
662
+ { isSemiring = isSemiring
663
+ ; integral = integral
664
+ }
665
+
666
+
652
667
record IsKleeneAlgebra (+ * : Op₂ A) (⋆ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where
653
668
field
654
669
isIdempotentSemiring : IsIdempotentSemiring + * 0# 1#
@@ -982,6 +997,35 @@ record IsIntegralRing
982
997
}
983
998
984
999
1000
+ record IsIntegralCommutativeRing
1001
+ (+ * : Op₂ A) (- : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where
1002
+ field
1003
+ isCommutativeRing : IsCommutativeRing + * - 0# 1#
1004
+ integral : Integral 1# 0# *
1005
+
1006
+ open IsCommutativeRing isCommutativeRing public
1007
+
1008
+ isIntegralCommutativeSemiring : IsIntegralCommutativeSemiring + * 0# 1#
1009
+ isIntegralCommutativeSemiring = record
1010
+ { isCommutativeSemiring = isCommutativeSemiring
1011
+ ; integral = integral
1012
+ }
1013
+
1014
+
1015
+ record IsIntegralDomain
1016
+ (+ * : Op₂ A) (- : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where
1017
+ field
1018
+ isIntegralCommutativeRing : IsIntegralCommutativeRing + * - 0# 1#
1019
+
1020
+ open IsIntegralCommutativeRing isIntegralCommutativeRing public
1021
+
1022
+ field
1023
+ nonTrivial : ¬ (1# ≈ 0#)
1024
+
1025
+ noZeroDivisors : NoZeroDivisors 0# *
1026
+ noZeroDivisors = Consequences.integral⇒noZeroDivisors _≈_ integral nonTrivial
1027
+
1028
+
985
1029
------------------------------------------------------------------------
986
1030
-- Structures with 3 binary operations
987
1031
------------------------------------------------------------------------
0 commit comments