hoelzl@57025: theory NthRoot_Limits hoelzl@57025: imports Complex_Main "~~/src/HOL/Number_Theory/Binomial" hoelzl@57025: begin hoelzl@57025: hoelzl@57025: text {* hoelzl@57025: hoelzl@57025: This does not fit into @{text Complex_Main}, as it depends on @{text Binomial} hoelzl@57025: hoelzl@57025: *} hoelzl@57025: hoelzl@57025: lemma LIMSEQ_root: "(\n. root n n) ----> 1" hoelzl@57025: proof - hoelzl@57025: def x \ "\n. root n n - 1" hoelzl@57025: have "x ----> sqrt 0" hoelzl@57025: proof (rule tendsto_sandwich[OF _ _ tendsto_const]) hoelzl@57025: show "(\x. sqrt (2 / x)) ----> sqrt 0" hoelzl@57025: by (intro tendsto_intros tendsto_divide_0[OF tendsto_const] filterlim_mono[OF filterlim_real_sequentially]) hoelzl@57025: (simp_all add: at_infinity_eq_at_top_bot) hoelzl@57025: { fix n :: nat assume "2 < n" hoelzl@57025: have "1 + (real (n - 1) * n) / 2 * x n^2 = 1 + of_nat (n choose 2) * x n^2" hoelzl@57025: using `2 < n` unfolding gbinomial_def binomial_gbinomial hoelzl@57025: by (simp add: atLeast0AtMost atMost_Suc field_simps real_of_nat_diff numeral_2_eq_2 real_eq_of_nat[symmetric]) hoelzl@57025: also have "\ \ (\k\{0, 2}. of_nat (n choose k) * x n^k)" hoelzl@57025: by (simp add: x_def) hoelzl@57025: also have "\ \ (\k=0..n. of_nat (n choose k) * x n^k)" hoelzl@57025: using `2 < n` by (intro setsum_mono2) (auto intro!: mult_nonneg_nonneg zero_le_power simp: x_def le_diff_eq) hoelzl@57025: also have "\ = (x n + 1) ^ n" hoelzl@57025: by (simp add: binomial_ring) hoelzl@57025: also have "\ = n" hoelzl@57025: using `2 < n` by (simp add: x_def) hoelzl@57025: finally have "real (n - 1) * (real n / 2 * (x n)\<^sup>2) \ real (n - 1) * 1" hoelzl@57025: by simp hoelzl@57025: then have "(x n)\<^sup>2 \ 2 / real n" hoelzl@57025: using `2 < n` unfolding mult_le_cancel_left by (simp add: field_simps) hoelzl@57025: from real_sqrt_le_mono[OF this] have "x n \ sqrt (2 / real n)" hoelzl@57025: by simp } hoelzl@57025: then show "eventually (\n. x n \ sqrt (2 / real n)) sequentially" hoelzl@57025: by (auto intro!: exI[of _ 3] simp: eventually_sequentially) hoelzl@57025: show "eventually (\n. sqrt 0 \ x n) sequentially" hoelzl@57025: by (auto intro!: exI[of _ 1] simp: eventually_sequentially le_diff_eq x_def) hoelzl@57025: qed hoelzl@57025: from tendsto_add[OF this tendsto_const[of 1]] show ?thesis hoelzl@57025: by (simp add: x_def) hoelzl@57025: qed hoelzl@57025: hoelzl@57025: lemma LIMSEQ_root_const: hoelzl@57025: assumes "0 < c" hoelzl@57025: shows "(\n. root n c) ----> 1" hoelzl@57025: proof - hoelzl@57025: { fix c :: real assume "1 \ c" hoelzl@57025: def x \ "\n. root n c - 1" hoelzl@57025: have "x ----> 0" hoelzl@57025: proof (rule tendsto_sandwich[OF _ _ tendsto_const]) hoelzl@57025: show "(\n. c / n) ----> 0" hoelzl@57025: by (intro tendsto_divide_0[OF tendsto_const] filterlim_mono[OF filterlim_real_sequentially]) hoelzl@57025: (simp_all add: at_infinity_eq_at_top_bot) hoelzl@57025: { fix n :: nat assume "1 < n" hoelzl@57025: have "1 + x n * n = 1 + of_nat (n choose 1) * x n^1" hoelzl@57025: using `1 < n` unfolding gbinomial_def binomial_gbinomial by (simp add: real_eq_of_nat[symmetric]) hoelzl@57025: also have "\ \ (\k\{0, 1}. of_nat (n choose k) * x n^k)" hoelzl@57025: by (simp add: x_def) hoelzl@57025: also have "\ \ (\k=0..n. of_nat (n choose k) * x n^k)" hoelzl@57025: using `1 < n` `1 \ c` by (intro setsum_mono2) (auto intro!: mult_nonneg_nonneg zero_le_power simp: x_def le_diff_eq) hoelzl@57025: also have "\ = (x n + 1) ^ n" hoelzl@57025: by (simp add: binomial_ring) hoelzl@57025: also have "\ = c" hoelzl@57025: using `1 < n` `1 \ c` by (simp add: x_def) hoelzl@57025: finally have "x n \ c / n" hoelzl@57025: using `1 \ c` `1 < n` by (simp add: field_simps) } hoelzl@57025: then show "eventually (\n. x n \ c / n) sequentially" hoelzl@57025: by (auto intro!: exI[of _ 3] simp: eventually_sequentially) hoelzl@57025: show "eventually (\n. 0 \ x n) sequentially" hoelzl@57025: using `1 \ c` by (auto intro!: exI[of _ 1] simp: eventually_sequentially le_diff_eq x_def) hoelzl@57025: qed hoelzl@57025: from tendsto_add[OF this tendsto_const[of 1]] have "(\n. root n c) ----> 1" hoelzl@57025: by (simp add: x_def) } hoelzl@57025: note ge_1 = this hoelzl@57025: hoelzl@57025: show ?thesis hoelzl@57025: proof cases hoelzl@57025: assume "1 \ c" with ge_1 show ?thesis by blast hoelzl@57025: next hoelzl@57025: assume "\ 1 \ c" hoelzl@57025: with `0 < c` have "1 \ 1 / c" hoelzl@57025: by simp hoelzl@57025: then have "(\n. 1 / root n (1 / c)) ----> 1 / 1" hoelzl@57025: by (intro tendsto_divide tendsto_const ge_1 `1 \ 1 / c` one_neq_zero) hoelzl@57025: then show ?thesis hoelzl@57025: by (rule filterlim_cong[THEN iffD1, rotated 3]) hoelzl@57025: (auto intro!: exI[of _ 1] simp: eventually_sequentially real_root_divide) hoelzl@57025: qed hoelzl@57025: qed hoelzl@57025: hoelzl@57025: end