Skip to content

Commit 1cca718

Browse files
committed
add: remaining Structures and Bundles
1 parent 50bc64a commit 1cca718

File tree

3 files changed

+95
-5
lines changed

3 files changed

+95
-5
lines changed

CHANGELOG.md

Lines changed: 11 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -69,13 +69,16 @@ Additions to existing modules
6969

7070
* In `Algebra.Bundles`:
7171
```agda
72-
IntegralRing : (c ℓ : Level) → Set _
73-
IntegralSemiring : (c ℓ : Level) → Set _
72+
IntegralCommutativeRing : (c ℓ : Level) → Set _
73+
IntegralCommutativeSemiring : (c ℓ : Level) → Set _
74+
IntegralDomain : (c ℓ : Level) → Set _
75+
IntegralRing : (c ℓ : Level) → Set _
76+
IntegralSemiring : (c ℓ : Level) → Set _
7477
```
7578

7679
* In `Algebra.Consequences.Base`:
7780
```agda
78-
integral⇒noZeroDivisors : Integral _≈_ 1# 0# _∙_ → ¬ (1# 0#)
81+
integral⇒noZeroDivisors : Integral _≈_ 1# 0# _∙_ → 1# 0# →
7982
NoZeroDivisors _≈_ 0# _∙_
8083
noZeroDivisors⇒x≉0∧y≉0⇒xẏ≉0 : NoZeroDivisors _≈_ 0# _∙_ →
8184
x ≉ 0# → y ≉ 0# → (x ∙ y) ≉ 0#
@@ -125,6 +128,9 @@ Additions to existing modules
125128

126129
* In `Algebra.Structures`:
127130
```agda
128-
IsIntegralSemiring : (+ * : Op₂ A) (0# 1# : A) → Set _
129-
IsIntegralRing : (+ * : Op₂ A) (- : Op₁ A) (0# 1# : A) → Set _
131+
IsIntegralCommutativeSemiring : (+ * : Op₂ A) (0# 1# : A) → Set _
132+
IsIntegralCommutativeRing : (+ * : Op₂ A) (- : Op₁ A) (0# 1# : A) → Set _
133+
IsIntegralDomain : (+ * : Op₂ A) (- : Op₁ A) (0# 1# : A) → Set _
134+
IsIntegralSemiring : (+ * : Op₂ A) (0# 1# : A) → Set _
135+
IsIntegralRing : (+ * : Op₂ A) (- : Op₁ A) (0# 1# : A) → Set _
130136
```

src/Algebra/Bundles.agda

Lines changed: 40 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1194,6 +1194,46 @@ record IntegralRing c ℓ : Set (suc (c ⊔ ℓ)) where
11941194
)
11951195

11961196

1197+
record IntegralCommutativeRing c ℓ : Set (suc (c ⊔ ℓ)) where
1198+
infix 8 -_
1199+
infixl 7 _*_
1200+
infixl 6 _+_
1201+
infix 4 _≈_
1202+
field
1203+
Carrier : Set c
1204+
_≈_ : Rel Carrier ℓ
1205+
_+_ : Op₂ Carrier
1206+
_*_ : Op₂ Carrier
1207+
-_ : Op₁ Carrier
1208+
0# : Carrier
1209+
1# : Carrier
1210+
isIntegralCommutativeRing : IsIntegralCommutativeRing _≈_ _+_ _*_ -_ 0# 1#
1211+
1212+
open IsIntegralCommutativeRing isIntegralCommutativeRing public
1213+
1214+
1215+
record IntegralDomain c ℓ : Set (suc (c ⊔ ℓ)) where
1216+
infix 8 -_
1217+
infixl 7 _*_
1218+
infixl 6 _+_
1219+
infix 4 _≈_
1220+
field
1221+
Carrier : Set c
1222+
_≈_ : Rel Carrier ℓ
1223+
_+_ : Op₂ Carrier
1224+
_*_ : Op₂ Carrier
1225+
-_ : Op₁ Carrier
1226+
0# : Carrier
1227+
1# : Carrier
1228+
isIntegralDomain : IsIntegralDomain _≈_ _+_ _*_ -_ 0# 1#
1229+
1230+
open IsIntegralDomain isIntegralDomain public
1231+
1232+
integralCommutativeRing : IntegralCommutativeRing _ _
1233+
integralCommutativeRing = record
1234+
{ isIntegralCommutativeRing = isIntegralCommutativeRing }
1235+
1236+
11971237
------------------------------------------------------------------------
11981238
-- Bundles with 3 binary operations
11991239
------------------------------------------------------------------------

src/Algebra/Structures.agda

Lines changed: 44 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -26,6 +26,7 @@ open import Algebra.Definitions _≈_
2626
import Algebra.Consequences.Setoid as Consequences
2727
open import Data.Product.Base using (_,_; proj₁; proj₂)
2828
open import Level using (_⊔_)
29+
open import Relation.Nullary.Negation.Core using (¬_)
2930

3031
------------------------------------------------------------------------
3132
-- Structures with 1 unary operation & 1 element
@@ -649,6 +650,20 @@ record IsIntegralSemiring (+ * : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ) where
649650
open IsSemiring isSemiring public
650651

651652

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+
652667
record IsKleeneAlgebra (+ * : Op₂ A) (⋆ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where
653668
field
654669
isIdempotentSemiring : IsIdempotentSemiring + * 0# 1#
@@ -982,6 +997,35 @@ record IsIntegralRing
982997
}
983998

984999

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+
9851029
------------------------------------------------------------------------
9861030
-- Structures with 3 binary operations
9871031
------------------------------------------------------------------------

0 commit comments

Comments
 (0)