| author | paulson | 
| Thu, 04 Mar 2004 12:06:07 +0100 | |
| changeset 14430 | 5cb24165a2e1 | 
| parent 14421 | ee97b6463cb4 | 
| child 14443 | 75910c7557c5 | 
| 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 | |
| 14323 | 10 | theory Complex = HLog: | 
| 13957 | 11 | |
| 14373 | 12 | datatype complex = Complex real real | 
| 13957 | 13 | |
| 14323 | 14 | instance complex :: zero .. | 
| 15 | instance complex :: one .. | |
| 16 | instance complex :: plus .. | |
| 17 | instance complex :: times .. | |
| 18 | instance complex :: minus .. | |
| 19 | instance complex :: inverse .. | |
| 20 | instance complex :: power .. | |
| 13957 | 21 | |
| 22 | consts | |
| 14373 | 23 |   "ii"    :: complex    ("\<i>")
 | 
| 24 | ||
| 25 | consts Re :: "complex => real" | |
| 26 | primrec "Re (Complex x y) = x" | |
| 27 | ||
| 28 | consts Im :: "complex => real" | |
| 29 | primrec "Im (Complex x y) = y" | |
| 30 | ||
| 31 | lemma complex_surj [simp]: "Complex (Re z) (Im z) = z" | |
| 32 | by (induct z) simp | |
| 13957 | 33 | |
| 34 | constdefs | |
| 35 | ||
| 36 | (*----------- modulus ------------*) | |
| 37 | ||
| 14323 | 38 | cmod :: "complex => real" | 
| 39 | "cmod z == sqrt(Re(z) ^ 2 + Im(z) ^ 2)" | |
| 13957 | 40 | |
| 14323 | 41 | (*----- injection from reals -----*) | 
| 42 | ||
| 43 | complex_of_real :: "real => complex" | |
| 14373 | 44 | "complex_of_real r == Complex r 0" | 
| 14323 | 45 | |
| 13957 | 46 | (*------- complex conjugate ------*) | 
| 47 | ||
| 14323 | 48 | cnj :: "complex => complex" | 
| 14373 | 49 | "cnj z == Complex (Re z) (-Im z)" | 
| 13957 | 50 | |
| 14323 | 51 | (*------------ Argand -------------*) | 
| 13957 | 52 | |
| 14323 | 53 | sgn :: "complex => complex" | 
| 13957 | 54 | "sgn z == z / complex_of_real(cmod z)" | 
| 55 | ||
| 14323 | 56 | arg :: "complex => real" | 
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 57 | "arg z == @a. Re(sgn z) = cos a & Im(sgn z) = sin a & -pi < a & a \<le> pi" | 
| 14323 | 58 | |
| 13957 | 59 | |
| 14323 | 60 | defs (overloaded) | 
| 61 | ||
| 62 | complex_zero_def: | |
| 14373 | 63 | "0 == Complex 0 0" | 
| 13957 | 64 | |
| 14323 | 65 | complex_one_def: | 
| 14373 | 66 | "1 == Complex 1 0" | 
| 14323 | 67 | |
| 14373 | 68 | i_def: "ii == Complex 0 1" | 
| 14323 | 69 | |
| 14373 | 70 | complex_minus_def: "- z == Complex (- Re z) (- Im z)" | 
| 14323 | 71 | |
| 72 | complex_inverse_def: | |
| 14373 | 73 | "inverse z == | 
| 74 | Complex (Re z / ((Re z)\<twosuperior> + (Im z)\<twosuperior>)) (- Im z / ((Re z)\<twosuperior> + (Im z)\<twosuperior>))" | |
| 13957 | 75 | |
| 14323 | 76 | complex_add_def: | 
| 14373 | 77 | "z + w == Complex (Re z + Re w) (Im z + Im w)" | 
| 13957 | 78 | |
| 14323 | 79 | complex_diff_def: | 
| 14373 | 80 | "z - w == z + - (w::complex)" | 
| 13957 | 81 | |
| 14374 | 82 | complex_mult_def: | 
| 14373 | 83 | "z * w == Complex (Re z * Re w - Im z * Im w) (Re z * Im w + Im z * Re w)" | 
| 13957 | 84 | |
| 14373 | 85 | complex_divide_def: "w / (z::complex) == w * inverse z" | 
| 14323 | 86 | |
| 13957 | 87 | |
| 88 | constdefs | |
| 89 | ||
| 90 | (* abbreviation for (cos a + i sin a) *) | |
| 14323 | 91 | cis :: "real => complex" | 
| 14377 | 92 | "cis a == Complex (cos a) (sin a)" | 
| 13957 | 93 | |
| 94 | (* abbreviation for r*(cos a + i sin a) *) | |
| 14323 | 95 | rcis :: "[real, real] => complex" | 
| 13957 | 96 | "rcis r a == complex_of_real r * cis a" | 
| 97 | ||
| 98 | (* e ^ (x + iy) *) | |
| 14323 | 99 | expi :: "complex => complex" | 
| 13957 | 100 | "expi z == complex_of_real(exp (Re z)) * cis (Im z)" | 
| 14323 | 101 | |
| 102 | ||
| 14373 | 103 | lemma complex_equality [intro?]: "Re z = Re w ==> Im z = Im w ==> z = w" | 
| 104 | by (induct z, induct w) simp | |
| 14323 | 105 | |
| 14374 | 106 | lemma Re [simp]: "Re(Complex x y) = x" | 
| 14373 | 107 | by simp | 
| 14323 | 108 | |
| 14374 | 109 | lemma Im [simp]: "Im(Complex x y) = y" | 
| 14373 | 110 | by simp | 
| 14323 | 111 | |
| 112 | lemma complex_Re_Im_cancel_iff: "(w=z) = (Re(w) = Re(z) & Im(w) = Im(z))" | |
| 14373 | 113 | by (induct w, induct z, simp) | 
| 14323 | 114 | |
| 14374 | 115 | lemma complex_Re_zero [simp]: "Re 0 = 0" | 
| 116 | by (simp add: complex_zero_def) | |
| 117 | ||
| 118 | lemma complex_Im_zero [simp]: "Im 0 = 0" | |
| 14373 | 119 | by (simp add: complex_zero_def) | 
| 14323 | 120 | |
| 14374 | 121 | lemma complex_Re_one [simp]: "Re 1 = 1" | 
| 122 | by (simp add: complex_one_def) | |
| 14323 | 123 | |
| 14374 | 124 | lemma complex_Im_one [simp]: "Im 1 = 0" | 
| 14373 | 125 | by (simp add: complex_one_def) | 
| 14323 | 126 | |
| 14374 | 127 | lemma complex_Re_i [simp]: "Re(ii) = 0" | 
| 14373 | 128 | by (simp add: i_def) | 
| 14323 | 129 | |
| 14374 | 130 | lemma complex_Im_i [simp]: "Im(ii) = 1" | 
| 14373 | 131 | by (simp add: i_def) | 
| 14323 | 132 | |
| 14374 | 133 | lemma Re_complex_of_real [simp]: "Re(complex_of_real z) = z" | 
| 14373 | 134 | by (simp add: complex_of_real_def) | 
| 14323 | 135 | |
| 14374 | 136 | lemma Im_complex_of_real [simp]: "Im(complex_of_real z) = 0" | 
| 14373 | 137 | by (simp add: complex_of_real_def) | 
| 14323 | 138 | |
| 139 | ||
| 14374 | 140 | subsection{*Unary Minus*}
 | 
| 14323 | 141 | |
| 14377 | 142 | lemma complex_minus [simp]: "- (Complex x y) = Complex (-x) (-y)" | 
| 14373 | 143 | by (simp add: complex_minus_def) | 
| 14323 | 144 | |
| 14374 | 145 | lemma complex_Re_minus [simp]: "Re (-z) = - Re z" | 
| 14373 | 146 | by (simp add: complex_minus_def) | 
| 14323 | 147 | |
| 14374 | 148 | lemma complex_Im_minus [simp]: "Im (-z) = - Im z" | 
| 149 | by (simp add: complex_minus_def) | |
| 14323 | 150 | |
| 151 | ||
| 152 | subsection{*Addition*}
 | |
