Documentation

Mathlib.Order.Bounds.Image

Images of upper/lower bounds under monotone functions #

In this file we prove various results about the behaviour of bounds under monotone/antitone maps.

theorem MonotoneOn.mem_upperBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} {t : Set α} {a : α} (Hf : MonotoneOn f t) (Hst : s t) (Has : a upperBounds s) (Hat : a t) :
f a upperBounds (f '' s)
theorem MonotoneOn.mem_upperBounds_image_self {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {t : Set α} {a : α} (Hf : MonotoneOn f t) :
a upperBounds ta tf a upperBounds (f '' t)
theorem MonotoneOn.mem_lowerBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} {t : Set α} {a : α} (Hf : MonotoneOn f t) (Hst : s t) (Has : a lowerBounds s) (Hat : a t) :
f a lowerBounds (f '' s)
theorem MonotoneOn.mem_lowerBounds_image_self {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {t : Set α} {a : α} (Hf : MonotoneOn f t) :
a lowerBounds ta tf a lowerBounds (f '' t)
theorem MonotoneOn.image_upperBounds_subset_upperBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} {t : Set α} (Hf : MonotoneOn f t) (Hst : s t) :
theorem MonotoneOn.image_lowerBounds_subset_lowerBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} {t : Set α} (Hf : MonotoneOn f t) (Hst : s t) :
theorem MonotoneOn.map_bddAbove {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} {t : Set α} (Hf : MonotoneOn f t) (Hst : s t) :
(upperBounds s t).NonemptyBddAbove (f '' s)

The image under a monotone function on a set t of a subset which has an upper bound in t is bounded above.

theorem MonotoneOn.map_bddBelow {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} {t : Set α} (Hf : MonotoneOn f t) (Hst : s t) :
(lowerBounds s t).NonemptyBddBelow (f '' s)

The image under a monotone function on a set t of a subset which has a lower bound in t is bounded below.

theorem MonotoneOn.map_isLeast {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {t : Set α} {a : α} (Hf : MonotoneOn f t) (Ha : IsLeast t a) :
IsLeast (f '' t) (f a)

A monotone map sends a least element of a set to a least element of its image.

theorem MonotoneOn.map_isGreatest {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {t : Set α} {a : α} (Hf : MonotoneOn f t) (Ha : IsGreatest t a) :
IsGreatest (f '' t) (f a)

A monotone map sends a greatest element of a set to a greatest element of its image.

theorem AntitoneOn.mem_upperBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} {t : Set α} {a : α} (Hf : AntitoneOn f t) (Hst : s t) (Has : a lowerBounds s) :
a tf a upperBounds (f '' s)
theorem AntitoneOn.mem_upperBounds_image_self {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {t : Set α} {a : α} (Hf : AntitoneOn f t) :
a lowerBounds ta tf a upperBounds (f '' t)
theorem AntitoneOn.mem_lowerBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} {t : Set α} {a : α} (Hf : AntitoneOn f t) (Hst : s t) :
a upperBounds sa tf a lowerBounds (f '' s)
theorem AntitoneOn.mem_lowerBounds_image_self {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {t : Set α} {a : α} (Hf : AntitoneOn f t) :
a upperBounds ta tf a lowerBounds (f '' t)
theorem AntitoneOn.image_lowerBounds_subset_upperBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} {t : Set α} (Hf : AntitoneOn f t) (Hst : s t) :
theorem AntitoneOn.image_upperBounds_subset_lowerBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} {t : Set α} (Hf : AntitoneOn f t) (Hst : s t) :
theorem AntitoneOn.map_bddAbove {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} {t : Set α} (Hf : AntitoneOn f t) (Hst : s t) :
(upperBounds s t).NonemptyBddBelow (f '' s)

The image under an antitone function of a set which is bounded above is bounded below.

theorem AntitoneOn.map_bddBelow {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} {t : Set α} (Hf : AntitoneOn f t) (Hst : s t) :
(lowerBounds s t).NonemptyBddAbove (f '' s)

The image under an antitone function of a set which is bounded below is bounded above.

