Finite intervals in a sigma type #
This file provides the LocallyFiniteOrder
instance for the disjoint sum of orders Σ i, α i
and
calculates the cardinality of its finite intervals.
TODO #
Do the same for the lexicographical order
Disjoint sum of orders #
instance
Sigma.instLocallyFiniteOrder
{ι : Type u_1}
{α : ι → Type u_2}
[DecidableEq ι]
[(i : ι) → Preorder (α i)]
[(i : ι) → LocallyFiniteOrder (α i)]
:
LocallyFiniteOrder ((i : ι) × α i)
Equations
- One or more equations did not get rendered due to their size.
theorem
Sigma.card_Icc
{ι : Type u_1}
{α : ι → Type u_2}
[DecidableEq ι]
[(i : ι) → Preorder (α i)]
[(i : ι) → LocallyFiniteOrder (α i)]
(a b : (i : ι) × α i)
:
theorem
Sigma.card_Ico
{ι : Type u_1}
{α : ι → Type u_2}
[DecidableEq ι]
[(i : ι) → Preorder (α i)]
[(i : ι) → LocallyFiniteOrder (α i)]
(a b : (i : ι) × α i)
:
theorem
Sigma.card_Ioc
{ι : Type u_1}
{α : ι → Type u_2}
[DecidableEq ι]
[(i : ι) → Preorder (α i)]
[(i : ι) → LocallyFiniteOrder (α i)]
(a b : (i : ι) × α i)
:
theorem
Sigma.card_Ioo
{ι : Type u_1}
{α : ι → Type u_2}
[DecidableEq ι]
[(i : ι) → Preorder (α i)]
[(i : ι) → LocallyFiniteOrder (α i)]
(a b : (i : ι) × α i)
:
@[simp]
theorem
Sigma.Icc_mk_mk
{ι : Type u_1}
{α : ι → Type u_2}
[DecidableEq ι]
[(i : ι) → Preorder (α i)]
[(i : ι) → LocallyFiniteOrder (α i)]
(i : ι)
(a b : α i)
:
@[simp]
theorem
Sigma.Ico_mk_mk
{ι : Type u_1}
{α : ι → Type u_2}
[DecidableEq ι]
[(i : ι) → Preorder (α i)]
[(i : ι) → LocallyFiniteOrder (α i)]
(i : ι)
(a b : α i)
:
@[simp]
theorem
Sigma.Ioc_mk_mk
{ι : Type u_1}
{α : ι → Type u_2}
[DecidableEq ι]
[(i : ι) → Preorder (α i)]
[(i : ι) → LocallyFiniteOrder (α i)]
(i : ι)
(a b : α i)
:
@[simp]
theorem
Sigma.Ioo_mk_mk
{ι : Type u_1}
{α : ι → Type u_2}
[DecidableEq ι]
[(i : ι) → Preorder (α i)]
[(i : ι) → LocallyFiniteOrder (α i)]
(i : ι)
(a b : α i)
: