7917
|
1 |
(* Title: HOL/Real/HahnBanach/VectorSpace.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Gertrud Bauer, TU Munich
|
|
4 |
*)
|
|
5 |
|
|
6 |
header {* Vector spaces *};
|
|
7 |
|
|
8 |
theory VectorSpace = Bounds + Aux:;
|
|
9 |
|
|
10 |
subsection {* Signature *};
|
|
11 |
|
7927
|
12 |
text {* For the definition of real vector spaces a type $\alpha$
|
7978
|
13 |
of the sort $\{ \idt{plus}, \idt{minus}\}$ is considered, on which a
|
|
14 |
real scalar multiplication $\mult$, and a zero
|
|
15 |
element $\zero$ is defined. *};
|
7917
|
16 |
|
|
17 |
consts
|
|
18 |
prod :: "[real, 'a] => 'a" (infixr "<*>" 70)
|
|
19 |
zero :: 'a ("<0>");
|
|
20 |
|
|
21 |
syntax (symbols)
|
|
22 |
prod :: "[real, 'a] => 'a" (infixr "\<prod>" 70)
|
|
23 |
zero :: 'a ("\<zero>");
|
|
24 |
|
7927
|
25 |
(* text {* The unary and binary minus can be considered as
|
7917
|
26 |
abbreviations: *};
|
7927
|
27 |
*)
|
7917
|
28 |
|
|
29 |
(***
|
|
30 |
constdefs
|
|
31 |
negate :: "'a => 'a" ("- _" [100] 100)
|
|
32 |
"- x == (- 1r) <*> x"
|
|
33 |
diff :: "'a => 'a => 'a" (infixl "-" 68)
|
|
34 |
"x - y == x + - y";
|
|
35 |
***)
|
|
36 |
|
|
37 |
subsection {* Vector space laws *};
|
|
38 |
|
7927
|
39 |
text {* A \emph{vector space} is a non-empty set $V$ of elements from
|
|
40 |
$\alpha$ with the following vector space laws: The set $V$ is closed
|
|
41 |
under addition and scalar multiplication, addition is associative
|
|
42 |
and commutative; $\minus x$ is the inverse of $x$ w.~r.~t.~addition
|
|
43 |
and $\zero$ is the neutral element of addition. Addition and
|
|
44 |
multiplication are distributive; scalar multiplication is
|
7978
|
45 |
associative and the real number $1$ is the neutral element of scalar
|
7927
|
46 |
multiplication.
|
7917
|
47 |
*};
|
|
48 |
|
|
49 |
constdefs
|
|
50 |
is_vectorspace :: "('a::{plus,minus}) set => bool"
|
|
51 |
"is_vectorspace V == V ~= {}
|
|
52 |
& (ALL x:V. ALL y:V. ALL z:V. ALL a b.
|
|
53 |
x + y : V
|
|
54 |
& a <*> x : V
|
7978
|
55 |
& (x + y) + z = x + (y + z)
|
7917
|
56 |
& x + y = y + x
|
|
57 |
& x - x = <0>
|
|
58 |
& <0> + x = x
|
|
59 |
& a <*> (x + y) = a <*> x + a <*> y
|
|
60 |
& (a + b) <*> x = a <*> x + b <*> x
|
|
61 |
& (a * b) <*> x = a <*> b <*> x
|
|
62 |
& 1r <*> x = x
|
|
63 |
& - x = (- 1r) <*> x
|
|
64 |
& x - y = x + - y)";
|
|
65 |
|
|
66 |
text_raw {* \medskip *};
|
|
67 |
text {* The corresponding introduction rule is:*};
|
|
68 |
|
|
69 |
lemma vsI [intro]:
|
|
70 |
"[| <0>:V;
|
|
71 |
ALL x:V. ALL y:V. x + y : V;
|
7978
|
72 |
ALL x:V. ALL a. a <*> x : V;
|
|
73 |
ALL x:V. ALL y:V. ALL z:V. (x + y) + z = x + (y + z);
|
7917
|
74 |
ALL x:V. ALL y:V. x + y = y + x;
|
|
75 |
ALL x:V. x - x = <0>;
|
|
76 |
ALL x:V. <0> + x = x;
|
|
77 |
ALL x:V. ALL y:V. ALL a. a <*> (x + y) = a <*> x + a <*> y;
|
|
78 |
ALL x:V. ALL a b. (a + b) <*> x = a <*> x + b <*> x;
|
|
79 |
ALL x:V. ALL a b. (a * b) <*> x = a <*> b <*> x;
|
|
80 |
ALL x:V. 1r <*> x = x;
|
|
81 |
ALL x:V. - x = (- 1r) <*> x;
|
7978
|
82 |
ALL x:V. ALL y:V. x - y = x + - y |] ==> is_vectorspace V";
|
7917
|
83 |
proof (unfold is_vectorspace_def, intro conjI ballI allI);
|
|
84 |
fix x y z;
|
|
85 |
assume "x:V" "y:V" "z:V"
|
|
86 |
"ALL x:V. ALL y:V. ALL z:V. x + y + z = x + (y + z)";
|
|
87 |
thus "x + y + z = x + (y + z)"; by (elim bspec[elimify]);
|
|
88 |
qed force+;
|
|
89 |
|
|
90 |
text_raw {* \medskip *};
|
|
91 |
text {* The corresponding destruction rules are: *};
|
|
92 |
|
|
93 |
lemma negate_eq1:
|
|
94 |
"[| is_vectorspace V; x:V |] ==> - x = (- 1r) <*> x";
|
|
95 |
by (unfold is_vectorspace_def) simp;
|
|
96 |
|
|
97 |
lemma diff_eq1:
|
|
98 |
"[| is_vectorspace V; x:V; y:V |] ==> x - y = x + - y";
|
|
99 |
by (unfold is_vectorspace_def) simp;
|
|
100 |
|
|
101 |
lemma negate_eq2:
|
|
102 |
"[| is_vectorspace V; x:V |] ==> (- 1r) <*> x = - x";
|
|
103 |
by (unfold is_vectorspace_def) simp;
|
|
104 |
|
|
105 |
lemma diff_eq2:
|
|
106 |
"[| is_vectorspace V; x:V; y:V |] ==> x + - y = x - y";
|
|
107 |
by (unfold is_vectorspace_def) simp;
|
|
108 |
|
|
109 |
lemma vs_not_empty [intro !!]: "is_vectorspace V ==> (V ~= {})";
|
|
110 |
by (unfold is_vectorspace_def) simp;
|
|
111 |
|
|
112 |
lemma vs_add_closed [simp, intro!!]:
|
7978
|
113 |
"[| is_vectorspace V; x:V; y:V |] ==> x + y : V";
|
7917
|
114 |
by (unfold is_vectorspace_def) simp;
|
|
115 |
|
|
116 |
lemma vs_mult_closed [simp, intro!!]:
|
|
117 |
"[| is_vectorspace V; x:V |] ==> a <*> x : V";
|
|
118 |
by (unfold is_vectorspace_def) simp;
|
|
119 |
|
|
120 |
lemma vs_diff_closed [simp, intro!!]:
|
7978
|
121 |
"[| is_vectorspace V; x:V; y:V |] ==> x - y : V";
|
7917
|
122 |
by (simp add: diff_eq1 negate_eq1);
|
|
123 |
|
|
124 |
lemma vs_neg_closed [simp, intro!!]:
|
|
125 |
"[| is_vectorspace V; x:V |] ==> - x : V";
|
|
126 |
by (simp add: negate_eq1);
|
|
127 |
|
|
128 |
lemma vs_add_assoc [simp]:
|
7978
|
129 |
"[| is_vectorspace V; x:V; y:V; z:V |]
|
|
130 |
==> (x + y) + z = x + (y + z)";
|
7917
|
131 |
by (unfold is_vectorspace_def) fast;
|
|
132 |
|
|
133 |
lemma vs_add_commute [simp]:
|
|
134 |
"[| is_vectorspace V; x:V; y:V |] ==> y + x = x + y";
|
|
135 |
by (unfold is_vectorspace_def) simp;
|
|
136 |
|
|
137 |
lemma vs_add_left_commute [simp]:
|
|
138 |
"[| is_vectorspace V; x:V; y:V; z:V |]
|
|
139 |
==> x + (y + z) = y + (x + z)";
|
|
140 |
proof -;
|
|
141 |
assume "is_vectorspace V" "x:V" "y:V" "z:V";
|
|
142 |
hence "x + (y + z) = (x + y) + z";
|
|
143 |
by (simp only: vs_add_assoc);
|
|
144 |
also; have "... = (y + x) + z"; by (simp! only: vs_add_commute);
|
|
145 |
also; have "... = y + (x + z)"; by (simp! only: vs_add_assoc);
|
|
146 |
finally; show ?thesis; .;
|
|
147 |
qed;
|
|
148 |
|
|
149 |
theorems vs_add_ac = vs_add_assoc vs_add_commute vs_add_left_commute;
|
|
150 |
|
|
151 |
lemma vs_diff_self [simp]:
|
|
152 |
"[| is_vectorspace V; x:V |] ==> x - x = <0>";
|
|
153 |
by (unfold is_vectorspace_def) simp;
|
|
154 |
|
7978
|
155 |
text {* The existence of the zero element of a vector space
|
|
156 |
follows from the non-emptiness of carrier set. *};
|
7917
|
157 |
|
|
158 |
lemma zero_in_vs [simp, intro]: "is_vectorspace V ==> <0>:V";
|
|
159 |
proof -;
|
|
160 |
assume "is_vectorspace V";
|
|
161 |
have "V ~= {}"; ..;
|
|
162 |
hence "EX x. x:V"; by force;
|
|
163 |
thus ?thesis;
|
|
164 |
proof;
|
|
165 |
fix x; assume "x:V";
|
|
166 |
have "<0> = x - x"; by (simp!);
|
|
167 |
also; have "... : V"; by (simp! only: vs_diff_closed);
|
|
168 |
finally; show ?thesis; .;
|
|
169 |
qed;
|
|
170 |
qed;
|
|
171 |
|
|
172 |
lemma vs_add_zero_left [simp]:
|
|
173 |
"[| is_vectorspace V; x:V |] ==> <0> + x = x";
|
|
174 |
by (unfold is_vectorspace_def) simp;
|
|
175 |
|
|
176 |
lemma vs_add_zero_right [simp]:
|
|
177 |
"[| is_vectorspace V; x:V |] ==> x + <0> = x";
|
|
178 |
proof -;
|
|
179 |
assume "is_vectorspace V" "x:V";
|
|
180 |
hence "x + <0> = <0> + x"; by simp;
|
|
181 |
also; have "... = x"; by (simp!);
|
|
182 |
finally; show ?thesis; .;
|
|
183 |
qed;
|
|
184 |
|
|
185 |
lemma vs_add_mult_distrib1:
|
|
186 |
"[| is_vectorspace V; x:V; y:V |]
|
|
187 |
==> a <*> (x + y) = a <*> x + a <*> y";
|
|
188 |
by (unfold is_vectorspace_def) simp;
|
|
189 |
|
|
190 |
lemma vs_add_mult_distrib2:
|
|
191 |
"[| is_vectorspace V; x:V |]
|
|
192 |
==> (a + b) <*> x = a <*> x + b <*> x";
|
|
193 |
by (unfold is_vectorspace_def) simp;
|
|
194 |
|
|
195 |
lemma vs_mult_assoc:
|
|
196 |
"[| is_vectorspace V; x:V |] ==> (a * b) <*> x = a <*> (b <*> x)";
|
|
197 |
by (unfold is_vectorspace_def) simp;
|
|
198 |
|
|
199 |
lemma vs_mult_assoc2 [simp]:
|
|
200 |
"[| is_vectorspace V; x:V |] ==> a <*> b <*> x = (a * b) <*> x";
|
|
201 |
by (simp only: vs_mult_assoc);
|
|
202 |
|
|
203 |
lemma vs_mult_1 [simp]:
|
|
204 |
"[| is_vectorspace V; x:V |] ==> 1r <*> x = x";
|
|
205 |
by (unfold is_vectorspace_def) simp;
|
|
206 |
|
|
207 |
lemma vs_diff_mult_distrib1:
|
|
208 |
"[| is_vectorspace V; x:V; y:V |]
|
|
209 |
==> a <*> (x - y) = a <*> x - a <*> y";
|
|
210 |
by (simp add: diff_eq1 negate_eq1 vs_add_mult_distrib1);
|
|
211 |
|
|
212 |
lemma vs_diff_mult_distrib2:
|
|
213 |
"[| is_vectorspace V; x:V |]
|
|
214 |
==> (a - b) <*> x = a <*> x - (b <*> x)";
|
|
215 |
proof -;
|
|
216 |
assume "is_vectorspace V" "x:V";
|
|
217 |
have " (a - b) <*> x = (a + - b ) <*> x";
|
|
218 |
by (unfold real_diff_def, simp);
|
|
219 |
also; have "... = a <*> x + (- b) <*> x";
|
|
220 |
by (rule vs_add_mult_distrib2);
|
|
221 |
also; have "... = a <*> x + - (b <*> x)";
|
|
222 |
by (simp! add: negate_eq1);
|
|
223 |
also; have "... = a <*> x - (b <*> x)";
|
|
224 |
by (simp! add: diff_eq1);
|
|
225 |
finally; show ?thesis; .;
|
|
226 |
qed;
|
|
227 |
|
7978
|
228 |
(*text_raw {* \paragraph {Further derived laws.} *};*)
|
7917
|
229 |
text_raw {* \medskip *};
|
|
230 |
text{* Further derived laws: *};
|
|
231 |
|
|
232 |
lemma vs_mult_zero_left [simp]:
|
7978
|
233 |
"[| is_vectorspace V; x:V |] ==> 0r <*> x = <0>";
|
7917
|
234 |
proof -;
|
|
235 |
assume "is_vectorspace V" "x:V";
|
|
236 |
have "0r <*> x = (1r - 1r) <*> x"; by (simp only: real_diff_self);
|
|
237 |
also; have "... = (1r + - 1r) <*> x"; by simp;
|
|
238 |
also; have "... = 1r <*> x + (- 1r) <*> x";
|
|
239 |
by (rule vs_add_mult_distrib2);
|
|
240 |
also; have "... = x + (- 1r) <*> x"; by (simp!);
|
|
241 |
also; have "... = x + - x"; by (simp! add: negate_eq2);;
|
|
242 |
also; have "... = x - x"; by (simp! add: diff_eq2);
|
|
243 |
also; have "... = <0>"; by (simp!);
|
|
244 |
finally; show ?thesis; .;
|
|
245 |
qed;
|
|
246 |
|
|
247 |
lemma vs_mult_zero_right [simp]:
|
|
248 |
"[| is_vectorspace (V:: 'a::{plus, minus} set) |]
|
|
249 |
==> a <*> <0> = (<0>::'a)";
|
|
250 |
proof -;
|
|
251 |
assume "is_vectorspace V";
|
|
252 |
have "a <*> <0> = a <*> (<0> - (<0>::'a))"; by (simp!);
|
|
253 |
also; have "... = a <*> <0> - a <*> <0>";
|
|
254 |
by (rule vs_diff_mult_distrib1) (simp!)+;
|
|
255 |
also; have "... = <0>"; by (simp!);
|
|
256 |
finally; show ?thesis; .;
|
|
257 |
qed;
|
|
258 |
|
|
259 |
lemma vs_minus_mult_cancel [simp]:
|
|
260 |
"[| is_vectorspace V; x:V |] ==> (- a) <*> - x = a <*> x";
|
|
261 |
by (simp add: negate_eq1);
|
|
262 |
|
|
263 |
lemma vs_add_minus_left_eq_diff:
|
|
264 |
"[| is_vectorspace V; x:V; y:V |] ==> - x + y = y - x";
|
|
265 |
proof -;
|
|
266 |
assume "is_vectorspace V" "x:V" "y:V";
|
|
267 |
have "- x + y = y + - x";
|
|
268 |
by (simp! add: vs_add_commute [RS sym, of V "- x"]);
|
|
269 |
also; have "... = y - x"; by (simp! add: diff_eq1);
|
|
270 |
finally; show ?thesis; .;
|
|
271 |
qed;
|
|
272 |
|
|
273 |
lemma vs_add_minus [simp]:
|
|
274 |
"[| is_vectorspace V; x:V |] ==> x + - x = <0>";
|
|
275 |
by (simp! add: diff_eq2);
|
|
276 |
|
|
277 |
lemma vs_add_minus_left [simp]:
|
|
278 |
"[| is_vectorspace V; x:V |] ==> - x + x = <0>";
|
|
279 |
by (simp! add: diff_eq2);
|
|
280 |
|
|
281 |
lemma vs_minus_minus [simp]:
|
|
282 |
"[| is_vectorspace V; x:V |] ==> - (- x) = x";
|
|
283 |
by (simp add: negate_eq1);
|
|
284 |
|
|
285 |
lemma vs_minus_zero [simp]:
|
|
286 |
"is_vectorspace (V::'a::{minus, plus} set) ==> - (<0>::'a) = <0>";
|
|
287 |
by (simp add: negate_eq1);
|
|
288 |
|
|
289 |
lemma vs_minus_zero_iff [simp]:
|
|
290 |
"[| is_vectorspace V; x:V |] ==> (- x = <0>) = (x = <0>)"
|
|
291 |
(concl is "?L = ?R");
|
|
292 |
proof -;
|
|
293 |
assume "is_vectorspace V" "x:V";
|
|
294 |
show "?L = ?R";
|
|
295 |
proof;
|
|
296 |
have "x = - (- x)"; by (rule vs_minus_minus [RS sym]);
|
|
297 |
also; assume ?L;
|
|
298 |
also; have "- ... = <0>"; by (rule vs_minus_zero);
|
|
299 |
finally; show ?R; .;
|
|
300 |
qed (simp!);
|
|
301 |
qed;
|
|
302 |
|
|
303 |
lemma vs_add_minus_cancel [simp]:
|
7978
|
304 |
"[| is_vectorspace V; x:V; y:V |] ==> x + (- x + y) = y";
|
7917
|
305 |
by (simp add: vs_add_assoc [RS sym] del: vs_add_commute);
|
|
306 |
|
|
307 |
lemma vs_minus_add_cancel [simp]:
|
7978
|
308 |
"[| is_vectorspace V; x:V; y:V |] ==> - x + (x + y) = y";
|
7917
|
309 |
by (simp add: vs_add_assoc [RS sym] del: vs_add_commute);
|
|
310 |
|
|
311 |
lemma vs_minus_add_distrib [simp]:
|
|
312 |
"[| is_vectorspace V; x:V; y:V |]
|
|
313 |
==> - (x + y) = - x + - y";
|
|
314 |
by (simp add: negate_eq1 vs_add_mult_distrib1);
|
|
315 |
|
|
316 |
lemma vs_diff_zero [simp]:
|
|
317 |
"[| is_vectorspace V; x:V |] ==> x - <0> = x";
|
|
318 |
by (simp add: diff_eq1);
|
|
319 |
|
|
320 |
lemma vs_diff_zero_right [simp]:
|
|
321 |
"[| is_vectorspace V; x:V |] ==> <0> - x = - x";
|
|
322 |
by (simp add:diff_eq1);
|
|
323 |
|
|
324 |
lemma vs_add_left_cancel:
|
7978
|
325 |
"[| is_vectorspace V; x:V; y:V; z:V |]
|
7917
|
326 |
==> (x + y = x + z) = (y = z)"
|
|
327 |
(concl is "?L = ?R");
|
|
328 |
proof;
|
|
329 |
assume "is_vectorspace V" "x:V" "y:V" "z:V";
|
|
330 |
have "y = <0> + y"; by (simp!);
|
|
331 |
also; have "... = - x + x + y"; by (simp!);
|
|
332 |
also; have "... = - x + (x + y)";
|
|
333 |
by (simp! only: vs_add_assoc vs_neg_closed);
|
|
334 |
also; assume ?L;
|
|
335 |
also; have "- x + ... = - x + x + z";
|
|
336 |
by (rule vs_add_assoc [RS sym]) (simp!)+;
|
|
337 |
also; have "... = z"; by (simp!);
|
7978
|
338 |
finally; show ?R; .;
|
7917
|
339 |
qed force;
|
|
340 |
|
|
341 |
lemma vs_add_right_cancel:
|
|
342 |
"[| is_vectorspace V; x:V; y:V; z:V |]
|
|
343 |
==> (y + x = z + x) = (y = z)";
|
|
344 |
by (simp only: vs_add_commute vs_add_left_cancel);
|
|
345 |
|
|
346 |
lemma vs_add_assoc_cong:
|
|
347 |
"[| is_vectorspace V; x:V; y:V; x':V; y':V; z:V |]
|
|
348 |
==> x + y = x' + y' ==> x + (y + z) = x' + (y' + z)";
|
|
349 |
by (simp only: vs_add_assoc [RS sym]);
|
|
350 |
|
|
351 |
lemma vs_mult_left_commute:
|
|
352 |
"[| is_vectorspace V; x:V; y:V; z:V |]
|
|
353 |
==> x <*> y <*> z = y <*> x <*> z";
|
|
354 |
by (simp add: real_mult_commute);
|
|
355 |
|
|
356 |
lemma vs_mult_zero_uniq :
|
|
357 |
"[| is_vectorspace V; x:V; a <*> x = <0>; x ~= <0> |] ==> a = 0r";
|
|
358 |
proof (rule classical);
|
|
359 |
assume "is_vectorspace V" "x:V" "a <*> x = <0>" "x ~= <0>";
|
|
360 |
assume "a ~= 0r";
|
|
361 |
have "x = (rinv a * a) <*> x"; by (simp!);
|
|
362 |
also; have "... = rinv a <*> (a <*> x)"; by (rule vs_mult_assoc);
|
|
363 |
also; have "... = rinv a <*> <0>"; by (simp!);
|
|
364 |
also; have "... = <0>"; by (simp!);
|
|
365 |
finally; have "x = <0>"; .;
|
|
366 |
thus "a = 0r"; by contradiction;
|
|
367 |
qed;
|
|
368 |
|
|
369 |
lemma vs_mult_left_cancel:
|
|
370 |
"[| is_vectorspace V; x:V; y:V; a ~= 0r |] ==>
|
|
371 |
(a <*> x = a <*> y) = (x = y)"
|
|
372 |
(concl is "?L = ?R");
|
|
373 |
proof;
|
|
374 |
assume "is_vectorspace V" "x:V" "y:V" "a ~= 0r";
|
|
375 |
have "x = 1r <*> x"; by (simp!);
|
|
376 |
also; have "... = (rinv a * a) <*> x"; by (simp!);
|
|
377 |
also; have "... = rinv a <*> (a <*> x)";
|
|
378 |
by (simp! only: vs_mult_assoc);
|
|
379 |
also; assume ?L;
|
|
380 |
also; have "rinv a <*> ... = y"; by (simp!);
|
7978
|
381 |
finally; show ?R; .;
|
7917
|
382 |
qed simp;
|
|
383 |
|
|
384 |
lemma vs_mult_right_cancel: (*** forward ***)
|
|
385 |
"[| is_vectorspace V; x:V; x ~= <0> |]
|
|
386 |
==> (a <*> x = b <*> x) = (a = b)" (concl is "?L = ?R");
|
|
387 |
proof;
|
|
388 |
assume "is_vectorspace V" "x:V" "x ~= <0>";
|
|
389 |
have "(a - b) <*> x = a <*> x - b <*> x";
|
|
390 |
by (simp! add: vs_diff_mult_distrib2);
|
|
391 |
also; assume ?L; hence "a <*> x - b <*> x = <0>"; by (simp!);
|
|
392 |
finally; have "(a - b) <*> x = <0>"; .;
|
|
393 |
hence "a - b = 0r"; by (simp! add: vs_mult_zero_uniq);
|
|
394 |
thus "a = b"; by (rule real_add_minus_eq);
|
|
395 |
qed simp; (***
|
|
396 |
|
|
397 |
backward :
|
|
398 |
lemma vs_mult_right_cancel:
|
|
399 |
"[| is_vectorspace V; x:V; x ~= <0> |] ==>
|
|
400 |
(a <*> x = b <*> x) = (a = b)"
|
|
401 |
(concl is "?L = ?R");
|
|
402 |
proof;
|
|
403 |
assume "is_vectorspace V" "x:V" "x ~= <0>";
|
|
404 |
assume l: ?L;
|
|
405 |
show "a = b";
|
|
406 |
proof (rule real_add_minus_eq);
|
|
407 |
show "a - b = 0r";
|
|
408 |
proof (rule vs_mult_zero_uniq);
|
|
409 |
have "(a - b) <*> x = a <*> x - b <*> x";
|
|
410 |
by (simp! add: vs_diff_mult_distrib2);
|
|
411 |
also; from l; have "a <*> x - b <*> x = <0>"; by (simp!);
|
|
412 |
finally; show "(a - b) <*> x = <0>"; .;
|
|
413 |
qed;
|
|
414 |
qed;
|
|
415 |
next;
|
|
416 |
assume ?R;
|
|
417 |
thus ?L; by simp;
|
|
418 |
qed;
|
|
419 |
**)
|
|
420 |
|
|
421 |
lemma vs_eq_diff_eq:
|
|
422 |
"[| is_vectorspace V; x:V; y:V; z:V |] ==>
|
|
423 |
(x = z - y) = (x + y = z)"
|
|
424 |
(concl is "?L = ?R" );
|
|
425 |
proof -;
|
|
426 |
assume vs: "is_vectorspace V" "x:V" "y:V" "z:V";
|
|
427 |
show "?L = ?R";
|
|
428 |
proof;
|
|
429 |
assume ?L;
|
|
430 |
hence "x + y = z - y + y"; by simp;
|
|
431 |
also; have "... = z + - y + y"; by (simp! add: diff_eq1);
|
|
432 |
also; have "... = z + (- y + y)";
|
|
433 |
by (rule vs_add_assoc) (simp!)+;
|
|
434 |
also; from vs; have "... = z + <0>";
|
|
435 |
by (simp only: vs_add_minus_left);
|
|
436 |
also; from vs; have "... = z"; by (simp only: vs_add_zero_right);
|
7978
|
437 |
finally; show ?R; .;
|
7917
|
438 |
next;
|
|
439 |
assume ?R;
|
|
440 |
hence "z - y = (x + y) - y"; by simp;
|
|
441 |
also; from vs; have "... = x + y + - y";
|
|
442 |
by (simp add: diff_eq1);
|
|
443 |
also; have "... = x + (y + - y)";
|
|
444 |
by (rule vs_add_assoc) (simp!)+;
|
|
445 |
also; have "... = x"; by (simp!);
|
|
446 |
finally; show ?L; by (rule sym);
|
|
447 |
qed;
|
|
448 |
qed;
|
|
449 |
|
|
450 |
lemma vs_add_minus_eq_minus:
|
7978
|
451 |
"[| is_vectorspace V; x:V; y:V; x + y = <0> |] ==> x = - y";
|
7917
|
452 |
proof -;
|
|
453 |
assume "is_vectorspace V" "x:V" "y:V";
|
|
454 |
have "x = (- y + y) + x"; by (simp!);
|
|
455 |
also; have "... = - y + (x + y)"; by (simp!);
|
|
456 |
also; assume "x + y = <0>";
|
|
457 |
also; have "- y + <0> = - y"; by (simp!);
|
|
458 |
finally; show "x = - y"; .;
|
|
459 |
qed;
|
|
460 |
|
|
461 |
lemma vs_add_minus_eq:
|
|
462 |
"[| is_vectorspace V; x:V; y:V; x - y = <0> |] ==> x = y";
|
|
463 |
proof -;
|
|
464 |
assume "is_vectorspace V" "x:V" "y:V" "x - y = <0>";
|
|
465 |
assume "x - y = <0>";
|
|
466 |
hence e: "x + - y = <0>"; by (simp! add: diff_eq1);
|
|
467 |
with _ _ _; have "x = - (- y)";
|
|
468 |
by (rule vs_add_minus_eq_minus) (simp!)+;
|
|
469 |
thus "x = y"; by (simp!);
|
|
470 |
qed;
|
|
471 |
|
|
472 |
lemma vs_add_diff_swap:
|
7978
|
473 |
"[| is_vectorspace V; a:V; b:V; c:V; d:V; a + b = c + d |]
|
7917
|
474 |
==> a - c = d - b";
|
|
475 |
proof -;
|
|
476 |
assume vs: "is_vectorspace V" "a:V" "b:V" "c:V" "d:V"
|
7978
|
477 |
and eq: "a + b = c + d";
|
7917
|
478 |
have "- c + (a + b) = - c + (c + d)";
|
|
479 |
by (simp! add: vs_add_left_cancel);
|
|
480 |
also; have "... = d"; by (rule vs_minus_add_cancel);
|
|
481 |
finally; have eq: "- c + (a + b) = d"; .;
|
|
482 |
from vs; have "a - c = (- c + (a + b)) + - b";
|
|
483 |
by (simp add: vs_add_ac diff_eq1);
|
|
484 |
also; from eq; have "... = d + - b";
|
|
485 |
by (simp! add: vs_add_right_cancel);
|
|
486 |
also; have "... = d - b"; by (simp! add : diff_eq2);
|
|
487 |
finally; show "a - c = d - b"; .;
|
|
488 |
qed;
|
|
489 |
|
|
490 |
lemma vs_add_cancel_21:
|
7978
|
491 |
"[| is_vectorspace V; x:V; y:V; z:V; u:V |]
|
7917
|
492 |
==> (x + (y + z) = y + u) = ((x + z) = u)"
|
7978
|
493 |
(concl is "?L = ?R");
|
7917
|
494 |
proof -;
|
|
495 |
assume "is_vectorspace V" "x:V" "y:V""z:V" "u:V";
|
|
496 |
show "?L = ?R";
|
|
497 |
proof;
|
|
498 |
have "x + z = - y + y + (x + z)"; by (simp!);
|
|
499 |
also; have "... = - y + (y + (x + z))";
|
|
500 |
by (rule vs_add_assoc) (simp!)+;
|
|
501 |
also; have "y + (x + z) = x + (y + z)"; by (simp!);
|
|
502 |
also; assume ?L;
|
|
503 |
also; have "- y + (y + u) = u"; by (simp!);
|
|
504 |
finally; show ?R; .;
|
|
505 |
qed (simp! only: vs_add_left_commute [of V x]);
|
|
506 |
qed;
|
|
507 |
|
|
508 |
lemma vs_add_cancel_end:
|
|
509 |
"[| is_vectorspace V; x:V; y:V; z:V |]
|
|
510 |
==> (x + (y + z) = y) = (x = - z)"
|
|
511 |
(concl is "?L = ?R" );
|
|
512 |
proof -;
|
|
513 |
assume "is_vectorspace V" "x:V" "y:V" "z:V";
|
|
514 |
show "?L = ?R";
|
|
515 |
proof;
|
|
516 |
assume l: ?L;
|
|
517 |
have "x + z = <0>";
|
|
518 |
proof (rule vs_add_left_cancel [RS iffD1]);
|
|
519 |
have "y + (x + z) = x + (y + z)"; by (simp!);
|
|
520 |
also; note l;
|
|
521 |
also; have "y = y + <0>"; by (simp!);
|
|
522 |
finally; show "y + (x + z) = y + <0>"; .;
|
|
523 |
qed (simp!)+;
|
|
524 |
thus "x = - z"; by (simp! add: vs_add_minus_eq_minus);
|
|
525 |
next;
|
|
526 |
assume r: ?R;
|
|
527 |
hence "x + (y + z) = - z + (y + z)"; by simp;
|
|
528 |
also; have "... = y + (- z + z)";
|
|
529 |
by (simp! only: vs_add_left_commute);
|
|
530 |
also; have "... = y"; by (simp!);
|
|
531 |
finally; show ?L; .;
|
|
532 |
qed;
|
|
533 |
qed;
|
|
534 |
|
|
535 |
end; |