diff --git a/Mathlib.lean b/Mathlib.lean index 72071e3040e97..38c49b4e2ec07 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -342,6 +342,7 @@ import Mathlib.Analysis.Convex.Segment import Mathlib.Analysis.Convex.Slope import Mathlib.Analysis.Convex.Star import Mathlib.Analysis.Convex.Strict +import Mathlib.Analysis.Convex.Topology import Mathlib.Analysis.Hofer import Mathlib.Analysis.LocallyConvex.BalancedCoreHull import Mathlib.Analysis.LocallyConvex.Basic diff --git a/Mathlib/Analysis/Convex/Topology.lean b/Mathlib/Analysis/Convex/Topology.lean new file mode 100644 index 0000000000000..9d4cdcdbd1da9 --- /dev/null +++ b/Mathlib/Analysis/Convex/Topology.lean @@ -0,0 +1,374 @@ +/- +Copyright (c) 2020 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Alexander Bentkamp, Yury Kudryashov + +! This file was ported from Lean 3 source module analysis.convex.topology +! leanprover-community/mathlib commit 0e3aacdc98d25e0afe035c452d876d28cbffaa7e +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathlib.Analysis.Convex.Combination +import Mathlib.Analysis.Convex.Strict +import Mathlib.Topology.PathConnected +import Mathlib.Topology.Algebra.Affine +import Mathlib.Topology.Algebra.Module.Basic + +/-! +# Topological properties of convex sets + +We prove the following facts: + +* `Convex.interior` : interior of a convex set is convex; +* `Convex.closure` : closure of a convex set is convex; +* `Set.Finite.isCompact_convexHull` : convex hull of a finite set is compact; +* `Set.Finite.isClosed_convexHull` : convex hull of a finite set is closed. +-/ + +-- Porting note: this does not exist in Lean 4: +--assert_not_exists Norm + +open Metric Set + +open Pointwise Convex + +variable {ΞΉ π•œ E : Type _} + +theorem Real.convex_iff_isPreconnected {s : Set ℝ} : Convex ℝ s ↔ IsPreconnected s := + convex_iff_ordConnected.trans isPreconnected_iff_ordConnected.symm +#align real.convex_iff_is_preconnected Real.convex_iff_isPreconnected + +alias Real.convex_iff_isPreconnected ↔ _ IsPreconnected.convex +#align is_preconnected.convex IsPreconnected.convex + +/-! ### Standard simplex -/ + + +section stdSimplex + +variable [Fintype ΞΉ] + +/-- Every vector in `stdSimplex π•œ ΞΉ` has `max`-norm at most `1`. -/ +theorem stdSimplex_subset_closedBall : stdSimplex ℝ ΞΉ βŠ† Metric.closedBall 0 1 := by + intro f hf + rw [Metric.mem_closedBall, dist_pi_le_iff zero_le_one] + intro x + rw [Pi.zero_apply, Real.dist_0_eq_abs, abs_of_nonneg <| hf.1 x] + exact (mem_Icc_of_mem_stdSimplex hf x).2 +#align std_simplex_subset_closed_ball stdSimplex_subset_closedBall + +variable (ΞΉ) + +/-- `stdSimplex ℝ ΞΉ` is bounded. -/ +theorem bounded_stdSimplex : Metric.Bounded (stdSimplex ℝ ΞΉ) := + (Metric.bounded_iff_subset_ball 0).2 ⟨1, stdSimplex_subset_closedBall⟩ +#align bounded_std_simplex bounded_stdSimplex + +/-- `stdSimplex ℝ ΞΉ` is closed. -/ +theorem isClosed_stdSimplex : IsClosed (stdSimplex ℝ ΞΉ) := + (stdSimplex_eq_inter ℝ ΞΉ).symm β–Έ + IsClosed.inter (isClosed_interα΅’ fun i => isClosed_le continuous_const (continuous_apply i)) + (isClosed_eq (continuous_finset_sum _ fun x _ => continuous_apply x) continuous_const) +#align is_closed_std_simplex isClosed_stdSimplex + +/-- `std_simplex ℝ ΞΉ` is compact. -/ +theorem isCompact_stdSimplex : IsCompact (stdSimplex ℝ ΞΉ) := + Metric.isCompact_iff_isClosed_bounded.2 ⟨isClosed_stdSimplex ΞΉ, bounded_stdSimplex ι⟩ +#align is_compact_std_simplex isCompact_stdSimplex + +end stdSimplex + +/-! ### Topological vector space -/ + + +section TopologicalSpace + +variable [LinearOrderedRing π•œ] [DenselyOrdered π•œ] [TopologicalSpace π•œ] [OrderTopology π•œ] + [AddCommGroup E] [TopologicalSpace E] [ContinuousAdd E] [Module π•œ E] [ContinuousSMul π•œ E] + {x y : E} + +theorem segment_subset_closure_openSegment : [x -[π•œ] y] βŠ† closure (openSegment π•œ x y) := by + rw [segment_eq_image, openSegment_eq_image, ← closure_Ioo (zero_ne_one' π•œ)] + exact image_closure_subset_closure_image (by continuity) +#align segment_subset_closure_open_segment segment_subset_closure_openSegment + +end TopologicalSpace + +section PseudoMetricSpace + +variable [LinearOrderedRing π•œ] [DenselyOrdered π•œ] [PseudoMetricSpace π•œ] [OrderTopology π•œ] + [ProperSpace π•œ] [CompactIccSpace π•œ] [AddCommGroup E] [TopologicalSpace E] [T2Space E] + [ContinuousAdd E] [Module π•œ E] [ContinuousSMul π•œ E] + +@[simp] +theorem closure_openSegment (x y : E) : closure (openSegment π•œ x y) = [x -[π•œ] y] := by + rw [segment_eq_image, openSegment_eq_image, ← closure_Ioo (zero_ne_one' π•œ)] + exact + (image_closure_of_isCompact (bounded_Ioo _ _).isCompact_closure <| + Continuous.continuousOn <| by continuity).symm +#align closure_open_segment closure_openSegment + +end PseudoMetricSpace + +section ContinuousConstSMul + +variable [LinearOrderedField π•œ] [AddCommGroup E] [Module π•œ E] [TopologicalSpace E] + [TopologicalAddGroup E] [ContinuousConstSMul π•œ E] + +/-- If `s` is a convex set, then `a β€’ interior s + b β€’ closure s βŠ† interior s` for all `0 < a`, +`0 ≀ b`, `a + b = 1`. See also `Convex.combo_interior_self_subset_interior` for a weaker version. -/ +theorem Convex.combo_interior_closure_subset_interior {s : Set E} (hs : Convex π•œ s) {a b : π•œ} + (ha : 0 < a) (hb : 0 ≀ b) (hab : a + b = 1) : a β€’ interior s + b β€’ closure s βŠ† interior s := + interior_smulβ‚€ ha.ne' s β–Έ + calc + interior (a β€’ s) + b β€’ closure s βŠ† interior (a β€’ s) + closure (b β€’ s) := + add_subset_add Subset.rfl (smul_closure_subset b s) + _ = interior (a β€’ s) + b β€’ s := by rw [isOpen_interior.add_closure (b β€’ s)] + _ βŠ† interior (a β€’ s + b β€’ s) := subset_interior_add_left + _ βŠ† interior s := interior_mono <| hs.set_combo_subset ha.le hb hab + +#align convex.combo_interior_closure_subset_interior Convex.combo_interior_closure_subset_interior + +/-- If `s` is a convex set, then `a β€’ interior s + b β€’ s βŠ† interior s` for all `0 < a`, `0 ≀ b`, +`a + b = 1`. See also `Convex.combo_interior_closure_subset_interior` for a stronger version. -/ +theorem Convex.combo_interior_self_subset_interior {s : Set E} (hs : Convex π•œ s) {a b : π•œ} + (ha : 0 < a) (hb : 0 ≀ b) (hab : a + b = 1) : a β€’ interior s + b β€’ s βŠ† interior s := + calc + a β€’ interior s + b β€’ s βŠ† a β€’ interior s + b β€’ closure s := + add_subset_add Subset.rfl <| image_subset _ subset_closure + _ βŠ† interior s := hs.combo_interior_closure_subset_interior ha hb hab + +#align convex.combo_interior_self_subset_interior Convex.combo_interior_self_subset_interior + +/-- If `s` is a convex set, then `a β€’ closure s + b β€’ interior s βŠ† interior s` for all `0 ≀ a`, +`0 < b`, `a + b = 1`. See also `Convex.combo_self_interior_subset_interior` for a weaker version. -/ +theorem Convex.combo_closure_interior_subset_interior {s : Set E} (hs : Convex π•œ s) {a b : π•œ} + (ha : 0 ≀ a) (hb : 0 < b) (hab : a + b = 1) : a β€’ closure s + b β€’ interior s βŠ† interior s := by + rw [add_comm] + exact hs.combo_interior_closure_subset_interior hb ha (add_comm a b β–Έ hab) +#align convex.combo_closure_interior_subset_interior Convex.combo_closure_interior_subset_interior + +/-- If `s` is a convex set, then `a β€’ s + b β€’ interior s βŠ† interior s` for all `0 ≀ a`, `0 < b`, +`a + b = 1`. See also `Convex.combo_closure_interior_subset_interior` for a stronger version. -/ +theorem Convex.combo_self_interior_subset_interior {s : Set E} (hs : Convex π•œ s) {a b : π•œ} + (ha : 0 ≀ a) (hb : 0 < b) (hab : a + b = 1) : a β€’ s + b β€’ interior s βŠ† interior s := by + rw [add_comm] + exact hs.combo_interior_self_subset_interior hb ha (add_comm a b β–Έ hab) +#align convex.combo_self_interior_subset_interior Convex.combo_self_interior_subset_interior + +theorem Convex.combo_interior_closure_mem_interior {s : Set E} (hs : Convex π•œ s) {x y : E} + (hx : x ∈ interior s) (hy : y ∈ closure s) {a b : π•œ} (ha : 0 < a) (hb : 0 ≀ b) + (hab : a + b = 1) : a β€’ x + b β€’ y ∈ interior s := + hs.combo_interior_closure_subset_interior ha hb hab <| + add_mem_add (smul_mem_smul_set hx) (smul_mem_smul_set hy) +#align convex.combo_interior_closure_mem_interior Convex.combo_interior_closure_mem_interior + +theorem Convex.combo_interior_self_mem_interior {s : Set E} (hs : Convex π•œ s) {x y : E} + (hx : x ∈ interior s) (hy : y ∈ s) {a b : π•œ} (ha : 0 < a) (hb : 0 ≀ b) (hab : a + b = 1) : + a β€’ x + b β€’ y ∈ interior s := + hs.combo_interior_closure_mem_interior hx (subset_closure hy) ha hb hab +#align convex.combo_interior_self_mem_interior Convex.combo_interior_self_mem_interior + +theorem Convex.combo_closure_interior_mem_interior {s : Set E} (hs : Convex π•œ s) {x y : E} + (hx : x ∈ closure s) (hy : y ∈ interior s) {a b : π•œ} (ha : 0 ≀ a) (hb : 0 < b) + (hab : a + b = 1) : a β€’ x + b β€’ y ∈ interior s := + hs.combo_closure_interior_subset_interior ha hb hab <| + add_mem_add (smul_mem_smul_set hx) (smul_mem_smul_set hy) +#align convex.combo_closure_interior_mem_interior Convex.combo_closure_interior_mem_interior + +theorem Convex.combo_self_interior_mem_interior {s : Set E} (hs : Convex π•œ s) {x y : E} (hx : x ∈ s) + (hy : y ∈ interior s) {a b : π•œ} (ha : 0 ≀ a) (hb : 0 < b) (hab : a + b = 1) : + a β€’ x + b β€’ y ∈ interior s := + hs.combo_closure_interior_mem_interior (subset_closure hx) hy ha hb hab +#align convex.combo_self_interior_mem_interior Convex.combo_self_interior_mem_interior + +theorem Convex.openSegment_interior_closure_subset_interior {s : Set E} (hs : Convex π•œ s) {x y : E} + (hx : x ∈ interior s) (hy : y ∈ closure s) : openSegment π•œ x y βŠ† interior s := by + rintro _ ⟨a, b, ha, hb, hab, rfl⟩ + exact hs.combo_interior_closure_mem_interior hx hy ha hb.le hab +#align convex.open_segment_interior_closure_subset_interior Convex.openSegment_interior_closure_subset_interior + +theorem Convex.openSegment_interior_self_subset_interior {s : Set E} (hs : Convex π•œ s) {x y : E} + (hx : x ∈ interior s) (hy : y ∈ s) : openSegment π•œ x y βŠ† interior s := + hs.openSegment_interior_closure_subset_interior hx (subset_closure hy) +#align convex.open_segment_interior_self_subset_interior Convex.openSegment_interior_self_subset_interior + +theorem Convex.openSegment_closure_interior_subset_interior {s : Set E} (hs : Convex π•œ s) {x y : E} + (hx : x ∈ closure s) (hy : y ∈ interior s) : openSegment π•œ x y βŠ† interior s := by + rintro _ ⟨a, b, ha, hb, hab, rfl⟩ + exact hs.combo_closure_interior_mem_interior hx hy ha.le hb hab +#align convex.open_segment_closure_interior_subset_interior Convex.openSegment_closure_interior_subset_interior + +theorem Convex.openSegment_self_interior_subset_interior {s : Set E} (hs : Convex π•œ s) {x y : E} + (hx : x ∈ s) (hy : y ∈ interior s) : openSegment π•œ x y βŠ† interior s := + hs.openSegment_closure_interior_subset_interior (subset_closure hx) hy +#align convex.open_segment_self_interior_subset_interior Convex.openSegment_self_interior_subset_interior + +/-- If `x ∈ closure s` and `y ∈ interior s`, then the segment `(x, y]` is included in `interior s`. +-/ +theorem Convex.add_smul_sub_mem_interior' {s : Set E} (hs : Convex π•œ s) {x y : E} + (hx : x ∈ closure s) (hy : y ∈ interior s) {t : π•œ} (ht : t ∈ Ioc (0 : π•œ) 1) : + x + t β€’ (y - x) ∈ interior s := by + simpa only [sub_smul, smul_sub, one_smul, add_sub, add_comm] using + hs.combo_interior_closure_mem_interior hy hx ht.1 (sub_nonneg.mpr ht.2) + (add_sub_cancel'_right _ _) +#align convex.add_smul_sub_mem_interior' Convex.add_smul_sub_mem_interior' + +/-- If `x ∈ s` and `y ∈ interior s`, then the segment `(x, y]` is included in `interior s`. -/ +theorem Convex.add_smul_sub_mem_interior {s : Set E} (hs : Convex π•œ s) {x y : E} (hx : x ∈ s) + (hy : y ∈ interior s) {t : π•œ} (ht : t ∈ Ioc (0 : π•œ) 1) : x + t β€’ (y - x) ∈ interior s := + hs.add_smul_sub_mem_interior' (subset_closure hx) hy ht +#align convex.add_smul_sub_mem_interior Convex.add_smul_sub_mem_interior + +/-- If `x ∈ closure s` and `x + y ∈ interior s`, then `x + t y ∈ interior s` for `t ∈ (0, 1]`. -/ +theorem Convex.add_smul_mem_interior' {s : Set E} (hs : Convex π•œ s) {x y : E} (hx : x ∈ closure s) + (hy : x + y ∈ interior s) {t : π•œ} (ht : t ∈ Ioc (0 : π•œ) 1) : x + t β€’ y ∈ interior s := by + simpa only [add_sub_cancel'] using hs.add_smul_sub_mem_interior' hx hy ht +#align convex.add_smul_mem_interior' Convex.add_smul_mem_interior' + +/-- If `x ∈ s` and `x + y ∈ interior s`, then `x + t y ∈ interior s` for `t ∈ (0, 1]`. -/ +theorem Convex.add_smul_mem_interior {s : Set E} (hs : Convex π•œ s) {x y : E} (hx : x ∈ s) + (hy : x + y ∈ interior s) {t : π•œ} (ht : t ∈ Ioc (0 : π•œ) 1) : x + t β€’ y ∈ interior s := + hs.add_smul_mem_interior' (subset_closure hx) hy ht +#align convex.add_smul_mem_interior Convex.add_smul_mem_interior + +/-- In a topological vector space, the interior of a convex set is convex. -/ +protected theorem Convex.interior {s : Set E} (hs : Convex π•œ s) : Convex π•œ (interior s) := + convex_iff_openSegment_subset.mpr fun _ hx _ hy => + hs.openSegment_closure_interior_subset_interior (interior_subset_closure hx) hy +#align convex.interior Convex.interior + +/-- In a topological vector space, the closure of a convex set is convex. -/ +protected theorem Convex.closure {s : Set E} (hs : Convex π•œ s) : Convex π•œ (closure s) := + fun x hx y hy a b ha hb hab => + let f : E β†’ E β†’ E := fun x' y' => a β€’ x' + b β€’ y' + have hf : Continuous (Function.uncurry f) := + (continuous_fst.const_smul _).add (continuous_snd.const_smul _) + show f x y ∈ closure s from map_mem_closureβ‚‚ hf hx hy fun _ hx' _ hy' => hs hx' hy' ha hb hab +#align convex.closure Convex.closure + +open AffineMap + +/-- A convex set `s` is strictly convex provided that for any two distinct points of +`s \ interior s`, the line passing through these points has nonempty intersection with +`interior s`. -/ +protected theorem Convex.strict_convex' {s : Set E} (hs : Convex π•œ s) + (h : (s \ interior s).Pairwise fun x y => βˆƒ c : π•œ, lineMap x y c ∈ interior s) : + StrictConvex π•œ s := by + refine' strictConvex_iff_openSegment_subset.2 _ + intro x hx y hy hne + by_cases hx' : x ∈ interior s; Β· exact hs.openSegment_interior_self_subset_interior hx' hy + by_cases hy' : y ∈ interior s; Β· exact hs.openSegment_self_interior_subset_interior hx hy' + rcases h ⟨hx, hx'⟩ ⟨hy, hy'⟩ hne with ⟨c, hc⟩ + refine' (openSegment_subset_union x y ⟨c, rfl⟩).trans (insert_subset.2 ⟨hc, union_subset _ _⟩) + exacts[hs.openSegment_self_interior_subset_interior hx hc, + hs.openSegment_interior_self_subset_interior hc hy] +#align convex.strict_convex' Convex.strict_convex' + +/-- A convex set `s` is strictly convex provided that for any two distinct points `x`, `y` of +`s \ interior s`, the segment with endpoints `x`, `y` has nonempty intersection with +`interior s`. -/ +protected theorem Convex.strictConvex {s : Set E} (hs : Convex π•œ s) + (h : (s \ interior s).Pairwise fun x y => ([x -[π•œ] y] \ frontier s).Nonempty) : + StrictConvex π•œ s := by + refine' hs.strict_convex' <| h.imp_on fun x hx y hy _ => _ + simp only [segment_eq_image_lineMap, ← self_diff_frontier] + rintro ⟨_, ⟨⟨c, hc, rfl⟩, hcs⟩⟩ + refine' ⟨c, hs.segment_subset hx.1 hy.1 _, hcs⟩ + exact (segment_eq_image_lineMap π•œ x y).symm β–Έ mem_image_of_mem _ hc +#align convex.strict_convex Convex.strictConvex + +end ContinuousConstSMul + +section ContinuousSMul + +variable [AddCommGroup E] [Module ℝ E] [TopologicalSpace E] [TopologicalAddGroup E] + [ContinuousSMul ℝ E] + +/-- Convex hull of a finite set is compact. -/ +theorem Set.Finite.isCompact_convexHull {s : Set E} (hs : s.Finite) : + IsCompact (convexHull ℝ s) := by + rw [hs.convexHull_eq_image] + apply (@isCompact_stdSimplex _ hs.fintype).image + haveI := hs.fintype + apply LinearMap.continuous_on_pi +#align set.finite.compact_convex_hull Set.Finite.isCompact_convexHull + +/-- Convex hull of a finite set is closed. -/ +theorem Set.Finite.isClosed_convexHull [T2Space E] {s : Set E} (hs : s.Finite) : + IsClosed (convexHull ℝ s) := + hs.isCompact_convexHull.isClosed +#align set.finite.is_closed_convex_hull Set.Finite.isClosed_convexHull + +open AffineMap + +/-- If we dilate the interior of a convex set about a point in its interior by a scale `t > 1`, +the result includes the closure of the original set. + +TODO Generalise this from convex sets to sets that are balanced / star-shaped about `x`. -/ +theorem Convex.closure_subset_image_homothety_interior_of_one_lt {s : Set E} (hs : Convex ℝ s) + {x : E} (hx : x ∈ interior s) (t : ℝ) (ht : 1 < t) : closure s βŠ† homothety x t '' interior s := + by + intro y hy + have hne : t β‰  0 := (one_pos.trans ht).ne' + refine' + ⟨homothety x t⁻¹ y, hs.openSegment_interior_closure_subset_interior hx hy _, + (AffineEquiv.homothetyUnitsMulHom x (Units.mk0 t hne)).apply_symm_apply y⟩ + rw [openSegment_eq_image_lineMap, ← inv_one, ← inv_Ioi (zero_lt_one' ℝ), ← image_inv, image_image, + homothety_eq_lineMap] + exact mem_image_of_mem _ ht +#align convex.closure_subset_image_homothety_interior_of_one_lt Convex.closure_subset_image_homothety_interior_of_one_lt + +/-- If we dilate a convex set about a point in its interior by a scale `t > 1`, the interior of +the result includes the closure of the original set. + +TODO Generalise this from convex sets to sets that are balanced / star-shaped about `x`. -/ +theorem Convex.closure_subset_interior_image_homothety_of_one_lt {s : Set E} (hs : Convex ℝ s) + {x : E} (hx : x ∈ interior s) (t : ℝ) (ht : 1 < t) : + closure s βŠ† interior (homothety x t '' s) := + (hs.closure_subset_image_homothety_interior_of_one_lt hx t ht).trans <| + (homothety_isOpenMap x t (one_pos.trans ht).ne').image_interior_subset _ +#align convex.closure_subset_interior_image_homothety_of_one_lt Convex.closure_subset_interior_image_homothety_of_one_lt + +/-- If we dilate a convex set about a point in its interior by a scale `t > 1`, the interior of +the result includes the closure of the original set. + +TODO Generalise this from convex sets to sets that are balanced / star-shaped about `x`. -/ +theorem Convex.subset_interior_image_homothety_of_one_lt {s : Set E} (hs : Convex ℝ s) {x : E} + (hx : x ∈ interior s) (t : ℝ) (ht : 1 < t) : s βŠ† interior (homothety x t '' s) := + subset_closure.trans <| hs.closure_subset_interior_image_homothety_of_one_lt hx t ht +#align convex.subset_interior_image_homothety_of_one_lt Convex.subset_interior_image_homothety_of_one_lt + +/-- A nonempty convex set is path connected. -/ +protected theorem Convex.isPathConnected {s : Set E} (hconv : Convex ℝ s) (hne : s.Nonempty) : + IsPathConnected s := by + refine' isPathConnected_iff.mpr ⟨hne, _⟩ + intro x x_in y y_in + have H := hconv.segment_subset x_in y_in + rw [segment_eq_image_lineMap] at H + exact + JoinedIn.ofLine AffineMap.lineMap_continuous.continuousOn (lineMap_apply_zero _ _) + (lineMap_apply_one _ _) H +#align convex.is_path_connected Convex.isPathConnected + +/-- A nonempty convex set is connected. -/ +protected theorem Convex.isConnected {s : Set E} (h : Convex ℝ s) (hne : s.Nonempty) : + IsConnected s := + (h.isPathConnected hne).isConnected +#align convex.is_connected Convex.isConnected + +/-- A convex set is preconnected. -/ +protected theorem Convex.isPreconnected {s : Set E} (h : Convex ℝ s) : IsPreconnected s := + s.eq_empty_or_nonempty.elim (fun h => h.symm β–Έ isPreconnected_empty) fun hne => + (h.isConnected hne).isPreconnected +#align convex.is_preconnected Convex.isPreconnected + +/-- Every topological vector space over ℝ is path connected. + +Not an instance, because it creates enormous TC subproblems (turn on `pp.all`). +-/ +protected theorem TopologicalAddGroup.pathConnectedSpace : PathConnectedSpace E := + pathConnectedSpace_iff_univ.mpr <| convex_univ.isPathConnected ⟨(0 : E), trivial⟩ +#align topological_add_group.path_connected TopologicalAddGroup.pathConnectedSpace + +end ContinuousSMul