Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat(analysis/convex/function): helper lemmas and general cleanup #9438

Closed
wants to merge 5 commits into from
Closed
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
14 changes: 14 additions & 0 deletions src/analysis/convex/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -518,6 +518,20 @@ begin
exact h hx hy ha hb hab
end

lemma convex_iff_pairwise_on_pos :
convex 𝕜 s ↔ s.pairwise_on (λ x y, ∀ ⦃a b : 𝕜⦄, 0 < a → 0 < b → a + b = 1 → a • x + b • y ∈ s) :=
begin
refine ⟨λ h x hx y hy _ a b ha hb hab, h hx hy ha.le hb.le hab, _⟩,
intros h x y hx hy a b ha hb hab,
obtain rfl | ha' := ha.eq_or_lt,
{ rw [zero_add] at hab, rwa [hab, zero_smul, one_smul, zero_add] },
obtain rfl | hb' := hb.eq_or_lt,
{ rw [add_zero] at hab, rwa [hab, zero_smul, one_smul, add_zero] },
obtain rfl | hxy := eq_or_ne x y,
{ rwa convex.combo_self hab },
exact h _ hx _ hy hxy ha' hb' hab,
end

lemma convex_iff_open_segment_subset :
convex 𝕜 s ↔ ∀ ⦃x y⦄, x ∈ s → y ∈ s → open_segment 𝕜 x y ⊆ s :=
begin
Expand Down
16 changes: 8 additions & 8 deletions src/analysis/convex/exposed.lean
Original file line number Diff line number Diff line change
Expand Up @@ -74,10 +74,10 @@ begin
exact hx.1,
end

@[refl] lemma refl (A : set E) : is_exposed 𝕜 A A :=
@[refl] protected lemma refl (A : set E) : is_exposed 𝕜 A A :=
λ ⟨w, hw⟩, ⟨0, subset.antisymm (λ x hx, ⟨hx, λ y hy, by exact le_refl 0⟩) (λ x hx, hx.1)⟩

lemma antisymm (hB : is_exposed 𝕜 A B) (hA : is_exposed 𝕜 B A) :
protected lemma antisymm (hB : is_exposed 𝕜 A B) (hA : is_exposed 𝕜 B A) :
A = B :=
hA.subset.antisymm hB.subset

Expand Down Expand Up @@ -113,7 +113,7 @@ begin
(λ x hx, ⟨hx.1, λ y hy, (hw.2 y hy).trans hx.2⟩)⟩,
end

lemma inter (hB : is_exposed 𝕜 A B) (hC : is_exposed 𝕜 A C) :
protected lemma inter (hB : is_exposed 𝕜 A B) (hC : is_exposed 𝕜 A C) :
is_exposed 𝕜 A (B ∩ C) :=
begin
rintro ⟨w, hwB, hwC⟩,
Expand Down Expand Up @@ -186,18 +186,18 @@ begin
{ exact convex_empty },
obtain ⟨l, rfl⟩ := hAB hB,
exact λ x₁ x₂ hx₁ hx₂ a b ha hb hab, ⟨hA hx₁.1 hx₂.1 ha hb hab, λ y hy,
((l.to_linear_map.concave_on convex_univ).concave_ge _
((l.to_linear_map.concave_on convex_univ).convex_ge _
⟨mem_univ _, hx₁.2 y hy⟩ ⟨mem_univ _, hx₂.2 y hy⟩ ha hb hab).2⟩,
end

lemma is_closed [normed_space ℝ E] (hAB : is_exposed A B) (hA : is_closed A) :
protected lemma is_closed [order_closed_topology 𝕜] (hAB : is_exposed 𝕜 A B) (hA : is_closed A) :
is_closed B :=
begin
obtain ⟨l, a, rfl⟩ := hAB.eq_inter_halfspace,
exact hA.is_closed_le continuous_on_const l.continuous.continuous_on,
end

lemma is_compact [normed_space ℝ E] (hAB : is_exposed A B) (hA : is_compact A) :
protected lemma is_compact [order_closed_topology 𝕜] (hAB : is_exposed 𝕜 A B) (hA : is_compact A) :
is_compact B :=
compact_of_is_closed_subset hA (hAB.is_closed hA.is_closed) hAB.subset

Expand Down Expand Up @@ -237,7 +237,7 @@ begin
exact ⟨hl.1.1, l, λ y hy, ⟨hl.1.2 y hy, λ hxy, hl.2 y ⟨hy, λ z hz, (hl.1.2 z hz).trans hxy⟩⟩⟩,
end

lemma exposed_points_subset_extreme_points [normed_space ℝ E] :
A.exposed_points ⊆ A.extreme_points :=
lemma exposed_points_subset_extreme_points :
A.exposed_points 𝕜 ⊆ A.extreme_points 𝕜 :=
λ x hx, mem_extreme_points_iff_extreme_singleton.2
(mem_exposed_points_iff_exposed_singleton.1 hx).is_extreme
Loading