9374
|
1 |
(* Title: HOL/Real/HahnBanach/HahnBanach.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Gertrud Bauer, TU Munich
|
|
4 |
*)
|
|
5 |
|
|
6 |
header {* The Hahn-Banach Theorem *}
|
|
7 |
|
10687
|
8 |
theory HahnBanach = HahnBanachLemmas:
|
9374
|
9 |
|
|
10 |
text {*
|
10687
|
11 |
We present the proof of two different versions of the Hahn-Banach
|
9374
|
12 |
Theorem, closely following \cite[\S36]{Heuser:1986}.
|
|
13 |
*}
|
|
14 |
|
|
15 |
subsection {* The Hahn-Banach Theorem for vector spaces *}
|
|
16 |
|
|
17 |
text {*
|
10687
|
18 |
\textbf{Hahn-Banach Theorem.} Let @{text F} be a subspace of a real
|
|
19 |
vector space @{text E}, let @{text p} be a semi-norm on @{text E},
|
|
20 |
and @{text f} be a linear form defined on @{text F} such that @{text
|
|
21 |
f} is bounded by @{text p}, i.e. @{text "\<forall>x \<in> F. f x \<le> p x"}. Then
|
|
22 |
@{text f} can be extended to a linear form @{text h} on @{text E}
|
|
23 |
such that @{text h} is norm-preserving, i.e. @{text h} is also
|
|
24 |
bounded by @{text p}.
|
|
25 |
|
|
26 |
\bigskip
|
|
27 |
\textbf{Proof Sketch.}
|
|
28 |
\begin{enumerate}
|
|
29 |
|
|
30 |
\item Define @{text M} as the set of norm-preserving extensions of
|
|
31 |
@{text f} to subspaces of @{text E}. The linear forms in @{text M}
|
|
32 |
are ordered by domain extension.
|
9374
|
33 |
|
10687
|
34 |
\item We show that every non-empty chain in @{text M} has an upper
|
|
35 |
bound in @{text M}.
|
|
36 |
|
|
37 |
\item With Zorn's Lemma we conclude that there is a maximal function
|
|
38 |
@{text g} in @{text M}.
|
|
39 |
|
|
40 |
\item The domain @{text H} of @{text g} is the whole space @{text
|
|
41 |
E}, as shown by classical contradiction:
|
|
42 |
|
|
43 |
\begin{itemize}
|
|
44 |
|
|
45 |
\item Assuming @{text g} is not defined on whole @{text E}, it can
|
|
46 |
still be extended in a norm-preserving way to a super-space @{text
|
|
47 |
H'} of @{text H}.
|
|
48 |
|
|
49 |
\item Thus @{text g} can not be maximal. Contradiction!
|
|
50 |
|
|
51 |
\end{itemize}
|
|
52 |
|
|
53 |
\end{enumerate}
|
9374
|
54 |
*}
|
|
55 |
|
9256
|
56 |
theorem HahnBanach:
|
10687
|
57 |
"is_vectorspace E \<Longrightarrow> is_subspace F E \<Longrightarrow> is_seminorm E p
|
|
58 |
\<Longrightarrow> is_linearform F f \<Longrightarrow> \<forall>x \<in> F. f x \<le> p x
|
|
59 |
\<Longrightarrow> \<exists>h. is_linearform E h \<and> (\<forall>x \<in> F. h x = f x) \<and> (\<forall>x \<in> E. h x \<le> p x)"
|
|
60 |
-- {* Let @{text E} be a vector space, @{text F} a subspace of @{text E}, @{text p} a seminorm on @{text E}, *}
|
|
61 |
-- {* and @{text f} a linear form on @{text F} such that @{text f} is bounded by @{text p}, *}
|
|
62 |
-- {* then @{text f} can be extended to a linear form @{text h} on @{text E} in a norm-preserving way. \skp *}
|
9035
|
63 |
proof -
|
10687
|
64 |
assume "is_vectorspace E" "is_subspace F E" "is_seminorm E p"
|
|
65 |
and "is_linearform F f" "\<forall>x \<in> F. f x \<le> p x"
|
9256
|
66 |
-- {* Assume the context of the theorem. \skp *}
|
10687
|
67 |
def M \<equiv> "norm_pres_extensions E p F f"
|
|
68 |
-- {* Define @{text M} as the set of all norm-preserving extensions of @{text F}. \skp *}
|
9035
|
69 |
{
|
10687
|
70 |
fix c assume "c \<in> chain M" "\<exists>x. x \<in> c"
|
9503
|
71 |
have "\<Union>c \<in> M"
|
10687
|
72 |
-- {* Show that every non-empty chain @{text c} of @{text M} has an upper bound in @{text M}: *}
|
|
73 |
-- {* @{text "\<Union>c"} is greater than any element of the chain @{text c}, so it suffices to show @{text "\<Union>c \<in> M"}. *}
|
9035
|
74 |
proof (unfold M_def, rule norm_pres_extensionI)
|
9503
|
75 |
show "\<exists>H h. graph H h = \<Union>c
|
10687
|
76 |
\<and> is_linearform H h
|
|
77 |
\<and> is_subspace H E
|
|
78 |
\<and> is_subspace F H
|
9503
|
79 |
\<and> graph F f \<subseteq> graph H h
|
10687
|
80 |
\<and> (\<forall>x \<in> H. h x \<le> p x)"
|
9035
|
81 |
proof (intro exI conjI)
|
9503
|
82 |
let ?H = "domain (\<Union>c)"
|
|
83 |
let ?h = "funct (\<Union>c)"
|
8084
|
84 |
|
10687
|
85 |
show a: "graph ?H ?h = \<Union>c"
|
9035
|
86 |
proof (rule graph_domain_funct)
|
10687
|
87 |
fix x y z assume "(x, y) \<in> \<Union>c" "(x, z) \<in> \<Union>c"
|
9035
|
88 |
show "z = y" by (rule sup_definite)
|
|
89 |
qed
|
10687
|
90 |
show "is_linearform ?H ?h"
|
9035
|
91 |
by (simp! add: sup_lf a)
|
10687
|
92 |
show "is_subspace ?H E"
|
9374
|
93 |
by (rule sup_subE, rule a) (simp!)+
|
10687
|
94 |
show "is_subspace F ?H"
|
9374
|
95 |
by (rule sup_supF, rule a) (simp!)+
|
10687
|
96 |
show "graph F f \<subseteq> graph ?H ?h"
|
9374
|
97 |
by (rule sup_ext, rule a) (simp!)+
|
10687
|
98 |
show "\<forall>x \<in> ?H. ?h x \<le> p x"
|
9374
|
99 |
by (rule sup_norm_pres, rule a) (simp!)+
|
9035
|
100 |
qed
|
|
101 |
qed
|
9256
|
102 |
|
9035
|
103 |
}
|
10687
|
104 |
hence "\<exists>g \<in> M. \<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
|
|
105 |
-- {* With Zorn's Lemma we can conclude that there is a maximal element in @{text M}. \skp *}
|
9035
|
106 |
proof (rule Zorn's_Lemma)
|
10687
|
107 |
-- {* We show that @{text M} is non-empty: *}
|
9503
|
108 |
have "graph F f \<in> norm_pres_extensions E p F f"
|
9035
|
109 |
proof (rule norm_pres_extensionI2)
|
|
110 |
have "is_vectorspace F" ..
|
|
111 |
thus "is_subspace F F" ..
|
10687
|
112 |
qed (blast!)+
|
9503
|
113 |
thus "graph F f \<in> M" by (simp!)
|
9035
|
114 |
qed
|
|
115 |
thus ?thesis
|
|
116 |
proof
|
10687
|
117 |
fix g assume "g \<in> M" "\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
|
|
118 |
-- {* We consider such a maximal element @{text "g \<in> M"}. \skp *}
|
|
119 |
obtain H h where "graph H h = g" "is_linearform H h"
|
|
120 |
"is_subspace H E" "is_subspace F H" "graph F f \<subseteq> graph H h"
|
|
121 |
"\<forall>x \<in> H. h x \<le> p x"
|
|
122 |
-- {* @{text g} is a norm-preserving extension of @{text f}, in other words: *}
|
|
123 |
-- {* @{text g} is the graph of some linear form @{text h} defined on a subspace @{text H} of @{text E}, *}
|
|
124 |
-- {* and @{text h} is an extension of @{text f} that is again bounded by @{text p}. \skp *}
|
9475
|
125 |
proof -
|
10687
|
126 |
have "\<exists>H h. graph H h = g \<and> is_linearform H h
|
9503
|
127 |
\<and> is_subspace H E \<and> is_subspace F H
|
|
128 |
\<and> graph F f \<subseteq> graph H h
|
10687
|
129 |
\<and> (\<forall>x \<in> H. h x \<le> p x)"
|
9475
|
130 |
by (simp! add: norm_pres_extension_D)
|
10387
|
131 |
with that show ?thesis by blast
|
9475
|
132 |
qed
|
|
133 |
have h: "is_vectorspace H" ..
|
|
134 |
have "H = E"
|
10687
|
135 |
-- {* We show that @{text h} is defined on whole @{text E} by classical contradiction. \skp *}
|
9475
|
136 |
proof (rule classical)
|
9503
|
137 |
assume "H \<noteq> E"
|
10687
|
138 |
-- {* Assume @{text h} is not defined on whole @{text E}. Then show that @{text h} can be extended *}
|
|
139 |
-- {* in a norm-preserving way to a function @{text h'} with the graph @{text g'}. \skp *}
|
9503
|
140 |
have "\<exists>g' \<in> M. g \<subseteq> g' \<and> g \<noteq> g'"
|
9035
|
141 |
proof -
|
10687
|
142 |
obtain x' where "x' \<in> E" "x' \<notin> H"
|
|
143 |
-- {* Pick @{text "x' \<in> E - H"}. \skp *}
|
9475
|
144 |
proof -
|
9503
|
145 |
have "\<exists>x' \<in> E. x' \<notin> H"
|
9475
|
146 |
proof (rule set_less_imp_diff_not_empty)
|
9503
|
147 |
have "H \<subseteq> E" ..
|
|
148 |
thus "H \<subset> E" ..
|
9035
|
149 |
qed
|
10387
|
150 |
with that show ?thesis by blast
|
9475
|
151 |
qed
|
9503
|
152 |
have x': "x' \<noteq> 0"
|
9475
|
153 |
proof (rule classical)
|
|
154 |
presume "x' = 0"
|
9503
|
155 |
with h have "x' \<in> H" by simp
|
9475
|
156 |
thus ?thesis by contradiction
|
|
157 |
qed blast
|
10687
|
158 |
def H' \<equiv> "H + lin x'"
|
|
159 |
-- {* Define @{text H'} as the direct sum of @{text H} and the linear closure of @{text x'}. \skp *}
|
|
160 |
obtain xi where "\<forall>y \<in> H. - p (y + x') - h y \<le> xi
|
|
161 |
\<and> xi \<le> p (y + x') - h y"
|
|
162 |
-- {* Pick a real number @{text \<xi>} that fulfills certain inequations; this will *}
|
|
163 |
-- {* be used to establish that @{text h'} is a norm-preserving extension of @{text h}.
|
9475
|
164 |
\label{ex-xi-use}\skp *}
|
|
165 |
proof -
|
10687
|
166 |
from h have "\<exists>xi. \<forall>y \<in> H. - p (y + x') - h y \<le> xi
|
|
167 |
\<and> xi \<le> p (y + x') - h y"
|
9475
|
168 |
proof (rule ex_xi)
|
10687
|
169 |
fix u v assume "u \<in> H" "v \<in> H"
|
9475
|
170 |
from h have "h v - h u = h (v - u)"
|
|
171 |
by (simp! add: linearform_diff)
|
10687
|
172 |
also have "... \<le> p (v - u)"
|
9475
|
173 |
by (simp!)
|
|
174 |
also have "v - u = x' + - x' + v + - u"
|
|
175 |
by (simp! add: diff_eq1)
|
|
176 |
also have "... = v + x' + - (u + x')"
|
|
177 |
by (simp!)
|
|
178 |
also have "... = (v + x') - (u + x')"
|
|
179 |
by (simp! add: diff_eq1)
|
10687
|
180 |
also have "p ... \<le> p (v + x') + p (u + x')"
|
9475
|
181 |
by (rule seminorm_diff_subadditive) (simp_all!)
|
10687
|
182 |
finally have "h v - h u \<le> p (v + x') + p (u + x')" .
|
9261
|
183 |
|
10687
|
184 |
thus "- p (u + x') - h u \<le> p (v + x') - h v"
|
9475
|
185 |
by (rule real_diff_ineq_swap)
|
|
186 |
qed
|
10387
|
187 |
thus ?thesis ..
|
9475
|
188 |
qed
|
|
189 |
|
10687
|
190 |
def h' \<equiv> "\<lambda>x. let (y, a) = SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H
|
9560
|
191 |
in h y + a * xi"
|
10687
|
192 |
-- {* Define the extension @{text h'} of @{text h} to @{text H'} using @{text \<xi>}. \skp *}
|
9475
|
193 |
show ?thesis
|
|
194 |
proof
|
10687
|
195 |
show "g \<subseteq> graph H' h' \<and> g \<noteq> graph H' h'"
|
|
196 |
-- {* Show that @{text h'} is an extension of @{text h} \dots \skp *}
|
9475
|
197 |
proof
|
9503
|
198 |
show "g \<subseteq> graph H' h'"
|
9261
|
199 |
proof -
|
9503
|
200 |
have "graph H h \<subseteq> graph H' h'"
|
9475
|
201 |
proof (rule graph_extI)
|
10687
|
202 |
fix t assume "t \<in> H"
|
9503
|
203 |
have "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H)
|
9475
|
204 |
= (t, #0)"
|
|
205 |
by (rule decomp_H'_H) (assumption+, rule x')
|
|
206 |
thus "h t = h' t" by (simp! add: Let_def)
|
|
207 |
next
|
9503
|
208 |
show "H \<subseteq> H'"
|
9475
|
209 |
proof (rule subspace_subset)
|
|
210 |
show "is_subspace H H'"
|
|
211 |
proof (unfold H'_def, rule subspace_vs_sum1)
|
|
212 |
show "is_vectorspace H" ..
|
|
213 |
show "is_vectorspace (lin x')" ..
|
|
214 |
qed
|
|
215 |
qed
|
10687
|
216 |
qed
|
9475
|
217 |
thus ?thesis by (simp!)
|
9261
|
218 |
qed
|
9503
|
219 |
show "g \<noteq> graph H' h'"
|
9475
|
220 |
proof -
|
9503
|
221 |
have "graph H h \<noteq> graph H' h'"
|
9374
|
222 |
proof
|
9475
|
223 |
assume e: "graph H h = graph H' h'"
|
10687
|
224 |
have "x' \<in> H'"
|
9475
|
225 |
proof (unfold H'_def, rule vs_sumI)
|
|
226 |
show "x' = 0 + x'" by (simp!)
|
9503
|
227 |
from h show "0 \<in> H" ..
|
|
228 |
show "x' \<in> lin x'" by (rule x_lin_x)
|
9475
|
229 |
qed
|
9503
|
230 |
hence "(x', h' x') \<in> graph H' h'" ..
|
|
231 |
with e have "(x', h' x') \<in> graph H h" by simp
|
|
232 |
hence "x' \<in> H" ..
|
9475
|
233 |
thus False by contradiction
|
9035
|
234 |
qed
|
9475
|
235 |
thus ?thesis by (simp!)
|
9035
|
236 |
qed
|
|
237 |
qed
|
10687
|
238 |
show "graph H' h' \<in> M"
|
|
239 |
-- {* and @{text h'} is norm-preserving. \skp *}
|
9475
|
240 |
proof -
|
9503
|
241 |
have "graph H' h' \<in> norm_pres_extensions E p F f"
|
9475
|
242 |
proof (rule norm_pres_extensionI2)
|
|
243 |
show "is_linearform H' h'"
|
|
244 |
by (rule h'_lf) (simp! add: x')+
|
10687
|
245 |
show "is_subspace H' E"
|
|
246 |
by (unfold H'_def)
|
9475
|
247 |
(rule vs_sum_subspace [OF _ lin_subspace])
|
|
248 |
have "is_subspace F H" .
|
10687
|
249 |
also from h lin_vs
|
9897
|
250 |
have [folded H'_def]: "is_subspace H (H + lin x')" ..
|
10687
|
251 |
finally (subspace_trans [OF _ h])
|
9475
|
252 |
show f_h': "is_subspace F H'" .
|
10687
|
253 |
|
9503
|
254 |
show "graph F f \<subseteq> graph H' h'"
|
9475
|
255 |
proof (rule graph_extI)
|
9503
|
256 |
fix x assume "x \<in> F"
|
9475
|
257 |
have "f x = h x" ..
|
|
258 |
also have " ... = h x + #0 * xi" by simp
|
10687
|
259 |
also
|
|
260 |
have "... = (let (y, a) = (x, #0) in h y + a * xi)"
|
9475
|
261 |
by (simp add: Let_def)
|
10687
|
262 |
also have
|
9503
|
263 |
"(x, #0) = (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)"
|
9623
|
264 |
by (rule decomp_H'_H [symmetric]) (simp! add: x')+
|
10687
|
265 |
also have
|
|
266 |
"(let (y, a) = (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)
|
9475
|
267 |
in h y + a * xi) = h' x" by (simp!)
|
|
268 |
finally show "f x = h' x" .
|
|
269 |
next
|
9503
|
270 |
from f_h' show "F \<subseteq> H'" ..
|
9475
|
271 |
qed
|
10687
|
272 |
|
|
273 |
show "\<forall>x \<in> H'. h' x \<le> p x"
|
9475
|
274 |
by (rule h'_norm_pres) (assumption+, rule x')
|
|
275 |
qed
|
9503
|
276 |
thus "graph H' h' \<in> M" by (simp!)
|
9475
|
277 |
qed
|
9035
|
278 |
qed
|
9256
|
279 |
qed
|
10687
|
280 |
hence "\<not> (\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x)" by simp
|
|
281 |
-- {* So the graph @{text g} of @{text h} cannot be maximal. Contradiction! \skp *}
|
9475
|
282 |
thus "H = E" by contradiction
|
|
283 |
qed
|
10687
|
284 |
thus "\<exists>h. is_linearform E h \<and> (\<forall>x \<in> F. h x = f x)
|
|
285 |
\<and> (\<forall>x \<in> E. h x \<le> p x)"
|
9475
|
286 |
proof (intro exI conjI)
|
|
287 |
assume eq: "H = E"
|
|
288 |
from eq show "is_linearform E h" by (simp!)
|
9503
|
289 |
show "\<forall>x \<in> F. h x = f x"
|
9475
|
290 |
proof
|
10687
|
291 |
fix x assume "x \<in> F" have "f x = h x " ..
|
|
292 |
thus "h x = f x" ..
|
9035
|
293 |
qed
|
10687
|
294 |
from eq show "\<forall>x \<in> E. h x \<le> p x" by (blast!)
|
9035
|
295 |
qed
|
|
296 |
qed
|
9475
|
297 |
qed
|
9374
|
298 |
|
|
299 |
|
|
300 |
subsection {* Alternative formulation *}
|
|
301 |
|
10687
|
302 |
text {*
|
|
303 |
The following alternative formulation of the Hahn-Banach
|
|
304 |
Theorem\label{abs-HahnBanach} uses the fact that for a real linear
|
|
305 |
form @{text f} and a seminorm @{text p} the following inequations
|
|
306 |
are equivalent:\footnote{This was shown in lemma @{thm [source]
|
|
307 |
abs_ineq_iff} (see page \pageref{abs-ineq-iff}).}
|
|
308 |
\begin{center}
|
|
309 |
\begin{tabular}{lll}
|
|
310 |
@{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
|
|
311 |
@{text "\<forall>x \<in> H. h x \<le> p x"} \\
|
|
312 |
\end{tabular}
|
|
313 |
\end{center}
|
9374
|
314 |
*}
|
|
315 |
|
|
316 |
theorem abs_HahnBanach:
|
10687
|
317 |
"is_vectorspace E \<Longrightarrow> is_subspace F E \<Longrightarrow> is_linearform F f
|
|
318 |
\<Longrightarrow> is_seminorm E p \<Longrightarrow> \<forall>x \<in> F. \<bar>f x\<bar> \<le> p x
|
|
319 |
\<Longrightarrow> \<exists>g. is_linearform E g \<and> (\<forall>x \<in> F. g x = f x)
|
|
320 |
\<and> (\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x)"
|
9374
|
321 |
proof -
|
10687
|
322 |
assume "is_vectorspace E" "is_subspace F E" "is_seminorm E p"
|
|
323 |
"is_linearform F f" "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
|
|
324 |
have "\<forall>x \<in> F. f x \<le> p x" by (rule abs_ineq_iff [THEN iffD1])
|
|
325 |
hence "\<exists>g. is_linearform E g \<and> (\<forall>x \<in> F. g x = f x)
|
|
326 |
\<and> (\<forall>x \<in> E. g x \<le> p x)"
|
9475
|
327 |
by (simp! only: HahnBanach)
|
10687
|
328 |
thus ?thesis
|
9475
|
329 |
proof (elim exE conjE)
|
10687
|
330 |
fix g assume "is_linearform E g" "\<forall>x \<in> F. g x = f x"
|
|
331 |
"\<forall>x \<in> E. g x \<le> p x"
|
|
332 |
hence "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
|
9475
|
333 |
by (simp! add: abs_ineq_iff [OF subspace_refl])
|
|
334 |
thus ?thesis by (intro exI conjI)
|
|
335 |
qed
|
9374
|
336 |
qed
|
|
337 |
|
|
338 |
subsection {* The Hahn-Banach Theorem for normed spaces *}
|
|
339 |
|
10687
|
340 |
text {*
|
|
341 |
Every continuous linear form @{text f} on a subspace @{text F} of a
|
|
342 |
norm space @{text E}, can be extended to a continuous linear form
|
|
343 |
@{text g} on @{text E} such that @{text "\<parallel>f\<parallel> = \<parallel>g\<parallel>"}.
|
|
344 |
*}
|
9374
|
345 |
|
|
346 |
theorem norm_HahnBanach:
|
10687
|
347 |
"is_normed_vectorspace E norm \<Longrightarrow> is_subspace F E
|
|
348 |
\<Longrightarrow> is_linearform F f \<Longrightarrow> is_continuous F norm f
|
|
349 |
\<Longrightarrow> \<exists>g. is_linearform E g
|
|
350 |
\<and> is_continuous E norm g
|
|
351 |
\<and> (\<forall>x \<in> F. g x = f x)
|
9503
|
352 |
\<and> \<parallel>g\<parallel>E,norm = \<parallel>f\<parallel>F,norm"
|
9374
|
353 |
proof -
|
9475
|
354 |
assume e_norm: "is_normed_vectorspace E norm"
|
10687
|
355 |
assume f: "is_subspace F E" "is_linearform F f"
|
9475
|
356 |
assume f_cont: "is_continuous F norm f"
|
|
357 |
have e: "is_vectorspace E" ..
|
|
358 |
hence f_norm: "is_normed_vectorspace F norm" ..
|
|
359 |
|
10687
|
360 |
txt{*
|
|
361 |
We define a function @{text p} on @{text E} as follows:
|
|
362 |
@{text "p x = \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
|
9475
|
363 |
*}
|
|
364 |
|
10687
|
365 |
def p \<equiv> "\<lambda>x. \<parallel>f\<parallel>F,norm * norm x"
|
9475
|
366 |
|
10687
|
367 |
txt {* @{text p} is a seminorm on @{text E}: *}
|
9475
|
368 |
|
|
369 |
have q: "is_seminorm E p"
|
|
370 |
proof
|
10687
|
371 |
fix x y a assume "x \<in> E" "y \<in> E"
|
9475
|
372 |
|
10687
|
373 |
txt {* @{text p} is positive definite: *}
|
9374
|
374 |
|
10687
|
375 |
show "#0 \<le> p x"
|
9475
|
376 |
proof (unfold p_def, rule real_le_mult_order1a)
|
10687
|
377 |
from f_cont f_norm show "#0 \<le> \<parallel>f\<parallel>F,norm" ..
|
|
378 |
show "#0 \<le> norm x" ..
|
9475
|
379 |
qed
|
|
380 |
|
10687
|
381 |
txt {* @{text p} is absolutely homogenous: *}
|
9374
|
382 |
|
10687
|
383 |
show "p (a \<cdot> x) = \<bar>a\<bar> * p x"
|
|
384 |
proof -
|
9503
|
385 |
have "p (a \<cdot> x) = \<parallel>f\<parallel>F,norm * norm (a \<cdot> x)"
|
9475
|
386 |
by (simp!)
|
10687
|
387 |
also have "norm (a \<cdot> x) = \<bar>a\<bar> * norm x"
|
9475
|
388 |
by (rule normed_vs_norm_abs_homogenous)
|
10687
|
389 |
also have "\<parallel>f\<parallel>F,norm * (\<bar>a\<bar> * norm x )
|
|
390 |
= \<bar>a\<bar> * (\<parallel>f\<parallel>F,norm * norm x)"
|
9475
|
391 |
by (simp! only: real_mult_left_commute)
|
10687
|
392 |
also have "... = \<bar>a\<bar> * p x" by (simp!)
|
9475
|
393 |
finally show ?thesis .
|
|
394 |
qed
|
|
395 |
|
10687
|
396 |
txt {* Furthermore, @{text p} is subadditive: *}
|
9374
|
397 |
|
10687
|
398 |
show "p (x + y) \<le> p x + p y"
|
9475
|
399 |
proof -
|
9503
|
400 |
have "p (x + y) = \<parallel>f\<parallel>F,norm * norm (x + y)"
|
9475
|
401 |
by (simp!)
|
10687
|
402 |
also
|
|
403 |
have "... \<le> \<parallel>f\<parallel>F,norm * (norm x + norm y)"
|
9475
|
404 |
proof (rule real_mult_le_le_mono1a)
|
10687
|
405 |
from f_cont f_norm show "#0 \<le> \<parallel>f\<parallel>F,norm" ..
|
|
406 |
show "norm (x + y) \<le> norm x + norm y" ..
|
9475
|
407 |
qed
|
10687
|
408 |
also have "... = \<parallel>f\<parallel>F,norm * norm x
|
9503
|
409 |
+ \<parallel>f\<parallel>F,norm * norm y"
|
9475
|
410 |
by (simp! only: real_add_mult_distrib2)
|
|
411 |
finally show ?thesis by (simp!)
|
|
412 |
qed
|
|
413 |
qed
|
9374
|
414 |
|
10687
|
415 |
txt {* @{text f} is bounded by @{text p}. *}
|
9475
|
416 |
|
10687
|
417 |
have "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
|
9475
|
418 |
proof
|
9503
|
419 |
fix x assume "x \<in> F"
|
10687
|
420 |
from f_norm show "\<bar>f x\<bar> \<le> p x"
|
9475
|
421 |
by (simp! add: norm_fx_le_norm_f_norm_x)
|
|
422 |
qed
|
|
423 |
|
10687
|
424 |
txt {*
|
|
425 |
Using the fact that @{text p} is a seminorm and @{text f} is bounded
|
|
426 |
by @{text p} we can apply the Hahn-Banach Theorem for real vector
|
|
427 |
spaces. So @{text f} can be extended in a norm-preserving way to
|
|
428 |
some function @{text g} on the whole vector space @{text E}.
|
|
429 |
*}
|
9475
|
430 |
|
10687
|
431 |
with e f q
|
|
432 |
have "\<exists>g. is_linearform E g \<and> (\<forall>x \<in> F. g x = f x)
|
|
433 |
\<and> (\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x)"
|
9475
|
434 |
by (simp! add: abs_HahnBanach)
|
9374
|
435 |
|
9475
|
436 |
thus ?thesis
|
10687
|
437 |
proof (elim exE conjE)
|
9475
|
438 |
fix g
|
10687
|
439 |
assume "is_linearform E g" and a: "\<forall>x \<in> F. g x = f x"
|
|
440 |
and b: "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
|
9475
|
441 |
|
10687
|
442 |
show "\<exists>g. is_linearform E g
|
|
443 |
\<and> is_continuous E norm g
|
|
444 |
\<and> (\<forall>x \<in> F. g x = f x)
|
9503
|
445 |
\<and> \<parallel>g\<parallel>E,norm = \<parallel>f\<parallel>F,norm"
|
9475
|
446 |
proof (intro exI conjI)
|
|
447 |
|
10687
|
448 |
txt {*
|
|
449 |
We furthermore have to show that @{text g} is also continuous:
|
|
450 |
*}
|
9475
|
451 |
|
|
452 |
show g_cont: "is_continuous E norm g"
|
|
453 |
proof
|
9503
|
454 |
fix x assume "x \<in> E"
|
10687
|
455 |
with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>F,norm * norm x"
|
|
456 |
by (simp add: p_def)
|
|
457 |
qed
|
9374
|
458 |
|
10687
|
459 |
txt {*
|
|
460 |
To complete the proof, we show that
|
|
461 |
@{text "\<parallel>g\<parallel> = \<parallel>f\<parallel>"}. \label{order_antisym} *}
|
9475
|
462 |
|
9503
|
463 |
show "\<parallel>g\<parallel>E,norm = \<parallel>f\<parallel>F,norm"
|
9475
|
464 |
(is "?L = ?R")
|
|
465 |
proof (rule order_antisym)
|
|
466 |
|
10687
|
467 |
txt {*
|
|
468 |
First we show @{text "\<parallel>g\<parallel> \<le> \<parallel>f\<parallel>"}. The function norm @{text
|
|
469 |
"\<parallel>g\<parallel>"} is defined as the smallest @{text "c \<in> \<real>"} such that
|
|
470 |
\begin{center}
|
|
471 |
\begin{tabular}{l}
|
|
472 |
@{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
|
|
473 |
\end{tabular}
|
|
474 |
\end{center}
|
|
475 |
\noindent Furthermore holds
|
|
476 |
\begin{center}
|
|
477 |
\begin{tabular}{l}
|
|
478 |
@{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
|
|
479 |
\end{tabular}
|
|
480 |
\end{center}
|
9475
|
481 |
*}
|
10687
|
482 |
|
|
483 |
have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel>F,norm * norm x"
|
9475
|
484 |
proof
|
10687
|
485 |
fix x assume "x \<in> E"
|
|
486 |
show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>F,norm * norm x"
|
9374
|
487 |
by (simp!)
|
|
488 |
qed
|
|
489 |
|
10687
|
490 |
with g_cont e_norm show "?L \<le> ?R"
|
9475
|
491 |
proof (rule fnorm_le_ub)
|
10687
|
492 |
from f_cont f_norm show "#0 \<le> \<parallel>f\<parallel>F,norm" ..
|
9374
|
493 |
qed
|
|
494 |
|
10687
|
495 |
txt{* The other direction is achieved by a similar
|
9475
|
496 |
argument. *}
|
9374
|
497 |
|
10687
|
498 |
have "\<forall>x \<in> F. \<bar>f x\<bar> \<le> \<parallel>g\<parallel>E,norm * norm x"
|
9475
|
499 |
proof
|
10687
|
500 |
fix x assume "x \<in> F"
|
9475
|
501 |
from a have "g x = f x" ..
|
10687
|
502 |
hence "\<bar>f x\<bar> = \<bar>g x\<bar>" by simp
|
9475
|
503 |
also from g_cont
|
10687
|
504 |
have "... \<le> \<parallel>g\<parallel>E,norm * norm x"
|
9475
|
505 |
proof (rule norm_fx_le_norm_f_norm_x)
|
9503
|
506 |
show "x \<in> E" ..
|
9374
|
507 |
qed
|
10687
|
508 |
finally show "\<bar>f x\<bar> \<le> \<parallel>g\<parallel>E,norm * norm x" .
|
|
509 |
qed
|
|
510 |
thus "?R \<le> ?L"
|
9475
|
511 |
proof (rule fnorm_le_ub [OF f_cont f_norm])
|
10687
|
512 |
from g_cont show "#0 \<le> \<parallel>g\<parallel>E,norm" ..
|
9374
|
513 |
qed
|
|
514 |
qed
|
|
515 |
qed
|
9475
|
516 |
qed
|
|
517 |
qed
|
9374
|
518 |
|
9475
|
519 |
end
|