Documentation

Mathlib.Order.CompleteLattice.Lemmas

Theory of complete lattices #

This file contains results on complete lattices that need more theory to develop.

Naming conventions #

In lemma names,

Notation #

iSup and iInf under Type #

theorem iSup_bool_eq {α : Type u_1} [CompleteLattice α] {f : Boolα} :
⨆ (b : Bool), f b = f true f false
theorem iInf_bool_eq {α : Type u_1} [CompleteLattice α] {f : Boolα} :
⨅ (b : Bool), f b = f true f false
theorem sup_eq_iSup {α : Type u_1} [CompleteLattice α] (x y : α) :
x y = ⨆ (b : Bool), bif b then x else y
theorem inf_eq_iInf {α : Type u_1} [CompleteLattice α] (x y : α) :
x y = ⨅ (b : Bool), bif b then x else y

iSup and iInf under #

theorem iSup_ge_eq_iSup_nat_add {α : Type u_1} [CompleteLattice α] (u : α) (n : ) :
⨆ (i : ), ⨆ (_ : i n), u i = ⨆ (i : ), u (i + n)
theorem iInf_ge_eq_iInf_nat_add {α : Type u_1} [CompleteLattice α] (u : α) (n : ) :
⨅ (i : ), ⨅ (_ : i n), u i = ⨅ (i : ), u (i + n)
theorem Monotone.iSup_nat_add {α : Type u_1} [CompleteLattice α] {f : α} (hf : Monotone f) (k : ) :
⨆ (n : ), f (n + k) = ⨆ (n : ), f n
theorem Antitone.iInf_nat_add {α : Type u_1} [CompleteLattice α] {f : α} (hf : Antitone f) (k : ) :
⨅ (n : ), f (n + k) = ⨅ (n : ), f n
theorem iSup_iInf_ge_nat_add {α : Type u_1} [CompleteLattice α] (f : α) (k : ) :
⨆ (n : ), ⨅ (i : ), ⨅ (_ : i n), f (i + k) = ⨆ (n : ), ⨅ (i : ), ⨅ (_ : i n), f i
theorem iInf_iSup_ge_nat_add {α : Type u_1} [CompleteLattice α] (f : α) (k : ) :
⨅ (n : ), ⨆ (i : ), ⨆ (_ : i n), f (i + k) = ⨅ (n : ), ⨆ (i : ), ⨆ (_ : i n), f i
theorem sup_iSup_nat_succ {α : Type u_1} [CompleteLattice α] (u : α) :
u 0 ⨆ (i : ), u (i + 1) = ⨆ (i : ), u i
theorem inf_iInf_nat_succ {α : Type u_1} [CompleteLattice α] (u : α) :
u 0 ⨅ (i : ), u (i + 1) = ⨅ (i : ), u i
theorem iInf_nat_gt_zero_eq {α : Type u_1} [CompleteLattice α] (f : α) :
⨅ (i : ), ⨅ (_ : i > 0), f i = ⨅ (i : ), f (i + 1)
theorem iSup_nat_gt_zero_eq {α : Type u_1} [CompleteLattice α] (f : α) :
⨆ (i : ), ⨆ (_ : i > 0), f i = ⨆ (i : ), f (i + 1)

Instances #

theorem sup_sInf_le_iInf_sup {α : Type u_1} [CompleteLattice α] {a : α} {s : Set α} :
a sInf s bs, a b

This is a weaker version of sup_sInf_eq

theorem iSup_inf_le_inf_sSup {α : Type u_1} [CompleteLattice α] {a : α} {s : Set α} :
bs, a b a sSup s

This is a weaker version of inf_sSup_eq

theorem sInf_sup_le_iInf_sup {α : Type u_1} [CompleteLattice α] {a : α} {s : Set α} :
sInf s a bs, b a

This is a weaker version of sInf_sup_eq

theorem iSup_inf_le_sSup_inf {α : Type u_1} [CompleteLattice α] {a : α} {s : Set α} :
bs, b a sSup s a

This is a weaker version of sSup_inf_eq

theorem le_iSup_inf_iSup {α : Type u_1} {ι : Sort u_4} [CompleteLattice α] (f g : ια) :
⨆ (i : ι), f i g i (⨆ (i : ι), f i) ⨆ (i : ι), g i
theorem iInf_sup_iInf_le {α : Type u_1} {ι : Sort u_4} [CompleteLattice α] (f g : ια) :
(⨅ (i : ι), f i) ⨅ (i : ι), g i ⨅ (i : ι), f i g i
theorem disjoint_sSup_left {α : Type u_1} [CompleteLattice α] {a : Set α} {b : α} (d : Disjoint (sSup a) b) {i : α} (hi : i a) :
theorem disjoint_sSup_right {α : Type u_1} [CompleteLattice α] {a : Set α} {b : α} (d : Disjoint b (sSup a)) {i : α} (hi : i a) :
theorem disjoint_of_sSup_disjoint_of_le_of_le {α : Type u_1} [CompleteLattice α] {a b : α} {c d : Set α} (hs : ec, e a) (ht : ed, e b) (hd : Disjoint a b) (he : c d) :
theorem disjoint_of_sSup_disjoint {α : Type u_1} [CompleteLattice α] {a b : Set α} (hd : Disjoint (sSup a) (sSup b)) (he : a b) :
instance ULift.supSet {α : Type u_1} [SupSet α] :
Equations
theorem ULift.down_sSup {α : Type u_1} [SupSet α] (s : Set (ULift.{v, u_1} α)) :
theorem ULift.up_sSup {α : Type u_1} [SupSet α] (s : Set α) :
{ down := sSup s } = sSup (down ⁻¹' s)
instance ULift.infSet {α : Type u_1} [InfSet α] :
Equations
theorem ULift.down_sInf {α : Type u_1} [InfSet α] (s : Set (ULift.{v, u_1} α)) :
theorem ULift.up_sInf {α : Type u_1} [InfSet α] (s : Set α) :
{ down := sInf s } = sInf (down ⁻¹' s)
theorem ULift.down_iSup {α : Type u_1} {ι : Sort u_4} [SupSet α] (f : ιULift.{v, u_1} α) :
(⨆ (i : ι), f i).down = ⨆ (i : ι), (f i).down
theorem ULift.up_iSup {α : Type u_1} {ι : Sort u_4} [SupSet α] (f : ια) :
{ down := ⨆ (i : ι), f i } = ⨆ (i : ι), { down := f i }
theorem ULift.down_iInf {α : Type u_1} {ι : Sort u_4} [InfSet α] (f : ιULift.{v, u_1} α) :
(⨅ (i : ι), f i).down = ⨅ (i : ι), (f i).down
theorem ULift.up_iInf {α : Type u_1} {ι : Sort u_4} [InfSet α] (f : ια) :
{ down := ⨅ (i : ι), f i } = ⨅ (i : ι), { down := f i }
Equations
  • One or more equations did not get rendered due to their size.