diff --git a/src/analysis/convex/caratheodory.lean b/src/analysis/convex/caratheodory.lean index 13d9cca6e1825..336dc99488afd 100644 --- a/src/analysis/convex/caratheodory.lean +++ b/src/analysis/convex/caratheodory.lean @@ -5,7 +5,7 @@ Authors: Johan Commelin, Scott Morrison -/ import analysis.convex.combination import linear_algebra.affine_space.independent -import data.real.basic +import tactic.field_simp /-! # CarathΓ©odory's convexity theorem @@ -21,7 +21,7 @@ contains `x`. Thus the difference from the case of affine span is that the affin 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`. +in `π•œα΅ˆ` is the union of the convex hulls of the `(d + 1)`-tuples in `s`. ## Main results @@ -36,20 +36,19 @@ convex hull, caratheodory -/ -universes u - open set finset open_locale big_operators -variables {E : Type u} [add_comm_group E] [module ℝ E] +universes u +variables {π•œ : Type*} {E : Type u} [linear_ordered_field π•œ] [add_comm_group E] [module π•œ E] namespace caratheodory /-- 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`. -/ lemma mem_convex_hull_erase [decidable_eq E] {t : finset E} - (h : Β¬ affine_independent ℝ (coe : t β†’ E)) {x : E} (m : x ∈ convex_hull ℝ (↑t : set E)) : - βˆƒ (y : (↑t : set E)), x ∈ convex_hull ℝ (↑(t.erase y) : set E) := + (h : Β¬ affine_independent π•œ (coe : t β†’ E)) {x : E} (m : x ∈ convex_hull π•œ (↑t : set E)) : + βˆƒ (y : (↑t : set E)), x ∈ convex_hull π•œ (↑(t.erase y) : set E) := begin simp only [finset.convex_hull_eq, mem_set_of_eq] at m ⊒, obtain ⟨f, fpos, fsum, rfl⟩ := m, @@ -63,7 +62,7 @@ begin exact ⟨x, mem_filter.mpr ⟨hx, hgx⟩⟩ }, have hg : 0 < g iβ‚€ := by { rw mem_filter at mem, exact mem.2 }, have hiβ‚€ : iβ‚€ ∈ t := filter_subset _ _ mem, - let k : E β†’ ℝ := Ξ» z, f z - (f iβ‚€ / g iβ‚€) * g z, + let k : E β†’ π•œ := Ξ» z, f z - (f iβ‚€ / g iβ‚€) * g z, have hk : k iβ‚€ = 0 := by field_simp [k, ne_of_gt hg], have ksum : βˆ‘ e in t.erase iβ‚€, k e = 1, { calc βˆ‘ e in t.erase iβ‚€, k e = βˆ‘ e in t, k e : @@ -89,36 +88,36 @@ begin sub_zero, center_mass, fsum, inv_one, one_smul, id.def] } end -variables {s : set E} {x : E} (hx : x ∈ convex_hull ℝ s) +variables {s : set E} {x : E} (hx : x ∈ convex_hull π•œ s) include hx /-- 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`. -/ noncomputable def min_card_finset_of_mem_convex_hull : finset E := -function.argmin_on finset.card nat.lt_wf { t | ↑t βŠ† s ∧ x ∈ convex_hull ℝ (t : set E) } +function.argmin_on finset.card nat.lt_wf { t | ↑t βŠ† s ∧ x ∈ convex_hull π•œ (t : set E) } (by simpa only [convex_hull_eq_union_convex_hull_finite_subsets s, exists_prop, mem_Union] using hx) lemma min_card_finset_of_mem_convex_hull_subseteq : ↑(min_card_finset_of_mem_convex_hull hx) βŠ† s := -(function.argmin_on_mem _ _ { t : finset E | ↑t βŠ† s ∧ x ∈ convex_hull ℝ (t : set E) } _).1 +(function.argmin_on_mem _ _ { t : finset E | ↑t βŠ† s ∧ x ∈ convex_hull π•œ (t : set E) } _).1 lemma mem_min_card_finset_of_mem_convex_hull : - x ∈ convex_hull ℝ (min_card_finset_of_mem_convex_hull hx : set E) := -(function.argmin_on_mem _ _ { t : finset E | ↑t βŠ† s ∧ x ∈ convex_hull ℝ (t : set E) } _).2 + x ∈ convex_hull π•œ (min_card_finset_of_mem_convex_hull hx : set E) := +(function.argmin_on_mem _ _ { t : finset E | ↑t βŠ† s ∧ x ∈ convex_hull π•œ (t : set E) } _).2 lemma min_card_finset_of_mem_convex_hull_nonempty : (min_card_finset_of_mem_convex_hull hx).nonempty := begin - rw [← finset.coe_nonempty, ← @convex_hull_nonempty_iff ℝ], + rw [← finset.coe_nonempty, ← @convex_hull_nonempty_iff π•œ], exact ⟨x, mem_min_card_finset_of_mem_convex_hull hx⟩, end lemma min_card_finset_of_mem_convex_hull_card_le_card - {t : finset E} (ht₁ : ↑t βŠ† s) (htβ‚‚ : x ∈ convex_hull ℝ (t : set E)) : + {t : finset E} (ht₁ : ↑t βŠ† s) (htβ‚‚ : x ∈ convex_hull π•œ (t : set E)) : (min_card_finset_of_mem_convex_hull hx).card ≀ t.card := function.argmin_on_le _ _ _ ⟨ht₁, htβ‚‚βŸ© lemma affine_independent_min_card_finset_of_mem_convex_hull : - affine_independent ℝ (coe : min_card_finset_of_mem_convex_hull hx β†’ E) := + affine_independent π•œ (coe : min_card_finset_of_mem_convex_hull hx β†’ E) := begin let k := (min_card_finset_of_mem_convex_hull hx).card - 1, have hk : (min_card_finset_of_mem_convex_hull hx).card = k + 1, @@ -142,8 +141,8 @@ variables {s : set E} /-- **CarathΓ©odory's convexity theorem** -/ lemma convex_hull_eq_union : - convex_hull ℝ s = - ⋃ (t : finset E) (hss : ↑t βŠ† s) (hai : affine_independent ℝ (coe : t β†’ E)), convex_hull ℝ ↑t := + convex_hull π•œ s = + ⋃ (t : finset E) (hss : ↑t βŠ† s) (hai : affine_independent π•œ (coe : t β†’ E)), convex_hull π•œ ↑t := begin apply set.subset.antisymm, { intros x hx, @@ -157,9 +156,9 @@ begin end /-- A more explicit version of `convex_hull_eq_union`. -/ -theorem eq_pos_convex_span_of_mem_convex_hull {x : E} (hx : x ∈ convex_hull ℝ s) : - βˆƒ (ΞΉ : Sort (u+1)) (_ : fintype ΞΉ), by exactI βˆƒ (z : ΞΉ β†’ E) (w : ΞΉ β†’ ℝ) - (hss : set.range z βŠ† s) (hai : affine_independent ℝ z) +theorem eq_pos_convex_span_of_mem_convex_hull {x : E} (hx : x ∈ convex_hull π•œ s) : + βˆƒ (ΞΉ : Sort (u+1)) (_ : fintype ΞΉ), by exactI βˆƒ (z : ΞΉ β†’ E) (w : ΞΉ β†’ π•œ) + (hss : set.range z βŠ† s) (hai : affine_independent π•œ z) (hw : βˆ€ i, 0 < w i), βˆ‘ i, w i = 1 ∧ βˆ‘ i, w i β€’ z i = x := begin rw convex_hull_eq_union at hx, diff --git a/src/analysis/convex/extreme.lean b/src/analysis/convex/extreme.lean index bf9a7e3a6ae94..5142870d5581f 100644 --- a/src/analysis/convex/extreme.lean +++ b/src/analysis/convex/extreme.lean @@ -8,7 +8,7 @@ import analysis.convex.hull /-! # Extreme sets -This file defines extreme sets and extreme points for sets in a real vector space. +This file defines extreme sets and extreme points for sets in a module. An extreme set of `A` is a subset of `A` that is as far as it can get in any outward direction: If point `x` is in it and point `y ∈ A`, then the line passing through `x` and `y` leaves `A` at `x`. diff --git a/src/analysis/convex/independent.lean b/src/analysis/convex/independent.lean index ae1a514473c7a..d335c94bddb32 100644 --- a/src/analysis/convex/independent.lean +++ b/src/analysis/convex/independent.lean @@ -42,10 +42,10 @@ independence, convex position open_locale affine big_operators classical open finset function -variables (π•œ : Type*) {E : Type*} +variables {π•œ E ΞΉ : Type*} section ordered_semiring -variables [ordered_semiring π•œ] [add_comm_group E] [module π•œ E] {ΞΉ : Type*} {s t : set E} +variables (π•œ) [ordered_semiring π•œ] [add_comm_group E] [module π•œ E] {s t : set E} /-- An indexed family is said to be convex independent if every point only belongs to convex hulls of sets containing it. -/ @@ -54,34 +54,6 @@ def convex_independent (p : ΞΉ β†’ E) : Prop := variables {π•œ} -/-- To check convex independence, one only has to check finsets thanks to CarathΓ©odory's theorem. -/ -lemma convex_independent_iff_finset [module ℝ E] {p : ΞΉ β†’ E} : - convex_independent ℝ p - ↔ βˆ€ (s : finset ΞΉ) (x : ΞΉ), p x ∈ convex_hull ℝ (s.image p : set E) β†’ x ∈ s := -begin - refine ⟨λ hc s x hx, hc s x _, Ξ» h s x hx, _⟩, - { rwa finset.coe_image at hx }, - have hp : injective p, - { rintro a b hab, - rw ←mem_singleton, - refine h {b} a _, - rw [hab, image_singleton, coe_singleton, convex_hull_singleton], - exact set.mem_singleton _ }, - let s' : finset ΞΉ := - (caratheodory.min_card_finset_of_mem_convex_hull hx).preimage p (hp.inj_on _), - suffices hs' : x ∈ s', - { rw ←hp.mem_set_image, - rw [finset.mem_preimage, ←mem_coe] at hs', - exact caratheodory.min_card_finset_of_mem_convex_hull_subseteq hx hs' }, - refine h s' x _, - have : s'.image p = caratheodory.min_card_finset_of_mem_convex_hull hx, - { rw [image_preimage, filter_true_of_mem], - exact Ξ» y hy, set.image_subset_range _ s - (caratheodory.min_card_finset_of_mem_convex_hull_subseteq hx $ mem_coe.2 hy) }, - rw this, - exact caratheodory.mem_min_card_finset_of_mem_convex_hull hx, -end - /-- A family with at most one point is convex independent. -/ lemma subsingleton.convex_independent [subsingleton ΞΉ] (p : ΞΉ β†’ E) : convex_independent π•œ p := @@ -193,11 +165,39 @@ end end ordered_semiring -/-! ### Extreme points -/ - section linear_ordered_field variables [linear_ordered_field π•œ] [add_comm_group E] [module π•œ E] {s : set E} +/-- To check convex independence, one only has to check finsets thanks to CarathΓ©odory's theorem. -/ +lemma convex_independent_iff_finset {p : ΞΉ β†’ E} : + convex_independent π•œ p + ↔ βˆ€ (s : finset ΞΉ) (x : ΞΉ), p x ∈ convex_hull π•œ (s.image p : set E) β†’ x ∈ s := +begin + refine ⟨λ hc s x hx, hc s x _, Ξ» h s x hx, _⟩, + { rwa finset.coe_image at hx }, + have hp : injective p, + { rintro a b hab, + rw ←mem_singleton, + refine h {b} a _, + rw [hab, image_singleton, coe_singleton, convex_hull_singleton], + exact set.mem_singleton _ }, + let s' : finset ΞΉ := + (caratheodory.min_card_finset_of_mem_convex_hull hx).preimage p (hp.inj_on _), + suffices hs' : x ∈ s', + { rw ←hp.mem_set_image, + rw [finset.mem_preimage, ←mem_coe] at hs', + exact caratheodory.min_card_finset_of_mem_convex_hull_subseteq hx hs' }, + refine h s' x _, + have : s'.image p = caratheodory.min_card_finset_of_mem_convex_hull hx, + { rw [image_preimage, filter_true_of_mem], + exact Ξ» y hy, set.image_subset_range _ s + (caratheodory.min_card_finset_of_mem_convex_hull_subseteq hx $ mem_coe.2 hy) }, + rw this, + exact caratheodory.mem_min_card_finset_of_mem_convex_hull hx, +end + +/-! ### Extreme points -/ + lemma convex.convex_independent_extreme_points (hs : convex π•œ s) : convex_independent π•œ (Ξ» p, p : s.extreme_points π•œ β†’ E) := convex_independent_set_iff_not_mem_convex_hull_diff.2 $ Ξ» x hx h,