Skip to content

Commit

Permalink
feat(topology/algebra/order/compact): remove conditional completeness…
Browse files Browse the repository at this point in the history
… assumption in `is_compact.exists_forall_le` (#18991)
  • Loading branch information
ADedecker committed Jun 21, 2023
1 parent eba7871 commit 3efd324
Show file tree
Hide file tree
Showing 4 changed files with 172 additions and 137 deletions.
6 changes: 5 additions & 1 deletion src/order/directed.lean
Expand Up @@ -131,6 +131,10 @@ lemma directed_on_of_inf_mem [semilattice_inf α] {S : set α}
(H : ∀ ⦃i j⦄, i ∈ S → j ∈ S → i ⊓ j ∈ S) : directed_on (≥) S :=
λ a ha b hb, ⟨a ⊓ b, H ha hb, inf_le_left, inf_le_right⟩

lemma is_total.directed [is_total α r] (f : ι → α) :
directed r f :=
λ i j, or.cases_on (total_of r (f i) (f j)) (λ h, ⟨j, h, refl _⟩) (λ h, ⟨i, refl _, h⟩)

/-- `is_directed α r` states that for any elements `a`, `b` there exists an element `c` such that
`r a c` and `r b c`. -/
class is_directed (α : Type*) (r : α → α → Prop) : Prop :=
Expand All @@ -150,7 +154,7 @@ lemma directed_on_univ_iff : directed_on r set.univ ↔ is_directed α r :=

@[priority 100] -- see Note [lower instance priority]
instance is_total.to_is_directed [is_total α r] : is_directed α r :=
⟨λ a b, or.cases_on (total_of r a b) (λ h, ⟨b, h, refl _⟩) (λ h, ⟨a, refl _, h⟩)⟩
by rw ← directed_id_iff; exact is_total.directed _

lemma is_directed_mono [is_directed α r] (h : ∀ ⦃a b⦄, r a b → s a b) : is_directed α s :=
⟨λ a b, let ⟨c, ha, hb⟩ := is_directed.directed a b in ⟨c, h ha, h hb⟩⟩
Expand Down
219 changes: 144 additions & 75 deletions src/topology/algebra/order/compact.lean
Expand Up @@ -136,80 +136,39 @@ is_compact_iff_compact_space.mp is_compact_Icc
end

/-!
### Min and max elements of a compact set
### Extreme value theorem
-/

variables {α β γ : Type*} [conditionally_complete_linear_order α] [topological_space α]
[order_topology α] [topological_space β] [topological_space γ]

lemma is_compact.Inf_mem {s : set α} (hs : is_compact s) (ne_s : s.nonempty) :
Inf s ∈ s :=
hs.is_closed.cInf_mem ne_s hs.bdd_below

lemma is_compact.Sup_mem {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : Sup s ∈ s :=
@is_compact.Inf_mem αᵒᵈ _ _ _ _ hs ne_s

lemma is_compact.is_glb_Inf {s : set α} (hs : is_compact s) (ne_s : s.nonempty) :
is_glb s (Inf s) :=
is_glb_cInf ne_s hs.bdd_below

lemma is_compact.is_lub_Sup {s : set α} (hs : is_compact s) (ne_s : s.nonempty) :
is_lub s (Sup s) :=
@is_compact.is_glb_Inf αᵒᵈ _ _ _ _ hs ne_s

lemma is_compact.is_least_Inf {s : set α} (hs : is_compact s) (ne_s : s.nonempty) :
is_least s (Inf s) :=
⟨hs.Inf_mem ne_s, (hs.is_glb_Inf ne_s).1
section linear_order

lemma is_compact.is_greatest_Sup {s : set α} (hs : is_compact s) (ne_s : s.nonempty) :
is_greatest s (Sup s) :=
@is_compact.is_least_Inf αᵒᵈ _ _ _ _ hs ne_s
variables {α β γ : Type*} [linear_order α] [topological_space α]
[order_closed_topology α] [topological_space β] [topological_space γ]

lemma is_compact.exists_is_least {s : set α} (hs : is_compact s) (ne_s : s.nonempty) :
∃ x, is_least s x :=
⟨_, hs.is_least_Inf ne_s⟩
begin
haveI : nonempty s := ne_s.to_subtype,
suffices : (s ∩ ⋂ x ∈ s, Iic x).nonempty,
from ⟨this.some, this.some_spec.1, mem_Inter₂.mp this.some_spec.2⟩,
rw bInter_eq_Inter,
by_contra H,
rw not_nonempty_iff_eq_empty at H,
rcases hs.elim_directed_family_closed (λ x : s, Iic ↑x) (λ x, is_closed_Iic) H
((is_total.directed coe).mono_comp _ (λ _ _, Iic_subset_Iic.mpr)) with ⟨x, hx⟩,
exact not_nonempty_iff_eq_empty.mpr hx ⟨x, x.2, le_rfl⟩
end

lemma is_compact.exists_is_greatest {s : set α} (hs : is_compact s) (ne_s : s.nonempty) :
∃ x, is_greatest s x :=
⟨_, hs.is_greatest_Sup ne_s
@is_compact.exists_is_least αᵒᵈ _ _ _ _ hs ne_s

lemma is_compact.exists_is_glb {s : set α} (hs : is_compact s) (ne_s : s.nonempty) :
∃ x ∈ s, is_glb s x :=
⟨_, hs.Inf_mem ne_s, hs.is_glb_Inf ne_s
exists_imp_exists (λ x (hx : is_least s x), ⟨hx.1, hx.is_glb⟩) (hs.exists_is_least ne_s)

lemma is_compact.exists_is_lub {s : set α} (hs : is_compact s) (ne_s : s.nonempty) :
∃ x ∈ s, is_lub s x :=
⟨_, hs.Sup_mem ne_s, hs.is_lub_Sup ne_s⟩

lemma is_compact.exists_Inf_image_eq_and_le {s : set β} (hs : is_compact s) (ne_s : s.nonempty)
{f : β → α} (hf : continuous_on f s) :
∃ x ∈ s, Inf (f '' s) = f x ∧ ∀ y ∈ s, f x ≤ f y :=
let ⟨x, hxs, hx⟩ := (hs.image_of_continuous_on hf).Inf_mem (ne_s.image f)
in ⟨x, hxs, hx.symm, λ y hy,
hx.trans_le $ cInf_le (hs.image_of_continuous_on hf).bdd_below $ mem_image_of_mem f hy⟩

lemma is_compact.exists_Sup_image_eq_and_ge {s : set β} (hs : is_compact s) (ne_s : s.nonempty)
{f : β → α} (hf : continuous_on f s) :
∃ x ∈ s, Sup (f '' s) = f x ∧ ∀ y ∈ s, f y ≤ f x :=
@is_compact.exists_Inf_image_eq_and_le αᵒᵈ _ _ _ _ _ _ hs ne_s _ hf

lemma is_compact.exists_Inf_image_eq {s : set β} (hs : is_compact s) (ne_s : s.nonempty)
{f : β → α} (hf : continuous_on f s) :
∃ x ∈ s, Inf (f '' s) = f x :=
let ⟨x, hxs, hx, _⟩ := hs.exists_Inf_image_eq_and_le ne_s hf in ⟨x, hxs, hx⟩

lemma is_compact.exists_Sup_image_eq :
∀ {s : set β}, is_compact s → s.nonempty → ∀ {f : β → α}, continuous_on f s →
∃ x ∈ s, Sup (f '' s) = f x :=
@is_compact.exists_Inf_image_eq αᵒᵈ _ _ _ _ _

lemma eq_Icc_of_connected_compact {s : set α} (h₁ : is_connected s) (h₂ : is_compact s) :
s = Icc (Inf s) (Sup s) :=
eq_Icc_cInf_cSup_of_connected_bdd_closed h₁ h₂.bdd_below h₂.bdd_above h₂.is_closed

/-!
### Extreme value theorem
-/
@is_compact.exists_is_glb αᵒᵈ _ _ _ _ hs ne_s

/-- The **extreme value theorem**: a continuous function realizes its minimum on a compact set. -/
lemma is_compact.exists_forall_le {s : set β} (hs : is_compact s) (ne_s : s.nonempty)
Expand Down Expand Up @@ -277,6 +236,68 @@ lemma continuous.exists_forall_ge [nonempty β] {f : β → α}
∃ x, ∀ y, f y ≤ f x :=
@continuous.exists_forall_le αᵒᵈ _ _ _ _ _ _ _ hf hlim

/-- A continuous function with compact support has a global minimum. -/
@[to_additive "A continuous function with compact support has a global minimum."]
lemma continuous.exists_forall_le_of_has_compact_mul_support [nonempty β] [has_one α]
{f : β → α} (hf : continuous f) (h : has_compact_mul_support f) :
∃ (x : β), ∀ (y : β), f x ≤ f y :=
begin
obtain ⟨_, ⟨x, rfl⟩, hx⟩ := (h.is_compact_range hf).exists_is_least (range_nonempty _),
rw [mem_lower_bounds, forall_range_iff] at hx,
exact ⟨x, hx⟩,
end

/-- A continuous function with compact support has a global maximum. -/
@[to_additive "A continuous function with compact support has a global maximum."]
lemma continuous.exists_forall_ge_of_has_compact_mul_support [nonempty β] [has_one α]
{f : β → α} (hf : continuous f) (h : has_compact_mul_support f) :
∃ (x : β), ∀ (y : β), f y ≤ f x :=
@continuous.exists_forall_le_of_has_compact_mul_support αᵒᵈ _ _ _ _ _ _ _ _ hf h

/-- A compact set is bounded below -/
lemma is_compact.bdd_below [nonempty α] {s : set α} (hs : is_compact s) : bdd_below s :=
begin
cases s.eq_empty_or_nonempty,
{ rw h,
exact bdd_below_empty },
{ obtain ⟨a, ha, has⟩ := hs.exists_is_least h,
exact ⟨a, has⟩ },
end

/-- A compact set is bounded above -/
lemma is_compact.bdd_above [nonempty α] {s : set α} (hs : is_compact s) : bdd_above s :=
@is_compact.bdd_below αᵒᵈ _ _ _ _ _ hs

/-- A continuous function is bounded below on a compact set. -/
lemma is_compact.bdd_below_image [nonempty α] {f : β → α} {K : set β}
(hK : is_compact K) (hf : continuous_on f K) : bdd_below (f '' K) :=
(hK.image_of_continuous_on hf).bdd_below

/-- A continuous function is bounded above on a compact set. -/
lemma is_compact.bdd_above_image [nonempty α] {f : β → α} {K : set β}
(hK : is_compact K) (hf : continuous_on f K) : bdd_above (f '' K) :=
@is_compact.bdd_below_image αᵒᵈ _ _ _ _ _ _ _ _ hK hf

/-- A continuous function with compact support is bounded below. -/
@[to_additive /-" A continuous function with compact support is bounded below. "-/]
lemma continuous.bdd_below_range_of_has_compact_mul_support [has_one α] {f : β → α}
(hf : continuous f) (h : has_compact_mul_support f) : bdd_below (range f) :=
(h.is_compact_range hf).bdd_below

/-- A continuous function with compact support is bounded above. -/
@[to_additive /-" A continuous function with compact support is bounded above. "-/]
lemma continuous.bdd_above_range_of_has_compact_mul_support [has_one α]
{f : β → α} (hf : continuous f) (h : has_compact_mul_support f) :
bdd_above (range f) :=
@continuous.bdd_below_range_of_has_compact_mul_support αᵒᵈ _ _ _ _ _ _ _ hf h

end linear_order

section conditionally_complete_linear_order

variables {α β γ : Type*} [conditionally_complete_linear_order α] [topological_space α]
[order_closed_topology α] [topological_space β] [topological_space γ]

lemma is_compact.Sup_lt_iff_of_continuous {f : β → α}
{K : set β} (hK : is_compact K) (h0K : K.nonempty) (hf : continuous_on f K) (y : α) :
Sup (f '' K) < y ↔ ∀ x ∈ K, f x < y :=
Expand All @@ -294,23 +315,71 @@ lemma is_compact.lt_Inf_iff_of_continuous {α β : Type*}
y < Inf (f '' K) ↔ ∀ x ∈ K, y < f x :=
@is_compact.Sup_lt_iff_of_continuous αᵒᵈ β _ _ _ _ _ _ hK h0K hf y

/-- A continuous function with compact support has a global minimum. -/
@[to_additive "A continuous function with compact support has a global minimum."]
lemma continuous.exists_forall_le_of_has_compact_mul_support [nonempty β] [has_one α]
{f : β → α} (hf : continuous f) (h : has_compact_mul_support f) :
∃ (x : β), ∀ (y : β), f x ≤ f y :=
begin
obtain ⟨_, ⟨x, rfl⟩, hx⟩ := (h.is_compact_range hf).exists_is_least (range_nonempty _),
rw [mem_lower_bounds, forall_range_iff] at hx,
exact ⟨x, hx⟩,
end
end conditionally_complete_linear_order

/-- A continuous function with compact support has a global maximum. -/
@[to_additive "A continuous function with compact support has a global maximum."]
lemma continuous.exists_forall_ge_of_has_compact_mul_support [nonempty β] [has_one α]
{f : β → α} (hf : continuous f) (h : has_compact_mul_support f) :
∃ (x : β), ∀ (y : β), f y ≤ f x :=
@continuous.exists_forall_le_of_has_compact_mul_support αᵒᵈ _ _ _ _ _ _ _ _ hf h
/-!
### Min and max elements of a compact set
-/

section order_closed_topology

variables {α β γ : Type*} [conditionally_complete_linear_order α] [topological_space α]
[order_closed_topology α] [topological_space β] [topological_space γ]

lemma is_compact.Inf_mem {s : set α} (hs : is_compact s) (ne_s : s.nonempty) :
Inf s ∈ s :=
let ⟨a, ha⟩ := hs.exists_is_least ne_s in
ha.Inf_mem

lemma is_compact.Sup_mem {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : Sup s ∈ s :=
@is_compact.Inf_mem αᵒᵈ _ _ _ _ hs ne_s

lemma is_compact.is_glb_Inf {s : set α} (hs : is_compact s) (ne_s : s.nonempty) :
is_glb s (Inf s) :=
is_glb_cInf ne_s hs.bdd_below

lemma is_compact.is_lub_Sup {s : set α} (hs : is_compact s) (ne_s : s.nonempty) :
is_lub s (Sup s) :=
@is_compact.is_glb_Inf αᵒᵈ _ _ _ _ hs ne_s

lemma is_compact.is_least_Inf {s : set α} (hs : is_compact s) (ne_s : s.nonempty) :
is_least s (Inf s) :=
⟨hs.Inf_mem ne_s, (hs.is_glb_Inf ne_s).1

lemma is_compact.is_greatest_Sup {s : set α} (hs : is_compact s) (ne_s : s.nonempty) :
is_greatest s (Sup s) :=
@is_compact.is_least_Inf αᵒᵈ _ _ _ _ hs ne_s

lemma is_compact.exists_Inf_image_eq_and_le {s : set β} (hs : is_compact s) (ne_s : s.nonempty)
{f : β → α} (hf : continuous_on f s) :
∃ x ∈ s, Inf (f '' s) = f x ∧ ∀ y ∈ s, f x ≤ f y :=
let ⟨x, hxs, hx⟩ := (hs.image_of_continuous_on hf).Inf_mem (ne_s.image f)
in ⟨x, hxs, hx.symm, λ y hy,
hx.trans_le $ cInf_le (hs.image_of_continuous_on hf).bdd_below $ mem_image_of_mem f hy⟩

lemma is_compact.exists_Sup_image_eq_and_ge {s : set β} (hs : is_compact s) (ne_s : s.nonempty)
{f : β → α} (hf : continuous_on f s) :
∃ x ∈ s, Sup (f '' s) = f x ∧ ∀ y ∈ s, f y ≤ f x :=
@is_compact.exists_Inf_image_eq_and_le αᵒᵈ _ _ _ _ _ _ hs ne_s _ hf

lemma is_compact.exists_Inf_image_eq {s : set β} (hs : is_compact s) (ne_s : s.nonempty)
{f : β → α} (hf : continuous_on f s) :
∃ x ∈ s, Inf (f '' s) = f x :=
let ⟨x, hxs, hx, _⟩ := hs.exists_Inf_image_eq_and_le ne_s hf in ⟨x, hxs, hx⟩

lemma is_compact.exists_Sup_image_eq :
∀ {s : set β}, is_compact s → s.nonempty → ∀ {f : β → α}, continuous_on f s →
∃ x ∈ s, Sup (f '' s) = f x :=
@is_compact.exists_Inf_image_eq αᵒᵈ _ _ _ _ _

end order_closed_topology

variables {α β γ : Type*} [conditionally_complete_linear_order α] [topological_space α]
[order_topology α] [topological_space β] [topological_space γ]

lemma eq_Icc_of_connected_compact {s : set α} (h₁ : is_connected s) (h₂ : is_compact s) :
s = Icc (Inf s) (Sup s) :=
eq_Icc_cInf_cSup_of_connected_bdd_closed h₁ h₂.bdd_below h₂.bdd_above h₂.is_closed

lemma is_compact.continuous_Sup {f : γ → β → α}
{K : set β} (hK : is_compact K) (hf : continuous ↿f) :
Expand Down
42 changes: 0 additions & 42 deletions src/topology/order/basic.lean
Expand Up @@ -648,48 +648,6 @@ lemma dense.exists_between [densely_ordered α] {s : set α} (hs : dense s) {x y
∃ z ∈ s, z ∈ Ioo x y :=
hs.exists_mem_open is_open_Ioo (nonempty_Ioo.2 h)

variables [nonempty α] [topological_space β]

/-- A compact set is bounded below -/
lemma is_compact.bdd_below {s : set α} (hs : is_compact s) : bdd_below s :=
begin
by_contra H,
rcases hs.elim_finite_subcover_image (λ x (_ : x ∈ s), @is_open_Ioi _ _ _ _ x) _
with ⟨t, st, ft, ht⟩,
{ refine H (ft.bdd_below.imp $ λ C hC y hy, _),
rcases mem_Union₂.1 (ht hy) with ⟨x, hx, xy⟩,
exact le_trans (hC hx) (le_of_lt xy) },
{ refine λ x hx, mem_Union₂.2 (not_imp_comm.1 _ H),
exact λ h, ⟨x, λ y hy, le_of_not_lt (h.imp $ λ ys, ⟨_, hy, ys⟩)⟩ }
end

/-- A compact set is bounded above -/
lemma is_compact.bdd_above {s : set α} (hs : is_compact s) : bdd_above s :=
@is_compact.bdd_below αᵒᵈ _ _ _ _ _ hs

/-- A continuous function is bounded below on a compact set. -/
lemma is_compact.bdd_below_image {f : β → α} {K : set β}
(hK : is_compact K) (hf : continuous_on f K) : bdd_below (f '' K) :=
(hK.image_of_continuous_on hf).bdd_below

/-- A continuous function is bounded above on a compact set. -/
lemma is_compact.bdd_above_image {f : β → α} {K : set β}
(hK : is_compact K) (hf : continuous_on f K) : bdd_above (f '' K) :=
@is_compact.bdd_below_image αᵒᵈ _ _ _ _ _ _ _ _ hK hf

/-- A continuous function with compact support is bounded below. -/
@[to_additive /-" A continuous function with compact support is bounded below. "-/]
lemma continuous.bdd_below_range_of_has_compact_mul_support [has_one α] {f : β → α}
(hf : continuous f) (h : has_compact_mul_support f) : bdd_below (range f) :=
(h.is_compact_range hf).bdd_below

/-- A continuous function with compact support is bounded above. -/
@[to_additive /-" A continuous function with compact support is bounded above. "-/]
lemma continuous.bdd_above_range_of_has_compact_mul_support [has_one α]
{f : β → α} (hf : continuous f) (h : has_compact_mul_support f) :
bdd_above (range f) :=
@continuous.bdd_below_range_of_has_compact_mul_support αᵒᵈ _ _ _ _ _ _ _ _ hf h

end linear_order

end order_closed_topology
Expand Down
42 changes: 23 additions & 19 deletions src/topology/subset_properties.lean
Expand Up @@ -230,6 +230,19 @@ lemma is_compact.disjoint_nhds_set_right {l : filter α} (hs : is_compact s) :
disjoint l (𝓝ˢ s) ↔ ∀ x ∈ s, disjoint l (𝓝 x) :=
by simpa only [disjoint.comm] using hs.disjoint_nhds_set_left

/-- For every directed family of closed sets whose intersection avoids a compact set,
there exists a single element of the family which itself avoids this compact set. -/
lemma is_compact.elim_directed_family_closed {ι : Type v} [hι : nonempty ι] (hs : is_compact s)
(Z : ι → set α) (hZc : ∀ i, is_closed (Z i)) (hsZ : s ∩ (⋂ i, Z i) = ∅) (hdZ : directed (⊇) Z) :
∃ i : ι, s ∩ Z i = ∅ :=
let ⟨t, ht⟩ := hs.elim_directed_cover (compl ∘ Z) (λ i, (hZc i).is_open_compl)
(by simpa only [subset_def, not_forall, eq_empty_iff_forall_not_mem, mem_Union,
exists_prop, mem_inter_iff, not_and, iff_self, mem_Inter, mem_compl_iff] using hsZ)
(hdZ.mono_comp _ $ λ _ _, compl_subset_compl.mpr)
in
⟨t, by simpa only [subset_def, not_forall, eq_empty_iff_forall_not_mem, mem_Union,
exists_prop, mem_inter_iff, not_and, iff_self, mem_Inter, mem_compl_iff] using ht⟩

/-- For every family of closed sets whose intersection avoids a compact set,
there exists a finite subfamily whose intersection avoids this compact set. -/
lemma is_compact.elim_finite_subfamily_closed {s : set α} {ι : Type v} (hs : is_compact s)
Expand Down Expand Up @@ -273,25 +286,16 @@ lemma is_compact.nonempty_Inter_of_directed_nonempty_compact_closed
(hZn : ∀ i, (Z i).nonempty) (hZc : ∀ i, is_compact (Z i)) (hZcl : ∀ i, is_closed (Z i)) :
(⋂ i, Z i).nonempty :=
begin
apply hι.elim,
intro i₀,
let Z' := λ i, Z i ∩ Z i₀,
suffices : (⋂ i, Z' i).nonempty,
{ exact this.mono (Inter_mono $ λ i, inter_subset_left (Z i) (Z i₀)) },
rw nonempty_iff_ne_empty,
intro H,
obtain ⟨t, ht⟩ : ∃ (t : finset ι), ((Z i₀) ∩ ⋂ (i ∈ t), Z' i) = ∅,
from (hZc i₀).elim_finite_subfamily_closed Z'
(assume i, is_closed.inter (hZcl i) (hZcl i₀)) (by rw [H, inter_empty]),
obtain ⟨i₁, hi₁⟩ : ∃ i₁ : ι, Z i₁ ⊆ Z i₀ ∧ ∀ i ∈ t, Z i₁ ⊆ Z' i,
{ rcases directed.finset_le hZd t with ⟨i, hi⟩,
rcases hZd i i₀ with ⟨i₁, hi₁, hi₁₀⟩,
use [i₁, hi₁₀],
intros j hj,
exact subset_inter (subset.trans hi₁ (hi j hj)) hi₁₀ },
suffices : ((Z i₀) ∩ ⋂ (i ∈ t), Z' i).nonempty,
{ rw nonempty_iff_ne_empty at this, contradiction },
exact (hZn i₁).mono (subset_inter hi₁.left $ subset_Inter₂ hi₁.right),
let i₀ := hι.some,
suffices : (Z i₀ ∩ ⋂ i, Z i).nonempty,
by rwa inter_eq_right_iff_subset.mpr (Inter_subset _ i₀) at this,
simp only [nonempty_iff_ne_empty] at hZn ⊢,
apply mt ((hZc i₀).elim_directed_family_closed Z hZcl),
push_neg,
simp only [← nonempty_iff_ne_empty] at hZn ⊢,
refine ⟨hZd, λ i, _⟩,
rcases hZd i₀ i with ⟨j, hji₀, hji⟩,
exact (hZn j).mono (subset_inter hji₀ hji)
end

/-- Cantor's intersection theorem for sequences indexed by `ℕ`:
Expand Down

0 comments on commit 3efd324

Please sign in to comment.