| 153 | ||
| 14377 | 154 | lemma complex_add [simp]: | 
| 155 | "Complex x1 y1 + Complex x2 y2 = Complex (x1+x2) (y1+y2)" | |
| 14373 | 156 | by (simp add: complex_add_def) | 
| 14323 | 157 | |
| 14374 | 158 | lemma complex_Re_add [simp]: "Re(x + y) = Re(x) + Re(y)" | 
| 14373 | 159 | by (simp add: complex_add_def) | 
| 14323 | 160 | |
| 14374 | 161 | lemma complex_Im_add [simp]: "Im(x + y) = Im(x) + Im(y)" | 
| 14373 | 162 | by (simp add: complex_add_def) | 
| 14323 | 163 | |
| 164 | lemma complex_add_commute: "(u::complex) + v = v + u" | |
| 14373 | 165 | by (simp add: complex_add_def add_commute) | 
| 14323 | 166 | |
| 167 | lemma complex_add_assoc: "((u::complex) + v) + w = u + (v + w)" | |
| 14373 | 168 | by (simp add: complex_add_def add_assoc) | 
| 14323 | 169 | |
| 170 | lemma complex_add_zero_left: "(0::complex) + z = z" | |
| 14373 | 171 | by (simp add: complex_add_def complex_zero_def) | 
| 14323 | 172 | |
| 173 | lemma complex_add_zero_right: "z + (0::complex) = z" | |
| 14373 | 174 | by (simp add: complex_add_def complex_zero_def) | 
| 14323 | 175 | |
| 14373 | 176 | lemma complex_add_minus_left: "-z + z = (0::complex)" | 
| 177 | by (simp add: complex_add_def complex_minus_def complex_zero_def) | |
| 14323 | 178 | |
| 179 | lemma complex_diff: | |
| 14373 | 180 | "Complex x1 y1 - Complex x2 y2 = Complex (x1-x2) (y1-y2)" | 
| 181 | by (simp add: complex_add_def complex_minus_def complex_diff_def) | |
| 14323 | 182 | |
| 14374 | 183 | lemma complex_Re_diff [simp]: "Re(x - y) = Re(x) - Re(y)" | 
| 184 | by (simp add: complex_diff_def) | |
| 185 | ||
| 186 | lemma complex_Im_diff [simp]: "Im(x - y) = Im(x) - Im(y)" | |
| 187 | by (simp add: complex_diff_def) | |
| 188 | ||
| 189 | ||
| 14323 | 190 | subsection{*Multiplication*}
 | 
| 191 | ||
| 14377 | 192 | lemma complex_mult [simp]: | 
| 14373 | 193 | "Complex x1 y1 * Complex x2 y2 = Complex (x1*x2 - y1*y2) (x1*y2 + y1*x2)" | 
| 194 | by (simp add: complex_mult_def) | |
| 14323 | 195 | |
| 196 | lemma complex_mult_commute: "(w::complex) * z = z * w" | |
| 14373 | 197 | by (simp add: complex_mult_def mult_commute add_commute) | 
| 14323 | 198 | |
| 199 | lemma complex_mult_assoc: "((u::complex) * v) * w = u * (v * w)" | |
| 14374 | 200 | by (simp add: complex_mult_def mult_ac add_ac | 
| 14373 | 201 | right_diff_distrib right_distrib left_diff_distrib left_distrib) | 
| 14323 | 202 | |
| 203 | lemma complex_mult_one_left: "(1::complex) * z = z" | |
| 14373 | 204 | by (simp add: complex_mult_def complex_one_def) | 
| 14323 | 205 | |
| 206 | lemma complex_mult_one_right: "z * (1::complex) = z" | |
| 14373 | 207 | by (simp add: complex_mult_def complex_one_def) | 
| 14323 | 208 | |
| 209 | ||
| 210 | subsection{*Inverse*}
 | |
| 211 | ||
| 14377 | 212 | lemma complex_inverse [simp]: | 
| 14373 | 213 | "inverse (Complex x y) = Complex (x/(x ^ 2 + y ^ 2)) (-y/(x ^ 2 + y ^ 2))" | 
| 214 | by (simp add: complex_inverse_def) | |
| 14335 | 215 | |
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 216 | lemma complex_mult_inv_left: "z \<noteq> (0::complex) ==> inverse(z) * z = 1" | 
| 14374 | 217 | apply (induct z) | 
| 218 | apply (rename_tac x y) | |
| 219 | apply (auto simp add: complex_mult complex_inverse complex_one_def | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 220 | complex_zero_def add_divide_distrib [symmetric] power2_eq_square mult_ac) | 
| 14334 | 221 | apply (drule_tac y = y in real_sum_squares_not_zero) | 
| 222 | apply (drule_tac [2] x = x in real_sum_squares_not_zero2, auto) | |
| 14323 | 223 | done | 
| 224 | ||
| 14335 | 225 | |
| 226 | subsection {* The field of complex numbers *}
 | |
| 227 | ||
| 228 | instance complex :: field | |
| 229 | proof | |
| 230 | fix z u v w :: complex | |
| 231 | show "(u + v) + w = u + (v + w)" | |
| 14374 | 232 | by (rule complex_add_assoc) | 
| 14335 | 233 | show "z + w = w + z" | 
| 14374 | 234 | by (rule complex_add_commute) | 
| 14335 | 235 | show "0 + z = z" | 
| 14374 | 236 | by (rule complex_add_zero_left) | 
| 14335 | 237 | show "-z + z = 0" | 
| 14374 | 238 | by (rule complex_add_minus_left) | 
| 14335 | 239 | show "z - w = z + -w" | 
| 240 | by (simp add: complex_diff_def) | |
| 241 | show "(u * v) * w = u * (v * w)" | |
| 14374 | 242 | by (rule complex_mult_assoc) | 
| 14335 | 243 | show "z * w = w * z" | 
| 14374 | 244 | by (rule complex_mult_commute) | 
| 14335 | 245 | show "1 * z = z" | 
| 14374 | 246 | by (rule complex_mult_one_left) | 
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14335diff
changeset | 247 | show "0 \<noteq> (1::complex)" | 
| 14373 | 248 | by (simp add: complex_zero_def complex_one_def) | 
| 14335 | 249 | show "(u + v) * w = u * w + v * w" | 
| 14421 
ee97b6463cb4
new Ring_and_Field hierarchy, eliminating redundant axioms
 paulson parents: 
14387diff
changeset | 250 | by (simp add: complex_mult_def complex_add_def left_distrib | 
| 
ee97b6463cb4
new Ring_and_Field hierarchy, eliminating redundant axioms
 paulson parents: 
14387diff
changeset | 251 | diff_minus add_ac) | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 252 | show "z / w = z * inverse w" | 
| 14335 | 253 | by (simp add: complex_divide_def) | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 254 | assume "w \<noteq> 0" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 255 | thus "inverse w * w = 1" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 256 | by (simp add: complex_mult_inv_left) | 
| 14335 | 257 | qed | 
| 258 | ||
| 14373 | 259 | instance complex :: division_by_zero | 
| 260 | proof | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 261 | show "inverse 0 = (0::complex)" | 
| 14373 | 262 | by (simp add: complex_inverse_def complex_zero_def) | 
| 263 | qed | |
| 14335 | 264 | |
| 14323 | 265 | |
| 266 | subsection{*Embedding Properties for @{term complex_of_real} Map*}
 | |
| 267 | ||
| 14377 | 268 | lemma Complex_add_complex_of_real [simp]: | 
| 269 | "Complex x y + complex_of_real r = Complex (x+r) y" | |
| 270 | by (simp add: complex_of_real_def) | |
| 271 | ||
| 272 | lemma complex_of_real_add_Complex [simp]: | |
| 273 | "complex_of_real r + Complex x y = Complex (r+x) y" | |
| 274 | by (simp add: i_def complex_of_real_def) | |
| 275 | ||
| 276 | lemma Complex_mult_complex_of_real: | |
| 277 | "Complex x y * complex_of_real r = Complex (x*r) (y*r)" | |
| 278 | by (simp add: complex_of_real_def) | |
| 279 | ||
| 280 | lemma complex_of_real_mult_Complex: | |
| 281 | "complex_of_real r * Complex x y = Complex (r*x) (r*y)" | |
| 282 | by (simp add: i_def complex_of_real_def) | |
| 283 | ||
| 284 | lemma i_complex_of_real [simp]: "ii * complex_of_real r = Complex 0 r" | |
| 285 | by (simp add: i_def complex_of_real_def) | |
| 286 | ||
| 287 | lemma complex_of_real_i [simp]: "complex_of_real r * ii = Complex 0 r" | |
| 288 | by (simp add: i_def complex_of_real_def) | |
| 289 | ||
| 14374 | 290 | lemma complex_of_real_one [simp]: "complex_of_real 1 = 1" | 
| 14373 | 291 | by (simp add: complex_one_def complex_of_real_def) | 
| 14323 | 292 | |
| 14374 | 293 | lemma complex_of_real_zero [simp]: "complex_of_real 0 = 0" | 
| 14373 | 294 | by (simp add: complex_zero_def complex_of_real_def) | 
| 14323 | 295 | |
| 14374 | 296 | lemma complex_of_real_eq_iff [iff]: | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 297 | "(complex_of_real x = complex_of_real y) = (x = y)" | 
| 14374 | 298 | by (simp add: complex_of_real_def) | 
| 14323 | 299 | |
| 300 | lemma complex_of_real_minus: "complex_of_real(-x) = - complex_of_real x" | |
| 14373 | 301 | by (simp add: complex_of_real_def complex_minus) | 
| 14323 | 302 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 303 | lemma complex_of_real_inverse: | 
| 14374 | 304 | "complex_of_real(inverse x) = inverse(complex_of_real x)" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 305 | apply (case_tac "x=0", simp) | 
| 14374 | 306 | apply (simp add: complex_inverse complex_of_real_def real_divide_def | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 307 | inverse_mult_distrib power2_eq_square) | 
| 14323 | 308 | done | 
| 309 | ||
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 310 | lemma complex_of_real_add: | 
| 14373 | 311 | "complex_of_real x + complex_of_real y = complex_of_real (x + y)" | 
| 312 | by (simp add: complex_add complex_of_real_def) | |
| 14323 | 313 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 314 | lemma complex_of_real_diff: | 
| 14373 | 315 | "complex_of_real x - complex_of_real y = complex_of_real (x - y)" | 
| 14374 | 316 | by (simp add: complex_of_real_minus [symmetric] complex_diff_def | 
| 317 | complex_of_real_add) | |
| 14323 | 318 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 319 | lemma complex_of_real_mult: | 
| 14373 | 320 | "complex_of_real x * complex_of_real y = complex_of_real (x * y)" | 
| 321 | by (simp add: complex_mult complex_of_real_def) | |
| 14323 | 322 | |
| 323 | lemma complex_of_real_divide: | |
| 324 | "complex_of_real x / complex_of_real y = complex_of_real(x/y)" | |
| 14373 | 325 | apply (simp add: complex_divide_def) | 
| 326 | apply (case_tac "y=0", simp) | |
| 14374 | 327 | apply (simp add: complex_of_real_mult [symmetric] complex_of_real_inverse | 
| 328 | real_divide_def) | |
| 14323 | 329 | done | 
| 330 | ||
| 14377 | 331 | lemma complex_mod [simp]: "cmod (Complex x y) = sqrt(x ^ 2 + y ^ 2)" | 
| 14373 | 332 | by (simp add: cmod_def) | 
| 14323 | 333 | |
| 14374 | 334 | lemma complex_mod_zero [simp]: "cmod(0) = 0" | 
| 14373 | 335 | by (simp add: cmod_def) | 
| 14323 | 336 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 337 | lemma complex_mod_one [simp]: "cmod(1) = 1" | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 338 | by (simp add: cmod_def power2_eq_square) | 
| 14323 | 339 | |
| 14374 | 340 | lemma complex_mod_complex_of_real [simp]: "cmod(complex_of_real x) = abs x" | 
| 14373 | 341 | by (simp add: complex_of_real_def power2_eq_square complex_mod) | 
| 14323 | 342 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 343 | lemma complex_of_real_abs: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 344 | "complex_of_real (abs x) = complex_of_real(cmod(complex_of_real x))" | 
| 14373 | 345 | by simp | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 346 | |
| 14323 | 347 | |
| 14377 | 348 | subsection{*The Functions @{term Re} and @{term Im}*}
 | 
