Skip to content

Commit 6c3530b

Browse files
feat: port Analysis.Convex.Cone.Proper (#5607)
The remaining errors are about `coe` and `norm_cast`. I do not understand these well enough in Lean 4. Please feel free to push changes. Co-authored-by: Parcly Taxel <reddeloostw@gmail.com>
1 parent cef7c18 commit 6c3530b

File tree

2 files changed

+230
-0
lines changed

2 files changed

+230
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -579,6 +579,7 @@ import Mathlib.Analysis.Convex.Combination
579579
import Mathlib.Analysis.Convex.Complex
580580
import Mathlib.Analysis.Convex.Cone.Basic
581581
import Mathlib.Analysis.Convex.Cone.Dual
582+
import Mathlib.Analysis.Convex.Cone.Proper
582583
import Mathlib.Analysis.Convex.Contractible
583584
import Mathlib.Analysis.Convex.Exposed
584585
import Mathlib.Analysis.Convex.Extrema
Lines changed: 229 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,229 @@
1+
/-
2+
Copyright (c) 2022 Apurva Nakade All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Apurva Nakade
5+
6+
! This file was ported from Lean 3 source module analysis.convex.cone.proper
7+
! leanprover-community/mathlib commit 74f1d61944a5a793e8c939d47608178c0a0cb0c2
8+
! Please do not edit these lines, except to modify the commit id
9+
! if you have ported upstream changes.
10+
-/
11+
import Mathlib.Analysis.Convex.Cone.Dual
12+
13+
/-!
14+
# Proper cones
15+
16+
We define a proper cone as a nonempty, closed, convex cone. Proper cones are used in defining conic
17+
programs which generalize linear programs. A linear program is a conic program for the positive
18+
cone. We then prove Farkas' lemma for conic programs following the proof in the reference below.
19+
Farkas' lemma is equivalent to strong duality. So, once have the definitions of conic programs and
20+
linear programs, the results from this file can be used to prove duality theorems.
21+
22+
## TODO
23+
24+
The next steps are:
25+
- Prove the cone version of Farkas' lemma (2.3.4 in the reference).
26+
- Add comap, adjoint
27+
- Add convex_cone_class that extends set_like and replace the below instance
28+
- Define the positive cone as a proper cone.
29+
- Define primal and dual cone programs and prove weak duality.
30+
- Prove regular and strong duality for cone programs using Farkas' lemma (see reference).
31+
- Define linear programs and prove LP duality as a special case of cone duality.
32+
- Find a better reference (textbook instead of lecture notes).
33+
- Show submodules are (proper) cones.
34+
35+
## References
36+
37+
- [B. Gartner and J. Matousek, Cone Programming][gartnerMatousek]
38+
39+
-/
40+
41+
42+
namespace ConvexCone
43+
44+
variable {𝕜 : Type _} [OrderedSemiring 𝕜]
45+
46+
variable {E : Type _} [AddCommMonoid E] [TopologicalSpace E] [ContinuousAdd E] [SMul 𝕜 E]
47+
[ContinuousConstSMul 𝕜 E]
48+
49+
/-- The closure of a convex cone inside a topological space as a convex cone. This
50+
construction is mainly used for defining maps between proper cones. -/
51+
protected def closure (K : ConvexCone 𝕜 E) : ConvexCone 𝕜 E where
52+
carrier := closure ↑K
53+
smul_mem' c hc _ h₁ :=
54+
map_mem_closure (continuous_id'.const_smul c) h₁ fun _ h₂ => K.smul_mem hc h₂
55+
add_mem' _ h₁ _ h₂ := map_mem_closure₂ continuous_add h₁ h₂ K.add_mem
56+
#align convex_cone.closure ConvexCone.closure
57+
58+
@[simp, norm_cast]
59+
theorem coe_closure (K : ConvexCone 𝕜 E) : (K.closure : Set E) = closure K :=
60+
rfl
61+
#align convex_cone.coe_closure ConvexCone.coe_closure
62+
63+
@[simp]
64+
protected theorem mem_closure {K : ConvexCone 𝕜 E} {a : E} :
65+
a ∈ K.closure ↔ a ∈ closure (K : Set E) :=
66+
Iff.rfl
67+
#align convex_cone.mem_closure ConvexCone.mem_closure
68+
69+
@[simp]
70+
theorem closure_eq {K L : ConvexCone 𝕜 E} : K.closure = L ↔ closure (K : Set E) = L :=
71+
SetLike.ext'_iff
72+
#align convex_cone.closure_eq ConvexCone.closure_eq
73+
74+
end ConvexCone
75+
76+
/-- A proper cone is a convex cone `K` that is nonempty and closed. Proper cones have the nice
77+
property that the dual of the dual of a proper cone is itself. This makes them useful for defining
78+
cone programs and proving duality theorems. -/
79+
structure ProperCone (𝕜 : Type _) (E : Type _) [OrderedSemiring 𝕜] [AddCommMonoid E]
80+
[TopologicalSpace E] [SMul 𝕜 E] extends ConvexCone 𝕜 E where
81+
nonempty' : (carrier : Set E).Nonempty
82+
is_closed' : IsClosed (carrier : Set E)
83+
#align proper_cone ProperCone
84+
85+
namespace ProperCone
86+
87+
section SMul
88+
89+
variable {𝕜 : Type _} [OrderedSemiring 𝕜]
90+
91+
variable {E : Type _} [AddCommMonoid E] [TopologicalSpace E] [SMul 𝕜 E]
92+
93+
instance : Coe (ProperCone 𝕜 E) (ConvexCone 𝕜 E) :=
94+
fun K => K.1
95+
96+
-- Porting note: now a syntactic tautology
97+
-- @[simp]
98+
-- theorem toConvexCone_eq_coe (K : ProperCone 𝕜 E) : K.toConvexCone = K :=
99+
-- rfl
100+
-- #align proper_cone.to_convex_cone_eq_coe ProperCone.toConvexCone_eq_coe
101+
102+
theorem ext' : Function.Injective ((↑) : ProperCone 𝕜 E → ConvexCone 𝕜 E) := fun S T h => by
103+
cases S; cases T; congr
104+
#align proper_cone.ext' ProperCone.ext'
105+
106+
-- TODO: add `ConvexConeClass` that extends `SetLike` and replace the below instance
107+
instance : SetLike (ProperCone 𝕜 E) E where
108+
coe K := K.carrier
109+
coe_injective' _ _ h := ProperCone.ext' (SetLike.coe_injective h)
110+
111+
@[ext]
112+
theorem ext {S T : ProperCone 𝕜 E} (h : ∀ x, x ∈ S ↔ x ∈ T) : S = T :=
113+
SetLike.ext h
114+
#align proper_cone.ext ProperCone.ext
115+
116+
@[simp]
117+
theorem mem_coe {x : E} {K : ProperCone 𝕜 E} : x ∈ (K : ConvexCone 𝕜 E) ↔ x ∈ K :=
118+
Iff.rfl
119+
#align proper_cone.mem_coe ProperCone.mem_coe
120+
121+
protected theorem nonempty (K : ProperCone 𝕜 E) : (K : Set E).Nonempty :=
122+
K.nonempty'
123+
#align proper_cone.nonempty ProperCone.nonempty
124+
125+
protected theorem isClosed (K : ProperCone 𝕜 E) : IsClosed (K : Set E) :=
126+
K.is_closed'
127+
#align proper_cone.is_closed ProperCone.isClosed
128+
129+
end SMul
130+
131+
section Module
132+
133+
variable {𝕜 : Type _} [OrderedSemiring 𝕜]
134+
135+
variable {E : Type _} [AddCommMonoid E] [TopologicalSpace E] [T1Space E] [Module 𝕜 E]
136+
137+
instance : Zero (ProperCone 𝕜 E) :=
138+
⟨{ toConvexCone := 0
139+
nonempty' := ⟨0, rfl⟩
140+
is_closed' := isClosed_singleton }⟩
141+
142+
instance : Inhabited (ProperCone 𝕜 E) :=
143+
0
144+
145+
@[simp]
146+
theorem mem_zero (x : E) : x ∈ (0 : ProperCone 𝕜 E) ↔ x = 0 :=
147+
Iff.rfl
148+
#align proper_cone.mem_zero ProperCone.mem_zero
149+
150+
@[simp] -- Porting note: removed `norm_cast` (new-style structures)
151+
theorem coe_zero : ↑(0 : ProperCone 𝕜 E) = (0 : ConvexCone 𝕜 E) :=
152+
rfl
153+
#align proper_cone.coe_zero ProperCone.coe_zero
154+
155+
theorem pointed_zero : (0 : ProperCone 𝕜 E).Pointed := by simp [ConvexCone.pointed_zero]
156+
#align proper_cone.pointed_zero ProperCone.pointed_zero
157+
158+
end Module
159+
160+
section InnerProductSpace
161+
162+
variable {E : Type _} [NormedAddCommGroup E] [InnerProductSpace ℝ E]
163+
164+
variable {F : Type _} [NormedAddCommGroup F] [InnerProductSpace ℝ F]
165+
166+
protected theorem pointed (K : ProperCone ℝ E) : (K : ConvexCone ℝ E).Pointed :=
167+
(K : ConvexCone ℝ E).pointed_of_nonempty_of_isClosed K.nonempty' K.isClosed
168+
#align proper_cone.pointed ProperCone.pointed
169+
170+
/-- The closure of image of a proper cone under a continuous `ℝ`-linear map is a proper cone. We
171+
use continuous maps here so that the adjoint of f is also a map between proper cones. -/
172+
noncomputable def map (f : E →L[ℝ] F) (K : ProperCone ℝ E) : ProperCone ℝ F where
173+
toConvexCone := ConvexCone.closure (ConvexCone.map (f : E →ₗ[ℝ] F) ↑K)
174+
nonempty' :=
175+
0, subset_closure <| SetLike.mem_coe.2 <| ConvexCone.mem_map.20, K.pointed, map_zero _⟩⟩
176+
is_closed' := isClosed_closure
177+
#align proper_cone.map ProperCone.map
178+
179+
@[simp] -- Porting note: removed `norm_cast` (new-style structures)
180+
theorem coe_map (f : E →L[ℝ] F) (K : ProperCone ℝ E) :
181+
↑(K.map f) = (ConvexCone.map (f : E →ₗ[ℝ] F) ↑K).closure :=
182+
rfl
183+
#align proper_cone.coe_map ProperCone.coe_map
184+
185+
@[simp]
186+
theorem mem_map {f : E →L[ℝ] F} {K : ProperCone ℝ E} {y : F} :
187+
y ∈ K.map f ↔ y ∈ (ConvexCone.map (f : E →ₗ[ℝ] F) ↑K).closure :=
188+
Iff.rfl
189+
#align proper_cone.mem_map ProperCone.mem_map
190+
191+
@[simp]
192+
theorem map_id (K : ProperCone ℝ E) : K.map (ContinuousLinearMap.id ℝ E) = K :=
193+
ProperCone.ext' <| by simpa using IsClosed.closure_eq K.isClosed
194+
#align proper_cone.map_id ProperCone.map_id
195+
196+
/-- The inner dual cone of a proper cone is a proper cone. -/
197+
def dual (K : ProperCone ℝ E) : ProperCone ℝ E where
198+
toConvexCone := (K : Set E).innerDualCone
199+
nonempty' := ⟨0, pointed_innerDualCone _⟩
200+
is_closed' := isClosed_innerDualCone _
201+
#align proper_cone.dual ProperCone.dual
202+
203+
@[simp] -- Porting note: removed `norm_cast` (new-style structures)
204+
theorem coe_dual (K : ProperCone ℝ E) : ↑(dual K) = (K : Set E).innerDualCone :=
205+
rfl
206+
#align proper_cone.coe_dual ProperCone.coe_dual
207+
208+
@[simp]
209+
theorem mem_dual {K : ProperCone ℝ E} {y : E} : y ∈ dual K ↔ ∀ ⦃x⦄, x ∈ K → 0 ≤ ⟪x, y⟫_ℝ := by
210+
rw [← mem_coe, coe_dual, mem_innerDualCone _ _]; rfl
211+
#align proper_cone.mem_dual ProperCone.mem_dual
212+
213+
-- TODO: add comap, adjoint
214+
end InnerProductSpace
215+
216+
section CompleteSpace
217+
218+
variable {E : Type _} [NormedAddCommGroup E] [InnerProductSpace ℝ E] [CompleteSpace E]
219+
220+
/-- The dual of the dual of a proper cone is itself. -/
221+
@[simp]
222+
theorem dual_dual (K : ProperCone ℝ E) : K.dual.dual = K :=
223+
ProperCone.ext' <|
224+
(K : ConvexCone ℝ E).innerDualCone_of_innerDualCone_eq_self K.nonempty' K.isClosed
225+
#align proper_cone.dual_dual ProperCone.dual_dual
226+
227+
end CompleteSpace
228+
229+
end ProperCone

0 commit comments

Comments
 (0)