author | haftmann |
Wed, 07 Nov 2018 11:08:10 +0000 | |
changeset 69250 | 1011f0b46af7 |
parent 61879 | e4f9d8f094fe |
child 69597 | ff784d5a5bfb |
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 |
||
58889 | 5 |
section \<open>Vector spaces\<close> |
7917 | 6 |
|
31795
be3e1cc5005c
standard naming conventions for session and theories;
wenzelm
parents:
29252
diff
changeset
|
7 |
theory Vector_Space |
55018
2a526bd279ed
moved 'Zorn' into 'Main', since it's a BNF dependency
blanchet
parents:
54230
diff
changeset
|
8 |
imports Complex_Main Bounds |
27612 | 9 |
begin |
7917 | 10 |
|
58744 | 11 |
subsection \<open>Signature\<close> |
7917 | 12 |
|
58744 | 13 |
text \<open> |
61540 | 14 |
For the definition of real vector spaces a type @{typ 'a} of the sort |
15 |
\<open>{plus, minus, zero}\<close> is considered, on which a real scalar multiplication |
|
16 |
\<open>\<cdot>\<close> is declared. |
|
58744 | 17 |
\<close> |
7917 | 18 |
|
19 |
consts |
|
58745 | 20 |
prod :: "real \<Rightarrow> 'a::{plus,minus,zero} \<Rightarrow> 'a" (infixr "\<cdot>" 70) |
7917 | 21 |
|
22 |
||
58744 | 23 |
subsection \<open>Vector space laws\<close> |
7917 | 24 |
|
58744 | 25 |
text \<open> |
61879 | 26 |
A \<^emph>\<open>vector space\<close> is a non-empty set \<open>V\<close> of elements from @{typ 'a} with the |
27 |
following vector space laws: The set \<open>V\<close> is closed under addition and scalar |
|
28 |
multiplication, addition is associative and commutative; \<open>- x\<close> is the |
|
29 |
inverse of \<open>x\<close> wrt.\ addition and \<open>0\<close> is the neutral element of addition. |
|
30 |
Addition and multiplication are distributive; scalar multiplication is |
|
31 |
associative and the real number \<open>1\<close> is the neutral element of scalar |
|
32 |
multiplication. |
|
58744 | 33 |
\<close> |
7917 | 34 |
|
46867
0883804b67bb
modernized locale expression, with some fluctuation of arguments;
wenzelm
parents:
44887
diff
changeset
|
35 |
locale vectorspace = |
0883804b67bb
modernized locale expression, with some fluctuation of arguments;
wenzelm
parents:
44887
diff
changeset
|
36 |
fixes V |
13515 | 37 |
assumes non_empty [iff, intro?]: "V \<noteq> {}" |
38 |
and add_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y \<in> V" |
|
39 |
and mult_closed [iff]: "x \<in> V \<Longrightarrow> a \<cdot> x \<in> V" |
|
40 |
and add_assoc: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + y) + z = x + (y + z)" |
|
41 |
and add_commute: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y = y + x" |
|
42 |
and diff_self [simp]: "x \<in> V \<Longrightarrow> x - x = 0" |
|
43 |
and add_zero_left [simp]: "x \<in> V \<Longrightarrow> 0 + x = x" |
|
44 |
and add_mult_distrib1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y" |
|
45 |
and add_mult_distrib2: "x \<in> V \<Longrightarrow> (a + b) \<cdot> x = a \<cdot> x + b \<cdot> x" |
|
46 |
and mult_assoc: "x \<in> V \<Longrightarrow> (a * b) \<cdot> x = a \<cdot> (b \<cdot> x)" |
|
47 |
and mult_1 [simp]: "x \<in> V \<Longrightarrow> 1 \<cdot> x = x" |
|
48 |
and negate_eq1: "x \<in> V \<Longrightarrow> - x = (- 1) \<cdot> x" |
|
49 |
and diff_eq1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y = x + - y" |
|
44887 | 50 |
begin |
7917 | 51 |
|
44887 | 52 |
lemma negate_eq2: "x \<in> V \<Longrightarrow> (- 1) \<cdot> x = - x" |
13515 | 53 |
by (rule negate_eq1 [symmetric]) |
9013
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents:
8703
diff
changeset
|
54 |
|
44887 | 55 |
lemma negate_eq2a: "x \<in> V \<Longrightarrow> -1 \<cdot> x = - x" |
13515 | 56 |
by (simp add: negate_eq1) |
7917 | 57 |
|
44887 | 58 |
lemma diff_eq2: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + - y = x - y" |
13515 | 59 |
by (rule diff_eq1 [symmetric]) |
7917 | 60 |
|
44887 | 61 |
lemma diff_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y \<in> V" |
9035 | 62 |
by (simp add: diff_eq1 negate_eq1) |
7917 | 63 |
|
44887 | 64 |
lemma neg_closed [iff]: "x \<in> V \<Longrightarrow> - x \<in> V" |
9035 | 65 |
by (simp add: negate_eq1) |
7917 | 66 |
|
61540 | 67 |
lemma add_left_commute: |
68 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> x + (y + z) = y + (x + z)" |
|
9035 | 69 |
proof - |
13515 | 70 |
assume xyz: "x \<in> V" "y \<in> V" "z \<in> V" |
27612 | 71 |
then have "x + (y + z) = (x + y) + z" |
13515 | 72 |
by (simp only: add_assoc) |
27612 | 73 |
also from xyz have "\<dots> = (y + x) + z" by (simp only: add_commute) |
74 |
also from xyz have "\<dots> = y + (x + z)" by (simp only: add_assoc) |
|
9035 | 75 |
finally show ?thesis . |
76 |
qed |
|
7917 | 77 |
|
61337 | 78 |
lemmas add_ac = add_assoc add_commute add_left_commute |
7917 | 79 |
|
80 |
||
61879 | 81 |
text \<open> |
82 |
The existence of the zero element of a vector space follows from the |
|
83 |
non-emptiness of carrier set. |
|
84 |
\<close> |
|
7917 | 85 |
|
44887 | 86 |
lemma zero [iff]: "0 \<in> V" |
10687 | 87 |
proof - |
13515 | 88 |
from non_empty obtain x where x: "x \<in> V" by blast |
89 |
then have "0 = x - x" by (rule diff_self [symmetric]) |
|
27612 | 90 |
also from x x have "\<dots> \<in> V" by (rule diff_closed) |
11472 | 91 |
finally show ?thesis . |
9035 | 92 |
qed |
7917 | 93 |
|
44887 | 94 |
lemma add_zero_right [simp]: "x \<in> V \<Longrightarrow> x + 0 = x" |
9035 | 95 |
proof - |
13515 | 96 |
assume x: "x \<in> V" |
97 |
from this and zero have "x + 0 = 0 + x" by (rule add_commute) |
|
27612 | 98 |
also from x have "\<dots> = x" by (rule add_zero_left) |
9035 | 99 |
finally show ?thesis . |
100 |
qed |
|
7917 | 101 |
|
44887 | 102 |
lemma mult_assoc2: "x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = (a * b) \<cdot> x" |
13515 | 103 |
by (simp only: mult_assoc) |
7917 | 104 |
|
44887 | 105 |
lemma diff_mult_distrib1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x - y) = a \<cdot> x - a \<cdot> y" |
13515 | 106 |
by (simp add: diff_eq1 negate_eq1 add_mult_distrib1 mult_assoc2) |
7917 | 107 |
|
44887 | 108 |
lemma diff_mult_distrib2: "x \<in> V \<Longrightarrow> (a - b) \<cdot> x = a \<cdot> x - (b \<cdot> x)" |
9035 | 109 |
proof - |
13515 | 110 |
assume x: "x \<in> V" |
10687 | 111 |
have " (a - b) \<cdot> x = (a + - b) \<cdot> x" |
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
46867
diff
changeset
|
112 |
by simp |
27612 | 113 |
also from x have "\<dots> = a \<cdot> x + (- b) \<cdot> x" |
13515 | 114 |
by (rule add_mult_distrib2) |
27612 | 115 |
also from x have "\<dots> = a \<cdot> x + - (b \<cdot> x)" |
13515 | 116 |
by (simp add: negate_eq1 mult_assoc2) |
27612 | 117 |
also from x have "\<dots> = a \<cdot> x - (b \<cdot> x)" |
13515 | 118 |
by (simp add: diff_eq1) |
9035 | 119 |
finally show ?thesis . |
120 |
qed |
|
7917 | 121 |
|
44887 | 122 |
lemmas distrib = |
13515 | 123 |
add_mult_distrib1 add_mult_distrib2 |
124 |
diff_mult_distrib1 diff_mult_distrib2 |
|
125 |
||
10687 | 126 |
|
61486 | 127 |
text \<open>\<^medskip> Further derived laws:\<close> |
7917 | 128 |
|
44887 | 129 |
lemma mult_zero_left [simp]: "x \<in> V \<Longrightarrow> 0 \<cdot> x = 0" |
9035 | 130 |
proof - |
13515 | 131 |
assume x: "x \<in> V" |
132 |
have "0 \<cdot> x = (1 - 1) \<cdot> x" by simp |
|
27612 | 133 |
also have "\<dots> = (1 + - 1) \<cdot> x" by simp |
134 |
also from x have "\<dots> = 1 \<cdot> x + (- 1) \<cdot> x" |
|
13515 | 135 |
by (rule add_mult_distrib2) |
27612 | 136 |
also from x have "\<dots> = x + (- 1) \<cdot> x" by simp |
137 |
also from x have "\<dots> = x + - x" by (simp add: negate_eq2a) |
|
138 |
also from x have "\<dots> = x - x" by (simp add: diff_eq2) |
|
139 |
also from x have "\<dots> = 0" by simp |
|
9035 | 140 |
finally show ?thesis . |
141 |
qed |
|
7917 | 142 |
|
44887 | 143 |
lemma mult_zero_right [simp]: "a \<cdot> 0 = (0::'a)" |
9035 | 144 |
proof - |
13515 | 145 |
have "a \<cdot> 0 = a \<cdot> (0 - (0::'a))" by simp |
27612 | 146 |
also have "\<dots> = a \<cdot> 0 - a \<cdot> 0" |
13515 | 147 |
by (rule diff_mult_distrib1) simp_all |
27612 | 148 |
also have "\<dots> = 0" by simp |
9035 | 149 |
finally show ?thesis . |
150 |
qed |
|
7917 | 151 |
|
44887 | 152 |
lemma minus_mult_cancel [simp]: "x \<in> V \<Longrightarrow> (- a) \<cdot> - x = a \<cdot> x" |
13515 | 153 |
by (simp add: negate_eq1 mult_assoc2) |
7917 | 154 |
|
44887 | 155 |
lemma add_minus_left_eq_diff: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + y = y - x" |
10687 | 156 |
proof - |
13515 | 157 |
assume xy: "x \<in> V" "y \<in> V" |
27612 | 158 |
then have "- x + y = y + - x" by (simp add: add_commute) |
159 |
also from xy have "\<dots> = y - x" by (simp add: diff_eq1) |
|
9035 | 160 |
finally show ?thesis . |
161 |
qed |
|
7917 | 162 |
|
44887 | 163 |
lemma add_minus [simp]: "x \<in> V \<Longrightarrow> x + - x = 0" |
13515 | 164 |
by (simp add: diff_eq2) |
7917 | 165 |
|
44887 | 166 |
lemma add_minus_left [simp]: "x \<in> V \<Longrightarrow> - x + x = 0" |
13515 | 167 |
by (simp add: diff_eq2 add_commute) |
7917 | 168 |
|
44887 | 169 |
lemma minus_minus [simp]: "x \<in> V \<Longrightarrow> - (- x) = x" |
13515 | 170 |
by (simp add: negate_eq1 mult_assoc2) |
171 |
||
44887 | 172 |
lemma minus_zero [simp]: "- (0::'a) = 0" |
9035 | 173 |
by (simp add: negate_eq1) |
7917 | 174 |
|
44887 | 175 |
lemma minus_zero_iff [simp]: |
176 |
assumes x: "x \<in> V" |
|
177 |
shows "(- x = 0) = (x = 0)" |
|
13515 | 178 |
proof |
44887 | 179 |
from x have "x = - (- x)" by simp |
180 |
also assume "- x = 0" |
|
181 |
also have "- \<dots> = 0" by (rule minus_zero) |
|
182 |
finally show "x = 0" . |
|
183 |
next |
|
184 |
assume "x = 0" |
|
185 |
then show "- x = 0" by simp |
|
9035 | 186 |
qed |
7917 | 187 |
|
44887 | 188 |
lemma add_minus_cancel [simp]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + (- x + y) = y" |
189 |
by (simp add: add_assoc [symmetric]) |
|
7917 | 190 |
|
44887 | 191 |
lemma minus_add_cancel [simp]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + (x + y) = y" |
192 |
by (simp add: add_assoc [symmetric]) |
|
7917 | 193 |
|
44887 | 194 |
lemma minus_add_distrib [simp]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - (x + y) = - x + - y" |
13515 | 195 |
by (simp add: negate_eq1 add_mult_distrib1) |
7917 | 196 |
|
44887 | 197 |
lemma diff_zero [simp]: "x \<in> V \<Longrightarrow> x - 0 = x" |
13515 | 198 |
by (simp add: diff_eq1) |
199 |
||
44887 | 200 |
lemma diff_zero_right [simp]: "x \<in> V \<Longrightarrow> 0 - x = - x" |
10687 | 201 |
by (simp add: diff_eq1) |
7917 | 202 |
|
44887 | 203 |
lemma add_left_cancel: |
204 |
assumes x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V" |
|
205 |
shows "(x + y = x + z) = (y = z)" |
|
9035 | 206 |
proof |
44887 | 207 |
from y have "y = 0 + y" by simp |
208 |
also from x y have "\<dots> = (- x + x) + y" by simp |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
55018
diff
changeset
|
209 |
also from x y have "\<dots> = - x + (x + y)" by (simp add: add.assoc) |
44887 | 210 |
also assume "x + y = x + z" |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
55018
diff
changeset
|
211 |
also from x z have "- x + (x + z) = - x + x + z" by (simp add: add.assoc) |
44887 | 212 |
also from x z have "\<dots> = z" by simp |
213 |
finally show "y = z" . |
|
214 |
next |
|
215 |
assume "y = z" |
|
216 |
then show "x + y = x + z" by (simp only:) |
|
13515 | 217 |
qed |
7917 | 218 |
|
61540 | 219 |
lemma add_right_cancel: |
220 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (y + x = z + x) = (y = z)" |
|
13515 | 221 |
by (simp only: add_commute add_left_cancel) |
7917 | 222 |
|
44887 | 223 |
lemma add_assoc_cong: |
13515 | 224 |
"x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x' \<in> V \<Longrightarrow> y' \<in> V \<Longrightarrow> z \<in> V |
225 |
\<Longrightarrow> x + y = x' + y' \<Longrightarrow> x + (y + z) = x' + (y' + z)" |
|
226 |
by (simp only: add_assoc [symmetric]) |
|
7917 | 227 |
|
44887 | 228 |
lemma mult_left_commute: "x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = b \<cdot> a \<cdot> x" |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
55018
diff
changeset
|
229 |
by (simp add: mult.commute mult_assoc2) |
7917 | 230 |
|
44887 | 231 |
lemma mult_zero_uniq: |
232 |
assumes x: "x \<in> V" "x \<noteq> 0" and ax: "a \<cdot> x = 0" |
|
233 |
shows "a = 0" |
|
9035 | 234 |
proof (rule classical) |
13515 | 235 |
assume a: "a \<noteq> 0" |
236 |
from x a have "x = (inverse a * a) \<cdot> x" by simp |
|
58744 | 237 |
also from \<open>x \<in> V\<close> have "\<dots> = inverse a \<cdot> (a \<cdot> x)" by (rule mult_assoc) |
27612 | 238 |
also from ax have "\<dots> = inverse a \<cdot> 0" by simp |
239 |
also have "\<dots> = 0" by simp |
|
9374 | 240 |
finally have "x = 0" . |
58744 | 241 |
with \<open>x \<noteq> 0\<close> show "a = 0" by contradiction |
9035 | 242 |
qed |
7917 | 243 |
|
44887 | 244 |
lemma mult_left_cancel: |
245 |
assumes x: "x \<in> V" and y: "y \<in> V" and a: "a \<noteq> 0" |
|
246 |
shows "(a \<cdot> x = a \<cdot> y) = (x = y)" |
|
9035 | 247 |
proof |
13515 | 248 |
from x have "x = 1 \<cdot> x" by simp |
27612 | 249 |
also from a have "\<dots> = (inverse a * a) \<cdot> x" by simp |
250 |
also from x have "\<dots> = inverse a \<cdot> (a \<cdot> x)" |
|
13515 | 251 |
by (simp only: mult_assoc) |
252 |
also assume "a \<cdot> x = a \<cdot> y" |
|
27612 | 253 |
also from a y have "inverse a \<cdot> \<dots> = y" |
13515 | 254 |
by (simp add: mult_assoc2) |
255 |
finally show "x = y" . |
|
256 |
next |
|
257 |
assume "x = y" |
|
258 |
then show "a \<cdot> x = a \<cdot> y" by (simp only:) |
|
259 |
qed |
|
7917 | 260 |
|
44887 | 261 |
lemma mult_right_cancel: |
262 |
assumes x: "x \<in> V" and neq: "x \<noteq> 0" |
|
263 |
shows "(a \<cdot> x = b \<cdot> x) = (a = b)" |
|
9035 | 264 |
proof |
44887 | 265 |
from x have "(a - b) \<cdot> x = a \<cdot> x - b \<cdot> x" |
266 |
by (simp add: diff_mult_distrib2) |
|
267 |
also assume "a \<cdot> x = b \<cdot> x" |
|
268 |
with x have "a \<cdot> x - b \<cdot> x = 0" by simp |
|
269 |
finally have "(a - b) \<cdot> x = 0" . |
|
270 |
with x neq have "a - b = 0" by (rule mult_zero_uniq) |
|
271 |
then show "a = b" by simp |
|
272 |
next |
|
273 |
assume "a = b" |
|
274 |
then show "a \<cdot> x = b \<cdot> x" by (simp only:) |
|
13515 | 275 |
qed |
7917 | 276 |
|
44887 | 277 |
lemma eq_diff_eq: |
278 |
assumes x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V" |
|
279 |
shows "(x = z - y) = (x + y = z)" |
|
13515 | 280 |
proof |
44887 | 281 |
assume "x = z - y" |
282 |
then have "x + y = z - y + y" by simp |
|
283 |
also from y z have "\<dots> = z + - y + y" |
|
284 |
by (simp add: diff_eq1) |
|
285 |
also have "\<dots> = z + (- y + y)" |
|
286 |
by (rule add_assoc) (simp_all add: y z) |
|
287 |
also from y z have "\<dots> = z + 0" |
|
288 |
by (simp only: add_minus_left) |
|
289 |
also from z have "\<dots> = z" |
|
290 |
by (simp only: add_zero_right) |
|
291 |
finally show "x + y = z" . |
|
292 |
next |
|
293 |
assume "x + y = z" |
|
294 |
then have "z - y = (x + y) - y" by simp |
|
295 |
also from x y have "\<dots> = x + y + - y" |
|
296 |
by (simp add: diff_eq1) |
|
297 |
also have "\<dots> = x + (y + - y)" |
|
298 |
by (rule add_assoc) (simp_all add: x y) |
|
299 |
also from x y have "\<dots> = x" by simp |
|
300 |
finally show "x = z - y" .. |
|
9035 | 301 |
qed |
7917 | 302 |
|
44887 | 303 |
lemma add_minus_eq_minus: |
304 |
assumes x: "x \<in> V" and y: "y \<in> V" and xy: "x + y = 0" |
|
305 |
shows "x = - y" |
|
9035 | 306 |
proof - |
13515 | 307 |
from x y have "x = (- y + y) + x" by simp |
27612 | 308 |
also from x y have "\<dots> = - y + (x + y)" by (simp add: add_ac) |
44887 | 309 |
also note xy |
13515 | 310 |
also from y have "- y + 0 = - y" by simp |
9035 | 311 |
finally show "x = - y" . |
312 |
qed |
|
7917 | 313 |
|
44887 | 314 |
lemma add_minus_eq: |
315 |
assumes x: "x \<in> V" and y: "y \<in> V" and xy: "x - y = 0" |
|
316 |
shows "x = y" |
|
9035 | 317 |
proof - |
44887 | 318 |
from x y xy have eq: "x + - y = 0" by (simp add: diff_eq1) |
13515 | 319 |
with _ _ have "x = - (- y)" |
320 |
by (rule add_minus_eq_minus) (simp_all add: x y) |
|
321 |
with x y show "x = y" by simp |
|
9035 | 322 |
qed |
7917 | 323 |
|
44887 | 324 |
lemma add_diff_swap: |
325 |
assumes vs: "a \<in> V" "b \<in> V" "c \<in> V" "d \<in> V" |
|
326 |
and eq: "a + b = c + d" |
|
327 |
shows "a - c = d - b" |
|
10687 | 328 |
proof - |
44887 | 329 |
from assms have "- c + (a + b) = - c + (c + d)" |
13515 | 330 |
by (simp add: add_left_cancel) |
58744 | 331 |
also have "\<dots> = d" using \<open>c \<in> V\<close> \<open>d \<in> V\<close> by (rule minus_add_cancel) |
9035 | 332 |
finally have eq: "- c + (a + b) = d" . |
10687 | 333 |
from vs have "a - c = (- c + (a + b)) + - b" |
13515 | 334 |
by (simp add: add_ac diff_eq1) |
27612 | 335 |
also from vs eq have "\<dots> = d + - b" |
13515 | 336 |
by (simp add: add_right_cancel) |
27612 | 337 |
also from vs have "\<dots> = d - b" by (simp add: diff_eq2) |
9035 | 338 |
finally show "a - c = d - b" . |
339 |
qed |
|
7917 | 340 |
|
44887 | 341 |
lemma vs_add_cancel_21: |
342 |
assumes vs: "x \<in> V" "y \<in> V" "z \<in> V" "u \<in> V" |
|
343 |
shows "(x + (y + z) = y + u) = (x + z = u)" |
|
13515 | 344 |
proof |
44887 | 345 |
from vs have "x + z = - y + y + (x + z)" by simp |
346 |
also have "\<dots> = - y + (y + (x + z))" |
|
347 |
by (rule add_assoc) (simp_all add: vs) |
|
348 |
also from vs have "y + (x + z) = x + (y + z)" |
|
349 |
by (simp add: add_ac) |
|
350 |
also assume "x + (y + z) = y + u" |
|
351 |
also from vs have "- y + (y + u) = u" by simp |
|
352 |
finally show "x + z = u" . |
|
353 |
next |
|
354 |
assume "x + z = u" |
|
355 |
with vs show "x + (y + z) = y + u" |
|
356 |
by (simp only: add_left_commute [of x]) |
|
9035 | 357 |
qed |
7917 | 358 |
|
44887 | 359 |
lemma add_cancel_end: |
360 |
assumes vs: "x \<in> V" "y \<in> V" "z \<in> V" |
|
361 |
shows "(x + (y + z) = y) = (x = - z)" |
|
13515 | 362 |
proof |
44887 | 363 |
assume "x + (y + z) = y" |
364 |
with vs have "(x + z) + y = 0 + y" by (simp add: add_ac) |
|
365 |
with vs have "x + z = 0" by (simp only: add_right_cancel add_closed zero) |
|
366 |
with vs show "x = - z" by (simp add: add_minus_eq_minus) |
|
367 |
next |
|
368 |
assume eq: "x = - z" |
|
369 |
then have "x + (y + z) = - z + (y + z)" by simp |
|
370 |
also have "\<dots> = y + (- z + z)" by (rule add_left_commute) (simp_all add: vs) |
|
371 |
also from vs have "\<dots> = y" by simp |
|
372 |
finally show "x + (y + z) = y" . |
|
9035 | 373 |
qed |
7917 | 374 |
|
10687 | 375 |
end |
44887 | 376 |
|
377 |
end |