author | wenzelm |
Mon, 02 Nov 2015 11:10:28 +0100 | |
changeset 61539 | a29295dac1ca |
parent 61486 | 3590367b0ce9 |
child 61540 | f92bf6674699 |
permissions | -rw-r--r-- |
31795
be3e1cc5005c
standard naming conventions for session and theories;
wenzelm
parents:
29252
diff
changeset
|
1 |
(* Title: HOL/Hahn_Banach/Hahn_Banach_Sup_Lemmas.thy |
7917 | 2 |
Author: Gertrud Bauer, TU Munich |
3 |
*) |
|
4 |
||
58889 | 5 |
section \<open>The supremum w.r.t.~the function order\<close> |
7917 | 6 |
|
31795
be3e1cc5005c
standard naming conventions for session and theories;
wenzelm
parents:
29252
diff
changeset
|
7 |
theory Hahn_Banach_Sup_Lemmas |
be3e1cc5005c
standard naming conventions for session and theories;
wenzelm
parents:
29252
diff
changeset
|
8 |
imports Function_Norm Zorn_Lemma |
27612 | 9 |
begin |
9256 | 10 |
|
58744 | 11 |
text \<open> |
10687 | 12 |
This section contains some lemmas that will be used in the proof of |
13 |
the Hahn-Banach Theorem. In this section the following context is |
|
61539 | 14 |
presumed. Let \<open>E\<close> be a real vector space with a seminorm |
15 |
\<open>p\<close> on \<open>E\<close>. \<open>F\<close> is a subspace of \<open>E\<close> and |
|
16 |
\<open>f\<close> a linear form on \<open>F\<close>. We consider a chain \<open>c\<close> |
|
17 |
of norm-preserving extensions of \<open>f\<close>, such that \<open>\<Union>c = |
|
18 |
graph H h\<close>. We will show some properties about the limit function |
|
19 |
\<open>h\<close>, i.e.\ the supremum of the chain \<open>c\<close>. |
|
7917 | 20 |
|
61486 | 21 |
\<^medskip> |
61539 | 22 |
Let \<open>c\<close> be a chain of norm-preserving extensions of |
23 |
the function \<open>f\<close> and let \<open>graph H h\<close> be the supremum |
|
24 |
of \<open>c\<close>. Every element in \<open>H\<close> is member of one of the |
|
10687 | 25 |
elements of the chain. |
58744 | 26 |
\<close> |
61486 | 27 |
|
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
45605
diff
changeset
|
28 |
lemmas [dest?] = chainsD |
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
45605
diff
changeset
|
29 |
lemmas chainsE2 [elim?] = chainsD2 [elim_format] |
7917 | 30 |
|
31 |
lemma some_H'h't: |
|
13515 | 32 |
assumes M: "M = norm_pres_extensions E p F f" |
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
45605
diff
changeset
|
33 |
and cM: "c \<in> chains M" |
13515 | 34 |
and u: "graph H h = \<Union>c" |
35 |
and x: "x \<in> H" |
|
36 |
shows "\<exists>H' h'. graph H' h' \<in> c |
|
37 |
\<and> (x, h x) \<in> graph H' h' |
|
38 |
\<and> linearform H' h' \<and> H' \<unlhd> E |
|
39 |
\<and> F \<unlhd> H' \<and> graph F f \<subseteq> graph H' h' |
|
40 |
\<and> (\<forall>x \<in> H'. h' x \<le> p x)" |
|
9261 | 41 |
proof - |
13515 | 42 |
from x have "(x, h x) \<in> graph H h" .. |
43 |
also from u have "\<dots> = \<Union>c" . |
|
44 |
finally obtain g where gc: "g \<in> c" and gh: "(x, h x) \<in> g" by blast |
|
7917 | 45 |
|
13515 | 46 |
from cM have "c \<subseteq> M" .. |
47 |
with gc have "g \<in> M" .. |
|
48 |
also from M have "\<dots> = norm_pres_extensions E p F f" . |
|
49 |
finally obtain H' and h' where g: "g = graph H' h'" |
|
50 |
and * : "linearform H' h'" "H' \<unlhd> E" "F \<unlhd> H'" |
|
51 |
"graph F f \<subseteq> graph H' h'" "\<forall>x \<in> H'. h' x \<le> p x" .. |
|
52 |
||
53 |
from gc and g have "graph H' h' \<in> c" by (simp only:) |
|
54 |
moreover from gh and g have "(x, h x) \<in> graph H' h'" by (simp only:) |
|
55 |
ultimately show ?thesis using * by blast |
|
9261 | 56 |
qed |
7917 | 57 |
|
58744 | 58 |
text \<open> |
61486 | 59 |
\<^medskip> |
61539 | 60 |
Let \<open>c\<close> be a chain of norm-preserving extensions of |
61 |
the function \<open>f\<close> and let \<open>graph H h\<close> be the supremum |
|
62 |
of \<open>c\<close>. Every element in the domain \<open>H\<close> of the supremum |
|
63 |
function is member of the domain \<open>H'\<close> of some function \<open>h'\<close>, such that \<open>h\<close> extends \<open>h'\<close>. |
|
58744 | 64 |
\<close> |
7917 | 65 |
|
10687 | 66 |
lemma some_H'h': |
13515 | 67 |
assumes M: "M = norm_pres_extensions E p F f" |
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
45605
diff
changeset
|
68 |
and cM: "c \<in> chains M" |
13515 | 69 |
and u: "graph H h = \<Union>c" |
70 |
and x: "x \<in> H" |
|
71 |
shows "\<exists>H' h'. x \<in> H' \<and> graph H' h' \<subseteq> graph H h |
|
72 |
\<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H' |
|
73 |
\<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)" |
|
9261 | 74 |
proof - |
13515 | 75 |
from M cM u x obtain H' h' where |
76 |
x_hx: "(x, h x) \<in> graph H' h'" |
|
77 |
and c: "graph H' h' \<in> c" |
|
78 |
and * : "linearform H' h'" "H' \<unlhd> E" "F \<unlhd> H'" |
|
10687 | 79 |
"graph F f \<subseteq> graph H' h'" "\<forall>x \<in> H'. h' x \<le> p x" |
13515 | 80 |
by (rule some_H'h't [elim_format]) blast |
81 |
from x_hx have "x \<in> H'" .. |
|
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
45605
diff
changeset
|
82 |
moreover from cM u c have "graph H' h' \<subseteq> graph H h" by blast |
13515 | 83 |
ultimately show ?thesis using * by blast |
9261 | 84 |
qed |
7917 | 85 |
|
58744 | 86 |
text \<open> |
61486 | 87 |
\<^medskip> |
61539 | 88 |
Any two elements \<open>x\<close> and \<open>y\<close> in the domain |
89 |
\<open>H\<close> of the supremum function \<open>h\<close> are both in the domain |
|
90 |
\<open>H'\<close> of some function \<open>h'\<close>, such that \<open>h\<close> extends |
|
91 |
\<open>h'\<close>. |
|
58744 | 92 |
\<close> |
7917 | 93 |
|
10687 | 94 |
lemma some_H'h'2: |
13515 | 95 |
assumes M: "M = norm_pres_extensions E p F f" |
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
45605
diff
changeset
|
96 |
and cM: "c \<in> chains M" |
13515 | 97 |
and u: "graph H h = \<Union>c" |
98 |
and x: "x \<in> H" |
|
99 |
and y: "y \<in> H" |
|
100 |
shows "\<exists>H' h'. x \<in> H' \<and> y \<in> H' |
|
101 |
\<and> graph H' h' \<subseteq> graph H h |
|
102 |
\<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H' |
|
103 |
\<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)" |
|
9261 | 104 |
proof - |
61539 | 105 |
txt \<open>\<open>y\<close> is in the domain \<open>H''\<close> of some function \<open>h''\<close>, |
106 |
such that \<open>h\<close> extends \<open>h''\<close>.\<close> |
|
7917 | 107 |
|
13515 | 108 |
from M cM u and y obtain H' h' where |
109 |
y_hy: "(y, h y) \<in> graph H' h'" |
|
110 |
and c': "graph H' h' \<in> c" |
|
111 |
and * : |
|
112 |
"linearform H' h'" "H' \<unlhd> E" "F \<unlhd> H'" |
|
113 |
"graph F f \<subseteq> graph H' h'" "\<forall>x \<in> H'. h' x \<le> p x" |
|
114 |
by (rule some_H'h't [elim_format]) blast |
|
115 |
||
61539 | 116 |
txt \<open>\<open>x\<close> is in the domain \<open>H'\<close> of some function \<open>h'\<close>, |
117 |
such that \<open>h\<close> extends \<open>h'\<close>.\<close> |
|
7917 | 118 |
|
13515 | 119 |
from M cM u and x obtain H'' h'' where |
120 |
x_hx: "(x, h x) \<in> graph H'' h''" |
|
121 |
and c'': "graph H'' h'' \<in> c" |
|
122 |
and ** : |
|
123 |
"linearform H'' h''" "H'' \<unlhd> E" "F \<unlhd> H''" |
|
124 |
"graph F f \<subseteq> graph H'' h''" "\<forall>x \<in> H''. h'' x \<le> p x" |
|
125 |
by (rule some_H'h't [elim_format]) blast |
|
7917 | 126 |
|
61539 | 127 |
txt \<open>Since both \<open>h'\<close> and \<open>h''\<close> are elements of the chain, |
128 |
\<open>h''\<close> is an extension of \<open>h'\<close> or vice versa. Thus both |
|
129 |
\<open>x\<close> and \<open>y\<close> are contained in the greater |
|
58744 | 130 |
one. \label{cases1}\<close> |
7917 | 131 |
|
60458 | 132 |
from cM c'' c' consider "graph H'' h'' \<subseteq> graph H' h'" | "graph H' h' \<subseteq> graph H'' h''" |
133 |
by (blast dest: chainsD) |
|
13515 | 134 |
then show ?thesis |
60458 | 135 |
proof cases |
136 |
case 1 |
|
23378 | 137 |
have "(x, h x) \<in> graph H'' h''" by fact |
27612 | 138 |
also have "\<dots> \<subseteq> graph H' h'" by fact |
13515 | 139 |
finally have xh:"(x, h x) \<in> graph H' h'" . |
140 |
then have "x \<in> H'" .. |
|
141 |
moreover from y_hy have "y \<in> H'" .. |
|
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
45605
diff
changeset
|
142 |
moreover from cM u and c' have "graph H' h' \<subseteq> graph H h" by blast |
13515 | 143 |
ultimately show ?thesis using * by blast |
144 |
next |
|
60458 | 145 |
case 2 |
13515 | 146 |
from x_hx have "x \<in> H''" .. |
60458 | 147 |
moreover have "y \<in> H''" |
148 |
proof - |
|
23378 | 149 |
have "(y, h y) \<in> graph H' h'" by (rule y_hy) |
150 |
also have "\<dots> \<subseteq> graph H'' h''" by fact |
|
13515 | 151 |
finally have "(y, h y) \<in> graph H'' h''" . |
60458 | 152 |
then show ?thesis .. |
153 |
qed |
|
154 |
moreover from u c'' have "graph H'' h'' \<subseteq> graph H h" by blast |
|
13515 | 155 |
ultimately show ?thesis using ** by blast |
9261 | 156 |
qed |
157 |
qed |
|
7917 | 158 |
|
58744 | 159 |
text \<open> |
61486 | 160 |
\<^medskip> |
161 |
The relation induced by the graph of the supremum of a |
|
61539 | 162 |
chain \<open>c\<close> is definite, i.~e.~t is the graph of a function. |
58744 | 163 |
\<close> |
7917 | 164 |
|
10687 | 165 |
lemma sup_definite: |
13515 | 166 |
assumes M_def: "M \<equiv> norm_pres_extensions E p F f" |
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
45605
diff
changeset
|
167 |
and cM: "c \<in> chains M" |
13515 | 168 |
and xy: "(x, y) \<in> \<Union>c" |
169 |
and xz: "(x, z) \<in> \<Union>c" |
|
170 |
shows "z = y" |
|
10687 | 171 |
proof - |
13515 | 172 |
from cM have c: "c \<subseteq> M" .. |
173 |
from xy obtain G1 where xy': "(x, y) \<in> G1" and G1: "G1 \<in> c" .. |
|
174 |
from xz obtain G2 where xz': "(x, z) \<in> G2" and G2: "G2 \<in> c" .. |
|
7917 | 175 |
|
13515 | 176 |
from G1 c have "G1 \<in> M" .. |
177 |
then obtain H1 h1 where G1_rep: "G1 = graph H1 h1" |
|
27612 | 178 |
unfolding M_def by blast |
7917 | 179 |
|
13515 | 180 |
from G2 c have "G2 \<in> M" .. |
181 |
then obtain H2 h2 where G2_rep: "G2 = graph H2 h2" |
|
27612 | 182 |
unfolding M_def by blast |
7917 | 183 |
|
61539 | 184 |
txt \<open>\<open>G\<^sub>1\<close> is contained in \<open>G\<^sub>2\<close> |
185 |
or vice versa, since both \<open>G\<^sub>1\<close> and \<open>G\<^sub>2\<close> are members of \<open>c\<close>. \label{cases2}\<close> |
|
7917 | 186 |
|
60458 | 187 |
from cM G1 G2 consider "G1 \<subseteq> G2" | "G2 \<subseteq> G1" |
188 |
by (blast dest: chainsD) |
|
13515 | 189 |
then show ?thesis |
60458 | 190 |
proof cases |
191 |
case 1 |
|
13515 | 192 |
with xy' G2_rep have "(x, y) \<in> graph H2 h2" by blast |
27612 | 193 |
then have "y = h2 x" .. |
13515 | 194 |
also |
195 |
from xz' G2_rep have "(x, z) \<in> graph H2 h2" by (simp only:) |
|
27612 | 196 |
then have "z = h2 x" .. |
13515 | 197 |
finally show ?thesis . |
198 |
next |
|
60458 | 199 |
case 2 |
13515 | 200 |
with xz' G1_rep have "(x, z) \<in> graph H1 h1" by blast |
27612 | 201 |
then have "z = h1 x" .. |
13515 | 202 |
also |
203 |
from xy' G1_rep have "(x, y) \<in> graph H1 h1" by (simp only:) |
|
27612 | 204 |
then have "y = h1 x" .. |
13515 | 205 |
finally show ?thesis .. |
9261 | 206 |
qed |
207 |
qed |
|
7917 | 208 |
|
58744 | 209 |
text \<open> |
61486 | 210 |
\<^medskip> |
61539 | 211 |
The limit function \<open>h\<close> is linear. Every element |
212 |
\<open>x\<close> in the domain of \<open>h\<close> is in the domain of a function |
|
213 |
\<open>h'\<close> in the chain of norm preserving extensions. Furthermore, |
|
214 |
\<open>h\<close> is an extension of \<open>h'\<close> so the function values of |
|
215 |
\<open>x\<close> are identical for \<open>h'\<close> and \<open>h\<close>. Finally, the |
|
216 |
function \<open>h'\<close> is linear by construction of \<open>M\<close>. |
|
58744 | 217 |
\<close> |
7917 | 218 |
|
10687 | 219 |
lemma sup_lf: |
13515 | 220 |
assumes M: "M = norm_pres_extensions E p F f" |
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
45605
diff
changeset
|
221 |
and cM: "c \<in> chains M" |
13515 | 222 |
and u: "graph H h = \<Union>c" |
223 |
shows "linearform H h" |
|
224 |
proof |
|
225 |
fix x y assume x: "x \<in> H" and y: "y \<in> H" |
|
226 |
with M cM u obtain H' h' where |
|
227 |
x': "x \<in> H'" and y': "y \<in> H'" |
|
228 |
and b: "graph H' h' \<subseteq> graph H h" |
|
229 |
and linearform: "linearform H' h'" |
|
230 |
and subspace: "H' \<unlhd> E" |
|
231 |
by (rule some_H'h'2 [elim_format]) blast |
|
7917 | 232 |
|
13515 | 233 |
show "h (x + y) = h x + h y" |
234 |
proof - |
|
235 |
from linearform x' y' have "h' (x + y) = h' x + h' y" |
|
236 |
by (rule linearform.add) |
|
237 |
also from b x' have "h' x = h x" .. |
|
238 |
also from b y' have "h' y = h y" .. |
|
239 |
also from subspace x' y' have "x + y \<in> H'" |
|
240 |
by (rule subspace.add_closed) |
|
241 |
with b have "h' (x + y) = h (x + y)" .. |
|
242 |
finally show ?thesis . |
|
243 |
qed |
|
244 |
next |
|
245 |
fix x a assume x: "x \<in> H" |
|
246 |
with M cM u obtain H' h' where |
|
247 |
x': "x \<in> H'" |
|
248 |
and b: "graph H' h' \<subseteq> graph H h" |
|
249 |
and linearform: "linearform H' h'" |
|
250 |
and subspace: "H' \<unlhd> E" |
|
251 |
by (rule some_H'h' [elim_format]) blast |
|
7917 | 252 |
|
13515 | 253 |
show "h (a \<cdot> x) = a * h x" |
254 |
proof - |
|
255 |
from linearform x' have "h' (a \<cdot> x) = a * h' x" |
|
256 |
by (rule linearform.mult) |
|
257 |
also from b x' have "h' x = h x" .. |
|
258 |
also from subspace x' have "a \<cdot> x \<in> H'" |
|
259 |
by (rule subspace.mult_closed) |
|
260 |
with b have "h' (a \<cdot> x) = h (a \<cdot> x)" .. |
|
261 |
finally show ?thesis . |
|
9261 | 262 |
qed |
263 |
qed |
|
7917 | 264 |
|
58744 | 265 |
text \<open> |
61486 | 266 |
\<^medskip> |
267 |
The limit of a non-empty chain of norm preserving |
|
61539 | 268 |
extensions of \<open>f\<close> is an extension of \<open>f\<close>, since every |
269 |
element of the chain is an extension of \<open>f\<close> and the supremum |
|
10687 | 270 |
is an extension for every element of the chain. |
58744 | 271 |
\<close> |
7917 | 272 |
|
273 |
lemma sup_ext: |
|
13515 | 274 |
assumes graph: "graph H h = \<Union>c" |
275 |
and M: "M = norm_pres_extensions E p F f" |
|
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
45605
diff
changeset
|
276 |
and cM: "c \<in> chains M" |
13515 | 277 |
and ex: "\<exists>x. x \<in> c" |
278 |
shows "graph F f \<subseteq> graph H h" |
|
9261 | 279 |
proof - |
13515 | 280 |
from ex obtain x where xc: "x \<in> c" .. |
281 |
from cM have "c \<subseteq> M" .. |
|
282 |
with xc have "x \<in> M" .. |
|
283 |
with M have "x \<in> norm_pres_extensions E p F f" |
|
284 |
by (simp only:) |
|
285 |
then obtain G g where "x = graph G g" and "graph F f \<subseteq> graph G g" .. |
|
286 |
then have "graph F f \<subseteq> x" by (simp only:) |
|
287 |
also from xc have "\<dots> \<subseteq> \<Union>c" by blast |
|
288 |
also from graph have "\<dots> = graph H h" .. |
|
289 |
finally show ?thesis . |
|
9261 | 290 |
qed |
7917 | 291 |
|
58744 | 292 |
text \<open> |
61486 | 293 |
\<^medskip> |
61539 | 294 |
The domain \<open>H\<close> of the limit function is a superspace |
295 |
of \<open>F\<close>, since \<open>F\<close> is a subset of \<open>H\<close>. The |
|
296 |
existence of the \<open>0\<close> element in \<open>F\<close> and the closure |
|
297 |
properties follow from the fact that \<open>F\<close> is a vector space. |
|
58744 | 298 |
\<close> |
7917 | 299 |
|
10687 | 300 |
lemma sup_supF: |
13515 | 301 |
assumes graph: "graph H h = \<Union>c" |
302 |
and M: "M = norm_pres_extensions E p F f" |
|
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
45605
diff
changeset
|
303 |
and cM: "c \<in> chains M" |
13515 | 304 |
and ex: "\<exists>x. x \<in> c" |
305 |
and FE: "F \<unlhd> E" |
|
306 |
shows "F \<unlhd> H" |
|
307 |
proof |
|
308 |
from FE show "F \<noteq> {}" by (rule subspace.non_empty) |
|
309 |
from graph M cM ex have "graph F f \<subseteq> graph H h" by (rule sup_ext) |
|
310 |
then show "F \<subseteq> H" .. |
|
311 |
fix x y assume "x \<in> F" and "y \<in> F" |
|
312 |
with FE show "x + y \<in> F" by (rule subspace.add_closed) |
|
313 |
next |
|
314 |
fix x a assume "x \<in> F" |
|
315 |
with FE show "a \<cdot> x \<in> F" by (rule subspace.mult_closed) |
|
9261 | 316 |
qed |
7917 | 317 |
|
58744 | 318 |
text \<open> |
61486 | 319 |
\<^medskip> |
61539 | 320 |
The domain \<open>H\<close> of the limit function is a subspace of |
321 |
\<open>E\<close>. |
|
58744 | 322 |
\<close> |
7917 | 323 |
|
10687 | 324 |
lemma sup_subE: |
13515 | 325 |
assumes graph: "graph H h = \<Union>c" |
326 |
and M: "M = norm_pres_extensions E p F f" |
|
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
45605
diff
changeset
|
327 |
and cM: "c \<in> chains M" |
13515 | 328 |
and ex: "\<exists>x. x \<in> c" |
329 |
and FE: "F \<unlhd> E" |
|
330 |
and E: "vectorspace E" |
|
331 |
shows "H \<unlhd> E" |
|
332 |
proof |
|
333 |
show "H \<noteq> {}" |
|
334 |
proof - |
|
13547 | 335 |
from FE E have "0 \<in> F" by (rule subspace.zero) |
13515 | 336 |
also from graph M cM ex FE have "F \<unlhd> H" by (rule sup_supF) |
337 |
then have "F \<subseteq> H" .. |
|
338 |
finally show ?thesis by blast |
|
339 |
qed |
|
340 |
show "H \<subseteq> E" |
|
9261 | 341 |
proof |
13515 | 342 |
fix x assume "x \<in> H" |
343 |
with M cM graph |
|
44887 | 344 |
obtain H' where x: "x \<in> H'" and H'E: "H' \<unlhd> E" |
13515 | 345 |
by (rule some_H'h' [elim_format]) blast |
346 |
from H'E have "H' \<subseteq> E" .. |
|
347 |
with x show "x \<in> E" .. |
|
348 |
qed |
|
349 |
fix x y assume x: "x \<in> H" and y: "y \<in> H" |
|
350 |
show "x + y \<in> H" |
|
351 |
proof - |
|
352 |
from M cM graph x y obtain H' h' where |
|
353 |
x': "x \<in> H'" and y': "y \<in> H'" and H'E: "H' \<unlhd> E" |
|
354 |
and graphs: "graph H' h' \<subseteq> graph H h" |
|
355 |
by (rule some_H'h'2 [elim_format]) blast |
|
356 |
from H'E x' y' have "x + y \<in> H'" |
|
357 |
by (rule subspace.add_closed) |
|
358 |
also from graphs have "H' \<subseteq> H" .. |
|
359 |
finally show ?thesis . |
|
360 |
qed |
|
361 |
next |
|
362 |
fix x a assume x: "x \<in> H" |
|
363 |
show "a \<cdot> x \<in> H" |
|
364 |
proof - |
|
365 |
from M cM graph x |
|
366 |
obtain H' h' where x': "x \<in> H'" and H'E: "H' \<unlhd> E" |
|
367 |
and graphs: "graph H' h' \<subseteq> graph H h" |
|
368 |
by (rule some_H'h' [elim_format]) blast |
|
369 |
from H'E x' have "a \<cdot> x \<in> H'" by (rule subspace.mult_closed) |
|
370 |
also from graphs have "H' \<subseteq> H" .. |
|
371 |
finally show ?thesis . |
|
9261 | 372 |
qed |
373 |
qed |
|
7917 | 374 |
|
58744 | 375 |
text \<open> |
61486 | 376 |
\<^medskip> |
61539 | 377 |
The limit function is bounded by the norm \<open>p\<close> as |
378 |
well, since all elements in the chain are bounded by \<open>p\<close>. |
|
58744 | 379 |
\<close> |
7917 | 380 |
|
9374 | 381 |
lemma sup_norm_pres: |
13515 | 382 |
assumes graph: "graph H h = \<Union>c" |
383 |
and M: "M = norm_pres_extensions E p F f" |
|
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
45605
diff
changeset
|
384 |
and cM: "c \<in> chains M" |
13515 | 385 |
shows "\<forall>x \<in> H. h x \<le> p x" |
9261 | 386 |
proof |
9503 | 387 |
fix x assume "x \<in> H" |
13515 | 388 |
with M cM graph obtain H' h' where x': "x \<in> H'" |
389 |
and graphs: "graph H' h' \<subseteq> graph H h" |
|
10687 | 390 |
and a: "\<forall>x \<in> H'. h' x \<le> p x" |
13515 | 391 |
by (rule some_H'h' [elim_format]) blast |
392 |
from graphs x' have [symmetric]: "h' x = h x" .. |
|
393 |
also from a x' have "h' x \<le> p x " .. |
|
394 |
finally show "h x \<le> p x" . |
|
9261 | 395 |
qed |
7917 | 396 |
|
58744 | 397 |
text \<open> |
61486 | 398 |
\<^medskip> |
399 |
The following lemma is a property of linear forms on real |
|
61539 | 400 |
vector spaces. It will be used for the lemma \<open>abs_Hahn_Banach\<close> |
31795
be3e1cc5005c
standard naming conventions for session and theories;
wenzelm
parents:
29252
diff
changeset
|
401 |
(see page \pageref{abs-Hahn_Banach}). \label{abs-ineq-iff} For real |
10687 | 402 |
vector spaces the following inequations are equivalent: |
403 |
\begin{center} |
|
404 |
\begin{tabular}{lll} |
|
61539 | 405 |
\<open>\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x\<close> & and & |
406 |
\<open>\<forall>x \<in> H. h x \<le> p x\<close> \\ |
|
10687 | 407 |
\end{tabular} |
408 |
\end{center} |
|
58744 | 409 |
\<close> |
7917 | 410 |
|
10687 | 411 |
lemma abs_ineq_iff: |
27611 | 412 |
assumes "subspace H E" and "vectorspace E" and "seminorm E p" |
413 |
and "linearform H h" |
|
13515 | 414 |
shows "(\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x) = (\<forall>x \<in> H. h x \<le> p x)" (is "?L = ?R") |
415 |
proof |
|
29234 | 416 |
interpret subspace H E by fact |
417 |
interpret vectorspace E by fact |
|
418 |
interpret seminorm E p by fact |
|
419 |
interpret linearform H h by fact |
|
58744 | 420 |
have H: "vectorspace H" using \<open>vectorspace E\<close> .. |
13515 | 421 |
{ |
9261 | 422 |
assume l: ?L |
423 |
show ?R |
|
424 |
proof |
|
9503 | 425 |
fix x assume x: "x \<in> H" |
13515 | 426 |
have "h x \<le> \<bar>h x\<bar>" by arith |
427 |
also from l x have "\<dots> \<le> p x" .. |
|
10687 | 428 |
finally show "h x \<le> p x" . |
9261 | 429 |
qed |
430 |
next |
|
431 |
assume r: ?R |
|
432 |
show ?L |
|
10687 | 433 |
proof |
13515 | 434 |
fix x assume x: "x \<in> H" |
60595 | 435 |
show "\<bar>b\<bar> \<le> a" when "- a \<le> b" "b \<le> a" for a b :: real |
436 |
using that by arith |
|
58744 | 437 |
from \<open>linearform H h\<close> and H x |
23378 | 438 |
have "- h x = h (- x)" by (rule linearform.neg [symmetric]) |
14710 | 439 |
also |
440 |
from H x have "- x \<in> H" by (rule vectorspace.neg_closed) |
|
441 |
with r have "h (- x) \<le> p (- x)" .. |
|
442 |
also have "\<dots> = p x" |
|
58744 | 443 |
using \<open>seminorm E p\<close> \<open>vectorspace E\<close> |
14710 | 444 |
proof (rule seminorm.minus) |
445 |
from x show "x \<in> E" .. |
|
9261 | 446 |
qed |
14710 | 447 |
finally have "- h x \<le> p x" . |
448 |
then show "- p x \<le> h x" by simp |
|
13515 | 449 |
from r x show "h x \<le> p x" .. |
9261 | 450 |
qed |
13515 | 451 |
} |
10687 | 452 |
qed |
7917 | 453 |
|
10687 | 454 |
end |