Documentation

Mathlib.Analysis.Convex.Topology

Topological properties of convex sets #

We prove the following facts:

Alias of the reverse direction of Real.convex_iff_isPreconnected.

Standard simplex #

Every vector in stdSimplex 𝕜 ι has max-norm at most 1.

stdSimplex ℝ ι is bounded.

theorem isClosed_stdSimplex (ι : Type u_1) [Fintype ι] :

stdSimplex ℝ ι is closed.

stdSimplex ℝ ι is compact.

Topological vector space #

theorem segment_subset_closure_openSegment {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedRing 𝕜] [DenselyOrdered 𝕜] [TopologicalSpace 𝕜] [OrderTopology 𝕜] [AddCommGroup E] [TopologicalSpace E] [ContinuousAdd E] [Module 𝕜 E] [ContinuousSMul 𝕜 E] {x : E} {y : E} :
segment 𝕜 x y closure (openSegment 𝕜 x y)
@[simp]
theorem closure_openSegment {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedRing 𝕜] [DenselyOrdered 𝕜] [PseudoMetricSpace 𝕜] [OrderTopology 𝕜] [ProperSpace 𝕜] [CompactIccSpace 𝕜] [AddCommGroup E] [TopologicalSpace E] [T2Space E] [ContinuousAdd E] [Module 𝕜 E] [ContinuousSMul 𝕜 E] (x : E) (y : E) :
closure (openSegment 𝕜 x y) = segment 𝕜 x y
theorem Convex.combo_interior_closure_subset_interior {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {a : 𝕜} {b : 𝕜} (ha : 0 < a) (hb : 0 b) (hab : a + b = 1) :

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_self_subset_interior {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {a : 𝕜} {b : 𝕜} (ha : 0 < a) (hb : 0 b) (hab : a + b = 1) :

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_closure_interior_subset_interior {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {a : 𝕜} {b : 𝕜} (ha : 0 a) (hb : 0 < b) (hab : a + b = 1) :

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_self_interior_subset_interior {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {a : 𝕜} {b : 𝕜} (ha : 0 a) (hb : 0 < b) (hab : a + b = 1) :

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_interior_closure_mem_interior {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {x : E} {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
theorem Convex.combo_interior_self_mem_interior {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {x : E} {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
theorem Convex.combo_closure_interior_mem_interior {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {x : E} {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
theorem Convex.combo_self_interior_mem_interior {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {x : E} {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
theorem Convex.openSegment_interior_closure_subset_interior {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {x : E} {y : E} (hx : x interior s) (hy : y closure s) :
theorem Convex.openSegment_interior_self_subset_interior {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {x : E} {y : E} (hx : x interior s) (hy : y s) :
theorem Convex.openSegment_closure_interior_subset_interior {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {x : E} {y : E} (hx : x closure s) (hy : y interior s) :
theorem Convex.openSegment_self_interior_subset_interior {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {x : E} {y : E} (hx : x s) (hy : y interior s) :
theorem Convex.add_smul_sub_mem_interior' {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {x : E} {y : E} (hx : x closure s) (hy : y interior s) {t : 𝕜} (ht : t Set.Ioc 0 1) :
x + t (y - x) interior s

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 {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {x : E} {y : E} (hx : x s) (hy : y interior s) {t : 𝕜} (ht : t Set.Ioc 0 1) :
x + t (y - x) interior s

If x ∈ s and y ∈ interior s, then the segment (x, y] is included in interior s.

theorem Convex.add_smul_mem_interior' {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {x : E} {y : E} (hx : x closure s) (hy : x + y interior s) {t : 𝕜} (ht : t Set.Ioc 0 1) :
x + t y interior s

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 {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) {x : E} {y : E} (hx : x s) (hy : x + y interior s) {t : 𝕜} (ht : t Set.Ioc 0 1) :
x + t y interior s

If x ∈ s and x + y ∈ interior s, then x + t y ∈ interior s for t ∈ (0, 1].

theorem Convex.interior {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) :
Convex 𝕜 (interior s)

In a topological vector space, the interior of a convex set is convex.

theorem Convex.closure {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) :
Convex 𝕜 (closure s)

In a topological vector space, the closure of a convex set is convex.

theorem Convex.strictConvex' {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) (h : Set.Pairwise (s \ interior s) fun x y => c, ↑(AffineMap.lineMap x y) c interior s) :

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.

theorem Convex.strictConvex {𝕜 : Type u_2} {E : Type u_3} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousConstSMul 𝕜 E] {s : Set E} (hs : Convex 𝕜 s) (h : Set.Pairwise (s \ interior s) fun x y => Set.Nonempty (segment 𝕜 x y \ frontier s)) :

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.

Convex hull of a finite set is compact.

Convex hull of a finite set is closed.

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.

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 {E : Type u_3} [AddCommGroup E] [Module E] [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousSMul E] {s : Set E} (hs : Convex s) {x : E} (hx : x interior s) (t : ) (ht : 1 < t) :

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 JoinedIn_of_segment_subset {E : Type u_4} [AddCommGroup E] [Module E] [TopologicalSpace E] [ContinuousAdd E] [ContinuousSMul E] {x : E} {y : E} {s : Set E} (h : segment x y s) :
JoinedIn s x y

A nonempty convex set is path connected.

A nonempty convex set is connected.

A convex set is preconnected.

Every topological vector space over ℝ is path connected.

Not an instance, because it creates enormous TC subproblems (turn on pp.all).