paulson@12196: (* Title : NthRoot.thy paulson@12196: Author : Jacques D. Fleuriot paulson@12196: Copyright : 1998 University of Cambridge paulson@14477: Conversion to Isar and new proofs by Lawrence C Paulson, 2004 paulson@12196: *) paulson@12196: huffman@22956: header {* Nth Roots of Real Numbers *} paulson@14324: nipkow@15131: theory NthRoot huffman@23009: imports SEQ Parity Deriv nipkow@15131: begin paulson@14324: huffman@22956: subsection {* Existence of Nth Root *} huffman@20687: huffman@23009: text {* Existence follows from the Intermediate Value Theorem *} paulson@14324: huffman@23009: lemma realpow_pos_nth: huffman@23009: assumes n: "0 < n" huffman@23009: assumes a: "0 < a" huffman@23009: shows "\r>0. r ^ n = (a::real)" huffman@23009: proof - huffman@23009: have "\r\0. r \ (max 1 a) \ r ^ n = a" huffman@23009: proof (rule IVT) huffman@23009: show "0 ^ n \ a" using n a by (simp add: power_0_left) huffman@23009: show "0 \ max 1 a" by simp huffman@23009: from n have n1: "1 \ n" by simp huffman@23009: have "a \ max 1 a ^ 1" by simp huffman@23009: also have "max 1 a ^ 1 \ max 1 a ^ n" huffman@23009: using n1 by (rule power_increasing, simp) huffman@23009: finally show "a \ max 1 a ^ n" . huffman@23009: show "\r. 0 \ r \ r \ max 1 a \ isCont (\x. x ^ n) r" huffman@23009: by (simp add: isCont_power isCont_Id) huffman@23009: qed huffman@23009: then obtain r where r: "0 \ r \ r ^ n = a" by fast huffman@23009: with n a have "r \ 0" by (auto simp add: power_0_left) huffman@23009: with r have "0 < r \ r ^ n = a" by simp huffman@23009: thus ?thesis .. huffman@23009: qed paulson@14325: huffman@23009: text {* Uniqueness of nth positive root *} paulson@14324: paulson@14324: lemma realpow_pos_nth_unique: huffman@23009: "\0 < n; 0 < a\ \ \!r. 0 < r \ r ^ n = (a::real)" paulson@14324: apply (auto intro!: realpow_pos_nth) huffman@23009: apply (rule_tac n=n in power_eq_imp_eq_base, simp_all) paulson@14324: done paulson@14324: huffman@20687: subsection {* Nth Root *} huffman@20687: huffman@22956: text {* We define roots of negative reals such that huffman@22956: @{term "root n (- x) = - root n x"}. This allows huffman@22956: us to omit side conditions from many theorems. *} huffman@20687: huffman@22956: definition huffman@22956: root :: "[nat, real] \ real" where huffman@22956: "root n x = (if 0 < x then (THE u. 0 < u \ u ^ n = x) else huffman@22956: if x < 0 then - (THE u. 0 < u \ u ^ n = - x) else 0)" huffman@20687: huffman@22956: lemma real_root_zero [simp]: "root n 0 = 0" huffman@22956: unfolding root_def by simp huffman@22956: huffman@22956: lemma real_root_minus: "0 < n \ root n (- x) = - root n x" huffman@22956: unfolding root_def by simp huffman@22956: huffman@22956: lemma real_root_gt_zero: "\0 < n; 0 < x\ \ 0 < root n x" huffman@20687: apply (simp add: root_def) huffman@22956: apply (drule (1) realpow_pos_nth_unique) huffman@20687: apply (erule theI' [THEN conjunct1]) huffman@20687: done huffman@20687: huffman@22956: lemma real_root_pow_pos: (* TODO: rename *) huffman@22956: "\0 < n; 0 < x\ \ root n x ^ n = x" huffman@22956: apply (simp add: root_def) huffman@22956: apply (drule (1) realpow_pos_nth_unique) huffman@22956: apply (erule theI' [THEN conjunct2]) huffman@22956: done huffman@20687: huffman@22956: lemma real_root_pow_pos2 [simp]: (* TODO: rename *) huffman@22956: "\0 < n; 0 \ x\ \ root n x ^ n = x" huffman@22956: by (auto simp add: order_le_less real_root_pow_pos) huffman@22956: huffman@22956: lemma real_root_ge_zero: "\0 < n; 0 \ x\ \ 0 \ root n x" huffman@20687: by (auto simp add: order_le_less real_root_gt_zero) huffman@20687: huffman@22956: lemma real_root_power_cancel: "\0 < n; 0 \ x\ \ root n (x ^ n) = x" huffman@22956: apply (subgoal_tac "0 \ x ^ n") huffman@22956: apply (subgoal_tac "0 \ root n (x ^ n)") huffman@22956: apply (subgoal_tac "root n (x ^ n) ^ n = x ^ n") huffman@22956: apply (erule (3) power_eq_imp_eq_base) huffman@22956: apply (erule (1) real_root_pow_pos2) huffman@22956: apply (erule (1) real_root_ge_zero) huffman@22956: apply (erule zero_le_power) huffman@20687: done huffman@20687: huffman@22956: lemma real_root_pos_unique: huffman@22956: "\0 < n; 0 \ y; y ^ n = x\ \ root n x = y" huffman@22956: by (erule subst, rule real_root_power_cancel) huffman@22956: huffman@22956: lemma real_root_one [simp]: "0 < n \ root n 1 = 1" huffman@22956: by (simp add: real_root_pos_unique) huffman@22956: huffman@22956: text {* Root function is strictly monotonic, hence injective *} huffman@22956: huffman@22956: lemma real_root_less_mono_lemma: huffman@22956: "\0 < n; 0 \ x; x < y\ \ root n x < root n y" huffman@22856: apply (subgoal_tac "0 \ y") huffman@22956: apply (subgoal_tac "root n x ^ n < root n y ^ n") huffman@22956: apply (erule power_less_imp_less_base) huffman@22956: apply (erule (1) real_root_ge_zero) huffman@22956: apply simp huffman@22956: apply simp huffman@22721: done huffman@22721: huffman@22956: lemma real_root_less_mono: "\0 < n; x < y\ \ root n x < root n y" huffman@22956: apply (cases "0 \ x") huffman@22956: apply (erule (2) real_root_less_mono_lemma) huffman@22956: apply (cases "0 \ y") huffman@22956: apply (rule_tac y=0 in order_less_le_trans) huffman@22956: apply (subgoal_tac "0 < root n (- x)") huffman@22956: apply (simp add: real_root_minus) huffman@22956: apply (simp add: real_root_gt_zero) huffman@22956: apply (simp add: real_root_ge_zero) huffman@22956: apply (subgoal_tac "root n (- y) < root n (- x)") huffman@22956: apply (simp add: real_root_minus) huffman@22956: apply (simp add: real_root_less_mono_lemma) huffman@22721: done huffman@22721: huffman@22956: lemma real_root_le_mono: "\0 < n; x \ y\ \ root n x \ root n y" huffman@22956: by (auto simp add: order_le_less real_root_less_mono) huffman@22956: huffman@22721: lemma real_root_less_iff [simp]: huffman@22956: "0 < n \ (root n x < root n y) = (x < y)" huffman@22956: apply (cases "x < y") huffman@22956: apply (simp add: real_root_less_mono) huffman@22956: apply (simp add: linorder_not_less real_root_le_mono) huffman@22721: done huffman@22721: huffman@22721: lemma real_root_le_iff [simp]: huffman@22956: "0 < n \ (root n x \ root n y) = (x \ y)" huffman@22956: apply (cases "x \ y") huffman@22956: apply (simp add: real_root_le_mono) huffman@22956: apply (simp add: linorder_not_le real_root_less_mono) huffman@22721: done huffman@22721: huffman@22721: lemma real_root_eq_iff [simp]: huffman@22956: "0 < n \ (root n x = root n y) = (x = y)" huffman@22956: by (simp add: order_eq_iff) huffman@22956: huffman@22956: lemmas real_root_gt_0_iff [simp] = real_root_less_iff [where x=0, simplified] huffman@22956: lemmas real_root_lt_0_iff [simp] = real_root_less_iff [where y=0, simplified] huffman@22956: lemmas real_root_ge_0_iff [simp] = real_root_le_iff [where x=0, simplified] huffman@22956: lemmas real_root_le_0_iff [simp] = real_root_le_iff [where y=0, simplified] huffman@22956: lemmas real_root_eq_0_iff [simp] = real_root_eq_iff [where y=0, simplified] huffman@22721: huffman@22956: text {* Roots of multiplication and division *} huffman@22956: huffman@22956: lemma real_root_mult_lemma: huffman@22956: "\0 < n; 0 \ x; 0 \ y\ \ root n (x * y) = root n x * root n y" huffman@22956: by (simp add: real_root_pos_unique mult_nonneg_nonneg power_mult_distrib) huffman@22956: huffman@22956: lemma real_root_inverse_lemma: huffman@22956: "\0 < n; 0 \ x\ \ root n (inverse x) = inverse (root n x)" huffman@22956: by (simp add: real_root_pos_unique power_inverse [symmetric]) huffman@22721: huffman@22721: lemma real_root_mult: huffman@22956: assumes n: "0 < n" huffman@22956: shows "root n (x * y) = root n x * root n y" huffman@22956: proof (rule linorder_le_cases, rule_tac [!] linorder_le_cases) huffman@22956: assume "0 \ x" and "0 \ y" huffman@22956: thus ?thesis by (rule real_root_mult_lemma [OF n]) huffman@22956: next huffman@22956: assume "0 \ x" and "y \ 0" huffman@22956: hence "0 \ x" and "0 \ - y" by simp_all huffman@22956: hence "root n (x * - y) = root n x * root n (- y)" huffman@22956: by (rule real_root_mult_lemma [OF n]) huffman@22956: thus ?thesis by (simp add: real_root_minus [OF n]) huffman@22956: next huffman@22956: assume "x \ 0" and "0 \ y" huffman@22956: hence "0 \ - x" and "0 \ y" by simp_all huffman@22956: hence "root n (- x * y) = root n (- x) * root n y" huffman@22956: by (rule real_root_mult_lemma [OF n]) huffman@22956: thus ?thesis by (simp add: real_root_minus [OF n]) huffman@22956: next huffman@22956: assume "x \ 0" and "y \ 0" huffman@22956: hence "0 \ - x" and "0 \ - y" by simp_all huffman@22956: hence "root n (- x * - y) = root n (- x) * root n (- y)" huffman@22956: by (rule real_root_mult_lemma [OF n]) huffman@22956: thus ?thesis by (simp add: real_root_minus [OF n]) huffman@22956: qed huffman@22721: huffman@22721: lemma real_root_inverse: huffman@22956: assumes n: "0 < n" huffman@22956: shows "root n (inverse x) = inverse (root n x)" huffman@22956: proof (rule linorder_le_cases) huffman@22956: assume "0 \ x" huffman@22956: thus ?thesis by (rule real_root_inverse_lemma [OF n]) huffman@22956: next huffman@22956: assume "x \ 0" huffman@22956: hence "0 \ - x" by simp huffman@22956: hence "root n (inverse (- x)) = inverse (root n (- x))" huffman@22956: by (rule real_root_inverse_lemma [OF n]) huffman@22956: thus ?thesis by (simp add: real_root_minus [OF n]) huffman@22956: qed huffman@22721: huffman@22956: lemma real_root_divide: huffman@22956: "0 < n \ root n (x / y) = root n x / root n y" huffman@22956: by (simp add: divide_inverse real_root_mult real_root_inverse) huffman@22956: huffman@22956: lemma real_root_power: huffman@22956: "0 < n \ root n (x ^ k) = root n x ^ k" huffman@22956: by (induct k, simp_all add: real_root_mult) huffman@22721: huffman@23042: lemma real_root_abs: "0 < n \ root n \x\ = \root n x\" huffman@23042: by (simp add: abs_if real_root_minus) huffman@23042: huffman@23042: text {* Continuity and derivatives *} huffman@23042: huffman@23042: lemma isCont_root_pos: huffman@23042: assumes n: "0 < n" huffman@23042: assumes x: "0 < x" huffman@23042: shows "isCont (root n) x" huffman@23042: proof - huffman@23042: have "isCont (root n) (root n x ^ n)" huffman@23042: proof (rule isCont_inverse_function [where f="\a. a ^ n"]) huffman@23042: show "0 < root n x" using n x by simp huffman@23042: show "\z. \z - root n x\ \ root n x \ root n (z ^ n) = z" huffman@23042: by (simp add: abs_le_iff real_root_power_cancel n) huffman@23042: show "\z. \z - root n x\ \ root n x \ isCont (\a. a ^ n) z" huffman@23042: by (simp add: isCont_power isCont_Id) huffman@23042: qed huffman@23042: thus ?thesis using n x by simp huffman@23042: qed huffman@23042: huffman@23042: lemma isCont_root_neg: huffman@23042: "\0 < n; x < 0\ \ isCont (root n) x" huffman@23042: apply (subgoal_tac "isCont (\x. - root n (- x)) x") huffman@23042: apply (simp add: real_root_minus) huffman@23042: apply (rule isCont_o2 [OF isCont_minus [OF isCont_Id]]) huffman@23042: apply (simp add: isCont_minus isCont_root_pos) huffman@23042: done huffman@23042: huffman@23042: lemma isCont_root_zero: huffman@23042: "0 < n \ isCont (root n) 0" huffman@23042: unfolding isCont_def huffman@23042: apply (rule LIM_I) huffman@23042: apply (rule_tac x="r ^ n" in exI, safe) huffman@23042: apply (simp add: zero_less_power) huffman@23042: apply (simp add: real_root_abs [symmetric]) huffman@23042: apply (rule_tac n="n" in power_less_imp_less_base, simp_all) huffman@23042: done huffman@23042: huffman@23042: lemma isCont_real_root: "0 < n \ isCont (root n) x" huffman@23042: apply (rule_tac x=x and y=0 in linorder_cases) huffman@23042: apply (simp_all add: isCont_root_pos isCont_root_neg isCont_root_zero) huffman@23042: done huffman@23042: huffman@23042: lemma DERIV_real_root: huffman@23042: assumes n: "0 < n" huffman@23042: assumes x: "0 < x" huffman@23042: shows "DERIV (root n) x :> inverse (real n * root n x ^ (n - Suc 0))" huffman@23042: proof (rule DERIV_inverse_function) huffman@23044: show "0 < x" using x . huffman@23044: show "x < x + 1" by simp huffman@23044: show "\y. 0 < y \ y < x + 1 \ root n y ^ n = y" huffman@23042: using n by simp huffman@23042: show "DERIV (\x. x ^ n) (root n x) :> real n * root n x ^ (n - Suc 0)" huffman@23042: by (rule DERIV_pow) huffman@23042: show "real n * root n x ^ (n - Suc 0) \ 0" huffman@23042: using n x by simp huffman@23042: show "isCont (root n) x" huffman@23042: by (rule isCont_real_root) huffman@23042: qed huffman@23042: huffman@22956: subsection {* Square Root *} huffman@20687: huffman@22956: definition huffman@22956: sqrt :: "real \ real" where huffman@22956: "sqrt = root 2" huffman@20687: huffman@22956: lemma pos2: "0 < (2::nat)" by simp huffman@22956: huffman@22956: lemma real_sqrt_unique: "\y\ = x; 0 \ y\ \ sqrt x = y" huffman@22956: unfolding sqrt_def by (rule real_root_pos_unique [OF pos2]) huffman@20687: huffman@22956: lemma real_sqrt_abs [simp]: "sqrt (x\) = \x\" huffman@22956: apply (rule real_sqrt_unique) huffman@22956: apply (rule power2_abs) huffman@22956: apply (rule abs_ge_zero) huffman@22956: done huffman@20687: huffman@22956: lemma real_sqrt_pow2 [simp]: "0 \ x \ (sqrt x)\ = x" huffman@22956: unfolding sqrt_def by (rule real_root_pow_pos2 [OF pos2]) huffman@22856: huffman@22956: lemma real_sqrt_pow2_iff [simp]: "((sqrt x)\ = x) = (0 \ x)" huffman@22856: apply (rule iffI) huffman@22856: apply (erule subst) huffman@22856: apply (rule zero_le_power2) huffman@22856: apply (erule real_sqrt_pow2) huffman@20687: done huffman@20687: huffman@22956: lemma real_sqrt_zero [simp]: "sqrt 0 = 0" huffman@22956: unfolding sqrt_def by (rule real_root_zero) huffman@22956: huffman@22956: lemma real_sqrt_one [simp]: "sqrt 1 = 1" huffman@22956: unfolding sqrt_def by (rule real_root_one [OF pos2]) huffman@22956: huffman@22956: lemma real_sqrt_minus: "sqrt (- x) = - sqrt x" huffman@22956: unfolding sqrt_def by (rule real_root_minus [OF pos2]) huffman@22956: huffman@22956: lemma real_sqrt_mult: "sqrt (x * y) = sqrt x * sqrt y" huffman@22956: unfolding sqrt_def by (rule real_root_mult [OF pos2]) huffman@22956: huffman@22956: lemma real_sqrt_inverse: "sqrt (inverse x) = inverse (sqrt x)" huffman@22956: unfolding sqrt_def by (rule real_root_inverse [OF pos2]) huffman@22956: huffman@22956: lemma real_sqrt_divide: "sqrt (x / y) = sqrt x / sqrt y" huffman@22956: unfolding sqrt_def by (rule real_root_divide [OF pos2]) huffman@22956: huffman@22956: lemma real_sqrt_power: "sqrt (x ^ k) = sqrt x ^ k" huffman@22956: unfolding sqrt_def by (rule real_root_power [OF pos2]) huffman@22956: huffman@22956: lemma real_sqrt_gt_zero: "0 < x \ 0 < sqrt x" huffman@22956: unfolding sqrt_def by (rule real_root_gt_zero [OF pos2]) huffman@22956: huffman@22956: lemma real_sqrt_ge_zero: "0 \ x \ 0 \ sqrt x" huffman@22956: unfolding sqrt_def by (rule real_root_ge_zero [OF pos2]) huffman@20687: huffman@22956: lemma real_sqrt_less_mono: "x < y \ sqrt x < sqrt y" huffman@22956: unfolding sqrt_def by (rule real_root_less_mono [OF pos2]) huffman@22956: huffman@22956: lemma real_sqrt_le_mono: "x \ y \ sqrt x \ sqrt y" huffman@22956: unfolding sqrt_def by (rule real_root_le_mono [OF pos2]) huffman@22956: huffman@22956: lemma real_sqrt_less_iff [simp]: "(sqrt x < sqrt y) = (x < y)" huffman@22956: unfolding sqrt_def by (rule real_root_less_iff [OF pos2]) huffman@22956: huffman@22956: lemma real_sqrt_le_iff [simp]: "(sqrt x \ sqrt y) = (x \ y)" huffman@22956: unfolding sqrt_def by (rule real_root_le_iff [OF pos2]) huffman@22956: huffman@22956: lemma real_sqrt_eq_iff [simp]: "(sqrt x = sqrt y) = (x = y)" huffman@22956: unfolding sqrt_def by (rule real_root_eq_iff [OF pos2]) huffman@22956: huffman@22956: lemmas real_sqrt_gt_0_iff [simp] = real_sqrt_less_iff [where x=0, simplified] huffman@22956: lemmas real_sqrt_lt_0_iff [simp] = real_sqrt_less_iff [where y=0, simplified] huffman@22956: lemmas real_sqrt_ge_0_iff [simp] = real_sqrt_le_iff [where x=0, simplified] huffman@22956: lemmas real_sqrt_le_0_iff [simp] = real_sqrt_le_iff [where y=0, simplified] huffman@22956: lemmas real_sqrt_eq_0_iff [simp] = real_sqrt_eq_iff [where y=0, simplified] huffman@22956: huffman@22956: lemmas real_sqrt_gt_1_iff [simp] = real_sqrt_less_iff [where x=1, simplified] huffman@22956: lemmas real_sqrt_lt_1_iff [simp] = real_sqrt_less_iff [where y=1, simplified] huffman@22956: lemmas real_sqrt_ge_1_iff [simp] = real_sqrt_le_iff [where x=1, simplified] huffman@22956: lemmas real_sqrt_le_1_iff [simp] = real_sqrt_le_iff [where y=1, simplified] huffman@22956: lemmas real_sqrt_eq_1_iff [simp] = real_sqrt_eq_iff [where y=1, simplified] huffman@20687: huffman@23042: lemma isCont_real_sqrt: "isCont sqrt x" huffman@23042: unfolding sqrt_def by (rule isCont_real_root [OF pos2]) huffman@23042: huffman@23042: lemma DERIV_real_sqrt: huffman@23042: "0 < x \ DERIV sqrt x :> inverse (sqrt x) / 2" huffman@23042: unfolding sqrt_def by (rule DERIV_real_root [OF pos2, simplified]) huffman@23042: huffman@20687: lemma not_real_square_gt_zero [simp]: "(~ (0::real) < x*x) = (x = 0)" huffman@20687: apply auto huffman@20687: apply (cut_tac x = x and y = 0 in linorder_less_linear) huffman@20687: apply (simp add: zero_less_mult_iff) huffman@20687: done huffman@20687: huffman@20687: lemma real_sqrt_abs2 [simp]: "sqrt(x*x) = \x\" huffman@22856: apply (subst power2_eq_square [symmetric]) huffman@20687: apply (rule real_sqrt_abs) huffman@20687: done huffman@20687: huffman@20687: lemma real_sqrt_pow2_gt_zero: "0 < x ==> 0 < (sqrt x)\" huffman@22956: by simp (* TODO: delete *) huffman@20687: huffman@20687: lemma real_sqrt_not_eq_zero: "0 < x ==> sqrt x \ 0" huffman@22956: by simp (* TODO: delete *) huffman@20687: huffman@20687: lemma real_inv_sqrt_pow2: "0 < x ==> inverse (sqrt(x)) ^ 2 = inverse x" huffman@22856: by (simp add: power_inverse [symmetric]) huffman@20687: huffman@20687: lemma real_sqrt_eq_zero_cancel: "[| 0 \ x; sqrt(x) = 0|] ==> x = 0" huffman@22956: by simp huffman@20687: huffman@20687: lemma real_sqrt_ge_one: "1 \ x ==> 1 \ sqrt x" huffman@22956: by simp huffman@20687: huffman@22443: lemma sqrt_divide_self_eq: huffman@22443: assumes nneg: "0 \ x" huffman@22443: shows "sqrt x / x = inverse (sqrt x)" huffman@22443: proof cases huffman@22443: assume "x=0" thus ?thesis by simp huffman@22443: next huffman@22443: assume nz: "x\0" huffman@22443: hence pos: "0 0" by (simp add: divide_inverse nneg nz) huffman@22443: show "inverse (sqrt x) / (sqrt x / x) = 1" huffman@22443: by (simp add: divide_inverse mult_assoc [symmetric] huffman@22443: power2_eq_square [symmetric] real_inv_sqrt_pow2 pos nz) huffman@22443: qed huffman@22443: qed huffman@22443: huffman@22721: lemma real_divide_square_eq [simp]: "(((r::real) * a) / (r * r)) = a / r" huffman@22721: apply (simp add: divide_inverse) huffman@22721: apply (case_tac "r=0") huffman@22721: apply (auto simp add: mult_ac) huffman@22721: done huffman@22721: huffman@22856: subsection {* Square Root of Sum of Squares *} huffman@22856: huffman@22856: lemma real_sqrt_mult_self_sum_ge_zero [simp]: "0 \ sqrt(x*x + y*y)" huffman@22968: by (rule real_sqrt_ge_zero [OF sum_squares_ge_zero]) huffman@22856: huffman@22856: lemma real_sqrt_sum_squares_ge_zero [simp]: "0 \ sqrt (x\ + y\)" huffman@22961: by simp huffman@22856: huffman@22856: lemma real_sqrt_sum_squares_mult_ge_zero [simp]: huffman@22856: "0 \ sqrt ((x\ + y\)*(xa\ + ya\))" huffman@22856: by (auto intro!: real_sqrt_ge_zero simp add: zero_le_mult_iff) huffman@22856: huffman@22856: lemma real_sqrt_sum_squares_mult_squared_eq [simp]: huffman@22856: "sqrt ((x\ + y\) * (xa\ + ya\)) ^ 2 = (x\ + y\) * (xa\ + ya\)" huffman@22956: by (auto simp add: zero_le_mult_iff) huffman@22856: huffman@22856: lemma real_sqrt_sum_squares_ge1 [simp]: "x \ sqrt(x\ + y\)" huffman@22856: by (rule power2_le_imp_le, simp_all) huffman@22856: huffman@22856: lemma real_sqrt_sum_squares_ge2 [simp]: "y \ sqrt(x\ + y\)" huffman@22856: by (rule power2_le_imp_le, simp_all) huffman@22856: huffman@22858: lemma power2_sum: huffman@22858: fixes x y :: "'a::{number_ring,recpower}" huffman@22858: shows "(x + y)\ = x\ + y\ + 2 * x * y" huffman@22858: by (simp add: left_distrib right_distrib power2_eq_square) huffman@22858: huffman@22858: lemma power2_diff: huffman@22858: fixes x y :: "'a::{number_ring,recpower}" huffman@22858: shows "(x - y)\ = x\ + y\ - 2 * x * y" huffman@22858: by (simp add: left_diff_distrib right_diff_distrib power2_eq_square) huffman@22858: huffman@22858: lemma real_sqrt_sum_squares_triangle_ineq: huffman@22858: "sqrt ((a + c)\ + (b + d)\) \ sqrt (a\ + b\) + sqrt (c\ + d\)" huffman@22858: apply (rule power2_le_imp_le, simp) huffman@22858: apply (simp add: power2_sum) huffman@22858: apply (simp only: mult_assoc right_distrib [symmetric]) huffman@22858: apply (rule mult_left_mono) huffman@22858: apply (rule power2_le_imp_le) huffman@22858: apply (simp add: power2_sum power_mult_distrib) huffman@22858: apply (simp add: ring_distrib) huffman@22858: apply (subgoal_tac "0 \ b\ * c\ + a\ * d\ - 2 * (a * c) * (b * d)", simp) huffman@22858: apply (rule_tac b="(a * d - b * c)\" in ord_le_eq_trans) huffman@22858: apply (rule zero_le_power2) huffman@22858: apply (simp add: power2_diff power_mult_distrib) huffman@22858: apply (simp add: mult_nonneg_nonneg) huffman@22858: apply simp huffman@22858: apply (simp add: add_increasing) huffman@22858: done huffman@22858: huffman@22956: text "Legacy theorem names:" huffman@22956: lemmas real_root_pos2 = real_root_power_cancel huffman@22956: lemmas real_root_pos_pos = real_root_gt_zero [THEN order_less_imp_le] huffman@22956: lemmas real_root_pos_pos_le = real_root_ge_zero huffman@22956: lemmas real_sqrt_mult_distrib = real_sqrt_mult huffman@22956: lemmas real_sqrt_mult_distrib2 = real_sqrt_mult huffman@22956: lemmas real_sqrt_eq_zero_cancel_iff = real_sqrt_eq_0_iff huffman@22956: huffman@22956: (* needed for CauchysMeanTheorem.het_base from AFP *) huffman@22956: lemma real_root_pos: "0 < x \ root (Suc n) (x ^ (Suc n)) = x" huffman@22956: by (rule real_root_power_cancel [OF zero_less_Suc order_less_imp_le]) huffman@22956: huffman@22956: (* FIXME: the stronger version of real_root_less_iff huffman@22956: breaks CauchysMeanTheorem.list_gmean_gt_iff from AFP. *) huffman@22956: huffman@22956: declare real_root_less_iff [simp del] huffman@22956: lemma real_root_less_iff_nonneg [simp]: huffman@22956: "\0 < n; 0 \ x; 0 \ y\ \ (root n x < root n y) = (x < y)" huffman@22956: by (rule real_root_less_iff) huffman@22956: paulson@14324: end