author | paulson |
Fri, 09 Jan 2004 10:46:18 +0100 | |
changeset 14348 | 744c868ee0b7 |
parent 14341 | a09441bd4f1e |
child 14353 | 79f9fbef9106 |
permissions | -rw-r--r-- |
13957 | 1 |
(* Title: Complex.thy |
2 |
Author: Jacques D. Fleuriot |
|
3 |
Copyright: 2001 University of Edinburgh |
|
4 |
Description: Complex numbers |
|
5 |
*) |
|
6 |
||
14323 | 7 |
theory Complex = HLog: |
13957 | 8 |
|
9 |
typedef complex = "{p::(real*real). True}" |
|
14323 | 10 |
by blast |
13957 | 11 |
|
14323 | 12 |
instance complex :: zero .. |
13 |
instance complex :: one .. |
|
14 |
instance complex :: plus .. |
|
15 |
instance complex :: times .. |
|
16 |
instance complex :: minus .. |
|
17 |
instance complex :: inverse .. |
|
18 |
instance complex :: power .. |
|
13957 | 19 |
|
20 |
consts |
|
14323 | 21 |
"ii" :: complex ("ii") |
13957 | 22 |
|
23 |
constdefs |
|
24 |
||
25 |
(*--- real and Imaginary parts ---*) |
|
14323 | 26 |
|
27 |
Re :: "complex => real" |
|
13957 | 28 |
"Re(z) == fst(Rep_complex z)" |
29 |
||
14323 | 30 |
Im :: "complex => real" |
13957 | 31 |
"Im(z) == snd(Rep_complex z)" |
32 |
||
33 |
(*----------- modulus ------------*) |
|
34 |
||
14323 | 35 |
cmod :: "complex => real" |
36 |
"cmod z == sqrt(Re(z) ^ 2 + Im(z) ^ 2)" |
|
13957 | 37 |
|
14323 | 38 |
(*----- injection from reals -----*) |
39 |
||
40 |
complex_of_real :: "real => complex" |
|
13957 | 41 |
"complex_of_real r == Abs_complex(r,0::real)" |
14323 | 42 |
|
13957 | 43 |
(*------- complex conjugate ------*) |
44 |
||
14323 | 45 |
cnj :: "complex => complex" |
13957 | 46 |
"cnj z == Abs_complex(Re z, -Im z)" |
47 |
||
14323 | 48 |
(*------------ Argand -------------*) |
13957 | 49 |
|
14323 | 50 |
sgn :: "complex => complex" |
13957 | 51 |
"sgn z == z / complex_of_real(cmod z)" |
52 |
||
14323 | 53 |
arg :: "complex => real" |
13957 | 54 |
"arg z == @a. Re(sgn z) = cos a & Im(sgn z) = sin a & -pi < a & a <= pi" |
14323 | 55 |
|
13957 | 56 |
|
14323 | 57 |
defs (overloaded) |
58 |
||
59 |
complex_zero_def: |
|
13957 | 60 |
"0 == Abs_complex(0::real,0)" |
61 |
||
14323 | 62 |
complex_one_def: |
13957 | 63 |
"1 == Abs_complex(1,0::real)" |
64 |
||
14323 | 65 |
(*------ imaginary unit ----------*) |
66 |
||
67 |
i_def: |
|
13957 | 68 |
"ii == Abs_complex(0::real,1)" |
69 |
||
70 |
(*----------- negation -----------*) |
|
14323 | 71 |
|
72 |
complex_minus_def: |
|
73 |
"- (z::complex) == Abs_complex(-Re z, -Im z)" |
|
13957 | 74 |
|
14323 | 75 |
|
13957 | 76 |
(*----------- inverse -----------*) |
14323 | 77 |
complex_inverse_def: |
13957 | 78 |
"inverse (z::complex) == Abs_complex(Re(z)/(Re(z) ^ 2 + Im(z) ^ 2), |
79 |
-Im(z)/(Re(z) ^ 2 + Im(z) ^ 2))" |
|
80 |
||
14323 | 81 |
complex_add_def: |
13957 | 82 |
"w + (z::complex) == Abs_complex(Re(w) + Re(z),Im(w) + Im(z))" |
83 |
||
14323 | 84 |
complex_diff_def: |
13957 | 85 |
"w - (z::complex) == w + -(z::complex)" |
86 |
||
14323 | 87 |
complex_mult_def: |
13957 | 88 |
"w * (z::complex) == Abs_complex(Re(w) * Re(z) - Im(w) * Im(z), |
89 |
Re(w) * Im(z) + Im(w) * Re(z))" |
|
90 |
||
91 |
||
92 |
(*----------- division ----------*) |
|
14323 | 93 |
complex_divide_def: |
13957 | 94 |
"w / (z::complex) == w * inverse z" |
14323 | 95 |
|
13957 | 96 |
|
97 |
primrec |
|
14323 | 98 |
complexpow_0: "z ^ 0 = complex_of_real 1" |
99 |
complexpow_Suc: "z ^ (Suc n) = (z::complex) * (z ^ n)" |
|
13957 | 100 |
|
101 |
||
102 |
constdefs |
|
103 |
||
104 |
(* abbreviation for (cos a + i sin a) *) |
|
14323 | 105 |
cis :: "real => complex" |
13957 | 106 |
"cis a == complex_of_real(cos a) + ii * complex_of_real(sin a)" |
107 |
||
108 |
(* abbreviation for r*(cos a + i sin a) *) |
|
14323 | 109 |
rcis :: "[real, real] => complex" |
13957 | 110 |
"rcis r a == complex_of_real r * cis a" |
111 |
||
112 |
(* e ^ (x + iy) *) |
|
14323 | 113 |
expi :: "complex => complex" |
13957 | 114 |
"expi z == complex_of_real(exp (Re z)) * cis (Im z)" |
14323 | 115 |
|
116 |
||
117 |
lemma inj_Rep_complex: "inj Rep_complex" |
|
118 |
apply (rule inj_on_inverseI) |
|
119 |
apply (rule Rep_complex_inverse) |
|
120 |
done |
|
121 |
||
122 |
lemma inj_Abs_complex: "inj Abs_complex" |
|
123 |
apply (rule inj_on_inverseI) |
|
124 |
apply (rule Abs_complex_inverse) |
|
125 |
apply (simp (no_asm) add: complex_def) |
|
126 |
done |
|
127 |
declare inj_Abs_complex [THEN injD, simp] |
|
128 |
||
129 |
lemma Abs_complex_cancel_iff: "(Abs_complex x = Abs_complex y) = (x = y)" |
|
14334 | 130 |
by (auto dest: inj_Abs_complex [THEN injD]) |
14323 | 131 |
declare Abs_complex_cancel_iff [simp] |
132 |
||
133 |
lemma pair_mem_complex: "(x,y) : complex" |
|
14334 | 134 |
by (unfold complex_def, auto) |
14323 | 135 |
declare pair_mem_complex [simp] |
136 |
||
137 |
lemma Abs_complex_inverse2: "Rep_complex (Abs_complex (x,y)) = (x,y)" |
|
138 |
apply (simp (no_asm) add: Abs_complex_inverse) |
|
139 |
done |
|
140 |
declare Abs_complex_inverse2 [simp] |
|
141 |
||
142 |
lemma eq_Abs_complex: "(!!x y. z = Abs_complex(x,y) ==> P) ==> P" |
|
143 |
apply (rule_tac p = "Rep_complex z" in PairE) |
|
14334 | 144 |
apply (drule_tac f = Abs_complex in arg_cong) |
14323 | 145 |
apply (force simp add: Rep_complex_inverse) |
146 |
done |
|
147 |
||
148 |
lemma Re: "Re(Abs_complex(x,y)) = x" |
|
149 |
apply (unfold Re_def) |
|
150 |
apply (simp (no_asm)) |
|
151 |
done |
|
152 |
declare Re [simp] |
|
153 |
||
154 |
lemma Im: "Im(Abs_complex(x,y)) = y" |
|
155 |
apply (unfold Im_def) |
|
156 |
apply (simp (no_asm)) |
|
157 |
done |
|
158 |
declare Im [simp] |
|
159 |
||
160 |
lemma Abs_complex_cancel: "Abs_complex(Re(z),Im(z)) = z" |
|
14334 | 161 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 162 |
apply (simp (no_asm_simp)) |
163 |
done |
|
164 |
declare Abs_complex_cancel [simp] |
|
165 |
||
166 |
lemma complex_Re_Im_cancel_iff: "(w=z) = (Re(w) = Re(z) & Im(w) = Im(z))" |
|
14334 | 167 |
apply (rule_tac z = w in eq_Abs_complex) |
168 |
apply (rule_tac z = z in eq_Abs_complex) |
|
14323 | 169 |
apply (auto dest: inj_Abs_complex [THEN injD]) |
170 |
done |
|
171 |
||
172 |
lemma complex_Re_zero: "Re 0 = 0" |
|
173 |
apply (unfold complex_zero_def) |
|
174 |
apply (simp (no_asm)) |
|
175 |
done |
|
176 |
||
177 |
lemma complex_Im_zero: "Im 0 = 0" |
|
178 |
apply (unfold complex_zero_def) |
|
179 |
apply (simp (no_asm)) |
|
180 |
done |
|
181 |
declare complex_Re_zero [simp] complex_Im_zero [simp] |
|
182 |
||
183 |
lemma complex_Re_one: "Re 1 = 1" |
|
184 |
apply (unfold complex_one_def) |
|
185 |
apply (simp (no_asm)) |
|
186 |
done |
|
187 |
declare complex_Re_one [simp] |
|
188 |
||
189 |
lemma complex_Im_one: "Im 1 = 0" |
|
190 |
apply (unfold complex_one_def) |
|
191 |
apply (simp (no_asm)) |
|
192 |
done |
|
193 |
declare complex_Im_one [simp] |
|
194 |
||
195 |
lemma complex_Re_i: "Re(ii) = 0" |
|
14334 | 196 |
by (unfold i_def, auto) |
14323 | 197 |
declare complex_Re_i [simp] |
198 |
||
199 |
lemma complex_Im_i: "Im(ii) = 1" |
|
14334 | 200 |
by (unfold i_def, auto) |
14323 | 201 |
declare complex_Im_i [simp] |
202 |
||
203 |
lemma Re_complex_of_real_zero: "Re(complex_of_real 0) = 0" |
|
204 |
apply (unfold complex_of_real_def) |
|
205 |
apply (simp (no_asm)) |
|
206 |
done |
|
207 |
declare Re_complex_of_real_zero [simp] |
|
208 |
||
209 |
lemma Im_complex_of_real_zero: "Im(complex_of_real 0) = 0" |
|
210 |
apply (unfold complex_of_real_def) |
|
211 |
apply (simp (no_asm)) |
|
212 |
done |
|
213 |
declare Im_complex_of_real_zero [simp] |
|
214 |
||
215 |
lemma Re_complex_of_real_one: "Re(complex_of_real 1) = 1" |
|
216 |
apply (unfold complex_of_real_def) |
|
217 |
apply (simp (no_asm)) |
|
218 |
done |
|
219 |
declare Re_complex_of_real_one [simp] |
|
220 |
||
221 |
lemma Im_complex_of_real_one: "Im(complex_of_real 1) = 0" |
|
222 |
apply (unfold complex_of_real_def) |
|
223 |
apply (simp (no_asm)) |
|
224 |
done |
|
225 |
declare Im_complex_of_real_one [simp] |
|
226 |
||
227 |
lemma Re_complex_of_real: "Re(complex_of_real z) = z" |
|
14334 | 228 |
by (unfold complex_of_real_def, auto) |
14323 | 229 |
declare Re_complex_of_real [simp] |
230 |
||
231 |
lemma Im_complex_of_real: "Im(complex_of_real z) = 0" |
|
14334 | 232 |
by (unfold complex_of_real_def, auto) |
14323 | 233 |
declare Im_complex_of_real [simp] |
234 |
||
235 |
||
236 |
subsection{*Negation*} |
|
237 |
||
238 |
lemma complex_minus: "- Abs_complex(x,y) = Abs_complex(-x,-y)" |
|
239 |
apply (unfold complex_minus_def) |
|
240 |
apply (simp (no_asm)) |
|
241 |
done |
|
242 |
||
243 |
lemma complex_Re_minus: "Re (-z) = - Re z" |
|
244 |
apply (unfold Re_def) |
|
14334 | 245 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 246 |
apply (auto simp add: complex_minus) |
247 |
done |
|
248 |
||
249 |
lemma complex_Im_minus: "Im (-z) = - Im z" |
|
250 |
apply (unfold Im_def) |
|
14334 | 251 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 252 |
apply (auto simp add: complex_minus) |
253 |
done |
|
254 |
||
255 |
lemma complex_minus_minus: "- (- z) = (z::complex)" |
|
256 |
apply (unfold complex_minus_def) |
|
257 |
apply (simp (no_asm)) |
|
258 |
done |
|
259 |
declare complex_minus_minus [simp] |
|
260 |
||
261 |
lemma inj_complex_minus: "inj(%r::complex. -r)" |
|
262 |
apply (rule inj_onI) |
|
14334 | 263 |
apply (drule_tac f = uminus in arg_cong, simp) |
14323 | 264 |
done |
265 |
||
266 |
lemma complex_minus_zero: "-(0::complex) = 0" |
|
267 |
apply (unfold complex_zero_def) |
|
268 |
apply (simp (no_asm) add: complex_minus) |
|
269 |
done |
|
270 |
declare complex_minus_zero [simp] |
|
271 |
||
272 |
lemma complex_minus_zero_iff: "(-x = 0) = (x = (0::complex))" |
|
14334 | 273 |
apply (rule_tac z = x in eq_Abs_complex) |
14323 | 274 |
apply (auto dest: inj_Abs_complex [THEN injD] |
275 |
simp add: complex_zero_def complex_minus) |
|
276 |
done |
|
277 |
declare complex_minus_zero_iff [simp] |
|
278 |
||
279 |
lemma complex_minus_zero_iff2: "(0 = -x) = (x = (0::real))" |
|
14334 | 280 |
by (auto dest: sym) |
14323 | 281 |
declare complex_minus_zero_iff2 [simp] |
282 |
||
283 |
lemma complex_minus_not_zero_iff: "(-x ~= 0) = (x ~= (0::complex))" |
|
14334 | 284 |
by auto |
14323 | 285 |
|
286 |
||
287 |
subsection{*Addition*} |
|
288 |
||
289 |
lemma complex_add: |
|
290 |
"Abs_complex(x1,y1) + Abs_complex(x2,y2) = Abs_complex(x1+x2,y1+y2)" |
|
291 |
apply (unfold complex_add_def) |
|
292 |
apply (simp (no_asm)) |
|
293 |
done |
|
294 |
||
295 |
lemma complex_Re_add: "Re(x + y) = Re(x) + Re(y)" |
|
296 |
apply (unfold Re_def) |
|
14334 | 297 |
apply (rule_tac z = x in eq_Abs_complex) |
298 |
apply (rule_tac z = y in eq_Abs_complex) |
|
14323 | 299 |
apply (auto simp add: complex_add) |
300 |
done |
|
301 |
||
302 |
lemma complex_Im_add: "Im(x + y) = Im(x) + Im(y)" |
|
303 |
apply (unfold Im_def) |
|
14334 | 304 |
apply (rule_tac z = x in eq_Abs_complex) |
305 |
apply (rule_tac z = y in eq_Abs_complex) |
|
14323 | 306 |
apply (auto simp add: complex_add) |
307 |
done |
|
308 |
||
309 |
lemma complex_add_commute: "(u::complex) + v = v + u" |
|
310 |
apply (unfold complex_add_def) |
|
311 |
apply (simp (no_asm) add: real_add_commute) |
|
312 |
done |
|
313 |
||
314 |
lemma complex_add_assoc: "((u::complex) + v) + w = u + (v + w)" |
|
315 |
apply (unfold complex_add_def) |
|
316 |
apply (simp (no_asm) add: real_add_assoc) |
|
317 |
done |
|
318 |
||
319 |
lemma complex_add_left_commute: "(x::complex) + (y + z) = y + (x + z)" |
|
320 |
apply (unfold complex_add_def) |
|
14334 | 321 |
apply (simp (no_asm) add: add_left_commute) |
14323 | 322 |
done |
323 |
||
324 |
lemmas complex_add_ac = complex_add_assoc complex_add_commute |
|
325 |
complex_add_left_commute |
|
326 |
||
327 |
lemma complex_add_zero_left: "(0::complex) + z = z" |
|
328 |
apply (unfold complex_add_def complex_zero_def) |
|
329 |
apply (simp (no_asm)) |
|
330 |
done |
|
331 |
declare complex_add_zero_left [simp] |
|
332 |
||
333 |
lemma complex_add_zero_right: "z + (0::complex) = z" |
|
334 |
apply (unfold complex_add_def complex_zero_def) |
|
335 |
apply (simp (no_asm)) |
|
336 |
done |
|
337 |
declare complex_add_zero_right [simp] |
|
338 |
||
339 |
lemma complex_add_minus_right_zero: |
|
340 |
"z + -z = (0::complex)" |
|
341 |
apply (unfold complex_add_def complex_minus_def complex_zero_def) |
|
342 |
apply (simp (no_asm)) |
|
343 |
done |
|
344 |
declare complex_add_minus_right_zero [simp] |
|
345 |
||
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
346 |
lemma complex_add_minus_left: |
14323 | 347 |
"-z + z = (0::complex)" |
348 |
apply (unfold complex_add_def complex_minus_def complex_zero_def) |
|
349 |
apply (simp (no_asm)) |
|
350 |
done |
|
351 |
||
352 |
lemma complex_add_minus_cancel: "z + (- z + w) = (w::complex)" |
|
353 |
apply (simp (no_asm) add: complex_add_assoc [symmetric]) |
|
354 |
done |
|
355 |
||
356 |
lemma complex_minus_add_cancel: "(-z) + (z + w) = (w::complex)" |
|
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
357 |
by (simp add: complex_add_minus_left complex_add_assoc [symmetric]) |
14323 | 358 |
|
359 |
declare complex_add_minus_cancel [simp] complex_minus_add_cancel [simp] |
|
360 |
||
361 |
lemma complex_add_minus_eq_minus: "x + y = (0::complex) ==> x = -y" |
|
14334 | 362 |
by (auto simp add: complex_Re_Im_cancel_iff complex_Re_add complex_Im_add complex_Re_minus complex_Im_minus) |
14323 | 363 |
|
364 |
lemma complex_minus_add_distrib: "-(x + y) = -x + -(y::complex)" |
|
14334 | 365 |
apply (rule_tac z = x in eq_Abs_complex) |
366 |
apply (rule_tac z = y in eq_Abs_complex) |
|
14323 | 367 |
apply (auto simp add: complex_minus complex_add) |
368 |
done |
|
369 |
declare complex_minus_add_distrib [simp] |
|
370 |
||
371 |
lemma complex_add_left_cancel: "((x::complex) + y = x + z) = (y = z)" |
|
14334 | 372 |
apply safe |
14323 | 373 |
apply (drule_tac f = "%t.-x + t" in arg_cong) |
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
374 |
apply (simp add: complex_add_minus_left complex_add_assoc [symmetric]) |
14323 | 375 |
done |
376 |
declare complex_add_left_cancel [iff] |
|
377 |
||
378 |
lemma complex_add_right_cancel: "(y + (x::complex)= z + x) = (y = z)" |
|
379 |
apply (simp (no_asm) add: complex_add_commute) |
|
380 |
done |
|
381 |
declare complex_add_right_cancel [simp] |
|
382 |
||
383 |
lemma complex_eq_minus_iff: "((x::complex) = y) = (0 = x + - y)" |
|
14334 | 384 |
apply safe |
385 |
apply (rule_tac [2] x1 = "-y" in complex_add_right_cancel [THEN iffD1], auto) |
|
14323 | 386 |
done |
387 |
||
388 |
lemma complex_eq_minus_iff2: "((x::complex) = y) = (x + - y = 0)" |
|
14334 | 389 |
apply safe |
390 |
apply (rule_tac [2] x1 = "-y" in complex_add_right_cancel [THEN iffD1], auto) |
|
14323 | 391 |
done |
392 |
||
393 |
lemma complex_diff_0: "(0::complex) - x = -x" |
|
394 |
apply (simp (no_asm) add: complex_diff_def) |
|
395 |
done |
|
396 |
||
397 |
lemma complex_diff_0_right: "x - (0::complex) = x" |
|
398 |
apply (simp (no_asm) add: complex_diff_def) |
|
399 |
done |
|
400 |
||
401 |
lemma complex_diff_self: "x - x = (0::complex)" |
|
402 |
apply (simp (no_asm) add: complex_diff_def) |
|
403 |
done |
|
404 |
||
405 |
declare complex_diff_0 [simp] complex_diff_0_right [simp] complex_diff_self [simp] |
|
406 |
||
407 |
lemma complex_diff: |
|
408 |
"Abs_complex(x1,y1) - Abs_complex(x2,y2) = Abs_complex(x1-x2,y1-y2)" |
|
409 |
apply (unfold complex_diff_def) |
|
410 |
apply (simp (no_asm) add: complex_add complex_minus) |
|
411 |
done |
|
412 |
||
413 |
lemma complex_diff_eq_eq: "((x::complex) - y = z) = (x = z + y)" |
|
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
414 |
by (auto simp add: complex_add_minus_left complex_diff_def complex_add_assoc) |
14323 | 415 |
|
416 |
||
417 |
subsection{*Multiplication*} |
|
418 |
||
419 |
lemma complex_mult: |
|
420 |
"Abs_complex(x1,y1) * Abs_complex(x2,y2) = |
|
421 |
Abs_complex(x1*x2 - y1*y2,x1*y2 + y1*x2)" |
|
422 |
apply (unfold complex_mult_def) |
|
423 |
apply (simp (no_asm)) |
|
424 |
done |
|
425 |
||
426 |
lemma complex_mult_commute: "(w::complex) * z = z * w" |
|
427 |
apply (unfold complex_mult_def) |
|
428 |
apply (simp (no_asm) add: real_mult_commute real_add_commute) |
|
429 |
done |
|
430 |
||
431 |
lemma complex_mult_assoc: "((u::complex) * v) * w = u * (v * w)" |
|
432 |
apply (unfold complex_mult_def) |
|
14334 | 433 |
apply (simp (no_asm) add: complex_Re_Im_cancel_iff real_mult_assoc right_diff_distrib right_distrib left_diff_distrib left_distrib mult_left_commute) |
14323 | 434 |
done |
435 |
||
436 |
lemma complex_mult_left_commute: "(x::complex) * (y * z) = y * (x * z)" |
|
437 |
apply (unfold complex_mult_def) |
|
14334 | 438 |
apply (simp (no_asm) add: complex_Re_Im_cancel_iff mult_left_commute right_diff_distrib right_distrib) |
14323 | 439 |
done |
440 |
||
441 |
lemmas complex_mult_ac = complex_mult_assoc complex_mult_commute |
|
442 |
complex_mult_left_commute |
|
443 |
||
444 |
lemma complex_mult_one_left: "(1::complex) * z = z" |
|
445 |
apply (unfold complex_one_def) |
|
14334 | 446 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 447 |
apply (simp (no_asm_simp) add: complex_mult) |
448 |
done |
|
449 |
declare complex_mult_one_left [simp] |
|
450 |
||
451 |
lemma complex_mult_one_right: "z * (1::complex) = z" |
|
452 |
apply (simp (no_asm) add: complex_mult_commute) |
|
453 |
done |
|
454 |
declare complex_mult_one_right [simp] |
|
455 |
||
456 |
lemma complex_mult_zero_left: "(0::complex) * z = 0" |
|
457 |
apply (unfold complex_zero_def) |
|
14334 | 458 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 459 |
apply (simp (no_asm_simp) add: complex_mult) |
460 |
done |
|
461 |
declare complex_mult_zero_left [simp] |
|
462 |
||
463 |
lemma complex_mult_zero_right: "z * 0 = (0::complex)" |
|
464 |
apply (simp (no_asm) add: complex_mult_commute) |
|
465 |
done |
|
466 |
declare complex_mult_zero_right [simp] |
|
467 |
||
468 |
lemma complex_divide_zero: "0 / z = (0::complex)" |
|
14334 | 469 |
by (unfold complex_divide_def, auto) |
14323 | 470 |
declare complex_divide_zero [simp] |
471 |
||
472 |
lemma complex_minus_mult_eq1: "-(x * y) = -x * (y::complex)" |
|
14334 | 473 |
apply (rule_tac z = x in eq_Abs_complex) |
474 |
apply (rule_tac z = y in eq_Abs_complex) |
|
14323 | 475 |
apply (auto simp add: complex_mult complex_minus real_diff_def) |
476 |
done |
|
477 |
||
478 |
lemma complex_minus_mult_eq2: "-(x * y) = x * -(y::complex)" |
|
14334 | 479 |
apply (rule_tac z = x in eq_Abs_complex) |
480 |
apply (rule_tac z = y in eq_Abs_complex) |
|
14323 | 481 |
apply (auto simp add: complex_mult complex_minus real_diff_def) |
482 |
done |
|
483 |
||
484 |
lemma complex_add_mult_distrib: "((z1::complex) + z2) * w = (z1 * w) + (z2 * w)" |
|
14334 | 485 |
apply (rule_tac z = z1 in eq_Abs_complex) |
486 |
apply (rule_tac z = z2 in eq_Abs_complex) |
|
487 |
apply (rule_tac z = w in eq_Abs_complex) |
|
488 |
apply (auto simp add: complex_mult complex_add left_distrib real_diff_def add_ac) |
|
14323 | 489 |
done |
490 |
||
491 |
lemma complex_add_mult_distrib2: "(w::complex) * (z1 + z2) = (w * z1) + (w * z2)" |
|
492 |
apply (rule_tac z1 = "z1 + z2" in complex_mult_commute [THEN ssubst]) |
|
493 |
apply (simp (no_asm) add: complex_add_mult_distrib) |
|
494 |
apply (simp (no_asm) add: complex_mult_commute) |
|
495 |
done |
|
496 |
||
497 |
lemma complex_zero_not_eq_one: "(0::complex) ~= 1" |
|
498 |
apply (unfold complex_zero_def complex_one_def) |
|
499 |
apply (simp (no_asm) add: complex_Re_Im_cancel_iff) |
|
500 |
done |
|
501 |
declare complex_zero_not_eq_one [simp] |
|
502 |
declare complex_zero_not_eq_one [THEN not_sym, simp] |
|
503 |
||
504 |
||
505 |
subsection{*Inverse*} |
|
506 |
||
507 |
lemma complex_inverse: "inverse (Abs_complex(x,y)) = |
|
508 |
Abs_complex(x/(x ^ 2 + y ^ 2),-y/(x ^ 2 + y ^ 2))" |
|
509 |
apply (unfold complex_inverse_def) |
|
510 |
apply (simp (no_asm)) |
|
511 |
done |
|
512 |
||
513 |
lemma COMPLEX_INVERSE_ZERO: "inverse 0 = (0::complex)" |
|
14334 | 514 |
by (unfold complex_inverse_def complex_zero_def, auto) |
14323 | 515 |
|
516 |
lemma COMPLEX_DIVISION_BY_ZERO: "a / (0::complex) = 0" |
|
517 |
apply (simp (no_asm) add: complex_divide_def COMPLEX_INVERSE_ZERO) |
|
518 |
done |
|
519 |
||
14335 | 520 |
instance complex :: division_by_zero |
521 |
proof |
|
522 |
fix x :: complex |
|
523 |
show "inverse 0 = (0::complex)" by (rule COMPLEX_INVERSE_ZERO) |
|
524 |
show "x/0 = 0" by (rule COMPLEX_DIVISION_BY_ZERO) |
|
525 |
qed |
|
526 |
||
14323 | 527 |
lemma complex_mult_inv_left: "z ~= (0::complex) ==> inverse(z) * z = 1" |
14334 | 528 |
apply (rule_tac z = z in eq_Abs_complex) |
529 |
apply (auto simp add: complex_mult complex_inverse complex_one_def |
|
530 |
complex_zero_def add_divide_distrib [symmetric] real_power_two mult_ac) |
|
531 |
apply (drule_tac y = y in real_sum_squares_not_zero) |
|
532 |
apply (drule_tac [2] x = x in real_sum_squares_not_zero2, auto) |
|
14323 | 533 |
done |
534 |
declare complex_mult_inv_left [simp] |
|
535 |
||
536 |
lemma complex_mult_inv_right: "z ~= (0::complex) ==> z * inverse(z) = 1" |
|
14334 | 537 |
by (auto intro: complex_mult_commute [THEN subst]) |
14323 | 538 |
declare complex_mult_inv_right [simp] |
539 |
||
14335 | 540 |
|
541 |
subsection {* The field of complex numbers *} |
|
542 |
||
543 |
instance complex :: field |
|
544 |
proof |
|
545 |
fix z u v w :: complex |
|
546 |
show "(u + v) + w = u + (v + w)" |
|
547 |
by (rule complex_add_assoc) |
|
548 |
show "z + w = w + z" |
|
549 |
by (rule complex_add_commute) |
|
550 |
show "0 + z = z" |
|
551 |
by (rule complex_add_zero_left) |
|
552 |
show "-z + z = 0" |
|
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
553 |
by (rule complex_add_minus_left) |
14335 | 554 |
show "z - w = z + -w" |
555 |
by (simp add: complex_diff_def) |
|
556 |
show "(u * v) * w = u * (v * w)" |
|
557 |
by (rule complex_mult_assoc) |
|
558 |
show "z * w = w * z" |
|
559 |
by (rule complex_mult_commute) |
|
560 |
show "1 * z = z" |
|
561 |
by (rule complex_mult_one_left) |
|
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
562 |
show "0 \<noteq> (1::complex)" |
14335 | 563 |
by (rule complex_zero_not_eq_one) |
564 |
show "(u + v) * w = u * w + v * w" |
|
565 |
by (rule complex_add_mult_distrib) |
|
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
566 |
show "z+u = z+v ==> u=v" |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
567 |
proof - |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
568 |
assume eq: "z+u = z+v" |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
569 |
hence "(-z + z) + u = (-z + z) + v" by (simp only: eq complex_add_assoc) |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
570 |
thus "u = v" by (simp add: complex_add_minus_left) |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
571 |
qed |
14335 | 572 |
assume neq: "w \<noteq> 0" |
573 |
thus "z / w = z * inverse w" |
|
574 |
by (simp add: complex_divide_def) |
|
575 |
show "inverse w * w = 1" |
|
576 |
by (simp add: neq complex_mult_inv_left) |
|
577 |
qed |
|
578 |
||
579 |
||
580 |
lemma complex_minus_mult_commute: "-x * y = x * -(y::complex)" |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
581 |
apply (simp) |
14323 | 582 |
done |
583 |
||
584 |
subsection{*Embedding Properties for @{term complex_of_real} Map*} |
|
585 |
||
586 |
lemma inj_complex_of_real: "inj complex_of_real" |
|
587 |
apply (rule inj_onI) |
|
588 |
apply (auto dest: inj_Abs_complex [THEN injD] simp add: complex_of_real_def) |
|
589 |
done |
|
590 |
||
591 |
lemma complex_of_real_one: |
|
592 |
"complex_of_real 1 = 1" |
|
593 |
apply (unfold complex_one_def complex_of_real_def) |
|
594 |
apply (rule refl) |
|
595 |
done |
|
596 |
declare complex_of_real_one [simp] |
|
597 |
||
598 |
lemma complex_of_real_zero: |
|
599 |
"complex_of_real 0 = 0" |
|
600 |
apply (unfold complex_zero_def complex_of_real_def) |
|
601 |
apply (rule refl) |
|
602 |
done |
|
603 |
declare complex_of_real_zero [simp] |
|
604 |
||
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
605 |
lemma complex_of_real_eq_iff: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
606 |
"(complex_of_real x = complex_of_real y) = (x = y)" |
14334 | 607 |
by (auto dest: inj_complex_of_real [THEN injD]) |
14323 | 608 |
declare complex_of_real_eq_iff [iff] |
609 |
||
610 |
lemma complex_of_real_minus: "complex_of_real(-x) = - complex_of_real x" |
|
611 |
apply (simp (no_asm) add: complex_of_real_def complex_minus) |
|
612 |
done |
|
613 |
||
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
614 |
lemma complex_of_real_inverse: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
615 |
"complex_of_real(inverse x) = inverse(complex_of_real x)" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
616 |
apply (case_tac "x=0", simp) |
14334 | 617 |
apply (simp add: complex_inverse complex_of_real_def real_divide_def |
618 |
inverse_mult_distrib real_power_two) |
|
14323 | 619 |
done |
620 |
||
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
621 |
lemma complex_of_real_add: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
622 |
"complex_of_real x + complex_of_real y = complex_of_real (x + y)" |
14323 | 623 |
apply (simp (no_asm) add: complex_add complex_of_real_def) |
624 |
done |
|
625 |
||
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
626 |
lemma complex_of_real_diff: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
627 |
"complex_of_real x - complex_of_real y = complex_of_real (x - y)" |
14323 | 628 |
apply (simp (no_asm) add: complex_of_real_minus [symmetric] complex_diff_def complex_of_real_add) |
629 |
done |
|
630 |
||
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
631 |
lemma complex_of_real_mult: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
632 |
"complex_of_real x * complex_of_real y = complex_of_real (x * y)" |
14323 | 633 |
apply (simp (no_asm) add: complex_mult complex_of_real_def) |
634 |
done |
|
635 |
||
636 |
lemma complex_of_real_divide: |
|
637 |
"complex_of_real x / complex_of_real y = complex_of_real(x/y)" |
|
638 |
apply (unfold complex_divide_def) |
|
639 |
apply (case_tac "y=0") |
|
640 |
apply (simp (no_asm_simp) add: DIVISION_BY_ZERO COMPLEX_INVERSE_ZERO) |
|
641 |
apply (simp (no_asm_simp) add: complex_of_real_mult [symmetric] complex_of_real_inverse real_divide_def) |
|
642 |
done |
|
643 |
||
644 |
lemma complex_of_real_pow: "complex_of_real (x ^ n) = (complex_of_real x) ^ n" |
|
645 |
apply (induct_tac "n") |
|
646 |
apply (auto simp add: complex_of_real_mult [symmetric]) |
|
647 |
done |
|
648 |
||
649 |
lemma complex_mod: "cmod (Abs_complex(x,y)) = sqrt(x ^ 2 + y ^ 2)" |
|
650 |
apply (unfold cmod_def) |
|
651 |
apply (simp (no_asm)) |
|
652 |
done |
|
653 |
||
654 |
lemma complex_mod_zero: "cmod(0) = 0" |
|
655 |
apply (unfold cmod_def) |
|
656 |
apply (simp (no_asm)) |
|
657 |
done |
|
658 |
declare complex_mod_zero [simp] |
|
659 |
||
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
660 |
lemma complex_mod_one [simp]: "cmod(1) = 1" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
661 |
by (simp add: cmod_def real_power_two) |
14323 | 662 |
|
663 |
lemma complex_mod_complex_of_real: "cmod(complex_of_real x) = abs x" |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
664 |
apply (simp add: complex_of_real_def real_power_two complex_mod) |
14323 | 665 |
done |
666 |
declare complex_mod_complex_of_real [simp] |
|
667 |
||
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
668 |
lemma complex_of_real_abs: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
669 |
"complex_of_real (abs x) = complex_of_real(cmod(complex_of_real x))" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
670 |
by (simp) |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
671 |
|
14323 | 672 |
|
673 |
||
674 |
subsection{*Conjugation is an Automorphism*} |
|
675 |
||
676 |
lemma complex_cnj: "cnj (Abs_complex(x,y)) = Abs_complex(x,-y)" |
|
677 |
apply (unfold cnj_def) |
|
678 |
apply (simp (no_asm)) |
|
679 |
done |
|
680 |
||
681 |
lemma inj_cnj: "inj cnj" |
|
682 |
apply (rule inj_onI) |
|
683 |
apply (auto simp add: cnj_def Abs_complex_cancel_iff complex_Re_Im_cancel_iff) |
|
684 |
done |
|
685 |
||
686 |
lemma complex_cnj_cancel_iff: "(cnj x = cnj y) = (x = y)" |
|
14334 | 687 |
by (auto dest: inj_cnj [THEN injD]) |
14323 | 688 |
declare complex_cnj_cancel_iff [simp] |
689 |
||
690 |
lemma complex_cnj_cnj: "cnj (cnj z) = z" |
|
691 |
apply (unfold cnj_def) |
|
692 |
apply (simp (no_asm)) |
|
693 |
done |
|
694 |
declare complex_cnj_cnj [simp] |
|
695 |
||
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
696 |
lemma complex_cnj_complex_of_real: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
697 |
"cnj (complex_of_real x) = complex_of_real x" |
14323 | 698 |
apply (unfold complex_of_real_def) |
699 |
apply (simp (no_asm) add: complex_cnj) |
|
700 |
done |
|
701 |
declare complex_cnj_complex_of_real [simp] |
|
702 |
||
703 |
lemma complex_mod_cnj: "cmod (cnj z) = cmod z" |
|
14334 | 704 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 705 |
apply (simp (no_asm_simp) add: complex_cnj complex_mod real_power_two) |
706 |
done |
|
707 |
declare complex_mod_cnj [simp] |
|
708 |
||
709 |
lemma complex_cnj_minus: "cnj (-z) = - cnj z" |
|
710 |
apply (unfold cnj_def) |
|
711 |
apply (simp (no_asm) add: complex_minus complex_Re_minus complex_Im_minus) |
|
712 |
done |
|
713 |
||
714 |
lemma complex_cnj_inverse: "cnj(inverse z) = inverse(cnj z)" |
|
14334 | 715 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 716 |
apply (simp (no_asm_simp) add: complex_cnj complex_inverse real_power_two) |
717 |
done |
|
718 |
||
719 |
lemma complex_cnj_add: "cnj(w + z) = cnj(w) + cnj(z)" |
|
14334 | 720 |
apply (rule_tac z = w in eq_Abs_complex) |
721 |
apply (rule_tac z = z in eq_Abs_complex) |
|
14323 | 722 |
apply (simp (no_asm_simp) add: complex_cnj complex_add) |
723 |
done |
|
724 |
||
725 |
lemma complex_cnj_diff: "cnj(w - z) = cnj(w) - cnj(z)" |
|
726 |
apply (unfold complex_diff_def) |
|
727 |
apply (simp (no_asm) add: complex_cnj_add complex_cnj_minus) |
|
728 |
done |
|
729 |
||
730 |
lemma complex_cnj_mult: "cnj(w * z) = cnj(w) * cnj(z)" |
|
14334 | 731 |
apply (rule_tac z = w in eq_Abs_complex) |
732 |
apply (rule_tac z = z in eq_Abs_complex) |
|
14323 | 733 |
apply (simp (no_asm_simp) add: complex_cnj complex_mult) |
734 |
done |
|
735 |
||
736 |
lemma complex_cnj_divide: "cnj(w / z) = (cnj w)/(cnj z)" |
|
737 |
apply (unfold complex_divide_def) |
|
738 |
apply (simp (no_asm) add: complex_cnj_mult complex_cnj_inverse) |
|
739 |
done |
|
740 |
||
741 |
lemma complex_cnj_one: "cnj 1 = 1" |
|
742 |
apply (unfold cnj_def complex_one_def) |
|
743 |
apply (simp (no_asm)) |
|
744 |
done |
|
745 |
declare complex_cnj_one [simp] |
|
746 |
||
747 |
lemma complex_cnj_pow: "cnj(z ^ n) = cnj(z) ^ n" |
|
748 |
apply (induct_tac "n") |
|
749 |
apply (auto simp add: complex_cnj_mult) |
|
750 |
done |
|
751 |
||
752 |
lemma complex_add_cnj: "z + cnj z = complex_of_real (2 * Re(z))" |
|
14334 | 753 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 754 |
apply (simp (no_asm_simp) add: complex_add complex_cnj complex_of_real_def) |
755 |
done |
|
756 |
||
757 |
lemma complex_diff_cnj: "z - cnj z = complex_of_real (2 * Im(z)) * ii" |
|
14334 | 758 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 759 |
apply (simp (no_asm_simp) add: complex_add complex_cnj complex_of_real_def complex_diff_def complex_minus i_def complex_mult) |
760 |
done |
|
761 |
||
762 |
lemma complex_cnj_zero: "cnj 0 = 0" |
|
14334 | 763 |
by (simp add: cnj_def complex_zero_def) |
14323 | 764 |
declare complex_cnj_zero [simp] |
765 |
||
766 |
lemma complex_cnj_zero_iff: "(cnj z = 0) = (z = 0)" |
|
14334 | 767 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 768 |
apply (auto simp add: complex_zero_def complex_cnj) |
769 |
done |
|
770 |
declare complex_cnj_zero_iff [iff] |
|
771 |
||
772 |
lemma complex_mult_cnj: "z * cnj z = complex_of_real (Re(z) ^ 2 + Im(z) ^ 2)" |
|
14334 | 773 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 774 |
apply (auto simp add: complex_cnj complex_mult complex_of_real_def real_power_two) |
775 |
done |
|
776 |
||
777 |
||
778 |
subsection{*Algebra*} |
|
779 |
||
780 |
lemma complex_add_left_cancel_zero: "(x + y = x) = (y = (0::complex))" |
|
781 |
apply (unfold complex_zero_def) |
|
14334 | 782 |
apply (rule_tac z = x in eq_Abs_complex) |
783 |
apply (rule_tac z = y in eq_Abs_complex) |
|
14323 | 784 |
apply (auto simp add: complex_add) |
785 |
done |
|
786 |
declare complex_add_left_cancel_zero [simp] |
|
787 |
||
788 |
lemma complex_diff_mult_distrib: |
|
789 |
"((z1::complex) - z2) * w = (z1 * w) - (z2 * w)" |
|
790 |
apply (unfold complex_diff_def) |
|
791 |
apply (simp (no_asm) add: complex_add_mult_distrib) |
|
792 |
done |
|
793 |
||
794 |
lemma complex_diff_mult_distrib2: |
|
795 |
"(w::complex) * (z1 - z2) = (w * z1) - (w * z2)" |
|
796 |
apply (unfold complex_diff_def) |
|
797 |
apply (simp (no_asm) add: complex_add_mult_distrib2) |
|
798 |
done |
|
799 |
||
800 |
||
801 |
subsection{*Modulus*} |
|
802 |
||
803 |
lemma complex_mod_eq_zero_cancel: "(cmod x = 0) = (x = 0)" |
|
14334 | 804 |
apply (rule_tac z = x in eq_Abs_complex) |
14323 | 805 |
apply (auto intro: real_sum_squares_cancel real_sum_squares_cancel2 simp add: complex_mod complex_zero_def real_power_two) |
806 |
done |
|
807 |
declare complex_mod_eq_zero_cancel [simp] |
|
808 |
||
809 |
lemma complex_mod_complex_of_real_of_nat: "cmod (complex_of_real(real (n::nat))) = real n" |
|
810 |
apply (simp (no_asm)) |
|
811 |
done |
|
812 |
declare complex_mod_complex_of_real_of_nat [simp] |
|
813 |
||
814 |
lemma complex_mod_minus: "cmod (-x) = cmod(x)" |
|
14334 | 815 |
apply (rule_tac z = x in eq_Abs_complex) |
14323 | 816 |
apply (simp (no_asm_simp) add: complex_mod complex_minus real_power_two) |
817 |
done |
|
818 |
declare complex_mod_minus [simp] |
|
819 |
||
820 |
lemma complex_mod_mult_cnj: "cmod(z * cnj(z)) = cmod(z) ^ 2" |
|
14334 | 821 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 822 |
apply (simp (no_asm_simp) add: complex_mod complex_cnj complex_mult); |
823 |
apply (simp (no_asm) add: real_power_two real_abs_def) |
|
824 |
done |
|
825 |
||
826 |
lemma complex_mod_squared: "cmod(Abs_complex(x,y)) ^ 2 = x ^ 2 + y ^ 2" |
|
14334 | 827 |
by (unfold cmod_def, auto) |
14323 | 828 |
|
829 |
lemma complex_mod_ge_zero: "0 <= cmod x" |
|
830 |
apply (unfold cmod_def) |
|
831 |
apply (auto intro: real_sqrt_ge_zero) |
|
832 |
done |
|
833 |
declare complex_mod_ge_zero [simp] |
|
834 |
||
835 |
lemma abs_cmod_cancel: "abs(cmod x) = cmod x" |
|
14334 | 836 |
by (auto intro: abs_eqI1) |
14323 | 837 |
declare abs_cmod_cancel [simp] |
838 |
||
839 |
lemma complex_mod_mult: "cmod(x*y) = cmod(x) * cmod(y)" |
|
14334 | 840 |
apply (rule_tac z = x in eq_Abs_complex) |
841 |
apply (rule_tac z = y in eq_Abs_complex) |
|
14323 | 842 |
apply (auto simp add: complex_mult complex_mod real_sqrt_mult_distrib2 [symmetric] simp del: realpow_Suc) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
843 |
apply (rule_tac n = 1 in power_inject_base) |
14323 | 844 |
apply (auto simp add: real_power_two [symmetric] simp del: realpow_Suc) |
14334 | 845 |
apply (auto simp add: real_diff_def real_power_two right_distrib left_distrib add_ac mult_ac) |
14323 | 846 |
done |
847 |
||
848 |
lemma complex_mod_add_squared_eq: "cmod(x + y) ^ 2 = cmod(x) ^ 2 + cmod(y) ^ 2 + 2 * Re(x * cnj y)" |
|
14334 | 849 |
apply (rule_tac z = x in eq_Abs_complex) |
850 |
apply (rule_tac z = y in eq_Abs_complex) |
|
14323 | 851 |
apply (auto simp add: complex_add complex_mod_squared complex_mult complex_cnj real_diff_def simp del: realpow_Suc) |
14334 | 852 |
apply (auto simp add: right_distrib left_distrib real_power_two mult_ac add_ac) |
14323 | 853 |
done |
854 |
||
855 |
lemma complex_Re_mult_cnj_le_cmod: "Re(x * cnj y) <= cmod(x * cnj y)" |
|
14334 | 856 |
apply (rule_tac z = x in eq_Abs_complex) |
857 |
apply (rule_tac z = y in eq_Abs_complex) |
|
14323 | 858 |
apply (auto simp add: complex_mod complex_mult complex_cnj real_diff_def simp del: realpow_Suc) |
859 |
done |
|
860 |
declare complex_Re_mult_cnj_le_cmod [simp] |
|
861 |
||
862 |
lemma complex_Re_mult_cnj_le_cmod2: "Re(x * cnj y) <= cmod(x * y)" |
|
14334 | 863 |
apply (cut_tac x = x and y = y in complex_Re_mult_cnj_le_cmod) |
14323 | 864 |
apply (simp add: complex_mod_mult) |
865 |
done |
|
866 |
declare complex_Re_mult_cnj_le_cmod2 [simp] |
|
867 |
||
868 |
lemma real_sum_squared_expand: "((x::real) + y) ^ 2 = x ^ 2 + y ^ 2 + 2 * x * y" |
|
14334 | 869 |
apply (simp (no_asm) add: left_distrib right_distrib real_power_two) |
14323 | 870 |
done |
871 |
||
872 |
lemma complex_mod_triangle_squared: "cmod (x + y) ^ 2 <= (cmod(x) + cmod(y)) ^ 2" |
|
873 |
apply (simp (no_asm) add: real_sum_squared_expand complex_mod_add_squared_eq real_mult_assoc complex_mod_mult [symmetric]) |
|
874 |
done |
|
875 |
declare complex_mod_triangle_squared [simp] |
|
876 |
||
877 |
lemma complex_mod_minus_le_complex_mod: "- cmod x <= cmod x" |
|
878 |
apply (rule order_trans [OF _ complex_mod_ge_zero]) |
|
879 |
apply (simp (no_asm)) |
|
880 |
done |
|
881 |
declare complex_mod_minus_le_complex_mod [simp] |
|
882 |
||
883 |
lemma complex_mod_triangle_ineq: "cmod (x + y) <= cmod(x) + cmod(y)" |
|
14334 | 884 |
apply (rule_tac n = 1 in realpow_increasing) |
14323 | 885 |
apply (auto intro: order_trans [OF _ complex_mod_ge_zero] |
886 |
simp add: real_power_two [symmetric]) |
|
887 |
done |
|
888 |
declare complex_mod_triangle_ineq [simp] |
|
889 |
||
890 |
lemma complex_mod_triangle_ineq2: "cmod(b + a) - cmod b <= cmod a" |
|
14334 | 891 |
apply (cut_tac x1 = b and y1 = a and c = "-cmod b" |
892 |
in complex_mod_triangle_ineq [THEN add_right_mono]) |
|
14323 | 893 |
apply (simp (no_asm)) |
894 |
done |
|
895 |
declare complex_mod_triangle_ineq2 [simp] |
|
896 |
||
897 |
lemma complex_mod_diff_commute: "cmod (x - y) = cmod (y - x)" |
|
14334 | 898 |
apply (rule_tac z = x in eq_Abs_complex) |
899 |
apply (rule_tac z = y in eq_Abs_complex) |
|
900 |
apply (auto simp add: complex_diff complex_mod right_diff_distrib real_power_two left_diff_distrib add_ac mult_ac) |
|
14323 | 901 |
done |
902 |
||
903 |
lemma complex_mod_add_less: "[| cmod x < r; cmod y < s |] ==> cmod (x + y) < r + s" |
|
14334 | 904 |
by (auto intro: order_le_less_trans complex_mod_triangle_ineq) |
14323 | 905 |
|
906 |
lemma complex_mod_mult_less: "[| cmod x < r; cmod y < s |] ==> cmod (x * y) < r * s" |
|
14334 | 907 |
by (auto intro: real_mult_less_mono' simp add: complex_mod_mult) |
14323 | 908 |
|
909 |
lemma complex_mod_diff_ineq: "cmod(a) - cmod(b) <= cmod(a + b)" |
|
910 |
apply (rule linorder_cases [of "cmod(a)" "cmod (b)"]) |
|
911 |
apply auto |
|
14334 | 912 |
apply (rule order_trans [of _ 0], rule order_less_imp_le) |
913 |
apply (simp add: compare_rls, simp) |
|
14323 | 914 |
apply (simp add: compare_rls) |
915 |
apply (rule complex_mod_minus [THEN subst]) |
|
916 |
apply (rule order_trans) |
|
917 |
apply (rule_tac [2] complex_mod_triangle_ineq) |
|
918 |
apply (auto simp add: complex_add_ac) |
|
919 |
done |
|
920 |
declare complex_mod_diff_ineq [simp] |
|
921 |
||
922 |
lemma complex_Re_le_cmod: "Re z <= cmod z" |
|
14334 | 923 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 924 |
apply (auto simp add: complex_mod simp del: realpow_Suc) |
925 |
done |
|
926 |
declare complex_Re_le_cmod [simp] |
|
927 |
||
928 |
lemma complex_mod_gt_zero: "z ~= 0 ==> 0 < cmod z" |
|
14334 | 929 |
apply (cut_tac x = z in complex_mod_ge_zero) |
930 |
apply (drule order_le_imp_less_or_eq, auto) |
|
14323 | 931 |
done |
932 |
||
933 |
||
934 |
subsection{*A Few More Theorems*} |
|
935 |
||
936 |
lemma complex_mod_complexpow: "cmod(x ^ n) = cmod(x) ^ n" |
|
937 |
apply (induct_tac "n") |
|
938 |
apply (auto simp add: complex_mod_mult) |
|
939 |
done |
|
940 |
||
941 |
lemma complexpow_minus: "(-x::complex) ^ n = (if even n then (x ^ n) else -(x ^ n))" |
|
14334 | 942 |
by (induct_tac "n", auto) |
14323 | 943 |
|
944 |
lemma complex_inverse_minus: "inverse (-x) = - inverse (x::complex)" |
|
14334 | 945 |
apply (rule_tac z = x in eq_Abs_complex) |
14323 | 946 |
apply (simp (no_asm_simp) add: complex_inverse complex_minus real_power_two) |
947 |
done |
|
948 |
||
949 |
lemma complex_divide_one: "x / (1::complex) = x" |
|
950 |
apply (unfold complex_divide_def) |
|
951 |
apply (simp (no_asm)) |
|
952 |
done |
|
953 |
declare complex_divide_one [simp] |
|
954 |
||
955 |
lemma complex_mod_inverse: "cmod(inverse x) = inverse(cmod x)" |
|
956 |
apply (case_tac "x=0", simp add: COMPLEX_INVERSE_ZERO) |
|
957 |
apply (rule_tac c1 = "cmod x" in real_mult_left_cancel [THEN iffD1]) |
|
958 |
apply (auto simp add: complex_mod_mult [symmetric]) |
|
959 |
done |
|
960 |
||
961 |
lemma complex_mod_divide: |
|
962 |
"cmod(x/y) = cmod(x)/(cmod y)" |
|
963 |
apply (unfold complex_divide_def real_divide_def) |
|
964 |
apply (auto simp add: complex_mod_mult complex_mod_inverse) |
|
965 |
done |
|
966 |
||
967 |
lemma complex_inverse_divide: |
|
968 |
"inverse(x/y) = y/(x::complex)" |
|
969 |
apply (unfold complex_divide_def) |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
970 |
apply (auto simp add: inverse_mult_distrib complex_mult_commute) |
14323 | 971 |
done |
972 |
declare complex_inverse_divide [simp] |
|
973 |
||
974 |
lemma complexpow_mult: "((r::complex) * s) ^ n = (r ^ n) * (s ^ n)" |
|
975 |
apply (induct_tac "n") |
|
976 |
apply (auto simp add: complex_mult_ac) |
|
977 |
done |
|
978 |
||
979 |
||
980 |
subsection{*More Exponentiation*} |
|
981 |
||
982 |
lemma complexpow_zero: "(0::complex) ^ (Suc n) = 0" |
|
14334 | 983 |
by auto |
14323 | 984 |
declare complexpow_zero [simp] |
985 |
||
986 |
lemma complexpow_not_zero [rule_format (no_asm)]: "r ~= (0::complex) --> r ^ n ~= 0" |
|
987 |
apply (induct_tac "n") |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
988 |
apply (auto simp add: ) |
14323 | 989 |
done |
990 |
declare complexpow_not_zero [simp] |
|
991 |
declare complexpow_not_zero [intro] |
|
992 |
||
993 |
lemma complexpow_zero_zero: "r ^ n = (0::complex) ==> r = 0" |
|
14334 | 994 |
by (blast intro: ccontr dest: complexpow_not_zero) |
14323 | 995 |
|
996 |
lemma complexpow_i_squared: "ii ^ 2 = -(1::complex)" |
|
997 |
apply (unfold i_def) |
|
998 |
apply (auto simp add: complex_mult complex_one_def complex_minus numeral_2_eq_2) |
|
999 |
done |
|
1000 |
declare complexpow_i_squared [simp] |
|
1001 |
||
1002 |
lemma complex_i_not_zero: "ii ~= 0" |
|
14334 | 1003 |
by (unfold i_def complex_zero_def, auto) |
14323 | 1004 |
declare complex_i_not_zero [simp] |
1005 |
||
1006 |
lemma complex_mult_eq_zero_cancel1: "x * y ~= (0::complex) ==> x ~= 0" |
|
14334 | 1007 |
by auto |
14323 | 1008 |
|
1009 |
lemma complex_mult_eq_zero_cancel2: "x * y ~= 0 ==> y ~= (0::complex)" |
|
14334 | 1010 |
by auto |
14323 | 1011 |
|
1012 |
lemma complex_mult_not_eq_zero_iff: "(x * y ~= 0) = (x ~= 0 & y ~= (0::complex))" |
|
14334 | 1013 |
by auto |
14323 | 1014 |
declare complex_mult_not_eq_zero_iff [iff] |
1015 |
||
1016 |
lemma complexpow_inverse: "inverse ((r::complex) ^ n) = (inverse r) ^ n" |
|
1017 |
apply (induct_tac "n") |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
1018 |
apply (auto simp add: inverse_mult_distrib) |
14323 | 1019 |
done |
1020 |
||
1021 |
(*---------------------------------------------------------------------------*) |
|
1022 |
(* sgn *) |
|
1023 |
(*---------------------------------------------------------------------------*) |
|
1024 |
||
1025 |
lemma sgn_zero: "sgn 0 = 0" |
|
1026 |
||
1027 |
apply (unfold sgn_def) |
|
1028 |
apply (simp (no_asm)) |
|
1029 |
done |
|
1030 |
declare sgn_zero [simp] |
|
1031 |
||
1032 |
lemma sgn_one: "sgn 1 = 1" |
|
1033 |
apply (unfold sgn_def) |
|
1034 |
apply (simp (no_asm)) |
|
1035 |
done |
|
1036 |
declare sgn_one [simp] |
|
1037 |
||
1038 |
lemma sgn_minus: "sgn (-z) = - sgn(z)" |
|
14334 | 1039 |
by (unfold sgn_def, auto) |
14323 | 1040 |
|
1041 |
lemma sgn_eq: |
|
1042 |
"sgn z = z / complex_of_real (cmod z)" |
|
1043 |
apply (unfold sgn_def) |
|
1044 |
apply (simp (no_asm)) |
|
1045 |
done |
|
1046 |
||
1047 |
lemma complex_split: "EX x y. z = complex_of_real(x) + ii * complex_of_real(y)" |
|
14334 | 1048 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 1049 |
apply (auto simp add: complex_of_real_def i_def complex_mult complex_add) |
1050 |
done |
|
1051 |
||
1052 |
lemma Re_complex_i: "Re(complex_of_real(x) + ii * complex_of_real(y)) = x" |
|
14334 | 1053 |
by (auto simp add: complex_of_real_def i_def complex_mult complex_add) |
14323 | 1054 |
declare Re_complex_i [simp] |
1055 |
||
1056 |
lemma Im_complex_i: "Im(complex_of_real(x) + ii * complex_of_real(y)) = y" |
|
14334 | 1057 |
by (auto simp add: complex_of_real_def i_def complex_mult complex_add) |
14323 | 1058 |
declare Im_complex_i [simp] |
1059 |
||
1060 |
lemma i_mult_eq: "ii * ii = complex_of_real (-1)" |
|
1061 |
apply (unfold i_def complex_of_real_def) |
|
1062 |
apply (auto simp add: complex_mult complex_add) |
|
1063 |
done |
|
1064 |
||
1065 |
lemma i_mult_eq2: "ii * ii = -(1::complex)" |
|
1066 |
apply (unfold i_def complex_one_def) |
|
1067 |
apply (simp (no_asm) add: complex_mult complex_minus) |
|
1068 |
done |
|
1069 |
declare i_mult_eq2 [simp] |
|
1070 |
||
1071 |
lemma cmod_i: "cmod (complex_of_real(x) + ii * complex_of_real(y)) = |
|
1072 |
sqrt (x ^ 2 + y ^ 2)" |
|
1073 |
apply (auto simp add: complex_mult complex_add i_def complex_of_real_def cmod_def) |
|
1074 |
done |
|
1075 |
||
1076 |
lemma complex_eq_Re_eq: |
|
1077 |
"complex_of_real xa + ii * complex_of_real ya = |
|
1078 |
complex_of_real xb + ii * complex_of_real yb |
|
1079 |
==> xa = xb" |
|
1080 |
apply (unfold complex_of_real_def i_def) |
|
1081 |
apply (auto simp add: complex_mult complex_add) |
|
1082 |
done |
|
1083 |
||
1084 |
lemma complex_eq_Im_eq: |
|
1085 |
"complex_of_real xa + ii * complex_of_real ya = |
|
1086 |
complex_of_real xb + ii * complex_of_real yb |
|
1087 |
==> ya = yb" |
|
1088 |
apply (unfold complex_of_real_def i_def) |
|
1089 |
apply (auto simp add: complex_mult complex_add) |
|
1090 |
done |
|
1091 |
||
1092 |
lemma complex_eq_cancel_iff: "(complex_of_real xa + ii * complex_of_real ya = |
|
1093 |
complex_of_real xb + ii * complex_of_real yb) = ((xa = xb) & (ya = yb))" |
|
1094 |
apply (auto intro: complex_eq_Im_eq complex_eq_Re_eq) |
|
1095 |
done |
|
1096 |
declare complex_eq_cancel_iff [iff] |
|
1097 |
||
1098 |
lemma complex_eq_cancel_iffA: "(complex_of_real xa + complex_of_real ya * ii = |
|
1099 |
complex_of_real xb + complex_of_real yb * ii ) = ((xa = xb) & (ya = yb))" |
|
1100 |
apply (auto simp add: complex_mult_commute) |
|
1101 |
done |
|
1102 |
declare complex_eq_cancel_iffA [iff] |
|
1103 |
||
1104 |
lemma complex_eq_cancel_iffB: "(complex_of_real xa + complex_of_real ya * ii = |
|
1105 |
complex_of_real xb + ii * complex_of_real yb) = ((xa = xb) & (ya = yb))" |
|
1106 |
apply (auto simp add: complex_mult_commute) |
|
1107 |
done |
|
1108 |
declare complex_eq_cancel_iffB [iff] |
|
1109 |
||
1110 |
lemma complex_eq_cancel_iffC: "(complex_of_real xa + ii * complex_of_real ya = |
|
1111 |
complex_of_real xb + complex_of_real yb * ii) = ((xa = xb) & (ya = yb))" |
|
1112 |
apply (auto simp add: complex_mult_commute) |
|
1113 |
done |
|
1114 |
declare complex_eq_cancel_iffC [iff] |
|
1115 |
||
1116 |
lemma complex_eq_cancel_iff2: "(complex_of_real x + ii * complex_of_real y = |
|
1117 |
complex_of_real xa) = (x = xa & y = 0)" |
|
14334 | 1118 |
apply (cut_tac xa = x and ya = y and xb = xa and yb = 0 in complex_eq_cancel_iff) |
14323 | 1119 |
apply (simp del: complex_eq_cancel_iff) |
1120 |
done |
|
1121 |
declare complex_eq_cancel_iff2 [simp] |
|
1122 |
||
1123 |
lemma complex_eq_cancel_iff2a: "(complex_of_real x + complex_of_real y * ii = |
|
1124 |
complex_of_real xa) = (x = xa & y = 0)" |
|
1125 |
apply (auto simp add: complex_mult_commute) |
|
1126 |
done |
|
1127 |
declare complex_eq_cancel_iff2a [simp] |
|
1128 |
||
1129 |
lemma complex_eq_cancel_iff3: "(complex_of_real x + ii * complex_of_real y = |
|
1130 |
ii * complex_of_real ya) = (x = 0 & y = ya)" |
|
14334 | 1131 |
apply (cut_tac xa = x and ya = y and xb = 0 and yb = ya in complex_eq_cancel_iff) |
14323 | 1132 |
apply (simp del: complex_eq_cancel_iff) |
1133 |
done |
|
1134 |
declare complex_eq_cancel_iff3 [simp] |
|
1135 |
||
1136 |
lemma complex_eq_cancel_iff3a: "(complex_of_real x + complex_of_real y * ii = |
|
1137 |
ii * complex_of_real ya) = (x = 0 & y = ya)" |
|
1138 |
apply (auto simp add: complex_mult_commute) |
|
1139 |
done |
|
1140 |
declare complex_eq_cancel_iff3a [simp] |
|
1141 |
||
1142 |
lemma complex_split_Re_zero: |
|
1143 |
"complex_of_real x + ii * complex_of_real y = 0 |
|
1144 |
==> x = 0" |
|
1145 |
apply (unfold complex_of_real_def i_def complex_zero_def) |
|
1146 |
apply (auto simp add: complex_mult complex_add) |
|
1147 |
done |
|
1148 |
||
1149 |
lemma complex_split_Im_zero: |
|
1150 |
"complex_of_real x + ii * complex_of_real y = 0 |
|
1151 |
==> y = 0" |
|
1152 |
apply (unfold complex_of_real_def i_def complex_zero_def) |
|
1153 |
apply (auto simp add: complex_mult complex_add) |
|
1154 |
done |
|
1155 |
||
1156 |
lemma Re_sgn: |
|
1157 |
"Re(sgn z) = Re(z)/cmod z" |
|
1158 |
apply (unfold sgn_def complex_divide_def) |
|
14334 | 1159 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 1160 |
apply (auto simp add: complex_of_real_inverse [symmetric]) |
1161 |
apply (auto simp add: complex_of_real_def complex_mult real_divide_def) |
|
1162 |
done |
|
1163 |
declare Re_sgn [simp] |
|
1164 |
||
1165 |
lemma Im_sgn: |
|
1166 |
"Im(sgn z) = Im(z)/cmod z" |
|
1167 |
apply (unfold sgn_def complex_divide_def) |
|
14334 | 1168 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 1169 |
apply (auto simp add: complex_of_real_inverse [symmetric]) |
1170 |
apply (auto simp add: complex_of_real_def complex_mult real_divide_def) |
|
1171 |
done |
|
1172 |
declare Im_sgn [simp] |
|
1173 |
||
1174 |
lemma complex_inverse_complex_split: |
|
1175 |
"inverse(complex_of_real x + ii * complex_of_real y) = |
|
1176 |
complex_of_real(x/(x ^ 2 + y ^ 2)) - |
|
1177 |
ii * complex_of_real(y/(x ^ 2 + y ^ 2))" |
|
1178 |
apply (unfold complex_of_real_def i_def) |
|
1179 |
apply (auto simp add: complex_mult complex_add complex_diff_def complex_minus complex_inverse real_divide_def) |
|
1180 |
done |
|
1181 |
||
1182 |
(*----------------------------------------------------------------------------*) |
|
1183 |
(* Many of the theorems below need to be moved elsewhere e.g. Transc. Also *) |
|
1184 |
(* many of the theorems are not used - so should they be kept? *) |
|
1185 |
(*----------------------------------------------------------------------------*) |
|
1186 |
||
1187 |
lemma Re_mult_i_eq: |
|
1188 |
"Re (ii * complex_of_real y) = 0" |
|
1189 |
apply (unfold i_def complex_of_real_def) |
|
1190 |
apply (auto simp add: complex_mult) |
|
1191 |
done |
|
1192 |
declare Re_mult_i_eq [simp] |
|
1193 |
||
1194 |
lemma Im_mult_i_eq: |
|
1195 |
"Im (ii * complex_of_real y) = y" |
|
1196 |
apply (unfold i_def complex_of_real_def) |
|
1197 |
apply (auto simp add: complex_mult) |
|
1198 |
done |
|
1199 |
declare Im_mult_i_eq [simp] |
|
1200 |
||
1201 |
lemma complex_mod_mult_i: |
|
1202 |
"cmod (ii * complex_of_real y) = abs y" |
|
1203 |
apply (unfold i_def complex_of_real_def) |
|
1204 |
apply (auto simp add: complex_mult complex_mod real_power_two) |
|
1205 |
done |
|
1206 |
declare complex_mod_mult_i [simp] |
|
1207 |
||
1208 |
lemma cos_arg_i_mult_zero: |
|
1209 |
"0 < y ==> cos (arg(ii * complex_of_real y)) = 0" |
|
1210 |
apply (unfold arg_def) |
|
1211 |
apply (auto simp add: abs_eqI2) |
|
14334 | 1212 |
apply (rule_tac a = "pi/2" in someI2, auto) |
1213 |
apply (rule order_less_trans [of _ 0], auto) |
|
14323 | 1214 |
done |
1215 |
declare cos_arg_i_mult_zero [simp] |
|
1216 |
||
1217 |
lemma cos_arg_i_mult_zero2: |
|
1218 |
"y < 0 ==> cos (arg(ii * complex_of_real y)) = 0" |
|
1219 |
apply (unfold arg_def) |
|
1220 |
apply (auto simp add: abs_minus_eqI2) |
|
14334 | 1221 |
apply (rule_tac a = "- pi/2" in someI2, auto) |
1222 |
apply (rule order_trans [of _ 0], auto) |
|
14323 | 1223 |
done |
1224 |
declare cos_arg_i_mult_zero2 [simp] |
|
1225 |
||
1226 |
lemma complex_of_real_not_zero_iff: |
|
1227 |
"(complex_of_real y ~= 0) = (y ~= 0)" |
|
14334 | 1228 |
apply (unfold complex_zero_def complex_of_real_def, auto) |
14323 | 1229 |
done |
1230 |
declare complex_of_real_not_zero_iff [simp] |
|
1231 |
||
1232 |
lemma complex_of_real_zero_iff: "(complex_of_real y = 0) = (y = 0)" |
|
1233 |
apply auto |
|
14334 | 1234 |
apply (rule ccontr, drule complex_of_real_not_zero_iff [THEN iffD2], simp) |
14323 | 1235 |
done |
1236 |
declare complex_of_real_zero_iff [simp] |
|
1237 |
||
1238 |
lemma cos_arg_i_mult_zero3: "y ~= 0 ==> cos (arg(ii * complex_of_real y)) = 0" |
|
14334 | 1239 |
by (cut_tac x = y and y = 0 in linorder_less_linear, auto) |
14323 | 1240 |
declare cos_arg_i_mult_zero3 [simp] |
1241 |
||
1242 |
||
1243 |
subsection{*Finally! Polar Form for Complex Numbers*} |
|
1244 |
||
1245 |
lemma complex_split_polar: "EX r a. z = complex_of_real r * |
|
1246 |
(complex_of_real(cos a) + ii * complex_of_real(sin a))" |
|
14334 | 1247 |
apply (cut_tac z = z in complex_split) |
14323 | 1248 |
apply (auto simp add: polar_Ex complex_add_mult_distrib2 complex_of_real_mult complex_mult_ac) |
1249 |
done |
|
1250 |
||
1251 |
lemma rcis_Ex: "EX r a. z = rcis r a" |
|
1252 |
apply (unfold rcis_def cis_def) |
|
1253 |
apply (rule complex_split_polar) |
|
1254 |
done |
|
1255 |
||
1256 |
lemma Re_complex_polar: "Re(complex_of_real r * |
|
1257 |
(complex_of_real(cos a) + ii * complex_of_real(sin a))) = r * cos a" |
|
1258 |
apply (auto simp add: complex_add_mult_distrib2 complex_of_real_mult complex_mult_ac) |
|
1259 |
done |
|
1260 |
declare Re_complex_polar [simp] |
|
1261 |
||
1262 |
lemma Re_rcis: "Re(rcis r a) = r * cos a" |
|
14334 | 1263 |
by (unfold rcis_def cis_def, auto) |
14323 | 1264 |
declare Re_rcis [simp] |
1265 |
||
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
1266 |
lemma Im_complex_polar [simp]: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
1267 |
"Im(complex_of_real r * |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
1268 |
(complex_of_real(cos a) + ii * complex_of_real(sin a))) = |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
1269 |
r * sin a" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
1270 |
by (auto simp add: complex_add_mult_distrib2 complex_of_real_mult mult_ac) |
14323 | 1271 |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
1272 |
lemma Im_rcis [simp]: "Im(rcis r a) = r * sin a" |
14334 | 1273 |
by (unfold rcis_def cis_def, auto) |
14323 | 1274 |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
1275 |
lemma complex_mod_complex_polar [simp]: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
1276 |
"cmod (complex_of_real r * |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
1277 |
(complex_of_real(cos a) + ii * complex_of_real(sin a))) = |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
1278 |
abs r" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
1279 |
by (auto simp add: complex_add_mult_distrib2 cmod_i complex_of_real_mult |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
1280 |
right_distrib [symmetric] power_mult_distrib mult_ac |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
1281 |
simp del: realpow_Suc) |
14323 | 1282 |
|
1283 |
lemma complex_mod_rcis: "cmod(rcis r a) = abs r" |
|
14334 | 1284 |
by (unfold rcis_def cis_def, auto) |
14323 | 1285 |
declare complex_mod_rcis [simp] |
1286 |
||
1287 |
lemma complex_mod_sqrt_Re_mult_cnj: "cmod z = sqrt (Re (z * cnj z))" |
|
1288 |
apply (unfold cmod_def) |
|
1289 |
apply (rule real_sqrt_eq_iff [THEN iffD2]) |
|
1290 |
apply (auto simp add: complex_mult_cnj) |
|
1291 |
done |
|
1292 |
||
1293 |
lemma complex_Re_cnj: "Re(cnj z) = Re z" |
|
14334 | 1294 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 1295 |
apply (auto simp add: complex_cnj) |
1296 |
done |
|
1297 |
declare complex_Re_cnj [simp] |
|
1298 |
||
1299 |
lemma complex_Im_cnj: "Im(cnj z) = - Im z" |
|
14334 | 1300 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 1301 |
apply (auto simp add: complex_cnj) |
1302 |
done |
|
1303 |
declare complex_Im_cnj [simp] |
|
1304 |
||
1305 |
lemma complex_In_mult_cnj_zero: "Im (z * cnj z) = 0" |
|
14334 | 1306 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 1307 |
apply (auto simp add: complex_cnj complex_mult) |
1308 |
done |
|
1309 |
declare complex_In_mult_cnj_zero [simp] |
|
1310 |
||
1311 |
lemma complex_Re_mult: "[| Im w = 0; Im z = 0 |] ==> Re(w * z) = Re(w) * Re(z)" |
|
14334 | 1312 |
apply (rule_tac z = z in eq_Abs_complex) |
1313 |
apply (rule_tac z = w in eq_Abs_complex) |
|
14323 | 1314 |
apply (auto simp add: complex_mult) |
1315 |
done |
|
1316 |
||
1317 |
lemma complex_Re_mult_complex_of_real: "Re (z * complex_of_real c) = Re(z) * c" |
|
1318 |
apply (unfold complex_of_real_def) |
|
14334 | 1319 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 1320 |
apply (auto simp add: complex_mult) |
1321 |
done |
|
1322 |
declare complex_Re_mult_complex_of_real [simp] |
|
1323 |
||
1324 |
lemma complex_Im_mult_complex_of_real: "Im (z * complex_of_real c) = Im(z) * c" |
|
1325 |
apply (unfold complex_of_real_def) |
|
14334 | 1326 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 1327 |
apply (auto simp add: complex_mult) |
1328 |
done |
|
1329 |
declare complex_Im_mult_complex_of_real [simp] |
|
1330 |
||
1331 |
lemma complex_Re_mult_complex_of_real2: "Re (complex_of_real c * z) = c * Re(z)" |
|
1332 |
apply (unfold complex_of_real_def) |
|
14334 | 1333 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 1334 |
apply (auto simp add: complex_mult) |
1335 |
done |
|
1336 |
declare complex_Re_mult_complex_of_real2 [simp] |
|
1337 |
||
1338 |
lemma complex_Im_mult_complex_of_real2: "Im (complex_of_real c * z) = c * Im(z)" |
|
1339 |
apply (unfold complex_of_real_def) |
|
14334 | 1340 |
apply (rule_tac z = z in eq_Abs_complex) |
14323 | 1341 |
apply (auto simp add: complex_mult) |
1342 |
done |
|
1343 |
declare complex_Im_mult_complex_of_real2 [simp] |
|
1344 |
||
1345 |
(*---------------------------------------------------------------------------*) |
|
1346 |
(* (r1 * cis a) * (r2 * cis b) = r1 * r2 * cis (a + b) *) |
|
1347 |
(*---------------------------------------------------------------------------*) |
|
1348 |
||
1349 |
lemma cis_rcis_eq: "cis a = rcis 1 a" |
|
1350 |
apply (unfold rcis_def) |
|
1351 |
apply (simp (no_asm)) |
|
1352 |
done |
|
1353 |
||
1354 |
lemma rcis_mult: |
|
1355 |
"rcis r1 a * rcis r2 b = rcis (r1*r2) (a + b)" |
|
1356 |
apply (unfold rcis_def cis_def) |
|
1357 |
apply (auto simp add: cos_add sin_add complex_add_mult_distrib2 complex_add_mult_distrib complex_mult_ac complex_add_ac) |
|
1358 |
apply (auto simp add: complex_add_mult_distrib2 [symmetric] complex_mult_assoc [symmetric] complex_of_real_mult complex_of_real_add complex_add_assoc [symmetric] i_mult_eq simp del: i_mult_eq2) |
|
1359 |
apply (auto simp add: complex_add_ac) |
|
14334 | 1360 |
apply (auto simp add: complex_add_assoc [symmetric] complex_of_real_add right_distrib real_diff_def mult_ac add_ac) |
14323 | 1361 |
done |
1362 |
||
1363 |
lemma cis_mult: "cis a * cis b = cis (a + b)" |
|
1364 |
apply (simp (no_asm) add: cis_rcis_eq rcis_mult) |
|
1365 |
done |
|
1366 |
||
1367 |
lemma cis_zero: "cis 0 = 1" |
|
14334 | 1368 |
by (unfold cis_def, auto) |
14323 | 1369 |
declare cis_zero [simp] |
1370 |
||
1371 |
lemma cis_zero2: "cis 0 = complex_of_real 1" |
|
14334 | 1372 |
by (unfold cis_def, auto) |
14323 | 1373 |
declare cis_zero2 [simp] |
1374 |
||
1375 |
lemma rcis_zero_mod: "rcis 0 a = 0" |
|
1376 |
apply (unfold rcis_def) |
|
1377 |
apply (simp (no_asm)) |
|
1378 |
done |
|
1379 |
declare rcis_zero_mod [simp] |
|
1380 |
||
1381 |
lemma rcis_zero_arg: "rcis r 0 = complex_of_real r" |
|
1382 |
apply (unfold rcis_def) |
|
1383 |
apply (simp (no_asm)) |
|
1384 |
done |
|
1385 |
declare rcis_zero_arg [simp] |
|
1386 |
||
1387 |
lemma complex_of_real_minus_one: |
|
1388 |
"complex_of_real (-(1::real)) = -(1::complex)" |
|
1389 |
apply (unfold complex_of_real_def complex_one_def) |
|
1390 |
apply (simp (no_asm) add: complex_minus) |
|
1391 |
done |
|
1392 |
||
1393 |
lemma complex_i_mult_minus: "ii * (ii * x) = - x" |
|
1394 |
apply (simp (no_asm) add: complex_mult_assoc [symmetric]) |
|
1395 |
done |
|
1396 |
declare complex_i_mult_minus [simp] |
|
1397 |
||
1398 |
lemma complex_i_mult_minus2: "ii * ii * x = - x" |
|
1399 |
apply (simp (no_asm)) |
|
1400 |
done |
|
1401 |
declare complex_i_mult_minus2 [simp] |
|
1402 |
||
1403 |
lemma cis_real_of_nat_Suc_mult: |
|
1404 |
"cis (real (Suc n) * a) = cis a * cis (real n * a)" |
|
1405 |
apply (unfold cis_def) |
|
14334 | 1406 |
apply (auto simp add: real_of_nat_Suc left_distrib cos_add sin_add complex_add_mult_distrib complex_add_mult_distrib2 complex_of_real_add complex_of_real_mult complex_mult_ac complex_add_ac) |
14323 | 1407 |
apply (auto simp add: complex_add_mult_distrib2 [symmetric] complex_mult_assoc [symmetric] i_mult_eq complex_of_real_mult complex_of_real_add complex_add_assoc [symmetric] complex_of_real_minus [symmetric] real_diff_def mult_ac simp del: i_mult_eq2) |
1408 |
done |
|
1409 |
||
1410 |
lemma DeMoivre: "(cis a) ^ n = cis (real n * a)" |
|
1411 |
apply (induct_tac "n") |
|
1412 |
apply (auto simp add: cis_real_of_nat_Suc_mult) |
|
1413 |
done |
|
1414 |
||
1415 |
lemma DeMoivre2: |
|
1416 |
"(rcis r a) ^ n = rcis (r ^ n) (real n * a)" |
|
1417 |
apply (unfold rcis_def) |
|
1418 |
apply (auto simp add: complexpow_mult DeMoivre complex_of_real_pow) |
|
1419 |
done |
|
1420 |
||
1421 |
lemma cis_inverse: "inverse(cis a) = cis (-a)" |
|
1422 |
apply (unfold cis_def) |
|
1423 |
apply (auto simp add: complex_inverse_complex_split complex_of_real_minus complex_diff_def) |
|
1424 |
done |
|
1425 |
declare cis_inverse [simp] |
|
1426 |
||
1427 |
lemma rcis_inverse: "inverse(rcis r a) = rcis (1/r) (-a)" |
|
1428 |
apply (case_tac "r=0") |
|
1429 |
apply (simp (no_asm_simp) add: DIVISION_BY_ZERO COMPLEX_INVERSE_ZERO) |
|
14334 | 1430 |
apply (auto simp add: complex_inverse_complex_split complex_add_mult_distrib2 complex_of_real_mult rcis_def cis_def real_power_two complex_mult_ac mult_ac) |
1431 |
apply (auto simp add: right_distrib [symmetric] complex_of_real_minus complex_diff_def) |
|
14323 | 1432 |
done |
1433 |
||
1434 |
lemma cis_divide: "cis a / cis b = cis (a - b)" |
|
1435 |
apply (unfold complex_divide_def) |
|
1436 |
apply (auto simp add: cis_mult real_diff_def) |
|
1437 |
done |
|
1438 |
||
1439 |
lemma rcis_divide: |
|
1440 |
"rcis r1 a / rcis r2 b = rcis (r1/r2) (a - b)" |
|
1441 |
apply (unfold complex_divide_def) |
|
1442 |
apply (case_tac "r2=0") |
|
1443 |
apply (simp (no_asm_simp) add: DIVISION_BY_ZERO COMPLEX_INVERSE_ZERO) |
|
1444 |
apply (auto simp add: rcis_inverse rcis_mult real_diff_def) |
|
1445 |
done |
|
1446 |
||
1447 |
lemma Re_cis: "Re(cis a) = cos a" |
|
14334 | 1448 |
by (unfold cis_def, auto) |
14323 | 1449 |
declare Re_cis [simp] |
1450 |
||
1451 |
lemma Im_cis: "Im(cis a) = sin a" |
|
14334 | 1452 |
by (unfold cis_def, auto) |
14323 | 1453 |
declare Im_cis [simp] |
1454 |
||
1455 |
lemma cos_n_Re_cis_pow_n: "cos (real n * a) = Re(cis a ^ n)" |
|
14334 | 1456 |
by (auto simp add: DeMoivre) |
14323 | 1457 |
|
1458 |
lemma sin_n_Im_cis_pow_n: "sin (real n * a) = Im(cis a ^ n)" |
|
14334 | 1459 |
by (auto simp add: DeMoivre) |
14323 | 1460 |
|
1461 |
lemma expi_Im_split: |
|
1462 |
"expi (ii * complex_of_real y) = |
|
1463 |
complex_of_real (cos y) + ii * complex_of_real (sin y)" |
|
14334 | 1464 |
apply (unfold expi_def cis_def, auto) |
14323 | 1465 |
done |
1466 |
||
1467 |
lemma expi_Im_cis: |
|
1468 |
"expi (ii * complex_of_real y) = cis y" |
|
14334 | 1469 |
apply (unfold expi_def, auto) |
14323 | 1470 |
done |
1471 |
||
1472 |
lemma expi_add: "expi(a + b) = expi(a) * expi(b)" |
|
1473 |
apply (unfold expi_def) |
|
1474 |
apply (auto simp add: complex_Re_add exp_add complex_Im_add cis_mult [symmetric] complex_of_real_mult complex_mult_ac) |
|
1475 |
done |
|
1476 |
||
1477 |
lemma expi_complex_split: |
|
1478 |
"expi(complex_of_real x + ii * complex_of_real y) = |
|
1479 |
complex_of_real (exp(x)) * cis y" |
|
14334 | 1480 |
apply (unfold expi_def, auto) |
14323 | 1481 |
done |
1482 |
||
1483 |
lemma expi_zero: "expi (0::complex) = 1" |
|
14334 | 1484 |
by (unfold expi_def, auto) |
14323 | 1485 |
declare expi_zero [simp] |
1486 |
||
1487 |
lemma complex_Re_mult_eq: "Re (w * z) = Re w * Re z - Im w * Im z" |
|
14334 | 1488 |
apply (rule_tac z = z in eq_Abs_complex) |
1489 |
apply (rule_tac z = w in eq_Abs_complex) |
|
14323 | 1490 |
apply (auto simp add: complex_mult) |
1491 |
done |
|
1492 |
||
1493 |
lemma complex_Im_mult_eq: |
|
1494 |
"Im (w * z) = Re w * Im z + Im w * Re z" |
|
14334 | 1495 |
apply (rule_tac z = z in eq_Abs_complex) |
1496 |
apply (rule_tac z = w in eq_Abs_complex) |
|
14323 | 1497 |
apply (auto simp add: complex_mult) |
1498 |
done |
|
1499 |
||
1500 |
lemma complex_expi_Ex: |
|
1501 |
"EX a r. z = complex_of_real r * expi a" |
|
14334 | 1502 |
apply (cut_tac z = z in rcis_Ex) |
14323 | 1503 |
apply (auto simp add: expi_def rcis_def complex_mult_assoc [symmetric] complex_of_real_mult) |
14334 | 1504 |
apply (rule_tac x = "ii * complex_of_real a" in exI, auto) |
14323 | 1505 |
done |
1506 |
||
1507 |
||
1508 |
(**** |
|
1509 |
Goal "[| - pi < a; a <= pi |] ==> (-pi < a & a <= 0) | (0 <= a & a <= pi)" |
|
14334 | 1510 |
by Auto_tac |
14323 | 1511 |
qed "lemma_split_interval"; |
1512 |
||
1513 |
Goalw [arg_def] |
|
1514 |
"[| r ~= 0; - pi < a; a <= pi |] \ |
|
1515 |
\ ==> arg(complex_of_real r * \ |
|
1516 |
\ (complex_of_real(cos a) + ii * complex_of_real(sin a))) = a"; |
|
14334 | 1517 |
by Auto_tac |
14323 | 1518 |
by (cut_inst_tac [("x","0"),("y","r")] linorder_less_linear 1); |
1519 |
by (auto_tac (claset(),simpset() addsimps (map (full_rename_numerals thy) |
|
1520 |
[rabs_eqI2,rabs_minus_eqI2,real_minus_rinv]) [real_divide_def, |
|
14334 | 1521 |
minus_mult_right RS sym] mult_ac)); |
14323 | 1522 |
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc RS sym])); |
14334 | 1523 |
by (dtac lemma_split_interval 1 THEN safe) |
14323 | 1524 |
****) |
1525 |
||
1526 |
||
1527 |
ML |
|
1528 |
{* |
|
1529 |
val complex_zero_def = thm"complex_zero_def"; |
|
1530 |
val complex_one_def = thm"complex_one_def"; |
|
1531 |
val complex_minus_def = thm"complex_minus_def"; |
|
1532 |
val complex_diff_def = thm"complex_diff_def"; |
|
1533 |
val complex_divide_def = thm"complex_divide_def"; |
|
1534 |
val complex_mult_def = thm"complex_mult_def"; |
|
1535 |
val complex_add_def = thm"complex_add_def"; |
|
1536 |
val complex_of_real_def = thm"complex_of_real_def"; |
|
1537 |
val i_def = thm"i_def"; |
|
1538 |
val expi_def = thm"expi_def"; |
|
1539 |
val cis_def = thm"cis_def"; |
|
1540 |
val rcis_def = thm"rcis_def"; |
|
1541 |
val cmod_def = thm"cmod_def"; |
|
1542 |
val cnj_def = thm"cnj_def"; |
|
1543 |
val sgn_def = thm"sgn_def"; |
|
1544 |
val arg_def = thm"arg_def"; |
|
1545 |
val complexpow_0 = thm"complexpow_0"; |
|
1546 |
val complexpow_Suc = thm"complexpow_Suc"; |
|
1547 |
||
1548 |
val inj_Rep_complex = thm"inj_Rep_complex"; |
|
1549 |
val inj_Abs_complex = thm"inj_Abs_complex"; |
|
1550 |
val Abs_complex_cancel_iff = thm"Abs_complex_cancel_iff"; |
|
1551 |
val pair_mem_complex = thm"pair_mem_complex"; |
|
1552 |
val Abs_complex_inverse2 = thm"Abs_complex_inverse2"; |
|
1553 |
val eq_Abs_complex = thm"eq_Abs_complex"; |
|
1554 |
val Re = thm"Re"; |
|
1555 |
val Im = thm"Im"; |
|
1556 |
val Abs_complex_cancel = thm"Abs_complex_cancel"; |
|
1557 |
val complex_Re_Im_cancel_iff = thm"complex_Re_Im_cancel_iff"; |
|
1558 |
val complex_Re_zero = thm"complex_Re_zero"; |
|
1559 |
val complex_Im_zero = thm"complex_Im_zero"; |
|
1560 |
val complex_Re_one = thm"complex_Re_one"; |
|
1561 |
val complex_Im_one = thm"complex_Im_one"; |
|
1562 |
val complex_Re_i = thm"complex_Re_i"; |
|
1563 |
val complex_Im_i = thm"complex_Im_i"; |
|
1564 |
val Re_complex_of_real_zero = thm"Re_complex_of_real_zero"; |
|
1565 |
val Im_complex_of_real_zero = thm"Im_complex_of_real_zero"; |
|
1566 |
val Re_complex_of_real_one = thm"Re_complex_of_real_one"; |
|
1567 |
val Im_complex_of_real_one = thm"Im_complex_of_real_one"; |
|
1568 |
val Re_complex_of_real = thm"Re_complex_of_real"; |
|
1569 |
val Im_complex_of_real = thm"Im_complex_of_real"; |
|
1570 |
val complex_minus = thm"complex_minus"; |
|
1571 |
val complex_Re_minus = thm"complex_Re_minus"; |
|
1572 |
val complex_Im_minus = thm"complex_Im_minus"; |
|
1573 |
val complex_minus_minus = thm"complex_minus_minus"; |
|
1574 |
val inj_complex_minus = thm"inj_complex_minus"; |
|
1575 |
val complex_minus_zero = thm"complex_minus_zero"; |
|
1576 |
val complex_minus_zero_iff = thm"complex_minus_zero_iff"; |
|
1577 |
val complex_minus_zero_iff2 = thm"complex_minus_zero_iff2"; |
|
1578 |
val complex_minus_not_zero_iff = thm"complex_minus_not_zero_iff"; |
|
1579 |
val complex_add = thm"complex_add"; |
|
1580 |
val complex_Re_add = thm"complex_Re_add"; |
|
1581 |
val complex_Im_add = thm"complex_Im_add"; |
|
1582 |
val complex_add_commute = thm"complex_add_commute"; |
|
1583 |
val complex_add_assoc = thm"complex_add_assoc"; |
|
1584 |
val complex_add_left_commute = thm"complex_add_left_commute"; |
|
1585 |
val complex_add_zero_left = thm"complex_add_zero_left"; |
|
1586 |
val complex_add_zero_right = thm"complex_add_zero_right"; |
|
1587 |
val complex_add_minus_right_zero = thm"complex_add_minus_right_zero"; |
|
1588 |
val complex_add_minus_cancel = thm"complex_add_minus_cancel"; |
|
1589 |
val complex_minus_add_cancel = thm"complex_minus_add_cancel"; |
|
1590 |
val complex_add_minus_eq_minus = thm"complex_add_minus_eq_minus"; |
|
1591 |
val complex_minus_add_distrib = thm"complex_minus_add_distrib"; |
|
1592 |
val complex_add_left_cancel = thm"complex_add_left_cancel"; |
|
1593 |
val complex_add_right_cancel = thm"complex_add_right_cancel"; |
|
1594 |
val complex_eq_minus_iff = thm"complex_eq_minus_iff"; |
|
1595 |
val complex_eq_minus_iff2 = thm"complex_eq_minus_iff2"; |
|
1596 |
val complex_diff_0 = thm"complex_diff_0"; |
|
1597 |
val complex_diff_0_right = thm"complex_diff_0_right"; |
|
1598 |
val complex_diff_self = thm"complex_diff_self"; |
|
1599 |
val complex_diff = thm"complex_diff"; |
|
1600 |
val complex_diff_eq_eq = thm"complex_diff_eq_eq"; |
|
1601 |
val complex_mult = thm"complex_mult"; |
|
1602 |
val complex_mult_commute = thm"complex_mult_commute"; |
|
1603 |
val complex_mult_assoc = thm"complex_mult_assoc"; |
|
1604 |
val complex_mult_left_commute = thm"complex_mult_left_commute"; |
|
1605 |
val complex_mult_one_left = thm"complex_mult_one_left"; |
|
1606 |
val complex_mult_one_right = thm"complex_mult_one_right"; |
|
1607 |
val complex_mult_zero_left = thm"complex_mult_zero_left"; |
|
1608 |
val complex_mult_zero_right = thm"complex_mult_zero_right"; |
|
1609 |
val complex_divide_zero = thm"complex_divide_zero"; |
|
1610 |
val complex_minus_mult_eq1 = thm"complex_minus_mult_eq1"; |
|
1611 |
val complex_minus_mult_eq2 = thm"complex_minus_mult_eq2"; |
|
1612 |
val complex_minus_mult_commute = thm"complex_minus_mult_commute"; |
|
1613 |
val complex_add_mult_distrib = thm"complex_add_mult_distrib"; |
|
1614 |
val complex_add_mult_distrib2 = thm"complex_add_mult_distrib2"; |
|
1615 |
val complex_zero_not_eq_one = thm"complex_zero_not_eq_one"; |
|
1616 |
val complex_inverse = thm"complex_inverse"; |
|
1617 |
val COMPLEX_INVERSE_ZERO = thm"COMPLEX_INVERSE_ZERO"; |
|
1618 |
val COMPLEX_DIVISION_BY_ZERO = thm"COMPLEX_DIVISION_BY_ZERO"; |
|
1619 |
val complex_mult_inv_left = thm"complex_mult_inv_left"; |
|
1620 |
val complex_mult_inv_right = thm"complex_mult_inv_right"; |
|
1621 |
val inj_complex_of_real = thm"inj_complex_of_real"; |
|
1622 |
val complex_of_real_one = thm"complex_of_real_one"; |
|
1623 |
val complex_of_real_zero = thm"complex_of_real_zero"; |
|
1624 |
val complex_of_real_eq_iff = thm"complex_of_real_eq_iff"; |
|
1625 |
val complex_of_real_minus = thm"complex_of_real_minus"; |
|
1626 |
val complex_of_real_inverse = thm"complex_of_real_inverse"; |
|
1627 |
val complex_of_real_add = thm"complex_of_real_add"; |
|
1628 |
val complex_of_real_diff = thm"complex_of_real_diff"; |
|
1629 |
val complex_of_real_mult = thm"complex_of_real_mult"; |
|
1630 |
val complex_of_real_divide = thm"complex_of_real_divide"; |
|
1631 |
val complex_of_real_pow = thm"complex_of_real_pow"; |
|
1632 |
val complex_mod = thm"complex_mod"; |
|
1633 |
val complex_mod_zero = thm"complex_mod_zero"; |
|
1634 |
val complex_mod_one = thm"complex_mod_one"; |
|
1635 |
val complex_mod_complex_of_real = thm"complex_mod_complex_of_real"; |
|
1636 |
val complex_of_real_abs = thm"complex_of_real_abs"; |
|
1637 |
val complex_cnj = thm"complex_cnj"; |
|
1638 |
val inj_cnj = thm"inj_cnj"; |
|
1639 |
val complex_cnj_cancel_iff = thm"complex_cnj_cancel_iff"; |
|
1640 |
val complex_cnj_cnj = thm"complex_cnj_cnj"; |
|
1641 |
val complex_cnj_complex_of_real = thm"complex_cnj_complex_of_real"; |
|
1642 |
val complex_mod_cnj = thm"complex_mod_cnj"; |
|
1643 |
val complex_cnj_minus = thm"complex_cnj_minus"; |
|
1644 |
val complex_cnj_inverse = thm"complex_cnj_inverse"; |
|
1645 |
val complex_cnj_add = thm"complex_cnj_add"; |
|
1646 |
val complex_cnj_diff = thm"complex_cnj_diff"; |
|
1647 |
val complex_cnj_mult = thm"complex_cnj_mult"; |
|
1648 |
val complex_cnj_divide = thm"complex_cnj_divide"; |
|
1649 |
val complex_cnj_one = thm"complex_cnj_one"; |
|
1650 |
val complex_cnj_pow = thm"complex_cnj_pow"; |
|
1651 |
val complex_add_cnj = thm"complex_add_cnj"; |
|
1652 |
val complex_diff_cnj = thm"complex_diff_cnj"; |
|
1653 |
val complex_cnj_zero = thm"complex_cnj_zero"; |
|
1654 |
val complex_cnj_zero_iff = thm"complex_cnj_zero_iff"; |
|
1655 |
val complex_mult_cnj = thm"complex_mult_cnj"; |
|
1656 |
val complex_add_left_cancel_zero = thm"complex_add_left_cancel_zero"; |
|
1657 |
val complex_diff_mult_distrib = thm"complex_diff_mult_distrib"; |
|
1658 |
val complex_diff_mult_distrib2 = thm"complex_diff_mult_distrib2"; |
|
1659 |
val complex_mod_eq_zero_cancel = thm"complex_mod_eq_zero_cancel"; |
|
1660 |
val complex_mod_complex_of_real_of_nat = thm"complex_mod_complex_of_real_of_nat"; |
|
1661 |
val complex_mod_minus = thm"complex_mod_minus"; |
|
1662 |
val complex_mod_mult_cnj = thm"complex_mod_mult_cnj"; |
|
1663 |
val complex_mod_squared = thm"complex_mod_squared"; |
|
1664 |
val complex_mod_ge_zero = thm"complex_mod_ge_zero"; |
|
1665 |
val abs_cmod_cancel = thm"abs_cmod_cancel"; |
|
1666 |
val complex_mod_mult = thm"complex_mod_mult"; |
|
1667 |
val complex_mod_add_squared_eq = thm"complex_mod_add_squared_eq"; |
|
1668 |
val complex_Re_mult_cnj_le_cmod = thm"complex_Re_mult_cnj_le_cmod"; |
|
1669 |
val complex_Re_mult_cnj_le_cmod2 = thm"complex_Re_mult_cnj_le_cmod2"; |
|
1670 |
val real_sum_squared_expand = thm"real_sum_squared_expand"; |
|
1671 |
val complex_mod_triangle_squared = thm"complex_mod_triangle_squared"; |
|
1672 |
val complex_mod_minus_le_complex_mod = thm"complex_mod_minus_le_complex_mod"; |
|
1673 |
val complex_mod_triangle_ineq = thm"complex_mod_triangle_ineq"; |
|
1674 |
val complex_mod_triangle_ineq2 = thm"complex_mod_triangle_ineq2"; |
|
1675 |
val complex_mod_diff_commute = thm"complex_mod_diff_commute"; |
|
1676 |
val complex_mod_add_less = thm"complex_mod_add_less"; |
|
1677 |
val complex_mod_mult_less = thm"complex_mod_mult_less"; |
|
1678 |
val complex_mod_diff_ineq = thm"complex_mod_diff_ineq"; |
|
1679 |
val complex_Re_le_cmod = thm"complex_Re_le_cmod"; |
|
1680 |
val complex_mod_gt_zero = thm"complex_mod_gt_zero"; |
|
1681 |
val complex_mod_complexpow = thm"complex_mod_complexpow"; |
|
1682 |
val complexpow_minus = thm"complexpow_minus"; |
|
1683 |
val complex_inverse_minus = thm"complex_inverse_minus"; |
|
1684 |
val complex_divide_one = thm"complex_divide_one"; |
|
1685 |
val complex_mod_inverse = thm"complex_mod_inverse"; |
|
1686 |
val complex_mod_divide = thm"complex_mod_divide"; |
|
1687 |
val complex_inverse_divide = thm"complex_inverse_divide"; |
|
1688 |
val complexpow_mult = thm"complexpow_mult"; |
|
1689 |
val complexpow_zero = thm"complexpow_zero"; |
|
1690 |
val complexpow_not_zero = thm"complexpow_not_zero"; |
|
1691 |
val complexpow_zero_zero = thm"complexpow_zero_zero"; |
|
1692 |
val complexpow_i_squared = thm"complexpow_i_squared"; |
|
1693 |
val complex_i_not_zero = thm"complex_i_not_zero"; |
|
1694 |
val complex_mult_eq_zero_cancel1 = thm"complex_mult_eq_zero_cancel1"; |
|
1695 |
val complex_mult_eq_zero_cancel2 = thm"complex_mult_eq_zero_cancel2"; |
|
1696 |
val complex_mult_not_eq_zero_iff = thm"complex_mult_not_eq_zero_iff"; |
|
1697 |
val complexpow_inverse = thm"complexpow_inverse"; |
|
1698 |
val sgn_zero = thm"sgn_zero"; |
|
1699 |
val sgn_one = thm"sgn_one"; |
|
1700 |
val sgn_minus = thm"sgn_minus"; |
|
1701 |
val sgn_eq = thm"sgn_eq"; |
|
1702 |
val complex_split = thm"complex_split"; |
|
1703 |
val Re_complex_i = thm"Re_complex_i"; |
|
1704 |
val Im_complex_i = thm"Im_complex_i"; |
|
1705 |
val i_mult_eq = thm"i_mult_eq"; |
|
1706 |
val i_mult_eq2 = thm"i_mult_eq2"; |
|
1707 |
val cmod_i = thm"cmod_i"; |
|
1708 |
val complex_eq_Re_eq = thm"complex_eq_Re_eq"; |
|
1709 |
val complex_eq_Im_eq = thm"complex_eq_Im_eq"; |
|
1710 |
val complex_eq_cancel_iff = thm"complex_eq_cancel_iff"; |
|
1711 |
val complex_eq_cancel_iffA = thm"complex_eq_cancel_iffA"; |
|
1712 |
val complex_eq_cancel_iffB = thm"complex_eq_cancel_iffB"; |
|
1713 |
val complex_eq_cancel_iffC = thm"complex_eq_cancel_iffC"; |
|
1714 |
val complex_eq_cancel_iff2 = thm"complex_eq_cancel_iff2"; |
|
1715 |
val complex_eq_cancel_iff2a = thm"complex_eq_cancel_iff2a"; |
|
1716 |
val complex_eq_cancel_iff3 = thm"complex_eq_cancel_iff3"; |
|
1717 |
val complex_eq_cancel_iff3a = thm"complex_eq_cancel_iff3a"; |
|
1718 |
val complex_split_Re_zero = thm"complex_split_Re_zero"; |
|
1719 |
val complex_split_Im_zero = thm"complex_split_Im_zero"; |
|
1720 |
val Re_sgn = thm"Re_sgn"; |
|
1721 |
val Im_sgn = thm"Im_sgn"; |
|
1722 |
val complex_inverse_complex_split = thm"complex_inverse_complex_split"; |
|
1723 |
val Re_mult_i_eq = thm"Re_mult_i_eq"; |
|
1724 |
val Im_mult_i_eq = thm"Im_mult_i_eq"; |
|
1725 |
val complex_mod_mult_i = thm"complex_mod_mult_i"; |
|
1726 |
val cos_arg_i_mult_zero = thm"cos_arg_i_mult_zero"; |
|
1727 |
val cos_arg_i_mult_zero2 = thm"cos_arg_i_mult_zero2"; |
|
1728 |
val complex_of_real_not_zero_iff = thm"complex_of_real_not_zero_iff"; |
|
1729 |
val complex_of_real_zero_iff = thm"complex_of_real_zero_iff"; |
|
1730 |
val cos_arg_i_mult_zero3 = thm"cos_arg_i_mult_zero3"; |
|
1731 |
val complex_split_polar = thm"complex_split_polar"; |
|
1732 |
val rcis_Ex = thm"rcis_Ex"; |
|
1733 |
val Re_complex_polar = thm"Re_complex_polar"; |
|
1734 |
val Re_rcis = thm"Re_rcis"; |
|
1735 |
val Im_complex_polar = thm"Im_complex_polar"; |
|
1736 |
val Im_rcis = thm"Im_rcis"; |
|
1737 |
val complex_mod_complex_polar = thm"complex_mod_complex_polar"; |
|
1738 |
val complex_mod_rcis = thm"complex_mod_rcis"; |
|
1739 |
val complex_mod_sqrt_Re_mult_cnj = thm"complex_mod_sqrt_Re_mult_cnj"; |
|
1740 |
val complex_Re_cnj = thm"complex_Re_cnj"; |
|
1741 |
val complex_Im_cnj = thm"complex_Im_cnj"; |
|
1742 |
val complex_In_mult_cnj_zero = thm"complex_In_mult_cnj_zero"; |
|
1743 |
val complex_Re_mult = thm"complex_Re_mult"; |
|
1744 |
val complex_Re_mult_complex_of_real = thm"complex_Re_mult_complex_of_real"; |
|
1745 |
val complex_Im_mult_complex_of_real = thm"complex_Im_mult_complex_of_real"; |
|
1746 |
val complex_Re_mult_complex_of_real2 = thm"complex_Re_mult_complex_of_real2"; |
|
1747 |
val complex_Im_mult_complex_of_real2 = thm"complex_Im_mult_complex_of_real2"; |
|
1748 |
val cis_rcis_eq = thm"cis_rcis_eq"; |
|
1749 |
val rcis_mult = thm"rcis_mult"; |
|
1750 |
val cis_mult = thm"cis_mult"; |
|
1751 |
val cis_zero = thm"cis_zero"; |
|
1752 |
val cis_zero2 = thm"cis_zero2"; |
|
1753 |
val rcis_zero_mod = thm"rcis_zero_mod"; |
|
1754 |
val rcis_zero_arg = thm"rcis_zero_arg"; |
|
1755 |
val complex_of_real_minus_one = thm"complex_of_real_minus_one"; |
|
1756 |
val complex_i_mult_minus = thm"complex_i_mult_minus"; |
|
1757 |
val complex_i_mult_minus2 = thm"complex_i_mult_minus2"; |
|
1758 |
val cis_real_of_nat_Suc_mult = thm"cis_real_of_nat_Suc_mult"; |
|
1759 |
val DeMoivre = thm"DeMoivre"; |
|
1760 |
val DeMoivre2 = thm"DeMoivre2"; |
|
1761 |
val cis_inverse = thm"cis_inverse"; |
|
1762 |
val rcis_inverse = thm"rcis_inverse"; |
|
1763 |
val cis_divide = thm"cis_divide"; |
|
1764 |
val rcis_divide = thm"rcis_divide"; |
|
1765 |
val Re_cis = thm"Re_cis"; |
|
1766 |
val Im_cis = thm"Im_cis"; |
|
1767 |
val cos_n_Re_cis_pow_n = thm"cos_n_Re_cis_pow_n"; |
|
1768 |
val sin_n_Im_cis_pow_n = thm"sin_n_Im_cis_pow_n"; |
|
1769 |
val expi_Im_split = thm"expi_Im_split"; |
|
1770 |
val expi_Im_cis = thm"expi_Im_cis"; |
|
1771 |
val expi_add = thm"expi_add"; |
|
1772 |
val expi_complex_split = thm"expi_complex_split"; |
|
1773 |
val expi_zero = thm"expi_zero"; |
|
1774 |
val complex_Re_mult_eq = thm"complex_Re_mult_eq"; |
|
1775 |
val complex_Im_mult_eq = thm"complex_Im_mult_eq"; |
|
1776 |
val complex_expi_Ex = thm"complex_expi_Ex"; |
|
1777 |
||
1778 |
val complex_add_ac = thms"complex_add_ac"; |
|
1779 |
val complex_mult_ac = thms"complex_mult_ac"; |
|
1780 |
*} |
|
1781 |
||
13957 | 1782 |
end |
1783 |
||
1784 |