| author | haftmann | 
| Sat, 29 Sep 2007 08:58:56 +0200 | |
| changeset 24751 | dbb34a03af5a | 
| parent 24520 | 40b220403257 | 
| child 25502 | 9200b36280c0 | 
| permissions | -rw-r--r-- | 
| 13957 | 1 | (* Title: Complex.thy | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 2 | ID: $Id$ | 
| 13957 | 3 | Author: Jacques D. Fleuriot | 
| 4 | Copyright: 2001 University of Edinburgh | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 5 | Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4 | 
| 13957 | 6 | *) | 
| 7 | ||
| 14377 | 8 | header {* Complex Numbers: Rectangular and Polar Representations *}
 | 
| 14373 | 9 | |
| 15131 | 10 | theory Complex | 
| 22655 | 11 | imports "../Hyperreal/Transcendental" | 
| 15131 | 12 | begin | 
| 13957 | 13 | |
| 14373 | 14 | datatype complex = Complex real real | 
| 13957 | 15 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 16 | consts Re :: "complex \<Rightarrow> real" | 
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 17 | primrec Re: "Re (Complex x y) = x" | 
| 14373 | 18 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 19 | consts Im :: "complex \<Rightarrow> real" | 
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 20 | primrec Im: "Im (Complex x y) = y" | 
| 14373 | 21 | |
| 22 | lemma complex_surj [simp]: "Complex (Re z) (Im z) = z" | |
| 23 | by (induct z) simp | |
| 13957 | 24 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 25 | lemma complex_equality [intro?]: "\<lbrakk>Re x = Re y; Im x = Im y\<rbrakk> \<Longrightarrow> x = y" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 26 | by (induct x, induct y) simp | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 27 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 28 | lemma expand_complex_eq: "(x = y) = (Re x = Re y \<and> Im x = Im y)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 29 | by (induct x, induct y) simp | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 30 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 31 | lemmas complex_Re_Im_cancel_iff = expand_complex_eq | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 32 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 33 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 34 | subsection {* Addition and Subtraction *}
 | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 35 | |
