Lexicographic order on a sigma type #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
This defines the lexicographical order of two arbitrary relations on a sigma type and proves some
lemmas about psigma.lex
, which is defined in core Lean.
Given a relation in the index type and a relation on each summand, the lexicographical order on the
sigma type relates a
and b
if their summands are related or they are in the same summand and
related by the summand's relation.
See also #
Related files are:
data.finset.colex
: Colexicographic order on finite sets.data.list.lex
: Lexicographic order on lists.data.sigma.order
: Lexicographic order onΣ i, α i
per say.data.psigma.order
: Lexicographic order onΣ' i, α i
.data.prod.lex
: Lexicographic order onα × β
. Can be thought of as the special case ofsigma.lex
where all summands are the same
- left : ∀ {ι : Type u_1} {α : ι → Type u_2} {r : ι → ι → Prop} {s : Π (i : ι), α i → α i → Prop} {i j : ι} (a : α i) (b : α j), r i j → sigma.lex r s ⟨i, a⟩ ⟨j, b⟩
- right : ∀ {ι : Type u_1} {α : ι → Type u_2} {r : ι → ι → Prop} {s : Π (i : ι), α i → α i → Prop} {i : ι} (a b : α i), s i a b → sigma.lex r s ⟨i, a⟩ ⟨i, b⟩
The lexicographical order on a sigma type. It takes in a relation on the index type and a
relation for each summand. a
is related to b
iff their summands are related or they are in the
same summand and are related through the summand's relation.
Equations
- sigma.lex.decidable r s = λ (a b : Σ (i : ι), α i), decidable_of_decidable_of_iff infer_instance _
Equations
- psigma.lex.decidable r s = λ (a b : Σ' (i : ι), α i), decidable_of_decidable_of_iff infer_instance _