@@ -796,13 +796,13 @@ begin
796
796
exact ⟨z i, hz i hit, Hi⟩
797
797
end
798
798
799
- lemma set.finite. convex_hull_eq { s : set E} (hs : finite s ) :
800
- convex_hull s = {x : E | ∃ (w : E → ℝ) (hw₀ : ∀ y ∈ s, 0 ≤ w y) (hw₁ : hs.to_finset .sum w = 1 ),
801
- hs.to_finset .center_mass w id = x} :=
799
+ lemma finset. convex_hull_eq ( s : finset E ) :
800
+ convex_hull ↑ s = {x : E | ∃ (w : E → ℝ) (hw₀ : ∀ y ∈ s, 0 ≤ w y) (hw₁ : s .sum w = 1 ),
801
+ s .center_mass w id = x} :=
802
802
begin
803
803
refine subset.antisymm (convex_hull_min _ _) _,
804
804
{ intros x hx,
805
- replace hx : x ∈ hs.to_finset, from finite.mem_to_finset. 2 hx,
805
+ rw [finset.mem_coe] at hx,
806
806
refine ⟨_, _, _, finset.center_mass_ite_eq _ _ _ hx⟩,
807
807
{ intros, split_ifs, exacts [zero_le_one, le_refl 0 ] },
808
808
{ rw [finset.sum_ite_eq, if_pos hx] } },
@@ -814,10 +814,16 @@ begin
814
814
apply_rules [add_nonneg, mul_nonneg, hwx₀, hwy₀], },
815
815
{ simp only [finset.sum_add_distrib, finset.mul_sum.symm, mul_one, *] } },
816
816
{ rintros _ ⟨w, hw₀, hw₁, rfl⟩,
817
- exact hs.to_finset. center_mass_mem_convex_hull (λ x hx, hw₀ _ $ finite.mem_to_finset. 1 hx)
818
- (hw₁.symm ▸ zero_lt_one) (λ x hx, finite.mem_to_finset. 1 hx) }
817
+ exact s. center_mass_mem_convex_hull (λ x hx, hw₀ _ hx)
818
+ (hw₁.symm ▸ zero_lt_one) (λ x hx, hx) }
819
819
end
820
820
821
+ lemma set.finite.convex_hull_eq {s : set E} (hs : finite s) :
822
+ convex_hull s = {x : E | ∃ (w : E → ℝ) (hw₀ : ∀ y ∈ s, 0 ≤ w y) (hw₁ : hs.to_finset.sum w = 1 ),
823
+ hs.to_finset.center_mass w id = x} :=
824
+ by simpa only [set.finite.coe_to_finset, set.finite.mem_to_finset, exists_prop]
825
+ using hs.to_finset.convex_hull_eq
826
+
821
827
lemma convex_hull_eq_union_convex_hull_finite_subsets (s : set E) :
822
828
convex_hull s = ⋃ (t : finset E) (w : ↑t ⊆ s), convex_hull ↑t :=
823
829
begin
0 commit comments