------------------------------------------------------------------------ -- The Agda standard library -- -- Lexicographic products of binary relations ------------------------------------------------------------------------ -- The definition of lexicographic product used here is suitable if -- the left-hand relation is a strict partial order. module Relation.Binary.Product.StrictLex where open import Function open import Data.Product open import Data.Sum open import Data.Empty open import Level open import Relation.Nullary.Product open import Relation.Nullary.Sum open import Relation.Binary open import Relation.Binary.Consequences open import Relation.Binary.Product.Pointwise as Pointwise using (_×-Rel_) open import Relation.Nullary module _ {a₁ a₂ ℓ₁ ℓ₂} {A₁ : Set a₁} {A₂ : Set a₂} where ×-Lex : (_≈₁_ _<₁_ : Rel A₁ ℓ₁) → (_≤₂_ : Rel A₂ ℓ₂) → Rel (A₁ × A₂) _ ×-Lex _≈₁_ _<₁_ _≤₂_ = (_<₁_ on proj₁) -⊎- (_≈₁_ on proj₁) -×- (_≤₂_ on proj₂) -- Some properties which are preserved by ×-Lex (under certain -- assumptions). ×-reflexive : ∀ _≈₁_ _∼₁_ {_≈₂_ : Rel A₂ ℓ₂} _≤₂_ → _≈₂_ ⇒ _≤₂_ → (_≈₁_ ×-Rel _≈₂_) ⇒ (×-Lex _≈₁_ _∼₁_ _≤₂_) ×-reflexive _ _ _ refl₂ = λ x≈y → inj₂ (proj₁ x≈y , refl₂ (proj₂ x≈y)) _×-irreflexive_ : ∀ {_≈₁_ _<₁_} → Irreflexive _≈₁_ _<₁_ → ∀ {_≈₂_ _<₂_ : Rel A₂ ℓ₂} → Irreflexive _≈₂_ _<₂_ → Irreflexive (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_) (ir₁ ×-irreflexive ir₂) x≈y (inj₁ x₁y₁ = inj₂ (inj₁ x₁>y₁) ×-compare : {_≈₁_ _<₁_ : Rel A₁ ℓ₁} → Symmetric _≈₁_ → Trichotomous _≈₁_ _<₁_ → {_≈₂_ _<₂_ : Rel A₂ ℓ₂} → Trichotomous _≈₂_ _<₂_ → Trichotomous (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_) ×-compare {_≈₁_} {_<₁_} sym₁ compare₁ {_≈₂_} {_<₂_} compare₂ = cmp where cmp″ : ∀ {x₁ y₁ x₂ y₂} → ¬ (x₁ <₁ y₁) → x₁ ≈₁ y₁ → ¬ (y₁ <₁ x₁) → Tri (x₂ <₂ y₂) (x₂ ≈₂ y₂) (y₂ <₂ x₂) → Tri (×-Lex _≈₁_ _<₁_ _<₂_ (x₁ , x₂) (y₁ , y₂)) ((_≈₁_ ×-Rel _≈₂_) (x₁ , x₂) (y₁ , y₂)) (×-Lex _≈₁_ _<₁_ _<₂_ (y₁ , y₂) (x₁ , x₂)) cmp″ x₁≮y₁ x₁≈y₁ x₁≯y₁ (tri< x₂ x₂≮y₂ x₂≉y₂ x₂>y₂) = tri> [ x₁≮y₁ , x₂≮y₂ ∘ proj₂ ] (x₂≉y₂ ∘ proj₂) (inj₂ (sym₁ x₁≈y₁ , x₂>y₂)) cmp″ x₁≮y₁ x₁≈y₁ x₁≯y₁ (tri≈ x₂≮y₂ x₂≈y₂ x₂≯y₂) = tri≈ [ x₁≮y₁ , x₂≮y₂ ∘ proj₂ ] (x₁≈y₁ , x₂≈y₂) [ x₁≯y₁ , x₂≯y₂ ∘ proj₂ ] cmp′ : ∀ {x₁ y₁} → Tri (x₁ <₁ y₁) (x₁ ≈₁ y₁) (y₁ <₁ x₁) → ∀ x₂ y₂ → Tri (×-Lex _≈₁_ _<₁_ _<₂_ (x₁ , x₂) (y₁ , y₂)) ((_≈₁_ ×-Rel _≈₂_) (x₁ , x₂) (y₁ , y₂)) (×-Lex _≈₁_ _<₁_ _<₂_ (y₁ , y₂) (x₁ , x₂)) cmp′ (tri< x₁ x₁≮y₁ x₁≉y₁ x₁>y₁) x₂ y₂ = tri> [ x₁≮y₁ , x₁≉y₁ ∘ proj₁ ] (x₁≉y₁ ∘ proj₁) (inj₁ x₁>y₁) cmp′ (tri≈ x₁≮y₁ x₁≈y₁ x₁≯y₁) x₂ y₂ = cmp″ x₁≮y₁ x₁≈y₁ x₁≯y₁ (compare₂ x₂ y₂) cmp : Trichotomous (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_) cmp (x₁ , x₂) (y₁ , y₂) = cmp′ (compare₁ x₁ y₁) x₂ y₂ -- Some collections of properties which are preserved by ×-Lex. _×-isPreorder_ : ∀ {_≈₁_ _∼₁_} → IsPreorder _≈₁_ _∼₁_ → ∀ {_≈₂_ _∼₂_} → IsPreorder _≈₂_ _∼₂_ → IsPreorder (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _∼₁_ _∼₂_) _×-isPreorder_ {_≈₁_ = _≈₁_} {_∼₁_ = _∼₁_} pre₁ {_∼₂_ = _∼₂_} pre₂ = record { isEquivalence = Pointwise._×-isEquivalence_ (isEquivalence pre₁) (isEquivalence pre₂) ; reflexive = λ {x y} → ×-reflexive _≈₁_ _∼₁_ _∼₂_ (reflexive pre₂) {x} {y} ; trans = λ {x y z} → ×-transitive (isEquivalence pre₁) (∼-resp-≈ pre₁) (trans pre₁) {_≤₂_ = _∼₂_} (trans pre₂) {x} {y} {z} } where open IsPreorder _×-isStrictPartialOrder_ : ∀ {_≈₁_ _<₁_} → IsStrictPartialOrder _≈₁_ _<₁_ → ∀ {_≈₂_ _<₂_} → IsStrictPartialOrder _≈₂_ _<₂_ → IsStrictPartialOrder (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_) _×-isStrictPartialOrder_ {_<₁_ = _<₁_} spo₁ {_<₂_ = _<₂_} spo₂ = record { isEquivalence = Pointwise._×-isEquivalence_ (isEquivalence spo₁) (isEquivalence spo₂) ; irrefl = λ {x y} → _×-irreflexive_ {_<₁_ = _<₁_} (irrefl spo₁) {_<₂_ = _<₂_} (irrefl spo₂) {x} {y} ; trans = λ {x y z} → ×-transitive {_<₁_ = _<₁_} (isEquivalence spo₁) (<-resp-≈ spo₁) (trans spo₁) {_≤₂_ = _<₂_} (trans spo₂) {x} {y} {z} ; <-resp-≈ = ×-≈-respects₂ (isEquivalence spo₁) (<-resp-≈ spo₁) (<-resp-≈ spo₂) } where open IsStrictPartialOrder _×-isStrictTotalOrder_ : ∀ {_≈₁_ _<₁_} → IsStrictTotalOrder _≈₁_ _<₁_ → ∀ {_≈₂_ _<₂_} → IsStrictTotalOrder _≈₂_ _<₂_ → IsStrictTotalOrder (_≈₁_ ×-Rel _≈₂_) (×-Lex _≈₁_ _<₁_ _<₂_) _×-isStrictTotalOrder_ {_<₁_ = _<₁_} spo₁ {_<₂_ = _<₂_} spo₂ = record { isEquivalence = Pointwise._×-isEquivalence_ (isEquivalence spo₁) (isEquivalence spo₂) ; trans = λ {x y z} → ×-transitive {_<₁_ = _<₁_} (isEquivalence spo₁) (<-resp-≈ spo₁) (trans spo₁) {_≤₂_ = _<₂_} (trans spo₂) {x} {y} {z} ; compare = ×-compare (Eq.sym spo₁) (compare spo₁) (compare spo₂) ; <-resp-≈ = ×-≈-respects₂ (isEquivalence spo₁) (<-resp-≈ spo₁) (<-resp-≈ spo₂) } where open IsStrictTotalOrder -- "Packages" (e.g. preorders) can also be combined. _×-preorder_ : ∀ {p₁ p₂ p₃ p₄} → Preorder p₁ p₂ _ → Preorder p₃ p₄ _ → Preorder _ _ _ p₁ ×-preorder p₂ = record { isPreorder = isPreorder p₁ ×-isPreorder isPreorder p₂ } where open Preorder _×-strictPartialOrder_ : ∀ {s₁ s₂ s₃ s₄} → StrictPartialOrder s₁ s₂ _ → StrictPartialOrder s₃ s₄ _ → StrictPartialOrder _ _ _ s₁ ×-strictPartialOrder s₂ = record { isStrictPartialOrder = isStrictPartialOrder s₁ ×-isStrictPartialOrder isStrictPartialOrder s₂ } where open StrictPartialOrder _×-strictTotalOrder_ : ∀ {s₁ s₂ s₃ s₄} → StrictTotalOrder s₁ s₂ _ → StrictTotalOrder s₃ s₄ _ → StrictTotalOrder _ _ _ s₁ ×-strictTotalOrder s₂ = record { isStrictTotalOrder = isStrictTotalOrder s₁ ×-isStrictTotalOrder isStrictTotalOrder s₂ } where open StrictTotalOrder