author | haftmann |
Mon, 19 Jul 2010 16:09:44 +0200 | |
changeset 37887 | 2ae085b07f2f |
parent 36778 | 739a9379e29b |
child 41413 | 64cd30d6b0b8 |
permissions | -rw-r--r-- |
31795
be3e1cc5005c
standard naming conventions for session and theories;
wenzelm
parents:
29252
diff
changeset
|
1 |
(* Title: HOL/Hahn_Banach/Vector_Space.thy |
7917 | 2 |
Author: Gertrud Bauer, TU Munich |
3 |
*) |
|
4 |
||
9035 | 5 |
header {* Vector spaces *} |
7917 | 6 |
|
31795
be3e1cc5005c
standard naming conventions for session and theories;
wenzelm
parents:
29252
diff
changeset
|
7 |
theory Vector_Space |
27612 | 8 |
imports Real Bounds Zorn |
9 |
begin |
|
7917 | 10 |
|
9035 | 11 |
subsection {* Signature *} |
7917 | 12 |
|
10687 | 13 |
text {* |
14 |
For the definition of real vector spaces a type @{typ 'a} of the |
|
15 |
sort @{text "{plus, minus, zero}"} is considered, on which a real |
|
16 |
scalar multiplication @{text \<cdot>} is declared. |
|
17 |
*} |
|
7917 | 18 |
|
19 |
consts |
|
10687 | 20 |
prod :: "real \<Rightarrow> 'a::{plus, minus, zero} \<Rightarrow> 'a" (infixr "'(*')" 70) |
7917 | 21 |
|
21210 | 22 |
notation (xsymbols) |
19736 | 23 |
prod (infixr "\<cdot>" 70) |
21210 | 24 |
notation (HTML output) |
19736 | 25 |
prod (infixr "\<cdot>" 70) |
7917 | 26 |
|
27 |
||
9035 | 28 |
subsection {* Vector space laws *} |
7917 | 29 |
|
10687 | 30 |
text {* |
31 |
A \emph{vector space} is a non-empty set @{text V} of elements from |
|
32 |
@{typ 'a} with the following vector space laws: The set @{text V} is |
|
33 |
closed under addition and scalar multiplication, addition is |
|
34 |
associative and commutative; @{text "- x"} is the inverse of @{text |
|
35 |
x} w.~r.~t.~addition and @{text 0} is the neutral element of |
|
36 |
addition. Addition and multiplication are distributive; scalar |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11704
diff
changeset
|
37 |
multiplication is associative and the real number @{text "1"} is |
10687 | 38 |
the neutral element of scalar multiplication. |
9035 | 39 |
*} |
7917 | 40 |
|
29234 | 41 |
locale var_V = fixes V |
42 |
||
43 |
locale vectorspace = var_V + |
|
13515 | 44 |
assumes non_empty [iff, intro?]: "V \<noteq> {}" |
45 |
and add_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y \<in> V" |
|
46 |
and mult_closed [iff]: "x \<in> V \<Longrightarrow> a \<cdot> x \<in> V" |
|
47 |
and add_assoc: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + y) + z = x + (y + z)" |
|
48 |
and add_commute: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y = y + x" |
|
49 |
and diff_self [simp]: "x \<in> V \<Longrightarrow> x - x = 0" |
|
50 |
and add_zero_left [simp]: "x \<in> V \<Longrightarrow> 0 + x = x" |
|
51 |
and add_mult_distrib1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y" |
|
52 |
and add_mult_distrib2: "x \<in> V \<Longrightarrow> (a + b) \<cdot> x = a \<cdot> x + b \<cdot> x" |
|
53 |
and mult_assoc: "x \<in> V \<Longrightarrow> (a * b) \<cdot> x = a \<cdot> (b \<cdot> x)" |
|
54 |
and mult_1 [simp]: "x \<in> V \<Longrightarrow> 1 \<cdot> x = x" |
|
55 |
and negate_eq1: "x \<in> V \<Longrightarrow> - x = (- 1) \<cdot> x" |
|
56 |
and diff_eq1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y = x + - y" |
|
7917 | 57 |
|
13515 | 58 |
lemma (in vectorspace) negate_eq2: "x \<in> V \<Longrightarrow> (- 1) \<cdot> x = - x" |
59 |
by (rule negate_eq1 [symmetric]) |
|
9013
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents:
8703
diff
changeset
|
60 |
|
13515 | 61 |
lemma (in vectorspace) negate_eq2a: "x \<in> V \<Longrightarrow> -1 \<cdot> x = - x" |
62 |
by (simp add: negate_eq1) |
|
7917 | 63 |
|
13515 | 64 |
lemma (in vectorspace) diff_eq2: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + - y = x - y" |
65 |
by (rule diff_eq1 [symmetric]) |
|
7917 | 66 |
|
13515 | 67 |
lemma (in vectorspace) diff_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y \<in> V" |
9035 | 68 |
by (simp add: diff_eq1 negate_eq1) |
7917 | 69 |
|
13515 | 70 |
lemma (in vectorspace) neg_closed [iff]: "x \<in> V \<Longrightarrow> - x \<in> V" |
9035 | 71 |
by (simp add: negate_eq1) |
7917 | 72 |
|
13515 | 73 |
lemma (in vectorspace) add_left_commute: |
74 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> x + (y + z) = y + (x + z)" |
|
9035 | 75 |
proof - |
13515 | 76 |
assume xyz: "x \<in> V" "y \<in> V" "z \<in> V" |
27612 | 77 |
then have "x + (y + z) = (x + y) + z" |
13515 | 78 |
by (simp only: add_assoc) |
27612 | 79 |
also from xyz have "\<dots> = (y + x) + z" by (simp only: add_commute) |
80 |
also from xyz have "\<dots> = y + (x + z)" by (simp only: add_assoc) |
|
9035 | 81 |
finally show ?thesis . |
82 |
qed |
|
7917 | 83 |
|
13515 | 84 |
theorems (in vectorspace) add_ac = |
85 |
add_assoc add_commute add_left_commute |
|
7917 | 86 |
|
87 |
||
7978 | 88 |
text {* The existence of the zero element of a vector space |
13515 | 89 |
follows from the non-emptiness of carrier set. *} |
7917 | 90 |
|
13515 | 91 |
lemma (in vectorspace) zero [iff]: "0 \<in> V" |
10687 | 92 |
proof - |
13515 | 93 |
from non_empty obtain x where x: "x \<in> V" by blast |
94 |
then have "0 = x - x" by (rule diff_self [symmetric]) |
|
27612 | 95 |
also from x x have "\<dots> \<in> V" by (rule diff_closed) |
11472 | 96 |
finally show ?thesis . |
9035 | 97 |
qed |
7917 | 98 |
|
13515 | 99 |
lemma (in vectorspace) add_zero_right [simp]: |
100 |
"x \<in> V \<Longrightarrow> x + 0 = x" |
|
9035 | 101 |
proof - |
13515 | 102 |
assume x: "x \<in> V" |
103 |
from this and zero have "x + 0 = 0 + x" by (rule add_commute) |
|
27612 | 104 |
also from x have "\<dots> = x" by (rule add_zero_left) |
9035 | 105 |
finally show ?thesis . |
106 |
qed |
|
7917 | 107 |
|
13515 | 108 |
lemma (in vectorspace) mult_assoc2: |
109 |
"x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = (a * b) \<cdot> x" |
|
110 |
by (simp only: mult_assoc) |
|
7917 | 111 |
|
13515 | 112 |
lemma (in vectorspace) diff_mult_distrib1: |
113 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x - y) = a \<cdot> x - a \<cdot> y" |
|
114 |
by (simp add: diff_eq1 negate_eq1 add_mult_distrib1 mult_assoc2) |
|
7917 | 115 |
|
13515 | 116 |
lemma (in vectorspace) diff_mult_distrib2: |
117 |
"x \<in> V \<Longrightarrow> (a - b) \<cdot> x = a \<cdot> x - (b \<cdot> x)" |
|
9035 | 118 |
proof - |
13515 | 119 |
assume x: "x \<in> V" |
10687 | 120 |
have " (a - b) \<cdot> x = (a + - b) \<cdot> x" |
37887 | 121 |
by (simp add: diff_minus) |
27612 | 122 |
also from x have "\<dots> = a \<cdot> x + (- b) \<cdot> x" |
13515 | 123 |
by (rule add_mult_distrib2) |
27612 | 124 |
also from x have "\<dots> = a \<cdot> x + - (b \<cdot> x)" |
13515 | 125 |
by (simp add: negate_eq1 mult_assoc2) |
27612 | 126 |
also from x have "\<dots> = a \<cdot> x - (b \<cdot> x)" |
13515 | 127 |
by (simp add: diff_eq1) |
9035 | 128 |
finally show ?thesis . |
129 |
qed |
|
7917 | 130 |
|
13515 | 131 |
lemmas (in vectorspace) distrib = |
132 |
add_mult_distrib1 add_mult_distrib2 |
|
133 |
diff_mult_distrib1 diff_mult_distrib2 |
|
134 |
||
10687 | 135 |
|
136 |
text {* \medskip Further derived laws: *} |
|
7917 | 137 |
|
13515 | 138 |
lemma (in vectorspace) mult_zero_left [simp]: |
139 |
"x \<in> V \<Longrightarrow> 0 \<cdot> x = 0" |
|
9035 | 140 |
proof - |
13515 | 141 |
assume x: "x \<in> V" |
142 |
have "0 \<cdot> x = (1 - 1) \<cdot> x" by simp |
|
27612 | 143 |
also have "\<dots> = (1 + - 1) \<cdot> x" by simp |
144 |
also from x have "\<dots> = 1 \<cdot> x + (- 1) \<cdot> x" |
|
13515 | 145 |
by (rule add_mult_distrib2) |
27612 | 146 |
also from x have "\<dots> = x + (- 1) \<cdot> x" by simp |
147 |
also from x have "\<dots> = x + - x" by (simp add: negate_eq2a) |
|
148 |
also from x have "\<dots> = x - x" by (simp add: diff_eq2) |
|
149 |
also from x have "\<dots> = 0" by simp |
|
9035 | 150 |
finally show ?thesis . |
151 |
qed |
|
7917 | 152 |
|
13515 | 153 |
lemma (in vectorspace) mult_zero_right [simp]: |
154 |
"a \<cdot> 0 = (0::'a)" |
|
9035 | 155 |
proof - |
13515 | 156 |
have "a \<cdot> 0 = a \<cdot> (0 - (0::'a))" by simp |
27612 | 157 |
also have "\<dots> = a \<cdot> 0 - a \<cdot> 0" |
13515 | 158 |
by (rule diff_mult_distrib1) simp_all |
27612 | 159 |
also have "\<dots> = 0" by simp |
9035 | 160 |
finally show ?thesis . |
161 |
qed |
|
7917 | 162 |
|
13515 | 163 |
lemma (in vectorspace) minus_mult_cancel [simp]: |
164 |
"x \<in> V \<Longrightarrow> (- a) \<cdot> - x = a \<cdot> x" |
|
165 |
by (simp add: negate_eq1 mult_assoc2) |
|
7917 | 166 |
|
13515 | 167 |
lemma (in vectorspace) add_minus_left_eq_diff: |
168 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + y = y - x" |
|
10687 | 169 |
proof - |
13515 | 170 |
assume xy: "x \<in> V" "y \<in> V" |
27612 | 171 |
then have "- x + y = y + - x" by (simp add: add_commute) |
172 |
also from xy have "\<dots> = y - x" by (simp add: diff_eq1) |
|
9035 | 173 |
finally show ?thesis . |
174 |
qed |
|
7917 | 175 |
|
13515 | 176 |
lemma (in vectorspace) add_minus [simp]: |
177 |
"x \<in> V \<Longrightarrow> x + - x = 0" |
|
178 |
by (simp add: diff_eq2) |
|
7917 | 179 |
|
13515 | 180 |
lemma (in vectorspace) add_minus_left [simp]: |
181 |
"x \<in> V \<Longrightarrow> - x + x = 0" |
|
182 |
by (simp add: diff_eq2 add_commute) |
|
7917 | 183 |
|
13515 | 184 |
lemma (in vectorspace) minus_minus [simp]: |
185 |
"x \<in> V \<Longrightarrow> - (- x) = x" |
|
186 |
by (simp add: negate_eq1 mult_assoc2) |
|
187 |
||
188 |
lemma (in vectorspace) minus_zero [simp]: |
|
189 |
"- (0::'a) = 0" |
|
9035 | 190 |
by (simp add: negate_eq1) |
7917 | 191 |
|
13515 | 192 |
lemma (in vectorspace) minus_zero_iff [simp]: |
193 |
"x \<in> V \<Longrightarrow> (- x = 0) = (x = 0)" |
|
194 |
proof |
|
195 |
assume x: "x \<in> V" |
|
196 |
{ |
|
197 |
from x have "x = - (- x)" by (simp add: minus_minus) |
|
198 |
also assume "- x = 0" |
|
27612 | 199 |
also have "- \<dots> = 0" by (rule minus_zero) |
13515 | 200 |
finally show "x = 0" . |
201 |
next |
|
202 |
assume "x = 0" |
|
203 |
then show "- x = 0" by simp |
|
204 |
} |
|
9035 | 205 |
qed |
7917 | 206 |
|
13515 | 207 |
lemma (in vectorspace) add_minus_cancel [simp]: |
208 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + (- x + y) = y" |
|
209 |
by (simp add: add_assoc [symmetric] del: add_commute) |
|
7917 | 210 |
|
13515 | 211 |
lemma (in vectorspace) minus_add_cancel [simp]: |
212 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + (x + y) = y" |
|
213 |
by (simp add: add_assoc [symmetric] del: add_commute) |
|
7917 | 214 |
|
13515 | 215 |
lemma (in vectorspace) minus_add_distrib [simp]: |
216 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - (x + y) = - x + - y" |
|
217 |
by (simp add: negate_eq1 add_mult_distrib1) |
|
7917 | 218 |
|
13515 | 219 |
lemma (in vectorspace) diff_zero [simp]: |
220 |
"x \<in> V \<Longrightarrow> x - 0 = x" |
|
221 |
by (simp add: diff_eq1) |
|
222 |
||
223 |
lemma (in vectorspace) diff_zero_right [simp]: |
|
224 |
"x \<in> V \<Longrightarrow> 0 - x = - x" |
|
10687 | 225 |
by (simp add: diff_eq1) |
7917 | 226 |
|
13515 | 227 |
lemma (in vectorspace) add_left_cancel: |
228 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + y = x + z) = (y = z)" |
|
9035 | 229 |
proof |
13515 | 230 |
assume x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V" |
231 |
{ |
|
232 |
from y have "y = 0 + y" by simp |
|
27612 | 233 |
also from x y have "\<dots> = (- x + x) + y" by simp |
234 |
also from x y have "\<dots> = - x + (x + y)" |
|
13515 | 235 |
by (simp add: add_assoc neg_closed) |
236 |
also assume "x + y = x + z" |
|
237 |
also from x z have "- x + (x + z) = - x + x + z" |
|
238 |
by (simp add: add_assoc [symmetric] neg_closed) |
|
27612 | 239 |
also from x z have "\<dots> = z" by simp |
13515 | 240 |
finally show "y = z" . |
241 |
next |
|
242 |
assume "y = z" |
|
243 |
then show "x + y = x + z" by (simp only:) |
|
244 |
} |
|
245 |
qed |
|
7917 | 246 |
|
13515 | 247 |
lemma (in vectorspace) add_right_cancel: |
248 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (y + x = z + x) = (y = z)" |
|
249 |
by (simp only: add_commute add_left_cancel) |
|
7917 | 250 |
|
13515 | 251 |
lemma (in vectorspace) add_assoc_cong: |
252 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x' \<in> V \<Longrightarrow> y' \<in> V \<Longrightarrow> z \<in> V |
|
253 |
\<Longrightarrow> x + y = x' + y' \<Longrightarrow> x + (y + z) = x' + (y' + z)" |
|
254 |
by (simp only: add_assoc [symmetric]) |
|
7917 | 255 |
|
13515 | 256 |
lemma (in vectorspace) mult_left_commute: |
257 |
"x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = b \<cdot> a \<cdot> x" |
|
36778
739a9379e29b
avoid using real-specific versions of generic lemmas
huffman
parents:
31795
diff
changeset
|
258 |
by (simp add: mult_commute mult_assoc2) |
7917 | 259 |
|
13515 | 260 |
lemma (in vectorspace) mult_zero_uniq: |
261 |
"x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> a \<cdot> x = 0 \<Longrightarrow> a = 0" |
|
9035 | 262 |
proof (rule classical) |
13515 | 263 |
assume a: "a \<noteq> 0" |
264 |
assume x: "x \<in> V" "x \<noteq> 0" and ax: "a \<cdot> x = 0" |
|
265 |
from x a have "x = (inverse a * a) \<cdot> x" by simp |
|
27612 | 266 |
also from `x \<in> V` have "\<dots> = inverse a \<cdot> (a \<cdot> x)" by (rule mult_assoc) |
267 |
also from ax have "\<dots> = inverse a \<cdot> 0" by simp |
|
268 |
also have "\<dots> = 0" by simp |
|
9374 | 269 |
finally have "x = 0" . |
23378 | 270 |
with `x \<noteq> 0` show "a = 0" by contradiction |
9035 | 271 |
qed |
7917 | 272 |
|
13515 | 273 |
lemma (in vectorspace) mult_left_cancel: |
274 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (a \<cdot> x = a \<cdot> y) = (x = y)" |
|
9035 | 275 |
proof |
13515 | 276 |
assume x: "x \<in> V" and y: "y \<in> V" and a: "a \<noteq> 0" |
277 |
from x have "x = 1 \<cdot> x" by simp |
|
27612 | 278 |
also from a have "\<dots> = (inverse a * a) \<cdot> x" by simp |
279 |
also from x have "\<dots> = inverse a \<cdot> (a \<cdot> x)" |
|
13515 | 280 |
by (simp only: mult_assoc) |
281 |
also assume "a \<cdot> x = a \<cdot> y" |
|
27612 | 282 |
also from a y have "inverse a \<cdot> \<dots> = y" |
13515 | 283 |
by (simp add: mult_assoc2) |
284 |
finally show "x = y" . |
|
285 |
next |
|
286 |
assume "x = y" |
|
287 |
then show "a \<cdot> x = a \<cdot> y" by (simp only:) |
|
288 |
qed |
|
7917 | 289 |
|
13515 | 290 |
lemma (in vectorspace) mult_right_cancel: |
291 |
"x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> (a \<cdot> x = b \<cdot> x) = (a = b)" |
|
9035 | 292 |
proof |
13515 | 293 |
assume x: "x \<in> V" and neq: "x \<noteq> 0" |
294 |
{ |
|
295 |
from x have "(a - b) \<cdot> x = a \<cdot> x - b \<cdot> x" |
|
296 |
by (simp add: diff_mult_distrib2) |
|
297 |
also assume "a \<cdot> x = b \<cdot> x" |
|
298 |
with x have "a \<cdot> x - b \<cdot> x = 0" by simp |
|
299 |
finally have "(a - b) \<cdot> x = 0" . |
|
300 |
with x neq have "a - b = 0" by (rule mult_zero_uniq) |
|
27612 | 301 |
then show "a = b" by simp |
13515 | 302 |
next |
303 |
assume "a = b" |
|
304 |
then show "a \<cdot> x = b \<cdot> x" by (simp only:) |
|
305 |
} |
|
306 |
qed |
|
7917 | 307 |
|
13515 | 308 |
lemma (in vectorspace) eq_diff_eq: |
309 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x = z - y) = (x + y = z)" |
|
310 |
proof |
|
311 |
assume x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V" |
|
312 |
{ |
|
313 |
assume "x = z - y" |
|
27612 | 314 |
then have "x + y = z - y + y" by simp |
315 |
also from y z have "\<dots> = z + - y + y" |
|
13515 | 316 |
by (simp add: diff_eq1) |
27612 | 317 |
also have "\<dots> = z + (- y + y)" |
13515 | 318 |
by (rule add_assoc) (simp_all add: y z) |
27612 | 319 |
also from y z have "\<dots> = z + 0" |
13515 | 320 |
by (simp only: add_minus_left) |
27612 | 321 |
also from z have "\<dots> = z" |
13515 | 322 |
by (simp only: add_zero_right) |
323 |
finally show "x + y = z" . |
|
9035 | 324 |
next |
13515 | 325 |
assume "x + y = z" |
27612 | 326 |
then have "z - y = (x + y) - y" by simp |
327 |
also from x y have "\<dots> = x + y + - y" |
|
9035 | 328 |
by (simp add: diff_eq1) |
27612 | 329 |
also have "\<dots> = x + (y + - y)" |
13515 | 330 |
by (rule add_assoc) (simp_all add: x y) |
27612 | 331 |
also from x y have "\<dots> = x" by simp |
13515 | 332 |
finally show "x = z - y" .. |
333 |
} |
|
9035 | 334 |
qed |
7917 | 335 |
|
13515 | 336 |
lemma (in vectorspace) add_minus_eq_minus: |
337 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y = 0 \<Longrightarrow> x = - y" |
|
9035 | 338 |
proof - |
13515 | 339 |
assume x: "x \<in> V" and y: "y \<in> V" |
340 |
from x y have "x = (- y + y) + x" by simp |
|
27612 | 341 |
also from x y have "\<dots> = - y + (x + y)" by (simp add: add_ac) |
9374 | 342 |
also assume "x + y = 0" |
13515 | 343 |
also from y have "- y + 0 = - y" by simp |
9035 | 344 |
finally show "x = - y" . |
345 |
qed |
|
7917 | 346 |
|
13515 | 347 |
lemma (in vectorspace) add_minus_eq: |
348 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y = 0 \<Longrightarrow> x = y" |
|
9035 | 349 |
proof - |
13515 | 350 |
assume x: "x \<in> V" and y: "y \<in> V" |
9374 | 351 |
assume "x - y = 0" |
13515 | 352 |
with x y have eq: "x + - y = 0" by (simp add: diff_eq1) |
353 |
with _ _ have "x = - (- y)" |
|
354 |
by (rule add_minus_eq_minus) (simp_all add: x y) |
|
355 |
with x y show "x = y" by simp |
|
9035 | 356 |
qed |
7917 | 357 |
|
13515 | 358 |
lemma (in vectorspace) add_diff_swap: |
359 |
"a \<in> V \<Longrightarrow> b \<in> V \<Longrightarrow> c \<in> V \<Longrightarrow> d \<in> V \<Longrightarrow> a + b = c + d |
|
360 |
\<Longrightarrow> a - c = d - b" |
|
10687 | 361 |
proof - |
13515 | 362 |
assume vs: "a \<in> V" "b \<in> V" "c \<in> V" "d \<in> V" |
9035 | 363 |
and eq: "a + b = c + d" |
13515 | 364 |
then have "- c + (a + b) = - c + (c + d)" |
365 |
by (simp add: add_left_cancel) |
|
27612 | 366 |
also have "\<dots> = d" using `c \<in> V` `d \<in> V` by (rule minus_add_cancel) |
9035 | 367 |
finally have eq: "- c + (a + b) = d" . |
10687 | 368 |
from vs have "a - c = (- c + (a + b)) + - b" |
13515 | 369 |
by (simp add: add_ac diff_eq1) |
27612 | 370 |
also from vs eq have "\<dots> = d + - b" |
13515 | 371 |
by (simp add: add_right_cancel) |
27612 | 372 |
also from vs have "\<dots> = d - b" by (simp add: diff_eq2) |
9035 | 373 |
finally show "a - c = d - b" . |
374 |
qed |
|
7917 | 375 |
|
13515 | 376 |
lemma (in vectorspace) vs_add_cancel_21: |
377 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> u \<in> V |
|
378 |
\<Longrightarrow> (x + (y + z) = y + u) = (x + z = u)" |
|
379 |
proof |
|
380 |
assume vs: "x \<in> V" "y \<in> V" "z \<in> V" "u \<in> V" |
|
381 |
{ |
|
382 |
from vs have "x + z = - y + y + (x + z)" by simp |
|
27612 | 383 |
also have "\<dots> = - y + (y + (x + z))" |
13515 | 384 |
by (rule add_assoc) (simp_all add: vs) |
385 |
also from vs have "y + (x + z) = x + (y + z)" |
|
386 |
by (simp add: add_ac) |
|
387 |
also assume "x + (y + z) = y + u" |
|
388 |
also from vs have "- y + (y + u) = u" by simp |
|
389 |
finally show "x + z = u" . |
|
390 |
next |
|
391 |
assume "x + z = u" |
|
392 |
with vs show "x + (y + z) = y + u" |
|
393 |
by (simp only: add_left_commute [of x]) |
|
394 |
} |
|
9035 | 395 |
qed |
7917 | 396 |
|
13515 | 397 |
lemma (in vectorspace) add_cancel_end: |
398 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + (y + z) = y) = (x = - z)" |
|
399 |
proof |
|
400 |
assume vs: "x \<in> V" "y \<in> V" "z \<in> V" |
|
401 |
{ |
|
402 |
assume "x + (y + z) = y" |
|
403 |
with vs have "(x + z) + y = 0 + y" |
|
404 |
by (simp add: add_ac) |
|
405 |
with vs have "x + z = 0" |
|
406 |
by (simp only: add_right_cancel add_closed zero) |
|
407 |
with vs show "x = - z" by (simp add: add_minus_eq_minus) |
|
9035 | 408 |
next |
13515 | 409 |
assume eq: "x = - z" |
27612 | 410 |
then have "x + (y + z) = - z + (y + z)" by simp |
411 |
also have "\<dots> = y + (- z + z)" |
|
13515 | 412 |
by (rule add_left_commute) (simp_all add: vs) |
27612 | 413 |
also from vs have "\<dots> = y" by simp |
13515 | 414 |
finally show "x + (y + z) = y" . |
415 |
} |
|
9035 | 416 |
qed |
7917 | 417 |
|
10687 | 418 |
end |