@[class]
- le : α → α → Prop
- lt : α → α → Prop
- le_refl : ∀ (a : α), a ≤ a
- le_trans : ∀ (a b c : α), a ≤ b → b ≤ c → a ≤ c
- lt_iff_le_not_le : (∀ (a b : α), a < b ↔ a ≤ b ∧ ¬b ≤ a) . "order_laws_tac"
A preorder is a reflexive, transitive relation ≤
with a < b
defined in the obvious way.
Instances of this typeclass
- partial_order.to_preorder
- order_dual.preorder
- pi.preorder
- subtype.preorder
- prod.preorder
- with_bot.preorder
- with_top.preorder
- order_hom.preorder
- with_zero.preorder
- units.preorder
- add_units.preorder
- multiplicative.preorder
- additive.preorder
- rat.preorder
- top_hom.preorder
- bot_hom.preorder
- associates.preorder
- prod.lex.preorder
- finsupp.preorder
- sum.preorder
- sum.lex.preorder
- lift.subfield_with_hom.preorder
- cau_seq.preorder
- real.preorder
- measure_theory.simple_func.preorder
- filter.germ.preorder
- measure_theory.ae_eq_fun.preorder
- bitvec.preorder
- category_theory.thin_skeleton.preorder
- Preord.preorder
- with_lower_topology.preorder
- continuous_order_hom.preorder
- category_theory.grothendieck_topology.cover.preorder
- dfinsupp.preorder
- nonempty_interval.preorder
- interval.preorder
- order.partial_iso.preorder
- zsqrtd.preorder
- pgame.preorder
- onote.preorder
- nonote.preorder
- tropical.preorder
- order_ring_hom.preorder
- sigma.preorder
- sigma.lex.preorder
- psigma.lex.preorder
- counterexample.from_Bhavik.K.preorder
Instances of other typeclasses for preorder
- preorder.has_sizeof_inst
<
is decidable if ≤
is.
Equations
- decidable_lt_of_decidable_le a b = dite (a ≤ b) (λ (hab : a ≤ b), dite (b ≤ a) (λ (hba : b ≤ a), decidable.is_false _) (λ (hba : ¬b ≤ a), decidable.is_true _)) (λ (hab : ¬a ≤ b), decidable.is_false _)
Definition of partial_order
and lemmas about types with a partial order #
@[class]
- le : α → α → Prop
- lt : α → α → Prop
- le_refl : ∀ (a : α), a ≤ a
- le_trans : ∀ (a b c : α), a ≤ b → b ≤ c → a ≤ c
- lt_iff_le_not_le : (∀ (a b : α), a < b ↔ a ≤ b ∧ ¬b ≤ a) . "order_laws_tac"
- le_antisymm : ∀ (a b : α), a ≤ b → b ≤ a → a = b
A partial order is a reflexive, transitive, antisymmetric relation ≤
.
Instances of this typeclass
- omega_complete_partial_order.to_partial_order
- linear_order.to_partial_order
- semilattice_sup.to_partial_order
- semilattice_inf.to_partial_order
- ordered_comm_monoid.to_partial_order
- ordered_add_comm_monoid.to_partial_order
- ordered_cancel_add_comm_monoid.to_partial_order
- ordered_cancel_comm_monoid.to_partial_order
- ordered_add_comm_group.to_partial_order
- ordered_comm_group.to_partial_order
- complete_semilattice_Sup.to_partial_order
- complete_semilattice_Inf.to_partial_order
- set_like.partial_order
- order_dual.partial_order
- pi.partial_order
- subtype.partial_order
- prod.partial_order
- Prop.partial_order
- complementeds.partial_order
- with_bot.partial_order
- with_top.partial_order
- order_hom.partial_order
- with_zero.partial_order
- units.partial_order
- add_units.partial_order
- multiplicative.partial_order
- additive.partial_order
- fin.partial_order
- multiset.partial_order
- finset.partial_order
- pi.lex.partial_order
- rat.partial_order
- top_hom.partial_order
- bot_hom.partial_order
- setoid.partial_order
- associates.partial_order
- antisymmetrization.partial_order
- part.partial_order
- omega_complete_partial_order.continuous_hom.partial_order
- prod.lex.partial_order
- con.partial_order
- add_con.partial_order
- Sup_hom.partial_order
- Inf_hom.partial_order
- frame_hom.partial_order
- set_semiring.partial_order
- part_enat.partial_order
- cardinal.partial_order
- finsupp.partial_order
- localization.partial_order
- add_localization.partial_order
- sum.partial_order
- sum.lex.partial_order
- ordinal.partial_order
- pequiv.partial_order
- intermediate_field.lifts.partial_order
- filter.partial_order
- topological_space.partial_order
- uniform_space.partial_order
- group_topology.partial_order
- add_group_topology.partial_order
- ring_topology.partial_order
- real.partial_order
- group_seminorm.partial_order
- add_group_seminorm.partial_order
- nonarch_add_group_seminorm.partial_order
- group_norm.partial_order
- add_group_norm.partial_order
- nonarch_add_group_norm.partial_order
- measurable_space.partial_order
- measurable_space.dynkin_system.partial_order
- measure_theory.outer_measure.outer_measure.partial_order
- measurable_set.subtype.partial_order
- measure_theory.measure.partial_order
- measure_theory.simple_func.partial_order
- seminorm.partial_order
- box_integral.box.partial_order
- box_integral.prepartition.partial_order
- box_integral.integration_params.partial_order
- filter.germ.partial_order
- continuous_map.partial_order
- bounded_continuous_function.partial_order
- measure_theory.ae_eq_fun.partial_order
- category_theory.thin_skeleton.thin_skeleton_partial_order
- measure_theory.vector_measure.partial_order
- category_theory.subobject.partial_order
- PartOrd.partial_order
- FinPartOrd.partial_order
- shrinking_lemma.partial_refinement.partial_order
- continuous_order_hom.partial_order
- topological_space.open_nhds.partial_order
- category_theory.grothendieck_topology.partial_order
- category_theory.pretopology.partial_order
- category_theory.grothendieck_topology.subpresheaf.partial_order
- dfinsupp.partial_order
- dfinsupp.lex.partial_order
- finsupp.lex.partial_order
- enorm.partial_order
- is_Lprojection.subtype.partial_order
- structure_groupoid.partial_order
- prime_spectrum.partial_order
- homogeneous_ideal.partial_order
- lie_subalgebra.partial_order
- nonempty_interval.partial_order
- interval.partial_order
- order.ideal.partial_order
- order.pfilter.partial_order
- finpartition.partial_order
- heyting_hom.partial_order
- coheyting_hom.partial_order
- biheyting_hom.partial_order
- BddOrd.partial_order
- game.partial_order
- tropical.partial_order
- order_ring_hom.partial_order
- many_one_degree.partial_order
- setoid.partition.partial_order
- semiquot.partial_order
- sigma.partial_order
- sigma.lex.partial_order
- psigma.lex.partial_order
- projective_spectrum.partial_order
- counterexample.Nxzmod_2.preN2
Instances of other typeclasses for partial_order
- partial_order.has_sizeof_inst
@[instance]
Instances for partial_order.to_preorder
Equality is decidable if ≤
is.
Equations
- decidable_eq_of_decidable_le a b = dite (a ≤ b) (λ (hab : a ≤ b), dite (b ≤ a) (λ (hba : b ≤ a), decidable.is_true _) (λ (hba : ¬b ≤ a), decidable.is_false _)) (λ (hab : ¬a ≤ b), decidable.is_false _)
theorem
decidable.lt_or_eq_of_le
{α : Type u}
[partial_order α]
[decidable_rel has_le.le]
{a b : α}
(hab : a ≤ b) :
theorem
decidable.eq_or_lt_of_le
{α : Type u}
[partial_order α]
[decidable_rel has_le.le]
{a b : α}
(hab : a ≤ b) :
theorem
decidable.le_iff_lt_or_eq
{α : Type u}
[partial_order α]
[decidable_rel has_le.le]
{a b : α} :
Definition of linear_order
and lemmas about types with a linear order #
Default definition of max
.
Equations
- max_default a b = ite (a ≤ b) b a
Default definition of min
.
Equations
- min_default a b = ite (a ≤ b) a b
@[class]
- le : α → α → Prop
- lt : α → α → Prop
- le_refl : ∀ (a : α), a ≤ a
- le_trans : ∀ (a b c : α), a ≤ b → b ≤ c → a ≤ c
- lt_iff_le_not_le : (∀ (a b : α), a < b ↔ a ≤ b ∧ ¬b ≤ a) . "order_laws_tac"
- le_antisymm : ∀ (a b : α), a ≤ b → b ≤ a → a = b
- le_total : ∀ (a b : α), a ≤ b ∨ b ≤ a
- decidable_le : decidable_rel has_le.le
- decidable_eq : decidable_eq α
- decidable_lt : decidable_rel has_lt.lt
- max : α → α → α
- max_def : linear_order.max = max_default . "reflexivity"
- min : α → α → α
- min_def : linear_order.min = min_default . "reflexivity"
A linear order is reflexive, transitive, antisymmetric and total relation ≤
.
We assume that every linear ordered type has decidable (≤)
, (<)
, and (=)
.
Instances of this typeclass
- linear_ordered_add_comm_monoid.to_linear_order
- linear_ordered_comm_monoid.to_linear_order
- linear_ordered_add_comm_group.to_linear_order
- linear_ordered_comm_group.to_linear_order
- canonically_linear_ordered_add_monoid.to_linear_order
- canonically_linear_ordered_monoid.to_linear_order
- linear_ordered_ring.to_linear_order
- complete_linear_order.to_linear_order
- conditionally_complete_linear_order.to_linear_order
- nonempty_fin_lin_ord.to_linear_order
- nat.linear_order
- int.linear_order
- bool.linear_order
- order_dual.linear_order
- subtype.linear_order
- punit.linear_order
- as_linear_order.linear_order
- with_bot.linear_order
- with_top.linear_order
- Prop.linear_order
- with_zero.linear_order
- units.linear_order
- add_units.linear_order
- multiplicative.linear_order
- additive.linear_order
- fin.linear_order
- list.linear_order
- pi.lex.linear_order
- pnat.linear_order
- rat.linear_order
- antisymmetrization.linear_order
- flag.linear_order
- prod.lex.linear_order
- enat.linear_order
- part_enat.linear_order
- cardinal.linear_order
- sum.lex.linear_order
- linear_order_out
- ordinal.linear_order
- real.linear_order
- sign_type.linear_order
- char.linear_order
- string.linear_order
- pos_num.linear_order
- znum.linear_order
- LinOrd.linear_order
- dfinsupp.lex.linear_order
- finsupp.lex.linear_order
- valuation_ring.ideal.linear_order
- valuation_subring.linear_order_overring
- filter.germ.linear_order
- well_order_extension.linear_order
- linear_extension.linear_order
- zsqrtd.linear_order
- nat_ordinal.linear_order
- nonote.linear_order
- tropical.linear_order
- colex.finset.colex.linear_order
- sigma.lex.linear_order
- psigma.lex.linear_order
- counterexample.int_with_epsilon.linear_order
- counterexample.F.linear_order
- counterexample.foo.linear_order
Instances of other typeclasses for linear_order
- linear_order.has_sizeof_inst
@[instance]
Instances for linear_order.to_partial_order
@[protected, instance]
Equations
@[protected, instance]
Equations
@[protected, instance]
Equations
- eq.decidable a b = linear_order.decidable_eq a b
@[protected, instance]
@[protected, instance]
@[protected, instance]
def
lt_by_cases
{α : Type u}
[linear_order α]
(x y : α)
{P : Sort u_1}
(h₁ : x < y → P)
(h₂ : x = y → P)
(h₃ : y < x → P) :
P
Perform a case-split on the ordering of x
and y
in a decidable linear order.
theorem
le_imp_le_of_lt_imp_lt
{α : Type u}
[linear_order α]
{β : Type u_1}
[preorder α]
[linear_order β]
{a b : α}
{c d : β}
(H : d < c → b < a)
(h : a ≤ b) :
c ≤ d