Skip to content

Commit

Permalink
add: remaining Structures and Bundles
Browse files Browse the repository at this point in the history
  • Loading branch information
jamesmckinna committed Jan 24, 2025
1 parent 50bc64a commit 1cca718
Show file tree
Hide file tree
Showing 3 changed files with 95 additions and 5 deletions.
16 changes: 11 additions & 5 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -69,13 +69,16 @@ Additions to existing modules

* In `Algebra.Bundles`:
```agda
IntegralRing : (c ℓ : Level) → Set _
IntegralSemiring : (c ℓ : Level) → Set _
IntegralCommutativeRing : (c ℓ : Level) → Set _
IntegralCommutativeSemiring : (c ℓ : Level) → Set _
IntegralDomain : (c ℓ : Level) → Set _
IntegralRing : (c ℓ : Level) → Set _
IntegralSemiring : (c ℓ : Level) → Set _
```

* In `Algebra.Consequences.Base`:
```agda
integral⇒noZeroDivisors : Integral _≈_ 1# 0# _∙_ → ¬ (1# 0#)
integral⇒noZeroDivisors : Integral _≈_ 1# 0# _∙_ → 1# 0# →
NoZeroDivisors _≈_ 0# _∙_
noZeroDivisors⇒x≉0∧y≉0⇒xẏ≉0 : NoZeroDivisors _≈_ 0# _∙_ →
x ≉ 0# → y ≉ 0# → (x ∙ y) ≉ 0#
Expand Down Expand Up @@ -125,6 +128,9 @@ Additions to existing modules

* In `Algebra.Structures`:
```agda
IsIntegralSemiring : (+ * : Op₂ A) (0# 1# : A) → Set _
IsIntegralRing : (+ * : Op₂ A) (- : Op₁ A) (0# 1# : A) → Set _
IsIntegralCommutativeSemiring : (+ * : Op₂ A) (0# 1# : A) → Set _
IsIntegralCommutativeRing : (+ * : Op₂ A) (- : Op₁ A) (0# 1# : A) → Set _
IsIntegralDomain : (+ * : Op₂ A) (- : Op₁ A) (0# 1# : A) → Set _
IsIntegralSemiring : (+ * : Op₂ A) (0# 1# : A) → Set _
IsIntegralRing : (+ * : Op₂ A) (- : Op₁ A) (0# 1# : A) → Set _
```
40 changes: 40 additions & 0 deletions src/Algebra/Bundles.agda
Original file line number Diff line number Diff line change
Expand Up @@ -1194,6 +1194,46 @@ record IntegralRing c ℓ : Set (suc (c ⊔ ℓ)) where
)


record IntegralCommutativeRing c ℓ : Set (suc (c ⊔ ℓ)) where
infix 8 -_
infixl 7 _*_
infixl 6 _+_
infix 4 _≈_
field
Carrier : Set c
_≈_ : Rel Carrier ℓ
_+_ : Op₂ Carrier
_*_ : Op₂ Carrier
-_ : Op₁ Carrier
0# : Carrier
1# : Carrier
isIntegralCommutativeRing : IsIntegralCommutativeRing _≈_ _+_ _*_ -_ 0# 1#

open IsIntegralCommutativeRing isIntegralCommutativeRing public


record IntegralDomain c ℓ : Set (suc (c ⊔ ℓ)) where
infix 8 -_
infixl 7 _*_
infixl 6 _+_
infix 4 _≈_
field
Carrier : Set c
_≈_ : Rel Carrier ℓ
_+_ : Op₂ Carrier
_*_ : Op₂ Carrier
-_ : Op₁ Carrier
0# : Carrier
1# : Carrier
isIntegralDomain : IsIntegralDomain _≈_ _+_ _*_ -_ 0# 1#

open IsIntegralDomain isIntegralDomain public

integralCommutativeRing : IntegralCommutativeRing _ _
integralCommutativeRing = record
{ isIntegralCommutativeRing = isIntegralCommutativeRing }


------------------------------------------------------------------------
-- Bundles with 3 binary operations
------------------------------------------------------------------------
Expand Down
44 changes: 44 additions & 0 deletions src/Algebra/Structures.agda
Original file line number Diff line number Diff line change
Expand Up @@ -26,6 +26,7 @@ open import Algebra.Definitions _≈_
import Algebra.Consequences.Setoid as Consequences
open import Data.Product.Base using (_,_; proj₁; proj₂)
open import Level using (_⊔_)
open import Relation.Nullary.Negation.Core using (¬_)

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


record IsIntegralCommutativeSemiring (+ * : Op₂ A) (0# 1# : A) : Set (a ⊔ ℓ) where
field
isCommutativeSemiring : IsCommutativeSemiring + * 0# 1#
integral : Integral 1# 0# *

open IsCommutativeSemiring isCommutativeSemiring public

isIntegralSemiring : IsIntegralSemiring + * 0# 1#
isIntegralSemiring = record
{ isSemiring = isSemiring
; integral = integral
}


record IsKleeneAlgebra (+ * : Op₂ A) (⋆ : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where
field
isIdempotentSemiring : IsIdempotentSemiring + * 0# 1#
Expand Down Expand Up @@ -982,6 +997,35 @@ record IsIntegralRing
}


record IsIntegralCommutativeRing
(+ * : Op₂ A) (- : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where
field
isCommutativeRing : IsCommutativeRing + * - 0# 1#
integral : Integral 1# 0# *

open IsCommutativeRing isCommutativeRing public

isIntegralCommutativeSemiring : IsIntegralCommutativeSemiring + * 0# 1#
isIntegralCommutativeSemiring = record
{ isCommutativeSemiring = isCommutativeSemiring
; integral = integral
}


record IsIntegralDomain
(+ * : Op₂ A) (- : Op₁ A) (0# 1# : A) : Set (a ⊔ ℓ) where
field
isIntegralCommutativeRing : IsIntegralCommutativeRing + * - 0# 1#

open IsIntegralCommutativeRing isIntegralCommutativeRing public

field
nonTrivial : ¬ (1# ≈ 0#)

noZeroDivisors : NoZeroDivisors 0# *
noZeroDivisors = Consequences.integral⇒noZeroDivisors _≈_ integral nonTrivial


------------------------------------------------------------------------
-- Structures with 3 binary operations
------------------------------------------------------------------------
Expand Down

0 comments on commit 1cca718

Please sign in to comment.