------------------------------------------------------------------------ -- The Agda standard library -- -- Conversion of < to ≤, along with a number of properties ------------------------------------------------------------------------ -- Possible TODO: Prove that a conversion ≤ → < → ≤ returns a -- relation equivalent to the original one (and similarly for -- < → ≤ → <). open import Relation.Binary module Relation.Binary.StrictToNonStrict {a ℓ₁ ℓ₂} {A : Set a} (_≈_ : Rel A ℓ₁) (_<_ : Rel A ℓ₂) where open import Relation.Nullary open import Relation.Binary.Consequences open import Function open import Data.Product open import Data.Sum open import Data.Empty ------------------------------------------------------------------------ -- Conversion -- _<_ can be turned into _≤_ as follows: _≤_ : Rel A _ x ≤ y = (x < y) ⊎ (x ≈ y) ------------------------------------------------------------------------ -- The converted relations have certain properties -- (if the original relations have certain other properties) reflexive : _≈_ ⇒ _≤_ reflexive = inj₂ antisym : IsEquivalence _≈_ → Transitive _<_ → Irreflexive _≈_ _<_ → Antisymmetric _≈_ _≤_ antisym eq trans irrefl = as where module Eq = IsEquivalence eq as : Antisymmetric _≈_ _≤_ as (inj₂ x≈y) _ = x≈y as (inj₁ _) (inj₂ y≈x) = Eq.sym y≈x as (inj₁ x x≮y x≉y x>y = inj₂ (inj₁ x>y) decidable : Decidable _≈_ → Decidable _<_ → Decidable _≤_ decidable ≈-dec <-dec x y with ≈-dec x y | <-dec x y ... | yes x≈y | _ = yes (inj₂ x≈y) ... | no x≉y | yes x x≮y x≉y _ = no helper where helper : x ≤ y → ⊥ helper (inj₁ x