# HG changeset patch # User huffman # Date 1178586203 -7200 # Node ID eb0e0582092aa5e67d9a77d96536af31575431ab # Parent eb364358878138425662c5a59fcf91dd3585e365 cleaned up diff -r eb3643588781 -r eb0e0582092a src/HOL/Hyperreal/NthRoot.thy --- a/src/HOL/Hyperreal/NthRoot.thy Tue May 08 01:40:30 2007 +0200 +++ b/src/HOL/Hyperreal/NthRoot.thy Tue May 08 03:03:23 2007 +0200 @@ -241,20 +241,12 @@ lemma real_root_less_mono: "[| 0 \ x; x < y |] ==> root(Suc n) x < root(Suc n) y" -apply (frule order_le_less_trans, assumption) -apply (frule_tac n1 = n in real_root_pow_pos2 [THEN ssubst]) -apply (rotate_tac 1, assumption) -apply (frule_tac n1 = n in real_root_pow_pos [THEN ssubst]) -apply (rotate_tac 3, assumption) -apply (drule_tac y = "root (Suc n) y ^ Suc n" in order_less_imp_le) -apply (frule_tac n = n in real_root_pos_pos_le) -apply (frule_tac n = n in real_root_pos_pos) -apply (drule_tac x = "root (Suc n) x" and y = "root (Suc n) y" in realpow_increasing) -apply (assumption, assumption) -apply (drule_tac x = "root (Suc n) x" in order_le_imp_less_or_eq) -apply auto -apply (drule_tac f = "%x. x ^ (Suc n)" in arg_cong) -apply (auto simp add: real_root_pow_pos2 simp del: realpow_Suc) +apply (subgoal_tac "0 \ y") +apply (rule_tac n="Suc n" in power_less_imp_less_base) +apply (simp only: real_root_pow_pos2) +apply (erule real_root_pos_pos_le) +apply (erule order_trans) +apply (erule order_less_imp_le) done lemma real_root_le_mono: @@ -318,8 +310,7 @@ text{*needed because 2 is a binary numeral!*} lemma root_2_eq [simp]: "root 2 = root (Suc (Suc 0))" -by (simp del: nat_numeral_0_eq_0 nat_numeral_1_eq_1 - add: nat_numeral_0_eq_0 [symmetric]) +by (simp only: numeral_2_eq_2) lemma real_sqrt_zero [simp]: "sqrt 0 = 0" by (simp add: sqrt_def) @@ -327,40 +318,40 @@ lemma real_sqrt_one [simp]: "sqrt 1 = 1" by (simp add: sqrt_def) +lemma real_sqrt_pow2 [simp]: "0 \ x ==> (sqrt x)\ = x" +unfolding sqrt_def numeral_2_eq_2 +by (rule real_root_pow_pos2) + lemma real_sqrt_pow2_iff [iff]: "((sqrt x)\ = x) = (0 \ x)" -apply (simp add: sqrt_def) -apply (rule iffI) - apply (cut_tac r = "root 2 x" in realpow_two_le) - apply (simp add: numeral_2_eq_2) -apply (subst numeral_2_eq_2) -apply (erule real_root_pow_pos2) +apply (rule iffI) +apply (erule subst) +apply (rule zero_le_power2) +apply (erule real_sqrt_pow2) done -lemma [simp]: "(sqrt(u2\ + v2\))\ = u2\ + v2\" -by (rule realpow_two_le_add_order [THEN real_sqrt_pow2_iff [THEN iffD2]]) - -lemma real_sqrt_pow2 [simp]: "0 \ x ==> (sqrt x)\ = x" -by (simp) +lemma real_sqrt_abs_abs [simp]: "(sqrt \x\)\ = \x\" (* TODO: delete *) +by simp -lemma real_sqrt_abs_abs [simp]: "sqrt\x\ ^ 2 = \x\" -by (rule real_sqrt_pow2_iff [THEN iffD2], arith) +lemma sqrt_eqI: "\r\ = a; 0 \ r\ \ sqrt a = r" +unfolding sqrt_def numeral_2_eq_2 +by (erule subst, erule real_root_pos2) -lemma real_pow_sqrt_eq_sqrt_pow: - "0 \ x ==> (sqrt x)\ = sqrt(x\)" -apply (simp add: sqrt_def) -apply (simp only: numeral_2_eq_2 real_root_pow_pos2 real_root_pos2) +lemma real_sqrt_abs [simp]: "sqrt (x\) = \x\" +apply (rule sqrt_eqI) +apply (rule power2_abs) +apply (rule abs_ge_zero) done -lemma real_pow_sqrt_eq_sqrt_abs_pow2: - "0 \ x ==> (sqrt x)\ = sqrt(\x\ ^ 2)" -by (simp add: real_pow_sqrt_eq_sqrt_pow [symmetric]) +lemma real_pow_sqrt_eq_sqrt_pow: (* TODO: delete *) + "0 \ x ==> (sqrt x)\ = sqrt(x\)" +by simp -lemma real_sqrt_pow_abs: "0 \ x ==> (sqrt x)\ = \x\" -apply (rule real_sqrt_abs_abs [THEN subst]) -apply (rule_tac x1 = x in real_pow_sqrt_eq_sqrt_abs_pow2 [THEN ssubst]) -apply (rule_tac [2] real_pow_sqrt_eq_sqrt_pow [symmetric]) -apply (assumption, arith) -done +lemma real_pow_sqrt_eq_sqrt_abs_pow2: (* TODO: delete *) + "0 \ x ==> (sqrt x)\ = sqrt(\x\ ^ 2)" +by simp + +lemma real_sqrt_pow_abs: "0 \ x ==> (sqrt x)\ = \x\" (* TODO: delete *) +by simp lemma not_real_square_gt_zero [simp]: "(~ (0::real) < x*x) = (x = 0)" apply auto @@ -374,9 +365,6 @@ lemma real_sqrt_ge_zero: "0 \ x ==> 0 \ sqrt(x)" by (auto intro: real_sqrt_gt_zero simp add: order_le_less) -lemma real_sqrt_mult_self_sum_ge_zero [simp]: "0 \ sqrt(x*x + y*y)" -by (rule real_sqrt_ge_zero [OF real_mult_self_sum_ge_zero]) - (*we need to prove something like this: lemma "[|r ^ n = a; 0 0 < r|] ==> root n a = r" @@ -386,44 +374,15 @@ apply (auto simp del: realpow_Suc dest: power_inject_base) *) -lemma sqrt_eqI: "[|r\ = a; 0 \ r|] ==> sqrt a = r" -apply (erule subst) -apply (simp add: sqrt_def numeral_2_eq_2 del: realpow_Suc) -apply (erule real_root_pos2) -done - lemma real_sqrt_mult_distrib: "[| 0 \ x; 0 \ y |] ==> sqrt(x*y) = sqrt(x) * sqrt(y)" -apply (rule sqrt_eqI) -apply (simp add: power_mult_distrib) -apply (simp add: zero_le_mult_iff real_sqrt_ge_zero) -done - -lemma real_sqrt_mult_distrib2: - "[|0\x; 0\y |] ==> sqrt(x*y) = sqrt(x) * sqrt(y)" -by (auto intro: real_sqrt_mult_distrib simp add: order_le_less) - -lemma real_sqrt_sum_squares_ge_zero [simp]: "0 \ sqrt (x\ + y\)" -by (auto intro!: real_sqrt_ge_zero) +unfolding sqrt_def numeral_2_eq_2 +by (rule real_root_mult) -lemma real_sqrt_sum_squares_mult_ge_zero [simp]: - "0 \ sqrt ((x\ + y\)*(xa\ + ya\))" -by (auto intro!: real_sqrt_ge_zero simp add: zero_le_mult_iff) - -lemma real_sqrt_sum_squares_mult_squared_eq [simp]: - "sqrt ((x\ + y\) * (xa\ + ya\)) ^ 2 = (x\ + y\) * (xa\ + ya\)" -by (auto simp add: zero_le_mult_iff simp del: realpow_Suc) - -lemma real_sqrt_abs [simp]: "sqrt(x\) = \x\" -apply (rule abs_realpow_two [THEN subst]) -apply (rule real_sqrt_abs_abs [THEN subst]) -apply (subst real_pow_sqrt_eq_sqrt_pow) -apply (auto simp add: numeral_2_eq_2) -done +lemmas real_sqrt_mult_distrib2 = real_sqrt_mult_distrib lemma real_sqrt_abs2 [simp]: "sqrt(x*x) = \x\" -apply (rule realpow_two [THEN subst]) -apply (subst numeral_2_eq_2 [symmetric]) +apply (subst power2_eq_square [symmetric]) apply (rule real_sqrt_abs) done @@ -436,7 +395,7 @@ done lemma real_inv_sqrt_pow2: "0 < x ==> inverse (sqrt(x)) ^ 2 = inverse x" -by (cut_tac n = 2 and a = "sqrt x" in power_inverse [symmetric], auto) +by (simp add: power_inverse [symmetric]) lemma real_sqrt_eq_zero_cancel: "[| 0 \ x; sqrt(x) = 0|] ==> x = 0" apply (drule real_le_imp_less_or_eq) @@ -446,22 +405,9 @@ lemma real_sqrt_eq_zero_cancel_iff [simp]: "0 \ x ==> ((sqrt x = 0) = (x=0))" by (auto simp add: real_sqrt_eq_zero_cancel) -lemma real_sqrt_sum_squares_ge1 [simp]: "x \ sqrt(x\ + y\)" -apply (subgoal_tac "x \ 0 | 0 \ x", safe) -apply (rule real_le_trans) -apply (auto simp del: realpow_Suc) -apply (rule_tac n = 1 in realpow_increasing) -apply (auto simp add: numeral_2_eq_2 [symmetric] simp del: realpow_Suc) -done - -lemma real_sqrt_sum_squares_ge2 [simp]: "y \ sqrt(z\ + y\)" -apply (simp (no_asm) add: real_add_commute del: realpow_Suc) -done - lemma real_sqrt_ge_one: "1 \ x ==> 1 \ sqrt x" -apply (rule_tac n = 1 in realpow_increasing) -apply (auto simp add: numeral_2_eq_2 [symmetric] real_sqrt_ge_zero simp - del: realpow_Suc) +apply (rule power2_le_imp_le, simp) +apply (simp add: real_sqrt_ge_zero) done lemma sqrt_divide_self_eq: @@ -500,22 +446,45 @@ "[| 0 \ x; 0 \ y |] ==> (sqrt(x) = sqrt(y)) = (x = y)" by (simp add: sqrt_def) -lemma real_sqrt_sos_less_one_iff [simp]: "(sqrt(x\ + y\) < 1) = (x\ + y\ < 1)" -apply (rule real_sqrt_one [THEN subst], safe) -apply (rule_tac [2] real_sqrt_less_mono) -apply (drule real_sqrt_less_iff [THEN [2] rev_iffD1], auto) -done - -lemma real_sqrt_sos_eq_one_iff [simp]: "(sqrt(x\ + y\) = 1) = (x\ + y\ = 1)" -apply (rule real_sqrt_one [THEN subst], safe) -apply (drule real_sqrt_eq_iff [THEN [2] rev_iffD1], auto) -done - lemma real_divide_square_eq [simp]: "(((r::real) * a) / (r * r)) = a / r" apply (simp add: divide_inverse) apply (case_tac "r=0") apply (auto simp add: mult_ac) done +subsection {* Square Root of Sum of Squares *} + +lemma "(sqrt (x\ + y\))\ = x\ + y\" +by (rule realpow_two_le_add_order [THEN real_sqrt_pow2_iff [THEN iffD2]]) + +lemma real_sqrt_mult_self_sum_ge_zero [simp]: "0 \ sqrt(x*x + y*y)" +by (rule real_sqrt_ge_zero [OF real_mult_self_sum_ge_zero]) + +lemma real_sqrt_sum_squares_ge_zero [simp]: "0 \ sqrt (x\ + y\)" +by (auto intro!: real_sqrt_ge_zero) + +lemma real_sqrt_sum_squares_mult_ge_zero [simp]: + "0 \ sqrt ((x\ + y\)*(xa\ + ya\))" +by (auto intro!: real_sqrt_ge_zero simp add: zero_le_mult_iff) + +lemma real_sqrt_sum_squares_mult_squared_eq [simp]: + "sqrt ((x\ + y\) * (xa\ + ya\)) ^ 2 = (x\ + y\) * (xa\ + ya\)" +by (auto simp add: zero_le_mult_iff simp del: realpow_Suc) + +lemma real_sqrt_sum_squares_ge1 [simp]: "x \ sqrt(x\ + y\)" +by (rule power2_le_imp_le, simp_all) + +lemma real_sqrt_sum_squares_ge2 [simp]: "y \ sqrt(x\ + y\)" +by (rule power2_le_imp_le, simp_all) + +lemma real_sqrt_sos_less_one_iff [simp]: "(sqrt(x\ + y\) < 1) = (x\ + y\ < 1)" +apply (subst real_sqrt_one [symmetric]) +apply (rule real_sqrt_less_iff, auto) +done + +lemma real_sqrt_sos_eq_one_iff [simp]: "(sqrt(x\ + y\) = 1) = (x\ + y\ = 1)" +apply (subst real_sqrt_one [symmetric]) +apply (rule real_sqrt_eq_iff, auto) +done end