# HG changeset patch # User wenzelm # Date 1468356877 -7200 # Node ID f3781c5fb03f445926406000d61ccca7cdb51d49 # Parent 2100fbbdc3f1d66e43a319e0882dad08521601d1 misc tuning and modernization; diff -r 2100fbbdc3f1 -r f3781c5fb03f src/HOL/NthRoot.thy --- a/src/HOL/NthRoot.thy Tue Jul 12 21:53:56 2016 +0200 +++ b/src/HOL/NthRoot.thy Tue Jul 12 22:54:37 2016 +0200 @@ -1,70 +1,87 @@ -(* Title : NthRoot.thy - Author : Jacques D. Fleuriot - Copyright : 1998 University of Cambridge - Conversion to Isar and new proofs by Lawrence C Paulson, 2004 +(* Title: HOL/NthRoot.thy + Author: Jacques D. Fleuriot, 1998 + Author: Lawrence C Paulson, 2004 *) section \Nth Roots of Real Numbers\ theory NthRoot -imports Deriv Binomial + imports Deriv Binomial begin + subsection \Existence of Nth Root\ text \Existence follows from the Intermediate Value Theorem\ lemma realpow_pos_nth: + fixes a :: real assumes n: "0 < n" - assumes a: "0 < a" - shows "\r>0. r ^ n = (a::real)" + and a: "0 < a" + shows "\r>0. r ^ n = a" proof - have "\r\0. r \ (max 1 a) \ r ^ n = a" proof (rule IVT) - show "0 ^ n \ a" using n a by (simp add: power_0_left) - show "0 \ max 1 a" by simp - from n have n1: "1 \ n" by simp - have "a \ max 1 a ^ 1" by simp + show "0 ^ n \ a" + using n a by (simp add: power_0_left) + show "0 \ max 1 a" + by simp + from n have n1: "1 \ n" + by simp + have "a \ max 1 a ^ 1" + by simp also have "max 1 a ^ 1 \ max 1 a ^ n" - using n1 by (rule power_increasing, simp) + using n1 by (rule power_increasing) simp finally show "a \ max 1 a ^ n" . show "\r. 0 \ r \ r \ max 1 a \ isCont (\x. x ^ n) r" by simp qed - then obtain r where r: "0 \ r \ r ^ n = a" by fast - with n a have "r \ 0" by (auto simp add: power_0_left) - with r have "0 < r \ r ^ n = a" by simp - thus ?thesis .. + then obtain r where r: "0 \ r \ r ^ n = a" + by fast + with n a have "r \ 0" + by (auto simp add: power_0_left) + with r have "0 < r \ r ^ n = a" + by simp + then show ?thesis .. qed (* Used by Integration/RealRandVar.thy in AFP *) lemma realpow_pos_nth2: "(0::real) < a \ \r>0. r ^ Suc n = a" -by (blast intro: realpow_pos_nth) + by (blast intro: realpow_pos_nth) -text \Uniqueness of nth positive root\ +text \Uniqueness of nth positive root.\ +lemma realpow_pos_nth_unique: "0 < n \ 0 < a \ \!r. 0 < r \ r ^ n = a" for a :: real + by (auto intro!: realpow_pos_nth simp: power_eq_iff_eq_base) -lemma realpow_pos_nth_unique: "\0 < n; 0 < a\ \ \!r. 0 < r \ r ^ n = (a::real)" - by (auto intro!: realpow_pos_nth simp: power_eq_iff_eq_base) subsection \Nth Root\ -text \We define roots of negative reals such that - @{term "root n (- x) = - root n x"}. This allows - us to omit side conditions from many theorems.\ +text \ + We define roots of negative reals such that \root n (- x) = - root n x\. + This allows us to omit side conditions from many theorems. +\ -lemma inj_sgn_power: assumes "0 < n" shows "inj (\y. sgn y * \y\^n :: real)" (is "inj ?f") +lemma inj_sgn_power: + assumes "0 < n" + shows "inj (\y. sgn y * \y\^n :: real)" + (is "inj ?f") proof (rule injI) - have x: "\a b :: real. (0 < a \ b < 0) \ (a < 0 \ 0 < b) \ a \ b" by auto - fix x y assume "?f x = ?f y" with power_eq_iff_eq_base[of n "\x\" "\y\"] \0 show "x = y" + have x: "(0 < a \ b < 0) \ (a < 0 \ 0 < b) \ a \ b" for a b :: real + by auto + fix x y + assume "?f x = ?f y" + with power_eq_iff_eq_base[of n "\x\" "\y\"] \0 < n\ show "x = y" by (cases rule: linorder_cases[of 0 x, case_product linorder_cases[of 0 y]]) (simp_all add: x) qed -lemma sgn_power_injE: "sgn a * \a\ ^ n = x \ x = sgn b * \b\ ^ n \ 0 < n \ a = (b::real)" +lemma sgn_power_injE: + "sgn a * \a\ ^ n = x \ x = sgn b * \b\ ^ n \ 0 < n \ a = b" + for a b :: real using inj_sgn_power[THEN injD, of n a b] by simp -definition root :: "nat \ real \ real" where - "root n x = (if n = 0 then 0 else the_inv (\y. sgn y * \y\^n) x)" +definition root :: "nat \ real \ real" + where "root n x = (if n = 0 then 0 else the_inv (\y. sgn y * \y\^n) x)" lemma root_0 [simp]: "root 0 x = 0" by (simp add: root_def) @@ -73,16 +90,24 @@ using the_inv_f_f[OF inj_sgn_power] by (simp add: root_def) lemma sgn_power_root: - assumes "0 < n" shows "sgn (root n x) * \(root n x)\^n = x" (is "?f (root n x) = x") -proof cases - assume "x \ 0" - with realpow_pos_nth[OF \0 < n\, of "\x\"] obtain r where "0 < r" "r ^ n = \x\" by auto + assumes "0 < n" + shows "sgn (root n x) * \(root n x)\^n = x" + (is "?f (root n x) = x") +proof (cases "x = 0") + case True + with assms root_sgn_power[of n 0] show ?thesis + by simp +next + case False + with realpow_pos_nth[OF \0 < n\, of "\x\"] + obtain r where "0 < r" "r ^ n = \x\" + by auto with \x \ 0\ have S: "x \ range ?f" by (intro image_eqI[of _ _ "sgn x * r"]) (auto simp: abs_mult sgn_mult power_mult_distrib abs_sgn_eq mult_sgn_abs) from \0 < n\ f_the_inv_into_f[OF inj_sgn_power[OF \0 < n\] this] show ?thesis by (simp add: root_def) -qed (insert \0 < n\ root_sgn_power[of n 0], simp) +qed lemma split_root: "P (root n x) \ (n = 0 \ P 0) \ (0 < n \ (\y. sgn y * \y\^n = x \ P y))" apply (cases "n = 0") @@ -96,72 +121,74 @@ lemma real_root_minus: "root n (- x) = - root n x" by (clarsimp split: split_root elim!: sgn_power_injE simp: sgn_minus) -lemma real_root_less_mono: "\0 < n; x < y\ \ root n x < root n y" +lemma real_root_less_mono: "0 < n \ x < y \ root n x < root n y" proof (clarsimp split: split_root) - have x: "\a b :: real. (0 < b \ a < 0) \ \ a > b" by auto - fix a b :: real assume "0 < n" "sgn a * \a\ ^ n < sgn b * \b\ ^ n" then show "a < b" - using power_less_imp_less_base[of a n b] power_less_imp_less_base[of "-b" n "-a"] - by (simp add: sgn_real_def x [of "a ^ n" "- ((- b) ^ n)"] split: if_split_asm) + have *: "0 < b \ a < 0 \ \ a > b" for a b :: real + by auto + fix a b :: real + assume "0 < n" "sgn a * \a\ ^ n < sgn b * \b\ ^ n" + then show "a < b" + using power_less_imp_less_base[of a n b] + power_less_imp_less_base[of "- b" n "- a"] + by (simp add: sgn_real_def * [of "a ^ n" "- ((- b) ^ n)"] + split: if_split_asm) qed -lemma real_root_gt_zero: "\0 < n; 0 < x\ \ 0 < root n x" +lemma real_root_gt_zero: "0 < n \ 0 < x \ 0 < root n x" using real_root_less_mono[of n 0 x] by simp lemma real_root_ge_zero: "0 \ x \ 0 \ root n x" - using real_root_gt_zero[of n x] by (cases "n = 0") (auto simp add: le_less) + using real_root_gt_zero[of n x] + by (cases "n = 0") (auto simp add: le_less) -lemma real_root_pow_pos: (* TODO: rename *) - "\0 < n; 0 < x\ \ root n x ^ n = x" +lemma real_root_pow_pos: "0 < n \ 0 < x \ root n x ^ n = x" (* TODO: rename *) using sgn_power_root[of n x] real_root_gt_zero[of n x] by simp -lemma real_root_pow_pos2 [simp]: (* TODO: rename *) - "\0 < n; 0 \ x\ \ root n x ^ n = x" -by (auto simp add: order_le_less real_root_pow_pos) +lemma real_root_pow_pos2 [simp]: "0 < n \ 0 \ x \ root n x ^ n = x" (* TODO: rename *) + by (auto simp add: order_le_less real_root_pow_pos) lemma sgn_root: "0 < n \ sgn (root n x) = sgn x" by (auto split: split_root simp: sgn_real_def) lemma odd_real_root_pow: "odd n \ root n x ^ n = x" - using sgn_power_root[of n x] by (simp add: odd_pos sgn_real_def split: if_split_asm) + using sgn_power_root[of n x] + by (simp add: odd_pos sgn_real_def split: if_split_asm) -lemma real_root_power_cancel: "\0 < n; 0 \ x\ \ root n (x ^ n) = x" +lemma real_root_power_cancel: "0 < n \ 0 \ x \ root n (x ^ n) = x" using root_sgn_power[of n x] by (auto simp add: le_less power_0_left) lemma odd_real_root_power_cancel: "odd n \ root n (x ^ n) = x" - using root_sgn_power[of n x] by (simp add: odd_pos sgn_real_def power_0_left split: if_split_asm) + using root_sgn_power[of n x] + by (simp add: odd_pos sgn_real_def power_0_left split: if_split_asm) -lemma real_root_pos_unique: "\0 < n; 0 \ y; y ^ n = x\ \ root n x = y" +lemma real_root_pos_unique: "0 < n \ 0 \ y \ y ^ n = x \ root n x = y" using root_sgn_power[of n y] by (auto simp add: le_less power_0_left) -lemma odd_real_root_unique: - "\odd n; y ^ n = x\ \ root n x = y" -by (erule subst, rule odd_real_root_power_cancel) +lemma odd_real_root_unique: "odd n \ y ^ n = x \ root n x = y" + by (erule subst, rule odd_real_root_power_cancel) lemma real_root_one [simp]: "0 < n \ root n 1 = 1" -by (simp add: real_root_pos_unique) + by (simp add: real_root_pos_unique) -text \Root function is strictly monotonic, hence injective\ +text \Root function is strictly monotonic, hence injective.\ -lemma real_root_le_mono: "\0 < n; x \ y\ \ root n x \ root n y" +lemma real_root_le_mono: "0 < n \ x \ y \ root n x \ root n y" by (auto simp add: order_le_less real_root_less_mono) -lemma real_root_less_iff [simp]: - "0 < n \ (root n x < root n y) = (x < y)" -apply (cases "x < y") -apply (simp add: real_root_less_mono) -apply (simp add: linorder_not_less real_root_le_mono) -done +lemma real_root_less_iff [simp]: "0 < n \ root n x < root n y \ x < y" + apply (cases "x < y") + apply (simp add: real_root_less_mono) + apply (simp add: linorder_not_less real_root_le_mono) + done -lemma real_root_le_iff [simp]: - "0 < n \ (root n x \ root n y) = (x \ y)" -apply (cases "x \ y") -apply (simp add: real_root_le_mono) -apply (simp add: linorder_not_le real_root_less_mono) -done +lemma real_root_le_iff [simp]: "0 < n \ root n x \ root n y \ x \ y" + apply (cases "x \ y") + apply (simp add: real_root_le_mono) + apply (simp add: linorder_not_le real_root_less_mono) + done -lemma real_root_eq_iff [simp]: - "0 < n \ (root n x = root n y) = (x = y)" -by (simp add: order_eq_iff) +lemma real_root_eq_iff [simp]: "0 < n \ root n x = root n y \ x = y" + by (simp add: order_eq_iff) lemmas real_root_gt_0_iff [simp] = real_root_less_iff [where x=0, simplified] lemmas real_root_lt_0_iff [simp] = real_root_less_iff [where y=0, simplified] @@ -169,28 +196,31 @@ lemmas real_root_le_0_iff [simp] = real_root_le_iff [where y=0, simplified] lemmas real_root_eq_0_iff [simp] = real_root_eq_iff [where y=0, simplified] -lemma real_root_gt_1_iff [simp]: "0 < n \ (1 < root n y) = (1 < y)" -by (insert real_root_less_iff [where x=1], simp) +lemma real_root_gt_1_iff [simp]: "0 < n \ 1 < root n y \ 1 < y" + using real_root_less_iff [where x=1] by simp -lemma real_root_lt_1_iff [simp]: "0 < n \ (root n x < 1) = (x < 1)" -by (insert real_root_less_iff [where y=1], simp) +lemma real_root_lt_1_iff [simp]: "0 < n \ root n x < 1 \ x < 1" + using real_root_less_iff [where y=1] by simp + +lemma real_root_ge_1_iff [simp]: "0 < n \ 1 \ root n y \ 1 \ y" + using real_root_le_iff [where x=1] by simp -lemma real_root_ge_1_iff [simp]: "0 < n \ (1 \ root n y) = (1 \ y)" -by (insert real_root_le_iff [where x=1], simp) +lemma real_root_le_1_iff [simp]: "0 < n \ root n x \ 1 \ x \ 1" + using real_root_le_iff [where y=1] by simp -lemma real_root_le_1_iff [simp]: "0 < n \ (root n x \ 1) = (x \ 1)" -by (insert real_root_le_iff [where y=1], simp) +lemma real_root_eq_1_iff [simp]: "0 < n \ root n x = 1 \ x = 1" + using real_root_eq_iff [where y=1] by simp -lemma real_root_eq_1_iff [simp]: "0 < n \ (root n x = 1) = (x = 1)" -by (insert real_root_eq_iff [where y=1], simp) -text \Roots of multiplication and division\ +text \Roots of multiplication and division.\ lemma real_root_mult: "root n (x * y) = root n x * root n y" - by (auto split: split_root elim!: sgn_power_injE simp: sgn_mult abs_mult power_mult_distrib) + by (auto split: split_root elim!: sgn_power_injE + simp: sgn_mult abs_mult power_mult_distrib) lemma real_root_inverse: "root n (inverse x) = inverse (root n x)" - by (auto split: split_root elim!: sgn_power_injE simp: inverse_sgn power_inverse) + by (auto split: split_root elim!: sgn_power_injE + simp: inverse_sgn power_inverse) lemma real_root_divide: "root n (x / y) = root n x / root n y" by (simp add: divide_inverse real_root_mult real_root_inverse) @@ -201,124 +231,138 @@ lemma real_root_power: "0 < n \ root n (x ^ k) = root n x ^ k" by (induct k) (simp_all add: real_root_mult) -text \Roots of roots\ + +text \Roots of roots.\ lemma real_root_Suc_0 [simp]: "root (Suc 0) x = x" -by (simp add: odd_real_root_unique) + by (simp add: odd_real_root_unique) lemma real_root_mult_exp: "root (m * n) x = root m (root n x)" by (auto split: split_root elim!: sgn_power_injE - simp: sgn_zero_iff sgn_mult power_mult[symmetric] abs_mult power_mult_distrib abs_sgn_eq) + simp: sgn_zero_iff sgn_mult power_mult[symmetric] + abs_mult power_mult_distrib abs_sgn_eq) lemma real_root_commute: "root m (root n x) = root n (root m x)" by (simp add: real_root_mult_exp [symmetric] mult.commute) -text \Monotonicity in first argument\ + +text \Monotonicity in first argument.\ -lemma real_root_strict_decreasing: - "\0 < n; n < N; 1 < x\ \ root N x < root n x" -apply (subgoal_tac "root n (root N x) ^ n < root N (root n x) ^ N", simp) -apply (simp add: real_root_commute power_strict_increasing - del: real_root_pow_pos2) -done +lemma real_root_strict_decreasing: "0 < n \ n < N \ 1 < x \ root N x < root n x" + apply (subgoal_tac "root n (root N x) ^ n < root N (root n x) ^ N") + apply simp + apply (simp add: real_root_commute power_strict_increasing del: real_root_pow_pos2) + done -lemma real_root_strict_increasing: - "\0 < n; n < N; 0 < x; x < 1\ \ root n x < root N x" -apply (subgoal_tac "root N (root n x) ^ N < root n (root N x) ^ n", simp) -apply (simp add: real_root_commute power_strict_decreasing - del: real_root_pow_pos2) -done +lemma real_root_strict_increasing: "0 < n \ n < N \ 0 < x \ x < 1 \ root n x < root N x" + apply (subgoal_tac "root N (root n x) ^ N < root n (root N x) ^ n") + apply simp + apply (simp add: real_root_commute power_strict_decreasing del: real_root_pow_pos2) + done -lemma real_root_decreasing: - "\0 < n; n < N; 1 \ x\ \ root N x \ root n x" -by (auto simp add: order_le_less real_root_strict_decreasing) +lemma real_root_decreasing: "0 < n \ n < N \ 1 \ x \ root N x \ root n x" + by (auto simp add: order_le_less real_root_strict_decreasing) -lemma real_root_increasing: - "\0 < n; n < N; 0 \ x; x \ 1\ \ root n x \ root N x" -by (auto simp add: order_le_less real_root_strict_increasing) +lemma real_root_increasing: "0 < n \ n < N \ 0 \ x \ x \ 1 \ root n x \ root N x" + by (auto simp add: order_le_less real_root_strict_increasing) -text \Continuity and derivatives\ + +text \Continuity and derivatives.\ lemma isCont_real_root: "isCont (root n) x" -proof cases - assume n: "0 < n" +proof (cases "n > 0") + case True let ?f = "\y::real. sgn y * \y\^n" have "continuous_on ({0..} \ {.. 0}) (\x. if 0 < x then x ^ n else - ((-x) ^ n) :: real)" - using n by (intro continuous_on_If continuous_intros) auto + using True by (intro continuous_on_If continuous_intros) auto then have "continuous_on UNIV ?f" - by (rule continuous_on_cong[THEN iffD1, rotated 2]) (auto simp: not_less le_less n) - then have [simp]: "\x. isCont ?f x" + by (rule continuous_on_cong[THEN iffD1, rotated 2]) (auto simp: not_less le_less True) + then have [simp]: "isCont ?f x" for x by (simp add: continuous_on_eq_continuous_at) - have "isCont (root n) (?f (root n x))" - by (rule isCont_inverse_function [where f="?f" and d=1]) (auto simp: root_sgn_power n) + by (rule isCont_inverse_function [where f="?f" and d=1]) (auto simp: root_sgn_power True) then show ?thesis - by (simp add: sgn_power_root n) -qed (simp add: root_def[abs_def]) + by (simp add: sgn_power_root True) +next + case False + then show ?thesis + by (simp add: root_def[abs_def]) +qed -lemma tendsto_real_root[tendsto_intros]: +lemma tendsto_real_root [tendsto_intros]: "(f \ x) F \ ((\x. root n (f x)) \ root n x) F" using isCont_tendsto_compose[OF isCont_real_root, of f x F] . -lemma continuous_real_root[continuous_intros]: +lemma continuous_real_root [continuous_intros]: "continuous F f \ continuous F (\x. root n (f x))" unfolding continuous_def by (rule tendsto_real_root) -lemma continuous_on_real_root[continuous_intros]: +lemma continuous_on_real_root [continuous_intros]: "continuous_on s f \ continuous_on s (\x. root n (f x))" unfolding continuous_on_def by (auto intro: tendsto_real_root) lemma DERIV_real_root: assumes n: "0 < n" - assumes x: "0 < x" + and x: "0 < x" shows "DERIV (root n) x :> inverse (real n * root n x ^ (n - Suc 0))" proof (rule DERIV_inverse_function) - show "0 < x" using x . - show "x < x + 1" by simp + show "0 < x" + using x . + show "x < x + 1" + by simp show "\y. 0 < y \ y < x + 1 \ root n y ^ n = y" using n by simp show "DERIV (\x. x ^ n) (root n x) :> real n * root n x ^ (n - Suc 0)" by (rule DERIV_pow) show "real n * root n x ^ (n - Suc 0) \ 0" using n x by simp -qed (rule isCont_real_root) + show "isCont (root n) x" + by (rule isCont_real_root) +qed lemma DERIV_odd_real_root: assumes n: "odd n" - assumes x: "x \ 0" + and x: "x \ 0" shows "DERIV (root n) x :> inverse (real n * root n x ^ (n - Suc 0))" proof (rule DERIV_inverse_function) - show "x - 1 < x" by simp - show "x < x + 1" by simp + show "x - 1 < x" + by simp + show "x < x + 1" + by simp show "\y. x - 1 < y \ y < x + 1 \ root n y ^ n = y" using n by (simp add: odd_real_root_pow) show "DERIV (\x. x ^ n) (root n x) :> real n * root n x ^ (n - Suc 0)" by (rule DERIV_pow) show "real n * root n x ^ (n - Suc 0) \ 0" using odd_pos [OF n] x by simp -qed (rule isCont_real_root) + show "isCont (root n) x" + by (rule isCont_real_root) +qed lemma DERIV_even_real_root: - assumes n: "0 < n" and "even n" - assumes x: "x < 0" + assumes n: "0 < n" + and "even n" + and x: "x < 0" shows "DERIV (root n) x :> inverse (- real n * root n x ^ (n - Suc 0))" proof (rule DERIV_inverse_function) - show "x - 1 < x" by simp - show "x < 0" using x . -next + show "x - 1 < x" + by simp + show "x < 0" + using x . show "\y. x - 1 < y \ y < 0 \ - (root n y ^ n) = y" proof (rule allI, rule impI, erule conjE) fix y assume "x - 1 < y" and "y < 0" - hence "root n (-y) ^ n = -y" using \0 < n\ by simp + then have "root n (-y) ^ n = -y" using \0 < n\ by simp with real_root_minus and \even n\ show "- (root n y ^ n) = y" by simp qed -next show "DERIV (\x. - (x ^ n)) (root n x) :> - real n * root n x ^ (n - Suc 0)" by (auto intro!: derivative_eq_intros) show "- real n * root n x ^ (n - Suc 0) \ 0" using n x by simp -qed (rule isCont_real_root) + show "isCont (root n) x" + by (rule isCont_real_root) +qed lemma DERIV_real_root_generic: assumes "0 < n" and "x \ 0" @@ -326,84 +370,87 @@ and "\ even n ; x < 0 \ \ D = - inverse (real n * root n x ^ (n - Suc 0))" and "odd n \ D = inverse (real n * root n x ^ (n - Suc 0))" shows "DERIV (root n) x :> D" -using assms by (cases "even n", cases "0 < x", - auto intro: DERIV_real_root[THEN DERIV_cong] + using assms + by (cases "even n", cases "0 < x", + auto intro: DERIV_real_root[THEN DERIV_cong] DERIV_odd_real_root[THEN DERIV_cong] DERIV_even_real_root[THEN DERIV_cong]) + subsection \Square Root\ -definition sqrt :: "real \ real" where - "sqrt = root 2" +definition sqrt :: "real \ real" + where "sqrt = root 2" -lemma pos2: "0 < (2::nat)" by simp +lemma pos2: "0 < (2::nat)" + by simp -lemma real_sqrt_unique: "\y\<^sup>2 = x; 0 \ y\ \ sqrt x = y" -unfolding sqrt_def by (rule real_root_pos_unique [OF pos2]) +lemma real_sqrt_unique: "y\<^sup>2 = x \ 0 \ y \ sqrt x = y" + unfolding sqrt_def by (rule real_root_pos_unique [OF pos2]) lemma real_sqrt_abs [simp]: "sqrt (x\<^sup>2) = \x\" -apply (rule real_sqrt_unique) -apply (rule power2_abs) -apply (rule abs_ge_zero) -done + apply (rule real_sqrt_unique) + apply (rule power2_abs) + apply (rule abs_ge_zero) + done lemma real_sqrt_pow2 [simp]: "0 \ x \ (sqrt x)\<^sup>2 = x" -unfolding sqrt_def by (rule real_root_pow_pos2 [OF pos2]) + unfolding sqrt_def by (rule real_root_pow_pos2 [OF pos2]) lemma real_sqrt_pow2_iff [simp]: "((sqrt x)\<^sup>2 = x) = (0 \ x)" -apply (rule iffI) -apply (erule subst) -apply (rule zero_le_power2) -apply (erule real_sqrt_pow2) -done + apply (rule iffI) + apply (erule subst) + apply (rule zero_le_power2) + apply (erule real_sqrt_pow2) + done lemma real_sqrt_zero [simp]: "sqrt 0 = 0" -unfolding sqrt_def by (rule real_root_zero) + unfolding sqrt_def by (rule real_root_zero) lemma real_sqrt_one [simp]: "sqrt 1 = 1" -unfolding sqrt_def by (rule real_root_one [OF pos2]) + unfolding sqrt_def by (rule real_root_one [OF pos2]) lemma real_sqrt_four [simp]: "sqrt 4 = 2" using real_sqrt_abs[of 2] by simp lemma real_sqrt_minus: "sqrt (- x) = - sqrt x" -unfolding sqrt_def by (rule real_root_minus) + unfolding sqrt_def by (rule real_root_minus) lemma real_sqrt_mult: "sqrt (x * y) = sqrt x * sqrt y" -unfolding sqrt_def by (rule real_root_mult) + unfolding sqrt_def by (rule real_root_mult) lemma real_sqrt_mult_self[simp]: "sqrt a * sqrt a = \a\" using real_sqrt_abs[of a] unfolding power2_eq_square real_sqrt_mult . lemma real_sqrt_inverse: "sqrt (inverse x) = inverse (sqrt x)" -unfolding sqrt_def by (rule real_root_inverse) + unfolding sqrt_def by (rule real_root_inverse) lemma real_sqrt_divide: "sqrt (x / y) = sqrt x / sqrt y" -unfolding sqrt_def by (rule real_root_divide) + unfolding sqrt_def by (rule real_root_divide) lemma real_sqrt_power: "sqrt (x ^ k) = sqrt x ^ k" -unfolding sqrt_def by (rule real_root_power [OF pos2]) + unfolding sqrt_def by (rule real_root_power [OF pos2]) lemma real_sqrt_gt_zero: "0 < x \ 0 < sqrt x" -unfolding sqrt_def by (rule real_root_gt_zero [OF pos2]) + unfolding sqrt_def by (rule real_root_gt_zero [OF pos2]) lemma real_sqrt_ge_zero: "0 \ x \ 0 \ sqrt x" -unfolding sqrt_def by (rule real_root_ge_zero) + unfolding sqrt_def by (rule real_root_ge_zero) lemma real_sqrt_less_mono: "x < y \ sqrt x < sqrt y" -unfolding sqrt_def by (rule real_root_less_mono [OF pos2]) + unfolding sqrt_def by (rule real_root_less_mono [OF pos2]) lemma real_sqrt_le_mono: "x \ y \ sqrt x \ sqrt y" -unfolding sqrt_def by (rule real_root_le_mono [OF pos2]) + unfolding sqrt_def by (rule real_root_le_mono [OF pos2]) lemma real_sqrt_less_iff [simp]: "(sqrt x < sqrt y) = (x < y)" -unfolding sqrt_def by (rule real_root_less_iff [OF pos2]) + unfolding sqrt_def by (rule real_root_less_iff [OF pos2]) lemma real_sqrt_le_iff [simp]: "(sqrt x \ sqrt y) = (x \ y)" -unfolding sqrt_def by (rule real_root_le_iff [OF pos2]) + unfolding sqrt_def by (rule real_root_le_iff [OF pos2]) lemma real_sqrt_eq_iff [simp]: "(sqrt x = sqrt y) = (x = y)" -unfolding sqrt_def by (rule real_root_eq_iff [OF pos2]) + unfolding sqrt_def by (rule real_root_eq_iff [OF pos2]) lemma real_less_lsqrt: "0 \ x \ 0 \ y \ x < y\<^sup>2 \ sqrt x < y" using real_sqrt_less_iff[of x "y\<^sup>2"] by simp @@ -417,7 +464,7 @@ lemma real_less_rsqrt: "x\<^sup>2 < y \ x < sqrt y" using real_sqrt_less_mono[of "x\<^sup>2" y] by simp -lemma sqrt_le_D: "sqrt x \ y \ x \ y^2" +lemma sqrt_le_D: "sqrt x \ y \ x \ y\<^sup>2" by (meson not_le real_less_rsqrt) lemma sqrt_even_pow2: @@ -446,73 +493,75 @@ lemma sqrt_add_le_add_sqrt: assumes "0 \ x" "0 \ y" shows "sqrt (x + y) \ sqrt x + sqrt y" -by (rule power2_le_imp_le) (simp_all add: power2_sum assms) + by (rule power2_le_imp_le) (simp_all add: power2_sum assms) lemma isCont_real_sqrt: "isCont sqrt x" -unfolding sqrt_def by (rule isCont_real_root) + unfolding sqrt_def by (rule isCont_real_root) -lemma tendsto_real_sqrt[tendsto_intros]: +lemma tendsto_real_sqrt [tendsto_intros]: "(f \ x) F \ ((\x. sqrt (f x)) \ sqrt x) F" unfolding sqrt_def by (rule tendsto_real_root) -lemma continuous_real_sqrt[continuous_intros]: +lemma continuous_real_sqrt [continuous_intros]: "continuous F f \ continuous F (\x. sqrt (f x))" unfolding sqrt_def by (rule continuous_real_root) -lemma continuous_on_real_sqrt[continuous_intros]: +lemma continuous_on_real_sqrt [continuous_intros]: "continuous_on s f \ continuous_on s (\x. sqrt (f x))" unfolding sqrt_def by (rule continuous_on_real_root) lemma DERIV_real_sqrt_generic: assumes "x \ 0" - assumes "x > 0 \ D = inverse (sqrt x) / 2" - assumes "x < 0 \ D = - inverse (sqrt x) / 2" + and "x > 0 \ D = inverse (sqrt x) / 2" + and "x < 0 \ D = - inverse (sqrt x) / 2" shows "DERIV sqrt x :> D" using assms unfolding sqrt_def by (auto intro!: DERIV_real_root_generic) -lemma DERIV_real_sqrt: - "0 < x \ DERIV sqrt x :> inverse (sqrt x) / 2" +lemma DERIV_real_sqrt: "0 < x \ DERIV sqrt x :> inverse (sqrt x) / 2" using DERIV_real_sqrt_generic by simp declare DERIV_real_sqrt_generic[THEN DERIV_chain2, derivative_intros] DERIV_real_root_generic[THEN DERIV_chain2, derivative_intros] -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 not_real_square_gt_zero [simp]: "\ 0 < x * x \ x = 0" for x :: real + 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_abs2 [simp]: "sqrt(x*x) = \x\" -apply (subst power2_eq_square [symmetric]) -apply (rule real_sqrt_abs) -done +lemma real_sqrt_abs2 [simp]: "sqrt (x * x) = \x\" + apply (subst power2_eq_square [symmetric]) + apply (rule real_sqrt_abs) + done -lemma real_inv_sqrt_pow2: "0 < x ==> (inverse (sqrt x))\<^sup>2 = inverse x" -by (simp add: power_inverse) +lemma real_inv_sqrt_pow2: "0 < x \ (inverse (sqrt x))\<^sup>2 = inverse x" + by (simp add: power_inverse) -lemma real_sqrt_eq_zero_cancel: "[| 0 \ x; sqrt(x) = 0|] ==> x = 0" -by simp +lemma real_sqrt_eq_zero_cancel: "0 \ x \ sqrt x = 0 \ x = 0" + by simp -lemma real_sqrt_ge_one: "1 \ x ==> 1 \ sqrt x" -by simp +lemma real_sqrt_ge_one: "1 \ x \ 1 \ sqrt x" + by simp lemma sqrt_divide_self_eq: assumes nneg: "0 \ x" shows "sqrt x / x = inverse (sqrt x)" -proof cases - assume "x=0" thus ?thesis by simp +proof (cases "x = 0") + case True + then show ?thesis by simp next - assume nz: "x\0" - hence pos: "0 0" by (simp add: divide_inverse nneg nz) + proof (rule right_inverse_eq [THEN iffD1, symmetric]) + show "sqrt x / x \ 0" + by (simp add: divide_inverse nneg False) show "inverse (sqrt x) / (sqrt x / x) = 1" by (simp add: divide_inverse mult.assoc [symmetric] - power2_eq_square [symmetric] real_inv_sqrt_pow2 pos nz) + power2_eq_square [symmetric] real_inv_sqrt_pow2 pos False) qed qed @@ -523,44 +572,44 @@ apply (simp add: field_simps) 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: ac_simps) -done +lemma real_divide_square_eq [simp]: "(r * a) / (r * r) = a / r" for a r :: real + apply (simp add: divide_inverse) + apply (case_tac "r = 0") + apply (auto simp add: ac_simps) + done -lemma lemma_real_divide_sqrt_less: "0 < u ==> u / sqrt 2 < u" -by (simp add: divide_less_eq) +lemma lemma_real_divide_sqrt_less: "0 < u \ u / sqrt 2 < u" + by (simp add: divide_less_eq) -lemma four_x_squared: - fixes x::real - shows "4 * x\<^sup>2 = (2 * x)\<^sup>2" -by (simp add: power2_eq_square) +lemma four_x_squared: "4 * x\<^sup>2 = (2 * x)\<^sup>2" for x :: real + by (simp add: power2_eq_square) lemma sqrt_at_top: "LIM x at_top. sqrt x :: real :> at_top" by (rule filterlim_at_top_at_top[where Q="\x. True" and P="\x. 0 < x" and g="power2"]) (auto intro: eventually_gt_at_top) + subsection \Square Root of Sum of Squares\ -lemma sum_squares_bound: - fixes x:: "'a::linordered_field" - shows "2*x*y \ x^2 + y^2" +lemma sum_squares_bound: "2 * x * y \ x\<^sup>2 + y\<^sup>2" for x y :: "'a::linordered_field" proof - - have "(x-y)^2 = x*x - 2*x*y + y*y" + have "(x - y)\<^sup>2 = x * x - 2 * x * y + y * y" by algebra - then have "0 \ x^2 - 2*x*y + y^2" + then have "0 \ x\<^sup>2 - 2 * x * y + y\<^sup>2" by (metis sum_power2_ge_zero zero_le_double_add_iff_zero_le_single_add power2_eq_square) then show ?thesis by arith qed lemma arith_geo_mean: - fixes u:: "'a::linordered_field" assumes "u\<^sup>2 = x*y" "x\0" "y\0" shows "u \ (x + y)/2" - apply (rule power2_le_imp_le) - using sum_squares_bound assms - apply (auto simp: zero_le_mult_iff) - by (auto simp: algebra_simps power2_eq_square) + fixes u :: "'a::linordered_field" + assumes "u\<^sup>2 = x * y" "x \ 0" "y \ 0" + shows "u \ (x + y)/2" + apply (rule power2_le_imp_le) + using sum_squares_bound assms + apply (auto simp: zero_le_mult_iff) + apply (auto simp: algebra_simps power2_eq_square) + done lemma arith_geo_mean_sqrt: fixes x::real assumes "x\0" "y\0" shows "sqrt(x*y) \ (x + y)/2" @@ -569,75 +618,77 @@ apply (auto simp: zero_le_mult_iff) done -lemma real_sqrt_sum_squares_mult_ge_zero [simp]: - "0 \ sqrt ((x\<^sup>2 + y\<^sup>2)*(xa\<^sup>2 + ya\<^sup>2))" +lemma real_sqrt_sum_squares_mult_ge_zero [simp]: "0 \ sqrt ((x\<^sup>2 + y\<^sup>2)*(xa\<^sup>2 + ya\<^sup>2))" by (metis real_sqrt_ge_0_iff split_mult_pos_le sum_power2_ge_zero) lemma real_sqrt_sum_squares_mult_squared_eq [simp]: - "(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)" + "(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)" by (simp add: zero_le_mult_iff) lemma real_sqrt_sum_squares_eq_cancel: "sqrt (x\<^sup>2 + y\<^sup>2) = x \ y = 0" -by (drule_tac f = "%x. x\<^sup>2" in arg_cong, simp) + by (drule arg_cong [where f = "\x. x\<^sup>2"]) simp lemma real_sqrt_sum_squares_eq_cancel2: "sqrt (x\<^sup>2 + y\<^sup>2) = y \ x = 0" -by (drule_tac f = "%x. x\<^sup>2" in arg_cong, simp) + by (drule arg_cong [where f = "\x. x\<^sup>2"]) simp lemma real_sqrt_sum_squares_ge1 [simp]: "x \ sqrt (x\<^sup>2 + y\<^sup>2)" -by (rule power2_le_imp_le, simp_all) + by (rule power2_le_imp_le) simp_all lemma real_sqrt_sum_squares_ge2 [simp]: "y \ sqrt (x\<^sup>2 + y\<^sup>2)" -by (rule power2_le_imp_le, simp_all) + by (rule power2_le_imp_le) simp_all lemma real_sqrt_ge_abs1 [simp]: "\x\ \ sqrt (x\<^sup>2 + y\<^sup>2)" -by (rule power2_le_imp_le, simp_all) + by (rule power2_le_imp_le) simp_all lemma real_sqrt_ge_abs2 [simp]: "\y\ \ sqrt (x\<^sup>2 + y\<^sup>2)" -by (rule power2_le_imp_le, simp_all) + by (rule power2_le_imp_le) simp_all lemma le_real_sqrt_sumsq [simp]: "x \ sqrt (x * x + y * y)" -by (simp add: power2_eq_square [symmetric]) + by (simp add: power2_eq_square [symmetric]) lemma real_sqrt_sum_squares_triangle_ineq: "sqrt ((a + c)\<^sup>2 + (b + d)\<^sup>2) \ sqrt (a\<^sup>2 + b\<^sup>2) + sqrt (c\<^sup>2 + d\<^sup>2)" -apply (rule power2_le_imp_le, simp) -apply (simp add: power2_sum) -apply (simp only: mult.assoc distrib_left [symmetric]) -apply (rule mult_left_mono) -apply (rule power2_le_imp_le) -apply (simp add: power2_sum power_mult_distrib) -apply (simp add: ring_distribs) -apply (subgoal_tac "0 \ b\<^sup>2 * c\<^sup>2 + a\<^sup>2 * d\<^sup>2 - 2 * (a * c) * (b * d)", simp) -apply (rule_tac b="(a * d - b * c)\<^sup>2" in ord_le_eq_trans) -apply (rule zero_le_power2) -apply (simp add: power2_diff power_mult_distrib) -apply (simp) -apply simp -apply (simp add: add_increasing) -done + apply (rule power2_le_imp_le) + apply simp + apply (simp add: power2_sum) + apply (simp only: mult.assoc distrib_left [symmetric]) + apply (rule mult_left_mono) + apply (rule power2_le_imp_le) + apply (simp add: power2_sum power_mult_distrib) + apply (simp add: ring_distribs) + apply (subgoal_tac "0 \ b\<^sup>2 * c\<^sup>2 + a\<^sup>2 * d\<^sup>2 - 2 * (a * c) * (b * d)") + apply simp + apply (rule_tac b="(a * d - b * c)\<^sup>2" in ord_le_eq_trans) + apply (rule zero_le_power2) + apply (simp add: power2_diff power_mult_distrib) + apply simp + apply simp + apply (simp add: add_increasing) + done -lemma real_sqrt_sum_squares_less: - "\\x\ < u / sqrt 2; \y\ < u / sqrt 2\ \ sqrt (x\<^sup>2 + y\<^sup>2) < u" -apply (rule power2_less_imp_less, simp) -apply (drule power_strict_mono [OF _ abs_ge_zero pos2]) -apply (drule power_strict_mono [OF _ abs_ge_zero pos2]) -apply (simp add: power_divide) -apply (drule order_le_less_trans [OF abs_ge_zero]) -apply (simp add: zero_less_divide_iff) -done +lemma real_sqrt_sum_squares_less: "\x\ < u / sqrt 2 \ \y\ < u / sqrt 2 \ sqrt (x\<^sup>2 + y\<^sup>2) < u" + apply (rule power2_less_imp_less) + apply simp + apply (drule power_strict_mono [OF _ abs_ge_zero pos2]) + apply (drule power_strict_mono [OF _ abs_ge_zero pos2]) + apply (simp add: power_divide) + apply (drule order_le_less_trans [OF abs_ge_zero]) + apply (simp add: zero_less_divide_iff) + done lemma sqrt2_less_2: "sqrt 2 < (2::real)" - by (metis not_less not_less_iff_gr_or_eq numeral_less_iff real_sqrt_four real_sqrt_le_iff semiring_norm(75) semiring_norm(78) semiring_norm(85)) + by (metis not_less not_less_iff_gr_or_eq numeral_less_iff real_sqrt_four + real_sqrt_le_iff semiring_norm(75) semiring_norm(78) semiring_norm(85)) -text\Needed for the infinitely close relation over the nonstandard - complex numbers\ +text \Needed for the infinitely close relation over the nonstandard complex numbers.\ lemma lemma_sqrt_hcomplex_capprox: - "[| 0 < u; x < u/2; y < u/2; 0 \ x; 0 \ y |] ==> sqrt (x\<^sup>2 + y\<^sup>2) < u" + "0 < u \ x < u/2 \ y < u/2 \ 0 \ x \ 0 \ y \ sqrt (x\<^sup>2 + y\<^sup>2) < u" apply (rule real_sqrt_sum_squares_less) apply (auto simp add: abs_if field_simps) apply (rule le_less_trans [where y = "x*2"]) - using less_eq_real_def sqrt2_less_2 apply force + using less_eq_real_def sqrt2_less_2 + apply force apply assumption apply (rule le_less_trans [where y = "y*2"]) using less_eq_real_def sqrt2_less_2 mult_le_cancel_left @@ -652,13 +703,15 @@ show "(\x. sqrt (2 / x)) \ sqrt 0" by (intro tendsto_intros tendsto_divide_0[OF tendsto_const] filterlim_mono[OF filterlim_real_sequentially]) (simp_all add: at_infinity_eq_at_top_bot) - { fix n :: nat assume "2 < n" - have "1 + (real (n - 1) * n) / 2 * x n^2 = 1 + of_nat (n choose 2) * x n^2" + have "x n \ sqrt (2 / real n)" if "2 < n" for n :: nat + proof - + have "1 + (real (n - 1) * n) / 2 * (x n)\<^sup>2 = 1 + of_nat (n choose 2) * (x n)\<^sup>2" by (auto simp add: choose_two of_nat_div mod_eq_0_iff_dvd) also have "\ \ (\k\{0, 2}. of_nat (n choose k) * x n^k)" by (simp add: x_def) also have "\ \ (\k=0..n. of_nat (n choose k) * x n^k)" - using \2 < n\ by (intro setsum_mono2) (auto intro!: mult_nonneg_nonneg zero_le_power simp: x_def le_diff_eq) + using \2 < n\ + by (intro setsum_mono2) (auto intro!: mult_nonneg_nonneg zero_le_power simp: x_def le_diff_eq) also have "\ = (x n + 1) ^ n" by (simp add: binomial_ring) also have "\ = n" @@ -667,8 +720,9 @@ by simp then have "(x n)\<^sup>2 \ 2 / real n" using \2 < n\ unfolding mult_le_cancel_left by (simp add: field_simps) - from real_sqrt_le_mono[OF this] have "x n \ sqrt (2 / real n)" - by simp } + from real_sqrt_le_mono[OF this] show ?thesis + by simp + qed then show "eventually (\n. x n \ sqrt (2 / real n)) sequentially" by (auto intro!: exI[of _ 3] simp: eventually_sequentially) show "eventually (\n. sqrt 0 \ x n) sequentially" @@ -682,47 +736,53 @@ assumes "0 < c" shows "(\n. root n c) \ 1" proof - - { fix c :: real assume "1 \ c" + have ge_1: "(\n. root n c) \ 1" if "1 \ c" for c :: real + proof - define x where "x n = root n c - 1" for n have "x \ 0" proof (rule tendsto_sandwich[OF _ _ tendsto_const]) show "(\n. c / n) \ 0" by (intro tendsto_divide_0[OF tendsto_const] filterlim_mono[OF filterlim_real_sequentially]) - (simp_all add: at_infinity_eq_at_top_bot) - { fix n :: nat assume "1 < n" + (simp_all add: at_infinity_eq_at_top_bot) + have "x n \ c / n" if "1 < n" for n :: nat + proof - have "1 + x n * n = 1 + of_nat (n choose 1) * x n^1" by (simp add: choose_one) also have "\ \ (\k\{0, 1}. of_nat (n choose k) * x n^k)" by (simp add: x_def) also have "\ \ (\k=0..n. of_nat (n choose k) * x n^k)" - using \1 < n\ \1 \ c\ by (intro setsum_mono2) (auto intro!: mult_nonneg_nonneg zero_le_power simp: x_def le_diff_eq) + using \1 < n\ \1 \ c\ + by (intro setsum_mono2) + (auto intro!: mult_nonneg_nonneg zero_le_power simp: x_def le_diff_eq) also have "\ = (x n + 1) ^ n" by (simp add: binomial_ring) also have "\ = c" using \1 < n\ \1 \ c\ by (simp add: x_def) - finally have "x n \ c / n" - using \1 \ c\ \1 < n\ by (simp add: field_simps) } + finally show ?thesis + using \1 \ c\ \1 < n\ by (simp add: field_simps) + qed then show "eventually (\n. x n \ c / n) sequentially" by (auto intro!: exI[of _ 3] simp: eventually_sequentially) show "eventually (\n. 0 \ x n) sequentially" - using \1 \ c\ by (auto intro!: exI[of _ 1] simp: eventually_sequentially le_diff_eq x_def) + using \1 \ c\ + by (auto intro!: exI[of _ 1] simp: eventually_sequentially le_diff_eq x_def) qed - from tendsto_add[OF this tendsto_const[of 1]] have "(\n. root n c) \ 1" - by (simp add: x_def) } - note ge_1 = this - + from tendsto_add[OF this tendsto_const[of 1]] show ?thesis + by (simp add: x_def) + qed show ?thesis - proof cases - assume "1 \ c" with ge_1 show ?thesis by blast + proof (cases "1 \ c") + case True + with ge_1 show ?thesis by blast next - assume "\ 1 \ c" + case False with \0 < c\ have "1 \ 1 / c" by simp then have "(\n. 1 / root n (1 / c)) \ 1 / 1" by (intro tendsto_divide tendsto_const ge_1 \1 \ 1 / c\ one_neq_zero) then show ?thesis by (rule filterlim_cong[THEN iffD1, rotated 3]) - (auto intro!: exI[of _ 1] simp: eventually_sequentially real_root_divide) + (auto intro!: exI[of _ 1] simp: eventually_sequentially real_root_divide) qed qed diff -r 2100fbbdc3f1 -r f3781c5fb03f src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Tue Jul 12 21:53:56 2016 +0200 +++ b/src/HOL/Transcendental.thy Tue Jul 12 22:54:37 2016 +0200 @@ -12,10 +12,10 @@ text \A fact theorem on reals.\ -lemma square_fact_le_2_fact: - shows "fact n * fact n \ (fact (2 * n) :: real)" +lemma square_fact_le_2_fact: "fact n * fact n \ (fact (2 * n) :: real)" proof (induct n) - case 0 then show ?case by simp + case 0 + then show ?case by simp next case (Suc n) have "(fact (Suc n)) * (fact (Suc n)) = of_nat (Suc n) * of_nat (Suc n) * (fact n * fact n :: real)" @@ -28,7 +28,6 @@ finally show ?case . qed - lemma fact_in_Reals: "fact n \ \" by (induction n) auto @@ -38,33 +37,33 @@ lemma pochhammer_of_real: "pochhammer (of_real x) n = of_real (pochhammer x n)" by (simp add: pochhammer_setprod) -lemma norm_fact [simp]: - "norm (fact n :: 'a :: {real_normed_algebra_1}) = fact n" +lemma norm_fact [simp]: "norm (fact n :: 'a::real_normed_algebra_1) = fact n" proof - - have "(fact n :: 'a) = of_real (fact n)" by simp - also have "norm \ = fact n" by (subst norm_of_real) simp + have "(fact n :: 'a) = of_real (fact n)" + by simp + also have "norm \ = fact n" + by (subst norm_of_real) simp finally show ?thesis . qed lemma root_test_convergence: fixes f :: "nat \ 'a::banach" assumes f: "(\n. root n (norm (f n))) \ x" \ "could be weakened to lim sup" - assumes "x < 1" + and "x < 1" shows "summable f" proof - have "0 \ x" by (rule LIMSEQ_le[OF tendsto_const f]) (auto intro!: exI[of _ 1]) from \x < 1\ obtain z where z: "x < z" "z < 1" by (metis dense) - from f \x < z\ - have "eventually (\n. root n (norm (f n)) < z) sequentially" + from f \x < z\ have "eventually (\n. root n (norm (f n)) < z) sequentially" by (rule order_tendstoD) then have "eventually (\n. norm (f n) \ z^n) sequentially" using eventually_ge_at_top proof eventually_elim - fix n assume less: "root n (norm (f n)) < z" and n: "1 \ n" - from power_strict_mono[OF less, of n] n - show "norm (f n) \ z ^ n" + fix n + assume less: "root n (norm (f n)) < z" and n: "1 \ n" + from power_strict_mono[OF less, of n] n show "norm (f n) \ z ^ n" by simp qed then show "summable f" @@ -72,30 +71,30 @@ using z \0 \ x\ by (auto intro!: summable_comparison_test[OF _ summable_geometric]) qed + subsection \Properties of Power Series\ -lemma powser_zero [simp]: - fixes f :: "nat \ 'a::real_normed_algebra_1" - shows "(\n. f n * 0 ^ n) = f 0" +lemma powser_zero [simp]: "(\n. f n * 0 ^ n) = f 0" + for f :: "nat \ 'a::real_normed_algebra_1" proof - have "(\n<1. f n * 0 ^ n) = (\n. f n * 0 ^ n)" by (subst suminf_finite[where N="{0}"]) (auto simp: power_0_left) - thus ?thesis unfolding One_nat_def by simp + then show ?thesis + by (simp add: One_nat_def) qed -lemma powser_sums_zero: - fixes a :: "nat \ 'a::real_normed_div_algebra" - shows "(\n. a n * 0^n) sums a 0" - using sums_finite [of "{0}" "\n. a n * 0 ^ n"] - by simp - -lemma powser_sums_zero_iff [simp]: - fixes a :: "nat \ 'a::real_normed_div_algebra" - shows "(\n. a n * 0^n) sums x \ a 0 = x" -using powser_sums_zero sums_unique2 by blast - -text\Power series has a circle or radius of convergence: if it sums for @{term - x}, then it sums absolutely for @{term z} with @{term "\z\ < \x\"}.\ +lemma powser_sums_zero: "(\n. a n * 0^n) sums a 0" + for a :: "nat \ 'a::real_normed_div_algebra" + using sums_finite [of "{0}" "\n. a n * 0 ^ n"] + by simp + +lemma powser_sums_zero_iff [simp]: "(\n. a n * 0^n) sums x \ a 0 = x" + for a :: "nat \ 'a::real_normed_div_algebra" + using powser_sums_zero sums_unique2 by blast + +text \ + Power series has a circle or radius of convergence: if it sums for \x\, + then it sums absolutely for \z\ with @{term "\z\ < \x\"}.\ lemma powser_insidea: fixes x z :: "'a::real_normed_div_algebra"