diff --git a/CHANGELOG.md b/CHANGELOG.md index 468a6e6915..507e4956d8 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -10,7 +10,7 @@ Bug-fixes --------- * In `Algebra.Apartness.Structures`, renamed `sym` from `IsApartnessRelation` - to `#-sym` in order to avoid overloaded projection. + to `#-sym` in order to avoid overloaded projection. The field names `irrefl` and `cotrans` are similarly renamed for the sake of consistency. Non-backwards compatible changes @@ -21,6 +21,20 @@ Non-backwards compatible changes significantly faster. However, its reduction behaviour on open terms may have changed. +* The definitions of `Algebra.Structures.IsHeyting*` and + `Algebra.Structures.IsHeyting*` have been refactored, together + with that of `Relation.Binary.Definitions.Tight` on which they depend. + Specifically: + - `Tight _≈_ _#_` has been redefined as `∀ x y → ¬ x # y → x ≈ y`, + dropping the redundant second conjunct, because it is equivalent to + `Irreflexive _≈_ _#_`. + - new definitions: `(Is)TightApartnessRelation` structure/bundle, exploiting + the above redefinition. + - the definition of `HeytingCommutativeRing` now drops the properties of + invertibility, in favour of moving them to `HeytingField`. + - both `Heyting*` algebraic structure/bundles have been redefined to base + off an underlying `TightApartnessRelation`. + Minor improvements ------------------ @@ -30,6 +44,12 @@ Deprecated modules Deprecated names ---------------- +* In `Algebra.Apartness.Properties.HeytingCommutativeRing`: + ```agda + x-0≈x ↦ Algebra.Properties.Ring.x-0#≈x + #-sym ↦ Algebra.Apartness.Structures.IsHeytingCommutativeRing.#-sym + ``` + * In `Algebra.Definitions.RawMagma`: ```agda _∣∣_ ↦ _∥_ @@ -95,6 +115,9 @@ Deprecated names New modules ----------- +* `Algebra.Apartness.Properties.HeytingField`, refactoring the existing + `Algebra.Apartness.Properties.HeytingCommutativeRing`. + * `Data.List.Base.{and|or|any|all}` have been lifted out into `Data.Bool.ListAction`. * `Data.List.Base.{sum|product}` and their properties have been lifted out into `Data.Nat.ListAction` and `Data.Nat.ListAction.Properties`. @@ -102,6 +125,16 @@ New modules Additions to existing modules ----------------------------- +* In `Algebra.Apartness.Bundles.HeytingCommutativeRing`: + ```agda + TightApartnessRelation c ℓ₁ ℓ₂ : Set _ + ``` + +* In `Algebra.Apartness.Structures.IsHeytingCommutativeRing`: + ```agda + IsTightApartnessRelation _≈_ _#_ : Set _ + ``` + * In `Algebra.Construct.Pointwise`: ```agda isNearSemiring : IsNearSemiring _≈_ _+_ _*_ 0# → @@ -131,6 +164,16 @@ Additions to existing modules commutativeRing : CommutativeRing c ℓ → CommutativeRing (a ⊔ c) (a ⊔ ℓ) ``` +* In `Algebra.Properties.AbelianGroup`: + ```agda + x-ε≈x : RightIdentity ε _-_ + ``` + +* In `Algebra.Properties.RingWithoutOne`: + ```agda + x-0#≈x : RightIdentity 0# _-_ + ``` + * In `Data.List.Properties`: ```agda map-applyUpTo : ∀ (f : ℕ → A) (g : A → B) n → map g (applyUpTo f n) ≡ applyUpTo (g ∘ f) n @@ -143,3 +186,21 @@ Additions to existing modules ```agda filter-↭ : ∀ (P? : Pred.Decidable P) → xs ↭ ys → filter P? xs ↭ filter P? ys ``` + +* In `Data.Rational.Unnormalised.Properties`: + ```agda + ≄-apartnessRelation : ApartnessRelation _ _ _ + ≄-isTightApartnessRelation : IsTightApartnessRelation _≃_ _≄_ + ≄-tightApartnessRelation : TightApartnessRelation _ _ _ + ``` + +* In `Relation.Binary.Properties.DecSetoid`: + ```agda + ≉-isTightApartnessRelation : IsTightApartnessRelation _≈_ _#_ + ≉-tightApartnessRelation : TightApartnessRelation _ _ _ + ``` + Additionally, + ```agda + ≉-tight : Tight _≈_ _≉_ + ``` + has been redefined to reflect the change in the definition of `Tight`. diff --git a/src/Algebra/Apartness/Bundles.agda b/src/Algebra/Apartness/Bundles.agda index 9a55794051..a6a4addcdd 100644 --- a/src/Algebra/Apartness/Bundles.agda +++ b/src/Algebra/Apartness/Bundles.agda @@ -10,7 +10,7 @@ module Algebra.Apartness.Bundles where open import Level using (_⊔_; suc) open import Relation.Binary.Core using (Rel) -open import Relation.Binary.Bundles using (ApartnessRelation) +open import Relation.Binary.Bundles using (TightApartnessRelation) open import Algebra.Core using (Op₁; Op₂) open import Algebra.Bundles using (CommutativeRing) open import Algebra.Apartness.Structures @@ -36,8 +36,11 @@ record HeytingCommutativeRing c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ commutativeRing : CommutativeRing c ℓ₁ commutativeRing = record { isCommutativeRing = isCommutativeRing } - apartnessRelation : ApartnessRelation c ℓ₁ ℓ₂ - apartnessRelation = record { isApartnessRelation = isApartnessRelation } + tightApartnessRelation : TightApartnessRelation c ℓ₁ ℓ₂ + tightApartnessRelation = record { isTightApartnessRelation = isTightApartnessRelation } + + open TightApartnessRelation tightApartnessRelation public + using (apartnessRelation) record HeytingField c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where @@ -61,5 +64,5 @@ record HeytingField c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where heytingCommutativeRing : HeytingCommutativeRing c ℓ₁ ℓ₂ heytingCommutativeRing = record { isHeytingCommutativeRing = isHeytingCommutativeRing } - apartnessRelation : ApartnessRelation c ℓ₁ ℓ₂ - apartnessRelation = record { isApartnessRelation = isApartnessRelation } + open HeytingCommutativeRing heytingCommutativeRing public + using (commutativeRing; tightApartnessRelation; apartnessRelation) diff --git a/src/Algebra/Apartness/Properties/HeytingCommutativeRing.agda b/src/Algebra/Apartness/Properties/HeytingCommutativeRing.agda index 713edba00c..e9d5523839 100644 --- a/src/Algebra/Apartness/Properties/HeytingCommutativeRing.agda +++ b/src/Algebra/Apartness/Properties/HeytingCommutativeRing.agda @@ -11,74 +11,29 @@ open import Algebra.Apartness.Bundles using (HeytingCommutativeRing) module Algebra.Apartness.Properties.HeytingCommutativeRing {c ℓ₁ ℓ₂} (HCR : HeytingCommutativeRing c ℓ₁ ℓ₂) where -open import Function.Base using (_∘_) -open import Data.Product.Base using (_,_; proj₁; proj₂) -open import Algebra using (CommutativeRing; RightIdentity; Invertible; LeftInvertible; RightInvertible) +open import Algebra.Bundles using (CommutativeRing) -open HeytingCommutativeRing HCR -open CommutativeRing commutativeRing using (ring; *-commutativeMonoid) +open HeytingCommutativeRing HCR using (commutativeRing) +open CommutativeRing commutativeRing using (ring) +open import Algebra.Properties.Ring ring using (x-0#≈x) -open import Algebra.Properties.Ring ring - using (-0#≈0#; -‿distribˡ-*; -‿distribʳ-*; -‿anti-homo-+; -‿involutive) -open import Relation.Binary.Definitions using (Symmetric) -import Relation.Binary.Reasoning.Setoid as ≈-Reasoning -open import Algebra.Properties.CommutativeMonoid *-commutativeMonoid -private variable - x y z : Carrier - -invertibleˡ⇒# : LeftInvertible _≈_ 1# _*_ (x - y) → x # y -invertibleˡ⇒# = invertible⇒# ∘ invertibleˡ⇒invertible - -invertibleʳ⇒# : RightInvertible _≈_ 1# _*_ (x - y) → x # y -invertibleʳ⇒# = invertible⇒# ∘ invertibleʳ⇒invertible - -x-0≈x : RightIdentity _≈_ 0# _-_ -x-0≈x x = trans (+-congˡ -0#≈0#) (+-identityʳ x) - -1#0 : 1# # 0# -1#0 = invertibleˡ⇒# (1# , 1*[x-0]≈x) - where - 1*[x-0]≈x : 1# * (x - 0#) ≈ x - 1*[x-0]≈x {x} = trans (*-identityˡ (x - 0#)) (x-0≈x x) - -x#0y#0→xy#0 : x # 0# → y # 0# → x * y # 0# -x#0y#0→xy#0 {x} {y} x#0 y#0 = helper (#⇒invertible x#0) (#⇒invertible y#0) - where - - helper : Invertible _≈_ 1# _*_ (x - 0#) → Invertible _≈_ 1# _*_ (y - 0#) → x * y # 0# - helper (x⁻¹ , x⁻¹*x≈1 , x*x⁻¹≈1) (y⁻¹ , y⁻¹*y≈1 , y*y⁻¹≈1) - = invertibleˡ⇒# (y⁻¹ * x⁻¹ , y⁻¹*x⁻¹*x*y≈1) - where - open ≈-Reasoning setoid - - y⁻¹*x⁻¹*x*y≈1 : y⁻¹ * x⁻¹ * (x * y - 0#) ≈ 1# - y⁻¹*x⁻¹*x*y≈1 = begin - y⁻¹ * x⁻¹ * (x * y - 0#) ≈⟨ *-congˡ (x-0≈x (x * y)) ⟩ - y⁻¹ * x⁻¹ * (x * y) ≈⟨ *-assoc y⁻¹ x⁻¹ (x * y) ⟩ - y⁻¹ * (x⁻¹ * (x * y)) ≈⟨ *-congˡ (*-assoc x⁻¹ x y) ⟨ - y⁻¹ * ((x⁻¹ * x) * y) ≈⟨ *-congˡ (*-congʳ (*-congˡ (x-0≈x x))) ⟨ - y⁻¹ * ((x⁻¹ * (x - 0#)) * y) ≈⟨ *-congˡ (*-congʳ x⁻¹*x≈1) ⟩ - y⁻¹ * (1# * y) ≈⟨ *-congˡ (*-identityˡ y) ⟩ - y⁻¹ * y ≈⟨ *-congˡ (x-0≈x y) ⟨ - y⁻¹ * (y - 0#) ≈⟨ y⁻¹*y≈1 ⟩ - 1# ∎ - -#-congʳ : x ≈ y → x # z → y # z -#-congʳ {x} {y} {z} x≈y x#z = helper (#⇒invertible x#z) - where - - helper : Invertible _≈_ 1# _*_ (x - z) → y # z - helper (x-z⁻¹ , x-z⁻¹*x-z≈1# , x-z*x-z⁻¹≈1#) - = invertibleˡ⇒# (x-z⁻¹ , x-z⁻¹*y-z≈1) - where - open ≈-Reasoning setoid - - x-z⁻¹*y-z≈1 : x-z⁻¹ * (y - z) ≈ 1# - x-z⁻¹*y-z≈1 = begin - x-z⁻¹ * (y - z) ≈⟨ *-congˡ (+-congʳ x≈y) ⟨ - x-z⁻¹ * (x - z) ≈⟨ x-z⁻¹*x-z≈1# ⟩ - 1# ∎ - -#-congˡ : y ≈ z → x # y → x # z -#-congˡ y≈z x#y = #-sym (#-congʳ y≈z (#-sym x#y)) +------------------------------------------------------------------------ +-- DEPRECATED NAMES +------------------------------------------------------------------------ +-- Please use the new names as continuing support for the old names is +-- not guaranteed. + +-- Version 2.3 + +x-0≈x = x-0#≈x +{-# WARNING_ON_USAGE x-0≈x +"Warning: x-0≈x was deprecated in v2.3. +Please use Algebra.Properties.Ring.x-0#≈x instead." +#-} + +open HeytingCommutativeRing HCR public using (#-sym) +{-# WARNING_ON_USAGE #-sym +"Warning: #-sym was deprecated in v2.3. +Please use Algebra.Apartness.Structures.IsHeytingCommutativeRing.#-sym instead." +#-} diff --git a/src/Algebra/Apartness/Properties/HeytingField.agda b/src/Algebra/Apartness/Properties/HeytingField.agda new file mode 100644 index 0000000000..6cae113896 --- /dev/null +++ b/src/Algebra/Apartness/Properties/HeytingField.agda @@ -0,0 +1,85 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- Properties of Heyting Fields +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +open import Algebra.Apartness.Bundles using (HeytingField) + +module Algebra.Apartness.Properties.HeytingField + {c ℓ₁ ℓ₂} (HF : HeytingField c ℓ₁ ℓ₂) where + +open import Function.Base using (_∘_) +open import Data.Product.Base using (_,_; proj₁; proj₂) +open import Algebra.Bundles using (CommutativeRing) + +open HeytingField HF +open CommutativeRing commutativeRing using (ring; *-commutativeMonoid) + +open import Algebra.Definitions _≈_ + using (Invertible; LeftInvertible; RightInvertible) +open import Algebra.Properties.CommutativeMonoid *-commutativeMonoid + using (invertibleˡ⇒invertible; invertibleʳ⇒invertible) +open import Algebra.Properties.Ring ring + using (x-0#≈x; -‿distribˡ-*; -‿distribʳ-*; -‿anti-homo-+; -‿involutive) +open import Relation.Binary.Definitions using (Symmetric) +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning + +private + variable + x y z : Carrier + + +invertibleˡ⇒# : LeftInvertible 1# _*_ (x - y) → x # y +invertibleˡ⇒# = invertible⇒# ∘ invertibleˡ⇒invertible + +invertibleʳ⇒# : RightInvertible 1# _*_ (x - y) → x # y +invertibleʳ⇒# = invertible⇒# ∘ invertibleʳ⇒invertible + +1#0 : 1# # 0# +1#0 = invertibleˡ⇒# (1# , 1*[x-0]≈x) + where + 1*[x-0]≈x : 1# * (x - 0#) ≈ x + 1*[x-0]≈x {x} = trans (*-identityˡ (x - 0#)) (x-0#≈x x) + +x#0y#0→xy#0 : x # 0# → y # 0# → x * y # 0# +x#0y#0→xy#0 {x} {y} x#0 y#0 = helper (#⇒invertible x#0) (#⇒invertible y#0) + where + open ≈-Reasoning setoid + + helper : Invertible 1# _*_ (x - 0#) → Invertible 1# _*_ (y - 0#) → x * y # 0# + helper (x⁻¹ , x⁻¹*x≈1 , x*x⁻¹≈1) (y⁻¹ , y⁻¹*y≈1 , y*y⁻¹≈1) + = invertibleˡ⇒# (y⁻¹ * x⁻¹ , y⁻¹*x⁻¹*x*y≈1) + where + + y⁻¹*x⁻¹*x*y≈1 : y⁻¹ * x⁻¹ * (x * y - 0#) ≈ 1# + y⁻¹*x⁻¹*x*y≈1 = begin + y⁻¹ * x⁻¹ * (x * y - 0#) ≈⟨ *-congˡ (x-0#≈x (x * y)) ⟩ + y⁻¹ * x⁻¹ * (x * y) ≈⟨ *-assoc y⁻¹ x⁻¹ (x * y) ⟩ + y⁻¹ * (x⁻¹ * (x * y)) ≈⟨ *-congˡ (*-assoc x⁻¹ x y) ⟨ + y⁻¹ * ((x⁻¹ * x) * y) ≈⟨ *-congˡ (*-congʳ (*-congˡ (x-0#≈x x))) ⟨ + y⁻¹ * ((x⁻¹ * (x - 0#)) * y) ≈⟨ *-congˡ (*-congʳ x⁻¹*x≈1) ⟩ + y⁻¹ * (1# * y) ≈⟨ *-congˡ (*-identityˡ y) ⟩ + y⁻¹ * y ≈⟨ *-congˡ (x-0#≈x y) ⟨ + y⁻¹ * (y - 0#) ≈⟨ y⁻¹*y≈1 ⟩ + 1# ∎ + +#-congʳ : x ≈ y → x # z → y # z +#-congʳ {x} {y} {z} x≈y = helper ∘ #⇒invertible + where + open ≈-Reasoning setoid + + helper : Invertible 1# _*_ (x - z) → y # z + helper (x-z⁻¹ , x-z⁻¹*x-z≈1# , x-z*x-z⁻¹≈1#) + = invertibleˡ⇒# (x-z⁻¹ , x-z⁻¹*y-z≈1) + where + x-z⁻¹*y-z≈1 : x-z⁻¹ * (y - z) ≈ 1# + x-z⁻¹*y-z≈1 = begin + x-z⁻¹ * (y - z) ≈⟨ *-congˡ (+-congʳ x≈y) ⟨ + x-z⁻¹ * (x - z) ≈⟨ x-z⁻¹*x-z≈1# ⟩ + 1# ∎ + +#-congˡ : y ≈ z → x # y → x # z +#-congˡ y≈z = #-sym ∘ #-congʳ y≈z ∘ #-sym diff --git a/src/Algebra/Apartness/Structures.agda b/src/Algebra/Apartness/Structures.agda index 6ebee66787..6bc71d835e 100644 --- a/src/Algebra/Apartness/Structures.agda +++ b/src/Algebra/Apartness/Structures.agda @@ -17,22 +17,22 @@ module Algebra.Apartness.Structures where open import Level using (_⊔_; suc) -open import Data.Product.Base using (∃-syntax; _×_; _,_; proj₂) open import Algebra.Definitions _≈_ using (Invertible) open import Algebra.Structures _≈_ using (IsCommutativeRing) -open import Relation.Binary.Structures using (IsEquivalence; IsApartnessRelation) -open import Relation.Binary.Definitions using (Tight) -open import Relation.Nullary.Negation using (¬_) +open import Relation.Binary.Structures + using (IsEquivalence; IsApartnessRelation; IsTightApartnessRelation) import Relation.Binary.Properties.ApartnessRelation as AR record IsHeytingCommutativeRing : Set (c ⊔ ℓ₁ ⊔ ℓ₂) where field - isCommutativeRing : IsCommutativeRing _+_ _*_ -_ 0# 1# - isApartnessRelation : IsApartnessRelation _≈_ _#_ + isCommutativeRing : IsCommutativeRing _+_ _*_ -_ 0# 1# + isTightApartnessRelation : IsTightApartnessRelation _≈_ _#_ open IsCommutativeRing isCommutativeRing public + open IsTightApartnessRelation isTightApartnessRelation public + using (isApartnessRelation; tight) open IsApartnessRelation isApartnessRelation public renaming ( irrefl to #-irrefl @@ -40,10 +40,6 @@ record IsHeytingCommutativeRing : Set (c ⊔ ℓ₁ ⊔ ℓ₂) where ; cotrans to #-cotrans ) - field - #⇒invertible : ∀ {x y} → x # y → Invertible 1# _*_ (x - y) - invertible⇒# : ∀ {x y} → Invertible 1# _*_ (x - y) → x # y - ¬#-isEquivalence : IsEquivalence _¬#_ ¬#-isEquivalence = AR.¬#-isEquivalence refl isApartnessRelation @@ -52,6 +48,10 @@ record IsHeytingField : Set (c ⊔ ℓ₁ ⊔ ℓ₂) where field isHeytingCommutativeRing : IsHeytingCommutativeRing - tight : Tight _≈_ _#_ open IsHeytingCommutativeRing isHeytingCommutativeRing public + + field + #⇒invertible : ∀ {x y} → x # y → Invertible 1# _*_ (x - y) + invertible⇒# : ∀ {x y} → Invertible 1# _*_ (x - y) → x # y + diff --git a/src/Algebra/Properties/AbelianGroup.agda b/src/Algebra/Properties/AbelianGroup.agda index b5eb290081..2b52c6c755 100644 --- a/src/Algebra/Properties/AbelianGroup.agda +++ b/src/Algebra/Properties/AbelianGroup.agda @@ -6,12 +6,13 @@ {-# OPTIONS --cubical-compatible --safe #-} -open import Algebra +open import Algebra.Bundles using (AbelianGroup) module Algebra.Properties.AbelianGroup {a ℓ} (G : AbelianGroup a ℓ) where open AbelianGroup G +open import Algebra.Definitions _≈_ using (RightIdentity) open import Function.Base using (_$_) open import Relation.Binary.Reasoning.Setoid setoid @@ -39,3 +40,6 @@ xyx⁻¹≈y x y = begin x ⁻¹ ∙ y ⁻¹ ≈⟨ ⁻¹-anti-homo-∙ y x ⟨ (y ∙ x) ⁻¹ ≈⟨ ⁻¹-cong $ comm y x ⟩ (x ∙ y) ⁻¹ ∎ + +x-ε≈x : RightIdentity ε _-_ +x-ε≈x x = trans (∙-congˡ ε⁻¹≈ε) (identityʳ x) diff --git a/src/Algebra/Properties/RingWithoutOne.agda b/src/Algebra/Properties/RingWithoutOne.agda index d81ef58477..790136c57f 100644 --- a/src/Algebra/Properties/RingWithoutOne.agda +++ b/src/Algebra/Properties/RingWithoutOne.agda @@ -22,6 +22,7 @@ open import Relation.Binary.Reasoning.Setoid setoid open AbelianGroupProperties +-abelianGroup public renaming ( ε⁻¹≈ε to -0#≈0# + ; x-ε≈x to x-0#≈x ; ∙-cancelˡ to +-cancelˡ ; ∙-cancelʳ to +-cancelʳ ; ∙-cancel to +-cancel diff --git a/src/Data/Rational/Properties.agda b/src/Data/Rational/Properties.agda index 0341f53c45..cdc5edf3dc 100644 --- a/src/Data/Rational/Properties.agda +++ b/src/Data/Rational/Properties.agda @@ -1279,7 +1279,7 @@ module _ where using (+-group; zeroˡ; *-congʳ; isCommutativeRing) open GroupProperties +-group - open DecSetoidProperties ≡-decSetoid + open DecSetoidProperties ≡-decSetoid using (≉-isTightApartnessRelation) #⇒invertible : p ≢ q → Invertible 1ℚ _*_ (p - q) #⇒invertible {p} {q} p≢q = let r = p - q in 1/ r , *-inverseˡ r , *-inverseʳ r @@ -1299,15 +1299,14 @@ module _ where isHeytingCommutativeRing : IsHeytingCommutativeRing _≡_ _≢_ _+_ _*_ -_ 0ℚ 1ℚ isHeytingCommutativeRing = record { isCommutativeRing = isCommutativeRing - ; isApartnessRelation = ≉-isApartnessRelation - ; #⇒invertible = #⇒invertible - ; invertible⇒# = invertible⇒# + ; isTightApartnessRelation = ≉-isTightApartnessRelation } isHeytingField : IsHeytingField _≡_ _≢_ _+_ _*_ -_ 0ℚ 1ℚ isHeytingField = record { isHeytingCommutativeRing = isHeytingCommutativeRing - ; tight = ≉-tight + ; #⇒invertible = #⇒invertible + ; invertible⇒# = invertible⇒# } heytingCommutativeRing : HeytingCommutativeRing 0ℓ 0ℓ 0ℓ diff --git a/src/Data/Rational/Unnormalised/Properties.agda b/src/Data/Rational/Unnormalised/Properties.agda index 11e8d917b1..b1ba874008 100644 --- a/src/Data/Rational/Unnormalised/Properties.agda +++ b/src/Data/Rational/Unnormalised/Properties.agda @@ -35,11 +35,13 @@ open import Relation.Binary.Core using (_⇒_; _Preserves_⟶_; _Preserves₂_ open import Relation.Binary.Bundles using (Setoid; DecSetoid; Preorder; TotalPreorder; Poset; TotalOrder; DecTotalOrder; StrictPartialOrder; StrictTotalOrder; DenseLinearOrder) open import Relation.Binary.Structures - using (IsEquivalence; IsDecEquivalence; IsApartnessRelation; IsTotalPreorder; IsPreorder; IsPartialOrder; IsTotalOrder; IsDecTotalOrder; IsStrictPartialOrder; IsStrictTotalOrder; IsDenseLinearOrder) + using (IsEquivalence; IsDecEquivalence; IsApartnessRelation; IsTightApartnessRelation; IsTotalPreorder; IsPreorder; IsPartialOrder; IsTotalOrder; IsDecTotalOrder; IsStrictPartialOrder; IsStrictTotalOrder; IsDenseLinearOrder) open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive; Cotransitive; Tight; Decidable; Antisymmetric; Asymmetric; Dense; Total; Trans; Trichotomous; Irreflexive; Irrelevant; _Respectsˡ_; _Respectsʳ_; _Respects₂_; tri≈; tri<; tri>) import Relation.Binary.Consequences as BC open import Relation.Binary.PropositionalEquality +import Relation.Binary.Properties.Setoid as SetoidProperties +import Relation.Binary.Properties.DecSetoid as DecSetoidProperties import Relation.Binary.Properties.Poset as PosetProperties import Relation.Binary.Reasoning.Setoid as ≈-Reasoning open import Relation.Binary.Reasoning.Syntax @@ -106,18 +108,6 @@ p ≃? q = Dec.map′ *≡* drop-*≡* (↥ p ℤ.* ↧ q ℤ.≟ ↥ q ℤ.* 0≄1 : 0ℚᵘ ≄ 1ℚᵘ 0≄1 = Dec.from-no (0ℚᵘ ≃? 1ℚᵘ) -≃-≄-irreflexive : Irreflexive _≃_ _≄_ -≃-≄-irreflexive x≃y x≄y = x≄y x≃y - -≄-symmetric : Symmetric _≄_ -≄-symmetric x≄y y≃x = x≄y (≃-sym y≃x) - -≄-cotransitive : Cotransitive _≄_ -≄-cotransitive {x} {y} x≄y z with x ≃? z | z ≃? y -... | no x≄z | _ = inj₁ x≄z -... | yes _ | no z≄y = inj₂ z≄y -... | yes x≃z | yes z≃y = contradiction (≃-trans x≃z z≃y) x≄y - ≃-isEquivalence : IsEquivalence _≃_ ≃-isEquivalence = record { refl = ≃-refl @@ -131,17 +121,6 @@ p ≃? q = Dec.map′ *≡* drop-*≡* (↥ p ℤ.* ↧ q ℤ.≟ ↥ q ℤ.* ; _≟_ = _≃?_ } -≄-isApartnessRelation : IsApartnessRelation _≃_ _≄_ -≄-isApartnessRelation = record - { irrefl = ≃-≄-irreflexive - ; sym = ≄-symmetric - ; cotrans = ≄-cotransitive - } - -≄-tight : Tight _≃_ _≄_ -proj₁ (≄-tight p q) ¬p≄q = Dec.decidable-stable (p ≃? q) ¬p≄q -proj₂ (≄-tight p q) p≃q p≄q = p≄q p≃q - ≃-setoid : Setoid 0ℓ 0ℓ ≃-setoid = record { isEquivalence = ≃-isEquivalence @@ -152,6 +131,22 @@ proj₂ (≄-tight p q) p≃q p≄q = p≄q p≃q { isDecEquivalence = ≃-isDecEquivalence } +open SetoidProperties ≃-setoid public + renaming + ( ≉-sym to ≄-symmetric + ; ≉-irrefl to ≃-≄-irreflexive + ) + +open DecSetoidProperties ≃-decSetoid public + renaming + ( ≉-cotrans to ≄-cotransitive + ; ≉-tight to ≄-tight + ; ≉-isApartnessRelation to ≄-isApartnessRelation + ; ≉-apartnessRelation to ≄-ApartnessRelation + ; ≉-isTightApartnessRelation to ≄-isTightApartnessRelation + ; ≉-tightApartnessRelation to ≄-tightApartnessRelation + ) + module ≃-Reasoning = ≈-Reasoning ≃-setoid ↥p≡0⇒p≃0 : ∀ p → ↥ p ≡ 0ℤ → p ≃ 0ℚᵘ @@ -1399,15 +1394,14 @@ nonNeg*nonNeg⇒nonNeg p q = nonNegative +-*-isHeytingCommutativeRing : IsHeytingCommutativeRing _≃_ _≄_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ +-*-isHeytingCommutativeRing = record { isCommutativeRing = +-*-isCommutativeRing - ; isApartnessRelation = ≄-isApartnessRelation - ; #⇒invertible = ≄⇒invertible - ; invertible⇒# = invertible⇒≄ + ; isTightApartnessRelation = ≄-isTightApartnessRelation } +-*-isHeytingField : IsHeytingField _≃_ _≄_ _+_ _*_ -_ 0ℚᵘ 1ℚᵘ +-*-isHeytingField = record { isHeytingCommutativeRing = +-*-isHeytingCommutativeRing - ; tight = ≄-tight + ; #⇒invertible = ≄⇒invertible + ; invertible⇒# = invertible⇒≄ } ------------------------------------------------------------------------ diff --git a/src/Relation/Binary/Bundles.agda b/src/Relation/Binary/Bundles.agda index f471cf7bd4..f3edd76ce7 100644 --- a/src/Relation/Binary/Bundles.agda +++ b/src/Relation/Binary/Bundles.agda @@ -424,3 +424,18 @@ record ApartnessRelation c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) w renaming (_≁_ to _¬#_; _∼ᵒ_ to _#ᵒ_; _≁ᵒ_ to _¬#ᵒ_) hiding (Carrier; _≈_ ; _∼_) +record TightApartnessRelation c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where + infix 4 _≈_ _#_ + field + Carrier : Set c + _≈_ : Rel Carrier ℓ₁ + _#_ : Rel Carrier ℓ₂ + isTightApartnessRelation : IsTightApartnessRelation _≈_ _#_ + + open IsTightApartnessRelation isTightApartnessRelation public + using (isApartnessRelation; tight) + + apartnessRelation : ApartnessRelation _ _ _ + apartnessRelation = record { isApartnessRelation = isApartnessRelation } + + open ApartnessRelation apartnessRelation public diff --git a/src/Relation/Binary/Definitions.agda b/src/Relation/Binary/Definitions.agda index 167b63a8cb..210ebd0f4e 100644 --- a/src/Relation/Binary/Definitions.agda +++ b/src/Relation/Binary/Definitions.agda @@ -147,7 +147,7 @@ Cotransitive : Rel A ℓ → Set _ Cotransitive _#_ = ∀ {x y} → x # y → ∀ z → (x # z) ⊎ (z # y) Tight : Rel A ℓ₁ → Rel A ℓ₂ → Set _ -Tight _≈_ _#_ = ∀ x y → (¬ x # y → x ≈ y) × (x ≈ y → ¬ x # y) +Tight _≈_ _#_ = ∀ x y → ¬ x # y → x ≈ y -- Properties of order morphisms, aka order-preserving maps diff --git a/src/Relation/Binary/Properties/DecSetoid.agda b/src/Relation/Binary/Properties/DecSetoid.agda index 7f9b323d85..888091a80f 100644 --- a/src/Relation/Binary/Properties/DecSetoid.agda +++ b/src/Relation/Binary/Properties/DecSetoid.agda @@ -6,22 +6,22 @@ {-# OPTIONS --cubical-compatible --safe #-} -open import Relation.Binary.Bundles using (DecSetoid; ApartnessRelation) +open import Relation.Binary.Bundles + using (DecSetoid; ApartnessRelation; TightApartnessRelation) module Relation.Binary.Properties.DecSetoid {c ℓ} (S : DecSetoid c ℓ) where -open import Data.Product using (_,_) open import Data.Sum using (inj₁; inj₂) open import Relation.Binary.Definitions using (Cotransitive; Tight) import Relation.Binary.Properties.Setoid as SetoidProperties open import Relation.Binary.Structures - using (IsApartnessRelation; IsDecEquivalence) + using (IsApartnessRelation; IsTightApartnessRelation) open import Relation.Nullary.Decidable.Core using (yes; no; decidable-stable) open DecSetoid S using (_≈_; _≉_; _≟_; setoid; trans) -open SetoidProperties setoid +open SetoidProperties setoid using (≉-sym; ≉-irrefl) ≉-cotrans : Cotransitive _≉_ ≉-cotrans {x} {y} x≉y z with x ≟ z | z ≟ y @@ -36,8 +36,19 @@ open SetoidProperties setoid ; cotrans = ≉-cotrans } +≉-tight : Tight _≈_ _≉_ +≉-tight x y = decidable-stable (x ≟ y) + +≉-isTightApartnessRelation : IsTightApartnessRelation _≈_ _≉_ +≉-isTightApartnessRelation = record + { isApartnessRelation = ≉-isApartnessRelation + ; tight = ≉-tight + } + ≉-apartnessRelation : ApartnessRelation c ℓ ℓ ≉-apartnessRelation = record { isApartnessRelation = ≉-isApartnessRelation } -≉-tight : Tight _≈_ _≉_ -≉-tight x y = decidable-stable (x ≟ y) , ≉-irrefl +≉-tightApartnessRelation : TightApartnessRelation c ℓ ℓ +≉-tightApartnessRelation = record + { isTightApartnessRelation = ≉-isTightApartnessRelation } + diff --git a/src/Relation/Binary/Structures.agda b/src/Relation/Binary/Structures.agda index 3bbfc6db27..f0d2665593 100644 --- a/src/Relation/Binary/Structures.agda +++ b/src/Relation/Binary/Structures.agda @@ -328,3 +328,10 @@ record IsApartnessRelation (_#_ : Rel A ℓ₂) : Set (a ⊔ ℓ ⊔ ℓ₂) whe _¬#_ : A → A → Set _ x ¬# y = ¬ (x # y) + +record IsTightApartnessRelation (_#_ : Rel A ℓ₂) : Set (a ⊔ ℓ ⊔ ℓ₂) where + field + isApartnessRelation : IsApartnessRelation _#_ + tight : Tight _≈_ _#_ + + open IsApartnessRelation isApartnessRelation public