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 haftmann@28952: imports Parity Deriv nipkow@15131: begin paulson@14324: hoelzl@51483: lemma abs_sgn_eq: "abs (sgn x :: real) = (if x = 0 then 0 else 1)" hoelzl@51483: by (simp add: sgn_real_def) hoelzl@51483: hoelzl@51483: lemma inverse_sgn: "sgn (inverse a) = inverse (sgn a :: real)" hoelzl@51483: by (simp add: sgn_real_def) hoelzl@51483: hoelzl@51483: lemma power_eq_iff_eq_base: hoelzl@51483: fixes a b :: "_ :: linordered_semidom" hoelzl@51483: shows "0 < n \ 0 \ a \ 0 \ b \ a ^ n = b ^ n \ a = b" hoelzl@51483: using power_eq_imp_eq_base[of a n b] by auto hoelzl@51483: 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@44289: by simp 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@23047: (* Used by Integration/RealRandVar.thy in AFP *) huffman@23047: lemma realpow_pos_nth2: "(0::real) < a \ \r>0. r ^ Suc n = a" huffman@23047: by (blast intro: realpow_pos_nth) huffman@23047: huffman@23009: text {* Uniqueness of nth positive root *} paulson@14324: hoelzl@51483: lemma realpow_pos_nth_unique: "\0 < n; 0 < a\ \ \!r. 0 < r \ r ^ n = (a::real)" hoelzl@51483: by (auto intro!: realpow_pos_nth simp: power_eq_iff_eq_base) 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: hoelzl@51483: lemma inj_sgn_power: assumes "0 < n" shows "inj (\y. sgn y * \y\^n :: real)" (is "inj ?f") hoelzl@51483: proof (rule injI) hoelzl@51483: have x: "\a b :: real. (0 < a \ b < 0) \ (a < 0 \ 0 < b) \ a \ b" by auto hoelzl@51483: fix x y assume "?f x = ?f y" with power_eq_iff_eq_base[of n "\x\" "\y\"] `0a\ ^ n = x \ x = sgn b * \b\ ^ n \ 0 < n \ a = (b::real)" hoelzl@51483: using inj_sgn_power[THEN injD, of n a b] by simp hoelzl@51483: hoelzl@51483: definition root :: "nat \ real \ real" where hoelzl@51483: "root n x = (if n = 0 then 0 else the_inv (\y. sgn y * \y\^n) x)" hoelzl@51483: hoelzl@51483: lemma root_0 [simp]: "root 0 x = 0" hoelzl@51483: by (simp add: root_def) hoelzl@51483: hoelzl@51483: lemma root_sgn_power: "0 < n \ root n (sgn y * \y\^n) = y" hoelzl@51483: using the_inv_f_f[OF inj_sgn_power] by (simp add: root_def) hoelzl@51483: hoelzl@51483: lemma sgn_power_root: hoelzl@51483: assumes "0 < n" shows "sgn (root n x) * \(root n x)\^n = x" (is "?f (root n x) = x") hoelzl@51483: proof cases hoelzl@51483: assume "x \ 0" hoelzl@51483: with realpow_pos_nth[OF `0 < n`, of "\x\"] obtain r where "0 < r" "r ^ n = \x\" by auto hoelzl@51483: with `x \ 0` have S: "x \ range ?f" hoelzl@51483: by (intro image_eqI[of _ _ "sgn x * r"]) hoelzl@51483: (auto simp: abs_mult sgn_mult power_mult_distrib abs_sgn_eq mult_sgn_abs) hoelzl@51483: from `0 < n` f_the_inv_into_f[OF inj_sgn_power[OF `0 < n`] this] show ?thesis hoelzl@51483: by (simp add: root_def) hoelzl@51483: qed (insert `0 < n` root_sgn_power[of n 0], simp) hoelzl@51483: hoelzl@51483: lemma split_root: "P (root n x) \ (n = 0 \ P 0) \ (0 < n \ (\y. sgn y * \y\^n = x \ P y))" hoelzl@51483: apply (cases "n = 0") hoelzl@51483: apply simp_all hoelzl@51483: apply (metis root_sgn_power sgn_power_root) hoelzl@51483: done huffman@20687: huffman@22956: lemma real_root_zero [simp]: "root n 0 = 0" hoelzl@51483: by (simp split: split_root add: sgn_zero_iff) hoelzl@51483: hoelzl@51483: lemma real_root_minus: "root n (- x) = - root n x" hoelzl@51483: by (clarsimp split: split_root elim!: sgn_power_injE simp: sgn_minus) huffman@22956: hoelzl@51483: lemma real_root_less_mono: "\0 < n; x < y\ \ root n x < root n y" hoelzl@51483: proof (clarsimp split: split_root) hoelzl@51483: have x: "\a b :: real. (0 < b \ a < 0) \ \ a > b" by auto hoelzl@51483: fix a b :: real assume "0 < n" "sgn a * \a\ ^ n < sgn b * \b\ ^ n" then show "a < b" hoelzl@51483: using power_less_imp_less_base[of a n b] power_less_imp_less_base[of "-b" n "-a"] hoelzl@51483: by (simp add: sgn_real_def power_less_zero_eq x[of "a ^ n" "- ((- b) ^ n)"] split: split_if_asm) hoelzl@51483: qed huffman@22956: huffman@22956: lemma real_root_gt_zero: "\0 < n; 0 < x\ \ 0 < root n x" hoelzl@51483: using real_root_less_mono[of n 0 x] by simp hoelzl@51483: hoelzl@51483: lemma real_root_ge_zero: "0 \ x \ 0 \ root n x" hoelzl@51483: using real_root_gt_zero[of n x] by (cases "n = 0") (auto simp add: le_less) huffman@20687: huffman@22956: lemma real_root_pow_pos: (* TODO: rename *) huffman@22956: "\0 < n; 0 < x\ \ root n x ^ n = x" hoelzl@51483: using sgn_power_root[of n x] real_root_gt_zero[of n x] by simp 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: hoelzl@51483: lemma sgn_root: "0 < n \ sgn (root n x) = sgn x" hoelzl@51483: by (auto split: split_root simp: sgn_real_def power_less_zero_eq) hoelzl@51483: huffman@23046: lemma odd_real_root_pow: "odd n \ root n x ^ n = x" hoelzl@51483: using sgn_power_root[of n x] by (simp add: odd_pos sgn_real_def split: split_if_asm) huffman@20687: huffman@22956: lemma real_root_power_cancel: "\0 < n; 0 \ x\ \ root n (x ^ n) = x" hoelzl@51483: using root_sgn_power[of n x] by (auto simp add: le_less power_0_left) huffman@20687: huffman@23046: lemma odd_real_root_power_cancel: "odd n \ root n (x ^ n) = x" hoelzl@51483: using root_sgn_power[of n x] by (simp add: odd_pos sgn_real_def power_0_left split: split_if_asm) huffman@23046: hoelzl@51483: lemma real_root_pos_unique: "\0 < n; 0 \ y; y ^ n = x\ \ root n x = y" hoelzl@51483: using root_sgn_power[of n y] by (auto simp add: le_less power_0_left) huffman@22956: huffman@23046: lemma odd_real_root_unique: huffman@23046: "\odd n; y ^ n = x\ \ root n x = y" huffman@23046: by (erule subst, rule odd_real_root_power_cancel) huffman@23046: 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_le_mono: "\0 < n; x \ y\ \ root n x \ root n y" hoelzl@51483: 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@23257: lemma real_root_gt_1_iff [simp]: "0 < n \ (1 < root n y) = (1 < y)" huffman@23257: by (insert real_root_less_iff [where x=1], simp) huffman@23257: huffman@23257: lemma real_root_lt_1_iff [simp]: "0 < n \ (root n x < 1) = (x < 1)" huffman@23257: by (insert real_root_less_iff [where y=1], simp) huffman@23257: huffman@23257: lemma real_root_ge_1_iff [simp]: "0 < n \ (1 \ root n y) = (1 \ y)" huffman@23257: by (insert real_root_le_iff [where x=1], simp) huffman@23257: huffman@23257: lemma real_root_le_1_iff [simp]: "0 < n \ (root n x \ 1) = (x \ 1)" huffman@23257: by (insert real_root_le_iff [where y=1], simp) huffman@23257: huffman@23257: lemma real_root_eq_1_iff [simp]: "0 < n \ (root n x = 1) = (x = 1)" huffman@23257: by (insert real_root_eq_iff [where y=1], simp) huffman@23257: hoelzl@51483: text {* Roots of multiplication and division *} hoelzl@51483: hoelzl@51483: lemma real_root_mult: "root n (x * y) = root n x * root n y" hoelzl@51483: by (auto split: split_root elim!: sgn_power_injE simp: sgn_mult abs_mult power_mult_distrib) hoelzl@51483: hoelzl@51483: lemma real_root_inverse: "root n (inverse x) = inverse (root n x)" hoelzl@51483: by (auto split: split_root elim!: sgn_power_injE simp: inverse_sgn power_inverse) hoelzl@51483: hoelzl@51483: lemma real_root_divide: "root n (x / y) = root n x / root n y" hoelzl@51483: by (simp add: divide_inverse real_root_mult real_root_inverse) hoelzl@51483: hoelzl@51483: lemma real_root_abs: "0 < n \ root n \x\ = \root n x\" hoelzl@51483: by (simp add: abs_if real_root_minus) hoelzl@51483: hoelzl@51483: lemma real_root_power: "0 < n \ root n (x ^ k) = root n x ^ k" hoelzl@51483: by (induct k) (simp_all add: real_root_mult) hoelzl@51483: huffman@23257: text {* Roots of roots *} huffman@23257: huffman@23257: lemma real_root_Suc_0 [simp]: "root (Suc 0) x = x" huffman@23257: by (simp add: odd_real_root_unique) huffman@23257: hoelzl@51483: lemma real_root_mult_exp: "root (m * n) x = root m (root n x)" hoelzl@51483: by (auto split: split_root elim!: sgn_power_injE hoelzl@51483: simp: sgn_zero_iff sgn_mult power_mult[symmetric] abs_mult power_mult_distrib abs_sgn_eq) huffman@23257: hoelzl@51483: lemma real_root_commute: "root m (root n x) = root n (root m x)" haftmann@57512: by (simp add: real_root_mult_exp [symmetric] mult.commute) huffman@23257: huffman@23257: text {* Monotonicity in first argument *} huffman@23257: huffman@23257: lemma real_root_strict_decreasing: huffman@23257: "\0 < n; n < N; 1 < x\ \ root N x < root n x" huffman@23257: apply (subgoal_tac "root n (root N x) ^ n < root N (root n x) ^ N", simp) huffman@23257: apply (simp add: real_root_commute power_strict_increasing huffman@23257: del: real_root_pow_pos2) huffman@23257: done huffman@23257: huffman@23257: lemma real_root_strict_increasing: huffman@23257: "\0 < n; n < N; 0 < x; x < 1\ \ root n x < root N x" huffman@23257: apply (subgoal_tac "root N (root n x) ^ N < root n (root N x) ^ n", simp) huffman@23257: apply (simp add: real_root_commute power_strict_decreasing huffman@23257: del: real_root_pow_pos2) huffman@23257: done huffman@23257: huffman@23257: lemma real_root_decreasing: huffman@23257: "\0 < n; n < N; 1 \ x\ \ root N x \ root n x" huffman@23257: by (auto simp add: order_le_less real_root_strict_decreasing) huffman@23257: huffman@23257: lemma real_root_increasing: huffman@23257: "\0 < n; n < N; 0 \ x; x \ 1\ \ root n x \ root N x" huffman@23257: by (auto simp add: order_le_less real_root_strict_increasing) huffman@23257: huffman@23042: text {* Continuity and derivatives *} huffman@23042: hoelzl@51483: lemma isCont_real_root: "isCont (root n) x" hoelzl@51483: proof cases hoelzl@51483: assume n: "0 < n" hoelzl@51483: let ?f = "\y::real. sgn y * \y\^n" hoelzl@51483: have "continuous_on ({0..} \ {.. 0}) (\x. if 0 < x then x ^ n else - ((-x) ^ n) :: real)" hoelzl@56371: using n by (intro continuous_on_If continuous_intros) auto hoelzl@51483: then have "continuous_on UNIV ?f" hoelzl@51483: by (rule continuous_on_cong[THEN iffD1, rotated 2]) (auto simp: not_less real_sgn_neg le_less n) hoelzl@51483: then have [simp]: "\x. isCont ?f x" hoelzl@51483: by (simp add: continuous_on_eq_continuous_at) huffman@23042: hoelzl@51483: have "isCont (root n) (?f (root n x))" hoelzl@51483: by (rule isCont_inverse_function [where f="?f" and d=1]) (auto simp: root_sgn_power n) hoelzl@51483: then show ?thesis hoelzl@51483: by (simp add: sgn_power_root n) hoelzl@51483: qed (simp add: root_def[abs_def]) huffman@23042: hoelzl@51478: lemma tendsto_real_root[tendsto_intros]: hoelzl@51483: "(f ---> x) F \ ((\x. root n (f x)) ---> root n x) F" hoelzl@51483: using isCont_tendsto_compose[OF isCont_real_root, of f x F] . hoelzl@51478: hoelzl@51478: lemma continuous_real_root[continuous_intros]: hoelzl@51483: "continuous F f \ continuous F (\x. root n (f x))" hoelzl@51478: unfolding continuous_def by (rule tendsto_real_root) hoelzl@51478: hoelzl@56371: lemma continuous_on_real_root[continuous_intros]: hoelzl@51483: "continuous_on s f \ continuous_on s (\x. root n (f x))" hoelzl@51478: unfolding continuous_on_def by (auto intro: tendsto_real_root) hoelzl@51478: 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 hoelzl@51483: qed (rule isCont_real_root) huffman@23042: huffman@23046: lemma DERIV_odd_real_root: huffman@23046: assumes n: "odd n" huffman@23046: assumes x: "x \ 0" huffman@23046: shows "DERIV (root n) x :> inverse (real n * root n x ^ (n - Suc 0))" huffman@23046: proof (rule DERIV_inverse_function) huffman@23046: show "x - 1 < x" by simp huffman@23046: show "x < x + 1" by simp huffman@23046: show "\y. x - 1 < y \ y < x + 1 \ root n y ^ n = y" huffman@23046: using n by (simp add: odd_real_root_pow) huffman@23046: show "DERIV (\x. x ^ n) (root n x) :> real n * root n x ^ (n - Suc 0)" huffman@23046: by (rule DERIV_pow) huffman@23046: show "real n * root n x ^ (n - Suc 0) \ 0" huffman@23046: using odd_pos [OF n] x by simp hoelzl@51483: qed (rule isCont_real_root) huffman@23046: hoelzl@31880: lemma DERIV_even_real_root: hoelzl@31880: assumes n: "0 < n" and "even n" hoelzl@31880: assumes x: "x < 0" hoelzl@31880: shows "DERIV (root n) x :> inverse (- real n * root n x ^ (n - Suc 0))" hoelzl@31880: proof (rule DERIV_inverse_function) hoelzl@31880: show "x - 1 < x" by simp hoelzl@31880: show "x < 0" using x . hoelzl@31880: next hoelzl@31880: show "\y. x - 1 < y \ y < 0 \ - (root n y ^ n) = y" hoelzl@31880: proof (rule allI, rule impI, erule conjE) hoelzl@31880: fix y assume "x - 1 < y" and "y < 0" hoelzl@31880: hence "root n (-y) ^ n = -y" using `0 < n` by simp hoelzl@51483: with real_root_minus and `even n` hoelzl@31880: show "- (root n y ^ n) = y" by simp hoelzl@31880: qed hoelzl@31880: next hoelzl@31880: show "DERIV (\x. - (x ^ n)) (root n x) :> - real n * root n x ^ (n - Suc 0)" hoelzl@56381: by (auto intro!: derivative_eq_intros simp: real_of_nat_def) hoelzl@31880: show "- real n * root n x ^ (n - Suc 0) \ 0" hoelzl@31880: using n x by simp hoelzl@51483: qed (rule isCont_real_root) hoelzl@31880: hoelzl@31880: lemma DERIV_real_root_generic: hoelzl@31880: assumes "0 < n" and "x \ 0" wenzelm@49753: and "\ even n ; 0 < x \ \ D = inverse (real n * root n x ^ (n - Suc 0))" wenzelm@49753: and "\ even n ; x < 0 \ \ D = - inverse (real n * root n x ^ (n - Suc 0))" wenzelm@49753: and "odd n \ D = inverse (real n * root n x ^ (n - Suc 0))" hoelzl@31880: shows "DERIV (root n) x :> D" hoelzl@31880: using assms by (cases "even n", cases "0 < x", hoelzl@31880: auto intro: DERIV_real_root[THEN DERIV_cong] hoelzl@31880: DERIV_odd_real_root[THEN DERIV_cong] hoelzl@31880: DERIV_even_real_root[THEN DERIV_cong]) hoelzl@31880: huffman@22956: subsection {* Square Root *} huffman@20687: hoelzl@51483: definition sqrt :: "real \ real" where huffman@22956: "sqrt = root 2" huffman@20687: huffman@22956: lemma pos2: "0 < (2::nat)" by simp huffman@22956: wenzelm@53015: lemma real_sqrt_unique: "\y\<^sup>2 = x; 0 \ y\ \ sqrt x = y" huffman@22956: unfolding sqrt_def by (rule real_root_pos_unique [OF pos2]) huffman@20687: wenzelm@53015: lemma real_sqrt_abs [simp]: "sqrt (x\<^sup>2) = \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: wenzelm@53015: lemma real_sqrt_pow2 [simp]: "0 \ x \ (sqrt x)\<^sup>2 = x" huffman@22956: unfolding sqrt_def by (rule real_root_pow_pos2 [OF pos2]) huffman@22856: wenzelm@53015: lemma real_sqrt_pow2_iff [simp]: "((sqrt x)\<^sup>2 = 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: hoelzl@56889: lemma real_sqrt_four [simp]: "sqrt 4 = 2" hoelzl@56889: using real_sqrt_abs[of 2] by simp hoelzl@56889: huffman@22956: lemma real_sqrt_minus: "sqrt (- x) = - sqrt x" hoelzl@51483: unfolding sqrt_def by (rule real_root_minus) huffman@22956: huffman@22956: lemma real_sqrt_mult: "sqrt (x * y) = sqrt x * sqrt y" hoelzl@51483: unfolding sqrt_def by (rule real_root_mult) huffman@22956: hoelzl@56889: lemma real_sqrt_mult_self[simp]: "sqrt a * sqrt a = \a\" hoelzl@56889: using real_sqrt_abs[of a] unfolding power2_eq_square real_sqrt_mult . hoelzl@56889: huffman@22956: lemma real_sqrt_inverse: "sqrt (inverse x) = inverse (sqrt x)" hoelzl@51483: unfolding sqrt_def by (rule real_root_inverse) huffman@22956: huffman@22956: lemma real_sqrt_divide: "sqrt (x / y) = sqrt x / sqrt y" hoelzl@51483: unfolding sqrt_def by (rule real_root_divide) 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" hoelzl@51483: unfolding sqrt_def by (rule real_root_ge_zero) 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: hoelzl@54413: lemma real_le_lsqrt: "0 \ x \ 0 \ y \ x \ y\<^sup>2 \ sqrt x \ y" hoelzl@54413: using real_sqrt_le_iff[of x "y\<^sup>2"] by simp hoelzl@54413: hoelzl@54413: lemma real_le_rsqrt: "x\<^sup>2 \ y \ x \ sqrt y" hoelzl@54413: using real_sqrt_le_mono[of "x\<^sup>2" y] by simp hoelzl@54413: hoelzl@54413: lemma real_less_rsqrt: "x\<^sup>2 < y \ x < sqrt y" hoelzl@54413: using real_sqrt_less_mono[of "x\<^sup>2" y] by simp hoelzl@54413: hoelzl@54413: lemma sqrt_even_pow2: hoelzl@54413: assumes n: "even n" hoelzl@54413: shows "sqrt (2 ^ n) = 2 ^ (n div 2)" hoelzl@54413: proof - hoelzl@54413: from n obtain m where m: "n = 2 * m" hoelzl@54413: unfolding even_mult_two_ex .. hoelzl@54413: from m have "sqrt (2 ^ n) = sqrt ((2 ^ m)\<^sup>2)" haftmann@57512: by (simp only: power_mult[symmetric] mult.commute) hoelzl@54413: then show ?thesis hoelzl@54413: using m by simp hoelzl@54413: qed hoelzl@54413: huffman@53594: lemmas real_sqrt_gt_0_iff [simp] = real_sqrt_less_iff [where x=0, unfolded real_sqrt_zero] huffman@53594: lemmas real_sqrt_lt_0_iff [simp] = real_sqrt_less_iff [where y=0, unfolded real_sqrt_zero] huffman@53594: lemmas real_sqrt_ge_0_iff [simp] = real_sqrt_le_iff [where x=0, unfolded real_sqrt_zero] huffman@53594: lemmas real_sqrt_le_0_iff [simp] = real_sqrt_le_iff [where y=0, unfolded real_sqrt_zero] huffman@53594: lemmas real_sqrt_eq_0_iff [simp] = real_sqrt_eq_iff [where y=0, unfolded real_sqrt_zero] huffman@22956: huffman@53594: lemmas real_sqrt_gt_1_iff [simp] = real_sqrt_less_iff [where x=1, unfolded real_sqrt_one] huffman@53594: lemmas real_sqrt_lt_1_iff [simp] = real_sqrt_less_iff [where y=1, unfolded real_sqrt_one] huffman@53594: lemmas real_sqrt_ge_1_iff [simp] = real_sqrt_le_iff [where x=1, unfolded real_sqrt_one] huffman@53594: lemmas real_sqrt_le_1_iff [simp] = real_sqrt_le_iff [where y=1, unfolded real_sqrt_one] huffman@53594: lemmas real_sqrt_eq_1_iff [simp] = real_sqrt_eq_iff [where y=1, unfolded real_sqrt_one] huffman@20687: huffman@23042: lemma isCont_real_sqrt: "isCont sqrt x" hoelzl@51483: unfolding sqrt_def by (rule isCont_real_root) huffman@23042: hoelzl@51478: lemma tendsto_real_sqrt[tendsto_intros]: hoelzl@51478: "(f ---> x) F \ ((\x. sqrt (f x)) ---> sqrt x) F" hoelzl@51483: unfolding sqrt_def by (rule tendsto_real_root) hoelzl@51478: hoelzl@51478: lemma continuous_real_sqrt[continuous_intros]: hoelzl@51478: "continuous F f \ continuous F (\x. sqrt (f x))" hoelzl@51483: unfolding sqrt_def by (rule continuous_real_root) hoelzl@51478: hoelzl@56371: lemma continuous_on_real_sqrt[continuous_intros]: hoelzl@57155: "continuous_on s f \ continuous_on s (\x. sqrt (f x))" hoelzl@51483: unfolding sqrt_def by (rule continuous_on_real_root) hoelzl@51478: hoelzl@31880: lemma DERIV_real_sqrt_generic: hoelzl@31880: assumes "x \ 0" hoelzl@31880: assumes "x > 0 \ D = inverse (sqrt x) / 2" hoelzl@31880: assumes "x < 0 \ D = - inverse (sqrt x) / 2" hoelzl@31880: shows "DERIV sqrt x :> D" hoelzl@31880: using assms unfolding sqrt_def hoelzl@31880: by (auto intro!: DERIV_real_root_generic) hoelzl@31880: huffman@23042: lemma DERIV_real_sqrt: huffman@23042: "0 < x \ DERIV sqrt x :> inverse (sqrt x) / 2" hoelzl@31880: using DERIV_real_sqrt_generic by simp hoelzl@31880: hoelzl@31880: declare hoelzl@56381: DERIV_real_sqrt_generic[THEN DERIV_chain2, derivative_intros] hoelzl@56381: DERIV_real_root_generic[THEN DERIV_chain2, derivative_intros] 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: wenzelm@53076: lemma real_inv_sqrt_pow2: "0 < x ==> (inverse (sqrt x))\<^sup>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" haftmann@57512: 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: hoelzl@54413: lemma real_div_sqrt: "0 \ x \ x / sqrt x = sqrt x" hoelzl@54413: apply (cases "x = 0") hoelzl@54413: apply simp_all hoelzl@54413: using sqrt_divide_self_eq[of x] hoelzl@54413: apply (simp add: inverse_eq_divide field_simps) hoelzl@54413: done hoelzl@54413: 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") haftmann@57514: apply (auto simp add: ac_simps) huffman@22721: done huffman@22721: huffman@23049: lemma lemma_real_divide_sqrt_less: "0 < u ==> u / sqrt 2 < u" huffman@35216: by (simp add: divide_less_eq) huffman@23049: huffman@23049: lemma four_x_squared: huffman@23049: fixes x::real wenzelm@53015: shows "4 * x\<^sup>2 = (2 * x)\<^sup>2" huffman@23049: by (simp add: power2_eq_square) huffman@23049: hoelzl@57275: lemma sqrt_at_top: "LIM x at_top. sqrt x :: real :> at_top" hoelzl@57275: by (rule filterlim_at_top_at_top[where Q="\x. True" and P="\x. 0 < x" and g="power2"]) hoelzl@57275: (auto intro: eventually_gt_at_top) hoelzl@57275: huffman@22856: subsection {* Square Root of Sum of Squares *} huffman@22856: lp15@55967: lemma sum_squares_bound: lp15@55967: fixes x:: "'a::linordered_field" lp15@55967: shows "2*x*y \ x^2 + y^2" lp15@55967: proof - lp15@55967: have "(x-y)^2 = x*x - 2*x*y + y*y" lp15@55967: by algebra lp15@55967: then have "0 \ x^2 - 2*x*y + y^2" lp15@55967: by (metis sum_power2_ge_zero zero_le_double_add_iff_zero_le_single_add power2_eq_square) lp15@55967: then show ?thesis lp15@55967: by arith lp15@55967: qed huffman@22856: lp15@55967: lemma arith_geo_mean: lp15@55967: fixes u:: "'a::linordered_field" assumes "u\<^sup>2 = x*y" "x\0" "y\0" shows "u \ (x + y)/2" lp15@55967: apply (rule power2_le_imp_le) lp15@55967: using sum_squares_bound assms lp15@55967: apply (auto simp: zero_le_mult_iff) lp15@55967: by (auto simp: algebra_simps power2_eq_square) lp15@55967: lp15@55967: lemma arith_geo_mean_sqrt: lp15@55967: fixes x::real assumes "x\0" "y\0" shows "sqrt(x*y) \ (x + y)/2" lp15@55967: apply (rule arith_geo_mean) lp15@55967: using assms lp15@55967: apply (auto simp: zero_le_mult_iff) lp15@55967: done huffman@23049: huffman@22856: lemma real_sqrt_sum_squares_mult_ge_zero [simp]: wenzelm@53015: "0 \ sqrt ((x\<^sup>2 + y\<^sup>2)*(xa\<^sup>2 + ya\<^sup>2))" lp15@55967: by (metis real_sqrt_ge_0_iff split_mult_pos_le sum_power2_ge_zero) huffman@22856: huffman@22856: lemma real_sqrt_sum_squares_mult_squared_eq [simp]: wenzelm@53076: "(sqrt ((x\<^sup>2 + y\<^sup>2) * (xa\<^sup>2 + ya\<^sup>2)))\<^sup>2 = (x\<^sup>2 + y\<^sup>2) * (xa\<^sup>2 + ya\<^sup>2)" huffman@44320: by (simp add: zero_le_mult_iff) huffman@22856: wenzelm@53015: lemma real_sqrt_sum_squares_eq_cancel: "sqrt (x\<^sup>2 + y\<^sup>2) = x \ y = 0" wenzelm@53015: by (drule_tac f = "%x. x\<^sup>2" in arg_cong, simp) huffman@23049: wenzelm@53015: lemma real_sqrt_sum_squares_eq_cancel2: "sqrt (x\<^sup>2 + y\<^sup>2) = y \ x = 0" wenzelm@53015: by (drule_tac f = "%x. x\<^sup>2" in arg_cong, simp) huffman@23049: wenzelm@53015: lemma real_sqrt_sum_squares_ge1 [simp]: "x \ sqrt (x\<^sup>2 + y\<^sup>2)" huffman@22856: by (rule power2_le_imp_le, simp_all) huffman@22856: wenzelm@53015: lemma real_sqrt_sum_squares_ge2 [simp]: "y \ sqrt (x\<^sup>2 + y\<^sup>2)" huffman@23049: by (rule power2_le_imp_le, simp_all) huffman@23049: wenzelm@53015: lemma real_sqrt_ge_abs1 [simp]: "\x\ \ sqrt (x\<^sup>2 + y\<^sup>2)" huffman@22856: by (rule power2_le_imp_le, simp_all) huffman@22856: wenzelm@53015: lemma real_sqrt_ge_abs2 [simp]: "\y\ \ sqrt (x\<^sup>2 + y\<^sup>2)" huffman@23049: by (rule power2_le_imp_le, simp_all) huffman@23049: huffman@23049: lemma le_real_sqrt_sumsq [simp]: "x \ sqrt (x * x + y * y)" huffman@23049: by (simp add: power2_eq_square [symmetric]) huffman@23049: huffman@22858: lemma real_sqrt_sum_squares_triangle_ineq: wenzelm@53015: "sqrt ((a + c)\<^sup>2 + (b + d)\<^sup>2) \ sqrt (a\<^sup>2 + b\<^sup>2) + sqrt (c\<^sup>2 + d\<^sup>2)" huffman@22858: apply (rule power2_le_imp_le, simp) huffman@22858: apply (simp add: power2_sum) haftmann@57512: apply (simp only: mult.assoc distrib_left [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) nipkow@23477: apply (simp add: ring_distribs) wenzelm@53015: apply (subgoal_tac "0 \ b\<^sup>2 * c\<^sup>2 + a\<^sup>2 * d\<^sup>2 - 2 * (a * c) * (b * d)", simp) wenzelm@53015: apply (rule_tac b="(a * d - b * c)\<^sup>2" in ord_le_eq_trans) huffman@22858: apply (rule zero_le_power2) huffman@22858: apply (simp add: power2_diff power_mult_distrib) nipkow@56536: apply (simp) huffman@22858: apply simp huffman@22858: apply (simp add: add_increasing) huffman@22858: done huffman@22858: huffman@23122: lemma real_sqrt_sum_squares_less: wenzelm@53015: "\\x\ < u / sqrt 2; \y\ < u / sqrt 2\ \ sqrt (x\<^sup>2 + y\<^sup>2) < u" huffman@23122: apply (rule power2_less_imp_less, simp) huffman@23122: apply (drule power_strict_mono [OF _ abs_ge_zero pos2]) huffman@23122: apply (drule power_strict_mono [OF _ abs_ge_zero pos2]) huffman@23122: apply (simp add: power_divide) huffman@23122: apply (drule order_le_less_trans [OF abs_ge_zero]) huffman@23122: apply (simp add: zero_less_divide_iff) huffman@23122: done huffman@23122: huffman@23049: text{*Needed for the infinitely close relation over the nonstandard huffman@23049: complex numbers*} huffman@23049: lemma lemma_sqrt_hcomplex_capprox: wenzelm@53015: "[| 0 < u; x < u/2; y < u/2; 0 \ x; 0 \ y |] ==> sqrt (x\<^sup>2 + y\<^sup>2) < u" huffman@23049: apply (rule_tac y = "u/sqrt 2" in order_le_less_trans) huffman@23049: apply (erule_tac [2] lemma_real_divide_sqrt_less) huffman@23049: apply (rule power2_le_imp_le) huffman@44349: apply (auto simp add: zero_le_divide_iff power_divide) wenzelm@53015: apply (rule_tac t = "u\<^sup>2" in real_sum_of_halves [THEN subst]) huffman@23049: apply (rule add_mono) huffman@30273: apply (auto simp add: four_x_squared intro: power_mono) huffman@23049: done huffman@23049: 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: paulson@14324: end