mathlib documentation

order.complete_lattice

Theory of complete lattices #

Main definitions #

Naming conventions #

In lemma names,

Notation #

def supr {α : Type u_1} [has_Sup α] {ι : Sort u_2} (s : ι → α) :
α

Indexed supremum

Equations
def infi {α : Type u_1} [has_Inf α] {ι : Sort u_2} (s : ι → α) :
α

Indexed infimum

Equations
@[protected, instance]
def has_Inf_to_nonempty (α : Type u_1) [has_Inf α] :
@[protected, instance]
def has_Sup_to_nonempty (α : Type u_1) [has_Sup α] :
@[protected, instance]
def order_dual.has_Sup (α : Type u_1) [has_Inf α] :
Equations
@[protected, instance]
def order_dual.has_Inf (α : Type u_1) [has_Sup α] :
Equations
@[class]
structure complete_semilattice_Sup (α : Type u_8) :
Type u_8

Note that we rarely use complete_semilattice_Sup (in fact, any such object is always a complete_lattice, so it's usually best to start there).

Nevertheless it is sometimes a useful intermediate step in constructions.

Instances
@[instance]
theorem le_Sup {α : Type u_1} [complete_semilattice_Sup α] {s : set α} {a : α} :
a sa Sup s
theorem Sup_le {α : Type u_1} [complete_semilattice_Sup α] {s : set α} {a : α} :
(∀ (b : α), b sb a)Sup s a
theorem is_lub_Sup {α : Type u_1} [complete_semilattice_Sup α] (s : set α) :
is_lub s (Sup s)
theorem is_lub.Sup_eq {α : Type u_1} [complete_semilattice_Sup α] {s : set α} {a : α} (h : is_lub s a) :
Sup s = a
theorem le_Sup_of_le {α : Type u_1} [complete_semilattice_Sup α] {s : set α} {a b : α} (hb : b s) (h : a b) :
a Sup s
theorem Sup_le_Sup {α : Type u_1} [complete_semilattice_Sup α] {s t : set α} (h : s t) :
Sup s Sup t
@[simp]
theorem Sup_le_iff {α : Type u_1} [complete_semilattice_Sup α] {s : set α} {a : α} :
Sup s a ∀ (b : α), b sb a
theorem le_Sup_iff {α : Type u_1} [complete_semilattice_Sup α] {s : set α} {a : α} :
a Sup s ∀ (b : α), b upper_bounds sa b
theorem le_supr_iff {α : Type u_1} {ι : Sort u_4} [complete_semilattice_Sup α] {a : α} {s : ι → α} :
a supr s ∀ (b : α), (∀ (i : ι), s i b)a b
theorem Sup_le_Sup_of_forall_exists_le {α : Type u_1} [complete_semilattice_Sup α] {s t : set α} (h : ∀ (x : α), x s(∃ (y : α) (H : y t), x y)) :
Sup s Sup t
theorem Sup_singleton {α : Type u_1} [complete_semilattice_Sup α] {a : α} :
Sup {a} = a
@[instance]
@[class]
structure complete_semilattice_Inf (α : Type u_8) :
Type u_8

Note that we rarely use complete_semilattice_Inf (in fact, any such object is always a complete_lattice, so it's usually best to start there).

Nevertheless it is sometimes a useful intermediate step in constructions.

Instances
theorem Inf_le {α : Type u_1} [complete_semilattice_Inf α] {s : set α} {a : α} :
a sInf s a
theorem le_Inf {α : Type u_1} [complete_semilattice_Inf α] {s : set α} {a : α} :
(∀ (b : α), b sa b)a Inf s
theorem is_glb_Inf {α : Type u_1} [complete_semilattice_Inf α] (s : set α) :
is_glb s (Inf s)
theorem is_glb.Inf_eq {α : Type u_1} [complete_semilattice_Inf α] {s : set α} {a : α} (h : is_glb s a) :
Inf s = a
theorem Inf_le_of_le {α : Type u_1} [complete_semilattice_Inf α] {s : set α} {a b : α} (hb : b s) (h : b a) :
Inf s a
theorem Inf_le_Inf {α : Type u_1} [complete_semilattice_Inf α] {s t : set α} (h : s t) :
Inf t Inf s
@[simp]
theorem le_Inf_iff {α : Type u_1} [complete_semilattice_Inf α] {s : set α} {a : α} :
a Inf s ∀ (b : α), b sa b
theorem Inf_le_iff {α : Type u_1} [complete_semilattice_Inf α] {s : set α} {a : α} :
Inf s a ∀ (b : α), b lower_bounds sb a
theorem infi_le_iff {α : Type u_1} {ι : Sort u_4} [complete_semilattice_Inf α] {a : α} {s : ι → α} :
infi s a ∀ (b : α), (∀ (i : ι), b s i)b a
theorem Inf_le_Inf_of_forall_exists_le {α : Type u_1} [complete_semilattice_Inf α] {s t : set α} (h : ∀ (x : α), x s(∃ (y : α) (H : y t), y x)) :
Inf t Inf s
theorem Inf_singleton {α : Type u_1} [complete_semilattice_Inf α] {a : α} :
Inf {a} = a
@[class]
structure complete_lattice (α : Type u_8) :
Type u_8
  • sup : α → α → α
  • le : α → α → Prop
  • lt : α → α → Prop
  • le_refl : ∀ (a : α), a a
  • le_trans : ∀ (a b c : α), a bb ca c
  • lt_iff_le_not_le : (∀ (a b : α), a < b a b ¬b a) . "order_laws_tac"
  • le_antisymm : ∀ (a b : α), a bb aa = b
  • le_sup_left : ∀ (a b : α), a a b
  • le_sup_right : ∀ (a b : α), b a b
  • sup_le : ∀ (a b c : α), a cb ca b c
  • inf : α → α → α
  • inf_le_left : ∀ (a b : α), a b a
  • inf_le_right : ∀ (a b : α), a b b
  • le_inf : ∀ (a b c : α), a ba ca b c
  • Sup : set α → α
  • le_Sup : ∀ (s : set α) (a : α), a sa complete_lattice.Sup s
  • Sup_le : ∀ (s : set α) (a : α), (∀ (b : α), b sb a)complete_lattice.Sup s a
  • Inf : set α → α
  • Inf_le : ∀ (s : set α) (a : α), a scomplete_lattice.Inf s a
  • le_Inf : ∀ (s : set α) (a : α), (∀ (b : α), b sa b)a complete_lattice.Inf s
  • top : α
  • bot : α
  • le_top : ∀ (x : α), x
  • bot_le : ∀ (x : α), x

A complete lattice is a bounded lattice which has suprema and infima for every subset.

Instances
@[instance]
def complete_lattice.to_lattice (α : Type u_8) [self : complete_lattice α] :
@[instance]
def complete_lattice.to_has_top (α : Type u_8) [self : complete_lattice α] :
@[instance]
def complete_lattice.to_has_bot (α : Type u_8) [self : complete_lattice α] :
def complete_lattice_of_Inf (α : Type u_1) [H1 : partial_order α] [H2 : has_Inf α] (is_glb_Inf : ∀ (s : set α), is_glb s (Inf s)) :

Create a complete_lattice from a partial_order and Inf function that returns the greatest lower bound of a set. Usually this constructor provides poor definitional equalities. If other fields are known explicitly, they should be provided; for example, if inf is known explicitly, construct the complete_lattice instance as

instance : complete_lattice my_T :=
{ inf := better_inf,
  le_inf := ...,
  inf_le_right := ...,
  inf_le_left := ...
  -- don't care to fix sup, Sup, bot, top
  ..complete_lattice_of_Inf my_T _ }
Equations

Any complete_semilattice_Inf is in fact a complete_lattice.

Note that this construction has bad definitional properties: see the doc-string on complete_lattice_of_Inf.

Equations
def complete_lattice_of_Sup (α : Type u_1) [H1 : partial_order α] [H2 : has_Sup α] (is_lub_Sup : ∀ (s : set α), is_lub s (Sup s)) :

Create a complete_lattice from a partial_order and Sup function that returns the least upper bound of a set. Usually this constructor provides poor definitional equalities. If other fields are known explicitly, they should be provided; for example, if inf is known explicitly, construct the complete_lattice instance as

instance : complete_lattice my_T :=
{ inf := better_inf,
  le_inf := ...,
  inf_le_right := ...,
  inf_le_left := ...
  -- don't care to fix sup, Inf, bot, top
  ..complete_lattice_of_Sup my_T _ }
Equations

Any complete_semilattice_Sup is in fact a complete_lattice.

Note that this construction has bad definitional properties: see the doc-string on complete_lattice_of_Sup.

Equations
@[class]
structure complete_linear_order (α : Type u_8) :
Type u_8

A complete linear order is a linear order whose lattice structure is complete.

Instances
@[instance]
theorem Inf_le_Sup {α : Type u_1} [complete_lattice α] {s : set α} (hs : s.nonempty) :
Inf s Sup s
theorem Sup_union {α : Type u_1} [complete_lattice α] {s t : set α} :
Sup (s t) = Sup s Sup t
theorem Sup_inter_le {α : Type u_1} [complete_lattice α] {s t : set α} :
Sup (s t) Sup s Sup t
theorem Inf_union {α : Type u_1} [complete_lattice α] {s t : set α} :
Inf (s t) = Inf s Inf t
theorem le_Inf_inter {α : Type u_1} [complete_lattice α] {s t : set α} :
Inf s Inf t Inf (s t)
@[simp]
theorem Sup_empty {α : Type u_1} [complete_lattice α] :
@[simp]
theorem Inf_empty {α : Type u_1} [complete_lattice α] :
@[simp]
theorem Sup_univ {α : Type u_1} [complete_lattice α] :
@[simp]
theorem Inf_univ {α : Type u_1} [complete_lattice α] :
@[simp]
theorem Sup_insert {α : Type u_1} [complete_lattice α] {a : α} {s : set α} :
Sup (insert a s) = a Sup s
@[simp]
theorem Inf_insert {α : Type u_1} [complete_lattice α] {a : α} {s : set α} :
Inf (insert a s) = a Inf s
theorem Sup_le_Sup_of_subset_insert_bot {α : Type u_1} [complete_lattice α] {s t : set α} (h : s insert t) :
Sup s Sup t
theorem Inf_le_Inf_of_subset_insert_top {α : Type u_1} [complete_lattice α] {s t : set α} (h : s insert t) :
Inf t Inf s
theorem Sup_pair {α : Type u_1} [complete_lattice α] {a b : α} :
Sup {a, b} = a b
theorem Inf_pair {α : Type u_1} [complete_lattice α] {a b : α} :
Inf {a, b} = a b
@[simp]
theorem Inf_eq_top {α : Type u_1} [complete_lattice α] {s : set α} :
Inf s = ∀ (a : α), a sa =
theorem eq_singleton_top_of_Inf_eq_top_of_nonempty {α : Type u_1} [complete_lattice α] {s : set α} (h_inf : Inf s = ) (hne : s.nonempty) :
s = {}
@[simp]
theorem Sup_eq_bot {α : Type u_1} [complete_lattice α] {s : set α} :
Sup s = ∀ (a : α), a sa =
theorem eq_singleton_bot_of_Sup_eq_bot_of_nonempty {α : Type u_1} [complete_lattice α] {s : set α} (h_sup : Sup s = ) (hne : s.nonempty) :
s = {}
theorem Sup_eq_of_forall_le_of_forall_lt_exists_gt {α : Type u_1} [complete_lattice α] {s : set α} {b : α} (_x : ∀ (a : α), a sa b) (H : ∀ (w : α), w < b(∃ (a : α) (H : a s), w < a)) :
Sup s = b

Introduction rule to prove that b is the supremum of s: it suffices to check that b is larger than all elements of s, and that this is not the case of any w < b. See cSup_eq_of_forall_le_of_forall_lt_exists_gt for a version in conditionally complete lattices.

theorem Inf_eq_of_forall_ge_of_forall_gt_exists_lt {α : Type u_1} [complete_lattice α] {s : set α} {b : α} :
(∀ (a : α), a sb a)(∀ (w : α), b < w(∃ (a : α) (H : a s), a < w))Inf s = b

Introduction rule to prove that b is the infimum of s: it suffices to check that b is smaller than all elements of s, and that this is not the case of any w > b. See cInf_eq_of_forall_ge_of_forall_gt_exists_lt for a version in conditionally complete lattices.

theorem Inf_lt_iff {α : Type u_1} [complete_linear_order α] {s : set α} {b : α} :
Inf s < b ∃ (a : α) (H : a s), a < b
theorem lt_Sup_iff {α : Type u_1} [complete_linear_order α] {s : set α} {b : α} :
b < Sup s ∃ (a : α) (H : a s), b < a
theorem Sup_eq_top {α : Type u_1} [complete_linear_order α] {s : set α} :
Sup s = ∀ (b : α), b < (∃ (a : α) (H : a s), b < a)
theorem Inf_eq_bot {α : Type u_1} [complete_linear_order α] {s : set α} :
Inf s = ∀ (b : α), b > (∃ (a : α) (H : a s), a < b)
theorem lt_supr_iff {α : Type u_1} {ι : Sort u_4} [complete_linear_order α] {a : α} {f : ι → α} :
a < supr f ∃ (i : ι), a < f i
theorem infi_lt_iff {α : Type u_1} {ι : Sort u_4} [complete_linear_order α] {a : α} {f : ι → α} :
infi f < a ∃ (i : ι), f i < a
theorem Sup_range {α : Type u_1} {ι : Sort u_4} [has_Sup α] {f : ι → α} :
theorem Sup_eq_supr' {α : Type u_1} [has_Sup α] (s : set α) :
Sup s = ⨆ (a : s), a
theorem supr_congr {α : Type u_1} {ι : Sort u_4} [has_Sup α] {f g : ι → α} (h : ∀ (i : ι), f i = g i) :
(⨆ (i : ι), f i) = ⨆ (i : ι), g i
theorem function.surjective.supr_comp {α : Type u_1} {ι : Sort u_4} {ι' : Sort u_5} [has_Sup α] {f : ι → ι'} (hf : function.surjective f) (g : ι' → α) :
(⨆ (x : ι), g (f x)) = ⨆ (y : ι'), g y
@[protected]
theorem function.surjective.supr_congr {α : Type u_1} {ι : Sort u_4} {ι' : Sort u_5} [has_Sup α] {f : ι → α} {g : ι' → α} (h : ι → ι') (h1 : function.surjective h) (h2 : ∀ (x : ι), g (h x) = f x) :
(⨆ (x : ι), f x) = ⨆ (y : ι'), g y
theorem supr_congr_Prop {α : Type u_1} [has_Sup α] {p q : Prop} {f₁ : p → α} {f₂ : q → α} (pq : p q) (f : ∀ (x : q), f₁ _ = f₂ x) :
supr f₁ = supr f₂
theorem supr_range' {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [has_Sup α] (g : β → α) (f : ι → β) :
(⨆ (b : (set.range f)), g b) = ⨆ (i : ι), g (f i)
theorem Sup_image' {α : Type u_1} {β : Type u_2} [has_Sup α] {s : set β} {f : β → α} :
Sup (f '' s) = ⨆ (a : s), f a
theorem Inf_range {α : Type u_1} {ι : Sort u_4} [has_Inf α] {f : ι → α} :
theorem Inf_eq_infi' {α : Type u_1} [has_Inf α] (s : set α) :
Inf s = ⨅ (a : s), a
theorem infi_congr {α : Type u_1} {ι : Sort u_4} [has_Inf α] {f g : ι → α} (h : ∀ (i : ι), f i = g i) :
(⨅ (i : ι), f i) = ⨅ (i : ι), g i
theorem function.surjective.infi_comp {α : Type u_1} {ι : Sort u_4} {ι' : Sort u_5} [has_Inf α] {f : ι → ι'} (hf : function.surjective f) (g : ι' → α) :
(⨅ (x : ι), g (f x)) = ⨅ (y : ι'), g y
theorem function.surjective.infi_congr {α : Type u_1} {ι : Sort u_4} {ι' : Sort u_5} [has_Inf α] {f : ι → α} {g : ι' → α} (h : ι → ι') (h1 : function.surjective h) (h2 : ∀ (x : ι), g (h x) = f x) :
(⨅ (x : ι), f x) = ⨅ (y : ι'), g y
theorem infi_congr_Prop {α : Type u_1} [has_Inf α] {p q : Prop} {f₁ : p → α} {f₂ : q → α} (pq : p q) (f : ∀ (x : q), f₁ _ = f₂ x) :
infi f₁ = infi f₂
theorem infi_range' {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [has_Inf α] (g : β → α) (f : ι → β) :
(⨅ (b : (set.range f)), g b) = ⨅ (i : ι), g (f i)
theorem Inf_image' {α : Type u_1} {β : Type u_2} [has_Inf α] {s : set β} {f : β → α} :
Inf (f '' s) = ⨅ (a : s), f a
theorem le_supr {α : Type u_1} {ι : Sort u_4} [complete_lattice α] (f : ι → α) (i : ι) :
f i supr f
theorem infi_le {α : Type u_1} {ι : Sort u_4} [complete_lattice α] (f : ι → α) (i : ι) :
infi f f i
theorem le_supr' {α : Type u_1} {ι : Sort u_4} [complete_lattice α] (f : ι → α) (i : ι) :
(:f i supr f:)
theorem infi_le' {α : Type u_1} {ι : Sort u_4} [complete_lattice α] (f : ι → α) (i : ι) :
(:infi f f i:)
theorem is_lub_supr {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f : ι → α} :
is_lub (set.range f) (⨆ (j : ι), f j)
theorem is_glb_infi {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f : ι → α} :
is_glb (set.range f) (⨅ (j : ι), f j)
theorem is_lub.supr_eq {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f : ι → α} {a : α} (h : is_lub (set.range f) a) :
(⨆ (j : ι), f j) = a
theorem is_glb.infi_eq {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f : ι → α} {a : α} (h : is_glb (set.range f) a) :
(⨅ (j : ι), f j) = a
theorem le_supr_of_le {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f : ι → α} {a : α} (i : ι) (h : a f i) :
a supr f
theorem infi_le_of_le {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f : ι → α} {a : α} (i : ι) (h : f i a) :
infi f a
theorem le_supr₂ {α : Type u_1} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] {f : Π (i : ι), κ i → α} (i : ι) (j : κ i) :
f i j ⨆ (i : ι) (j : κ i), f i j
theorem infi₂_le {α : Type u_1} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] {f : Π (i : ι), κ i → α} (i : ι) (j : κ i) :
(⨅ (i : ι) (j : κ i), f i j) f i j
theorem le_supr₂_of_le {α : Type u_1} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] {a : α} {f : Π (i : ι), κ i → α} (i : ι) (j : κ i) (h : a f i j) :
a ⨆ (i : ι) (j : κ i), f i j
theorem infi₂_le_of_le {α : Type u_1} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] {a : α} {f : Π (i : ι), κ i → α} (i : ι) (j : κ i) (h : f i j a) :
(⨅ (i : ι) (j : κ i), f i j) a
theorem supr_le {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f : ι → α} {a : α} (h : ∀ (i : ι), f i a) :
supr f a
theorem le_infi {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f : ι → α} {a : α} (h : ∀ (i : ι), a f i) :
a infi f
theorem supr₂_le {α : Type u_1} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] {a : α} {f : Π (i : ι), κ i → α} (h : ∀ (i : ι) (j : κ i), f i j a) :
(⨆ (i : ι) (j : κ i), f i j) a
theorem le_infi₂ {α : Type u_1} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] {a : α} {f : Π (i : ι), κ i → α} (h : ∀ (i : ι) (j : κ i), a f i j) :
a ⨅ (i : ι) (j : κ i), f i j
theorem supr₂_le_supr {α : Type u_1} {ι : Sort u_4} [complete_lattice α] (κ : ι → Sort u_2) (f : ι → α) :
(⨆ (i : ι) (j : κ i), f i) ⨆ (i : ι), f i
theorem infi_le_infi₂ {α : Type u_1} {ι : Sort u_4} [complete_lattice α] (κ : ι → Sort u_2) (f : ι → α) :
(⨅ (i : ι), f i) ⨅ (i : ι) (j : κ i), f i
theorem supr_mono {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f g : ι → α} (h : ∀ (i : ι), f i g i) :
theorem infi_mono {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f g : ι → α} (h : ∀ (i : ι), f i g i) :
theorem supr₂_mono {α : Type u_1} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] {f g : Π (i : ι), κ i → α} (h : ∀ (i : ι) (j : κ i), f i j g i j) :
(⨆ (i : ι) (j : κ i), f i j) ⨆ (i : ι) (j : κ i), g i j
theorem infi₂_mono {α : Type u_1} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] {f g : Π (i : ι), κ i → α} (h : ∀ (i : ι) (j : κ i), f i j g i j) :
(⨅ (i : ι) (j : κ i), f i j) ⨅ (i : ι) (j : κ i), g i j
theorem supr_mono' {α : Type u_1} {ι : Sort u_4} {ι' : Sort u_5} [complete_lattice α] {f : ι → α} {g : ι' → α} (h : ∀ (i : ι), ∃ (i' : ι'), f i g i') :
theorem infi_mono' {α : Type u_1} {ι : Sort u_4} {ι' : Sort u_5} [complete_lattice α] {f : ι → α} {g : ι' → α} (h : ∀ (i' : ι'), ∃ (i : ι), f i g i') :
theorem supr₂_mono' {α : Type u_1} {ι : Sort u_4} {ι' : Sort u_5} {κ : ι → Sort u_6} {κ' : ι' → Sort u_7} [complete_lattice α] {f : Π (i : ι), κ i → α} {g : Π (i' : ι'), κ' i' → α} (h : ∀ (i : ι) (j : κ i), ∃ (i' : ι') (j' : κ' i'), f i j g i' j') :
(⨆ (i : ι) (j : κ i), f i j) ⨆ (i : ι') (j : κ' i), g i j
theorem infi₂_mono' {α : Type u_1} {ι : Sort u_4} {ι' : Sort u_5} {κ : ι → Sort u_6} {κ' : ι' → Sort u_7} [complete_lattice α] {f : Π (i : ι), κ i → α} {g : Π (i' : ι'), κ' i' → α} (h : ∀ (i : ι') (j : κ' i), ∃ (i' : ι) (j' : κ i'), f i' j' g i j) :
(⨅ (i : ι) (j : κ i), f i j) ⨅ (i : ι') (j : κ' i), g i j
theorem supr_const_mono {α : Type u_1} {ι : Sort u_4} {ι' : Sort u_5} [complete_lattice α] {a : α} (h : ι → ι') :
(⨆ (i : ι), a) ⨆ (j : ι'), a
theorem infi_const_mono {α : Type u_1} {ι : Sort u_4} {ι' : Sort u_5} [complete_lattice α] {a : α} (h : ι' → ι) :
(⨅ (i : ι), a) ⨅ (j : ι'), a
theorem bsupr_mono {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f : ι → α} {p q : ι → Prop} (hpq : ∀ (i : ι), p iq i) :
(⨆ (i : ι) (h : p i), f i) ⨆ (i : ι) (h : q i), f i
theorem binfi_mono {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f : ι → α} {p q : ι → Prop} (hpq : ∀ (i : ι), p iq i) :
(⨅ (i : ι) (h : q i), f i) ⨅ (i : ι) (h : p i), f i
@[simp]
theorem supr_le_iff {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f : ι → α} {a : α} :
supr f a ∀ (i : ι), f i a
@[simp]
theorem le_infi_iff {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f : ι → α} {a : α} :
a infi f ∀ (i : ι), a f i
@[simp]
theorem supr₂_le_iff {α : Type u_1} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] {a : α} {f : Π (i : ι), κ i → α} :
(⨆ (i : ι) (j : κ i), f i j) a ∀ (i : ι) (j : κ i), f i j a
@[simp]
theorem le_infi₂_iff {α : Type u_1} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] {a : α} {f : Π (i : ι), κ i → α} :
(a ⨅ (i : ι) (j : κ i), f i j) ∀ (i : ι) (j : κ i), a f i j
theorem supr_lt_iff {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f : ι → α} {a : α} :
supr f < a ∃ (b : α), b < a ∀ (i : ι), f i b
theorem lt_infi_iff {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f : ι → α} {a : α} :
a < infi f ∃ (b : α), a < b ∀ (i : ι), b f i
theorem Sup_eq_supr {α : Type u_1} [complete_lattice α] {s : set α} :
Sup s = ⨆ (a : α) (H : a s), a
theorem Inf_eq_infi {α : Type u_1} [complete_lattice α] {s : set α} :
Inf s = ⨅ (a : α) (H : a s), a
theorem Sup_sUnion {α : Type u_1} [complete_lattice α] (s : set (set α)) :
Sup (⋃₀s) = ⨆ (t : set α) (H : t s), Sup t
theorem Inf_sUnion {α : Type u_1} [complete_lattice α] (s : set (set α)) :
Inf (⋃₀s) = ⨅ (t : set α) (H : t s), Inf t
theorem monotone.le_map_supr {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [complete_lattice α] {s : ι → α} [complete_lattice β] {f : α → β} (hf : monotone f) :
(⨆ (i : ι), f (s i)) f (supr s)
theorem monotone.le_map_supr₂ {α : Type u_1} {β : Type u_2} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] [complete_lattice β] {f : α → β} (hf : monotone f) (s : Π (i : ι), κ i → α) :
(⨆ (i : ι) (j : κ i), f (s i j)) f (⨆ (i : ι) (j : κ i), s i j)
theorem monotone.le_map_Sup {α : Type u_1} {β : Type u_2} [complete_lattice α] [complete_lattice β] {s : set α} {f : α → β} (hf : monotone f) :
(⨆ (a : α) (H : a s), f a) f (Sup s)
theorem antitone.le_map_infi {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [complete_lattice α] {s : ι → α} [complete_lattice β] {f : α → β} (hf : antitone f) :
(⨆ (i : ι), f (s i)) f (infi s)
theorem antitone.le_map_infi₂ {α : Type u_1} {β : Type u_2} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] [complete_lattice β] {f : α → β} (hf : antitone f) (s : Π (i : ι), κ i → α) :
(⨆ (i : ι) (j : κ i), f (s i j)) f (⨅ (i : ι) (j : κ i), s i j)
theorem antitone.le_map_Inf {α : Type u_1} {β : Type u_2} [complete_lattice α] [complete_lattice β] {s : set α} {f : α → β} (hf : antitone f) :
(⨆ (a : α) (H : a s), f a) f (Inf s)
theorem order_iso.map_supr {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [complete_lattice α] [complete_lattice β] (f : α ≃o β) (x : ι → α) :
f (⨆ (i : ι), x i) = ⨆ (i : ι), f (x i)
theorem order_iso.map_Sup {α : Type u_1} {β : Type u_2} [complete_lattice α] [complete_lattice β] (f : α ≃o β) (s : set α) :
f (Sup s) = ⨆ (a : α) (H : a s), f a
theorem supr_comp_le {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {ι' : Sort u_2} (f : ι' → α) (g : ι → ι') :
(⨆ (x : ι), f (g x)) ⨆ (y : ι'), f y
theorem monotone.supr_comp_eq {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [complete_lattice α] [preorder β] {f : β → α} (hf : monotone f) {s : ι → β} (hs : ∀ (x : β), ∃ (i : ι), x s i) :
(⨆ (x : ι), f (s x)) = ⨆ (y : β), f y
theorem monotone.map_infi_le {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [complete_lattice α] {s : ι → α} [complete_lattice β] {f : α → β} (hf : monotone f) :
f (infi s) ⨅ (i : ι), f (s i)
theorem monotone.map_infi₂_le {α : Type u_1} {β : Type u_2} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] [complete_lattice β] {f : α → β} (hf : monotone f) (s : Π (i : ι), κ i → α) :
f (⨅ (i : ι) (j : κ i), s i j) ⨅ (i : ι) (j : κ i), f (s i j)
theorem monotone.map_Inf_le {α : Type u_1} {β : Type u_2} [complete_lattice α] [complete_lattice β] {s : set α} {f : α → β} (hf : monotone f) :
f (Inf s) ⨅ (a : α) (H : a s), f a
theorem antitone.map_supr_le {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [complete_lattice α] {s : ι → α} [complete_lattice β] {f : α → β} (hf : antitone f) :
f (supr s) ⨅ (i : ι), f (s i)
theorem antitone.map_supr₂_le {α : Type u_1} {β : Type u_2} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] [complete_lattice β] {f : α → β} (hf : antitone f) (s : Π (i : ι), κ i → α) :
f (⨆ (i : ι) (j : κ i), s i j) ⨅ (i : ι) (j : κ i), f (s i j)
theorem antitone.map_Sup_le {α : Type u_1} {β : Type u_2} [complete_lattice α] [complete_lattice β] {s : set α} {f : α → β} (hf : antitone f) :
f (Sup s) ⨅ (a : α) (H : a s), f a
theorem order_iso.map_infi {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [complete_lattice α] [complete_lattice β] (f : α ≃o β) (x : ι → α) :
f (⨅ (i : ι), x i) = ⨅ (i : ι), f (x i)
theorem order_iso.map_Inf {α : Type u_1} {β : Type u_2} [complete_lattice α] [complete_lattice β] (f : α ≃o β) (s : set α) :
f (Inf s) = ⨅ (a : α) (H : a s), f a
theorem le_infi_comp {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {ι' : Sort u_2} (f : ι' → α) (g : ι → ι') :
(⨅ (y : ι'), f y) ⨅ (x : ι), f (g x)
theorem monotone.infi_comp_eq {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [complete_lattice α] [preorder β] {f : β → α} (hf : monotone f) {s : ι → β} (hs : ∀ (x : β), ∃ (i : ι), s i x) :
(⨅ (x : ι), f (s x)) = ⨅ (y : β), f y
theorem supr_const_le {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {x : α} :
(⨆ (h : ι), x) x
theorem le_infi_const {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {x : α} :
x ⨅ (h : ι), x
theorem infi_const {α : Type u_1} {ι : Sort u_4} [complete_lattice α] [nonempty ι] {a : α} :
(⨅ (b : ι), a) = a
theorem supr_const {α : Type u_1} {ι : Sort u_4} [complete_lattice α] [nonempty ι] {a : α} :
(⨆ (b : ι), a) = a
@[simp]
theorem infi_top {α : Type u_1} {ι : Sort u_4} [complete_lattice α] :
(⨅ (i : ι), ) =
@[simp]
theorem supr_bot {α : Type u_1} {ι : Sort u_4} [complete_lattice α] :
(⨆ (i : ι), ) =
@[simp]
theorem supr_eq_bot {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {s : ι → α} :
supr s = ∀ (i : ι), s i =
@[simp]
theorem infi_eq_top {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {s : ι → α} :
infi s = ∀ (i : ι), s i =
@[simp]
theorem supr₂_eq_bot {α : Type u_1} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] {f : Π (i : ι), κ i → α} :
(⨆ (i : ι) (j : κ i), f i j) = ∀ (i : ι) (j : κ i), f i j =
@[simp]
theorem infi₂_eq_top {α : Type u_1} {ι : Sort u_4} {κ : ι → Sort u_6} [complete_lattice α] {f : Π (i : ι), κ i → α} :
(⨅ (i : ι) (j : κ i), f i j) = ∀ (i : ι) (j : κ i), f i j =
@[simp]
theorem supr_pos {α : Type u_1} [complete_lattice α] {p : Prop} {f : p → α} (hp : p) :
(⨆ (h : p), f h) = f hp
@[simp]
theorem infi_pos {α : Type u_1} [complete_lattice α] {p : Prop} {f : p → α} (hp : p) :
(⨅ (h : p), f h) = f hp
@[simp]
theorem supr_neg {α : Type u_1} [complete_lattice α] {p : Prop} {f : p → α} (hp : ¬p) :
(⨆ (h : p), f h) =
@[simp]
theorem infi_neg {α : Type u_1} [complete_lattice α] {p : Prop} {f : p → α} (hp : ¬p) :
(⨅ (h : p), f h) =
theorem supr_eq_of_forall_le_of_forall_lt_exists_gt {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {b : α} {f : ι → α} (h₁ : ∀ (i : ι), f i b) (h₂ : ∀ (w : α), w < b(∃ (i : ι), w < f i)) :
(⨆ (i : ι), f i) = b

Introduction rule to prove that b is the supremum of f: it suffices to check that b is larger than f i for all i, and that this is not the case of any w<b. See csupr_eq_of_forall_le_of_forall_lt_exists_gt for a version in conditionally complete lattices.

theorem infi_eq_of_forall_ge_of_forall_gt_exists_lt {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {b : α} {f : ι → α} (h₁ : ∀ (i : ι), b f i) (h₂ : ∀ (w : α), b < w(∃ (i : ι), f i < w)) :
(⨅ (i : ι), f i) = b

Introduction rule to prove that b is the infimum of f: it suffices to check that b is smaller than f i for all i, and that this is not the case of any w>b. See cinfi_eq_of_forall_ge_of_forall_gt_exists_lt for a version in conditionally complete lattices.

theorem supr_eq_dif {α : Type u_1} [complete_lattice α] {p : Prop} [decidable p] (a : p → α) :
(⨆ (h : p), a h) = dite p (λ (h : p), a h) (λ (h : ¬p), )
theorem supr_eq_if {α : Type u_1} [complete_lattice α] {p : Prop} [decidable p] (a : α) :
(⨆ (h : p), a) = ite p a
theorem infi_eq_dif {α : Type u_1} [complete_lattice α] {p : Prop} [decidable p] (a : p → α) :
(⨅ (h : p), a h) = dite p (λ (h : p), a h) (λ (h : ¬p), )
theorem infi_eq_if {α : Type u_1} [complete_lattice α] {p : Prop} [decidable p] (a : α) :
(⨅ (h : p), a) = ite p a
theorem supr_comm {α : Type u_1} {ι : Sort u_4} {ι' : Sort u_5} [complete_lattice α] {f : ι → ι' → α} :
(⨆ (i : ι) (j : ι'), f i j) = ⨆ (j : ι') (i : ι), f i j
theorem infi_comm {α : Type u_1} {ι : Sort u_4} {ι' : Sort u_5} [complete_lattice α] {f : ι → ι' → α} :
(⨅ (i : ι) (j : ι'), f i j) = ⨅ (j : ι') (i : ι), f i j
@[simp]
theorem infi_infi_eq_left {α : Type u_1} {β : Type u_2} [complete_lattice α] {b : β} {f : Π (x : β), x = b → α} :
(⨅ (x : β) (h : x = b), f x h) = f b rfl
@[simp]
theorem infi_infi_eq_right {α : Type u_1} {β : Type u_2} [complete_lattice α] {b : β} {f : Π (x : β), b = x → α} :
(⨅ (x : β) (h : b = x), f x h) = f b rfl
@[simp]
theorem supr_supr_eq_left {α : Type u_1} {β : Type u_2} [complete_lattice α] {b : β} {f : Π (x : β), x = b → α} :
(⨆ (x : β) (h : x = b), f x h) = f b rfl
@[simp]
theorem supr_supr_eq_right {α : Type u_1} {β : Type u_2} [complete_lattice α] {b : β} {f : Π (x : β), b = x → α} :
(⨆ (x : β) (h : b = x), f x h) = f b rfl
theorem infi_subtype {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {p : ι → Prop} {f : subtype p → α} :
(⨅ (x : subtype p), f x) = ⨅ (i : ι) (h : p i), f i, h⟩
theorem infi_subtype' {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {p : ι → Prop} {f : Π (i : ι), p i → α} :
(⨅ (i : ι) (h : p i), f i h) = ⨅ (x : subtype p), f x _
theorem infi_subtype'' {α : Type u_1} [complete_lattice α] {ι : Type u_2} (s : set ι) (f : ι → α) :
(⨅ (i : s), f i) = ⨅ (t : ι) (H : t s), f t
theorem infi_inf_eq {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f g : ι → α} :
(⨅ (x : ι), f x g x) = (⨅ (x : ι), f x) ⨅ (x : ι), g x
theorem infi_inf {α : Type u_1} {ι : Sort u_4} [complete_lattice α] [h : nonempty ι] {f : ι → α} {a : α} :
(⨅ (x : ι), f x) a = ⨅ (x : ι), f x a
theorem inf_infi {α : Type u_1} {ι : Sort u_4} [complete_lattice α] [nonempty ι] {f : ι → α} {a : α} :
(a ⨅ (x : ι), f x) = ⨅ (x : ι), a f x
theorem binfi_inf {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {p : ι → Prop} {f : Π (i : ι), p i → α} {a : α} (h : ∃ (i : ι), p i) :
(⨅ (i : ι) (h : p i), f i h) a = ⨅ (i : ι) (h : p i), f i h a
theorem inf_binfi {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {p : ι → Prop} {f : Π (i : ι), p i → α} {a : α} (h : ∃ (i : ι), p i) :
(a ⨅ (i : ι) (h : p i), f i h) = ⨅ (i : ι) (h : p i), a f i h
theorem supr_sup_eq {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {f g : ι → α} :
(⨆ (x : ι), f x g x) = (⨆ (x : ι), f x) ⨆ (x : ι), g x
theorem supr_sup {α : Type u_1} {ι : Sort u_4} [complete_lattice α] [h : nonempty ι] {f : ι → α} {a : α} :
(⨆ (x : ι), f x) a = ⨆ (x : ι), f x a
theorem sup_supr {α : Type u_1} {ι : Sort u_4} [complete_lattice α] [nonempty ι] {f : ι → α} {a : α} :
(a ⨆ (x : ι), f x) = ⨆ (x : ι), a f x

supr and infi under Prop #

@[simp]
theorem infi_false {α : Type u_1} [complete_lattice α] {s : false → α} :
@[simp]
theorem supr_false {α : Type u_1} [complete_lattice α] {s : false → α} :
theorem supr_true {α : Type u_1} [complete_lattice α] {s : true → α} :
theorem infi_true {α : Type u_1} [complete_lattice α] {s : true → α} :
@[simp]
theorem infi_exists {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {p : ι → Prop} {f : Exists p → α} :
(⨅ (x : Exists p), f x) = ⨅ (i : ι) (h : p i), f _
@[simp]
theorem supr_exists {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {p : ι → Prop} {f : Exists p → α} :
(⨆ (x : Exists p), f x) = ⨆ (i : ι) (h : p i), f _
theorem infi_and {α : Type u_1} [complete_lattice α] {p q : Prop} {s : p q → α} :
infi s = ⨅ (h₁ : p) (h₂ : q), s _
theorem infi_and' {α : Type u_1} [complete_lattice α] {p q : Prop} {s : p → q → α} :
(⨅ (h₁ : p) (h₂ : q), s h₁ h₂) = ⨅ (h : p q), s _ _

The symmetric case of infi_and, useful for rewriting into a infimum over a conjunction

theorem supr_and {α : Type u_1} [complete_lattice α] {p q : Prop} {s : p q → α} :
supr s = ⨆ (h₁ : p) (h₂ : q), s _
theorem supr_and' {α : Type u_1} [complete_lattice α] {p q : Prop} {s : p → q → α} :
(⨆ (h₁ : p) (h₂ : q), s h₁ h₂) = ⨆ (h : p q), s _ _

The symmetric case of supr_and, useful for rewriting into a supremum over a conjunction

theorem infi_or {α : Type u_1} [complete_lattice α] {p q : Prop} {s : p q → α} :
infi s = (⨅ (h : p), s _) ⨅ (h : q), s _
theorem supr_or {α : Type u_1} [complete_lattice α] {p q : Prop} {s : p q → α} :
(⨆ (x : p q), s x) = (⨆ (i : p), s _) ⨆ (j : q), s _
theorem supr_dite {α : Type u_1} {ι : Sort u_4} [complete_lattice α] (p : ι → Prop) [decidable_pred p] (f : Π (i : ι), p i → α) (g : Π (i : ι), ¬p i → α) :
(⨆ (i : ι), dite (p i) (λ (h : p i), f i h) (λ (h : ¬p i), g i h)) = (⨆ (i : ι) (h : p i), f i h) ⨆ (i : ι) (h : ¬p i), g i h
theorem supr_ite {α : Type u_1} {ι : Sort u_4} [complete_lattice α] (p : ι → Prop) [decidable_pred p] (f g : ι → α) :
(⨆ (i : ι), ite (p i) (f i) (g i)) = (⨆ (i : ι) (h : p i), f i) ⨆ (i : ι) (h : ¬p i), g i
theorem infi_dite {α : Type u_1} {ι : Sort u_4} [complete_lattice α] (p : ι → Prop) [decidable_pred p] (f : Π (i : ι), p i → α) (g : Π (i : ι), ¬p i → α) :
(⨅ (i : ι), dite (p i) (λ (h : p i), f i h) (λ (h : ¬p i), g i h)) = (⨅ (i : ι) (h : p i), f i h) ⨅ (i : ι) (h : ¬p i), g i h
theorem infi_ite {α : Type u_1} {ι : Sort u_4} [complete_lattice α] (p : ι → Prop) [decidable_pred p] (f g : ι → α) :
(⨅ (i : ι), ite (p i) (f i) (g i)) = (⨅ (i : ι) (h : p i), f i) ⨅ (i : ι) (h : ¬p i), g i
theorem infi_range {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [complete_lattice α] {g : β → α} {f : ι → β} :
(⨅ (b : β) (H : b set.range f), g b) = ⨅ (i : ι), g (f i)
theorem supr_range {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [complete_lattice α] {g : β → α} {f : ι → β} :
(⨆ (b : β) (H : b set.range f), g b) = ⨆ (i : ι), g (f i)
theorem Inf_image {α : Type u_1} {β : Type u_2} [complete_lattice α] {s : set β} {f : β → α} :
Inf (f '' s) = ⨅ (a : β) (H : a s), f a
theorem Sup_image {α : Type u_1} {β : Type u_2} [complete_lattice α] {s : set β} {f : β → α} :
Sup (f '' s) = ⨆ (a : β) (H : a s), f a
theorem infi_emptyset {α : Type u_1} {β : Type u_2} [complete_lattice α] {f : β → α} :
(⨅ (x : β) (H : x ), f x) =
theorem supr_emptyset {α : Type u_1} {β : Type u_2} [complete_lattice α] {f : β → α} :
(⨆ (x : β) (H : x ), f x) =
theorem infi_univ {α : Type u_1} {β : Type u_2} [complete_lattice α] {f : β → α} :
(⨅ (x : β) (H : x set.univ), f x) = ⨅ (x : β), f x
theorem supr_univ {α : Type u_1} {β : Type u_2} [complete_lattice α] {f : β → α} :
(⨆ (x : β) (H : x set.univ), f x) = ⨆ (x : β), f x
theorem infi_union {α : Type u_1} {β : Type u_2} [complete_lattice α] {f : β → α} {s t : set β} :
(⨅ (x : β) (H : x s t), f x) = (⨅ (x : β) (H : x s), f x) ⨅ (x : β) (H : x t), f x
theorem infi_split {α : Type u_1} {β : Type u_2} [complete_lattice α] (f : β → α) (p : β → Prop) :
(⨅ (i : β), f i) = (⨅ (i : β) (h : p i), f i) ⨅ (i : β) (h : ¬p i), f i
theorem infi_split_single {α : Type u_1} {β : Type u_2} [complete_lattice α] (f : β → α) (i₀ : β) :
(⨅ (i : β), f i) = f i₀ ⨅ (i : β) (h : i i₀), f i
theorem infi_le_infi_of_subset {α : Type u_1} {β : Type u_2} [complete_lattice α] {f : β → α} {s t : set β} (h : s t) :
(⨅ (x : β) (H : x t), f x) ⨅ (x : β) (H : x s), f x
theorem supr_union {α : Type u_1} {β : Type u_2} [complete_lattice α] {f : β → α} {s t : set β} :
(⨆ (x : β) (H : x s t), f x) = (⨆ (x : β) (H : x s), f x) ⨆ (x : β) (H : x t), f x
theorem supr_split {α : Type u_1} {β : Type u_2} [complete_lattice α] (f : β → α) (p : β → Prop) :
(⨆ (i : β), f i) = (⨆ (i : β) (h : p i), f i) ⨆ (i : β) (h : ¬p i), f i
theorem supr_split_single {α : Type u_1} {β : Type u_2} [complete_lattice α] (f : β → α) (i₀ : β) :
(⨆ (i : β), f i) = f i₀ ⨆ (i : β) (h : i i₀), f i
theorem supr_le_supr_of_subset {α : Type u_1} {β : Type u_2} [complete_lattice α] {f : β → α} {s t : set β} (h : s t) :
(⨆ (x : β) (H : x s), f x) ⨆ (x : β) (H : x t), f x
theorem infi_insert {α : Type u_1} {β : Type u_2} [complete_lattice α] {f : β → α} {s : set β} {b : β} :
(⨅ (x : β) (H : x insert b s), f x) = f b ⨅ (x : β) (H : x s), f x
theorem supr_insert {α : Type u_1} {β : Type u_2} [complete_lattice α] {f : β → α} {s : set β} {b : β} :
(⨆ (x : β) (H : x insert b s), f x) = f b ⨆ (x : β) (H : x s), f x
theorem infi_singleton {α : Type u_1} {β : Type u_2} [complete_lattice α] {f : β → α} {b : β} :
(⨅ (x : β) (H : x {b}), f x) = f b
theorem infi_pair {α : Type u_1} {β : Type u_2} [complete_lattice α] {f : β → α} {a b : β} :
(⨅ (x : β) (H : x {a, b}), f x) = f a f b
theorem supr_singleton {α : Type u_1} {β : Type u_2} [complete_lattice α] {f : β → α} {b : β} :
(⨆ (x : β) (H : x {b}), f x) = f b
theorem supr_pair {α : Type u_1} {β : Type u_2} [complete_lattice α] {f : β → α} {a b : β} :
(⨆ (x : β) (H : x {a, b}), f x) = f a f b
theorem infi_image {α : Type u_1} {β : Type u_2} [complete_lattice α] {γ : Type u_3} {f : β → γ} {g : γ → α} {t : set β} :
(⨅ (c : γ) (H : c f '' t), g c) = ⨅ (b : β) (H : b t), g (f b)
theorem supr_image {α : Type u_1} {β : Type u_2} [complete_lattice α] {γ : Type u_3} {f : β → γ} {g : γ → α} {t : set β} :
(⨆ (c : γ) (H : c f '' t), g c) = ⨆ (b : β) (H : b t), g (f b)
theorem supr_extend_bot {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [complete_lattice α] {e : ι → β} (he : function.injective e) (f : ι → α) :
(⨆ (j : β), function.extend e f j) = ⨆ (i : ι), f i

supr and infi under Type #

theorem supr_of_empty' {α : Type u_1} {ι : Sort u_2} [has_Sup α] [is_empty ι] (f : ι → α) :
theorem supr_of_empty {α : Type u_1} {ι : Sort u_4} [complete_lattice α] [is_empty ι] (f : ι → α) :
theorem infi_of_empty' {α : Type u_1} {ι : Sort u_2} [has_Inf α] [is_empty ι] (f : ι → α) :
theorem infi_of_empty {α : Type u_1} {ι : Sort u_4} [complete_lattice α] [is_empty ι] (f : ι → α) :
theorem supr_bool_eq {α : Type u_1} [complete_lattice α] {f : bool → α} :
(⨆ (b : bool), f b) = f tt f ff
theorem infi_bool_eq {α : Type u_1} [complete_lattice α] {f : bool → α} :
(⨅ (b : bool), f b) = f tt f ff
theorem sup_eq_supr {α : Type u_1} [complete_lattice α] (x y : α) :
x y = ⨆ (b : bool), cond b x y
theorem inf_eq_infi {α : Type u_1} [complete_lattice α] (x y : α) :
x y = ⨅ (b : bool), cond b x y
theorem is_glb_binfi {α : Type u_1} {β : Type u_2} [complete_lattice α] {s : set β} {f : β → α} :
is_glb (f '' s) (⨅ (x : β) (H : x s), f x)
theorem supr_subtype {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {p : ι → Prop} {f : subtype p → α} :
(⨆ (x : subtype p), f x) = ⨆ (i : ι) (h : p i), f i, h⟩
theorem supr_subtype' {α : Type u_1} {ι : Sort u_4} [complete_lattice α] {p : ι → Prop} {f : Π (i : ι), p i → α} :
(⨆ (i : ι) (h : p i), f i h) = ⨆ (x : subtype p), f x _
theorem supr_subtype'' {α : Type u_1} [complete_lattice α] {ι : Type u_2} (s : set ι) (f : ι → α) :
(⨆ (i : s), f i) = ⨆ (t : ι) (H : t s), f t
theorem is_lub_bsupr {α : Type u_1} {β : Type u_2} [complete_lattice α] {s : set β} {f : β → α} :
is_lub (f '' s) (⨆ (x : β) (H : x s), f x)
theorem infi_sigma {α : Type u_1} {β : Type u_2} [complete_lattice α] {p : β → Type u_3} {f : sigma p → α} :
(⨅ (x : sigma p), f x) = ⨅ (i : β) (h : p i), f i, h⟩
theorem supr_sigma {α : Type u_1} {β : Type u_2} [complete_lattice α] {p : β → Type u_3} {f : sigma p → α} :
(⨆ (x : sigma p), f x) = ⨆ (i : β) (h : p i), f i, h⟩
theorem infi_prod {α : Type u_1} {β : Type u_2} [complete_lattice α] {γ : Type u_3} {f : β × γ → α} :
(⨅ (x : β × γ), f x) = ⨅ (i : β) (j : γ), f (i, j)
theorem supr_prod {α : Type u_1} {β : Type u_2} [complete_lattice α] {γ : Type u_3} {f : β × γ → α} :
(⨆ (x : β × γ), f x) = ⨆ (i : β) (j : γ), f (i, j)
theorem infi_sum {α : Type u_1} {β : Type u_2} [complete_lattice α] {γ : Type u_3} {f : β γ → α} :
(⨅ (x : β γ), f x) = (⨅ (i : β), f (sum.inl i)) ⨅ (j : γ), f (sum.inr j)
theorem supr_sum {α : Type u_1} {β : Type u_2} [complete_lattice α] {γ : Type u_3} {f : β γ → α} :
(⨆ (x : β γ), f x) = (⨆ (i : β), f (sum.inl i)) ⨆ (j : γ), f (sum.inr j)
theorem supr_option {α : Type u_1} {β : Type u_2} [complete_lattice α] (f : option β → α) :
(⨆ (o : option β), f o) = f none ⨆ (b : β), f (some b)
theorem infi_option {α : Type u_1} {β : Type u_2} [complete_lattice α] (f : option β → α) :
(⨅ (o : option β), f o) = f none ⨅ (b : β), f (some b)
theorem supr_option_elim {α : Type u_1} {β : Type u_2} [complete_lattice α] (a : α) (f : β → α) :
(⨆ (o : option β), o.elim a f) = a ⨆ (b : β), f b

A version of supr_option useful for rewriting right-to-left.

theorem infi_option_elim {α : Type u_1} {β : Type u_2} [complete_lattice α] (a : α) (f : β → α) :
(⨅ (o : option β), o.elim a f) = a ⨅ (b : β), f b

A version of infi_option useful for rewriting right-to-left.

theorem supr_ne_bot_subtype {α : Type u_1} {ι : Sort u_4} [complete_lattice α] (f : ι → α) :
(⨆ (i : {i // f i }), f i) = ⨆ (i : ι), f i

When taking the supremum of f : ι → α, the elements of ι on which f gives can be dropped, without changing the result.

theorem infi_ne_top_subtype {α : Type u_1} {ι : Sort u_4} [complete_lattice α] (f : ι → α) :
(⨅ (i : {i // f i }), f i) = ⨅ (i : ι), f i

When taking the infimum of f : ι → α, the elements of ι on which f gives can be dropped, without changing the result.

supr and infi under #

theorem supr_ge_eq_supr_nat_add {α : Type u_1} [complete_lattice α] {u : → α} (n : ) :
(⨆ (i : ) (H : i n), u i) = ⨆ (i : ), u (i + n)
theorem infi_ge_eq_infi_nat_add {α : Type u_1} [complete_lattice α] {u : → α} (n : ) :
(⨅ (i : ) (H : i n), u i) = ⨅ (i : ), u (i + n)
theorem monotone.supr_nat_add {α : Type u_1} [complete_lattice α] {f : → α} (hf : monotone f) (k : ) :
(⨆ (n : ), f (n + k)) = ⨆ (n : ), f n
@[simp]
theorem supr_infi_ge_nat_add {α : Type u_1} [complete_lattice α] (f : → α) (k : ) :
(⨆ (n : ), ⨅ (i : ) (H : i n), f (i + k)) = ⨆ (n : ), ⨅ (i : ) (H : i n), f i
theorem sup_supr_nat_succ {α : Type u_1} [complete_lattice α] (u : → α) :
(u 0 ⨆ (i : ), u (i + 1)) = ⨆ (i : ), u i
theorem inf_infi_nat_succ {α : Type u_1} [complete_lattice α] (u : → α) :
(u 0 ⨅ (i : ), u (i + 1)) = ⨅ (i : ), u i
theorem supr_eq_top {α : Type u_1} {ι : Sort u_4} [complete_linear_order α] (f : ι → α) :
supr f = ∀ (b : α), b < (∃ (i : ι), b < f i)
theorem infi_eq_bot {α : Type u_1} {ι : Sort u_4} [complete_linear_order α] (f : ι → α) :
infi f = ∀ (b : α), b > (∃ (i : ι), f i < b)

Instances #

@[protected, instance]
Equations
@[simp]
theorem Sup_Prop_eq {s : set Prop} :
Sup s = ∃ (p : Prop) (H : p s), p
@[simp]
theorem Inf_Prop_eq {s : set Prop} :
Inf s = ∀ (p : Prop), p s → p
@[simp]
theorem infi_Prop_eq {ι : Sort u_4} {p : ι → Prop} :
(⨅ (i : ι), p i) = ∀ (i : ι), p i
@[simp]
theorem supr_Prop_eq {ι : Sort u_4} {p : ι → Prop} :
(⨆ (i : ι), p i) = ∃ (i : ι), p i
@[protected, instance]
def pi.has_Sup {α : Type u_1} {β : α → Type u_2} [Π (i : α), has_Sup (β i)] :
has_Sup (Π (i : α), β i)
Equations
@[protected, instance]
def pi.has_Inf {α : Type u_1} {β : α → Type u_2} [Π (i : α), has_Inf (β i)] :
has_Inf (Π (i : α), β i)
Equations
@[protected, instance]
def pi.complete_lattice {α : Type u_1} {β : α → Type u_2} [Π (i : α), complete_lattice (β i)] :
complete_lattice (Π (i : α), β i)
Equations
theorem Inf_apply {α : Type u_1} {β : α → Type u_2} [Π (i : α), has_Inf (β i)] {s : set (Π (a : α), β a)} {a : α} :
Inf s a = ⨅ (f : s), f a
@[simp]
theorem infi_apply {α : Type u_1} {β : α → Type u_2} {ι : Sort u_3} [Π (i : α), has_Inf (β i)] {f : ι → Π (a : α), β a} {a : α} :
(⨅ (i : ι), f i) a = ⨅ (i : ι), f i a
theorem Sup_apply {α : Type u_1} {β : α → Type u_2} [Π (i : α), has_Sup (β i)] {s : set (Π (a : α), β a)} {a : α} :
Sup s a = ⨆ (f : s), f a
theorem unary_relation_Sup_iff {α : Type u_1} (s : set (α → Prop)) {a : α} :
Sup s a ∃ (r : α → Prop), r s r a
theorem binary_relation_Sup_iff {α : Type u_1} {β : Type u_2} (s : set (α → β → Prop)) {a : α} {b : β} :
Sup s a b ∃ (r : α → β → Prop), r s r a b
@[simp]
theorem supr_apply {α : Type u_1} {β : α → Type u_2} {ι : Sort u_3} [Π (i : α), has_Sup (β i)] {f : ι → Π (a : α), β a} {a : α} :
(⨆ (i : ι), f i) a = ⨆ (i : ι), f i a
theorem monotone_Sup_of_monotone {α : Type u_1} {β : Type u_2} [preorder α] [complete_lattice β] {s : set (α → β)} (m_s : ∀ (f : α → β), f smonotone f) :
theorem monotone_Inf_of_monotone {α : Type u_1} {β : Type u_2} [preorder α] [complete_lattice β] {s : set (α → β)} (m_s : ∀ (f : α → β), f smonotone f) :
@[protected, instance]
def prod.has_Inf (α : Type u_1) (β : Type u_2) [has_Inf α] [has_Inf β] :
has_Inf × β)
Equations
@[protected, instance]
def prod.has_Sup (α : Type u_1) (β : Type u_2) [has_Sup α] [has_Sup β] :
has_Sup × β)
Equations
@[protected, instance]
def prod.complete_lattice (α : Type u_1) (β : Type u_2) [complete_lattice α] [complete_lattice β] :
Equations
theorem sup_Inf_le_infi_sup {α : Type u_1} [complete_lattice α] {a : α} {s : set α} :
a Inf s ⨅ (b : α) (H : b s), a b

This is a weaker version of sup_Inf_eq

theorem Inf_sup_le_infi_sup {α : Type u_1} [complete_lattice α] {a : α} {s : set α} :
Inf s a ⨅ (b : α) (H : b s), b a

This is a weaker version of Inf_sup_eq

theorem supr_inf_le_inf_Sup {α : Type u_1} [complete_lattice α] {a : α} {s : set α} :
(⨆ (b : α) (H : b s), a b) a Sup s

This is a weaker version of inf_Sup_eq

theorem supr_inf_le_Sup_inf {α : Type u_1} [complete_lattice α] {a : α} {s : set α} :
(⨆ (b : α) (H : b s), b a) Sup s a

This is a weaker version of Sup_inf_eq

theorem disjoint_Sup_left {α : Type u_1} [complete_lattice α] {a : set α} {b : α} (d : disjoint (Sup a) b) {i : α} (hi : i a) :
theorem disjoint_Sup_right {α : Type u_1} [complete_lattice α] {a : set α} {b : α} (d : disjoint b (Sup a)) {i : α} (hi : i a) :
@[protected]
def function.injective.complete_lattice {α : Type u_1} {β : Type u_2} [has_sup α] [has_inf α] [has_Sup α] [has_Inf α] [has_top α] [has_bot α] [complete_lattice β] (f : α → β) (hf : function.injective f) (map_sup : ∀ (a b : α), f (a b) = f a f b) (map_inf : ∀ (a b : α), f (a b) = f a f b) (map_Sup : ∀ (s : set α), f (Sup s) = ⨆ (a : α) (H : a s), f a) (map_Inf : ∀ (s : set α), f (Inf s) = ⨅ (a : α) (H : a s), f a) (map_top : f = ) (map_bot : f = ) :

Pullback a complete_lattice along an injection.

Equations