|
7917
|
1 |
(* Title: HOL/Real/HahnBanach/HahnBanachExtLemmas.thy
|
|
|
2 |
ID: $Id$
|
|
|
3 |
Author: Gertrud Bauer, TU Munich
|
|
|
4 |
*)
|
|
|
5 |
|
|
10007
|
6 |
header {* Extending non-maximal functions *}
|
|
7917
|
7 |
|
|
10007
|
8 |
theory HahnBanachExtLemmas = FunctionNorm:
|
|
7917
|
9 |
|
|
10687
|
10 |
text {*
|
|
|
11 |
In this section the following context is presumed. Let @{text E} be
|
|
|
12 |
a real vector space with a seminorm @{text q} on @{text E}. @{text
|
|
|
13 |
F} is a subspace of @{text E} and @{text f} a linear function on
|
|
|
14 |
@{text F}. We consider a subspace @{text H} of @{text E} that is a
|
|
|
15 |
superspace of @{text F} and a linear form @{text h} on @{text
|
|
|
16 |
H}. @{text H} is a not equal to @{text E} and @{text "x\<^sub>0"} is
|
|
|
17 |
an element in @{text "E - H"}. @{text H} is extended to the direct
|
|
|
18 |
sum @{text "H' = H + lin x\<^sub>0"}, so for any @{text "x \<in> H'"}
|
|
|
19 |
the decomposition of @{text "x = y + a \<cdot> x"} with @{text "y \<in> H"} is
|
|
|
20 |
unique. @{text h'} is defined on @{text H'} by
|
|
|
21 |
@{text "h' x = h y + a \<cdot> \<xi>"} for a certain @{text \<xi>}.
|
|
7917
|
22 |
|
|
10687
|
23 |
Subsequently we show some properties of this extension @{text h'} of
|
|
|
24 |
@{text h}.
|
|
|
25 |
*}
|
|
7917
|
26 |
|
|
10687
|
27 |
text {*
|
|
|
28 |
This lemma will be used to show the existence of a linear extension
|
|
|
29 |
of @{text f} (see page \pageref{ex-xi-use}). It is a consequence of
|
|
|
30 |
the completeness of @{text \<real>}. To show
|
|
|
31 |
\begin{center}
|
|
|
32 |
\begin{tabular}{l}
|
|
|
33 |
@{text "\<exists>\<xi>. \<forall>y \<in> F. a y \<le> \<xi> \<and> \<xi> \<le> b y"}
|
|
|
34 |
\end{tabular}
|
|
|
35 |
\end{center}
|
|
|
36 |
\noindent it suffices to show that
|
|
|
37 |
\begin{center}
|
|
|
38 |
\begin{tabular}{l}
|
|
|
39 |
@{text "\<forall>u \<in> F. \<forall>v \<in> F. a u \<le> b v"}
|
|
|
40 |
\end{tabular}
|
|
|
41 |
\end{center}
|
|
|
42 |
*}
|
|
7917
|
43 |
|
|
10687
|
44 |
lemma ex_xi:
|
|
|
45 |
"is_vectorspace F \<Longrightarrow> (\<And>u v. u \<in> F \<Longrightarrow> v \<in> F \<Longrightarrow> a u \<le> b v)
|
|
|
46 |
\<Longrightarrow> \<exists>xi::real. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y"
|
|
10007
|
47 |
proof -
|
|
|
48 |
assume vs: "is_vectorspace F"
|
|
10687
|
49 |
assume r: "(\<And>u v. u \<in> F \<Longrightarrow> v \<in> F \<Longrightarrow> a u \<le> (b v::real))"
|
|
7917
|
50 |
|
|
|
51 |
txt {* From the completeness of the reals follows:
|
|
10687
|
52 |
The set @{text "S = {a u. u \<in> F}"} has a supremum, if
|
|
10007
|
53 |
it is non-empty and has an upper bound. *}
|
|
7917
|
54 |
|
|
10007
|
55 |
let ?S = "{a u :: real | u. u \<in> F}"
|
|
7917
|
56 |
|
|
10687
|
57 |
have "\<exists>xi. isLub UNIV ?S xi"
|
|
10007
|
58 |
proof (rule reals_complete)
|
|
7917
|
59 |
|
|
10687
|
60 |
txt {* The set @{text S} is non-empty, since @{text "a 0 \<in> S"}: *}
|
|
|
61 |
|
|
|
62 |
from vs have "a 0 \<in> ?S" by blast
|
|
10007
|
63 |
thus "\<exists>X. X \<in> ?S" ..
|
|
7917
|
64 |
|
|
10687
|
65 |
txt {* @{text "b 0"} is an upper bound of @{text S}: *}
|
|
7917
|
66 |
|
|
10687
|
67 |
show "\<exists>Y. isUb UNIV ?S Y"
|
|
|
68 |
proof
|
|
10007
|
69 |
show "isUb UNIV ?S (b 0)"
|
|
|
70 |
proof (intro isUbI setleI ballI)
|
|
|
71 |
show "b 0 \<in> UNIV" ..
|
|
|
72 |
next
|
|
7917
|
73 |
|
|
10687
|
74 |
txt {* Every element @{text "y \<in> S"} is less than @{text "b 0"}: *}
|
|
7917
|
75 |
|
|
10687
|
76 |
fix y assume y: "y \<in> ?S"
|
|
10007
|
77 |
from y have "\<exists>u \<in> F. y = a u" by fast
|
|
10687
|
78 |
thus "y \<le> b 0"
|
|
10007
|
79 |
proof
|
|
10687
|
80 |
fix u assume "u \<in> F"
|
|
10007
|
81 |
assume "y = a u"
|
|
10687
|
82 |
also have "a u \<le> b 0" by (rule r) (simp!)+
|
|
10007
|
83 |
finally show ?thesis .
|
|
|
84 |
qed
|
|
|
85 |
qed
|
|
|
86 |
qed
|
|
|
87 |
qed
|
|
7917
|
88 |
|
|
10687
|
89 |
thus "\<exists>xi. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y"
|
|
10007
|
90 |
proof (elim exE)
|
|
10687
|
91 |
fix xi assume "isLub UNIV ?S xi"
|
|
10007
|
92 |
show ?thesis
|
|
10687
|
93 |
proof (intro exI conjI ballI)
|
|
|
94 |
|
|
|
95 |
txt {* For all @{text "y \<in> F"} holds @{text "a y \<le> \<xi>"}: *}
|
|
|
96 |
|
|
10007
|
97 |
fix y assume y: "y \<in> F"
|
|
10687
|
98 |
show "a y \<le> xi"
|
|
|
99 |
proof (rule isUbD)
|
|
10007
|
100 |
show "isUb UNIV ?S xi" ..
|
|
10687
|
101 |
qed (blast!)
|
|
10007
|
102 |
next
|
|
7917
|
103 |
|
|
10687
|
104 |
txt {* For all @{text "y \<in> F"} holds @{text "\<xi> \<le> b y"}: *}
|
|
7917
|
105 |
|
|
10007
|
106 |
fix y assume "y \<in> F"
|
|
10687
|
107 |
show "xi \<le> b y"
|
|
10007
|
108 |
proof (intro isLub_le_isUb isUbI setleI)
|
|
|
109 |
show "b y \<in> UNIV" ..
|
|
10687
|
110 |
show "\<forall>ya \<in> ?S. ya \<le> b y"
|
|
10007
|
111 |
proof
|
|
|
112 |
fix au assume au: "au \<in> ?S "
|
|
|
113 |
hence "\<exists>u \<in> F. au = a u" by fast
|
|
10687
|
114 |
thus "au \<le> b y"
|
|
10007
|
115 |
proof
|
|
10687
|
116 |
fix u assume "u \<in> F" assume "au = a u"
|
|
|
117 |
also have "... \<le> b y" by (rule r)
|
|
10007
|
118 |
finally show ?thesis .
|
|
|
119 |
qed
|
|
|
120 |
qed
|
|
10687
|
121 |
qed
|
|
10007
|
122 |
qed
|
|
|
123 |
qed
|
|
|
124 |
qed
|
|
7917
|
125 |
|
|
10687
|
126 |
text {*
|
|
|
127 |
\medskip The function @{text h'} is defined as a
|
|
|
128 |
@{text "h' x = h y + a \<cdot> \<xi>"} where @{text "x = y + a \<cdot> \<xi>"} is a
|
|
|
129 |
linear extension of @{text h} to @{text H'}. *}
|
|
7917
|
130 |
|
|
10687
|
131 |
lemma h'_lf:
|
|
|
132 |
"h' \<equiv> \<lambda>x. let (y, a) = SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi
|
|
|
133 |
\<Longrightarrow> H' \<equiv> H + lin x0 \<Longrightarrow> is_subspace H E \<Longrightarrow> is_linearform H h \<Longrightarrow> x0 \<notin> H
|
|
|
134 |
\<Longrightarrow> x0 \<in> E \<Longrightarrow> x0 \<noteq> 0 \<Longrightarrow> is_vectorspace E
|
|
|
135 |
\<Longrightarrow> is_linearform H' h'"
|
|
10007
|
136 |
proof -
|
|
10687
|
137 |
assume h'_def:
|
|
|
138 |
"h' \<equiv> (\<lambda>x. let (y, a) = SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H
|
|
7917
|
139 |
in h y + a * xi)"
|
|
10687
|
140 |
and H'_def: "H' \<equiv> H + lin x0"
|
|
|
141 |
and vs: "is_subspace H E" "is_linearform H h" "x0 \<notin> H"
|
|
|
142 |
"x0 \<noteq> 0" "x0 \<in> E" "is_vectorspace E"
|
|
7917
|
143 |
|
|
10687
|
144 |
have h': "is_vectorspace H'"
|
|
10007
|
145 |
proof (unfold H'_def, rule vs_sum_vs)
|
|
|
146 |
show "is_subspace (lin x0) E" ..
|
|
10687
|
147 |
qed
|
|
7917
|
148 |
|
|
10007
|
149 |
show ?thesis
|
|
|
150 |
proof
|
|
10687
|
151 |
fix x1 x2 assume x1: "x1 \<in> H'" and x2: "x2 \<in> H'"
|
|
7917
|
152 |
|
|
10687
|
153 |
txt {* We now have to show that @{text h'} is additive, i.~e.\
|
|
|
154 |
@{text "h' (x\<^sub>1 + x\<^sub>2) = h' x\<^sub>1 + h' x\<^sub>2"} for
|
|
|
155 |
@{text "x\<^sub>1, x\<^sub>2 \<in> H"}. *}
|
|
7917
|
156 |
|
|
10687
|
157 |
have x1x2: "x1 + x2 \<in> H'"
|
|
|
158 |
by (rule vs_add_closed, rule h')
|
|
|
159 |
from x1
|
|
|
160 |
have ex_x1: "\<exists>y1 a1. x1 = y1 + a1 \<cdot> x0 \<and> y1 \<in> H"
|
|
10007
|
161 |
by (unfold H'_def vs_sum_def lin_def) fast
|
|
10687
|
162 |
from x2
|
|
|
163 |
have ex_x2: "\<exists>y2 a2. x2 = y2 + a2 \<cdot> x0 \<and> y2 \<in> H"
|
|
10007
|
164 |
by (unfold H'_def vs_sum_def lin_def) fast
|
|
10687
|
165 |
from x1x2
|
|
10007
|
166 |
have ex_x1x2: "\<exists>y a. x1 + x2 = y + a \<cdot> x0 \<and> y \<in> H"
|
|
|
167 |
by (unfold H'_def vs_sum_def lin_def) fast
|
|
7917
|
168 |
|
|
10007
|
169 |
from ex_x1 ex_x2 ex_x1x2
|
|
|
170 |
show "h' (x1 + x2) = h' x1 + h' x2"
|
|
|
171 |
proof (elim exE conjE)
|
|
|
172 |
fix y1 y2 y a1 a2 a
|
|
9503
|
173 |
assume y1: "x1 = y1 + a1 \<cdot> x0" and y1': "y1 \<in> H"
|
|
10687
|
174 |
and y2: "x2 = y2 + a2 \<cdot> x0" and y2': "y2 \<in> H"
|
|
|
175 |
and y: "x1 + x2 = y + a \<cdot> x0" and y': "y \<in> H"
|
|
9379
|
176 |
txt {* \label{decomp-H-use}*}
|
|
10687
|
177 |
have ya: "y1 + y2 = y \<and> a1 + a2 = a"
|
|
9379
|
178 |
proof (rule decomp_H')
|
|
10687
|
179 |
show "y1 + y2 + (a1 + a2) \<cdot> x0 = y + a \<cdot> x0"
|
|
10007
|
180 |
by (simp! add: vs_add_mult_distrib2 [of E])
|
|
|
181 |
show "y1 + y2 \<in> H" ..
|
|
|
182 |
qed
|
|
7917
|
183 |
|
|
10007
|
184 |
have "h' (x1 + x2) = h y + a * xi"
|
|
10687
|
185 |
by (rule h'_definite)
|
|
|
186 |
also have "... = h (y1 + y2) + (a1 + a2) * xi"
|
|
10007
|
187 |
by (simp add: ya)
|
|
10687
|
188 |
also from vs y1' y2'
|
|
|
189 |
have "... = h y1 + h y2 + a1 * xi + a2 * xi"
|
|
|
190 |
by (simp add: linearform_add [of H]
|
|
10007
|
191 |
real_add_mult_distrib)
|
|
10687
|
192 |
also have "... = (h y1 + a1 * xi) + (h y2 + a2 * xi)"
|
|
10007
|
193 |
by simp
|
|
|
194 |
also have "h y1 + a1 * xi = h' x1"
|
|
|
195 |
by (rule h'_definite [symmetric])
|
|
10687
|
196 |
also have "h y2 + a2 * xi = h' x2"
|
|
10007
|
197 |
by (rule h'_definite [symmetric])
|
|
|
198 |
finally show ?thesis .
|
|
|
199 |
qed
|
|
7917
|
200 |
|
|
10687
|
201 |
txt {* We further have to show that @{text h'} is multiplicative,
|
|
|
202 |
i.~e.\ @{text "h' (c \<cdot> x\<^sub>1) = c \<cdot> h' x\<^sub>1"} for @{text "x \<in> H"}
|
|
|
203 |
and @{text "c \<in> \<real>"}. *}
|
|
|
204 |
|
|
|
205 |
next
|
|
|
206 |
fix c x1 assume x1: "x1 \<in> H'"
|
|
10007
|
207 |
have ax1: "c \<cdot> x1 \<in> H'"
|
|
|
208 |
by (rule vs_mult_closed, rule h')
|
|
10687
|
209 |
from x1
|
|
|
210 |
have ex_x: "\<And>x. x\<in> H' \<Longrightarrow> \<exists>y a. x = y + a \<cdot> x0 \<and> y \<in> H"
|
|
10007
|
211 |
by (unfold H'_def vs_sum_def lin_def) fast
|
|
7917
|
212 |
|
|
10007
|
213 |
from x1 have ex_x1: "\<exists>y1 a1. x1 = y1 + a1 \<cdot> x0 \<and> y1 \<in> H"
|
|
|
214 |
by (unfold H'_def vs_sum_def lin_def) fast
|
|
|
215 |
with ex_x [of "c \<cdot> x1", OF ax1]
|
|
10687
|
216 |
show "h' (c \<cdot> x1) = c * (h' x1)"
|
|
10007
|
217 |
proof (elim exE conjE)
|
|
10687
|
218 |
fix y1 y a1 a
|
|
9503
|
219 |
assume y1: "x1 = y1 + a1 \<cdot> x0" and y1': "y1 \<in> H"
|
|
10687
|
220 |
and y: "c \<cdot> x1 = y + a \<cdot> x0" and y': "y \<in> H"
|
|
7917
|
221 |
|
|
10687
|
222 |
have ya: "c \<cdot> y1 = y \<and> c * a1 = a"
|
|
|
223 |
proof (rule decomp_H')
|
|
|
224 |
show "c \<cdot> y1 + (c * a1) \<cdot> x0 = y + a \<cdot> x0"
|
|
10007
|
225 |
by (simp! add: vs_add_mult_distrib1)
|
|
|
226 |
show "c \<cdot> y1 \<in> H" ..
|
|
|
227 |
qed
|
|
7917
|
228 |
|
|
10687
|
229 |
have "h' (c \<cdot> x1) = h y + a * xi"
|
|
|
230 |
by (rule h'_definite)
|
|
10007
|
231 |
also have "... = h (c \<cdot> y1) + (c * a1) * xi"
|
|
|
232 |
by (simp add: ya)
|
|
10687
|
233 |
also from vs y1' have "... = c * h y1 + c * a1 * xi"
|
|
|
234 |
by (simp add: linearform_mult [of H])
|
|
|
235 |
also from vs y1' have "... = c * (h y1 + a1 * xi)"
|
|
|
236 |
by (simp add: real_add_mult_distrib2 real_mult_assoc)
|
|
|
237 |
also have "h y1 + a1 * xi = h' x1"
|
|
10007
|
238 |
by (rule h'_definite [symmetric])
|
|
|
239 |
finally show ?thesis .
|
|
|
240 |
qed
|
|
|
241 |
qed
|
|
|
242 |
qed
|
|
7917
|
243 |
|
|
10687
|
244 |
text {* \medskip The linear extension @{text h'} of @{text h}
|
|
|
245 |
is bounded by the seminorm @{text p}. *}
|
|
7917
|
246 |
|
|
9374
|
247 |
lemma h'_norm_pres:
|
|
10687
|
248 |
"h' \<equiv> \<lambda>x. let (y, a) = SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi
|
|
|
249 |
\<Longrightarrow> H' \<equiv> H + lin x0 \<Longrightarrow> x0 \<notin> H \<Longrightarrow> x0 \<in> E \<Longrightarrow> x0 \<noteq> 0 \<Longrightarrow> is_vectorspace E
|
|
|
250 |
\<Longrightarrow> is_subspace H E \<Longrightarrow> is_seminorm E p \<Longrightarrow> is_linearform H h
|
|
|
251 |
\<Longrightarrow> \<forall>y \<in> H. h y \<le> p y
|
|
|
252 |
\<Longrightarrow> \<forall>y \<in> H. - p (y + x0) - h y \<le> xi \<and> xi \<le> p (y + x0) - h y
|
|
|
253 |
\<Longrightarrow> \<forall>x \<in> H'. h' x \<le> p x"
|
|
|
254 |
proof
|
|
|
255 |
assume h'_def:
|
|
|
256 |
"h' \<equiv> (\<lambda>x. let (y, a) = SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H
|
|
7917
|
257 |
in (h y) + a * xi)"
|
|
10687
|
258 |
and H'_def: "H' \<equiv> H + lin x0"
|
|
|
259 |
and vs: "x0 \<notin> H" "x0 \<in> E" "x0 \<noteq> 0" "is_vectorspace E"
|
|
|
260 |
"is_subspace H E" "is_seminorm E p" "is_linearform H h"
|
|
|
261 |
and a: "\<forall>y \<in> H. h y \<le> p y"
|
|
|
262 |
presume a1: "\<forall>ya \<in> H. - p (ya + x0) - h ya \<le> xi"
|
|
|
263 |
presume a2: "\<forall>ya \<in> H. xi \<le> p (ya + x0) - h ya"
|
|
|
264 |
fix x assume "x \<in> H'"
|
|
|
265 |
have ex_x:
|
|
|
266 |
"\<And>x. x \<in> H' \<Longrightarrow> \<exists>y a. x = y + a \<cdot> x0 \<and> y \<in> H"
|
|
10007
|
267 |
by (unfold H'_def vs_sum_def lin_def) fast
|
|
|
268 |
have "\<exists>y a. x = y + a \<cdot> x0 \<and> y \<in> H"
|
|
|
269 |
by (rule ex_x)
|
|
10687
|
270 |
thus "h' x \<le> p x"
|
|
10007
|
271 |
proof (elim exE conjE)
|
|
|
272 |
fix y a assume x: "x = y + a \<cdot> x0" and y: "y \<in> H"
|
|
|
273 |
have "h' x = h y + a * xi"
|
|
|
274 |
by (rule h'_definite)
|
|
7917
|
275 |
|
|
10687
|
276 |
txt {* Now we show @{text "h y + a \<cdot> \<xi> \<le> p (y + a \<cdot> x\<^sub>0)"}
|
|
|
277 |
by case analysis on @{text a}. *}
|
|
7917
|
278 |
|
|
10687
|
279 |
also have "... \<le> p (y + a \<cdot> x0)"
|
|
10007
|
280 |
proof (rule linorder_cases)
|
|
7917
|
281 |
|
|
10687
|
282 |
assume z: "a = #0"
|
|
10007
|
283 |
with vs y a show ?thesis by simp
|
|
7917
|
284 |
|
|
10687
|
285 |
txt {* In the case @{text "a < 0"}, we use @{text "a\<^sub>1"}
|
|
|
286 |
with @{text ya} taken as @{text "y / a"}: *}
|
|
7917
|
287 |
|
|
10007
|
288 |
next
|
|
|
289 |
assume lz: "a < #0" hence nz: "a \<noteq> #0" by simp
|
|
10687
|
290 |
from a1
|
|
|
291 |
have "- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y) \<le> xi"
|
|
10007
|
292 |
by (rule bspec) (simp!)
|
|
7917
|
293 |
|
|
10687
|
294 |
txt {* The thesis for this case now follows by a short
|
|
|
295 |
calculation. *}
|
|
|
296 |
hence "a * xi \<le> a * (- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
|
|
10007
|
297 |
by (rule real_mult_less_le_anti [OF lz])
|
|
10687
|
298 |
also
|
|
10606
|
299 |
have "... = - a * (p (inverse a \<cdot> y + x0)) - a * (h (inverse a \<cdot> y))"
|
|
10007
|
300 |
by (rule real_mult_diff_distrib)
|
|
10687
|
301 |
also from lz vs y
|
|
10606
|
302 |
have "- a * (p (inverse a \<cdot> y + x0)) = p (a \<cdot> (inverse a \<cdot> y + x0))"
|
|
10007
|
303 |
by (simp add: seminorm_abs_homogenous abs_minus_eqI2)
|
|
|
304 |
also from nz vs y have "... = p (y + a \<cdot> x0)"
|
|
|
305 |
by (simp add: vs_add_mult_distrib1)
|
|
10606
|
306 |
also from nz vs y have "a * (h (inverse a \<cdot> y)) = h y"
|
|
10007
|
307 |
by (simp add: linearform_mult [symmetric])
|
|
10687
|
308 |
finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
|
|
7917
|
309 |
|
|
10687
|
310 |
hence "h y + a * xi \<le> h y + p (y + a \<cdot> x0) - h y"
|
|
10007
|
311 |
by (simp add: real_add_left_cancel_le)
|
|
|
312 |
thus ?thesis by simp
|
|
7917
|
313 |
|
|
10687
|
314 |
txt {* In the case @{text "a > 0"}, we use @{text "a\<^sub>2"}
|
|
|
315 |
with @{text ya} taken as @{text "y / a"}: *}
|
|
7978
|
316 |
|
|
10687
|
317 |
next
|
|
10007
|
318 |
assume gz: "#0 < a" hence nz: "a \<noteq> #0" by simp
|
|
10687
|
319 |
from a2 have "xi \<le> p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y)"
|
|
10007
|
320 |
by (rule bspec) (simp!)
|
|
7917
|
321 |
|
|
7978
|
322 |
txt {* The thesis for this case follows by a short
|
|
10007
|
323 |
calculation: *}
|
|
7917
|
324 |
|
|
10687
|
325 |
with gz
|
|
|
326 |
have "a * xi \<le> a * (p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
|
|
10007
|
327 |
by (rule real_mult_less_le_mono)
|
|
10606
|
328 |
also have "... = a * p (inverse a \<cdot> y + x0) - a * h (inverse a \<cdot> y)"
|
|
10687
|
329 |
by (rule real_mult_diff_distrib2)
|
|
|
330 |
also from gz vs y
|
|
10606
|
331 |
have "a * p (inverse a \<cdot> y + x0) = p (a \<cdot> (inverse a \<cdot> y + x0))"
|
|
10007
|
332 |
by (simp add: seminorm_abs_homogenous abs_eqI2)
|
|
|
333 |
also from nz vs y have "... = p (y + a \<cdot> x0)"
|
|
|
334 |
by (simp add: vs_add_mult_distrib1)
|
|
10606
|
335 |
also from nz vs y have "a * h (inverse a \<cdot> y) = h y"
|
|
10687
|
336 |
by (simp add: linearform_mult [symmetric])
|
|
|
337 |
finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
|
|
|
338 |
|
|
|
339 |
hence "h y + a * xi \<le> h y + (p (y + a \<cdot> x0) - h y)"
|
|
10007
|
340 |
by (simp add: real_add_left_cancel_le)
|
|
|
341 |
thus ?thesis by simp
|
|
|
342 |
qed
|
|
|
343 |
also from x have "... = p x" by simp
|
|
|
344 |
finally show ?thesis .
|
|
|
345 |
qed
|
|
10687
|
346 |
qed blast+
|
|
7917
|
347 |
|
|
10007
|
348 |
end
|