Unbundled relation classes
In this file we prove some properties of is_*
classes defined in init.algebra.classes
. The main
difference between these classes and the usual order classes (preorder
etc) is that usual classes
extend has_le
and/or has_lt
while these classes take a relation as an explicit argument.
Equations
Equations
- ge.is_refl = _
Equations
Equations
- ge.is_trans = _
Equations
Equations
Equations
Equations
Equations
Equations
- gt.is_trans = _
Equations
Equations
- gt.is_asymm = _
Equations
Equations
Equations
Equations
Equations
Equations
Equations
Equations
Equations
Equations
Equations
- ge.is_total = _
Equations
Equations
Equations
Equations
Equations
Construct a partial order from a is_strict_order
relation
Equations
- partial_order_of_SO r = {le := λ (x y : α), x = y ∨ r x y, lt := r, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _}
- to_is_trichotomous : is_trichotomous α lt
- to_is_strict_order : is_strict_order α lt
This is basically the same as is_strict_total_order
, but that definition is
in Type (probably by mistake) and also has redundant assumptions.
Construct a linear order from a is_strict_total_order'
relation
Equations
- linear_order_of_STO' r = {le := partial_order.le (partial_order_of_SO r), lt := partial_order.lt (partial_order_of_SO r), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_total := _}
Construct a decidable linear order from a is_strict_total_order'
relation
Equations
- decidable_linear_order_of_STO' r = let LO : linear_order α := linear_order_of_STO' r in {le := linear_order.le LO, lt := linear_order.lt LO, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_total := _, decidable_le := λ (x y : α), decidable_of_iff (¬r y x) not_lt, decidable_eq := decidable_eq_of_decidable_le (λ (x y : α), decidable_of_iff (¬r y x) not_lt), decidable_lt := decidable_lt_of_decidable_le (λ (x y : α), decidable_of_iff (¬r y x) not_lt)}
- conn : ∀ (a b c : α), lt a c → lt a b ∨ lt b c
A connected order is one satisfying the condition a < c → a < b ∨ b < c
.
This is recognizable as an intuitionistic substitute for a ≤ b ∨ b ≤ a
on
the constructive reals, and is also known as negative transitivity,
since the contrapositive asserts transitivity of the relation ¬ a < b
.
Equations
An extensional relation is one in which an element is determined by its set
of predecessors. It is named for the x ∈ y
relation in set theory, whose
extensionality is one of the first axioms of ZFC.
Equations
- to_is_strict_total_order' : is_strict_total_order' α r
- wf : well_founded r
A well order is a well-founded linear order.
Equations
Equations
Equations
Equations
Equations
Equations
Construct a decidable linear order from a well-founded linear order.
Equations
Equations
Equations
Equations
If r
is a well-founded relation, then any nonempty set has a minimal element
with respect to r
.
A minimal element of a nonempty set in a well-founded order
Equations
- H.min p h = classical.some _
The supremum of a bounded, well-founded order
A successor of an element x
in a well-founded order is a minimal element y
such that
x < y
if one exists. Otherwise it is x
itself.