| 349 | ||
| 350 | lemma complex_Re_mult_eq: "Re (w * z) = Re w * Re z - Im w * Im z" | |
| 351 | by (induct z, induct w, simp add: complex_mult) | |
| 352 | ||
| 353 | lemma complex_Im_mult_eq: "Im (w * z) = Re w * Im z + Im w * Re z" | |
| 354 | by (induct z, induct w, simp add: complex_mult) | |
| 355 | ||
| 356 | lemma Re_i_times [simp]: "Re(ii * z) = - Im z" | |
| 357 | by (simp add: complex_Re_mult_eq) | |
| 358 | ||
| 359 | lemma Re_times_i [simp]: "Re(z * ii) = - Im z" | |
| 360 | by (simp add: complex_Re_mult_eq) | |
| 361 | ||
| 362 | lemma Im_i_times [simp]: "Im(ii * z) = Re z" | |
| 363 | by (simp add: complex_Im_mult_eq) | |
| 364 | ||
| 365 | lemma Im_times_i [simp]: "Im(z * ii) = Re z" | |
| 366 | by (simp add: complex_Im_mult_eq) | |
| 367 | ||
| 368 | lemma complex_Re_mult: "[| Im w = 0; Im z = 0 |] ==> Re(w * z) = Re(w) * Re(z)" | |
| 369 | by (simp add: complex_Re_mult_eq) | |
| 370 | ||
| 371 | lemma complex_Re_mult_complex_of_real [simp]: | |
| 372 | "Re (z * complex_of_real c) = Re(z) * c" | |
| 373 | by (simp add: complex_Re_mult_eq) | |
| 374 | ||
| 375 | lemma complex_Im_mult_complex_of_real [simp]: | |
| 376 | "Im (z * complex_of_real c) = Im(z) * c" | |
| 377 | by (simp add: complex_Im_mult_eq) | |
| 378 | ||
| 379 | lemma complex_Re_mult_complex_of_real2 [simp]: | |
| 380 | "Re (complex_of_real c * z) = c * Re(z)" | |
| 381 | by (simp add: complex_Re_mult_eq) | |
| 382 | ||
| 383 | lemma complex_Im_mult_complex_of_real2 [simp]: | |
| 384 | "Im (complex_of_real c * z) = c * Im(z)" | |
| 385 | by (simp add: complex_Im_mult_eq) | |
| 386 | ||
| 387 | ||
| 14323 | 388 | subsection{*Conjugation is an Automorphism*}
 | 
| 389 | ||
| 14373 | 390 | lemma complex_cnj: "cnj (Complex x y) = Complex x (-y)" | 
| 391 | by (simp add: cnj_def) | |
| 14323 | 392 | |
| 14374 | 393 | lemma complex_cnj_cancel_iff [simp]: "(cnj x = cnj y) = (x = y)" | 
| 14373 | 394 | by (simp add: cnj_def complex_Re_Im_cancel_iff) | 
| 14323 | 395 | |
| 14374 | 396 | lemma complex_cnj_cnj [simp]: "cnj (cnj z) = z" | 
| 14373 | 397 | by (simp add: cnj_def) | 
| 14323 | 398 | |
| 14374 | 399 | lemma complex_cnj_complex_of_real [simp]: | 
| 14373 | 400 | "cnj (complex_of_real x) = complex_of_real x" | 
| 401 | by (simp add: complex_of_real_def complex_cnj) | |
| 14323 | 402 | |
| 14374 | 403 | lemma complex_mod_cnj [simp]: "cmod (cnj z) = cmod z" | 
| 14373 | 404 | by (induct z, simp add: complex_cnj complex_mod power2_eq_square) | 
| 14323 | 405 | |
| 406 | lemma complex_cnj_minus: "cnj (-z) = - cnj z" | |
| 14373 | 407 | by (simp add: cnj_def complex_minus complex_Re_minus complex_Im_minus) | 
| 14323 | 408 | |
| 409 | lemma complex_cnj_inverse: "cnj(inverse z) = inverse(cnj z)" | |
| 14373 | 410 | by (induct z, simp add: complex_cnj complex_inverse power2_eq_square) | 
| 14323 | 411 | |
| 412 | lemma complex_cnj_add: "cnj(w + z) = cnj(w) + cnj(z)" | |
| 14373 | 413 | by (induct w, induct z, simp add: complex_cnj complex_add) | 
| 14323 | 414 | |
| 415 | lemma complex_cnj_diff: "cnj(w - z) = cnj(w) - cnj(z)" | |
| 14373 | 416 | by (simp add: complex_diff_def complex_cnj_add complex_cnj_minus) | 
| 14323 | 417 | |
| 418 | lemma complex_cnj_mult: "cnj(w * z) = cnj(w) * cnj(z)" | |
| 14373 | 419 | by (induct w, induct z, simp add: complex_cnj complex_mult) | 
| 14323 | 420 | |
| 421 | lemma complex_cnj_divide: "cnj(w / z) = (cnj w)/(cnj z)" | |
| 14373 | 422 | by (simp add: complex_divide_def complex_cnj_mult complex_cnj_inverse) | 
| 14323 | 423 | |
| 14374 | 424 | lemma complex_cnj_one [simp]: "cnj 1 = 1" | 
| 14373 | 425 | by (simp add: cnj_def complex_one_def) | 
| 14323 | 426 | |
| 427 | lemma complex_add_cnj: "z + cnj z = complex_of_real (2 * Re(z))" | |
| 14373 | 428 | by (induct z, simp add: complex_add complex_cnj complex_of_real_def) | 
| 14323 | 429 | |
| 430 | lemma complex_diff_cnj: "z - cnj z = complex_of_real (2 * Im(z)) * ii" | |
| 14373 | 431 | apply (induct z) | 
| 14374 | 432 | apply (simp add: complex_add complex_cnj complex_of_real_def complex_diff_def | 
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 433 | complex_minus i_def complex_mult) | 
| 14323 | 434 | done | 
| 435 | ||
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 436 | lemma complex_cnj_zero [simp]: "cnj 0 = 0" | 
| 14334 | 437 | by (simp add: cnj_def complex_zero_def) | 
| 14323 | 438 | |
| 14374 | 439 | lemma complex_cnj_zero_iff [iff]: "(cnj z = 0) = (z = 0)" | 
| 14373 | 440 | by (induct z, simp add: complex_zero_def complex_cnj) | 
| 14323 | 441 | |
| 442 | lemma complex_mult_cnj: "z * cnj z = complex_of_real (Re(z) ^ 2 + Im(z) ^ 2)" | |
| 14374 | 443 | by (induct z, | 
| 444 | simp add: complex_cnj complex_mult complex_of_real_def power2_eq_square) | |
| 14323 | 445 | |
| 446 | ||
| 447 | subsection{*Modulus*}
 | |
