# HG changeset patch # User huffman # Date 1159319128 -7200 # Node ID 72e20198f83450d61abadab9779d59210690276e # Parent a1a8ba09e0ead223bdffd402322f006463da7dbf instance complex :: real_normed_field; cleaned up diff -r a1a8ba09e0ea -r 72e20198f834 src/HOL/Complex/Complex.thy --- a/src/HOL/Complex/Complex.thy Wed Sep 27 03:04:35 2006 +0200 +++ b/src/HOL/Complex/Complex.thy Wed Sep 27 03:05:28 2006 +0200 @@ -159,7 +159,7 @@ lemma complex_mult_inv_left: "z \ (0::complex) ==> inverse(z) * z = 1" apply (induct z) apply (rename_tac x y) -apply (auto simp add: times_divide_eq complex_mult complex_inverse +apply (auto simp add: complex_one_def complex_zero_def add_divide_distrib [symmetric] power2_eq_square mult_ac) apply (simp_all add: real_sum_squares_not_zero real_sum_squares_not_zero2) @@ -213,7 +213,7 @@ defs (overloaded) complex_scaleR_def: "r *# x == Complex r 0 * x" -instance complex :: real_algebra_1 +instance complex :: real_field proof fix a b :: real fix x y :: complex @@ -269,41 +269,34 @@ lemma complex_of_real_i [simp]: "complex_of_real r * ii = Complex 0 r" by (simp add: i_def complex_of_real_def) -(* TODO: generalize and move to Real/RealVector.thy *) -lemma complex_of_real_inverse [simp]: +lemma complex_of_real_inverse: "complex_of_real(inverse x) = inverse(complex_of_real x)" -apply (case_tac "x=0", simp) -apply (simp add: complex_of_real_def divide_inverse power2_eq_square) -done +by (rule of_real_inverse) -(* TODO: generalize and move to Real/RealVector.thy *) -lemma complex_of_real_divide [simp]: +lemma complex_of_real_divide: "complex_of_real(x/y) = complex_of_real x / complex_of_real y" -apply (simp add: complex_divide_def) -apply (case_tac "y=0", simp) -apply (simp add: divide_inverse) -done +by (rule of_real_divide) subsection{*The Functions @{term Re} and @{term Im}*} lemma complex_Re_mult_eq: "Re (w * z) = Re w * Re z - Im w * Im z" -by (induct z, induct w, simp add: complex_mult) +by (induct z, induct w, simp) lemma complex_Im_mult_eq: "Im (w * z) = Re w * Im z + Im w * Re z" -by (induct z, induct w, simp add: complex_mult) +by (induct z, induct w, simp) lemma Re_i_times [simp]: "Re(ii * z) = - Im z" -by (simp add: complex_Re_mult_eq) +by (simp add: complex_Re_mult_eq) lemma Re_times_i [simp]: "Re(z * ii) = - Im z" -by (simp add: complex_Re_mult_eq) +by (simp add: complex_Re_mult_eq) lemma Im_i_times [simp]: "Im(ii * z) = Re z" -by (simp add: complex_Im_mult_eq) +by (simp add: complex_Im_mult_eq) lemma Im_times_i [simp]: "Im(z * ii) = Re z" -by (simp add: complex_Im_mult_eq) +by (simp add: complex_Im_mult_eq) lemma complex_Re_mult: "[| Im w = 0; Im z = 0 |] ==> Re(w * z) = Re(w) * Re(z)" by (simp add: complex_Re_mult_eq) @@ -345,19 +338,19 @@ by (simp add: complex_of_real_def complex_cnj) lemma complex_cnj_minus: "cnj (-z) = - cnj z" -by (simp add: cnj_def complex_minus complex_Re_minus complex_Im_minus) +by (simp add: cnj_def) lemma complex_cnj_inverse: "cnj(inverse z) = inverse(cnj z)" -by (induct z, simp add: complex_cnj complex_inverse power2_eq_square) +by (induct z, simp add: complex_cnj power2_eq_square) lemma complex_cnj_add: "cnj(w + z) = cnj(w) + cnj(z)" -by (induct w, induct z, simp add: complex_cnj complex_add) +by (induct w, induct z, simp add: complex_cnj) lemma complex_cnj_diff: "cnj(w - z) = cnj(w) - cnj(z)" by (simp add: diff_minus complex_cnj_add complex_cnj_minus) lemma complex_cnj_mult: "cnj(w * z) = cnj(w) * cnj(z)" -by (induct w, induct z, simp add: complex_cnj complex_mult) +by (induct w, induct z, simp add: complex_cnj) lemma complex_cnj_divide: "cnj(w / z) = (cnj w)/(cnj z)" by (simp add: complex_divide_def complex_cnj_mult complex_cnj_inverse) @@ -366,7 +359,7 @@ by (simp add: cnj_def complex_one_def) lemma complex_add_cnj: "z + cnj z = complex_of_real (2 * Re(z))" -by (induct z, simp add: complex_add complex_cnj complex_of_real_def) +by (induct z, simp add: complex_cnj complex_of_real_def) lemma complex_diff_cnj: "z - cnj z = complex_of_real (2 * Im(z)) * ii" apply (induct z) @@ -381,8 +374,7 @@ by (induct z, simp add: complex_zero_def complex_cnj) lemma complex_mult_cnj: "z * cnj z = complex_of_real (Re(z) ^ 2 + Im(z) ^ 2)" -by (induct z, - simp add: complex_cnj complex_mult complex_of_real_def power2_eq_square) +by (induct z, simp add: complex_cnj complex_of_real_def power2_eq_square) subsection{*Modulus*} @@ -408,7 +400,7 @@ by (simp add: cmod_def power2_eq_square) lemma complex_mod_complex_of_real [simp]: "cmod(complex_of_real x) = abs x" -by (simp add: complex_of_real_def power2_eq_square complex_mod) +by (simp add: complex_of_real_def power2_eq_square) lemma complex_of_real_abs: "complex_of_real (abs x) = complex_of_real(cmod(complex_of_real x))" @@ -426,10 +418,10 @@ by simp lemma complex_mod_minus [simp]: "cmod (-x) = cmod(x)" -by (induct x, simp add: complex_mod complex_minus power2_eq_square) +by (induct x, simp add: power2_eq_square) lemma complex_mod_cnj [simp]: "cmod (cnj z) = cmod z" -by (induct z, simp add: complex_cnj complex_mod power2_eq_square) +by (induct z, simp add: complex_cnj power2_eq_square) lemma complex_mod_mult_cnj: "cmod(z * cnj(z)) = cmod(z) ^ 2" apply (induct z, simp add: complex_mod complex_cnj complex_mult) @@ -464,13 +456,13 @@ lemma complex_mod_add_squared_eq: "cmod(x + y) ^ 2 = cmod(x) ^ 2 + cmod(y) ^ 2 + 2 * Re(x * cnj y)" apply (induct x, induct y) -apply (auto simp add: complex_add complex_mod_squared complex_mult complex_cnj real_diff_def simp del: realpow_Suc) +apply (auto simp add: complex_mod_squared complex_cnj real_diff_def simp del: realpow_Suc) apply (auto simp add: right_distrib left_distrib power2_eq_square mult_ac add_ac) done lemma complex_Re_mult_cnj_le_cmod [simp]: "Re(x * cnj y) \ cmod(x * cnj y)" apply (induct x, induct y) -apply (auto simp add: complex_mod complex_mult complex_cnj real_diff_def simp del: realpow_Suc) +apply (auto simp add: complex_mod complex_cnj diff_def simp del: realpow_Suc) done lemma complex_Re_mult_cnj_le_cmod2 [simp]: "Re(x * cnj y) \ cmod(x * y)" @@ -493,7 +485,7 @@ simp add: add_increasing power2_eq_square [symmetric]) done -instance complex :: real_normed_div_algebra +instance complex :: real_normed_field proof fix r :: real fix x y :: complex @@ -524,19 +516,15 @@ by (auto intro: real_mult_less_mono' simp add: complex_mod_mult) lemma complex_mod_diff_ineq [simp]: "cmod(a) - cmod(b) \ cmod(a + b)" -apply (rule linorder_cases [of "cmod(a)" "cmod (b)"]) -apply auto -apply (rule order_trans [of _ 0], rule order_less_imp_le) -apply (simp add: compare_rls, simp) -apply (simp add: compare_rls) -apply (rule complex_mod_minus [THEN subst]) -apply (rule order_trans) -apply (rule_tac [2] complex_mod_triangle_ineq) -apply (auto simp add: add_ac) -done +proof - + have "cmod a - cmod b = cmod a - cmod (- b)" by simp + also have "\ \ cmod (a - - b)" by (rule norm_triangle_ineq2) + also have "\ = cmod (a + b)" by simp + finally show ?thesis . +qed lemma complex_Re_le_cmod [simp]: "Re z \ cmod z" -by (induct z, simp add: complex_mod del: realpow_Suc) +by (induct z, simp) lemma complex_mod_gt_zero: "z \ 0 ==> 0 < cmod z" by (rule zero_less_norm_iff [THEN iffD2]) @@ -545,7 +533,7 @@ by (rule norm_inverse) lemma complex_mod_divide: "cmod(x/y) = cmod(x)/(cmod y)" -by (simp add: divide_inverse norm_mult norm_inverse) +by (rule norm_divide) subsection{*Exponentiation*} @@ -565,9 +553,7 @@ lemma complex_of_real_pow: "complex_of_real (x ^ n) = (complex_of_real x) ^ n" -apply (induct_tac "n") -apply (auto simp add: of_real_mult [symmetric]) -done +by (rule of_real_power) lemma complex_cnj_pow: "cnj(z ^ n) = cnj(z) ^ n" apply (induct_tac "n") @@ -575,12 +561,10 @@ done lemma complex_mod_complexpow: "cmod(x ^ n) = cmod(x) ^ n" -apply (induct_tac "n") -apply (auto simp add: complex_mod_mult) -done +by (rule norm_power) lemma complexpow_i_squared [simp]: "ii ^ 2 = -(1::complex)" -by (simp add: i_def complex_mult complex_one_def complex_minus numeral_2_eq_2) +by (simp add: i_def complex_one_def numeral_2_eq_2) lemma complex_i_not_zero [simp]: "ii \ 0" by (simp add: i_def complex_zero_def) @@ -610,17 +594,13 @@ by (simp add: sgn_def) lemma i_mult_eq: "ii * ii = complex_of_real (-1)" -by (simp add: i_def complex_of_real_def complex_mult complex_add) +by (simp add: i_def complex_of_real_def) lemma i_mult_eq2 [simp]: "ii * ii = -(1::complex)" -by (simp add: i_def complex_one_def complex_mult complex_minus) +by (simp add: i_def complex_one_def) lemma complex_eq_cancel_iff2 [simp]: "(Complex x y = complex_of_real xa) = (x = xa & y = 0)" -by (simp add: complex_of_real_def) - -lemma complex_eq_cancel_iff2a [simp]: - "(Complex x y = complex_of_real xa) = (x = xa & y = 0)" by (simp add: complex_of_real_def) lemma Complex_eq_0 [simp]: "(Complex x y = 0) = (x = 0 & y = 0)" @@ -657,16 +637,15 @@ "inverse(complex_of_real x + ii * complex_of_real y) = complex_of_real(x/(x ^ 2 + y ^ 2)) - ii * complex_of_real(y/(x ^ 2 + y ^ 2))" -by (simp add: complex_of_real_def i_def complex_mult complex_add - diff_minus complex_minus complex_inverse divide_inverse) +by (simp add: complex_of_real_def i_def diff_minus divide_inverse) (*----------------------------------------------------------------------------*) (* Many of the theorems below need to be moved elsewhere e.g. Transc. Also *) (* many of the theorems are not used - so should they be kept? *) (*----------------------------------------------------------------------------*) -lemma complex_of_real_zero_iff [simp]: "(complex_of_real y = 0) = (y = 0)" -by (auto simp add: complex_zero_def complex_of_real_def) +lemma complex_of_real_zero_iff: "(complex_of_real y = 0) = (y = 0)" +by (rule of_real_eq_0_iff) lemma cos_arg_i_mult_zero_pos: "0 < y ==> cos (arg(Complex 0 y)) = 0" @@ -705,12 +684,12 @@ lemma complex_split_polar: "\r a. z = complex_of_real r * (Complex (cos a) (sin a))" -apply (induct z) +apply (induct z) apply (auto simp add: polar_Ex complex_of_real_mult_Complex) done lemma rcis_Ex: "\r a. z = rcis r a" -apply (induct z) +apply (induct z) apply (simp add: rcis_def cis_def polar_Ex complex_of_real_mult_Complex) done @@ -723,7 +702,7 @@ lemma sin_cos_squared_add2_mult: "(r * cos a)\ + (r * sin a)\ = r\" proof - have "(r * cos a)\ + (r * sin a)\ = r\ * ((cos a)\ + (sin a)\)" - by (simp only: power_mult_distrib right_distrib) + by (simp only: power_mult_distrib right_distrib) thus ?thesis by simp qed @@ -733,7 +712,8 @@ lemma complex_mod_sqrt_Re_mult_cnj: "cmod z = sqrt (Re (z * cnj z))" apply (simp add: cmod_def) apply (rule real_sqrt_eq_iff [THEN iffD2]) -apply (auto simp add: complex_mult_cnj) +apply (auto simp add: complex_mult_cnj + simp del: of_real_add) done lemma complex_Re_cnj [simp]: "Re(cnj z) = Re z" @@ -771,7 +751,7 @@ lemma complex_of_real_minus_one: "complex_of_real (-(1::real)) = -(1::complex)" -by (simp add: complex_of_real_def complex_one_def complex_minus) +by (simp add: complex_of_real_def complex_one_def) lemma complex_i_mult_minus [simp]: "ii * (ii * x) = - x" by (simp add: complex_mult_assoc [symmetric]) @@ -790,8 +770,7 @@ by (simp add: rcis_def power_mult_distrib DeMoivre complex_of_real_pow) lemma cis_inverse [simp]: "inverse(cis a) = cis (-a)" -by (simp add: cis_def complex_inverse_complex_split of_real_minus - diff_minus) +by (simp add: cis_def complex_inverse_complex_split diff_minus) lemma rcis_inverse: "inverse(rcis r a) = rcis (1/r) (-a)" by (simp add: divide_inverse rcis_def complex_of_real_inverse) @@ -818,8 +797,7 @@ by (auto simp add: DeMoivre) lemma expi_add: "expi(a + b) = expi(a) * expi(b)" -by (simp add: expi_def complex_Re_add exp_add complex_Im_add - cis_mult [symmetric] of_real_mult mult_ac) +by (simp add: expi_def exp_add cis_mult [symmetric] mult_ac) lemma expi_zero [simp]: "expi (0::complex) = 1" by (simp add: expi_def) @@ -840,7 +818,7 @@ --{*the type constraint is essential!*} instance complex :: number_ring -by (intro_classes, simp add: complex_number_of_def) +by (intro_classes, simp add: complex_number_of_def) text{*Collapse applications of @{term complex_of_real} to @{term number_of}*} @@ -855,7 +833,7 @@ lemma iszero_complex_number_of [simp]: "iszero (number_of w :: complex) = iszero (number_of w :: real)" by (simp only: complex_of_real_zero_iff complex_number_of [symmetric] - iszero_def) + iszero_def) lemma complex_number_of_cnj [simp]: "cnj(number_of v :: complex) = number_of v" by (simp only: complex_number_of [symmetric] complex_cnj_complex_of_real) @@ -908,5 +886,3 @@ *) end - -