------------------------------------------------------------------------ -- The Agda standard library -- -- Some properties imply others ------------------------------------------------------------------------ module Relation.Binary.Consequences where open import Relation.Binary.Core open import Relation.Nullary using (yes; no) open import Relation.Unary using (∁) open import Function using (_∘_; flip) open import Data.Sum using (inj₁; inj₂) open import Data.Product using (_,_) open import Data.Empty using (⊥-elim) ------------------------------------------------------------------------ -- Substitutive properties module _ {a ℓ p} {A : Set a} {_∼_ : Rel A ℓ} (P : Rel A p) where subst⟶respˡ : Substitutive _∼_ p → P Respectsˡ _∼_ subst⟶respˡ subst {y} x'∼x Px'y = subst (flip P y) x'∼x Px'y subst⟶respʳ : Substitutive _∼_ p → P Respectsʳ _∼_ subst⟶respʳ subst {x} y'∼y Pxy' = subst (P x) y'∼y Pxy' subst⟶resp₂ : Substitutive _∼_ p → P Respects₂ _∼_ subst⟶resp₂ subst = subst⟶respʳ subst , subst⟶respˡ subst module _ {a ℓ p} {A : Set a} {∼ : Rel A ℓ} {P : A → Set p} where P-resp⟶¬P-resp : Symmetric ∼ → P Respects ∼ → (∁ P) Respects ∼ P-resp⟶¬P-resp sym resp x∼y ¬Px Py = ¬Px (resp (sym x∼y) Py) ------------------------------------------------------------------------ -- Proofs for non-strict orders module _ {a ℓ₁ ℓ₂} {A : Set a} {_≈_ : Rel A ℓ₁} {_≤_ : Rel A ℓ₂} where total⟶refl : _≤_ Respects₂ _≈_ → Symmetric _≈_ → Total _≤_ → _≈_ ⇒ _≤_ total⟶refl (respʳ , respˡ) sym total {x} {y} x≈y with total x y ... | inj₁ x∼y = x∼y ... | inj₂ y∼x = respʳ x≈y (respˡ (sym x≈y) y∼x) total+dec⟶dec : _≈_ ⇒ _≤_ → Antisymmetric _≈_ _≤_ → Total _≤_ → Decidable _≈_ → Decidable _≤_ total+dec⟶dec refl antisym total _≟_ x y with total x y ... | inj₁ x≤y = yes x≤y ... | inj₂ y≤x with x ≟ y ... | yes x≈y = yes (refl x≈y) ... | no x≉y = no (λ x≤y → x≉y (antisym x≤y y≤x)) ------------------------------------------------------------------------ -- Proofs for strict orders module _ {a ℓ₁ ℓ₂} {A : Set a} {_≈_ : Rel A ℓ₁} {_<_ : Rel A ℓ₂} where trans∧irr⟶asym : Reflexive _≈_ → Transitive _<_ → Irreflexive _≈_ _<_ → Asymmetric _<_ trans∧irr⟶asym refl trans irrefl xy with tri x y ... | tri< _ _ x≯y = x≯y x>y ... | tri≈ _ _ x≯y = x≯y x>y ... | tri> x≮y _ _ = x≮y x x≮y x≉y y _ x≉y _ = no x≉y tri⟶dec< : Trichotomous _≈_ _<_ → Decidable _<_ tri⟶dec< compare x y with compare x y ... | tri< x x≮y _ _ = no x≮y trans∧tri⟶respʳ≈ : Symmetric _≈_ → Transitive _≈_ → Transitive _<_ → Trichotomous _≈_ _<_ → _<_ Respectsʳ _≈_ trans∧tri⟶respʳ≈ sym ≈-tr <-tr tri {x} {y} {z} y≈z x _ _ z _ _ z