diff --git a/thys/Continued_Fractions/Continued_Fraction_Approximation.thy b/thys/Continued_Fractions/Continued_Fraction_Approximation.thy new file mode 100644 --- /dev/null +++ b/thys/Continued_Fractions/Continued_Fraction_Approximation.thy @@ -0,0 +1,190 @@ +(* + File: Continued_Fractions/Continued_Fraction_Approximation.thy + Author: Manuel Eberl, University of Innsbruck +*) +section \Computing continued fraction expansions through interval arithmetic\ +theory Continued_Fraction_Approximation +imports + Complex_Main + "HOL-Decision_Procs.Approximation" + "Coinductive.Coinductive_List" + "HOL-Library.Code_Lazy" + "HOL-Library.Code_Target_Numeral" + Continued_Fractions +keywords "approximate_cfrac" :: diag +begin + +text \ + The approximation package allows us to compute an enclosing interval for a given + real constant. From this, we are able to compute an initial fragment of the continued + fraction expansion of the number. + + The algorithm essentially works by computing the continued fraction expansion of the lower + and upper bound simultaneously and stopping when the results start to diverge. + + This algorithm terminates because the lower and upper bounds, being rational numbers, have + a finite continued fraction expansion. +\ + +definition float_to_rat :: "float \ int \ int" where + "float_to_rat f = (if exponent f \ 0 then + (mantissa f * 2 ^ nat (exponent f), 1) else (mantissa f, 2 ^ nat (-exponent f)))" + +lemma float_to_rat: "fst (float_to_rat f) / snd (float_to_rat f) = real_of_float f" + by (auto simp: float_to_rat_def mantissa_exponent powr_int) + +lemma snd_float_to_rat_pos [simp]: "snd (float_to_rat f) > 0" + by (simp add: float_to_rat_def) + + +function cfrac_from_approx :: "int \ int \ int \ int \ int list" where + "cfrac_from_approx (nl, dl) (nu, du) = + (if nl = 0 \ nu = 0 \ dl = 0 \ du = 0 then [] + else let l = nl div dl; u = nu div du + in if l \ u then [] + else l # (let m = nl mod dl in if m = 0 then [] else + cfrac_from_approx (du, nu mod du) (dl, m)))" + by auto +termination proof (relation "measure (\((nl, dl), (nu, du)). nat (abs dl + abs du))", goal_cases) + case (2 nl dl nu du) + hence "\nl mod dl\ + \nu mod du\ < \dl\ + \du\" + by (intro add_strict_mono) (auto simp: abs_mod_less) + thus ?case using 2 by simp +qed auto + +lemmas [simp del] = cfrac_from_approx.simps + +lemma cfrac_from_approx_correct: + assumes "x \ {fst l / snd l..fst u / snd u}" and "snd l > 0" and "snd u > 0" + assumes "i < length (cfrac_from_approx l u)" + shows "cfrac_nth (cfrac_of_real x) i = cfrac_from_approx l u ! i" + using assms +proof (induction l u arbitrary: i x rule: cfrac_from_approx.induct) + case (1 nl dl nu du i x) + from "1.prems" have *: "nl div dl = nu div du" "nl \ 0" "nu \ 0" "dl > 0" "du > 0" + by (auto simp: cfrac_from_approx.simps Let_def split: if_splits) + have "\nl / dl\ \ \x\" "\x\ \ \nu / du\" + using "1.prems"(1) by (intro floor_mono; simp)+ + hence "nl div dl \ \x\" "\x\ \ nu div du" + by (simp_all add: floor_divide_of_int_eq) + with * have "\x\ = nu div du" + by linarith + + show ?case + proof (cases i) + case 0 + with 0 and \\x\ = _\ show ?thesis using "1.prems" + by (auto simp: Let_def cfrac_from_approx.simps) + next + case [simp]: (Suc i') + from "1.prems" * have "nl mod dl \ 0" + by (subst (asm) cfrac_from_approx.simps) (auto split: if_splits) + have frac_eq: "frac x = x - nu div du" + using \\x\ = _\ by (simp add: frac_def) + + have "frac x \ nl / dl - nl div dl" + using * "1.prems" by (simp add: frac_eq) + also have "nl / dl - nl div dl = (nl - dl * (nl div dl)) / dl" + using * by (simp add: field_simps) + also have "nl - dl * (nl div dl) = nl mod dl" + by (subst minus_div_mult_eq_mod [symmetric]) auto + finally have "frac x \ (nl mod dl) / dl" . + + have "nl mod dl \ 0" + using * by (intro pos_mod_sign) auto + with \nl mod dl \ 0\ have "nl mod dl > 0" + by linarith + hence "0 < (nl mod dl) / dl" + using * by (intro divide_pos_pos) auto + also have "\ \ frac x" + by fact + finally have "frac x > 0" . + + have "frac x \ nu / du - nu div du" + using * "1.prems" by (simp add: frac_eq) + also have "\ = (nu - du * (nu div du)) / du" + using * by (simp add: field_simps) + also have "nu - du * (nu div du) = nu mod du" + by (subst minus_div_mult_eq_mod [symmetric]) auto + finally have "frac x \ real_of_int (nu mod du) / real_of_int du" . + + have "0 < frac x" + by fact + also have "\ \ (nu mod du) / du" + by fact + finally have "nu mod du > 0" + using * by (auto simp: field_simps) + + have "cfrac_nth (cfrac_of_real x) i = cfrac_nth (cfrac_tl (cfrac_of_real x)) i'" + by simp + also have "cfrac_tl (cfrac_of_real x) = cfrac_of_real (1 / frac x)" + using \frac x > 0\ by (intro cfrac_tl_of_real) auto + also have "cfrac_nth (cfrac_of_real (1 / frac x)) i' = + cfrac_from_approx (du, nu mod du) (dl, nl mod dl) ! i'" + proof (rule "1.IH"[OF _ refl refl _ refl]) + show "\ (nl = 0 \ nu = 0 \ dl = 0 \ du = 0)" "\ nl div dl \ nu div du" + using "1.prems" by (auto split: if_splits simp: Let_def cfrac_from_approx.simps) + next + show "i' < length (cfrac_from_approx (du, nu mod du) (dl, nl mod dl))" using "1.prems" + by (subst (asm) cfrac_from_approx.simps) (auto split: if_splits simp: Let_def) + next + have "1 / frac x \ dl / (nl mod dl)" + using \frac x > 0\ and \nl mod dl > 0\ and \frac x \ (nl mod dl) / dl\ and * + by (auto simp: field_simps) + moreover have "1 / frac x \ du / (nu mod du)" + using \frac x > 0\ and \nu mod du > 0\ and \frac x \ (nu mod du) / du\ and * + by (auto simp: field_simps) + ultimately show + "1 / frac x \ {real_of_int (fst (du, nu mod du)) / real_of_int (snd (du, nu mod du)).. + real_of_int (fst (dl, nl mod dl)) / real_of_int (snd (dl, nl mod dl))}" + by simp + show "snd (du, nu mod du) > 0" "snd (dl, nl mod dl) > 0" and "nl mod dl \ 0" + using \nu mod du > 0\ and \nl mod dl > 0\ by simp_all + qed + also have "cfrac_from_approx (du, nu mod du) (dl, nl mod dl) ! i' = + cfrac_from_approx (nl, dl) (nu, du) ! i" + using "1.prems" * \nl mod dl \ 0\ by (subst (2) cfrac_from_approx.simps) auto + finally show ?thesis . + qed +qed + +definition cfrac_from_approx' :: "float \ float \ int list" where + "cfrac_from_approx' l u = cfrac_from_approx (float_to_rat l) (float_to_rat u)" + +lemma cfrac_from_approx'_correct: + assumes "x \ {real_of_float l..real_of_float u}" + assumes "i < length (cfrac_from_approx' l u)" + shows "cfrac_nth (cfrac_of_real x) i = cfrac_from_approx' l u ! i" + using assms unfolding cfrac_from_approx'_def + by (intro cfrac_from_approx_correct) (auto simp: float_to_rat cfrac_from_approx'_def) + +definition approx_cfrac :: "nat \ floatarith \ int list" where + "approx_cfrac prec e = + (case approx' prec e [] of + None \ [] + | Some ivl \ cfrac_from_approx' (lower ivl) (upper ivl))" + +ML_file \approximation_cfrac.ML\ + + +text \Now let us do some experiments:\ + +value "let prec = 34; c = cfrac_from_approx' (lb_pi prec) (ub_pi prec) in c" +value "let prec = 34; c = cfrac_from_approx' (lb_pi prec) (ub_pi prec) + in map (\n. (conv_num_fun ((!) c) n, conv_denom_fun ((!) c) n)) [0..Continued Fractions\ +theory Continued_Fractions +imports + Complex_Main + "Coinductive.Lazy_LList" + "Coinductive.Coinductive_Nat" + "HOL-Number_Theory.Fib" + "HOL-Library.BNF_Corec" + "Coinductive.Coinductive_Stream" +begin + +subsection \Auxiliary results\ + +(* TODO: Move? *) +coinductive linfinite :: "'a llist \ bool" where + "linfinite xs \ linfinite (LCons x xs)" + +lemma llength_llist_of_stream [simp]: "llength (llist_of_stream xs) = \" + by (simp add: not_lfinite_llength) + +lemma linfinite_conv_llength: "linfinite xs \ llength xs = \" +proof + assume "linfinite xs" + thus "llength xs = \" + proof (coinduction arbitrary: xs rule: enat_coinduct2) + fix xs :: "'a llist" + assume "llength xs \ 0" "linfinite xs" + thus "(\xs'::'a llist. epred (llength xs) = llength xs' \ epred \ = \ \ linfinite xs') \ + epred (llength xs) = epred \" + by (intro disjI1 exI[of _ "ltl xs"]) (auto simp: linfinite.simps[of xs]) + next + fix xs :: "'a llist" assume "linfinite xs"thus "(llength xs = 0) \ (\ = (0::enat))" + by (subst (asm) linfinite.simps) auto + qed +next + assume "llength xs = \" + thus "linfinite xs" + proof (coinduction arbitrary: xs) + case linfinite + thus "\xsa x. + xs = LCons x xsa \ + ((\xs. xsa = xs \ llength xs = \) \ + linfinite xsa)" + by (cases xs) (auto simp: eSuc_eq_infinity_iff) + qed +qed + +definition lnth_default :: "'a \ 'a llist \ nat \ 'a" where + "lnth_default dflt xs n = (if n < llength xs then lnth xs n else dflt)" + +lemma lnth_default_code [code]: + "lnth_default dflt xs n = + (if lnull xs then dflt else if n = 0 then lhd xs else lnth_default dflt (ltl xs) (n - 1))" +proof (induction n arbitrary: xs) + case 0 + thus ?case + by (cases xs) (auto simp: lnth_default_def simp flip: zero_enat_def) +next + case (Suc n) + show ?case + proof (cases xs) + case LNil + thus ?thesis + by (auto simp: lnth_default_def) + next + case (LCons x xs') + thus ?thesis + by (auto simp: lnth_default_def Suc_ile_eq) + qed +qed + +lemma enat_le_iff: + "enat n \ m \ m = \ \ (\m'. m = enat m' \ n \ m')" + by (cases m) auto + +lemma enat_less_iff: + "enat n < m \ m = \ \ (\m'. m = enat m' \ n < m')" + by (cases m) auto + +lemma real_of_int_divide_in_Ints_iff: + "real_of_int a / real_of_int b \ \ \ b dvd a \ b = 0" +proof safe + assume "real_of_int a / real_of_int b \ \" "b \ 0" + then obtain n where "real_of_int a / real_of_int b = real_of_int n" + by (auto simp: Ints_def) + hence "real_of_int b * real_of_int n = real_of_int a" + using \b \ 0\ by (auto simp: field_simps) + also have "real_of_int b * real_of_int n = real_of_int (b * n)" + by simp + finally have "b * n = a" + by linarith + thus "b dvd a" + by auto +qed auto + +lemma frac_add_of_nat: "frac (of_nat y + x) = frac x" + unfolding frac_def by simp + +lemma frac_add_of_int: "frac (of_int y + x) = frac x" + unfolding frac_def by simp + +lemma frac_fraction: "frac (real_of_int a / real_of_int b) = (a mod b) / b" +proof - + have "frac (a / b) = frac ((a mod b + b * (a div b)) / b)" + by (subst mod_mult_div_eq) auto + also have "(a mod b + b * (a div b)) / b = of_int (a div b) + a mod b / b" + unfolding of_int_add by (subst add_divide_distrib) auto + also have "frac \ = frac (a mod b / b)" + by (rule frac_add_of_int) + also have "\ = a mod b / b" + by (simp add: floor_divide_of_int_eq frac_def) + finally show ?thesis . +qed + +lemma Suc_fib_ge: "Suc (fib n) \ n" +proof (induction n rule: fib.induct) + case (3 n) + show ?case + proof (cases "n < 2") + case True + thus ?thesis by (cases n) auto + next + case False + hence "Suc (Suc (Suc n)) \ Suc n + n" by simp + also have "\ \ Suc (fib (Suc n)) + Suc (fib n)" + by (intro add_mono 3) + also have "\ = Suc (Suc (fib (Suc (Suc n))))" + by simp + finally show ?thesis by (simp only: Suc_le_eq) + qed +qed auto + +lemma fib_ge: "fib n \ n - 1" + using Suc_fib_ge[of n] by simp + +lemma frac_diff_of_nat_right [simp]: "frac (x - of_nat y) = frac x" + using floor_diff_of_int[of x "int y"] by (simp add: frac_def) + +lemma funpow_cycle: + assumes "m > 0" + assumes "(f ^^ m) x = x" + shows "(f ^^ k) x = (f ^^ (k mod m)) x" +proof (induction k rule: less_induct) + case (less k) + show ?case + proof (cases "k < m") + case True + thus ?thesis using \m > 0\ by simp + next + case False + hence "k = (k - m) + m" by simp + also have "(f ^^ \) x = (f ^^ (k - m)) ((f ^^ m) x)" + by (simp add: funpow_add) + also have "(f ^^ m) x = x" by fact + also have "(f ^^ (k - m)) x = (f ^^ (k mod m)) x" + using assms False by (subst less.IH) (auto simp: mod_geq) + finally show ?thesis . + qed +qed + +lemma of_nat_ge_1_iff: "of_nat n \ (1 :: 'a :: linordered_semidom) \ n > 0" + using of_nat_le_iff[of 1 n] unfolding of_nat_1 by auto + +lemma not_frac_less_0: "\frac x < 0" + by (simp add: frac_def not_less) + +lemma frac_le_1: "frac x \ 1" + unfolding frac_def by linarith + +lemma divide_in_Rats_iff1: + "(x::real) \ \ \ x \ 0 \ x / y \ \ \ y \ \" +proof safe + assume *: "x \ \" "x \ 0" "x / y \ \" + from *(1,3) have "x / (x / y) \ \" + by (rule Rats_divide) + also from * have "x / (x / y) = y" by simp + finally show "y \ \" . +qed (auto intro: Rats_divide) + +lemma divide_in_Rats_iff2: + "(y::real) \ \ \ y \ 0 \ x / y \ \ \ x \ \" +proof safe + assume *: "y \ \" "y \ 0" "x / y \ \" + from *(3,1) have "x / y * y \ \" + by (rule Rats_mult) + also from * have "x / y * y = x" by simp + finally show "x \ \" . +qed (auto intro: Rats_divide) + +lemma add_in_Rats_iff1: "x \ \ \ x + y \ \ \ y \ \" + using Rats_diff[of "x + y" x] by auto + +lemma add_in_Rats_iff2: "y \ \ \ x + y \ \ \ x \ \" + using Rats_diff[of "x + y" y] by auto + +lemma diff_in_Rats_iff1: "x \ \ \ x - y \ \ \ y \ \" + using Rats_diff[of x "x - y"] by auto + +lemma diff_in_Rats_iff2: "y \ \ \ x - y \ \ \ x \ \" + using Rats_add[of "x - y" y] by auto + +lemma frac_in_Rats_iff [simp]: "frac x \ \ \ x \ \" + by (simp add: frac_def diff_in_Rats_iff2) + +lemma filterlim_sequentially_shift: + "filterlim (\n. f (n + m)) F sequentially \ filterlim f F sequentially" +proof (induction m) + case (Suc m) + have "filterlim (\n. f (n + Suc m)) F at_top \ + filterlim (\n. f (Suc n + m)) F at_top" by simp + also have "\ \ filterlim (\n. f (n + m)) F at_top" + by (rule filterlim_sequentially_Suc) + also have "\ \ filterlim f F at_top" + by (rule Suc.IH) + finally show ?case . +qed simp_all + + +subsection \Bounds on alternating decreasing sums\ + +lemma alternating_decreasing_sum_bounds: + fixes f :: "nat \ 'a :: {linordered_ring, ring_1}" + assumes "m \ n" "\k. k \ {m..n} \ f k \ 0" + "\k. k \ {m.. f (Suc k) \ f k" + defines "S \ (\m. (\k=m..n. (-1) ^ k * f k))" + shows "if even m then S m \ {0..f m} else S m \ {-f m..0}" + using assms(1) +proof (induction rule: inc_induct) + case (step m') + have [simp]: "-a \ b \ a + b \ (0 :: 'a)" for a b + by (metis le_add_same_cancel1 minus_add_cancel) + have [simp]: "S m' = (-1) ^ m' * f m' + S (Suc m')" + using step.hyps unfolding S_def + by (subst sum.atLeast_Suc_atMost) simp_all + from step.hyps have nonneg: "f m' \ 0" + by (intro assms) auto + from step.hyps have mono: "f (Suc m') \ f m'" + by (intro assms) auto + show ?case + proof (cases "even m'") + case True + hence "0 \ f (Suc m') + S (Suc m')" + using step.IH by simp + also note mono + finally show ?thesis using True step.IH by auto + next + case False + with step.IH have "S (Suc m') \ f (Suc m')" + by simp + also note mono + finally show ?thesis using step.IH False by auto + qed +qed (insert assms, auto) + +lemma alternating_decreasing_sum_bounds': + fixes f :: "nat \ 'a :: {linordered_ring, ring_1}" + assumes "m < n" "\k. k \ {m..n-1} \ f k \ 0" + "\k. k \ {m.. f (Suc k) \ f k" + defines "S \ (\m. (\k=m.. {0..f m} else S m \ {-f m..0}" +proof (cases n) + case 0 + thus ?thesis using assms by auto +next + case (Suc n') + hence "if even m then (\k=m..n-1. (-1) ^ k * f k) \ {0..f m} + else (\k=m..n-1. (-1) ^ k * f k) \ {-f m..0}" + using assms by (intro alternating_decreasing_sum_bounds) auto + also have "(\k=m..n-1. (-1) ^ k * f k) = S m" + unfolding S_def by (intro sum.cong) (auto simp: Suc) + finally show ?thesis . +qed + +lemma alternating_decreasing_sum_upper_bound: + fixes f :: "nat \ 'a :: {linordered_ring, ring_1}" + assumes "m \ n" "\k. k \ {m..n} \ f k \ 0" + "\k. k \ {m.. f (Suc k) \ f k" + shows "(\k=m..n. (-1) ^ k * f k) \ f m" + using alternating_decreasing_sum_bounds[of m n f, OF assms] assms(1) + by (auto split: if_splits intro: order.trans[OF _ assms(2)]) + +lemma alternating_decreasing_sum_upper_bound': + fixes f :: "nat \ 'a :: {linordered_ring, ring_1}" + assumes "m < n" "\k. k \ {m..n-1} \ f k \ 0" + "\k. k \ {m.. f (Suc k) \ f k" + shows "(\k=m.. f m" + using alternating_decreasing_sum_bounds'[of m n f, OF assms] assms(1) + by (auto split: if_splits intro: order.trans[OF _ assms(2)]) + +lemma abs_alternating_decreasing_sum_upper_bound: + fixes f :: "nat \ 'a :: {linordered_ring, ring_1}" + assumes "m \ n" "\k. k \ {m..n} \ f k \ 0" + "\k. k \ {m.. f (Suc k) \ f k" + shows "\(\k=m..n. (-1) ^ k * f k)\ \ f m" (is "abs ?S \ _") + using alternating_decreasing_sum_bounds[of m n f, OF assms] + by (auto split: if_splits simp: minus_le_iff) + +lemma abs_alternating_decreasing_sum_upper_bound': + fixes f :: "nat \ 'a :: {linordered_ring, ring_1}" + assumes "m < n" "\k. k \ {m..n-1} \ f k \ 0" + "\k. k \ {m.. f (Suc k) \ f k" + shows "\(\k=m.. \ f m" + using alternating_decreasing_sum_bounds'[of m n f, OF assms] + by (auto split: if_splits simp: minus_le_iff) + +lemma abs_alternating_decreasing_sum_lower_bound: + fixes f :: "nat \ 'a :: {linordered_ring, ring_1}" + assumes "m < n" "\k. k \ {m..n} \ f k \ 0" + "\k. k \ {m.. f (Suc k) \ f k" + shows "\(\k=m..n. (-1) ^ k * f k)\ \ f m - f (Suc m)" +proof - + have "(\k=m..n. (-1) ^ k * f k) = (\k\insert m {m<..n}. (-1) ^ k * f k)" + using assms by (intro sum.cong) auto + also have "\ = (-1) ^ m * f m + (\k\{m<..n}. (-1) ^ k * f k)" + by auto + also have "(\k\{m<..n}. (-1) ^ k * f k) = (\k\{m..i. i - 1"]) auto + also have "(-1)^m * f m + \ = (-1)^m * f m - (\k\{m..\\ \ \(-1)^m * f m\ - \(\k\{m.." + by (rule abs_triangle_ineq2) + also have "\(-1)^m * f m\ = f m" + using assms by (cases "even m") auto + finally have "f m - \\k = m.. + \ \\k = m..n. (- 1) ^ k * f k\" . + moreover have "f m - \(\k\{m.. \ f m - f (Suc m)" + using assms by (intro diff_mono abs_alternating_decreasing_sum_upper_bound') auto + ultimately show ?thesis by (rule order.trans[rotated]) +qed + +lemma abs_alternating_decreasing_sum_lower_bound': + fixes f :: "nat \ 'a :: {linordered_ring, ring_1}" + assumes "m+1 < n" "\k. k \ {m..n} \ f k \ 0" + "\k. k \ {m.. f (Suc k) \ f k" + shows "\(\k=m.. \ f m - f (Suc m)" +proof (cases n) + case 0 + thus ?thesis using assms by auto +next + case (Suc n') + hence "\(\k=m..n-1. (-1) ^ k * f k)\ \ f m - f (Suc m)" + using assms by (intro abs_alternating_decreasing_sum_lower_bound) auto + also have "(\k=m..n-1. (-1) ^ k * f k) = (\k=m..k. f k \ (0 :: real)" "\k. f (Suc k) \ f k" + "f \ 0" + shows "(\k. (-1) ^ k * f k) \ {f 0 - f 1..f 0}" +proof - + have "summable (\k. (-1) ^ k * f k)" + by (intro summable_Leibniz' assms) + hence lim: "(\n. \k\n. (-1) ^ k * f k) \ (\k. (-1) ^ k * f k)" + by (auto dest: summable_LIMSEQ') + have bounds: "(\k=0..n. (- 1) ^ k * f k) \ {f 0 - f 1..f 0}" + if "n > 0" for n + using alternating_decreasing_sum_bounds[of 1 n f] assms that + by (subst sum.atLeast_Suc_atMost) auto + note [simp] = atLeast0AtMost + note [intro!] = eventually_mono[OF eventually_gt_at_top[of 0]] + + from lim have "(\k. (-1) ^ k * f k) \ f 0 - f 1" + by (rule tendsto_lowerbound) (insert bounds, auto) + moreover from lim have "(\k. (-1) ^ k * f k) \ f 0" + by (rule tendsto_upperbound) (use bounds in auto) + ultimately show ?thesis by simp +qed + +lemma + assumes "\k. k \ m \ f k \ (0 :: real)" + "\k. k \ m \ f (Suc k) \ f k" "f \ 0" + defines "S \ (\k. (-1) ^ (k + m) * f (k + m))" + shows summable_alternating_decreasing: "summable (\k. (-1) ^ (k + m) * f (k + m))" + and alternating_decreasing_suminf_bounds': + "if even m then S \ {f m - f (Suc m) .. f m} + else S \ {-f m..f (Suc m) - f m}" (is ?th1) + and abs_alternating_decreasing_suminf: + "abs S \ {f m - f (Suc m)..f m}" (is ?th2) +proof - + have summable: "summable (\k. (-1) ^ k * f (k + m))" + using assms by (intro summable_Leibniz') (auto simp: filterlim_sequentially_shift) + thus "summable (\k. (-1) ^ (k + m) * f (k + m))" + by (subst add.commute) (auto simp: power_add mult.assoc intro: summable_mult) + have "S = (\k. (-1) ^ m * ((-1) ^ k * f (k + m)))" + by (simp add: S_def power_add mult_ac) + also have "\ = (-1) ^ m * (\k. (-1) ^ k * f (k + m))" + using summable by (rule suminf_mult) + finally have "S = (- 1) ^ m * (\k. (- 1) ^ k * f (k + m))" . + moreover have "(\k. (-1) ^ k * f (k + m)) \ + {f (0 + m) - f (1 + m) .. f (0 + m)}" + using assms + by (intro alternating_decreasing_suminf_bounds) + (auto simp: filterlim_sequentially_shift) + ultimately show ?th1 by (auto split: if_splits) + thus ?th2 using assms(2)[of m] by (auto split: if_splits) +qed + +lemma + assumes "\k. k \ m \ f k \ (0 :: real)" + "\k. k \ m \ f (Suc k) < f k" "f \ 0" + defines "S \ (\k. (-1) ^ (k + m) * f (k + m))" + shows alternating_decreasing_suminf_bounds_strict': + "if even m then S \ {f m - f (Suc m)<.. {-f m<.. {f m - f (Suc m)<..k. (-1) ^ (k + Suc (Suc m)) * f (k + Suc (Suc m)))" + have "(\k. (-1) ^ (k + m) * f (k + m)) sums S" using assms unfolding S_def + by (intro summable_sums summable_Leibniz' summable_alternating_decreasing) + (auto simp: less_eq_real_def) + from sums_split_initial_segment[OF this, of 2] + have S': "S' = S - (-1) ^ m * (f m - f (Suc m))" + by (simp_all add: sums_iff S'_def algebra_simps lessThan_nat_numeral) + have "if even (Suc (Suc m)) then S' \ {f (Suc (Suc m)) - f (Suc (Suc (Suc m)))..f (Suc (Suc m))} + else S' \ {- f (Suc (Suc m))..f (Suc (Suc (Suc m))) - f (Suc (Suc m))}" unfolding S'_def + using assms by (intro alternating_decreasing_suminf_bounds') (auto simp: less_eq_real_def) + thus ?th1 using assms(2)[of "Suc m"] assms(2)[of "Suc (Suc m)"] + unfolding S' by (auto simp: algebra_simps) + thus ?th2 using assms(2)[of m] by (auto split: if_splits) +qed + + +datatype cfrac = CFrac int "nat llist" + +quickcheck_generator cfrac constructors: CFrac + +lemma type_definition_cfrac': + "type_definition (\x. case x of CFrac a b \ (a, b)) (\(x,y). CFrac x y) UNIV" + by (auto simp: type_definition_def split: cfrac.splits) + +setup_lifting type_definition_cfrac' + +lift_definition cfrac_of_int :: "int \ cfrac" is + "\n. (n, LNil)" . + +lemma cfrac_of_int_code [code]: "cfrac_of_int n = CFrac n LNil" + by (auto simp: cfrac_of_int_def) + +lift_definition cfrac_of_stream :: "int stream \ cfrac" is + "\xs. (shd xs, llist_of_stream (smap (\x. nat (x - 1)) (stl xs)))" . + +instantiation cfrac :: zero +begin +definition zero_cfrac where "0 = cfrac_of_int 0" +instance .. +end + +instantiation cfrac :: one +begin +definition one_cfrac where "1 = cfrac_of_int 1" +instance .. +end + +lift_definition cfrac_tl :: "cfrac \ cfrac" is + "\(_, bs) \ case bs of LNil \ (1, LNil) | LCons b bs' \ (int b + 1, bs')" . + +lemma cfrac_tl_code [code]: + "cfrac_tl (CFrac a bs) = + (case bs of LNil \ CFrac 1 LNil | LCons b bs' \ CFrac (int b + 1) bs')" + by (auto simp: cfrac_tl_def split: llist.splits) + +definition cfrac_drop :: "nat \ cfrac \ cfrac" where + "cfrac_drop n c = (cfrac_tl ^^ n) c" + +lemma cfrac_drop_Suc_right: "cfrac_drop (Suc n) c = cfrac_drop n (cfrac_tl c)" + by (simp add: cfrac_drop_def funpow_Suc_right del: funpow.simps) + +lemma cfrac_drop_Suc_left: "cfrac_drop (Suc n) c = cfrac_tl (cfrac_drop n c)" + by (simp add: cfrac_drop_def) + +lemma cfrac_drop_add: "cfrac_drop (m + n) c = cfrac_drop m (cfrac_drop n c)" + by (simp add: cfrac_drop_def funpow_add) + +lemma cfrac_drop_0 [simp]: "cfrac_drop 0 = (\x. x)" + by (simp add: fun_eq_iff cfrac_drop_def) + +lemma cfrac_drop_1 [simp]: "cfrac_drop 1 = cfrac_tl" + by (simp add: fun_eq_iff cfrac_drop_def) + +lift_definition cfrac_length :: "cfrac \ enat" is + "\(_, bs) \ llength bs" . + +lemma cfrac_length_code [code]: "cfrac_length (CFrac a bs) = llength bs" + by (simp add: cfrac_length_def) + +lemma cfrac_length_tl [simp]: "cfrac_length (cfrac_tl c) = cfrac_length c - 1" + by transfer (auto split: llist.splits) + +lemma enat_diff_Suc_right [simp]: "m - enat (Suc n) = m - n - 1" + by (auto simp: diff_enat_def enat_1_iff split: enat.splits) + +lemma cfrac_length_drop [simp]: "cfrac_length (cfrac_drop n c) = cfrac_length c - n" + by (induction n) (auto simp: cfrac_drop_def) + +lemma cfrac_length_of_stream [simp]: "cfrac_length (cfrac_of_stream xs) = \" + by transfer auto + +lift_definition cfrac_nth :: "cfrac \ nat \ int" is + "\(a :: int, bs :: nat llist). \(n :: nat). + if n = 0 then a + else if n \ llength bs then int (lnth bs (n - 1)) + 1 else 1" . + +lemma cfrac_nth_code [code]: + "cfrac_nth (CFrac a bs) n = (if n = 0 then a else lnth_default 0 bs (n - 1) + 1)" +proof - + have "n > 0 \ enat (n - Suc 0) < llength bs \ enat n \ llength bs" + by (metis Suc_ile_eq Suc_pred) + thus ?thesis by (auto simp: cfrac_nth_def lnth_default_def) +qed + +lemma cfrac_nth_nonneg [simp, intro]: "n > 0 \ cfrac_nth c n \ 0" + by transfer auto + +lemma cfrac_nth_nonzero [simp]: "n > 0 \ cfrac_nth c n \ 0" + by transfer (auto split: if_splits) + +lemma cfrac_nth_pos[simp, intro]: "n > 0 \ cfrac_nth c n > 0" + by transfer auto + +lemma cfrac_nth_ge_1[simp, intro]: "n > 0 \ cfrac_nth c n \ 1" + by transfer auto + +lemma cfrac_nth_not_less_1[simp, intro]: "n > 0 \ \cfrac_nth c n < 1" + by transfer (auto split: if_splits) + +lemma cfrac_nth_tl [simp]: "cfrac_nth (cfrac_tl c) n = cfrac_nth c (Suc n)" + apply transfer + apply (auto split: llist.splits nat.splits simp: Suc_ile_eq lnth_LCons enat_0_iff + simp flip: zero_enat_def) + done + +lemma cfrac_nth_drop [simp]: "cfrac_nth (cfrac_drop n c) m = cfrac_nth c (m + n)" + by (induction n arbitrary: m) (auto simp: cfrac_drop_def) + +lemma cfrac_nth_0_of_int [simp]: "cfrac_nth (cfrac_of_int n) 0 = n" + by transfer auto + +lemma cfrac_nth_gt0_of_int [simp]: "m > 0 \ cfrac_nth (cfrac_of_int n) m = 1" + by transfer (auto simp: enat_0_iff) + +lemma cfrac_nth_of_stream: + assumes "sset (stl xs) \ {0<..}" + shows "cfrac_nth (cfrac_of_stream xs) n = snth xs n" + using assms +proof (transfer', goal_cases) + case (1 xs n) + thus ?case + by (cases xs; cases n) (auto simp: subset_iff) +qed + + +lift_definition cfrac :: "(nat \ int) \ cfrac" is + "\f. (f 0, inf_llist (\n. nat (f (Suc n) - 1)))" . + +definition is_cfrac :: "(nat \ int) \ bool" where "is_cfrac f \ (\n>0. f n > 0)" + +lemma cfrac_nth_cfrac [simp]: + assumes "is_cfrac f" + shows "cfrac_nth (cfrac f) n = f n" + using assms unfolding is_cfrac_def by transfer auto + +lemma llength_eq_infty_lnth: "llength b = \ \ inf_llist (lnth b) = b" + by (simp add: llength_eq_infty_conv_lfinite) + +lemma cfrac_cfrac_nth [simp]: "cfrac_length c = \ \ cfrac (cfrac_nth c) = c" + by transfer (auto simp: llength_eq_infty_lnth) + +lemma cfrac_length_cfrac [simp]: "cfrac_length (cfrac f) = \" + by transfer auto + + +lift_definition cfrac_of_list :: "int list \ cfrac" is + "\xs. if xs = [] then (0, LNil) else (hd xs, llist_of (map (\n. nat n - 1) (tl xs)))" . + +lemma cfrac_length_of_list [simp]: "cfrac_length (cfrac_of_list xs) = length xs - 1" + by transfer (auto simp: zero_enat_def) + +lemma cfrac_of_list_Nil [simp]: "cfrac_of_list [] = 0" + unfolding zero_cfrac_def by transfer auto + +lemma cfrac_nth_of_list [simp]: + assumes "n < length xs" and "\i\{0<.. 0" + shows "cfrac_nth (cfrac_of_list xs) n = xs ! n" + using assms +proof (transfer, goal_cases) + case (1 n xs) + show ?case + proof (cases n) + case (Suc n') + with 1 have "xs ! n > 0" + using 1 by auto + hence "int (nat (tl xs ! n') - Suc 0) + 1 = xs ! Suc n'" + using 1(1) Suc by (auto simp: nth_tl of_nat_diff) + thus ?thesis + using Suc 1(1) by (auto simp: hd_conv_nth zero_enat_def) + qed (use 1 in \auto simp: hd_conv_nth\) +qed + + +primcorec cfrac_of_real_aux :: "real \ nat llist" where + "cfrac_of_real_aux x = + (if x \ {0<..<1} then LCons (nat \1/x\ - 1) (cfrac_of_real_aux (frac (1/x))) else LNil)" + +lemma cfrac_of_real_aux_code [code]: + "cfrac_of_real_aux x = + (if x > 0 \ x < 1 then LCons (nat \1/x\ - 1) (cfrac_of_real_aux (frac (1/x))) else LNil)" + by (subst cfrac_of_real_aux.code) auto + + +lemma cfrac_of_real_aux_LNil [simp]: "x \ {0<..<1} \ cfrac_of_real_aux x = LNil" + by (subst cfrac_of_real_aux.code) auto + +lemma cfrac_of_real_aux_0 [simp]: "cfrac_of_real_aux 0 = LNil" + by (subst cfrac_of_real_aux.code) auto + +lemma cfrac_of_real_aux_eq_LNil_iff [simp]: "cfrac_of_real_aux x = LNil \ x \ {0<..<1}" + by (subst cfrac_of_real_aux.code) auto + +lemma lnth_cfrac_of_real_aux: + assumes "n < llength (cfrac_of_real_aux x)" + shows "lnth (cfrac_of_real_aux x) (Suc n) = lnth (cfrac_of_real_aux (frac (1/x))) n" + using assms + apply (induction n arbitrary: x) + apply (subst cfrac_of_real_aux.code) + apply auto [] + apply (subst cfrac_of_real_aux.code) + apply (auto) + done + +lift_definition cfrac_of_real :: "real \ cfrac" is + "\x. (\x\, cfrac_of_real_aux (frac x))" . + +lemma cfrac_of_real_code [code]: "cfrac_of_real x = CFrac \x\ (cfrac_of_real_aux (frac x))" + by (simp add: cfrac_of_real_def) + +lemma eq_epred_iff: "m = epred n \ m = 0 \ n = 0 \ n = eSuc m" + by (cases m; cases n) (auto simp: enat_0_iff enat_eSuc_iff infinity_eq_eSuc_iff) + +lemma epred_eq_iff: "epred n = m \ m = 0 \ n = 0 \ n = eSuc m" + by (cases m; cases n) (auto simp: enat_0_iff enat_eSuc_iff infinity_eq_eSuc_iff) + +lemma epred_less: "n > 0 \ n \ \ \ epred n < n" + by (cases n) (auto simp: enat_0_iff) + +lemma cfrac_nth_of_real_0 [simp]: + "cfrac_nth (cfrac_of_real x) 0 = \x\" + by transfer auto + +lemma frac_eq_0 [simp]: "x \ \ \ frac x = 0" + by simp + +lemma cfrac_tl_of_real: + assumes "x \ \" + shows "cfrac_tl (cfrac_of_real x) = cfrac_of_real (1 / frac x)" + using assms +proof (transfer, goal_cases) + case (1 x) + hence "int (nat \1 / frac x\ - Suc 0) + 1 = \1 / frac x\" + by (subst of_nat_diff) (auto simp: le_nat_iff frac_le_1) + with \x \ \\ show ?case + by (subst cfrac_of_real_aux.code) (auto split: llist.splits simp: frac_lt_1) +qed + +lemma cfrac_nth_of_real_Suc: + assumes "x \ \" + shows "cfrac_nth (cfrac_of_real x) (Suc n) = cfrac_nth (cfrac_of_real (1 / frac x)) n" +proof - + have "cfrac_nth (cfrac_of_real x) (Suc n) = + cfrac_nth (cfrac_tl (cfrac_of_real x)) n" + by simp + also have "cfrac_tl (cfrac_of_real x) = cfrac_of_real (1 / frac x)" + by (simp add: cfrac_tl_of_real assms) + finally show ?thesis . +qed + + +fun conv :: "cfrac \ nat \ real" where + "conv c 0 = real_of_int (cfrac_nth c 0)" +| "conv c (Suc n) = real_of_int (cfrac_nth c 0) + 1 / conv (cfrac_tl c) n" + + +text \ + The numerator and denominator of a convergent: +\ +fun conv_num :: "cfrac \ nat \ int" where + "conv_num c 0 = cfrac_nth c 0" +| "conv_num c (Suc 0) = cfrac_nth c 1 * cfrac_nth c 0 + 1" +| "conv_num c (Suc (Suc n)) = cfrac_nth c (Suc (Suc n)) * conv_num c (Suc n) + conv_num c n" + +fun conv_denom :: "cfrac \ nat \ int" where + "conv_denom c 0 = 1" +| "conv_denom c (Suc 0) = cfrac_nth c 1" +| "conv_denom c (Suc (Suc n)) = cfrac_nth c (Suc (Suc n)) * conv_denom c (Suc n) + conv_denom c n" + +lemma conv_num_rec: + "n \ 2 \ conv_num c n = cfrac_nth c n * conv_num c (n - 1) + conv_num c (n - 2)" + by (cases n; cases "n - 1") auto + +lemma conv_denom_rec: + "n \ 2 \ conv_denom c n = cfrac_nth c n * conv_denom c (n - 1) + conv_denom c (n - 2)" + by (cases n; cases "n - 1") auto + + +fun conv' :: "cfrac \ nat \ real \ real" where + "conv' c 0 z = z" +| "conv' c (Suc n) z = conv' c n (real_of_int (cfrac_nth c n) + 1 / z)" + +text \ + Occasionally, it can be useful to extend the domain of @{const conv_num} and @{const conv_denom} + to \-1\ and \-2\. +\ +definition conv_num_int :: "cfrac \ int \ int" where + "conv_num_int c n = (if n = -1 then 1 else if n < 0 then 0 else conv_num c (nat n))" + +definition conv_denom_int :: "cfrac \ int \ int" where + "conv_denom_int c n = (if n = -2 then 1 else if n < 0 then 0 else conv_denom c (nat n))" + +lemma conv_num_int_rec: + assumes "n \ 0" + shows "conv_num_int c n = cfrac_nth c (nat n) * conv_num_int c (n - 1) + conv_num_int c (n - 2)" +proof (cases "n \ 2") + case True + define n' where "n' = nat (n - 2)" + have n: "n = int (Suc (Suc n'))" + using True by (simp add: n'_def) + show ?thesis + by (simp add: n conv_num_int_def nat_add_distrib) +qed (use assms in \auto simp: conv_num_int_def\) + +lemma conv_denom_int_rec: + assumes "n \ 0" + shows "conv_denom_int c n = cfrac_nth c (nat n) * conv_denom_int c (n - 1) + conv_denom_int c (n - 2)" +proof - + consider "n = 0" | "n = 1" | "n \ 2" + using assms by force + thus ?thesis + proof cases + assume "n \ 2" + define n' where "n' = nat (n - 2)" + have n: "n = int (Suc (Suc n'))" + using \n \ 2\ by (simp add: n'_def) + show ?thesis + by (simp add: n conv_denom_int_def nat_add_distrib) + qed (use assms in \auto simp: conv_denom_int_def\) +qed + +text \ + The number \[a\<^sub>0; a\<^sub>1, a\<^sub>2, \]\ that the infinite continued fraction converges to: +\ +definition cfrac_lim :: "cfrac \ real" where + "cfrac_lim c = + (case cfrac_length c of \ \ lim (conv c) | enat l \ conv c l)" + +lemma cfrac_lim_code [code]: + "cfrac_lim c = + (case cfrac_length c of enat l \ conv c l + | _ \ Code.abort (STR ''Cannot compute infinite continued fraction'') (\_. cfrac_lim c))" + by (simp add: cfrac_lim_def split: enat.splits) + +definition cfrac_remainder where "cfrac_remainder c n = cfrac_lim (cfrac_drop n c)" + +lemmas conv'_Suc_right = conv'.simps(2) + +lemma conv'_Suc_left: + assumes "z > 0" + shows "conv' c (Suc n) z = + real_of_int (cfrac_nth c 0) + 1 / conv' (cfrac_tl c) n z" + using assms +proof (induction n arbitrary: z) + case (Suc n z) + have "conv' c (Suc (Suc n)) z = + conv' c (Suc n) (real_of_int (cfrac_nth c (Suc n)) + 1 / z)" + by simp + also have "\ = cfrac_nth c 0 + 1 / conv' (cfrac_tl c) (Suc n) z" + using Suc.prems by (subst Suc.IH) (auto intro!: add_nonneg_pos cfrac_nth_nonneg) + finally show ?case . +qed simp_all + +lemmas [simp del] = conv'.simps(2) + +lemma conv'_left_induct: + assumes "\c. P c 0 z" "\c n. P (cfrac_tl c) n z \ P c (Suc n) z" + shows "P c n z" + using assms by (rule conv.induct) + +lemma enat_less_diff_conv [simp]: + assumes "a = \ \ b < \ \ c < \" + shows "a < c - (b :: enat) \ a + b < c" + using assms by (cases a; cases b; cases c) auto + +lemma conv_eq_conv': "conv c n = conv' c n (cfrac_nth c n)" +proof (cases "n = 0") + case False + hence "cfrac_nth c n > 0" by (auto intro!: cfrac_nth_pos) + thus ?thesis + by (induction c n rule: conv.induct) (simp_all add: conv'_Suc_left) +qed simp_all + +lemma conv_num_pos': + assumes "cfrac_nth c 0 > 0" + shows "conv_num c n > 0" + using assms by (induction n rule: fib.induct) (auto simp: intro!: add_pos_nonneg) + +lemma conv_num_nonneg: "cfrac_nth c 0 \ 0 \ conv_num c n \ 0" + by (induction c n rule: conv_num.induct) + (auto simp: intro!: mult_nonneg_nonneg add_nonneg_nonneg + intro: cfrac_nth_nonneg) + +lemma conv_num_pos: + "cfrac_nth c 0 \ 0 \ n > 0 \ conv_num c n > 0" + by (induction c n rule: conv_num.induct) + (auto intro!: mult_pos_pos mult_nonneg_nonneg add_pos_nonneg conv_num_nonneg cfrac_nth_pos + intro: cfrac_nth_nonneg simp: enat_le_iff) + +lemma conv_denom_pos [simp, intro]: "conv_denom c n > 0" + by (induction c n rule: conv_num.induct) + (auto intro!: add_nonneg_pos mult_nonneg_nonneg cfrac_nth_nonneg + simp: enat_le_iff) + +lemma conv_denom_not_nonpos [simp]: "\conv_denom c n \ 0" + using conv_denom_pos[of c n] by linarith + +lemma conv_denom_not_neg [simp]: "\conv_denom c n < 0" + using conv_denom_pos[of c n] by linarith + +lemma conv_denom_nonzero [simp]: "conv_denom c n \ 0" + using conv_denom_pos[of c n] by linarith + +lemma conv_denom_nonneg [simp, intro]: "conv_denom c n \ 0" + using conv_denom_pos[of c n] by linarith + +lemma conv_num_int_neg1 [simp]: "conv_num_int c (-1) = 1" + by (simp add: conv_num_int_def) + +lemma conv_num_int_neg [simp]: "n < 0 \ n \ -1 \ conv_num_int c n = 0" + by (simp add: conv_num_int_def) + +lemma conv_num_int_of_nat [simp]: "conv_num_int c (int n) = conv_num c n" + by (simp add: conv_num_int_def) + +lemma conv_num_int_nonneg [simp]: "n \ 0 \ conv_num_int c n = conv_num c (nat n)" + by (simp add: conv_num_int_def) + +lemma conv_denom_int_neg2 [simp]: "conv_denom_int c (-2) = 1" + by (simp add: conv_denom_int_def) + +lemma conv_denom_int_neg [simp]: "n < 0 \ n \ -2 \ conv_denom_int c n = 0" + by (simp add: conv_denom_int_def) + +lemma conv_denom_int_of_nat [simp]: "conv_denom_int c (int n) = conv_denom c n" + by (simp add: conv_denom_int_def) + +lemma conv_denom_int_nonneg [simp]: "n \ 0 \ conv_denom_int c n = conv_denom c (nat n)" + by (simp add: conv_denom_int_def) + +lemmas conv_Suc [simp del] = conv.simps(2) + + +lemma conv'_gt_1: + assumes "cfrac_nth c 0 > 0" "x > 1" + shows "conv' c n x > 1" + using assms +proof (induction n arbitrary: c x) + case (Suc n c x) + from Suc.prems have pos: "cfrac_nth c n > 0" using cfrac_nth_pos[of n c] + by (cases "n = 0") (auto simp: enat_le_iff) + have "1 < 1 + 1 / x" + using Suc.prems by simp + also have "\ \ cfrac_nth c n + 1 / x" using pos + by (intro add_right_mono) (auto simp: of_nat_ge_1_iff) + finally show ?case + by (subst conv'_Suc_right, intro Suc.IH) + (use Suc.prems in \auto simp: enat_le_iff\) +qed auto + +lemma enat_eq_iff: "a = enat b \ (\a'. a = enat a' \ a' = b)" + by (cases a) auto + +lemma eq_enat_iff: "enat a = b \ (\b'. b = enat b' \ a = b')" + by (cases b) auto + +lemma enat_diff_one [simp]: "enat a - 1 = enat (a - 1)" + by (cases "enat (a - 1)") (auto simp flip: idiff_enat_enat) + +lemma conv'_eqD: + assumes "conv' c n x = conv' c' n x" "x > 1" "m < n" + shows "cfrac_nth c m = cfrac_nth c' m" + using assms +proof (induction n arbitrary: m c c') + case (Suc n m c c') + have gt: "conv' (cfrac_tl c) n x > 1" "conv' (cfrac_tl c') n x > 1" + by (rule conv'_gt_1; + use Suc.prems in \force intro: cfrac_nth_pos simp: enat_le_iff\)+ + have eq: "cfrac_nth c 0 + 1 / conv' (cfrac_tl c) n x = + cfrac_nth c' 0 + 1 / conv' (cfrac_tl c') n x" + using Suc.prems by (subst (asm) (1 2) conv'_Suc_left) auto + hence "\cfrac_nth c 0 + 1 / conv' (cfrac_tl c) n x\ = + \cfrac_nth c' 0 + 1 / conv' (cfrac_tl c') n x\" + by (simp only: ) + also from gt have "floor (cfrac_nth c 0 + 1 / conv' (cfrac_tl c) n x) = cfrac_nth c 0" + by (intro floor_unique) auto + also from gt have "floor (cfrac_nth c' 0 + 1 / conv' (cfrac_tl c') n x) = cfrac_nth c' 0" + by (intro floor_unique) auto + finally have [simp]: "cfrac_nth c 0 = cfrac_nth c' 0" by simp + + show ?case + proof (cases m) + case (Suc m') + from eq and gt have "conv' (cfrac_tl c) n x = conv' (cfrac_tl c') n x" + by simp + hence "cfrac_nth (cfrac_tl c) m' = cfrac_nth (cfrac_tl c') m'" + using Suc.prems + by (intro Suc.IH[of "cfrac_tl c" "cfrac_tl c'"]) (auto simp: o_def Suc enat_le_iff) + with Suc show ?thesis by simp + qed simp_all +qed simp_all + +context + fixes c :: cfrac and h k + defines "h \ conv_num c" and "k \ conv_denom c" +begin + +lemma conv'_num_denom_aux: + assumes z: "z > 0" + shows "conv' c (Suc (Suc n)) z * (z * k (Suc n) + k n) = + (z * h (Suc n) + h n)" + using z +proof (induction n arbitrary: z) + case 0 + hence "1 + z * cfrac_nth c 1 > 0" + by (intro add_pos_nonneg) (auto simp: cfrac_nth_nonneg) + with 0 show ?case + by (auto simp add: h_def k_def field_simps conv'_Suc_right max_def not_le) +next + case (Suc n) + have [simp]: "h (Suc (Suc n)) = cfrac_nth c (n+2) * h (n+1) + h n" + by (simp add: h_def) + have [simp]: "k (Suc (Suc n)) = cfrac_nth c (n+2) * k (n+1) + k n" + by (simp add: k_def) + define z' where "z' = cfrac_nth c (n+2) + 1 / z" + from \z > 0\ have "z' > 0" + by (auto simp: z'_def intro!: add_nonneg_pos cfrac_nth_nonneg) + + have "z * real_of_int (h (Suc (Suc n))) + real_of_int (h (Suc n)) = + z * (z' * h (Suc n) + h n)" + using \z > 0\ by (simp add: algebra_simps z'_def) + also have "\ = z * (conv' c (Suc (Suc n)) z' * (z' * k (Suc n) + k n))" + using \z' > 0\ by (subst Suc.IH [symmetric]) auto + also have "\ = conv' c (Suc (Suc (Suc n))) z * + (z * k (Suc (Suc n)) + k (Suc n))" + unfolding z'_def using \z > 0\ + by (subst (2) conv'_Suc_right) (simp add: algebra_simps) + finally show ?case .. +qed + +lemma conv'_num_denom: + assumes "z > 0" + shows "conv' c (Suc (Suc n)) z = + (z * h (Suc n) + h n) / (z * k (Suc n) + k n)" +proof - + have "z * real_of_int (k (Suc n)) + real_of_int (k n) > 0" + using assms by (intro add_pos_nonneg mult_pos_pos) (auto simp: k_def) + with conv'_num_denom_aux[of z n] assms show ?thesis + by (simp add: divide_simps) +qed + +lemma conv_num_denom: "conv c n = h n / k n" +proof - + consider "n = 0" | "n = Suc 0" | m where "n = Suc (Suc m)" + using not0_implies_Suc by blast + thus ?thesis + proof cases + assume "n = Suc 0" + thus ?thesis + by (auto simp: h_def k_def field_simps max_def conv_Suc) + next + fix m assume [simp]: "n = Suc (Suc m)" + have "conv c n = conv' c (Suc (Suc m)) (cfrac_nth c (Suc (Suc m)))" + by (subst conv_eq_conv') simp_all + also have "\ = h n / k n" + by (subst conv'_num_denom) (simp_all add: h_def k_def) + finally show ?thesis . + qed (auto simp: h_def k_def) +qed + +lemma conv'_num_denom': + assumes "z > 0" and "n \ 2" + shows "conv' c n z = (z * h (n - 1) + h (n - 2)) / (z * k (n - 1) + k (n - 2))" + using assms conv'_num_denom[of z "n - 2"] + by (auto simp: eval_nat_numeral Suc_diff_Suc) + +lemma conv'_num_denom_int: + assumes "z > 0" + shows "conv' c n z = + (z * conv_num_int c (int n - 1) + conv_num_int c (int n - 2)) / + (z * conv_denom_int c (int n - 1) + conv_denom_int c (int n - 2))" +proof - + consider "n = 0" | "n = 1" | "n \ 2" by force + thus ?thesis + proof cases + case 1 + thus ?thesis using conv_num_int_neg1 by auto + next + case 2 + thus ?thesis using assms by (auto simp: conv'_Suc_right field_simps) + next + case 3 + thus ?thesis using conv'_num_denom'[OF assms(1), of "nat n"] + by (auto simp: nat_diff_distrib h_def k_def) + qed +qed + +lemma conv_nonneg: "cfrac_nth c 0 \ 0 \ conv c n \ 0" + by (subst conv_num_denom) + (auto intro!: divide_nonneg_nonneg conv_num_nonneg simp: h_def k_def) + +lemma conv_pos: + assumes "cfrac_nth c 0 > 0" + shows "conv c n > 0" +proof - + have "conv c n = h n / k n" + using assms by (intro conv_num_denom) + also from assms have "\ > 0" unfolding h_def k_def + by (intro divide_pos_pos) (auto intro!: conv_num_pos') + finally show ?thesis . +qed + +lemma conv_num_denom_prod_diff: + "k n * h (Suc n) - k (Suc n) * h n = (-1) ^ n" + by (induction c n rule: conv_num.induct) + (auto simp: k_def h_def algebra_simps) + +lemma conv_num_denom_prod_diff': + "k (Suc n) * h n - k n * h (Suc n) = (-1) ^ Suc n" + by (induction c n rule: conv_num.induct) + (auto simp: k_def h_def algebra_simps) + +lemma + fixes n :: int + assumes "n \ -2" + shows conv_num_denom_int_prod_diff: + "conv_denom_int c n * conv_num_int c (n + 1) - + conv_denom_int c (n + 1) * conv_num_int c n = (-1) ^ (nat (n + 2))" (is ?th1) + and conv_num_denom_int_prod_diff': + "conv_denom_int c (n + 1) * conv_num_int c n - + conv_denom_int c n * conv_num_int c (n + 1) = (-1) ^ (nat (n + 3))" (is ?th2) +proof - + from assms consider "n = -2" | "n = -1" | "n \ 0" by force + thus ?th1 using conv_num_denom_prod_diff[of "nat n"] + by cases (auto simp: h_def k_def nat_add_distrib) + moreover from assms have "nat (n + 3) = Suc (nat (n + 2))" by (simp add: nat_add_distrib) + ultimately show ?th2 by simp +qed + +lemma coprime_conv_num_denom: "coprime (h n) (k n)" +proof (cases n) + case [simp]: (Suc m) + { + fix d :: int + assume "d dvd h n" and "d dvd k n" + hence "abs d dvd abs (k n * h (Suc n) - k (Suc n) * h n)" + by simp + also have "\ = 1" + by (subst conv_num_denom_prod_diff) auto + finally have "is_unit d" by simp + } + thus ?thesis by (rule coprimeI) +qed (auto simp: h_def k_def) + +lemma coprime_conv_num_denom_int: + assumes "n \ -2" + shows "coprime (conv_num_int c n) (conv_denom_int c n)" +proof - + from assms consider "n = -2" | "n = -1" | "n \ 0" by force + thus ?thesis by cases (insert coprime_conv_num_denom[of "nat n"], auto simp: h_def k_def) +qed + +lemma mono_conv_num: + assumes "cfrac_nth c 0 \ 0" + shows "mono h" +proof (rule incseq_SucI) + show "h n \ h (Suc n)" for n + proof (cases n) + case 0 + have "1 * cfrac_nth c 0 + 1 \ cfrac_nth c (Suc 0) * cfrac_nth c 0 + 1" + using assms by (intro add_mono mult_right_mono) auto + thus ?thesis using assms by (simp add: le_Suc_eq Suc_le_eq h_def 0) + next + case (Suc m) + have "1 * h (Suc m) + 0 \ cfrac_nth c (Suc (Suc m)) * h (Suc m) + h m" + using assms + by (intro add_mono mult_right_mono) + (auto simp: Suc_le_eq h_def intro!: conv_num_nonneg) + with Suc show ?thesis by (simp add: h_def) + qed +qed + +lemma mono_conv_denom: "mono k" +proof (rule incseq_SucI) + show "k n \ k (Suc n)" for n + proof (cases n) + case 0 + thus ?thesis by (simp add: le_Suc_eq Suc_le_eq k_def) + next + case (Suc m) + have "1 * k (Suc m) + 0 \ cfrac_nth c (Suc (Suc m)) * k (Suc m) + k m" + by (intro add_mono mult_right_mono) (auto simp: Suc_le_eq k_def) + with Suc show ?thesis by (simp add: k_def) + qed +qed + +lemma conv_num_leI: "cfrac_nth c 0 \ 0 \ m \ n \ h m \ h n" + using mono_conv_num by (auto simp: mono_def) + +lemma conv_denom_leI: "m \ n \ k m \ k n" + using mono_conv_denom by (auto simp: mono_def) + +lemma conv_denom_lessI: + assumes "m < n" "1 < n" + shows "k m < k n" +proof (cases n) + case [simp]: (Suc n') + show ?thesis + proof (cases n') + case [simp]: (Suc n'') + from assms have "k m \ 1 * k n' + 0" + by (auto intro: conv_denom_leI simp: less_Suc_eq) + also have "\ \ cfrac_nth c n * k n' + 0" + using assms by (intro add_mono mult_mono) (auto simp: Suc_le_eq k_def) + also have "\ < cfrac_nth c n * k n' + k n''" unfolding k_def + by (intro add_strict_left_mono conv_denom_pos assms) + also have "\ = k n" by (simp add: k_def) + finally show ?thesis . + qed (insert assms, auto simp: k_def) +qed (insert assms, auto) + +lemma conv_num_lower_bound: + assumes "cfrac_nth c 0 \ 0" + shows "h n \ fib n" unfolding h_def + using assms +proof (induction c n rule: conv_denom.induct) + case (3 c n) + hence "conv_num c (Suc (Suc n)) \ 1 * int (fib (Suc n)) + int (fib n)" + using "3.prems" unfolding conv_num.simps + by (intro add_mono mult_mono "3.IH") auto + thus ?case by simp +qed auto + +lemma conv_denom_lower_bound: "k n \ fib (Suc n)" + unfolding k_def +proof (induction c n rule: conv_denom.induct) + case (3 c n) + hence "conv_denom c (Suc (Suc n)) \ 1 * int (fib (Suc (Suc n))) + int (fib (Suc n))" + using "3.prems" unfolding conv_denom.simps + by (intro add_mono mult_mono "3.IH") auto + thus ?case by simp +qed (auto simp: Suc_le_eq) + +lemma conv_diff_eq: "conv c (Suc n) - conv c n = (-1) ^ n / (k n * k (Suc n))" +proof - + have pos: "k n > 0" "k (Suc n) > 0" unfolding k_def + by (intro conv_denom_pos)+ + have "conv c (Suc n) - conv c n = + (k n * h (Suc n) - k (Suc n) * h n) / (k n * k (Suc n))" + using pos by (subst (1 2) conv_num_denom) (simp add: conv_num_denom field_simps) + also have "k n * h (Suc n) - k (Suc n) * h n = (-1) ^ n" + by (rule conv_num_denom_prod_diff) + finally show ?thesis by simp +qed + +lemma conv_telescope: + assumes "m \ n" + shows "conv c m + (\i=m..i=m..i=m.. = conv c n" + using assms by (induction rule: dec_induct) simp_all + finally show ?thesis . +qed + +lemma fib_at_top: "filterlim fib at_top at_top" +proof (rule filterlim_at_top_mono) + show "eventually (\n. fib n \ n - 1) at_top" + by (intro always_eventually fib_ge allI) + show "filterlim (\n::nat. n - 1) at_top at_top" + by (subst filterlim_sequentially_Suc [symmetric]) + (simp_all add: filterlim_ident) +qed + +lemma conv_denom_at_top: "filterlim k at_top at_top" +proof (rule filterlim_at_top_mono) + show "filterlim (\n. int (fib (Suc n))) at_top at_top" + by (rule filterlim_compose[OF filterlim_int_sequentially]) + (simp add: fib_at_top filterlim_sequentially_Suc) + show "eventually (\n. fib (Suc n) \ k n) at_top" + by (intro always_eventually conv_denom_lower_bound allI) +qed + +lemma + shows summable_conv_telescope: + "summable (\i. (-1) ^ i / (k i * k (Suc i)))" (is ?th1) + and cfrac_remainder_bounds: + "\(\i. (-1) ^ (i + m) / (k (i + m) * k (Suc i + m)))\ \ + {1/(k m * (k m + k (Suc m))) <..< 1 / (k m * k (Suc m))}" (is ?th2) +proof - + have [simp]: "k n > 0" "k n \ 0" "\k n = 0" for n + by (auto simp: k_def) + have k_rec: "k (Suc (Suc n)) = cfrac_nth c (Suc (Suc n)) * k (Suc n) + k n" for n + by (simp add: k_def) + have [simp]: "a + b = 0 \ a = 0 \ b = 0" if "a \ 0" "b \ 0" for a b :: real + using that by linarith + + define g where "g = (\i. inverse (real_of_int (k i * k (Suc i))))" + + { + fix m :: nat + have "filterlim (\n. k n) at_top at_top" and "filterlim (\n. k (Suc n)) at_top at_top" + by (force simp: filterlim_sequentially_Suc intro: conv_denom_at_top)+ + hence lim: "g \ 0" + unfolding g_def of_int_mult + by (intro tendsto_inverse_0_at_top filterlim_at_top_mult_at_top + filterlim_compose[OF filterlim_real_of_int_at_top]) + from lim have A: "summable (\n. (-1) ^ (n + m) * g (n + m))" unfolding g_def + by (intro summable_alternating_decreasing) + (auto intro!: conv_denom_leI mult_nonneg_nonneg) + + have "1 / (k m * (real_of_int (k (Suc m)) + k m / 1)) \ + 1 / (k m * (k (Suc m) + k m / cfrac_nth c (m+2)))" + by (intro divide_left_mono mult_left_mono add_left_mono mult_pos_pos add_pos_pos divide_pos_pos) + (auto simp: of_nat_ge_1_iff) + also have "\ = g m - g (Suc m)" + by (simp add: g_def k_rec field_simps add_pos_pos) + finally have le: "1 / (k m * (real_of_int (k (Suc m)) + k m / 1)) \ g m - g (Suc m)" by simp + have *: "\(\i. (-1) ^ (i + m) * g (i + m))\ \ {g m - g (Suc m) <..< g m}" + using lim unfolding g_def + by (intro abs_alternating_decreasing_suminf_strict) (auto intro!: conv_denom_lessI) + also from le have "\ \ {1 / (k m * (k (Suc m) + k m)) <..< g m}" + by (subst greaterThanLessThan_subseteq_greaterThanLessThan) auto + finally have B: "\\i. (- 1) ^ (i + m) * g (i + m)\ \ \" . + note A B + } note AB = this + + from AB(1)[of 0] show ?th1 by (simp add: field_simps g_def) + from AB(2)[of m] show ?th2 by (simp add: g_def divide_inverse add_ac) +qed + +lemma convergent_conv: "convergent (conv c)" +proof - + have "convergent (\n. conv c 0 + (\i = conv c" + by (rule ext, subst (2) conv_telescope [of 0, symmetric]) (simp_all add: atLeast0LessThan) + finally show ?thesis . +qed + +lemma LIMSEQ_cfrac_lim: "cfrac_length c = \ \ conv c \ cfrac_lim c" + using convergent_conv by (auto simp: convergent_LIMSEQ_iff cfrac_lim_def) + +lemma cfrac_lim_nonneg: + assumes "cfrac_nth c 0 \ 0" + shows "cfrac_lim c \ 0" +proof (cases "cfrac_length c") + case infinity + have "conv c \ cfrac_lim c" + by (rule LIMSEQ_cfrac_lim) fact + thus ?thesis + by (rule tendsto_lowerbound) + (auto intro!: conv_nonneg always_eventually assms) +next + case (enat l) + thus ?thesis using assms + by (auto simp: cfrac_lim_def conv_nonneg) +qed + +lemma sums_cfrac_lim_minus_conv: + assumes "cfrac_length c = \" + shows "(\i. (-1) ^ (i + m) / (k (i + m) * k (Suc i + m))) sums (cfrac_lim c - conv c m)" +proof - + have "(\n. conv c (n + m) - conv c m) \ cfrac_lim c - conv c m" + by (auto intro!: tendsto_diff LIMSEQ_cfrac_lim simp: filterlim_sequentially_shift assms) + also have "(\n. conv c (n + m) - conv c m) = + (\n. (\i=0 + m.. = (\n. (\i cfrac_length c" + shows "\cfrac_lim c - conv c m\ \ 1 / (k m * k (Suc m))" +proof (cases "cfrac_length c") + case infinity + have "cfrac_lim c - conv c m = (\i. (-1) ^ (i + m) / (k (i + m) * k (Suc i + m)))" + using sums_cfrac_lim_minus_conv infinity by (simp add: sums_iff) + also note cfrac_remainder_bounds[of m] + finally show ?thesis by simp +next + case [simp]: (enat l) + show ?thesis + proof (cases "l = m") + case True + thus ?thesis by (auto simp: cfrac_lim_def k_def) + next + case False + let ?S = "(\i=m.. 0" "k n > 0" for n + by (simp_all add: k_def) + hence "cfrac_lim c - conv c m = conv c l - conv c m" + by (simp add: cfrac_lim_def) + also have "\ = ?S" + using assms by (subst conv_telescope [symmetric, of m]) auto + finally have "cfrac_lim c - conv c m = ?S" . + moreover have "\?S\ \ 1 / real_of_int (k m * k (Suc m))" + unfolding of_int_mult using assms False + by (intro abs_alternating_decreasing_sum_upper_bound' divide_nonneg_nonneg frac_le mult_mono) + (simp_all add: conv_denom_leI del: conv_denom.simps) + ultimately show ?thesis by simp + qed +qed + +lemma cfrac_lim_minus_conv_lower_bound: + assumes "m < cfrac_length c" + shows "\cfrac_lim c - conv c m\ \ 1 / (k m * (k m + k (Suc m)))" +proof (cases "cfrac_length c") + case infinity + have "cfrac_lim c - conv c m = (\i. (-1) ^ (i + m) / (k (i + m) * k (Suc i + m)))" + using sums_cfrac_lim_minus_conv infinity by (simp add: sums_iff) + also note cfrac_remainder_bounds[of m] + finally show ?thesis by simp +next + case [simp]: (enat l) + let ?S = "(\i=m.. 0" "k n > 0" for n + by (simp_all add: k_def) + hence "cfrac_lim c - conv c m = conv c l - conv c m" + by (simp add: cfrac_lim_def) + also have "\ = ?S" + using assms by (subst conv_telescope [symmetric, of m]) (auto simp: split: enat.splits) + finally have "cfrac_lim c - conv c m = ?S" . + + moreover have "\?S\ \ 1 / (k m * (k m + k (Suc m)))" + proof (cases "m < cfrac_length c - 1") + case False + hence [simp]: "m = l - 1" and "l > 0" using assms + by (auto simp: not_less) + have "1 / (k m * (k m + k (Suc m))) \ 1 / (k m * k (Suc m))" + unfolding of_int_mult + by (intro divide_left_mono mult_mono mult_pos_pos) (auto intro!: add_pos_pos) + also from \l > 0\ have "{m..?S\" + by simp + finally show ?thesis . + next + case True + with assms have less: "m < l - 1" + by auto + have "k m + k (Suc m) > 0" + by (intro add_pos_pos) (auto simp: k_def) + hence "1 / (k m * (k m + k (Suc m))) \ 1 / (k m * k (Suc m)) - 1 / (k (Suc m) * k (Suc (Suc m)))" + by (simp add: divide_simps) (auto simp: k_def algebra_simps) + also have "\ \ \?S\" + unfolding of_int_mult using less + by (intro abs_alternating_decreasing_sum_lower_bound' divide_nonneg_nonneg frac_le mult_mono) + (simp_all add: conv_denom_leI del: conv_denom.simps) + finally show ?thesis . + qed + ultimately show ?thesis by simp +qed + +lemma cfrac_lim_minus_conv_bounds: + assumes "m < cfrac_length c" + shows "\cfrac_lim c - conv c m\ \ {1 / (k m * (k m + k (Suc m)))..1 / (k m * k (Suc m))}" + using cfrac_lim_minus_conv_lower_bound[of m] cfrac_lim_minus_conv_upper_bound[of m] assms + by auto + +end + + +lemma conv_pos': + assumes "n > 0" "cfrac_nth c 0 \ 0" + shows "conv c n > 0" + using assms by (cases n) (auto simp: conv_Suc intro!: add_nonneg_pos conv_pos) + +lemma conv_in_Rats [intro]: "conv c n \ \" + by (induction c n rule: conv.induct) (auto simp: conv_Suc o_def) + +lemma + assumes "0 < z1" "z1 \ z2" + shows conv'_even_mono: "even n \ conv' c n z1 \ conv' c n z2" + and conv'_odd_mono: "odd n \ conv' c n z1 \ conv' c n z2" +proof - + let ?P = "(\n (f::nat\real\real). + if even n then f n z1 \ f n z2 else f n z1 \ f n z2)" + have "?P n (conv' c)" using assms + proof (induction n arbitrary: z1 z2) + case (Suc n) + note z12 = Suc.prems + consider "n = 0" | "even n" "n > 0" | "odd n" by force + thus ?case + proof cases + assume "n = 0" + thus ?thesis using Suc by (simp add: conv'_Suc_right field_simps) + next + assume n: "even n" "n > 0" + with Suc.IH have IH: "conv' c n z1 \ conv' c n z2" + if "0 < z1" "z1 \ z2" for z1 z2 using that by auto + show ?thesis using Suc.prems n z12 + by (auto simp: conv'_Suc_right field_simps intro!: IH add_pos_nonneg mult_nonneg_nonneg) + next + assume n: "odd n" + hence [simp]: "n > 0" by (auto intro!: Nat.gr0I) + from n and Suc.IH have IH: "conv' c n z1 \ conv' c n z2" + if "0 < z1" "z1 \ z2" for z1 z2 using that by auto + show ?thesis using Suc.prems n + by (auto simp: conv'_Suc_right field_simps + intro!: IH add_pos_nonneg mult_nonneg_nonneg) + qed + qed auto + thus "even n \ conv' c n z1 \ conv' c n z2" + "odd n \ conv' c n z1 \ conv' c n z2" by auto +qed + +lemma + shows conv_even_mono: "even n \ n \ m \ conv c n \ conv c m" + and conv_odd_mono: "odd n \ n \ m \ conv c n \ conv c m" +proof - + assume "even n" + have A: "conv c n \ conv c (Suc (Suc n))" if "even n" for n + proof (cases "n = 0") + case False + with \even n\ show ?thesis + by (auto simp add: conv_eq_conv' conv'_Suc_right intro: conv'_even_mono) + qed (auto simp: conv_Suc) + + have B: "conv c n \ conv c (Suc n)" if "even n" for n + proof (cases "n = 0") + case False + with \even n\ show ?thesis + by (auto simp add: conv_eq_conv' conv'_Suc_right intro: conv'_even_mono) + qed (auto simp: conv_Suc) + + show "conv c n \ conv c m" if "n \ m" for m + using that + proof (induction m rule: less_induct) + case (less m) + from \n \ m\ consider "m = n" | "even m" "m > n" | "odd m" "m > n" + by force + thus ?case + proof cases + assume m: "even m" "m > n" + with \even n\ have m': "m - 2 \ n" by presburger + with m have "conv c n \ conv c (m - 2)" + by (intro less.IH) auto + also have "\ \ conv c (Suc (Suc (m - 2)))" + using m m' by (intro A) auto + also have "Suc (Suc (m - 2)) = m" + using m by presburger + finally show ?thesis . + next + assume m: "odd m" "m > n" + hence "conv c n \ conv c (m - 1)" + by (intro less.IH) auto + also have "\ \ conv c (Suc (m - 1))" + using m by (intro B) auto + also have "Suc (m - 1) = m" + using m by simp + finally show ?thesis . + qed simp_all + qed +next + assume "odd n" + have A: "conv c n \ conv c (Suc (Suc n))" if "odd n" for n + using that + by (auto simp add: conv_eq_conv' conv'_Suc_right odd_pos intro!: conv'_odd_mono) + have B: "conv c n \ conv c (Suc n)" if "odd n" for n using that + by (auto simp add: conv_eq_conv' conv'_Suc_right odd_pos intro!: conv'_odd_mono) + + show "conv c n \ conv c m" if "n \ m" for m + using that + proof (induction m rule: less_induct) + case (less m) + from \n \ m\ consider "m = n" | "even m" "m > n" | "odd m" "m > n" + by force + thus ?case + proof cases + assume m: "odd m" "m > n" + with \odd n\ have m': "m - 2 \ n" "m \ 2" by presburger+ + from m and \odd n\ have "m = Suc (Suc (m - 2))" by presburger + also have "conv c \ \ conv c (m - 2)" + using m m' by (intro A) auto + also have "\ \ conv c n" + using m m' by (intro less.IH) auto + finally show ?thesis . + next + assume m: "even m" "m > n" + from m have "m = Suc (m - 1)" by presburger + also have "conv c \ \ conv c (m - 1)" + using m by (intro B) auto + also have "\ \ conv c n" + using m by (intro less.IH) auto + finally show ?thesis . + qed simp_all + qed +qed + +lemma + assumes "m \ cfrac_length c" + shows conv_le_cfrac_lim: "even m \ conv c m \ cfrac_lim c" + and conv_ge_cfrac_lim: "odd m \ conv c m \ cfrac_lim c" +proof - + have "if even m then conv c m \ cfrac_lim c else conv c m \ cfrac_lim c" + proof (cases "cfrac_length c") + case [simp]: infinity + show ?thesis + proof (cases "even m") + case True + have "eventually (\i. conv c m \ conv c i) at_top" + using eventually_ge_at_top[of m] by eventually_elim (rule conv_even_mono[OF True]) + hence "conv c m \ cfrac_lim c" + by (intro tendsto_lowerbound[OF LIMSEQ_cfrac_lim]) auto + thus ?thesis using True by simp + next + case False + have "eventually (\i. conv c m \ conv c i) at_top" + using eventually_ge_at_top[of m] by eventually_elim (rule conv_odd_mono[OF False]) + hence "conv c m \ cfrac_lim c" + by (intro tendsto_upperbound[OF LIMSEQ_cfrac_lim]) auto + thus ?thesis using False by simp + qed + next + case [simp]: (enat l) + show ?thesis + using conv_even_mono[of m l c] conv_odd_mono[of m l c] assms + by (auto simp: cfrac_lim_def) + qed + thus "even m \ conv c m \ cfrac_lim c" and "odd m \ conv c m \ cfrac_lim c" + by auto +qed + +lemma cfrac_lim_ge_first: "cfrac_lim c \ cfrac_nth c 0" + using conv_le_cfrac_lim[of 0 c] by (auto simp: less_eq_enat_def split: enat.splits) + +lemma cfrac_lim_pos: "cfrac_nth c 0 > 0 \ cfrac_lim c > 0" + by (rule less_le_trans[OF _ cfrac_lim_ge_first]) auto + +lemma conv'_eq_iff: + assumes "0 \ z1 \ 0 \ z2" + shows "conv' c n z1 = conv' c n z2 \ z1 = z2" +proof + assume "conv' c n z1 = conv' c n z2" + thus "z1 = z2" using assms + proof (induction n arbitrary: z1 z2) + case (Suc n) + show ?case + proof (cases "n = 0") + case True + thus ?thesis using Suc by (auto simp: conv'_Suc_right) + next + case False + have "conv' c n (real_of_int (cfrac_nth c n) + 1 / z1) = + conv' c n (real_of_int (cfrac_nth c n) + 1 / z2)" using Suc.prems + by (simp add: conv'_Suc_right) + hence "real_of_int (cfrac_nth c n) + 1 / z1 = real_of_int (cfrac_nth c n) + 1 / z2" + by (rule Suc.IH) + (insert Suc.prems False, auto intro!: add_nonneg_pos add_nonneg_nonneg) + with Suc.prems show "z1 = z2" by simp + qed + qed auto +qed auto + +lemma conv_even_mono_strict: + assumes "even n" "n < m" + shows "conv c n < conv c m" +proof (cases "m = n + 1") + case [simp]: True + show ?thesis + proof (cases "n = 0") + case True + thus ?thesis using assms by (auto simp: conv_Suc) + next + case False + hence "conv' c n (real_of_int (cfrac_nth c n)) \ + conv' c n (real_of_int (cfrac_nth c n) + 1 / real_of_int (cfrac_nth c (Suc n)))" + by (subst conv'_eq_iff) auto + with assms have "conv c n \ conv c m" + by (auto simp: conv_eq_conv' conv'_eq_iff conv'_Suc_right field_simps) + moreover from assms have "conv c n \ conv c m" + by (intro conv_even_mono) auto + + ultimately show ?thesis by simp + qed +next + case False + show ?thesis + proof (cases "n = 0") + case True + thus ?thesis using assms + by (cases m) (auto simp: conv_Suc conv_pos) + next + case False + have "1 + real_of_int (cfrac_nth c (n+1)) * cfrac_nth c (n+2) > 0" + by (intro add_pos_nonneg) auto + with assms have "conv c n \ conv c (Suc (Suc n))" + unfolding conv_eq_conv' conv'_Suc_right using False + by (subst conv'_eq_iff) (auto simp: field_simps) + moreover from assms have "conv c n \ conv c (Suc (Suc n))" + by (intro conv_even_mono) auto + ultimately have "conv c n < conv c (Suc (Suc n))" by simp + also have "\ \ conv c m" using assms \m \ n + 1\ + by (intro conv_even_mono) auto + finally show ?thesis . + qed +qed + +lemma conv_odd_mono_strict: + assumes "odd n" "n < m" + shows "conv c n > conv c m" +proof (cases "m = n + 1") + case [simp]: True + from assms have "n > 0" by (intro Nat.gr0I) auto + hence "conv' c n (real_of_int (cfrac_nth c n)) \ + conv' c n (real_of_int (cfrac_nth c n) + 1 / real_of_int (cfrac_nth c (Suc n)))" + by (subst conv'_eq_iff) auto + hence "conv c n \ conv c m" + by (simp add: conv_eq_conv' conv'_Suc_right) + moreover from assms have "conv c n \ conv c m" + by (intro conv_odd_mono) auto + ultimately show ?thesis by simp +next + case False + from assms have "n > 0" by (intro Nat.gr0I) auto + have "1 + real_of_int (cfrac_nth c (n+1)) * cfrac_nth c (n+2) > 0" + by (intro add_pos_nonneg) auto + with assms \n > 0\ have "conv c n \ conv c (Suc (Suc n))" + unfolding conv_eq_conv' conv'_Suc_right + by (subst conv'_eq_iff) (auto simp: field_simps) + moreover from assms have "conv c n \ conv c (Suc (Suc n))" + by (intro conv_odd_mono) auto + ultimately have "conv c n > conv c (Suc (Suc n))" by simp + moreover have "conv c (Suc (Suc n)) \ conv c m" using assms False + by (intro conv_odd_mono) auto + ultimately show ?thesis by linarith +qed + +lemma conv_less_cfrac_lim: + assumes "even n" "n < cfrac_length c" + shows "conv c n < cfrac_lim c" +proof (cases "cfrac_length c") + case (enat l) + with assms show ?thesis by (auto simp: cfrac_lim_def conv_even_mono_strict) +next + case [simp]: infinity + from assms have "conv c n < conv c (n + 2)" + by (intro conv_even_mono_strict) auto + also from assms have "\ \ cfrac_lim c" + by (intro conv_le_cfrac_lim) auto + finally show ?thesis . +qed + +lemma conv_gt_cfrac_lim: + assumes "odd n" "n < cfrac_length c" + shows "conv c n > cfrac_lim c" +proof (cases "cfrac_length c") + case (enat l) + with assms show ?thesis by (auto simp: cfrac_lim_def conv_odd_mono_strict) +next + case [simp]: infinity + from assms have "cfrac_lim c \ conv c (n + 2)" + by (intro conv_ge_cfrac_lim) auto + also from assms have "\ < conv c n" + by (intro conv_odd_mono_strict) auto + finally show ?thesis . +qed + +lemma conv_neq_cfrac_lim: + assumes "n < cfrac_length c" + shows "conv c n \ cfrac_lim c" + using conv_gt_cfrac_lim[OF _ assms] conv_less_cfrac_lim[OF _ assms] + by (cases "even n") auto + +lemma conv_ge_first: "conv c n \ cfrac_nth c 0" + using conv_even_mono[of 0 n c] by simp + + +definition cfrac_is_zero :: "cfrac \ bool" where "cfrac_is_zero c \ c = 0" + +lemma cfrac_is_zero_code [code]: "cfrac_is_zero (CFrac n xs) \ lnull xs \ n = 0" + unfolding cfrac_is_zero_def lnull_def zero_cfrac_def cfrac_of_int_def + by (auto simp: cfrac_length_def) + + +definition cfrac_is_int where "cfrac_is_int c \ cfrac_length c = 0" + +lemma cfrac_is_int_code [code]: "cfrac_is_int (CFrac n xs) \ lnull xs" + unfolding cfrac_is_int_def lnull_def by (auto simp: cfrac_length_def) + +lemma cfrac_length_of_int [simp]: "cfrac_length (cfrac_of_int n) = 0" + by transfer auto + +lemma cfrac_is_int_of_int [simp, intro]: "cfrac_is_int (cfrac_of_int n)" + unfolding cfrac_is_int_def by simp + +lemma cfrac_is_int_iff: "cfrac_is_int c \ (\n. c = cfrac_of_int n)" +proof - + have "c = cfrac_of_int (cfrac_nth c 0)" if "cfrac_is_int c" + using that unfolding cfrac_is_int_def by transfer auto + thus ?thesis + by auto +qed + +lemma cfrac_lim_reduce: + assumes "\cfrac_is_int c" + shows "cfrac_lim c = cfrac_nth c 0 + 1 / cfrac_lim (cfrac_tl c)" +proof (cases "cfrac_length c") + case [simp]: infinity + have "0 < cfrac_nth (cfrac_tl c) 0" + by simp + also have "\ \ cfrac_lim (cfrac_tl c)" + by (rule cfrac_lim_ge_first) + finally have "(\n. real_of_int (cfrac_nth c 0) + 1 / conv (cfrac_tl c) n) \ + real_of_int (cfrac_nth c 0) + 1 / cfrac_lim (cfrac_tl c)" + by (intro tendsto_intros LIMSEQ_cfrac_lim) auto + also have "(\n. real_of_int (cfrac_nth c 0) + 1 / conv (cfrac_tl c) n) = conv c \ Suc" + by (simp add: o_def conv_Suc) + finally have *: "conv c \ real_of_int (cfrac_nth c 0) + 1 / cfrac_lim (cfrac_tl c)" + by (simp add: o_def filterlim_sequentially_Suc) + show ?thesis + by (rule tendsto_unique[OF _ LIMSEQ_cfrac_lim *]) auto +next + case [simp]: (enat l) + from assms obtain l' where [simp]: "l = Suc l'" + by (cases l) (auto simp: cfrac_is_int_def zero_enat_def) + thus ?thesis + by (auto simp: cfrac_lim_def conv_Suc) +qed + +lemma cfrac_lim_tl: + assumes "\cfrac_is_int c" + shows "cfrac_lim (cfrac_tl c) = 1 / (cfrac_lim c - cfrac_nth c 0)" + using cfrac_lim_reduce[OF assms] by simp + + + +lemma cfrac_remainder_Suc': + assumes "n < cfrac_length c" + shows "cfrac_remainder c (Suc n) * (cfrac_remainder c n - cfrac_nth c n) = 1" +proof - + have "0 < real_of_int (cfrac_nth c (Suc n))" by simp + also have "cfrac_nth c (Suc n) \ cfrac_remainder c (Suc n)" + using cfrac_lim_ge_first[of "cfrac_drop (Suc n) c"] + by (simp add: cfrac_remainder_def) + finally have "\ > 0" . + + have "cfrac_remainder c (Suc n) = cfrac_lim (cfrac_tl (cfrac_drop n c))" + by (simp add: o_def cfrac_remainder_def cfrac_drop_Suc_left) + also have "\ = 1 / (cfrac_remainder c n - cfrac_nth c n)" using assms + by (subst cfrac_lim_tl) (auto simp: cfrac_remainder_def cfrac_is_int_def enat_less_iff enat_0_iff) + finally show ?thesis + using \cfrac_remainder c (Suc n) > 0\ + by (auto simp add: cfrac_remainder_def field_simps) +qed + +lemma cfrac_remainder_Suc: + assumes "n < cfrac_length c" + shows "cfrac_remainder c (Suc n) = 1 / (cfrac_remainder c n - cfrac_nth c n)" +proof - + have "cfrac_remainder c (Suc n) = cfrac_lim (cfrac_tl (cfrac_drop n c))" + by (simp add: o_def cfrac_remainder_def cfrac_drop_Suc_left) + also have "\ = 1 / (cfrac_remainder c n - cfrac_nth c n)" using assms + by (subst cfrac_lim_tl) (auto simp: cfrac_remainder_def cfrac_is_int_def enat_less_iff enat_0_iff) + finally show ?thesis . +qed + +lemma cfrac_remainder_0 [simp]: "cfrac_remainder c 0 = cfrac_lim c" + by (simp add: cfrac_remainder_def) + +context + fixes c h k x + defines "h \ conv_num c" and "k \ conv_denom c" and "x \ cfrac_remainder c" +begin + +lemma cfrac_lim_eq_num_denom_remainder_aux: + assumes "Suc (Suc n) \ cfrac_length c" + shows "cfrac_lim c * (k (Suc n) * x (Suc (Suc n)) + k n) = h (Suc n) * x (Suc (Suc n)) + h n" + using assms +proof (induction n) + case 0 + have "cfrac_lim c \ cfrac_nth c 0" + using conv_neq_cfrac_lim[of 0 c] 0 by (auto simp: enat_le_iff) + moreover have "cfrac_nth c 1 * (cfrac_lim c - cfrac_nth c 0) \ 1" + using conv_neq_cfrac_lim[of 1 c] 0 + by (auto simp: enat_le_iff conv_Suc field_simps) + ultimately show ?case using assms + by (auto simp: cfrac_remainder_Suc divide_simps x_def h_def k_def enat_le_iff) + (auto simp: field_simps) +next + case (Suc n) + have less: "enat (Suc (Suc n)) < cfrac_length c" + using Suc.prems by (cases "cfrac_length c") auto + have *: "x (Suc (Suc n)) \ real_of_int (cfrac_nth c (Suc (Suc n)))" + using conv_neq_cfrac_lim[of 0 "cfrac_drop (n+2) c"] Suc.prems + by (cases "cfrac_length c") (auto simp: x_def cfrac_remainder_def) + hence "cfrac_lim c * (k (Suc (Suc n)) * x (Suc (Suc (Suc n))) + k (Suc n)) = + (cfrac_lim c * (k (Suc n) * x (Suc (Suc n)) + k n)) / (x (Suc (Suc n)) - cfrac_nth c (Suc (Suc n)))" + unfolding x_def k_def h_def using less + by (subst cfrac_remainder_Suc) (auto simp: field_simps) + also have "cfrac_lim c * (k (Suc n) * x (Suc (Suc n)) + k n) = + h (Suc n) * x (Suc (Suc n)) + h n" using less + by (intro Suc.IH) auto + also have "(h (Suc n) * x (Suc (Suc n)) + h n) / (x (Suc (Suc n)) - cfrac_nth c (Suc (Suc n))) = + h (Suc (Suc n)) * x (Suc (Suc (Suc n))) + h (Suc n)" using * + unfolding x_def k_def h_def using less + by (subst (3) cfrac_remainder_Suc) (auto simp: field_simps) + finally show ?case . +qed + +lemma cfrac_remainder_nonneg: "cfrac_nth c n \ 0 \ cfrac_remainder c n \ 0" + unfolding cfrac_remainder_def by (rule cfrac_lim_nonneg) auto + +lemma cfrac_remainder_pos: "cfrac_nth c n > 0 \ cfrac_remainder c n > 0" + unfolding cfrac_remainder_def by (rule cfrac_lim_pos) auto + +lemma cfrac_lim_eq_num_denom_remainder: + assumes "Suc (Suc n) < cfrac_length c" + shows "cfrac_lim c = (h (Suc n) * x (Suc (Suc n)) + h n) / (k (Suc n) * x (Suc (Suc n)) + k n)" +proof - + have "k (Suc n) * x (Suc (Suc n)) + k n > 0" + by (intro add_nonneg_pos mult_nonneg_nonneg) + (auto simp: k_def x_def intro!: conv_denom_pos cfrac_remainder_nonneg) + with cfrac_lim_eq_num_denom_remainder_aux[of n] assms show ?thesis + by (auto simp add: field_simps h_def k_def x_def) +qed + +lemma abs_diff_successive_convs: + shows "\conv c (Suc n) - conv c n\ = 1 / (k n * k (Suc n))" +proof - + have [simp]: "k n \ 0" for n :: nat + unfolding k_def using conv_denom_pos[of c n] by auto + have "conv c (Suc n) - conv c n = h (Suc n) / k (Suc n) - h n / k n" + by (simp add: conv_num_denom k_def h_def) + also have "\ = (k n * h (Suc n) - k (Suc n) * h n) / (k n * k (Suc n))" + by (simp add: field_simps) + also have "k n * h (Suc n) - k (Suc n) * h n = (-1) ^ n" + unfolding h_def k_def by (intro conv_num_denom_prod_diff) + finally show ?thesis by (simp add: k_def) +qed + +lemma conv_denom_plus2_ratio_ge: "k (Suc (Suc n)) \ 2 * k n" +proof - + have "1 * k n + k n \ cfrac_nth c (Suc (Suc n)) * k (Suc n) + k n" + by (intro add_mono mult_mono) + (auto simp: k_def Suc_le_eq intro!: conv_denom_leI) + thus ?thesis by (simp add: k_def) +qed + +end + +lemma conv'_cfrac_remainder: + assumes "n < cfrac_length c" + shows "conv' c n (cfrac_remainder c n) = cfrac_lim c" + using assms +proof (induction n arbitrary: c) + case (Suc n c) + have "conv' c (Suc n) (cfrac_remainder c (Suc n)) = + cfrac_nth c 0 + 1 / conv' (cfrac_tl c) n (cfrac_remainder c (Suc n))" + using Suc.prems + by (subst conv'_Suc_left) (auto intro!: cfrac_remainder_pos) + also have "cfrac_remainder c (Suc n) = cfrac_remainder (cfrac_tl c) n" + by (simp add: cfrac_remainder_def cfrac_drop_Suc_right) + also have "conv' (cfrac_tl c) n \ = cfrac_lim (cfrac_tl c)" + using Suc.prems by (subst Suc.IH) (auto simp: cfrac_remainder_def enat_less_iff) + also have "cfrac_nth c 0 + 1 / \ = cfrac_lim c" + using Suc.prems by (intro cfrac_lim_reduce [symmetric]) (auto simp: cfrac_is_int_def) + finally show ?case by (simp add: cfrac_remainder_def cfrac_drop_Suc_right) +qed auto + +lemma cfrac_lim_rational [intro]: + assumes "cfrac_length c < \" + shows "cfrac_lim c \ \" + using assms by (cases "cfrac_length c") (auto simp: cfrac_lim_def) + +lemma linfinite_cfrac_of_real_aux: + "x \ \ \ x \ {0<..<1} \ linfinite (cfrac_of_real_aux x)" +proof (coinduction arbitrary: x) + case (linfinite x) + hence "1 / x \ \" using Rats_divide[of 1 "1 / x"] by auto + thus ?case using linfinite Ints_subset_Rats + by (intro disjI1 exI[of _ "nat \1/x\ - 1"] exI[of _ "cfrac_of_real_aux (frac (1/x))"] + exI[of _ "frac (1/x)"] conjI) + (auto simp: cfrac_of_real_aux.code[of x] frac_lt_1) +qed + +lemma cfrac_length_of_real_irrational: + assumes "x \ \" + shows "cfrac_length (cfrac_of_real x) = \" +proof (insert assms, transfer, clarify) + fix x :: real assume "x \ \" + thus "llength (cfrac_of_real_aux (frac x)) = \" + using linfinite_cfrac_of_real_aux[of "frac x"] Ints_subset_Rats + by (auto simp: linfinite_conv_llength frac_lt_1) +qed + +lemma cfrac_length_of_real_reduce: + assumes "x \ \" + shows "cfrac_length (cfrac_of_real x) = eSuc (cfrac_length (cfrac_of_real (1 / frac x)))" + using assms + by (transfer, subst cfrac_of_real_aux.code) (auto simp: frac_lt_1) + +lemma cfrac_length_of_real_int [simp]: "x \ \ \ cfrac_length (cfrac_of_real x) = 0" + by transfer auto + +lemma conv_cfrac_of_real_le_ge: + assumes "n \ cfrac_length (cfrac_of_real x)" + shows "if even n then conv (cfrac_of_real x) n \ x else conv (cfrac_of_real x) n \ x" + using assms +proof (induction n arbitrary: x) + case (Suc n x) + hence [simp]: "x \ \" + using Suc by (auto simp: enat_0_iff) + let ?x' = "1 / frac x" + have "enat n \ cfrac_length (cfrac_of_real (1 / frac x))" + using Suc.prems by (auto simp: cfrac_length_of_real_reduce simp flip: eSuc_enat) + hence IH: "if even n then conv (cfrac_of_real ?x') n \ ?x' else ?x' \ conv (cfrac_of_real ?x') n" + using Suc.prems by (intro Suc.IH) auto + have remainder_pos: "conv (cfrac_of_real ?x') n > 0" + by (rule conv_pos) (auto simp: frac_le_1) + show ?case + proof (cases "even n") + case True + have "x \ real_of_int \x\ + frac x" + by (simp add: frac_def) + also have "frac x \ 1 / conv (cfrac_of_real ?x') n" + using IH True remainder_pos frac_gt_0_iff[of x] by (simp add: field_simps) + finally show ?thesis using True + by (auto simp: conv_Suc cfrac_tl_of_real) + next + case False + have "real_of_int \x\ + 1 / conv (cfrac_of_real ?x') n \ real_of_int \x\ + frac x" + using IH False remainder_pos frac_gt_0_iff[of x] by (simp add: field_simps) + also have "\ = x" + by (simp add: frac_def) + finally show ?thesis using False + by (auto simp: conv_Suc cfrac_tl_of_real) + qed +qed auto + +lemma cfrac_lim_of_real [simp]: "cfrac_lim (cfrac_of_real x) = x" +proof (cases "cfrac_length (cfrac_of_real x)") + case (enat l) + hence "conv (cfrac_of_real x) l = x" + proof (induction l arbitrary: x) + case 0 + hence "x \ \" + using cfrac_length_of_real_reduce zero_enat_def by fastforce + thus ?case by (auto elim: Ints_cases) + next + case (Suc l x) + hence [simp]: "x \ \" + by (auto simp: enat_0_iff) + have "eSuc (cfrac_length (cfrac_of_real (1 / frac x))) = enat (Suc l)" + using Suc.prems by (auto simp: cfrac_length_of_real_reduce) + hence "conv (cfrac_of_real (1 / frac x)) l = 1 / frac x" + by (intro Suc.IH) (auto simp flip: eSuc_enat) + thus ?case + by (simp add: conv_Suc cfrac_tl_of_real frac_def) + qed + thus ?thesis by (simp add: enat cfrac_lim_def) +next + case [simp]: infinity + have lim: "conv (cfrac_of_real x) \ cfrac_lim (cfrac_of_real x)" + by (simp add: LIMSEQ_cfrac_lim) + have "cfrac_lim (cfrac_of_real x) \ x" + proof (rule tendsto_upperbound) + show "(\n. conv (cfrac_of_real x) (n * 2)) \ cfrac_lim (cfrac_of_real x)" + by (intro filterlim_compose[OF lim] mult_nat_right_at_top) auto + show "eventually (\n. conv (cfrac_of_real x) (n * 2) \ x) at_top" + using conv_cfrac_of_real_le_ge[of "n * 2" x for n] by (intro always_eventually) auto + qed auto + moreover have "cfrac_lim (cfrac_of_real x) \ x" + proof (rule tendsto_lowerbound) + show "(\n. conv (cfrac_of_real x) (Suc (n * 2))) \ cfrac_lim (cfrac_of_real x)" + by (intro filterlim_compose[OF lim] filterlim_compose[OF filterlim_Suc] + mult_nat_right_at_top) auto + show "eventually (\n. conv (cfrac_of_real x) (Suc (n * 2)) \ x) at_top" + using conv_cfrac_of_real_le_ge[of "Suc (n * 2)" x for n] by (intro always_eventually) auto + qed auto + ultimately show ?thesis by (rule antisym) +qed + +lemma Ints_add_left_cancel: "x \ \ \ x + y \ \ \ y \ \" + using Ints_diff[of "x + y" x] by auto + +lemma Ints_add_right_cancel: "y \ \ \ x + y \ \ \ x \ \" + using Ints_diff[of "x + y" y] by auto + +lemma cfrac_of_real_conv': + fixes m n :: nat + assumes "x > 1" "m < n" + shows "cfrac_nth (cfrac_of_real (conv' c n x)) m = cfrac_nth c m" + using assms +proof (induction n arbitrary: c m) + case (Suc n c m) + from Suc.prems have gt_1: "1 < conv' (cfrac_tl c) n x" + by (intro conv'_gt_1) (auto simp: enat_le_iff intro: cfrac_nth_pos) + show ?case + proof (cases m) + case 0 + thus ?thesis using gt_1 Suc.prems + by (simp add: conv'_Suc_left nat_add_distrib floor_eq_iff) + next + case (Suc m') + from gt_1 have "1 / conv' (cfrac_tl c) n x \ {0<..<1}" + by auto + have "1 / conv' (cfrac_tl c) n x \ \" + proof + assume "1 / conv' (cfrac_tl c) n x \ \" + then obtain k :: int where k: "1 / conv' (cfrac_tl c) n x = of_int k" + by (elim Ints_cases) + have "real_of_int k \ {0<..<1}" + using gt_1 by (subst k [symmetric]) auto + thus False by auto + qed + hence not_int: "real_of_int (cfrac_nth c 0) + 1 / conv' (cfrac_tl c) n x \ \" + by (subst Ints_add_left_cancel) (auto simp: field_simps elim!: Ints_cases) + have "cfrac_nth (cfrac_of_real (conv' c (Suc n) x)) m = + cfrac_nth (cfrac_of_real (of_int (cfrac_nth c 0) + 1 / conv' (cfrac_tl c) n x)) (Suc m')" + using \x > 1\ by (subst conv'_Suc_left) (auto simp: Suc) + also have "\ = cfrac_nth (cfrac_of_real (1 / frac (1 / conv' (cfrac_tl c) n x))) m'" + using \x > 1\ Suc not_int by (subst cfrac_nth_of_real_Suc) (auto simp: frac_add_of_int) + also have "1 / conv' (cfrac_tl c) n x \ {0<..<1}" using gt_1 + by (auto simp: field_simps) + hence "frac (1 / conv' (cfrac_tl c) n x) = 1 / conv' (cfrac_tl c) n x" + by (subst frac_eq) auto + hence "1 / frac (1 / conv' (cfrac_tl c) n x) = conv' (cfrac_tl c) n x" + by simp + also have "cfrac_nth (cfrac_of_real \) m' = cfrac_nth c m" + using Suc.prems by (subst Suc.IH) (auto simp: Suc enat_le_iff) + finally show ?thesis . + qed +qed simp_all + +lemma cfrac_lim_irrational: + assumes [simp]: "cfrac_length c = \" + shows "cfrac_lim c \ \" +proof + assume "cfrac_lim c \ \" + then obtain a :: int and b :: nat where ab: "b > 0" "cfrac_lim c = a / b" + by (auto simp: Rats_eq_int_div_nat) + define h and k where "h = conv_num c" and "k = conv_denom c" + + have "filterlim (\m. conv_denom c (Suc m)) at_top at_top" + using conv_denom_at_top filterlim_Suc by (rule filterlim_compose) + then obtain m where m: "conv_denom c (Suc m) \ b + 1" + by (auto simp: filterlim_at_top eventually_at_top_linorder) + + have *: "(a * k m - b * h m) / (k m * b) = a / b - h m / k m" + using \b > 0\ by (simp add: field_simps k_def) + have "\cfrac_lim c - conv c m\ = \(a * k m - b * h m) / (k m * b)\" + by (subst *) (auto simp: ab h_def k_def conv_num_denom) + also have "\ = \a * k m - b * h m\ / (k m * b)" + by (simp add: k_def) + finally have eq: "\cfrac_lim c - conv c m\ = of_int \a * k m - b * h m\ / of_int (k m * b)" . + + have "\cfrac_lim c - conv c m\ * (k m * b) \ 0" + using conv_neq_cfrac_lim[of m c] \b > 0\ by (auto simp: k_def) + also have "\cfrac_lim c - conv c m\ * (k m * b) = of_int \a * k m - b * h m\" + using \b > 0\ by (subst eq) (auto simp: k_def) + finally have "\a * k m - b * h m\ \ 1" by linarith + hence "real_of_int \a * k m - b * h m\ \ 1" by linarith + hence "1 / of_int (k m * b) \ of_int \a * k m - b * h m\ / real_of_int (k m * b)" + using \b > 0\ by (intro divide_right_mono) (auto simp: k_def) + also have "\ = \cfrac_lim c - conv c m\" + by (rule eq [symmetric]) + also have "\ \ 1 / real_of_int (conv_denom c m * conv_denom c (Suc m))" + by (intro cfrac_lim_minus_conv_upper_bound) auto + also have "\ = 1 / (real_of_int (k m) * real_of_int (k (Suc m)))" + by (simp add: k_def) + also have "\ < 1 / (real_of_int (k m) * real b)" + using m \b > 0\ + by (intro divide_strict_left_mono mult_strict_left_mono) (auto simp: k_def) + finally show False by simp +qed + +lemma cfrac_infinite_iff: "cfrac_length c = \ \ cfrac_lim c \ \" + using cfrac_lim_irrational[of c] cfrac_lim_rational[of c] by auto + +lemma cfrac_lim_rational_iff: "cfrac_lim c \ \ \ cfrac_length c \ \" + using cfrac_lim_irrational[of c] cfrac_lim_rational[of c] by auto + +lemma cfrac_of_real_infinite_iff [simp]: "cfrac_length (cfrac_of_real x) = \ \ x \ \" + by (simp add: cfrac_infinite_iff) + +lemma cfrac_remainder_rational_iff [simp]: + "cfrac_remainder c n \ \ \ cfrac_length c < \" +proof - + have "cfrac_remainder c n \ \ \ cfrac_lim (cfrac_drop n c) \ \" + by (simp add: cfrac_remainder_def) + also have "\ \ cfrac_length c \ \" + by (cases "cfrac_length c") (auto simp add: cfrac_lim_rational_iff) + finally show ?thesis by simp +qed + +lift_definition cfrac_cons :: "int \ cfrac \ cfrac" is + "\a bs. case bs of (b, bs) \ if b \ 0 then (1, LNil) else (a, LCons (nat (b - 1)) bs)" . + +lemma cfrac_nth_cons: + assumes "cfrac_nth x 0 \ 1" + shows "cfrac_nth (cfrac_cons a x) n = (if n = 0 then a else cfrac_nth x (n - 1))" + using assms +proof (transfer, goal_cases) + case (1 x a n) + obtain b bs where [simp]: "x = (b, bs)" + by (cases x) + show ?case using 1 + by (cases "llength bs") (auto simp: lnth_LCons eSuc_enat le_imp_diff_is_add split: nat.splits) +qed + +lemma cfrac_length_cons [simp]: + assumes "cfrac_nth x 0 \ 1" + shows "cfrac_length (cfrac_cons a x) = eSuc (cfrac_length x)" + using assms by transfer auto + +lemma cfrac_tl_cons [simp]: + assumes "cfrac_nth x 0 \ 1" + shows "cfrac_tl (cfrac_cons a x) = x" + using assms by transfer auto + +lemma cfrac_cons_tl: + assumes "\cfrac_is_int x" + shows "cfrac_cons (cfrac_nth x 0) (cfrac_tl x) = x" + using assms unfolding cfrac_is_int_def + by transfer (auto split: llist.splits) + + +subsection \Non-canonical continued fractions\ + +text \ + As we will show later, every irrational number has a unique continued fraction + expansion. Every rational number \x\, however, has two different expansions: + The canonical one ends with some number \n\ (which is not equal to 1 unless \x = 1\) + and a non-canonical one which ends with $n-1, 1$. + + We now define this non-canonical expansion analogously to the canonical one before + and show its characteristic properties: + + \<^item> The length of the non-canonical expansion is one greater than that of the + canonical one. + + \<^item> If the expansion is infinite, the non-canonical and the canonical one coincide. + + \<^item> The coefficients of the expansions are all equal except for the last two. + The last coefficient of the non-canonical expansion is always 1, and the + second to last one is the last of the canonical one minus 1. +\ + +lift_definition cfrac_canonical :: "cfrac \ bool" is + "\(x, xs). \lfinite xs \ lnull xs \ llast xs \ 0" . + +lemma cfrac_canonical [code]: + "cfrac_canonical (CFrac x xs) \ lnull xs \ llast xs \ 0 \ \lfinite xs" + by (auto simp add: cfrac_canonical_def) + +lemma cfrac_canonical_iff: + "cfrac_canonical c \ + cfrac_length c \ {0, \} \ cfrac_nth c (the_enat (cfrac_length c)) \ 1" +proof (transfer, clarify, goal_cases) + case (1 x xs) + show ?case + by (cases "llength xs") + (auto simp: llast_def enat_0 lfinite_conv_llength_enat split: nat.splits) +qed + +lemma llast_cfrac_of_real_aux_nonzero: + assumes "lfinite (cfrac_of_real_aux x)" "\lnull (cfrac_of_real_aux x)" + shows "llast (cfrac_of_real_aux x) \ 0" + using assms +proof (induction "cfrac_of_real_aux x" arbitrary: x rule: lfinite_induct) + case (LCons x) + from LCons.prems have "x \ {0<..<1}" + by (subst (asm) cfrac_of_real_aux.code) (auto split: if_splits) + show ?case + proof (cases "1 / x \ \") + case False + thus ?thesis using LCons + by (auto simp: llast_LCons frac_lt_1 cfrac_of_real_aux.code[of x]) + next + case True + then obtain n where n: "1 / x = of_int n" + by (elim Ints_cases) + have "1 / x > 1" using \x \ _\ by auto + with n have "n > 1" by simp + from n have "x = 1 / of_int n" + using \n > 1\ \x \ _\ by (simp add: field_simps) + with \n > 1\ show ?thesis + using LCons cfrac_of_real_aux.code[of x] by (auto simp: llast_LCons frac_lt_1) + qed +qed auto + +lemma cfrac_canonical_of_real [intro]: "cfrac_canonical (cfrac_of_real x)" + by (transfer fixing: x) (use llast_cfrac_of_real_aux_nonzero[of "frac x"] in force) + +primcorec cfrac_of_real_alt_aux :: "real \ nat llist" where + "cfrac_of_real_alt_aux x = + (if x \ {0<..<1} then + if 1 / x \ \ then + LCons (nat \1/x\ - 2) (LCons 0 LNil) + else LCons (nat \1/x\ - 1) (cfrac_of_real_alt_aux (frac (1/x))) + else LNil)" + +lemma cfrac_of_real_aux_alt_LNil [simp]: "x \ {0<..<1} \ cfrac_of_real_alt_aux x = LNil" + by (subst cfrac_of_real_alt_aux.code) auto + +lemma cfrac_of_real_aux_alt_0 [simp]: "cfrac_of_real_alt_aux 0 = LNil" + by (subst cfrac_of_real_alt_aux.code) auto + +lemma cfrac_of_real_aux_alt_eq_LNil_iff [simp]: "cfrac_of_real_alt_aux x = LNil \ x \ {0<..<1}" + by (subst cfrac_of_real_alt_aux.code) auto + +lift_definition cfrac_of_real_alt :: "real \ cfrac" is + "\x. if x \ \ then (\x\ - 1, LCons 0 LNil) else (\x\, cfrac_of_real_alt_aux (frac x))" . + +lemma cfrac_tl_of_real_alt: + assumes "x \ \" + shows "cfrac_tl (cfrac_of_real_alt x) = cfrac_of_real_alt (1 / frac x)" + using assms +proof (transfer, goal_cases) + case (1 x) + show ?case + proof (cases "1 / frac x \ \") + case False + from 1 have "int (nat \1 / frac x\ - Suc 0) + 1 = \1 / frac x\" + by (subst of_nat_diff) (auto simp: le_nat_iff frac_le_1) + with False show ?thesis + using \x \ \\ + by (subst cfrac_of_real_alt_aux.code) (auto split: llist.splits simp: frac_lt_1) + next + case True + then obtain n where "1 / frac x = of_int n" + by (auto simp: Ints_def) + moreover have "1 / frac x > 1" + using 1 by (auto simp: divide_simps frac_lt_1) + ultimately have "1 / frac x \ 2" + by simp + hence "int (nat \1 / frac x\ - 2) + 2 = \1 / frac x\" + by (subst of_nat_diff) (auto simp: le_nat_iff frac_le_1) + thus ?thesis + using \x \ \\ + by (subst cfrac_of_real_alt_aux.code) (auto split: llist.splits simp: frac_lt_1) + qed +qed + +lemma cfrac_nth_of_real_alt_Suc: + assumes "x \ \" + shows "cfrac_nth (cfrac_of_real_alt x) (Suc n) = cfrac_nth (cfrac_of_real_alt (1 / frac x)) n" +proof - + have "cfrac_nth (cfrac_of_real_alt x) (Suc n) = + cfrac_nth (cfrac_tl (cfrac_of_real_alt x)) n" + by simp + also have "cfrac_tl (cfrac_of_real_alt x) = cfrac_of_real_alt (1 / frac x)" + by (simp add: cfrac_tl_of_real_alt assms) + finally show ?thesis . +qed + +lemma cfrac_nth_gt0_of_real_int [simp]: + "m > 0 \ cfrac_nth (cfrac_of_real (of_int n)) m = 1" + by transfer (auto simp: lnth_LCons eSuc_def enat_0_iff split: nat.splits) + +lemma cfrac_nth_0_of_real_alt_int [simp]: + "cfrac_nth (cfrac_of_real_alt (of_int n)) 0 = n - 1" + by transfer auto + +lemma cfrac_nth_gt0_of_real_alt_int [simp]: + "m > 0 \ cfrac_nth (cfrac_of_real_alt (of_int n)) m = 1" + by transfer (auto simp: lnth_LCons eSuc_def split: nat.splits) + +lemma llength_cfrac_of_real_alt_aux: + assumes "x \ {0<..<1}" + shows "llength (cfrac_of_real_alt_aux x) = eSuc (llength (cfrac_of_real_aux x))" + using assms +proof (coinduction arbitrary: x rule: enat_coinduct) + case (Eq_enat x) + show ?case + proof (cases "1 / x \ \") + case False + with Eq_enat have "frac (1 / x) \ {0<..<1}" + by (auto intro: frac_lt_1) + hence "\x'. llength (cfrac_of_real_alt_aux (frac (1 / x))) = + llength (cfrac_of_real_alt_aux x') \ + llength (cfrac_of_real_aux (frac (1 / x))) = llength (cfrac_of_real_aux x') \ + 0 < x' \ x' < 1" + by (intro exI[of _ "frac (1 / x)"]) auto + thus ?thesis using False Eq_enat + by (auto simp: cfrac_of_real_alt_aux.code[of x] cfrac_of_real_aux.code[of x]) + qed (use Eq_enat in \auto simp: cfrac_of_real_alt_aux.code[of x] cfrac_of_real_aux.code[of x]\) +qed + +lemma cfrac_length_of_real_alt: + "cfrac_length (cfrac_of_real_alt x) = eSuc (cfrac_length (cfrac_of_real x))" + by transfer (auto simp: llength_cfrac_of_real_alt_aux frac_lt_1) + +lemma cfrac_of_real_alt_aux_eq_regular: + assumes "x \ {0<..<1}" "llength (cfrac_of_real_aux x) = \" + shows "cfrac_of_real_alt_aux x = cfrac_of_real_aux x" + using assms +proof (coinduction arbitrary: x) + case (Eq_llist x) + hence "\x'. cfrac_of_real_aux (frac (1 / x)) = + cfrac_of_real_aux x' \ + cfrac_of_real_alt_aux (frac (1 / x)) = + cfrac_of_real_alt_aux x' \ 0 < x' \ x' < 1 \ llength (cfrac_of_real_aux x') = \" + by (intro exI[of _ "frac (1 / x)"]) + (auto simp: cfrac_of_real_aux.code[of x] cfrac_of_real_alt_aux.code[of x] + eSuc_eq_infinity_iff frac_lt_1) + with Eq_llist show ?case + by (auto simp: eSuc_eq_infinity_iff) +qed + +lemma cfrac_of_real_alt_irrational [simp]: + assumes "x \ \" + shows "cfrac_of_real_alt x = cfrac_of_real x" +proof - + from assms have "cfrac_length (cfrac_of_real x) = \" + using cfrac_length_of_real_irrational by blast + with assms show ?thesis + by transfer + (use Ints_subset_Rats in + \auto intro!: cfrac_of_real_alt_aux_eq_regular simp: frac_lt_1 llength_cfrac_of_real_alt_aux\) +qed + +lemma cfrac_nth_of_real_alt_0: + "cfrac_nth (cfrac_of_real_alt x) 0 = (if x \ \ then \x\ - 1 else \x\)" + by transfer auto + +lemma cfrac_nth_of_real_alt: + fixes n :: nat and x :: real + defines "c \ cfrac_of_real x" + defines "c' \ cfrac_of_real_alt x" + defines "l \ cfrac_length c" + shows "cfrac_nth c' n = + (if enat n = l then + cfrac_nth c n - 1 + else if enat n = l + 1 then + 1 + else + cfrac_nth c n)" + unfolding c_def c'_def l_def +proof (induction n arbitrary: x rule: less_induct) + case (less n) + consider "x \ \" | "x \ \" | "n = 0" "x \ \ - \" | n' where "n = Suc n'" "x \ \ - \" + by (cases n) auto + thus ?case + proof cases + assume "x \ \" + thus ?thesis + by (auto simp: cfrac_length_of_real_irrational) + next + assume "x \ \" + thus ?thesis + by (auto simp: Ints_def one_enat_def zero_enat_def) + next + assume *: "n = 0" "x \ \ - \" + have "enat 0 \ cfrac_length (cfrac_of_real x) + 1" + using zero_enat_def by auto + moreover have "enat 0 \ cfrac_length (cfrac_of_real x)" + using * cfrac_length_of_real_reduce zero_enat_def by auto + ultimately show ?thesis using * + by (auto simp: cfrac_nth_of_real_alt_0) + next + fix n' assume *: "n = Suc n'" "x \ \ - \" + from less.IH [of n' "1 / frac x"] and * show ?thesis + by (auto simp: cfrac_nth_of_real_Suc cfrac_nth_of_real_alt_Suc cfrac_length_of_real_reduce + eSuc_def one_enat_def enat_0_iff split: enat.splits) + qed +qed + +lemma cfrac_of_real_length_eq_0_iff: "cfrac_length (cfrac_of_real x) = 0 \ x \ \" + by transfer (auto simp: frac_lt_1) + +lemma conv'_cong: + assumes "(\k. k < n \ cfrac_nth c k = cfrac_nth c' k)" "n = n'" "x = y" + shows "conv' c n x = conv' c' n' y" + using assms(1) unfolding assms(2,3) [symmetric] + by (induction n arbitrary: x) (auto simp: conv'_Suc_right) + +lemma conv_cong: + assumes "(\k. k \ n \ cfrac_nth c k = cfrac_nth c' k)" "n = n'" + shows "conv c n = conv c' n'" + using assms(1) unfolding assms(2) [symmetric] + by (induction n arbitrary: c c') (auto simp: conv_Suc) + +lemma conv'_cfrac_of_real_alt: + assumes "enat n \ cfrac_length (cfrac_of_real x)" + shows "conv' (cfrac_of_real_alt x) n y = conv' (cfrac_of_real x) n y" +proof (cases "cfrac_length (cfrac_of_real x)") + case infinity + thus ?thesis by auto +next + case [simp]: (enat l') + with assms show ?thesis + by (intro conv'_cong refl; subst cfrac_nth_of_real_alt) (auto simp: one_enat_def) +qed + +lemma cfrac_lim_of_real_alt [simp]: "cfrac_lim (cfrac_of_real_alt x) = x" +proof (cases "cfrac_length (cfrac_of_real x)") + case infinity + thus ?thesis by auto +next + case (enat l) + thus ?thesis + proof (induction l arbitrary: x) + case 0 + hence "x \ \" + using cfrac_of_real_length_eq_0_iff zero_enat_def by auto + thus ?case + by (auto simp: Ints_def cfrac_lim_def cfrac_length_of_real_alt eSuc_def conv_Suc) + next + case (Suc l x) + hence *: "\cfrac_is_int (cfrac_of_real_alt x)" "x \ \" + by (auto simp: cfrac_is_int_def cfrac_length_of_real_alt Ints_def zero_enat_def eSuc_def) + hence "cfrac_lim (cfrac_of_real_alt x) = + of_int \x\ + 1 / cfrac_lim (cfrac_tl (cfrac_of_real_alt x))" + by (subst cfrac_lim_reduce) (auto simp: cfrac_nth_of_real_alt_0) + also have "cfrac_length (cfrac_of_real (1 / frac x)) = l" + using Suc.prems * by (metis cfrac_length_of_real_reduce eSuc_enat eSuc_inject) + hence "1 / cfrac_lim (cfrac_tl (cfrac_of_real_alt x)) = frac x" + by (subst cfrac_tl_of_real_alt[OF *(2)], subst Suc) (use Suc.prems * in auto) + also have "real_of_int \x\ + frac x = x" + by (simp add: frac_def) + finally show ?case . + qed +qed + +lemma cfrac_eqI: + assumes "cfrac_length c = cfrac_length c'" and "\n. cfrac_nth c n = cfrac_nth c' n" + shows "c = c'" +proof (use assms in transfer, safe, goal_cases) + case (1 a xs b ys) + from 1(2)[of 0] show ?case + by auto +next + case (2 a xs b ys) + define f where "f = (\xs n. if enat (Suc n) \ llength xs then int (lnth xs n) + 1 else 1)" + have "\n. f xs n = f ys n" + using 2(2)[of "Suc n" for n] by (auto simp: f_def cong: if_cong) + with 2(1) show "xs = ys" + proof (coinduction arbitrary: xs ys) + case (Eq_llist xs ys) + show ?case + proof (cases "lnull xs \ lnull ys") + case False + from False have *: "enat (Suc 0) \ llength ys" + using Suc_ile_eq zero_enat_def by auto + have "llength (ltl xs) = llength (ltl ys)" + using Eq_llist by (cases xs; cases ys) auto + moreover have "lhd xs = lhd ys" + using False * Eq_llist(1) spec[OF Eq_llist(2), of 0] + by (auto simp: f_def lnth_0_conv_lhd) + moreover have "f (ltl xs) n = f (ltl ys) n" for n + using Eq_llist(1) * spec[OF Eq_llist(2), of "Suc n"] + by (cases xs; cases ys) (auto simp: f_def Suc_ile_eq split: if_splits) + ultimately show ?thesis + using False by auto + next + case True + thus ?thesis + using Eq_llist(1) by auto + qed + qed +qed + +lemma cfrac_eq_0I: + assumes "cfrac_lim c = 0" "cfrac_nth c 0 \ 0" + shows "c = 0" +proof - + have *: "cfrac_is_int c" + proof (rule ccontr) + assume *: "\cfrac_is_int c" + from * have "conv c 0 < cfrac_lim c" + by (intro conv_less_cfrac_lim) (auto simp: cfrac_is_int_def simp flip: zero_enat_def) + hence "cfrac_nth c 0 < 0" + using assms by simp + thus False + using assms by simp + qed + from * assms have "cfrac_nth c 0 = 0" + by (auto simp: cfrac_lim_def cfrac_is_int_def) + from * and this show "c = 0" + unfolding zero_cfrac_def cfrac_is_int_def by transfer auto +qed + +lemma cfrac_eq_1I: + assumes "cfrac_lim c = 1" "cfrac_nth c 0 \ 0" + shows "c = 1" +proof - + have *: "cfrac_is_int c" + proof (rule ccontr) + assume *: "\cfrac_is_int c" + from * have "conv c 0 < cfrac_lim c" + by (intro conv_less_cfrac_lim) (auto simp: cfrac_is_int_def simp flip: zero_enat_def) + hence "cfrac_nth c 0 < 0" + using assms by simp + + have "cfrac_lim c = real_of_int (cfrac_nth c 0) + 1 / cfrac_lim (cfrac_tl c)" + using * by (subst cfrac_lim_reduce) auto + also have "real_of_int (cfrac_nth c 0) < 0" + using \cfrac_nth c 0 < 0\ by simp + also have "1 / cfrac_lim (cfrac_tl c) \ 1" + proof - + have "1 \ cfrac_nth (cfrac_tl c) 0" + by auto + also have "\ \ cfrac_lim (cfrac_tl c)" + by (rule cfrac_lim_ge_first) + finally show ?thesis by simp + qed + finally show False + using assms by simp + qed + + from * assms have "cfrac_nth c 0 = 1" + by (auto simp: cfrac_lim_def cfrac_is_int_def) + from * and this show "c = 1" + unfolding one_cfrac_def cfrac_is_int_def by transfer auto +qed + +lemma cfrac_coinduct [coinduct type: cfrac]: + assumes "R c1 c2" + assumes IH: "\c1 c2. R c1 c2 \ + cfrac_is_int c1 = cfrac_is_int c2 \ + cfrac_nth c1 0 = cfrac_nth c2 0 \ + (\cfrac_is_int c1 \ \cfrac_is_int c2 \ R (cfrac_tl c1) (cfrac_tl c2))" + shows "c1 = c2" +proof (rule cfrac_eqI) + show "cfrac_nth c1 n = cfrac_nth c2 n" for n + using assms(1) + proof (induction n arbitrary: c1 c2) + case 0 + from IH[OF this] show ?case + by auto + next + case (Suc n) + thus ?case + using IH by (metis cfrac_is_int_iff cfrac_nth_0_of_int cfrac_nth_tl) + qed +next + show "cfrac_length c1 = cfrac_length c2" + using assms(1) + proof (coinduction arbitrary: c1 c2 rule: enat_coinduct) + case (Eq_enat c1 c2) + show ?case + proof (cases "cfrac_is_int c1") + case True + thus ?thesis + using IH[OF Eq_enat(1)] by (auto simp: cfrac_is_int_def) + next + case False + with IH[OF Eq_enat(1)] have **: "\cfrac_is_int c1" "R (cfrac_tl c1) (cfrac_tl c2)" + by auto + have *: "(cfrac_length c1 = 0) = (cfrac_length c2 = 0)" + using IH[OF Eq_enat(1)] by (auto simp: cfrac_is_int_def) + show ?thesis + by (intro conjI impI disjI1 *, rule exI[of _ "cfrac_tl c1"], rule exI[of _ "cfrac_tl c2"]) + (use ** in \auto simp: epred_conv_minus\) + qed + qed +qed + +lemma cfrac_nth_0_cases: + "cfrac_nth c 0 = \cfrac_lim c\ \ cfrac_nth c 0 = \cfrac_lim c\ - 1 \ cfrac_tl c = 1" +proof (cases "cfrac_is_int c") + case True + hence "cfrac_nth c 0 = \cfrac_lim c\" + by (auto simp: cfrac_lim_def cfrac_is_int_def) + thus ?thesis by blast +next + case False + note not_int = this + have bounds: "1 / cfrac_lim (cfrac_tl c) \ 0 \ 1 / cfrac_lim (cfrac_tl c) \ 1" + proof - + have "1 \ cfrac_nth (cfrac_tl c) 0" + by simp + also have "\ \ cfrac_lim (cfrac_tl c)" + by (rule cfrac_lim_ge_first) + finally show ?thesis + using False by (auto simp: cfrac_lim_nonneg) + qed + + thus ?thesis + proof (cases "cfrac_lim (cfrac_tl c) = 1") + case False + have "\cfrac_lim c\ = cfrac_nth c 0 + \1 / cfrac_lim (cfrac_tl c)\" + using not_int by (subst cfrac_lim_reduce) auto + also have "1 / cfrac_lim (cfrac_tl c) \ 0 \ 1 / cfrac_lim (cfrac_tl c) < 1" + using bounds False by (auto simp: divide_simps) + hence "\1 / cfrac_lim (cfrac_tl c)\ = 0" + by linarith + finally show ?thesis by simp + next + case True + have "cfrac_nth c 0 = \cfrac_lim c\ - 1" + using not_int True by (subst cfrac_lim_reduce) auto + moreover have "cfrac_tl c = 1" + using True by (intro cfrac_eq_1I) auto + ultimately show ?thesis by blast + qed +qed + +lemma cfrac_length_1 [simp]: "cfrac_length 1 = 0" + unfolding one_cfrac_def by simp + +lemma cfrac_nth_1 [simp]: "cfrac_nth 1 m = 1" + unfolding one_cfrac_def by transfer (auto simp: enat_0_iff) + +lemma cfrac_lim_1 [simp]: "cfrac_lim 1 = 1" + by (auto simp: cfrac_lim_def) + + +lemma cfrac_nth_0_not_int: + assumes "cfrac_lim c \ \" + shows "cfrac_nth c 0 = \cfrac_lim c\" +proof - + have "cfrac_tl c \ 1" + proof + assume eq: "cfrac_tl c = 1" + have "\cfrac_is_int c" + using assms by (auto simp: cfrac_lim_def cfrac_is_int_def) + hence "cfrac_lim c = of_int \cfrac_nth c 0\ + 1" + using eq by (subst cfrac_lim_reduce) auto + hence "cfrac_lim c \ \" + by auto + with assms show False by auto + qed + with cfrac_nth_0_cases[of c] show ?thesis by auto +qed + +lemma cfrac_of_real_cfrac_lim_irrational: + assumes "cfrac_lim c \ \" + shows "cfrac_of_real (cfrac_lim c) = c" +proof (rule cfrac_eqI) + from assms show "cfrac_length (cfrac_of_real (cfrac_lim c)) = cfrac_length c" + using cfrac_lim_rational_iff by auto +next + fix n + show "cfrac_nth (cfrac_of_real (cfrac_lim c)) n = cfrac_nth c n" + using assms + proof (induction n arbitrary: c) + case (0 c) + thus ?case + using Ints_subset_Rats by (subst cfrac_nth_0_not_int) auto + next + case (Suc n c) + from Suc.prems have [simp]: "cfrac_lim c \ \" + using Ints_subset_Rats by blast + have "cfrac_nth (cfrac_of_real (cfrac_lim c)) (Suc n) = + cfrac_nth (cfrac_tl (cfrac_of_real (cfrac_lim c))) n" + by (simp flip: cfrac_nth_tl) + also have "cfrac_tl (cfrac_of_real (cfrac_lim c)) = cfrac_of_real (1 / frac (cfrac_lim c))" + using Suc.prems Ints_subset_Rats by (subst cfrac_tl_of_real) auto + also have "1 / frac (cfrac_lim c) = cfrac_lim (cfrac_tl c)" + using Suc.prems by (subst cfrac_lim_tl) (auto simp: frac_def cfrac_is_int_def cfrac_nth_0_not_int) + also have "cfrac_nth (cfrac_of_real (cfrac_lim (cfrac_tl c))) n = cfrac_nth c (Suc n)" + using Suc.prems by (subst Suc.IH) (auto simp: cfrac_lim_rational_iff) + finally show ?case . + qed +qed + +lemma cfrac_eqI_first: + assumes "\cfrac_is_int c" "\cfrac_is_int c'" + assumes "cfrac_nth c 0 = cfrac_nth c' 0" and "cfrac_tl c = cfrac_tl c'" + shows "c = c'" + using assms unfolding cfrac_is_int_def + by transfer (auto split: llist.splits) + +lemma cfrac_is_int_of_real_iff: "cfrac_is_int (cfrac_of_real x) \ x \ \" + unfolding cfrac_is_int_def by transfer (use frac_lt_1 in auto) + +lemma cfrac_not_is_int_of_real_alt: "\cfrac_is_int (cfrac_of_real_alt x)" + unfolding cfrac_is_int_def by transfer (auto simp: frac_lt_1) + +lemma cfrac_tl_of_real_alt_of_int [simp]: "cfrac_tl (cfrac_of_real_alt (of_int n)) = 1" + unfolding one_cfrac_def by transfer auto + +lemma cfrac_is_intI: + assumes "cfrac_nth c 0 \ \cfrac_lim c\" and "cfrac_lim c \ \" + shows "cfrac_is_int c" +proof (rule ccontr) + assume *: "\cfrac_is_int c" + from * have "conv c 0 < cfrac_lim c" + by (intro conv_less_cfrac_lim) (auto simp: cfrac_is_int_def simp flip: zero_enat_def) + with assms show False + by (auto simp: Ints_def) +qed + +lemma cfrac_eq_of_intI: + assumes "cfrac_nth c 0 \ \cfrac_lim c\" and "cfrac_lim c \ \" + shows "c = cfrac_of_int \cfrac_lim c\" +proof - + from assms have int: "cfrac_is_int c" + by (intro cfrac_is_intI) auto + have [simp]: "cfrac_lim c = cfrac_nth c 0" + using int by (simp add: cfrac_lim_def cfrac_is_int_def) + from int have "c = cfrac_of_int (cfrac_nth c 0)" + unfolding cfrac_is_int_def by transfer auto + also from assms have "cfrac_nth c 0 = \cfrac_lim c\" + using int by auto + finally show ?thesis . +qed + +lemma cfrac_lim_of_int [simp]: "cfrac_lim (cfrac_of_int n) = of_int n" + by (simp add: cfrac_lim_def) + +lemma cfrac_of_real_of_int [simp]: "cfrac_of_real (of_int n) = cfrac_of_int n" + by transfer auto + +lemma cfrac_of_real_of_nat [simp]: "cfrac_of_real (of_nat n) = cfrac_of_int (int n)" + by transfer auto + +lemma cfrac_int_cases: + assumes "cfrac_lim c = of_int n" + shows "c = cfrac_of_int n \ c = cfrac_of_real_alt (of_int n)" +proof - + from cfrac_nth_0_cases[of c] show ?thesis + proof (rule disj_forward) + assume eq: "cfrac_nth c 0 = \cfrac_lim c\" + have "c = cfrac_of_int \cfrac_lim c\" + using assms eq by (intro cfrac_eq_of_intI) auto + with assms eq show "c = cfrac_of_int n" + by simp + next + assume *: "cfrac_nth c 0 = \cfrac_lim c\ - 1 \ cfrac_tl c = 1" + have "\cfrac_is_int c" + using * by (auto simp: cfrac_is_int_def cfrac_lim_def) + hence "cfrac_length c = eSuc (cfrac_length (cfrac_tl c))" + by (subst cfrac_length_tl; cases "cfrac_length c") + (auto simp: cfrac_is_int_def eSuc_def enat_0_iff split: enat.splits) + also have "cfrac_tl c = 1" + using * by auto + finally have "cfrac_length c = 1" + by (simp add: eSuc_def one_enat_def) + show "c = cfrac_of_real_alt (of_int n)" + by (rule cfrac_eqI_first) + (use \\cfrac_is_int c\ * assms in \auto simp: cfrac_not_is_int_of_real_alt\) + qed +qed + +lemma cfrac_cases: + "c \ {cfrac_of_real (cfrac_lim c), cfrac_of_real_alt (cfrac_lim c)}" +proof (cases "cfrac_length c") + case infinity + hence "cfrac_lim c \ \" + by (simp add: cfrac_lim_irrational) + thus ?thesis + using cfrac_of_real_cfrac_lim_irrational by simp +next + case (enat l) + thus ?thesis + proof (induction l arbitrary: c) + case (0 c) + hence "c = cfrac_of_real (cfrac_nth c 0)" + by transfer (auto simp flip: zero_enat_def) + with 0 show ?case by (auto simp: cfrac_lim_def) + next + case (Suc l c) + show ?case + proof (cases "cfrac_lim c \ \") + case True + thus ?thesis + using cfrac_int_cases[of c] by (force simp: Ints_def) + next + case [simp]: False + have "\cfrac_is_int c" + using Suc.prems by (auto simp: cfrac_is_int_def enat_0_iff) + show ?thesis + using cfrac_nth_0_cases[of c] + proof (elim disjE conjE) + assume *: "cfrac_nth c 0 = \cfrac_lim c\ - 1" "cfrac_tl c = 1" + hence "cfrac_lim c \ \" + using \\cfrac_is_int c\ by (subst cfrac_lim_reduce) auto + thus ?thesis + by (auto simp: cfrac_int_cases) + next + assume eq: "cfrac_nth c 0 = \cfrac_lim c\" + have "cfrac_tl c = cfrac_of_real (cfrac_lim (cfrac_tl c)) \ + cfrac_tl c = cfrac_of_real_alt (cfrac_lim (cfrac_tl c))" + using Suc.IH[of "cfrac_tl c"] Suc.prems by auto + hence "c = cfrac_of_real (cfrac_lim c) \ + c = cfrac_of_real_alt (cfrac_lim c)" + proof (rule disj_forward) + assume eq': "cfrac_tl c = cfrac_of_real (cfrac_lim (cfrac_tl c))" + show "c = cfrac_of_real (cfrac_lim c)" + by (rule cfrac_eqI_first) + (use \\cfrac_is_int c\ eq eq' in + \auto simp: cfrac_is_int_of_real_iff cfrac_tl_of_real cfrac_lim_tl frac_def\) + next + assume eq': "cfrac_tl c = cfrac_of_real_alt (cfrac_lim (cfrac_tl c))" + have eq'': "cfrac_nth (cfrac_of_real_alt (cfrac_lim c)) 0 = \cfrac_lim c\" + using Suc.prems by (subst cfrac_nth_of_real_alt_0) auto + show "c = cfrac_of_real_alt (cfrac_lim c)" + by (rule cfrac_eqI_first) + (use \\cfrac_is_int c\ eq eq' eq'' in + \auto simp: cfrac_not_is_int_of_real_alt cfrac_tl_of_real_alt cfrac_lim_tl frac_def\) + qed + thus ?thesis by simp + qed + qed + qed +qed + +lemma cfrac_lim_eq_iff: + assumes "cfrac_length c = \ \ cfrac_length c' = \" + shows "cfrac_lim c = cfrac_lim c' \ c = c'" +proof + assume *: "cfrac_lim c = cfrac_lim c'" + hence "cfrac_of_real (cfrac_lim c) = cfrac_of_real (cfrac_lim c')" + by (simp only:) + thus "c = c'" + using assms * + by (subst (asm) (1 2) cfrac_of_real_cfrac_lim_irrational) + (auto simp: cfrac_infinite_iff) +qed auto + +lemma floor_cfrac_remainder: + assumes "cfrac_length c = \" + shows "\cfrac_remainder c n\ = cfrac_nth c n" + by (metis add.left_neutral assms cfrac_length_drop cfrac_lim_eq_iff idiff_infinity + cfrac_lim_of_real cfrac_nth_drop cfrac_nth_of_real_0 cfrac_remainder_def) + + +subsection \Approximation properties\ + +text \ + In this section, we will show that convergents of the continued fraction expansion of a + number \x\ are good approximations of \x\, and in a certain sense, the reverse holds as well. +\ + +lemma sgn_of_int: + "sgn (of_int x :: 'a :: {linordered_idom}) = of_int (sgn x)" + by (auto simp: sgn_if) + +lemma conv_ge_one: "cfrac_nth c 0 > 0 \ conv c n \ 1" + by (rule order.trans[OF _ conv_ge_first]) auto + +context + fixes c h k + defines "h \ conv_num c" and "k \ conv_denom c" +begin + +lemma abs_diff_le_abs_add: + fixes x y :: real + assumes "x \ 0 \ y \ 0 \ x \ 0 \ y \ 0" + shows "\x - y\ \ \x + y\" + using assms by linarith + +lemma abs_diff_less_abs_add: + fixes x y :: real + assumes "x > 0 \ y > 0 \ x < 0 \ y < 0" + shows "\x - y\ < \x + y\" + using assms by linarith + +lemma abs_diff_le_imp_same_sign: + assumes "\x - y\ \ d" "d < \y\" + shows "sgn x = sgn (y::real)" + using assms by (auto simp: sgn_if) + +lemma conv_nonpos: + assumes "cfrac_nth c 0 < 0" + shows "conv c n \ 0" +proof (cases n) + case 0 + thus ?thesis using assms by auto +next + case [simp]: (Suc n') + have "conv c n = real_of_int (cfrac_nth c 0) + 1 / conv (cfrac_tl c) n'" + by (simp add: conv_Suc) + also have "\ \ -1 + 1 / 1" + using assms by (intro add_mono divide_left_mono) (auto intro!: conv_pos conv_ge_one) + finally show ?thesis by simp +qed + +lemma cfrac_lim_nonpos: + assumes "cfrac_nth c 0 < 0" + shows "cfrac_lim c \ 0" +proof (cases "cfrac_length c") + case infinity + show ?thesis using LIMSEQ_cfrac_lim[OF infinity] + by (rule tendsto_upperbound) (use assms in \auto simp: conv_nonpos\) +next + case (enat l) + thus ?thesis by (auto simp: cfrac_lim_def conv_nonpos assms) +qed + +lemma conv_num_nonpos: + assumes "cfrac_nth c 0 < 0" + shows "h n \ 0" +proof (induction n rule: fib.induct) + case 2 + have "cfrac_nth c (Suc 0) * cfrac_nth c 0 \ 1 * cfrac_nth c 0" + using assms by (intro mult_right_mono_neg) auto + also have "\ + 1 \ 0" using assms by auto + finally show ?case by (auto simp: h_def) +next + case (3 n) + have "cfrac_nth c (Suc (Suc n)) * h (Suc n) \ 0" + using 3 by (simp add: mult_nonneg_nonpos) + also have "\ + h n \ 0" + using 3 by simp + finally show ?case + by (auto simp: h_def) +qed (use assms in \auto simp: h_def\) + +lemma conv_best_approximation_aux: + "cfrac_lim c \ 0 \ h n \ 0 \ cfrac_lim c \ 0 \ h n \ 0" +proof (cases "cfrac_nth c 0 \ 0") + case True + from True have "0 \ conv c 0" + by simp + also have "\ \ cfrac_lim c" + by (rule conv_le_cfrac_lim) (auto simp: enat_0) + finally have "cfrac_lim c \ 0" . + moreover from True have "h n \ 0" + unfolding h_def by (intro conv_num_nonneg) + ultimately show ?thesis by (simp add: sgn_if) +next + case False + thus ?thesis + using cfrac_lim_nonpos conv_num_nonpos[of n] by (auto simp: h_def) +qed + +lemma conv_best_approximation_ex: + fixes a b :: int and x :: real + assumes "n \ cfrac_length c" + assumes "0 < b" and "b \ k n" and "coprime a b" and "n > 0" + assumes "(a, b) \ (h n, k n)" + assumes "\(cfrac_length c = 1 \ n = 0)" + assumes "Suc n \ cfrac_length c \ cfrac_canonical c" + defines "x \ cfrac_lim c" + shows "\k n * x - h n\ < \b * x - a\" +proof (cases "\a\ = \h n\ \ b = k n") + case True + with assms have [simp]: "a = -h n" + by (auto simp: abs_if split: if_splits) + have "k n > 0" + by (auto simp: k_def) + show ?thesis + proof (cases "x = 0") + case True + hence "c = cfrac_of_real 0 \ c = cfrac_of_real_alt 0" + unfolding x_def by (metis cfrac_cases empty_iff insert_iff) + hence False + proof + assume "c = cfrac_of_real 0" + thus False + using assms by (auto simp: enat_0_iff h_def k_def) + next + assume [simp]: "c = cfrac_of_real_alt 0" + hence "n = 0 \ n = 1" + using assms by (auto simp: cfrac_length_of_real_alt enat_0_iff k_def h_def eSuc_def) + thus False + using assms True + by (elim disjE) (auto simp: cfrac_length_of_real_alt enat_0_iff k_def h_def eSuc_def + cfrac_nth_of_real_alt one_enat_def split: if_splits) + qed + thus ?thesis .. + next + case False + have "h n \ 0" + using True assms(6) h_def by auto + hence "x > 0 \ h n > 0 \ x < 0 \ h n < 0" + using \x \ 0\ conv_best_approximation_aux[of n] unfolding x_def by auto + hence "\real_of_int (k n) * x - real_of_int (h n)\ < \real_of_int (k n) * x + real_of_int (h n)\" + using \k n > 0\ + by (intro abs_diff_less_abs_add) (auto simp: not_le zero_less_mult_iff mult_less_0_iff) + thus ?thesis using True by auto + qed +next + case False + note * = this + show ?thesis + proof (cases "n = cfrac_length c") + case True + hence "x = conv c n" + by (auto simp: cfrac_lim_def x_def split: enat.splits) + also have "\ = h n / k n" + by (auto simp: h_def k_def conv_num_denom) + finally have x: "x = h n / k n" . + hence "\k n * x - h n\ = 0" + by (simp add: k_def) + also have "b * x \ a" + proof + assume "b * x = a" + hence "of_int (h n) * of_int b = of_int (k n) * (of_int a :: real)" + using assms True by (auto simp: field_simps k_def x) + hence "of_int (h n * b) = (of_int (k n * a) :: real)" + by (simp only: of_int_mult) + hence "h n * b = k n * a" + by linarith + hence "h n = a \ k n = b" + using assms by (subst (asm) coprime_crossproduct') + (auto simp: h_def k_def coprime_conv_num_denom) + thus False using True False by simp + qed + hence "0 < \b * x - a\" + by simp + finally show ?thesis . + next + case False + + define s where "s = (-1) ^ n * (a * k n - b * h n)" + define r where "r = (-1) ^ n * (b * h (Suc n) - a * k (Suc n))" + have "k n \ k (Suc n)" + unfolding k_def by (intro conv_denom_leI) auto + + have "r * h n + s * h (Suc n) = + (-1) ^ Suc n * a * (k (Suc n) * h n - k n * h (Suc n))" + by (simp add: s_def r_def algebra_simps h_def k_def) + also have "\ = a" using assms unfolding h_def k_def + by (subst conv_num_denom_prod_diff') (auto simp: algebra_simps) + finally have eq1: "r * h n + s * h (Suc n) = a" . + + have "r * k n + s * k (Suc n) = + (-1) ^ Suc n * b * (k (Suc n) * h n - k n * h (Suc n))" + by (simp add: s_def r_def algebra_simps h_def k_def) + also have "\ = b" using assms unfolding h_def k_def + by (subst conv_num_denom_prod_diff') (auto simp: algebra_simps) + finally have eq2: "r * k n + s * k (Suc n) = b" . + + have "k n < k (Suc n)" + using \n > 0\ by (auto simp: k_def intro: conv_denom_lessI) + + have "r \ 0" + proof + assume "r = 0" + hence "a * k (Suc n) = b * h (Suc n)" by (simp add: r_def) + hence "abs (a * k (Suc n)) = abs (h (Suc n) * b)" by (simp only: mult_ac) + hence *: "abs (h (Suc n)) = abs a \ k (Suc n) = b" + unfolding abs_mult h_def k_def using coprime_conv_num_denom assms + by (subst (asm) coprime_crossproduct_int) auto + with \k n < k (Suc n)\ and \b \ k n\ show False by auto + qed + + have "s \ 0" + proof + assume "s = 0" + hence "a * k n = b * h n" by (simp add: s_def) + hence "abs (a * k n) = abs (h n * b)" by (simp only: mult_ac) + hence "b = k n \ \a\ = \h n\" unfolding abs_mult h_def k_def using coprime_conv_num_denom assms + by (subst (asm) coprime_crossproduct_int) auto + with * show False by simp + qed + + have "r * k n + s * k (Suc n) = b" by fact + also have "\ \ {0<..k n < k (Suc n)\ by auto + finally have *: "r * k n + s * k (Suc n) \ \" . + + have opposite_signs1: "r > 0 \ s < 0 \ r < 0 \ s > 0" + proof (cases "r \ 0"; cases "s \ 0") + assume "r \ 0" "s \ 0" + hence "0 * (k n) + 1 * (k (Suc n)) \ r * k n + s * k (Suc n)" + using \s \ 0\ by (intro add_mono mult_mono) (auto simp: k_def) + with * show ?thesis by auto + next + assume "\(r \ 0)" "\(s \ 0)" + hence "r * k n + s * k (Suc n) \ 0" + by (intro add_nonpos_nonpos mult_nonpos_nonneg) (auto simp: k_def) + with * show ?thesis by auto + qed (insert \r \ 0\ \s \ 0\, auto) + + have "r \ 1" + proof + assume [simp]: "r = 1" + have "b = r * k n + s * k (Suc n)" + using \r * k n + s * k (Suc n) = b\ .. + also have "s * k (Suc n) \ (-1) * k (Suc n)" + using opposite_signs1 by (intro mult_right_mono) (auto simp: k_def) + also have "r * k n + (-1) * k (Suc n) = k n - k (Suc n)" + by simp + also have "\ \ 0" + unfolding k_def by (auto intro!: conv_denom_leI) + finally show False using \b > 0\ by simp + qed + + have "enat n \ cfrac_length c" "enat (Suc n) \ cfrac_length c" + using assms False by (cases "cfrac_length c"; simp)+ + hence "conv c n \ x \ conv c (Suc n) \ x \ conv c n \ x \ conv c (Suc n) \ x" + using conv_ge_cfrac_lim[of n c] conv_ge_cfrac_lim[of "Suc n" c] + conv_le_cfrac_lim[of n c] conv_le_cfrac_lim[of "Suc n" c] assms + by (cases "even n") auto + hence opposite_signs2: "k n * x - h n \ 0 \ k (Suc n) * x - h (Suc n) \ 0 \ + k n * x - h n \ 0 \ k (Suc n) * x - h (Suc n) \ 0" + using assms conv_denom_pos[of c n] conv_denom_pos[of c "Suc n"] + by (auto simp: k_def h_def conv_num_denom field_simps) + + from opposite_signs1 opposite_signs2 have same_signs: + "r * (k n * x - h n) \ 0 \ s * (k (Suc n) * x - h (Suc n)) \ 0 \ + r * (k n * x - h n) \ 0 \ s * (k (Suc n) * x - h (Suc n)) \ 0" + by (auto intro: mult_nonpos_nonneg mult_nonneg_nonpos mult_nonneg_nonneg mult_nonpos_nonpos) + + show ?thesis + proof (cases "Suc n = cfrac_length c") + case True + have x: "x = h (Suc n) / k (Suc n)" + using True[symmetric] by (auto simp: cfrac_lim_def h_def k_def conv_num_denom x_def) + have "r \ -1" + proof + assume [simp]: "r = -1" + have "r * k n + s * k (Suc n) = b" + by fact + also have "b < k (Suc n)" + using \b \ k n\ and \k n < k (Suc n)\ by simp + finally have "(s - 1) * k (Suc n) < k n" + by (simp add: algebra_simps) + also have "k n \ 1 * k (Suc n)" + by (simp add: k_def conv_denom_leI) + finally have "s < 2" + by (subst (asm) mult_less_cancel_right) (auto simp: k_def) + moreover from opposite_signs1 have "s > 0" by auto + ultimately have [simp]: "s = 1" by simp + + have "b = (cfrac_nth c (Suc n) - 1) * k n + k (n - 1)" + using eq2 \n > 0\ by (cases n) (auto simp: k_def algebra_simps) + also have "cfrac_nth c (Suc n) > 1" + proof - + have "cfrac_canonical c" + using assms True by auto + hence "cfrac_nth c (Suc n) \ 1" + using True[symmetric] by (auto simp: cfrac_canonical_iff enat_0_iff) + moreover have "cfrac_nth c (Suc n) > 0" + by auto + ultimately show "cfrac_nth c (Suc n) > 1" + by linarith + qed + hence "(cfrac_nth c (Suc n) - 1) * k n + k (n - 1) \ 1 * k n + k (n - 1)" + by (intro add_mono mult_right_mono) (auto simp: k_def) + finally have "b > k n" + using conv_denom_pos[of c "n - 1"] unfolding k_def by linarith + with assms show False by simp + qed + with \r \ 1\ \r \ 0\ have "\r\ > 1" + by auto + + from \s \ 0\ have "k n * x \ h n" + using conv_num_denom_prod_diff[of c n] + by (auto simp: x field_simps k_def h_def simp flip: of_int_mult) + hence "1 * \k n * x - h n\ < \r\ * \k n * x - h n\" + using \\r\ > 1\ by (intro mult_strict_right_mono) auto + also have "\ = \r\ * \k n * x - h n\ + 0" by simp + also have "\ \ \r * (k n * x - h n)\ + \s * (k (Suc n) * x - h (Suc n))\" + unfolding abs_mult of_int_abs using conv_denom_pos[of c "Suc n"] \s \ 0\ + by (intro add_left_mono mult_nonneg_nonneg) (auto simp: field_simps k_def) + also have "\ = \r * (k n * x - h n) + s * (k (Suc n) * x - h (Suc n))\" + using same_signs by auto + also have "\ = \(r * k n + s * k (Suc n)) * x - (r * h n + s * h (Suc n))\" + by (simp add: algebra_simps) + also have "\ = \b * x - a\" + unfolding eq1 eq2 by simp + finally show ?thesis by simp + next + case False + from assms have "Suc n < cfrac_length c" + using False \Suc n \ cfrac_length c\ by force + have "1 * \k n * x - h n\ \ \r\ * \k n * x - h n\" + using \r \ 0\ by (intro mult_right_mono) auto + also have "\ = \r\ * \k n * x - h n\ + 0" by simp + also have "x \ h (Suc n) / k (Suc n)" + using conv_neq_cfrac_lim[of "Suc n" c] \Suc n < cfrac_length c\ + by (auto simp: conv_num_denom h_def k_def x_def) + hence "\s * (k (Suc n) * x - h (Suc n))\ > 0" + using \s \ 0\ by (auto simp: field_simps k_def) + also have "\r\ * \k n * x - h n\ + \ \ + \r * (k n * x - h n)\ + \s * (k (Suc n) * x - h (Suc n))\" + unfolding abs_mult of_int_abs by (intro add_left_mono mult_nonneg_nonneg) auto + also have "\ = \r * (k n * x - h n) + s * (k (Suc n) * x - h (Suc n))\" + using same_signs by auto + also have "\ = \(r * k n + s * k (Suc n)) * x - (r * h n + s * h (Suc n))\" + by (simp add: algebra_simps) + also have "\ = \b * x - a\" + unfolding eq1 eq2 by simp + finally show ?thesis by simp + qed + qed +qed + +lemma conv_best_approximation_ex_weak: + fixes a b :: int and x :: real + assumes "n \ cfrac_length c" + assumes "0 < b" and "b < k (Suc n)" and "coprime a b" + defines "x \ cfrac_lim c" + shows "\k n * x - h n\ \ \b * x - a\" +proof (cases "\a\ = \h n\ \ b = k n") + case True + note * = this + show ?thesis + proof (cases "sgn a = sgn (h n)") + case True + with * have [simp]: "a = h n" + by (auto simp: abs_if split: if_splits) + thus ?thesis using * by auto + next + case False + with True have [simp]: "a = -h n" + by (auto simp: abs_if split: if_splits) + have "\real_of_int (k n) * x - real_of_int (h n)\ \ \real_of_int (k n) * x + real_of_int (h n)\" + unfolding x_def using conv_best_approximation_aux[of n] + by (intro abs_diff_le_abs_add) (auto simp: k_def not_le zero_less_mult_iff) + thus ?thesis using True by auto + qed +next + case False + note * = this + show ?thesis + proof (cases "n = cfrac_length c") + case True + hence "x = conv c n" + by (auto simp: cfrac_lim_def x_def split: enat.splits) + also have "\ = h n / k n" + by (auto simp: h_def k_def conv_num_denom) + finally show ?thesis by (auto simp: k_def) + next + case False + + define s where "s = (-1) ^ n * (a * k n - b * h n)" + define r where "r = (-1) ^ n * (b * h (Suc n) - a * k (Suc n))" + + have "r * h n + s * h (Suc n) = + (-1) ^ Suc n * a * (k (Suc n) * h n - k n * h (Suc n))" + by (simp add: s_def r_def algebra_simps h_def k_def) + also have "\ = a" using assms unfolding h_def k_def + by (subst conv_num_denom_prod_diff') (auto simp: algebra_simps) + finally have eq1: "r * h n + s * h (Suc n) = a" . + + have "r * k n + s * k (Suc n) = + (-1) ^ Suc n * b * (k (Suc n) * h n - k n * h (Suc n))" + by (simp add: s_def r_def algebra_simps h_def k_def) + also have "\ = b" using assms unfolding h_def k_def + by (subst conv_num_denom_prod_diff') (auto simp: algebra_simps) + finally have eq2: "r * k n + s * k (Suc n) = b" . + + have "r \ 0" + proof + assume "r = 0" + hence "a * k (Suc n) = b * h (Suc n)" by (simp add: r_def) + hence "abs (a * k (Suc n)) = abs (h (Suc n) * b)" by (simp only: mult_ac) + hence "b = k (Suc n)" unfolding abs_mult h_def k_def using coprime_conv_num_denom assms + by (subst (asm) coprime_crossproduct_int) auto + with assms show False by simp + qed + + have "s \ 0" + proof + assume "s = 0" + hence "a * k n = b * h n" by (simp add: s_def) + hence "abs (a * k n) = abs (h n * b)" by (simp only: mult_ac) + hence "b = k n \ \a\ = \h n\" unfolding abs_mult h_def k_def using coprime_conv_num_denom assms + by (subst (asm) coprime_crossproduct_int) auto + with * show False by simp + qed + + have "r * k n + s * k (Suc n) = b" by fact + also have "\ \ {0<.. \" . + + have opposite_signs1: "r > 0 \ s < 0 \ r < 0 \ s > 0" + proof (cases "r \ 0"; cases "s \ 0") + assume "r \ 0" "s \ 0" + hence "0 * (k n) + 1 * (k (Suc n)) \ r * k n + s * k (Suc n)" + using \s \ 0\ by (intro add_mono mult_mono) (auto simp: k_def) + with * show ?thesis by auto + next + assume "\(r \ 0)" "\(s \ 0)" + hence "r * k n + s * k (Suc n) \ 0" + by (intro add_nonpos_nonpos mult_nonpos_nonneg) (auto simp: k_def) + with * show ?thesis by auto + qed (insert \r \ 0\ \s \ 0\, auto) + + have "enat n \ cfrac_length c" "enat (Suc n) \ cfrac_length c" + using assms False by (cases "cfrac_length c"; simp)+ + hence "conv c n \ x \ conv c (Suc n) \ x \ conv c n \ x \ conv c (Suc n) \ x" + using conv_ge_cfrac_lim[of n c] conv_ge_cfrac_lim[of "Suc n" c] + conv_le_cfrac_lim[of n c] conv_le_cfrac_lim[of "Suc n" c] assms + by (cases "even n") auto + hence opposite_signs2: "k n * x - h n \ 0 \ k (Suc n) * x - h (Suc n) \ 0 \ + k n * x - h n \ 0 \ k (Suc n) * x - h (Suc n) \ 0" + using assms conv_denom_pos[of c n] conv_denom_pos[of c "Suc n"] + by (auto simp: k_def h_def conv_num_denom field_simps) + + from opposite_signs1 opposite_signs2 have same_signs: + "r * (k n * x - h n) \ 0 \ s * (k (Suc n) * x - h (Suc n)) \ 0 \ + r * (k n * x - h n) \ 0 \ s * (k (Suc n) * x - h (Suc n)) \ 0" + by (auto intro: mult_nonpos_nonneg mult_nonneg_nonpos mult_nonneg_nonneg mult_nonpos_nonpos) + + have "1 * \k n * x - h n\ \ \r\ * \k n * x - h n\" + using \r \ 0\ by (intro mult_right_mono) auto + also have "\ = \r\ * \k n * x - h n\ + 0" by simp + also have "\ \ \r * (k n * x - h n)\ + \s * (k (Suc n) * x - h (Suc n))\" + unfolding abs_mult of_int_abs using conv_denom_pos[of c "Suc n"] \s \ 0\ + by (intro add_left_mono mult_nonneg_nonneg) (auto simp: field_simps k_def) + also have "\ = \r * (k n * x - h n) + s * (k (Suc n) * x - h (Suc n))\" + using same_signs by auto + also have "\ = \(r * k n + s * k (Suc n)) * x - (r * h n + s * h (Suc n))\" + by (simp add: algebra_simps) + also have "\ = \b * x - a\" + unfolding eq1 eq2 by simp + finally show ?thesis by simp + qed +qed + +lemma cfrac_canonical_reduce: + "cfrac_canonical c \ + cfrac_is_int c \ \cfrac_is_int c \ cfrac_tl c \ 1 \ cfrac_canonical (cfrac_tl c)" + unfolding cfrac_is_int_def one_cfrac_def + by transfer (auto simp: cfrac_canonical_def llast_LCons split: if_splits split: llist.splits) + +lemma cfrac_nth_0_conv_floor: + assumes "cfrac_canonical c \ cfrac_length c \ 1" + shows "cfrac_nth c 0 = \cfrac_lim c\" +proof (cases "cfrac_is_int c") + case True + thus ?thesis + by (auto simp: cfrac_lim_def cfrac_is_int_def) +next + case False + show ?thesis + proof (cases "cfrac_length c = 1") + case True + hence "cfrac_canonical c" using assms by auto + hence "cfrac_tl c \ 1" using False + by (subst (asm) cfrac_canonical_reduce) auto + thus ?thesis + using cfrac_nth_0_cases[of c] by auto + next + case False + hence "cfrac_length c > 1" + using \\cfrac_is_int c\ + by (cases "cfrac_length c") (auto simp: cfrac_is_int_def one_enat_def zero_enat_def) + have "cfrac_tl c \ 1" + proof + assume "cfrac_tl c = 1" + have "0 < cfrac_length c - 1" + proof (cases "cfrac_length c") + case [simp]: (enat l) + have "cfrac_length c - 1 = enat (l - 1)" + by auto + also have "\ > enat 0" + using \cfrac_length c > 1\ by (simp add: one_enat_def) + finally show ?thesis by (simp add: zero_enat_def) + qed auto + also have "\ = cfrac_length (cfrac_tl c)" + by simp + also have "cfrac_tl c = 1" + by fact + finally show False by simp + qed + thus ?thesis using cfrac_nth_0_cases[of c] by auto + qed +qed + +lemma conv_best_approximation_ex_nat: + fixes a b :: nat and x :: real + assumes "n \ cfrac_length c" "0 < b" "b < k (Suc n)" "coprime a b" + shows "\k n * cfrac_lim c - h n\ \ \b * cfrac_lim c - a\" + using conv_best_approximation_ex_weak[OF assms(1), of b a] assms by auto + +lemma abs_mult_nonneg_left: + assumes "x \ (0 :: 'a :: {ordered_ab_group_add_abs, idom_abs_sgn})" + shows "x * \y\ = \x * y\" +proof - + from assms have "x = \x\" by simp + also have "\ * \y\ = \x * y\" by (simp add: abs_mult) + finally show ?thesis . +qed + +text \ + Any convergent of the continued fraction expansion of \x\ is a best approximation of \x\, + i.e. there is no other number with a smaller denominator that approximates it better. +\ +lemma conv_best_approximation: + fixes a b :: int and x :: real + assumes "n \ cfrac_length c" + assumes "0 < b" and "b < k n" and "coprime a b" + defines "x \ cfrac_lim c" + shows "\x - conv c n\ \ \x - a / b\" +proof - + have "b < k n" by fact + also have "k n \ k (Suc n)" + unfolding k_def by (intro conv_denom_leI) auto + finally have *: "b < k (Suc n)" by simp + have "\x - conv c n\ = \k n * x - h n\ / k n" + using conv_denom_pos[of c n] assms(1) + by (auto simp: conv_num_denom field_simps k_def h_def) + also have "\ \ \b * x - a\ / k n" unfolding x_def using assms * + by (intro divide_right_mono conv_best_approximation_ex_weak) auto + also from assms have "\ \ \b * x - a\ / b" + by (intro divide_left_mono) auto + also have "\ = \x - a / b\" using assms by (simp add: field_simps) + finally show ?thesis . +qed + +lemma conv_denom_partition: + assumes "y > 0" + shows "\!n. y \ {k n..n. k n \ y + 1" + by (auto simp: k_def filterlim_at_top eventually_at_top_linorder) + define n where "n = (LEAST n. k n \ y + 1)" + from LeastI_ex[OF *] have n: "k n > y" by (simp add: Suc_le_eq n_def) + from n and assms have "n > 0" by (intro Nat.gr0I) (auto simp: k_def) + + have "k (n - 1) \ y" + proof (rule ccontr) + assume "\k (n - 1) \ y" + hence "k (n - 1) \ y + 1" by auto + hence "n - 1 \ n" unfolding n_def by (rule Least_le) + with \n > 0\ show False by simp + qed + with n and \n > 0\ have "y \ {k (n - 1)..n. y \ {k n.. {k m.. {k n.. n" + with le have "k (Suc m) \ k n" + unfolding k_def by (intro conv_denom_leI assms) auto + with le show False by auto + qed + qed auto +qed + +text \ + A fraction that approximates a real number \x\ sufficiently well (in a certain sense) + is a convergent of its continued fraction expansion. +\ +lemma frac_is_convergentI: + fixes a b :: int and x :: real + defines "x \ cfrac_lim c" + assumes "b > 0" and "coprime a b" and "\x - a / b\ < 1 / (2 * b\<^sup>2)" + shows "\n. enat n \ cfrac_length c \ (a, b) = (h n, k n)" +proof (cases "a = 0") + case True + with assms have [simp]: "a = 0" "b = 1" + by auto + + show ?thesis + proof (cases x "0 :: real" rule: linorder_cases) + case greater + hence "0 < x" "x < 1/2" + using assms by auto + hence "x \ \" + by (auto simp: Ints_def) + hence "cfrac_nth c 0 = \x\" + using assms by (subst cfrac_nth_0_not_int) (auto simp: x_def) + also from \x > 0\ \x < 1/2\ have "\ = 0" + by linarith + finally have "(a, b) = (h 0, k 0)" + by (auto simp: h_def k_def) + thus ?thesis by (intro exI[of _ 0]) (auto simp flip: zero_enat_def) + next + case less + hence "x < 0" "x > -1/2" + using assms by auto + hence "x \ \" + by (auto simp: Ints_def) + hence not_int: "\cfrac_is_int c" + by (auto simp: cfrac_is_int_def x_def cfrac_lim_def) + have "cfrac_nth c 0 = \x\" + using \x \ \\ assms by (subst cfrac_nth_0_not_int) (auto simp: x_def) + also from \x < 0\ \x > -1/2\ have "\ = -1" + by linarith + finally have [simp]: "cfrac_nth c 0 = -1" . + have "cfrac_nth c (Suc 0) = cfrac_nth (cfrac_tl c) 0" + by simp + have "cfrac_lim (cfrac_tl c) = 1 / (x + 1)" + using not_int by (subst cfrac_lim_tl) (auto simp: x_def) + also from \x < 0\ \x > -1/2\ have "\ \ {1<..<2}" + by (auto simp: divide_simps) + finally have *: "cfrac_lim (cfrac_tl c) \ {1<..<2}" . + have "cfrac_nth (cfrac_tl c) 0 = \cfrac_lim (cfrac_tl c)\" + using * by (subst cfrac_nth_0_not_int) (auto simp: Ints_def) + also have "\ = 1" + using * by (simp, linarith?) + finally have "(a, b) = (h 1, k 1)" + by (auto simp: h_def k_def) + moreover have "cfrac_length c \ 1" + using not_int + by (cases "cfrac_length c") (auto simp: cfrac_is_int_def one_enat_def zero_enat_def) + ultimately show ?thesis by (intro exI[of _ 1]) (auto simp: one_enat_def) + next + case equal + show ?thesis + using cfrac_nth_0_cases[of c] + proof + assume "cfrac_nth c 0 = \cfrac_lim c\" + with equal have "(a, b) = (h 0, k 0)" + by (simp add: x_def h_def k_def) + thus ?thesis by (intro exI[of _ 0]) (auto simp flip: zero_enat_def) + next + assume *: "cfrac_nth c 0 = \cfrac_lim c\ - 1 \ cfrac_tl c = 1" + have [simp]: "cfrac_nth c 0 = -1" + using * equal by (auto simp: x_def) + from * have "\cfrac_is_int c" + by (auto simp: cfrac_is_int_def cfrac_lim_def floor_minus) + have "cfrac_nth c 1 = cfrac_nth (cfrac_tl c) 0" + by auto + also have "cfrac_tl c = 1" + using * by auto + finally have "cfrac_nth c 1 = 1" + by simp + hence "(a, b) = (h 1, k 1)" + by (auto simp: h_def k_def) + moreover from \\cfrac_is_int c\ have "cfrac_length c \ 1" + by (cases "cfrac_length c") (auto simp: one_enat_def zero_enat_def cfrac_is_int_def) + ultimately show ?thesis + by (intro exI[of _ 1]) (auto simp: one_enat_def) + qed + qed +next + case False + hence a_nz: "a \ 0" by auto + + have "x \ 0" + proof + assume [simp]: "x = 0" + hence "\a\ / b < 1 / (2 * b ^ 2)" + using assms by simp + hence "\a\ < 1 / (2 * b)" + using assms by (simp add: field_simps power2_eq_square) + also have "\ \ 1 / 2" + using assms by (intro divide_left_mono) auto + finally have "a = 0" by auto + with \a \ 0\ show False by simp + qed + + show ?thesis + proof (rule ccontr) + assume no_convergent: "\n. enat n \ cfrac_length c \ (a, b) = (h n, k n)" + from assms have "\!r. b \ {k r.. {k r.. 0" + using conv_denom_pos[of c r] assms by (auto simp: k_def) + + show False + proof (cases "enat r \ cfrac_length c") + case False + then obtain l where l: "cfrac_length c = enat l" + by (cases "cfrac_length c") auto + have "k l \ k r" + using False l unfolding k_def by (intro conv_denom_leI) auto + also have "\ \ b" + using r by simp + finally have "b \ k l" . + + have "x = conv c l" + by (auto simp: x_def cfrac_lim_def l) + hence x_eq: "x = h l / k l" + by (auto simp: conv_num_denom h_def k_def) + have "k l > 0" + by (simp add: k_def) + + have "b * k l * \h l / k l - a / b\ < k l / (2*b)" + using assms x_eq \k l > 0\ by (auto simp: field_simps power2_eq_square) + also have "b * k l * \h l / k l - a / b\ = \b * k l * (h l / k l - a / b)\" + using \b > 0\ \k l > 0\ by (subst abs_mult) auto + also have "\ = of_int \b * h l - a * k l\" + using \b > 0\ \k l > 0\ by (simp add: algebra_simps) + also have "k l / (2 * b) < 1" + using \b \ k l\ \b > 0\ by auto + finally have "a * k l = b * h l" + by linarith + moreover have "coprime (h l) (k l)" + unfolding h_def k_def by (simp add: coprime_conv_num_denom) + ultimately have "(a, b) = (h l, k l)" + using \coprime a b\ using a_nz \b > 0\ \k l > 0\ + by (subst (asm) coprime_crossproduct') (auto simp: coprime_commute) + with no_convergent and l show False + by auto + + next + + case True + have "k r * \x - h r / k r\ = \k r * x - h r\" + using \k r > 0\ by (simp add: field_simps) + also have "\k r * x - h r\ \ \b * x - a\" + using assms r True unfolding x_def by (intro conv_best_approximation_ex_weak) auto + also have "\ = b * \x - a / b\" + using \b > 0\ by (simp add: field_simps) + also have "\ < b * (1 / (2 * b\<^sup>2))" + using \b > 0\ by (intro mult_strict_left_mono assms) auto + finally have less: "\x - conv c r\ < 1 / (2 * b * k r)" + using \k r > 0\ and \b > 0\ and assms + by (simp add: field_simps power2_eq_square conv_num_denom h_def k_def) + + have "\x - a / b\ < 1 / (2 * b\<^sup>2)" by fact + also have "\ = 1 / (2 * b) * (1 / b)" + by (simp add: power2_eq_square) + also have "\ \ 1 / (2 * b) * (\a\ / b)" + using a_nz assms by (intro mult_left_mono divide_right_mono) auto + also have "\ < 1 / 1 * (\a\ / b)" + using a_nz assms + by (intro mult_strict_right_mono divide_left_mono divide_strict_left_mono) auto + also have "\ = \a / b\" using assms by simp + finally have "sgn x = sgn (a / b)" + by (auto simp: sgn_if split: if_splits) + hence "sgn x = sgn a" using assms by (auto simp: sgn_of_int) + hence "a \ 0 \ x \ 0 \ a \ 0 \ x \ 0" + by (auto simp: sgn_if split: if_splits) + moreover have "h r \ 0 \ x \ 0 \ h r \ 0 \ x \ 0" + using conv_best_approximation_aux[of r] by (auto simp: h_def x_def) + ultimately have signs: "h r \ 0 \ a \ 0 \ h r \ 0 \ a \ 0" + using \x \ 0\ by auto + + with no_convergent assms assms True have "\h r\ \ \a\ \ b \ k r" + by (auto simp: h_def k_def) + + hence "\h r\ * \b\ \ \a\ * \k r\" unfolding h_def k_def + using assms coprime_conv_num_denom[of c r] + by (subst coprime_crossproduct_int) auto + hence "\h r\ * b \ \a\ * k r" using assms by (simp add: k_def) + hence "k r * a - h r * b \ 0" + using signs by (auto simp: algebra_simps) + hence "\k r * a - h r * b\ \ 1" by presburger + hence "real_of_int 1 / (k r * b) \ \k r * a - h r * b\ / (k r * b)" + using assms + by (intro divide_right_mono, subst of_int_le_iff) (auto simp: k_def) + also have "\ = \(real_of_int (k r) * a - h r * b) / (k r * b)\" + using assms by (simp add: k_def) + also have "(real_of_int (k r) * a - h r * b) / (k r * b) = a / b - conv c r" + using assms \k r > 0\ by (simp add: h_def k_def conv_num_denom field_simps) + also have "\a / b - conv c r\ = \(x - conv c r) - (x - a / b)\" + by (simp add: algebra_simps) + also have "\ \ \x - conv c r\ + \x - a / b\" + by (rule abs_triangle_ineq4) + also have "\ < 1 / (2 * b * k r) + 1 / (2 * b\<^sup>2)" + by (intro add_strict_mono assms less) + finally have "k r > b" + using \b > 0\ and \k r > 0\ by (simp add: power2_eq_square field_simps) + with r show False by auto + qed + qed +qed + +end + + +subsection \Efficient code for convergents\ + + +function conv_gen :: "(nat \ int) \ int \ int \ nat \ nat \ int" where + "conv_gen c (a, b, n) N = + (if n > N then b else conv_gen c (b, b * c n + a, Suc n) N)" + by auto +termination by (relation "measure (\(_, (_, _, n), N). Suc N - n)") auto + +lemmas [simp del] = conv_gen.simps + +lemma conv_gen_aux_simps [simp]: + "n > N \ conv_gen c (a, b, n) N = b" + "n \ N \ conv_gen c (a, b, n) N = conv_gen c (b, b * c n + a, Suc n) N" + by (subst conv_gen.simps, simp)+ + +lemma conv_num_eq_conv_gen_aux: + "Suc n \ N \ conv_num c n = b * cfrac_nth c n + a \ + conv_num c (Suc n) = conv_num c n * cfrac_nth c (Suc n) + b \ + conv_num c N = conv_gen (cfrac_nth c) (a, b, n) N" +proof (induction "cfrac_nth c" "(a, b, n)" N arbitrary: c a b n rule: conv_gen.induct) + case (1 a b n N c) + show ?case + proof (cases "Suc (Suc n) \ N") + case True + thus ?thesis + by (subst 1) (insert "1.prems", auto) + next + case False + thus ?thesis using 1 + by (auto simp: not_le less_Suc_eq) + qed +qed + +lemma conv_denom_eq_conv_gen_aux: + "Suc n \ N \ conv_denom c n = b * cfrac_nth c n + a \ + conv_denom c (Suc n) = conv_denom c n * cfrac_nth c (Suc n) + b \ + conv_denom c N = conv_gen (cfrac_nth c) (a, b, n) N" +proof (induction "cfrac_nth c" "(a, b, n)" N arbitrary: c a b n rule: conv_gen.induct) + case (1 a b n N c) + show ?case + proof (cases "Suc (Suc n) \ N") + case True + thus ?thesis + by (subst 1) (insert "1.prems", auto) + next + case False + thus ?thesis using 1 + by (auto simp: not_le less_Suc_eq) + qed +qed + +lemma conv_num_code [code]: "conv_num c n = conv_gen (cfrac_nth c) (0, 1, 0) n" + using conv_num_eq_conv_gen_aux[of 0 n c 1 0] by (cases n) simp_all + +lemma conv_denom_code [code]: "conv_denom c n = conv_gen (cfrac_nth c) (1, 0, 0) n" + using conv_denom_eq_conv_gen_aux[of 0 n c 0 1] by (cases n) simp_all + +definition conv_num_fun where "conv_num_fun c = conv_gen c (0, 1, 0)" +definition conv_denom_fun where "conv_denom_fun c = conv_gen c (1, 0, 0)" + +lemma + assumes "is_cfrac c" + shows conv_num_fun_eq: "conv_num_fun c n = conv_num (cfrac c) n" + and conv_denom_fun_eq: "conv_denom_fun c n = conv_denom (cfrac c) n" +proof - + from assms have "cfrac_nth (cfrac c) = c" + by (intro ext) simp_all + thus "conv_num_fun c n = conv_num (cfrac c) n" and "conv_denom_fun c n = conv_denom (cfrac c) n" + by (simp_all add: conv_num_fun_def conv_num_code conv_denom_fun_def conv_denom_code) +qed + + +subsection \Computing the continued fraction expansion of a rational number\ + + +function cfrac_list_of_rat :: "int \ int \ int list" where + "cfrac_list_of_rat (a, b) = + (if b = 0 then [0] + else a div b # (if a mod b = 0 then [] else cfrac_list_of_rat (b, a mod b)))" + by auto +termination + by (relation "measure (\(a,b). nat (abs b))") (auto simp: abs_mod_less) + +lemmas [simp del] = cfrac_list_of_rat.simps + +lemma cfrac_list_of_rat_correct: + "(let xs = cfrac_list_of_rat (a, b); c = cfrac_of_real (a / b) + in length xs = cfrac_length c + 1 \ (\i (\i 0" + using l False cfrac_of_real_length_eq_0_iff[of "a / b"] \b \ 0\ + by (auto simp: c_def zero_enat_def real_of_int_divide_in_Ints_iff intro!: Nat.gr0I) + have c': "c' = cfrac_tl c" + using False \b \ 0\ unfolding c'_def c_def + by (subst cfrac_tl_of_real) (auto simp: real_of_int_divide_in_Ints_iff frac_fraction) + from 1 have "enat (length xs') = cfrac_length c' + 1" + and xs': "\ib \ 0\ \\b dvd a\ by (auto simp: Let_def xs'_def c'_def) + + have "enat (length xs') = cfrac_length c' + 1" + by fact + also have "\ = enat l - 1 + 1" + using c' l by simp + also have "\ = enat (l - 1 + 1)" + by (metis enat_diff_one one_enat_def plus_enat_simps(1)) + also have "l - 1 + 1 = l" + using \l \ 0\ by simp + finally have [simp]: "length xs' = l" + by simp + + from xs' show ?thesis + by (auto simp: xs_def nth_Cons c' split: nat.splits) + qed + thus ?thesis using l False + by (subst cfrac_list_of_rat.simps) (simp_all add: xs_def xs'_def c_def one_enat_def) + qed +qed + +lemma conv_num_cong: + assumes "(\k. k \ n \ cfrac_nth c k = cfrac_nth c' k)" "n = n'" + shows "conv_num c n = conv_num c' n" +proof - + have "conv_num c n = conv_num c' n" + using assms(1) + by (induction n arbitrary: rule: conv_num.induct) simp_all + thus ?thesis using assms(2) + by simp +qed + +lemma conv_denom_cong: + assumes "(\k. k \ n \ cfrac_nth c k = cfrac_nth c' k)" "n = n'" + shows "conv_denom c n = conv_denom c' n'" +proof - + have "conv_denom c n = conv_denom c' n" + using assms(1) + by (induction n arbitrary: rule: conv_denom.induct) simp_all + thus ?thesis using assms(2) + by simp +qed + +lemma cfrac_lim_diff_le: + assumes "\k\Suc n. cfrac_nth c1 k = cfrac_nth c2 k" + assumes "n \ cfrac_length c1" "n \ cfrac_length c2" + shows "\cfrac_lim c1 - cfrac_lim c2\ \ 2 / (conv_denom c1 n * conv_denom c1 (Suc n))" +proof - + define d where "d = (\k. conv_denom c1 k)" + have "\cfrac_lim c1 - cfrac_lim c2\ \ \cfrac_lim c1 - conv c1 n\ + \cfrac_lim c2 - conv c1 n\" + by linarith + also have "\cfrac_lim c1 - conv c1 n\ \ 1 / (d n * d (Suc n))" + unfolding d_def using assms + by (intro cfrac_lim_minus_conv_upper_bound) auto + also have "conv c1 n = conv c2 n" + using assms by (intro conv_cong) auto + also have "\cfrac_lim c2 - conv c2 n\ \ 1 / (conv_denom c2 n * conv_denom c2 (Suc n))" + using assms unfolding d_def by (intro cfrac_lim_minus_conv_upper_bound) auto + also have "conv_denom c2 n = d n" + unfolding d_def using assms by (intro conv_denom_cong) auto + also have "conv_denom c2 (Suc n) = d (Suc n)" + unfolding d_def using assms by (intro conv_denom_cong) auto + also have "1 / (d n * d (Suc n)) + 1 / (d n * d (Suc n)) = 2 / (d n * d (Suc n))" + by simp + finally show ?thesis + by (simp add: d_def) +qed + +lemma of_int_leI: "n \ m \ (of_int n :: 'a :: linordered_idom) \ of_int m" + by simp + +lemma cfrac_lim_diff_le': + assumes "\k\Suc n. cfrac_nth c1 k = cfrac_nth c2 k" + assumes "n \ cfrac_length c1" "n \ cfrac_length c2" + shows "\cfrac_lim c1 - cfrac_lim c2\ \ 2 / (fib (n+1) * fib (n+2))" +proof - + have "\cfrac_lim c1 - cfrac_lim c2\ \ 2 / (conv_denom c1 n * conv_denom c1 (Suc n))" + by (rule cfrac_lim_diff_le) (use assms in auto) + also have "\ \ 2 / (int (fib (Suc n)) * int (fib (Suc (Suc n))))" + unfolding of_nat_mult of_int_mult + by (intro divide_left_mono mult_mono mult_pos_pos of_int_leI conv_denom_lower_bound) + (auto intro!: fib_neq_0_nat simp del: fib.simps) + also have "\ = 2 / (fib (n+1) * fib (n+2))" + by simp + finally show ?thesis . +qed + +end \ No newline at end of file diff --git a/thys/Continued_Fractions/E_CFrac.thy b/thys/Continued_Fractions/E_CFrac.thy new file mode 100644 --- /dev/null +++ b/thys/Continued_Fractions/E_CFrac.thy @@ -0,0 +1,369 @@ +(* + File: Continued_Fractions/E_CFrac.thy + Author: Manuel Eberl, University of Innsbruck +*) +section \The continued fraction expansion of $e$\ +theory E_CFrac +imports + "HOL-Analysis.Analysis" + "Continued_Fractions" + "Quadratic_Irrationals" +begin + +lemma fact_real_at_top: "filterlim (fact :: nat \ real) at_top at_top" +proof (rule filterlim_at_top_mono) + have "real n \ real (fact n)" for n + unfolding of_nat_le_iff by (rule fact_ge_self) + thus "eventually (\n. real n \ fact n) at_top" by simp +qed (fact filterlim_real_sequentially) + +lemma filterlim_div_nat_at_top: + assumes "filterlim f at_top F" "m > 0" + shows "filterlim (\x. f x div m :: nat) at_top F" + unfolding filterlim_at_top +proof + fix C :: nat + from assms(1) have "eventually (\x. f x \ C * m) F" + by (auto simp: filterlim_at_top) + thus "eventually (\x. f x div m \ C) F" + proof eventually_elim + case (elim x) + hence "(C * m) div m \ f x div m" + by (intro div_le_mono) + thus ?case using \m > 0\ by simp + qed +qed + +text \ + The continued fraction expansion of $e$ has the form + $[2; 1, 2, 1, 1, 4, 1, 1, 6, 1, 1, 8, 1, \ldots]$: +\ +definition e_cfrac where + "e_cfrac = cfrac (\n. if n = 0 then 2 else if n mod 3 = 2 then 2 * (Suc n div 3) else 1)" + +lemma cfrac_nth_e: + "cfrac_nth e_cfrac n = (if n = 0 then 2 else if n mod 3 = 2 then 2 * (Suc n div 3) else 1)" + unfolding e_cfrac_def by (subst cfrac_nth_cfrac) (auto simp: is_cfrac_def) + +lemma cfrac_length_e [simp]: "cfrac_length e_cfrac = \" + by (simp add: e_cfrac_def) + + +text \ + The formalised proof follows the one from Proof Wiki~\cite{proofwiki}. +\ + +context + fixes A B C :: "nat \ real" and p q :: "nat \ int" and a :: "nat \ int" + defines "A \ (\n. integral {0..1} (\x. exp x * x ^ n * (x - 1) ^ n / fact n))" + and "B \ (\n. integral {0..1} (\x. exp x * x ^ Suc n * (x - 1) ^ n / fact n))" + and "C \ (\n. integral {0..1} (\x. exp x * x ^ n * (x - 1) ^ Suc n / fact n))" + and "p \ (\n. if n \ 1 then 1 else conv_num e_cfrac (n - 2))" + and "q \ (\n. if n = 0 then 1 else if n = 1 then 0 else conv_denom e_cfrac (n - 2))" + and "a \ (\n. if n mod 3 = 2 then 2 * (Suc n div 3) else 1)" +begin + +lemma + assumes "n \ 2" + shows p_rec: "p n = a (n - 2) * p (n - 1) + p (n - 2)" (is ?th1) + and q_rec: "q n = a (n - 2) * q (n - 1) + q (n - 2)" (is ?th2) +proof - + have n_minus_3: "n - 3 = n - Suc (Suc (Suc 0))" + by (simp add: numeral_3_eq_3) + consider "n = 2" | "n = 3" | "n \ 4" + using assms by force + hence "?th1 \ ?th2" + by cases (auto simp: p_def q_def cfrac_nth_e a_def conv_num_rec conv_denom_rec n_minus_3) + thus ?th1 ?th2 by blast+ +qed + +lemma + assumes "n \ 1" + shows p_rec0: "p (3 * n) = p (3 * n - 1) + p (3 * n - 2)" + and q_rec0: "q (3 * n) = q (3 * n - 1) + q (3 * n - 2)" +proof - + define n' where "n' = n - 1" + from assms have "(3 * n' + 1) mod 3 \ 2" + by presburger + also have "(3 * n' + 1) = 3 * n - 2" + using assms by (simp add: n'_def) + finally show "p (3 * n) = p (3 * n - 1) + p (3 * n - 2)" + "q (3 * n) = q (3 * n - 1) + q (3 * n - 2)" + using assms by (subst p_rec q_rec; simp add: a_def)+ +qed + +lemma + assumes "n \ 1" + shows p_rec1: "p (3 * n + 1) = 2 * int n * p (3 * n) + p (3 * n - 1)" + and q_rec1: "q (3 * n + 1) = 2 * int n * q (3 * n) + q (3 * n - 1)" +proof - + define n' where "n' = n - 1" + from assms have "(3 * n' + 2) mod 3 = 2" + by presburger + also have "(3 * n' + 2) = 3 * n - 1" + using assms by (simp add: n'_def) + finally show "p (3 * n + 1) = 2 * int n * p (3 * n) + p (3 * n - 1)" + "q (3 * n + 1) = 2 * int n * q (3 * n) + q (3 * n - 1)" + using assms by (subst p_rec q_rec; simp add: a_def)+ +qed + +lemma p_rec2: "p (3 * n + 2) = p (3 * n + 1) + p (3 * n)" + and q_rec2: "q (3 * n + 2) = q (3 * n + 1) + q (3 * n)" + by (subst p_rec q_rec; simp add: a_def nat_mult_distrib nat_add_distrib)+ + +lemma A_0: "A 0 = exp 1 - 1" and B_0: "B 0 = 1" and C_0: "C 0 = 2 - exp 1" +proof - + have "(exp has_integral (exp 1 - exp 0)) {0..1::real}" + by (intro fundamental_theorem_of_calculus) + (auto intro!: derivative_eq_intros + simp flip: has_real_derivative_iff_has_vector_derivative) + thus "A 0 = exp 1 - 1" by (simp add: A_def has_integral_iff) + + have "((\x. exp x * x) has_integral (exp 1 * (1 - 1) - exp 0 * (0 - 1))) {0..1::real}" + by (intro fundamental_theorem_of_calculus) + (auto intro!: derivative_eq_intros + simp flip: has_real_derivative_iff_has_vector_derivative simp: algebra_simps) + thus "B 0 = 1" by (simp add: B_def has_integral_iff) + + have "((\x. exp x * (x - 1)) has_integral (exp 1 * (1 - 2) - exp 0 * (0 - 2))) {0..1::real}" + by (intro fundamental_theorem_of_calculus) + (auto intro!: derivative_eq_intros + simp flip: has_real_derivative_iff_has_vector_derivative simp: algebra_simps) + thus "C 0 = 2 - exp 1" by (simp add: C_def has_integral_iff) +qed + +lemma A_bound: "norm (A n) \ exp 1 / fact n" +proof - + have "norm (exp t * t ^ n * (t - 1) ^ n / fact n) \ exp 1 * 1 ^ n * 1 ^ n / fact n" + if "t \ {0..1}" for t :: real using that unfolding norm_mult norm_divide norm_power norm_fact + by (intro mult_mono divide_right_mono power_mono) auto + hence "norm (A n) \ exp 1 / fact n * (1 - 0)" + unfolding A_def by (intro integral_bound) (auto intro!: continuous_intros) + thus ?thesis by simp +qed + +lemma B_bound: "norm (B n) \ exp 1 / fact n" +proof - + have "norm (exp t * t ^ Suc n * (t - 1) ^ n / fact n) \ exp 1 * 1 ^ Suc n * 1 ^ n / fact n" + if "t \ {0..1}" for t :: real using that unfolding norm_mult norm_divide norm_power norm_fact + by (intro mult_mono divide_right_mono power_mono) auto + hence "norm (B n) \ exp 1 / fact n * (1 - 0)" + unfolding B_def by (intro integral_bound) (auto intro!: continuous_intros) + thus ?thesis by simp +qed + +lemma C_bound: "norm (C n) \ exp 1 / fact n" +proof - + have "norm (exp t * t ^ n * (t - 1) ^ Suc n / fact n) \ exp 1 * 1 ^ n * 1 ^ Suc n / fact n" + if "t \ {0..1}" for t :: real using that unfolding norm_mult norm_divide norm_power norm_fact + by (intro mult_mono divide_right_mono power_mono) auto + hence "norm (C n) \ exp 1 / fact n * (1 - 0)" + unfolding C_def by (intro integral_bound) (auto intro!: continuous_intros) + thus ?thesis by simp +qed + +lemma A_Suc: "A (Suc n) = -B n - C n" +proof - + let ?g = "\x. x ^ Suc n * (x - 1) ^ Suc n * exp x / fact (Suc n)" + have pos: "fact n + real n * fact n > 0" by (intro add_pos_nonneg) auto + have "A (Suc n) + B n + C n = + integral {0..1} (\x. exp x * x ^ Suc n * (x - 1) ^ Suc n / fact (Suc n) + + exp x * x ^ Suc n * (x - 1) ^ n / fact n + exp x * x ^ n * (x - 1) ^ Suc n / fact n)" + unfolding A_def B_def C_def + apply (subst integral_add [symmetric]) + subgoal + by (auto intro!: integrable_continuous_real continuous_intros) + subgoal + by (auto intro!: integrable_continuous_real continuous_intros) + apply (subst integral_add [symmetric]) + apply (auto intro!: integrable_continuous_real continuous_intros) + done + also have "\ = integral {0..1} (\x. exp x / fact (Suc n) * + (x ^ Suc n * (x - 1) ^ Suc n + Suc n * x ^ Suc n * (x - 1) ^ n + + Suc n * x ^ n * (x - 1) ^ Suc n))" + (is "_ = integral _ ?f") + apply (simp add: divide_simps) + apply (simp add: field_simps)? + done + also have "(?f has_integral (?g 1 - ?g 0)) {0..1}" + apply (intro fundamental_theorem_of_calculus) + subgoal + by simp + unfolding has_real_derivative_iff_has_vector_derivative [symmetric] + apply (rule derivative_eq_intros refl | simp)+ + apply (simp add: algebra_simps)? + done + hence "integral {0..1} ?f = 0" + by (simp add: has_integral_iff) + finally show ?thesis by simp +qed + +lemma B_Suc: "B (Suc n) = -2 * Suc n * A (Suc n) + C n" +proof - + let ?g = "\x. x ^ Suc n * (x - 1) ^ (n+2) * exp x / fact (Suc n)" + have pos: "fact n + real n * fact n > 0" by (intro add_pos_nonneg) auto + have "B (Suc n) + 2 * Suc n * A (Suc n) - C n = + integral {0..1} (\x. exp x * x^(n+2) * (x - 1)^(n+1) / fact (Suc n) + 2 * Suc n * + exp x * x ^ Suc n * (x - 1) ^ Suc n / fact (Suc n) - exp x * x ^ n * (x - 1) ^ Suc n / fact n)" + unfolding A_def B_def C_def integral_mult_right [symmetric] + apply (subst integral_add [symmetric]) + subgoal + by (auto intro!: integrable_continuous_real continuous_intros) + subgoal + by (auto intro!: integrable_continuous_real continuous_intros) + apply (subst integral_diff [symmetric]) + apply (auto intro!: integrable_continuous_real continuous_intros simp: mult_ac) + done + also have "\ = integral {0..1} (\x. exp x / fact (Suc n) * + (x^(n+2) * (x - 1)^(n+1) + 2 * Suc n * x ^ Suc n * (x - 1) ^ Suc n - + Suc n * x ^ n * (x - 1) ^ Suc n))" + (is "_ = integral _ ?f") + apply (simp add: divide_simps) + apply (simp add: field_simps)? + done + also have "(?f has_integral (?g 1 - ?g 0)) {0..1}" + apply (intro fundamental_theorem_of_calculus) + apply (simp; fail) + unfolding has_real_derivative_iff_has_vector_derivative [symmetric] + apply (rule derivative_eq_intros refl | simp)+ + apply (simp add: algebra_simps)? + done + hence "integral {0..1} ?f = 0" + by (simp add: has_integral_iff) + finally show ?thesis by (simp add: algebra_simps) +qed + +lemma C_Suc: "C n = B n - A n" + unfolding A_def B_def C_def + by (subst integral_diff [symmetric]) + (auto intro!: integrable_continuous_real continuous_intros simp: field_simps) + +lemma unfold_add_numeral: "c * n + numeral b = Suc (c * n + pred_numeral b)" + by simp + +lemma ABC: + "A n = q (3 * n) * exp 1 - p (3 * n) \ + B n = p (Suc (3 * n)) - q (Suc (3 * n)) * exp 1 \ + C n = p (Suc (Suc (3 * n))) - q (Suc (Suc (3 * n))) * exp 1" +proof (induction n) + case 0 + thus ?case by (simp add: A_0 B_0 C_0 a_def p_def q_def cfrac_nth_e) +next + case (Suc n) + note [simp] = + conjunct1[OF Suc.IH] conjunct1[OF conjunct2[OF Suc.IH]] conjunct2[OF conjunct2[OF Suc.IH]] + have [simp]: "3 + m = Suc (Suc (Suc m))" for m by simp + + have A': "A (Suc n) = of_int (q (3 * Suc n)) * exp 1 - of_int (p (3 * Suc n))" + unfolding A_Suc + by (subst p_rec0 q_rec0, simp)+ (auto simp: algebra_simps) + have B': "B (Suc n) = of_int (p (3 * Suc n + 1)) - of_int (q (3 * Suc n + 1)) * exp 1" + unfolding B_Suc + by (subst p_rec1 q_rec1 p_rec0 q_rec0, simp)+ (auto simp: algebra_simps A_Suc) + have C': "C (Suc n) = of_int (p (3*Suc n+2)) - of_int (q (3*Suc n+2)) * exp 1" + unfolding A_Suc B_Suc C_Suc using p_rec2[of n] q_rec2[of n] + by ((subst p_rec2 q_rec2)+, (subst p_rec0 q_rec0 p_rec1 q_rec1, simp)+) + (auto simp: algebra_simps A_Suc B_Suc) + from A' B' C' show ?case by simp +qed + +lemma q_pos: "q n > 0" if "n \ 1" + using that by (auto simp: q_def) + +lemma conv_diff_exp_bound: "norm (exp 1 - p n / q n) \ exp 1 / fact (n div 3)" +proof (cases "n = 1") + case False + define n' where "n' = n div 3" + consider "n mod 3 = 0" | "n mod 3 = 1" | "n mod 3 = 2" + by force + hence diff [unfolded n'_def]: "q n * exp 1 - p n = + (if n mod 3 = 0 then A n' else if n mod 3 = 1 then -B n' else -C n')" + proof cases + assume "n mod 3 = 0" + hence "3 * n' = n" unfolding n'_def by presburger + with ABC[of n'] show ?thesis by auto + next + assume *: "n mod 3 = 1" + hence "Suc (3 * n') = n" unfolding n'_def by presburger + with * ABC[of n'] show ?thesis by auto + next + assume *: "n mod 3 = 2" + hence "Suc (Suc (3 * n')) = n" unfolding n'_def by presburger + with * ABC[of n'] show ?thesis by auto + qed + + note [[linarith_split_limit = 0]] + have "norm ((q n * exp 1 - p n) / q n) \ exp 1 / fact (n div 3) / 1" unfolding diff norm_divide + using A_bound[of "n div 3"] B_bound[of "n div 3"] C_bound[of "n div 3"] q_pos[OF \n \ 1\] + by (subst frac_le) (auto simp: of_nat_ge_1_iff) + also have "(q n * exp 1 - p n) / q n = exp 1 - p n / q n" + using q_pos[OF \n \ 1\] by (simp add: divide_simps) + finally show ?thesis by simp +qed (auto simp: p_def q_def) + +theorem e_cfrac: "cfrac_lim e_cfrac = exp 1" +proof - + have num: "conv_num e_cfrac n = p (n + 2)" + and denom: "conv_denom e_cfrac n = q (n + 2)" for n + by (simp_all add: p_def q_def) + + have "(\n. exp 1 - p n / q n) \ 0" + proof (rule Lim_null_comparison) + show "eventually (\n. norm (exp 1 - p n / q n) \ exp 1 / fact (n div 3)) at_top" + using conv_diff_exp_bound by (intro always_eventually) auto + show "(\n. exp 1 / fact (n div 3) :: real) \ 0" + by (rule real_tendsto_divide_at_top tendsto_const filterlim_div_nat_at_top + filterlim_ident filterlim_compose[OF fact_real_at_top])+ auto + qed + moreover have "eventually (\n. exp 1 - p n / q n = exp 1 - conv e_cfrac (n - 2)) at_top" + using eventually_ge_at_top[of 2] + proof eventually_elim + case (elim n) + with num[of "n - 2"] denom[of "n - 2"] wf show ?case + by (simp add: eval_nat_numeral Suc_diff_Suc conv_num_denom) + qed + ultimately have "(\n. exp 1 - conv e_cfrac (n - 2)) \ 0" + using Lim_transform_eventually by fast + hence "(\n. exp 1 - (exp 1 - conv e_cfrac (Suc (Suc n) - 2))) \ exp 1 - 0" + by (subst filterlim_sequentially_Suc)+ (intro tendsto_diff tendsto_const) + hence "conv e_cfrac \ exp 1" by simp + moreover have "conv e_cfrac \ cfrac_lim e_cfrac" + by (intro LIMSEQ_cfrac_lim wf) auto + ultimately have "exp 1 = cfrac_lim e_cfrac" + by (rule LIMSEQ_unique) + thus ?thesis .. +qed + +corollary e_cfrac_altdef: "e_cfrac = cfrac_of_real (exp 1)" + by (metis e_cfrac cfrac_infinite_iff cfrac_length_e cfrac_of_real_cfrac_lim_irrational) + +text \ + This also provides us with a nice proof that $e$ is not rational and not a quadratic irrational + either. +\ +corollary exp1_irrational: "(exp 1 :: real) \ \" + by (metis cfrac_length_e e_cfrac cfrac_infinite_iff) + +corollary exp1_not_quadratic_irrational: "\quadratic_irrational (exp 1 :: real)" +proof - + have "range (\n. 2 * (int n + 1)) \ range (cfrac_nth e_cfrac)" + proof safe + fix n :: nat + have "cfrac_nth e_cfrac (3*n+2) \ range (cfrac_nth e_cfrac)" + by blast + also have "(3 * n + 2) mod 3 = 2" + by presburger + hence "cfrac_nth e_cfrac (3*n+2) = 2 * (int n + 1)" + by (simp add: cfrac_nth_e) + finally show "2 * (int n + 1) \ range (cfrac_nth e_cfrac)" . + qed + moreover have "infinite (range (\n. 2 * (int n + 1)))" + by (subst finite_image_iff) (auto intro!: injI) + ultimately have "infinite (range (cfrac_nth e_cfrac))" + using finite_subset by blast + thus ?thesis using quadratic_irrational_cfrac_nth_range_finite[of e_cfrac] + by (auto simp: e_cfrac) +qed + +end +end \ No newline at end of file diff --git a/thys/Continued_Fractions/Pell_Continued_Fraction.thy b/thys/Continued_Fractions/Pell_Continued_Fraction.thy new file mode 100644 --- /dev/null +++ b/thys/Continued_Fractions/Pell_Continued_Fraction.thy @@ -0,0 +1,421 @@ +(* + File: Continued_Fractions/Pell_Continued_Fraction.thy + Author: Manuel Eberl, University of Innsbruck +*) +section \The Connection between the continued fraction expansion of square roots and Pell's equation\ +theory Pell_Continued_Fraction +imports + Sqrt_Nat_Cfrac + Pell.Pell_Algorithm + Polynomial_Factorization.Prime_Factorization + Pell_Lifting +begin + +lemma irrational_times_int_eq_intD: + assumes "p * real_of_int a = real_of_int b" + assumes "p \ \" + shows "a = 0 \ b = 0" +proof - + have "a = 0" + proof (rule ccontr) + assume "a \ 0" + with assms(1) have "p = b / a" by (auto simp: field_simps) + also have "\ \ \" by auto + finally show False using assms(2) by contradiction + qed + with assms show ?thesis by simp +qed + + +text \ + The solutions to Pell's equation for some non-square \D\ are linked to the continued + fraction expansion of $\sqrt{D}$, which we shall show here. +\ + +context + fixes D :: nat and c h k P Q l + assumes nonsquare: "\is_square D" + defines "c \ cfrac_of_real (sqrt D)" + defines "h \ conv_num c" and "k \ conv_denom c" + defines "P \ fst \ sqrt_remainder_surd D" and "Q \ snd \ sqrt_remainder_surd D" + defines "l \ sqrt_nat_period_length D" +begin + +interpretation pell D + by unfold_locales fact+ + +lemma cfrac_length_infinite [simp]: "cfrac_length c = \" +proof - + have "sqrt D \ \" + using nonsquare by (simp add: irrat_sqrt_nonsquare) + thus ?thesis + by (simp add: c_def) +qed + +lemma conv_num_denom_pell: + "h 0 ^ 2 - D * k 0 ^ 2 < 0" + "m > 0 \ h m ^ 2 - D * k m ^ 2 = (-1) ^ Suc m * Q m" +proof - + define D' where "D' = Discrete.sqrt D" + have "h 0 ^ 2 - D * k 0 ^ 2 = int (D' ^ 2) - int D" + by (simp_all add: h_def k_def c_def Discrete_sqrt_altdef D'_def) + also { + have "int (D' ^ 2) - int D \ 0" + using Discrete.sqrt_power2_le[of D] by (simp add: D'_def) + moreover have "D \ D' ^ 2" using nonsquare by auto + ultimately have "int (D' ^ 2) - int D < 0" by linarith + } + finally show "h 0 ^ 2 - D * k 0 ^ 2 < 0" . +next + assume "m > 0" + define n where "n = m - 1" + define \ where "\ = cfrac_remainder c" + define \' where "\' = sqrt_remainder_surd D" + have m: "m = Suc n" using \m > 0\ by (simp add: n_def) + from nonsquare have "D > 1" + by (cases D) (auto intro!: Nat.gr0I) + from nonsquare have irrat: "sqrt D \ \" + using irrat_sqrt_nonsquare by blast + have [simp]: "cfrac_lim c = sqrt D" + using irrat \D > 1\ by (simp add: c_def) + have \_pos: "\ n > 0" for n + unfolding \_def using wf \D > 1\ cfrac_remainder_pos[of c n] + by (cases "n = 0") auto + have \': "\' n = (P n, Q n)" for n by (simp add: \'_def P_def Q_def) + have Q_pos: "Q n > 0" for n + using snd_sqrt_remainder_surd_pos[OF nonsquare] by (simp add: Q_def) + have k_pos: "k n > 0" for n + by (auto simp: k_def intro!: conv_denom_pos) + have k_nonneg: "k n \ 0" for n + by (auto simp: k_def intro!: conv_denom_nonneg) + + let ?A = "(sqrt D + P (n + 1)) * h (n + 1) + Q (n + 1) * h n" + let ?B = "(sqrt D + P (n + 1)) * k (n + 1) + Q (n + 1) * k n" + have "?B > 0" using k_pos Q_pos k_nonneg + by (intro add_nonneg_pos mult_nonneg_nonneg add_nonneg_nonneg) auto + + have "sqrt D = conv' c (Suc (Suc n)) (\ (Suc (Suc n)))" + unfolding \_def by (subst conv'_cfrac_remainder) auto + also have "\ = (\ (n + 2) * h (n + 1) + h n) / (\ (n + 2) * k (n + 1) + k n)" + using wf \_pos by (subst conv'_num_denom) (simp_all add: h_def k_def) + also have "\ (n + 2) = surd_to_real D (\' (Suc n))" + using surd_to_real_sqrt_remainder_surd[OF nonsquare, of "Suc n"] + by (simp add: \'_def \_def c_def) + also have "\ = (sqrt D + P (Suc n)) / Q (Suc n)" (is "_ = ?\") + by (simp add: \' surd_to_real_def) + also have "?\ * h (n + 1) + h n = + 1 / Q (n + 1) * ((sqrt D + P (n + 1)) * h (n + 1) + Q (n + 1) * h n)" + using Q_pos by (simp add: field_simps) + also have "?\ * k (n + 1) + k n = + 1 / Q (n + 1) * ((sqrt D + P (n + 1)) * k (n + 1) + Q (n + 1) * k n)" + (is "_ = ?f k") using Q_pos by (simp add: field_simps) + also have "?f h / ?f k = ((sqrt D + P (n + 1)) * h (n + 1) + Q (n + 1) * h n) / + ((sqrt D + P (n + 1)) * k (n + 1) + Q (n + 1) * k n)" + (is "_ = ?A / ?B") using Q_pos by (intro mult_divide_mult_cancel_left) auto + finally have "sqrt D * ?B = ?A" + using \?B > 0\ by (simp add: divide_simps) + moreover have "sqrt D * sqrt D = D" by simp + ultimately have "sqrt D * (P (n + 1) * k (n + 1) + Q (n + 1) * k n - h (n + 1)) = + P (n + 1) * h (n + 1) + Q (n + 1) * h n - k (n + 1) * D" + unfolding of_int_add of_int_mult of_int_diff of_int_of_nat_eq of_nat_mult of_nat_add + by Groebner_Basis.algebra + from irrational_times_int_eq_intD[OF this] irrat + have 1: "h (Suc n) = P (Suc n) * k (Suc n) + Q (Suc n) * k n" + and 2: "D * k (Suc n) = P (Suc n) * h (Suc n) + Q (Suc n) * h n" + by (simp_all del: of_nat_add of_nat_mult) + + have "h (Suc n) * h (Suc n) - D * k (Suc n) * k (Suc n) = + Q (Suc n) * (k n * h (Suc n) - k (Suc n) * h n)" + by (subst 1, subst 2) (simp add: algebra_simps) + also have "k n * h (Suc n) - k (Suc n) * h n = (-1) ^ n" + unfolding h_def k_def by (rule conv_num_denom_prod_diff) + finally have "h (Suc n) ^ 2 - D * k (Suc n) ^ 2 = (-1) ^ n * Q (Suc n)" + by (simp add: power2_eq_square algebra_simps) + thus "h m ^ 2 - D * k m ^ 2 = (-1) ^ Suc m * Q m" + by (simp add: m) +qed + +text \ + Every non-trivial solution to Pell's equation is a convergent in the expansion of $\sqrt{D}$: +\ +theorem pell_solution_is_conv: + assumes "x\<^sup>2 = Suc (D * y\<^sup>2)" and "y > 0" + shows "(int x, int y) \ range (\n. (conv_num c n, conv_denom c n))" +proof - + have "\n. enat n \ cfrac_length c \ (int x, int y) = (conv_num c n, conv_denom c n)" + proof (rule frac_is_convergentI) + have "gcd (x\<^sup>2) (y\<^sup>2) = 1" unfolding assms(1) + using gcd_add_mult[of "y\<^sup>2" D 1] by (simp add: gcd.commute) + thus "coprime (int x) (int y)" + by (simp add: coprime_iff_gcd_eq_1) + next + from assms have "D > 1" + using nonsquare by (cases D) (auto intro!: Nat.gr0I) + hence pos: "x + y * sqrt D > 0" using assms + by (intro add_nonneg_pos) auto + + from assms have "real (x\<^sup>2) = real (Suc (D * y\<^sup>2))" + by (simp only: of_nat_eq_iff) + hence "1 = real x ^ 2 - D * real y ^ 2" + unfolding of_nat_power by simp + also have "\ = (x - y * sqrt D) * (x + y * sqrt D)" + by (simp add: field_simps power2_eq_square) + finally have *: "x - y * sqrt D = 1 / (x + y * sqrt D)" + using pos by (simp add: field_simps) + + from pos have "0 < 1 / (x + y * sqrt D)" + by (intro divide_pos_pos) auto + also have "\ = x - y * sqrt D" by (rule * [symmetric]) + finally have less: "y * sqrt D < x" by simp + + have "sqrt D - x / y = -((x - y * sqrt D) / y)" + using \y > 0\ by (simp add: field_simps) + also have "\\\ = (x - y * sqrt D) / y" + using less by simp + also have "(x - y * sqrt D) / y = 1 / (y * (x + y * sqrt D))" + using \y > 0\ by (subst *) auto + also have "\ \ 1 / (y * (y * sqrt D + y * sqrt D))" + using \y > 0\ \D > 1\ pos less + by (intro divide_left_mono mult_left_mono add_right_mono mult_pos_pos) auto + also have "\ = 1 / (2 * y\<^sup>2 * sqrt D)" + by (simp add: power2_eq_square) + also have "\ < 1 / (real (2 * y\<^sup>2) * 1)" using \y > 0\ \D > 1\ + by (intro divide_strict_left_mono mult_strict_left_mono mult_pos_pos) auto + finally show "\cfrac_lim c - int x / int y\ < 1 / (2 * int y ^ 2)" + unfolding c_def using irrat_sqrt_nonsquare[of D] \\is_square D\ by simp + qed (insert assms irrat_sqrt_nonsquare[of D], auto simp: c_def) + thus ?thesis by auto +qed + +text \ + Let \l\ be the length of the period in the continued fraction expansion of $\sqrt{D}$ and let + $h_i$ and $k_i$ be the numerator and denominator of the \i\-th convergent. + + Then the non-trivial solutions of Pell's equation are exactly the pairs of the form + $(h_{lm -1}, k_{lm-1})$ for any \m\ such that $lm$ is even. +\ +lemma nontriv_solution_iff_conv_num_denom: + "nontriv_solution (x, y) \ + (\m>0. int x = h (l * m - 1) \ int y = k (l * m - 1) \ even (l * m))" +proof safe + fix m assume xy: "x = h (l * m - 1)" "y = k (l * m - 1)" + and lm: "even (l * m)" and m: "m > 0" + have l: "l > 0" using period_nonempty[OF nonsquare] by (auto simp: l_def) + from lm have "l * m \ 1" by (intro notI) auto + with l m have lm': "l * m > 1" by (cases "l * m") auto + + have "(h (l * m - 1))\<^sup>2 - D * (k (l * m - 1))\<^sup>2 = + (- 1) ^ Suc (l * m - 1) * int (Q (l * m - 1))" + using lm' by (intro conv_num_denom_pell) auto + also have "(- 1) ^ Suc (l * m - 1) = (1 :: int)" + using lm l m by (subst neg_one_even_power) auto + also have "Q (l * m - 1) = Q ((l * m - 1) mod l)" + unfolding Q_def l_def o_def by (subst sqrt_remainder_surd_periodic[OF nonsquare]) simp + also { + have "l * m - 1 = (m - 1) * l + (l - 1)" + using m l lm' by (cases m) (auto simp: mult_ac) + also have "\ mod l = (l - 1) mod l" + by simp + also have "\ = l - 1" + using l by (intro mod_less) auto + also have "Q \ = 1" + using sqrt_remainder_surd_last[OF nonsquare] by (simp add: Q_def l_def) + finally have "Q ((l * m - 1) mod l) = 1" . + } + finally have "h (l * m - 1) ^ 2 = D * k (l * m - 1) ^ 2 + 1" + unfolding of_nat_Suc by (simp add: algebra_simps) + hence "h (l * m - 1) ^ 2 = D * k (l * m - 1) ^ 2 + 1" + by (simp only: of_nat_eq_iff) + moreover have "k (l * m - 1) > 0" + unfolding k_def by (intro conv_denom_pos) + ultimately have "nontriv_solution (int x, int y)" + using xy by (simp add: nontriv_solution_def) + thus "nontriv_solution (x, y)" + by simp +next + assume "nontriv_solution (x, y)" + hence asm: "x ^ 2 = Suc (D * y ^ 2)" "y > 0" + by (auto simp: nontriv_solution_def abs_square_eq_1 intro!: Nat.gr0I) + from asm have asm': "int x ^ 2 = int D * int y ^ 2 + 1" + by (metis add.commute of_nat_Suc of_nat_mult of_nat_power_eq_of_nat_cancel_iff) + have l: "l > 0" using period_nonempty[OF nonsquare] by (auto simp: l_def) + from pell_solution_is_conv[OF asm] obtain m where + xy: "h m = x" "k m = y" by (auto simp: c_def h_def k_def) + + have m: "m > 0" + using asm' conv_num_denom_pell(1) xy by (intro Nat.gr0I) auto + have "1 = h m ^ 2 - D * k m ^ 2" + using asm' xy by simp + also have "\ = (- 1) ^ Suc m * int (Q m)" + using conv_num_denom_pell(2)[OF m] . + finally have *: "(- 1) ^ Suc m * int (Q m) = 1" .. + from * have m': "odd m \ Q m = 1" + by (cases "even m") auto + + define n where "n = Suc m div l" + have "l dvd Suc m" + proof (rule ccontr) + assume *: "\(l dvd Suc m)" + have "Q m = Q (m mod l)" + unfolding Q_def l_def o_def by (subst sqrt_remainder_surd_periodic[OF nonsquare]) simp + also { + have "m mod l < l" using \l > 0\ by simp + moreover have "Suc (m mod l) \ l" using * l \m > 0\ + using mod_Suc[of m l] by auto + ultimately have "m mod l < l - 1" by simp + hence "Q (m mod l) > 1" unfolding Q_def o_def l_def + by (rule snd_sqrt_remainder_surd_gt_1[OF nonsquare]) + } + finally show False using m' by simp + qed + hence m_eq: "Suc m = n * l" "m = n * l - 1" + by (simp_all add: n_def) + hence "n > 0" by (auto intro!: Nat.gr0I) + thus "\n>0. int x = h (l * n - 1) \ int y = k (l * n - 1) \ even (l * n)" + using xy m_eq m' by (intro exI[of _ n]) (auto simp: mult_ac) +qed + +text \ + Consequently, the fundamental solution is $(h_n, k_n)$ where $n = l - 1$ + if \l\ is even and $n = 2l-1$ otherwise: +\ +lemma fund_sol_conv_num_denom: + defines "n \ if even l then l - 1 else 2 * l - 1" + shows "fund_sol = (nat (h n), nat (k n))" +proof (rule fund_sol_eq_sndI) + have [simp]: "h n \ 0" "k n \ 0" for n + by (auto simp: h_def k_def c_def intro!: conv_num_nonneg) + show "nontriv_solution (nat (h n), nat (k n))" + by (subst nontriv_solution_iff_conv_num_denom, rule exI[of _ "if even l then 1 else 2"]) + (simp_all add: n_def mult_ac) +next + fix x y :: nat assume "nontriv_solution (x, y)" + then obtain m where m: "m > 0" "x = h (l * m - 1)" "y = k (l * m - 1)" "even (l * m)" + by (subst (asm) nontriv_solution_iff_conv_num_denom) auto + have l: "l > 0" using period_nonempty[OF nonsquare] by (auto simp: l_def) + from m l have "Suc n \ l * m" by (auto simp: n_def) + hence "n \ l * m - 1" by simp + hence "k n \ k (l * m - 1)" + unfolding k_def c_def using irrat_sqrt_nonsquare[OF nonsquare] + by (intro conv_denom_leI) auto + with m show "nat (k n) \ y" by simp +qed + +end + + +text \ + The following algorithm computes the fundamental solution (or the dummy result \(0, 0)\ if + \D\ is a square) fairly quickly by computing the continued fraction expansion of $\sqrt{D}$ + and then computing the fundamental solution as the appropriate convergent. +\ +lemma find_fund_sol_code [code]: + "find_fund_sol D = + (let info = sqrt_cfrac_info_array D; + l = fst info + in if l = 0 then (0, 0) else + let + c = cfrac_sqrt_nth info; + n = if even l then l - 1 else 2 * l - 1 + in + (nat (conv_num_fun c n), nat (conv_denom_fun c n)))" +proof - + have *: "is_cfrac (cfrac_sqrt_nth (sqrt_cfrac_info_array D))" if "\is_square D" + using that cfrac_sqrt_nth[of D] unfolding is_cfrac_def + by (metis cfrac_nth_nonzero neq0_conv of_nat_0 of_nat_0_less_iff) + have **: "cfrac (\x. int (cfrac_sqrt_nth (sqrt_cfrac_info_array D) x)) = cfrac_of_real (sqrt D)" + if "\is_square D" + using that cfrac_sqrt_nth[of D] * by (intro cfrac_eqI) auto + show ?thesis using * ** + by (auto simp: square_test_correct find_fund_sol_correct conv_num_fun_eq conv_denom_fun_eq + Let_def cfrac_sqrt_nth fund_sol_conv_num_denom conv_num_nonneg) +qed + +lemma find_nth_solution_square [simp]: "is_square D \ find_nth_solution D n = (0, 0)" + by (simp add: find_nth_solution_def) + +lemma fst_find_fund_sol_eq_0_iff [simp]: "fst (find_fund_sol D) = 0 \ is_square D" +proof (cases "is_square D") + case False + then interpret pell D by unfold_locales + from False have "find_fund_sol D = fund_sol" by (simp add: find_fund_sol_correct) + moreover from fund_sol_is_nontriv_solution have "fst fund_sol > 0" + by (auto simp: nontriv_solution_def intro!: Nat.gr0I) + ultimately show ?thesis using False + by (simp add: find_fund_sol_def square_test_correct split: if_splits) +qed (auto simp: find_fund_sol_def square_test_correct) + +text \ + Arbitrary solutions can now be computed as powers of the fundamental solution. +\ +lemma find_nth_solution_code [code]: + "find_nth_solution D n = + (let xy = find_fund_sol D + in if fst xy = 0 then (0, 0) else efficient_pell_power D xy n)" +proof (cases "is_square D") + case False + then interpret pell D by unfold_locales + from fund_sol_is_nontriv_solution have "fst fund_sol > 0" + by (auto simp: nontriv_solution_def intro!: Nat.gr0I) + thus ?thesis using False + by (simp add: find_nth_solution_correct Let_def nth_solution_def pell_power_def + pell_mul_commutes[of _ "fund_sol"] find_fund_sol_correct) +qed auto + +lemma nth_solution_code [code]: + "pell.nth_solution D n = + (let info = sqrt_cfrac_info_array D; + l = fst info + in if l = 0 then + Code.abort (STR ''nth_solution is undefined for perfect square parameter.'') + (\_. pell.nth_solution D n) + else + let + c = cfrac_sqrt_nth info; + m = if even l then l - 1 else 2 * l - 1; + fund_sol = (nat (conv_num_fun c m), nat (conv_denom_fun c m)) + in + efficient_pell_power D fund_sol n)" +proof (cases "is_square D") + case False + then interpret pell by unfold_locales + have *: "is_cfrac (cfrac_sqrt_nth (sqrt_cfrac_info_array D))" + using False cfrac_sqrt_nth[of D] unfolding is_cfrac_def + by (metis cfrac_nth_nonzero neq0_conv of_nat_0 of_nat_0_less_iff) + have **: "cfrac (\x. int (cfrac_sqrt_nth (sqrt_cfrac_info_array D) x)) = cfrac_of_real (sqrt D)" + using False cfrac_sqrt_nth[of D] * by (intro cfrac_eqI) auto + + from False * ** show ?thesis + by (auto simp: Let_def cfrac_sqrt_nth fund_sol_conv_num_denom nth_solution_def + pell_power_def pell_mul_commutes[of _ "(_, _)"] + conv_num_fun_eq conv_denom_fun_eq conv_num_nonneg) +qed auto + +lemma fund_sol_code [code]: + "pell.fund_sol D = (let info = sqrt_cfrac_info_array D; + l = fst info + in if l = 0 then + Code.abort (STR ''fund_sol is undefined for perfect square parameter.'') + (\_. pell.fund_sol D) + else + let + c = cfrac_sqrt_nth info; + n = if even l then l - 1 else 2 * l - 1 + in + (nat (conv_num_fun c n), nat (conv_denom_fun c n)))" +proof (cases "is_square D") + case False + then interpret pell by unfold_locales + have *: "is_cfrac (cfrac_sqrt_nth (sqrt_cfrac_info_array D))" + using False cfrac_sqrt_nth[of D] unfolding is_cfrac_def + by (metis cfrac_nth_nonzero neq0_conv of_nat_0 of_nat_0_less_iff) + have **: "cfrac (\x. int (cfrac_sqrt_nth (sqrt_cfrac_info_array D) x)) = cfrac_of_real (sqrt D)" + using False cfrac_sqrt_nth[of D] * by (intro cfrac_eqI) auto + + from False * ** show ?thesis + by (auto simp: Let_def cfrac_sqrt_nth fund_sol_conv_num_denom nth_solution_def + pell_power_def pell_mul_commutes[of _ "(_, _)"] + conv_num_fun_eq conv_denom_fun_eq conv_num_nonneg) +qed auto + +end \ No newline at end of file diff --git a/thys/Continued_Fractions/Pell_Continued_Fraction_Tests.thy b/thys/Continued_Fractions/Pell_Continued_Fraction_Tests.thy new file mode 100644 --- /dev/null +++ b/thys/Continued_Fractions/Pell_Continued_Fraction_Tests.thy @@ -0,0 +1,98 @@ +(* + File: Continued_Fractions/Pell_Continued_Fraction_Tests.thy + Author: Manuel Eberl, University of Innsbruck +*) +section \Tests for Continued Fractions of Square Roots and Pell's Equation\ +theory Pell_Continued_Fraction_Tests +imports + Pell.Efficient_Discrete_Sqrt + "HOL-Library.Code_Lazy" + "HOL-Library.Code_Target_Numeral" + Pell_Continued_Fraction + Pell_Lifting +begin + +code_lazy_type stream + +(* TODO Move *) +lemma lnth_code [code]: + "lnth xs 0 = (if lnull xs then undefined (0 :: nat) else lhd xs)" + "lnth xs (Suc n) = (if lnull xs then undefined (Suc n) else lnth (ltl xs) n)" + by (auto simp: lnth.simps split: llist.splits) + +value "let c = sqrt_cfrac 1339 in map (cfrac_nth c) [0..<30]" + + +fun arg_max_list where + "arg_max_list _ [] = undefined" +| "arg_max_list f (x # xs) = + foldl (\(x, y) x'. let y' = f x' in if y' > y then (x', y') else (x, y)) (x, f x) xs" + + +value [code] "sqrt_cfrac_info 17" +value [code] "sqrt_cfrac_info 1339" +value [code] "sqrt_cfrac_info 121" +value [code] "sqrt_nat_period_length 410286423278424" + +text \ + For which number $D < 100000$ does $\sqrt{D}$ have the longest period? +\ +value [code] "arg_max_list sqrt_nat_period_length [0..<100000]" + + +subsection \Fundamental solutions of Pell's equation\ + +value [code] "pell.fund_sol 12" +value [code] "pell.fund_sol 13" +value [code] "pell.fund_sol 61" +value [code] "pell.fund_sol 661" +value [code] "pell.fund_sol 6661" +value [code] "pell.fund_sol 4729494" + +text \ + Project Euler problem \#66: + For which $D < 1000$ does Pell's equation have the largest fundamental solution? +\ + +value [code] "arg_max_list (fst \ find_fund_sol) [0..<1001]" + + +text \ + The same for $D < 100000$: +\ + +value [code] "arg_max_list (fst \ find_fund_sol) [0..<100000]" + + +text \ + The solution to the next example, which is at the core of Archimedes' cattle problem, + is so big that termifying the result takes extremely long. + Therefore, we simply compute the number of decimal digits in the result instead. +\ + +fun log10_aux :: "nat \ nat \ nat" where + "log10_aux acc n = + (if n \ 10000000000 then log10_aux (acc + 10) (n div 10000000000) + else if n = 0 then acc else log10_aux (Suc acc) (n div 10))" + +definition log10 where "log10 = log10_aux 0" + +value [code] "map_prod log10 log10 (pell.fund_sol 410286423278424)" + +text \ + Factoring out the square factor \9314\<^sup>2\ does yield a significant speed-up in this case: +\ +value [code] "map_prod log10 log10 (find_fund_sol_fast 410286423278424)" + + +subsection \Tests for other operations\ + +value [code] "pell.nth_solution 13 100" +value [code] "pell.nth_solution 4729494 3" + +value [code] "stake 10 (pell_solutions 13)" +value [code] "stake 10 (pell_solutions 61)" + +value [code] "pell.nth_solution 23 8" + +end \ No newline at end of file diff --git a/thys/Continued_Fractions/Pell_Lifting.thy b/thys/Continued_Fractions/Pell_Lifting.thy new file mode 100644 --- /dev/null +++ b/thys/Continued_Fractions/Pell_Lifting.thy @@ -0,0 +1,395 @@ +(* + File: Continued_Fractions/Pell_Lifting.thy + Author: Manuel Eberl, University of Innsbruck +*) +section \Lifting solutions of Pell's Equation\ +theory Pell_Lifting + imports Pell.Pell Pell.Pell_Algorithm +begin + +subsection \Auxiliary material\ + +(* TODO Move *) +lemma (in pell) snth_pell_solutions: "snth (pell_solutions D) n = nth_solution n" + by (simp add: pell_solutions_def Let_def find_fund_sol_correct nonsquare_D nth_solution_def + pell_power_def pell_mul_commutes[of _ fund_sol]) + +definition square_squarefree_part_nat :: "nat \ nat \ nat" where + "square_squarefree_part_nat n = (square_part n, squarefree_part n)" + +lemma prime_factorization_squarefree_part: + assumes "x \ 0" + shows "prime_factorization (squarefree_part x) = + mset_set {p \ prime_factors x. odd (multiplicity p x)}" (is "?lhs = ?rhs") +proof (rule multiset_eqI) + fix p show "count ?lhs p = count ?rhs p" + proof (cases "prime p") + case False + thus ?thesis by (auto simp: count_prime_factorization) + next + case True + have "finite (prime_factors x)" by simp + hence "finite {p. p dvd x \ prime p}" using assms + by (subst (asm) prime_factors_dvd) (auto simp: conj_commute) + hence "finite {p. p dvd x \ prime p \ odd (multiplicity p x)}" + by (rule finite_subset [rotated]) auto + moreover have "odd (n :: nat) \ n mod 2 = Suc 0" for n by presburger + ultimately show ?thesis using assms + by (cases "p dvd x"; cases "even (multiplicity p x)") + (auto simp: count_prime_factorization prime_multiplicity_squarefree_part + in_prime_factors_iff not_dvd_imp_multiplicity_0) + qed +qed + +lemma squarefree_part_nat: + "squarefree_part (n :: nat) = (\{p \ prime_factors n. odd (multiplicity p n)})" +proof (cases "n = 0") + case False + hence "(\{p \ prime_factors n. odd (multiplicity p n)}) = + prod_mset (prime_factorization (squarefree_part n))" + by (subst prime_factorization_squarefree_part) (auto simp: prod_unfold_prod_mset) + also have "\ = squarefree_part n" + by (intro prod_mset_prime_factorization_nat Nat.gr0I) auto + finally show ?thesis .. +qed auto + +lemma prime_factorization_square_part: + assumes "x \ 0" + shows "prime_factorization (square_part x) = + (\p \ prime_factors x. replicate_mset (multiplicity p x div 2) p)" (is "?lhs = ?rhs") +proof (rule multiset_eqI) + fix p show "count ?lhs p = count ?rhs p" + proof (cases "prime p \ p dvd x") + case False + thus ?thesis by (auto simp: count_prime_factorization count_sum + prime_multiplicity_square_part not_dvd_imp_multiplicity_0) + next + case True + thus ?thesis using assms + by (cases "p dvd x") + (auto simp: count_prime_factorization prime_multiplicity_squarefree_part + in_prime_factors_iff count_sum prime_multiplicity_square_part) + qed +qed + +lemma prod_mset_sum: "prod_mset (sum f A) = (\x\A. prod_mset (f x))" + by (induction A rule: infinite_finite_induct) auto + +lemma square_part_nat: + assumes "n > 0" + shows "square_part (n :: nat) = (\p \ prime_factors n. p ^ (multiplicity p n div 2))" +proof - + have "(\p \ prime_factors n. p ^ (multiplicity p n div 2)) = + prod_mset (prime_factorization (square_part n))" using assms + by (subst prime_factorization_square_part) (auto simp: prod_unfold_prod_mset prod_mset_sum) + also have "\ = square_part n" using assms + by (intro prod_mset_prime_factorization_nat Nat.gr0I) auto + finally show ?thesis .. +qed + +lemma square_squarefree_part_nat_code [code]: + "square_squarefree_part_nat n = (if n = 0 then (0, 1) + else let ps = prime_factorization n + in ((\p\set_mset ps. p ^ (count ps p div 2)), + \(Set.filter (\p. odd (count ps p)) (set_mset ps))))" + by (cases "n = 0") + (auto simp: Let_def square_squarefree_part_nat_def squarefree_part_nat Set.filter_def + count_prime_factorization square_part_nat intro!: prod.cong) + +lemma square_part_nat_code [code_unfold]: + "square_part (n :: nat) = (if n = 0 then 0 + else let ps = prime_factorization n in (\p\set_mset ps. p ^ (count ps p div 2)))" + using square_squarefree_part_nat_code[of n] + by (simp add: square_squarefree_part_nat_def Let_def split: if_splits) + +lemma squarefree_part_nat_code [code_unfold]: + "squarefree_part (n :: nat) = (if n = 0 then 1 + else let ps = prime_factorization n in (\(Set.filter (\p. odd (count ps p)) (set_mset ps))))" + using square_squarefree_part_nat_code[of n] + by (simp add: square_squarefree_part_nat_def Let_def split: if_splits) + +lemma is_nth_power_mult_nth_powerD: + assumes "is_nth_power n (a * b ^ n)" "b > 0" "n > 0" + shows "is_nth_power n (a::nat)" +proof - + from assms obtain k where k: "k ^ n = a * b ^ n" + by (auto elim: is_nth_powerE) + with assms(2,3) have "b dvd k" + by (metis dvd_triv_right pow_divides_pow_iff) + then obtain l where "k = b * l" + by auto + with k have "a = l ^ n" using assms(2) + by (simp add: power_mult_distrib) + thus ?thesis by auto +qed + +lemma (in pell) fund_sol_eq_fstI: + assumes "nontriv_solution (x, y)" + assumes "\x' y'. nontriv_solution (x', y') \ x \ x'" + shows "fund_sol = (x, y)" +proof - + have "x = fst fund_sol" + using fund_sol_is_nontriv_solution assms(1) fund_sol_minimal''[of "(x, y)"] + by (auto intro!: antisym assms(2)[of "fst fund_sol" "snd fund_sol"]) + moreover from this have "y = snd fund_sol" + using assms(1) solutions_linorder_strict[of x y "fst fund_sol" "snd fund_sol"] + fund_sol_is_nontriv_solution + by (auto simp: nontriv_solution_imp_solution prod_eq_iff) + ultimately show ?thesis by simp +qed + +lemma (in pell) fund_sol_eqI_fst': + assumes "nontriv_solution xy" + assumes "\x' y'. nontriv_solution (x', y') \ fst xy \ x'" + shows "fund_sol = xy" + using fund_sol_eq_fstI[of "fst xy" "snd xy"] assms by simp + +lemma (in pell) fund_sol_eq_sndI: + assumes "nontriv_solution (x, y)" + assumes "\x' y'. nontriv_solution (x', y') \ y \ y'" + shows "fund_sol = (x, y)" +proof - + have "y = snd fund_sol" + using fund_sol_is_nontriv_solution assms(1) fund_sol_minimal''[of "(x, y)"] + by (auto intro!: antisym assms(2)[of "fst fund_sol" "snd fund_sol"]) + moreover from this have "x = fst fund_sol" + using assms(1) solutions_linorder_strict[of x y "fst fund_sol" "snd fund_sol"] + fund_sol_is_nontriv_solution + by (auto simp: nontriv_solution_imp_solution prod_eq_iff) + ultimately show ?thesis by simp +qed + +lemma (in pell) fund_sol_eqI_snd': + assumes "nontriv_solution xy" + assumes "\x' y'. nontriv_solution (x', y') \ snd xy \ y'" + shows "fund_sol = xy" + using fund_sol_eq_sndI[of "fst xy" "snd xy"] assms by simp + + +subsection \The lifting mechanism\ + +text \ + The solutions of Pell's equations for parameters \D\ and \a\<^sup>2 D\ stand in correspondence to + one another: every solution \(x, y)\ for parameter \D\ can be lowered to a solution \(x, ay)\ + for \a\<^sup>2 D\, and every solution of the form \(x, ay)\ for parameter \a\<^sup>2 D\ can be lifted to a + solution \(x, y)\ for parameter \D\. +\ + +locale pell_lift = pell + + fixes a D' :: nat + assumes nz: "a > 0" + defines "D' \ D * a\<^sup>2" +begin + +lemma nonsquare_D': "\is_square D'" + using nonsquare_D is_nth_power_mult_nth_powerD[of 2 D a] nz by (auto simp: D'_def) + +definition lift_solution :: "nat \ nat \ nat \ nat" where + "lift_solution = (\(x, y). (x, y div a))" + +definition lower_solution :: "nat \ nat \ nat \ nat" where + "lower_solution = (\(x, y). (x, y * a))" + +definition liftable_solution :: "nat \ nat \ bool" where + "liftable_solution = (\(x, y). a dvd y)" + +sublocale lift: pell D' + by unfold_locales (fact nonsquare_D') + +lemma lift_solution_iff: "lift.solution xy \ solution (lower_solution xy)" + unfolding solution_def lift.solution_def + by (auto simp: lower_solution_def D'_def case_prod_unfold power_mult_distrib) + +lemma lift_solution: + assumes "solution xy" "liftable_solution xy" + shows "lift.solution (lift_solution xy)" + using assms unfolding solution_def lift.solution_def + by (auto simp: liftable_solution_def lift_solution_def D'_def case_prod_unfold power_mult_distrib + elim!: dvdE) + +text \ + In particular, the fundamental solution for \a\<^sup>2 D\ is the smallest liftable solution for \D\: +\ +lemma lift_fund_sol: + assumes "\n. 0 < n \ n < m \ \liftable_solution (nth_solution n)" + assumes "liftable_solution (nth_solution m)" "m > 0" + shows "lift.fund_sol = lift_solution (nth_solution m)" +proof (rule lift.fund_sol_eqI_fst') + from assms have "nontriv_solution (nth_solution m)" + by (intro nth_solution_sound') + hence "lift_solution (nth_solution m) \ (1, 0)" using nz assms(2) + by (auto simp: lift_solution_def case_prod_unfold nontriv_solution_def liftable_solution_def) + with assms show "lift.nontriv_solution (lift_solution (nth_solution m))" + by (auto simp: lift.nontriv_solution_altdef intro: lift_solution) +next + fix x' y' :: nat + assume *: "lift.nontriv_solution (x', y')" + hence nz': "x' \ 1" using nonsquare_D' + by (auto simp: lift.nontriv_solution_altdef lift.solution_def) + from * have "solution (lower_solution (x', y'))" + by (simp add: lift_solution_iff lift.nontriv_solution_altdef) + hence "lower_solution (x', y') \ range nth_solution" by (rule nth_solution_complete) + then obtain n where n: "nth_solution n = lower_solution (x', y')" by auto + with nz' have "n > 0" by (auto intro!: Nat.gr0I simp: nth_solution_def lower_solution_def) + with n have "liftable_solution (nth_solution n)" + by (auto simp: liftable_solution_def lower_solution_def) + with \n > 0\ and assms(1)[of n] have "n \ m" by (cases "n \ m") auto + hence "fst (nth_solution m) \ fst (nth_solution n)" + using strict_mono_less_eq[OF strict_mono_nth_solution(1)] by simp + thus "fst (lift_solution (nth_solution m)) \ x'" + by (simp add: lift_solution_def lower_solution_def n case_prod_unfold) +qed + +end + + +subsection \Accelerated computation of the fundamental solution for non-squarefree inputs\ + +text \ + Solving Pell's equation for some \D\ of the form \a\<^sup>2 D'\ can be done by solving + it for \D'\ and then lifting the solution. Thus, if \D\ is not squarefree, we can compute + its squarefree decomposition \a\<^sup>2 D'\ with \D'\ squarefree and thus speed up the computation + (since \D'\ is smaller than \D\). + + The squarefree decomposition can only be computed (according to current knowledge in mathematics) + through the prime decomposition. However, given how big the solutions are for even + moderate values of \D\, it is usually worth doing it if \D\ is not squarefree. +\ + +lemma squarefree_part_of_square [simp]: + assumes "is_square (x :: 'a :: {factorial_semiring, normalization_semidom_multiplicative})" + assumes "x \ 0" + shows "squarefree_part x = unit_factor x" +proof - + from assms obtain y where [simp]: "x = y ^ 2" + by (auto simp: is_nth_power_def) + have "unit_factor x * normalize x = squarefree_part x * square_part x ^ 2" + by (subst squarefree_decompose [symmetric]) auto + also have "\ = squarefree_part x * normalize x" + by (simp add: square_part_even_power normalize_power) + finally show ?thesis using assms + by (subst (asm) mult_cancel_right) auto +qed + +lemma squarefree_part_1_imp_square: + assumes "squarefree_part x = 1" + shows "is_square x" +proof - + have "is_square (square_part x ^ 2)" + by auto + also have "square_part x ^ 2 = squarefree_part x * square_part x ^ 2" + using assms by simp + also have "\ = x" + by (rule squarefree_decompose [symmetric]) + finally show ?thesis . +qed + + +definition find_fund_sol_fast where + "find_fund_sol_fast D = + (let (a, D') = square_squarefree_part_nat D + in + if D' = 0 \ D' = 1 then (0, 0) + else if a = 1 then pell.fund_sol D + else map_prod id (\y. y div a) + (shd (sdrop_while (\(_, y). y = 0 \ \a dvd y) (pell_solutions D'))))" + +lemma find_fund_sol_fast: "find_fund_sol D = find_fund_sol_fast D" +proof (cases "is_square D \ square_part D = 1") + case True + thus ?thesis + using squarefree_part_1_imp_square[of D] + by (cases "D = 0") + (auto simp: find_fund_sol_correct find_fund_sol_fast_def + square_squarefree_part_nat_def square_test_correct unit_factor_nat_def) +next + case False + define D' a where "D' = squarefree_part D" and "a = square_part D" + have "D > 0" + using False by (intro Nat.gr0I) auto + have "a > 0" + using \D > 0\ by (intro Nat.gr0I) (auto simp: a_def) + moreover have "\is_square D'" + unfolding D'_def + by (metis False is_nth_power_mult is_nth_power_nth_power squarefree_decompose) + ultimately interpret lift: pell_lift D' a D + using False \D > 0\ + by unfold_locales (auto simp: D'_def a_def squarefree_decompose [symmetric]) + + define i where "i = (LEAST i. case lift.nth_solution i of (_, y) \ y > 0 \ a dvd y)" + have ex: "\i. case lift.nth_solution i of (_, y) \ y > 0 \ a dvd y" + proof - + define sol where "sol = lift.lift.fund_sol" + have is_sol: "lift.solution (lift.lower_solution sol)" + unfolding sol_def using lift.lift.fund_sol_is_nontriv_solution lift.lift_solution_iff by blast + then obtain j where j: "lift.lower_solution sol = lift.nth_solution j" + using lift.solution_iff_nth_solution by blast + have "snd (lift.lower_solution sol) > 0" + proof (rule Nat.gr0I) + assume *: "snd (lift.lower_solution sol) = 0" + have "lift.solution (fst (lift.lower_solution sol), snd (lift.lower_solution sol))" + using is_sol by simp + hence "fst (lift.lower_solution sol) = 1" + by (subst (asm) *) simp + with * have "lift.lower_solution sol = (1, 0)" + by (cases "lift.lower_solution sol") auto + hence "fst sol = 1" + unfolding lift.lower_solution_def by (auto simp: lift.lower_solution_def case_prod_unfold) + thus False + unfolding sol_def + using lift.lift.fund_sol_is_nontriv_solution \D > 0\ + by (auto simp: lift.lift.nontriv_solution_def) + qed + moreover have "a dvd snd (lift.lower_solution sol)" + by (auto simp: lift.lower_solution_def case_prod_unfold) + ultimately show ?thesis + using j by (auto simp: case_prod_unfold) + qed + + define sol where "sol = lift.nth_solution i" + have sol: "snd sol > 0" "a dvd snd sol" + using LeastI_ex[OF ex] by (simp_all add: sol_def i_def case_prod_unfold) + have "i > 0" + using sol by (intro Nat.gr0I) (auto simp: sol_def lift.nth_solution_def) + + have "find_fund_sol_fast D = map_prod id (\y. y div a) + (shd (sdrop_while (\(_, y). y = 0 \ \a dvd y) (pell_solutions D')))" + unfolding D'_def a_def find_fund_sol_fast_def using False squarefree_part_1_imp_square[of D] + by (auto simp: square_squarefree_part_nat_def) + also have "sdrop_while (\(_, y). y = 0 \ \a dvd y) (pell_solutions D') = + sdrop_while (Not \ (\(_, y). y > 0 \ a dvd y)) (pell_solutions D')" + by (simp add: o_def case_prod_unfold) + also have "\ = sdrop i (pell_solutions D')" + using ex by (subst sdrop_while_sdrop_LEAST) (simp_all add: lift.snth_pell_solutions i_def) + also have "shd \ = sol" + by (simp add: lift.snth_pell_solutions sol_def) + finally have eq: "find_fund_sol_fast D = map_prod id (\y. y div a) sol" . + + have "lift.lift.fund_sol = lift.lift_solution sol" + unfolding sol_def + proof (rule lift.lift_fund_sol) + show "i > 0" by fact + show "lift.liftable_solution (lift.nth_solution i)" + using sol by (simp add: sol_def lift.liftable_solution_def case_prod_unfold) + next + fix j :: nat assume j: "j > 0" "j < i" + show "\lift.liftable_solution (lift.nth_solution j)" + proof + assume liftable: "lift.liftable_solution (lift.nth_solution j)" + have "snd (lift.nth_solution j) > 0" + using \j > 0\ by (metis gr0I lift.nontriv_solution_altdef lift.nth_solution_sound' + lift.solution_0_snd_nat_iff prod.collapse) + hence "case lift.nth_solution j of (_, y) \ y > 0 \ a dvd y" + using \j > 0\ liftable by (auto simp: lift.liftable_solution_def) + hence "i \ j" + unfolding i_def by (rule Least_le) + thus False using \j < i\ by simp + qed + qed + also have "\ = find_fund_sol_fast D" + by (simp add: eq lift.lift_solution_def case_prod_unfold map_prod_def) + finally show ?thesis + using \D > 0\ False by (simp add: find_fund_sol_correct) +qed + +end \ No newline at end of file diff --git a/thys/Continued_Fractions/Quadratic_Irrationals.thy b/thys/Continued_Fractions/Quadratic_Irrationals.thy new file mode 100644 --- /dev/null +++ b/thys/Continued_Fractions/Quadratic_Irrationals.thy @@ -0,0 +1,1018 @@ +(* + File: Continued_Fractions/Quadratic Irrationals.thy + Author: Manuel Eberl, University of Innsbruck +*) +section \Quadratic Irrationals\ +theory Quadratic_Irrationals +imports + Continued_Fractions + "HOL-Computational_Algebra.Computational_Algebra" + "HOL-Library.Discrete" + "Coinductive.Coinductive_Stream" +begin + +lemma snth_cycle: + assumes "xs \ []" + shows "snth (cycle xs) n = xs ! (n mod length xs)" +proof (induction n rule: less_induct) + case (less n) + have "snth (shift xs (cycle xs)) n = xs ! (n mod length xs)" + proof (cases "n < length xs") + case True + thus ?thesis + by (subst shift_snth_less) auto + next + case False + have "0 < length xs" + using assms by simp + also have "\ \ n" + using False by simp + finally have "n > 0" . + + from False have "snth (shift xs (cycle xs)) n = snth (cycle xs) (n - length xs)" + by (subst shift_snth_ge) auto + also have "\ = xs ! ((n - length xs) mod length xs)" + using assms \n > 0\ by (intro less) auto + also have "(n - length xs) mod length xs = n mod length xs" + using False by (simp add: mod_if) + finally show ?thesis . + qed + also have "shift xs (cycle xs) = cycle xs" + by (rule cycle_decomp [symmetric]) fact + finally show ?case . +qed + +subsection \Basic results on rationality of square roots\ + +lemma inverse_in_Rats_iff [simp]: "inverse (x :: real) \ \ \ x \ \" + by (auto simp: inverse_eq_divide divide_in_Rats_iff1) + +lemma nonneg_sqrt_nat_or_irrat: + assumes "x ^ 2 = real a" and "x \ 0" + shows "x \ \ \ x \ \" +proof safe + assume "x \ \" and "x \ \" + from Rats_abs_nat_div_natE[OF this(2)] + obtain p q :: nat where q_nz [simp]: "q \ 0" and "abs x = p / q" and coprime: "coprime p q" . + with \x \ 0\ have x: "x = p / q" + by simp + with assms have "real (q ^ 2) * real a = real (p ^ 2)" + by (simp add: field_simps) + also have "real (q ^ 2) * real a = real (q ^ 2 * a)" + by simp + finally have "p ^ 2 = q ^ 2 * a" + by (subst (asm) of_nat_eq_iff) auto + hence "q ^ 2 dvd p ^ 2" + by simp + hence "q dvd p" + by simp + with coprime have "q = 1" + by auto + with x and \x \ \\ show False + by simp +qed + +text \ + A square root of a natural number is either an integer or irrational. +\ +corollary sqrt_nat_or_irrat: + assumes "x ^ 2 = real a" + shows "x \ \ \ x \ \" +proof (cases "x \ 0") + case True + with nonneg_sqrt_nat_or_irrat[OF assms this] + show ?thesis by (auto simp: Nats_altdef2) +next + case False + from assms have "(-x) ^ 2 = real a" + by simp + moreover from False have "-x \ 0" + by simp + ultimately have "-x \ \ \ -x \ \" + by (rule nonneg_sqrt_nat_or_irrat) + thus ?thesis + by (auto simp: Nats_altdef2 minus_in_Ints_iff) +qed + +corollary sqrt_nat_or_irrat': + "sqrt (real a) \ \ \ sqrt (real a) \ \" + using nonneg_sqrt_nat_or_irrat[of "sqrt a" a] by auto + +text \ + The square root of a natural number \n\ is again a natural number iff \n is a perfect square.\ +\ +corollary sqrt_nat_iff_is_square: + "sqrt (real n) \ \ \ is_square n" +proof + assume "sqrt (real n) \ \" + then obtain k where "sqrt (real n) = real k" by (auto elim!: Nats_cases) + hence "sqrt (real n) ^ 2 = real (k ^ 2)" by (simp only: of_nat_power) + also have "sqrt (real n) ^ 2 = real n" by simp + finally have "n = k ^ 2" by (simp only: of_nat_eq_iff) + thus "is_square n" by blast +qed (auto elim!: is_nth_powerE) + +corollary irrat_sqrt_nonsquare: "\is_square n \ sqrt (real n) \ \" + using sqrt_nat_or_irrat'[of n] by (auto simp: sqrt_nat_iff_is_square) + +lemma sqrt_of_nat_in_Rats_iff: "sqrt (real n) \ \ \ is_square n" + using irrat_sqrt_nonsquare[of n] sqrt_nat_iff_is_square[of n] Nats_subset_Rats by blast + +lemma Discrete_sqrt_altdef: "Discrete.sqrt n = nat \sqrt n\" +proof - + have "real (Discrete.sqrt n ^ 2) \ sqrt n ^ 2" + by simp + hence "Discrete.sqrt n \ sqrt n" + unfolding of_nat_power by (rule power2_le_imp_le) auto + moreover have "real (Suc (Discrete.sqrt n) ^ 2) > real n" + unfolding of_nat_less_iff by (rule Suc_sqrt_power2_gt) + hence "real (Discrete.sqrt n + 1) ^ 2 > sqrt n ^ 2" + unfolding of_nat_power by simp + hence "real (Discrete.sqrt n + 1) > sqrt n" + by (rule power2_less_imp_less) auto + hence "Discrete.sqrt n + 1 > sqrt n" by simp + ultimately show ?thesis by linarith +qed + + +subsection \Definition of quadratic irrationals\ + +text \ + Irrational real numbers $x$ that satisfy a quadratic equation $a x^2 + b x + c = 0$ + with \a\, \b\, \c\ not all equal to 0 are called \<^emph>\quadratic irrationals\. These are of the form + $p + q \sqrt{d}$ for rational numbers \p\, \q\ and a positive integer \d\. +\ +inductive quadratic_irrational :: "real \ bool" where + "x \ \ \ real_of_int a * x ^ 2 + real_of_int b * x + real_of_int c = 0 \ + a \ 0 \ b \ 0 \ c \ 0 \ quadratic_irrational x" + +lemma quadratic_irrational_sqrt [intro]: + assumes "\is_square n" + shows "quadratic_irrational (sqrt (real n))" + using irrat_sqrt_nonsquare[OF assms] + by (intro quadratic_irrational.intros[of "sqrt n" 1 0 "-int n"]) auto + +lemma quadratic_irrational_uminus [intro]: + assumes "quadratic_irrational x" + shows "quadratic_irrational (-x)" + using assms +proof induction + case (1 x a b c) + thus ?case by (intro quadratic_irrational.intros[of "-x" a "-b" c]) auto +qed + +lemma quadratic_irrational_uminus_iff [simp]: + "quadratic_irrational (-x) \ quadratic_irrational x" + using quadratic_irrational_uminus[of x] quadratic_irrational_uminus[of "-x"] by auto + +lemma quadratic_irrational_plus_int [intro]: + assumes "quadratic_irrational x" + shows "quadratic_irrational (x + of_int n)" + using assms +proof induction + case (1 x a b c) + define x' where "x' = x + of_int n" + define a' b' c' where + "a' = a" and "b' = b - 2 * of_int n * a" and + "c' = a * of_int n ^ 2 - b * of_int n + c" + from 1 have "0 = a * (x' - of_int n) ^ 2 + b * (x' - of_int n) + c" + by (simp add: x'_def) + also have "\ = a' * x' ^ 2 + b' * x' + c'" + by (simp add: algebra_simps a'_def b'_def c'_def power2_eq_square) + finally have "\ = 0" .. + moreover have "x' \ \" + using 1 by (auto simp: x'_def add_in_Rats_iff2) + moreover have "a' \ 0 \ b' \ 0 \ c' \ 0" + using 1 by (auto simp: a'_def b'_def c'_def) + ultimately show ?case + by (intro quadratic_irrational.intros[of "x + of_int n" a' b' c']) (auto simp: x'_def) +qed + +lemma quadratic_irrational_plus_int_iff [simp]: + "quadratic_irrational (x + of_int n) \ quadratic_irrational x" + using quadratic_irrational_plus_int[of x n] + quadratic_irrational_plus_int[of "x + of_int n" "-n"] by auto + +lemma quadratic_irrational_minus_int_iff [simp]: + "quadratic_irrational (x - of_int n) \ quadratic_irrational x" + using quadratic_irrational_plus_int_iff[of x "-n"] + by (simp del: quadratic_irrational_plus_int_iff) + +lemma quadratic_irrational_plus_nat_iff [simp]: + "quadratic_irrational (x + of_nat n) \ quadratic_irrational x" + using quadratic_irrational_plus_int_iff[of x "int n"] + by (simp del: quadratic_irrational_plus_int_iff) + +lemma quadratic_irrational_minus_nat_iff [simp]: + "quadratic_irrational (x - of_nat n) \ quadratic_irrational x" + using quadratic_irrational_plus_int_iff[of x "-int n"] + by (simp del: quadratic_irrational_plus_int_iff) + +lemma quadratic_irrational_plus_1_iff [simp]: + "quadratic_irrational (x + 1) \ quadratic_irrational x" + using quadratic_irrational_plus_int_iff[of x 1] + by (simp del: quadratic_irrational_plus_int_iff) + +lemma quadratic_irrational_minus_1_iff [simp]: + "quadratic_irrational (x - 1) \ quadratic_irrational x" + using quadratic_irrational_plus_int_iff[of x "-1"] + by (simp del: quadratic_irrational_plus_int_iff) + +lemma quadratic_irrational_plus_numeral_iff [simp]: + "quadratic_irrational (x + numeral n) \ quadratic_irrational x" + using quadratic_irrational_plus_int_iff[of x "numeral n"] + by (simp del: quadratic_irrational_plus_int_iff) + +lemma quadratic_irrational_minus_numeral_iff [simp]: + "quadratic_irrational (x - numeral n) \ quadratic_irrational x" + using quadratic_irrational_plus_int_iff[of x "-numeral n"] + by (simp del: quadratic_irrational_plus_int_iff) + +lemma quadratic_irrational_inverse: + assumes "quadratic_irrational x" + shows "quadratic_irrational (inverse x)" + using assms +proof induction + case (1 x a b c) + from 1 have "x \ 0" by auto + have "0 = (real_of_int a * x\<^sup>2 + real_of_int b * x + real_of_int c) / x ^ 2" + by (subst 1) simp + also have "\ = real_of_int c * (inverse x) ^ 2 + real_of_int b * inverse x + real_of_int a" + using \x \ 0\ by (simp add: field_simps power2_eq_square) + finally have "\ = 0" .. + thus ?case using 1 + by (intro quadratic_irrational.intros[of "inverse x" c b a]) auto +qed + +lemma quadratic_irrational_inverse_iff [simp]: + "quadratic_irrational (inverse x) \ quadratic_irrational x" + using quadratic_irrational_inverse[of x] quadratic_irrational_inverse[of "inverse x"] + by (cases "x = 0") auto + +lemma quadratic_irrational_cfrac_remainder_iff: + "quadratic_irrational (cfrac_remainder c n) \ quadratic_irrational (cfrac_lim c)" +proof (cases "cfrac_length c = \") + case False + thus ?thesis + by (auto simp: quadratic_irrational.simps) +next + case [simp]: True + show ?thesis + proof (induction n) + case (Suc n) + from Suc.prems have "cfrac_remainder c (Suc n) = + inverse (cfrac_remainder c n - of_int (cfrac_nth c n))" + by (subst cfrac_remainder_Suc) (auto simp: field_simps) + also have "quadratic_irrational \ \ quadratic_irrational (cfrac_remainder c n)" + by simp + also have "\ \ quadratic_irrational (cfrac_lim c)" + by (rule Suc.IH) + finally show ?case . + qed auto +qed + +subsection \Real solutions of quadratic equations\ + +text \ + For the next result, we need some basic properties of real solutions to quadratic equations. +\ +lemma quadratic_equation_reals: + fixes a b c :: real + defines "f \ (\x. a * x ^ 2 + b * x + c)" + defines "discr \ (b^2 - 4 * a * c)" + shows "{x. f x = 0} = + (if a = 0 then + (if b = 0 then if c = 0 then UNIV else {} else {-c/b}) + else if discr \ 0 then {(-b + sqrt discr) / (2 * a), (-b - sqrt discr) / (2 * a)} + else {})" (is ?th1) +proof (cases "a = 0") + case [simp]: True + show ?th1 + proof (cases "b = 0") + case [simp]: True + hence "{x. f x = 0} = (if c = 0 then UNIV else {})" + by (auto simp: f_def) + thus ?th1 by simp + next + case False + hence "{x. f x = 0} = {-c / b}" by (auto simp: f_def field_simps) + thus ?th1 using False by simp + qed +next + case [simp]: False + show ?th1 + proof (cases "discr \ 0") + case True + { + fix x :: real + have "f x = a * (x - (-b + sqrt discr) / (2 * a)) * (x - (-b - sqrt discr) / (2 * a))" + using True by (simp add: f_def field_simps discr_def power2_eq_square) + also have "\ = 0 \ x \ {(-b + sqrt discr) / (2 * a), (-b - sqrt discr) / (2 * a)}" + by simp + finally have "f x = 0 \ \" . + } + hence "{x. f x = 0} = {(-b + sqrt discr) / (2 * a), (-b - sqrt discr) / (2 * a)}" + by blast + thus ?th1 using True by simp + next + case False + { + fix x :: real + assume x: "f x = 0" + have "0 \ (x + b / (2 * a)) ^ 2" by simp + also have "f x = a * ((x + b / (2 * a)) ^ 2 - b ^ 2 / (4 * a ^ 2) + c / a)" + by (simp add: field_simps power2_eq_square f_def) + with x have "(x + b / (2 * a)) ^ 2 - b ^ 2 / (4 * a ^ 2) + c / a = 0" + by simp + hence "(x + b / (2 * a)) ^ 2 = b ^ 2 / (4 * a ^ 2) - c / a" + by (simp add: algebra_simps) + finally have "0 \ (b\<^sup>2 / (4 * a\<^sup>2) - c / a) * (4 * a\<^sup>2)" + by (intro mult_nonneg_nonneg) auto + also have "\ = b\<^sup>2 - 4 * a * c" by (simp add: field_simps power2_eq_square) + also have "\ < 0" using False by (simp add: discr_def) + finally have False by simp + } + hence "{x. f x = 0} = {}" by auto + thus ?th1 using False by simp + qed +qed + +lemma finite_quadratic_equation_solutions_reals: + fixes a b c :: real + defines "discr \ (b^2 - 4 * a * c)" + shows "finite {x. a * x ^ 2 + b * x + c = 0} \ a \ 0 \ b \ 0 \ c \ 0" + by (subst quadratic_equation_reals) + (auto simp: discr_def card_eq_0_iff infinite_UNIV_char_0 split: if_split) + +lemma card_quadratic_equation_solutions_reals: + fixes a b c :: real + defines "discr \ (b^2 - 4 * a * c)" + shows "card {x. a * x ^ 2 + b * x + c = 0} = + (if a = 0 then + (if b = 0 then 0 else 1) + else if discr \ 0 then if discr = 0 then 1 else 2 else 0)" (is ?th1) + by (subst quadratic_equation_reals) + (auto simp: discr_def card_eq_0_iff infinite_UNIV_char_0 split: if_split) + +lemma card_quadratic_equation_solutions_reals_le_2: + "card {x :: real. a * x ^ 2 + b * x + c = 0} \ 2" + by (subst card_quadratic_equation_solutions_reals) auto + +lemma quadratic_equation_solution_rat_iff: + fixes a b c :: int and x y :: real + defines "f \ (\x::real. a * x ^ 2 + b * x + c)" + defines "discr \ nat (b ^ 2 - 4 * a * c)" + assumes "a \ 0" "f x = 0" + shows "x \ \ \ is_square discr" +proof - + define discr' where "discr' \ real_of_int (b ^ 2 - 4 * a * c)" + from assms have "x \ {x. f x = 0}" by simp + with \a \ 0\ have "discr' \ 0" unfolding discr'_def f_def of_nat_diff + by (subst (asm) quadratic_equation_reals) (auto simp: discr_def split: if_splits) + hence *: "sqrt (discr') = sqrt (real discr)" unfolding of_int_0_le_iff discr_def discr'_def + by (simp add: algebra_simps nat_diff_distrib) + from \x \ {x. f x = 0}\ have "x = (-b + sqrt discr) / (2 * a) \ x = (-b - sqrt discr) / (2 * a)" + using \a \ 0\ * unfolding discr'_def f_def + by (subst (asm) quadratic_equation_reals) (auto split: if_splits) + thus ?thesis using \a \ 0\ + by (auto simp: sqrt_of_nat_in_Rats_iff divide_in_Rats_iff2 diff_in_Rats_iff2 diff_in_Rats_iff1) +qed + + +subsection \Periodic continued fractions and quadratic irrationals\ + +text \ + We now show the main result: A positive irrational number has a periodic continued + fraction expansion iff it is a quadratic irrational. + + In principle, this statement naturally also holds for negative numbers, but the current + formalisation of continued fractions only supports non-negative numbers. It also holds for + rational numbers in some sense, since their continued fraction expansion is finite to begin with. +\ +theorem periodic_cfrac_imp_quadratic_irrational: + assumes [simp]: "cfrac_length c = \" + and period: "l > 0" "\k. k \ N \ cfrac_nth c (k + l) = cfrac_nth c k" + shows "quadratic_irrational (cfrac_lim c)" +proof - + define h' and k' where "h' = conv_num_int (cfrac_drop N c)" + and "k' = conv_denom_int (cfrac_drop N c)" + define x' where "x' = cfrac_remainder c N" + + have c_pos: "cfrac_nth c n > 0" if "n \ N" for n + proof - + from assms(1,2) have "cfrac_nth c (n + l) > 0" by auto + with assms(3)[OF that] show ?thesis by simp + qed + have k'_pos: "k' n > 0" if "n \ -1" "n \ -2" for n + using that by (auto simp: k'_def conv_denom_int_def intro!: conv_denom_pos) + have k'_nonneg: "k' n \ 0" if "n \ -2" for n + using that by (auto simp: k'_def conv_denom_int_def intro!: conv_denom_pos) + have "cfrac_nth c (n + (N + l)) = cfrac_nth c (n + N)" for n + using period(2)[of "n + N"] by (simp add: add_ac) + have "cfrac_drop (N + l) c = cfrac_drop N c" + by (rule cfrac_eqI) (use period(2)[of "n + N" for n] in \auto simp: algebra_simps\) + hence x'_altdef: "x' = cfrac_remainder c (N + l)" + by (simp add: x'_def cfrac_remainder_def) + have x'_pos: "x' > 0" unfolding x'_def + using c_pos by (intro cfrac_remainder_pos) auto + + define A where "A = (k' (int l - 1))" + define B where "B = k' (int l - 2) - h' (int l - 1)" + define C where "C = -(h' (int l - 2))" + + have pos: "(k' (int l - 1) * x' + k' (int l - 2)) > 0" + using x'_pos \l > 0\ + by (intro add_pos_nonneg mult_pos_pos) (auto intro!: k'_pos k'_nonneg) + have "cfrac_remainder c N = conv' (cfrac_drop N c) l (cfrac_remainder c (l + N))" + unfolding cfrac_remainder_def cfrac_drop_add + by (subst (2) cfrac_remainder_def [symmetric]) (auto simp: conv'_cfrac_remainder) + hence "x' = conv' (cfrac_drop N c) l x'" + by (subst (asm) add.commute) (simp only: x'_def [symmetric] x'_altdef [symmetric]) + also have "\ = (h' (int l - 1) * x' + h' (int l - 2)) / (k' (int l - 1) * x' + k' (int l - 2))" + using conv'_num_denom_int[OF x'_pos, of _ l] unfolding h'_def k'_def + by (simp add: mult_ac) + finally have "x' * (k' (int l - 1) * x' + k' (int l - 2)) = (h' (int l - 1) * x' + h' (int l - 2))" + using pos by (simp add: divide_simps) + hence quadratic: "A * x' ^ 2 + B * x' + C = 0" + by (simp add: algebra_simps power2_eq_square A_def B_def C_def) + moreover have "x' \ \" unfolding x'_def + by auto + moreover have "A > 0" using \l > 0\ by (auto simp: A_def intro!: k'_pos) + ultimately have "quadratic_irrational x'" using \x' \ \\ + by (intro quadratic_irrational.intros[of x' A B C]) simp_all + thus ?thesis + using assms by (simp add: x'_def quadratic_irrational_cfrac_remainder_iff) +qed + + +lift_definition pperiodic_cfrac :: "nat list \ cfrac" is + "\xs. if xs = [] then (0, LNil) else + (int (hd xs), llist_of_stream (cycle (map (\n. n- 1) (tl xs @ [hd xs]))))" . + +definition periodic_cfrac :: "int list \ int list \ cfrac" where + "periodic_cfrac xs ys = cfrac_of_stream (Stream.shift xs (Stream.cycle ys))" + +lemma periodic_cfrac_Nil [simp]: "pperiodic_cfrac [] = 0" + unfolding zero_cfrac_def by transfer auto + +lemma cfrac_length_pperiodic_cfrac [simp]: + "xs \ [] \ cfrac_length (pperiodic_cfrac xs) = \" + by transfer auto + +lemma cfrac_nth_pperiodic_cfrac: + assumes "xs \ []" and "0 \ set xs" + shows "cfrac_nth (pperiodic_cfrac xs) n = xs ! (n mod length xs)" + using assms +proof (transfer, goal_cases) + case (1 xs n) + show ?case + proof (cases n) + case (Suc n') + have "int (cycle (tl (map (\n. n - 1) xs) @ [hd (map (\n. n - 1) xs)]) !! n') + 1 = + int (stl (cycle (map (\n. n - 1) xs)) !! n') + 1" + by (subst cycle.sel(2) [symmetric]) (rule refl) + also have "\ = int (cycle (map (\n. n - 1) xs) !! n) + 1" + by (simp add: Suc del: cycle.sel) + also have "\ = int (xs ! (n mod length xs) - 1) + 1" + by (simp add: snth_cycle \xs \ []\) + also have "xs ! (n mod length xs) \ set xs" + using \xs \ []\ by (auto simp: set_conv_nth) + with 1 have "xs ! (n mod length xs) > 0" + by (intro Nat.gr0I) auto + hence "int (xs ! (n mod length xs) - 1) + 1 = int (xs ! (n mod length xs))" + by simp + finally show ?thesis + using Suc 1 by (simp add: hd_conv_nth map_tl) + qed (use 1 in \auto simp: hd_conv_nth\) +qed + +definition pperiodic_cfrac_info :: "nat list \ int \ int \ int"where + "pperiodic_cfrac_info xs = + (let l = length xs; + h = conv_num_fun (\n. xs ! n); + k = conv_denom_fun (\n. xs ! n); + A = k (l - 1); + B = h (l - 1) - (if l = 1 then 0 else k (l - 2)); + C = (if l = 1 then -1 else -h (l - 2)) + in (B^2-4*A*C, B, 2 * A))" + +lemma conv_gen_cong: + assumes "\k\{n..N}. f k = f' k" + shows "conv_gen f (a,b,n) N = conv_gen f' (a,b,n) N" + using assms +proof (induction "N - n" arbitrary: a b n N) + case (Suc d n N a b) + have "conv_gen f (b, b * f n + a, Suc n) N = conv_gen f' (b, b * f n + a, Suc n) N" + using Suc(2,3) by (intro Suc) auto + moreover have "f n = f' n" + using bspec[OF Suc.prems, of n] Suc(2) by auto + ultimately show ?case + by (subst (1 2) conv_gen.simps) auto +qed (auto simp: conv_gen.simps) + +lemma + assumes "\k\n. c k = cfrac_nth c' k" + shows conv_num_fun_eq': "conv_num_fun c n = conv_num c' n" + and conv_denom_fun_eq': "conv_denom_fun c n = conv_denom c' n" +proof - + have "conv_num c' n = conv_gen (cfrac_nth c') (0, 1, 0) n" + unfolding conv_num_code .. + also have "\ = conv_gen c (0, 1, 0) n" + unfolding conv_num_fun_def using assms by (intro conv_gen_cong) auto + finally show "conv_num_fun c n = conv_num c' n" + by (simp add: conv_num_fun_def) +next + have "conv_denom c' n = conv_gen (cfrac_nth c') (1, 0, 0) n" + unfolding conv_denom_code .. + also have "\ = conv_gen c (1, 0, 0) n" + unfolding conv_denom_fun_def using assms by (intro conv_gen_cong) auto + finally show "conv_denom_fun c n = conv_denom c' n" + by (simp add: conv_denom_fun_def) +qed + +lemma gcd_minus_commute_left: "gcd (a - b :: 'a :: ring_gcd) c = gcd (b - a) c" + by (metis gcd.commute gcd_neg2 minus_diff_eq) + +lemma gcd_minus_commute_right: "gcd c (a - b :: 'a :: ring_gcd) = gcd c (b - a)" + by (metis gcd_neg2 minus_diff_eq) + +lemma periodic_cfrac_info_aux: + fixes D E F :: int + assumes "pperiodic_cfrac_info xs = (D, E, F)" + assumes "xs \ []" "0 \ set xs" + shows "cfrac_lim (pperiodic_cfrac xs) = (sqrt D + E) / F" + and "D > 0" and "F > 0" +proof - + define c where "c = pperiodic_cfrac xs" + have [simp]: "cfrac_length c = \" + using assms by (simp add: c_def) + define h and k where "h = conv_num_int c" and "k = conv_denom_int c" + define x where "x = cfrac_lim c" + define l where "l = length xs" + + define A where "A = (k (int l - 1))" + define B where "B = k (int l - 2) - h (int l - 1)" + define C where "C = -(h (int l - 2))" + define discr where "discr = B ^ 2 - 4 * A * C" + + have "l > 0" + using assms by (simp add: l_def) + have c_pos: "cfrac_nth c n > 0" for n + using assms by (auto simp: c_def cfrac_nth_pperiodic_cfrac set_conv_nth) + have x_pos: "x > 0" + unfolding x_def by (intro cfrac_lim_pos c_pos) + have h_pos: "h n > 0" if "n > -2" for n + using that unfolding h_def by (auto simp: conv_num_int_def intro: conv_num_pos' c_pos) + have k_pos: "k n > 0" if "n > -1" for n + using that unfolding k_def by (auto simp: conv_denom_int_def) + have k_nonneg: "k n \ 0" for n + unfolding k_def by (auto simp: conv_denom_int_def) + + have pos: "(k (int l - 1) * x + k (int l - 2)) > 0" + using x_pos \l > 0\ + by (intro add_pos_nonneg mult_pos_pos) (auto intro!: k_pos k_nonneg) + have "cfrac_drop l c = c" + using assms by (intro cfrac_eqI) (auto simp: c_def cfrac_nth_pperiodic_cfrac l_def) + + have "x = conv' c l (cfrac_remainder c l)" + unfolding x_def by (rule conv'_cfrac_remainder[symmetric]) auto + also have "\ = conv' c l x" + unfolding cfrac_remainder_def \cfrac_drop l c = c\ x_def .. + finally have "x = conv' c l x" . + also have "\ = (h (int l - 1) * x + h (int l - 2)) / (k (int l - 1) * x + k (int l - 2))" + using conv'_num_denom_int[OF x_pos, of _ l] unfolding h_def k_def + by (simp add: mult_ac) + finally have "x * (k (int l - 1) * x + k (int l - 2)) = (h (int l - 1) * x + h (int l - 2))" + using pos by (simp add: divide_simps) + hence quadratic: "A * x ^ 2 + B * x + C = 0" + by (simp add: algebra_simps power2_eq_square A_def B_def C_def) + + have "A > 0" using \l > 0\ by (auto simp: A_def intro!: k_pos) + have discr_altdef: "discr = (k (int l-2) - h (int l-1)) ^ 2 + 4 * k (int l-1) * h (int l-2)" + by (simp add: discr_def A_def B_def C_def) + + have "0 < 0 + 4 * A * 1" + using \A > 0\ by simp + also have "0 + 4 * A * 1 \ discr" + unfolding discr_altdef A_def using h_pos[of "int l - 2"] \l > 0\ + by (intro add_mono mult_mono order.refl k_nonneg mult_nonneg_nonneg) auto + finally have "discr > 0" . + + have "x \ {x. A * x ^ 2 + B * x + C = 0}" + using quadratic by simp + hence x_cases: "x = (-B - sqrt discr) / (2 * A) \ x = (-B + sqrt discr) / (2 * A)" + unfolding quadratic_equation_reals of_int_diff using \A > 0\ + by (auto split: if_splits simp: discr_def) + + have "B ^ 2 < discr" + unfolding discr_def by (auto intro!: mult_pos_pos k_pos h_pos \l > 0\ simp: A_def C_def) + hence "\B\ < sqrt discr" + using \discr > 0\ by (simp add: real_less_rsqrt) + + have "x = (if x \ 0 then (sqrt discr - B) / (2 * A) else -(sqrt discr + B) / (2 * A))" + using x_cases + proof + assume x: "x = (-B - sqrt discr) / (2 * A)" + have "(-B - sqrt discr) / (2 * A) < 0" + using \\B\ < sqrt discr\ \A > 0\ by (intro divide_neg_pos) auto + also note x[symmetric] + finally show ?thesis using x by simp + next + assume x: "x = (-B + sqrt discr) / (2 * A)" + have "(-B + sqrt discr) / (2 * A) > 0" + using \\B\ < sqrt discr\ \A > 0\ by (intro divide_pos_pos) auto + also note x[symmetric] + finally show ?thesis using x by simp + qed + also have "x \ 0 \ floor x \ 0" + by auto + also have "floor x = floor (cfrac_lim c)" + by (simp add: x_def) + also have "\ = cfrac_nth c 0" + by (subst cfrac_nth_0_conv_floor) auto + also have "\ = int (hd xs)" + using assms unfolding c_def by (subst cfrac_nth_pperiodic_cfrac) (auto simp: hd_conv_nth) + finally have x_eq: "x = (sqrt discr - B) / (2 * A)" + by simp + + + define h' where "h' = conv_num_fun (\n. int (xs ! n))" + define k' where "k' = conv_denom_fun (\n. int (xs ! n))" + have num_eq: "h' i = h i" + if "i < l" for i using that assms unfolding h'_def h_def + by (subst conv_num_fun_eq'[where c' = c]) (auto simp: c_def l_def cfrac_nth_pperiodic_cfrac) + have denom_eq: "k' i = k i" + if "i < l" for i using that assms unfolding k'_def k_def + by (subst conv_denom_fun_eq'[where c' = c]) (auto simp: c_def l_def cfrac_nth_pperiodic_cfrac) + + have 1: "h (int l - 1) = h' (l - 1)" + by (subst num_eq) (use \l > 0\ in \auto simp: of_nat_diff\) + have 2: "k (int l - 1) = k' (l - 1)" + by (subst denom_eq) (use \l > 0\ in \auto simp: of_nat_diff\) + have 3: "h (int l - 2) = (if l = 1 then 1 else h' (l - 2))" + using \l > 0\ num_eq[of "l - 2"] by (auto simp: h_def nat_diff_distrib) + have 4: "k (int l - 2) = (if l = 1 then 0 else k' (l - 2))" + using \l > 0\ denom_eq[of "l - 2"] by (auto simp: k_def nat_diff_distrib) + + have "pperiodic_cfrac_info xs = + (let A = k (int l - 1); + B = h (int l - 1) - (if l = 1 then 0 else k (int l - 2)); + C = (if l = 1 then -1 else - h (int l - 2)) + in (B\<^sup>2 - 4 * A * C, B, 2 * A))" + unfolding pperiodic_cfrac_info_def Let_def using 1 2 3 4 \l > 0\ + by (auto simp: num_eq denom_eq h'_def k'_def l_def of_nat_diff) + also have "\ = (B\<^sup>2 - 4 * A * C, -B, 2 * A)" + by (simp add: Let_def A_def B_def C_def h_def k_def algebra_simps power2_commute) + finally have per_eq: "pperiodic_cfrac_info xs = (discr, -B, 2 * A)" + by (simp add: discr_def) + + show "x = (sqrt (real_of_int D) + real_of_int E) / real_of_int F" + using per_eq assms by (simp add: x_eq) + show "D > 0" "F > 0" + using assms per_eq \discr > 0\ \A > 0\ by auto +qed + +text \ + We can now compute surd representations for (purely) periodic continued fractions, e.g. + $[1, 1, 1, \ldots] = \frac{\sqrt{5} + 1}{2}$: +\ +value "pperiodic_cfrac_info [1]" + +text \ + We can now compute surd representations for periodic continued fractions, e.g. + $[\overline{1,1,1,1,6}] = \frac{\sqrt{13} + 3}{4}$: +\ +value "pperiodic_cfrac_info [1,1,1,1,6]" + + + +text \ + With a little bit of work, one could also easily derive from this a version for + non-purely periodic continued fraction. +\ + + +text \ + Next, we show that any quadratic irrational has a periodic continued fraction expansion. +\ +theorem quadratic_irrational_imp_periodic_cfrac: + assumes "quadratic_irrational (cfrac_lim e)" + obtains N l where "l > 0" and "\n m. n \ N \ cfrac_nth e (n + m * l) = cfrac_nth e n" + and "cfrac_remainder e (N + l) = cfrac_remainder e N" + and "cfrac_length e = \" +proof - + have [simp]: "cfrac_length e = \" + using assms by (auto simp: quadratic_irrational.simps) + note [intro] = assms(1) + define x where "x = cfrac_lim e" + from assms obtain a b c :: int where + nontrivial: "a \ 0 \ b \ 0 \ c \ 0" and + root: "a * x^2 + b * x + c = 0" (is "?f x = 0") + by (auto simp: quadratic_irrational.simps x_def) + + define f where "f = ?f" + define h and k where "h = conv_num e" and "k = conv_denom e" + define X where "X = cfrac_remainder e" + have [simp]: "k i > 0" "k i \ 0" for i + using conv_denom_pos[of e i] by (auto simp: k_def) + have k_leI: "k i \ k j" if "i \ j" for i j + by (auto simp: k_def intro!: conv_denom_leI that) + have k_nonneg: "k n \ 0" for n + by (auto simp: k_def) + have k_ge_1: "k n \ 1" for n + using k_leI[of 0 n] by (simp add: k_def) + + define R where "R = conv e" + define A where "A = (\n. a * h (n - 1) ^ 2 + b * h (n - 1) * k (n - 1) + c * k (n - 1) ^ 2)" + define B where "B = (\n. 2 * a * h (n - 1) * h (n - 2) + b * (h (n - 1) * k (n - 2) + h (n - 2) * k (n - 1)) + 2 * c * k (n - 1) * k (n - 2))" + define C where "C = (\n. a * h (n - 2) ^ 2 + b * h (n - 2) * k (n - 2) + c * k (n - 2) ^ 2)" + + define A' where "A' = nat \2 * \a\ * \x\ + \a\ + \b\\" + define B' where "B' = nat \(3 / 2) * (2 * \a\ * \x\ + \b\) + 9 / 4 * \a\\" + + have [simp]: "X n \ \" for n unfolding X_def + by simp + from this[of 0] have [simp]: "x \ \" + unfolding X_def by (simp add: x_def) + + have "a \ 0" + proof + assume "a = 0" + with root and nontrivial have "x = 0 \ x = -c / b" + by (auto simp: divide_simps add_eq_0_iff) + hence "x \ \" by (auto simp del: \x \ \\) + thus False by simp + qed + + have bounds: "(A n, B n, C n) \ {-A'..A'} \ {-B'..B'} \ {-A'..A'}" + and X_root: "A n * X n ^ 2 + B n * X n + C n = 0" if n: "n \ 2" for n + proof - + define n' where "n' = n - 2" + have n': "n = Suc (Suc n')" using \n \ 2\ unfolding n'_def by simp + have *: "of_int (k (n - Suc 0)) * X n + of_int (k (n - 2)) \ 0" + proof + assume "of_int (k (n - Suc 0)) * X n + of_int (k (n - 2)) = 0" + hence "X n = -k (n - 2) / k (n - 1)" by (auto simp: divide_simps mult_ac) + also have "\ \ \" by auto + finally show False by simp + qed + + let ?denom = "(k (n - 1) * X n + k (n - 2))" + have "0 = 0 * ?denom ^ 2" by simp + also have "0 * ?denom ^ 2 = (a * x ^ 2 + b * x + c) * ?denom ^ 2" using root by simp + also have "\ = a * (x * ?denom) ^ 2 + b * ?denom * (x * ?denom) + c * ?denom * ?denom" + by (simp add: algebra_simps power2_eq_square) + also have "x * ?denom = h (n - 1) * X n + h (n - 2)" + using cfrac_lim_eq_num_denom_remainder_aux[of "n - 2" e] \n \ 2\ + by (simp add: numeral_2_eq_2 Suc_diff_Suc x_def k_def h_def X_def) + also have "a * \ ^ 2 + b * ?denom * \ + c * ?denom * ?denom = A n * X n ^ 2 + B n * X n + C n" + by (simp add: A_def B_def C_def power2_eq_square algebra_simps) + finally show "A n * X n ^ 2 + B n * X n + C n = 0" .. + + have f_abs_bound: "\f (R n)\ \ (2 * \a\ * \x\ + \b\) * (1 / (k n * k (Suc n))) + + \a\ * (1 / (k n * k (Suc n))) ^ 2" for n + proof - + have "\f (R n)\ = \?f (R n) - ?f x\" by (simp add: root f_def) + also have "?f (R n) - ?f x = (R n - x) * (2 * a * x + b) + (R n - x) ^ 2 * a" + by (simp add: power2_eq_square algebra_simps) + also have "\\\ \ \(R n - x) * (2 * a * x + b)\ + \(R n - x) ^ 2 * a\" + by (rule abs_triangle_ineq) + also have "\ = \2 * a * x + b\ * \R n - x\ + \a\ * \R n - x\ ^ 2" + by (simp add: abs_mult) + also have "\ \ \2 * a * x + b\ * (1 / (k n * k (Suc n))) + \a\ * (1 / (k n * k (Suc n))) ^ 2" + unfolding x_def R_def using cfrac_lim_minus_conv_bounds[of n e] + by (intro add_mono mult_left_mono power_mono) (auto simp: k_def) + also have "\2 * a * x + b\ \ 2 * \a\ * \x\ + \b\" + by (rule order.trans[OF abs_triangle_ineq]) (auto simp: abs_mult) + hence "\2 * a * x + b\ * (1 / (k n * k (Suc n))) + \a\ * (1 / (k n * k (Suc n))) ^ 2 \ + \ * (1 / (k n * k (Suc n))) + \a\ * (1 / (k n * k (Suc n))) ^ 2" + by (intro add_mono mult_right_mono) (auto intro!: mult_nonneg_nonneg k_nonneg) + finally show "\f (R n)\ \ \" + by (simp add: mult_right_mono add_mono divide_left_mono) + qed + + have h_eq_conv_k: "h i = R i * k i" for i + using conv_denom_pos[of e i] unfolding R_def + by (subst conv_num_denom) (auto simp: h_def k_def) + + have "A n = k (n - 1) ^ 2 * f (R (n - 1))" for n + by (simp add: algebra_simps A_def n' k_def power2_eq_square h_eq_conv_k f_def) + have A_bound: "\A i\ \ A'" if "i > 0" for i + proof - + have "k i > 0" + by simp + hence "k i \ 1" + by linarith + have "A i = k (i - 1) ^ 2 * f (R (i - 1))" + by (simp add: algebra_simps A_def k_def power2_eq_square h_eq_conv_k f_def) + also have "\\\ = k (i - 1) ^ 2 * \f (R (i - 1))\" + by (simp add: abs_mult f_def) + also have "\ \ k (i - 1) ^ 2 * ((2 * \a\ * \x\ + \b\) * (1 / (k (i - 1) * k (Suc (i - 1)))) + + \a\ * (1 / (k (i - 1) * k (Suc (i - 1)))) ^ 2)" + by (intro mult_left_mono f_abs_bound) auto + also have "\ = k (i - 1) / k i * (2 * \a\ * \x\ + \b\) + \a\ / k i ^ 2" using \i > 0\ + by (simp add: power2_eq_square field_simps) + also have "\ \ 1 * (2 * \a\ * \x\ + \b\) + \a\ / 1" using \i > 0\ \k i \ 1\ + by (intro add_mono divide_left_mono mult_right_mono) + (auto intro!: k_leI one_le_power simp: of_nat_ge_1_iff) + also have "\ = 2 * \a\ * \x\ + \a\ + \b\" by simp + finally show ?thesis unfolding A'_def by linarith + qed + + have "C n = A (n - 1)" by (simp add: A_def C_def n') + hence C_bound: "\C n\ \ A'" using A_bound[of "n - 1"] n by simp + + have "B n = k (n - 1) * k (n - 2) * + (f (R (n - 1)) + f (R (n - 2)) - a * (R (n - 1) - R (n - 2)) ^ 2)" + by (simp add: B_def h_eq_conv_k algebra_simps power2_eq_square f_def) + also have "\\\ = k (n - 1) * k (n - 2) * + \f (R (n - 1)) + f (R (n - 2)) - a * (R (n - 1) - R (n - 2)) ^ 2\" + by (simp add: abs_mult k_nonneg) + also have "\ \ k (n - 1) * k (n - 2) * + (((2 * \a\ * \x\ + \b\) * (1 / (k (n - 1) * k (Suc (n - 1)))) + + \a\ * (1 / (k (n - 1) * k (Suc (n - 1)))) ^ 2) + + ((2 * \a\ * \x\ + \b\) * (1 / (k (n - 2) * k (Suc (n - 2)))) + + \a\ * (1 / (k (n - 2) * k (Suc (n - 2)))) ^ 2) + + \a\ * \R (Suc (n - 2)) - R (n - 2)\ ^ 2)" (is "_ \ _ * (?S1 + ?S2 + ?S3)") + by (intro mult_left_mono order.trans[OF abs_triangle_ineq4] order.trans[OF abs_triangle_ineq] + add_mono f_abs_bound order.refl) + (insert n, auto simp: abs_mult Suc_diff_Suc numeral_2_eq_2 k_nonneg) + also have "\R (Suc (n - 2)) - R (n - 2)\ = 1 / (k (n - 2) * k (Suc (n - 2)))" + unfolding R_def k_def by (rule abs_diff_successive_convs) + also have "of_int (k (n - 1) * k (n - 2)) * (?S1 + ?S2 + \a\ * \ ^ 2) = + (k (n - 2) / k n + 1) * (2 * \a\ * \x\ + \b\) + + \a\ * (k (n - 2) / (k (n - 1) * k n ^ 2) + 2 / (k (n - 1) * k (n - 2)))" + (is "_ = ?S") using n by (simp add: field_simps power2_eq_square numeral_2_eq_2 Suc_diff_Suc) + also { + have A: "2 * real_of_int (k (n - 2)) \ of_int (k n)" + using conv_denom_plus2_ratio_ge[of e "n - 2"] n + by (simp add: numeral_2_eq_2 Suc_diff_Suc k_def) + have "fib (Suc 2) \ k 2" unfolding k_def by (intro conv_denom_lower_bound) + also have "\ \ k n" by (intro k_leI n) + finally have "k n \ 2" by (simp add: numeral_3_eq_3) + hence B: "of_int (k (n - 2)) * 2 ^ 2 \ (of_int (k (n - 1)) * (of_int (k n))\<^sup>2 :: real)" + by (intro mult_mono power_mono) (auto intro: k_leI k_nonneg) + have C: "1 * 1 \ real_of_int (k (n - 1)) * of_int (k (n - 2))" using k_ge_1 + by (intro mult_mono) (auto simp: Suc_le_eq of_nat_ge_1_iff k_nonneg) + note A B C + } + hence "?S \ (1 / 2 + 1) * (2 * \a\ * \x\ + \b\) + \a\ * (1 / 4 + 2)" + by (intro add_mono mult_right_mono mult_left_mono) (auto simp: field_simps) + also have "\ = (3 / 2) * (2 * \a\ * \x\ + \b\) + 9 / 4 * \a\" by simp + finally have B_bound: "\B n\ \ B'" unfolding B'_def by linarith + from A_bound[of n] B_bound C_bound n + show "(A n, B n, C n) \ {-A'..A'} \ {-B'..B'} \ {-A'..A'}" by auto + qed + + have A_nz: "A n \ 0" if "n \ 1" for n + using that + proof (induction n rule: dec_induct) + case base + show ?case + proof + assume "A 1 = 0" + hence "real_of_int (A 1) = 0" by simp + also have "real_of_int (A 1) = + real_of_int a * of_int (cfrac_nth e 0) ^ 2 + + real_of_int b * cfrac_nth e 0 + real_of_int c" + by (simp add: A_def h_def k_def) + finally have root': "\ = 0" . + + have "cfrac_nth e 0 \ \" by auto + also from root' and \a \ 0\ have "?this \ is_square (nat (b\<^sup>2 - 4 * a * c))" + by (intro quadratic_equation_solution_rat_iff) auto + also from root and \a \ 0\ have "\ \ x \ \" + by (intro quadratic_equation_solution_rat_iff [symmetric]) auto + finally show False using \x \ \\ by contradiction + qed + next + case (step m) + hence nz: "C (Suc m) \ 0" by (simp add: C_def A_def) + show "A (Suc m) \ 0" + proof + assume [simp]: "A (Suc m) = 0" + have "X (Suc m) > 0" unfolding X_def + by (intro cfrac_remainder_pos) auto + with X_root[of "Suc m"] step.hyps nz have "X (Suc m) = -C (Suc m) / B (Suc m)" + by (auto simp: divide_simps mult_ac) + also have "\ \ \" by auto + finally show False by simp + qed + qed + + have "finite ({-A'..A'} \ {-B'..B'} \ {-A'..A'})" by auto + from this and bounds have "finite ((\n. (A n, B n, C n)) ` {2..})" + by (blast intro: finite_subset) + moreover have "infinite ({2..} :: nat set)" by (simp add: infinite_Ici) + ultimately have "\k1\{2..}. infinite {n \ {2..}. (A n, B n, C n) = (A k1, B k1, C k1)}" + by (intro pigeonhole_infinite) + then obtain k0 where k0: "k0 \ 2" "infinite {n \ {2..}. (A n, B n, C n) = (A k0, B k0, C k0)}" + by auto + from infinite_countable_subset[OF this(2)] obtain g :: "nat \ _" + where g: "inj g" "range g \ {n\{2..}. (A n, B n, C n) = (A k0, B k0, C k0)}" by blast + hence g_ge_2: "g k \ 2" for k by auto + from g have [simp]: "A (g k) = A k0" "B (g k) = B k0" "C (g k) = C k0" for k + by auto + + from g(1) have [simp]: "g k1 = g k2 \ k1 = k2" for k1 k2 by (auto simp: inj_def) + define z where "z = (A k0, B k0, C k0)" + let ?h = "\k. (A (g k), B (g k), C (g k))" + from g have g': "distinct [g 1, g 2, g 3]" "?h 0 = z" "?h 1 = z" "?h 2 = z" + by (auto simp: z_def) + have fin: "finite {x :: real. A k0 * x ^ 2 + B k0 * x + C k0 = 0}" using A_nz[of k0] k0(1) + by (subst finite_quadratic_equation_solutions_reals) auto + from X_root[of "g 0"] X_root[of "g 1"] X_root[of "g 2"] g_ge_2 g + have "(X \ g) ` {0, 1, 2} \ {x. A k0 * x ^ 2 + B k0 * x + C k0 = 0}" + by auto + hence "card ((X \ g) ` {0, 1, 2}) \ card \" + by (intro card_mono fin) auto + also have "\ \ 2" + by (rule card_quadratic_equation_solutions_reals_le_2) + also have "\ < card {0, 1, 2 :: nat}" by simp + finally have "\inj_on (X \ g) {0, 1, 2}" + by (rule pigeonhole) + then obtain m1 m2 where + m12: "m1 \ {0, 1, 2}" "m2 \ {0, 1, 2}" "X (g m1) = X (g m2)" "m1 \ m2" + unfolding inj_on_def o_def by blast + define n and l where "n = min (g m1) (g m2)" and "l = nat \int (g m1) - g m2\" + with m12 g' have l: "l > 0" "X (n + l) = X n" + by (auto simp: min_def nat_diff_distrib split: if_splits) + + from l have "cfrac_lim (cfrac_drop (n + l) e) = cfrac_lim (cfrac_drop n e)" + by (simp add: X_def cfrac_remainder_def) + hence "cfrac_drop (n + l) e = cfrac_drop n e" + by (simp add: cfrac_lim_eq_iff) + hence "cfrac_nth (cfrac_drop (n + l) e) = cfrac_nth (cfrac_drop n e)" + by (simp only:) + hence period: "cfrac_nth e (n + l + k) = cfrac_nth e (n + k)" for k + by (simp add: fun_eq_iff add_ac) + have period: "cfrac_nth e (k + l) = cfrac_nth e k" if "k \ n" for k + using period[of "k - n"] that by (simp add: add_ac) + have period: "cfrac_nth e (k + m * l) = cfrac_nth e k" if "k \ n" for k m + using that + proof (induction m) + case (Suc m) + have "cfrac_nth e (k + Suc m * l) = cfrac_nth e (k + m * l + l)" + by (simp add: algebra_simps) + also have "\ = cfrac_nth e (k + m * l)" + using Suc.prems by (intro period) auto + also have "\ = cfrac_nth e k" + using Suc.prems by (intro Suc.IH) auto + finally show ?case . + qed simp_all + + from this and l and that[of l n] show ?thesis by (simp add: X_def) +qed + +theorem periodic_cfrac_iff_quadratic_irrational: + assumes "x \ \" "x \ 0" + shows "quadratic_irrational x \ + (\N l. l > 0 \ (\n\N. cfrac_nth (cfrac_of_real x) (n + l) = + cfrac_nth (cfrac_of_real x) n))" +proof safe + assume *: "quadratic_irrational x" + with assms have **: "quadratic_irrational (cfrac_lim (cfrac_of_real x))" by auto + obtain N l where Nl: "l > 0" + "\n m. N \ n \ cfrac_nth (cfrac_of_real x) (n + m * l) = cfrac_nth (cfrac_of_real x) n" + "cfrac_remainder (cfrac_of_real x) (N + l) = cfrac_remainder (cfrac_of_real x) N" + "cfrac_length (cfrac_of_real x) = \" + using quadratic_irrational_imp_periodic_cfrac [OF **] by metis + show "\N l. l > 0 \ (\n\N. cfrac_nth (cfrac_of_real x) (n + l) = cfrac_nth (cfrac_of_real x) n)" + by (rule exI[of _ N], rule exI[of _ l]) (insert Nl(1) Nl(2)[of _ 1], auto) +next + fix N l assume "l > 0" "\n\N. cfrac_nth (cfrac_of_real x) (n + l) = cfrac_nth (cfrac_of_real x) n" + hence "quadratic_irrational (cfrac_lim (cfrac_of_real x))" using assms + by (intro periodic_cfrac_imp_quadratic_irrational[of _ l N]) auto + with assms show "quadratic_irrational x" + by simp +qed + +text \ + The following result can e.g. be used to show that a number is \<^emph>\not\ a quadratic + irrational. +\ +lemma quadratic_irrational_cfrac_nth_range_finite: + assumes "quadratic_irrational (cfrac_lim e)" + shows "finite (range (cfrac_nth e))" +proof - + from quadratic_irrational_imp_periodic_cfrac[OF assms] obtain l N + where period: "l > 0" "\m n. n \ N \ cfrac_nth e (n + m * l) = cfrac_nth e n" + by metis + have "cfrac_nth e k \ cfrac_nth e ` {.. cfrac_nth e ` {..l > 0\ by (intro imageI) (auto simp: n_def) + also have "cfrac_nth e n = cfrac_nth e (n + m * l)" + by (subst period) (auto simp: n_def) + also have "n + m * l = k" + using False by (simp add: n_def m_def) + finally show ?thesis . + qed auto + hence "range (cfrac_nth e) \ cfrac_nth e ` {..Continued fraction expansions for square roots of naturals\ +theory Sqrt_Nat_Cfrac +imports + Quadratic_Irrationals + "HOL-Library.While_Combinator" + "HOL-Library.IArray" +begin + +text \ + In this section, we shall explore the continued fraction expansion of $\sqrt{D}$, where $D$ + is a natural number. +\ + +lemma butlast_nth [simp]: "n < length xs - 1 \ butlast xs ! n = xs ! n" + by (induction xs arbitrary: n) (auto simp: nth_Cons split: nat.splits) + +text \ + The following is the length of the period in the continued fraction expansion of + $\sqrt{D}$ for a natural number $D$. +\ +definition sqrt_nat_period_length :: "nat \ nat" where + "sqrt_nat_period_length D = + (if is_square D then 0 + else (LEAST l. l > 0 \ (\n. cfrac_nth (cfrac_of_real (sqrt D)) (Suc n + l) = + cfrac_nth (cfrac_of_real (sqrt D)) (Suc n))))" + +text \ + Next, we define a more workable representation for the continued fraction expansion of + $\sqrt{D}$ consisting of the period length, the natural number $\lfloor\sqrt{D}\rfloor$, and + the content of the period. +\ +definition sqrt_cfrac_info :: "nat \ nat \ nat \ nat list" where + "sqrt_cfrac_info D = + (sqrt_nat_period_length D, Discrete.sqrt D, + map (\n. nat (cfrac_nth (cfrac_of_real (sqrt D)) (Suc n))) [0.. sqrt_nat_period_length D = 0" + by (auto simp: sqrt_nat_period_length_def) + +definition sqrt_cfrac :: "nat \ cfrac" + where "sqrt_cfrac D = cfrac_of_real (sqrt (real D))" + +context + fixes D D' :: nat + defines "D' \ nat \sqrt D\" +begin + +text \ + A number $\alpha = \frac{\sqrt D + p}{q}$ for \p, q \ \\ is called a \<^emph>\reduced quadratic surd\ + if $\alpha > 1$ and $bar\alpha \in (-1;0)$, where $\bar\alpha$ denotes the conjugate + $\frac{-\sqrt D + p}{q}$. + + It is furthermore called \<^emph>\associated\ to $D$ if \q\ divides \D - p\<^sup>2\. +\ +definition red_assoc :: "nat \ nat \ bool" where + "red_assoc = (\(p, q). + q > 0 \ q dvd (D - p\<^sup>2) \ (sqrt D + p) / q > 1 \ (-sqrt D + p) / q \ {-1<..<0})" + +text \ + The following two functions convert between a surd represented as a pair of natural numbers + and the actual real number and its conjugate: +\ +definition surd_to_real :: "nat \ nat \ real" + where "surd_to_real = (\(p, q). (sqrt D + p) / q)" + +definition surd_to_real_cnj :: "nat \ nat \ real" + where "surd_to_real_cnj = (\(p, q). (-sqrt D + p) / q)" + +text \ + The next function performs a single step in the continued fraction expansion of $\sqrt{D}$. +\ +definition sqrt_remainder_step :: "nat \ nat \ nat \ nat" where + "sqrt_remainder_step = (\(p, q). let X = (p + D') div q; p' = X * q - p in (p', (D - p'\<^sup>2) div q))" + +text \ + If we iterate this step function starting from the surd + $\frac{1}{\sqrt{D} - \lfloor\sqrt{D}\rfloor}$, we get the entire expansion. +\ +definition sqrt_remainder_surd :: "nat \ nat \ nat" + where "sqrt_remainder_surd = (\n. (sqrt_remainder_step ^^ n) (D', D - D'\<^sup>2))" + +context + fixes sqrt_cfrac_nth :: "nat \ nat" and l + assumes nonsquare: "\is_square D" + defines "sqrt_cfrac_nth \ (\n. case sqrt_remainder_surd n of (p, q) \ (D' + p) div q)" + defines "l \ sqrt_nat_period_length D" +begin + +lemma D'_pos: "D' > 0" + using nonsquare by (auto simp: D'_def of_nat_ge_1_iff intro: Nat.gr0I) + +lemma D'_sqr_less_D: "D'\<^sup>2 < D" +proof - + have "D' \ sqrt D" by (auto simp: D'_def) + hence "real D' ^ 2 \ sqrt D ^ 2" by (intro power_mono) auto + also have "\ = D" by simp + finally have "D'\<^sup>2 \ D" by simp + moreover from nonsquare have "D \ D'\<^sup>2" by auto + ultimately show ?thesis by simp +qed + +lemma red_assoc_imp_irrat: + assumes "red_assoc pq" + shows "surd_to_real pq \ \" +proof + assume rat: "surd_to_real pq \ \" + with assms rat show False using irrat_sqrt_nonsquare[OF nonsquare] + by (auto simp: field_simps red_assoc_def surd_to_real_def divide_in_Rats_iff2 add_in_Rats_iff1) +qed + +lemma surd_to_real_cnj_irrat: + assumes "red_assoc pq" + shows "surd_to_real_cnj pq \ \" +proof + assume rat: "surd_to_real_cnj pq \ \" + with assms rat show False using irrat_sqrt_nonsquare[OF nonsquare] + by (auto simp: field_simps red_assoc_def surd_to_real_cnj_def divide_in_Rats_iff2 diff_in_Rats_iff1) +qed + +lemma surd_to_real_nonneg [intro]: "surd_to_real pq \ 0" + by (auto simp: surd_to_real_def case_prod_unfold divide_simps intro!: divide_nonneg_nonneg) + +lemma surd_to_real_pos [intro]: "red_assoc pq \ surd_to_real pq > 0" + by (auto simp: surd_to_real_def case_prod_unfold divide_simps red_assoc_def + intro!: divide_nonneg_nonneg) + +lemma surd_to_real_nz [simp]: "red_assoc pq \ surd_to_real pq \ 0" + by (auto simp: surd_to_real_def case_prod_unfold divide_simps red_assoc_def + intro!: divide_nonneg_nonneg) + +lemma surd_to_real_cnj_nz [simp]: "red_assoc pq \ surd_to_real_cnj pq \ 0" + using surd_to_real_cnj_irrat[of pq] by auto + +lemma red_assoc_step: + assumes "red_assoc pq" + defines "X \ (D' + fst pq) div snd pq" + defines "pq' \ sqrt_remainder_step pq" + shows "red_assoc pq'" + "surd_to_real pq' = 1 / frac (surd_to_real pq)" + "surd_to_real_cnj pq' = 1 / (surd_to_real_cnj pq - X)" + "X > 0" "X * snd pq \ 2 * D'" "X = nat \surd_to_real pq\" + "X = nat \-1 / surd_to_real_cnj pq'\" +proof - + obtain p q where [simp]: "pq = (p, q)" by (cases pq) + obtain p' q' where [simp]: "pq' = (p', q')" by (cases pq') + define \ where "\ = (sqrt D + p) / q" + define \' where "\' = 1 / frac \" + define cnj_\' where "cnj_\' = (-sqrt D + (X * q - int p)) / ((D - (X * q - int p)\<^sup>2) div q)" + from assms(1) have "\ > 0" "q > 0" + by (auto simp: \_def red_assoc_def) + from assms(1) nonsquare have "\ \ \" + by (auto simp: \_def red_assoc_def divide_in_Rats_iff2 add_in_Rats_iff2 irrat_sqrt_nonsquare) + hence \'_pos: "frac \ > 0" using Ints_subset_Rats by auto + from \pq' = (p', q')\ have p'_def: "p' = X * q - p" and q'_def: "q' = (D - p'\<^sup>2) div q" + unfolding pq'_def sqrt_remainder_step_def X_def by (auto simp: Let_def add_ac) + + have "D' + p = \sqrt D + p\" + by (auto simp: D'_def) + also have "\ div int q = \(sqrt D + p) / q\" + by (subst floor_divide_real_eq_div [symmetric]) auto + finally have X_altdef: "X = nat \(sqrt D + p) / q\" + unfolding X_def zdiv_int [symmetric] by auto + + have nz: "sqrt (real D) + (X * q - real p) \ 0" + proof + assume "sqrt (real D) + (X * q - real p) = 0" + hence "sqrt (real D) = real p - X * q" by (simp add: algebra_simps) + also have "\ \ \" by auto + finally show False using irrat_sqrt_nonsquare nonsquare by blast + qed + + from assms(1) have "real (p ^ 2) \ sqrt D ^ 2" + unfolding of_nat_power by (intro power_mono) (auto simp: red_assoc_def field_simps) + also have "sqrt D ^ 2 = D" by simp + finally have "p\<^sup>2 \ D" by (subst (asm) of_nat_le_iff) + + have "frac \ = \ - X" + by (simp add: X_altdef frac_def \_def) + also have "\ = (sqrt D - (X * q - int p)) / q" + using \q > 0\ by (simp add: field_simps \_def) + finally have "1 / frac \ = q / (sqrt D - (X * q - int p))" + by simp + also have "\ = q * (sqrt D + (X * q - int p)) / + ((sqrt D - (X * q - int p)) * (sqrt D + (X * q - int p)))" (is "_ = ?A / ?B") + using nz by (subst mult_divide_mult_cancel_right) auto + also have "?B = real_of_int (D - int p ^ 2 + 2 * X * p * q - int X ^ 2 * q ^ 2)" + by (auto simp: algebra_simps power2_eq_square) + also have "q dvd (D - p ^ 2)" using assms(1) by (auto simp: red_assoc_def) + with \p\<^sup>2 \ D\ have "int q dvd (int D - int p ^ 2)" + unfolding of_nat_power [symmetric] by (subst of_nat_diff [symmetric]) auto + hence "D - int p ^ 2 + 2 * X * p * q - int X ^ 2 * q ^ 2 = q * ((D - (X * q - int p)\<^sup>2) div q)" + by (auto simp: power2_eq_square algebra_simps) + also have "?A / \ = (sqrt D + (X * q - int p)) / ((D - (X * q - int p)\<^sup>2) div q)" + unfolding of_int_mult of_int_of_nat_eq + by (rule mult_divide_mult_cancel_left) (insert \q > 0\, auto) + finally have \': "\' = \" by (simp add: \'_def) + + have dvd: "q dvd (D - (X * q - int p)\<^sup>2)" + using assms(1) \int q dvd (int D - int p ^ 2)\ + by (auto simp: power2_eq_square algebra_simps) + + have "X \ (sqrt D + p) / q" unfolding X_altdef by simp + moreover have "X \ (sqrt D + p) / q" + proof + assume "X = (sqrt D + p) / q" + hence "sqrt D = q * X - real p" using \q > 0\ by (auto simp: field_simps) + also have "\ \ \" by auto + finally show False using irrat_sqrt_nonsquare[OF nonsquare] by simp + qed + ultimately have "X < (sqrt D + p) / q" by simp + hence *: "(X * q - int p) < sqrt D" + using \q > 0\ by (simp add: field_simps) + moreover + have pos: "real_of_int (int D - (int X * int q - int p)\<^sup>2) > 0" + proof (cases "X * q \ p") + case True + hence "real p \ real X * real q" unfolding of_nat_mult [symmetric] of_nat_le_iff . + hence "real_of_int ((X * q - int p) ^ 2) < sqrt D ^ 2" using * + unfolding of_int_power by (intro power_strict_mono) auto + also have "\ = D" by simp + finally show ?thesis by simp + next + case False + hence less: "real X * real q < real p" + unfolding of_nat_mult [symmetric] of_nat_less_iff by auto + have "(real X * real q - real p)\<^sup>2 = (real p - real X * real q)\<^sup>2" + by (simp add: power2_eq_square algebra_simps) + also have "\ \ real p ^ 2" using less by (intro power_mono) auto + also have "\ < sqrt D ^ 2" + using \q > 0\ assms(1) unfolding of_int_power + by (intro power_strict_mono) (auto simp: red_assoc_def field_simps) + also have "\ = D" by simp + finally show ?thesis by simp + qed + hence pos': "int D - (int X * int q - int p)\<^sup>2 > 0" + by (subst (asm) of_int_0_less_iff) + from pos have "real_of_int ((int D - (int X * int q - int p)\<^sup>2) div q) > 0" + using \q > 0\ dvd by (subst real_of_int_div) (auto intro!: divide_pos_pos) + ultimately have cnj_neg: "cnj_\' < 0" unfolding cnj_\'_def using dvd + unfolding of_int_0_less_iff by (intro divide_neg_pos) auto + + have "(p - sqrt D) / q < 0" + using assms(1) by (auto simp: red_assoc_def X_altdef le_nat_iff) + also have "X \ 1" + using assms(1) by (auto simp: red_assoc_def X_altdef le_nat_iff) + hence "0 \ real X - 1" by simp + finally have "q < sqrt D + int q * X - p" + using \q > 0\ by (simp add: field_simps) + hence "q * (sqrt D - (int q * X - p)) < (sqrt D + (int q * X - p)) * (sqrt D - (int q * X - p))" + using * by (intro mult_strict_right_mono) (auto simp: red_assoc_def X_altdef field_simps) + also have "\ = D - (int q * X - p) ^ 2" + by (simp add: power2_eq_square algebra_simps) + finally have "cnj_\' > -1" + using dvd pos \q > 0\ by (simp add: real_of_int_div field_simps cnj_\'_def) + + from cnj_neg and this have "cnj_\' \ {-1<..<0}" by auto + have "\' > 1" using \frac \ > 0\ + by (auto simp: \'_def field_simps frac_lt_1) + + have "0 = 1 + (-1 :: real)" + by simp + also have "1 + -1 < \' + cnj_\'" + using \cnj_\' > -1\ and \\' > 1\ by (intro add_strict_mono) + also have "\' + cnj_\' = 2 * (real X * q - real p) / ((int D - (int X * q - int p)\<^sup>2) div int q)" + by (simp add: \' cnj_\'_def add_divide_distrib [symmetric]) + finally have "real X * q - real p > 0" using pos dvd \q > 0\ + by (subst (asm) zero_less_divide_iff, subst (asm) (1 2 3) real_of_int_div) + (auto simp: field_simps) + hence "real (X * q) > real p" unfolding of_nat_mult by simp + hence p_less_Xq: "p < X * q" by (simp only: of_nat_less_iff) + + from pos' and p_less_Xq have "int D > int ((X * q - p)\<^sup>2)" + by (subst of_nat_power) (auto simp: of_nat_diff) + hence pos'': "D > (X * q - p)\<^sup>2" unfolding of_nat_less_iff . + + from dvd have "int q dvd int (D - (X * q - p)\<^sup>2)" + using p_less_Xq pos'' by (subst of_nat_diff) (auto simp: of_nat_diff) + with dvd have dvd': "q dvd (D - (X * q - p)\<^sup>2)" + by simp + + have \'_altdef: "\' = (sqrt D + p') / q'" + using dvd dvd' pos'' p_less_Xq \' + by (simp add: real_of_int_div p'_def q'_def real_of_nat_div mult_ac of_nat_diff) + have cnj_\'_altdef: "cnj_\' = (-sqrt D + p') / q'" + using dvd dvd' pos'' p_less_Xq unfolding cnj_\'_def + by (simp add: real_of_int_div p'_def q'_def real_of_nat_div mult_ac of_nat_diff) + from dvd' have dvd'': "q' dvd (D - p'\<^sup>2)" + by (auto simp: mult_ac p'_def q'_def) + have "real ((D - p'\<^sup>2) div q) > 0" unfolding p'_def + by (subst real_of_nat_div[OF dvd'], rule divide_pos_pos) (insert \q > 0\ pos'', auto) + hence "q' > 0" unfolding q'_def of_nat_0_less_iff . + + show "red_assoc pq'" using \\' > 1\ and \cnj_\' \ _\ and dvd'' and \q' > 0\ + by (auto simp: red_assoc_def \'_altdef cnj_\'_altdef) + + from assms(1) have "real p < sqrt D" + by (auto simp add: field_simps red_assoc_def) + hence "p \ D'" unfolding D'_def by linarith + with * have "real (X * q) < sqrt (real D) + D'" + by simp + thus "X * snd pq \ 2 * D'" unfolding D'_def \pq = (p, q)\ snd_conv by linarith + + have "(sqrt D + p') / q' = \'" + by (rule \'_altdef [symmetric]) + also have "\' = 1 / frac ((sqrt D + p) / q)" + by (simp add: \'_def \_def) + finally show "surd_to_real pq' = 1 / frac (surd_to_real pq)" by (simp add: surd_to_real_def) + from \X \ 1\ show "X > 0" by simp + from X_altdef show "X = nat \surd_to_real pq\" by (simp add: surd_to_real_def) + + have "sqrt (real D) < real p + 1 * real q" + using assms(1) by (auto simp: red_assoc_def field_simps) + also have "\ \ real p + real X * real q" + using \X > 0\ by (intro add_left_mono mult_right_mono) (auto simp: of_nat_ge_1_iff) + finally have "sqrt (real D) < \" . + + have "real p < sqrt D" + using assms(1) by (auto simp add: field_simps red_assoc_def) + also have "\ \ sqrt D + q * X" + by linarith + finally have less: "real p < sqrt D + X * q" by (simp add: algebra_simps) + moreover have "D + p * p' + X * q * sqrt D = q * q' + p * sqrt D + p' * sqrt D + X * p' * q" + using dvd' pos'' p_less_Xq \q > 0\ unfolding p'_def q'_def of_nat_mult of_nat_add + by (simp add: power2_eq_square field_simps of_nat_diff real_of_nat_div) + ultimately show *: "surd_to_real_cnj pq' = 1 / (surd_to_real_cnj pq - X)" + using \q > 0\ \q' > 0\ by (auto simp: surd_to_real_cnj_def field_simps) + + have **: "a = nat \y\" if "x \ 0" "x < 1" "real a + x = y" for a :: nat and x y :: real + using that by linarith + from assms(1) have surd_to_real_cnj: "surd_to_real_cnj (p, q) \ {-1<..<0}" + by (auto simp: surd_to_real_cnj_def red_assoc_def) + have "surd_to_real_cnj (p, q) < X" + using assms(1) less by (auto simp: surd_to_real_cnj_def field_simps red_assoc_def) + hence "real X = surd_to_real_cnj (p, q) - 1 / surd_to_real_cnj (p', q')" using * + using surd_to_real_cnj_irrat assms(1) \red_assoc pq'\ by (auto simp: field_simps) + thus "X = nat \-1 / surd_to_real_cnj pq'\" using surd_to_real_cnj + by (intro **[of "-surd_to_real_cnj (p, q)"]) auto +qed + +lemma red_assoc_denom_2D: + assumes "red_assoc (p, q)" + defines "X \ (D' + p) div q" + assumes "X > D'" + shows "q = 1" +proof - + have "X * q \ 2 * D'" "X > 0" + using red_assoc_step(4,5)[OF assms(1)] by (simp_all add: X_def) + note this(1) + also have "2 * D' < 2 * X" + by (intro mult_strict_left_mono assms) auto + finally have "q < 2" using \X > 0\ by simp + moreover from assms(1) have "q > 0" by (auto simp: red_assoc_def) + ultimately show ?thesis by simp +qed + +lemma red_assoc_denom_1: + assumes "red_assoc (p, 1)" + shows "p = D'" +proof - + from assms have "sqrt D > p" "sqrt D < real p + 1" + by (auto simp: red_assoc_def) + thus "p = D'" unfolding D'_def + by linarith +qed + +lemma red_assoc_begin: + "red_assoc (D', D - D'\<^sup>2)" + "surd_to_real (D', D - D'\<^sup>2) = 1 / frac (sqrt D)" + "surd_to_real_cnj (D', D - D'\<^sup>2) = -1 / (sqrt D + D')" +proof - + have pos: "D > 0" "D' > 0" + using nonsquare by (auto simp: D'_def of_nat_ge_1_iff intro!: Nat.gr0I) + + have "sqrt D \ D'" + using irrat_sqrt_nonsquare[OF nonsquare] by auto + moreover have "sqrt D \ 0" by simp + hence "D' \ sqrt D" unfolding D'_def by linarith + ultimately have less: "D' < sqrt D" by simp + + have "sqrt D \ D' + 1" + using irrat_sqrt_nonsquare[OF nonsquare] by auto + moreover have "sqrt D \ 0" by simp + hence "D' \ sqrt D - 1" unfolding D'_def by linarith + ultimately have gt: "D' > sqrt D - 1" by simp + + from less have "real D' ^ 2 < sqrt D ^ 2" by (intro power_strict_mono) auto + also have "\ = D" by simp + finally have less': "D'\<^sup>2 < D" unfolding of_nat_power [symmetric] of_nat_less_iff . + + moreover have "real D' * (real D' - 1) < sqrt D * (sqrt D - 1)" + using less pos + by (intro mult_strict_mono diff_strict_right_mono) (auto simp: of_nat_ge_1_iff) + hence "D'\<^sup>2 + sqrt D < D' + D" + by (simp add: field_simps power2_eq_square) + moreover have "(sqrt D - 1) * sqrt D < real D' * (real D' + 1)" + using pos gt by (intro mult_strict_mono) auto + hence "D < sqrt D + D'\<^sup>2 + D'" by (simp add: power2_eq_square field_simps) + ultimately show "red_assoc (D', D - D'\<^sup>2)" + by (auto simp: red_assoc_def field_simps of_nat_diff less) + + have frac: "frac (sqrt D) = sqrt D - D'" unfolding frac_def D'_def + by auto + show "surd_to_real (D', D - D'\<^sup>2) = 1 / frac (sqrt D)" unfolding surd_to_real_def + using less less' pos by (subst frac) (auto simp: of_nat_diff power2_eq_square field_simps) + + have "surd_to_real_cnj (D', D - D'\<^sup>2) = -((sqrt D - D') / (D - D'\<^sup>2))" + using less less' pos by (auto simp: surd_to_real_cnj_def field_simps) + also have "real (D - D'\<^sup>2) = (sqrt D - D') * (sqrt D + D')" + using less' by (simp add: power2_eq_square algebra_simps of_nat_diff) + also have "(sqrt D - D') / \ = 1 / (sqrt D + D')" + using less by (subst nonzero_divide_mult_cancel_left) auto + finally show "surd_to_real_cnj (D', D - D'\<^sup>2) = -1 / (sqrt D + D')" by simp +qed + +lemma cfrac_remainder_surd_to_real: + assumes "red_assoc pq" + shows "cfrac_remainder (cfrac_of_real (surd_to_real pq)) n = + surd_to_real ((sqrt_remainder_step ^^ n) pq)" + using assms(1) +proof (induction n arbitrary: pq) + case 0 + hence "cfrac_lim (cfrac_of_real (surd_to_real pq)) = surd_to_real pq" + by (intro cfrac_lim_of_real red_assoc_imp_irrat 0) + thus ?case using 0 + by auto +next + case (Suc n) + obtain p q where [simp]: "pq = (p, q)" by (cases pq) + have "surd_to_real ((sqrt_remainder_step ^^ Suc n) pq) = + surd_to_real ((sqrt_remainder_step ^^ n) (sqrt_remainder_step (p, q)))" + by (subst funpow_Suc_right) auto + also have "\ = cfrac_remainder (cfrac_of_real (surd_to_real (sqrt_remainder_step (p, q)))) n" + using red_assoc_step(1)[of "(p, q)"] Suc.prems + by (intro Suc.IH [symmetric]) (auto simp: sqrt_remainder_step_def Let_def add_ac) + also have "surd_to_real (sqrt_remainder_step (p, q)) = 1 / frac (surd_to_real (p, q))" + using red_assoc_step(2)[of "(p, q)"] Suc.prems + by (auto simp: sqrt_remainder_step_def Let_def add_ac surd_to_real_def) + also have "cfrac_of_real \ = cfrac_tl (cfrac_of_real (surd_to_real (p, q)))" + using Suc.prems Ints_subset_Rats red_assoc_imp_irrat by (subst cfrac_tl_of_real) auto + also have "cfrac_remainder \ n = cfrac_remainder (cfrac_of_real (surd_to_real (p, q))) (Suc n)" + by (simp add: cfrac_drop_Suc_right cfrac_remainder_def) + finally show ?case by simp +qed + +lemma red_assoc_step' [intro]: "red_assoc pq \ red_assoc (sqrt_remainder_step pq)" + using red_assoc_step(1)[of pq] + by (simp add: sqrt_remainder_step_def case_prod_unfold add_ac Let_def) + +lemma red_assoc_steps [intro]: "red_assoc pq \ red_assoc ((sqrt_remainder_step ^^ n) pq)" + by (induction n) auto + +lemma floor_sqrt_less_sqrt: "D' < sqrt D" +proof - + have "D' \ sqrt D" unfolding D'_def by auto + moreover have "sqrt D \ D'" + using irrat_sqrt_nonsquare[OF nonsquare] by auto + ultimately show ?thesis by auto +qed + +lemma red_assoc_bounds: + assumes "red_assoc pq" + shows "pq \ (SIGMA p:{0<..D'}. {Suc D' - p..D' + p})" +proof - + obtain p q where [simp]: "pq = (p, q)" by (cases pq) + from assms have *: "p < sqrt D" + by (auto simp: red_assoc_def field_simps) + hence p: "p \ D'" unfolding D'_def by linarith + from assms have "p > 0" by (auto intro!: Nat.gr0I simp: red_assoc_def) + + have "q > sqrt D - p" "q < sqrt D + p" + using assms by (auto simp: red_assoc_def field_simps) + hence "q \ D' + 1 - p" "q \ D' + p" + unfolding D'_def by linarith+ + with p \p > 0\ show ?thesis by simp +qed + +lemma surd_to_real_cnj_eq_iff: + assumes "red_assoc pq" "red_assoc pq'" + shows "surd_to_real_cnj pq = surd_to_real_cnj pq' \ pq = pq'" +proof + assume eq: "surd_to_real_cnj pq = surd_to_real_cnj pq'" + from assms have pos: "snd pq > 0" "snd pq' > 0" by (auto simp: red_assoc_def) + have "snd pq = snd pq'" + proof (rule ccontr) + assume "snd pq \ snd pq'" + with eq have "sqrt D = (real (fst pq' * snd pq) - fst pq * snd pq') / (real (snd pq) - snd pq')" + using pos by (auto simp: field_simps surd_to_real_cnj_def case_prod_unfold) + also have "\ \ \" by auto + finally show False using irrat_sqrt_nonsquare[OF nonsquare] by auto + qed + moreover from this eq pos have "fst pq = fst pq'" + by (auto simp: surd_to_real_cnj_def case_prod_unfold) + ultimately show "pq = pq'" by (simp add: prod_eq_iff) +qed auto + +lemma red_assoc_sqrt_remainder_surd [intro]: "red_assoc (sqrt_remainder_surd n)" + by (auto simp: sqrt_remainder_surd_def intro!: red_assoc_begin) + +lemma surd_to_real_sqrt_remainder_surd: + "surd_to_real (sqrt_remainder_surd n) = cfrac_remainder (cfrac_of_real (sqrt D)) (Suc n)" +proof (induction n) + case 0 + from nonsquare have "D > 0" by (auto intro!: Nat.gr0I) + with red_assoc_begin show ?case using nonsquare irrat_sqrt_nonsquare[OF nonsquare] + using Ints_subset_Rats cfrac_drop_Suc_right cfrac_remainder_def cfrac_tl_of_real + sqrt_remainder_surd_def by fastforce +next + case (Suc n) + have "surd_to_real (sqrt_remainder_surd (Suc n)) = + surd_to_real (sqrt_remainder_step (sqrt_remainder_surd n))" + by (simp add: sqrt_remainder_surd_def) + also have "\ = 1 / frac (surd_to_real (sqrt_remainder_surd n))" + using red_assoc_step[OF red_assoc_sqrt_remainder_surd[of n]] by simp + also have "surd_to_real (sqrt_remainder_surd n) = + cfrac_remainder (cfrac_of_real (sqrt D)) (Suc n)" (is "_ = ?X") + by (rule Suc.IH) + also have "\cfrac_remainder (cfrac_of_real (sqrt (real D))) (Suc n)\ = + cfrac_nth (cfrac_of_real (sqrt (real D))) (Suc n)" + using irrat_sqrt_nonsquare[OF nonsquare] by (intro floor_cfrac_remainder) auto + hence "1 / frac ?X = cfrac_remainder (cfrac_of_real (sqrt D)) (Suc (Suc n))" + using irrat_sqrt_nonsquare[OF nonsquare] + by (subst cfrac_remainder_Suc[of "Suc n"]) + (simp_all add: frac_def cfrac_length_of_real_irrational) + finally show ?case . +qed + +lemma sqrt_cfrac: "sqrt_cfrac_nth n = cfrac_nth (cfrac_of_real (sqrt D)) (Suc n)" +proof - + have "cfrac_nth (cfrac_of_real (sqrt D)) (Suc n) = + \cfrac_remainder (cfrac_of_real (sqrt D)) (Suc n)\" + using irrat_sqrt_nonsquare[OF nonsquare] by (subst floor_cfrac_remainder) auto + also have "cfrac_remainder (cfrac_of_real (sqrt D)) (Suc n) = surd_to_real (sqrt_remainder_surd n)" + by (rule surd_to_real_sqrt_remainder_surd [symmetric]) + also have "nat \surd_to_real (sqrt_remainder_surd n)\ = sqrt_cfrac_nth n" + unfolding sqrt_cfrac_nth_def using red_assoc_step(6)[OF red_assoc_sqrt_remainder_surd[of n]] + by (simp add: case_prod_unfold) + finally show ?thesis + by (simp add: nat_eq_iff) +qed + +lemma sqrt_cfrac_pos: "sqrt_cfrac_nth k > 0" + using red_assoc_step(4)[OF red_assoc_sqrt_remainder_surd[of k]] + by (simp add: sqrt_cfrac_nth_def case_prod_unfold) + +lemma snd_sqrt_remainder_surd_pos: "snd (sqrt_remainder_surd n) > 0" + using red_assoc_sqrt_remainder_surd[of n] by (auto simp: red_assoc_def) + + +lemma + shows period_nonempty: "l > 0" + and period_length_le_aux: "l \ D' * (D' + 1)" + and sqrt_remainder_surd_periodic: "\n. sqrt_remainder_surd n = sqrt_remainder_surd (n mod l)" + and sqrt_cfrac_periodic: "\n. sqrt_cfrac_nth n = sqrt_cfrac_nth (n mod l)" + and sqrt_remainder_surd_smallest_period: + "\n. n \ {0<.. sqrt_remainder_surd n \ sqrt_remainder_surd 0" + and snd_sqrt_remainder_surd_gt_1: "\n. n < l - 1 \ snd (sqrt_remainder_surd n) > 1" + and sqrt_cfrac_le: "\n. n < l - 1 \ sqrt_cfrac_nth n \ D'" + and sqrt_remainder_surd_last: "sqrt_remainder_surd (l - 1) = (D', 1)" + and sqrt_cfrac_last: "sqrt_cfrac_nth (l - 1) = 2 * D'" + and sqrt_cfrac_palindrome: "\n. n < l - 1 \ sqrt_cfrac_nth (l - n - 2) = sqrt_cfrac_nth n" + and sqrt_cfrac_smallest_period: + "\l'. l' > 0 \ (\k. sqrt_cfrac_nth (k + l') = sqrt_cfrac_nth k) \ l' \ l" +proof - + note [simp] = sqrt_remainder_surd_def + define f where "f = sqrt_remainder_surd" + have *[intro]: "red_assoc (f n)" for n + unfolding f_def by (rule red_assoc_sqrt_remainder_surd) + + define S where "S = (SIGMA p:{0<..D'}. {Suc D' - p..D' + p})" + have [intro]: "finite S" by (simp add: S_def) + have "card S = (\p=1..D'. 2 * p)" unfolding S_def + by (subst card_SigmaI) (auto intro!: sum.cong) + also have "\ = D' * (D' + 1)" + by (induction D') (auto simp: power2_eq_square) + finally have [simp]: "card S = D' * (D' + 1)" . + + have "D' * (D' + 1) + 1 = card {..D' * (D' + 1)}" by simp + define k1 where + "k1 = (LEAST k1. k1 \ D' * (D' + 1) \ (\k2. k2 \ D' * (D' + 1) \ k1 \ k2 \ f k1 = f k2))" + define k2 where + "k2 = (LEAST k2. k2 \ D' * (D' + 1) \ k1 \ k2 \ f k1 = f k2)" + + have "f ` {..D' * (D' + 1)} \ S" unfolding S_def + using red_assoc_bounds[OF *] by blast + hence "card (f ` {..D' * (D' + 1)}) \ card S" + by (intro card_mono) auto + also have "card S = D' * (D' + 1)" by simp + also have "\ < card {..D' * (D' + 1)}" by simp + finally have "\inj_on f {..D' * (D' + 1)}" + by (rule pigeonhole) + hence "\k1. k1 \ D' * (D' + 1) \ (\k2. k2 \ D' * (D' + 1) \ k1 \ k2 \ f k1 = f k2)" + by (auto simp: inj_on_def) + from LeastI_ex[OF this, folded k1_def] + have "k1 \ D' * (D' + 1)" "\k2\D' * (D' + 1). k1 \ k2 \ f k1 = f k2" by auto + moreover from LeastI_ex[OF this(2), folded k2_def] + have "k2 \ D' * (D' + 1)" "k1 \ k2" "f k1 = f k2" by auto + moreover have "k1 \ k2" + proof (rule ccontr) + assume "\(k1 \ k2)" + hence "k2 \ D' * (D' + 1) \ (\k2'. k2' \ D' * (D' + 1) \ k2 \ k2' \ f k2 = f k2')" + using \k1 \ D' * (D' + 1)\ and \k1 \ k2\ and \f k1 = f k2\ by auto + hence "k1 \ k2" unfolding k1_def by (rule Least_le) + with \\(k1 \ k2)\ show False by simp + qed + ultimately have k12: "k1 < k2" "k2 \ D' * (D' + 1)" "f k1 = f k2" by auto + + have [simp]: "k1 = 0" + proof (cases k1) + case (Suc k1') + define k2' where "k2' = k2 - 1" + have Suc': "k2 = Suc k2'" using k12 by (simp add: k2'_def) + have nz: "surd_to_real_cnj (sqrt_remainder_step (f k1')) \ 0" + "surd_to_real_cnj (sqrt_remainder_step (f k2')) \ 0" + using surd_to_real_cnj_nz[OF *[of k2]] surd_to_real_cnj_nz[OF *[of k1]] + by (simp_all add: f_def Suc Suc') + + define a where "a = (D' + fst (f k1)) div snd (f k1)" + define a' where "a' = (D' + fst (f k1')) div snd (f k1')" + define a'' where "a'' = (D' + fst (f k2')) div snd (f k2')" + have "a' = nat \- 1 / surd_to_real_cnj (sqrt_remainder_step (f k1'))\" + using red_assoc_step[OF *[of k1']] by (simp add: a'_def) + also have "sqrt_remainder_step (f k1') = f k1" + by (simp add: Suc f_def) + also have "f k1 = f k2" by fact + also have "f k2 = sqrt_remainder_step (f k2')" by (simp add: Suc' f_def) + also have "nat \- 1 / surd_to_real_cnj (sqrt_remainder_step (f k2'))\ = a''" + using red_assoc_step[OF *[of k2']] by (simp add: a''_def) + finally have a'_a'': "a' = a''" . + + have "surd_to_real_cnj (f k2') \ a''" + using surd_to_real_cnj_irrat[OF *[of k2']] by auto + hence "surd_to_real_cnj (f k2') = 1 / surd_to_real_cnj (sqrt_remainder_step (f k2')) + a''" + using red_assoc_step(3)[OF *[of k2'], folded a''_def] nz + by (simp add: field_simps) + also have "\ = 1 / surd_to_real_cnj (sqrt_remainder_step (f k1')) + a'" + using k12 by (simp add: a'_a'' k12 Suc Suc' f_def) + also have nz': "surd_to_real_cnj (f k1') \ a'" + using surd_to_real_cnj_irrat[OF *[of k1']] by auto + hence "1 / surd_to_real_cnj (sqrt_remainder_step (f k1')) + a' = surd_to_real_cnj (f k1')" + using red_assoc_step(3)[OF *[of k1'], folded a'_def] nz nz' + by (simp add: field_simps) + finally have "f k1' = f k2'" + by (subst (asm) surd_to_real_cnj_eq_iff) auto + with k12 have "k1' \ D' * (D' + 1) \ (\k2\D' * (D' + 1). k1' \ k2 \ f k1' = f k2)" + by (auto simp: Suc Suc' intro!: exI[of _ k2']) + hence "k1 \ k1'" unfolding k1_def by (rule Least_le) + thus "k1 = 0" by (simp add: Suc) + qed auto + + have smallest_period: "f k \ f 0" if "k \ {0<.. D' * (D' + 1) \ k1 \ k \ f k1 = f k" + using k12 that by auto + hence "k2 \ k" unfolding k2_def by (rule Least_le) + with that show False by auto + qed + + have snd_f_gt_1: "snd (f k) > 1" if "k < k2 - 1" for k + proof - + have "snd (f k) \ 1" + proof + assume "snd (f k) = 1" + hence "f k = (D', 1)" using red_assoc_denom_1[of "fst (f k)"] *[of k] + by (cases "f k") auto + hence "sqrt_remainder_step (f k) = (D', D - D'\<^sup>2)" by (auto simp: sqrt_remainder_step_def) + hence "f (Suc k) = f 0" by (simp add: f_def) + moreover have "f (Suc k) \ f 0" + using that by (intro smallest_period) auto + ultimately show False by contradiction + qed + moreover have "snd (f k) > 0" using *[of k] by (auto simp: red_assoc_def) + ultimately show ?thesis by simp + qed + + have sqrt_cfrac_le: "sqrt_cfrac_nth k \ D'" if "k < k2 - 1" for k + proof - + define p and q where "p = fst (f k)" and "q = snd (f k)" + have "q \ 2" using snd_f_gt_1[of k] that by (auto simp: q_def) + also have "sqrt_cfrac_nth k * q \ D' * 2" + using red_assoc_step(5)[OF *[of k]] + by (simp add: sqrt_cfrac_nth_def p_def q_def case_prod_unfold f_def) + finally show ?thesis by simp + qed + + have last: "f (k2 - 1) = (D', 1)" + proof - + define p and q where "p = fst (f (k2 - 1))" and "q = snd (f (k2 - 1))" + have pq: "f (k2 - 1) = (p, q)" by (simp add: p_def q_def) + have "sqrt_remainder_step (f (k2 - 1)) = f (Suc (k2 - 1))" + by (simp add: f_def) + also from k12 have "Suc (k2 - 1) = k2" by simp + also have "f k2 = f 0" + using k12 by simp + also have "f 0 = (D', D - D'\<^sup>2)" by (simp add: f_def) + finally have eq: "sqrt_remainder_step (f (k2 - 1)) = (D', D - D'\<^sup>2)" . + + hence "(D - D'\<^sup>2) div q = D - D'\<^sup>2" unfolding sqrt_remainder_step_def Let_def pq + by auto + moreover have "q > 0" using *[of "k2 - 1"] + by (auto simp: red_assoc_def q_def) + ultimately have "q = 1" using D'_sqr_less_D + by (subst (asm) div_eq_dividend_iff) auto + hence "p = D'" + using red_assoc_denom_1[of p] *[of "k2 - 1"] unfolding pq by auto + with \q = 1\ show "f (k2 - 1) = (D', 1)" unfolding pq by simp + qed + + have period: "sqrt_remainder_surd n = sqrt_remainder_surd (n mod k2)" for n + unfolding sqrt_remainder_surd_def using k12 by (intro funpow_cycle) (auto simp: f_def) + have period': "sqrt_cfrac_nth k = sqrt_cfrac_nth (k mod k2)" for k + using period[of k] by (simp add: sqrt_cfrac_nth_def) + + have k2_le: "l \ k2" if "l > 0" "\k. sqrt_cfrac_nth (k + l) = sqrt_cfrac_nth k" for l + proof (rule ccontr) + assume *: "\(l \ k2)" + hence "sqrt_cfrac_nth (k2 - Suc l) = sqrt_cfrac_nth (k2 - 1)" + using that(2)[of "k2 - Suc l"] by simp + also have "\ = 2 * D'" + using last by (simp add: sqrt_cfrac_nth_def f_def) + finally have "2 * D' = sqrt_cfrac_nth (k2 - Suc l)" .. + also have "\ \ D'" using k12 that * + by (intro sqrt_cfrac_le diff_less_mono2) auto + finally show False using D'_pos by simp + qed + + have "l = (LEAST l. 0 < l \ (\n. int (sqrt_cfrac_nth (n + l)) = int (sqrt_cfrac_nth n)))" + using nonsquare unfolding sqrt_cfrac_def + by (simp add: l_def sqrt_nat_period_length_def sqrt_cfrac) + hence l_altdef: "l = (LEAST l. 0 < l \ (\n. sqrt_cfrac_nth (n + l) = sqrt_cfrac_nth n))" + by simp + + have [simp]: "D \ 0" using nonsquare by (auto intro!: Nat.gr0I) + have "\l. l > 0 \ (\k. sqrt_cfrac_nth (k + l) = sqrt_cfrac_nth k)" + proof (rule exI, safe) + fix k show "sqrt_cfrac_nth (k + k2) = sqrt_cfrac_nth k" + using period'[of k] period'[of "k + k2"] k12 by simp + qed (insert k12, auto) + from LeastI_ex[OF this, folded l_altdef] + have l: "l > 0" "\k. sqrt_cfrac_nth (k + l) = sqrt_cfrac_nth k" + by (simp_all add: sqrt_cfrac) + + have "l \ k2" unfolding l_altdef + by (rule Least_le) (subst (1 2) period', insert k12, auto) + moreover have "k2 \ l" using k2_le l by blast + ultimately have [simp]: "l = k2" by auto + + define x' where "x' = (\k. -1 / surd_to_real_cnj (f k))" + { + fix k :: nat + have nz: "surd_to_real_cnj (f k) \ 0" "surd_to_real_cnj (f (Suc k)) \ 0" + using surd_to_real_cnj_nz[OF *, of k] surd_to_real_cnj_nz[OF *, of "Suc k"] + by (simp_all add: f_def) + + have "surd_to_real_cnj (f k) \ sqrt_cfrac_nth k" + using surd_to_real_cnj_irrat[OF *[of k]] by auto + hence "x' (Suc k) = sqrt_cfrac_nth k + 1 / x' k" + using red_assoc_step(3)[OF *[of k]] nz + by (simp add: field_simps sqrt_cfrac_nth_def case_prod_unfold f_def x'_def) + } note x'_Suc = this + + have x'_nz: "x' k \ 0" for k + using surd_to_real_cnj_nz[OF *[of k]] by (auto simp: x'_def) + have x'_0: "x' 0 = real D' + sqrt D" + using red_assoc_begin by (simp add: x'_def f_def) + + define c' where "c' = cfrac (\n. sqrt_cfrac_nth (l - Suc n))" + define c'' where "c'' = cfrac (\n. if n = 0 then 2 * D' else sqrt_cfrac_nth (n - 1))" + have nth_c' [simp]: "cfrac_nth c' n = sqrt_cfrac_nth (l - Suc n)" for n + unfolding c'_def by (subst cfrac_nth_cfrac) (auto simp: is_cfrac_def intro!: sqrt_cfrac_pos) + have nth_c'' [simp]: "cfrac_nth c'' n = (if n = 0 then 2 * D' else sqrt_cfrac_nth (n - 1))" for n + unfolding c''_def by (subst cfrac_nth_cfrac) (auto simp: is_cfrac_def intro!: sqrt_cfrac_pos) + + have "conv' c' n (x' (l - n)) = x' l" if "n \ l" for n + using that + proof (induction n) + case (Suc n) + have "x' l = conv' c' n (x' (l - n))" + using Suc.prems by (intro Suc.IH [symmetric]) auto + also have "l - n = Suc (l - Suc n)" + using Suc.prems by simp + also have "x' \ = cfrac_nth c' n + 1 / x' (l - Suc n)" + by (subst x'_Suc) simp + also have "conv' c' n \ = conv' c' (Suc n) (x' (l - Suc n))" + by (simp add: conv'_Suc_right) + finally show ?case .. + qed simp_all + from this[of l] have conv'_x'_0: "conv' c' l (x' 0) = x' 0" + using k12 by (simp add: x'_def) + + have "cfrac_nth (cfrac_of_real (x' 0)) n = cfrac_nth c'' n" for n + proof (cases n) + case 0 + thus ?thesis by (simp add: x'_0 D'_def) + next + case (Suc n') + have "sqrt D \ \" + using red_assoc_begin(1) red_assoc_begin(2) by auto + hence "cfrac_nth (cfrac_of_real (real D' + sqrt (real D))) (Suc n') = + cfrac_nth (cfrac_of_real (sqrt (real D))) (Suc n')" + by (simp add: cfrac_tl_of_real frac_add_of_nat Ints_add_left_cancel flip: cfrac_nth_tl) + thus ?thesis using x'_nz[of 0] + by (simp add: x'_0 sqrt_cfrac Suc) + qed + + show "sqrt_cfrac_nth (l - n - 2) = sqrt_cfrac_nth n" if "n < l - 1" for n + proof - + have "D > 1" using nonsquare by (cases D) (auto intro!: Nat.gr0I) + hence "D' + sqrt D > 0 + 1" using D'_pos by (intro add_strict_mono) auto + hence "x' 0 > 1" by (auto simp: x'_0) + hence "cfrac_nth c' (Suc n) = cfrac_nth (cfrac_of_real (conv' c' l (x' 0))) (Suc n)" + using \n < l - 1\ using cfrac_of_real_conv' by auto + also have "\ = cfrac_nth (cfrac_of_real (x' 0)) (Suc n)" + by (subst conv'_x'_0) auto + also have "\ = cfrac_nth c'' (Suc n)" by fact + finally show "sqrt_cfrac_nth (l - n - 2) = sqrt_cfrac_nth n" + by simp + qed + + show "l > 0" "l \ D' * (D' + 1)" using k12 by simp_all + show "sqrt_remainder_surd n = sqrt_remainder_surd (n mod l)" + "sqrt_cfrac_nth n = sqrt_cfrac_nth (n mod l)" for n + using period[of n] period'[of n] by simp_all + show "sqrt_remainder_surd n \ sqrt_remainder_surd 0" if "n \ {0<.. 1" if "n < l - 1" for n + using that snd_f_gt_1[of n] by (simp add: f_def) + show "f (l - 1) = (D', 1)" and "sqrt_cfrac_nth (l - 1) = 2 * D'" + using last by (simp_all add: sqrt_cfrac_nth_def f_def) + show "sqrt_cfrac_nth k \ D'" if "k < l - 1" for k + using sqrt_cfrac_le[of k] that by simp + show "l' \ l" if "l' > 0" "\k. sqrt_cfrac_nth (k + l') = sqrt_cfrac_nth k" for l' + using k2_le[of l'] that by auto +qed + +theorem cfrac_sqrt_periodic: + "cfrac_nth (cfrac_of_real (sqrt D)) (Suc n) = + cfrac_nth (cfrac_of_real (sqrt D)) (Suc (n mod l))" + using sqrt_cfrac_periodic[of n] by (metis sqrt_cfrac) + +theorem cfrac_sqrt_le: "n \ {0<.. cfrac_nth (cfrac_of_real (sqrt D)) n \ D'" + using sqrt_cfrac_le[of "n - 1"] + by (metis Suc_less_eq Suc_pred add.right_neutral greaterThanLessThan_iff of_nat_mono + period_nonempty plus_1_eq_Suc sqrt_cfrac) + +theorem cfrac_sqrt_last: "cfrac_nth (cfrac_of_real (sqrt D)) l = 2 * D'" + using sqrt_cfrac_last by (metis One_nat_def Suc_pred period_nonempty sqrt_cfrac) + +theorem cfrac_sqrt_palindrome: + assumes "n \ {0<.. = sqrt_cfrac_nth (n - 1)" + using assms by (subst sqrt_cfrac_palindrome [symmetric]) auto + also have "\ = cfrac_nth (cfrac_of_real (sqrt D)) n" + using assms by (subst sqrt_cfrac) auto + finally show ?thesis . +qed + +lemma sqrt_cfrac_info_palindrome: + assumes "sqrt_cfrac_info D = (a, b, cs)" + shows "rev (butlast cs) = butlast cs" +proof (rule List.nth_equalityI; safe?) + fix i assume "i < length (rev (butlast cs))" + with period_nonempty have "Suc i < length cs" by simp + thus "rev (butlast cs) ! i = butlast cs ! i" + using assms cfrac_sqrt_palindrome[of "Suc i"] period_nonempty unfolding l_def + by (auto simp: sqrt_cfrac_info_def rev_nth algebra_simps Suc_diff_Suc simp del: cfrac.simps) +qed simp_all + +lemma sqrt_cfrac_info_last: + assumes "sqrt_cfrac_info D = (a, b, cs)" + shows "last cs = 2 * Discrete.sqrt D" +proof - + from assms show ?thesis using period_nonempty cfrac_sqrt_last + by (auto simp: sqrt_cfrac_info_def last_map l_def D'_def Discrete_sqrt_altdef) +qed + +text \ + The following lemmas allow us to compute the period of the expansion of the square root: +\ +lemma while_option_sqrt_cfrac: + defines "step' \ (\(as, pq). ((D' + fst pq) div snd pq # as, sqrt_remainder_step pq))" + defines "b \ (\(_, pq). snd pq \ 1)" + defines "initial \ ([] :: nat list, (D', D - D'\<^sup>2))" + shows "while_option b step' initial = + Some (rev (map sqrt_cfrac_nth [0..(as, pq). let n = length as + in n < l \ pq = sqrt_remainder_surd n \ as = rev (map sqrt_cfrac_nth [0.. :: "nat list \ (nat \ nat) \ nat" where "\ = (\(as, _). l - length as)" + have [simp]: "P initial" using period_nonempty + by (auto simp: initial_def P_def sqrt_remainder_surd_def) + have step': "P (step' s) \ Suc (length (fst s)) < l" if "P s" "b s" for s + proof (cases s) + case (fields as p q) + define n where "n = length as" + from that fields sqrt_remainder_surd_last have "Suc n \ l" + by (auto simp: b_def P_def Let_def n_def [symmetric]) + moreover from that fields sqrt_remainder_surd_last have "Suc n \ l" + by (auto simp: b_def P_def Let_def n_def [symmetric]) + ultimately have "Suc n < l" by auto + with that fields sqrt_remainder_surd_last show "P (step' s) \ Suc (length (fst s)) < l" + by (simp add: b_def P_def Let_def n_def step'_def sqrt_cfrac_nth_def + sqrt_remainder_surd_def case_prod_unfold) + qed + have [simp]: "length (fst (step' s)) = Suc (length (fst s))" for s + by (simp add: step'_def case_prod_unfold) + + have "\x. while_option b step' initial = Some x" + proof (rule measure_while_option_Some) + fix s assume *: "P s" "b s" + from step'[OF *] show "P (step' s) \ \ (step' s) < \ s" + by (auto simp: b_def \_def case_prod_unfold intro!: diff_less_mono2) + qed auto + then obtain x where x: "while_option b step' initial = Some x" .. + have "P x" by (rule while_option_rule[OF _ x]) (insert step', auto) + have "\b x" using while_option_stop[OF x] by auto + + obtain as p q where [simp]: "x = (as, (p, q))" by (cases x) + define n where "n = length as" + have [simp]: "q = 1" using \\b x\ by (auto simp: b_def) + have [simp]: "p = D'" using \P x\ + using red_assoc_denom_1[of p] by (auto simp: P_def Let_def) + have "n < l" "sqrt_remainder_surd (length as) = (D', Suc 0)" + and as: "as = rev (map sqrt_cfrac_nth [0..P x\ + by (auto simp: P_def Let_def n_def) + hence "\(n < l - 1)" + using snd_sqrt_remainder_surd_gt_1[of n] by (intro notI) auto + with \n < l\ have [simp]: "n = l - 1" by auto + show ?thesis by (simp add: as x) +qed + +lemma while_option_sqrt_cfrac_info: + defines "step' \ (\(as, pq). ((D' + fst pq) div snd pq # as, sqrt_remainder_step pq))" + defines "b \ (\(_, pq). snd pq \ 1)" + defines "initial \ ([], (D', D - D'\<^sup>2))" + shows "sqrt_cfrac_info D = + (case while_option b step' initial of + Some (as, _) \ (Suc (length as), D', rev ((2 * D') # as)))" +proof - + have "nat (cfrac_nth (cfrac_of_real (sqrt (real D))) (Suc k)) = sqrt_cfrac_nth k" for k + by (metis nat_int sqrt_cfrac) + thus ?thesis unfolding assms while_option_sqrt_cfrac + using period_nonempty sqrt_cfrac_last + by (cases l) (auto simp: sqrt_cfrac_info_def D'_def l_def Discrete_sqrt_altdef) +qed + +end +end + +lemma sqrt_nat_period_length_le: "sqrt_nat_period_length D \ nat \sqrt D\ * (nat \sqrt D\ + 1)" + by (cases "is_square D") (use period_length_le_aux[of D] in auto) + +lemma sqrt_nat_period_length_0_iff [simp]: + "sqrt_nat_period_length D = 0 \ is_square D" + using period_nonempty[of D] by (cases "is_square D") auto + +lemma sqrt_nat_period_length_pos_iff [simp]: + "sqrt_nat_period_length D > 0 \ \is_square D" + using period_nonempty[of D] by (cases "is_square D") auto + +lemma sqrt_cfrac_info_code [code]: + "sqrt_cfrac_info D = + (let D' = Discrete.sqrt D + in if D'\<^sup>2 = D then (0, D', []) + else + case while_option + (\(_, pq). snd pq \ 1) + (\(as, (p, q)). let X = (p + D') div q; p' = X * q - p + in (X # as, p', (D - p'\<^sup>2) div q)) + ([], D', D - D'\<^sup>2) + of Some (as, _) \ (Suc (length as), D', rev ((2 * D') # as)))" +proof - + define D' where "D' = Discrete.sqrt D" + show ?thesis + proof (cases "is_square D") + case True + hence "D' ^ 2 = D" by (auto simp: D'_def elim!: is_nth_powerE) + thus ?thesis using True + by (simp add: D'_def Let_def sqrt_cfrac_info_def sqrt_nat_period_length_def) + next + case False + hence "D' ^ 2 \ D" by (subst eq_commute) auto + thus ?thesis using while_option_sqrt_cfrac_info[OF False] + by (simp add: sqrt_cfrac_info_def D'_def Let_def + case_prod_unfold Discrete_sqrt_altdef add_ac sqrt_remainder_step_def) + qed +qed + +lemma sqrt_nat_period_length_code [code]: + "sqrt_nat_period_length D = fst (sqrt_cfrac_info D)" + by (simp add: sqrt_cfrac_info_def) + +text \ + For efficiency reasons, it is often better to use an array instead of a list: +\ +definition sqrt_cfrac_info_array where + "sqrt_cfrac_info_array D = (case sqrt_cfrac_info D of (a, b, c) \ (a, b, IArray c))" + +lemma fst_sqrt_cfrac_info_array [simp]: "fst (sqrt_cfrac_info_array D) = sqrt_nat_period_length D" + by (simp add: sqrt_cfrac_info_array_def sqrt_cfrac_info_def) + +lemma snd_sqrt_cfrac_info_array [simp]: "fst (snd (sqrt_cfrac_info_array D)) = Discrete.sqrt D" + by (simp add: sqrt_cfrac_info_array_def sqrt_cfrac_info_def) + + +definition cfrac_sqrt_nth :: "nat \ nat \ nat iarray \ nat \ nat" where + "cfrac_sqrt_nth info n = + (case info of (l, a0, as) \ if n = 0 then a0 else as !! ((n - 1) mod l))" + +lemma cfrac_sqrt_nth: + assumes "\is_square D" + shows "cfrac_nth (cfrac_of_real (sqrt D)) n = + int (cfrac_sqrt_nth (sqrt_cfrac_info_array D) n)" (is "?lhs = ?rhs") +proof (cases n) + case (Suc n') + define l where "l = sqrt_nat_period_length D" + from period_nonempty[OF assms] have "l > 0" by (simp add: l_def) + have "cfrac_nth (cfrac_of_real (sqrt D)) (Suc n') = + cfrac_nth (cfrac_of_real (sqrt D)) (Suc (n' mod l))" unfolding l_def + using cfrac_sqrt_periodic[OF assms, of n'] by simp + also have "\ = map (\n. nat (cfrac_nth (cfrac_of_real (sqrt D)) (Suc n))) [0..l > 0\ by (subst nth_map) auto + finally show ?thesis using Suc + by (simp add: sqrt_cfrac_info_array_def sqrt_cfrac_info_def l_def cfrac_sqrt_nth_def) +qed (simp_all add: sqrt_cfrac_info_def sqrt_cfrac_info_array_def + Discrete_sqrt_altdef cfrac_sqrt_nth_def) + +lemma sqrt_cfrac_code [code]: + "sqrt_cfrac D = + (let info = sqrt_cfrac_info_array D; + (l, a0, _) = info + in if l = 0 then cfrac_of_int (int a0) else cfrac (cfrac_sqrt_nth info))" +proof (cases "is_square D") + case True + hence "sqrt (real D) = of_int (Discrete.sqrt D)" + by (auto elim!: is_nth_powerE) + thus ?thesis using True + by (auto simp: Let_def sqrt_cfrac_info_array_def sqrt_cfrac_info_def sqrt_cfrac_def) +next + case False + have "cfrac_sqrt_nth (sqrt_cfrac_info_array D) n > 0" if "n > 0" for n + proof - + have "int (cfrac_sqrt_nth (sqrt_cfrac_info_array D) n) > 0" + using False that by (subst cfrac_sqrt_nth [symmetric]) auto + thus ?thesis by simp + qed + moreover have "sqrt D \ \" + using False irrat_sqrt_nonsquare by blast + ultimately have "sqrt_cfrac D = cfrac (cfrac_sqrt_nth (sqrt_cfrac_info_array D))" + using cfrac_sqrt_nth[OF False] + by (intro cfrac_eqI) (auto simp: sqrt_cfrac_def is_cfrac_def) + thus ?thesis + using False by (simp add: Let_def sqrt_cfrac_info_array_def sqrt_cfrac_info_def) +qed + +text \ + As a test, we determine the continued fraction expansion of $\sqrt{129}$, which is + $[11; \overline{2, 1, 3, 1, 6, 1, 3, 1, 2, 22}]$ (a period length of 10): +\ +value "let info = sqrt_cfrac_info_array 129 in info" +value "sqrt_nat_period_length 129" + +text \ + We can also compute convergents of $\sqrt{129}$ and observe that the difference between + the square of the convergents and 129 vanishes quickly:: +\ +value "map (conv (sqrt_cfrac 129)) [0..<10]" +value "map (\n. \conv (sqrt_cfrac 129) n ^ 2 - 129\) [0..<20]" + +end \ No newline at end of file diff --git a/thys/Continued_Fractions/approximation_cfrac.ML b/thys/Continued_Fractions/approximation_cfrac.ML new file mode 100644 --- /dev/null +++ b/thys/Continued_Fractions/approximation_cfrac.ML @@ -0,0 +1,80 @@ +val _ = \ \Trusting the oracle \@{oracle_name "holds_by_evaluation"} +signature CFRAC_APPROXIMATION_COMPUTATION = sig +val approx_cfrac: Proof.context -> term -> term +end + +structure Cfrac_Approximation_Computation : CFRAC_APPROXIMATION_COMPUTATION = struct + + + + val mk_int = HOLogic.mk_number \<^typ>\int\ o @{code integer_of_int}; + + val term_of_int_list = map mk_int #> HOLogic.mk_list HOLogic.intT + + val approx_cfrac = @{computation "int list" + terms: "0 :: nat" "1 :: nat" "2 :: nat" "3 :: nat" Suc + "(+)::nat\nat\nat" "(-)::nat\nat\nat" "(*)::nat\nat\nat" + "0 :: int" "1 :: int" "2 :: int" "3 :: int" "-1 :: int" + "(+)::int\int\int" "(-)::int\int\int" "(*)::int\int\int" "uminus::int\int" + approx_cfrac + datatypes: int integer "int list" num floatarith float} + (fn _ => fn x => case x of SOME lst => term_of_int_list lst + | NONE => error "Computation approx_cfrac failed.") + +end + + +signature CFRAC_APPROXIMATION = sig + val approx_cfrac : int -> Proof.context -> term -> term +end + + +structure Cfrac_Approximation : CFRAC_APPROXIMATION = struct + +local +open Approximation + +fun mk_approx_cfrac prec t = + \<^const>\approx_cfrac\ $ (HOLogic.mk_number HOLogic.natT prec) $ t + +in + +fun approx_cfrac prec ctxt t = + realify t + |> Thm.cterm_of ctxt + |> (preproc_form_conv ctxt then_conv reify_form_conv ctxt) + |> Thm.prop_of + |> Logic.dest_equals |> snd + |> dest_interpret |> fst + |> mk_approx_cfrac prec + |> Cfrac_Approximation_Computation.approx_cfrac ctxt +end + +fun approximate_cfrac_cmd prec modes raw_t state = + let + val ctxt = Toplevel.context_of state; + val t = + raw_t + |> Syntax.parse_term ctxt + |> Type.constraint \<^typ>\real\ + |> Syntax.check_term ctxt; + val t' = approx_cfrac prec ctxt t; + val ty' = Term.type_of t'; + val ctxt' = Proof_Context.augment t' ctxt; + in + Print_Mode.with_modes modes (fn () => + Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk, + Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) () + end |> Pretty.writeln; + +val opt_modes = + Scan.optional (Args.$$$ "prec" |-- Args.colon |-- Parse.nat) 100 -- + Scan.optional (\<^keyword>\(\ |-- Parse.!!! (Scan.repeat1 Parse.name --| \<^keyword>\)\)) []; + +val _ = + Outer_Syntax.command \<^command_keyword>\approximate_cfrac\ + "print initial fragment of continued fraction of a real number by approximation" + (opt_modes -- Parse.term + >> (fn ((prec, modes), t) => Toplevel.keep (approximate_cfrac_cmd prec modes t))); + +end \ No newline at end of file diff --git a/thys/Continued_Fractions/document/root.bib b/thys/Continued_Fractions/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/Continued_Fractions/document/root.bib @@ -0,0 +1,14 @@ +@misc {proofwiki, + author="{Proof Wiki}", + url="https://proofwiki.org/wiki/Continued_Fraction_Expansion_of_Euler%27s_Number" +} + +@book{khinchin, + title={Continued Fractions}, + author={Khinchin, A.I.A. and Eagle, H.}, + isbn={9780486696300}, + lccn={97008056}, + series={Dover books on mathematics}, + year={1997}, + publisher={Dover Publications} +} \ No newline at end of file diff --git a/thys/Continued_Fractions/document/root.tex b/thys/Continued_Fractions/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/Continued_Fractions/document/root.tex @@ -0,0 +1,54 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} +\usepackage{amsfonts, amsmath, amssymb} + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + + +\begin{document} + +\title{Continued Fractions} +\author{Manuel Eberl} +\maketitle + +\begin{abstract} +This article provides a formalisation of continued fractions of real numbers and their basic properties. +It also contains a proof of the classic result that the irrational numbers with periodic continued + fraction expansions are precisely the quadratic irrationals, i.\,e.\ real numbers that fulfil a non-trivial quadratic equation $a x^2 + b x + c = 0$ with integer coefficients. + +Particular attention is given to the continued fraction expansion of $\sqrt{D}$ for a non-square + natural number $D$. Basic results about the length and structure of its period are provided, + along with an executable algorithm to compute the period (and from it, the entire expansion). + +This is then also used to provide a fairly efficient, executable, and fully formalised algorithm to +compute solutions to Pell's equation $x^2 - D y^2 = 1$. The performance is sufficiently good to +find the solution to Archimedes's cattle problem in less than a second on a typical computer. +This involves the value $D = 410286423278424$, for which the solution has over 200000 decimals. + +Lastly, a derivation of the continued fraction expansions of Euler's number $e$ and an executable +function to compute continued fraction expansions using interval arithmetic is also provided. +\end{abstract} + +\newpage +\tableofcontents +\newpage +\parindent 0pt\parskip 0.5ex + +\input{session} + +\nocite{khinchin} + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,806 +1,807 @@ ABY3_Protocols ADS_Functor AI_Planning_Languages_Semantics AODV AOT AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Ackermanns_not_PR Actuarial_Mathematics Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximate_Model_Counting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BD_Security_Compositional BNF_CC BNF_Operations BTree Balog_Szemeredi_Gowers Banach_Steinhaus Belief_Revision Bell_Numbers_Spivey BenOr_Kozen_Reif Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binary_Code_Imprimitive Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Birkhoff_Finite_Distributive_Lattices Blue_Eyes Bondy Boolean_Expression_Checkers Boolos_Curious_Inference Boolos_Curious_Inference_Automated Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CHERI-C_Memory_Model CISC-Kernel Concentration_Inequalities CRDT CRYSTALS-Kyber CRYSTALS-Kyber_Security CSP_RefTK CVP_Hardness CYK CZH_Elementary_Categories CZH_Foundations CZH_Universal_Constructions CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cardinality_Continuum Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Catoids Cauchy Cayley_Hamilton Certification_Monads Ceva Chandy_Lamport Chebyshev_Polynomials Chord_Segments Circus Clean Clique_and_Monotone_Circuits ClockSynchInst Closest_Pair_Points CoCon CoSMeDis CoSMed CofGroups Coinductive Coinductive_Languages Collections Combinable_Wands Combinatorial_Enumeration_Algorithms Combinatorics_Words Combinatorics_Words_Graph_Lemma Combinatorics_Words_Lyndon CommCSL Commuting_Hermitian Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Bounded_Operators Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Conditional_Simplification Conditional_Transfer_Rule Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers +Continued_Fractions Cook_Levin CoreC++ Core_DOM Core_SC_DOM Correctness_Algebras Cotangent_PFD_Formula Count_Complex_Roots Coupledsim_Contrasim CryptHOL CryptoBasedCompositionalProperties Crypto_Standards CubicalCategories Cubic_Quartic_Equations DCR-ExecutionEquivalence DFS_Framework DOM_Components DPRM_Theorem DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Dedekind_Real Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Design_Theory Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic DigitsInBase Digit_Expansions Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Directed_Sets Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation Disintegration DiskPaxos Distributed_Distinct_Elements Dominance_CHK DynamicArchitectures Dynamic_Tables E_Transcendental Earley_Parser Echelon_Form EdmondsKarp_Maxflow Edwards_Elliptic_Curves_Group Efficient-Mergesort Efficient_Weighted_Path_Order Elimination_Of_Repeated_Factors Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Equivalence_Relation_Enumeration Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Euler_Polyhedron_Formula Eudoxus_Reals Eval_FO Example-Submission Executable_Randomized_Algorithms Expander_Graphs Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Axiomatic FOL_Harrison FOL_Seq_Calc1 FOL_Seq_Calc2 FOL_Seq_Calc3 FO_Theory_Rewriting FSM_Tests Factor_Algebraic_Polynomial Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF Finite_Fields Finitely_Generated_Abelian_Groups First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Fishers_Inequality Fixed_Length_Vector Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Foundation_of_geometry Fourier Free-Boolean-Algebra Free-Groups Frequency_Moments Fresh_Identifiers FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC GaleStewart_Games Gale_Shapley Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic Given_Clause_Loops Go GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Gray_Codes Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Grothendieck_Schemes Group-Ring-Module HOL-CSP HOL-CSPM HOL-CSP_OpSem HOLCF-Prelude HRB-Slicing Hahn_Jordan_Decomposition Hales_Jewett Heard_Of Hello_World HereditarilyFinite Hermite Hermite_Lindemann Hidden_Markov_Models Higher_Order_Terms HoareForDivergence Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL HyperHoareLogic Hyperdual Hypergraph_Basics Hypergraph_Colourings IEEE_Floating_Point IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IMP_Compiler IMP_Compiler_Reuse Interval_Analysis IO_Language_Conformance IP_Addresses Imperative_Insertion_Sort Implicational_Logic Impossible_Geometry IMP_Noninterference Incompleteness Incredible_Proof_Machine Independence_CH Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpolation_Polynomials_HOL_Algebra Interpreter_Optimizations Interval_Arithmetic_Word32 Intro_Dest_Elim Involutions2Squares Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Irrationals_From_THEBOOK IsaGeoCoq IsaNet Isabelle_C Isabelle_hoops Isabelle_Marries_Dirac Isabelle_Meta_Model Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD Karatsuba KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Khovanskii_Theorem Kleene_Algebra Kneser_Cauchy_Davenport Knights_Tour Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt KnuthMorrisPratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LP_Duality LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_Series Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Laws_of_Large_Numbers Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option Lifting_the_Exponent LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Logging_Independent_Anonymity Lovasz_Local Lowe_Ontological_Argument Lower_Semicontinuous Lp Lucas_Theorem MDP-Algorithms MDP-Rewards MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Martingales Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Maximum_Segment_Sum Median_Method Median_Of_Medians_Selection Menger Mereology Mersenne_Primes Metalogic_ProofChecker MHComputation MiniML MiniSail Minimal_SSA Minkowskis_Theorem Minsky_Machines MLSS_Decision_Proc ML_Unification Modal_Logics_for_NTS Modular_Assembly_Kit_Security Modular_arithmetic_LLL_and_HNF_algorithms Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multi_Party_Computation Multirelations Multirelations_Heterogeneous Multiset_Ordering_NPC Multitape_To_Singletape_TM Myhill-Nerode Name_Carrying_Type_Inference Nano_JSON Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers No_FTL_observers_Gen_Rel Nominal2 Nominal_Myhill_Nerode Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Number_Theoretic_Transform Octonions OmegaCatoidsQuantales OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker PAL PAPP_Impossibility PCF PLM POPLmark-deBruijn PSemigroupsConvolution Package_logic Padic_Field Padic_Ints Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect_Fields Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Pluennecke_Ruzsa_Inequality Poincare_Bendixson Poincare_Disc Polygonal_Number_Theorem Polylog Polynomial_Crit_Geometry Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Prefix_Free_Code_Combinators Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Probability_Inequality_Completeness Program-Conflict-Analysis Progress_Tracking Projective_Geometry Projective_Measurements Promela Proof_Strategy_Language PropResPI Propositional_Logic_Class Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem Public_Announcement_Logic Q0_Metatheory Q0_Soundness QBF_Solver_Verification QHLProver QR_Decomposition Quantales Quantales_Converse Quantifier_Elimination_Hybrid Quasi_Borel_Spaces Quaternions Query_Optimization Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Real_Power Real_Time_Deque Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Region_Quadtrees Registers Regression_Test_Selection Regular-Sets Regular_Algebras Regular_Tree_Relations Relation_Algebra Relational-Incorrectness-Logic Relational_Cardinality Relational_Disjoint_Set_Forests Relational_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rensets Rep_Fin_Groups ResiduatedTransitionSystem Residuated_Lattices Resolution_FOL Rewrite_Properties_Reduction Rewriting_Z Ribbon_Proofs Risk_Free_Lending Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Roth_Arithmetic_Progressions Routing Roy_Floyd_Warshall SATSolverVerification SCC_Bloemen_Sequential SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 S_Finite_Measure_Monad Safe_Distance Safe_OCL Safe_Range_RC Saturation_Framework Saturation_Framework_Extensions Sauer_Shelah_Lemma Schutz_Spacetime Schwartz_Zippel Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL Separation_Logic_Unbounded SequentInvertibility Shadow_DOM Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Clause_Learning Simple_Firewall Simplex Simplicial_complexes_and_boolean_functions SimplifiedOntologicalArgument Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Solidity Sophomores_Dream Sort_Encodings Source_Coding_Theorem SpecCheck Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Stalnaker_Logic Standard_Borel_Spaces Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code StrictOmegaCategories Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras Sumcheck_Protocol SumSquares Sunflowers SuperCalc Suppes_Theorem Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Synthetic_Completeness Szemeredi_Regularity Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Three_Circles Three_Squares Timed_Automata Topological_Semantics Topology TortoiseHare TsirelsonBound Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Transitive_Models Transport Treaps Tree-Automata Tree_Decomposition Tree_Enumeration Triangle Trie Turans_Graph_Theorem Twelvefold_Way Two_Generated_Word_Monoids_Intersection Tycon Types_Tableaus_and_Goedels_God Types_To_Sets_Extension UPF UPF_Firewall UTP Undirected_Graph_Theory Universal_Hash_Families Universal_Turing_Machine UpDown_Scheme VYDRA_MDL Valuation Van_Emde_Boas_Trees Van_der_Waerden VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves Virtual_Substitution VolpanoSmith WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Weighted_Arithmetic_Geometric_Mean Weighted_Path_Order Well_Quasi_Orders Wetzels_Problem Wieferich_Kempner Winding_Number_Eval Word_Lib WorkerWrapper X86_Semantics XML Youngs_Inequality ZFC_in_HOL Zeckendorf Zeta_3_Irrational Zeta_Function pGCL Labeled_Transition_Systems Pushdown_Systems