------------------------------------------------------------------------ -- The Agda standard library -- -- Lexicographic ordering of lists ------------------------------------------------------------------------ -- The definition of lexicographic ordering used here is suitable if -- the argument order is a strict partial order. The lexicographic -- ordering itself can be either strict or non-strict, depending on -- the value of a parameter. module Relation.Binary.List.StrictLex where open import Data.Empty open import Data.Unit using (⊤; tt) open import Function open import Data.Product open import Data.Sum open import Data.List open import Level open import Relation.Nullary open import Relation.Binary open import Relation.Binary.Consequences open import Relation.Binary.List.Pointwise as Pointwise using ([]; _∷_; head; tail) module _ {A : Set} where data Lex (P : Set) (_≈_ _<_ : Rel A zero) : Rel (List A) zero where base : P → Lex P _≈_ _<_ [] [] halt : ∀ {y ys} → Lex P _≈_ _<_ [] (y ∷ ys) this : ∀ {x xs y ys} (x _ _ y (λ ()) (λ ()) halt cmp (x ∷ xs) (y ∷ ys) with tri x y ... | tri< x ¬x (¬≤-this ¬x≈y ¬x ¬xs (¬≤-next ¬x