Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 80a52e9

Browse files
committed
chore(analysis/convex/basic): add finset.convex_hull_eq (#2956)
1 parent 2ceb7f7 commit 80a52e9

File tree

1 file changed

+12
-6
lines changed

1 file changed

+12
-6
lines changed

src/analysis/convex/basic.lean

Lines changed: 12 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -796,13 +796,13 @@ begin
796796
exact ⟨z i, hz i hit, Hi⟩
797797
end
798798

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} :=
802802
begin
803803
refine subset.antisymm (convex_hull_min _ _) _,
804804
{ intros x hx,
805-
replace hx : x ∈ hs.to_finset, from finite.mem_to_finset.2 hx,
805+
rw [finset.mem_coe] at hx,
806806
refine ⟨_, _, _, finset.center_mass_ite_eq _ _ _ hx⟩,
807807
{ intros, split_ifs, exacts [zero_le_one, le_refl 0] },
808808
{ rw [finset.sum_ite_eq, if_pos hx] } },
@@ -814,10 +814,16 @@ begin
814814
apply_rules [add_nonneg, mul_nonneg, hwx₀, hwy₀], },
815815
{ simp only [finset.sum_add_distrib, finset.mul_sum.symm, mul_one, *] } },
816816
{ 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) }
819819
end
820820

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+
821827
lemma convex_hull_eq_union_convex_hull_finite_subsets (s : set E) :
822828
convex_hull s = ⋃ (t : finset E) (w : ↑t ⊆ s), convex_hull ↑t :=
823829
begin

0 commit comments

Comments
 (0)