Skip to content

Commit 8a95474

Browse files
committed
chore: split Analysis.Convex.Cone.Basic (#8357)
Splits `Mathlib.Analysis.Convex.Cone.Basic`, to move Riesz extension and Hahn-Banach out of the basic file about definitions. Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
1 parent 7a22fac commit 8a95474

File tree

6 files changed

+208
-188
lines changed

6 files changed

+208
-188
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -659,6 +659,7 @@ import Mathlib.Analysis.Convex.Combination
659659
import Mathlib.Analysis.Convex.Complex
660660
import Mathlib.Analysis.Convex.Cone.Basic
661661
import Mathlib.Analysis.Convex.Cone.Dual
662+
import Mathlib.Analysis.Convex.Cone.Extension
662663
import Mathlib.Analysis.Convex.Cone.Pointed
663664
import Mathlib.Analysis.Convex.Cone.Proper
664665
import Mathlib.Analysis.Convex.Contractible

Mathlib/Analysis/Convex/Cone/Basic.lean

Lines changed: 6 additions & 184 deletions
Original file line numberDiff line numberDiff line change
@@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Yury Kudryashov, Frédéric Dupuis
55
-/
66
import Mathlib.Analysis.Convex.Hull
7-
import Mathlib.Data.Real.Archimedean
8-
import Mathlib.LinearAlgebra.LinearPMap
97

108
#align_import analysis.convex.cone.basic from "leanprover-community/mathlib"@"915591b2bb3ea303648db07284a161a7f2a9e3d4"
119

@@ -23,29 +21,11 @@ We define `Convex.toCone` to be the minimal cone that includes a given convex se
2321
2422
## Main statements
2523
26-
We prove two extension theorems:
27-
* `riesz_extension`:
28-
[M. Riesz extension theorem](https://en.wikipedia.org/wiki/M._Riesz_extension_theorem) says that
29-
if `s` is a convex cone in a real vector space `E`, `p` is a submodule of `E`
30-
such that `p + s = E`, and `f` is a linear function `p → ℝ` which is
31-
nonnegative on `p ∩ s`, then there exists a globally defined linear function
32-
`g : E → ℝ` that agrees with `f` on `p`, and is nonnegative on `s`.
33-
* `exists_extension_of_le_sublinear`:
34-
Hahn-Banach theorem: if `N : E → ℝ` is a sublinear map, `f` is a linear map
35-
defined on a subspace of `E`, and `f x ≤ N x` for all `x` in the domain of `f`,
36-
then `f` can be extended to the whole space to a linear map `g` such that `g x ≤ N x`
37-
for all `x`
38-
39-
In `Mathlib/Analysis/Convex/Cone/Dual.lean`, we prove the following theorems:
40-
* `ConvexCone.hyperplane_separation_of_nonempty_of_isClosed_of_nmem`:
41-
This variant of the
42-
[hyperplane separation theorem](https://en.wikipedia.org/wiki/Hyperplane_separation_theorem)
43-
states that given a nonempty, closed, convex cone `K` in a complete, real inner product space `H`
44-
and a point `b` disjoint from it, there is a vector `y` which separates `b` from `K` in the sense
45-
that for all points `x` in `K`, `0 ≤ ⟪x, y⟫_ℝ` and `⟪y, b⟫_ℝ < 0`. This is also a geometric
46-
interpretation of the
47-
[Farkas lemma](https://en.wikipedia.org/wiki/Farkas%27_lemma#Geometric_interpretation).
48-
* `ConvexCone.inner_dual_cone_of_inner_dual_cone_eq_self`:
24+
In `Mathlib/Analysis/Convex/Cone/Extension.lean` we prove
25+
the M. Riesz extension theorem and a form of the Hahn-Banach theorem.
26+
27+
In `Mathlib/Analysis/Convex/Cone/Dual.lean` we prove
28+
a variant of the hyperplane separation theorem.
4929
5030
## Implementation notes
5131
@@ -60,6 +40,7 @@ While `Convex 𝕜` is a predicate on sets, `ConvexCone 𝕜 E` is a bundled con
6040

6141

6242
assert_not_exists NormedSpace
43+
assert_not_exists Real
6344

6445
open Set LinearMap
6546

@@ -699,162 +680,3 @@ theorem convexHull_toCone_eq_sInf (s : Set E) :
699680
#align convex_hull_to_cone_eq_Inf convexHull_toCone_eq_sInf
700681

701682
end ConeFromConvex
702-
703-
/-!
704-
### M. Riesz extension theorem
705-
706-
Given a convex cone `s` in a vector space `E`, a submodule `p`, and a linear `f : p → ℝ`, assume
707-
that `f` is nonnegative on `p ∩ s` and `p + s = E`. Then there exists a globally defined linear
708-
function `g : E → ℝ` that agrees with `f` on `p`, and is nonnegative on `s`.
709-
710-
We prove this theorem using Zorn's lemma. `RieszExtension.step` is the main part of the proof.
711-
It says that if the domain `p` of `f` is not the whole space, then `f` can be extended to a larger
712-
subspace `p ⊔ span ℝ {y}` without breaking the non-negativity condition.
713-
714-
In `RieszExtension.exists_top` we use Zorn's lemma to prove that we can extend `f`
715-
to a linear map `g` on `⊤ : Submodule E`. Mathematically this is the same as a linear map on `E`
716-
but in Lean `⊤ : Submodule E` is isomorphic but is not equal to `E`. In `riesz_extension`
717-
we use this isomorphism to prove the theorem.
718-
-/
719-
720-
721-
variable [AddCommGroup E] [Module ℝ E]
722-
723-
namespace RieszExtension
724-
725-
open Submodule
726-
727-
variable (s : ConvexCone ℝ E) (f : E →ₗ.[ℝ] ℝ)
728-
729-
/-- Induction step in M. Riesz extension theorem. Given a convex cone `s` in a vector space `E`,
730-
a partially defined linear map `f : f.domain → ℝ`, assume that `f` is nonnegative on `f.domain ∩ p`
731-
and `p + s = E`. If `f` is not defined on the whole `E`, then we can extend it to a larger
732-
submodule without breaking the non-negativity condition. -/
733-
theorem step (nonneg : ∀ x : f.domain, (x : E) ∈ s → 0 ≤ f x)
734-
(dense : ∀ y, ∃ x : f.domain, (x : E) + y ∈ s) (hdom : f.domain ≠ ⊤) :
735-
∃ g, f < g ∧ ∀ x : g.domain, (x : E) ∈ s → 0 ≤ g x := by
736-
obtain ⟨y, -, hy⟩ : ∃ y ∈ ⊤, y ∉ f.domain :=
737-
@SetLike.exists_of_lt (Submodule ℝ E) _ _ _ _ (lt_top_iff_ne_top.2 hdom)
738-
obtain ⟨c, le_c, c_le⟩ :
739-
∃ c, (∀ x : f.domain, -(x : E) - y ∈ s → f x ≤ c) ∧
740-
∀ x : f.domain, (x : E) + y ∈ s → c ≤ f x := by
741-
set Sp := f '' { x : f.domain | (x : E) + y ∈ s }
742-
set Sn := f '' { x : f.domain | -(x : E) - y ∈ s }
743-
suffices (upperBounds Sn ∩ lowerBounds Sp).Nonempty by
744-
simpa only [Set.Nonempty, upperBounds, lowerBounds, ball_image_iff] using this
745-
refine' exists_between_of_forall_le (Nonempty.image f _) (Nonempty.image f (dense y)) _
746-
· rcases dense (-y) with ⟨x, hx⟩
747-
rw [← neg_neg x, AddSubgroupClass.coe_neg, ← sub_eq_add_neg] at hx
748-
exact ⟨_, hx⟩
749-
rintro a ⟨xn, hxn, rfl⟩ b ⟨xp, hxp, rfl⟩
750-
have := s.add_mem hxp hxn
751-
rw [add_assoc, add_sub_cancel'_right, ← sub_eq_add_neg, ← AddSubgroupClass.coe_sub] at this
752-
replace := nonneg _ this
753-
rwa [f.map_sub, sub_nonneg] at this
754-
-- Porting note: removed an unused `have`
755-
refine' ⟨f.supSpanSingleton y (-c) hy, _, _⟩
756-
· refine' lt_iff_le_not_le.2 ⟨f.left_le_sup _ _, fun H => _⟩
757-
replace H := LinearPMap.domain_mono.monotone H
758-
rw [LinearPMap.domain_supSpanSingleton, sup_le_iff, span_le, singleton_subset_iff] at H
759-
exact hy H.2
760-
· rintro ⟨z, hz⟩ hzs
761-
rcases mem_sup.1 hz with ⟨x, hx, y', hy', rfl⟩
762-
rcases mem_span_singleton.1 hy' with ⟨r, rfl⟩
763-
simp only [Subtype.coe_mk] at hzs
764-
erw [LinearPMap.supSpanSingleton_apply_mk _ _ _ _ _ hx, smul_neg, ← sub_eq_add_neg, sub_nonneg]
765-
rcases lt_trichotomy r 0 with (hr | hr | hr)
766-
· have : -(r⁻¹ • x) - y ∈ s := by
767-
rwa [← s.smul_mem_iff (neg_pos.2 hr), smul_sub, smul_neg, neg_smul, neg_neg, smul_smul,
768-
mul_inv_cancel hr.ne, one_smul, sub_eq_add_neg, neg_smul, neg_neg]
769-
-- Porting note: added type annotation and `by exact`
770-
replace : f (r⁻¹ • ⟨x, hx⟩) ≤ c := le_c (r⁻¹ • ⟨x, hx⟩) (by exact this)
771-
rwa [← mul_le_mul_left (neg_pos.2 hr), neg_mul, neg_mul, neg_le_neg_iff, f.map_smul,
772-
smul_eq_mul, ← mul_assoc, mul_inv_cancel hr.ne, one_mul] at this
773-
· subst r
774-
simp only [zero_smul, add_zero] at hzs ⊢
775-
apply nonneg
776-
exact hzs
777-
· have : r⁻¹ • x + y ∈ s := by
778-
rwa [← s.smul_mem_iff hr, smul_add, smul_smul, mul_inv_cancel hr.ne', one_smul]
779-
-- Porting note: added type annotation and `by exact`
780-
replace : c ≤ f (r⁻¹ • ⟨x, hx⟩) := c_le (r⁻¹ • ⟨x, hx⟩) (by exact this)
781-
rwa [← mul_le_mul_left hr, f.map_smul, smul_eq_mul, ← mul_assoc, mul_inv_cancel hr.ne',
782-
one_mul] at this
783-
#align riesz_extension.step RieszExtension.step
784-
785-
theorem exists_top (p : E →ₗ.[ℝ] ℝ) (hp_nonneg : ∀ x : p.domain, (x : E) ∈ s → 0 ≤ p x)
786-
(hp_dense : ∀ y, ∃ x : p.domain, (x : E) + y ∈ s) :
787-
∃ q ≥ p, q.domain = ⊤ ∧ ∀ x : q.domain, (x : E) ∈ s → 0 ≤ q x := by
788-
set S := { p : E →ₗ.[ℝ] ℝ | ∀ x : p.domain, (x : E) ∈ s → 0 ≤ p x }
789-
have hSc : ∀ c, c ⊆ S → IsChain (· ≤ ·) c → ∀ y ∈ c, ∃ ub ∈ S, ∀ z ∈ c, z ≤ ub
790-
· intro c hcs c_chain y hy
791-
clear hp_nonneg hp_dense p
792-
have cne : c.Nonempty := ⟨y, hy⟩
793-
have hcd : DirectedOn (· ≤ ·) c := c_chain.directedOn
794-
refine' ⟨LinearPMap.sSup c hcd, _, fun _ ↦ LinearPMap.le_sSup hcd⟩
795-
rintro ⟨x, hx⟩ hxs
796-
have hdir : DirectedOn (· ≤ ·) (LinearPMap.domain '' c) :=
797-
directedOn_image.2 (hcd.mono fun h ↦ LinearPMap.domain_mono.monotone h)
798-
rcases (mem_sSup_of_directed (cne.image _) hdir).1 hx with ⟨_, ⟨f, hfc, rfl⟩, hfx⟩
799-
have : f ≤ LinearPMap.sSup c hcd := LinearPMap.le_sSup _ hfc
800-
convert ← hcs hfc ⟨x, hfx⟩ hxs using 1
801-
exact this.2 rfl
802-
obtain ⟨q, hqs, hpq, hq⟩ := zorn_nonempty_partialOrder₀ S hSc p hp_nonneg
803-
· refine' ⟨q, hpq, _, hqs⟩
804-
contrapose! hq
805-
have hqd : ∀ y, ∃ x : q.domain, (x : E) + y ∈ s := fun y ↦
806-
let ⟨x, hx⟩ := hp_dense y
807-
⟨ofLe hpq.left x, hx⟩
808-
rcases step s q hqs hqd hq with ⟨r, hqr, hr⟩
809-
exact ⟨r, hr, hqr.le, hqr.ne'⟩
810-
#align riesz_extension.exists_top RieszExtension.exists_top
811-
812-
end RieszExtension
813-
814-
/-- M. **Riesz extension theorem**: given a convex cone `s` in a vector space `E`, a submodule `p`,
815-
and a linear `f : p → ℝ`, assume that `f` is nonnegative on `p ∩ s` and `p + s = E`. Then
816-
there exists a globally defined linear function `g : E → ℝ` that agrees with `f` on `p`,
817-
and is nonnegative on `s`. -/
818-
theorem riesz_extension (s : ConvexCone ℝ E) (f : E →ₗ.[ℝ] ℝ)
819-
(nonneg : ∀ x : f.domain, (x : E) ∈ s → 0 ≤ f x)
820-
(dense : ∀ y, ∃ x : f.domain, (x : E) + y ∈ s) :
821-
∃ g : E →ₗ[ℝ] ℝ, (∀ x : f.domain, g x = f x) ∧ ∀ x ∈ s, 0 ≤ g x := by
822-
rcases RieszExtension.exists_top s f nonneg dense
823-
with ⟨⟨g_dom, g⟩, ⟨-, hfg⟩, rfl : g_dom = ⊤, hgs⟩
824-
refine' ⟨g.comp (LinearMap.id.codRestrict ⊤ fun _ ↦ trivial), _, _⟩
825-
· exact fun x => (hfg rfl).symm
826-
· exact fun x hx => hgs ⟨x, _⟩ hx
827-
#align riesz_extension riesz_extension
828-
829-
/-- **Hahn-Banach theorem**: if `N : E → ℝ` is a sublinear map, `f` is a linear map
830-
defined on a subspace of `E`, and `f x ≤ N x` for all `x` in the domain of `f`,
831-
then `f` can be extended to the whole space to a linear map `g` such that `g x ≤ N x`
832-
for all `x`. -/
833-
theorem exists_extension_of_le_sublinear (f : E →ₗ.[ℝ] ℝ) (N : E → ℝ)
834-
(N_hom : ∀ c : ℝ, 0 < c → ∀ x, N (c • x) = c * N x) (N_add : ∀ x y, N (x + y) ≤ N x + N y)
835-
(hf : ∀ x : f.domain, f x ≤ N x) :
836-
∃ g : E →ₗ[ℝ] ℝ, (∀ x : f.domain, g x = f x) ∧ ∀ x, g x ≤ N x := by
837-
let s : ConvexCone ℝ (E × ℝ) :=
838-
{ carrier := { p : E × ℝ | N p.1 ≤ p.2 }
839-
smul_mem' := fun c hc p hp =>
840-
calc
841-
N (c • p.1) = c * N p.1 := N_hom c hc p.1
842-
_ ≤ c * p.2 := mul_le_mul_of_nonneg_left hp hc.le
843-
add_mem' := fun x hx y hy => (N_add _ _).trans (add_le_add hx hy) }
844-
set f' := (-f).coprod (LinearMap.id.toPMap ⊤)
845-
have hf'_nonneg : ∀ x : f'.domain, x.1 ∈ s → 0 ≤ f' x := fun x (hx : N x.1.1 ≤ x.1.2) ↦ by
846-
simpa using le_trans (hf ⟨x.1.1, x.2.1⟩) hx
847-
have hf'_dense : ∀ y : E × ℝ, ∃ x : f'.domain, ↑x + y ∈ s
848-
· rintro ⟨x, y⟩
849-
refine' ⟨⟨(0, N x - y), ⟨f.domain.zero_mem, trivial⟩⟩, _⟩
850-
simp only [ConvexCone.mem_mk, mem_setOf_eq, Prod.fst_add, Prod.snd_add, zero_add,
851-
sub_add_cancel, le_rfl]
852-
obtain ⟨g, g_eq, g_nonneg⟩ := riesz_extension s f' hf'_nonneg hf'_dense
853-
replace g_eq : ∀ (x : f.domain) (y : ℝ), g (x, y) = y - f x := fun x y ↦
854-
(g_eq ⟨(x, y), ⟨x.2, trivial⟩⟩).trans (sub_eq_neg_add _ _).symm
855-
refine ⟨-g.comp (inl ℝ E ℝ), fun x ↦ ?_, fun x ↦ ?_⟩
856-
· simp [g_eq x 0]
857-
· calc -g (x, 0) = g (0, N x) - g (x, N x) := by simp [← map_sub, ← map_neg]
858-
_ = N x - g (x, N x) := by simpa using g_eq 0 (N x)
859-
_ ≤ N x := by simpa using g_nonneg ⟨x, N x⟩ (le_refl (N x))
860-
#align exists_extension_of_le_sublinear exists_extension_of_le_sublinear

Mathlib/Analysis/Convex/Cone/Dual.lean

Lines changed: 8 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -19,10 +19,16 @@ all points `x` in a given set `0 ≤ ⟪ x, y ⟫`.
1919
We prove the following theorems:
2020
* `ConvexCone.innerDualCone_of_innerDualCone_eq_self`:
2121
The `innerDualCone` of the `innerDualCone` of a nonempty, closed, convex cone is itself.
22-
22+
* `ConvexCone.hyperplane_separation_of_nonempty_of_isClosed_of_nmem`:
23+
This variant of the
24+
[hyperplane separation theorem](https://en.wikipedia.org/wiki/Hyperplane_separation_theorem)
25+
states that given a nonempty, closed, convex cone `K` in a complete, real inner product space `H`
26+
and a point `b` disjoint from it, there is a vector `y` which separates `b` from `K` in the sense
27+
that for all points `x` in `K`, `0 ≤ ⟪x, y⟫_ℝ` and `⟪y, b⟫_ℝ < 0`. This is also a geometric
28+
interpretation of the
29+
[Farkas lemma](https://en.wikipedia.org/wiki/Farkas%27_lemma#Geometric_interpretation).
2330
-/
2431

25-
2632
open Set LinearMap
2733

2834
open Classical Pointwise

0 commit comments

Comments
 (0)