@@ -805,6 +805,13 @@ New modules
805805 Algebra.Morphism.Construct.Identity
806806 ```
807807
808+ * Ordered algebraic structures (pomonoids, posemigroups, etc.)
809+ ```
810+ Algebra.Ordered
811+ Algebra.Ordered.Bundles
812+ Algebra.Ordered.Structures
813+ ```
814+
808815* 'Optimised' tail-recursive exponentiation properties:
809816 ```
810817 Algebra.Properties.Semiring.Exp.TailRecursiveOptimised
@@ -1554,10 +1561,18 @@ Other minor changes
15541561
15551562* Added new proofs to ` Relation.Binary.Lattice.Properties.{Join,Meet}Semilattice ` :
15561563 ``` agda
1564+ isPosemigroup : IsPosemigroup _≈_ _≤_ _∨_
1565+ posemigroup : Posemigroup c ℓ₁ ℓ₂
15571566 ≈-dec⇒≤-dec : Decidable _≈_ → Decidable _≤_
15581567 ≈-dec⇒isDecPartialOrder : Decidable _≈_ → IsDecPartialOrder _≈_ _≤_
15591568 ```
15601569
1570+ * Added new proofs to ` Relation.Binary.Lattice.Properties.Bounded{Join,Meet}Semilattice ` :
1571+ ``` agda
1572+ isCommutativePomonoid : IsCommutativePomonoid _≈_ _≤_ _∨_ ⊥
1573+ commutativePomonoid : CommutativePomonoid c ℓ₁ ℓ₂
1574+ ```
1575+
15611576* Added new proofs to ` Relation.Binary.Properties.Poset ` :
15621577 ``` agda
15631578 ≤-dec⇒≈-dec : Decidable _≤_ → Decidable _≈_
@@ -1654,6 +1669,14 @@ Other minor changes
16541669 ```
16551670 Cotransitive _#_ = ∀ {x y} → x # y → ∀ z → (x # z) ⊎ (z # y)
16561671 Tight _≈_ _#_ = ∀ x y → (¬ x # y → x ≈ y) × (x ≈ y → ¬ x # y)
1672+
1673+ Monotonic₁ _≤_ _⊑_ f = f Preserves _≤_ ⟶ _⊑_
1674+ Antitonic₁ _≤_ _⊑_ f = f Preserves (flip _≤_) ⟶ _⊑_
1675+ Monotonic₂ _≤_ _⊑_ _≼_ ∙ = ∙ Preserves₂ _≤_ ⟶ _⊑_ ⟶ _≼_
1676+ MonotonicAntitonic _≤_ _⊑_ _≼_ ∙ = ∙ Preserves₂ _≤_ ⟶ (flip _⊑_) ⟶ _≼_
1677+ AntitonicMonotonic _≤_ _⊑_ _≼_ ∙ = ∙ Preserves₂ (flip _≤_) ⟶ _⊑_ ⟶ _≼_
1678+ Antitonic₂ _≤_ _⊑_ _≼_ ∙ = ∙ Preserves₂ (flip _≤_) ⟶ (flip _⊑_) ⟶ _≼_
1679+ Adjoint _≤_ _⊑_ f g = ∀ {x y} → (f x ⊑ y → x ≤ g y) × (x ≤ g y → f x ⊑ y)
16571680 ```
16581681
16591682* Added new definitions in ` Relation.Binary.Bundles ` :
@@ -1671,6 +1694,9 @@ Other minor changes
16711694 sym⇒¬-sym : Symmetric _∼_ → Symmetric _≁_
16721695 cotrans⇒¬-trans : Cotransitive _∼_ → Transitive _≁_
16731696 irrefl⇒¬-refl : Reflexive _≈_ → Irreflexive _≈_ _∼_ → Reflexive _≁_
1697+ mono₂⇒cong₂ : Symmetric ≈₁ → ≈₁ ⇒ ≤₁ → Antisymmetric ≈₂ ≤₂ → ∀ {f} →
1698+ f Preserves₂ ≤₁ ⟶ ≤₁ ⟶ ≤₂ →
1699+ f Preserves₂ ≈₁ ⟶ ≈₁ ⟶ ≈₂
16741700 ```
16751701
16761702* Added new operations in ` Relation.Binary.PropositionalEquality.Properties ` :
0 commit comments