------------------------------------------------------------------------ -- The Agda standard library -- -- Some properties imply others ------------------------------------------------------------------------ module Relation.Binary.Consequences where open import Relation.Binary.Core hiding (refl) open import Relation.Nullary open import Relation.Binary.PropositionalEquality.Core open import Function open import Data.Sum open import Data.Product open import Data.Empty -- Some of the definitions can be found in the following module: open import Relation.Binary.Consequences.Core public trans∧irr⟶asym : ∀ {a ℓ₁ ℓ₂} {A : Set a} → {_≈_ : Rel A ℓ₁} {_<_ : Rel A ℓ₂} → 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 _ = no x≉y tri⟶dec< : ∀ {a ℓ₁ ℓ₂} {A : Set a} {_≈_ : Rel A ℓ₁} {_<_ : Rel A ℓ₂} → Trichotomous _≈_ _<_ → Decidable _<_ tri⟶dec< compare x y with compare x y ... | tri< x x≮y _ _ = no x≮y map-NonEmpty : ∀ {a b p q} {A : Set a} {B : Set b} {P : REL A B p} {Q : REL A B q} → P ⇒ Q → NonEmpty P → NonEmpty Q map-NonEmpty f x = nonEmpty (f (NonEmpty.proof x))