Documentation

Mathlib.Order.ConditionallyCompleteLattice.Indexed

Indexed sup / inf in conditionally complete lattices #

This file proves lemmas about iSup and iInf for functions valued in a conditionally complete, rather than complete, lattice. We add a prefix c to distinguish them from the versions for complete lattices, giving names ciSup_xxx or ciInf_xxx.

Extension of iSup and iInf from a preorder α to WithTop α and WithBot α

@[simp]
theorem WithTop.iInf_empty {α : Type u_1} {ι : Sort u_4} [Preorder α] [IsEmpty ι] [InfSet α] (f : ιWithTop α) :
⨅ (i : ι), f i =
theorem WithTop.coe_iInf {α : Type u_1} {ι : Sort u_4} [Preorder α] [Nonempty ι] [InfSet α] {f : ια} (hf : BddBelow (Set.range f)) :
(⨅ (i : ι), f i) = ⨅ (i : ι), (f i)
theorem WithTop.coe_iSup {α : Type u_1} {ι : Sort u_4} [Preorder α] [SupSet α] (f : ια) (h : BddAbove (Set.range f)) :
(⨆ (i : ι), f i) = ⨆ (i : ι), (f i)
@[simp]
theorem WithBot.ciSup_empty {α : Type u_1} {ι : Sort u_4} [Preorder α] [IsEmpty ι] [SupSet α] (f : ιWithBot α) :
⨆ (i : ι), f i =
theorem WithBot.coe_iSup {α : Type u_1} {ι : Sort u_4} [Preorder α] [Nonempty ι] [SupSet α] {f : ια} (hf : BddAbove (Set.range f)) :
(⨆ (i : ι), f i) = ⨆ (i : ι), (f i)
theorem WithBot.coe_iInf {α : Type u_1} {ι : Sort u_4} [Preorder α] [InfSet α] (f : ια) (h : BddBelow (Set.range f)) :
(⨅ (i : ι), f i) = ⨅ (i : ι), (f i)
theorem isLUB_ciSup {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Nonempty ι] {f : ια} (H : BddAbove (Set.range f)) :
IsLUB (Set.range f) (⨆ (i : ι), f i)
theorem isLUB_ciSup_set {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] {f : βα} {s : Set β} (H : BddAbove (f '' s)) (Hne : s.Nonempty) :
IsLUB (f '' s) (⨆ (i : s), f i)
theorem isGLB_ciInf {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Nonempty ι] {f : ια} (H : BddBelow (Set.range f)) :
IsGLB (Set.range f) (⨅ (i : ι), f i)
theorem isGLB_ciInf_set {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] {f : βα} {s : Set β} (H : BddBelow (f '' s)) (Hne : s.Nonempty) :
IsGLB (f '' s) (⨅ (i : s), f i)
theorem ciSup_le_iff {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Nonempty ι] {f : ια} {a : α} (hf : BddAbove (Set.range f)) :
iSup f a ∀ (i : ι), f i a
theorem le_ciInf_iff {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Nonempty ι] {f : ια} {a : α} (hf : BddBelow (Set.range f)) :
a iInf f ∀ (i : ι), a f i
theorem ciSup_set_le_iff {α : Type u_1} [ConditionallyCompleteLattice α] {ι : Type u_5} {s : Set ι} {f : ια} {a : α} (hs : s.Nonempty) (hf : BddAbove (f '' s)) :
⨆ (i : s), f i a is, f i a
theorem le_ciInf_set_iff {α : Type u_1} [ConditionallyCompleteLattice α] {ι : Type u_5} {s : Set ι} {f : ια} {a : α} (hs : s.Nonempty) (hf : BddBelow (f '' s)) :
a ⨅ (i : s), f i is, a f i
theorem IsLUB.ciSup_eq {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] {a : α} [Nonempty ι] {f : ια} (H : IsLUB (Set.range f) a) :
⨆ (i : ι), f i = a
theorem IsLUB.ciSup_set_eq {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] {a : α} {s : Set β} {f : βα} (H : IsLUB (f '' s) a) (Hne : s.Nonempty) :
⨆ (i : s), f i = a
theorem IsGLB.ciInf_eq {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] {a : α} [Nonempty ι] {f : ια} (H : IsGLB (Set.range f) a) :
⨅ (i : ι), f i = a
theorem IsGLB.ciInf_set_eq {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] {a : α} {s : Set β} {f : βα} (H : IsGLB (f '' s) a) (Hne : s.Nonempty) :
⨅ (i : s), f i = a
theorem ciSup_le {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Nonempty ι] {f : ια} {c : α} (H : ∀ (x : ι), f x c) :
iSup f c

