diff --git a/src/analysis/convex/basic.lean b/src/analysis/convex/basic.lean index 8d5852c54ad30..f9628f68a4489 100644 --- a/src/analysis/convex/basic.lean +++ b/src/analysis/convex/basic.lean @@ -796,13 +796,13 @@ begin exact ⟨z i, hz i hit, Hi⟩ end -lemma set.finite.convex_hull_eq {s : set E} (hs : finite s) : - convex_hull s = {x : E | ∃ (w : E → ℝ) (hw₀ : ∀ y ∈ s, 0 ≤ w y) (hw₁ : hs.to_finset.sum w = 1), - hs.to_finset.center_mass w id = x} := +lemma finset.convex_hull_eq (s : finset E) : + convex_hull ↑s = {x : E | ∃ (w : E → ℝ) (hw₀ : ∀ y ∈ s, 0 ≤ w y) (hw₁ : s.sum w = 1), + s.center_mass w id = x} := begin refine subset.antisymm (convex_hull_min _ _) _, { intros x hx, - replace hx : x ∈ hs.to_finset, from finite.mem_to_finset.2 hx, + rw [finset.mem_coe] at hx, refine ⟨_, _, _, finset.center_mass_ite_eq _ _ _ hx⟩, { intros, split_ifs, exacts [zero_le_one, le_refl 0] }, { rw [finset.sum_ite_eq, if_pos hx] } }, @@ -814,10 +814,16 @@ begin apply_rules [add_nonneg, mul_nonneg, hwx₀, hwy₀], }, { simp only [finset.sum_add_distrib, finset.mul_sum.symm, mul_one, *] } }, { rintros _ ⟨w, hw₀, hw₁, rfl⟩, - exact hs.to_finset.center_mass_mem_convex_hull (λ x hx, hw₀ _ $ finite.mem_to_finset.1 hx) - (hw₁.symm ▸ zero_lt_one) (λ x hx, finite.mem_to_finset.1 hx) } + exact s.center_mass_mem_convex_hull (λ x hx, hw₀ _ hx) + (hw₁.symm ▸ zero_lt_one) (λ x hx, hx) } end +lemma set.finite.convex_hull_eq {s : set E} (hs : finite s) : + convex_hull s = {x : E | ∃ (w : E → ℝ) (hw₀ : ∀ y ∈ s, 0 ≤ w y) (hw₁ : hs.to_finset.sum w = 1), + hs.to_finset.center_mass w id = x} := +by simpa only [set.finite.coe_to_finset, set.finite.mem_to_finset, exists_prop] + using hs.to_finset.convex_hull_eq + lemma convex_hull_eq_union_convex_hull_finite_subsets (s : set E) : convex_hull s = ⋃ (t : finset E) (w : ↑t ⊆ s), convex_hull ↑t := begin