| 448 | ||
| 14374 | 449 | lemma complex_mod_eq_zero_cancel [simp]: "(cmod x = 0) = (x = 0)" | 
| 14373 | 450 | apply (induct x) | 
| 14374 | 451 | apply (auto intro: real_sum_squares_cancel real_sum_squares_cancel2 | 
| 14373 | 452 | simp add: complex_mod complex_zero_def power2_eq_square) | 
| 14323 | 453 | done | 
| 454 | ||
| 14374 | 455 | lemma complex_mod_complex_of_real_of_nat [simp]: | 
| 14373 | 456 | "cmod (complex_of_real(real (n::nat))) = real n" | 
| 457 | by simp | |
| 14323 | 458 | |
| 14374 | 459 | lemma complex_mod_minus [simp]: "cmod (-x) = cmod(x)" | 
| 14373 | 460 | by (induct x, simp add: complex_mod complex_minus power2_eq_square) | 
| 14323 | 461 | |
| 462 | lemma complex_mod_mult_cnj: "cmod(z * cnj(z)) = cmod(z) ^ 2" | |
| 14373 | 463 | apply (induct z, simp add: complex_mod complex_cnj complex_mult) | 
| 464 | apply (simp add: power2_eq_square real_abs_def) | |
| 14323 | 465 | done | 
| 466 | ||
| 14373 | 467 | lemma complex_mod_squared: "cmod(Complex x y) ^ 2 = x ^ 2 + y ^ 2" | 
| 468 | by (simp add: cmod_def) | |
| 14323 | 469 | |
| 14374 | 470 | lemma complex_mod_ge_zero [simp]: "0 \<le> cmod x" | 
| 14373 | 471 | by (simp add: cmod_def) | 
| 14323 | 472 | |
| 14374 | 473 | lemma abs_cmod_cancel [simp]: "abs(cmod x) = cmod x" | 
| 474 | by (simp add: abs_if linorder_not_less) | |
| 14323 | 475 | |
| 476 | lemma complex_mod_mult: "cmod(x*y) = cmod(x) * cmod(y)" | |
| 14373 | 477 | apply (induct x, induct y) | 
| 14377 | 478 | apply (auto simp add: complex_mult complex_mod real_sqrt_mult_distrib2[symmetric]) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 479 | apply (rule_tac n = 1 in power_inject_base) | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 480 | apply (auto simp add: power2_eq_square [symmetric] simp del: realpow_Suc) | 
| 14374 | 481 | apply (auto simp add: real_diff_def power2_eq_square right_distrib left_distrib | 
| 482 | add_ac mult_ac) | |
| 14323 | 483 | done | 
| 484 | ||
| 14377 | 485 | lemma cmod_unit_one [simp]: "cmod (Complex (cos a) (sin a)) = 1" | 
| 486 | by (simp add: cmod_def) | |
| 487 | ||
| 488 | lemma cmod_complex_polar [simp]: | |
| 489 | "cmod (complex_of_real r * Complex (cos a) (sin a)) = abs r" | |
| 490 | by (simp only: cmod_unit_one complex_mod_mult, simp) | |
| 491 | ||
| 14374 | 492 | lemma complex_mod_add_squared_eq: | 
| 493 | "cmod(x + y) ^ 2 = cmod(x) ^ 2 + cmod(y) ^ 2 + 2 * Re(x * cnj y)" | |
| 14373 | 494 | apply (induct x, induct y) | 
| 14323 | 495 | apply (auto simp add: complex_add complex_mod_squared complex_mult complex_cnj real_diff_def simp del: realpow_Suc) | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 496 | apply (auto simp add: right_distrib left_distrib power2_eq_square mult_ac add_ac) | 
| 14323 | 497 | done | 
| 498 | ||
| 14374 | 499 | lemma complex_Re_mult_cnj_le_cmod [simp]: "Re(x * cnj y) \<le> cmod(x * cnj y)" | 
| 14373 | 500 | apply (induct x, induct y) | 
| 14323 | 501 | apply (auto simp add: complex_mod complex_mult complex_cnj real_diff_def simp del: realpow_Suc) | 
| 502 | done | |
| 503 | ||
| 14374 | 504 | lemma complex_Re_mult_cnj_le_cmod2 [simp]: "Re(x * cnj y) \<le> cmod(x * y)" | 
| 14373 | 505 | by (insert complex_Re_mult_cnj_le_cmod [of x y], simp add: complex_mod_mult) | 
| 14323 | 506 | |
| 14374 | 507 | lemma real_sum_squared_expand: | 
| 508 | "((x::real) + y) ^ 2 = x ^ 2 + y ^ 2 + 2 * x * y" | |
| 14373 | 509 | by (simp add: left_distrib right_distrib power2_eq_square) | 
| 14323 | 510 | |
| 14374 | 511 | lemma complex_mod_triangle_squared [simp]: | 
| 512 | "cmod (x + y) ^ 2 \<le> (cmod(x) + cmod(y)) ^ 2" | |
| 14373 | 513 | by (simp add: real_sum_squared_expand complex_mod_add_squared_eq real_mult_assoc complex_mod_mult [symmetric]) | 
| 14323 | 514 | |
| 14374 | 515 | lemma complex_mod_minus_le_complex_mod [simp]: "- cmod x \<le> cmod x" | 
| 14373 | 516 | by (rule order_trans [OF _ complex_mod_ge_zero], simp) | 
| 14323 | 517 | |
| 14374 | 518 | lemma complex_mod_triangle_ineq [simp]: "cmod (x + y) \<le> cmod(x) + cmod(y)" | 
| 14334 | 519 | apply (rule_tac n = 1 in realpow_increasing) | 
| 14323 | 520 | apply (auto intro: order_trans [OF _ complex_mod_ge_zero] | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 521 | simp add: power2_eq_square [symmetric]) | 
| 14323 | 522 | done | 
| 523 | ||
| 14374 | 524 | lemma complex_mod_triangle_ineq2 [simp]: "cmod(b + a) - cmod b \<le> cmod a" | 
| 14373 | 525 | by (insert complex_mod_triangle_ineq [THEN add_right_mono, of b a"-cmod b"], simp) | 
| 14323 | 526 | |
| 527 | lemma complex_mod_diff_commute: "cmod (x - y) = cmod (y - x)" | |
| 14373 | 528 | apply (induct x, induct y) | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 529 | apply (auto simp add: complex_diff complex_mod right_diff_distrib power2_eq_square left_diff_distrib add_ac mult_ac) | 
| 14323 | 530 | done | 
| 531 | ||
| 14374 | 532 | lemma complex_mod_add_less: | 
| 533 | "[| cmod x < r; cmod y < s |] ==> cmod (x + y) < r + s" | |
| 14334 | 534 | by (auto intro: order_le_less_trans complex_mod_triangle_ineq) | 
| 14323 | 535 | |
| 14374 | 536 | lemma complex_mod_mult_less: | 
| 537 | "[| cmod x < r; cmod y < s |] ==> cmod (x * y) < r * s" | |
| 14334 | 538 | by (auto intro: real_mult_less_mono' simp add: complex_mod_mult) | 
| 14323 | 539 | |
| 14374 | 540 | lemma complex_mod_diff_ineq [simp]: "cmod(a) - cmod(b) \<le> cmod(a + b)" | 
| 14323 | 541 | apply (rule linorder_cases [of "cmod(a)" "cmod (b)"]) | 
| 542 | apply auto | |
| 14334 | 543 | apply (rule order_trans [of _ 0], rule order_less_imp_le) | 
| 14374 | 544 | apply (simp add: compare_rls, simp) | 
| 14323 | 545 | apply (simp add: compare_rls) | 
| 546 | apply (rule complex_mod_minus [THEN subst]) | |
| 547 | apply (rule order_trans) | |
| 548 | apply (rule_tac [2] complex_mod_triangle_ineq) | |
| 14373 | 549 | apply (auto simp add: add_ac) | 
| 14323 | 550 | done | 
| 551 | ||
| 14374 | 552 | lemma complex_Re_le_cmod [simp]: "Re z \<le> cmod z" | 
| 14373 | 553 | by (induct z, simp add: complex_mod del: realpow_Suc) | 
| 14323 | 554 | |
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 555 | lemma complex_mod_gt_zero: "z \<noteq> 0 ==> 0 < cmod z" | 
| 14373 | 556 | apply (insert complex_mod_ge_zero [of z]) | 
| 14334 | 557 | apply (drule order_le_imp_less_or_eq, auto) | 
| 14323 | 558 | done | 
| 559 | ||
| 560 | ||
| 561 | subsection{*A Few More Theorems*}
 | |
| 562 | ||
| 563 | lemma complex_mod_inverse: "cmod(inverse x) = inverse(cmod x)" | |
| 14373 | 564 | apply (case_tac "x=0", simp) | 
| 14323 | 565 | apply (rule_tac c1 = "cmod x" in real_mult_left_cancel [THEN iffD1]) | 
| 566 | apply (auto simp add: complex_mod_mult [symmetric]) | |
| 567 | done | |
| 568 | ||
| 14373 | 569 | lemma complex_mod_divide: "cmod(x/y) = cmod(x)/(cmod y)" | 
| 14377 | 570 | by (simp add: complex_divide_def real_divide_def complex_mod_mult complex_mod_inverse) | 
| 14323 | 571 | |
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 572 | |
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 573 | subsection{*Exponentiation*}
 | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 574 | |
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 575 | primrec | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 576 | complexpow_0: "z ^ 0 = 1" | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 577 | complexpow_Suc: "z ^ (Suc n) = (z::complex) * (z ^ n)" | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 578 | |
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 579 | |
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 580 | instance complex :: ringpower | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 581 | proof | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 582 | fix z :: complex | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 583 | fix n :: nat | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 584 | show "z^0 = 1" by simp | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 585 | show "z^(Suc n) = z * (z^n)" by simp | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 586 | qed | 
| 14323 | 587 | |
| 588 | ||
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 589 | lemma complex_of_real_pow: "complex_of_real (x ^ n) = (complex_of_real x) ^ n" | 
| 14323 | 590 | apply (induct_tac "n") | 
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 591 | apply (auto simp add: complex_of_real_mult [symmetric]) | 
| 14323 | 592 | done | 
| 593 | ||
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 594 | lemma complex_cnj_pow: "cnj(z ^ n) = cnj(z) ^ n" | 
| 14323 | 595 | apply (induct_tac "n") | 
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 596 | apply (auto simp add: complex_cnj_mult) | 
| 14323 | 597 | done | 
| 598 | ||
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 599 | lemma complex_mod_complexpow: "cmod(x ^ n) = cmod(x) ^ n" | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 600 | apply (induct_tac "n") | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 601 | apply (auto simp add: complex_mod_mult) | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 602 | done | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 603 | |
| 14374 | 604 | lemma complexpow_minus: | 
| 605 | "(-x::complex) ^ n = (if even n then (x ^ n) else -(x ^ n))" | |
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 606 | by (induct_tac "n", auto) | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 607 | |
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 608 | lemma complexpow_i_squared [simp]: "ii ^ 2 = -(1::complex)" | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 609 | by (simp add: i_def complex_mult complex_one_def complex_minus numeral_2_eq_2) | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 610 | |
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 611 | lemma complex_i_not_zero [simp]: "ii \<noteq> 0" | 
| 14373 | 612 | by (simp add: i_def complex_zero_def) | 
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 613 | |
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 614 | |
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 615 | subsection{*The Function @{term sgn}*}
 | 
| 14323 | 616 | |
| 14374 | 617 | lemma sgn_zero [simp]: "sgn 0 = 0" | 
| 14373 | 618 | by (simp add: sgn_def) | 
| 14323 | 619 | |
| 14374 | 620 | lemma sgn_one [simp]: "sgn 1 = 1" | 
| 14373 | 621 | by (simp add: sgn_def) | 
| 14323 | 622 | |
| 623 | lemma sgn_minus: "sgn (-z) = - sgn(z)" | |
| 14373 | 624 | by (simp add: sgn_def) | 
| 14323 | 625 | |
| 14374 | 626 | lemma sgn_eq: "sgn z = z / complex_of_real (cmod z)" | 
| 14377 | 627 | by (simp add: sgn_def) | 
| 14323 | 628 | |
| 629 | lemma i_mult_eq: "ii * ii = complex_of_real (-1)" | |
| 14373 | 630 | by (simp add: i_def complex_of_real_def complex_mult complex_add) | 
| 14323 | 631 | |
| 14374 | 632 | lemma i_mult_eq2 [simp]: "ii * ii = -(1::complex)" | 
| 14373 | 633 | by (simp add: i_def complex_one_def complex_mult complex_minus) | 
| 14323 | 634 | |
| 14374 | 635 | lemma complex_eq_cancel_iff2 [simp]: | 
| 14377 | 636 | "(Complex x y = complex_of_real xa) = (x = xa & y = 0)" | 
| 637 | by (simp add: complex_of_real_def) | |
| 14323 | 638 | |
| 14374 | 639 | lemma complex_eq_cancel_iff2a [simp]: | 
| 14377 | 640 | "(Complex x y = complex_of_real xa) = (x = xa & y = 0)" | 
| 641 | by (simp add: complex_of_real_def) | |
| 14323 | 642 | |
| 14377 | 643 | lemma Complex_eq_0 [simp]: "(Complex x y = 0) = (x = 0 & y = 0)" | 
| 644 | by (simp add: complex_zero_def) | |
| 14323 | 645 | |
| 14377 | 646 | lemma Complex_eq_1 [simp]: "(Complex x y = 1) = (x = 1 & y = 0)" | 
| 647 | by (simp add: complex_one_def) | |
| 14323 | 648 | |
| 14377 | 649 | lemma Complex_eq_i [simp]: "(Complex x y = ii) = (x = 0 & y = 1)" | 
| 650 | by (simp add: i_def) | |
| 14323 | 651 | |
| 14374 | 652 | lemma Re_sgn [simp]: "Re(sgn z) = Re(z)/cmod z" | 
| 14373 | 653 | apply (induct z) | 
| 654 | apply (simp add: sgn_def complex_divide_def complex_of_real_inverse [symmetric]) | |
| 655 | apply (simp add: complex_of_real_def complex_mult real_divide_def) | |
| 14323 | 656 | done | 
| 657 | ||
| 14374 | 658 | lemma Im_sgn [simp]: "Im(sgn z) = Im(z)/cmod z" | 
| 14373 | 659 | apply (induct z) | 
| 660 | apply (simp add: sgn_def complex_divide_def complex_of_real_inverse [symmetric]) | |
| 661 | apply (simp add: complex_of_real_def complex_mult real_divide_def) | |
| 14323 | 662 | done | 
| 663 | ||
| 664 | lemma complex_inverse_complex_split: | |
| 665 | "inverse(complex_of_real x + ii * complex_of_real y) = | |
| 666 | complex_of_real(x/(x ^ 2 + y ^ 2)) - | |
| 667 | ii * complex_of_real(y/(x ^ 2 + y ^ 2))" | |
| 14374 | 668 | by (simp add: complex_of_real_def i_def complex_mult complex_add | 
| 14373 | 669 | complex_diff_def complex_minus complex_inverse real_divide_def) | 
| 14323 | 670 | |
| 671 | (*----------------------------------------------------------------------------*) | |
| 672 | (* Many of the theorems below need to be moved elsewhere e.g. Transc. Also *) | |
| 673 | (* many of the theorems are not used - so should they be kept? *) | |
| 674 | (*----------------------------------------------------------------------------*) | |
| 675 | ||
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 676 | lemma complex_of_real_zero_iff [simp]: "(complex_of_real y = 0) = (y = 0)" | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 677 | by (auto simp add: complex_zero_def complex_of_real_def) | 
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 678 | |
| 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 679 | lemma cos_arg_i_mult_zero_pos: | 
| 14377 | 680 | "0 < y ==> cos (arg(Complex 0 y)) = 0" | 
| 14373 | 681 | apply (simp add: arg_def abs_if) | 
| 14334 | 682 | apply (rule_tac a = "pi/2" in someI2, auto) | 
| 683 | apply (rule order_less_trans [of _ 0], auto) | |
| 14323 | 684 | done | 
| 685 | ||
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 686 | lemma cos_arg_i_mult_zero_neg: | 
| 14377 | 687 | "y < 0 ==> cos (arg(Complex 0 y)) = 0" | 
| 14373 | 688 | apply (simp add: arg_def abs_if) | 
| 14334 | 689 | apply (rule_tac a = "- pi/2" in someI2, auto) | 
| 690 | apply (rule order_trans [of _ 0], auto) | |
| 14323 | 691 | done | 
| 692 | ||
| 14374 | 693 | lemma cos_arg_i_mult_zero [simp]: | 
| 14377 | 694 | "y \<noteq> 0 ==> cos (arg(Complex 0 y)) = 0" | 
| 695 | by (auto simp add: linorder_neq_iff cos_arg_i_mult_zero_pos cos_arg_i_mult_zero_neg) | |
| 14323 | 696 | |
| 697 | ||
| 698 | subsection{*Finally! Polar Form for Complex Numbers*}
 | |
| 699 | ||
| 14374 | 700 | lemma complex_split_polar: | 
| 14377 | 701 | "\<exists>r a. z = complex_of_real r * (Complex (cos a) (sin a))" | 
| 702 | apply (induct z) | |
| 703 | apply (auto simp add: polar_Ex complex_of_real_mult_Complex) | |
| 14323 | 704 | done | 
| 705 | ||
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 706 | lemma rcis_Ex: "\<exists>r a. z = rcis r a" | 
| 14377 | 707 | apply (induct z) | 
| 708 | apply (simp add: rcis_def cis_def polar_Ex complex_of_real_mult_Complex) | |
| 14323 | 709 | done | 
| 710 | ||
| 14374 | 711 | lemma Re_rcis [simp]: "Re(rcis r a) = r * cos a" | 
| 14373 | 712 | by (simp add: rcis_def cis_def) | 
| 14323 | 713 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 714 | lemma Im_rcis [simp]: "Im(rcis r a) = r * sin a" | 
| 14373 | 715 | by (simp add: rcis_def cis_def) | 
| 14323 | 716 | |
| 14377 | 717 | lemma sin_cos_squared_add2_mult: "(r * cos a)\<twosuperior> + (r * sin a)\<twosuperior> = r\<twosuperior>" | 
| 718 | proof - | |
| 719 | have "(r * cos a)\<twosuperior> + (r * sin a)\<twosuperior> = r\<twosuperior> * ((cos a)\<twosuperior> + (sin a)\<twosuperior>)" | |
| 720 | by (simp only: power_mult_distrib right_distrib) | |
| 721 | thus ?thesis by simp | |
| 722 | qed | |
| 14323 | 723 | |
| 14374 | 724 | lemma complex_mod_rcis [simp]: "cmod(rcis r a) = abs r" | 
| 14377 | 725 | by (simp add: rcis_def cis_def sin_cos_squared_add2_mult) | 
| 14323 | 726 | |
| 727 | lemma complex_mod_sqrt_Re_mult_cnj: "cmod z = sqrt (Re (z * cnj z))" | |
| 14373 | 728 | apply (simp add: cmod_def) | 
| 14323 | 729 | apply (rule real_sqrt_eq_iff [THEN iffD2]) | 
| 730 | apply (auto simp add: complex_mult_cnj) | |
| 731 | done | |
| 732 | ||
| 14374 | 733 | lemma complex_Re_cnj [simp]: "Re(cnj z) = Re z" | 
| 14373 | 734 | by (induct z, simp add: complex_cnj) | 
| 14323 | 735 | |
| 14374 | 736 | lemma complex_Im_cnj [simp]: "Im(cnj z) = - Im z" | 
| 737 | by (induct z, simp add: complex_cnj) | |
| 738 | ||
| 739 | lemma complex_In_mult_cnj_zero [simp]: "Im (z * cnj z) = 0" | |
| 14373 | 740 | by (induct z, simp add: complex_cnj complex_mult) | 
| 14323 | 741 | |
| 742 | ||
| 743 | (*---------------------------------------------------------------------------*) | |
| 744 | (* (r1 * cis a) * (r2 * cis b) = r1 * r2 * cis (a + b) *) | |
| 745 | (*---------------------------------------------------------------------------*) | |
| 746 | ||
| 747 | lemma cis_rcis_eq: "cis a = rcis 1 a" | |
| 14373 | 748 | by (simp add: rcis_def) | 
| 14323 | 749 | |
| 14374 | 750 | lemma rcis_mult: "rcis r1 a * rcis r2 b = rcis (r1*r2) (a + b)" | 
| 14377 | 751 | by (simp add: rcis_def cis_def complex_of_real_mult_Complex cos_add sin_add right_distrib right_diff_distrib) | 
| 14323 | 752 | |
| 753 | lemma cis_mult: "cis a * cis b = cis (a + b)" | |
| 14373 | 754 | by (simp add: cis_rcis_eq rcis_mult) | 
| 14323 | 755 | |
| 14374 | 756 | lemma cis_zero [simp]: "cis 0 = 1" | 
| 14377 | 757 | by (simp add: cis_def complex_one_def) | 
| 14323 | 758 | |
| 14374 | 759 | lemma rcis_zero_mod [simp]: "rcis 0 a = 0" | 
| 14373 | 760 | by (simp add: rcis_def) | 
| 14323 | 761 | |
| 14374 | 762 | lemma rcis_zero_arg [simp]: "rcis r 0 = complex_of_real r" | 
| 14373 | 763 | by (simp add: rcis_def) | 
| 14323 | 764 | |
| 765 | lemma complex_of_real_minus_one: | |
| 766 | "complex_of_real (-(1::real)) = -(1::complex)" | |
| 14377 | 767 | by (simp add: complex_of_real_def complex_one_def complex_minus) | 
| 14323 | 768 | |
| 14374 | 769 | lemma complex_i_mult_minus [simp]: "ii * (ii * x) = - x" | 
| 14373 | 770 | by (simp add: complex_mult_assoc [symmetric]) | 
| 14323 | 771 | |
| 772 | ||
| 773 | lemma cis_real_of_nat_Suc_mult: | |
| 774 | "cis (real (Suc n) * a) = cis a * cis (real n * a)" | |
| 14377 | 775 | by (simp add: cis_def real_of_nat_Suc left_distrib cos_add sin_add right_distrib) | 
| 14323 | 776 | |
| 777 | lemma DeMoivre: "(cis a) ^ n = cis (real n * a)" | |
| 778 | apply (induct_tac "n") | |
| 779 | apply (auto simp add: cis_real_of_nat_Suc_mult) | |
| 780 | done | |
| 781 | ||
| 14374 | 782 | lemma DeMoivre2: "(rcis r a) ^ n = rcis (r ^ n) (real n * a)" | 
| 783 | by (simp add: rcis_def power_mult_distrib DeMoivre complex_of_real_pow) | |
| 14323 | 784 | |
| 14374 | 785 | lemma cis_inverse [simp]: "inverse(cis a) = cis (-a)" | 
| 786 | by (simp add: cis_def complex_inverse_complex_split complex_of_real_minus | |
| 787 | complex_diff_def) | |
| 14323 | 788 | |
| 789 | lemma rcis_inverse: "inverse(rcis r a) = rcis (1/r) (-a)" | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 790 | by (simp add: divide_inverse rcis_def complex_of_real_inverse) | 
| 14323 | 791 | |
| 792 | lemma cis_divide: "cis a / cis b = cis (a - b)" | |
| 14373 | 793 | by (simp add: complex_divide_def cis_mult real_diff_def) | 
| 14323 | 794 | |
| 14354 
988aa4648597
types complex and hcomplex are now instances of class ringpower:
 paulson parents: 
14353diff
changeset | 795 | lemma rcis_divide: "rcis r1 a / rcis r2 b = rcis (r1/r2) (a - b)" | 
| 14373 | 796 | apply (simp add: complex_divide_def) | 
| 797 | apply (case_tac "r2=0", simp) | |
| 798 | apply (simp add: rcis_inverse rcis_mult real_diff_def) | |
| 14323 | 799 | done | 
| 800 | ||
| 14374 | 801 | lemma Re_cis [simp]: "Re(cis a) = cos a" | 
| 14373 | 802 | by (simp add: cis_def) | 
| 14323 | 803 | |
| 14374 | 804 | lemma Im_cis [simp]: "Im(cis a) = sin a" | 
| 14373 | 805 | by (simp add: cis_def) | 
| 14323 | 806 | |
| 807 | lemma cos_n_Re_cis_pow_n: "cos (real n * a) = Re(cis a ^ n)" | |
| 14334 | 808 | by (auto simp add: DeMoivre) | 
| 14323 | 809 | |
| 810 | lemma sin_n_Im_cis_pow_n: "sin (real n * a) = Im(cis a ^ n)" | |
| 14334 | 811 | by (auto simp add: DeMoivre) | 
| 14323 | 812 | |
| 813 | lemma expi_add: "expi(a + b) = expi(a) * expi(b)" | |
| 14374 | 814 | by (simp add: expi_def complex_Re_add exp_add complex_Im_add | 
| 815 | cis_mult [symmetric] complex_of_real_mult mult_ac) | |
| 14323 | 816 | |
| 14374 | 817 | lemma expi_zero [simp]: "expi (0::complex) = 1" | 
| 14373 | 818 | by (simp add: expi_def) | 
| 14323 | 819 | |
| 14374 | 820 | lemma complex_expi_Ex: "\<exists>a r. z = complex_of_real r * expi a" | 
| 14373 | 821 | apply (insert rcis_Ex [of z]) | 
| 14323 | 822 | apply (auto simp add: expi_def rcis_def complex_mult_assoc [symmetric] complex_of_real_mult) | 
| 14334 | 823 | apply (rule_tac x = "ii * complex_of_real a" in exI, auto) | 
| 14323 | 824 | done | 
| 825 | ||
| 826 | ||
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 827 | subsection{*Numerals and Arithmetic*}
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 828 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 829 | instance complex :: number .. | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 830 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 831 | primrec (*the type constraint is essential!*) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 832 | number_of_Pls: "number_of bin.Pls = 0" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 833 | number_of_Min: "number_of bin.Min = - (1::complex)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 834 | number_of_BIT: "number_of(w BIT x) = (if x then 1 else 0) + | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 835 | (number_of w) + (number_of w)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 836 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 837 | declare number_of_Pls [simp del] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 838 | number_of_Min [simp del] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 839 | number_of_BIT [simp del] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 840 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 841 | instance complex :: number_ring | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 842 | proof | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 843 | show "Numeral0 = (0::complex)" by (rule number_of_Pls) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 844 | show "-1 = - (1::complex)" by (rule number_of_Min) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 845 | fix w :: bin and x :: bool | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 846 | show "(number_of (w BIT x) :: complex) = | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 847 | (if x then 1 else 0) + number_of w + number_of w" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 848 | by (rule number_of_BIT) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 849 | qed | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 850 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 851 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 852 | text{*Collapse applications of @{term complex_of_real} to @{term number_of}*}
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 853 | lemma complex_number_of [simp]: "complex_of_real (number_of w) = number_of w" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 854 | apply (induct w) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 855 | apply (simp_all only: number_of complex_of_real_add [symmetric] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 856 | complex_of_real_minus, simp_all) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 857 | done | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 858 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 859 | text{*This theorem is necessary because theorems such as
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 860 |    @{text iszero_number_of_0} only hold for ordered rings. They cannot
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 861 | be generalized to fields in general because they fail for finite fields. | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 862 | They work for type complex because the reals can be embedded in them.*} | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 863 | lemma iszero_complex_number_of [simp]: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 864 | "iszero (number_of w :: complex) = iszero (number_of w :: real)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 865 | by (simp only: complex_of_real_zero_iff complex_number_of [symmetric] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 866 | iszero_def) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 867 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 868 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 869 | (*These allow simplification of expressions involving mixed numbers. | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 870 | Convert??? | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 871 | Goalw [complex_number_of_def] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 872 | "((number_of xa :: complex) + ii * number_of ya = | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 873 | number_of xb) = | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 874 | (((number_of xa :: complex) = number_of xb) & | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 875 | ((number_of ya :: complex) = 0))" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 876 | by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff2, | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 877 | complex_of_real_zero_iff])); | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 878 | qed "complex_number_of_eq_cancel_iff2"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 879 | Addsimps [complex_number_of_eq_cancel_iff2]; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 880 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 881 | Goalw [complex_number_of_def] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 882 | "((number_of xa :: complex) + number_of ya * ii = \ | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 883 | \ number_of xb) = \ | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 884 | \ (((number_of xa :: complex) = number_of xb) & \ | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 885 | \ ((number_of ya :: complex) = 0))"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 886 | by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff2a, | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 887 | complex_of_real_zero_iff])); | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 888 | qed "complex_number_of_eq_cancel_iff2a"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 889 | Addsimps [complex_number_of_eq_cancel_iff2a]; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 890 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 891 | Goalw [complex_number_of_def] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 892 | "((number_of xa :: complex) + ii * number_of ya = \ | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 893 | \ ii * number_of yb) = \ | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 894 | \ (((number_of xa :: complex) = 0) & \ | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 895 | \ ((number_of ya :: complex) = number_of yb))"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 896 | by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff3, | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 897 | complex_of_real_zero_iff])); | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 898 | qed "complex_number_of_eq_cancel_iff3"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 899 | Addsimps [complex_number_of_eq_cancel_iff3]; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 900 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 901 | Goalw [complex_number_of_def] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 902 | "((number_of xa :: complex) + number_of ya * ii= \ | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 903 | \ ii * number_of yb) = \ | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 904 | \ (((number_of xa :: complex) = 0) & \ | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 905 | \ ((number_of ya :: complex) = number_of yb))"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 906 | by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff3a, | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 907 | complex_of_real_zero_iff])); | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 908 | qed "complex_number_of_eq_cancel_iff3a"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 909 | Addsimps [complex_number_of_eq_cancel_iff3a]; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 910 | *) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 911 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 912 | lemma complex_number_of_cnj [simp]: "cnj(number_of v :: complex) = number_of v" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 913 | apply (subst complex_number_of [symmetric]) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 914 | apply (rule complex_cnj_complex_of_real) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 915 | done | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 916 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 917 | lemma complex_number_of_cmod: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 918 | "cmod(number_of v :: complex) = abs (number_of v :: real)" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 919 | by (simp only: complex_number_of [symmetric] complex_mod_complex_of_real) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 920 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 921 | lemma complex_number_of_Re [simp]: "Re(number_of v :: complex) = number_of v" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 922 | by (simp only: complex_number_of [symmetric] Re_complex_of_real) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 923 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 924 | lemma complex_number_of_Im [simp]: "Im(number_of v :: complex) = 0" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 925 | by (simp only: complex_number_of [symmetric] Im_complex_of_real) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 926 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 927 | lemma expi_two_pi_i [simp]: "expi((2::complex) * complex_of_real pi * ii) = 1" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 928 | by (simp add: expi_def complex_Re_mult_eq complex_Im_mult_eq cis_def) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 929 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 930 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 931 | (*examples: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 932 | print_depth 22 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 933 | set timing; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 934 | set trace_simp; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 935 | fun test s = (Goal s, by (Simp_tac 1)); | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 936 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 937 | test "23 * ii + 45 * ii= (x::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 938 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 939 | test "5 * ii + 12 - 45 * ii= (x::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 940 | test "5 * ii + 40 - 12 * ii + 9 = (x::complex) + 89 * ii"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 941 | test "5 * ii + 40 - 12 * ii + 9 - 78 = (x::complex) + 89 * ii"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 942 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 943 | test "l + 10 * ii + 90 + 3*l + 9 + 45 * ii= (x::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 944 | test "87 + 10 * ii + 90 + 3*7 + 9 + 45 * ii= (x::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 945 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 946 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 947 | fun test s = (Goal s; by (Asm_simp_tac 1)); | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 948 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 949 | test "x*k = k*(y::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 950 | test "k = k*(y::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 951 | test "a*(b*c) = (b::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 952 | test "a*(b*c) = d*(b::complex)*(x*a)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 953 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 954 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 955 | test "(x*k) / (k*(y::complex)) = (uu::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 956 | test "(k) / (k*(y::complex)) = (uu::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 957 | test "(a*(b*c)) / ((b::complex)) = (uu::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 958 | test "(a*(b*c)) / (d*(b::complex)*(x*a)) = (uu::complex)"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 959 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 960 | (*FIXME: what do we do about this?*) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 961 | test "a*(b*c)/(y*z) = d*(b::complex)*(x*a)/z"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 962 | *) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14377diff
changeset | 963 | |
| 14323 | 964 | |
| 965 | ML | |
| 966 | {*
 | |
| 967 | val complex_zero_def = thm"complex_zero_def"; | |
| 968 | val complex_one_def = thm"complex_one_def"; | |
| 969 | val complex_minus_def = thm"complex_minus_def"; | |
| 970 | val complex_diff_def = thm"complex_diff_def"; | |
| 971 | val complex_divide_def = thm"complex_divide_def"; | |
| 972 | val complex_mult_def = thm"complex_mult_def"; | |
| 973 | val complex_add_def = thm"complex_add_def"; | |
| 974 | val complex_of_real_def = thm"complex_of_real_def"; | |
| 975 | val i_def = thm"i_def"; | |
| 976 | val expi_def = thm"expi_def"; | |
| 977 | val cis_def = thm"cis_def"; | |
| 978 | val rcis_def = thm"rcis_def"; | |
| 979 | val cmod_def = thm"cmod_def"; | |
| 980 | val cnj_def = thm"cnj_def"; | |
| 981 | val sgn_def = thm"sgn_def"; | |
| 982 | val arg_def = thm"arg_def"; | |
| 983 | val complexpow_0 = thm"complexpow_0"; | |
| 984 | val complexpow_Suc = thm"complexpow_Suc"; | |
| 985 | ||
| 986 | val Re = thm"Re"; | |
| 987 | val Im = thm"Im"; | |
| 988 | val complex_Re_Im_cancel_iff = thm"complex_Re_Im_cancel_iff"; | |
| 989 | val complex_Re_zero = thm"complex_Re_zero"; | |
| 990 | val complex_Im_zero = thm"complex_Im_zero"; | |
| 991 | val complex_Re_one = thm"complex_Re_one"; | |
| 992 | val complex_Im_one = thm"complex_Im_one"; | |
| 993 | val complex_Re_i = thm"complex_Re_i"; | |
| 994 | val complex_Im_i = thm"complex_Im_i"; | |
| 995 | val Re_complex_of_real = thm"Re_complex_of_real"; | |
| 996 | val Im_complex_of_real = thm"Im_complex_of_real"; | |
| 997 | val complex_minus = thm"complex_minus"; | |
| 998 | val complex_Re_minus = thm"complex_Re_minus"; | |
| 999 | val complex_Im_minus = thm"complex_Im_minus"; | |
| 1000 | val complex_add = thm"complex_add"; | |
| 1001 | val complex_Re_add = thm"complex_Re_add"; | |
| 1002 | val complex_Im_add = thm"complex_Im_add"; | |
| 1003 | val complex_add_commute = thm"complex_add_commute"; | |
| 1004 | val complex_add_assoc = thm"complex_add_assoc"; | |
| 1005 | val complex_add_zero_left = thm"complex_add_zero_left"; | |
| 1006 | val complex_add_zero_right = thm"complex_add_zero_right"; | |
| 1007 | val complex_diff = thm"complex_diff"; | |
| 1008 | val complex_mult = thm"complex_mult"; | |
| 1009 | val complex_mult_one_left = thm"complex_mult_one_left"; | |
| 1010 | val complex_mult_one_right = thm"complex_mult_one_right"; | |
| 1011 | val complex_inverse = thm"complex_inverse"; | |
| 1012 | val complex_of_real_one = thm"complex_of_real_one"; | |
| 1013 | val complex_of_real_zero = thm"complex_of_real_zero"; | |
| 1014 | val complex_of_real_eq_iff = thm"complex_of_real_eq_iff"; | |
| 1015 | val complex_of_real_minus = thm"complex_of_real_minus"; | |
| 1016 | val complex_of_real_inverse = thm"complex_of_real_inverse"; | |
| 1017 | val complex_of_real_add = thm"complex_of_real_add"; | |
| 1018 | val complex_of_real_diff = thm"complex_of_real_diff"; | |
| 1019 | val complex_of_real_mult = thm"complex_of_real_mult"; | |
| 1020 | val complex_of_real_divide = thm"complex_of_real_divide"; | |
| 1021 | val complex_of_real_pow = thm"complex_of_real_pow"; | |
| 1022 | val complex_mod = thm"complex_mod"; | |
| 1023 | val complex_mod_zero = thm"complex_mod_zero"; | |
| 1024 | val complex_mod_one = thm"complex_mod_one"; | |
| 1025 | val complex_mod_complex_of_real = thm"complex_mod_complex_of_real"; | |
| 1026 | val complex_of_real_abs = thm"complex_of_real_abs"; | |
| 1027 | val complex_cnj = thm"complex_cnj"; | |
| 1028 | val complex_cnj_cancel_iff = thm"complex_cnj_cancel_iff"; | |
| 1029 | val complex_cnj_cnj = thm"complex_cnj_cnj"; | |
| 1030 | val complex_cnj_complex_of_real = thm"complex_cnj_complex_of_real"; | |
| 1031 | val complex_mod_cnj = thm"complex_mod_cnj"; | |
| 1032 | val complex_cnj_minus = thm"complex_cnj_minus"; | |
| 1033 | val complex_cnj_inverse = thm"complex_cnj_inverse"; | |
| 1034 | val complex_cnj_add = thm"complex_cnj_add"; | |
| 1035 | val complex_cnj_diff = thm"complex_cnj_diff"; | |
| 1036 | val complex_cnj_mult = thm"complex_cnj_mult"; | |
| 1037 | val complex_cnj_divide = thm"complex_cnj_divide"; | |
| 1038 | val complex_cnj_one = thm"complex_cnj_one"; | |
| 1039 | val complex_cnj_pow = thm"complex_cnj_pow"; | |
| 1040 | val complex_add_cnj = thm"complex_add_cnj"; | |
| 1041 | val complex_diff_cnj = thm"complex_diff_cnj"; | |
| 1042 | val complex_cnj_zero = thm"complex_cnj_zero"; | |
| 1043 | val complex_cnj_zero_iff = thm"complex_cnj_zero_iff"; | |
| 1044 | val complex_mult_cnj = thm"complex_mult_cnj"; | |
| 1045 | val complex_mod_eq_zero_cancel = thm"complex_mod_eq_zero_cancel"; | |
| 1046 | val complex_mod_complex_of_real_of_nat = thm"complex_mod_complex_of_real_of_nat"; | |
| 1047 | val complex_mod_minus = thm"complex_mod_minus"; | |
| 1048 | val complex_mod_mult_cnj = thm"complex_mod_mult_cnj"; | |
| 1049 | val complex_mod_squared = thm"complex_mod_squared"; | |
| 1050 | val complex_mod_ge_zero = thm"complex_mod_ge_zero"; | |
| 1051 | val abs_cmod_cancel = thm"abs_cmod_cancel"; | |
| 1052 | val complex_mod_mult = thm"complex_mod_mult"; | |
| 1053 | val complex_mod_add_squared_eq = thm"complex_mod_add_squared_eq"; | |
| 1054 | val complex_Re_mult_cnj_le_cmod = thm"complex_Re_mult_cnj_le_cmod"; | |
| 1055 | val complex_Re_mult_cnj_le_cmod2 = thm"complex_Re_mult_cnj_le_cmod2"; | |
| 1056 | val real_sum_squared_expand = thm"real_sum_squared_expand"; | |
| 1057 | val complex_mod_triangle_squared = thm"complex_mod_triangle_squared"; | |
| 1058 | val complex_mod_minus_le_complex_mod = thm"complex_mod_minus_le_complex_mod"; | |
| 1059 | val complex_mod_triangle_ineq = thm"complex_mod_triangle_ineq"; | |
| 1060 | val complex_mod_triangle_ineq2 = thm"complex_mod_triangle_ineq2"; | |
| 1061 | val complex_mod_diff_commute = thm"complex_mod_diff_commute"; | |
| 1062 | val complex_mod_add_less = thm"complex_mod_add_less"; | |
| 1063 | val complex_mod_mult_less = thm"complex_mod_mult_less"; | |
| 1064 | val complex_mod_diff_ineq = thm"complex_mod_diff_ineq"; | |
| 1065 | val complex_Re_le_cmod = thm"complex_Re_le_cmod"; | |
| 1066 | val complex_mod_gt_zero = thm"complex_mod_gt_zero"; | |
| 1067 | val complex_mod_complexpow = thm"complex_mod_complexpow"; | |
| 1068 | val complexpow_minus = thm"complexpow_minus"; | |
| 1069 | val complex_mod_inverse = thm"complex_mod_inverse"; | |
| 1070 | val complex_mod_divide = thm"complex_mod_divide"; | |
| 1071 | val complexpow_i_squared = thm"complexpow_i_squared"; | |
| 1072 | val complex_i_not_zero = thm"complex_i_not_zero"; | |
| 1073 | val sgn_zero = thm"sgn_zero"; | |
| 1074 | val sgn_one = thm"sgn_one"; | |
| 1075 | val sgn_minus = thm"sgn_minus"; | |
| 1076 | val sgn_eq = thm"sgn_eq"; | |
| 1077 | val i_mult_eq = thm"i_mult_eq"; | |
| 1078 | val i_mult_eq2 = thm"i_mult_eq2"; | |
| 1079 | val Re_sgn = thm"Re_sgn"; | |
| 1080 | val Im_sgn = thm"Im_sgn"; | |
| 1081 | val complex_inverse_complex_split = thm"complex_inverse_complex_split"; | |
| 1082 | val cos_arg_i_mult_zero = thm"cos_arg_i_mult_zero"; | |
| 1083 | val complex_of_real_zero_iff = thm"complex_of_real_zero_iff"; | |
| 1084 | val rcis_Ex = thm"rcis_Ex"; | |
| 1085 | val Re_rcis = thm"Re_rcis"; | |
| 1086 | val Im_rcis = thm"Im_rcis"; | |
| 1087 | val complex_mod_rcis = thm"complex_mod_rcis"; | |
| 1088 | val complex_mod_sqrt_Re_mult_cnj = thm"complex_mod_sqrt_Re_mult_cnj"; | |
| 1089 | val complex_Re_cnj = thm"complex_Re_cnj"; | |
| 1090 | val complex_Im_cnj = thm"complex_Im_cnj"; | |
| 1091 | val complex_In_mult_cnj_zero = thm"complex_In_mult_cnj_zero"; | |
| 1092 | val complex_Re_mult = thm"complex_Re_mult"; | |
| 1093 | val complex_Re_mult_complex_of_real = thm"complex_Re_mult_complex_of_real"; | |
| 1094 | val complex_Im_mult_complex_of_real = thm"complex_Im_mult_complex_of_real"; | |
| 1095 | val complex_Re_mult_complex_of_real2 = thm"complex_Re_mult_complex_of_real2"; | |
| 1096 | val complex_Im_mult_complex_of_real2 = thm"complex_Im_mult_complex_of_real2"; | |
| 1097 | val cis_rcis_eq = thm"cis_rcis_eq"; | |
| 1098 | val rcis_mult = thm"rcis_mult"; | |
| 1099 | val cis_mult = thm"cis_mult"; | |
| 1100 | val cis_zero = thm"cis_zero"; | |
| 1101 | val rcis_zero_mod = thm"rcis_zero_mod"; | |
| 1102 | val rcis_zero_arg = thm"rcis_zero_arg"; | |
| 1103 | val complex_of_real_minus_one = thm"complex_of_real_minus_one"; | |
| 1104 | val complex_i_mult_minus = thm"complex_i_mult_minus"; | |
| 1105 | val cis_real_of_nat_Suc_mult = thm"cis_real_of_nat_Suc_mult"; | |
| 1106 | val DeMoivre = thm"DeMoivre"; | |
| 1107 | val DeMoivre2 = thm"DeMoivre2"; | |
| 1108 | val cis_inverse = thm"cis_inverse"; | |
| 1109 | val rcis_inverse = thm"rcis_inverse"; | |
| 1110 | val cis_divide = thm"cis_divide"; | |
| 1111 | val rcis_divide = thm"rcis_divide"; | |
| 1112 | val Re_cis = thm"Re_cis"; | |
| 1113 | val Im_cis = thm"Im_cis"; | |
| 1114 | val cos_n_Re_cis_pow_n = thm"cos_n_Re_cis_pow_n"; | |
| 1115 | val sin_n_Im_cis_pow_n = thm"sin_n_Im_cis_pow_n"; | |
| 1116 | val expi_add = thm"expi_add"; | |
| 1117 | val expi_zero = thm"expi_zero"; | |
| 1118 | val complex_Re_mult_eq = thm"complex_Re_mult_eq"; | |
| 1119 | val complex_Im_mult_eq = thm"complex_Im_mult_eq"; | |
| 1120 | val complex_expi_Ex = thm"complex_expi_Ex"; | |
| 1121 | *} | |
| 1122 | ||
| 13957 | 1123 | end | 
| 1124 | ||
| 1125 |