theorem AntitoneOn.map_isGreatest {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {t : Set α} {a : α} (Hf : AntitoneOn f t) :
IsGreatest t aIsLeast (f '' t) (f a)

An antitone map sends a greatest element of a set to a least element of its image.

theorem AntitoneOn.map_isLeast {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {t : Set α} {a : α} (Hf : AntitoneOn f t) :
IsLeast t aIsGreatest (f '' t) (f a)

An antitone map sends a least element of a set to a greatest element of its image.

theorem Monotone.mem_upperBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (Hf : Monotone f) {a : α} {s : Set α} (Ha : a upperBounds s) :
f a upperBounds (f '' s)
theorem Monotone.mem_lowerBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (Hf : Monotone f) {a : α} {s : Set α} (Ha : a lowerBounds s) :
f a lowerBounds (f '' s)
theorem Monotone.image_upperBounds_subset_upperBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (Hf : Monotone f) {s : Set α} :
theorem Monotone.image_lowerBounds_subset_lowerBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (Hf : Monotone f) {s : Set α} :
theorem Monotone.map_bddAbove {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (Hf : Monotone f) {s : Set α} :
BddAbove sBddAbove (f '' s)

The image under a monotone function of a set which is bounded above is bounded above. See also BddAbove.image2.

theorem Monotone.map_bddBelow {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (Hf : Monotone f) {s : Set α} :
BddBelow sBddBelow (f '' s)

The image under a monotone function of a set which is bounded below is bounded below. See also BddBelow.image2.

theorem Monotone.map_isLeast {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (Hf : Monotone f) {a : α} {s : Set α} (Ha : IsLeast s a) :
IsLeast (f '' s) (f a)

A monotone map sends a least element of a set to a least element of its image.

theorem Monotone.map_isGreatest {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (Hf : Monotone f) {a : α} {s : Set α} (Ha : IsGreatest s a) :
IsGreatest (f '' s) (f a)

A monotone map sends a greatest element of a set to a greatest element of its image.

theorem Antitone.mem_upperBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (hf : Antitone f) {a : α} {s : Set α} :
a lowerBounds sf a upperBounds (f '' s)
theorem Antitone.mem_lowerBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (hf : Antitone f) {a : α} {s : Set α} :
a upperBounds sf a lowerBounds (f '' s)
theorem Antitone.image_lowerBounds_subset_upperBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (hf : Antitone f) {s : Set α} :
theorem Antitone.image_upperBounds_subset_lowerBounds_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (hf : Antitone f) {s : Set α} :
theorem Antitone.map_bddAbove {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (hf : Antitone f) {s : Set α} :
BddAbove sBddBelow (f '' s)

The image under an antitone function of a set which is bounded above is bounded below.

theorem Antitone.map_bddBelow {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (hf : Antitone f) {s : Set α} :
BddBelow sBddAbove (f '' s)

The image under an antitone function of a set which is bounded below is bounded above.

theorem Antitone.map_isGreatest {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (hf : Antitone f) {a : α} {s : Set α} :
IsGreatest s aIsLeast (f '' s) (f a)

An antitone map sends a greatest element of a set to a least element of its image.

theorem Antitone.map_isLeast {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (hf : Antitone f) {a : α} {s : Set α} :
IsLeast s aIsGreatest (f '' s) (f a)

An antitone map sends a least element of a set to a greatest element of its image.

theorem mem_upperBounds_image2 {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) (ha : a upperBounds s) (hb : b upperBounds t) :
f a b upperBounds (Set.image2 f s t)
theorem mem_lowerBounds_image2 {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) (ha : a lowerBounds s) (hb : b lowerBounds t) :
f a b lowerBounds (Set.image2 f s t)
theorem image2_upperBounds_upperBounds_subset {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) :
theorem image2_lowerBounds_lowerBounds_subset {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) :
theorem BddAbove.image2 {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) :
BddAbove sBddAbove tBddAbove (Set.image2 f s t)

See also Monotone.map_bddAbove.

theorem BddBelow.image2 {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) :
BddBelow sBddBelow tBddBelow (Set.image2 f s t)

See also Monotone.map_bddBelow.

theorem IsGreatest.image2 {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) (ha : IsGreatest s a) (hb : IsGreatest t b) :
IsGreatest (Set.image2 f s t) (f a b)
theorem IsLeast.image2 {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) (ha : IsLeast s a) (hb : IsLeast t b) :
IsLeast (Set.image2 f s t) (f a b)
theorem mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) (ha : a upperBounds s) (hb : b lowerBounds t) :
f a b upperBounds (Set.image2 f s t)
theorem mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) (ha : a lowerBounds s) (hb : b upperBounds t) :
f a b lowerBounds (Set.image2 f s t)
theorem image2_upperBounds_lowerBounds_subset_upperBounds_image2 {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) :
theorem image2_lowerBounds_upperBounds_subset_lowerBounds_image2 {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) :
theorem BddAbove.bddAbove_image2_of_bddBelow {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) :
BddAbove sBddBelow tBddAbove (Set.image2 f s t)
theorem BddBelow.bddBelow_image2_of_bddAbove {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) :
BddBelow sBddAbove tBddBelow (Set.image2 f s t)
theorem IsGreatest.isGreatest_image2_of_isLeast {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) (ha : IsGreatest s a) (hb : IsLeast t b) :
IsGreatest (Set.image2 f s t) (f a b)
theorem IsLeast.isLeast_image2_of_isGreatest {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Monotone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) (ha : IsLeast s a) (hb : IsGreatest t b) :
IsLeast (Set.image2 f s t) (f a b)
theorem mem_upperBounds_image2_of_mem_lowerBounds {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) (ha : a lowerBounds s) (hb : b lowerBounds t) :
f a b upperBounds (Set.image2 f s t)
theorem mem_lowerBounds_image2_of_mem_upperBounds {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) (ha : a upperBounds s) (hb : b upperBounds t) :
f a b lowerBounds (Set.image2 f s t)
theorem image2_upperBounds_upperBounds_subset_upperBounds_image2 {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) :
theorem image2_lowerBounds_lowerBounds_subset_lowerBounds_image2 {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) :
theorem BddBelow.image2_bddAbove {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) :
BddBelow sBddBelow tBddAbove (Set.image2 f s t)
theorem BddAbove.image2_bddBelow {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) :
BddAbove sBddAbove tBddBelow (Set.image2 f s t)
theorem IsLeast.isGreatest_image2 {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) (ha : IsLeast s a) (hb : IsLeast t b) :
IsGreatest (Set.image2 f s t) (f a b)
theorem IsGreatest.isLeast_image2 {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Antitone (f a)) (ha : IsGreatest s a) (hb : IsGreatest t b) :
IsLeast (Set.image2 f s t) (f a b)
theorem mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) (ha : a lowerBounds s) (hb : b upperBounds t) :
f a b upperBounds (Set.image2 f s t)
theorem mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) (ha : a upperBounds s) (hb : b lowerBounds t) :
f a b lowerBounds (Set.image2 f s t)
theorem image2_lowerBounds_upperBounds_subset_upperBounds_image2 {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) :
theorem image2_upperBounds_lowerBounds_subset_lowerBounds_image2 {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) :
theorem BddBelow.bddAbove_image2_of_bddAbove {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) :
BddBelow sBddAbove tBddAbove (Set.image2 f s t)
theorem BddAbove.bddBelow_image2_of_bddAbove {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) :
BddAbove sBddBelow tBddBelow (Set.image2 f s t)
theorem IsLeast.isGreatest_image2_of_isGreatest {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) (ha : IsLeast s a) (hb : IsGreatest t b) :
IsGreatest (Set.image2 f s t) (f a b)
theorem IsGreatest.isLeast_image2_of_isLeast {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {f : αβγ} {s : Set α} {t : Set β} {a : α} {b : β} (h₀ : ∀ (b : β), Antitone (Function.swap f b)) (h₁ : ∀ (a : α), Monotone (f a)) (ha : IsGreatest s a) (hb : IsLeast t b) :
IsLeast (Set.image2 f s t) (f a b)
theorem bddAbove_prod {α : Type u_1} {β : Type u_2} [Preorder α] [Preorder β] {s : Set (α × β)} :
BddAbove s BddAbove (Prod.fst '' s) BddAbove (Prod.snd '' s)
theorem bddBelow_prod {α : Type u_1} {β : Type u_2} [Preorder α] [Preorder β] {s : Set (α × β)} :
BddBelow s BddBelow (Prod.fst '' s) BddBelow (Prod.snd '' s)
theorem bddAbove_range_prod {ι : Sort x} {α : Type u_1} {β : Type u_2} [Preorder α] [Preorder β] {F : ια × β} :
BddAbove (Set.range F) BddAbove (Set.range (Prod.fst F)) BddAbove (Set.range (Prod.snd F))
theorem bddBelow_range_prod {ι : Sort x} {α : Type u_1} {β : Type u_2} [Preorder α] [Preorder β] {F : ια × β} :
BddBelow (Set.range F) BddBelow (Set.range (Prod.fst F)) BddBelow (Set.range (Prod.snd F))
theorem isLUB_prod {α : Type u_1} {β : Type u_2} [Preorder α] [Preorder β] {s : Set (α × β)} (p : α × β) :
IsLUB s p IsLUB (Prod.fst '' s) p.1 IsLUB (Prod.snd '' s) p.2
theorem isGLB_prod {α : Type u_1} {β : Type u_2} [Preorder α] [Preorder β] {s : Set (α × β)} (p : α × β) :
IsGLB s p IsGLB (Prod.fst '' s) p.1 IsGLB (Prod.snd '' s) p.2
theorem bddAbove_pi {α : Type u} {π : αType u_1} [(a : α) → Preorder (π a)] {s : Set ((a : α) → π a)} :
BddAbove s ∀ (a : α), BddAbove (Function.eval a '' s)
theorem bddBelow_pi {α : Type u} {π : αType u_1} [(a : α) → Preorder (π a)] {s : Set ((a : α) → π a)} :
BddBelow s ∀ (a : α), BddBelow (Function.eval a '' s)
theorem bddAbove_range_pi {α : Type u} {ι : Sort x} {π : αType u_1} [(a : α) → Preorder (π a)] {F : ι(a : α) → π a} :
BddAbove (Set.range F) ∀ (a : α), BddAbove (Set.range fun (i : ι) => F i a)
theorem bddBelow_range_pi {α : Type u} {ι : Sort x} {π : αType u_1} [(a : α) → Preorder (π a)] {F : ι(a : α) → π a} :
BddBelow (Set.range F) ∀ (a : α), BddBelow (Set.range fun (i : ι) => F i a)
theorem isLUB_pi {α : Type u} {π : αType u_1} [(a : α) → Preorder (π a)] {s : Set ((a : α) → π a)} {f : (a : α) → π a} :
IsLUB s f ∀ (a : α), IsLUB (Function.eval a '' s) (f a)
theorem isGLB_pi {α : Type u} {π : αType u_1} [(a : α) → Preorder (π a)] {s : Set ((a : α) → π a)} {f : (a : α) → π a} :
IsGLB s f ∀ (a : α), IsGLB (Function.eval a '' s) (f a)
theorem IsGLB.of_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (hf : ∀ {x y : α}, f x f y x y) {s : Set α} {x : α} (hx : IsGLB (f '' s) (f x)) :
IsGLB s x
theorem IsLUB.of_image {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (hf : ∀ {x y : α}, f x f y x y) {s : Set α} {x : α} (hx : IsLUB (f '' s) (f x)) :
IsLUB s x
theorem BddAbove.range_mono {α : Type u} {β : Type v} [Preorder β] {f : αβ} (g : αβ) (h : ∀ (a : α), f a g a) (hbdd : BddAbove (Set.range g)) :
theorem BddBelow.range_mono {α : Type u} {β : Type v} [Preorder β] (f : αβ) {g : αβ} (h : ∀ (a : α), f a g a) (hbdd : BddBelow (Set.range f)) :
theorem BddAbove.range_comp {α : Type u} {β : Type v} {γ : Type u_1} [Preorder β] [Preorder γ] {f : αβ} {g : βγ} (hf : BddAbove (Set.range f)) (hg : Monotone g) :
BddAbove (Set.range fun (x : α) => g (f x))
theorem BddBelow.range_comp {α : Type u} {β : Type v} {γ : Type u_1} [Preorder β] [Preorder γ] {f : αβ} {g : βγ} (hf : BddBelow (Set.range f)) (hg : Monotone g) :
BddBelow (Set.range fun (x : α) => g (f x))