# HG changeset patch # User huffman # Date 1159059419 -7200 # Node ID fedb901be39236b21330621adc1e593fd079ea3c # Parent e3d2b881ed0fa7a0eb649ce6ed646b576b5b9be4 move root and sqrt stuff from Transcendental to NthRoot diff -r e3d2b881ed0f -r fedb901be392 src/HOL/Hyperreal/NthRoot.thy --- a/src/HOL/Hyperreal/NthRoot.thy Sun Sep 24 01:04:44 2006 +0200 +++ b/src/HOL/Hyperreal/NthRoot.thy Sun Sep 24 02:56:59 2006 +0200 @@ -10,6 +10,15 @@ imports SEQ begin +definition + + root :: "[nat, real] \ real" + "root n x = (THE u. (0 < x \ 0 < u) \ (u ^ n = x))" + + sqrt :: "real \ real" + "sqrt x = root 2 x" + + text {* Various lemmas needed for this result. We follow the proof given by John Lindsay Orr (\texttt{jorr@math.unl.edu}) in his Analysis @@ -170,4 +179,211 @@ apply (drule_tac [4] x = y in realpow_less, auto) done +subsection {* Nth Root *} + +lemma real_root_zero [simp]: "root (Suc n) 0 = 0" +apply (simp add: root_def) +apply (safe intro!: the_equality power_0_Suc elim!: realpow_zero_zero) +done + +lemma real_root_pow_pos: + "0 < x ==> (root (Suc n) x) ^ (Suc n) = x" +apply (simp add: root_def del: realpow_Suc) +apply (drule_tac n="Suc n" in realpow_pos_nth_unique, simp) +apply (erule theI' [THEN conjunct2]) +done + +lemma real_root_pow_pos2: "0 \ x ==> (root (Suc n) x) ^ (Suc n) = x" +by (auto dest!: real_le_imp_less_or_eq dest: real_root_pow_pos) + +lemma real_root_pos: + "0 < x ==> root(Suc n) (x ^ (Suc n)) = x" +apply (simp add: root_def) +apply (rule the_equality) +apply (frule_tac [2] n = n in zero_less_power) +apply (auto simp add: zero_less_mult_iff) +apply (rule_tac x = u and y = x in linorder_cases) +apply (drule_tac n1 = n and x = u in zero_less_Suc [THEN [3] realpow_less]) +apply (drule_tac [4] n1 = n and x = x in zero_less_Suc [THEN [3] realpow_less]) +apply (auto) +done + +lemma real_root_pos2: "0 \ x ==> root(Suc n) (x ^ (Suc n)) = x" +by (auto dest!: real_le_imp_less_or_eq real_root_pos) + +lemma real_root_gt_zero: + "0 < x ==> 0 < root (Suc n) x" +apply (simp add: root_def del: realpow_Suc) +apply (drule_tac n="Suc n" in realpow_pos_nth_unique, simp) +apply (erule theI' [THEN conjunct1]) +done + +lemma real_root_pos_pos: + "0 < x ==> 0 \ root(Suc n) x" +by (rule real_root_gt_zero [THEN order_less_imp_le]) + +lemma real_root_pos_pos_le: "0 \ x ==> 0 \ root(Suc n) x" +by (auto simp add: order_le_less real_root_gt_zero) + +lemma real_root_one [simp]: "root (Suc n) 1 = 1" +apply (simp add: root_def) +apply (rule the_equality, auto) +apply (rule ccontr) +apply (rule_tac x = u and y = 1 in linorder_cases) +apply (drule_tac n = n in realpow_Suc_less_one) +apply (drule_tac [4] n = n in power_gt1_lemma) +apply (auto) +done + + +subsection{*Square Root*} + +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]) + +lemma real_sqrt_zero [simp]: "sqrt 0 = 0" +by (simp add: sqrt_def) + +lemma real_sqrt_one [simp]: "sqrt 1 = 1" +by (simp add: sqrt_def) + +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) +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\ ^ 2 = \x\" +by (rule real_sqrt_pow2_iff [THEN iffD2], arith) + +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) +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_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 not_real_square_gt_zero [simp]: "(~ (0::real) < x*x) = (x = 0)" +apply auto +apply (cut_tac x = x and y = 0 in linorder_less_linear) +apply (simp add: zero_less_mult_iff) +done + +lemma real_sqrt_gt_zero: "0 < x ==> 0 < sqrt(x)" +by (simp add: sqrt_def real_root_gt_zero) + +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" +apply (case_tac n, simp) +apply (simp add: root_def) +apply (rule someI2 [of _ r], safe) +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) + +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 + +lemma real_sqrt_abs2 [simp]: "sqrt(x*x) = \x\" +apply (rule realpow_two [THEN subst]) +apply (subst numeral_2_eq_2 [symmetric]) +apply (rule real_sqrt_abs) +done + +lemma real_sqrt_pow2_gt_zero: "0 < x ==> 0 < (sqrt x)\" +by simp + +lemma real_sqrt_not_eq_zero: "0 < x ==> sqrt x \ 0" +apply (frule real_sqrt_pow2_gt_zero) +apply (auto simp add: numeral_2_eq_2) +done + +lemma real_inv_sqrt_pow2: "0 < x ==> inverse (sqrt(x)) ^ 2 = inverse x" +by (cut_tac n1 = 2 and a1 = "sqrt x" in power_inverse [symmetric], auto) + +lemma real_sqrt_eq_zero_cancel: "[| 0 \ x; sqrt(x) = 0|] ==> x = 0" +apply (drule real_le_imp_less_or_eq) +apply (auto dest: real_sqrt_not_eq_zero) +done + +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) +done + end diff -r e3d2b881ed0f -r fedb901be392 src/HOL/Hyperreal/Transcendental.thy --- a/src/HOL/Hyperreal/Transcendental.thy Sun Sep 24 01:04:44 2006 +0200 +++ b/src/HOL/Hyperreal/Transcendental.thy Sun Sep 24 02:56:59 2006 +0200 @@ -12,11 +12,6 @@ begin definition - root :: "[nat,real] => real" - "root n x = (THE u. ((0::real) < x --> 0 < u) & (u ^ n = x))" - - sqrt :: "real => real" - "sqrt x = root 2 x" exp :: "real => real" "exp x = (\n. inverse(real (fact n)) * (x ^ n))" @@ -51,211 +46,6 @@ "arctan y = (SOME x. -(pi/2) < x & x < pi/2 & tan x = y)" -lemma real_root_zero [simp]: "root (Suc n) 0 = 0" -apply (simp add: root_def) -apply (safe intro!: the_equality power_0_Suc elim!: realpow_zero_zero) -done - -lemma real_root_pow_pos: - "0 < x ==> (root (Suc n) x) ^ (Suc n) = x" -apply (simp add: root_def del: realpow_Suc) -apply (drule_tac n="Suc n" in realpow_pos_nth_unique, simp) -apply (erule theI' [THEN conjunct2]) -done - -lemma real_root_pow_pos2: "0 \ x ==> (root (Suc n) x) ^ (Suc n) = x" -by (auto dest!: real_le_imp_less_or_eq dest: real_root_pow_pos) - -lemma real_root_pos: - "0 < x ==> root(Suc n) (x ^ (Suc n)) = x" -apply (simp add: root_def) -apply (rule the_equality) -apply (frule_tac [2] n = n in zero_less_power) -apply (auto simp add: zero_less_mult_iff) -apply (rule_tac x = u and y = x in linorder_cases) -apply (drule_tac n1 = n and x = u in zero_less_Suc [THEN [3] realpow_less]) -apply (drule_tac [4] n1 = n and x = x in zero_less_Suc [THEN [3] realpow_less]) -apply (auto) -done - -lemma real_root_pos2: "0 \ x ==> root(Suc n) (x ^ (Suc n)) = x" -by (auto dest!: real_le_imp_less_or_eq real_root_pos) - -lemma real_root_gt_zero: - "0 < x ==> 0 < root (Suc n) x" -apply (simp add: root_def del: realpow_Suc) -apply (drule_tac n="Suc n" in realpow_pos_nth_unique, simp) -apply (erule theI' [THEN conjunct1]) -done - -lemma real_root_pos_pos: - "0 < x ==> 0 \ root(Suc n) x" -by (rule real_root_gt_zero [THEN order_less_imp_le]) - -lemma real_root_pos_pos_le: "0 \ x ==> 0 \ root(Suc n) x" -by (auto simp add: order_le_less real_root_gt_zero) - -lemma real_root_one [simp]: "root (Suc n) 1 = 1" -apply (simp add: root_def) -apply (rule the_equality, auto) -apply (rule ccontr) -apply (rule_tac x = u and y = 1 in linorder_cases) -apply (drule_tac n = n in realpow_Suc_less_one) -apply (drule_tac [4] n = n in power_gt1_lemma) -apply (auto) -done - - -subsection{*Square Root*} - -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]) - -lemma real_sqrt_zero [simp]: "sqrt 0 = 0" -by (simp add: sqrt_def) - -lemma real_sqrt_one [simp]: "sqrt 1 = 1" -by (simp add: sqrt_def) - -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) -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\ ^ 2 = \x\" -by (rule real_sqrt_pow2_iff [THEN iffD2], arith) - -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) -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_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 not_real_square_gt_zero [simp]: "(~ (0::real) < x*x) = (x = 0)" -apply auto -apply (cut_tac x = x and y = 0 in linorder_less_linear) -apply (simp add: zero_less_mult_iff) -done - -lemma real_sqrt_gt_zero: "0 < x ==> 0 < sqrt(x)" -by (simp add: sqrt_def real_root_gt_zero) - -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" -apply (case_tac n, simp) -apply (simp add: root_def) -apply (rule someI2 [of _ r], safe) -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) - -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 - -lemma real_sqrt_abs2 [simp]: "sqrt(x*x) = \x\" -apply (rule realpow_two [THEN subst]) -apply (subst numeral_2_eq_2 [symmetric]) -apply (rule real_sqrt_abs) -done - -lemma real_sqrt_pow2_gt_zero: "0 < x ==> 0 < (sqrt x)\" -by simp - -lemma real_sqrt_not_eq_zero: "0 < x ==> sqrt x \ 0" -apply (frule real_sqrt_pow2_gt_zero) -apply (auto simp add: numeral_2_eq_2) -done - -lemma real_inv_sqrt_pow2: "0 < x ==> inverse (sqrt(x)) ^ 2 = inverse x" -by (cut_tac n1 = 2 and a1 = "sqrt x" in power_inverse [symmetric], auto) - -lemma real_sqrt_eq_zero_cancel: "[| 0 \ x; sqrt(x) = 0|] ==> x = 0" -apply (drule real_le_imp_less_or_eq) -apply (auto dest: real_sqrt_not_eq_zero) -done - -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) -done - subsection{*Exponential Function*}