Comparison #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
This file provides basic results about orderings and comparison in linear orders.
Definitions #
- cmp_le: An- orderingfrom- ≤.
- ordering.compares: Turns an- orderinginto- <and- =propositions.
- linear_order_of_compares: Constructs a- linear_orderinstance from the fact that any two elements that are not one strictly less than the other either way are equal.
Like cmp, but uses a ≤ on the type instead of <. Given two elements x and y, returns a
three-way comparison result ordering.
Equations
- cmp_le x y = ite (x ≤ y) (ite (y ≤ x) ordering.eq ordering.lt) ordering.gt
    
theorem
cmp_le_eq_cmp
    {α : Type u_1}
    [preorder α]
    [is_total α has_le.le]
    [decidable_rel has_le.le]
    [decidable_rel has_lt.lt]
    (x y : α) :
@[simp]
compares o a b means that a and b have the ordering relation o between them, assuming
that the relation a < b is defined.
Equations
- ordering.gt.compares a b = (a > b)
- ordering.eq.compares a b = (a = b)
- ordering.lt.compares a b = (a < b)
Alias of the forward direction of ordering.compares_swap.
Alias of the reverse direction of ordering.compares_swap.
    
theorem
ordering.compares.eq_gt
    {α : Type u_1}
    [preorder α]
    {o : ordering}
    {a b : α}
    (h : o.compares a b) :
o = ordering.gt ↔ b < a
    
theorem
ordering.compares.ne_gt
    {α : Type u_1}
    [preorder α]
    {o : ordering}
    {a b : α}
    (h : o.compares a b) :
o ≠ ordering.gt ↔ a ≤ b
    
theorem
ordering.or_else_eq_lt
    (o₁ o₂ : ordering) :
o₁.or_else o₂ = ordering.lt ↔ o₁ = ordering.lt ∨ o₁ = ordering.eq ∧ o₂ = ordering.lt
@[simp]
    
theorem
to_dual_compares_to_dual
    {α : Type u_1}
    [has_lt α]
    {a b : α}
    {o : ordering} :
o.compares (⇑order_dual.to_dual a) (⇑order_dual.to_dual b) ↔ o.compares b a
@[simp]
    
theorem
of_dual_compares_of_dual
    {α : Type u_1}
    [has_lt α]
    {a b : αᵒᵈ}
    {o : ordering} :
o.compares (⇑order_dual.of_dual a) (⇑order_dual.of_dual b) ↔ o.compares b a
    
theorem
ordering.compares.cmp_eq
    {α : Type u_1}
    [linear_order α]
    {a b : α}
    {o : ordering}
    (h : o.compares a b) :
@[simp]
    
theorem
cmp_le_to_dual
    {α : Type u_1}
    [has_le α]
    [decidable_rel has_le.le]
    (x y : α) :
cmp_le (⇑order_dual.to_dual x) (⇑order_dual.to_dual y) = cmp_le y x
@[simp]
    
theorem
cmp_le_of_dual
    {α : Type u_1}
    [has_le α]
    [decidable_rel has_le.le]
    (x y : αᵒᵈ) :
cmp_le (⇑order_dual.of_dual x) (⇑order_dual.of_dual y) = cmp_le y x
@[simp]
    
theorem
cmp_to_dual
    {α : Type u_1}
    [has_lt α]
    [decidable_rel has_lt.lt]
    (x y : α) :
cmp (⇑order_dual.to_dual x) (⇑order_dual.to_dual y) = cmp y x
@[simp]
    
theorem
cmp_of_dual
    {α : Type u_1}
    [has_lt α]
    [decidable_rel has_lt.lt]
    (x y : αᵒᵈ) :
cmp (⇑order_dual.of_dual x) (⇑order_dual.of_dual y) = cmp y x
    
def
linear_order_of_compares
    {α : Type u_1}
    [preorder α]
    (cmp : α → α → ordering)
    (h : ∀ (a b : α), (cmp a b).compares a b) :
Generate a linear order structure from a preorder and cmp function.
Equations
- linear_order_of_compares cmp h = {le := preorder.le _inst_1, lt := preorder.lt _inst_1, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_total := _, decidable_le := λ (a b : α), decidable_of_iff (cmp a b ≠ ordering.gt) _, decidable_eq := λ (a b : α), decidable_of_iff (cmp a b = ordering.eq) _, decidable_lt := λ (a b : α), decidable_of_iff (cmp a b = ordering.lt) _, max := max_default (λ (a b : α), decidable_of_iff (cmp a b ≠ ordering.gt) _), max_def := _, min := min_default (λ (a b : α), decidable_of_iff (cmp a b ≠ ordering.gt) _), min_def := _}
@[simp]
@[simp]
@[simp]
@[simp]
    
theorem
cmp_eq_cmp_symm
    {α : Type u_1}
    [linear_order α]
    {x y : α}
    {β : Type u_2}
    [linear_order β]
    {x' y' : β} :
    
theorem
lt_iff_lt_of_cmp_eq_cmp
    {α : Type u_1}
    [linear_order α]
    {x y : α}
    {β : Type u_2}
    [linear_order β]
    {x' y' : β}
    (h : cmp x y = cmp x' y') :
    
theorem
le_iff_le_of_cmp_eq_cmp
    {α : Type u_1}
    [linear_order α]
    {x y : α}
    {β : Type u_2}
    [linear_order β]
    {x' y' : β}
    (h : cmp x y = cmp x' y') :
    
theorem
eq_iff_eq_of_cmp_eq_cmp
    {α : Type u_1}
    [linear_order α]
    {x y : α}
    {β : Type u_2}
    [linear_order β]
    {x' y' : β}
    (h : cmp x y = cmp x' y') :
    
theorem
has_lt.lt.cmp_eq_lt
    {α : Type u_1}
    [linear_order α]
    {x y : α}
    (h : x < y) :
cmp x y = ordering.lt
    
theorem
has_lt.lt.cmp_eq_gt
    {α : Type u_1}
    [linear_order α]
    {x y : α}
    (h : x < y) :
cmp y x = ordering.gt