File tree Expand file tree Collapse file tree 1 file changed +23
-1
lines changed
Mathlib/Analysis/Convex/Cone Expand file tree Collapse file tree 1 file changed +23
-1
lines changed Original file line number Diff line number Diff line change @@ -21,7 +21,6 @@ linear programs, the results from this file can be used to prove duality theorem
21
21
22
22
The next steps are:
23
23
- Add convex_cone_class that extends set_like and replace the below instance
24
- - Define the positive cone as a proper cone.
25
24
- Define primal and dual cone programs and prove weak duality.
26
25
- Prove regular and strong duality for cone programs using Farkas' lemma (see reference).
27
26
- Define linear programs and prove LP duality as a special case of cone duality.
@@ -125,6 +124,29 @@ protected theorem isClosed (K : ProperCone 𝕜 E) : IsClosed (K : Set E) :=
125
124
126
125
end SMul
127
126
127
+ section PositiveCone
128
+
129
+ variable (𝕜 E)
130
+ variable [OrderedSemiring 𝕜] [OrderedAddCommGroup E] [Module 𝕜 E] [OrderedSMul 𝕜 E]
131
+ [TopologicalSpace E] [OrderClosedTopology E]
132
+
133
+ /-- The positive cone is the proper cone formed by the set of nonnegative elements in an ordered
134
+ module. -/
135
+ def positive : ProperCone 𝕜 E where
136
+ toConvexCone := ConvexCone.positive 𝕜 E
137
+ nonempty' := ⟨0 , ConvexCone.pointed_positive _ _⟩
138
+ is_closed' := isClosed_Ici
139
+
140
+ @[simp]
141
+ theorem mem_positive {x : E} : x ∈ positive 𝕜 E ↔ 0 ≤ x :=
142
+ Iff.rfl
143
+
144
+ @[simp]
145
+ theorem coe_positive : ↑(positive 𝕜 E) = ConvexCone.positive 𝕜 E :=
146
+ rfl
147
+
148
+ end PositiveCone
149
+
128
150
section Module
129
151
130
152
variable {𝕜 : Type _} [OrderedSemiring 𝕜]
You can’t perform that action at this time.
0 commit comments