The indexed supremum of a function is bounded above by a uniform bound

theorem le_ciSup {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] {f : ια} (H : BddAbove (Set.range f)) (c : ι) :
f c iSup f

The indexed supremum of a function is bounded below by the value taken at one point

theorem le_ciSup_of_le {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] {a : α} {f : ια} (H : BddAbove (Set.range f)) (c : ι) (h : a f c) :
a iSup f
theorem ciSup_mono {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] {f : ια} {g : ια} (B : BddAbove (Set.range g)) (H : ∀ (x : ι), f x g x) :

The indexed suprema of two functions are comparable if the functions are pointwise comparable

theorem le_ciSup_set {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] {f : βα} {s : Set β} (H : BddAbove (f '' s)) {c : β} (hc : c s) :
f c ⨆ (i : s), f i
theorem ciInf_mono {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] {f : ια} {g : ια} (B : BddBelow (Set.range f)) (H : ∀ (x : ι), f x g x) :

The indexed infimum of two functions are comparable if the functions are pointwise comparable

theorem le_ciInf {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Nonempty ι] {f : ια} {c : α} (H : ∀ (x : ι), c f x) :
c iInf f

The indexed minimum of a function is bounded below by a uniform lower bound

theorem ciInf_le {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] {f : ια} (H : BddBelow (Set.range f)) (c : ι) :
iInf f f c

The indexed infimum of a function is bounded above by the value taken at one point