| 23124 | 36 | instance complex :: zero | 
| 14323 | 37 | complex_zero_def: | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 38 | "0 \<equiv> Complex 0 0" .. | 
| 23124 | 39 | |
| 40 | instance complex :: plus | |
| 41 | complex_add_def: | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 42 | "x + y \<equiv> Complex (Re x + Re y) (Im x + Im y)" .. | 
| 14323 | 43 | |
| 23124 | 44 | instance complex :: minus | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 45 | complex_minus_def: | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 46 | "- x \<equiv> Complex (- Re x) (- Im x)" | 
| 23124 | 47 | complex_diff_def: | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 48 | "x - y \<equiv> x + - y" .. | 
| 14323 | 49 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 50 | lemma Complex_eq_0 [simp]: "(Complex a b = 0) = (a = 0 \<and> b = 0)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 51 | by (simp add: complex_zero_def) | 
| 14323 | 52 | |
| 14374 | 53 | lemma complex_Re_zero [simp]: "Re 0 = 0" | 
| 54 | by (simp add: complex_zero_def) | |
| 55 | ||
| 56 | lemma complex_Im_zero [simp]: "Im 0 = 0" | |
| 14373 | 57 | by (simp add: complex_zero_def) | 
| 14323 | 58 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 59 | lemma complex_add [simp]: | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 60 | "Complex a b + Complex c d = Complex (a + c) (b + d)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 61 | by (simp add: complex_add_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 62 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 63 | lemma complex_Re_add [simp]: "Re (x + y) = Re x + Re y" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 64 | by (simp add: complex_add_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 65 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 66 | lemma complex_Im_add [simp]: "Im (x + y) = Im x + Im y" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 67 | by (simp add: complex_add_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 68 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 69 | lemma complex_minus [simp]: "- (Complex a b) = Complex (- a) (- b)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 70 | by (simp add: complex_minus_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 71 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 72 | lemma complex_Re_minus [simp]: "Re (- x) = - Re x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 73 | by (simp add: complex_minus_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 74 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 75 | lemma complex_Im_minus [simp]: "Im (- x) = - Im x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 76 | by (simp add: complex_minus_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 77 | |
| 23275 | 78 | lemma complex_diff [simp]: | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 79 | "Complex a b - Complex c d = Complex (a - c) (b - d)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 80 | by (simp add: complex_diff_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 81 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 82 | lemma complex_Re_diff [simp]: "Re (x - y) = Re x - Re y" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 83 | by (simp add: complex_diff_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 84 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 85 | lemma complex_Im_diff [simp]: "Im (x - y) = Im x - Im y" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 86 | by (simp add: complex_diff_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 87 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 88 | instance complex :: ab_group_add | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 89 | proof | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 90 | fix x y z :: complex | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 91 | show "(x + y) + z = x + (y + z)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 92 | by (simp add: expand_complex_eq add_assoc) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 93 | show "x + y = y + x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 94 | by (simp add: expand_complex_eq add_commute) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 95 | show "0 + x = x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 96 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 97 | show "- x + x = 0" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 98 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 99 | show "x - y = x + - y" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 100 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 101 | qed | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 102 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 103 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 104 | subsection {* Multiplication and Division *}
 | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 105 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 106 | instance complex :: one | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 107 | complex_one_def: | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 108 | "1 \<equiv> Complex 1 0" .. | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 109 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 110 | instance complex :: times | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 111 | complex_mult_def: | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 112 | "x * y \<equiv> Complex (Re x * Re y - Im x * Im y) (Re x * Im y + Im x * Re y)" .. | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 113 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 114 | instance complex :: inverse | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 115 | complex_inverse_def: | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 116 | "inverse x \<equiv> | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 117 | Complex (Re x / ((Re x)\<twosuperior> + (Im x)\<twosuperior>)) (- Im x / ((Re x)\<twosuperior> + (Im x)\<twosuperior>))" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 118 | complex_divide_def: | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 119 | "x / y \<equiv> x * inverse y" .. | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 120 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 121 | lemma Complex_eq_1 [simp]: "(Complex a b = 1) = (a = 1 \<and> b = 0)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 122 | by (simp add: complex_one_def) | 
| 22861 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 123 | |
| 14374 | 124 | lemma complex_Re_one [simp]: "Re 1 = 1" | 
| 125 | by (simp add: complex_one_def) | |
| 14323 | 126 | |
| 14374 | 127 | lemma complex_Im_one [simp]: "Im 1 = 0" | 
| 14373 | 128 | by (simp add: complex_one_def) | 
| 14323 | 129 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 130 | lemma complex_mult [simp]: | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 131 | "Complex a b * Complex c d = Complex (a * c - b * d) (a * d + b * c)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 132 | by (simp add: complex_mult_def) | 
| 14323 | 133 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 134 | lemma complex_Re_mult [simp]: "Re (x * y) = Re x * Re y - Im x * Im y" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 135 | by (simp add: complex_mult_def) | 
| 14323 | 136 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 137 | lemma complex_Im_mult [simp]: "Im (x * y) = Re x * Im y + Im x * Re y" | 
| 14373 | 138 | by (simp add: complex_mult_def) | 
| 14323 | 139 | |
| 14377 | 140 | lemma complex_inverse [simp]: | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 141 | "inverse (Complex a b) = Complex (a / (a\<twosuperior> + b\<twosuperior>)) (- b / (a\<twosuperior> + b\<twosuperior>))" | 
| 14373 | 142 | by (simp add: complex_inverse_def) | 
| 14335 | 143 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 144 | lemma complex_Re_inverse: | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 145 | "Re (inverse x) = Re x / ((Re x)\<twosuperior> + (Im x)\<twosuperior>)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 146 | by (simp add: complex_inverse_def) | 
| 14323 | 147 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 148 | lemma complex_Im_inverse: | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 149 | "Im (inverse x) = - Im x / ((Re x)\<twosuperior> + (Im x)\<twosuperior>)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 150 | by (simp add: complex_inverse_def) | 
| 14335 | 151 | |
| 152 | instance complex :: field | |
| 153 | proof | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 154 | fix x y z :: complex | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 155 | show "(x * y) * z = x * (y * z)" | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23275diff
changeset | 156 | by (simp add: expand_complex_eq ring_simps) | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 157 | show "x * y = y * x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 158 | by (simp add: expand_complex_eq mult_commute add_commute) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 159 | show "1 * x = x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 160 | by (simp add: expand_complex_eq) | 
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14335diff
changeset | 161 | show "0 \<noteq> (1::complex)" | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 162 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 163 | show "(x + y) * z = x * z + y * z" | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23275diff
changeset | 164 | by (simp add: expand_complex_eq ring_simps) | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 165 | show "x / y = x * inverse y" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 166 | by (simp only: complex_divide_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 167 | show "x \<noteq> 0 \<Longrightarrow> inverse x * x = 1" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 168 | by (induct x, simp add: power2_eq_square add_divide_distrib [symmetric]) | 
| 14335 | 169 | qed | 
| 170 | ||
| 14373 | 171 | instance complex :: division_by_zero | 
| 172 | proof | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 173 | show "inverse 0 = (0::complex)" | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 174 | by (simp add: complex_inverse_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 175 | qed | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 176 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 177 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 178 | subsection {* Exponentiation *}
 | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 179 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 180 | instance complex :: power .. | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 181 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 182 | primrec | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 183 | complexpow_0: "z ^ 0 = 1" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 184 | complexpow_Suc: "z ^ (Suc n) = (z::complex) * (z ^ n)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 185 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 186 | instance complex :: recpower | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 187 | proof | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 188 | fix x :: complex and n :: nat | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 189 | show "x ^ 0 = 1" by simp | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 190 | show "x ^ Suc n = x * x ^ n" by simp | 
| 14373 | 191 | qed | 
| 14335 | 192 | |
| 14323 | 193 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 194 | subsection {* Numerals and Arithmetic *}
 | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 195 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 196 | instance complex :: number | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 197 | complex_number_of_def: | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 198 | "number_of w \<equiv> of_int w" .. | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 199 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 200 | instance complex :: number_ring | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 201 | by (intro_classes, simp only: complex_number_of_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 202 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 203 | lemma complex_Re_of_nat [simp]: "Re (of_nat n) = of_nat n" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 204 | by (induct n) simp_all | 
| 20556 
2e8227b81bf1
add instance for real_algebra_1 and real_normed_div_algebra
 huffman parents: 
20485diff
changeset | 205 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 206 | lemma complex_Im_of_nat [simp]: "Im (of_nat n) = 0" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 207 | by (induct n) simp_all | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 208 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 209 | lemma complex_Re_of_int [simp]: "Re (of_int z) = of_int z" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 210 | by (cases z rule: int_diff_cases) simp | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 211 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 212 | lemma complex_Im_of_int [simp]: "Im (of_int z) = 0" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 213 | by (cases z rule: int_diff_cases) simp | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 214 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 215 | lemma complex_Re_number_of [simp]: "Re (number_of v) = number_of v" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 216 | unfolding number_ring_class.axioms by (rule complex_Re_of_int) | 
| 20556 
2e8227b81bf1
add instance for real_algebra_1 and real_normed_div_algebra
 huffman parents: 
20485diff
changeset | 217 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 218 | lemma complex_Im_number_of [simp]: "Im (number_of v) = 0" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 219 | unfolding number_ring_class.axioms by (rule complex_Im_of_int) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 220 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 221 | lemma Complex_eq_number_of [simp]: | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 222 | "(Complex a b = number_of w) = (a = number_of w \<and> b = 0)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 223 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 224 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 225 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 226 | subsection {* Scalar Multiplication *}
 | 
| 20556 
2e8227b81bf1
add instance for real_algebra_1 and real_normed_div_algebra
 huffman parents: 
20485diff
changeset | 227 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 228 | instance complex :: scaleR | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 229 | complex_scaleR_def: | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 230 | "scaleR r x \<equiv> Complex (r * Re x) (r * Im x)" .. | 
| 22972 
3e96b98d37c6
generalized sgn function to work on any real normed vector space
 huffman parents: 
22968diff
changeset | 231 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 232 | lemma complex_scaleR [simp]: | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 233 | "scaleR r (Complex a b) = Complex (r * a) (r * b)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 234 | unfolding complex_scaleR_def by simp | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 235 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 236 | lemma complex_Re_scaleR [simp]: "Re (scaleR r x) = r * Re x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 237 | unfolding complex_scaleR_def by simp | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 238 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 239 | lemma complex_Im_scaleR [simp]: "Im (scaleR r x) = r * Im x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 240 | unfolding complex_scaleR_def by simp | 
| 22972 
3e96b98d37c6
generalized sgn function to work on any real normed vector space
 huffman parents: 
22968diff
changeset | 241 | |
| 20725 
72e20198f834
instance complex :: real_normed_field; cleaned up
 huffman parents: 
20560diff
changeset | 242 | instance complex :: real_field | 
| 20556 
2e8227b81bf1
add instance for real_algebra_1 and real_normed_div_algebra
 huffman parents: 
20485diff
changeset | 243 | proof | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 244 | fix a b :: real and x y :: complex | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 245 | show "scaleR a (x + y) = scaleR a x + scaleR a y" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 246 | by (simp add: expand_complex_eq right_distrib) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 247 | show "scaleR (a + b) x = scaleR a x + scaleR b x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 248 | by (simp add: expand_complex_eq left_distrib) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 249 | show "scaleR a (scaleR b x) = scaleR (a * b) x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 250 | by (simp add: expand_complex_eq mult_assoc) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 251 | show "scaleR 1 x = x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 252 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 253 | show "scaleR a x * y = scaleR a (x * y)" | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23275diff
changeset | 254 | by (simp add: expand_complex_eq ring_simps) | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 255 | show "x * scaleR a y = scaleR a (x * y)" | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23275diff
changeset | 256 | by (simp add: expand_complex_eq ring_simps) | 
| 20556 
2e8227b81bf1
add instance for real_algebra_1 and real_normed_div_algebra
 huffman parents: 
20485diff
changeset | 257 | qed | 
| 
2e8227b81bf1
add instance for real_algebra_1 and real_normed_div_algebra
 huffman parents: 
20485diff
changeset | 258 | |
| 
2e8227b81bf1
add instance for real_algebra_1 and real_normed_div_algebra
 huffman parents: 
20485diff
changeset | 259 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 260 | subsection{* Properties of Embedding from Reals *}
 | 
| 14323 | 261 | |
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 262 | abbreviation | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 263 | complex_of_real :: "real \<Rightarrow> complex" where | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 264 | "complex_of_real \<equiv> of_real" | 
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 265 | |
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 266 | lemma complex_of_real_def: "complex_of_real r = Complex r 0" | 
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 267 | by (simp add: of_real_def complex_scaleR_def) | 
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 268 | |
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 269 | lemma Re_complex_of_real [simp]: "Re (complex_of_real z) = z" | 
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 270 | by (simp add: complex_of_real_def) | 
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 271 | |
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 272 | lemma Im_complex_of_real [simp]: "Im (complex_of_real z) = 0" | 
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 273 | by (simp add: complex_of_real_def) | 
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 274 | |
| 14377 | 275 | lemma Complex_add_complex_of_real [simp]: | 
| 276 | "Complex x y + complex_of_real r = Complex (x+r) y" | |
| 277 | by (simp add: complex_of_real_def) | |
| 278 | ||
| 279 | lemma complex_of_real_add_Complex [simp]: | |
| 280 | "complex_of_real r + Complex x y = Complex (r+x) y" | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 281 | by (simp add: complex_of_real_def) | 
| 14377 | 282 | |
| 283 | lemma Complex_mult_complex_of_real: | |
| 284 | "Complex x y * complex_of_real r = Complex (x*r) (y*r)" | |
| 285 | by (simp add: complex_of_real_def) | |
| 286 | ||
| 287 | lemma complex_of_real_mult_Complex: | |
| 288 | "complex_of_real r * Complex x y = Complex (r*x) (r*y)" | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 289 | by (simp add: complex_of_real_def) | 
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 290 | |
| 14377 | 291 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 292 | subsection {* Vector Norm *}
 | 
| 14323 | 293 | |
| 22861 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 294 | instance complex :: norm | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 295 | complex_norm_def: | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 296 | "norm z \<equiv> sqrt ((Re z)\<twosuperior> + (Im z)\<twosuperior>)" .. | 
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 297 | |
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 298 | abbreviation | 
| 22861 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 299 | cmod :: "complex \<Rightarrow> real" where | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 300 | "cmod \<equiv> norm" | 
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 301 | |
| 24506 | 302 | instance complex :: sgn | 
| 303 | complex_sgn_def: "sgn x == x /\<^sub>R cmod x" .. | |
| 304 | ||
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 305 | lemmas cmod_def = complex_norm_def | 
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 306 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 307 | lemma complex_norm [simp]: "cmod (Complex x y) = sqrt (x\<twosuperior> + y\<twosuperior>)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 308 | by (simp add: complex_norm_def) | 
| 22852 | 309 | |
| 20725 
72e20198f834
instance complex :: real_normed_field; cleaned up
 huffman parents: 
20560diff
changeset | 310 | instance complex :: real_normed_field | 
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 311 | proof | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 312 | fix r :: real and x y :: complex | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 313 | show "0 \<le> norm x" | 
| 22861 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 314 | by (induct x) simp | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 315 | show "(norm x = 0) = (x = 0)" | 
| 22861 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 316 | by (induct x) simp | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 317 | show "norm (x + y) \<le> norm x + norm y" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 318 | by (induct x, induct y) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 319 | (simp add: real_sqrt_sum_squares_triangle_ineq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 320 | show "norm (scaleR r x) = \<bar>r\<bar> * norm x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 321 | by (induct x) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 322 | (simp add: power_mult_distrib right_distrib [symmetric] real_sqrt_mult) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 323 | show "norm (x * y) = norm x * norm y" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 324 | by (induct x, induct y) | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23275diff
changeset | 325 | (simp add: real_sqrt_mult [symmetric] power2_eq_square ring_simps) | 
| 24506 | 326 | show "sgn x = x /\<^sub>R cmod x" by(simp add: complex_sgn_def) | 
| 24520 | 327 | qed | 
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 328 | |
| 22861 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 329 | lemma cmod_unit_one [simp]: "cmod (Complex (cos a) (sin a)) = 1" | 
| 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 330 | by simp | 
| 14323 | 331 | |
| 22861 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 332 | lemma cmod_complex_polar [simp]: | 
| 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 333 | "cmod (complex_of_real r * Complex (cos a) (sin a)) = abs r" | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 334 | by (simp add: norm_mult) | 
| 22861 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 335 | |
| 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 336 | lemma complex_Re_le_cmod: "Re x \<le> cmod x" | 
| 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 337 | unfolding complex_norm_def | 
| 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 338 | by (rule real_sqrt_sum_squares_ge1) | 
| 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 339 | |
| 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 340 | lemma complex_mod_minus_le_complex_mod [simp]: "- cmod x \<le> cmod x" | 
| 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 341 | by (rule order_trans [OF _ norm_ge_zero], simp) | 
| 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 342 | |
| 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 343 | lemma complex_mod_triangle_ineq2 [simp]: "cmod(b + a) - cmod b \<le> cmod a" | 
| 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 344 | by (rule ord_le_eq_trans [OF norm_triangle_ineq2], simp) | 
| 14323 | 345 | |
| 22861 
8ec47039614e
clean up complex norm proofs, remove redundant lemmas
 huffman parents: 
22852diff
changeset | 346 | lemmas real_sum_squared_expand = power2_sum [where 'a=real] | 
| 14323 | 347 | |
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 348 | |
| 23123 | 349 | subsection {* Completeness of the Complexes *}
 | 
| 350 | ||
| 351 | interpretation Re: bounded_linear ["Re"] | |
| 352 | apply (unfold_locales, simp, simp) | |
| 353 | apply (rule_tac x=1 in exI) | |
| 354 | apply (simp add: complex_norm_def) | |
| 355 | done | |
| 356 | ||
| 357 | interpretation Im: bounded_linear ["Im"] | |
| 358 | apply (unfold_locales, simp, simp) | |
| 359 | apply (rule_tac x=1 in exI) | |
| 360 | apply (simp add: complex_norm_def) | |
| 361 | done | |
| 362 | ||
| 363 | lemma LIMSEQ_Complex: | |
| 364 | "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. Complex (X n) (Y n)) ----> Complex a b" | |
| 365 | apply (rule LIMSEQ_I) | |
| 366 | apply (subgoal_tac "0 < r / sqrt 2") | |
| 367 | apply (drule_tac r="r / sqrt 2" in LIMSEQ_D, safe) | |
| 368 | apply (drule_tac r="r / sqrt 2" in LIMSEQ_D, safe) | |
| 369 | apply (rename_tac M N, rule_tac x="max M N" in exI, safe) | |
| 370 | apply (simp add: real_sqrt_sum_squares_less) | |
| 371 | apply (simp add: divide_pos_pos) | |
| 372 | done | |
| 373 | ||
| 374 | instance complex :: banach | |
| 375 | proof | |
| 376 | fix X :: "nat \<Rightarrow> complex" | |
| 377 | assume X: "Cauchy X" | |
| 378 | from Re.Cauchy [OF X] have 1: "(\<lambda>n. Re (X n)) ----> lim (\<lambda>n. Re (X n))" | |
| 379 | by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) | |
| 380 | from Im.Cauchy [OF X] have 2: "(\<lambda>n. Im (X n)) ----> lim (\<lambda>n. Im (X n))" | |
| 381 | by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) | |
| 382 | have "X ----> Complex (lim (\<lambda>n. Re (X n))) (lim (\<lambda>n. Im (X n)))" | |
| 383 | using LIMSEQ_Complex [OF 1 2] by simp | |
| 384 | thus "convergent X" | |
| 385 | by (rule convergentI) | |
| 386 | qed | |
| 387 | ||
| 388 | ||
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 389 | subsection {* The Complex Number @{term "\<i>"} *}
 | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 390 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 391 | definition | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 392 |   "ii" :: complex  ("\<i>") where
 | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 393 | i_def: "ii \<equiv> Complex 0 1" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 394 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 395 | lemma complex_Re_i [simp]: "Re ii = 0" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 396 | by (simp add: i_def) | 
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 397 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 398 | lemma complex_Im_i [simp]: "Im ii = 1" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 399 | by (simp add: i_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 400 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 401 | lemma Complex_eq_i [simp]: "(Complex x y = ii) = (x = 0 \<and> y = 1)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 402 | by (simp add: i_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 403 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 404 | lemma complex_i_not_zero [simp]: "ii \<noteq> 0" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 405 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 406 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 407 | lemma complex_i_not_one [simp]: "ii \<noteq> 1" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 408 | by (simp add: expand_complex_eq) | 
| 23124 | 409 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 410 | lemma complex_i_not_number_of [simp]: "ii \<noteq> number_of w" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 411 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 412 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 413 | lemma i_mult_Complex [simp]: "ii * Complex a b = Complex (- b) a" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 414 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 415 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 416 | lemma Complex_mult_i [simp]: "Complex a b * ii = Complex (- b) a" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 417 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 418 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 419 | lemma i_complex_of_real [simp]: "ii * complex_of_real r = Complex 0 r" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 420 | by (simp add: i_def complex_of_real_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 421 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 422 | lemma complex_of_real_i [simp]: "complex_of_real r * ii = Complex 0 r" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 423 | by (simp add: i_def complex_of_real_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 424 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 425 | lemma i_squared [simp]: "ii * ii = -1" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 426 | by (simp add: i_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 427 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 428 | lemma power2_i [simp]: "ii\<twosuperior> = -1" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 429 | by (simp add: power2_eq_square) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 430 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 431 | lemma inverse_i [simp]: "inverse ii = - ii" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 432 | by (rule inverse_unique, simp) | 
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 433 | |
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 434 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 435 | subsection {* Complex Conjugation *}
 | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 436 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 437 | definition | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 438 | cnj :: "complex \<Rightarrow> complex" where | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 439 | "cnj z = Complex (Re z) (- Im z)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 440 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 441 | lemma complex_cnj [simp]: "cnj (Complex a b) = Complex a (- b)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 442 | by (simp add: cnj_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 443 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 444 | lemma complex_Re_cnj [simp]: "Re (cnj x) = Re x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 445 | by (simp add: cnj_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 446 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 447 | lemma complex_Im_cnj [simp]: "Im (cnj x) = - Im x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 448 | by (simp add: cnj_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 449 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 450 | lemma complex_cnj_cancel_iff [simp]: "(cnj x = cnj y) = (x = y)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 451 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 452 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 453 | lemma complex_cnj_cnj [simp]: "cnj (cnj z) = z" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 454 | by (simp add: cnj_def) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 455 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 456 | lemma complex_cnj_zero [simp]: "cnj 0 = 0" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 457 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 458 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 459 | lemma complex_cnj_zero_iff [iff]: "(cnj z = 0) = (z = 0)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 460 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 461 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 462 | lemma complex_cnj_add: "cnj (x + y) = cnj x + cnj y" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 463 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 464 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 465 | lemma complex_cnj_diff: "cnj (x - y) = cnj x - cnj y" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 466 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 467 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 468 | lemma complex_cnj_minus: "cnj (- x) = - cnj x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 469 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 470 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 471 | lemma complex_cnj_one [simp]: "cnj 1 = 1" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 472 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 473 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 474 | lemma complex_cnj_mult: "cnj (x * y) = cnj x * cnj y" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 475 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 476 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 477 | lemma complex_cnj_inverse: "cnj (inverse x) = inverse (cnj x)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 478 | by (simp add: complex_inverse_def) | 
| 14323 | 479 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 480 | lemma complex_cnj_divide: "cnj (x / y) = cnj x / cnj y" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 481 | by (simp add: complex_divide_def complex_cnj_mult complex_cnj_inverse) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 482 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 483 | lemma complex_cnj_power: "cnj (x ^ n) = cnj x ^ n" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 484 | by (induct n, simp_all add: complex_cnj_mult) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 485 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 486 | lemma complex_cnj_of_nat [simp]: "cnj (of_nat n) = of_nat n" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 487 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 488 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 489 | lemma complex_cnj_of_int [simp]: "cnj (of_int z) = of_int z" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 490 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 491 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 492 | lemma complex_cnj_number_of [simp]: "cnj (number_of w) = number_of w" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 493 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 494 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 495 | lemma complex_cnj_scaleR: "cnj (scaleR r x) = scaleR r (cnj x)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 496 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 497 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 498 | lemma complex_mod_cnj [simp]: "cmod (cnj z) = cmod z" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 499 | by (simp add: complex_norm_def) | 
| 14323 | 500 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 501 | lemma complex_cnj_complex_of_real [simp]: "cnj (of_real x) = of_real x" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 502 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 503 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 504 | lemma complex_cnj_i [simp]: "cnj ii = - ii" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 505 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 506 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 507 | lemma complex_add_cnj: "z + cnj z = complex_of_real (2 * Re z)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 508 | by (simp add: expand_complex_eq) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 509 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 510 | lemma complex_diff_cnj: "z - cnj z = complex_of_real (2 * Im z) * ii" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 511 | by (simp add: expand_complex_eq) | 
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 512 | |
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 513 | lemma complex_mult_cnj: "z * cnj z = complex_of_real ((Re z)\<twosuperior> + (Im z)\<twosuperior>)" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 514 | by (simp add: expand_complex_eq power2_eq_square) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 515 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 516 | lemma complex_mod_mult_cnj: "cmod (z * cnj z) = (cmod z)\<twosuperior>" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 517 | by (simp add: norm_mult power2_eq_square) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 518 | |
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 519 | interpretation cnj: bounded_linear ["cnj"] | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 520 | apply (unfold_locales) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 521 | apply (rule complex_cnj_add) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 522 | apply (rule complex_cnj_scaleR) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 523 | apply (rule_tac x=1 in exI, simp) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 524 | done | 
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 525 | |
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 526 | |
| 22972 
3e96b98d37c6
generalized sgn function to work on any real normed vector space
 huffman parents: 
22968diff
changeset | 527 | subsection{*The Functions @{term sgn} and @{term arg}*}
 | 
| 14323 | 528 | |
| 22972 
3e96b98d37c6
generalized sgn function to work on any real normed vector space
 huffman parents: 
22968diff
changeset | 529 | text {*------------ Argand -------------*}
 | 
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 530 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20763diff
changeset | 531 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20763diff
changeset | 532 | arg :: "complex => real" where | 
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 533 | "arg z = (SOME a. Re(sgn z) = cos a & Im(sgn z) = sin a & -pi < a & a \<le> pi)" | 
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 534 | |
| 14374 | 535 | lemma sgn_eq: "sgn z = z / complex_of_real (cmod z)" | 
| 24506 | 536 | by (simp add: complex_sgn_def divide_inverse scaleR_conv_of_real mult_commute) | 
| 14323 | 537 | |
| 538 | lemma i_mult_eq: "ii * ii = complex_of_real (-1)" | |
| 20725 
72e20198f834
instance complex :: real_normed_field; cleaned up
 huffman parents: 
20560diff
changeset | 539 | by (simp add: i_def complex_of_real_def) | 
| 14323 | 540 | |
| 14374 | 541 | lemma i_mult_eq2 [simp]: "ii * ii = -(1::complex)" | 
| 20725 
72e20198f834
instance complex :: real_normed_field; cleaned up
 huffman parents: 
20560diff
changeset | 542 | by (simp add: i_def complex_one_def) | 
| 14323 | 543 | |
| 14374 | 544 | lemma complex_eq_cancel_iff2 [simp]: | 
| 14377 | 545 | "(Complex x y = complex_of_real xa) = (x = xa & y = 0)" | 
| 546 | by (simp add: complex_of_real_def) | |
| 14323 | 547 | |
| 14374 | 548 | lemma Re_sgn [simp]: "Re(sgn z) = Re(z)/cmod z" | 
| 24506 | 549 | by (simp add: complex_sgn_def divide_inverse) | 
| 14323 | 550 | |
| 14374 | 551 | lemma Im_sgn [simp]: "Im(sgn z) = Im(z)/cmod z" | 
| 24506 | 552 | by (simp add: complex_sgn_def divide_inverse) | 
| 14323 | 553 | |
| 554 | lemma complex_inverse_complex_split: | |
| 555 | "inverse(complex_of_real x + ii * complex_of_real y) = | |
| 556 | complex_of_real(x/(x ^ 2 + y ^ 2)) - | |
| 557 | ii * complex_of_real(y/(x ^ 2 + y ^ 2))" | |
| 20725 
72e20198f834
instance complex :: real_normed_field; cleaned up
 huffman parents: 
20560diff
changeset | 558 | by (simp add: complex_of_real_def i_def diff_minus divide_inverse) | 
| 14323 | 559 | |
| 560 | (*----------------------------------------------------------------------------*) | |
| 561 | (* Many of the theorems below need to be moved elsewhere e.g. Transc. Also *) | |
| 562 | (* many of the theorems are not used - so should they be kept? *) | |
| 563 | (*----------------------------------------------------------------------------*) | |
| 564 | ||
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 565 | lemma cos_arg_i_mult_zero_pos: | 
| 14377 | 566 | "0 < y ==> cos (arg(Complex 0 y)) = 0" | 
| 14373 | 567 | apply (simp add: arg_def abs_if) | 
| 14334 | 568 | apply (rule_tac a = "pi/2" in someI2, auto) | 
| 569 | apply (rule order_less_trans [of _ 0], auto) | |
| 14323 | 570 | done | 
| 571 | ||
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 572 | lemma cos_arg_i_mult_zero_neg: | 
| 14377 | 573 | "y < 0 ==> cos (arg(Complex 0 y)) = 0" | 
| 14373 | 574 | apply (simp add: arg_def abs_if) | 
| 14334 | 575 | apply (rule_tac a = "- pi/2" in someI2, auto) | 
| 576 | apply (rule order_trans [of _ 0], auto) | |
| 14323 | 577 | done | 
| 578 | ||
| 14374 | 579 | lemma cos_arg_i_mult_zero [simp]: | 
| 14377 | 580 | "y \<noteq> 0 ==> cos (arg(Complex 0 y)) = 0" | 
| 581 | by (auto simp add: linorder_neq_iff cos_arg_i_mult_zero_pos cos_arg_i_mult_zero_neg) | |
| 14323 | 582 | |
| 583 | ||
| 584 | subsection{*Finally! Polar Form for Complex Numbers*}
 | |
| 585 | ||
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 586 | definition | 
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 587 | |
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 588 | (* abbreviation for (cos a + i sin a) *) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20763diff
changeset | 589 | cis :: "real => complex" where | 
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 590 | "cis a = Complex (cos a) (sin a)" | 
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 591 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20763diff
changeset | 592 | definition | 
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 593 | (* abbreviation for r*(cos a + i sin a) *) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20763diff
changeset | 594 | rcis :: "[real, real] => complex" where | 
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 595 | "rcis r a = complex_of_real r * cis a" | 
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 596 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20763diff
changeset | 597 | definition | 
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 598 | (* e ^ (x + iy) *) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20763diff
changeset | 599 | expi :: "complex => complex" where | 
| 20557 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 600 | "expi z = complex_of_real(exp (Re z)) * cis (Im z)" | 
| 
81dd3679f92c
complex_of_real abbreviates of_real::real=>complex;
 huffman parents: 
20556diff
changeset | 601 | |
| 14374 | 602 | lemma complex_split_polar: | 
| 14377 | 603 | "\<exists>r a. z = complex_of_real r * (Complex (cos a) (sin a))" | 
| 20725 
72e20198f834
instance complex :: real_normed_field; cleaned up
 huffman parents: 
20560diff
changeset | 604 | apply (induct z) | 
| 14377 | 605 | apply (auto simp add: polar_Ex complex_of_real_mult_Complex) | 
| 14323 | 606 | done | 
| 607 | ||
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 608 | lemma rcis_Ex: "\<exists>r a. z = rcis r a" | 
| 20725 
72e20198f834
instance complex :: real_normed_field; cleaned up
 huffman parents: 
20560diff
changeset | 609 | apply (induct z) | 
| 14377 | 610 | apply (simp add: rcis_def cis_def polar_Ex complex_of_real_mult_Complex) | 
| 14323 | 611 | done | 
| 612 | ||
| 14374 | 613 | lemma Re_rcis [simp]: "Re(rcis r a) = r * cos a" | 
| 14373 | 614 | by (simp add: rcis_def cis_def) | 
| 14323 | 615 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 616 | lemma Im_rcis [simp]: "Im(rcis r a) = r * sin a" | 
| 14373 | 617 | by (simp add: rcis_def cis_def) | 
| 14323 | 618 | |
| 14377 | 619 | lemma sin_cos_squared_add2_mult: "(r * cos a)\<twosuperior> + (r * sin a)\<twosuperior> = r\<twosuperior>" | 
| 620 | proof - | |
| 621 | have "(r * cos a)\<twosuperior> + (r * sin a)\<twosuperior> = r\<twosuperior> * ((cos a)\<twosuperior> + (sin a)\<twosuperior>)" | |
| 20725 
72e20198f834
instance complex :: real_normed_field; cleaned up
 huffman parents: 
20560diff
changeset | 622 | by (simp only: power_mult_distrib right_distrib) | 
| 14377 | 623 | thus ?thesis by simp | 
| 624 | qed | |
| 14323 | 625 | |
| 14374 | 626 | lemma complex_mod_rcis [simp]: "cmod(rcis r a) = abs r" | 
| 14377 | 627 | by (simp add: rcis_def cis_def sin_cos_squared_add2_mult) | 
| 14323 | 628 | |
| 14374 | 629 | lemma complex_Re_cnj [simp]: "Re(cnj z) = Re z" | 
| 14373 | 630 | by (induct z, simp add: complex_cnj) | 
| 14323 | 631 | |
| 14374 | 632 | lemma complex_Im_cnj [simp]: "Im(cnj z) = - Im z" | 
| 633 | by (induct z, simp add: complex_cnj) | |
| 634 | ||
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 635 | lemma complex_mod_sqrt_Re_mult_cnj: "cmod z = sqrt (Re (z * cnj z))" | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 636 | by (simp add: cmod_def power2_eq_square) | 
| 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 637 | |
| 14374 | 638 | lemma complex_In_mult_cnj_zero [simp]: "Im (z * cnj z) = 0" | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 639 | by simp | 
| 14323 | 640 | |
| 641 | ||
| 642 | (*---------------------------------------------------------------------------*) | |
| 643 | (* (r1 * cis a) * (r2 * cis b) = r1 * r2 * cis (a + b) *) | |
| 644 | (*---------------------------------------------------------------------------*) | |
| 645 | ||
| 646 | lemma cis_rcis_eq: "cis a = rcis 1 a" | |
| 14373 | 647 | by (simp add: rcis_def) | 
| 14323 | 648 | |
| 14374 | 649 | lemma rcis_mult: "rcis r1 a * rcis r2 b = rcis (r1*r2) (a + b)" | 
| 15013 | 650 | by (simp add: rcis_def cis_def cos_add sin_add right_distrib right_diff_distrib | 
| 651 | complex_of_real_def) | |
| 14323 | 652 | |
| 653 | lemma cis_mult: "cis a * cis b = cis (a + b)" | |
| 14373 | 654 | by (simp add: cis_rcis_eq rcis_mult) | 
| 14323 | 655 | |
| 14374 | 656 | lemma cis_zero [simp]: "cis 0 = 1" | 
| 14377 | 657 | by (simp add: cis_def complex_one_def) | 
| 14323 | 658 | |
| 14374 | 659 | lemma rcis_zero_mod [simp]: "rcis 0 a = 0" | 
| 14373 | 660 | by (simp add: rcis_def) | 
| 14323 | 661 | |
| 14374 | 662 | lemma rcis_zero_arg [simp]: "rcis r 0 = complex_of_real r" | 
| 14373 | 663 | by (simp add: rcis_def) | 
| 14323 | 664 | |
| 665 | lemma complex_of_real_minus_one: | |
| 666 | "complex_of_real (-(1::real)) = -(1::complex)" | |
| 20725 
72e20198f834
instance complex :: real_normed_field; cleaned up
 huffman parents: 
20560diff
changeset | 667 | by (simp add: complex_of_real_def complex_one_def) | 
| 14323 | 668 | |
| 14374 | 669 | lemma complex_i_mult_minus [simp]: "ii * (ii * x) = - x" | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 670 | by (simp add: mult_assoc [symmetric]) | 
| 14323 | 671 | |
| 672 | ||
| 673 | lemma cis_real_of_nat_Suc_mult: | |
| 674 | "cis (real (Suc n) * a) = cis a * cis (real n * a)" | |
| 14377 | 675 | by (simp add: cis_def real_of_nat_Suc left_distrib cos_add sin_add right_distrib) | 
| 14323 | 676 | |
| 677 | lemma DeMoivre: "(cis a) ^ n = cis (real n * a)" | |
| 678 | apply (induct_tac "n") | |
| 679 | apply (auto simp add: cis_real_of_nat_Suc_mult) | |
| 680 | done | |
| 681 | ||
| 14374 | 682 | lemma DeMoivre2: "(rcis r a) ^ n = rcis (r ^ n) (real n * a)" | 
| 22890 | 683 | by (simp add: rcis_def power_mult_distrib DeMoivre) | 
| 14323 | 684 | |
| 14374 | 685 | lemma cis_inverse [simp]: "inverse(cis a) = cis (-a)" | 
| 20725 
72e20198f834
instance complex :: real_normed_field; cleaned up
 huffman parents: 
20560diff
changeset | 686 | by (simp add: cis_def complex_inverse_complex_split diff_minus) | 
| 14323 | 687 | |
| 688 | lemma rcis_inverse: "inverse(rcis r a) = rcis (1/r) (-a)" | |
| 22884 | 689 | by (simp add: divide_inverse rcis_def) | 
| 14323 | 690 | |
| 691 | lemma cis_divide: "cis a / cis b = cis (a - b)" | |
| 14373 | 692 | by (simp add: complex_divide_def cis_mult real_diff_def) | 
| 14323 | 693 | |
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 694 | lemma rcis_divide: "rcis r1 a / rcis r2 b = rcis (r1/r2) (a - b)" | 
| 14373 | 695 | apply (simp add: complex_divide_def) | 
| 696 | apply (case_tac "r2=0", simp) | |
| 697 | apply (simp add: rcis_inverse rcis_mult real_diff_def) | |
| 14323 | 698 | done | 
| 699 | ||
| 14374 | 700 | lemma Re_cis [simp]: "Re(cis a) = cos a" | 
| 14373 | 701 | by (simp add: cis_def) | 
| 14323 | 702 | |
| 14374 | 703 | lemma Im_cis [simp]: "Im(cis a) = sin a" | 
| 14373 | 704 | by (simp add: cis_def) | 
| 14323 | 705 | |
| 706 | lemma cos_n_Re_cis_pow_n: "cos (real n * a) = Re(cis a ^ n)" | |
| 14334 | 707 | by (auto simp add: DeMoivre) | 
| 14323 | 708 | |
| 709 | lemma sin_n_Im_cis_pow_n: "sin (real n * a) = Im(cis a ^ n)" | |
| 14334 | 710 | by (auto simp add: DeMoivre) | 
| 14323 | 711 | |
| 712 | lemma expi_add: "expi(a + b) = expi(a) * expi(b)" | |
| 20725 
72e20198f834
instance complex :: real_normed_field; cleaned up
 huffman parents: 
20560diff
changeset | 713 | by (simp add: expi_def exp_add cis_mult [symmetric] mult_ac) | 
| 14323 | 714 | |
| 14374 | 715 | lemma expi_zero [simp]: "expi (0::complex) = 1" | 
| 14373 | 716 | by (simp add: expi_def) | 
| 14323 | 717 | |
| 14374 | 718 | lemma complex_expi_Ex: "\<exists>a r. z = complex_of_real r * expi a" | 
| 14373 | 719 | apply (insert rcis_Ex [of z]) | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 720 | apply (auto simp add: expi_def rcis_def mult_assoc [symmetric]) | 
| 14334 | 721 | apply (rule_tac x = "ii * complex_of_real a" in exI, auto) | 
| 14323 | 722 | done | 
| 723 | ||
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 724 | lemma expi_two_pi_i [simp]: "expi((2::complex) * complex_of_real pi * ii) = 1" | 
| 23125 
6f7b5b96241f
cleaned up proofs; reorganized sections; removed redundant lemmas
 huffman parents: 
23124diff
changeset | 725 | by (simp add: expi_def cis_def) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 726 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 727 | (*examples: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 728 | print_depth 22 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 729 | set timing; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 730 | set trace_simp; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 731 | fun test s = (Goal s, by (Simp_tac 1)); | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 732 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 733 | test "23 * ii + 45 * ii= (x::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 734 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 735 | test "5 * ii + 12 - 45 * ii= (x::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 736 | test "5 * ii + 40 - 12 * ii + 9 = (x::complex) + 89 * ii"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 737 | test "5 * ii + 40 - 12 * ii + 9 - 78 = (x::complex) + 89 * ii"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 738 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 739 | test "l + 10 * ii + 90 + 3*l + 9 + 45 * ii= (x::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 740 | test "87 + 10 * ii + 90 + 3*7 + 9 + 45 * ii= (x::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 741 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 742 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 743 | fun test s = (Goal s; by (Asm_simp_tac 1)); | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 744 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 745 | test "x*k = k*(y::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 746 | test "k = k*(y::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 747 | test "a*(b*c) = (b::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 748 | test "a*(b*c) = d*(b::complex)*(x*a)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 749 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 750 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 751 | test "(x*k) / (k*(y::complex)) = (uu::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 752 | test "(k) / (k*(y::complex)) = (uu::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 753 | test "(a*(b*c)) / ((b::complex)) = (uu::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 754 | test "(a*(b*c)) / (d*(b::complex)*(x*a)) = (uu::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 755 | |
| 15003 | 756 | FIXME: what do we do about this? | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 757 | test "a*(b*c)/(y*z) = d*(b::complex)*(x*a)/z"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 758 | *) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 759 | |
| 13957 | 760 | end |