Documentation

Mathlib.Analysis.Convex.Caratheodory

Carathéodory's convexity theorem #

Convex hull can be regarded as a refinement of affine span. Both are closure operators but whereas convex hull takes values in the lattice of convex subsets, affine span takes values in the much coarser sublattice of affine subspaces.

The cost of this refinement is that one no longer has bases. However Carathéodory's convexity theorem offers some compensation. Given a set s together with a point x in its convex hull, Carathéodory says that one may find an affine-independent family of elements s whose convex hull contains x. Thus the difference from the case of affine span is that the affine-independent family depends on x.

In particular, in finite dimensions Carathéodory's theorem implies that the convex hull of a set s in 𝕜ᵈ is the union of the convex hulls of the (d + 1)-tuples in s.

Main results #

Implementation details #

This theorem was formalized as part of the Sphere Eversion project.

Tags #

convex hull, caratheodory

theorem Caratheodory.mem_convexHull_erase {𝕜 : Type u_1} {E : Type u} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [DecidableEq E] {t : Finset E} (h : ¬AffineIndependent 𝕜 Subtype.val) {x : E} (m : x (convexHull 𝕜) t) :
∃ (y : t), x (convexHull 𝕜) (t.erase y)

If x is in the convex hull of some finset t whose elements are not affine-independent, then it is in the convex hull of a strict subset of t.

noncomputable def Caratheodory.minCardFinsetOfMemConvexHull {𝕜 : Type u_1} {E : Type u} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} {x : E} (hx : x (convexHull 𝕜) s) :

Given a point x in the convex hull of a set s, this is a finite subset of s of minimum cardinality, whose convex hull contains x.

Equations
Instances For
    theorem Caratheodory.mem_minCardFinsetOfMemConvexHull {𝕜 : Type u_1} {E : Type u} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} {x : E} (hx : x (convexHull 𝕜) s) :
    theorem Caratheodory.minCardFinsetOfMemConvexHull_nonempty {𝕜 : Type u_1} {E : Type u} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} {x : E} (hx : x (convexHull 𝕜) s) :
    theorem Caratheodory.minCardFinsetOfMemConvexHull_card_le_card {𝕜 : Type u_1} {E : Type u} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} {x : E} (hx : x (convexHull 𝕜) s) {t : Finset E} (ht₁ : t s) (ht₂ : x (convexHull 𝕜) t) :
    theorem Caratheodory.affineIndependent_minCardFinsetOfMemConvexHull {𝕜 : Type u_1} {E : Type u} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} {x : E} (hx : x (convexHull 𝕜) s) :
    AffineIndependent 𝕜 Subtype.val
    theorem convexHull_eq_union {𝕜 : Type u_1} {E : Type u} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} :
    (convexHull 𝕜) s = ⋃ (t : Finset E), ⋃ (_ : t s), ⋃ (_ : AffineIndependent 𝕜 Subtype.val), (convexHull 𝕜) t

    Carathéodory's convexity theorem

    theorem eq_pos_convex_span_of_mem_convexHull {𝕜 : Type u_1} {E : Type u} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} {x : E} (hx : x (convexHull 𝕜) s) :
    ∃ (ι : Type u) (x_1 : Fintype ι) (z : ιE) (w : ι𝕜), Set.range z s AffineIndependent 𝕜 z (∀ (i : ι), 0 < w i) (Finset.univ.sum fun (i : ι) => w i) = 1 (Finset.univ.sum fun (i : ι) => w i z i) = x

    A more explicit version of convexHull_eq_union.