theorem ciInf_le_of_le {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] {a : α} {f : ια} (H : BddBelow (Set.range f)) (c : ι) (h : f c a) :
iInf f a
theorem ciInf_set_le {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] {f : βα} {s : Set β} (H : BddBelow (f '' s)) {c : β} (hc : c s) :
⨅ (i : s), f i f c
@[simp]
theorem ciSup_const {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [hι : Nonempty ι] {a : α} :
⨆ (x : ι), a = a
@[simp]
theorem ciInf_const {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Nonempty ι] {a : α} :
⨅ (x : ι), a = a
@[simp]
theorem ciSup_unique {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Unique ι] {s : ια} :
⨆ (i : ι), s i = s default
@[simp]
theorem ciInf_unique {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Unique ι] {s : ια} :
⨅ (i : ι), s i = s default
theorem ciSup_subsingleton {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Subsingleton ι] (i : ι) (s : ια) :
⨆ (i : ι), s i = s i
theorem ciInf_subsingleton {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Subsingleton ι] (i : ι) (s : ια) :
⨅ (i : ι), s i = s i
@[simp]
theorem ciSup_pos {α : Type u_1} [ConditionallyCompleteLattice α] {p : Prop} {f : pα} (hp : p) :
⨆ (h : p), f h = f hp
@[simp]
theorem ciInf_pos {α : Type u_1} [ConditionallyCompleteLattice α] {p : Prop} {f : pα} (hp : p) :
⨅ (h : p), f h = f hp
theorem ciSup_neg {α : Type u_1} [ConditionallyCompleteLattice α] {p : Prop} {f : pα} (hp : ¬p) :
⨆ (h : p), f h = sSup
theorem ciInf_neg {α : Type u_1} [ConditionallyCompleteLattice α] {p : Prop} {f : pα} (hp : ¬p) :
⨅ (h : p), f h = sInf
theorem ciSup_eq_ite {α : Type u_1} [ConditionallyCompleteLattice α] {p : Prop} [Decidable p] {f : pα} :
⨆ (h : p), f h = if h : p then f h else sSup
theorem ciInf_eq_ite {α : Type u_1} [ConditionallyCompleteLattice α] {p : Prop} [Decidable p] {f : pα} :
⨅ (h : p), f h = if h : p then f h else sInf
theorem cbiSup_eq_of_forall {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] {p : ιProp} {f : Subtype pα} (hp : ∀ (i : ι), p i) :
⨆ (i : ι), ⨆ (h : p i), f i, h = iSup f
theorem cbiInf_eq_of_forall {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] {p : ιProp} {f : Subtype pα} (hp : ∀ (i : ι), p i) :
⨅ (i : ι), ⨅ (h : p i), f i, h = iInf f
theorem ciSup_eq_of_forall_le_of_forall_lt_exists_gt {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] {b : α} [Nonempty ι] {f : ια} (h₁ : ∀ (i : ι), f i b) (h₂ : 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 iSup_eq_of_forall_le_of_forall_lt_exists_gt for a version in complete lattices.

theorem ciInf_eq_of_forall_ge_of_forall_gt_exists_lt {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] {b : α} [Nonempty ι] {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 iInf_eq_of_forall_ge_of_forall_gt_exists_lt for a version in complete lattices.

theorem Monotone.ciSup_mem_iInter_Icc_of_antitone {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] [SemilatticeSup β] {f : βα} {g : βα} (hf : Monotone f) (hg : Antitone g) (h : f g) :
⨆ (n : β), f n ⋂ (n : β), Set.Icc (f n) (g n)

Nested intervals lemma: if f is a monotone sequence, g is an antitone sequence, and f n ≤ g n for all n, then ⨆ n, f n belongs to all the intervals [f n, g n].

theorem ciSup_mem_iInter_Icc_of_antitone_Icc {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] [SemilatticeSup β] {f : βα} {g : βα} (h : Antitone fun (n : β) => Set.Icc (f n) (g n)) (h' : ∀ (n : β), f n g n) :
⨆ (n : β), f n ⋂ (n : β), Set.Icc (f n) (g n)

Nested intervals lemma: if [f n, g n] is an antitone sequence of nonempty closed intervals, then ⨆ n, f n belongs to all the intervals [f n, g n].

theorem Set.Iic_ciInf {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Nonempty ι] {f : ια} (hf : BddBelow (Set.range f)) :
Set.Iic (⨅ (i : ι), f i) = ⋂ (i : ι), Set.Iic (f i)
theorem Set.Ici_ciSup {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Nonempty ι] {f : ια} (hf : BddAbove (Set.range f)) :
Set.Ici (⨆ (i : ι), f i) = ⋂ (i : ι), Set.Ici (f i)
theorem ciSup_subtype {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Nonempty ι] {p : ιProp} [Nonempty (Subtype p)] {f : Subtype pα} (hf : BddAbove (Set.range f)) (hf' : sSup iSup f) :
iSup f = ⨆ (i : ι), ⨆ (h : p i), f i, h
theorem ciInf_subtype {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Nonempty ι] {p : ιProp} [Nonempty (Subtype p)] {f : Subtype pα} (hf : BddBelow (Set.range f)) (hf' : iInf f sInf ) :
iInf f = ⨅ (i : ι), ⨅ (h : p i), f i, h
theorem ciSup_subtype' {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Nonempty ι] {p : ιProp} [Nonempty (Subtype p)] {f : (i : ι) → p iα} (hf : BddAbove (Set.range fun (i : Subtype p) => f i )) (hf' : sSup ⨆ (i : Subtype p), f i ) :
⨆ (i : ι), ⨆ (h : p i), f i h = ⨆ (x : Subtype p), f x
theorem ciInf_subtype' {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLattice α] [Nonempty ι] {p : ιProp} [Nonempty (Subtype p)] {f : (i : ι) → p iα} (hf : BddBelow (Set.range fun (i : Subtype p) => f i )) (hf' : ⨅ (i : Subtype p), f i sInf ) :
⨅ (i : ι), ⨅ (h : p i), f i h = ⨅ (x : Subtype p), f x
theorem ciSup_subtype'' {α : Type u_1} [ConditionallyCompleteLattice α] {ι : Type u_5} [Nonempty ι] {s : Set ι} (hs : s.Nonempty) {f : ια} (hf : BddAbove (Set.range fun (i : s) => f i)) (hf' : sSup ⨆ (i : s), f i) :
⨆ (i : s), f i = ts, f t
theorem ciInf_subtype'' {α : Type u_1} [ConditionallyCompleteLattice α] {ι : Type u_5} [Nonempty ι] {s : Set ι} (hs : s.Nonempty) {f : ια} (hf : BddBelow (Set.range fun (i : s) => f i)) (hf' : ⨅ (i : s), f i sInf ) :
⨅ (i : s), f i = ts, f t
theorem csSup_image {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] [Nonempty β] {s : Set β} (hs : s.Nonempty) {f : βα} (hf : BddAbove (Set.range fun (i : s) => f i)) (hf' : sSup ⨆ (i : s), f i) :
sSup (f '' s) = as, f a
theorem csInf_image {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] [Nonempty β] {s : Set β} (hs : s.Nonempty) {f : βα} (hf : BddBelow (Set.range fun (i : s) => f i)) (hf' : ⨅ (i : s), f i sInf ) :
sInf (f '' s) = as, f a
theorem ciSup_image {α : Type u_5} {ι : Type u_6} {ι' : Type u_7} [ConditionallyCompleteLattice α] [Nonempty ι] [Nonempty ι'] {s : Set ι} (hs : s.Nonempty) {f : ιι'} {g : ι'α} (hf : BddAbove (Set.range fun (i : s) => g (f i))) (hg' : sSup ⨆ (i : s), g (f i)) :
if '' s, g i = xs, g (f x)
theorem ciInf_image {α : Type u_5} {ι : Type u_6} {ι' : Type u_7} [ConditionallyCompleteLattice α] [Nonempty ι] [Nonempty ι'] {s : Set ι} (hs : s.Nonempty) {f : ιι'} {g : ι'α} (hf : BddBelow (Set.range fun (i : s) => g (f i))) (hg' : ⨅ (i : s), g (f i) sInf ) :
if '' s, g i = xs, g (f x)
theorem exists_lt_of_lt_ciSup {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrder α] {b : α} [Nonempty ι] {f : ια} (h : b < iSup f) :
∃ (i : ι), b < f i

Indexed version of exists_lt_of_lt_csSup. When b < iSup f, there is an element i such that b < f i.

theorem exists_lt_of_ciInf_lt {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrder α] {a : α} [Nonempty ι] {f : ια} (h : iInf f < a) :
∃ (i : ι), f i < a

Indexed version of exists_lt_of_csInf_lt. When iInf f < a, there is an element i such that f i < a.

theorem lt_ciSup_iff {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrder α] {a : α} [Nonempty ι] {f : ια} (hb : BddAbove (Set.range f)) :
a < iSup f ∃ (i : ι), a < f i
theorem ciInf_lt_iff {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrder α] {a : α} [Nonempty ι] {f : ια} (hb : BddBelow (Set.range f)) :
iInf f < a ∃ (i : ι), f i < a
theorem cbiSup_eq_of_not_forall {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrder α] {p : ιProp} {f : Subtype pα} (hp : ¬∀ (i : ι), p i) :
⨆ (i : ι), ⨆ (h : p i), f i, h = iSup f sSup
theorem cbiInf_eq_of_not_forall {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrder α] {p : ιProp} {f : Subtype pα} (hp : ¬∀ (i : ι), p i) :
⨅ (i : ι), ⨅ (h : p i), f i, h = iInf f sInf
theorem ciInf_eq_bot_of_bot_mem {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrder α] [OrderBot α] {f : ια} (hs : Set.range f) :
theorem ciInf_eq_top_of_top_mem {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrder α] [OrderTop α] {f : ια} (hs : Set.range f) :
theorem ciInf_mem {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrder α] [WellFoundedLT α] [Nonempty ι] (f : ια) :

Lemmas about a conditionally complete linear order with bottom element #

In this case we have Sup ∅ = ⊥, so we can drop some Nonempty/Set.Nonempty assumptions.

@[simp]
theorem ciSup_of_empty {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrderBot α] [IsEmpty ι] (f : ια) :
⨆ (i : ι), f i =
theorem ciSup_false {α : Type u_1} [ConditionallyCompleteLinearOrderBot α] (f : Falseα) :
⨆ (i : False), f i =
theorem le_ciSup_iff' {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrderBot α] {s : ια} {a : α} (h : BddAbove (Set.range s)) :
a iSup s ∀ (b : α), (∀ (i : ι), s i b)a b
theorem le_ciInf_iff' {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrderBot α] [Nonempty ι] {f : ια} {a : α} :
a iInf f ∀ (i : ι), a f i
theorem ciInf_le' {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrderBot α] (f : ια) (i : ι) :
iInf f f i
theorem ciInf_le_of_le' {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrderBot α] {f : ια} {a : α} (c : ι) :
f c aiInf f a
theorem ciSup_le_iff' {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrderBot α] {f : ια} (h : BddAbove (Set.range f)) {a : α} :
⨆ (i : ι), f i a ∀ (i : ι), f i a

In conditionally complete orders with a bottom element, the nonempty condition can be omitted from ciSup_le_iff.

theorem ciSup_le' {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrderBot α] {f : ια} {a : α} (h : ∀ (i : ι), f i a) :
⨆ (i : ι), f i a
theorem lt_ciSup_iff' {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrderBot α] {a : α} {f : ια} (h : BddAbove (Set.range f)) :
a < iSup f ∃ (i : ι), a < f i

In conditionally complete orders with a bottom element, the nonempty condition can be omitted from lt_ciSup_iff.

theorem exists_lt_of_lt_ciSup' {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrderBot α] {f : ια} {a : α} (h : a < ⨆ (i : ι), f i) :
∃ (i : ι), a < f i
theorem ciSup_mono' {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrderBot α] {ι' : Sort u_5} {f : ια} {g : ι'α} (hg : BddAbove (Set.range g)) (h : ∀ (i : ι), ∃ (i' : ι'), f i g i') :
theorem ciSup_or' {α : Type u_1} [ConditionallyCompleteLinearOrderBot α] (p : Prop) (q : Prop) (f : p qα) :
⨆ (h : p q), f h = (⨆ (h : p), f ) ⨆ (h : q), f
theorem GaloisConnection.l_csSup {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] {l : αβ} {u : βα} (gc : GaloisConnection l u) {s : Set α} (hne : s.Nonempty) (hbdd : BddAbove s) :
l (sSup s) = ⨆ (x : s), l x
theorem GaloisConnection.l_csSup' {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] {l : αβ} {u : βα} (gc : GaloisConnection l u) {s : Set α} (hne : s.Nonempty) (hbdd : BddAbove s) :
l (sSup s) = sSup (l '' s)
theorem GaloisConnection.l_ciSup {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] [Nonempty ι] {l : αβ} {u : βα} (gc : GaloisConnection l u) {f : ια} (hf : BddAbove (Set.range f)) :
l (⨆ (i : ι), f i) = ⨆ (i : ι), l (f i)
theorem GaloisConnection.l_ciSup_set {α : Type u_1} {β : Type u_2} {γ : Type u_3} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] {l : αβ} {u : βα} (gc : GaloisConnection l u) {s : Set γ} {f : γα} (hf : BddAbove (f '' s)) (hne : s.Nonempty) :
l (⨆ (i : s), f i) = ⨆ (i : s), l (f i)
theorem GaloisConnection.u_csInf {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] {l : αβ} {u : βα} (gc : GaloisConnection l u) {s : Set β} (hne : s.Nonempty) (hbdd : BddBelow s) :
u (sInf s) = ⨅ (x : s), u x
theorem GaloisConnection.u_csInf' {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] {l : αβ} {u : βα} (gc : GaloisConnection l u) {s : Set β} (hne : s.Nonempty) (hbdd : BddBelow s) :
u (sInf s) = sInf (u '' s)
theorem GaloisConnection.u_ciInf {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] [Nonempty ι] {l : αβ} {u : βα} (gc : GaloisConnection l u) {f : ιβ} (hf : BddBelow (Set.range f)) :
u (⨅ (i : ι), f i) = ⨅ (i : ι), u (f i)
theorem GaloisConnection.u_ciInf_set {α : Type u_1} {β : Type u_2} {γ : Type u_3} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] {l : αβ} {u : βα} (gc : GaloisConnection l u) {s : Set γ} {f : γβ} (hf : BddBelow (f '' s)) (hne : s.Nonempty) :
u (⨅ (i : s), f i) = ⨅ (i : s), u (f i)
theorem OrderIso.map_csSup {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] (e : α ≃o β) {s : Set α} (hne : s.Nonempty) (hbdd : BddAbove s) :
e (sSup s) = ⨆ (x : s), e x
theorem OrderIso.map_csSup' {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] (e : α ≃o β) {s : Set α} (hne : s.Nonempty) (hbdd : BddAbove s) :
e (sSup s) = sSup (e '' s)
theorem OrderIso.map_ciSup {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] [Nonempty ι] (e : α ≃o β) {f : ια} (hf : BddAbove (Set.range f)) :
e (⨆ (i : ι), f i) = ⨆ (i : ι), e (f i)
theorem OrderIso.map_ciSup_set {α : Type u_1} {β : Type u_2} {γ : Type u_3} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] (e : α ≃o β) {s : Set γ} {f : γα} (hf : BddAbove (f '' s)) (hne : s.Nonempty) :
e (⨆ (i : s), f i) = ⨆ (i : s), e (f i)
theorem OrderIso.map_csInf {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] (e : α ≃o β) {s : Set α} (hne : s.Nonempty) (hbdd : BddBelow s) :
e (sInf s) = ⨅ (x : s), e x
theorem OrderIso.map_csInf' {α : Type u_1} {β : Type u_2} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] (e : α ≃o β) {s : Set α} (hne : s.Nonempty) (hbdd : BddBelow s) :
e (sInf s) = sInf (e '' s)
theorem OrderIso.map_ciInf {α : Type u_1} {β : Type u_2} {ι : Sort u_4} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] [Nonempty ι] (e : α ≃o β) {f : ια} (hf : BddBelow (Set.range f)) :
e (⨅ (i : ι), f i) = ⨅ (i : ι), e (f i)
theorem OrderIso.map_ciInf_set {α : Type u_1} {β : Type u_2} {γ : Type u_3} [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] (e : α ≃o β) {s : Set γ} {f : γα} (hf : BddBelow (f '' s)) (hne : s.Nonempty) :
e (⨅ (i : s), f i) = ⨅ (i : s), e (f i)
theorem WithTop.iSup_coe_eq_top {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrderBot α] {f : ια} :
⨆ (x : ι), (f x) = ¬BddAbove (Set.range f)
theorem WithTop.iSup_coe_lt_top {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrderBot α] {f : ια} :
⨆ (x : ι), (f x) < BddAbove (Set.range f)
theorem WithTop.iInf_coe_eq_top {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrderBot α] {f : ια} :
⨅ (x : ι), (f x) = IsEmpty ι
theorem WithTop.iInf_coe_lt_top {α : Type u_1} {ι : Sort u_4} [ConditionallyCompleteLinearOrderBot α] {f : ια} :
⨅ (i : ι), (f i) < Nonempty ι