diff --git a/CHANGELOG.md b/CHANGELOG.md index 200f618201..520d05db99 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -18,7 +18,6 @@ Highlights * Improved the `solve` tactic in `Tactic.RingSolver` to work in a much wider range of situations. - Bug-fixes --------- @@ -1269,6 +1268,22 @@ Other minor changes anyUpTo? : ∀ (P? : U.Decidable P) (v : ℕ) → Dec (∃ λ n → n < v × P n) allUpTo? : ∀ (P? : U.Decidable P) (v : ℕ) → Dec (∀ {n} → n < v → P n) + + +-isPomagma : IsPomagma _+_ + +-isPosemigroup : IsPosemigroup _+_ + +-0-isPomonoid : IsPomonoid _+_ 0 + +-0-isCommutativePomonoid : IsCommutativePomonoid _+_ 0 + + +-0-commutativePomonoid : CommutativePomonoid 0ℓ 0ℓ 0ℓ + + *-isPomagma : IsPomagma _≤_ _*_ + *-isPosemigroup : IsPosemigroup _≤_ _*_ + *-1-isPomonoid : IsPomonoid _≤_ _*_ 1 + *-1-isCommutativePomonoid : IsCommutativePomonoid _*_ 1 + +-*-isPosemiring : IsPosemiring _+_ _*_ 0 1 + + *-1-commutativePomonoid : CommutativePomonoid 0ℓ 0ℓ 0ℓ + +-*-posemiring : Posemiring 0ℓ 0ℓ 0ℓ ``` * Added new functions in `Data.Nat`: diff --git a/src/Algebra/Ordered/Structures.agda b/src/Algebra/Ordered/Structures.agda index 34a64e6519..b53cc7f682 100644 --- a/src/Algebra/Ordered/Structures.agda +++ b/src/Algebra/Ordered/Structures.agda @@ -13,9 +13,9 @@ open import Relation.Binary.Core using (Rel; _⇒_) module Algebra.Ordered.Structures - {a ℓ₁ ℓ₂} {A : Set a} -- The underlying set - (_≈_ : Rel A ℓ₁) -- The underlying equality relation - (_≤_ : Rel A ℓ₂) -- The order + {a} {A : Set a} -- The underlying set + {ℓ₁} (_≈_ : Rel A ℓ₁) -- The underlying equality relation + {ℓ₂} (_≤_ : Rel A ℓ₂) -- The order where open import Algebra.Core diff --git a/src/Data/Nat/Properties.agda b/src/Data/Nat/Properties.agda index ae1fd5809f..bc63e36860 100644 --- a/src/Data/Nat/Properties.agda +++ b/src/Data/Nat/Properties.agda @@ -19,6 +19,7 @@ open import Algebra.Construct.NaturalChoice.Base import Algebra.Construct.NaturalChoice.MinMaxOp as MinMaxOp import Algebra.Lattice.Construct.NaturalChoice.MinMaxOp as LatticeMinMaxOp import Algebra.Properties.CommutativeSemigroup as CommSemigroupProperties +open import Algebra.Ordered.Bundles open import Data.Bool.Base using (Bool; false; true; T) open import Data.Bool.Properties using (T?) open import Data.Empty using (⊥) @@ -44,6 +45,7 @@ open import Algebra.Definitions {A = ℕ} _≡_ open import Algebra.Definitions using (LeftCancellative; RightCancellative; Cancellative) open import Algebra.Structures {A = ℕ} _≡_ +open import Algebra.Ordered.Structures {A = ℕ} _≡_ ------------------------------------------------------------------------ -- Properties of NonZero @@ -614,10 +616,6 @@ open ≤-Reasoning { isCommutativeMonoid = +-0-isCommutativeMonoid } -∸-magma : Magma 0ℓ 0ℓ -∸-magma = magma _∸_ - - ------------------------------------------------------------------------ -- Other properties of _+_ and _≡_ @@ -732,6 +730,41 @@ m+n≮n (suc m) (suc n) (s