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 LE and/or LT while these classes take a relation as an explicit argument.
A version of antisymm with r explicit.
This lemma matches the lemmas from lean core in Init.Algebra.Classes, but is missing there.
A version of antisymm' with r explicit.
This lemma matches the lemmas from lean core in Init.Algebra.Classes, but is missing there.
In a trichotomous irreflexive order, every element is determined by the set of predecessors.
Construct a partial order from an isStrictOrder relation.
See note [reducible non-instances].
Equations
- partialOrderOfSO r = PartialOrder.mk (_ : ∀ (x y : α), x ≤ y → y ≤ x → x = y)
Instances For
Construct a linear order from an IsStrictTotalOrder relation.
See note [reducible non-instances].
Equations
- One or more equations did not get rendered due to their size.
Instances For
Order connection #
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.
- 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.
Instances
Equations
- (_ : IsOrderConnected α r) = (_ : IsOrderConnected α r)
Equations
- (_ : IsStrictWeakOrder α r) = (_ : IsStrictWeakOrder α r)
Well-order #
A well-founded relation. Not to be confused with IsWellOrder.
- wf : WellFounded r
The relation is
WellFounded, as a proposition.
Instances
Equations
- (_ : IsWellFounded α WellFoundedRelation.rel) = (_ : IsWellFounded α WellFoundedRelation.rel)
Induction on a well-founded relation.
All values are accessible under the well-founded relation.
Creates data, given a way to generate a value from all that compare as less under a well-founded
relation. See also IsWellFounded.fix_eq.
Equations
- IsWellFounded.fix r = WellFounded.fix (_ : WellFounded fun (y x : α) => r y x)
Instances For
The value from IsWellFounded.fix is built from the previous ones as specified.
Derive a WellFoundedRelation instance from an isWellFounded instance.
Equations
- IsWellFounded.toWellFoundedRelation r = { rel := r, wf := (_ : WellFounded r) }
Instances For
Equations
- (_ : IsWellFounded α (Relation.TransGen r)) = (_ : IsWellFounded α (Relation.TransGen r))
A class for a well founded relation <.
Equations
- WellFoundedLT α = IsWellFounded α fun (x x_1 : α) => x < x_1
Instances For
A class for a well founded relation >.
Equations
- WellFoundedGT α = IsWellFounded α fun (x x_1 : α) => x > x_1
Instances For
Equations
- (_ : WellFoundedGT αᵒᵈ) = h
Equations
- (_ : WellFoundedLT αᵒᵈ) = h
A well order is a well-founded linear order.
Instances
Equations
- (_ : IsStrictTotalOrder α r) = (_ : IsStrictTotalOrder α r)
Equations
- (_ : IsTrichotomous α r) = (_ : IsTrichotomous α r)
Inducts on a well-founded < relation.
All values are accessible under the well-founded <.
Creates data, given a way to generate a value from all that compare as lesser. See also
WellFoundedLT.fix_eq.
Equations
- WellFoundedLT.fix = IsWellFounded.fix fun (x x_1 : α) => x < x_1
Instances For
The value from WellFoundedLT.fix is built from the previous ones as specified.
Derive a WellFoundedRelation instance from a WellFoundedLT instance.
Equations
- WellFoundedLT.toWellFoundedRelation = IsWellFounded.toWellFoundedRelation fun (x x_1 : α) => x < x_1
Instances For
Inducts on a well-founded > relation.
All values are accessible under the well-founded >.
Creates data, given a way to generate a value from all that compare as greater. See also
WellFoundedGT.fix_eq.
Equations
- WellFoundedGT.fix = IsWellFounded.fix fun (x x_1 : α) => x > x_1
Instances For
The value from WellFoundedGT.fix is built from the successive ones as specified.
Derive a WellFoundedRelation instance from a WellFoundedGT instance.
Equations
- WellFoundedGT.toWellFoundedRelation = IsWellFounded.toWellFoundedRelation fun (x x_1 : α) => x > x_1
Instances For
Construct a decidable linear order from a well-founded linear order.
Equations
Instances For
Derive a WellFoundedRelation instance from an IsWellOrder instance.
Equations
- IsWellOrder.toHasWellFounded = { rel := fun (x x_1 : α) => x < x_1, wf := (_ : WellFounded fun (x x_1 : α) => x < x_1) }
Instances For
Equations
- (_ : IsWellOrder α EmptyRelation) = (_ : IsWellOrder α EmptyRelation)
Equations
- (_ : IsWellOrder α r) = (_ : IsWellOrder α r)
Equations
- (_ : IsWellFounded (α × β) (Prod.Lex r s)) = (_ : IsWellFounded (α × β) (Prod.Lex r s))
Equations
- (_ : IsWellOrder (α × β) (Prod.Lex r s)) = (_ : IsWellOrder (α × β) (Prod.Lex r s))
Equations
- (_ : IsWellFounded β (InvImage r f)) = (_ : IsWellFounded β (InvImage r f))
Equations
- (_ : IsWellFounded α (InvImage (fun (x x_1 : ℕ) => x < x_1) f)) = (_ : IsWellFounded α (InvImage (fun (x x_1 : ℕ) => x < x_1) f))
Equations
- (_ : WellFoundedLT (α × β)) = (_ : IsWellFounded (α × β) fun (x x_1 : α × β) => x < x_1)
Equations
- (_ : WellFoundedGT (α × β)) = (_ : WellFoundedLT (αᵒᵈ × βᵒᵈ))
A bounded or final set. Not to be confused with Bornology.IsBounded.
Equations
- Set.Bounded r s = ∃ (a : α), ∀ (b : α), b ∈ s → r b a
Instances For
Strict-non strict relations #
An unbundled relation class stating that r is the nonstrict relation corresponding to the
strict relation s. Compare Preorder.lt_iff_le_not_le. This is mostly meant to provide dot
notation on (⊆) and (⊂).
The relation
ris the nonstrict relation corresponding to the strict relations.
Instances
A version of right_iff_left_not_left with explicit r and s.
⊆ and ⊂ #
Alias of subset_of_eq_of_subset.
Alias of subset_of_subset_of_eq.
Alias of subset_of_eq.
Alias of superset_of_eq.
Alias of subset_trans.
Alias of subset_antisymm.
Alias of superset_antisymm.
Alias of ssubset_of_eq_of_ssubset.
Alias of ssubset_of_ssubset_of_eq.
Alias of ssubset_irrfl.
Alias of ne_of_ssubset.
Alias of ne_of_ssuperset.
Alias of ssubset_trans.
Alias of ssubset_asymm.
Alias of subset_of_ssubset.
Alias of not_subset_of_ssubset.
Alias of not_ssubset_of_subset.
Alias of ssubset_of_subset_not_subset.
Alias of ssubset_of_subset_of_ssubset.
Alias of ssubset_of_ssubset_of_subset.
Alias of ssubset_of_subset_of_ne.
Alias of ssubset_of_ne_of_subset.
Alias of eq_or_ssubset_of_subset.
Alias of ssubset_or_eq_of_subset.
Conversion of bundled order typeclasses to unbundled relation typeclasses #
Equations
- (_ : IsPreorder α fun (x x_1 : α) => x ≤ x_1) = (_ : IsPreorder α fun (x x_1 : α) => x ≤ x_1)
Equations
- (_ : IsPreorder α fun (x x_1 : α) => x ≥ x_1) = (_ : IsPreorder α fun (x x_1 : α) => x ≥ x_1)
Equations
- (_ : IsAntisymm α fun (x x_1 : α) => x < x_1) = (_ : IsAntisymm α fun (x x_1 : α) => x < x_1)
Equations
- (_ : IsAntisymm α fun (x x_1 : α) => x > x_1) = (_ : IsAntisymm α fun (x x_1 : α) => x > x_1)
Equations
- (_ : IsStrictOrder α fun (x x_1 : α) => x < x_1) = (_ : IsStrictOrder α fun (x x_1 : α) => x < x_1)
Equations
- (_ : IsStrictOrder α fun (x x_1 : α) => x > x_1) = (_ : IsStrictOrder α fun (x x_1 : α) => x > x_1)
Equations
- (_ : IsNonstrictStrictOrder α (fun (x x_1 : α) => x ≤ x_1) fun (x x_1 : α) => x < x_1) = (_ : IsNonstrictStrictOrder α (fun (x x_1 : α) => x ≤ x_1) fun (x x_1 : α) => x < x_1)
Equations
- (_ : IsAntisymm α fun (x x_1 : α) => x ≤ x_1) = (_ : IsAntisymm α fun (x x_1 : α) => x ≤ x_1)
Equations
- (_ : IsAntisymm α fun (x x_1 : α) => x ≥ x_1) = (_ : IsAntisymm α (Function.swap fun (x x_1 : α) => x ≤ x_1))
Equations
- (_ : IsPartialOrder α fun (x x_1 : α) => x ≤ x_1) = (_ : IsPartialOrder α fun (x x_1 : α) => x ≤ x_1)
Equations
- (_ : IsPartialOrder α fun (x x_1 : α) => x ≥ x_1) = (_ : IsPartialOrder α fun (x x_1 : α) => x ≥ x_1)
Equations
- (_ : IsTotalPreorder α fun (x x_1 : α) => x ≤ x_1) = (_ : IsTotalPreorder α fun (x x_1 : α) => x ≤ x_1)
Equations
- (_ : IsTotalPreorder α fun (x x_1 : α) => x ≥ x_1) = (_ : IsTotalPreorder α fun (x x_1 : α) => x ≥ x_1)
Equations
- (_ : IsLinearOrder α fun (x x_1 : α) => x ≤ x_1) = (_ : IsLinearOrder α fun (x x_1 : α) => x ≤ x_1)
Equations
- (_ : IsLinearOrder α fun (x x_1 : α) => x ≥ x_1) = (_ : IsLinearOrder α fun (x x_1 : α) => x ≥ x_1)
Equations
- (_ : IsTrichotomous α fun (x x_1 : α) => x < x_1) = (_ : IsTrichotomous α fun (x x_1 : α) => x < x_1)
Equations
- (_ : IsTrichotomous α fun (x x_1 : α) => x > x_1) = (_ : IsTrichotomous α (Function.swap fun (x x_1 : α) => x < x_1))
Equations
- (_ : IsTrichotomous α fun (x x_1 : α) => x ≤ x_1) = (_ : IsTrichotomous α fun (x x_1 : α) => x ≤ x_1)
Equations
- (_ : IsTrichotomous α fun (x x_1 : α) => x ≥ x_1) = (_ : IsTrichotomous α fun (x x_1 : α) => x ≥ x_1)
Equations
- (_ : IsStrictTotalOrder α fun (x x_1 : α) => x < x_1) = (_ : IsStrictTotalOrder α fun (x x_1 : α) => x < x_1)
Equations
- (_ : IsOrderConnected α fun (x x_1 : α) => x < x_1) = (_ : IsOrderConnected α fun (x x_1 : α) => x < x_1)
Equations
- (_ : IsIncompTrans α fun (x x_1 : α) => x < x_1) = (_ : IsIncompTrans α fun (x x_1 : α) => x < x_1)
Equations
- (_ : IsStrictWeakOrder α fun (x x_1 : α) => x < x_1) = (_ : IsStrictWeakOrder α fun (x x_1 : α) => x < x_1)
Equations
- instWellFoundedLTNatInstLTNat = (_ : IsWellFounded ℕ fun (x x_1 : ℕ) => x < x_1)
Equations
- Nat.lt.isWellOrder = (_ : IsWellOrder ℕ fun (x x_1 : ℕ) => x < x_1)
Equations
- (_ : IsWellOrder αᵒᵈ fun (x x_1 : αᵒᵈ) => x > x_1) = h
Equations
- (_ : IsWellOrder αᵒᵈ fun (x x_1 : αᵒᵈ) => x < x_1) = h