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