diff --git a/thys/Algebraic_Numbers/Complex_Algebraic_Numbers.thy b/thys/Algebraic_Numbers/Complex_Algebraic_Numbers.thy --- a/thys/Algebraic_Numbers/Complex_Algebraic_Numbers.thy +++ b/thys/Algebraic_Numbers/Complex_Algebraic_Numbers.thy @@ -1,1130 +1,1130 @@ (* Author: Sebastiaan Joosten René Thiemann Akihisa Yamada License: BSD *) section \Complex Algebraic Numbers\ text \Since currently there is no immediate analog of Sturm's theorem for the complex numbers, we implement complex algebraic numbers via their real and imaginary part. The major algorithm in this theory is a factorization algorithm which factors a rational polynomial over the complex numbers. This algorithm is then combined with explicit root algorithms to try to factor arbitrary complex polymials.\ theory Complex_Algebraic_Numbers imports Real_Roots Complex_Roots_Real_Poly Compare_Complex Jordan_Normal_Form.Char_Poly Berlekamp_Zassenhaus.Code_Abort_Gcd Interval_Arithmetic begin subsection \Complex Roots\ abbreviation complex_of_int_poly :: "int poly \ complex poly" where "complex_of_int_poly \ map_poly of_int" abbreviation complex_of_rat_poly :: "rat poly \ complex poly" where "complex_of_rat_poly \ map_poly of_rat" lemma poly_complex_to_real: "(poly (complex_of_int_poly p) (complex_of_real x) = 0) = (poly (real_of_int_poly p) x = 0)" proof - have id: "of_int = complex_of_real o real_of_int" by auto interpret cr: semiring_hom complex_of_real by (unfold_locales, auto) show ?thesis unfolding id by (subst map_poly_map_poly[symmetric], force+) qed lemma represents_cnj: assumes "p represents x" shows "p represents (cnj x)" proof - from assms have p: "p \ 0" and "ipoly p x = 0" by auto hence rt: "poly (complex_of_int_poly p) x = 0" by auto have "poly (complex_of_int_poly p) (cnj x) = 0" by (rule complex_conjugate_root[OF _ rt], subst coeffs_map_poly, auto) with p show ?thesis by auto qed definition poly_2i :: "int poly" where "poly_2i \ [: 4, 0, 1:]" lemma represents_2i: "poly_2i represents (2 * \)" unfolding represents_def poly_2i_def by simp definition root_poly_Re :: "int poly \ int poly" where "root_poly_Re p = cf_pos_poly (poly_mult_rat (inverse 2) (poly_add p p))" lemma root_poly_Re_code[code]: "root_poly_Re p = (let fs = coeffs (poly_add p p); k = length fs in cf_pos_poly (poly_of_list (map (\(fi, i). fi * 2 ^ i) (zip fs [0.. int poly list" where "root_poly_Im p = (let fs = factors_of_int_poly (poly_add p (poly_uminus p)) in remdups ((if (\ f \ set fs. coeff f 0 = 0) then [[:0,1:]] else [])) @ [ cf_pos_poly (poly_div f poly_2i) . f \ fs, coeff f 0 \ 0])" lemma represents_root_poly: assumes "ipoly p x = 0" and p: "p \ 0" shows "(root_poly_Re p) represents (Re x)" and "\ q \ set (root_poly_Im p). q represents (Im x)" proof - let ?Rep = "root_poly_Re p" let ?Imp = "root_poly_Im p" from assms have ap: "p represents x" by auto from represents_cnj[OF this] have apc: "p represents (cnj x)" . from represents_mult_rat[OF _ represents_add[OF ap apc], of "inverse 2"] have "?Rep represents (1 / 2 * (x + cnj x))" unfolding root_poly_Re_def Let_def by (auto simp: hom_distribs) also have "1 / 2 * (x + cnj x) = of_real (Re x)" by (simp add: complex_add_cnj) finally have Rep: "?Rep \ 0" and rt: "ipoly ?Rep (complex_of_real (Re x)) = 0" unfolding represents_def by auto from rt[unfolded poly_complex_to_real] have "ipoly ?Rep (Re x) = 0" . with Rep show "?Rep represents (Re x)" by auto let ?q = "poly_add p (poly_uminus p)" from represents_add[OF ap, of "poly_uminus p" "- cnj x"] represents_uminus[OF apc] have apq: "?q represents (x - cnj x)" by auto from factors_int_poly_represents[OF this] obtain pi where pi: "pi \ set (factors_of_int_poly ?q)" and appi: "pi represents (x - cnj x)" and irr_pi: "irreducible pi" by auto have id: "inverse (2 * \) * (x - cnj x) = of_real (Im x)" apply (cases x) by (simp add: complex_split imaginary_unit.ctr legacy_Complex_simps) from represents_2i have 12: "poly_2i represents (2 * \)" by simp have "\ qi \ set ?Imp. qi represents (inverse (2 * \) * (x - cnj x))" proof (cases "x - cnj x = 0") case False have "poly poly_2i 0 \ 0" unfolding poly_2i_def by auto from represents_div[OF appi 12 this] represents_irr_non_0[OF irr_pi appi False, unfolded poly_0_coeff_0] pi show ?thesis unfolding root_poly_Im_def Let_def by (auto intro: bexI[of _ "cf_pos_poly (poly_div pi poly_2i)"]) next case True hence id2: "Im x = 0" by (simp add: complex_eq_iff) from appi[unfolded True represents_def] have "coeff pi 0 = 0" by (cases pi, auto) with pi have mem: "[:0,1:] \ set ?Imp" unfolding root_poly_Im_def Let_def by auto have "[:0,1:] represents (complex_of_real (Im x))" unfolding id2 represents_def by simp with mem show ?thesis unfolding id by auto qed then obtain qi where qi: "qi \ set ?Imp" "qi \ 0" and rt: "ipoly qi (complex_of_real (Im x)) = 0" unfolding id represents_def by auto from qi rt[unfolded poly_complex_to_real] show "\ qi \ set ?Imp. qi represents (Im x)" by auto qed text \Determine complex roots of a polynomial, intended for polynomials of degree 3 or higher, for lower degree polynomials use @{const roots1} or @{const croots2}\ hide_const (open) eq primrec remdups_gen :: "('a \ 'a \ bool) \ 'a list \ 'a list" where "remdups_gen eq [] = []" | "remdups_gen eq (x # xs) = (if (\ y \ set xs. eq x y) then remdups_gen eq xs else x # remdups_gen eq xs)" lemma real_of_3_remdups_equal_3[simp]: "real_of_3 ` set (remdups_gen equal_3 xs) = real_of_3 ` set xs" by (induct xs, auto simp: equal_3) lemma distinct_remdups_equal_3: "distinct (map real_of_3 (remdups_gen equal_3 xs))" by (induct xs, auto, auto simp: equal_3) lemma real_of_3_code [code]: "real_of_3 x = real_of (Real_Alg_Quotient x)" by (transfer, auto) definition "real_parts_3 p = roots_of_3 (root_poly_Re p)" definition "pos_imaginary_parts_3 p = remdups_gen equal_3 (filter (\ x. sgn_3 x = 1) (concat (map roots_of_3 (root_poly_Im p))))" lemma real_parts_3: assumes p: "p \ 0" and "ipoly p x = 0" shows "Re x \ real_of_3 ` set (real_parts_3 p)" unfolding real_parts_3_def using represents_root_poly(1)[OF assms(2,1)] roots_of_3(1) unfolding represents_def by auto lemma distinct_real_parts_3: "distinct (map real_of_3 (real_parts_3 p))" unfolding real_parts_3_def using roots_of_3(2) . lemma pos_imaginary_parts_3: assumes p: "p \ 0" and "ipoly p x = 0" and "Im x > 0" shows "Im x \ real_of_3 ` set (pos_imaginary_parts_3 p)" proof - from represents_root_poly(2)[OF assms(2,1)] obtain q where q: "q \ set (root_poly_Im p)" "q represents Im x" by auto from roots_of_3(1)[of q] have "Im x \ real_of_3 ` set (roots_of_3 q)" using q unfolding represents_def by auto then obtain i3 where i3: "i3 \ set (roots_of_3 q)" and id: "Im x = real_of_3 i3" by auto from \Im x > 0\ have "sgn (Im x) = 1" by simp hence sgn: "sgn_3 i3 = 1" unfolding id by (metis of_rat_eq_1_iff sgn_3) show ?thesis unfolding pos_imaginary_parts_3_def real_of_3_remdups_equal_3 id using sgn i3 q(1) by auto qed lemma distinct_pos_imaginary_parts_3: "distinct (map real_of_3 (pos_imaginary_parts_3 p))" unfolding pos_imaginary_parts_3_def by (rule distinct_remdups_equal_3) lemma remdups_gen_subset: "set (remdups_gen eq xs) \ set xs" by (induct xs, auto) lemma positive_pos_imaginary_parts_3: assumes "x \ set (pos_imaginary_parts_3 p)" shows "0 < real_of_3 x" proof - from subsetD[OF remdups_gen_subset assms[unfolded pos_imaginary_parts_3_def]] have "sgn_3 x = 1" by auto thus ?thesis using sgn_3[of x] by (simp add: sgn_1_pos) qed definition "pair_to_complex ri \ case ri of (r,i) \ Complex (real_of_3 r) (real_of_3 i)" fun get_itvl_2 :: "real_alg_2 \ real interval" where "get_itvl_2 (Irrational n (p,l,r)) = Interval (of_rat l) (of_rat r)" | "get_itvl_2 (Rational r) = (let rr = of_rat r in Interval rr rr)" lemma get_bounds_2: assumes "invariant_2 x" shows "real_of_2 x \\<^sub>i get_itvl_2 x" proof (cases x) case (Irrational n plr) with assms obtain p l r where plr: "plr = (p,l,r)" by (cases plr, auto) from assms Irrational plr have inv1: "invariant_1 (p,l,r)" and id: "real_of_2 x = real_of_1 (p,l,r)" by auto show ?thesis unfolding id using invariant_1D(1)[OF inv1] by (auto simp: plr Irrational) qed (insert assms, auto simp: Let_def) lift_definition get_itvl_3 :: "real_alg_3 \ real interval" is get_itvl_2 . lemma get_itvl_3: "real_of_3 x \\<^sub>i get_itvl_3 x" by (transfer, insert get_bounds_2, auto) fun tighten_bounds_2 :: "real_alg_2 \ real_alg_2" where "tighten_bounds_2 (Irrational n (p,l,r)) = (case tighten_poly_bounds p l r (sgn (ipoly p r)) of (l',r',_) \ Irrational n (p,l',r'))" | "tighten_bounds_2 (Rational r) = Rational r" lemma tighten_bounds_2: assumes inv: "invariant_2 x" shows "real_of_2 (tighten_bounds_2 x) = real_of_2 x" "invariant_2 (tighten_bounds_2 x)" "get_itvl_2 x = Interval l r \ get_itvl_2 (tighten_bounds_2 x) = Interval l' r' \ r' - l' = (r-l) / 2" proof (atomize(full), cases x) case (Irrational n plr) show "real_of_2 (tighten_bounds_2 x) = real_of_2 x \ invariant_2 (tighten_bounds_2 x) \ (get_itvl_2 x = Interval l r \ get_itvl_2 (tighten_bounds_2 x) = Interval l' r' \ r' - l' = (r - l) / 2)" proof - obtain p l r where plr: "plr = (p,l,r)" by (cases plr, auto) let ?tb = "tighten_poly_bounds p l r (sgn (ipoly p r))" obtain l' r' sr' where tb: "?tb = (l',r',sr')" by (cases ?tb, auto) have id: "tighten_bounds_2 x = Irrational n (p,l',r')" unfolding Irrational plr using tb by auto from inv[unfolded Irrational plr] have inv: "invariant_1_2 (p, l, r)" "n = card {y. y \ real_of_1 (p, l, r) \ ipoly p y = 0}" by auto have rof: "real_of_2 x = real_of_1 (p, l, r)" "real_of_2 (tighten_bounds_2 x) = real_of_1 (p, l', r')" using Irrational plr id by auto from inv have inv1: "invariant_1 (p, l, r)" and "poly_cond2 p" by auto hence rc: "\!x. root_cond (p, l, r) x" "poly_cond2 p" by auto note tb' = tighten_poly_bounds[OF tb rc refl] have eq: "real_of_1 (p, l, r) = real_of_1 (p, l', r')" using tb' inv1 using invariant_1_sub_interval(2) by presburger from inv1 tb' have "invariant_1 (p, l', r')" by (metis invariant_1_sub_interval(1)) hence inv2: "invariant_2 (tighten_bounds_2 x)" unfolding id using inv eq by auto thus ?thesis unfolding rof eq unfolding id unfolding Irrational plr using tb'(1-4) arg_cong[OF tb'(5), of real_of_rat] by (auto simp: hom_distribs) qed qed (auto simp: Let_def) lift_definition tighten_bounds_3 :: "real_alg_3 \ real_alg_3" is tighten_bounds_2 using tighten_bounds_2 by auto lemma tighten_bounds_3: "real_of_3 (tighten_bounds_3 x) = real_of_3 x" "get_itvl_3 x = Interval l r \ get_itvl_3 (tighten_bounds_3 x) = Interval l' r' \ r' - l' = (r-l) / 2" by (transfer, insert tighten_bounds_2, auto)+ partial_function (tailrec) filter_list_length :: "('a \ 'a) \ ('a \ bool) \ nat \ 'a list \ 'a list" where [code]: "filter_list_length f p n xs = (let ys = filter p xs in if length ys = n then ys else filter_list_length f p n (map f ys))" lemma filter_list_length: assumes "length (filter P xs) = n" and "\ i x. x \ set xs \ P x \ p ((f ^^ i) x)" and "\ x. x \ set xs \ \ P x \ \ i. \ p ((f ^^ i) x)" and g: "\ x. g (f x) = g x" and P: "\ x. P (f x) = P x" shows "map g (filter_list_length f p n xs) = map g (filter P xs)" proof - from assms(3) have "\ x. \ i. x \ set xs \ \ P x \ \ p ((f ^^ i) x)" by auto from choice[OF this] obtain i where i: "\ x. x \ set xs \ \ P x \ \ p ((f ^^ (i x)) x)" by auto define m where "m = max_list (map i xs)" have m: "\ x. x \ set xs \ \ P x \ \ i \ m. \ p ((f ^^ i) x)" using max_list[of _ "map i xs", folded m_def] i by auto show ?thesis using assms(1-2) m proof (induct m arbitrary: xs rule: less_induct) case (less m xs) define ys where "ys = filter p xs" have xs_ys: "filter P xs = filter P ys" unfolding ys_def filter_filter by (rule filter_cong[OF refl], insert less(3)[of _ 0], auto) have "filter (P \ f) ys = filter P ys" using P unfolding o_def by auto hence id3: "filter P (map f ys) = map f (filter P ys)" unfolding filter_map by simp hence id2: "map g (filter P (map f ys)) = map g (filter P ys)" by (simp add: g) show ?case proof (cases "length ys = n") case True hence id: "filter_list_length f p n xs = ys" unfolding ys_def filter_list_length.simps[of _ _ _ xs] Let_def by auto show ?thesis using True unfolding id xs_ys using less(2) by (metis filter_id_conv length_filter_less less_le xs_ys) next case False { assume "m = 0" from less(4)[unfolded this] have Pp: "x \ set xs \ \ P x \ \ p x" for x by auto with xs_ys False[folded less(2)] have False by (metis (mono_tags, lifting) filter_True mem_Collect_eq set_filter ys_def) } note m0 = this then obtain M where mM: "m = Suc M" by (cases m, auto) hence m: "M < m" by simp from False have id: "filter_list_length f p n xs = filter_list_length f p n (map f ys)" unfolding ys_def filter_list_length.simps[of _ _ _ xs] Let_def by auto show ?thesis unfolding id xs_ys id2[symmetric] proof (rule less(1)[OF m]) fix y assume "y \ set (map f ys)" then obtain x where x: "x \ set xs" "p x" and y: "y = f x" unfolding ys_def by auto { assume "\ P y" hence "\ P x" unfolding y P . from less(4)[OF x(1) this] obtain i where i: "i \ m" and p: "\ p ((f ^^ i) x)" by auto with x obtain j where ij: "i = Suc j" by (cases i, auto) with i have j: "j \ M" unfolding mM by auto have "\ p ((f ^^ j) y)" using p unfolding ij y funpow_Suc_right by simp thus "\i\ M. \ p ((f ^^ i) y)" using j by auto } { fix i assume "P y" hence "P x" unfolding y P . from less(3)[OF x(1) this, of "Suc i"] show "p ((f ^^ i) y)" unfolding y funpow_Suc_right by simp } next show "length (filter P (map f ys)) = n" unfolding id3 length_map using xs_ys less(2) by auto qed qed qed qed definition complex_roots_of_int_poly3 :: "int poly \ complex list" where "complex_roots_of_int_poly3 p \ let n = degree p; rrts = real_roots_of_int_poly p; nr = length rrts; crts = map (\ r. Complex r 0) rrts in if n = nr then crts else let nr_crts = n - nr in if nr_crts = 2 then let pp = real_of_int_poly p div (prod_list (map (\ x. [:-x,1:]) rrts)); cpp = map_poly (\ r. Complex r 0) pp in crts @ croots2 cpp else let nr_pos_crts = nr_crts div 2; rxs = real_parts_3 p; ixs = pos_imaginary_parts_3 p; rts = [(rx, ix). rx <- rxs, ix <- ixs]; crts' = map pair_to_complex (filter_list_length (map_prod tighten_bounds_3 tighten_bounds_3) (\ (r, i). 0 \\<^sub>c ipoly_complex_interval p (Complex_Interval (get_itvl_3 r) (get_itvl_3 i))) nr_pos_crts rts) in crts @ (concat (map (\ x. [x, cnj x]) crts'))" definition complex_roots_of_int_poly_all :: "int poly \ complex list" where "complex_roots_of_int_poly_all p = (let n = degree p in if n \ 3 then complex_roots_of_int_poly3 p else if n = 1 then [roots1 (map_poly of_int p)] else if n = 2 then croots2 (map_poly of_int p) else [])" lemma in_real_itvl_get_bounds_tighten: "real_of_3 x \\<^sub>i get_itvl_3 ((tighten_bounds_3 ^^ n) x)" proof (induct n arbitrary: x) case 0 thus ?case using get_itvl_3[of x] by simp next case (Suc n x) have id: "(tighten_bounds_3 ^^ (Suc n)) x = (tighten_bounds_3 ^^ n) (tighten_bounds_3 x)" by (metis comp_apply funpow_Suc_right) show ?case unfolding id tighten_bounds_3(1)[of x, symmetric] by (rule Suc) qed lemma sandwitch_real: fixes l r :: "nat \ real" assumes la: "l \ a" and ra: "r \ a" and lm: "\i. l i \ m i" and mr: "\i. m i \ r i" shows "m \ a" proof (rule LIMSEQ_I) fix e :: real assume "0 < e" hence e: "0 < e / 2" by simp from LIMSEQ_D[OF la e] obtain n1 where n1: "\ n. n \ n1 \ norm (l n - a) < e/2" by auto from LIMSEQ_D[OF ra e] obtain n2 where n2: "\ n. n \ n2 \ norm (r n - a) < e/2" by auto show "\no. \n\no. norm (m n - a) < e" proof (rule exI[of _ "max n1 n2"], intro allI impI) fix n assume "max n1 n2 \ n" with n1 n2 have *: "norm (l n - a) < e/2" "norm (r n - a) < e/2" by auto from lm[of n] mr[of n] have "norm (m n - a) \ norm (l n - a) + norm (r n - a)" by simp with * show "norm (m n - a) < e" by auto qed qed lemma real_of_tighten_bounds_many[simp]: "real_of_3 ((tighten_bounds_3 ^^ i) x) = real_of_3 x" apply (induct i) using tighten_bounds_3 by auto definition lower_3 where "lower_3 x i \ interval.lower (get_itvl_3 ((tighten_bounds_3 ^^ i) x))" definition upper_3 where "upper_3 x i \ interval.upper (get_itvl_3 ((tighten_bounds_3 ^^ i) x))" lemma interval_size_3: "upper_3 x i - lower_3 x i = (upper_3 x 0 - lower_3 x 0)/2^i" proof (induct i) case (Suc i) have "upper_3 x (Suc i) - lower_3 x (Suc i) = (upper_3 x i - lower_3 x i) / 2" unfolding upper_3_def lower_3_def using tighten_bounds_3 get_itvl_3 by auto with Suc show ?case by auto qed auto lemma interval_size_3_tendsto_0: "(\i. (upper_3 x i - lower_3 x i)) \ 0" by (subst interval_size_3, auto intro: LIMSEQ_divide_realpow_zero) lemma dist_tendsto_0_imp_tendsto: "(\i. \f i - a\ :: real) \ 0 \ f \ a" using LIM_zero_cancel tendsto_rabs_zero_iff by blast lemma upper_3_tendsto: "upper_3 x \ real_of_3 x" proof(rule dist_tendsto_0_imp_tendsto, rule sandwitch_real) fix i obtain l r where lr: "get_itvl_3 ((tighten_bounds_3 ^^ i) x) = Interval l r" by (metis interval.collapse) with get_itvl_3[of "(tighten_bounds_3 ^^ i) x"] show "\(upper_3 x) i - real_of_3 x\ \ (upper_3 x i - lower_3 x i)" unfolding upper_3_def lower_3_def by auto qed (insert interval_size_3_tendsto_0, auto) lemma lower_3_tendsto: "lower_3 x \ real_of_3 x" proof(rule dist_tendsto_0_imp_tendsto, rule sandwitch_real) fix i obtain l r where lr: "get_itvl_3 ((tighten_bounds_3 ^^ i) x) = Interval l r" by (metis interval.collapse) with get_itvl_3[of "(tighten_bounds_3 ^^ i) x"] show "\lower_3 x i - real_of_3 x\ \ (upper_3 x i - lower_3 x i)" unfolding upper_3_def lower_3_def by auto qed (insert interval_size_3_tendsto_0, auto) lemma tends_to_tight_bounds_3: "(\x. get_itvl_3 ((tighten_bounds_3 ^^ x) y)) \\<^sub>i real_of_3 y" using lower_3_tendsto[of y] upper_3_tendsto[of y] unfolding lower_3_def upper_3_def interval_tendsto_def o_def by auto lemma complex_roots_of_int_poly3: assumes p: "p \ 0" and sf: "square_free p" shows "set (complex_roots_of_int_poly3 p) = {x. ipoly p x = 0}" (is "?l = ?r") "distinct (complex_roots_of_int_poly3 p)" proof - interpret map_poly_inj_idom_hom of_real.. define q where "q = real_of_int_poly p" let ?q = "map_poly complex_of_real q" from p have q0: "q \ 0" unfolding q_def by auto hence q: "?q \ 0" by auto define rr where "rr = real_roots_of_int_poly p" define rrts where "rrts = map (\r. Complex r 0) rr" note d = complex_roots_of_int_poly3_def[of p, unfolded Let_def, folded rr_def, folded rrts_def] have rr: "set rr = {x. ipoly p x = 0}" unfolding rr_def using real_roots_of_int_poly(1)[OF p] . have rrts: "set rrts = {x. poly ?q x = 0 \ x \ \}" unfolding rrts_def set_map rr q_def complex_of_real_def[symmetric] by (auto elim: Reals_cases) have dist: "distinct rr" unfolding rr_def using real_roots_of_int_poly(2) . from dist have dist1: "distinct rrts" unfolding rrts_def distinct_map inj_on_def by auto have lrr: "length rr = card {x. poly (real_of_int_poly p) x = 0}" unfolding rr_def using real_roots_of_int_poly[of p] p distinct_card by fastforce have cr: "length rr = card {x. poly ?q x = 0 \ x \ \}" unfolding lrr q_def[symmetric] proof - have "card {x. poly q x = 0} \ card {x. poly (map_poly complex_of_real q) x = 0 \ x \ \}" (is "?l \ ?r") by (rule card_inj_on_le[of of_real], insert poly_roots_finite[OF q], auto simp: inj_on_def) moreover have "?l \ ?r" by (rule card_inj_on_le[of Re, OF _ _ poly_roots_finite[OF q0]], auto simp: inj_on_def elim!: Reals_cases) ultimately show "?l = ?r" by simp qed have conv: "\ x. ipoly p x = 0 \ poly ?q x = 0" unfolding q_def by (subst map_poly_map_poly, auto simp: o_def) have r: "?r = {x. poly ?q x = 0}" unfolding conv .. have "?l = {x. ipoly p x = 0} \ distinct (complex_roots_of_int_poly3 p)" proof (cases "degree p = length rr") case False note oFalse = this show ?thesis proof (cases "degree p - length rr = 2") case False let ?nr = "(degree p - length rr) div 2" define cpxI where "cpxI = pos_imaginary_parts_3 p" define cpxR where "cpxR = real_parts_3 p" let ?rts = "[(rx,ix). rx <- cpxR, ix <- cpxI]" define cpx where "cpx = map pair_to_complex (filter (\ c. ipoly p (pair_to_complex c) = 0) ?rts)" let ?LL = "cpx @ map cnj cpx" let ?LL' = "concat (map (\ x. [x,cnj x]) cpx)" let ?ll = "rrts @ ?LL" let ?ll' = "rrts @ ?LL'" have cpx: "set cpx \ ?r" unfolding cpx_def by auto have ccpx: "cnj ` set cpx \ ?r" using cpx unfolding r by (auto intro!: complex_conjugate_root[of ?q] simp: Reals_def) have "set ?ll \ ?r" using rrts cpx ccpx unfolding r by auto moreover { fix x :: complex assume rt: "ipoly p x = 0" { fix x assume rt: "ipoly p x = 0" and gt: "Im x > 0" define rx where "rx = Re x" let ?x = "Complex rx (Im x)" have x: "x = ?x" by (cases x, auto simp: rx_def) from rt x have rt': "ipoly p ?x = 0" by auto from real_parts_3[OF p rt, folded rx_def] pos_imaginary_parts_3[OF p rt gt] rt' have "?x \ set cpx" unfolding cpx_def cpxI_def cpxR_def by (force simp: pair_to_complex_def[abs_def]) hence "x \ set cpx" using x by simp } note gt = this have cases: "Im x = 0 \ Im x > 0 \ Im x < 0" by auto from rt have rt': "ipoly p (cnj x) = 0" unfolding conv by (intro complex_conjugate_root[of ?q x], auto simp: Reals_def) { assume "Im x > 0" from gt[OF rt this] have "x \ set ?ll" by auto } moreover { assume "Im x < 0" hence "Im (cnj x) > 0" by simp from gt[OF rt' this] have "cnj (cnj x) \ set ?ll" unfolding set_append set_map by blast hence "x \ set ?ll" by simp } moreover { assume "Im x = 0" hence "x \ \" using complex_is_Real_iff by blast with rt rrts have "x \ set ?ll" unfolding conv by auto } ultimately have "x \ set ?ll" using cases by blast } ultimately have lr: "set ?ll = {x. ipoly p x = 0}" by blast let ?rr = "map real_of_3 cpxR" let ?pi = "map real_of_3 cpxI" have dist2: "distinct ?rr" unfolding cpxR_def by (rule distinct_real_parts_3) have dist3: "distinct ?pi" unfolding cpxI_def by (rule distinct_pos_imaginary_parts_3) have idd: "concat (map (map pair_to_complex) (map (\rx. map (Pair rx) cpxI) cpxR)) = concat (map (\r. map (\ i. Complex (real_of_3 r) (real_of_3 i)) cpxI) cpxR)" unfolding pair_to_complex_def by (auto simp: o_def) have dist4: "distinct cpx" unfolding cpx_def proof (rule distinct_map_filter, unfold map_concat idd, unfold distinct_conv_nth, intro allI impI, goal_cases) case (1 i j) from nth_concat_diff[OF 1, unfolded length_map] dist2[unfolded distinct_conv_nth] dist3[unfolded distinct_conv_nth] show ?case by auto qed have dist5: "distinct (map cnj cpx)" using dist4 unfolding distinct_map by (auto simp: inj_on_def) { fix x :: complex have rrts: "x \ set rrts \ Im x = 0" unfolding rrts_def by auto have cpx: "\ x. x \ set cpx \ Im x > 0" unfolding cpx_def cpxI_def by (auto simp: pair_to_complex_def[abs_def] positive_pos_imaginary_parts_3) have cpx': "x \ cnj ` set cpx \ sgn (Im x) = -1" using cpx by auto have "x \ set rrts \ set cpx \ set rrts \ cnj ` set cpx \ set cpx \ cnj ` set cpx" using rrts cpx[of x] cpx' by auto } note dist6 = this have dist: "distinct ?ll" unfolding distinct_append using dist6 by (auto simp: dist1 dist4 dist5) let ?p = "complex_of_int_poly p" have pp: "?p \ 0" using p by auto from p square_free_of_int_poly[OF sf] square_free_rsquarefree have rsf:"rsquarefree ?p" by auto from dist lr have "length ?ll = card {x. poly ?p x = 0}" by (metis distinct_card) also have "\ = degree p" using rsf unfolding rsquarefree_card_degree[OF pp] by simp finally have deg_len: "degree p = length ?ll" by simp let ?P = "\ c. ipoly p (pair_to_complex c) = 0" let ?itvl = "\ r i. ipoly_complex_interval p (Complex_Interval (get_itvl_3 r) (get_itvl_3 i))" let ?itv = "\ (r,i). ?itvl r i" let ?p = "(\ (r,i). 0 \\<^sub>c (?itvl r i))" let ?tb = tighten_bounds_3 let ?f = "map_prod ?tb ?tb" have filter: "map pair_to_complex (filter_list_length ?f ?p ?nr ?rts) = map pair_to_complex (filter ?P ?rts)" proof (rule filter_list_length) have "length (filter ?P ?rts) = length cpx" unfolding cpx_def by simp also have "\ = ?nr" unfolding deg_len by (simp add: rrts_def) finally show "length (filter ?P ?rts) = ?nr" by auto next fix n x assume x: "?P x" obtain r i where xri: "x = (r,i)" by force have id: "(?f ^^ n) x = ((?tb ^^ n) r, (?tb ^^ n) i)" unfolding xri by (induct n, auto) have px: "pair_to_complex x = Complex (real_of_3 r) (real_of_3 i)" unfolding xri pair_to_complex_def by auto show "?p ((?f ^^ n) x)" unfolding id split by (rule ipoly_complex_interval[of "pair_to_complex x" _ p, unfolded x], unfold px, auto simp: in_complex_interval_def in_real_itvl_get_bounds_tighten) next fix x assume x: "x \ set ?rts" "\ ?P x" let ?x = "pair_to_complex x" obtain r i where xri: "x = (r,i)" by force have id: "(?f ^^ n) x = ((?tb ^^ n) r, (?tb ^^ n) i)" for n unfolding xri by (induct n, auto) have px: "?x = Complex (real_of_3 r) (real_of_3 i)" unfolding xri pair_to_complex_def by auto have cvg: "(\ n. ?itv ((?f ^^ n) x)) \\<^sub>c ipoly p ?x" unfolding id split px proof (rule ipoly_complex_interval_tendsto) show "(\ia. Complex_Interval (get_itvl_3 ((?tb ^^ ia) r)) (get_itvl_3 ((?tb ^^ ia) i))) \\<^sub>c Complex (real_of_3 r) (real_of_3 i)" unfolding complex_interval_tendsto_def by (simp add: tends_to_tight_bounds_3 o_def) qed from complex_interval_tendsto_neq[OF this x(2)] show "\ i. \ ?p ((?f ^^ i) x)" unfolding id by auto next show "pair_to_complex (?f x) = pair_to_complex x" for x by (cases x, auto simp: pair_to_complex_def tighten_bounds_3(1)) next show "?P (?f x) = ?P x" for x by (cases x, auto simp: pair_to_complex_def tighten_bounds_3(1)) qed have l: "complex_roots_of_int_poly3 p = ?ll'" unfolding d filter cpx_def[symmetric] cpxI_def[symmetric] cpxR_def[symmetric] using False oFalse by auto have "distinct ?ll' = (distinct rrts \ distinct ?LL' \ set rrts \ set ?LL' = {})" unfolding distinct_append .. also have "set ?LL' = set ?LL" by auto also have "distinct ?LL' = distinct ?LL" by (induct cpx, auto) finally have "distinct ?ll' = distinct ?ll" unfolding distinct_append by auto with dist have "distinct ?ll'" by auto with lr l show ?thesis by auto next case True let ?cr = "map_poly of_real :: real poly \ complex poly" define pp where "pp = complex_of_int_poly p" have id: "pp = map_poly of_real q" unfolding q_def pp_def by (subst map_poly_map_poly, auto simp: o_def) let ?rts = "map (\ x. [:-x,1:]) rr" define rts where "rts = prod_list ?rts" let ?c2 = "?cr (q div rts)" have pq: "\ x. ipoly p x = 0 \ poly q x = 0" unfolding q_def by simp from True have 2: "degree q - card {x. poly q x = 0} = 2" unfolding pq[symmetric] lrr unfolding q_def by simp from True have id: "degree p = length rr \ False" "degree p - length rr = 2 \ True" by auto have l: "?l = of_real ` {x. poly q x = 0} \ set (croots2 ?c2)" unfolding d rts_def id if_False if_True set_append rrts Reals_def by (fold complex_of_real_def q_def, auto) from dist have len_rr: "length rr = card {x. poly q x = 0}" unfolding rr[unfolded pq, symmetric] by (simp add: distinct_card) have rr': "\ r. r \ set rr \ poly q r = 0" using rr unfolding q_def by simp with dist have "q = q div prod_list ?rts * prod_list ?rts" proof (induct rr arbitrary: q) case (Cons r rr q) note dist = Cons(2) let ?p = "q div [:-r,1:]" from Cons.prems(2) have "poly q r = 0" by simp hence "[:-r,1:] dvd q" using poly_eq_0_iff_dvd by blast from dvd_mult_div_cancel[OF this] have "q = ?p * [:-r,1:]" by simp moreover have "?p = ?p div (\x\rr. [:- x, 1:]) * (\x\rr. [:- x, 1:])" proof (rule Cons.hyps) show "distinct rr" using dist by auto fix s assume "s \ set rr" with dist Cons(3) have "s \ r" "poly q s = 0" by auto hence "poly (?p * [:- 1 * r, 1:]) s = 0" using calculation by force thus "poly ?p s = 0" by (simp add: \s \ r\) qed ultimately have q: "q = ?p div (\x\rr. [:- x, 1:]) * (\x\rr. [:- x, 1:]) * [:-r,1:]" by auto also have "\ = (?p div (\x\rr. [:- x, 1:])) * (\x\r # rr. [:- x, 1:])" unfolding mult.assoc by simp also have "?p div (\x\rr. [:- x, 1:]) = q div (\x\r # rr. [:- x, 1:])" unfolding poly_div_mult_right[symmetric] by simp finally show ?case . qed simp hence q_div: "q = q div rts * rts" unfolding rts_def . from q_div q0 have "q div rts \ 0" "rts \ 0" by auto from degree_mult_eq[OF this] have "degree q = degree (q div rts) + degree rts" using q_div by simp also have "degree rts = length rr" unfolding rts_def by (rule degree_linear_factors) also have "\ = card {x. poly q x = 0}" unfolding len_rr by simp finally have deg2: "degree ?c2 = 2" using 2 by simp note croots2 = croots2[OF deg2, symmetric] have "?q = ?cr (q div rts * rts)" using q_div by simp also have "\ = ?cr rts * ?c2" unfolding hom_distribs by simp finally have q_prod: "?q = ?cr rts * ?c2" . from croots2 l have l: "?l = of_real ` {x. poly q x = 0} \ {x. poly ?c2 x = 0}" by simp from r[unfolded q_prod] have r: "?r = {x. poly (?cr rts) x = 0} \ {x. poly ?c2 x = 0}" by auto also have "?cr rts = (\x\rr. ?cr [:- x, 1:])" by (simp add: rts_def o_def of_real_poly_hom.hom_prod_list) also have "{x. poly \ x = 0} = of_real ` set rr" unfolding poly_prod_list_zero_iff by auto also have "set rr = {x. poly q x = 0}" unfolding rr q_def by simp finally have lr: "?l = ?r" unfolding l by simp show ?thesis proof (intro conjI[OF lr]) from sf have sf: "square_free q" unfolding q_def by (rule square_free_of_int_poly) { interpret field_hom_0' complex_of_real .. from sf have "square_free ?q" unfolding square_free_map_poly . } note sf = this have l: "complex_roots_of_int_poly3 p = rrts @ croots2 ?c2" unfolding d rts_def id if_False if_True set_append rrts q_def complex_of_real_def by auto have dist2: "distinct (croots2 ?c2)" unfolding croots2_def Let_def by auto { fix x assume x: "x \ set (croots2 ?c2)" "x \ set rrts" from x(1)[unfolded croots2] have x1: "poly ?c2 x = 0" by auto from x(2) have x2: "poly (?cr rts) x = 0" unfolding rrts_def rts_def complex_of_real_def[symmetric] by (auto simp: poly_prod_list_zero_iff o_def) from square_free_multD(1)[OF sf[unfolded q_prod], of "[: -x, 1:]"] x1 x2 have False unfolding poly_eq_0_iff_dvd by auto } note dist3 = this show "distinct (complex_roots_of_int_poly3 p)" unfolding l distinct_append by (intro conjI dist1 dist2, insert dist3, auto) qed qed next case True have "card {x. poly ?q x = 0} \ degree ?q" by (rule poly_roots_degree[OF q]) also have "\ = degree p" unfolding q_def by simp also have "\ = card {x. poly ?q x = 0 \ x \ \}" using True cr by simp finally have le: "card {x. poly ?q x = 0} \ card {x. poly ?q x = 0 \ x \ \}" by auto have "{x. poly ?q x = 0 \ x \ \} = {x. poly ?q x = 0}" by (rule card_seteq[OF _ _ le], insert poly_roots_finite[OF q], auto) with True rrts dist1 show ?thesis unfolding r d by auto qed thus "distinct (complex_roots_of_int_poly3 p)" "?l = ?r" by auto qed lemma complex_roots_of_int_poly_all: assumes sf: "degree p \ 3 \ square_free p" shows "p \ 0 \ set (complex_roots_of_int_poly_all p) = {x. ipoly p x = 0}" (is "_ \ set ?l = ?r") and "distinct (complex_roots_of_int_poly_all p)" proof - note d = complex_roots_of_int_poly_all_def Let_def have "(p \ 0 \ set ?l = ?r) \ (distinct (complex_roots_of_int_poly_all p))" proof (cases "degree p \ 3") case True hence p: "p \ 0" by auto from True complex_roots_of_int_poly3[OF p] sf show ?thesis unfolding d by auto next case False let ?p = "map_poly (of_int :: int \ complex) p" have deg: "degree ?p = degree p" by (simp add: degree_map_poly) show ?thesis proof (cases "degree p = 1") case True hence l: "?l = [roots1 ?p]" unfolding d by auto from True have "degree ?p = 1" unfolding deg by auto from roots1[OF this] show ?thesis unfolding l roots1_def by auto next case False show ?thesis proof (cases "degree p = 2") case True hence l: "?l = croots2 ?p" unfolding d by auto from True have "degree ?p = 2" unfolding deg by auto from croots2[OF this] show ?thesis unfolding l by (simp add: croots2_def Let_def) next case False with \degree p \ 1\ \degree p \ 2\ \\ (degree p \ 3)\ have True: "degree p = 0" by auto hence l: "?l = []" unfolding d by auto from True have "degree ?p = 0" unfolding deg by auto from roots0[OF _ this] show ?thesis unfolding l by simp qed qed qed thus "p \ 0 \ set ?l = ?r" "distinct (complex_roots_of_int_poly_all p)" by auto qed text \It now comes the preferred function to compute complex roots of a integer polynomial.\ definition complex_roots_of_int_poly :: "int poly \ complex list" where "complex_roots_of_int_poly p = ( let ps = (if degree p \ 3 then factors_of_int_poly p else [p]) in concat (map complex_roots_of_int_poly_all ps))" definition complex_roots_of_rat_poly :: "rat poly \ complex list" where "complex_roots_of_rat_poly p = complex_roots_of_int_poly (snd (rat_to_int_poly p))" lemma complex_roots_of_int_poly: shows "p \ 0 \ set (complex_roots_of_int_poly p) = {x. ipoly p x = 0}" (is "_ \ ?l = ?r") and "distinct (complex_roots_of_int_poly p)" proof - have "(p \ 0 \ ?l = ?r) \ (distinct (complex_roots_of_int_poly p))" proof (cases "degree p \ 3") case False hence "complex_roots_of_int_poly p = complex_roots_of_int_poly_all p" unfolding complex_roots_of_int_poly_def Let_def by auto with complex_roots_of_int_poly_all[of p] False show ?thesis by auto next case True { fix q assume "q \ set (factors_of_int_poly p)" from factors_of_int_poly(1)[OF refl this] irreducible_imp_square_free[of q] have 0: "q \ 0" and sf: "square_free q" by auto from complex_roots_of_int_poly_all(1)[OF sf 0] complex_roots_of_int_poly_all(2)[OF sf] have "set (complex_roots_of_int_poly_all q) = {x. ipoly q x = 0}" "distinct (complex_roots_of_int_poly_all q)" by auto } note all = this from True have "?l = (\ ((\ p. set (complex_roots_of_int_poly_all p)) ` set (factors_of_int_poly p)))" unfolding complex_roots_of_int_poly_def Let_def by auto also have "\ = (\ ((\ p. {x. ipoly p x = 0}) ` set (factors_of_int_poly p)))" using all by blast finally have l: "?l = (\ ((\ p. {x. ipoly p x = 0}) ` set (factors_of_int_poly p)))" . have lr: "p \ 0 \ ?l = ?r" using l factors_of_int_poly(2)[OF refl, of p] by auto show ?thesis proof (rule conjI[OF lr]) from True have id: "complex_roots_of_int_poly p = concat (map complex_roots_of_int_poly_all (factors_of_int_poly p))" unfolding complex_roots_of_int_poly_def Let_def by auto show "distinct (complex_roots_of_int_poly p)" unfolding id distinct_conv_nth proof (intro allI impI, goal_cases) case (1 i j) let ?fp = "factors_of_int_poly p" let ?rr = "complex_roots_of_int_poly_all" let ?cc = "concat (map ?rr (factors_of_int_poly p))" from nth_concat_diff[OF 1, unfolded length_map] obtain j1 k1 j2 k2 where *: "(j1,k1) \ (j2,k2)" "j1 < length ?fp" "j2 < length ?fp" and "k1 < length (map ?rr ?fp ! j1)" "k2 < length (map ?rr ?fp ! j2)" "?cc ! i = map ?rr ?fp ! j1 ! k1" "?cc ! j = map ?rr ?fp ! j2 ! k2" by blast hence **: "k1 < length (?rr (?fp ! j1))" "k2 < length (?rr (?fp ! j2))" "?cc ! i = ?rr (?fp ! j1) ! k1" "?cc ! j = ?rr (?fp ! j2) ! k2" by auto from * have mem: "?fp ! j1 \ set ?fp" "?fp ! j2 \ set ?fp" by auto show "?cc ! i \ ?cc ! j" proof (cases "j1 = j2") case True with * have "k1 \ k2" by auto with all(2)[OF mem(2)] **(1-2) show ?thesis unfolding **(3-4) unfolding True distinct_conv_nth by auto next case False from \degree p \ 3\ have p: "p \ 0" by auto note fip = factors_of_int_poly(2-3)[OF refl this] show ?thesis unfolding **(3-4) proof define x where "x = ?rr (?fp ! j2) ! k2" assume id: "?rr (?fp ! j1) ! k1 = ?rr (?fp ! j2) ! k2" from ** have x1: "x \ set (?rr (?fp ! j1))" unfolding x_def id[symmetric] by auto from ** have x2: "x \ set (?rr (?fp ! j2))" unfolding x_def by auto from all(1)[OF mem(1)] x1 have x1: "ipoly (?fp ! j1) x = 0" by auto from all(1)[OF mem(2)] x2 have x2: "ipoly (?fp ! j2) x = 0" by auto from False factors_of_int_poly(4)[OF refl, of p] have neq: "?fp ! j1 \ ?fp ! j2" using * unfolding distinct_conv_nth by auto have "poly (complex_of_int_poly p) x = 0" by (meson fip(1) mem(2) x2) from fip(2)[OF this] mem x1 x2 neq show False by blast qed qed qed qed qed thus "p \ 0 \ ?l = ?r" "distinct (complex_roots_of_int_poly p)" by auto qed lemma complex_roots_of_rat_poly: "p \ 0 \ set (complex_roots_of_rat_poly p) = {x. rpoly p x = 0}" (is "_ \ ?l = ?r") "distinct (complex_roots_of_rat_poly p)" proof - obtain c q where cq: "rat_to_int_poly p = (c,q)" by force from rat_to_int_poly[OF this] have pq: "p = smult (inverse (of_int c)) (of_int_poly q)" and c: "c \ 0" by auto show "distinct (complex_roots_of_rat_poly p)" unfolding complex_roots_of_rat_poly_def using complex_roots_of_int_poly(2) . assume p: "p \ 0" with pq c have q: "q \ 0" by auto have id: "{x. rpoly p x = (0 :: complex)} = {x. ipoly q x = 0}" - unfolding pq by (simp add: c of_rat_of_int_poly map_poly_map_poly o_def hom_distribs) + unfolding pq by (simp add: c of_rat_of_int_poly hom_distribs) show "?l = ?r" unfolding complex_roots_of_rat_poly_def cq snd_conv id complex_roots_of_int_poly(1)[OF q] .. qed definition roots_of_complex_main :: "complex poly \ complex list" where "roots_of_complex_main p \ let n = degree p in if n = 0 then [] else if n = 1 then [roots1 p] else if n = 2 then croots2 p else (complex_roots_of_rat_poly (map_poly to_rat p))" definition roots_of_complex_poly :: "complex poly \ complex list option" where "roots_of_complex_poly p \ let (c,pis) = yun_factorization gcd p in if (c \ 0 \ (\ (p,i) \ set pis. degree p \ 2 \ (\ x \ set (coeffs p). x \ \))) then Some (concat (map (roots_of_complex_main o fst) pis)) else None" lemma roots_of_complex_main: assumes p: "p \ 0" and deg: "degree p \ 2 \ set (coeffs p) \ \" shows "set (roots_of_complex_main p) = {x. poly p x = 0}" (is "set ?l = ?r") and "distinct (roots_of_complex_main p)" proof - note d = roots_of_complex_main_def Let_def have "set ?l = ?r \ distinct (roots_of_complex_main p)" proof (cases "degree p = 0") case True hence "?l = []" unfolding d by auto with roots0[OF p True] show ?thesis by auto next case False note 0 = this show ?thesis proof (cases "degree p = 1") case True hence "?l = [roots1 p]" unfolding d by auto with roots1[OF True] show ?thesis by auto next case False note 1 = this show ?thesis proof (cases "degree p = 2") case True hence "?l = croots2 p" unfolding d by auto with croots2[OF True] show ?thesis by (auto simp: croots2_def Let_def) next case False note 2 = this let ?q = "map_poly to_rat p" from 0 1 2 have l: "?l = complex_roots_of_rat_poly ?q" unfolding d by auto from deg 0 1 2 have rat: "set (coeffs p) \ \" by auto have "p = map_poly (of_rat o to_rat) p" by (rule sym, rule map_poly_idI, insert rat, auto) also have "\ = complex_of_rat_poly ?q" by (subst map_poly_map_poly, auto simp: to_rat) finally have id: "{x. poly p x = 0} = {x. poly (complex_of_rat_poly ?q) x = 0}" and q: "?q \ 0" using p by auto from complex_roots_of_rat_poly[of ?q, folded id l] q show ?thesis by auto qed qed qed thus "set ?l = ?r" "distinct ?l" by auto qed lemma roots_of_complex_poly: assumes rt: "roots_of_complex_poly p = Some xs" shows "set xs = {x. poly p x = 0}" proof - obtain c pis where yun: "yun_factorization gcd p = (c,pis)" by force from rt[unfolded roots_of_complex_poly_def yun split Let_def] have c: "c \ 0" and pis: "\ p i. (p, i)\ set pis \ degree p \ 2 \ (\x\set (coeffs p). x \ \)" and xs: "xs = concat (map (roots_of_complex_main \ fst) pis)" by (auto split: if_splits) note yun = square_free_factorizationD(1,2,4)[OF yun_factorization(1)[OF yun]] from yun(1) have p: "p = smult c (\(a, i)\set pis. a ^ Suc i)" . have "{x. poly p x = 0} = {x. poly (\(a, i)\set pis. a ^ Suc i) x = 0}" unfolding p using c by auto also have "\ = \ ((\ p. {x. poly p x = 0}) ` fst ` set pis)" (is "_ = ?r") by (subst poly_prod_0, force+) finally have r: "{x. poly p x = 0} = ?r" . { fix p i assume p: "(p,i) \ set pis" have "set (roots_of_complex_main p) = {x. poly p x = 0}" by (rule roots_of_complex_main, insert yun(2)[OF p] pis[OF p], auto) } note main = this have "set xs = \ ((\ (p, i). set (roots_of_complex_main p)) ` set pis)" unfolding xs o_def by auto also have "\ = ?r" using main by auto finally show ?thesis unfolding r by simp qed subsection \Factorization of Complex Polynomials\ definition factorize_complex_main :: "complex poly \ (complex \ (complex \ nat) list) option" where "factorize_complex_main p \ let (c,pis) = yun_factorization gcd p in if ((\ (p,i) \ set pis. degree p \ 2 \ (\ x \ set (coeffs p). x \ \))) then Some (c, concat (map (\ (p,i). map (\ r. (r,i)) (roots_of_complex_main p)) pis)) else None" definition factorize_complex_poly :: "complex poly \ (complex \ (complex poly \ nat) list) option" where "factorize_complex_poly p \ map_option (\ (c,ris). (c, map (\ (r,i). ([:-r,1:],Suc i)) ris)) (factorize_complex_main p)" lemma factorize_complex_main: assumes rt: "factorize_complex_main p = Some (c,xis)" shows "p = smult c (\(x, i)\xis. [:- x, 1:] ^ Suc i)" proof - obtain d pis where yun: "yun_factorization gcd p = (d,pis)" by force from rt[unfolded factorize_complex_main_def yun split Let_def] have pis: "\ p i. (p, i)\set pis \ degree p \ 2 \ (\x\set (coeffs p). x \ \)" and xis: "xis = concat (map (\(p, i). map (\r. (r, i)) (roots_of_complex_main p)) pis)" and d: "d = c" by (auto split: if_splits) note yun = yun_factorization[OF yun[unfolded d]] note yun = square_free_factorizationD[OF yun(1)] yun(2)[unfolded snd_conv] let ?exp = "\ pis. \(x, i)\concat (map (\(p, i). map (\r. (r, i)) (roots_of_complex_main p)) pis). [:- x, 1:] ^ Suc i" from yun(1) have p: "p = smult c (\(a, i)\set pis. a ^ Suc i)" . also have "(\(a, i)\set pis. a ^ Suc i) = (\(a, i)\pis. a ^ Suc i)" by (rule prod.distinct_set_conv_list[OF yun(5)]) also have "\ = ?exp pis" using pis yun(2,6) proof (induct pis) case (Cons pi pis) obtain p i where pi: "pi = (p,i)" by force let ?rts = "roots_of_complex_main p" note Cons = Cons[unfolded pi] have IH: "(\(a, i)\pis. a ^ Suc i) = (?exp pis)" by (rule Cons(1)[OF Cons(2-4)], auto) from Cons(2-4)[of p i] have deg: "degree p \ 2 \ (\x\set (coeffs p). x \ \)" and p: "square_free p" "degree p \ 0" "p \ 0" "monic p" by auto have "(\(a, i)\(pi # pis). a ^ Suc i) = p ^ Suc i * (\(a, i)\pis. a ^ Suc i)" unfolding pi by simp also have "(\(a, i)\pis. a ^ Suc i) = ?exp pis" by (rule IH) finally have id: "(\(a, i)\(pi # pis). a ^ Suc i) = p ^ Suc i * ?exp pis" by simp have "?exp (pi # pis) = ?exp [(p,i)] * ?exp pis" unfolding pi by simp also have "?exp [(p,i)] = (\(x, i)\ (map (\r. (r, i)) ?rts). [:- x, 1:] ^ Suc i)" by simp also have "\ = (\ x \ ?rts. [:- x, 1:])^Suc i" unfolding prod_list_power by (rule arg_cong[of _ _ prod_list], auto) also have "(\ x \ ?rts. [:- x, 1:]) = p" proof - from fundamental_theorem_algebra_factorized[of p, unfolded \monic p\] obtain as where as: "p = (\a\as. [:- a, 1:])" by auto also have "\ = (\a\set as. [:- a, 1:])" proof (rule sym, rule prod.distinct_set_conv_list, rule ccontr) assume "\ distinct as" from not_distinct_decomp[OF this] obtain as1 as2 as3 a where a: "as = as1 @ [a] @ as2 @ [a] @ as3" by blast define q where "q = (\a\as1 @ as2 @ as3. [:- a, 1:])" have "p = (\a\as. [:- a, 1:])" by fact also have "\ = (\a\([a] @ [a]). [:- a, 1:]) * q" unfolding q_def a map_append prod_list.append by (simp only: ac_simps) also have "\ = [:-a,1:] * [:-a,1:] * q" by simp finally have "p = ([:-a,1:] * [:-a,1:]) * q" by simp hence "[:-a,1:] * [:-a,1:] dvd p" unfolding dvd_def .. with \square_free p\[unfolded square_free_def, THEN conjunct2, rule_format, of "[:-a,1:]"] show False by auto qed also have "set as = {x. poly p x = 0}" unfolding as poly_prod_list by (simp add: o_def, induct as, auto) also have "\ = set ?rts" by (rule roots_of_complex_main(1)[symmetric], insert p deg, auto) also have "(\a\set ?rts. [:- a, 1:]) = (\a\?rts. [:- a, 1:])" by (rule prod.distinct_set_conv_list[OF roots_of_complex_main(2)], insert deg p, auto) finally show ?thesis by simp qed finally have id2: "?exp (pi # pis) = p ^ Suc i * ?exp pis" by simp show ?case unfolding id id2 .. qed simp also have "?exp pis = (\(x, i)\xis. [:- x, 1:] ^ Suc i)" unfolding xis .. finally show ?thesis unfolding p xis by simp qed lemma distinct_factorize_complex_main: assumes "factorize_complex_main p = Some fctrs" shows "distinct (map fst (snd fctrs))" proof - from assms have solvable: "\x\set (snd (yun_factorization gcd p)). degree (fst x) \ 2 \ (\x\set (coeffs (fst x)). x \ \)" by (auto simp add: factorize_complex_main_def case_prod_unfold Let_def map_concat o_def split: if_splits) have sqf: "square_free_factorization p (fst (yun_factorization gcd p), snd (yun_factorization gcd p))" by (rule yun_factorization) simp have "map fst (snd fctrs) = concat (map (\x. (roots_of_complex_main (fst x))) (snd (yun_factorization gcd p)))" using assms by (auto simp add: factorize_complex_main_def case_prod_unfold Let_def map_concat o_def split: if_splits) also have "distinct \" proof (rule distinct_concat, goal_cases) case 1 show ?case proof (subst distinct_map, safe) from square_free_factorizationD(5)[OF sqf] show "distinct (snd (yun_factorization gcd p))" . show "inj_on (\x. (roots_of_complex_main (fst x))) (set (snd (yun_factorization gcd p)))" proof (rule inj_onI, clarify, goal_cases) case (1 a1 b1 a2 b2) { assume neq: "(a1, b1) \ (a2, b2)" from 1(1,2)[THEN square_free_factorizationD(2)[OF sqf]] have "degree a1 \ 0" "degree a2 \ 0" by blast+ hence [simp]: "a1 \ 0" "a2 \ 0" by auto from square_free_factorizationD(3)[OF sqf 1(1,2) neq] have "coprime a1 a2" by simp from solvable 1(1) have "{z. poly a1 z = 0} = set (roots_of_complex_main a1)" by (intro roots_of_complex_main(1) [symmetric]) auto also have "set (roots_of_complex_main a1) = set (roots_of_complex_main a2)" using 1(3) by (subst (1 2) set_remdups [symmetric]) (simp only: fst_conv) also from solvable 1(2) have "\ = {z. poly a2 z = 0}" by (intro roots_of_complex_main) auto finally have "{z. poly a1 z = 0} = {z. poly a2 z = 0}" . with coprime_imp_no_common_roots \coprime a1 a2\ have "{z. poly a1 z = 0} = {}" by auto with fundamental_theorem_of_algebra constant_degree have "degree a1 = 0" by auto with \degree a1 \ 0\ have False by contradiction } thus ?case by blast qed qed next case (2 ys) then obtain f b where fb: "(f, b) \ set (snd (yun_factorization gcd p))" and ys: "ys = roots_of_complex_main f" by auto from square_free_factorizationD(2)[OF sqf fb] have 0: "f \ 0" by auto from solvable[rule_format, OF fb] have f: "degree f \ 2 \ (set (coeffs f) \ \)" by auto show ?case unfolding ys by (rule roots_of_complex_main[OF 0 f]) next case (3 ys zs) then obtain a1 b1 a2 b2 where ab: "(a1, b1) \ set (snd (yun_factorization gcd p))" "(a2, b2) \ set (snd (yun_factorization gcd p))" "ys = roots_of_complex_main a1" "zs = roots_of_complex_main a2" by auto with 3 have neq: "(a1,b1) \ (a2,b2)" by auto from ab(1,2)[THEN square_free_factorizationD(2)[OF sqf]] have [simp]: "a1 \ 0" "a2 \ 0" by auto from square_free_factorizationD(3)[OF sqf ab(1,2) neq] have "coprime a1 a2" by simp have "set ys = {z. poly a1 z = 0}" "set zs = {z. poly a2 z = 0}" by (insert solvable ab(1,2), subst ab, rule roots_of_complex_main; (auto) [])+ with coprime_imp_no_common_roots \coprime a1 a2\ show ?case by auto qed finally show ?thesis . qed lemma factorize_complex_poly: assumes fp: "factorize_complex_poly p = Some (c,qis)" shows "p = smult c (\(q, i)\qis. q ^ i)" "(q,i) \ set qis \ irreducible q \ i \ 0 \ monic q \ degree q = 1" proof - from fp[unfolded factorize_complex_poly_def] obtain pis where fp: "factorize_complex_main p = Some (c,pis)" and qis: "qis = map (\(r, i). ([:- r, 1:], Suc i)) pis" by auto from factorize_complex_main[OF fp] have p: "p = smult c (\(x, i)\pis. [:- x, 1:] ^ Suc i)" . show "p = smult c (\(q, i)\qis. q ^ i)" unfolding p qis by (rule arg_cong[of _ _ "\ p. smult c (prod_list p)"], auto) show "(q,i) \ set qis \ irreducible q \ i \ 0 \ monic q \ degree q = 1" using linear_irreducible_field[of q] unfolding qis by auto qed end diff --git a/thys/Algebraic_Numbers/Real_Algebraic_Numbers.thy b/thys/Algebraic_Numbers/Real_Algebraic_Numbers.thy --- a/thys/Algebraic_Numbers/Real_Algebraic_Numbers.thy +++ b/thys/Algebraic_Numbers/Real_Algebraic_Numbers.thy @@ -1,3538 +1,3540 @@ (* Author: Sebastiaan Joosten René Thiemann Akihisa Yamada License: BSD *) section \Real Algebraic Numbers\ text \Whereas we previously only proved the closure properties of algebraic numbers, this theory adds the numeric computations that are required to separate the roots, and to pick unique representatives of algebraic numbers. The development is split into three major parts. First, an ambiguous representation of algebraic numbers is used, afterwards another layer is used with special treatment of rational numbers which still does not admit unique representatives, and finally, a quotient type is created modulo the equivalence. The theory also contains a code-setup to implement real numbers via real algebraic numbers.\ text \The results are taken from the textbook \cite[pages 329ff]{AlgNumbers}.\ theory Real_Algebraic_Numbers imports "Abstract-Rewriting.SN_Order_Carrier" Deriving.Compare_Rat Deriving.Compare_Real Jordan_Normal_Form.Gauss_Jordan_IArray_Impl Algebraic_Numbers Sturm_Rat begin (*TODO: move *) lemma ex1_imp_Collect_singleton: "(\!x. P x) \ P x \ Collect P = {x}" proof(intro iffI conjI, unfold conj_imp_eq_imp_imp) assume "Ex1 P" "P x" then show "Collect P = {x}" by blast next assume Px: "Collect P = {x}" then have "P y \ x = y" for y by auto then show "Ex1 P" by auto from Px show "P x" by auto qed lemma ex1_Collect_singleton[consumes 2]: assumes "\!x. P x" and "P x" and "Collect P = {x} \ thesis" shows thesis by (rule assms(3), subst ex1_imp_Collect_singleton[symmetric], insert assms(1,2), auto) lemma ex1_iff_Collect_singleton: "P x \ (\!x. P x) \ Collect P = {x}" by (subst ex1_imp_Collect_singleton[symmetric], auto) context fixes f assumes bij: "bij f" begin lemma bij_imp_ex1_iff: "(\!x. P (f x)) \ (\!y. P y)" (is "?l = ?r") proof (intro iffI) assume l: ?l then obtain x where "P (f x)" by auto with l have *: "{x} = Collect (P o f)" by auto also have "f ` \ = {y. P (f (Hilbert_Choice.inv f y))}" using bij_image_Collect_eq[OF bij] by auto also have "\ = {y. P y}" proof- have "f (Hilbert_Choice.inv f y) = y" for y by (meson bij bij_inv_eq_iff) then show ?thesis by simp qed finally have "Collect P = {f x}" by auto then show ?r by (fold ex1_imp_Collect_singleton, auto) next assume r: ?r then obtain y where "P y" by auto with r have "{y} = Collect P" by auto also have "Hilbert_Choice.inv f ` \ = Collect (P \ f)" using bij_image_Collect_eq[OF bij_imp_bij_inv[OF bij]] bij by (auto simp: inv_inv_eq) finally have "Collect (P o f) = {Hilbert_Choice.inv f y}" by (simp add: o_def) then show ?l by (fold ex1_imp_Collect_singleton, auto) qed lemma bij_ex1_imp_the_shift: assumes ex1: "\!y. P y" shows "(THE x. P (f x)) = Hilbert_Choice.inv f (THE y. P y)" (is "?l = ?r") proof- from ex1 have "P (THE y. P y)" by (rule the1I2) moreover from ex1[folded bij_imp_ex1_iff] have "P (f (THE x. P (f x)))" by (rule the1I2) ultimately have "(THE y. P y) = f (THE x. P (f x))" using ex1 by auto also have "Hilbert_Choice.inv f \ = (THE x. P (f x))" using bij by (simp add: bij_is_inj) finally show "?l = ?r" by auto qed lemma bij_imp_Collect_image: "{x. P (f x)} = Hilbert_Choice.inv f ` {y. P y}" (is "?l = ?g ` _") proof- have "?l = ?g ` f ` ?l" by (simp add: image_comp inv_o_cancel[OF bij_is_inj[OF bij]]) also have "f ` ?l = {f x | x. P (f x)}" by auto also have "\ = {y. P y}" by (metis bij bij_iff) finally show ?thesis. qed lemma bij_imp_card_image: "card (f ` X) = card X" by (metis bij bij_iff card.infinite finite_imageD inj_onI inj_on_iff_eq_card) end lemma bij_imp_card: assumes bij: "bij f" shows "card {x. P (f x)} = card {x. P x}" unfolding bij_imp_Collect_image[OF bij] bij_imp_card_image[OF bij_imp_bij_inv[OF bij]].. lemma bij_add: "bij (\x. x + y :: 'a :: group_add)" (is ?g1) and bij_minus: "bij (\x. x - y :: 'a)" (is ?g2) and inv_add[simp]: "Hilbert_Choice.inv (\x. x + y) = (\x. x - y)" (is ?g3) and inv_minus[simp]: "Hilbert_Choice.inv (\x. x - y) = (\x. x + y)" (is ?g4) proof- have 1: "(\x. x - y) \ (\x. x + y) = id" and 2: "(\x. x + y) \ (\x. x - y) = id" by auto from o_bij[OF 1 2] show ?g1. from o_bij[OF 2 1] show ?g2. from inv_unique_comp[OF 2 1] show ?g3. from inv_unique_comp[OF 1 2] show ?g4. qed lemmas ex1_shift[simp] = bij_imp_ex1_iff[OF bij_add] bij_imp_ex1_iff[OF bij_minus] lemma ex1_the_shift: assumes ex1: "\!y :: 'a :: group_add. P y" shows "(THE x. P (x + d)) = (THE y. P y) - d" and "(THE x. P (x - d)) = (THE y. P y) + d" unfolding bij_ex1_imp_the_shift[OF bij_add ex1] bij_ex1_imp_the_shift[OF bij_minus ex1] by auto lemma card_shift_image[simp]: shows "card ((\x :: 'a :: group_add. x + d) ` X) = card X" and "card ((\x. x - d) ` X) = card X" by (auto simp: bij_imp_card_image[OF bij_add] bij_imp_card_image[OF bij_minus]) lemma irreducible_root_free: fixes p :: "'a :: {idom,comm_ring_1} poly" assumes irr: "irreducible p" shows "root_free p" proof (cases "degree p" "1::nat" rule: linorder_cases) case greater { fix x assume "poly p x = 0" hence "[:-x,1:] dvd p" using poly_eq_0_iff_dvd by blast then obtain r where p: "p = r * [:-x,1:]" by (elim dvdE, auto) have deg: "degree [:-x,1:] = 1" by simp have dvd: "\ [:-x,1:] dvd 1" by (auto simp: poly_dvd_1) from greater have "degree r \ 0" using degree_mult_le[of r "[:-x,1:]", unfolded deg, folded p] by auto then have "\ r dvd 1" by (auto simp: poly_dvd_1) with p irr irreducibleD[OF irr p] dvd have False by auto } thus ?thesis unfolding root_free_def by auto next case less then have deg: "degree p = 0" by auto from deg obtain p0 where p: "p = [:p0:]" using degree0_coeffs by auto with irr have "p \ 0" by auto with p have "poly p x \ 0" for x by auto thus ?thesis by (auto simp: root_free_def) qed (auto simp: root_free_def) (* **************************************************************** *) subsection \Real Algebraic Numbers -- Innermost Layer\ text \We represent a real algebraic number \\\ by a tuple (p,l,r): \\\ is the unique root in the interval [l,r] and l and r have the same sign. We always assume that p is normalized, i.e., p is the unique irreducible and positive content-free polynomial which represents the algebraic number. This representation clearly admits duplicate representations for the same number, e.g. (...,x-3, 3,3) is equivalent to (...,x-3,2,10).\ subsubsection \Basic Definitions\ type_synonym real_alg_1 = "int poly \ rat \ rat" fun poly_real_alg_1 :: "real_alg_1 \ int poly" where "poly_real_alg_1 (p,_,_) = p" fun rai_ub :: "real_alg_1 \ rat" where "rai_ub (_,_,r) = r" fun rai_lb :: "real_alg_1 \ rat" where "rai_lb (_,l,_) = l" abbreviation "roots_below p x \ {y :: real. y \ x \ ipoly p y = 0}" abbreviation(input) unique_root :: "real_alg_1 \ bool" where "unique_root plr \ (\! x. root_cond plr x)" abbreviation the_unique_root :: "real_alg_1 \ real" where "the_unique_root plr \ (THE x. root_cond plr x)" abbreviation real_of_1 where "real_of_1 \ the_unique_root" lemma root_condI[intro]: assumes "of_rat (rai_lb plr) \ x" and "x \ of_rat (rai_ub plr)" and "ipoly (poly_real_alg_1 plr) x = 0" shows "root_cond plr x" using assms by (auto simp: root_cond_def) lemma root_condE[elim]: assumes "root_cond plr x" and "of_rat (rai_lb plr) \ x \ x \ of_rat (rai_ub plr) \ ipoly (poly_real_alg_1 plr) x = 0 \ thesis" shows thesis using assms by (auto simp: root_cond_def) lemma assumes ur: "unique_root plr" defines "x \ the_unique_root plr" and "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" shows unique_rootD: "of_rat l \ x" "x \ of_rat r" "ipoly p x = 0" "root_cond plr x" "x = y \ root_cond plr y" "y = x \ root_cond plr y" and the_unique_root_eqI: "root_cond plr y \ y = x" "root_cond plr y \ x = y" proof - from ur show x: "root_cond plr x" unfolding x_def by (rule theI') have "plr = (p,l,r)" by (cases plr, auto simp: p_def l_def r_def) from x[unfolded this] show "of_rat l \ x" "x \ of_rat r" "ipoly p x = 0" by auto from x ur show "root_cond plr y \ y = x" and "root_cond plr y \ x = y" and "x = y \ root_cond plr y" and "y = x \ root_cond plr y" by auto qed lemma unique_rootE: assumes ur: "unique_root plr" defines "x \ the_unique_root plr" and "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes main: "of_rat l \ x \ x \ of_rat r \ ipoly p x = 0 \ root_cond plr x \ (\y. x = y \ root_cond plr y) \ (\y. y = x \ root_cond plr y) \ thesis" shows thesis by (rule main, unfold x_def p_def l_def r_def; rule unique_rootD[OF ur]) lemma unique_rootI: assumes "\ y. root_cond plr y \ x = y" "root_cond plr x" shows "unique_root plr" using assms by blast definition poly_cond :: "int poly \ bool" where "poly_cond p = (lead_coeff p > 0 \ irreducible p)" lemma poly_condI[intro]: assumes "lead_coeff p > 0" and "irreducible p" shows "poly_cond p" using assms by (auto simp: poly_cond_def) lemma poly_condD: assumes "poly_cond p" shows "irreducible p" and "lead_coeff p > 0" and "root_free p" and "square_free p" and "p \ 0" using assms unfolding poly_cond_def using irreducible_root_free irreducible_imp_square_free cf_pos_def by auto lemma poly_condE[elim]: assumes "poly_cond p" and "irreducible p \ lead_coeff p > 0 \ root_free p \ square_free p \ p \ 0 \ thesis" shows thesis using assms by (auto dest:poly_condD) definition invariant_1 :: "real_alg_1 \ bool" where "invariant_1 tup \ case tup of (p,l,r) \ unique_root (p,l,r) \ sgn l = sgn r \ poly_cond p" lemma invariant_1I: assumes "unique_root plr" and "sgn (rai_lb plr) = sgn (rai_ub plr)" and "poly_cond (poly_real_alg_1 plr)" shows "invariant_1 plr" using assms by (auto simp: invariant_1_def) lemma assumes "invariant_1 plr" defines "x \ the_unique_root plr" and "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" shows invariant_1D: "root_cond plr x" "sgn l = sgn r" "sgn x = of_rat (sgn r)" "unique_root plr" "poly_cond p" "degree p > 0" "primitive p" and invariant_1_root_cond: "\ y. root_cond plr y \ y = x" proof - let ?l = "of_rat l :: real" let ?r = "of_rat r :: real" have plr: "plr = (p,l,r)" by (cases plr, auto simp: p_def l_def r_def) from assms show ur: "unique_root plr" and sgn: "sgn l = sgn r" and pc: "poly_cond p" by (auto simp: invariant_1_def) from ur show rc: "root_cond plr x" by (auto simp add: x_def plr intro: theI') from this[unfolded plr] have x: "ipoly p x = 0" and bnd: "?l \ x" "x \ ?r" by auto show "sgn x = of_rat (sgn r)" proof (cases "0::real" "x" rule:linorder_cases) case less with bnd(2) have "0 < ?r" by arith thus ?thesis using less by simp next case equal with bnd have "?l \ 0" "?r \ 0" by auto hence "l \ 0" "r \ 0" by auto with \sgn l = sgn r\ have "l = 0" "r = 0" unfolding sgn_rat_def by (auto split: if_splits) with rc[unfolded plr] show ?thesis by auto next case greater with bnd(1) have "?l < 0" by arith thus ?thesis unfolding \sgn l = sgn r\[symmetric] using greater by simp qed from the_unique_root_eqI[OF ur] rc show "\ y. root_cond plr y \ y = x" by metis { assume "degree p = 0" with poly_zero[OF x, simplified] sgn bnd have "p = 0" by auto with pc have "False" by auto } then show "degree p > 0" by auto with pc show "primitive p" by (intro irreducible_imp_primitive, auto) qed lemma invariant_1E[elim]: assumes "invariant_1 plr" defines "x \ the_unique_root plr" and "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes main: "root_cond plr x \ sgn l = sgn r \ sgn x = of_rat (sgn r) \ unique_root plr \ poly_cond p \ degree p > 0 \ primitive p \ thesis" shows thesis apply (rule main) using assms(1) unfolding x_def p_def l_def r_def by (auto dest: invariant_1D) lemma invariant_1_realI: fixes plr :: real_alg_1 defines "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes x: "root_cond plr x" and "sgn l = sgn r" and ur: "unique_root plr" and "poly_cond p" shows "invariant_1 plr \ real_of_1 plr = x" using the_unique_root_eqI[OF ur x] assms by (cases plr, auto intro: invariant_1I) lemma real_of_1_0: assumes "invariant_1 (p,l,r)" shows [simp]: "the_unique_root (p,l,r) = 0 \ r = 0" and [dest]: "l = 0 \ r = 0" and [intro]: "r = 0 \ l = 0" using assms by (auto simp: sgn_0_0) lemma invariant_1_pos: assumes rc: "invariant_1 (p,l,r)" shows [simp]:"the_unique_root (p,l,r) > 0 \ r > 0" (is "?x > 0 \ _") and [simp]:"the_unique_root (p,l,r) < 0 \ r < 0" and [simp]:"the_unique_root (p,l,r) \ 0 \ r \ 0" and [simp]:"the_unique_root (p,l,r) \ 0 \ r \ 0" and [intro]: "r > 0 \ l > 0" and [dest]: "l > 0 \ r > 0" and [intro]: "r < 0 \ l < 0" and [dest]: "l < 0 \ r < 0" proof(atomize(full),goal_cases) case 1 let ?r = "real_of_rat" from assms[unfolded invariant_1_def] have ur: "unique_root (p,l,r)" and sgn: "sgn l = sgn r" by auto from unique_rootD(1-2)[OF ur] have le: "?r l \ ?x" "?x \ ?r r" by auto from rc show ?case proof (cases r "0::rat" rule:linorder_cases) case greater with sgn have "sgn l = 1" by simp hence l0: "l > 0" by (auto simp: sgn_1_pos) hence "?r l > 0" by auto hence "?x > 0" using le(1) by arith with greater l0 show ?thesis by auto next case equal with real_of_1_0[OF rc] show ?thesis by auto next case less hence "?r r < 0" by auto with le(2) have "?x < 0" by arith with less sgn show ?thesis by (auto simp: sgn_1_neg) qed qed definition invariant_1_2 where "invariant_1_2 rai \ invariant_1 rai \ degree (poly_real_alg_1 rai) > 1" definition poly_cond2 where "poly_cond2 p \ poly_cond p \ degree p > 1" lemma poly_cond2I[intro!]: "poly_cond p \ degree p > 1 \ poly_cond2 p" by (simp add: poly_cond2_def) lemma poly_cond2D: assumes "poly_cond2 p" shows "poly_cond p" and "degree p > 1" using assms by (auto simp: poly_cond2_def) lemma poly_cond2E[elim!]: assumes "poly_cond2 p" and "poly_cond p \ degree p > 1 \ thesis" shows thesis using assms by (auto simp: poly_cond2_def) lemma invariant_1_2_poly_cond2: "invariant_1_2 rai \ poly_cond2 (poly_real_alg_1 rai)" unfolding invariant_1_def invariant_1_2_def poly_cond2_def by auto lemma invariant_1_2I[intro!]: assumes "invariant_1 rai" and "degree (poly_real_alg_1 rai) > 1" shows "invariant_1_2 rai" using assms by (auto simp: invariant_1_2_def) lemma invariant_1_2E[elim!]: assumes "invariant_1_2 rai" and "invariant_1 rai \ degree (poly_real_alg_1 rai) > 1 \ thesis" shows thesis using assms[unfolded invariant_1_2_def] by auto lemma invariant_1_2_realI: fixes plr :: real_alg_1 defines "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes x: "root_cond plr x" and sgn: "sgn l = sgn r" and ur: "unique_root plr" and p: "poly_cond2 p" shows "invariant_1_2 plr \ real_of_1 plr = x" using invariant_1_realI[OF x] p sgn ur unfolding p_def l_def r_def by auto subsection \Real Algebraic Numbers = Rational + Irrational Real Algebraic Numbers\ text \In the next representation of real algebraic numbers, we distinguish between rational and irrational numbers. The advantage is that whenever we only work on rational numbers, there is not much overhead involved in comparison to the existing implementation of real numbers which just supports the rational numbers. For irrational numbers we additionally store the number of the root, counting from left to right. For instance $-\sqrt{2}$ and $\sqrt{2}$ would be root number 1 and 2 of $x^2 - 2$.\ subsubsection \Definitions and Algorithms on Raw Type\ datatype real_alg_2 = Rational rat | Irrational nat real_alg_1 fun invariant_2 :: "real_alg_2 \ bool" where "invariant_2 (Irrational n rai) = (invariant_1_2 rai \ n = card(roots_below (poly_real_alg_1 rai) (real_of_1 rai)))" | "invariant_2 (Rational r) = True" fun real_of_2 :: "real_alg_2 \ real" where "real_of_2 (Rational r) = of_rat r" | "real_of_2 (Irrational n rai) = real_of_1 rai" definition of_rat_2 :: "rat \ real_alg_2" where [code_unfold]: "of_rat_2 = Rational" lemma of_rat_2: "real_of_2 (of_rat_2 x) = of_rat x" "invariant_2 (of_rat_2 x)" by (auto simp: of_rat_2_def) (* Invariant type *) typedef real_alg_3 = "Collect invariant_2" morphisms rep_real_alg_3 Real_Alg_Invariant by (rule exI[of _ "Rational 0"], auto) setup_lifting type_definition_real_alg_3 lift_definition real_of_3 :: "real_alg_3 \ real" is real_of_2 . (* *************** *) subsubsection \Definitions and Algorithms on Quotient Type\ quotient_type real_alg = real_alg_3 / "\ x y. real_of_3 x = real_of_3 y" morphisms rep_real_alg Real_Alg_Quotient by (auto simp: equivp_def) metis (* real_of *) lift_definition real_of :: "real_alg \ real" is real_of_3 . lemma real_of_inj: "(real_of x = real_of y) = (x = y)" by (transfer, simp) (* ********************** *) subsubsection \Sign\ definition sgn_1 :: "real_alg_1 \ rat" where "sgn_1 x = sgn (rai_ub x)" lemma sgn_1: "invariant_1 x \ real_of_rat (sgn_1 x) = sgn (real_of_1 x)" unfolding sgn_1_def by auto lemma sgn_1_inj: "invariant_1 x \ invariant_1 y \ real_of_1 x = real_of_1 y \ sgn_1 x = sgn_1 y" by (auto simp: sgn_1_def elim!: invariant_1E) (* ********************** *) subsubsection \Normalization: Bounds Close Together\ lemma unique_root_lr: assumes ur: "unique_root plr" shows "rai_lb plr \ rai_ub plr" (is "?l \ ?r") proof - let ?p = "poly_real_alg_1 plr" from ur[unfolded root_cond_def] have ex1: "\! x :: real. of_rat ?l \ x \ x \ of_rat ?r \ ipoly ?p x = 0" by (cases plr, simp) then obtain x :: real where bnd: "of_rat ?l \ x" "x \ of_rat ?r" and rt: "ipoly ?p x = 0" by auto from bnd have "real_of_rat ?l \ of_rat ?r" by linarith thus "?l \ ?r" by (simp add: of_rat_less_eq) qed locale map_poly_zero_hom_0 = base: zero_hom_0 begin sublocale zero_hom_0 "map_poly hom" by (unfold_locales,auto) end interpretation of_int_poly_hom: map_poly_zero_hom_0 "of_int :: int \ 'a :: {ring_1, ring_char_0}" .. lemma ipoly_roots_finite: "p \ 0 \ finite {x :: 'a :: {idom, ring_char_0}. ipoly p x = 0}" by (rule poly_roots_finite, simp) lemma roots_below_the_unique_root: assumes ur: "unique_root (p,l,r)" shows "roots_below p (the_unique_root (p,l,r)) = roots_below p (of_rat r)" (is "roots_below p ?x = _") proof- from ur have rc: "root_cond (p,l,r) ?x" by (auto dest!: unique_rootD) with ur have x: "{x. root_cond (p,l,r) x} = {?x}" by (auto intro: the_unique_root_eqI) from rc have "?x \ {y. ?x \ y \ y \ of_rat r \ ipoly p y = 0}" by auto with rc have l1x: "... = {?x}" by (intro equalityI, fold x(1), force, simp add: x) have rb:"roots_below p (of_rat r) = roots_below p ?x \ {y. ?x < y \ y \ of_rat r \ ipoly p y = 0}" using rc by auto have emp: "\x. the_unique_root (p, l, r) < x \ x \ {ra. ?x \ ra \ ra \ real_of_rat r \ ipoly p ra = 0}" using l1x by auto with rb show ?thesis by auto qed lemma unique_root_sub_interval: assumes ur: "unique_root (p,l,r)" and rc: "root_cond (p,l',r') (the_unique_root (p,l,r))" and between: "l \ l'" "r' \ r" shows "unique_root (p,l',r')" and "the_unique_root (p,l',r') = the_unique_root (p,l,r)" proof - from between have ord: "real_of_rat l \ of_rat l'" "real_of_rat r' \ of_rat r" by (auto simp: of_rat_less_eq) from rc have lr': "real_of_rat l' \ of_rat r'" by auto with ord have lr: "real_of_rat l \ real_of_rat r" by auto show "\!x. root_cond (p, l', r') x" proof (rule, rule rc) fix y assume "root_cond (p,l',r') y" with ord have "root_cond (p,l,r) y" by (auto intro!:root_condI) from the_unique_root_eqI[OF ur this] show "y = the_unique_root (p,l,r)" by simp qed from the_unique_root_eqI[OF this rc] show "the_unique_root (p,l',r') = the_unique_root (p,l,r)" by simp qed lemma invariant_1_sub_interval: assumes rc: "invariant_1 (p,l,r)" and sub: "root_cond (p,l',r') (the_unique_root (p,l,r))" and between: "l \ l'" "r' \ r" shows "invariant_1 (p,l',r')" and "real_of_1 (p,l',r') = real_of_1 (p,l,r)" proof - let ?r = real_of_rat note rcD = invariant_1D[OF rc] from rc have ur: "unique_root (p, l', r')" and id: "the_unique_root (p, l', r') = the_unique_root (p, l, r)" by (atomize(full), intro conjI unique_root_sub_interval[OF _ sub between], auto) show "real_of_1 (p,l',r') = real_of_1 (p,l,r)" using id by simp from rcD(1)[unfolded split] have "?r l \ ?r r" by auto hence lr: "l \ r" by (auto simp: of_rat_less_eq) from unique_rootD[OF ur] have "?r l' \ ?r r'" by auto hence lr': "l' \ r'" by (auto simp: of_rat_less_eq) have "sgn l' = sgn r'" proof (cases "r" "0::rat" rule: linorder_cases) case less with lr lr' between have "l < 0" "l' < 0" "r' < 0" "r < 0" by auto thus ?thesis unfolding sgn_rat_def by auto next case equal with rcD(2) have "l = 0" using sgn_0_0 by auto with equal between lr' have "l' = 0" "r' = 0" by auto then show ?thesis by auto next case greater with rcD(4) have "sgn r = 1" unfolding sgn_rat_def by (cases "r = 0", auto) with rcD(2) have "sgn l = 1" by simp hence l: "l > 0" unfolding sgn_rat_def by (cases "l = 0"; cases "l < 0"; auto) with lr lr' between have "l > 0" "l' > 0" "r' > 0" "r > 0" by auto thus ?thesis unfolding sgn_rat_def by auto qed with between ur rc show "invariant_1 (p,l',r')" by (auto simp add: invariant_1_def id) qed lemma root_sign_change: assumes p0: "poly (p::real poly) x = 0" and pd_ne0: "poly (pderiv p) x \ 0" obtains d where "0 < d" "sgn (poly p (x - d)) \ sgn (poly p (x + d))" "sgn (poly p (x - d)) \ 0" "0 \ sgn (poly p (x + d))" "\ d' > 0. d' \ d \ sgn (poly p (x + d')) = sgn (poly p (x + d)) \ sgn (poly p (x - d')) = sgn (poly p (x - d))" proof - assume a:"(\d. 0 < d \ sgn (poly p (x - d)) \ sgn (poly p (x + d)) \ sgn (poly p (x - d)) \ 0 \ 0 \ sgn (poly p (x + d)) \ \d'>0. d' \ d \ sgn (poly p (x + d')) = sgn (poly p (x + d)) \ sgn (poly p (x - d')) = sgn (poly p (x - d)) \ thesis)" from pd_ne0 consider "poly (pderiv p) x > 0" | "poly (pderiv p) x < 0" by linarith thus ?thesis proof(cases) case 1 obtain d1 where d1:"\h. 0 h < d1 \ poly p (x - h) < 0" "d1 > 0" using DERIV_pos_inc_left[OF poly_DERIV 1] p0 by auto obtain d2 where d2:"\h. 0 h < d2 \ poly p (x + h) > 0" "d2 > 0" using DERIV_pos_inc_right[OF poly_DERIV 1] p0 by auto have g0:"0 < (min d1 d2) / 2" using d1 d2 by auto hence m1:"min d1 d2 / 2 < d1" and m2:"min d1 d2 / 2 < d2" by auto { fix d assume a1:"0 < d" and a2:"d < min d1 d2" have "sgn (poly p (x - d)) = -1" "sgn (poly p (x + d)) = 1" using d1(1)[OF a1] d2(1)[OF a1] a2 by auto } note d=this show ?thesis by(rule a[OF g0];insert d g0 m1 m2, simp) next case 2 obtain d1 where d1:"\h. 0 h < d1 \ poly p (x - h) > 0" "d1 > 0" using DERIV_neg_dec_left[OF poly_DERIV 2] p0 by auto obtain d2 where d2:"\h. 0 h < d2 \ poly p (x + h) < 0" "d2 > 0" using DERIV_neg_dec_right[OF poly_DERIV 2] p0 by auto have g0:"0 < (min d1 d2) / 2" using d1 d2 by auto hence m1:"min d1 d2 / 2 < d1" and m2:"min d1 d2 / 2 < d2" by auto { fix d assume a1:"0 < d" and a2:"d < min d1 d2" have "sgn (poly p (x - d)) = 1" "sgn (poly p (x + d)) = -1" using d1(1)[OF a1] d2(1)[OF a1] a2 by auto } note d=this show ?thesis by(rule a[OF g0];insert d g0 m1 m2, simp) qed qed lemma rational_root_free_degree_iff: assumes rf: "root_free (map_poly rat_of_int p)" and rt: "ipoly p x = 0" shows "(x \ \) = (degree p = 1)" proof assume "x \ \" then obtain y where x: "x = of_rat y" (is "_ = ?x") unfolding Rats_def by blast from rt[unfolded x] have "poly (map_poly rat_of_int p) y = 0" by simp with rf show "degree p = 1" unfolding root_free_def by auto next assume "degree p = 1" from degree1_coeffs[OF this] obtain a b where p: "p = [:a,b:]" and b: "b \ 0" by auto from rt[unfolded p hom_distribs] have "of_int a + x * of_int b = 0" by auto from arg_cong[OF this, of "\ x. (x - of_int a) / of_int b"] have "x = - of_rat (of_int a) / of_rat (of_int b)" using b by auto also have "\ = of_rat (- of_int a / of_int b)" unfolding of_rat_minus of_rat_divide .. finally show "x \ \" by auto qed lemma rational_poly_cond_iff: assumes "poly_cond p" and "ipoly p x = 0" and "degree p > 1" shows "(x \ \) = (degree p = 1)" proof (rule rational_root_free_degree_iff[OF _ assms(2)]) from poly_condD[OF assms(1)] irreducible_connect_rev[of p] assms(3) have p: "irreducible\<^sub>d p" by auto from irreducible\<^sub>d_int_rat[OF this] have "irreducible (map_poly rat_of_int p)" by simp thus "root_free (map_poly rat_of_int p)" by (rule irreducible_root_free) qed lemma poly_cond_degree_gt_1: assumes "poly_cond p" "degree p > 1" "ipoly p x = 0" shows "x \ \" using rational_poly_cond_iff[OF assms(1,3)] assms(2) by simp lemma poly_cond2_no_rat_root: assumes "poly_cond2 p" shows "ipoly p (real_of_rat x) \ 0" using poly_cond_degree_gt_1[of p "real_of_rat x"] assms by auto context fixes p :: "int poly" and x :: "rat" begin lemma gt_rat_sign_change: assumes ur: "unique_root plr" defines "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes p: "poly_cond2 p" and in_interval: "l \ y" "y \ r" shows "(sgn (ipoly p y) = sgn (ipoly p r)) = (of_rat y > the_unique_root plr)" (is "?gt = _") proof (rule ccontr) have plr[simp]: "plr = (p,l,r)" by (cases plr, auto simp: p_def l_def r_def) assume "?gt \ (real_of_rat y > the_unique_root plr)" note a = this[unfolded plr] from p have "irreducible p" by auto note nz = poly_cond2_no_rat_root[OF p] hence "p \ 0" unfolding irreducible_def by auto hence p0_real: "real_of_int_poly p \ (0::real poly)" by auto let ?p = "real_of_int_poly p" note urD = unique_rootD[OF ur, simplified] let ?ur = "the_unique_root (p, l, r)" let ?r = real_of_rat from poly_cond2_no_rat_root p have ru:"ipoly p y \ 0" by auto from in_interval have in':"?r l \ ?r y" "?r y \ ?r r" unfolding of_rat_less_eq by auto from p square_free_of_int_poly[of p] square_free_rsquarefree have rsf:"rsquarefree ?p" by auto have ur3:"poly ?p ?ur = 0" using urD(3) by simp from ur have "?ur \ of_rat r" by (auto elim!: unique_rootE) moreover from nz have "ipoly p (real_of_rat r) \ 0" by auto with ur3 have "real_of_rat r \ real_of_1 (p,l,r)" by force ultimately have "?ur < ?r r" by auto hence ur2: "0 < ?r r - ?ur" by linarith from rsquarefree_roots rsf ur3 have pd_nonz:"poly (pderiv ?p) ?ur \ 0" by auto obtain d where d':"\d'. d'>0 \ d' \ d \ sgn (poly ?p (?ur + d')) = sgn (poly ?p (?ur + d)) \ sgn (poly ?p (?ur - d')) = sgn (poly ?p (?ur - d))" "sgn (poly ?p (?ur - d)) \ sgn (poly ?p (?ur + d))" "sgn (poly ?p (?ur + d)) \ 0" and d_ge_0:"d > 0" by (metis root_sign_change[OF ur3 pd_nonz]) have sr:"sgn (poly ?p (?ur + d)) = sgn (poly ?p (?r r))" proof (cases "?r r - ?ur \ d") case True show ?thesis using d'(1)[OF ur2 True] by auto next case False hence less:"?ur + d < ?r r" by auto show ?thesis proof(rule no_roots_inbetween_imp_same_sign[OF less,rule_format],goal_cases) case (1 x) from ur 1 d_ge_0 have ran: "real_of_rat l \ x" "x \ real_of_rat r" by (auto elim!: unique_rootE) from 1 d_ge_0 have "the_unique_root (p, l, r) \ x" by auto with ur have "\ root_cond (p,l,r) x" by auto with ran show ?case by auto qed qed consider "?r l < ?ur - d" "?r l < ?ur" | "0 < ?ur - ?r l" "?ur - ?r l \ d" | "?ur = ?r l" using urD by argo hence sl:"sgn (poly ?p (?ur - d)) = sgn (poly ?p (?r l)) \ 0 = sgn (poly ?p (?r l))" proof (cases) case 1 have "sgn (poly ?p (?r l)) = sgn (poly ?p (?ur - d))" proof(rule no_roots_inbetween_imp_same_sign[OF 1(1),rule_format],goal_cases) case (1 x) from ur 1 d_ge_0 have ran: "real_of_rat l \ x" "x \ real_of_rat r" by (auto elim!: unique_rootE) from 1 d_ge_0 have "the_unique_root (p, l, r) \ x" by auto with ur have "\ root_cond (p,l,r) x" by auto with ran show ?case by auto qed thus ?thesis by auto next case 2 show ?thesis using d'(1)[OF 2] by simp qed (insert ur3,simp) have diff_sign: "sgn (ipoly p l) \ sgn (ipoly p r)" using d'(2-) sr sl real_of_rat_sgn by auto have ur':"\x. real_of_rat l \ x \ x \ real_of_rat y \ ipoly p x = 0 \ \ (?r y \ the_unique_root (p,l,r))" proof(standard+,goal_cases) case (1 x) { assume id: "the_unique_root (p,l,r) = ?r y" from nz[of y] id ur have False by (auto elim!: unique_rootE) } note neq = this have "root_cond (p, l, r) x" unfolding root_cond_def using 1 a ur by (auto elim!: unique_rootE) with conjunct2[OF 1(1)] 1(2-) the_unique_root_eqI[OF ur] show ?case by (auto intro!: neq) qed hence ur'':"\x. real_of_rat y \ x \ x \ real_of_rat r \ poly (real_of_int_poly p) x \ 0 \ \ (?r y \ the_unique_root (p,l,r))" using urD(2,3) by auto have "(sgn (ipoly p y) = sgn (ipoly p r)) = (?r y > the_unique_root (p,l,r))" proof(cases "sgn (ipoly p r) = sgn (ipoly p y)") case True have sgn:"sgn (poly ?p (real_of_rat l)) \ sgn (poly ?p (real_of_rat y))" using True diff_sign by (simp add: real_of_rat_sgn) have ly:"of_rat l < (of_rat y::real)" using in_interval True diff_sign less_eq_rat_def of_rat_less by auto with no_roots_inbetween_imp_same_sign[OF ly,of ?p] sgn ur' True show ?thesis by force next case False hence ne:"sgn (ipoly p (real_of_rat y)) \ sgn (ipoly p (real_of_rat r))" by (simp add: real_of_rat_sgn) have ry:"of_rat y < (of_rat r::real)" using in_interval False diff_sign less_eq_rat_def of_rat_less by auto obtain x where x:"real_of_rat y \ x" "x \ real_of_rat r" "ipoly p x = 0" using no_roots_inbetween_imp_same_sign[OF ry,of ?p] ne by auto hence lx:"real_of_rat l \ x" using in_interval using False a urD by auto have "?ur = x" using x lx ur by (intro the_unique_root_eqI, auto) then show ?thesis using False x by auto qed thus False using diff_sign(1) a ru by(cases "ipoly p r = 0";auto simp:sgn_0_0) qed definition tighten_poly_bounds :: "rat \ rat \ rat \ rat \ rat \ rat" where "tighten_poly_bounds l r sr = (let m = (l + r) / 2; sm = sgn (ipoly p m) in if sm = sr then (l,m,sm) else (m,r,sr))" lemma tighten_poly_bounds: assumes res: "tighten_poly_bounds l r sr = (l',r',sr')" and ur: "unique_root (p,l,r)" and p: "poly_cond2 p" and sr: "sr = sgn (ipoly p r)" shows "root_cond (p,l',r') (the_unique_root (p,l,r))" "l \ l'" "l' \ r'" "r' \ r" "(r' - l') = (r - l) / 2" "sr' = sgn (ipoly p r')" proof - let ?x = "the_unique_root (p,l,r)" let ?x' = "the_unique_root (p,l',r')" let ?m = "(l + r) / 2" note d = tighten_poly_bounds_def Let_def from unique_root_lr[OF ur] have lr: "l \ r" by auto thus "l \ l'" "l' \ r'" "r' \ r" "(r' - l') = (r - l) / 2" "sr' = sgn (ipoly p r')" using res sr unfolding d by (auto split: if_splits) hence "l \ ?m" "?m \ r" by auto note le = gt_rat_sign_change[OF ur,simplified,OF p this] note urD = unique_rootD[OF ur] show "root_cond (p,l',r') ?x" proof (cases "sgn (ipoly p ?m) = sgn (ipoly p r)") case *: False with res sr have id: "l' = ?m" "r' = r" unfolding d by auto from *[unfolded le] urD show ?thesis unfolding id by auto next case *: True with res sr have id: "l' = l" "r' = ?m" unfolding d by auto from *[unfolded le] urD show ?thesis unfolding id by auto qed qed partial_function (tailrec) tighten_poly_bounds_epsilon :: "rat \ rat \ rat \ rat \ rat \ rat" where [code]: "tighten_poly_bounds_epsilon l r sr = (if r - l \ x then (l,r,sr) else (case tighten_poly_bounds l r sr of (l',r',sr') \ tighten_poly_bounds_epsilon l' r' sr'))" partial_function (tailrec) tighten_poly_bounds_for_x :: "rat \ rat \ rat \ rat \ rat \ rat" where [code]: "tighten_poly_bounds_for_x l r sr = (if x < l \ r < x then (l, r, sr) else (case tighten_poly_bounds l r sr of (l',r',sr') \ tighten_poly_bounds_for_x l' r' sr'))" lemma tighten_poly_bounds_epsilon: assumes ur: "unique_root (p,l,r)" defines u: "u \ the_unique_root (p,l,r)" assumes p: "poly_cond2 p" and res: "tighten_poly_bounds_epsilon l r sr = (l',r',sr')" and sr: "sr = sgn (ipoly p r)" and x: "x > 0" shows "l \ l'" "r' \ r" "root_cond (p,l',r') u" "r' - l' \ x" "sr' = sgn (ipoly p r')" proof - let ?u = "the_unique_root (p,l,r)" define delta where "delta = x / 2" have delta: "delta > 0" unfolding delta_def using x by auto let ?dist = "\ (l,r,sr). r - l" let ?rel = "inv_image {(x, y). 0 \ y \ delta_gt delta x y} ?dist" note SN = SN_inv_image[OF delta_gt_SN[OF delta], of ?dist] note simps = res[unfolded tighten_poly_bounds_for_x.simps[of l r]] let ?P = "\ (l,r,sr). unique_root (p,l,r) \ u = the_unique_root (p,l,r) \ tighten_poly_bounds_epsilon l r sr = (l',r',sr') \ sr = sgn (ipoly p r) \ l \ l' \ r' \ r \ r' - l' \ x \ root_cond (p,l',r') u \ sr' = sgn (ipoly p r')" have "?P (l,r,sr)" proof (induct rule: SN_induct[OF SN]) case (1 lr) obtain l r sr where lr: "lr = (l,r,sr)" by (cases lr, auto) show ?case unfolding lr split proof (intro impI) assume ur: "unique_root (p, l, r)" and u: "u = the_unique_root (p, l, r)" and res: "tighten_poly_bounds_epsilon l r sr = (l', r', sr')" and sr: "sr = sgn (ipoly p r)" note tur = unique_rootD[OF ur] note simps = tighten_poly_bounds_epsilon.simps[of l r sr] show "l \ l' \ r' \ r \ r' - l' \ x \ root_cond (p, l', r') u \ sr' = sgn (ipoly p r')" proof (cases "r - l \ x") case True with res[unfolded simps] ur tur(4) u sr show ?thesis by auto next case False hence x: "r - l > x" by auto let ?tight = "tighten_poly_bounds l r sr" obtain L R SR where tight: "?tight = (L,R,SR)" by (cases ?tight, auto) note tighten = tighten_poly_bounds[OF tight[unfolded sr] ur p] from unique_root_sub_interval[OF ur tighten(1-2,4)] p have ur': "unique_root (p,L,R)" "u = the_unique_root (p,L,R)" unfolding u by auto from res[unfolded simps tight] False sr have "tighten_poly_bounds_epsilon L R SR = (l',r',sr')" by auto note IH = 1[of "(L,R,SR)", unfolded tight split lr, rule_format, OF _ ur' this] have "L \ l' \ r' \ R \ r' - l' \ x \ root_cond (p, l', r') u \ sr' = sgn (ipoly p r')" by (rule IH, insert tighten False, auto simp: delta_gt_def delta_def) thus ?thesis using tighten by auto qed qed qed from this[unfolded split u, rule_format, OF ur refl res sr] show "l \ l'" "r' \ r" "root_cond (p,l',r') u" "r' - l' \ x" "sr' = sgn (ipoly p r')" using u by auto qed lemma tighten_poly_bounds_for_x: assumes ur: "unique_root (p,l,r)" defines u: "u \ the_unique_root (p,l,r)" assumes p: "poly_cond2 p" and res: "tighten_poly_bounds_for_x l r sr = (l',r',sr')" and sr: "sr = sgn (ipoly p r)" shows "l \ l'" "l' \ r'" "r' \ r" "root_cond (p,l',r') u" "\ (l' \ x \ x \ r')" "sr' = sgn (ipoly p r')" "unique_root (p,l',r')" proof - let ?u = "the_unique_root (p,l,r)" let ?x = "real_of_rat x" define delta where "delta = abs ((u - ?x) / 2)" let ?p = "real_of_int_poly p" note ru = unique_rootD[OF ur] { assume "u = ?x" note u = this[unfolded u] from poly_cond2_no_rat_root[OF p] ur have False by (elim unique_rootE, auto simp: u) } hence delta: "delta > 0" unfolding delta_def by auto let ?dist = "\ (l,r,sr). real_of_rat (r - l)" let ?rel = "inv_image {(x, y). 0 \ y \ delta_gt delta x y} ?dist" note SN = SN_inv_image[OF delta_gt_SN[OF delta], of ?dist] note simps = res[unfolded tighten_poly_bounds_for_x.simps[of l r]] let ?P = "\ (l,r,sr). unique_root (p,l,r) \ u = the_unique_root (p,l,r) \ tighten_poly_bounds_for_x l r sr = (l',r',sr') \ sr = sgn (ipoly p r) \ l \ l' \ r' \ r \ \ (l' \ x \ x \ r') \ root_cond (p,l',r') u \ sr' = sgn (ipoly p r')" have "?P (l,r,sr)" proof (induct rule: SN_induct[OF SN]) case (1 lr) obtain l r sr where lr: "lr = (l,r,sr)" by (cases lr, auto) let ?l = "real_of_rat l" let ?r = "real_of_rat r" show ?case unfolding lr split proof (intro impI) assume ur: "unique_root (p, l, r)" and u: "u = the_unique_root (p, l, r)" and res: "tighten_poly_bounds_for_x l r sr = (l', r', sr')" and sr: "sr = sgn (ipoly p r)" note tur = unique_rootD[OF ur] note simps = tighten_poly_bounds_for_x.simps[of l r] show "l \ l' \ r' \ r \ \ (l' \ x \ x \ r') \ root_cond (p, l', r') u \ sr' = sgn (ipoly p r')" proof (cases "x < l \ r < x") case True with res[unfolded simps] ur tur(4) u sr show ?thesis by auto next case False hence x: "?l \ ?x" "?x \ ?r" by (auto simp: of_rat_less_eq) let ?tight = "tighten_poly_bounds l r sr" obtain L R SR where tight: "?tight = (L,R,SR)" by (cases ?tight, auto) note tighten = tighten_poly_bounds[OF tight ur p sr] from unique_root_sub_interval[OF ur tighten(1-2,4)] p have ur': "unique_root (p,L,R)" "u = the_unique_root (p,L,R)" unfolding u by auto from res[unfolded simps tight] False have "tighten_poly_bounds_for_x L R SR = (l',r',sr')" by auto note IH = 1[of ?tight, unfolded tight split lr, rule_format, OF _ ur' this] let ?DIFF = "real_of_rat (R - L)" let ?diff = "real_of_rat (r - l)" have diff0: "0 \ ?DIFF" using tighten(3) by (metis cancel_comm_monoid_add_class.diff_cancel diff_right_mono of_rat_less_eq of_rat_hom.hom_zero) have *: "r - l - (r - l) / 2 = (r - l) / 2" by (auto simp: field_simps) have "delta_gt delta ?diff ?DIFF = (abs (u - of_rat x) \ real_of_rat (r - l) * 1)" unfolding delta_gt_def tighten(5) delta_def of_rat_diff[symmetric] * by (simp add: hom_distribs) also have "real_of_rat (r - l) * 1 = ?r - ?l" unfolding of_rat_divide of_rat_mult of_rat_diff by auto also have "abs (u - of_rat x) \ ?r - ?l" using x ur by (elim unique_rootE, auto simp: u) finally have delta: "delta_gt delta ?diff ?DIFF" . have "L \ l' \ r' \ R \ \ (l' \ x \ x \ r') \ root_cond (p, l', r') u \ sr' = sgn (ipoly p r')" by (rule IH, insert delta diff0 tighten(6), auto) with \l \ L\ \R \ r\ show ?thesis by auto qed qed qed from this[unfolded split u, rule_format, OF ur refl res sr] show *: "l \ l'" "r' \ r" "root_cond (p,l',r') u" "\ (l' \ x \ x \ r')" "sr' = sgn (ipoly p r')" unfolding u by auto from *(3)[unfolded split] have "real_of_rat l' \ of_rat r'" by auto thus "l' \ r'" unfolding of_rat_less_eq . show "unique_root (p,l',r')" using ur *(1-3) p poly_condD(5) u unique_root_sub_interval(1) by blast qed end definition real_alg_precision :: rat where "real_alg_precision \ Rat.Fract 1 2" lemma real_alg_precision: "real_alg_precision > 0" by eval definition normalize_bounds_1_main :: "rat \ real_alg_1 \ real_alg_1" where "normalize_bounds_1_main eps rai = (case rai of (p,l,r) \ let (l',r',sr') = tighten_poly_bounds_epsilon p eps l r (sgn (ipoly p r)); fr = rat_of_int (floor r'); (l'',r'',_) = tighten_poly_bounds_for_x p fr l' r' sr' in (p,l'',r''))" definition normalize_bounds_1 :: "real_alg_1 \ real_alg_1" where "normalize_bounds_1 = (normalize_bounds_1_main real_alg_precision)" context fixes p q and l r :: rat assumes cong: "\ x. real_of_rat l \ x \ x \ of_rat r \ (ipoly p x = (0 :: real)) = (ipoly q x = 0)" begin lemma root_cond_cong: "root_cond (p,l,r) = root_cond (q,l,r)" by (intro ext, insert cong, auto simp: root_cond_def) lemma the_unique_root_cong: "the_unique_root (p,l,r) = the_unique_root (q,l,r)" unfolding root_cond_cong .. lemma unique_root_cong: "unique_root (p,l,r) = unique_root (q,l,r)" unfolding root_cond_cong .. end lemma normalize_bounds_1_main: assumes eps: "eps > 0" and rc: "invariant_1_2 x" defines y: "y \ normalize_bounds_1_main eps x" shows "invariant_1_2 y \ (real_of_1 y = real_of_1 x)" proof - obtain p l r where x: "x = (p,l,r)" by (cases x) auto note rc = rc[unfolded x] obtain l' r' sr' where tb: "tighten_poly_bounds_epsilon p eps l r (sgn (ipoly p r)) = (l',r',sr')" by (cases rule: prod_cases3, auto) let ?fr = "rat_of_int (floor r')" obtain l'' r'' sr'' where tbx: "tighten_poly_bounds_for_x p ?fr l' r' sr' = (l'',r'',sr'')" by (cases rule: prod_cases3, auto) from y[unfolded normalize_bounds_1_main_def x] tb tbx have y: "y = (p, l'', r'')" by (auto simp: Let_def) from rc have "unique_root (p, l, r)" and p2: "poly_cond2 p" by auto from tighten_poly_bounds_epsilon[OF this tb refl eps] have bnd: "l \ l'" "r' \ r" and rc': "root_cond (p, l', r') (the_unique_root (p, l, r))" and eps: "r' - l' \ eps" (* currently not relevant for lemma *) and sr': "sr' = sgn (ipoly p r')" by auto from invariant_1_sub_interval[OF _ rc' bnd] rc have inv': "invariant_1 (p, l', r')" and eq: "real_of_1 (p, l', r') = real_of_1 (p, l, r)" by auto have bnd: "l' \ l''" "r'' \ r'" and rc': "root_cond (p, l'', r'') (the_unique_root (p, l', r'))" by (rule tighten_poly_bounds_for_x[OF _ p2 tbx sr'], fact invariant_1D[OF inv'])+ from invariant_1_sub_interval[OF inv' rc' bnd] p2 eq show ?thesis unfolding y x by auto qed lemma normalize_bounds_1: assumes x: "invariant_1_2 x" shows "invariant_1_2 (normalize_bounds_1 x) \ (real_of_1 (normalize_bounds_1 x) = real_of_1 x)" proof(cases x) case xx:(fields p l r) let ?res = "(p,l,r)" have norm: "normalize_bounds_1 x = (normalize_bounds_1_main real_alg_precision ?res)" unfolding normalize_bounds_1_def by (simp add: xx) from x have x: "invariant_1_2 ?res" "real_of_1 ?res = real_of_1 x" unfolding xx by auto from normalize_bounds_1_main[OF real_alg_precision x(1)] x(2-) show ?thesis unfolding normalize_bounds_1_def xx by auto qed lemma normalize_bound_1_poly: "poly_real_alg_1 (normalize_bounds_1 rai) = poly_real_alg_1 rai" unfolding normalize_bounds_1_def normalize_bounds_1_main_def Let_def by (auto split: prod.splits) definition real_alg_2_main :: "root_info \ real_alg_1 \ real_alg_2" where "real_alg_2_main ri rai \ let p = poly_real_alg_1 rai in (if degree p = 1 then Rational (Rat.Fract (- coeff p 0) (coeff p 1)) else (case normalize_bounds_1 rai of (p',l,r) \ Irrational (root_info.number_root ri r) (p',l,r)))" definition real_alg_2 :: "real_alg_1 \ real_alg_2" where "real_alg_2 rai \ let p = poly_real_alg_1 rai in (if degree p = 1 then Rational (Rat.Fract (- coeff p 0) (coeff p 1)) else (case normalize_bounds_1 rai of (p',l,r) \ Irrational (root_info.number_root (root_info p) r) (p',l,r)))" lemma degree_1_ipoly: assumes "degree p = Suc 0" shows "ipoly p x = 0 \ (x = real_of_rat (Rat.Fract (- coeff p 0) (coeff p 1)))" proof - from roots1[of "map_poly real_of_int p"] assms have "ipoly p x = 0 \ x \ {roots1 (real_of_int_poly p)}" by auto also have "\ = (x = real_of_rat (Rat.Fract (- coeff p 0) (coeff p 1)))" unfolding Fract_of_int_quotient roots1_def hom_distribs by auto finally show ?thesis . qed lemma invariant_1_degree_0: assumes inv: "invariant_1 rai" shows "degree (poly_real_alg_1 rai) \ 0" (is "degree ?p \ 0") proof (rule notI) assume deg: "degree ?p = 0" from inv have "ipoly ?p (real_of_1 rai) = 0" by auto with deg have "?p = 0" by (meson less_Suc0 representsI represents_degree) with inv show False by auto qed lemma real_alg_2_main: assumes inv: "invariant_1 rai" defines [simp]: "p \ poly_real_alg_1 rai" assumes ric: "irreducible (poly_real_alg_1 rai) \ root_info_cond ri (poly_real_alg_1 rai)" shows "invariant_2 (real_alg_2_main ri rai)" "real_of_2 (real_alg_2_main ri rai) = real_of_1 rai" proof (atomize(full)) define l r where [simp]: "l \ rai_lb rai" and [simp]: "r \ rai_ub rai" show "invariant_2 (real_alg_2_main ri rai) \ real_of_2 (real_alg_2_main ri rai) = real_of_1 rai" unfolding id using invariant_1D proof (cases "degree p" "Suc 0" rule: linorder_cases) case deg: equal hence id: "real_alg_2_main ri rai = Rational (Rat.Fract (- coeff p 0) (coeff p 1))" unfolding real_alg_2_main_def Let_def by auto note rc = invariant_1D[OF inv] from degree_1_ipoly[OF deg, of "the_unique_root rai"] rc(1) show ?thesis unfolding id by auto next case deg: greater with inv have inv: "invariant_1_2 rai" unfolding p_def by auto define rai' where "rai' = normalize_bounds_1 rai" have rai': "real_of_1 rai = real_of_1 rai'" and inv': "invariant_1_2 rai'" unfolding rai'_def using normalize_bounds_1[OF inv] by auto obtain p' l' r' where "rai' = (p',l',r')" by (cases rai') with arg_cong[OF rai'_def, of poly_real_alg_1, unfolded normalize_bound_1_poly] split have split: "rai' = (p,l',r')" by auto from inv'[unfolded split] have "poly_cond p" by auto from poly_condD[OF this] have irr: "irreducible p" by simp from ric irr have ric: "root_info_cond ri p" by auto have id: "real_alg_2_main ri rai = (Irrational (root_info.number_root ri r') rai')" unfolding real_alg_2_main_def Let_def using deg split rai'_def by (auto simp: rai'_def rai') show ?thesis unfolding id using rai' root_info_condD(2)[OF ric] inv'[unfolded split] apply (elim invariant_1_2E invariant_1E) using inv' by(auto simp: split roots_below_the_unique_root) next case deg: less then have "degree p = 0" by auto from this invariant_1_degree_0[OF inv] have "p = 0" by simp with inv show ?thesis by auto qed qed lemma real_alg_2: assumes "invariant_1 rai" shows "invariant_2 (real_alg_2 rai)" "real_of_2 (real_alg_2 rai) = real_of_1 rai" proof - have deg: "0 < degree (poly_real_alg_1 rai)" using assms by auto have "real_alg_2 rai = real_alg_2_main (root_info (poly_real_alg_1 rai)) rai" unfolding real_alg_2_def real_alg_2_main_def Let_def by auto from real_alg_2_main[OF assms root_info, folded this, simplified] deg show "invariant_2 (real_alg_2 rai)" "real_of_2 (real_alg_2 rai) = real_of_1 rai" by auto qed lemma invariant_2_realI: fixes plr :: real_alg_1 defines "p \ poly_real_alg_1 plr" and "l \ rai_lb plr" and "r \ rai_ub plr" assumes x: "root_cond plr x" and sgn: "sgn l = sgn r" and ur: "unique_root plr" and p: "poly_cond p" shows "invariant_2 (real_alg_2 plr) \ real_of_2 (real_alg_2 plr) = x" using invariant_1_realI[OF x,folded p_def l_def r_def] sgn ur p real_alg_2[of plr] by auto (* ********************* *) subsubsection \Comparisons\ fun compare_rat_1 :: "rat \ real_alg_1 \ order" where "compare_rat_1 x (p,l,r) = (if x < l then Lt else if x > r then Gt else if sgn (ipoly p x) = sgn(ipoly p r) then Gt else Lt)" lemma compare_rat_1: assumes rai: "invariant_1_2 y" shows "compare_rat_1 x y = compare (of_rat x) (real_of_1 y)" proof- define p l r where "p \ poly_real_alg_1 y" "l \ rai_lb y" "r \ rai_ub y" then have y [simp]: "y = (p,l,r)" by (cases y, auto) from rai have ur: "unique_root y" by auto show ?thesis proof (cases "x < l \ x > r") case True { assume xl: "x < l" hence "real_of_rat x < of_rat l" unfolding of_rat_less by auto with rai have "of_rat x < the_unique_root y" by (auto elim!: invariant_1E) with xl rai have ?thesis by (cases y, auto simp: compare_real_def comparator_of_def) } moreover { assume xr: "\ x < l" "x > r" hence "real_of_rat x > of_rat r" unfolding of_rat_less by auto with rai have "of_rat x > the_unique_root y" by (auto elim!: invariant_1E) with xr rai have ?thesis by (cases y, auto simp: compare_real_def comparator_of_def) } ultimately show ?thesis using True by auto next case False have 0: "ipoly p (real_of_rat x) \ 0" by (rule poly_cond2_no_rat_root, insert rai, auto) with rai have diff: "real_of_1 y \ of_rat x" by (auto elim!: invariant_1E) have "\ P. (1 < degree (poly_real_alg_1 y) \ \!x. root_cond y x \ poly_cond p \ P) \ P" using poly_real_alg_1.simps y rai invariant_1_2E invariant_1E by metis from this[OF gt_rat_sign_change] False have left: "compare_rat_1 x y = (if real_of_rat x \ the_unique_root y then Lt else Gt)" by (auto simp:poly_cond2_def) also have "\ = compare (real_of_rat x) (real_of_1 y)" using diff by (auto simp: compare_real_def comparator_of_def) finally show ?thesis . qed qed lemma cf_pos_0[simp]: "\ cf_pos 0" unfolding cf_pos_def by auto (* ********************* *) subsubsection\Negation\ fun uminus_1 :: "real_alg_1 \ real_alg_1" where "uminus_1 (p,l,r) = (abs_int_poly (poly_uminus p), -r, -l)" lemma uminus_1: assumes x: "invariant_1 x" defines y: "y \ uminus_1 x" shows "invariant_1 y \ (real_of_1 y = - real_of_1 x)" proof (cases x) case plr: (fields p l r) from x plr have inv: "invariant_1 (p,l,r)" by auto note * = invariant_1D[OF this] from plr have x: "x = (p,l,r)" by simp let ?p = "poly_uminus p" let ?mp = "abs_int_poly ?p" have y: "y = (?mp, -r , -l)" unfolding y plr by (simp add: Let_def) { fix y assume "root_cond (?mp, - r, - l) y" hence mpy: "ipoly ?mp y = 0" and bnd: "- of_rat r \ y" "y \ - of_rat l" unfolding root_cond_def by (auto simp: of_rat_minus) from mpy have id: "ipoly p (- y) = 0" by auto from bnd have bnd: "of_rat l \ - y" "-y \ of_rat r" by auto from id bnd have "root_cond (p, l, r) (-y)" unfolding root_cond_def by auto with inv x have "real_of_1 x = -y" by (auto intro!: the_unique_root_eqI) then have "-real_of_1 x = y" by auto } note inj = this have rc: "root_cond (?mp, - r, - l) (- real_of_1 x)" using * unfolding root_cond_def y x by (auto simp: of_rat_minus sgn_minus_rat) from inj rc have ur': "unique_root (?mp, -r, -l)" by (auto intro: unique_rootI) with rc have the: "- real_of_1 x = the_unique_root (?mp, -r, -l)" by (auto intro: the_unique_root_eqI) have xp: "p represents (real_of_1 x)" using * unfolding root_cond_def split represents_def x by auto from * have mon: "lead_coeff ?mp > 0" by (unfold pos_poly_abs_poly, auto) from poly_uminus_irreducible * have mi: "irreducible ?mp" by auto from mi mon have pc': "poly_cond ?mp" by (auto simp: cf_pos_def) from poly_condD[OF pc'] have irr: "irreducible ?mp" by auto show ?thesis unfolding y apply (intro invariant_1_realI ur' rc) using pc' inv by auto qed lemma uminus_1_2: assumes x: "invariant_1_2 x" defines y: "y \ uminus_1 x" shows "invariant_1_2 y \ (real_of_1 y = - real_of_1 x)" proof - from x have "invariant_1 x" by auto from uminus_1[OF this] have *: "real_of_1 y = - real_of_1 x" "invariant_1 y" unfolding y by auto obtain p l r where id: "x = (p,l,r)" by (cases x) from x[unfolded id] have "degree p > 1" by auto moreover have "poly_real_alg_1 y = abs_int_poly (poly_uminus p)" unfolding y id uminus_1.simps split Let_def by auto ultimately have "degree (poly_real_alg_1 y) > 1" by simp with * show ?thesis by auto qed fun uminus_2 :: "real_alg_2 \ real_alg_2" where "uminus_2 (Rational r) = Rational (-r)" | "uminus_2 (Irrational n x) = real_alg_2 (uminus_1 x)" lemma uminus_2: assumes "invariant_2 x" shows "real_of_2 (uminus_2 x) = uminus (real_of_2 x)" "invariant_2 (uminus_2 x)" using assms real_alg_2 uminus_1 by (atomize(full), cases x, auto simp: hom_distribs) declare uminus_1.simps[simp del] lift_definition uminus_3 :: "real_alg_3 \ real_alg_3" is uminus_2 by (auto simp: uminus_2) lemma uminus_3: "real_of_3 (uminus_3 x) = - real_of_3 x" by (transfer, auto simp: uminus_2) instantiation real_alg :: uminus begin lift_definition uminus_real_alg :: "real_alg \ real_alg" is uminus_3 by (simp add: uminus_3) instance .. end lemma uminus_real_alg: "- (real_of x) = real_of (- x)" by (transfer, rule uminus_3[symmetric]) (* ********************* *) subsubsection\Inverse\ fun inverse_1 :: "real_alg_1 \ real_alg_2" where "inverse_1 (p,l,r) = real_alg_2 (abs_int_poly (reflect_poly p), inverse r, inverse l)" lemma invariant_1_2_of_rat: assumes rc: "invariant_1_2 rai" shows "real_of_1 rai \ of_rat x" proof - obtain p l r where rai: "rai = (p, l, r)" by (cases rai, auto) from rc[unfolded rai] have "poly_cond2 p" "ipoly p (the_unique_root (p, l, r)) = 0" by (auto elim!: invariant_1E) from poly_cond2_no_rat_root[OF this(1), of x] this(2) show ?thesis unfolding rai by auto qed lemma inverse_1: assumes rcx: "invariant_1_2 x" defines y: "y \ inverse_1 x" shows "invariant_2 y \ (real_of_2 y = inverse (real_of_1 x))" proof (cases x) case x: (fields p l r) from x rcx have rcx: "invariant_1_2 (p,l,r)" by auto from invariant_1_2_poly_cond2[OF rcx] have pc2: "poly_cond2 p" by simp have x0: "real_of_1 (p,l,r) \ 0" using invariant_1_2_of_rat[OF rcx, of 0] x by auto let ?x = "real_of_1 (p,l,r)" let ?mp = "abs_int_poly (reflect_poly p)" from x0 rcx have lr0: "l \ 0" and "r \ 0" by auto from x0 rcx have y: "y = real_alg_2 (?mp, inverse r, inverse l)" unfolding y x Let_def inverse_1.simps by auto from rcx have mon: "lead_coeff ?mp > 0" by (unfold lead_coeff_abs_int_poly, auto) { fix y assume "root_cond (?mp, inverse r, inverse l) y" hence mpy: "ipoly ?mp y = 0" and bnd: "inverse (of_rat r) \ y" "y \ inverse (of_rat l)" unfolding root_cond_def by (auto simp: of_rat_inverse) from sgn_real_mono[OF bnd(1)] sgn_real_mono[OF bnd(2)] have "sgn (of_rat r) \ sgn y" "sgn y \ sgn (of_rat l)" by (simp_all add: algebra_simps) with rcx have sgn: "sgn (inverse (of_rat r)) = sgn y" "sgn y = sgn (inverse (of_rat l))" unfolding sgn_inverse inverse_sgn by (auto simp add: real_of_rat_sgn intro: order_antisym) from sgn[simplified, unfolded real_of_rat_sgn] lr0 have "y \ 0" by (auto simp:sgn_0_0) with mpy have id: "ipoly p (inverse y) = 0" by (auto simp: ipoly_reflect_poly) from inverse_le_sgn[OF sgn(1) bnd(1)] inverse_le_sgn[OF sgn(2) bnd(2)] have bnd: "of_rat l \ inverse y" "inverse y \ of_rat r" by auto from id bnd have "root_cond (p,l,r) (inverse y)" unfolding root_cond_def by auto from rcx this x0 have "?x = inverse y" by auto then have "inverse ?x = y" by auto } note inj = this have rc: "root_cond (?mp, inverse r, inverse l) (inverse ?x)" using rcx x0 apply (elim invariant_1_2E invariant_1E) by (simp add: root_cond_def of_rat_inverse real_of_rat_sgn inverse_le_iff_sgn ipoly_reflect_poly) from inj rc have ur: "unique_root (?mp, inverse r, inverse l)" by (auto intro: unique_rootI) with rc have the: "the_unique_root (?mp, inverse r, inverse l) = inverse ?x" by (auto intro: the_unique_root_eqI) have xp: "p represents ?x" unfolding split represents_def using rcx by (auto elim!: invariant_1E) from reflect_poly_irreducible[OF _ xp x0] poly_condD rcx have mi: "irreducible ?mp" by auto from mi mon have un: "poly_cond ?mp" by (auto simp: poly_cond_def) show ?thesis using rcx rc ur unfolding y by (intro invariant_2_realI, auto simp: x y un) qed fun inverse_2 :: "real_alg_2 \ real_alg_2" where "inverse_2 (Rational r) = Rational (inverse r)" | "inverse_2 (Irrational n x) = inverse_1 x" lemma inverse_2: assumes "invariant_2 x" shows "real_of_2 (inverse_2 x) = inverse (real_of_2 x)" "invariant_2 (inverse_2 x)" using assms by (atomize(full), cases x, auto simp: real_alg_2 inverse_1 hom_distribs) lift_definition inverse_3 :: "real_alg_3 \ real_alg_3" is inverse_2 by (auto simp: inverse_2) lemma inverse_3: "real_of_3 (inverse_3 x) = inverse (real_of_3 x)" by (transfer, auto simp: inverse_2) (* ********************* *) subsubsection\Floor\ fun floor_1 :: "real_alg_1 \ int" where "floor_1 (p,l,r) = (let (l',r',sr') = tighten_poly_bounds_epsilon p (1/2) l r (sgn (ipoly p r)); fr = floor r'; fl = floor l'; fr' = rat_of_int fr in (if fr = fl then fr else let (l'',r'',sr'') = tighten_poly_bounds_for_x p fr' l' r' sr' in if fr' < l'' then fr else fl))" lemma floor_1: assumes "invariant_1_2 x" shows "floor (real_of_1 x) = floor_1 x" proof (cases x) case (fields p l r) obtain l' r' sr' where tbe: "tighten_poly_bounds_epsilon p (1 / 2) l r (sgn (ipoly p r)) = (l',r',sr')" by (cases rule: prod_cases3, auto) let ?fr = "floor r'" let ?fl = "floor l'" let ?fr' = "rat_of_int ?fr" obtain l'' r'' sr'' where tbx: "tighten_poly_bounds_for_x p ?fr' l' r' sr' = (l'',r'',sr'')" by (cases rule: prod_cases3, auto) note rc = assms[unfolded fields] hence rc1: "invariant_1 (p,l,r)" by auto have id: "floor_1 x = ((if ?fr = ?fl then ?fr else if ?fr' < l'' then ?fr else ?fl))" unfolding fields floor_1.simps tbe Let_def split tbx by simp let ?x = "real_of_1 x" have x: "?x = the_unique_root (p,l,r)" unfolding fields by simp have bnd: "l \ l'" "r' \ r" "r' - l' \ 1 / 2" and rc': "root_cond (p, l', r') (the_unique_root (p, l, r))" and sr': "sr' = sgn (ipoly p r')" by (atomize(full), intro conjI tighten_poly_bounds_epsilon[OF _ _ tbe refl],insert rc,auto elim!: invariant_1E) let ?r = real_of_rat from rc'[folded x, unfolded split] have ineq: "?r l' \ ?x" "?x \ ?r r'" "?r l' \ ?r r'" by auto hence lr': "l' \ r'" unfolding of_rat_less_eq by simp have flr: "?fl \ ?fr" by (rule floor_mono[OF lr']) from invariant_1_sub_interval[OF rc1 rc' bnd(1,2)] have rc': "invariant_1 (p, l', r')" and id': "the_unique_root (p, l', r') = the_unique_root (p, l, r)" by auto with rc have rc2': "invariant_1_2 (p, l', r')" by auto have x: "?x = the_unique_root (p,l',r')" unfolding fields using id' by simp { assume "?fr \ ?fl" with flr have flr: "?fl \ ?fr - 1" by simp have "?fr' \ r'" "l' \ ?fr'" using flr bnd by linarith+ } note fl_diff = this show ?thesis proof (cases "?fr = ?fl") case True hence id1: "floor_1 x = ?fr" unfolding id by auto - from True have id: "floor (?r l') = floor (?r r')" unfolding real_of_rat_floor by simp + from True have id: "floor (?r l') = floor (?r r')" + by simp have "floor ?x \ floor (?r r')" by (rule floor_mono[OF ineq(2)]) moreover have "floor (?r l') \ floor ?x" by (rule floor_mono[OF ineq(1)]) - ultimately have "floor ?x = floor (?r r')" unfolding id by simp - thus ?thesis unfolding id1 real_of_rat_floor . + ultimately have "floor ?x = floor (?r r')" + unfolding id by (simp add: id) + then show ?thesis by (simp add: id1) next case False with id have id: "floor_1 x = (if ?fr' < l'' then ?fr else ?fl)" by simp from rc2' have "unique_root (p,l',r')" "poly_cond2 p" by auto from tighten_poly_bounds_for_x[OF this tbx sr'] have ineq': "l' \ l''" "r'' \ r'" and lr'': "l'' \ r''" and rc'': "root_cond (p,l'',r'') ?x" and fr': "\ (l'' \ ?fr' \ ?fr' \ r'')" unfolding x by auto from rc''[unfolded split] have ineq'': "?r l'' \ ?x" "?x \ ?r r''" by auto from False have "?fr \ ?fl" by auto note fr = fl_diff[OF this] show ?thesis proof (cases "?fr' < l''") case True with id have id: "floor_1 x = ?fr" by simp have "floor ?x \ ?fr" using floor_mono[OF ineq(2)] by simp moreover from True have "?r ?fr' < ?r l''" unfolding of_rat_less . with ineq''(1) have "?r ?fr' \ ?x" by simp from floor_mono[OF this] have "?fr \ floor ?x" by simp ultimately show ?thesis unfolding id by auto next case False with id have id: "floor_1 x = ?fl" by simp from False have "l'' \ ?fr'" by auto from floor_mono[OF ineq(1)] have "?fl \ floor ?x" by simp moreover have "floor ?x \ ?fl" proof - from False fr' have fr': "r'' < ?fr'" by auto hence "floor r'' < ?fr" by linarith with floor_mono[OF ineq''(2)] have "floor ?x \ ?fr - 1" by auto also have "?fr - 1 = floor (r' - 1)" by simp also have "\ \ ?fl" by (rule floor_mono, insert bnd, auto) finally show ?thesis . qed ultimately show ?thesis unfolding id by auto qed qed qed (* ********************* *) subsubsection\Generic Factorization and Bisection Framework\ lemma card_1_Collect_ex1: assumes "card (Collect P) = 1" shows "\! x. P x" proof - from assms[unfolded card_eq_1_iff] obtain x where "Collect P = {x}" by auto thus ?thesis by (intro ex1I[of _ x], auto) qed fun sub_interval :: "rat \ rat \ rat \ rat \ bool" where "sub_interval (l,r) (l',r') = (l' \ l \ r \ r')" fun in_interval :: "rat \ rat \ real \ bool" where "in_interval (l,r) x = (of_rat l \ x \ x \ of_rat r)" definition converges_to :: "(nat \ rat \ rat) \ real \ bool" where "converges_to f x \ (\ n. in_interval (f n) x \ sub_interval (f (Suc n)) (f n)) \ (\ (eps :: real) > 0. \ n l r. f n = (l,r) \ of_rat r - of_rat l \ eps)" context fixes bnd_update :: "'a \ 'a" and bnd_get :: "'a \ rat \ rat" begin definition at_step :: "(nat \ rat \ rat) \ nat \ 'a \ bool" where "at_step f n a \ \ i. bnd_get ((bnd_update ^^ i) a) = f (n + i)" partial_function (tailrec) select_correct_factor_main :: "'a \ (int poly \ root_info)list \ (int poly \ root_info)list \ rat \ rat \ nat \ (int poly \ root_info) \ rat \ rat" where [code]: "select_correct_factor_main bnd todo old l r n = (case todo of Nil \ if n = 1 then (hd old, l, r) else let bnd' = bnd_update bnd in (case bnd_get bnd' of (l,r) \ select_correct_factor_main bnd' old [] l r 0) | Cons (p,ri) todo \ let m = root_info.l_r ri l r in if m = 0 then select_correct_factor_main bnd todo old l r n else select_correct_factor_main bnd todo ((p,ri) # old) l r (n + m))" definition select_correct_factor :: "'a \ (int poly \ root_info)list \ (int poly \ root_info) \ rat \ rat" where "select_correct_factor init polys = (case bnd_get init of (l,r) \ select_correct_factor_main init polys [] l r 0)" lemma select_correct_factor_main: assumes conv: "converges_to f x" and at: "at_step f i a" and res: "select_correct_factor_main a todo old l r n = ((q,ri_fin),(l_fin,r_fin))" and bnd: "bnd_get a = (l,r)" and ri: "\ q ri. (q,ri) \ set todo \ set old \ root_info_cond ri q" and q0: "\ q ri. (q,ri) \ set todo \ set old \ q \ 0" and ex: "\q. q \ fst ` set todo \ fst ` set old \ ipoly q x = 0" and dist: "distinct (map fst (todo @ old))" and old: "\ q ri. (q,ri) \ set old \ root_info.l_r ri l r \ 0" and un: "\ x :: real. (\q. q \ fst ` set todo \ fst ` set old \ ipoly q x = 0) \ \!q. q \ fst ` set todo \ fst ` set old \ ipoly q x = 0" and n: "n = sum_list (map (\ (q,ri). root_info.l_r ri l r) old)" shows "unique_root (q,l_fin,r_fin) \ (q,ri_fin) \ set todo \ set old \ x = the_unique_root (q,l_fin,r_fin)" proof - define orig where "orig = set todo \ set old" have orig: "set todo \ set old \ orig" unfolding orig_def by auto let ?rts = "{x :: real. \ q ri. (q,ri) \ orig \ ipoly q x = 0}" define rts where "rts = ?rts" let ?h = "\ (x,y). abs (x - y)" let ?r = real_of_rat have rts: "?rts = (\ ((\ (q,ri). {x. ipoly q x = 0}) ` set (todo @ old)))" unfolding orig_def by auto have "finite rts" unfolding rts rts_def using finite_ipoly_roots[OF q0] finite_set[of "todo @ old"] by auto hence fin: "finite (rts \ rts - Id)" by auto define diffs where "diffs = insert 1 {abs (x - y) | x y. x \ rts \ y \ rts \ x \ y}" have "finite {abs (x - y) | x y. x \ rts \ y \ rts \ x \ y}" by (rule subst[of _ _ finite, OF _ finite_imageI[OF fin, of ?h]], auto) hence diffs: "finite diffs" "diffs \ {}" unfolding diffs_def by auto define eps where "eps = Min diffs / 2" have "\ x. x \ diffs \ x > 0" unfolding diffs_def by auto with Min_gr_iff[OF diffs] have eps: "eps > 0" unfolding eps_def by auto note conv = conv[unfolded converges_to_def] from conv eps obtain N L R where N: "f N = (L,R)" "?r R - ?r L \ eps" by auto obtain pair where pair: "pair = (todo,i)" by auto define rel where "rel = measures [ \ (t,i). N - i, \ (t :: (int poly \ root_info) list,i). length t]" have wf: "wf rel" unfolding rel_def by simp show ?thesis using at res bnd ri q0 ex dist old un n pair orig proof (induct pair arbitrary: todo i old a l r n rule: wf_induct[OF wf]) case (1 pair todo i old a l r n) note IH = 1(1)[rule_format] note at = 1(2) note res = 1(3)[unfolded select_correct_factor_main.simps[of _ todo]] note bnd = 1(4) note ri = 1(5) note q0 = 1(6) note ex = 1(7) note dist = 1(8) note old = 1(9) note un = 1(10) note n = 1(11) note pair = 1(12) note orig = 1(13) from at[unfolded at_step_def, rule_format, of 0] bnd have fi: "f i = (l,r)" by auto with conv have inx: "in_interval (f i) x" by blast hence lxr: "?r l \ x" "x \ ?r r" unfolding fi by auto from order.trans[OF this] have lr: "l \ r" unfolding of_rat_less_eq . show ?case proof (cases todo) case (Cons rri tod) obtain s ri where rri: "rri = (s,ri)" by force with Cons have todo: "todo = (s,ri) # tod" by simp note res = res[unfolded todo list.simps split Let_def] from root_info_condD(1)[OF ri[of s ri, unfolded todo] lr] have ri': "root_info.l_r ri l r = card {x. root_cond (s, l, r) x}" by auto from q0 have s0: "s \ 0" unfolding todo by auto from finite_ipoly_roots[OF s0] have fins: "finite {x. root_cond (s, l, r) x}" unfolding root_cond_def by auto have rel: "((tod,i), pair) \ rel" unfolding rel_def pair todo by simp show ?thesis proof (cases "root_info.l_r ri l r = 0") case True with res have res: "select_correct_factor_main a tod old l r n = ((q, ri_fin), l_fin, r_fin)" by auto from ri'[symmetric, unfolded True] fins have empty: "{x. root_cond (s, l, r) x} = {}" by simp from ex lxr empty have ex': "(\q. q \ fst ` set tod \ fst ` set old \ ipoly q x = 0)" unfolding todo root_cond_def split by auto have "unique_root (q, l_fin, r_fin) \ (q, ri_fin) \ set tod \ set old \ x = the_unique_root (q, l_fin, r_fin)" proof (rule IH[OF rel at res bnd ri _ ex' _ _ _ n refl], goal_cases) case (5 y) thus ?case using un[of y] unfolding todo by auto next case 2 thus ?case using q0 unfolding todo by auto qed (insert dist old orig, auto simp: todo) thus ?thesis unfolding todo by auto next case False with res have res: "select_correct_factor_main a tod ((s, ri) # old) l r (n + root_info.l_r ri l r) = ((q, ri_fin), l_fin, r_fin)" by auto from ex have ex': "\q. q \ fst ` set tod \ fst ` set ((s, ri) # old) \ ipoly q x = 0" unfolding todo by auto from dist have dist: "distinct (map fst (tod @ (s, ri) # old))" unfolding todo by auto have id: "set todo \ set old = set tod \ set ((s, ri) # old)" unfolding todo by simp show ?thesis unfolding id proof (rule IH[OF rel at res bnd ri _ ex' dist], goal_cases) case 4 thus ?case using un unfolding todo by auto qed (insert old False orig, auto simp: q0 todo n) qed next case Nil note res = res[unfolded Nil list.simps Let_def] from ex[unfolded Nil] lxr obtain s where "s \ fst ` set old \ root_cond (s,l,r) x" unfolding root_cond_def by auto then obtain q1 ri1 old' where old': "old = (q1,ri1) # old'" using id by (cases old, auto) let ?ri = "root_info.l_r ri1 l r" from old[unfolded old'] have 0: "?ri \ 0" by auto from n[unfolded old'] 0 have n0: "n \ 0" by auto from ri[unfolded old'] have ri': "root_info_cond ri1 q1" by auto show ?thesis proof (cases "n = 1") case False with n0 have n1: "n > 1" by auto obtain l' r' where bnd': "bnd_get (bnd_update a) = (l',r')" by force with res False have res: "select_correct_factor_main (bnd_update a) old [] l' r' 0 = ((q, ri_fin), l_fin, r_fin)" by auto have at': "at_step f (Suc i) (bnd_update a)" unfolding at_step_def proof (intro allI, goal_cases) case (1 n) have id: "(bnd_update ^^ Suc n) a = (bnd_update ^^ n) (bnd_update a)" by (induct n, auto) from at[unfolded at_step_def, rule_format, of "Suc n"] show ?case unfolding id by simp qed from 0[unfolded root_info_condD(1)[OF ri' lr]] obtain y1 where y1: "root_cond (q1,l,r) y1" by (cases "Collect (root_cond (q1, l, r)) = {}", auto) from n1[unfolded n old'] have "?ri > 1 \ sum_list (map (\ (q,ri). root_info.l_r ri l r) old') \ 0" by (cases "sum_list (map (\ (q,ri). root_info.l_r ri l r) old')", auto) hence "\ q2 ri2 y2. (q2,ri2) \ set old \ root_cond (q2,l,r) y2 \ y1 \ y2" proof assume "?ri > 1" with root_info_condD(1)[OF ri' lr] have "card {x. root_cond (q1, l, r) x} > 1" by simp from card_gt_1D[OF this] y1 obtain y2 where "root_cond (q1,l,r) y2" and "y1 \ y2" by auto thus ?thesis unfolding old' by auto next assume "sum_list (map (\ (q,ri). root_info.l_r ri l r) old') \ 0" then obtain q2 ri2 where mem: "(q2,ri2) \ set old'" and ri2: "root_info.l_r ri2 l r \ 0" by auto with q0 ri have "root_info_cond ri2 q2" unfolding old' by auto from ri2[unfolded root_info_condD(1)[OF this lr]] obtain y2 where y2: "root_cond (q2,l,r) y2" by (cases "Collect (root_cond (q2, l, r)) = {}", auto) from dist[unfolded old'] split_list[OF mem] have diff: "q1 \ q2" by auto from y1 have q1: "q1 \ fst ` set todo \ fst ` set old \ ipoly q1 y1 = 0" unfolding old' root_cond_def by auto from y2 have q2: "q2 \ fst ` set todo \ fst ` set old \ ipoly q2 y2 = 0" unfolding old' root_cond_def using mem by force have "y1 \ y2" proof assume id: "y1 = y2" from q1 have "\ q1. q1 \ fst ` set todo \ fst ` set old \ ipoly q1 y1 = 0" by blast from un[OF this] q1 q2[folded id] have "q1 = q2" by auto with diff show False by simp qed with mem y2 show ?thesis unfolding old' by auto qed then obtain q2 ri2 y2 where mem2: "(q2,ri2) \ set old" and y2: "root_cond (q2,l,r) y2" and diff: "y1 \ y2" by auto from mem2 orig have "(q1,ri1) \ orig" "(q2,ri2) \ orig" unfolding old' by auto with y1 y2 diff have "abs (y1 - y2) \ diffs" unfolding diffs_def rts_def root_cond_def by auto from Min_le[OF diffs(1) this] have "abs (y1 - y2) \ 2 * eps" unfolding eps_def by auto with eps have eps: "abs (y1 - y2) > eps" by auto from y1 y2 have l: "of_rat l \ min y1 y2" unfolding root_cond_def by auto from y1 y2 have r: "of_rat r \ max y1 y2" unfolding root_cond_def by auto from l r eps have eps: "of_rat r - of_rat l > eps" by auto have "i < N" proof (rule ccontr) assume "\ i < N" hence "\ k. i = N + k" by presburger then obtain k where i: "i = N + k" by auto { fix k l r assume "f (N + k) = (l,r)" hence "of_rat r - of_rat l \ eps" proof (induct k arbitrary: l r) case 0 with N show ?case by auto next case (Suc k l r) obtain l' r' where f: "f (N + k) = (l',r')" by force from Suc(1)[OF this] have IH: "?r r' - ?r l' \ eps" by auto from f Suc(2) conv[THEN conjunct1, rule_format, of "N + k"] have "?r l \ ?r l'" "?r r \ ?r r'" by (auto simp: of_rat_less_eq) thus ?case using IH by auto qed } note * = this from at[unfolded at_step_def i, rule_format, of 0] bnd have "f (N + k) = (l,r)" by auto from *[OF this] eps show False by auto qed hence rel: "((old, Suc i), pair) \ rel" unfolding pair rel_def by auto from dist have dist: "distinct (map fst (old @ []))" unfolding Nil by auto have id: "set todo \ set old = set old \ set []" unfolding Nil by auto show ?thesis unfolding id proof (rule IH[OF rel at' res bnd' ri _ _ dist _ _ _ refl], goal_cases) case 2 thus ?case using q0 by auto qed (insert ex un orig Nil, auto) next case True with res old' have id: "q = q1" "ri_fin = ri1" "l_fin = l" "r_fin = r" by auto from n[unfolded True old'] 0 have 1: "?ri = 1" by (cases ?ri; cases "?ri - 1", auto) from root_info_condD(1)[OF ri' lr] 1 have "card {x. root_cond (q1,l,r) x} = 1" by auto from card_1_Collect_ex1[OF this] have unique: "unique_root (q1,l,r)" . from ex[unfolded Nil old'] consider (A) "ipoly q1 x = 0" | (B) q where "q \ fst ` set old'" "ipoly q x = 0" by auto hence "x = the_unique_root (q1,l,r)" proof (cases) case A with lxr have "root_cond (q1,l,r) x" unfolding root_cond_def by auto from the_unique_root_eqI[OF unique this] show ?thesis by simp next case (B q) with lxr have "root_cond (q,l,r) x" unfolding root_cond_def by auto hence empty: "{x. root_cond (q,l,r) x} \ {}" by auto from B(1) obtain ri' where mem: "(q,ri') \ set old'" by force from q0[unfolded old'] mem have q0: "q \ 0" by auto from finite_ipoly_roots[OF this] have "finite {x. root_cond (q,l,r) x}" unfolding root_cond_def by auto with empty have card: "card {x. root_cond (q,l,r) x} \ 0" by simp from ri[unfolded old'] mem have "root_info_cond ri' q" by auto from root_info_condD(1)[OF this lr] card have "root_info.l_r ri' l r \ 0" by auto with n[unfolded True old'] 1 split_list[OF mem] have False by auto thus ?thesis by simp qed thus ?thesis unfolding id using unique ri' unfolding old' by auto qed qed qed qed lemma select_correct_factor: assumes conv: "converges_to (\ i. bnd_get ((bnd_update ^^ i) init)) x" and res: "select_correct_factor init polys = ((q,ri),(l,r))" and ri: "\ q ri. (q,ri) \ set polys \ root_info_cond ri q" and q0: "\ q ri. (q,ri) \ set polys \ q \ 0" and ex: "\q. q \ fst ` set polys \ ipoly q x = 0" and dist: "distinct (map fst polys)" and un: "\ x :: real. (\q. q \ fst ` set polys \ ipoly q x = 0) \ \!q. q \ fst ` set polys \ ipoly q x = 0" shows "unique_root (q,l,r) \ (q,ri) \ set polys \ x = the_unique_root (q,l,r)" proof - obtain l' r' where init: "bnd_get init = (l',r')" by force from res[unfolded select_correct_factor_def init split] have res: "select_correct_factor_main init polys [] l' r' 0 = ((q, ri), l, r)" by auto have at: "at_step (\ i. bnd_get ((bnd_update ^^ i) init)) 0 init" unfolding at_step_def by auto have "unique_root (q,l,r) \ (q,ri) \ set polys \ set [] \ x = the_unique_root (q,l,r)" by (rule select_correct_factor_main[OF conv at res init ri], insert dist un ex q0, auto) thus ?thesis by auto qed definition real_alg_2' :: "root_info \ int poly \ rat \ rat \ real_alg_2" where [code del]: "real_alg_2' ri p l r = ( if degree p = 1 then Rational (Rat.Fract (- coeff p 0) (coeff p 1)) else real_alg_2_main ri (case tighten_poly_bounds_for_x p 0 l r (sgn (ipoly p r)) of (l',r',sr') \ (p, l', r')))" lemma real_alg_2'_code[code]: "real_alg_2' ri p l r = (if degree p = 1 then Rational (Rat.Fract (- coeff p 0) (coeff p 1)) else case normalize_bounds_1 (case tighten_poly_bounds_for_x p 0 l r (sgn (ipoly p r)) of (l', r', sr') \ (p, l', r')) of (p', l, r) \ Irrational (root_info.number_root ri r) (p', l, r))" unfolding real_alg_2'_def real_alg_2_main_def by (cases "tighten_poly_bounds_for_x p 0 l r (sgn (ipoly p r))", simp add: Let_def) definition real_alg_2'' :: "root_info \ int poly \ rat \ rat \ real_alg_2" where "real_alg_2'' ri p l r = (case normalize_bounds_1 (case tighten_poly_bounds_for_x p 0 l r (sgn (ipoly p r)) of (l', r', sr') \ (p, l', r')) of (p', l, r) \ Irrational (root_info.number_root ri r) (p', l, r))" lemma real_alg_2'': "degree p \ 1 \ real_alg_2'' ri p l r = real_alg_2' ri p l r" unfolding real_alg_2'_code real_alg_2''_def by auto lemma poly_cond_degree_0_imp_no_root: fixes x :: "'b :: {comm_ring_1,ring_char_0}" assumes pc: "poly_cond p" and deg: "degree p = 0" shows "ipoly p x \ 0" proof from pc have "p \ 0" by auto moreover assume "ipoly p x = 0" note poly_zero[OF this] ultimately show False using deg by auto qed lemma real_alg_2': assumes ur: "unique_root (q,l,r)" and pc: "poly_cond q" and ri: "root_info_cond ri q" shows "invariant_2 (real_alg_2' ri q l r) \ real_of_2 (real_alg_2' ri q l r) = the_unique_root (q,l,r)" (is "_ \ _ = ?x") proof (cases "degree q" "Suc 0" rule: linorder_cases) case deg: less then have "degree q = 0" by auto from poly_cond_degree_0_imp_no_root[OF pc this] ur have False by force then show ?thesis by auto next case deg: equal hence id: "real_alg_2' ri q l r = Rational (Rat.Fract (- coeff q 0) (coeff q 1))" unfolding real_alg_2'_def by auto show ?thesis unfolding id using degree_1_ipoly[OF deg] using unique_rootD(4)[OF ur] by auto next case deg: greater with pc have pc2: "poly_cond2 q" by auto let ?rai = "real_alg_2' ri q l r" let ?r = real_of_rat obtain l' r' sr' where tight: "tighten_poly_bounds_for_x q 0 l r (sgn (ipoly q r)) = (l',r',sr')" by (cases rule: prod_cases3, auto) let ?rai' = "(q, l', r')" have rai': "?rai = real_alg_2_main ri ?rai'" unfolding real_alg_2'_def using deg tight by auto hence rai: "real_of_1 ?rai' = the_unique_root (q,l',r')" by auto note tight = tighten_poly_bounds_for_x[OF ur pc2 tight refl] let ?x = "the_unique_root (q, l, r)" from tight have tight: "root_cond (q,l',r') ?x" "l \ l'" "l' \ r'" "r' \ r" "l' > 0 \ r' < 0" by auto from unique_root_sub_interval[OF ur tight(1) tight(2,4)] poly_condD[OF pc] have ur': "unique_root (q, l', r')" and x: "?x = the_unique_root (q,l',r')" by auto from tight(2-) have sgn: "sgn l' = sgn r'" by auto show ?thesis unfolding rai' using real_alg_2_main[of ?rai' ri] invariant_1_realI[of ?rai' ?x] by (auto simp: tight(1) sgn pc ri ur') qed definition select_correct_factor_int_poly :: "'a \ int poly \ real_alg_2" where "select_correct_factor_int_poly init p \ let qs = factors_of_int_poly p; polys = map (\ q. (q, root_info q)) qs; ((q,ri),(l,r)) = select_correct_factor init polys in real_alg_2' ri q l r" lemma select_correct_factor_int_poly: assumes conv: "converges_to (\ i. bnd_get ((bnd_update ^^ i) init)) x" and rai: "select_correct_factor_int_poly init p = rai" and x: "ipoly p x = 0" and p: "p \ 0" shows "invariant_2 rai \ real_of_2 rai = x" proof - obtain qs where fact: "factors_of_int_poly p = qs" by auto define polys where "polys = map (\ q. (q, root_info q)) qs" obtain q ri l r where res: "select_correct_factor init polys = ((q,ri),(l,r))" by (cases "select_correct_factor init polys", auto) have fst: "map fst polys = qs" "fst ` set polys = set qs" unfolding polys_def map_map o_def by force+ note fact' = factors_of_int_poly[OF fact] note rai = rai[unfolded select_correct_factor_int_poly_def Let_def fact, folded polys_def, unfolded res split] from fact' fst have dist: "distinct (map fst polys)" by auto from fact'(2)[OF p, of x] x fst have ex: "\q. q \ fst ` set polys \ ipoly q x = 0" by auto { fix q ri assume "(q,ri) \ set polys" hence ri: "ri = root_info q" and q: "q \ set qs" unfolding polys_def by auto from fact'(1)[OF q] have *: "lead_coeff q > 0" "irreducible q" "degree q > 0" by auto from * have q0: "q \ 0" by auto from root_info[OF *(2-3)] ri have ri: "root_info_cond ri q" by auto note ri q0 * } note polys = this have "unique_root (q, l, r) \ (q, ri) \ set polys \ x = the_unique_root (q, l, r)" by (rule select_correct_factor[OF conv res polys(1) _ ex dist, unfolded fst, OF _ _ fact'(3)[OF p]], insert fact'(2)[OF p] polys(2), auto) hence ur: "unique_root (q,l,r)" and mem: "(q,ri) \ set polys" and x: "x = the_unique_root (q,l,r)" by auto note polys = polys[OF mem] from polys(3-4) have ty: "poly_cond q" by (simp add: poly_cond_def) show ?thesis unfolding x rai[symmetric] by (intro real_alg_2' ur ty polys(1)) qed end (* ********************* *) subsubsection\Addition\ lemma ipoly_0_0[simp]: "ipoly f (0::'a::{comm_ring_1,ring_char_0}) = 0 \ poly f 0 = 0" unfolding poly_0_coeff_0 by simp lemma add_rat_roots_below[simp]: "roots_below (poly_add_rat r p) x = (\y. y + of_rat r) ` roots_below p (x - of_rat r)" proof (unfold add_rat_roots image_def, intro Collect_eqI, goal_cases) case (1 y) then show ?case by (auto intro: exI[of _ "y - real_of_rat r"]) qed lemma add_rat_root_cond: shows "root_cond (cf_pos_poly (poly_add_rat m p),l,r) x = root_cond (p, l - m, r - m) (x - of_rat m)" by (unfold root_cond_def, auto simp add: add_rat_roots hom_distribs) lemma add_rat_unique_root: "unique_root (cf_pos_poly (poly_add_rat m p), l, r) = unique_root (p, l-m, r-m)" by (auto simp: add_rat_root_cond) fun add_rat_1 :: "rat \ real_alg_1 \ real_alg_1" where "add_rat_1 r1 (p2,l2,r2) = ( let p = cf_pos_poly (poly_add_rat r1 p2); (l,r,sr) = tighten_poly_bounds_for_x p 0 (l2+r1) (r2+r1) (sgn (ipoly p (r2+r1))) in (p,l,r))" lemma poly_real_alg_1_add_rat[simp]: "poly_real_alg_1 (add_rat_1 r y) = cf_pos_poly (poly_add_rat r (poly_real_alg_1 y))" by (cases y, auto simp: Let_def split: prod.split) lemma sgn_cf_pos: assumes "lead_coeff p > 0" shows "sgn (ipoly (cf_pos_poly p) (x::'a::linordered_field)) = sgn (ipoly p x)" proof (cases "p = 0") case True with assms show ?thesis by auto next case False from cf_pos_poly_main False obtain d where p': "Polynomial.smult d (cf_pos_poly p) = p" by auto have "d > 0" proof (rule zero_less_mult_pos2) from False assms have "0 < lead_coeff p" by (auto simp: cf_pos_def) also from p' have "\ = d * lead_coeff (cf_pos_poly p)" by (metis lead_coeff_smult) finally show "0 < \". show "lead_coeff (cf_pos_poly p) > 0" using False by (unfold lead_coeff_cf_pos_poly) qed moreover from p' have "ipoly p x = of_int d * ipoly (cf_pos_poly p) x" by (fold poly_smult of_int_hom.map_poly_hom_smult, auto) ultimately show ?thesis by (auto simp: sgn_mult[where 'a='a]) qed lemma add_rat_1: fixes r1 :: rat assumes inv_y: "invariant_1_2 y" defines "z \ add_rat_1 r1 y" shows "invariant_1_2 z \ (real_of_1 z = of_rat r1 + real_of_1 y)" proof (cases y) case y_def: (fields p2 l2 r2) define p where "p \ cf_pos_poly (poly_add_rat r1 p2)" obtain l r sr where lr: "tighten_poly_bounds_for_x p 0 (l2+r1) (r2+r1) (sgn (ipoly p (r2+r1))) = (l,r,sr)" by (metis surj_pair) from lr have z: "z = (p,l,r)" by (auto simp: y_def z_def p_def Let_def) from inv_y have ur: "unique_root (p, l2 + r1, r2 + r1)" by (auto simp: p_def add_rat_root_cond y_def add_rat_unique_root) from inv_y[unfolded y_def invariant_1_2_def,simplified] have pc2: "poly_cond2 p" unfolding p_def apply (intro poly_cond2I poly_add_rat_irreducible poly_condI, unfold lead_coeff_cf_pos_poly) apply (auto elim!: invariant_1E) done note main = tighten_poly_bounds_for_x[OF ur pc2 lr refl, simplified] then have "sgn l = sgn r" unfolding sgn_if apply simp apply linarith done from invariant_1_2_realI[OF main(4) _ main(7), simplified, OF this pc2] main(1-3) ur show ?thesis by (auto simp: z p_def y_def add_rat_root_cond ex1_the_shift) qed fun tighten_poly_bounds_binary :: "int poly \ int poly \ (rat \ rat \ rat) \ rat \ rat \ rat \ (rat \ rat \ rat) \ rat \ rat \ rat" where "tighten_poly_bounds_binary cr1 cr2 ((l1,r1,sr1),(l2,r2,sr2)) = (tighten_poly_bounds cr1 l1 r1 sr1, tighten_poly_bounds cr2 l2 r2 sr2)" lemma tighten_poly_bounds_binary: assumes ur: "unique_root (p1,l1,r1)" "unique_root (p2,l2,r2)" and pt: "poly_cond2 p1" "poly_cond2 p2" defines "x \ the_unique_root (p1,l1,r1)" and "y \ the_unique_root (p2,l2,r2)" assumes bnd: "\ l1 r1 l2 r2 l r sr1 sr2. I l1 \ I l2 \ root_cond (p1,l1,r1) x \ root_cond (p2,l2,r2) y \ bnd ((l1,r1,sr1),(l2,r2,sr2)) = (l,r) \ of_rat l \ f x y \ f x y \ of_rat r" and approx: "\ l1 r1 l2 r2 l1' r1' l2' r2' l l' r r' sr1 sr2 sr1' sr2'. I l1 \ I l2 \ l1 \ r1 \ l2 \ r2 \ (l,r) = bnd ((l1,r1,sr1), (l2,r2,sr2)) \ (l',r') = bnd ((l1',r1',sr1'), (l2',r2',sr2')) \ (l1',r1') \ {(l1,(l1+r1)/2),((l1+r1)/2,r1)} \ (l2',r2') \ {(l2,(l2+r2)/2),((l2+r2)/2,r2)} \ (r' - l') \ 3/4 * (r - l) \ l \ l' \ r' \ r" and I_mono: "\ l l'. I l \ l \ l' \ I l'" and I: "I l1" "I l2" and sr: "sr1 = sgn (ipoly p1 r1)" "sr2 = sgn (ipoly p2 r2)" shows "converges_to (\ i. bnd ((tighten_poly_bounds_binary p1 p2 ^^ i) ((l1,r1,sr1),(l2,r2,sr2)))) (f x y)" proof - let ?upd = "tighten_poly_bounds_binary p1 p2" define upd where "upd = ?upd" define init where "init = ((l1, r1, sr1), l2, r2, sr2)" let ?g = "(\i. bnd ((upd ^^ i) init))" obtain l r where bnd_init: "bnd init = (l,r)" by force note ur1 = unique_rootD[OF ur(1)] note ur2 = unique_rootD[OF ur(2)] from ur1(4) ur2(4) x_def y_def have rc1: "root_cond (p1,l1,r1) x" and rc2: "root_cond (p2,l2,r2) y" by auto define g where "g = ?g" { fix i L1 R1 L2 R2 L R j SR1 SR2 assume "((upd ^^ i)) init = ((L1,R1,SR1),(L2,R2,SR2))" "g i = (L,R)" hence "I L1 \ I L2 \ root_cond (p1,L1,R1) x \ root_cond (p2,L2,R2) y \ unique_root (p1, L1, R1) \ unique_root (p2, L2, R2) \ in_interval (L,R) (f x y) \ (i = Suc j \ sub_interval (g i) (g j) \ (R - L \ 3/4 * (snd (g j) - fst (g j)))) \ SR1 = sgn (ipoly p1 R1) \ SR2 = sgn (ipoly p2 R2)" proof (induct i arbitrary: L1 R1 L2 R2 L R j SR1 SR2) case 0 thus ?case using I rc1 rc2 ur bnd[of l1 l2 r1 r2 sr1 sr2 L R] g_def sr unfolding init_def by auto next case (Suc i) obtain l1 r1 l2 r2 sr1 sr2 where updi: "(upd ^^ i) init = ((l1, r1, sr1), l2, r2, sr2)" by (cases "(upd ^^ i) init", auto) obtain l r where bndi: "bnd ((l1, r1, sr1), l2, r2, sr2) = (l,r)" by force hence gi: "g i = (l,r)" using updi unfolding g_def by auto have "(upd ^^ Suc i) init = upd ((l1, r1, sr1), l2, r2, sr2)" using updi by simp from Suc(2)[unfolded this] have upd: "upd ((l1, r1, sr1), l2, r2, sr2) = ((L1, R1, SR1), L2, R2, SR2)" . from upd updi Suc(3) have bndsi: "bnd ((L1, R1, SR1), L2, R2, SR2) = (L,R)" by (auto simp: g_def) from Suc(1)[OF updi gi] have I: "I l1" "I l2" and rc: "root_cond (p1,l1,r1) x" "root_cond (p2,l2,r2) y" and ur: "unique_root (p1, l1, r1)" "unique_root (p2, l2, r2)" and sr: "sr1 = sgn (ipoly p1 r1)" "sr2 = sgn (ipoly p2 r2)" by auto from upd[unfolded upd_def] have tight: "tighten_poly_bounds p1 l1 r1 sr1 = (L1, R1, SR1)" "tighten_poly_bounds p2 l2 r2 sr2 = (L2, R2, SR2)" by auto note tight1 = tighten_poly_bounds[OF tight(1) ur(1) pt(1) sr(1)] note tight2 = tighten_poly_bounds[OF tight(2) ur(2) pt(2) sr(2)] from tight1 have lr1: "l1 \ r1" by auto from tight2 have lr2: "l2 \ r2" by auto note ur1 = unique_rootD[OF ur(1)] note ur2 = unique_rootD[OF ur(2)] from tight1 I_mono[OF I(1)] have I1: "I L1" by auto from tight2 I_mono[OF I(2)] have I2: "I L2" by auto note ur1 = unique_root_sub_interval[OF ur(1) tight1(1,2,4)] note ur2 = unique_root_sub_interval[OF ur(2) tight2(1,2,4)] from rc(1) ur ur1 have x: "x = the_unique_root (p1,L1,R1)" by (auto intro!:the_unique_root_eqI) from rc(2) ur ur2 have y: "y = the_unique_root (p2,L2,R2)" by (auto intro!:the_unique_root_eqI) from unique_rootD[OF ur1(1)] x have x: "root_cond (p1,L1,R1) x" by auto from unique_rootD[OF ur2(1)] y have y: "root_cond (p2,L2,R2) y" by auto from tight(1) have half1: "(L1, R1) \ {(l1, (l1 + r1) / 2), ((l1 + r1) / 2, r1)}" unfolding tighten_poly_bounds_def Let_def by (auto split: if_splits) from tight(2) have half2: "(L2, R2) \ {(l2, (l2 + r2) / 2), ((l2 + r2) / 2, r2)}" unfolding tighten_poly_bounds_def Let_def by (auto split: if_splits) from approx[OF I lr1 lr2 bndi[symmetric] bndsi[symmetric] half1 half2] have "R - L \ 3 / 4 * (r - l) \ l \ L \ R \ r" . hence "sub_interval (g (Suc i)) (g i)" "R - L \ 3/4 * (snd (g i) - fst (g i))" unfolding gi Suc(3) by auto with bnd[OF I1 I2 x y bndsi] show ?case using I1 I2 x y ur1 ur2 tight1(6) tight2(6) by auto qed } note invariants = this define L where "L = (\ i. fst (g i))" define R where "R = (\ i. snd (g i))" { fix i obtain l1 r1 l2 r2 sr1 sr2 where updi: "(upd ^^ i) init = ((l1, r1, sr1), l2, r2, sr2)" by (cases "(upd ^^ i) init", auto) obtain l r where bnd': "bnd ((l1, r1, sr1), l2, r2, sr2) = (l,r)" by force have gi: "g i = (l,r)" unfolding g_def updi bnd' by auto hence id: "l = L i" "r = R i" unfolding L_def R_def by auto from invariants[OF updi gi[unfolded id]] have "in_interval (L i, R i) (f x y)" "\ j. i = Suc j \ sub_interval (g i) (g j) \ R i - L i \ 3 / 4 * (R j - L j)" unfolding L_def R_def by auto } note * = this { fix i from *(1)[of i] *(2)[of "Suc i", OF refl] have "in_interval (g i) (f x y)" "sub_interval (g (Suc i)) (g i)" "R (Suc i) - L (Suc i) \ 3 / 4 * (R i - L i)" unfolding L_def R_def by auto } note * = this show ?thesis unfolding upd_def[symmetric] init_def[symmetric] g_def[symmetric] unfolding converges_to_def proof (intro conjI allI impI, rule *(1), rule *(2)) fix eps :: real assume eps: "0 < eps" let ?r = real_of_rat define r where "r = (\ n. ?r (R n))" define l where "l = (\ n. ?r (L n))" define diff where "diff = (\ n. r n - l n)" { fix n from *(3)[of n] have "?r (R (Suc n) - L (Suc n)) \ ?r (3 / 4 * (R n - L n))" unfolding of_rat_less_eq by simp also have "?r (R (Suc n) - L (Suc n)) = (r (Suc n) - l (Suc n))" unfolding of_rat_diff r_def l_def by simp also have "?r (3 / 4 * (R n - L n)) = 3 / 4 * (r n - l n)" unfolding r_def l_def by (simp add: hom_distribs) finally have "diff (Suc n) \ 3 / 4 * diff n" unfolding diff_def . } note * = this { fix i have "diff i \ (3/4)^i * diff 0" proof (induct i) case (Suc i) from Suc *[of i] show ?case by auto qed auto } then obtain c where *: "\ i. diff i \ (3/4)^i * c" by auto have "\ n. diff n \ eps" proof (cases "c \ 0") case True with *[of 0] eps show ?thesis by (intro exI[of _ 0], auto) next case False hence c: "c > 0" by auto with eps have "inverse c * eps > 0" by auto from exp_tends_to_zero[of "3/4 :: real", OF _ _ this] obtain n where "(3/4) ^ n \ inverse c * eps" by auto from mult_right_mono[OF this, of c] c have "(3/4) ^ n * c \ eps" by (auto simp: field_simps) with *[of n] show ?thesis by (intro exI[of _ n], auto) qed then obtain n where "?r (R n) - ?r (L n) \ eps" unfolding l_def r_def diff_def by blast thus "\n l r. g n = (l, r) \ ?r r - ?r l \ eps" unfolding L_def R_def by (intro exI[of _ n], force) qed qed fun add_1 :: "real_alg_1 \ real_alg_1 \ real_alg_2" where "add_1 (p1,l1,r1) (p2,l2,r2) = ( select_correct_factor_int_poly (tighten_poly_bounds_binary p1 p2) (\ ((l1,r1,sr1),(l2,r2,sr2)). (l1 + l2, r1 + r2)) ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))) (poly_add p1 p2))" lemma add_1: assumes x: "invariant_1_2 x" and y: "invariant_1_2 y" defines z: "z \ add_1 x y" shows "invariant_2 z \ (real_of_2 z = real_of_1 x + real_of_1 y)" proof (cases x) case xt: (fields p1 l1 r1) show ?thesis proof (cases y) case yt: (fields p2 l2 r2) let ?x = "real_of_1 (p1, l1, r1)" let ?y = "real_of_1 (p2, l2, r2)" let ?p = "poly_add p1 p2" note x = x[unfolded xt] note y = y[unfolded yt] from x have ax: "p1 represents ?x" unfolding represents_def by (auto elim!: invariant_1E) from y have ay: "p2 represents ?y" unfolding represents_def by (auto elim!: invariant_1E) let ?bnd = "(\((l1, r1, sr1 :: rat), l2 :: rat, r2 :: rat, sr2 :: rat). (l1 + l2, r1 + r2))" define bnd where "bnd = ?bnd" have "invariant_2 z \ real_of_2 z = ?x + ?y" proof (intro select_correct_factor_int_poly) from represents_add[OF ax ay] show "?p \ 0" "ipoly ?p (?x + ?y) = 0" by auto from z[unfolded xt yt] show sel: "select_correct_factor_int_poly (tighten_poly_bounds_binary p1 p2) bnd ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))) (poly_add p1 p2) = z" by (auto simp: bnd_def) have ur1: "unique_root (p1,l1,r1)" "poly_cond2 p1" using x by auto have ur2: "unique_root (p2,l2,r2)" "poly_cond2 p2" using y by auto show "converges_to (\i. bnd ((tighten_poly_bounds_binary p1 p2 ^^ i) ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))))) (?x + ?y)" by (intro tighten_poly_bounds_binary ur1 ur2; force simp: bnd_def hom_distribs) qed thus ?thesis unfolding xt yt . qed qed declare add_rat_1.simps[simp del] declare add_1.simps[simp del] (* ********************* *) subsubsection\Multiplication\ context begin private fun mult_rat_1_pos :: "rat \ real_alg_1 \ real_alg_2" where "mult_rat_1_pos r1 (p2,l2,r2) = real_alg_2 (cf_pos_poly (poly_mult_rat r1 p2), l2*r1, r2*r1)" private fun mult_1_pos :: "real_alg_1 \ real_alg_1 \ real_alg_2" where "mult_1_pos (p1,l1,r1) (p2,l2,r2) = select_correct_factor_int_poly (tighten_poly_bounds_binary p1 p2) (\ ((l1,r1,sr1),(l2,r2,sr2)). (l1 * l2, r1 * r2)) ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))) (poly_mult p1 p2)" fun mult_rat_1 :: "rat \ real_alg_1 \ real_alg_2" where "mult_rat_1 x y = (if x < 0 then uminus_2 (mult_rat_1_pos (-x) y) else if x = 0 then Rational 0 else (mult_rat_1_pos x y))" fun mult_1 :: "real_alg_1 \ real_alg_1 \ real_alg_2" where "mult_1 x y = (case (x,y) of ((p1,l1,r1),(p2,l2,r2)) \ if r1 > 0 then if r2 > 0 then mult_1_pos x y else uminus_2 (mult_1_pos x (uminus_1 y)) else if r2 > 0 then uminus_2 (mult_1_pos (uminus_1 x) y) else mult_1_pos (uminus_1 x) (uminus_1 y))" lemma mult_rat_1_pos: fixes r1 :: rat assumes r1: "r1 > 0" and y: "invariant_1 y" defines z: "z \ mult_rat_1_pos r1 y" shows "invariant_2 z \ (real_of_2 z = of_rat r1 * real_of_1 y)" proof - obtain p2 l2 r2 where yt: "y = (p2,l2,r2)" by (cases y, auto) let ?x = "real_of_rat r1" let ?y = "real_of_1 (p2, l2, r2)" let ?p = "poly_mult_rat r1 p2" let ?mp = "cf_pos_poly ?p" note y = y[unfolded yt] note yD = invariant_1D[OF y] from yD r1 have p: "?p \ 0" and r10: "r1 \ 0" by auto hence mp: "?mp \ 0" by simp from yD(1) have rt: "ipoly p2 ?y = 0" and bnd: "of_rat l2 \ ?y" "?y \ of_rat r2" by auto from rt r1 have rt: "ipoly ?mp (?x * ?y) = 0" by (auto simp add: field_simps ipoly_mult_rat[OF r10]) from yD(5) have irr: "irreducible p2" unfolding represents_def using y unfolding root_cond_def split by auto from poly_mult_rat_irreducible[OF this _ r10] yD have irr: "irreducible ?mp" by simp from p have mon: "cf_pos ?mp" by auto obtain l r where lr: "l = l2 * r1" "r = r2 * r1" by force from bnd r1 have bnd: "of_rat l \ ?x * ?y" "?x * ?y \ of_rat r" unfolding lr of_rat_mult by auto with rt have rc: "root_cond (?mp,l,r) (?x * ?y)" unfolding root_cond_def by auto have ur: "unique_root (?mp,l,r)" proof (rule ex1I, rule rc) fix z assume "root_cond (?mp,l,r) z" from this[unfolded root_cond_def split] have bndz: "of_rat l \ z" "z \ of_rat r" and rt: "ipoly ?mp z = 0" by auto have "fst (quotient_of r1) \ 0" using quotient_of_div[of r1] r10 by (cases "quotient_of r1", auto) with rt have rt: "ipoly p2 (z * inverse ?x) = 0" by (auto simp: ipoly_mult_rat[OF r10]) from bndz r1 have "of_rat l2 \ z * inverse ?x" "z * inverse ?x \ of_rat r2" unfolding lr of_rat_mult by (auto simp: field_simps) with rt have "root_cond (p2,l2,r2) (z * inverse ?x)" unfolding root_cond_def by auto also note invariant_1_root_cond[OF y] finally have "?y = z * inverse ?x" by auto thus "z = ?x * ?y" using r1 by auto qed from r1 have sgnr: "sgn r = sgn r2" unfolding lr by (cases "r2 = 0"; cases "r2 < 0"; auto simp: mult_neg_pos mult_less_0_iff) from r1 have sgnl: "sgn l = sgn l2" unfolding lr by (cases "l2 = 0"; cases "l2 < 0"; auto simp: mult_neg_pos mult_less_0_iff) from the_unique_root_eqI[OF ur rc] have xy: "?x * ?y = the_unique_root (?mp,l,r)" by auto from z[unfolded yt, simplified, unfolded Let_def lr[symmetric] split] have z: "z = real_alg_2 (?mp, l, r)" by simp have yp2: "p2 represents ?y" using yD unfolding root_cond_def split represents_def by auto with irr mon have pc: "poly_cond ?mp" by (auto simp: poly_cond_def cf_pos_def) have rc: "invariant_1 (?mp, l, r)" unfolding z using yD(2) pc ur by (auto simp add: invariant_1_def ur mp sgnr sgnl) show ?thesis unfolding z using real_alg_2[OF rc] unfolding yt xy unfolding z by simp qed lemma mult_1_pos: assumes x: "invariant_1_2 x" and y: "invariant_1_2 y" defines z: "z \ mult_1_pos x y" assumes pos: "real_of_1 x > 0" "real_of_1 y > 0" shows "invariant_2 z \ (real_of_2 z = real_of_1 x * real_of_1 y)" proof - obtain p1 l1 r1 where xt: "x = (p1,l1,r1)" by (cases x, auto) obtain p2 l2 r2 where yt: "y = (p2,l2,r2)" by (cases y, auto) let ?x = "real_of_1 (p1, l1, r1)" let ?y = "real_of_1 (p2, l2, r2)" let ?r = "real_of_rat" let ?p = "poly_mult p1 p2" note x = x[unfolded xt] note y = y[unfolded yt] from x y have basic: "unique_root (p1, l1, r1)" "poly_cond2 p1" "unique_root (p2, l2, r2)" "poly_cond2 p2" by auto from basic have irr1: "irreducible p1" and irr2: "irreducible p2" by auto from x have ax: "p1 represents ?x" unfolding represents_def by (auto elim!:invariant_1E) from y have ay: "p2 represents ?y" unfolding represents_def by (auto elim!:invariant_1E) from ax ay pos[unfolded xt yt] have axy: "?p represents (?x * ?y)" by (intro represents_mult represents_irr_non_0[OF irr2], auto) from representsD[OF this] have p: "?p \ 0" and rt: "ipoly ?p (?x * ?y) = 0" . from x pos(1)[unfolded xt] have "?r r1 > 0" unfolding split by auto hence "sgn r1 = 1" unfolding sgn_rat_def by (auto split: if_splits) with x have "sgn l1 = 1" by auto hence l1_pos: "l1 > 0" unfolding sgn_rat_def by (cases "l1 = 0"; cases "l1 < 0"; auto) from y pos(2)[unfolded yt] have "?r r2 > 0" unfolding split by auto hence "sgn r2 = 1" unfolding sgn_rat_def by (auto split: if_splits) with y have "sgn l2 = 1" by auto hence l2_pos: "l2 > 0" unfolding sgn_rat_def by (cases "l2 = 0"; cases "l2 < 0"; auto) let ?bnd = "(\((l1, r1, sr1 :: rat), l2 :: rat, r2 :: rat, sr2 :: rat). (l1 * l2, r1 * r2))" define bnd where "bnd = ?bnd" obtain z' where sel: "select_correct_factor_int_poly (tighten_poly_bounds_binary p1 p2) bnd ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))) ?p = z'" by auto have main: "invariant_2 z' \ real_of_2 z' = ?x * ?y" proof (rule select_correct_factor_int_poly[OF _ sel rt p]) { fix l1 r1 l2 r2 l1' r1' l2' r2' l l' r r' :: rat let ?m1 = "(l1+r1)/2" let ?m2 = "(l2+r2)/2" define d1 where "d1 = r1 - l1" define d2 where "d2 = r2 - l2" let ?M1 = "l1 + d1/2" let ?M2 = "l2 + d2/2" assume le: "l1 > 0" "l2 > 0" "l1 \ r1" "l2 \ r2" and id: "(l, r) = (l1 * l2, r1 * r2)" "(l', r') = (l1' * l2', r1' * r2')" and mem: "(l1', r1') \ {(l1, ?m1), (?m1, r1)}" "(l2', r2') \ {(l2, ?m2), (?m2, r2)}" hence id: "l = l1 * l2" "r = (l1 + d1) * (l2 + d2)" "l' = l1' * l2'" "r' = r1' * r2'" "r1 = l1 + d1" "r2 = l2 + d2" and id': "?m1 = ?M1" "?m2 = ?M2" unfolding d1_def d2_def by (auto simp: field_simps) define l1d1 where "l1d1 = l1 + d1" from le have ge0: "d1 \ 0" "d2 \ 0" "l1 \ 0" "l2 \ 0" unfolding d1_def d2_def by auto have "4 * (r' - l') \ 3 * (r - l)" proof (cases "l1' = l1 \ r1' = ?M1 \ l2' = l2 \ r2' = ?M2") case True hence id2: "l1' = l1" "r1' = ?M1" "l2' = l2" "r2' = ?M2" by auto show ?thesis unfolding id id2 unfolding ring_distribs using ge0 by simp next case False note 1 = this show ?thesis proof (cases "l1' = l1 \ r1' = ?M1 \ l2' = ?M2 \ r2' = r2") case True hence id2: "l1' = l1" "r1' = ?M1" "l2' = ?M2" "r2' = r2" by auto show ?thesis unfolding id id2 unfolding ring_distribs using ge0 by simp next case False note 2 = this show ?thesis proof (cases "l1' = ?M1 \ r1' = r1 \ l2' = l2 \ r2' = ?M2") case True hence id2: "l1' = ?M1" "r1' = r1" "l2' = l2" "r2' = ?M2" by auto show ?thesis unfolding id id2 unfolding ring_distribs using ge0 by simp next case False note 3 = this from 1 2 3 mem have id2: "l1' = ?M1" "r1' = r1" "l2' = ?M2" "r2' = r2" unfolding id' by auto show ?thesis unfolding id id2 unfolding ring_distribs using ge0 by simp qed qed qed hence "r' - l' \ 3 / 4 * (r - l)" by simp } note decr = this show "converges_to (\i. bnd ((tighten_poly_bounds_binary p1 p2 ^^ i) ((l1,r1,sgn (ipoly p1 r1)),(l2,r2, sgn (ipoly p2 r2))))) (?x * ?y)" proof (intro tighten_poly_bounds_binary[where f = "(*)" and I = "\ l. l > 0"] basic l1_pos l2_pos, goal_cases) case (1 L1 R1 L2 R2 L R) hence "L = L1 * L2" "R = R1 * R2" unfolding bnd_def by auto hence id: "?r L = ?r L1 * ?r L2" "?r R = ?r R1 * ?r R2" by (auto simp: hom_distribs) from 1(3-4) have le: "?r L1 \ ?x" "?x \ ?r R1" "?r L2 \ ?y" "?y \ ?r R2" unfolding root_cond_def by auto from 1(1-2) have lt: "0 < ?r L1" "0 < ?r L2" by auto from mult_mono[OF le(1,3), folded id] lt le have L: "?r L \ ?x * ?y" by linarith have R: "?x * ?y \ ?r R" by (rule mult_mono[OF le(2,4), folded id], insert lt le, linarith+) show ?case using L R by blast next case (2 l1 r1 l2 r2 l1' r1' l2' r2' l l' r r') from 2(5-6) have lr: "l = l1 * l2" "r = r1 * r2" "l' = l1' * l2'" "r' = r1' * r2'" unfolding bnd_def by auto from 2(1-4) have le: "0 < l1" "0 < l2" "l1 \ r1" "l2 \ r2" by auto from 2(7-8) le have le': "l1 \ l1'" "r1' \ r1" "l2 \ l2'" "r2' \ r2" "0 < r2'" "0 < r2" by auto from mult_mono[OF le'(1,3), folded lr] le le' have l: "l \ l'" by auto have r: "r' \ r" by (rule mult_mono[OF le'(2,4), folded lr], insert le le', linarith+) have "r' - l' \ 3 / 4 * (r - l)" by (rule decr[OF _ _ _ _ _ _ 2(7-8)], insert le le' lr, auto) thus ?case using l r by blast qed auto qed have z': "z' = z" unfolding z[unfolded xt yt, simplified, unfolded bnd_def[symmetric] sel] by auto from main[unfolded this] show ?thesis unfolding xt yt by simp qed lemma mult_1: assumes x: "invariant_1_2 x" and y: "invariant_1_2 y" defines z[simp]: "z \ mult_1 x y" shows "invariant_2 z \ (real_of_2 z = real_of_1 x * real_of_1 y)" proof - obtain p1 l1 r1 where xt[simp]: "x = (p1,l1,r1)" by (cases x) obtain p2 l2 r2 where yt[simp]: "y = (p2,l2,r2)" by (cases y) let ?xt = "(p1, l1, r1)" let ?yt = "(p2, l2, r2)" let ?x = "real_of_1 ?xt" let ?y = "real_of_1 ?yt" let ?mxt = "uminus_1 ?xt" let ?myt = "uminus_1 ?yt" let ?mx = "real_of_1 ?mxt" let ?my = "real_of_1 ?myt" let ?r = "real_of_rat" from invariant_1_2_of_rat[OF x, of 0] have x0: "?x < 0 \ ?x > 0" by auto from invariant_1_2_of_rat[OF y, of 0] have y0: "?y < 0 \ ?y > 0" by auto from uminus_1_2[OF x] have mx: "invariant_1_2 ?mxt" and [simp]: "?mx = - ?x" by auto from uminus_1_2[OF y] have my: "invariant_1_2 ?myt" and [simp]: "?my = - ?y" by auto have id: "r1 > 0 \ ?x > 0" "r1 < 0 \ ?x < 0" "r2 > 0 \ ?y > 0" "r2 < 0 \ ?y < 0" using x y by auto show ?thesis proof (cases "?x > 0") case x0: True show ?thesis proof (cases "?y > 0") case y0: True with x y x0 mult_1_pos[OF x y] show ?thesis by auto next case False with y0 have y0: "?y < 0" by auto with x0 have z: "z = uminus_2 (mult_1_pos ?xt ?myt)" unfolding z xt yt mult_1.simps split id by simp from x0 y0 mult_1_pos[OF x my] uminus_2[of "mult_1_pos ?xt ?myt"] show ?thesis unfolding z by simp qed next case False with x0 have x0: "?x0 < 0" by simp show ?thesis proof (cases "?y > 0") case y0: True with x0 x y id have z: "z = uminus_2 (mult_1_pos ?mxt ?yt)" by simp from x0 y0 mult_1_pos[OF mx y] uminus_2[of "mult_1_pos ?mxt ?yt"] show ?thesis unfolding z by auto next case False with y0 have y0: "?y < 0" by simp with x0 x y have z: "z = mult_1_pos ?mxt ?myt" by auto with x0 y0 x y mult_1_pos[OF mx my] show ?thesis unfolding z by auto qed qed qed lemma mult_rat_1: fixes x assumes y: "invariant_1 y" defines z: "z \ mult_rat_1 x y" shows "invariant_2 z \ (real_of_2 z = of_rat x * real_of_1 y)" proof (cases y) case yt: (fields p2 l2 r2) let ?yt = "(p2, l2, r2)" let ?x = "real_of_rat x" let ?y = "real_of_1 ?yt" let ?myt = "mult_rat_1_pos (- x) ?yt" note y = y[unfolded yt] note z = z[unfolded yt] show ?thesis proof(cases x "0::rat" rule:linorder_cases) case x: greater with z have z: "z = mult_rat_1_pos x ?yt" by simp from mult_rat_1_pos[OF x y] show ?thesis unfolding yt z by auto next case less then have x: "- x > 0" by auto hence z: "z = uminus_2 ?myt" unfolding z by simp from mult_rat_1_pos[OF x y] have rc: "invariant_2 ?myt" and rr: "real_of_2 ?myt = - ?x * ?y" by (auto simp: hom_distribs) from uminus_2[OF rc] rr show ?thesis unfolding z[symmetric] unfolding yt[symmetric] by simp qed (auto simp: z) qed end declare mult_1.simps[simp del] declare mult_rat_1.simps[simp del] (* ********************* *) subsubsection\Root\ definition ipoly_root_delta :: "int poly \ real" where "ipoly_root_delta p = Min (insert 1 { abs (x - y) | x y. ipoly p x = 0 \ ipoly p y = 0 \ x \ y}) / 4" lemma ipoly_root_delta: assumes "p \ 0" shows "ipoly_root_delta p > 0" "2 \ card (Collect (root_cond (p, l, r))) \ ipoly_root_delta p \ real_of_rat (r - l) / 4" proof - let ?z = "0 :: real" let ?R = "{x. ipoly p x = ?z}" let ?set = "{ abs (x - y) | x y. ipoly p x = ?z \ ipoly p y = 0 \ x \ y}" define S where "S = insert 1 ?set" from finite_ipoly_roots[OF assms] have finR: "finite ?R" and fin: "finite (?R \ ?R)" by auto have "finite ?set" by (rule finite_subset[OF _ finite_imageI[OF fin, of "\ (x,y). abs (x - y)"]], force) hence fin: "finite S" and ne: "S \ {}" and pos: "\ x. x \ S \ x > 0" unfolding S_def by auto have delta: "ipoly_root_delta p = Min S / 4" unfolding ipoly_root_delta_def S_def .. have pos: "Min S > 0" using fin ne pos by auto show "ipoly_root_delta p > 0" unfolding delta using pos by auto let ?S = "Collect (root_cond (p, l, r))" assume "2 \ card ?S" hence 2: "Suc (Suc 0) \ card ?S" by simp from 2[unfolded card_le_Suc_iff[of _ ?S]] obtain x T where ST: "?S = insert x T" and xT: "x \ T" and 1: "Suc 0 \ card T" by auto from 1[unfolded card_le_Suc_iff[of _ T]] obtain y where yT: "y \ T" by auto from ST xT yT have x: "x \ ?S" and y: "y \ ?S" and xy: "x \ y" by auto hence "abs (x - y) \ S" unfolding S_def root_cond_def[abs_def] by auto with fin have "Min S \ abs (x - y)" by auto with pos have le: "Min S / 2 \ abs (x - y) / 2" by auto from x y have "abs (x - y) \ of_rat r - of_rat l" unfolding root_cond_def[abs_def] by auto also have "\ = of_rat (r - l)" by (auto simp: of_rat_diff) finally have "abs (x - y) / 2 \ of_rat (r - l) / 2" by auto with le show "ipoly_root_delta p \ real_of_rat (r - l) / 4" unfolding delta by auto qed lemma sgn_less_eq_1_rat: fixes a b :: rat shows "sgn a = 1 \ a \ b \ sgn b = 1" by (metis (no_types, hide_lams) not_less one_neq_neg_one one_neq_zero order_trans sgn_rat_def) lemma sgn_less_eq_1_real: fixes a b :: real shows "sgn a = 1 \ a \ b \ sgn b = 1" by (metis (no_types, hide_lams) not_less one_neq_neg_one one_neq_zero order_trans sgn_real_def) definition compare_1_rat :: "real_alg_1 \ rat \ order" where "compare_1_rat rai = (let p = poly_real_alg_1 rai in if degree p = 1 then let x = Rat.Fract (- coeff p 0) (coeff p 1) in (\ y. compare y x) else (\ y. compare_rat_1 y rai))" lemma compare_real_of_rat: "compare (real_of_rat x) (of_rat y) = compare x y" unfolding compare_rat_def compare_real_def comparator_of_def of_rat_less by auto lemma compare_1_rat: assumes rc: "invariant_1 y" shows "compare_1_rat y x = compare (of_rat x) (real_of_1 y)" proof (cases "degree (poly_real_alg_1 y)" "Suc 0" rule: linorder_cases) case less with invariant_1_degree_0[OF rc] show ?thesis by auto next case deg: greater with rc have rc: "invariant_1_2 y" by auto from deg compare_rat_1[OF rc, of x] show ?thesis unfolding compare_1_rat_def by auto next case deg: equal obtain p l r where y: "y = (p,l,r)" by (cases y) note rc = invariant_1D[OF rc[unfolded y]] from deg have p: "degree p = Suc 0" and id: "compare_1_rat y x = compare x (Rat.Fract (- coeff p 0) (coeff p 1))" unfolding compare_1_rat_def by (auto simp: Let_def y) from rc(1)[unfolded split] have "ipoly p (real_of_1 y) = 0" unfolding y by auto with degree_1_ipoly[OF p, of "real_of_1 y"] have id': "real_of_1 y = real_of_rat (Rat.Fract (- coeff p 0) (coeff p 1))" by simp show ?thesis unfolding id id' compare_real_of_rat .. qed context fixes n :: nat begin private definition initial_lower_bound :: "rat \ rat" where "initial_lower_bound l = (if l \ 1 then l else of_int (root_rat_floor n l))" private definition initial_upper_bound :: "rat \ rat" where "initial_upper_bound r = (of_int (root_rat_ceiling n r))" context fixes cmpx :: "rat \ order" begin fun tighten_bound_root :: "rat \ rat \ rat \ rat" where "tighten_bound_root (l',r') = (let m' = (l' + r') / 2; m = m' ^ n in case cmpx m of Eq \ (m',m') | Lt \ (m',r') | Gt \ (l',m'))" lemma tighten_bound_root: assumes sgn: "sgn il = 1" "real_of_1 x \ 0" and il: "real_of_rat il \ root n (real_of_1 x)" and ir: "root n (real_of_1 x) \ real_of_rat ir" and rai: "invariant_1 x" and cmpx: "cmpx = compare_1_rat x" and n: "n \ 0" shows "converges_to (\ i. (tighten_bound_root ^^ i) (il, ir)) (root n (real_of_1 x))" (is "converges_to ?f ?x") unfolding converges_to_def proof (intro conjI impI allI) { fix x :: real have "x \ 0 \ (root n x) ^ n = x" using n by simp } note root_exp_cancel = this { fix x :: real have "x \ 0 \ root n (x ^ n) = x" using n using real_root_pos_unique by blast } note root_exp_cancel' = this from il ir have "real_of_rat il \ of_rat ir" by auto hence ir_il: "il \ ir" by (auto simp: of_rat_less_eq) from n have n': "n > 0" by auto { fix i have "in_interval (?f i) ?x \ sub_interval (?f i) (il,ir) \ (i \ 0 \ sub_interval (?f i) (?f (i - 1))) \ snd (?f i) - fst (?f i) \ (ir - il) / 2^i" proof (induct i) case 0 show ?case using il ir by auto next case (Suc i) obtain l' r' where id: "(tighten_bound_root ^^ i) (il, ir) = (l',r')" by (cases "(tighten_bound_root ^^ i) (il, ir)", auto) let ?m' = "(l' + r') / 2" let ?m = "?m' ^ n" define m where "m = ?m" note IH = Suc[unfolded id split snd_conv fst_conv] from IH have "sub_interval (l', r') (il, ir)" by auto hence ill': "il \ l'" "r' \ ir" by auto with sgn have l'0: "l' > 0" using sgn_1_pos sgn_less_eq_1_rat by blast from IH have lr'x: "in_interval (l', r') ?x" by auto hence lr'': "real_of_rat l' \ of_rat r'" by auto hence lr': "l' \ r'" unfolding of_rat_less_eq . with l'0 have r'0: "r' > 0" by auto note compare = compare_1_rat[OF rai, of ?m, folded cmpx] from IH have *: "r' - l' \ (ir - il) / 2 ^ i" by auto have "r' - (l' + r') / 2 = (r' - l') / 2" by (simp add: field_simps) also have "\ \ (ir - il) / 2 ^ i / 2" using * by (rule divide_right_mono, auto) finally have size: "r' - (l' + r') / 2 \ (ir - il) / (2 * 2 ^ i)" by simp also have "r' - (l' + r') / 2 = (l' + r') / 2 - l'" by auto finally have size': "(l' + r') / 2 - l' \ (ir - il) / (2 * 2 ^ i)" by simp have "root n (real_of_rat ?m) = root n ((real_of_rat ?m') ^ n)" by (simp add: hom_distribs) also have "\ = real_of_rat ?m'" by (rule root_exp_cancel', insert l'0 lr', auto) finally have root: "root n (of_rat ?m) = of_rat ?m'" . show ?case proof (cases "cmpx ?m") case Eq from compare[unfolded Eq] have "real_of_1 x = of_rat ?m" unfolding compare_real_def comparator_of_def by (auto split: if_splits) from arg_cong[OF this, of "root n"] have "?x = root n (of_rat ?m)" . also have "\ = root n (real_of_rat ?m') ^ n" using n real_root_power by (auto simp: hom_distribs) also have "\ = of_rat ?m'" by (rule root_exp_cancel, insert IH sgn(2) l'0 r'0, auto) finally have x: "?x = of_rat ?m'" . show ?thesis using x id Eq lr' ill' ir_il by (auto simp: Let_def) next case Lt from compare[unfolded Lt] have lt: "of_rat ?m \ real_of_1 x" unfolding compare_real_def comparator_of_def by (auto split: if_splits) have id'': "?f (Suc i) = (?m',r')" "?f (Suc i - 1) = (l',r')" using Lt id by (auto simp add: Let_def) from real_root_le_mono[OF n' lt] have "of_rat ?m' \ ?x" unfolding root by simp with lr'x lr'' have ineq': "real_of_rat l' + real_of_rat r' \ ?x * 2" by (auto simp: hom_distribs) show ?thesis unfolding id'' by (auto simp: Let_def hom_distribs, insert size ineq' lr' ill' lr'x ir_il, auto) next case Gt from compare[unfolded Gt] have lt: "of_rat ?m \ real_of_1 x" unfolding compare_real_def comparator_of_def by (auto split: if_splits) have id'': "?f (Suc i) = (l',?m')" "?f (Suc i - 1) = (l',r')" using Gt id by (auto simp add: Let_def) from real_root_le_mono[OF n' lt] have "?x \ of_rat ?m'" unfolding root by simp with lr'x lr'' have ineq': "?x * 2 \ real_of_rat l' + real_of_rat r'" by (auto simp: hom_distribs) show ?thesis unfolding id'' by (auto simp: Let_def hom_distribs, insert size' ineq' lr' ill' lr'x ir_il, auto) qed qed } note main = this fix i from main[of i] show "in_interval (?f i) ?x" by auto from main[of "Suc i"] show "sub_interval (?f (Suc i)) (?f i)" by auto fix eps :: real assume eps: "0 < eps" define c where "c = eps / (max (real_of_rat (ir - il)) 1)" have c0: "c > 0" using eps unfolding c_def by auto from exp_tends_to_zero[OF _ _ this, of "1/2"] obtain i where c: "(1/2)^i \ c" by auto obtain l' r' where fi: "?f i = (l',r')" by force from main[of i, unfolded fi] have le: "r' - l' \ (ir - il) / 2 ^ i" by auto have iril: "real_of_rat (ir - il) \ 0" using ir_il by (auto simp: of_rat_less_eq) show "\n la ra. ?f n = (la, ra) \ real_of_rat ra - real_of_rat la \ eps" proof (intro conjI exI, rule fi) have "real_of_rat r' - of_rat l' = real_of_rat (r' - l')" by (auto simp: hom_distribs) also have "\ \ real_of_rat ((ir - il) / 2 ^ i)" using le unfolding of_rat_less_eq . also have "\ = (real_of_rat (ir - il)) * ((1/2) ^ i)" by (simp add: field_simps hom_distribs) also have "\ \ (real_of_rat (ir - il)) * c" by (rule mult_left_mono[OF c iril]) also have "\ \ eps" proof (cases "real_of_rat (ir - il) \ 1") case True hence "c = eps" unfolding c_def by (auto simp: hom_distribs) thus ?thesis using eps True by auto next case False hence "max (real_of_rat (ir - il)) 1 = real_of_rat (ir - il)" "real_of_rat (ir - il) \ 0" by (auto simp: hom_distribs) hence "(real_of_rat (ir - il)) * c = eps" unfolding c_def by auto thus ?thesis by simp qed finally show "real_of_rat r' - of_rat l' \ eps" . qed qed end private fun root_pos_1 :: "real_alg_1 \ real_alg_2" where "root_pos_1 (p,l,r) = ( (select_correct_factor_int_poly (tighten_bound_root (compare_1_rat (p,l,r))) (\ x. x) (initial_lower_bound l, initial_upper_bound r) (poly_nth_root n p)))" fun root_1 :: "real_alg_1 \ real_alg_2" where "root_1 (p,l,r) = ( if n = 0 \ r = 0 then Rational 0 else if r > 0 then root_pos_1 (p,l,r) else uminus_2 (root_pos_1 (uminus_1 (p,l,r))))" context assumes n: "n \ 0" begin lemma initial_upper_bound: assumes x: "x > 0" and xr: "x \ of_rat r" shows "sgn (initial_upper_bound r) = 1" "root n x \ of_rat (initial_upper_bound r)" proof - have n: "n > 0" using n by auto note d = initial_upper_bound_def let ?r = "initial_upper_bound r" from x xr have r0: "r > 0" by (meson not_less of_rat_le_0_iff order_trans) hence "of_rat r > (0 :: real)" by auto hence "root n (of_rat r) > 0" using n by simp hence "1 \ ceiling (root n (of_rat r))" by auto hence "(1 :: rat) \ of_int (ceiling (root n (of_rat r)))" by linarith also have "\ = ?r" unfolding d by simp finally show "sgn ?r = 1" unfolding sgn_rat_def by auto have "root n x \ root n (of_rat r)" unfolding real_root_le_iff[OF n] by (rule xr) also have "\ \ of_rat ?r" unfolding d by simp finally show "root n x \ of_rat ?r" . qed lemma initial_lower_bound: assumes l: "l > 0" and lx: "of_rat l \ x" shows "sgn (initial_lower_bound l) = 1" "of_rat (initial_lower_bound l) \ root n x" proof - have n: "n > 0" using n by auto note d = initial_lower_bound_def let ?l = "initial_lower_bound l" from l lx have x0: "x > 0" by (meson not_less of_rat_le_0_iff order_trans) have "sgn ?l = 1 \ of_rat ?l \ root n x" proof (cases "l \ 1") case True hence ll: "?l = l" and l0: "of_rat l \ (0 :: real)" and l1: "of_rat l \ (1 :: real)" using l unfolding True d by auto have sgn: "sgn ?l = 1" using l unfolding ll by auto have "of_rat ?l = of_rat l" unfolding ll by simp also have "of_rat l \ root n (of_rat l)" using real_root_increasing[OF _ _ l0 l1, of 1 n] n by (cases "n = 1", auto) also have "\ \ root n x" using lx unfolding real_root_le_iff[OF n] . finally show ?thesis using sgn by auto next case False hence l: "(1 :: real) \ of_rat l" and ll: "?l = of_int (floor (root n (of_rat l)))" unfolding d by auto hence "root n 1 \ root n (of_rat l)" unfolding real_root_le_iff[OF n] by auto hence "1 \ root n (of_rat l)" using n by auto from floor_mono[OF this] have "1 \ ?l" using one_le_floor unfolding ll by fastforce hence sgn: "sgn ?l = 1" by simp have "of_rat ?l \ root n (of_rat l)" unfolding ll by simp also have "\ \ root n x" using lx unfolding real_root_le_iff[OF n] . finally have "of_rat ?l \ root n x" . with sgn show ?thesis by auto qed thus "sgn ?l = 1" "of_rat ?l \ root n x" by auto qed lemma root_pos_1: assumes x: "invariant_1 x" and pos: "rai_ub x > 0" defines y: "y \ root_pos_1 x" shows "invariant_2 y \ real_of_2 y = root n (real_of_1 x)" proof (cases x) case (fields p l r) let ?l = "initial_lower_bound l" let ?r = "initial_upper_bound r" from x fields have rai: "invariant_1 (p,l,r)" by auto note * = invariant_1D[OF this] let ?x = "the_unique_root (p,l,r)" from pos[unfolded fields] * have sgnl: "sgn l = 1" by auto from sgnl have l0: "l > 0" by (unfold sgn_1_pos) hence ll0: "real_of_rat l > 0" by auto from * have lx: "of_rat l \ ?x" by auto with ll0 have x0: "?x > 0" by linarith note il = initial_lower_bound[OF l0 lx] from * have "?x \ of_rat r" by auto note iu = initial_upper_bound[OF x0 this] let ?p = "poly_nth_root n p" from x0 have id: "root n ?x ^ n = ?x" using n real_root_pow_pos by blast have rc: "root_cond (?p, ?l, ?r) (root n ?x)" using il iu * by (intro root_condI, auto simp: ipoly_nth_root id) hence root: "ipoly ?p (root n (real_of_1 x)) = 0" unfolding root_cond_def fields by auto from * have "p \ 0" by auto hence p': "?p \ 0" using poly_nth_root_0[of n p] n by auto have tbr: "0 \ real_of_1 x" "real_of_rat (initial_lower_bound l) \ root n (real_of_1 x)" "root n (real_of_1 x) \ real_of_rat (initial_upper_bound r)" using x0 il(2) iu(2) fields by auto from select_correct_factor_int_poly[OF tighten_bound_root[OF il(1)[folded fields] tbr x refl n] refl root p'] show ?thesis by (simp add: y fields) qed end lemma root_1: assumes x: "invariant_1 x" defines y: "y \ root_1 x" shows "invariant_2 y \ (real_of_2 y = root n (real_of_1 x))" proof (cases "n = 0 \ rai_ub x = 0") case True with x have "n = 0 \ real_of_1 x = 0" by (cases x, auto) then have "root n (real_of_1 x) = 0" by auto then show ?thesis unfolding y root_1.simps using x by (cases x, auto) next case False with x have n: "n \ 0" and x0: "real_of_1 x \ 0" by (simp, cases x, auto) note rt = root_pos_1 show ?thesis proof (cases "rai_ub x" "0::rat" rule:linorder_cases) case greater with rt[OF n x this] n show ?thesis by (unfold y, cases x, simp) next case less let ?um = "uminus_1" let ?rt = "root_pos_1" from n less y x0 have y: "y = uminus_2 (?rt (?um x))" by (cases x, auto) from uminus_1[OF x] have umx: "invariant_1 (?um x)" and umx2: "real_of_1 (?um x) = - real_of_1 x" by auto with x less have "0 < rai_ub (uminus_1 x)" by (cases x, auto simp: uminus_1.simps Let_def) from rt[OF n umx this] umx2 have rumx: "invariant_2 (?rt (?um x))" and rumx2: "real_of_2 (?rt (?um x)) = root n (- real_of_1 x)" by auto from uminus_2[OF rumx] rumx2 y real_root_minus show ?thesis by auto next case equal with x0 x show ?thesis by (cases x, auto) qed qed end declare root_1.simps[simp del] (* ********************** *) subsubsection \Embedding of Rational Numbers\ definition of_rat_1 :: "rat \ real_alg_1" where "of_rat_1 x \ (poly_rat x,x,x)" lemma of_rat_1: shows "invariant_1 (of_rat_1 x)" and "real_of_1 (of_rat_1 x) = of_rat x" unfolding of_rat_1_def by (atomize(full), intro invariant_1_realI unique_rootI poly_condI, auto ) fun info_2 :: "real_alg_2 \ rat + int poly \ nat" where "info_2 (Rational x) = Inl x" | "info_2 (Irrational n (p,l,r)) = Inr (p,n)" lemma info_2_card: assumes rc: "invariant_2 x" shows "info_2 x = Inr (p,n) \ poly_cond p \ ipoly p (real_of_2 x) = 0 \ degree p \ 2 \ card (roots_below p (real_of_2 x)) = n" "info_2 x = Inl y \ real_of_2 x = of_rat y" proof (atomize(full), goal_cases) case 1 show ?case proof (cases x) case (Irrational m rai) then obtain q l r where x: "x = Irrational m (q,l,r)" by (cases rai, auto) show ?thesis proof (cases "q = p \ m = n") case False thus ?thesis using x by auto next case True with x have x: "x = Irrational n (p,l,r)" by auto from rc[unfolded x, simplified] have inv: "invariant_1_2 (p,l,r)" and n: "card (roots_below p (real_of_2 x)) = n" and 1: "degree p \ 1" by (auto simp: x) from inv have "degree p \ 0" unfolding irreducible_def by auto with 1 have "degree p \ 2" by linarith thus ?thesis unfolding n using inv x by (auto elim!: invariant_1E) qed qed auto qed lemma real_of_2_Irrational: "invariant_2 (Irrational n rai) \ real_of_2 (Irrational n rai) \ of_rat x" proof assume "invariant_2 (Irrational n rai)" and rat: "real_of_2 (Irrational n rai) = real_of_rat x" hence "real_of_1 rai \ \" "invariant_1_2 rai" by auto from invariant_1_2_of_rat[OF this(2)] rat show False by auto qed lemma info_2: assumes ix: "invariant_2 x" and iy: "invariant_2 y" shows "info_2 x = info_2 y \ real_of_2 x = real_of_2 y" proof (cases x) case x: (Irrational n1 rai1) note ix = ix[unfolded x] show ?thesis proof (cases y) case (Rational y) with real_of_2_Irrational[OF ix, of y] show ?thesis unfolding x by (cases rai1, auto) next case y: (Irrational n2 rai2) obtain p1 l1 r1 where rai1: "rai1 = (p1,l1,r1)" by (cases rai1) obtain p2 l2 r2 where rai2: "rai2 = (p2,l2,r2)" by (cases rai2) let ?rx = "the_unique_root (p1,l1,r1)" let ?ry = "the_unique_root (p2,l2,r2)" have id: "(info_2 x = info_2 y) = (p1 = p2 \ n1 = n2)" "(real_of_2 x = real_of_2 y) = (?rx = ?ry)" unfolding x y rai1 rai2 by auto from ix[unfolded x rai1] have ix: "invariant_1 (p1, l1, r1)" and deg1: "degree p1 > 1" and n1: "n1 = card (roots_below p1 ?rx)" by auto note Ix = invariant_1D[OF ix] from deg1 have p1_0: "p1 \ 0" by auto from iy[unfolded y rai2] have iy: "invariant_1 (p2, l2, r2)" and "degree p2 > 1" and n2: "n2 = card (roots_below p2 ?ry)" by auto note Iy = invariant_1D[OF iy] show ?thesis unfolding id proof assume eq: "?rx = ?ry" from Ix have algx: "p1 represents ?rx \ irreducible p1 \ lead_coeff p1 > 0" unfolding represents_def by auto from iy have algy: "p2 represents ?rx \ irreducible p2 \ lead_coeff p2 > 0" unfolding represents_def eq by (auto elim!: invariant_1E) from algx have "algebraic ?rx" unfolding algebraic_altdef_ipoly by auto note unique = algebraic_imp_represents_unique[OF this] with algx algy have id: "p2 = p1" by auto from eq id n1 n2 show "p1 = p2 \ n1 = n2" by auto next assume "p1 = p2 \ n1 = n2" hence id: "p1 = p2" "n1 = n2" by auto hence card: "card (roots_below p1 ?rx) = card (roots_below p1 ?ry)" unfolding n1 n2 by auto show "?rx = ?ry" proof (cases ?rx ?ry rule: linorder_cases) case less have "roots_below p1 ?rx = roots_below p1 ?ry" proof (intro card_subset_eq finite_subset[OF _ ipoly_roots_finite] card) from less show "roots_below p1 ?rx \ roots_below p1 ?ry" by auto qed (insert p1_0, auto) then show ?thesis using id less unique_rootD(3)[OF Iy(4)] by (auto simp: less_eq_real_def) next case equal then show ?thesis by (simp add: id) next case greater have "roots_below p1 ?ry = roots_below p1 ?rx" proof (intro card_subset_eq card[symmetric] finite_subset[OF _ ipoly_roots_finite[OF p1_0]]) from greater show "roots_below p1 ?ry \ roots_below p1 ?rx" by auto qed auto hence "roots_below p2 ?ry = roots_below p2 ?rx" unfolding id by auto thus ?thesis using id greater unique_rootD(3)[OF Ix(4)] by (auto simp: less_eq_real_def) qed qed qed next case x: (Rational x) show ?thesis proof (cases y) case (Rational y) thus ?thesis using x by auto next case y: (Irrational n rai) with real_of_2_Irrational[OF iy[unfolded y], of x] show ?thesis unfolding x by (cases rai, auto) qed qed lemma info_2_unique: "invariant_2 x \ invariant_2 y \ real_of_2 x = real_of_2 y \ info_2 x = info_2 y" using info_2 by blast lemma info_2_inj: "invariant_2 x \ invariant_2 y \ info_2 x = info_2 y \ real_of_2 x = real_of_2 y" using info_2 by blast context fixes cr1 cr2 :: "rat \ rat \ nat" begin partial_function (tailrec) compare_1 :: "int poly \ int poly \ rat \ rat \ rat \ rat \ rat \ rat \ order" where [code]: "compare_1 p1 p2 l1 r1 sr1 l2 r2 sr2 = (if r1 < l2 then Lt else if r2 < l1 then Gt else let (l1',r1',sr1') = tighten_poly_bounds p1 l1 r1 sr1; (l2',r2',sr2') = tighten_poly_bounds p2 l2 r2 sr2 in compare_1 p1 p2 l1' r1' sr1' l2' r2' sr2') " lemma compare_1: assumes ur1: "unique_root (p1,l1,r1)" and ur2: "unique_root (p2,l2,r2)" and pc: "poly_cond2 p1" "poly_cond2 p2" and diff: "the_unique_root (p1,l1,r1) \ the_unique_root (p2,l2,r2)" and sr: "sr1 = sgn (ipoly p1 r1)" "sr2 = sgn (ipoly p2 r2)" shows "compare_1 p1 p2 l1 r1 sr1 l2 r2 sr2 = compare (the_unique_root (p1,l1,r1)) (the_unique_root (p2,l2,r2))" proof - let ?r = real_of_rat { fix d x y assume d: "d = (r1 - l1) + (r2 - l2)" and xy: "x = the_unique_root (p1,l1,r1)" "y = the_unique_root (p2,l2,r2)" define delta where "delta = abs (x - y) / 4" have delta: "delta > 0" and diff: "x \ y" unfolding delta_def using diff xy by auto let ?rel' = "{(x, y). 0 \ y \ delta_gt delta x y}" let ?rel = "inv_image ?rel' ?r" have SN: "SN ?rel" by (rule SN_inv_image[OF delta_gt_SN[OF delta]]) from d ur1 ur2 have ?thesis unfolding xy[symmetric] using xy sr proof (induct d arbitrary: l1 r1 l2 r2 sr1 sr2 rule: SN_induct[OF SN]) case (1 d l1 r1 l2 r2) note IH = 1(1) note d = 1(2) note ur = 1(3-4) note xy = 1(5-6) note sr = 1(7-8) note simps = compare_1.simps[of p1 p2 l1 r1 sr1 l2 r2 sr2] note urx = unique_rootD[OF ur(1), folded xy] note ury = unique_rootD[OF ur(2), folded xy] show ?case (is "?l = _") proof (cases "r1 < l2") case True hence l: "?l = Lt" and lt: "?r r1 < ?r l2" unfolding simps of_rat_less by auto show ?thesis unfolding l using lt True urx(2) ury(1) by (auto simp: compare_real_def comparator_of_def) next case False note le = this show ?thesis proof (cases "r2 < l1") case True with le have l: "?l = Gt" and lt: "?r r2 < ?r l1" unfolding simps of_rat_less by auto show ?thesis unfolding l using lt True ury(2) urx(1) by (auto simp: compare_real_def comparator_of_def) next case False obtain l1' r1' sr1' where tb1: "tighten_poly_bounds p1 l1 r1 sr1 = (l1',r1',sr1')" by (cases rule: prod_cases3, auto) obtain l2' r2' sr2' where tb2: "tighten_poly_bounds p2 l2 r2 sr2 = (l2',r2',sr2')" by (cases rule: prod_cases3, auto) from False le tb1 tb2 have l: "?l = compare_1 p1 p2 l1' r1' sr1' l2' r2' sr2'" unfolding simps by auto from tighten_poly_bounds[OF tb1 ur(1) pc(1) sr(1)] have rc1: "root_cond (p1, l1', r1') (the_unique_root (p1, l1, r1))" and bnd1: "l1 \ l1'" "l1' \ r1'" "r1' \ r1" and d1: "r1' - l1' = (r1 - l1) / 2" and sr1: "sr1' = sgn (ipoly p1 r1')" by auto from pc have "p1 \ 0" "p2 \ 0" by auto from unique_root_sub_interval[OF ur(1) rc1 bnd1(1,3)] xy ur this have ur1: "unique_root (p1, l1', r1')" and x: "x = the_unique_root (p1, l1', r1')" by (auto intro!: the_unique_root_eqI) from tighten_poly_bounds[OF tb2 ur(2) pc(2) sr(2)] have rc2: "root_cond (p2, l2', r2') (the_unique_root (p2, l2, r2))" and bnd2: "l2 \ l2'" "l2' \ r2'" "r2' \ r2" and d2: "r2' - l2' = (r2 - l2) / 2" and sr2: "sr2' = sgn (ipoly p2 r2')" by auto from unique_root_sub_interval[OF ur(2) rc2 bnd2(1,3)] xy ur pc have ur2: "unique_root (p2, l2', r2')" and y: "y = the_unique_root (p2, l2', r2')" by auto define d' where "d' = d/2" have d': "d' = r1' - l1' + (r2' - l2')" unfolding d'_def d d1 d2 by (simp add: field_simps) have d'0: "d' \ 0" using bnd1 bnd2 unfolding d' by auto have dd: "d - d' = d/2" unfolding d'_def by simp have "abs (x - y) \ 2 * ?r d" proof (rule ccontr) assume "\ ?thesis" hence lt: "2 * ?r d < abs (x - y)" by auto have "r1 - l1 \ d" "r2 - l2 \ d" unfolding d using bnd1 bnd2 by auto from this[folded of_rat_less_eq[where 'a = real]] lt have "?r (r1 - l1) < abs (x - y) / 2" "?r (r2 - l2) < abs (x - y) / 2" and dd: "?r r1 - ?r l1 \ ?r d" "?r r2 - ?r l2 \ ?r d" by (auto simp: of_rat_diff) from le have "r1 \ l2" by auto hence r1l2: "?r r1 \ ?r l2" unfolding of_rat_less_eq by auto from False have "r2 \ l1" by auto hence r2l1: "?r r2 \ ?r l1" unfolding of_rat_less_eq by auto show False proof (cases "x \ y") case True from urx(1-2) dd(1) have "?r r1 \ x + ?r d" by auto with r1l2 have "?r l2 \ x + ?r d" by auto with True lt ury(2) dd(2) show False by auto next case False from ury(1-2) dd(2) have "?r r2 \ y + ?r d" by auto with r2l1 have "?r l1 \ y + ?r d" by auto with False lt urx(2) dd(1) show False by auto qed qed hence dd': "delta_gt delta (?r d) (?r d')" unfolding delta_gt_def delta_def using dd by (auto simp: hom_distribs) show ?thesis unfolding l by (rule IH[OF _ d' ur1 ur2 x y sr1 sr2], insert d'0 dd', auto) qed qed qed } thus ?thesis by auto qed end (* **************************************************************** *) fun real_alg_1 :: "real_alg_2 \ real_alg_1" where "real_alg_1 (Rational r) = of_rat_1 r" | "real_alg_1 (Irrational n rai) = rai" lemma real_alg_1: "real_of_1 (real_alg_1 x) = real_of_2 x" by (cases x, auto simp: of_rat_1) definition root_2 :: "nat \ real_alg_2 \ real_alg_2" where "root_2 n x = root_1 n (real_alg_1 x)" lemma root_2: assumes "invariant_2 x" shows "real_of_2 (root_2 n x) = root n (real_of_2 x)" "invariant_2 (root_2 n x)" proof (atomize(full), cases x, goal_cases) case (1 y) from of_rat_1[of y] root_1[of "of_rat_1 y" n] assms 1 real_alg_2 show ?case by (simp add: root_2_def) next case (2 i rai) from root_1[of rai n] assms 2 real_alg_2 show ?case by (auto simp: root_2_def) qed fun add_2 :: "real_alg_2 \ real_alg_2 \ real_alg_2" where "add_2 (Rational r) (Rational q) = Rational (r + q)" | "add_2 (Rational r) (Irrational n x) = Irrational n (add_rat_1 r x)" | "add_2 (Irrational n x) (Rational q) = Irrational n (add_rat_1 q x)" | "add_2 (Irrational n x) (Irrational m y) = add_1 x y" lemma add_2: assumes x: "invariant_2 x" and y: "invariant_2 y" shows "invariant_2 (add_2 x y)" (is ?g1) and "real_of_2 (add_2 x y) = real_of_2 x + real_of_2 y" (is ?g2) using assms add_rat_1 add_1 by (atomize (full), (cases x; cases y), auto simp: hom_distribs) fun mult_2 :: "real_alg_2 \ real_alg_2 \ real_alg_2" where "mult_2 (Rational r) (Rational q) = Rational (r * q)" | "mult_2 (Rational r) (Irrational n y) = mult_rat_1 r y" | "mult_2 (Irrational n x) (Rational q) = mult_rat_1 q x" | "mult_2 (Irrational n x) (Irrational m y) = mult_1 x y" lemma mult_2: assumes "invariant_2 x" "invariant_2 y" shows "real_of_2 (mult_2 x y) = real_of_2 x * real_of_2 y" "invariant_2 (mult_2 x y)" using assms by (atomize(full), (cases x; cases y; auto simp: mult_rat_1 mult_1 hom_distribs)) fun to_rat_2 :: "real_alg_2 \ rat option" where "to_rat_2 (Rational r) = Some r" | "to_rat_2 (Irrational n rai) = None" lemma to_rat_2: assumes rc: "invariant_2 x" shows "to_rat_2 x = (if real_of_2 x \ \ then Some (THE q. real_of_2 x = of_rat q) else None)" proof (cases x) case (Irrational n rai) from real_of_2_Irrational[OF rc[unfolded this]] show ?thesis unfolding Irrational Rats_def by auto qed simp fun equal_2 :: "real_alg_2 \ real_alg_2 \ bool" where "equal_2 (Rational r) (Rational q) = (r = q)" | "equal_2 (Irrational n (p,_)) (Irrational m (q,_)) = (p = q \ n = m)" | "equal_2 (Rational r) (Irrational _ yy) = False" | "equal_2 (Irrational _ xx) (Rational q) = False" lemma equal_2[simp]: assumes rc: "invariant_2 x" "invariant_2 y" shows "equal_2 x y = (real_of_2 x = real_of_2 y)" using info_2[OF rc] by (cases x; cases y, auto) fun compare_2 :: "real_alg_2 \ real_alg_2 \ order" where "compare_2 (Rational r) (Rational q) = (compare r q)" | "compare_2 (Irrational n (p,l,r)) (Irrational m (q,l',r')) = (if p = q \ n = m then Eq else compare_1 p q l r (sgn (ipoly p r)) l' r' (sgn (ipoly q r')))" | "compare_2 (Rational r) (Irrational _ xx) = (compare_rat_1 r xx)" | "compare_2 (Irrational _ xx) (Rational r) = (invert_order (compare_rat_1 r xx))" lemma compare_2: assumes rc: "invariant_2 x" "invariant_2 y" shows "compare_2 x y = compare (real_of_2 x) (real_of_2 y)" proof (cases x) case (Rational r) note xx = this show ?thesis proof (cases y) case (Rational q) note yy = this show ?thesis unfolding xx yy by (simp add: compare_rat_def compare_real_def comparator_of_def of_rat_less) next case (Irrational n yy) note yy = this from compare_rat_1 rc show ?thesis unfolding xx yy by (simp add: of_rat_1) qed next case (Irrational n xx) note xx = this show ?thesis proof (cases y) case (Rational q) note yy = this from compare_rat_1 rc show ?thesis unfolding xx yy by simp next case (Irrational m yy) note yy = this obtain p l r where xxx: "xx = (p,l,r)" by (cases xx) obtain q l' r' where yyy: "yy = (q,l',r')" by (cases yy) note rc = rc[unfolded xx xxx yy yyy] from rc have I: "invariant_1_2 (p,l,r)" "invariant_1_2 (q,l',r')" by auto then have "unique_root (p,l,r)" "unique_root (q,l',r')" "poly_cond2 p" "poly_cond2 q" by auto from compare_1[OF this _ refl refl] show ?thesis using equal_2[OF rc] unfolding xx xxx yy yyy by simp qed qed fun sgn_2 :: "real_alg_2 \ rat" where "sgn_2 (Rational r) = sgn r" | "sgn_2 (Irrational n rai) = sgn_1 rai" lemma sgn_2: "invariant_2 x \ real_of_rat (sgn_2 x) = sgn (real_of_2 x)" using sgn_1 by (cases x, auto simp: real_of_rat_sgn) fun floor_2 :: "real_alg_2 \ int" where "floor_2 (Rational r) = floor r" | "floor_2 (Irrational n rai) = floor_1 rai" lemma floor_2: "invariant_2 x \ floor_2 x = floor (real_of_2 x)" by (cases x, auto simp: floor_1) (* *************** *) subsubsection \Definitions and Algorithms on Type with Invariant\ lift_definition of_rat_3 :: "rat \ real_alg_3" is of_rat_2 by (auto simp: of_rat_2) lemma of_rat_3: "real_of_3 (of_rat_3 x) = of_rat x" by (transfer, auto simp: of_rat_2) lift_definition root_3 :: "nat \ real_alg_3 \ real_alg_3" is root_2 by (auto simp: root_2) lemma root_3: "real_of_3 (root_3 n x) = root n (real_of_3 x)" by (transfer, auto simp: root_2) lift_definition equal_3 :: "real_alg_3 \ real_alg_3 \ bool" is equal_2 . lemma equal_3: "equal_3 x y = (real_of_3 x = real_of_3 y)" by (transfer, auto) lift_definition compare_3 :: "real_alg_3 \ real_alg_3 \ order" is compare_2 . lemma compare_3: "compare_3 x y = (compare (real_of_3 x) (real_of_3 y))" by (transfer, auto simp: compare_2) lift_definition add_3 :: "real_alg_3 \ real_alg_3 \ real_alg_3" is add_2 by (auto simp: add_2) lemma add_3: "real_of_3 (add_3 x y) = real_of_3 x + real_of_3 y" by (transfer, auto simp: add_2) lift_definition mult_3 :: "real_alg_3 \ real_alg_3 \ real_alg_3" is mult_2 by (auto simp: mult_2) lemma mult_3: "real_of_3 (mult_3 x y) = real_of_3 x * real_of_3 y" by (transfer, auto simp: mult_2) lift_definition sgn_3 :: "real_alg_3 \ rat" is sgn_2 . lemma sgn_3: "real_of_rat (sgn_3 x) = sgn (real_of_3 x)" by (transfer, auto simp: sgn_2) lift_definition to_rat_3 :: "real_alg_3 \ rat option" is to_rat_2 . lemma to_rat_3: "to_rat_3 x = (if real_of_3 x \ \ then Some (THE q. real_of_3 x = of_rat q) else None)" by (transfer, simp add: to_rat_2) lift_definition floor_3 :: "real_alg_3 \ int" is floor_2 . lemma floor_3: "floor_3 x = floor (real_of_3 x)" by (transfer, auto simp: floor_2) (* *************** *) (* info *) lift_definition info_3 :: "real_alg_3 \ rat + int poly \ nat" is info_2 . lemma info_3_fun: "real_of_3 x = real_of_3 y \ info_3 x = info_3 y" by (transfer, intro info_2_unique, auto) lift_definition info_real_alg :: "real_alg \ rat + int poly \ nat" is info_3 by (metis info_3_fun) lemma info_real_alg: "info_real_alg x = Inr (p,n) \ p represents (real_of x) \ card {y. y \ real_of x \ ipoly p y = 0} = n \ irreducible p" "info_real_alg x = Inl q \ real_of x = of_rat q" proof (atomize(full), transfer, transfer, goal_cases) case (1 x p n q) from 1 have x: "invariant_2 x" by auto note info = info_2_card[OF this] show ?case proof (cases x) case irr: (Irrational m rai) from info(1)[of p n] show ?thesis unfolding irr by (cases rai, auto simp: poly_cond_def) qed (insert 1 info, auto) qed (* add *) instantiation real_alg :: plus begin lift_definition plus_real_alg :: "real_alg \ real_alg \ real_alg" is add_3 by (simp add: add_3) instance .. end lemma plus_real_alg: "(real_of x) + (real_of y) = real_of (x + y)" by (transfer, rule add_3[symmetric]) (* minus *) instantiation real_alg :: minus begin definition minus_real_alg :: "real_alg \ real_alg \ real_alg" where "minus_real_alg x y = x + (-y)" instance .. end lemma minus_real_alg: "(real_of x) - (real_of y) = real_of (x - y)" unfolding minus_real_alg_def minus_real_def uminus_real_alg plus_real_alg .. (* of_rat *) lift_definition of_rat_real_alg :: "rat \ real_alg" is of_rat_3 . lemma of_rat_real_alg: "real_of_rat x = real_of (of_rat_real_alg x)" by (transfer, rule of_rat_3[symmetric]) (* zero *) instantiation real_alg :: zero begin definition zero_real_alg :: "real_alg" where "zero_real_alg \ of_rat_real_alg 0" instance .. end lemma zero_real_alg: "0 = real_of 0" unfolding zero_real_alg_def by (simp add: of_rat_real_alg[symmetric]) (* one *) instantiation real_alg :: one begin definition one_real_alg :: "real_alg" where "one_real_alg \ of_rat_real_alg 1" instance .. end lemma one_real_alg: "1 = real_of 1" unfolding one_real_alg_def by (simp add: of_rat_real_alg[symmetric]) (* times *) instantiation real_alg :: times begin lift_definition times_real_alg :: "real_alg \ real_alg \ real_alg" is mult_3 by (simp add: mult_3) instance .. end lemma times_real_alg: "(real_of x) * (real_of y) = real_of (x * y)" by (transfer, rule mult_3[symmetric]) (* inverse *) instantiation real_alg :: inverse begin lift_definition inverse_real_alg :: "real_alg \ real_alg" is inverse_3 by (simp add: inverse_3) definition divide_real_alg :: "real_alg \ real_alg \ real_alg" where "divide_real_alg x y = x * inverse y" (* TODO: better to use poly_div *) instance .. end lemma inverse_real_alg: "inverse (real_of x) = real_of (inverse x)" by (transfer, rule inverse_3[symmetric]) lemma divide_real_alg: "(real_of x) / (real_of y) = real_of (x / y)" unfolding divide_real_alg_def times_real_alg[symmetric] divide_real_def inverse_real_alg .. (* group *) instance real_alg :: ab_group_add apply intro_classes apply (transfer, unfold add_3, force) apply (unfold zero_real_alg_def, transfer, unfold add_3 of_rat_3, force) apply (transfer, unfold add_3 of_rat_3, force) apply (transfer, unfold add_3 uminus_3 of_rat_3, force) apply (unfold minus_real_alg_def, force) done (* field *) instance real_alg :: field apply intro_classes apply (transfer, unfold mult_3, force) apply (transfer, unfold mult_3, force) apply (unfold one_real_alg_def, transfer, unfold mult_3 of_rat_3, force) apply (transfer, unfold mult_3 add_3, force simp: field_simps) apply (unfold zero_real_alg_def, transfer, unfold of_rat_3, force) apply (transfer, unfold mult_3 inverse_3 of_rat_3, force simp: field_simps) apply (unfold divide_real_alg_def, force) apply (transfer, unfold inverse_3 of_rat_3, force) done (* numeral *) instance real_alg :: numeral .. (* root *) lift_definition root_real_alg :: "nat \ real_alg \ real_alg" is root_3 by (simp add: root_3) lemma root_real_alg: "root n (real_of x) = real_of (root_real_alg n x)" by (transfer, rule root_3[symmetric]) (* sgn *) lift_definition sgn_real_alg_rat :: "real_alg \ rat" is sgn_3 by (insert sgn_3, metis to_rat_of_rat) lemma sgn_real_alg_rat: "real_of_rat (sgn_real_alg_rat x) = sgn (real_of x)" by (transfer, auto simp: sgn_3) instantiation real_alg :: sgn begin definition sgn_real_alg :: "real_alg \ real_alg" where "sgn_real_alg x = of_rat_real_alg (sgn_real_alg_rat x)" instance .. end lemma sgn_real_alg: "sgn (real_of x) = real_of (sgn x)" unfolding sgn_real_alg_def of_rat_real_alg[symmetric] by (transfer, simp add: sgn_3) (* equal *) instantiation real_alg :: equal begin lift_definition equal_real_alg :: "real_alg \ real_alg \ bool" is equal_3 by (simp add: equal_3) instance proof fix x y :: real_alg show "equal_class.equal x y = (x = y)" by (transfer, simp add: equal_3) qed end lemma equal_real_alg: "HOL.equal (real_of x) (real_of y) = (x = y)" unfolding equal_real_def by (transfer, auto) (* comparisons *) instantiation real_alg :: ord begin definition less_real_alg :: "real_alg \ real_alg \ bool" where [code del]: "less_real_alg x y = (real_of x < real_of y)" definition less_eq_real_alg :: "real_alg \ real_alg \ bool" where [code del]: "less_eq_real_alg x y = (real_of x \ real_of y)" instance .. end lemma less_real_alg: "less (real_of x) (real_of y) = (x < y)" unfolding less_real_alg_def .. lemma less_eq_real_alg: "less_eq (real_of x) (real_of y) = (x \ y)" unfolding less_eq_real_alg_def .. instantiation real_alg :: compare_order begin lift_definition compare_real_alg :: "real_alg \ real_alg \ order" is compare_3 by (simp add: compare_3) lemma compare_real_alg: "compare (real_of x) (real_of y) = (compare x y)" by (transfer, simp add: compare_3) instance proof (intro_classes, unfold compare_real_alg[symmetric, abs_def]) show "le_of_comp (\x y. compare (real_of x) (real_of y)) = (\)" by (intro ext, auto simp: compare_real_def comparator_of_def le_of_comp_def less_eq_real_alg_def) show "lt_of_comp (\x y. compare (real_of x) (real_of y)) = (<)" by (intro ext, auto simp: compare_real_def comparator_of_def lt_of_comp_def less_real_alg_def) show "comparator (\x y. compare (real_of x) (real_of y))" unfolding comparator_def proof (intro conjI impI allI) fix x y z :: "real_alg" let ?r = real_of note rc = comparator_compare[where 'a = real, unfolded comparator_def] from rc show "invert_order (compare (?r x) (?r y)) = compare (?r y) (?r x)" by blast from rc show "compare (?r x) (?r y) = Lt \ compare (?r y) (?r z) = Lt \ compare (?r x) (?r z) = Lt" by blast assume "compare (?r x) (?r y) = Eq" with rc have "?r x = ?r y" by blast thus "x = y" unfolding real_of_inj . qed qed end lemma less_eq_real_alg_code[code]: "(less_eq :: real_alg \ real_alg \ bool) = le_of_comp compare" "(less :: real_alg \ real_alg \ bool) = lt_of_comp compare" by (rule ord_defs(1)[symmetric], rule ord_defs(2)[symmetric]) instantiation real_alg :: abs begin definition abs_real_alg :: "real_alg \ real_alg" where "abs_real_alg x = (if real_of x < 0 then uminus x else x)" instance .. end lemma abs_real_alg: "abs (real_of x) = real_of (abs x)" unfolding abs_real_alg_def abs_real_def if_distrib by (auto simp: uminus_real_alg) lemma sgn_real_alg_sound: "sgn x = (if x = 0 then 0 else if 0 < real_of x then 1 else - 1)" (is "_ = ?r") proof - have "real_of (sgn x) = sgn (real_of x)" by (simp add: sgn_real_alg) also have "\ = real_of ?r" unfolding sgn_real_def if_distrib by (auto simp: less_real_alg_def zero_real_alg_def one_real_alg_def of_rat_real_alg[symmetric] equal_real_alg[symmetric] equal_real_def uminus_real_alg[symmetric]) finally show "sgn x = ?r" unfolding equal_real_alg[symmetric] equal_real_def by simp qed lemma real_of_of_int: "real_of_rat (rat_of_int z) = real_of (of_int z)" proof (cases "z \ 0") case True define n where "n = nat z" from True have z: "z = int n" unfolding n_def by simp show ?thesis unfolding z by (induct n, auto simp: zero_real_alg plus_real_alg[symmetric] one_real_alg hom_distribs) next case False define n where "n = nat (-z)" from False have z: "z = - int n" unfolding n_def by simp show ?thesis unfolding z by (induct n, auto simp: zero_real_alg plus_real_alg[symmetric] one_real_alg uminus_real_alg[symmetric] minus_real_alg[symmetric] hom_distribs) qed instance real_alg :: linordered_field apply standard apply (unfold less_eq_real_alg_def plus_real_alg[symmetric], force) apply (unfold abs_real_alg_def less_real_alg_def zero_real_alg[symmetric], rule refl) apply (unfold less_real_alg_def times_real_alg[symmetric], force) apply (rule sgn_real_alg_sound) done instantiation real_alg :: floor_ceiling begin lift_definition floor_real_alg :: "real_alg \ int" is floor_3 by (auto simp: floor_3) lemma floor_real_alg: "floor (real_of x) = floor x" by (transfer, auto simp: floor_3) instance proof fix x :: real_alg show "of_int \x\ \ x \ x < of_int (\x\ + 1)" unfolding floor_real_alg[symmetric] using floor_correct[of "real_of x"] unfolding less_eq_real_alg_def less_real_alg_def real_of_of_int[symmetric] by (auto simp: hom_distribs) hence "x \ of_int (\x\ + 1)" by auto thus "\z. x \ of_int z" by blast qed end definition real_alg_of_real :: "real \ real_alg" where "real_alg_of_real x = (if (\ y. x = real_of y) then (THE y. x = real_of y) else 0)" lemma real_alg_of_real_code[code]: "real_alg_of_real (real_of x) = x" using real_of_inj unfolding real_alg_of_real_def by auto lift_definition to_rat_real_alg_main :: "real_alg \ rat option" is to_rat_3 by (simp add: to_rat_3) lemma to_rat_real_alg_main: "to_rat_real_alg_main x = (if real_of x \ \ then Some (THE q. real_of x = of_rat q) else None)" by (transfer, simp add: to_rat_3) definition to_rat_real_alg :: "real_alg \ rat" where "to_rat_real_alg x = (case to_rat_real_alg_main x of Some q \ q | None \ 0)" definition is_rat_real_alg :: "real_alg \ bool" where "is_rat_real_alg x = (case to_rat_real_alg_main x of Some q \ True | None \ False)" lemma is_rat_real_alg: "is_rat (real_of x) = (is_rat_real_alg x)" unfolding is_rat_real_alg_def is_rat to_rat_real_alg_main by auto lemma to_rat_real_alg: "to_rat (real_of x) = (to_rat_real_alg x)" unfolding to_rat to_rat_real_alg_def to_rat_real_alg_main by auto subsection \Real Algebraic Numbers as Implementation for Real Numbers\ lemmas real_alg_code_eqns = one_real_alg zero_real_alg uminus_real_alg root_real_alg minus_real_alg plus_real_alg times_real_alg inverse_real_alg divide_real_alg equal_real_alg less_real_alg less_eq_real_alg compare_real_alg sgn_real_alg abs_real_alg floor_real_alg is_rat_real_alg to_rat_real_alg code_datatype real_of declare [[code drop: "plus :: real \ real \ real" "uminus :: real \ real" "minus :: real \ real \ real" "times :: real \ real \ real" "inverse :: real \ real" "divide :: real \ real \ real" "floor :: real \ int" "HOL.equal :: real \ real \ bool" "compare :: real \ real \ order" "less_eq :: real \ real \ bool" "less :: real \ real \ bool" "0 :: real" "1 :: real" "sgn :: real \ real" "abs :: real \ real" root]] declare real_alg_code_eqns [code equation] lemma [code]: "Ratreal = real_of \ of_rat_real_alg" by (transfer, transfer) (simp add: fun_eq_iff of_rat_2) end diff --git a/thys/Gromov_Hyperbolicity/Morse_Gromov_Theorem.thy b/thys/Gromov_Hyperbolicity/Morse_Gromov_Theorem.thy --- a/thys/Gromov_Hyperbolicity/Morse_Gromov_Theorem.thy +++ b/thys/Gromov_Hyperbolicity/Morse_Gromov_Theorem.thy @@ -1,3435 +1,3436 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) theory Morse_Gromov_Theorem imports "HOL-Decision_Procs.Approximation" Gromov_Hyperbolicity Hausdorff_Distance begin hide_const (open) Approximation.Min hide_const (open) Approximation.Max section \Quasiconvexity\ text \In a Gromov-hyperbolic setting, convexity is not a well-defined notion as everything should be coarse. The good replacement is quasi-convexity: A set $X$ is $C$-quasi-convex if any pair of points in $X$ can be joined by a geodesic that remains within distance $C$ of $X$. One could also require this for all geodesics, up to changing $C$, as two geodesics between the same endpoints remain within uniformly bounded distance. We use the first definition to ensure that a geodesic is $0$-quasi-convex.\ definition quasiconvex::"real \ ('a::metric_space) set \ bool" where "quasiconvex C X = (C \ 0 \ (\x\X. \y\X. \G. geodesic_segment_between G x y \ (\z\G. infdist z X \ C)))" lemma quasiconvexD: assumes "quasiconvex C X" "x \ X" "y \ X" shows "\G. geodesic_segment_between G x y \ (\z\G. infdist z X \ C)" using assms unfolding quasiconvex_def by auto lemma quasiconvexC: assumes "quasiconvex C X" shows "C \ 0" using assms unfolding quasiconvex_def by auto lemma quasiconvexI: assumes "C \ 0" "\x y. x \ X \ y \ X \ (\G. geodesic_segment_between G x y \ (\z\G. infdist z X \ C))" shows "quasiconvex C X" using assms unfolding quasiconvex_def by auto lemma quasiconvex_of_geodesic: assumes "geodesic_segment G" shows "quasiconvex 0 G" proof (rule quasiconvexI, simp) fix x y assume *: "x \ G" "y \ G" obtain H where H: "H \ G" "geodesic_segment_between H x y" using geodesic_subsegment_exists[OF assms(1) *] by auto have "infdist z G \ 0" if "z \ H" for z using H(1) that by auto then show "\H. geodesic_segment_between H x y \ (\z\H. infdist z G \ 0)" using H(2) by auto qed lemma quasiconvex_empty: assumes "C \ 0" shows "quasiconvex C {}" unfolding quasiconvex_def using assms by auto lemma quasiconvex_mono: assumes "C \ D" "quasiconvex C G" shows "quasiconvex D G" using assms unfolding quasiconvex_def by (auto, fastforce) text \The $r$-neighborhood of a quasi-convex set is still quasi-convex in a hyperbolic space, for a constant that does not depend on $r$.\ lemma (in Gromov_hyperbolic_space_geodesic) quasiconvex_thickening: assumes "quasiconvex C (X::'a set)" "r \ 0" shows "quasiconvex (C + 8 *deltaG(TYPE('a))) (\x\X. cball x r)" proof (rule quasiconvexI) show "C + 8 *deltaG(TYPE('a)) \ 0" using quasiconvexC[OF assms(1)] by simp next fix y z assume *: "y \ (\x\X. cball x r)" "z \ (\x\X. cball x r)" have A: "infdist w (\x\X. cball x r) \ C + 8 * deltaG TYPE('a)" if "w \ {y--z}" for w proof - obtain py where py: "py \ X" "y \ cball py r" using * by auto obtain pz where pz: "pz \ X" "z \ cball pz r" using * by auto obtain G where G: "geodesic_segment_between G py pz" "(\p\G. infdist p X \ C)" using quasiconvexD[OF assms(1) \py \ X\ \pz \ X\] by auto have A: "infdist w ({y--py} \ G \ {pz--z}) \ 8 * deltaG(TYPE('a))" by (rule thin_quadrilaterals[OF _ G(1) _ _ \w \ {y--z}\, where ?x = y and ?t = z], auto) have "\u \ {y--py} \ G \ {pz--z}. infdist w ({y--py} \ G \ {pz--z}) = dist w u" apply (rule infdist_proper_attained, auto intro!: proper_Un simp add: geodesic_segment_topology(7)) by (meson G(1) geodesic_segmentI geodesic_segment_topology(7)) then obtain u where u: "u \ {y--py} \ G \ {pz--z}" "infdist w ({y--py} \ G \ {pz--z}) = dist w u" by auto then consider "u \ {y--py}" | "u \ G" | "u \ {pz--z}" by auto then have "infdist u (\x\X. cball x r) \ C" proof (cases) case 1 then have "dist py u \ dist py y" using geodesic_segment_dist_le local.some_geodesic_is_geodesic_segment(1) some_geodesic_commute some_geodesic_endpoints(1) by blast also have "... \ r" using py(2) by auto finally have "u \ cball py r" by auto then have "u \ (\x\X. cball x r)" using py(1) by auto then have "infdist u (\x\X. cball x r) = 0" by auto then show ?thesis using quasiconvexC[OF assms(1)] by auto next case 3 then have "dist pz u \ dist pz z" using geodesic_segment_dist_le local.some_geodesic_is_geodesic_segment(1) some_geodesic_commute some_geodesic_endpoints(1) by blast also have "... \ r" using pz(2) by auto finally have "u \ cball pz r" by auto then have "u \ (\x\X. cball x r)" using pz(1) by auto then have "infdist u (\x\X. cball x r) = 0" by auto then show ?thesis using quasiconvexC[OF assms(1)] by auto next case 2 have "infdist u (\x\X. cball x r) \ infdist u X" apply (rule infdist_mono) using assms(2) py(1) by auto then show ?thesis using 2 G(2) by auto qed moreover have "infdist w (\x\X. cball x r) \ infdist u (\x\X. cball x r) + dist w u" by (intro mono_intros) ultimately show ?thesis using A u(2) by auto qed show "\G. geodesic_segment_between G y z \ (\w\G. infdist w (\x\X. cball x r) \ C + 8 * deltaG TYPE('a))" apply (rule exI[of _ "{y--z}"]) using A by auto qed text \If $x$ has a projection $p$ on a quasi-convex set $G$, then all segments from a point in $G$ to $x$ go close to $p$, i.e., the triangular inequality $d(x,y) \leq d(x,p) + d(p,y)$ is essentially an equality, up to an additive constant.\ lemma (in Gromov_hyperbolic_space_geodesic) dist_along_quasiconvex: assumes "quasiconvex C G" "p \ proj_set x G" "y \ G" shows "dist x p + dist p y \ dist x y + 4 * deltaG(TYPE('a)) + 2 * C" proof - have *: "p \ G" using assms proj_setD by auto obtain H where H: "geodesic_segment_between H p y" "\q. q \ H \ infdist q G \ C" using quasiconvexD[OF assms(1) * assms(3)] by auto have "\m\H. infdist x H = dist x m" apply (rule infdist_proper_attained[of H x]) using geodesic_segment_topology[OF geodesic_segmentI[OF H(1)]] by auto then obtain m where m: "m \ H" "infdist x H = dist x m" by auto then have I: "dist x m \ Gromov_product_at x p y + 2 * deltaG(TYPE('a))" using infdist_triangle_side[OF H(1), of x] by auto have "dist x p - dist x m - C \ e" if "e > 0" for e proof - have "\r\G. dist m r < infdist m G + e" apply (rule infdist_almost_attained) using \e > 0\ assms(3) by auto then obtain r where r: "r \ G" "dist m r < infdist m G + e" by auto then have *: "dist m r \ C + e" using H(2)[OF \m \ H\] by auto have "dist x p \ dist x r" using \r \ G\ assms(2) proj_set_dist_le by blast also have "... \ dist x m + dist m r" by (intro mono_intros) finally show ?thesis using * by (auto simp add: metric_space_class.dist_commute) qed then have "dist x p - dist x m - C \ 0" using dense_ge by blast then show ?thesis using I unfolding Gromov_product_at_def by (auto simp add: algebra_simps divide_simps) qed text \The next lemma is~\cite[Proposition 10.2.1]{coornaert_delzant_papadopoulos} with better constants. It states that the distance between the projections on a quasi-convex set is controlled by the distance of the original points, with a gain given by the distances of the points to the set.\ lemma (in Gromov_hyperbolic_space_geodesic) proj_along_quasiconvex_contraction: assumes "quasiconvex C G" "px \ proj_set x G" "py \ proj_set y G" shows "dist px py \ max (5 * deltaG(TYPE('a)) + 2 * C) (dist x y - dist px x - dist py y + 10 * deltaG(TYPE('a)) + 4 * C)" proof - have "px \ G" "py \ G" using assms proj_setD by auto have "(dist x px + dist px py - 4 * deltaG(TYPE('a)) - 2 * C) + (dist y py + dist py px - 4 *deltaG(TYPE('a)) - 2 * C) \ dist x py + dist y px" apply (intro mono_intros) using dist_along_quasiconvex[OF assms(1) assms(2) \py \ G\] dist_along_quasiconvex[OF assms(1) assms(3) \px \ G\] by auto also have "... \ max (dist x y + dist py px) (dist x px + dist py y) + 2 * deltaG(TYPE('a))" by (rule hyperb_quad_ineq) finally have *: "dist x px + dist y py + 2 * dist px py \ max (dist x y + dist py px) (dist x px + dist py y) + 10 * deltaG(TYPE('a)) + 4 * C" by (auto simp add: metric_space_class.dist_commute) show ?thesis proof (cases "dist x y + dist py px \ dist x px + dist py y") case True then have "dist x px + dist y py + 2 * dist px py \ dist x y + dist py px + 10 * deltaG(TYPE('a)) + 4 * C" using * by auto then show ?thesis by (auto simp add: metric_space_class.dist_commute) next case False then have "dist x px + dist y py + 2 * dist px py \ dist x px + dist py y + 10 * deltaG(TYPE('a)) + 4 * C" using * by auto then show ?thesis by (simp add: metric_space_class.dist_commute) qed qed text \The projection on a quasi-convex set is $1$-Lipschitz up to an additive error.\ lemma (in Gromov_hyperbolic_space_geodesic) proj_along_quasiconvex_contraction': assumes "quasiconvex C G" "px \ proj_set x G" "py \ proj_set y G" shows "dist px py \ dist x y + 4 * deltaG(TYPE('a)) + 2 * C" proof (cases "dist y py \ dist x px") case True have "dist x px + dist px py \ dist x py + 4 * deltaG(TYPE('a)) + 2 * C" by (rule dist_along_quasiconvex[OF assms(1) assms(2) proj_setD(1)[OF assms(3)]]) also have "... \ (dist x y + dist y py) + 4 * deltaG(TYPE('a)) + 2 * C" by (intro mono_intros) finally show ?thesis using True by auto next case False have "dist y py + dist py px \ dist y px + 4 * deltaG(TYPE('a)) + 2 * C" by (rule dist_along_quasiconvex[OF assms(1) assms(3) proj_setD(1)[OF assms(2)]]) also have "... \ (dist y x + dist x px) + 4 * deltaG(TYPE('a)) + 2 * C" by (intro mono_intros) finally show ?thesis using False by (auto simp add: metric_space_class.dist_commute) qed text \We can in particular specialize the previous statements to geodesics, which are $0$-quasi-convex.\ lemma (in Gromov_hyperbolic_space_geodesic) dist_along_geodesic: assumes "geodesic_segment G" "p \ proj_set x G" "y \ G" shows "dist x p + dist p y \ dist x y + 4 * deltaG(TYPE('a))" using dist_along_quasiconvex[OF quasiconvex_of_geodesic[OF assms(1)] assms(2) assms(3)] by auto lemma (in Gromov_hyperbolic_space_geodesic) proj_along_geodesic_contraction: assumes "geodesic_segment G" "px \ proj_set x G" "py \ proj_set y G" shows "dist px py \ max (5 * deltaG(TYPE('a))) (dist x y - dist px x - dist py y + 10 * deltaG(TYPE('a)))" using proj_along_quasiconvex_contraction[OF quasiconvex_of_geodesic[OF assms(1)] assms(2) assms(3)] by auto lemma (in Gromov_hyperbolic_space_geodesic) proj_along_geodesic_contraction': assumes "geodesic_segment G" "px \ proj_set x G" "py \ proj_set y G" shows "dist px py \ dist x y + 4 * deltaG(TYPE('a))" using proj_along_quasiconvex_contraction'[OF quasiconvex_of_geodesic[OF assms(1)] assms(2) assms(3)] by auto text \If one projects a continuous curve on a quasi-convex set, the image does not have to be connected (the projection is discontinuous), but since the projections of nearby points are within uniformly bounded distance one can find in the projection a point with almost prescribed distance to the starting point, say. For further applications, we also pick the first such point, i.e., all the previous points are also close to the starting point.\ lemma (in Gromov_hyperbolic_space_geodesic) quasi_convex_projection_small_gaps: assumes "continuous_on {a..(b::real)} f" "a \ b" "quasiconvex C G" "\t. t \ {a..b} \ p t \ proj_set (f t) G" "delta > deltaG(TYPE('a))" "d \ {4 * delta + 2 * C..dist (p a) (p b)}" shows "\t \ {a..b}. (dist (p a) (p t) \ {d - 4 * delta - 2 * C .. d}) \ (\s \ {a..t}. dist (p a) (p s) \ d)" proof - have "delta > 0" using assms(5) local.delta_nonneg by linarith moreover have "C \ 0" using quasiconvexC[OF assms(3)] by simp ultimately have "d \ 0" using assms by auto text \The idea is to define the desired point as the last point $u$ for which there is a projection at distance at most $d$ of the starting point. Then the projection can not be much closer to the starting point, or one could point another such point further away by almost continuity, giving a contradiction. The technical implementation requires some care, as the "last point" may not satisfy the property, for lack of continuity. If it does, then fine. Otherwise, one should go just a little bit to its left to find the desired point.\ define I where "I = {t \ {a..b}. \s \ {a..t}. dist (p a) (p s) \ d}" have "a \ I" using \a \ b\ \d \ 0\ unfolding I_def by auto have "bdd_above I" unfolding I_def by auto define u where "u = Sup I" have "a \ u" unfolding u_def apply (rule cSup_upper) using \a \ I\ \bdd_above I\ by auto have "u \ b" unfolding u_def apply (rule cSup_least) using \a \ I\ apply auto unfolding I_def by auto have A: "dist (p a) (p s) \ d" if "s < u" "a \ s" for s proof - have "\t\I. s < t" unfolding u_def apply (subst less_cSup_iff[symmetric]) using \a \ I\ \bdd_above I\ using \s < u\ unfolding u_def by auto then obtain t where t: "t \ I" "s < t" by auto then have "s \ {a..t}" using \a \ s\ by auto then show ?thesis using t(1) unfolding I_def by auto qed have "continuous (at u within {a..b}) f" using assms(1) by (simp add: \a \ u\ \u \ b\ continuous_on_eq_continuous_within) then have "\i > 0. \s\{a..b}. dist u s < i \ dist (f u) (f s) < (delta - deltaG(TYPE('a)))" unfolding continuous_within_eps_delta using \deltaG(TYPE('a)) < delta\ by (auto simp add: metric_space_class.dist_commute) then obtain e0 where e0: "e0 > 0" "\s. s \ {a..b} \ dist u s < e0 \ dist (f u) (f s) < (delta - deltaG(TYPE('a)))" by auto show ?thesis proof (cases "dist (p a) (p u) > d") text \First, consider the case where $u$ does not satisfy the defining property. Then the desired point $t$ is taken slightly to its left.\ case True then have "u \ a" using \d \ 0\ by auto then have "a < u" using \a \ u\ by auto define e::real where "e = min (e0/2) ((u-a)/2)" then have "e > 0" using \a < u\ \e0 > 0\ by auto define t where "t = u - e" then have "t < u" using \e > 0\ by auto have "u - b \ e" "e \ u - a" using \e > 0\ \u \ b\ unfolding e_def by (auto simp add: min_def) then have "t \ {a..b}" "t \ {a..t}" unfolding t_def by auto have "dist u t < e0" unfolding t_def e_def dist_real_def using \e0 > 0\ \a \ u\ by auto have *: "\s \ {a..t}. dist (p a) (p s) \ d" using A \t < u\ by auto have "dist (p t) (p u) \ dist (f t) (f u) + 4 * deltaG(TYPE('a)) + 2 * C" apply (rule proj_along_quasiconvex_contraction'[OF \quasiconvex C G\]) using assms (4) \t \ {a..b}\ \a \ u\ \u \ b\ by auto also have "... \ (delta - deltaG(TYPE('a))) + 4 * deltaG(TYPE('a)) + 2 * C" apply (intro mono_intros) using e0(2)[OF \t \ {a..b}\ \dist u t < e0\] by (auto simp add: metric_space_class.dist_commute) finally have I: "dist (p t) (p u) \ 4 * delta + 2 * C" using \delta > deltaG(TYPE('a))\ by simp have "d \ dist (p a) (p u)" using True by auto also have "... \ dist (p a) (p t) + dist (p t) (p u)" by (intro mono_intros) also have "... \ dist (p a) (p t) + 4 * delta + 2 * C" using I by simp finally have **: "d - 4 * delta - 2 * C \ dist (p a) (p t)" by simp show ?thesis apply (rule bexI[OF _ \t \ {a..b}\]) using * ** \t \ {a..b}\ by auto next text \Next, consider the case where $u$ satisfies the defining property. Then we will take $t = u$. The only nontrivial point to check is that the distance of $f(u)$ to the starting point is not too small. For this, we need to separate the case where $u = b$ (in which case one argues directly) and the case where $u < b$, where one can use a point slightly to the right of $u$ which has a projection at distance $ > d$ of the starting point, and use almost continuity.\ case False have B: "dist (p a) (p s) \ d" if "s \ {a..u}" for s proof (cases "s = u") case True show ?thesis unfolding True using False by auto next case False then show ?thesis using that A by auto qed have C: "dist (p a) (p u) \ d - 4 *delta - 2 * C" proof (cases "u = b") case True have "d \ dist (p a) (p b)" using assms by auto also have "... \ dist (p a) (p u) + dist (p u) (p b)" by (intro mono_intros) also have "... \ dist (p a) (p u) + (dist (f u) (f b) + 4 * deltaG TYPE('a) + 2 * C)" apply (intro mono_intros proj_along_quasiconvex_contraction'[OF \quasiconvex C G\]) using assms \a \ u\ \u \ b\ by auto finally show ?thesis unfolding True using \deltaG(TYPE('a)) < delta\ by auto next case False then have "u < b" using \u \ b\ by auto define e::real where "e = min (e0/2) ((b-u)/2)" then have "e > 0" using \u < b\ \e0 > 0\ by auto define v where "v = u + e" then have "u < v" using \e > 0\ by auto have "e \ b - u" "a - u \ e" using \e > 0\ \a \ u\ unfolding e_def by (auto simp add: min_def) then have "v \ {a..b}" unfolding v_def by auto moreover have "v \ I" using \u < v\ \bdd_above I\ cSup_upper not_le unfolding u_def by auto ultimately have "\w \ {a..v}. dist (p a) (p w) > d" unfolding I_def by force then obtain w where w: "w \ {a..v}" "dist (p a) (p w) > d" by auto then have "w \ {a..u}" using B by force then have "u < w" using w(1) by auto have "w \ {a..b}" using w(1) \v \ {a..b}\ by auto have "dist u w = w - u" unfolding dist_real_def using \u < w\ by auto also have "... \ v - u" using w(1) by auto also have "... < e0" unfolding v_def e_def min_def using \e0 > 0\ by auto finally have "dist u w < e0" by simp have "dist (p u) (p w) \ dist (f u) (f w) + 4 * deltaG(TYPE('a)) + 2 * C" apply (rule proj_along_quasiconvex_contraction'[OF \quasiconvex C G\]) using assms \a \ u\ \u \ b\ \w \ {a..b}\ by auto also have "... \ (delta - deltaG(TYPE('a))) + 4 * deltaG(TYPE('a)) + 2 * C" apply (intro mono_intros) using e0(2)[OF \w \ {a..b}\ \dist u w < e0\] by (auto simp add: metric_space_class.dist_commute) finally have I: "dist (p u) (p w) \ 4 * delta + 2 * C" using \delta > deltaG(TYPE('a))\ by simp have "d \ dist (p a) (p u) + dist (p u) (p w)" using w(2) metric_space_class.dist_triangle[of "p a" "p w" "p u"] by auto also have "... \ dist (p a) (p u) + 4 * delta + 2 * C" using I by auto finally show ?thesis by simp qed show ?thesis apply (rule bexI[of _ u]) using B \a \ u\ \u \ b\ C by auto qed qed text \Same lemma, except that one exchanges the roles of the beginning and the end point.\ lemma (in Gromov_hyperbolic_space_geodesic) quasi_convex_projection_small_gaps': assumes "continuous_on {a..(b::real)} f" "a \ b" "quasiconvex C G" "\x. x \ {a..b} \ p x \ proj_set (f x) G" "delta > deltaG(TYPE('a))" "d \ {4 * delta + 2 * C..dist (p a) (p b)}" shows "\t \ {a..b}. dist (p b) (p t) \ {d - 4 * delta - 2 * C .. d} \ (\s \ {t..b}. dist (p b) (p s) \ d)" proof - have *: "continuous_on {-b..-a} (\t. f(-t))" using continuous_on_compose[of "{-b..-a}" "\t. -t" f] using assms(1) continuous_on_minus[OF continuous_on_id] by auto define q where "q = (\t. p(-t))" have "\t \ {-b..-a}. (dist (q (-b)) (q t) \ {d - 4 * delta - 2 * C .. d}) \ (\s \ {-b..t}. dist (q (-b)) (q s) \ d)" apply (rule quasi_convex_projection_small_gaps[where ?f = "\t. f(-t)" and ?G = G]) unfolding q_def using assms * by (auto simp add: metric_space_class.dist_commute) then obtain t where t: "t \ {-b..-a}" "dist (q (-b)) (q t) \ {d - 4 * delta - 2 * C .. d}" "\s. s \ {-b..t} \ dist (q (-b)) (q s) \ d" by blast have *: "dist (p b) (p s) \ d" if "s \ {-t..b}" for s using t(3)[of "-s"] that q_def by auto show ?thesis apply (rule bexI[of _ "-t"]) using t * q_def by auto qed section \The Morse-Gromov Theorem\ text \The goal of this section is to prove a central basic result in the theory of hyperbolic spaces, usually called the Morse Lemma. It is really a theorem, and we add the name Gromov the avoid the confusion with the other Morse lemma on the existence of good coordinates for $C^2$ functions with non-vanishing hessian. It states that a quasi-geodesic remains within bounded distance of a geodesic with the same endpoints, the error depending only on $\delta$ and on the parameters $(\lambda, C)$ of the quasi-geodesic, but not on its length. There are several proofs of this result. We will follow the one of Shchur~\cite{shchur}, which gets an optimal dependency in terms of the parameters of the quasi-isometry, contrary to all previous proofs. The price to pay is that the proof is more involved (relying in particular on the fact that the closest point projection on quasi-convex sets is exponentially contracting). We will also give afterwards for completeness the proof in~\cite{bridson_haefliger}, as it brings up interesting tools, although the dependency it gives is worse.\ text \The next lemma (for $C = 0$, Lemma 2 in~\cite{shchur}) asserts that, if two points are not too far apart (at distance at most $10 \delta$), and far enough from a given geodesic segment, then when one moves towards this geodesic segment by a fixed amount (here $5 \delta$), then the two points become closer (the new distance is at most $5 \delta$, gaining a factor of $2$). Later, we will iterate this lemma to show that the projection on a geodesic segment is exponentially contracting. For the application, we give a more general version involving an additional constant $C$. This lemma holds for $\delta$ the hyperbolicity constant. We will want to apply it with $\delta > 0$, so to avoid problems in the case $\delta = 0$ we formulate it not using the hyperbolicity constant of the given type, but any constant which is at least the hyperbolicity constant (this is to work around the fact that one can not say or use easily in Isabelle that a type with hyperbolicity $\delta$ is also hyperbolic for any larger constant $\delta'$.\ lemma (in Gromov_hyperbolic_space_geodesic) geodesic_projection_exp_contracting_aux: assumes "geodesic_segment G" "px \ proj_set x G" "py \ proj_set y G" "delta \ deltaG(TYPE('a))" "dist x y \ 10 * delta + C" "M \ 15/2 * delta" "dist px x \ M + 5 * delta + C/2" "dist py y \ M + 5 * delta + C/2" "C \ 0" shows "dist (geodesic_segment_param {px--x} px M) (geodesic_segment_param {py--y} py M) \ 5 * delta" proof - have "dist px x \ dist py x" using proj_setD(2)[OF assms(2)] infdist_le[OF proj_setD(1)[OF assms(3)], of x] by (simp add: metric_space_class.dist_commute) have "dist py y \ dist px y" using proj_setD(2)[OF assms(3)] infdist_le[OF proj_setD(1)[OF assms(2)], of y] by (simp add: metric_space_class.dist_commute) have "delta \ 0" using assms local.delta_nonneg by linarith then have M: "M \ 0" "M \ dist px x" "M \ dist px y" "M \ dist py x" "M \ dist py y" using assms \dist px x \ dist py x\ \dist py y \ dist px y \by auto have "px \ G" "py \ G" using assms proj_setD by auto define x' where "x' = geodesic_segment_param {px--x} px M" define y' where "y' = geodesic_segment_param {py--y} py M" text \First step: the distance between $px$ and $py$ is at most $5\delta$.\ have "dist px py \ max (5 * deltaG(TYPE('a))) (dist x y - dist px x - dist py y + 10 * deltaG(TYPE('a)))" by (rule proj_along_geodesic_contraction[OF assms(1) assms(2) assms(3)]) also have "... \ max (5 * deltaG(TYPE('a))) (5 * deltaG(TYPE('a)))" apply (intro mono_intros) using assms \delta \ 0\ by auto finally have "dist px py \ 5 * delta" using \delta \ deltaG(TYPE('a))\ by auto text \Second step: show that all the interesting Gromov products at bounded below by $M$.\ have *: "x' \ {px--x}" unfolding x'_def by (simp add: geodesic_segment_param_in_segment) have "px \ proj_set x' G" by (rule proj_set_geodesic_same_basepoint[OF \px \ proj_set x G\ _ *], auto) have "dist px x' = M" unfolding x'_def using M by auto have "dist px x' \ dist py x'" using proj_setD(2)[OF \px \ proj_set x' G\] infdist_le[OF proj_setD(1)[OF assms(3)], of x'] by (simp add: metric_space_class.dist_commute) have **: "dist px x = dist px x' + dist x' x" using geodesic_segment_dist[OF _ *, of px x] by auto have Ixx: "Gromov_product_at px x' x = M" unfolding Gromov_product_at_def ** x'_def using M by auto have "2 * M = dist px x' + dist px x - dist x' x" unfolding ** x'_def using M by auto also have "... \ dist py x' + dist py x - dist x' x" apply (intro mono_intros, auto) by fact+ also have "... = 2 * Gromov_product_at py x x'" unfolding Gromov_product_at_def by (auto simp add: metric_space_class.dist_commute) finally have Iyx: "Gromov_product_at py x x' \ M" by auto have *: "y' \ {py--y}" unfolding y'_def by (simp add: geodesic_segment_param_in_segment) have "py \ proj_set y' G" by (rule proj_set_geodesic_same_basepoint[OF \py \ proj_set y G\ _ *], auto) have "dist py y' = M" unfolding y'_def using M by auto have "dist py y' \ dist px y'" using proj_setD(2)[OF \py \ proj_set y' G\] infdist_le[OF proj_setD(1)[OF assms(2)], of y'] by (simp add: metric_space_class.dist_commute) have **: "dist py y = dist py y' + dist y' y" using geodesic_segment_dist[OF _ *, of py y] by auto have Iyy: "Gromov_product_at py y' y = M" unfolding Gromov_product_at_def ** y'_def using M by auto have "2 * M = dist py y' + dist py y - dist y' y" unfolding ** y'_def using M by auto also have "... \ dist px y' + dist px y - dist y' y" apply (intro mono_intros, auto) by fact+ also have "... = 2 * Gromov_product_at px y y'" unfolding Gromov_product_at_def by (auto simp add: metric_space_class.dist_commute) finally have Ixy: "Gromov_product_at px y y' \ M" by auto have "2 * M \ dist px x + dist py y - dist x y" using assms by auto also have "... \ dist px x + dist px y - dist x y" by (intro mono_intros, fact) also have "... = 2 * Gromov_product_at px x y" unfolding Gromov_product_at_def by auto finally have Ix: "Gromov_product_at px x y \ M" by auto have "2 * M \ dist px x + dist py y - dist x y" using assms by auto also have "... \ dist py x + dist py y - dist x y" by (intro mono_intros, fact) also have "... = 2 * Gromov_product_at py x y" unfolding Gromov_product_at_def by auto finally have Iy: "Gromov_product_at py x y \ M" by auto text \Third step: prove the estimate\ have "M - 2 * delta \ Min {Gromov_product_at px x' x, Gromov_product_at px x y, Gromov_product_at px y y'} - 2 * deltaG(TYPE('a))" using Ixx Ixy Ix \delta \ deltaG(TYPE('a))\ by auto also have "... \ Gromov_product_at px x' y'" by (intro mono_intros) finally have A: "M - 4 * delta + dist x' y' \ dist px y'" unfolding Gromov_product_at_def \dist px x' = M\ by auto have "M - 2 * delta \ Min {Gromov_product_at py x' x, Gromov_product_at py x y, Gromov_product_at py y y'} - 2 * deltaG(TYPE('a))" using Iyx Iyy Iy \delta \ deltaG(TYPE('a))\ by (auto simp add: Gromov_product_commute) also have "... \ Gromov_product_at py x' y'" by (intro mono_intros) finally have B: "M - 4 * delta + dist x' y' \ dist py x'" unfolding Gromov_product_at_def \dist py y' = M\ by auto have "dist px py \ 2 * M - 10 * delta" using assms \dist px py \ 5 * delta\ by auto have "2 * M - 8 * delta + 2 * dist x' y' \ dist px y' + dist py x'" using A B by auto also have "... \ max (dist px py + dist y' x') (dist px x' + dist y' py) + 2 * deltaG TYPE('a)" by (rule hyperb_quad_ineq) also have "... \ max (dist px py + dist y' x') (dist px x' + dist y' py) + 2 * delta" using \deltaG(TYPE('a)) \ delta\ by auto finally have "2 * M - 10 * delta + 2 * dist x' y' \ max (dist px py + dist y' x') (dist px x' + dist y' py)" by auto then have "2 * M - 10 * delta + 2 * dist x' y' \ dist px x' + dist py y'" apply (auto simp add: metric_space_class.dist_commute) using \0 \ delta\ \dist px py \ 2 * M - 10 * delta\ \dist px x' = M\ \dist py y' = M\ by auto then have "dist x' y' \ 5 * delta" unfolding \dist px x' = M\ \dist py y' = M\ by auto then show ?thesis unfolding x'_def y'_def by auto qed text \The next lemma (Lemma 10 in~\cite{shchur} for $C = 0$) asserts that the projection on a geodesic segment is an exponential contraction. More precisely, if a path of length $L$ is at distance at least $D$ of a geodesic segment $G$, then the projection of the path on $G$ has diameter at most $C L \exp(-c D/\delta)$, where $C$ and $c$ are universal constants. This is not completely true at one can not go below a fixed size, as always, so the correct bound is $K \max(\delta, L \exp(-c D/\delta))$. For the application, we give a slightly more general statement involving an additional constant $C$. This statement follows from the previous lemma: if one moves towards $G$ by $10 \delta$, then the distance between points is divided by $2$. Then one iterates this statement as many times as possible, gaining a factor $2$ each time and therefore an exponential factor in the end.\ lemma (in Gromov_hyperbolic_space_geodesic) geodesic_projection_exp_contracting: assumes "geodesic_segment G" "\x y. x \ {a..b} \ y \ {a..b} \ dist (f x) (f y) \ lambda * dist x y + C" "a \ b" "pa \ proj_set (f a) G" "pb \ proj_set (f b) G" "\t. t \ {a..b} \ infdist (f t) G \ D" "D \ 15/2 * delta + C/2" "delta > deltaG(TYPE('a))" "C \ 0" "lambda \ 0" shows "dist pa pb \ max (5 * deltaG(TYPE('a))) ((4 * exp(1/2 * ln 2)) * lambda * (b-a) * exp(-(D-C/2) * ln 2 / (5 * delta)))" proof - have "delta > 0" using assms using local.delta_nonneg by linarith have "exp(15/2/5 * ln 2) = exp(ln 2) * exp(1/2 * ln (2::real))" unfolding mult_exp_exp by simp also have "... = 2 * exp(1/2 * ln 2)" by auto finally have "exp(15/2/5 * ln 2) = 2 * exp(1/2 * ln (2::real))" by simp text \The idea of the proof is to start with a sequence of points separated by $10 \delta + C$ along the original path, and push them by a fixed distance towards $G$ to bring them at distance at most $5 \delta$, thanks to the previous lemma. Then, discard half the points, and start again. This is possible while one is far enough from $G$. In the first step of the proof, we formalize this in the case where the process can be iterated long enough that, at the end, the projections on $G$ are very close together. This is a simple induction, based on the previous lemma.\ have Main: "\c g p. (\i \ {0..2^k}. p i \ proj_set (g i) G) \ (\i \ {0..2^k}. dist (p i) (g i) \ 5 * delta * k + 15/2 * delta + c/2) \ (\i \ {0..<2^k}. dist (g i) (g (Suc i)) \ 10 * delta + c) \ c \ 0 \ dist (p 0) (p (2^k)) \ 5 * deltaG(TYPE('a))" for k proof (induction k) case 0 then have H: "p 0 \ proj_set (g 0) G" "p 1 \ proj_set (g 1) G" "dist (g 0) (g 1) \ 10 * delta + c" "dist (p 0) (g 0) \ 15/2 * delta + c/2" "dist (p 1) (g 1) \ 15/2 * delta + c/2" by auto have "dist (p 0) (p 1) \ max (5 * deltaG(TYPE('a))) (dist (g 0) (g 1) - dist (p 0) (g 0) - dist (p 1) (g 1) + 10 * deltaG(TYPE('a)))" by (rule proj_along_geodesic_contraction[OF \geodesic_segment G\ \p 0 \ proj_set (g 0) G\ \p 1 \ proj_set (g 1) G\]) also have "... \ max (5 * deltaG(TYPE('a))) (5 * deltaG(TYPE('a)))" apply (intro mono_intros) using H \delta > deltaG(TYPE('a))\ by auto finally show "dist (p 0) (p (2^0)) \ 5 * deltaG(TYPE('a))" by auto next case (Suc k) have *: "5 * delta * real (k + 1) + 5 * delta = 5 * delta * real (Suc k + 1)" by (simp add: algebra_simps) define h where "h = (\i. geodesic_segment_param {p i--g i} (p i) (5 * delta * k + 15/2 * delta))" have h_dist: "dist (h i) (h (Suc i)) \ 5 * delta" if "i \ {0..<2^(Suc k)}" for i unfolding h_def apply (rule geodesic_projection_exp_contracting_aux[OF \geodesic_segment G\ _ _ less_imp_le[OF \delta > deltaG(TYPE('a))\]]) unfolding * using Suc.prems that \delta > 0\ by (auto simp add: algebra_simps divide_simps) define g' where "g' = (\i. h (2 * i))" define p' where "p' = (\i. p (2 * i))" have "dist (p' 0) (p' (2^k)) \ 5 * deltaG(TYPE('a))" proof (rule Suc.IH[where ?g = g' and ?c = 0]) show "\i\{0..2 ^ k}. p' i \ proj_set (g' i) G" proof fix i::nat assume "i \ {0..2^k}" then have *: "2 * i \ {0..2^(Suc k)}" by auto show "p' i \ proj_set (g' i) G" unfolding p'_def g'_def h_def apply (rule proj_set_geodesic_same_basepoint[of _ "g (2 * i)" _ "{p(2 * i)--g(2 * i)}"]) using Suc * by (auto simp add: geodesic_segment_param_in_segment) qed show "\i\{0..2 ^ k}. 5 * delta * k + 15/2 * delta + 0/2 \ dist (p' i) (g' i)" proof fix i::nat assume "i \ {0..2^k}" then have *: "2 * i \ {0..2^(Suc k)}" by auto have "5 * delta * k + 15/2 * delta \ 5 * delta * Suc k + 15/2 * delta + c/2" using \delta > 0\ \c \ 0\ by (auto simp add: algebra_simps divide_simps) also have "... \ dist (p (2 * i)) (g (2 * i))" using Suc * by auto finally have *: "5 * delta * k + 15/2 * delta \ dist (p (2 * i)) (g (2 * i))" by simp have "dist (p' i) (g' i) = 5 * delta * k + 15/2 * delta" unfolding p'_def g'_def h_def apply (rule geodesic_segment_param_in_geodesic_spaces(6)) using * \delta > 0\ by auto then show "5 * delta * k + 15/2 * delta + 0/2 \ dist (p' i) (g' i)" by simp qed show "\i\{0..<2 ^ k}. dist (g' i) (g' (Suc i)) \ 10 * delta + 0" proof fix i::nat assume *: "i \ {0..<2 ^ k}" have "dist (g' i) (g' (Suc i)) = dist (h (2 * i)) (h (Suc (Suc (2 * i))))" unfolding g'_def by auto also have "... \ dist (h (2 * i)) (h (Suc (2 * i))) + dist (h (Suc (2 * i))) (h (Suc (Suc (2 * i))))" by (intro mono_intros) also have "... \ 5 * delta + 5 * delta" apply (intro mono_intros h_dist) using * by auto finally show "dist (g' i) (g' (Suc i)) \ 10 * delta + 0" by simp qed qed (simp) then show "dist (p 0) (p (2 ^ Suc k)) \ 5 * deltaG(TYPE('a))" unfolding p'_def by auto qed text \Now, we will apply the previous basic statement to points along our original path. We introduce $k$, the number of steps for which the pushing process can be done -- it only depends on the original distance $D$ to $G$. \ define k where "k = nat(floor((D - C/2 - 15/2 * delta)/(5 * delta)))" have "int k = floor((D - C/2 - 15/2 * delta)/(5 * delta))" unfolding k_def apply (rule nat_0_le) using \D \ 15/2 * delta + C/2\ \delta > 0\ by auto then have "k \ (D - C/2 - 15/2 * delta)/(5 * delta)" "(D - C/2 - 15/2 * delta)/(5 * delta) \ k + 1" by linarith+ then have k: "D \ 5 * delta * k + 15/2 * delta + C/2" "D \ 5 * delta * (k+1) + 15/2 * delta + C/2" using \delta > 0\ by (auto simp add: algebra_simps divide_simps) have "exp((D-C/2)/(5 * delta) * ln 2) * exp(-15/2/5 * ln 2) = exp(((D-C/2-15/2 * delta)/(5 * delta)) * ln 2)" unfolding mult_exp_exp using \delta > 0\ by (simp add: algebra_simps divide_simps) also have "... \ exp((k+1) * ln 2)" apply (intro mono_intros) using k(2) \delta > 0\ by (auto simp add: divide_simps algebra_simps) also have "... = 2^(k+1)" by (subst powr_realpow[symmetric], auto simp add: powr_def) also have "... = 2 * 2^k" by auto finally have k': "1/2^k \ 2 * exp(15/2/5 * ln 2) * exp(- ((D-C/2) * ln 2 / (5 * delta)))" by (auto simp add: algebra_simps divide_simps exp_minus) text \We separate the proof into two cases. If the path is not too long, then it can be covered by $2^k$ points at distance at most $10 \delta + C$. By the basic statement, it follows that the diameter of the projection is at most $5 \delta$. Otherwise, we subdivide the path into $2^N$ points at distance at most $10 \delta + C$, with $N \geq k$, and apply the basic statement to blocks of $2^k$ consecutive points. It follows that the projections of $g_0, g_{2^k}, g_{2\cdot 2^k},\dotsc$ are at distances at most $5 \delta$. Hence, the first and last projections are at distance at most $2^{N-k} \cdot 5 \delta$, which is the desired bound.\ show ?thesis proof (cases "lambda * (b-a) \ 10 * delta * 2^k") text \First, treat the case where the path is rather short.\ case True define g::"nat \ 'a" where "g = (\i. f(a + (b-a) * i/2^k))" have "g 0 = f a" "g(2^k) = f b" unfolding g_def by auto have *: "a + (b-a) * i/2^k \ {a..b}" if "i \ {0..2^k}" for i::nat proof - have "a + (b - a) * (real i / 2 ^ k) \ a + (b-a) * (2^k/2^k)" apply (intro mono_intros) using that \a \ b\ by auto then show ?thesis using \a \ b\ by auto qed have A: "dist (g i) (g (Suc i)) \ 10 * delta + C" if "i \ {0..<2^k}" for i proof - have "dist (g i) (g (Suc i)) \ lambda * dist (a + (b-a) * i/2^k) (a + (b-a) * (Suc i)/2^k) + C" unfolding g_def apply (intro assms(2) *) using that by auto also have "... = lambda * (b-a)/2^k + C" unfolding dist_real_def using \a \ b\ by (auto simp add: algebra_simps divide_simps) also have "... \ 10 * delta + C" using True by (simp add: divide_simps algebra_simps) finally show ?thesis by simp qed define p where "p = (\i. if i = 0 then pa else if i = 2^k then pb else SOME p. p \ proj_set (g i) G)" have B: "p i \ proj_set (g i) G" if "i \ {0..2^k}" for i proof (cases "i = 0 \ i = 2^k") case True then show ?thesis using \pa \ proj_set (f a) G\ \pb \ proj_set (f b) G\ unfolding p_def g_def by auto next case False then have "p i = (SOME p. p \ proj_set (g i) G)" unfolding p_def by auto moreover have "proj_set (g i) G \ {}" apply (rule proj_set_nonempty_of_proper) using geodesic_segment_topology[OF \geodesic_segment G\] by auto ultimately show ?thesis using some_in_eq by auto qed have C: "dist (p i) (g i) \ 5 * delta * k + 15/2 * delta + C/2" if "i \ {0..2^k}" for i proof - have "5 * delta * k + 15/2 * delta + C/2 \ D" using k(1) by simp also have "... \ infdist (g i) G" unfolding g_def apply (rule \\t. t \ {a..b} \ infdist (f t) G \ D\) using * that by auto also have "... = dist (p i) (g i)" using that proj_setD(2)[OF B[OF that]] by (simp add: metric_space_class.dist_commute) finally show ?thesis by simp qed have "dist (p 0) (p (2^k)) \ 5 * deltaG(TYPE('a))" apply (rule Main[where ?g = g and ?c = C]) using A B C \C \ 0\ by auto then show ?thesis unfolding p_def by auto next text \Now, the case where the path is long. We introduce $N$ such that it is roughly of length $2^N \cdot 10 \delta$.\ case False have *: "10 * delta * 2^k \ lambda * (b-a)" using False by simp have "lambda * (b-a) > 0" using \delta > 0\ False \0 \ lambda\ assms(3) less_eq_real_def mult_le_0_iff by auto then have "a < b" "lambda > 0" using \a \ b\ \lambda \ 0\ less_eq_real_def by auto define n where "n = nat(floor(log 2 (lambda * (b-a)/(10 * delta))))" have "log 2 (lambda * (b-a)/(10 * delta)) \ log 2 (2^k)" apply (subst log_le_cancel_iff) using * \delta > 0\ \a < b\ \lambda > 0\ by (auto simp add: divide_simps algebra_simps) moreover have "log 2 (2^k) = k" - by (simp add: log2_of_power_eq) + by simp ultimately have A: "log 2 (lambda * (b-a)/(10 * delta)) \ k" by auto have **: "int n = floor(log 2 (lambda * (b-a)/(10 * delta)))" unfolding n_def apply (rule nat_0_le) using A by auto then have "log 2 (2^n) \ log 2 (lambda * (b-a)/(10 * delta))" apply (subst log_nat_power, auto) by linarith then have I: "2^n \ lambda * (b-a)/(10 * delta)" - using \0 < lambda * (b - a)\ \0 < delta\ by auto + using \0 < lambda * (b - a)\ \0 < delta\ + by (simp add: le_log_iff powr_realpow) have "log 2 (lambda * (b-a)/(10 * delta)) \ log 2 (2^(n+1))" apply (subst log_nat_power, auto) using ** by linarith then have J: "lambda * (b-a)/(10 * delta) \ 2^(n+1)" using \0 < lambda * (b - a)\ \0 < delta\ by auto have K: "k \ n" using A ** by linarith define N where "N = n+1" have N: "k+1 \ N" "lambda * (b-a) / 2^N \ 10 *delta" "2 ^ N \ lambda * (b - a) / (5 * delta)" using I J K \delta > 0\ unfolding N_def by (auto simp add: divide_simps algebra_simps) then have "2 ^ k \ (0::real)" "k \ N" by auto then have "(2^(N-k)::real) = 2^N/2^k" by (metis (no_types) add_diff_cancel_left' le_Suc_ex nonzero_mult_div_cancel_left power_add) text \Define $2^N$ points along the path, separated by at most $10\delta$, and their projections.\ define g::"nat \ 'a" where "g = (\i. f(a + (b-a) * i/2^N))" have "g 0 = f a" "g(2^N) = f b" unfolding g_def by auto have *: "a + (b-a) * i/2^N \ {a..b}" if "i \ {0..2^N}" for i::nat proof - have "a + (b - a) * (real i / 2 ^ N) \ a + (b-a) * (2^N/2^N)" apply (intro mono_intros) using that \a \ b\ by auto then show ?thesis using \a \ b\ by auto qed have A: "dist (g i) (g (Suc i)) \ 10 * delta + C" if "i \ {0..<2^N}" for i proof - have "dist (g i) (g (Suc i)) \ lambda * dist (a + (b-a) * i/2^N) (a + (b-a) * (Suc i)/2^N) + C" unfolding g_def apply (intro assms(2) *) using that by auto also have "... = lambda * (b-a)/2^N + C" unfolding dist_real_def using \a \ b\ by (auto simp add: algebra_simps divide_simps) also have "... \ 10 * delta + C" using N by simp finally show ?thesis by simp qed define p where "p = (\i. if i = 0 then pa else if i = 2^N then pb else SOME p. p \ proj_set (g i) G)" have B: "p i \ proj_set (g i) G" if "i \ {0..2^N}" for i proof (cases "i = 0 \ i = 2^N") case True then show ?thesis using \pa \ proj_set (f a) G\ \pb \ proj_set (f b) G\ unfolding p_def g_def by auto next case False then have "p i = (SOME p. p \ proj_set (g i) G)" unfolding p_def by auto moreover have "proj_set (g i) G \ {}" apply (rule proj_set_nonempty_of_proper) using geodesic_segment_topology[OF \geodesic_segment G\] by auto ultimately show ?thesis using some_in_eq by auto qed have C: "dist (p i) (g i) \ 5 * delta * k + 15/2 * delta + C/2" if "i \ {0..2^N}" for i proof - have "5 * delta * k + 15/2 * delta + C/2 \ D" using k(1) by simp also have "... \ infdist (g i) G" unfolding g_def apply (rule \\t. t \ {a..b} \ infdist (f t) G \ D\) using * that by auto also have "... = dist (p i) (g i)" using that proj_setD(2)[OF B[OF that]] by (simp add: metric_space_class.dist_commute) finally show ?thesis by simp qed text \Use the basic statement to show that, along packets of size $2^k$, the projections are within $5\delta$ of each other.\ have I: "dist (p (2^k * j)) (p (2^k * (Suc j))) \ 5 * delta" if "j \ {0..<2^(N-k)}" for j proof - have I: "i + 2^k * j \ {0..2^N}" if "i \ {0..2^k}" for i proof - have "i + 2 ^ k * j \ 2^k + 2^k * (2^(N-k)-1)" apply (intro mono_intros) using that \j \ {0..<2^(N-k)}\ by auto also have "... = 2^N" using \k +1 \ N\ by (auto simp add: algebra_simps semiring_normalization_rules(26)) finally show ?thesis by auto qed have I': "i + 2^k * j \ {0..<2^N}" if "i \ {0..<2^k}" for i proof - have "i + 2 ^ k * j < 2^k + 2^k * (2^(N-k)-1)" apply (intro mono_intros) using that \j \ {0..<2^(N-k)}\ by auto also have "... = 2^N" using \k +1 \ N\ by (auto simp add: algebra_simps semiring_normalization_rules(26)) finally show ?thesis by auto qed define g' where "g' = (\i. g (i + 2^k * j))" define p' where "p' = (\i. p (i + 2^k * j))" have "dist (p' 0) (p' (2^k)) \ 5 * deltaG(TYPE('a))" apply (rule Main[where ?g = g' and ?c = C]) unfolding p'_def g'_def using A B C I I' \C \ 0\ by auto also have "... \ 5 * delta" using \deltaG(TYPE('a)) < delta\ by auto finally show ?thesis unfolding p'_def by auto qed text \Control the total distance by adding the contributions of blocks of size $2^k$.\ have *: "dist (p 0) (p(2^k * j)) \ (\i dist (p 0) (p(2^k * j)) + dist (p(2^k * j)) (p(2^k * (Suc j)))" by (intro mono_intros) also have "... \ (\iik +1 \ N\ by (auto simp add: semiring_normalization_rules(26)) also have "... \ (\i<2^(N-k). dist (p (2^k * i)) (p (2^k * (Suc i))))" using * by auto also have "... \ (\(i::nat)<2^(N-k). 5 * delta)" apply (rule sum_mono) using I by auto also have "... = 5 * delta * 2^(N-k)" by auto also have "... = 5 * delta * 2^N * (1/ 2^k)" unfolding \(2^(N-k)::real) = 2^N/2^k\ by simp also have "... \ 5 * delta * (2 * lambda * (b-a)/(10 * delta)) * (2 * exp(15/2/5 * ln 2) * exp(- ((D-C/2) * ln 2 / (5 * delta))))" apply (intro mono_intros) using \delta > 0\ \lambda > 0\ \a < b\ k' N by auto also have "... = (2 * exp(15/2/5 * ln 2)) * lambda * (b-a) * exp(-(D-C/2) * ln 2 / (5 * delta))" using \delta > 0\ by (auto simp add: algebra_simps divide_simps) finally show ?thesis unfolding \exp(15/2/5 * ln 2) = 2 * exp(1/2 * ln (2::real))\ by auto qed qed text \We deduce from the previous result that a projection on a quasiconvex set is also exponentially contracting. To do this, one uses the contraction of a projection on a geodesic, and one adds up the additional errors due to the quasi-convexity. In particular, the projections on the original quasiconvex set or the geodesic do not have to coincide, but they are within distance at most $C + 8 \delta$.\ lemma (in Gromov_hyperbolic_space_geodesic) quasiconvex_projection_exp_contracting: assumes "quasiconvex K G" "\x y. x \ {a..b} \ y \ {a..b} \ dist (f x) (f y) \ lambda * dist x y + C" "a \ b" "pa \ proj_set (f a) G" "pb \ proj_set (f b) G" "\t. t \ {a..b} \ infdist (f t) G \ D" "D \ 15/2 * delta + K + C/2" "delta > deltaG(TYPE('a))" "C \ 0" "lambda \ 0" shows "dist pa pb \ 2 * K + 8 * delta + max (5 * deltaG(TYPE('a))) ((4 * exp(1/2 * ln 2)) * lambda * (b-a) * exp(-(D - K - C/2) * ln 2 / (5 * delta)))" proof - obtain H where H: "geodesic_segment_between H pa pb" "\q. q \ H \ infdist q G \ K" using quasiconvexD[OF assms(1) proj_setD(1)[OF \pa \ proj_set (f a) G\] proj_setD(1)[OF \pb \ proj_set (f b) G\]] by auto obtain qa where qa: "qa \ proj_set (f a) H" using proj_set_nonempty_of_proper[of H "f a"] geodesic_segment_topology[OF geodesic_segmentI[OF H(1)]] by auto obtain qb where qb: "qb \ proj_set (f b) H" using proj_set_nonempty_of_proper[of H "f b"] geodesic_segment_topology[OF geodesic_segmentI[OF H(1)]] by auto have I: "infdist (f t) H \ D - K" if "t \ {a..b}" for t proof - have *: "D - K \ dist (f t) h" if "h \ H" for h proof - have "D - K - dist (f t) h \ e" if "e > 0" for e proof - have *: "infdist h G < K + e" using H(2)[OF \h \ H\] \e > 0\ by auto obtain g where g: "g \ G" "dist h g < K + e" using infdist_almost_attained[OF *] proj_setD(1)[OF \pa \ proj_set (f a) G\] by auto have "D \ dist (f t) g" using \\t. t \ {a..b} \ infdist (f t) G \ D\[OF \t \ {a..b}\] infdist_le[OF \g \ G\, of "f t"] by auto also have "... \ dist (f t) h + dist h g" by (intro mono_intros) also have "... \ dist (f t) h + K + e" using g(2) by auto finally show ?thesis by auto qed then have *: "D - K - dist (f t) h \ 0" using dense_ge by blast then show ?thesis by simp qed have "D - K \ Inf (dist (f t) ` H)" apply (rule cInf_greatest) using * H(1) by auto then show "D - K \ infdist (f t) H" apply (subst infdist_notempty) using H(1) by auto qed have Q: "dist qa qb \ max (5 * deltaG(TYPE('a))) ((4 * exp(1/2 * ln 2)) * lambda * (b-a) * exp(-((D - K)-C/2 ) * ln 2 / (5 * delta)))" apply (rule geodesic_projection_exp_contracting[OF geodesic_segmentI[OF \geodesic_segment_between H pa pb\] assms(2) assms(3)]) using qa qb I assms by auto have A: "dist pa qa \ 4 * delta + K" proof - have "dist (f a) pa - dist (f a) qa - K \ e" if "e > 0" for e::real proof - have *: "infdist qa G < K + e" using H(2)[OF proj_setD(1)[OF qa]] \e > 0\ by auto obtain g where g: "g \ G" "dist qa g < K + e" using infdist_almost_attained[OF *] proj_setD(1)[OF \pa \ proj_set (f a) G\] by auto have "dist (f a) pa \ dist (f a) g" unfolding proj_setD(2)[OF \pa \ proj_set (f a) G\] using infdist_le[OF \g \ G\, of "f a"] by simp also have "... \ dist (f a) qa + dist qa g" by (intro mono_intros) also have "... \ dist (f a) qa + K + e" using g(2) by auto finally show ?thesis by simp qed then have I: "dist (f a) pa - dist (f a) qa - K \ 0" using dense_ge by blast have "dist (f a) qa + dist qa pa \ dist (f a) pa + 4 * deltaG(TYPE('a))" apply (rule dist_along_geodesic[OF geodesic_segmentI[OF H(1)]]) using qa H(1) by auto also have "... \ dist (f a) qa + K + 4 * delta" using I assms by auto finally show ?thesis by (simp add: metric_space_class.dist_commute) qed have B: "dist qb pb \ 4 * delta + K" proof - have "dist (f b) pb - dist (f b) qb - K \ e" if "e > 0" for e::real proof - have *: "infdist qb G < K + e" using H(2)[OF proj_setD(1)[OF qb]] \e > 0\ by auto obtain g where g: "g \ G" "dist qb g < K + e" using infdist_almost_attained[OF *] proj_setD(1)[OF \pa \ proj_set (f a) G\] by auto have "dist (f b) pb \ dist (f b) g" unfolding proj_setD(2)[OF \pb \ proj_set (f b) G\] using infdist_le[OF \g \ G\, of "f b"] by simp also have "... \ dist (f b) qb + dist qb g" by (intro mono_intros) also have "... \ dist (f b) qb + K + e" using g(2) by auto finally show ?thesis by simp qed then have I: "dist (f b) pb - dist (f b) qb - K \ 0" using dense_ge by blast have "dist (f b) qb + dist qb pb \ dist (f b) pb + 4 * deltaG(TYPE('a))" apply (rule dist_along_geodesic[OF geodesic_segmentI[OF H(1)]]) using qb H(1) by auto also have "... \ dist (f b) qb + K + 4 * delta" using I assms by auto finally show ?thesis by simp qed have "dist pa pb \ dist pa qa + dist qa qb + dist qb pb" by (intro mono_intros) then show ?thesis using Q A B by auto qed text \The next statement is the main step in the proof of the Morse-Gromov theorem given by Shchur in~\cite{shchur}, asserting that a quasi-geodesic and a geodesic with the same endpoints are close. We show that a point on the quasi-geodesic is close to the geodesic -- the other inequality will follow easily later on. We also assume that the quasi-geodesic is parameterized by a Lipschitz map -- the general case will follow as any quasi-geodesic can be approximated by a Lipschitz map with good controls. Here is a sketch of the proof. Fix two large constants $L \leq D$ (that we will choose carefully to optimize the values of the constants at the end of the proof). Consider a quasi-geodesic $f$ between two points $f(u^-)$ and $f(u^+)$, and a geodesic segment $G$ between the same points. Fix $f(z)$. We want to find a bound on $d(f(z), G)$. 1 - If this distance is smaller than $L$, we are done (and the bound is $L$). 2 - Assume it is larger. Let $\pi_z$ be a projection of $f(z)$ on $G$ (at distance $d(f(z),G)$ of $f(z)$), and $H$ a geodesic between $z$ and $\pi_z$. The idea will be to project the image of $f$ on $H$, and record how much progress is made towards $f(z)$. In this proof, we will construct several points before and after $z$. When necessary, we put an exponent $-$ on the points before $z$, and $+$ on the points after $z$. To ease the reading, the points are ordered following the alphabetical order, i.e., $u^- \leq v \leq w \leq x \leq y^- \leq z$. One can find two points $f(y^-)$ and $f(y^+)$ on the left and the right of $f(z)$ that project on $H$ roughly at distance $L$ of $\pi_z$ (up to some $O(\delta)$ -- recall that the closest point projection is not uniquely defined, and not continuous, so we make some choice here). Let $d^-$ be the minimal distance of $f([u^-, y^-])$ to $H$, and let $d^+$ be the minimal distance of $f([y^+, u^+)]$ to $H$. 2.1 If the two distances $d^-$ and $d^+$ are less than $D$, then the distance between two points realizing the minimum (say $f(c^-)$ and $f(c^+)$) is at most $2D+L$, hence $c^+ - c^-$ is controlled (by $\lambda \cdot (2D+L) + C$) thanks to the quasi-isometry property. Therefore, $f(z)$ is not far away from $f(c^-)$ and $f(c^+)$ (again by the quasi-isometry property). Since the distance from these points to $\pi_z$ is controlled (by $D+L$), we get a good control on $d(f(z),\pi_z)$, as desired. 2.2 The interesting case is when $d^-$ and $d^+$ are both $ > D$. Assume also for instance $d^- \geq d^+$, as the other case is analogous. We will construct two points $f(v)$ and $f(x)$ with $u^- \leq v \leq x \leq y^-$ with the following property: \begin{equation} \label{eq:xvK} K_1 e^{K_2 d(f(v), H)} \leq x-v, \end{equation} where $K_1$ and $K_2$ are some explicit constants (depending on $\lambda$, $\delta$, $L$ and $D$). Let us show how this will conclude the proof. The distance from $f(v)$ to $f(c^+)$ is at most $d(f(v),H) + L + d^+ \leq 3 d(f(v), H)$. Therefore, $c^+ - v$ is also controlled by $K' d(f(v), H)$ by the quasi-isometry property. This gives \begin{align*} K &\leq K (x - v) e^{-K (c^+ - v)} \leq (e^{K (x-v)} - 1) \cdot e^{-K(c^+ - v)} \\& = e^{-K (c^+ - x)} - e^{-K (c^+ - v)} \leq e^{-K(c^+ - x)} - e^{-K (u^+ - u^-)}. \end{align*} This shows that, when one goes from the original quasi-geodesic $f([u^-, u^+])$ to the restricted quasi-geodesic $f([x, c^+])$, the quantity $e^{-K \cdot}$ decreases by a fixed amount. In particular, this process can only happen a uniformly bounded number of times, say $n$. Let $G'$ be a geodesic between $f(x)$ and $f(c^+)$. One checks geometrically that $d(f(z), G) \leq d(f(z), G') + (L + O(\delta))$, as both projections of $f(x)$ and $f(c^+)$ on $H$ are within distance $L$ of $\pi_z$. Iterating the process $n$ times, one gets finally $d(f(z), G) \leq O(1) + n (L + O(\delta))$. This is the desired bound for $d(f(z), G)$. To complete the proof, it remains to construct the points $f(v)$ and $f(x)$ satisfying~\eqref{eq:xvK}. This will be done through an inductive process. Assume first that there is a point $f(v)$ whose projection on $H$ is close to the projection of $f(u^-)$, and with $d(f(v), H) \leq 2 d^-$. Then the projections of $f(v)$ and $f(y^-)$ are far away (at distance at least $L + O(\delta)$). Since the portion of $f$ between $v$ and $y^-$ is everywhere at distance at least $d^-$ of $H$, the projection on $H$ contracts by a factor $e^{-d^-}$. It follows that this portion of $f$ has length at least $e^{d^-} \cdot (L+O(\delta))$. Therefore, by the quasi-isometry property, one gets $x - v \geq K e^{d^-}$. On the other hand, $d(v, H)$ is bounded above by $2 d^-$ by assumption. This gives the desired inequality~\eqref{eq:xvK} with $x = y^-$. Otherwise, all points $f(v)$ whose projection on $H$ is close to the projection of $f(u^-)$ are such that $d(f(v), H) \geq 2 d^-$. Consider $f(w_1)$ a point whose projection on $H$ is at distance roughly $10 \delta$ of the projection of $f(u^-)$. Let $V_1$ be the set of points at distance at most $d^-$ of $H$, i.e., the $d_1$-neighborhood of $H$. Then the distance between the projections of $f(u^-)$ and $f(w_1)$ on $V_1$ is very large (are there is an additional big contraction to go from $V_1$ to $H$). And moreover all the intermediate points $f(v)$ are at distance at least $2 d^-$ of $H$, and therefore at distance at least $d^-$ of $H$. Then one can play the same game as in the first case, where $y^-$ replaced by $w_1$ and $H$ replaced by $V_1$. If there is a point $f(v)$ whose projection on $V_1$ is close to the projection of $f(u^-)$, then the pair of points $v$ and $x = w_1$ works. Otherwise, one lifts everything to $V_2$, the neighborhood of size $2d^-$ of $V_1$, and one argues again in the same way. The induction goes on like this until one finds a suitable pair of points. The process has indeed to stop at one time, as it can only go on while $f(u^-)$ is outside of $V_k$, the $(2^k-1) d^-$ neighborhood of $H$). This concludes the sketch of the proof, modulo the adjustment of constants. Comments on the formalization below: \begin{itemize} \item The proof is written as an induction on $u^+ - u^-$. This makes it possible to either prove the bound directly (in the cases 1 and 2.1 above), or to use the bound on $d(z, G')$ in case 2.2 using the induction assumption, and conclude the proof. Of course, $u^+ - u^-$ is not integer-valued, but in the reduction to $G'$ it decays by a fixed amount, so one can easily write this down as a genuine induction. \item The main difficulty in the proof is to construct the pair $(v, x)$ in case 2.2. This is again written as an induction over $k$: either the required bound is true, or one can find a point $f(w)$ whose projection on $V_k$ is far enough from the projection of $f(u^-)$. Then, either one can use this point to prove the bound, or one can construct a point with the same property with respect to $V_{k+1}$, concluding the induction. \item Instead of writing $u^-$ and $u^+$ (which are not good variable names in Isabelle), we write $um$ and $uM$. Similarly for other variables. \item The proof only works when $\delta > 0$ (as one needs to divide by $\delta$ in the exponential gain). Hence, we formulate it for some $\delta$ which is strictly larger than the hyperbolicity constant. In a subsequent application of the lemma, we will deduce the same statement for the hyperbolicity constant by a limiting argument. \item To optimize the value of the constant in the end, there is an additional important trick with respect to the above sketch: in case 2.2, there is an exponential gain. One can spare a fraction $(1-\alpha)$ of this gain to improve the constants, and spend the remaining fraction $\alpha$ to make the argument work. This makes it possible to reduce the value of the constant roughly from $40000$ to $100$, so we do it in the proof below. The values of $L$, $D$ and $\alpha$ can be chosen freely, and have been chosen to get the best possible constant in the end. \item For another optimization, we do not induce in terms of the distance from $f(z)$ to the geodesic $G$, but rather in terms of the Gromov product $(f(u^-), f(u^+))_{f(z)}$ (which is the same up to $O(\delta)$. And we do not take for $H$ a geodesic from $f(z)$ to its projection on $G$, but rather a geodesic from $f(z)$ to the point $m$ on $[f(u^-), f(u^+)]$ opposite to $f(z)$ in the tripod, i.e., at distance $(f(z), f(u^+))_{f(u^-)}$ of $f(u^-)$, and at distance $(f(z), f(u^-))_{f(u^+)}$ of $f(u^+)$. Let $\pi_z$ denote the point on $[f(z), m]$ at distance $(f(u^-), f(u^+)_{f(z)}$ of $f(z)$. (It is within distance $2 \delta$ of $m$). In both approaches, what we want to control by induction is the distance from $f(z)$ to $\pi_z$. However, in the first approach, the points $f(u^-)$ and $f(u^+)$ project on $H$ between $\pi_z$ and $f(z)$, and since the location of their projection is only controlled up to $4\delta$ one loses essentially a $4\delta$-length of $L$ for the forthcoming argument. In the second approach, the projections on $H$ are on the other side of $\pi_z$ compared to $f(z)$, so one does not lose anything, and in the end it gives genuinely better bounds (making it possible to gain roughly $10 \delta$ in the final estimate). \end{itemize} \ lemma (in Gromov_hyperbolic_space_geodesic) Morse_Gromov_theorem_aux1: fixes f::"real \ 'a" assumes "continuous_on {a..b} f" "lambda C-quasi_isometry_on {a..b} f" "a \ b" "geodesic_segment_between G (f a) (f b)" "z \ {a..b}" "delta > deltaG(TYPE('a))" shows "infdist (f z) G \ lambda^2 * (11/2 * C + 91 * delta)" proof - have "C \ 0" "lambda \ 1" using quasi_isometry_onD assms by auto have "delta > 0" using assms delta_nonneg order_trans by linarith text \We give their values to the parameters $L$, $D$ and $\alpha$ that we will use in the proof. We also define two constants $K$ and $K_{mult}$ that appear in the precise formulation of the bounds. Their values have no precise meaning, they are just the outcome of the computation\ define alpha::real where "alpha = 12/100" have alphaaux:"alpha > 0" "alpha \ 1" unfolding alpha_def by auto define L::real where "L = 18 * delta" define D::real where "D = 55 * delta" define K where "K = alpha * ln 2 / (5 * (4 + (L + 2 * delta)/D) * delta * lambda)" have "K > 0" "L > 0" "D > 0" unfolding K_def L_def D_def using \delta > 0\ \lambda \ 1\ alpha_def by auto have Laux: "L \ 18 * delta" "D \ 50 * delta" "L \ D" "D \ 4 * L" unfolding L_def D_def using \delta > 0\ by auto have Daux: "8 * delta \ (1 - alpha) * D" unfolding alpha_def D_def using \delta > 0\ by auto define Kmult where "Kmult = ((L + 4 * delta)/(L - 13 * delta)) * ((4 * exp(1/2 * ln 2)) * lambda * exp (- (1 - alpha) * D * ln 2 / (5 * delta)) / K)" have "Kmult > 0" unfolding Kmult_def using Laux \delta > 0\ \K > 0\ \lambda \ 1\ by (auto simp add: divide_simps) text \We prove that, for any pair of points to the left and to the right of $f(z)$, the distance from $f(z)$ to a geodesic between these points is controlled. We prove this by reducing to a closer pair of points, i.e., this is an inductive argument over real numbers. However, we formalize it as an artificial induction over natural numbers, as this is how induction works best, and since in our reduction step the new pair of points is always significantly closer than the initial one, at least by an amount $\delta/\lambda$. The main inductive bound that we will prove is the following. In this bound, the first term is what comes from the trivial cases 1 and 2.1 in the description of the proof before the statement of the theorem, while the most interesting term is the second term, corresponding to the induction per se.\ have Main: "\um uM. um \ {a..z} \ uM \ {z..b} \ uM - um \ n * (1/4) * delta / lambda \ Gromov_product_at (f z) (f um) (f uM) \ lambda^2 * (D + (3/2) * L + delta + 11/2 * C) - 2 * delta + Kmult * (1 - exp(- K * (uM - um)))" for n::nat proof (induction n) text \Trivial base case of the induction\ case 0 then have *: "z = um" "z = uM" by auto then have "Gromov_product_at (f z) (f um) (f uM) = 0" by auto also have "... \ 1 * (D + (3/2) * L + delta + 11/2 * C) - 2 * delta + 0 * (1 - exp(- K * (uM - um)))" using Laux \C \ 0\ \delta > 0\ by auto also have "... \ lambda^2 * (D + (3/2) * L + delta + 11/2 * C) - 2 * delta + Kmult * (1 - exp(- K * (uM - um)))" apply (intro mono_intros) using \C \ 0\ \delta > 0\ Laux \D > 0\ \K > 0\ "0.prems" \lambda \ 1\ \Kmult > 0\ by auto finally show ?case by auto next case (Suc n) show ?case proof (cases "Gromov_product_at (f z) (f um) (f uM) \ L") text \If $f(z)$ is already close to the geodesic, there is nothing to do, and we do not need the induction assumption. This is case 1 in the description above.\ case True have "L \ 1 * (D + (3/2) * L + delta + 11/2 * C) - 2 * delta + 0 * (1 - exp(- K * (uM - um)))" using Laux \C \ 0\ \delta > 0\ by auto also have "... \ lambda^2 * (D + (3/2) * L + delta + 11/2 * C) - 2 * delta + Kmult * (1 - exp(- K * (uM - um)))" apply (intro mono_intros) using \C \ 0\ \delta > 0\ Laux \D > 0\ "Suc.prems" \K > 0\ \lambda \ 1\ \Kmult > 0\ by auto finally show ?thesis using True by auto next text \We come to the interesting case where $f(z)$ is far away from a geodesic between $f(um)$ and $f(uM)$. Let $m$ be close to a projection of $f(z)$ on such a geodesic (we use the opposite point of $f(z)$ on the corresponding tripod). On a geodesic between $f(z)$ and $m$, consider the point $pi_z$ at distance $(f(um), f(uM))_{f(z)}$ of $f(z)$. It is very close to $m$ (within distance $2 \delta$). We will push the points $f(um)$ and $f(uM)$ towards $f(z)$ by considering points whose projection on a geodesic $H$ between $m$ and $z$ is roughly at distance $L$ of $pi_z$.\ case False define m where "m = geodesic_segment_param {f um--f uM} (f um) (Gromov_product_at (f um) (f z) (f uM))" have "dist (f z) m \ Gromov_product_at (f z) (f um) (f uM) + 2 * deltaG(TYPE('a))" unfolding m_def by (rule dist_triangle_side_middle, auto) then have *: "dist (f z) m \ Gromov_product_at (f z) (f um) (f uM) + 2 * delta" using \deltaG(TYPE('a)) < delta\ by auto have "Gromov_product_at (f z) (f um) (f uM) \ infdist (f z) {f um--f uM}" by (intro mono_intros, auto) also have "... \ dist (f z) m" apply (rule infdist_le) unfolding m_def by auto finally have **: "Gromov_product_at (f z) (f um) (f uM) \ dist (f z) m" by auto define H where "H = {f z--m}" define pi_z where "pi_z = geodesic_segment_param H (f z) (Gromov_product_at (f z) (f um) (f uM))" have "pi_z \ H" "m \ H" "f z \ H" unfolding pi_z_def H_def by (auto simp add: geodesic_segment_param_in_segment) have H: "geodesic_segment_between H (f z) m" unfolding H_def by auto have Dpi_z: "dist (f z) pi_z = Gromov_product_at (f z) (f um) (f uM)" unfolding pi_z_def H_def by (rule geodesic_segment_param(6)[where ?y = m], auto simp add: **) moreover have "dist (f z) m = dist (f z) pi_z + dist pi_z m" apply (rule geodesic_segment_dist[of H, symmetric]) using \pi_z \ H\ unfolding H_def by auto ultimately have "dist pi_z m \ 2 * delta" using * by auto text \Introduce the notation $p$ for some projection on the geodesic $H$.\ define p where "p = (\r. SOME x. x \ proj_set (f r) H)" have p: "p x \ proj_set (f x) H" for x unfolding p_def using proj_set_nonempty_of_proper[of H "f x"] geodesic_segment_topology[OF geodesic_segmentI[OF H]] by (simp add: some_in_eq) then have pH: "p x \ H" for x using proj_setD(1) by auto have pz: "p z = f z" using p[of z] H by auto text \The projection of $f(um)$ on $H$ is close to $pi_z$ (but it does not have to be exactly $pi_z$). It is between $pi_z$ and $m$.\ have "dist (f um) (f z) \ dist (f um) (p um) + dist (p um) (f z)" by (intro mono_intros) also have "... \ dist (f um) m + dist (p um) (f z)" unfolding proj_setD(2)[OF p[of um]] H_def by (auto intro!: infdist_le) also have "... = Gromov_product_at (f um) (f z) (f uM) + dist (p um) (f z)" unfolding m_def by simp finally have A: "Gromov_product_at (f z) (f um) (f uM) \ dist (p um) (f z)" unfolding Gromov_product_at_def by (simp add: metric_space_class.dist_commute divide_simps) have "dist (p um) pi_z = abs(dist (p um) (f z) - dist pi_z (f z))" apply (rule dist_along_geodesic_wrt_endpoint[of H _ m]) using pH \pi_z \ H\ H_def by auto also have "... = dist (p um) (f z) - dist pi_z (f z)" using A Dpi_z by (simp add: metric_space_class.dist_commute) finally have Dum: "dist (p um) (f z) = dist (p um) pi_z + dist pi_z (f z)" by auto text \Choose a point $f(ym)$ whose projection on $H$ is roughly at distance $L$ of $pi_z$.\ have "\ym \ {um..z}. (dist (p um) (p ym) \ {(L + dist pi_z (p um)) - 4 * delta - 2 * 0 .. L + dist pi_z (p um)}) \ (\r \ {um..ym}. dist (p um) (p r) \ L + dist pi_z (p um))" proof (rule quasi_convex_projection_small_gaps[where ?f = f and ?G = H]) show "continuous_on {um..z} f" apply (rule continuous_on_subset[OF \continuous_on {a..b} f\]) using \um \ {a..z}\ \z \ {a..b}\ by auto show "um \ z" using \um \ {a..z}\ by auto show "quasiconvex 0 H" using quasiconvex_of_geodesic geodesic_segmentI H by auto show "deltaG TYPE('a) < delta" by fact have "L + dist pi_z (p um) \ dist (f z) pi_z + dist pi_z (p um)" using False Dpi_z by (simp add: metric_space_class.dist_commute) then have "L + dist pi_z (p um) \ dist (p um) (f z)" using Dum by (simp add: metric_space_class.dist_commute) then show "L + dist pi_z (p um) \ {4 * delta + 2 * 0..dist (p um) (p z)}" using \delta > 0\ False L_def pz by auto show "p ym \ proj_set (f ym) H" for ym using p by simp qed then obtain ym where ym : "ym \ {um..z}" "dist (p um) (p ym) \ {(L + dist pi_z (p um)) - 4 * delta - 2 * 0 .. L + dist pi_z (p um)}" "\r. r \ {um..ym} \ dist (p um) (p r) \ L + dist pi_z (p um)" by blast have *: "continuous_on {um..ym} (\r. infdist (f r) H)" using continuous_on_infdist[OF continuous_on_subset[OF \continuous_on {a..b} f\, of "{um..ym}"], of H] \ym \ {um..z}\ \um \ {a..z}\ \z \ {a..b}\ by auto text \Choose a point $cm$ between $f(um)$ and $f(ym)$ realizing the minimal distance to $H$. Call this distance $dm$.\ have "\closestm \ {um..ym}. \v \ {um..ym}. infdist (f closestm) H \ infdist (f v) H" apply (rule continuous_attains_inf) using ym(1) * by auto then obtain closestm where closestm: "closestm \ {um..ym}" "\v. v \ {um..ym} \ infdist (f closestm) H \ infdist (f v) H" by auto define dm where "dm = infdist (f closestm) H" have [simp]: "dm \ 0" unfolding dm_def using infdist_nonneg by auto text \Same things but in the interval $[z, uM]$.\ have I: "dist m (f uM) = dist (f um) (f uM) - dist (f um) m" "dist (f um) m = Gromov_product_at (f um) (f z) (f uM)" using geodesic_segment_dist[of "{f um--f uM}" "f um" "f uM" m] m_def by auto have "dist (f uM) (f z) \ dist (f uM) (p uM) + dist (p uM) (f z)" by (intro mono_intros) also have "... \ dist (f uM) m + dist (p uM) (f z)" unfolding proj_setD(2)[OF p[of uM]] H_def by (auto intro!: infdist_le) also have "... = Gromov_product_at (f uM) (f z) (f um) + dist (p uM) (f z)" using I unfolding Gromov_product_at_def by (simp add: divide_simps algebra_simps metric_space_class.dist_commute) finally have A: "Gromov_product_at (f z) (f um) (f uM) \ dist (p uM) (f z)" unfolding Gromov_product_at_def by (simp add: metric_space_class.dist_commute divide_simps) have "dist (p uM) pi_z = abs(dist (p uM) (f z) - dist pi_z (f z))" apply (rule dist_along_geodesic_wrt_endpoint[of H _ m]) using pH \pi_z \ H\ H_def by auto also have "... = dist (p uM) (f z) - dist pi_z (f z)" using A Dpi_z by (simp add: metric_space_class.dist_commute) finally have DuM: "dist (p uM) (f z) = dist (p uM) pi_z + dist pi_z (f z)" by auto text \Choose a point $f(yM)$ whose projection on $H$ is roughly at distance $L$ of $pi_z$.\ have "\yM \ {z..uM}. dist (p uM) (p yM) \ {(L + dist pi_z (p uM)) - 4* delta - 2 * 0 .. L + dist pi_z (p uM)} \ (\r \ {yM..uM}. dist (p uM) (p r) \ L + dist pi_z (p uM))" proof (rule quasi_convex_projection_small_gaps'[where ?f = f and ?G = H]) show "continuous_on {z..uM} f" apply (rule continuous_on_subset[OF \continuous_on {a..b} f\]) using \uM \ {z..b}\ \z \ {a..b}\ by auto show "z \ uM" using \uM \ {z..b}\ by auto show "quasiconvex 0 H" using quasiconvex_of_geodesic geodesic_segmentI H by auto show "deltaG TYPE('a) < delta" by fact have "L + dist pi_z (p uM) \ dist (f z) pi_z + dist pi_z (p uM)" using False Dpi_z by (simp add: metric_space_class.dist_commute) then have "L + dist pi_z (p uM) \ dist (p uM) (f z)" using DuM by (simp add: metric_space_class.dist_commute) then show "L + dist pi_z (p uM) \ {4 * delta + 2 * 0..dist (p z) (p uM)}" using \delta > 0\ False L_def pz by (auto simp add: metric_space_class.dist_commute) show "p yM \ proj_set (f yM) H" for yM using p by simp qed then obtain yM where yM: "yM \ {z..uM}" "dist (p uM) (p yM) \ {(L + dist pi_z (p uM)) - 4* delta - 2 * 0 .. L + dist pi_z (p uM)}" "\r. r \ {yM..uM} \ dist (p uM) (p r) \ L + dist pi_z (p uM)" by blast have *: "continuous_on {yM..uM} (\r. infdist (f r) H)" using continuous_on_infdist[OF continuous_on_subset[OF \continuous_on {a..b} f\, of "{yM..uM}"], of H] \yM \ {z..uM}\ \uM \ {z..b}\ \z \ {a..b}\ by auto have "\closestM \ {yM..uM}. \v \ {yM..uM}. infdist (f closestM) H \ infdist (f v) H" apply (rule continuous_attains_inf) using yM(1) * by auto then obtain closestM where closestM: "closestM \ {yM..uM}" "\v. v \ {yM..uM} \ infdist (f closestM) H \ infdist (f v) H" by auto define dM where "dM = infdist (f closestM) H" have [simp]: "dM \ 0" unfolding dM_def using infdist_nonneg by auto text \Points between $f(um)$ and $f(ym)$, or between $f(yM)$ and $f(uM)$, project within distance at most $L$ of $pi_z$ by construction.\ have P0: "dist m (p x) \ dist m pi_z + L" if "x \ {um..ym} \ {yM..uM}" for x proof (cases "x \ {um..ym}") case True have "dist m (f z) = dist m (p um) + dist (p um) pi_z + dist pi_z (f z)" using geodesic_segment_dist[OF H pH[of um]] Dum by (simp add: metric_space_class.dist_commute) moreover have "dist m (f z) = dist m pi_z + dist pi_z (f z)" using geodesic_segment_dist[OF H \pi_z \ H\] by (simp add: metric_space_class.dist_commute) ultimately have *: "dist m pi_z = dist m (p um) + dist (p um) pi_z" by auto have "dist (p um) (p x) \ L + dist pi_z (p um)" using ym(3)[OF \x \ {um..ym}\] by blast then show ?thesis using metric_space_class.dist_triangle[of m "p x" "p um"] * by (auto simp add: metric_space_class.dist_commute) next case False then have "x \ {yM..uM}" using that by auto have "dist m (f z) = dist m (p uM) + dist (p uM) pi_z + dist pi_z (f z)" using geodesic_segment_dist[OF H pH[of uM]] DuM by (simp add: metric_space_class.dist_commute) moreover have "dist m (f z) = dist m pi_z + dist pi_z (f z)" using geodesic_segment_dist[OF H \pi_z \ H\] by (simp add: metric_space_class.dist_commute) ultimately have *: "dist m pi_z = dist m (p uM) + dist (p uM) pi_z" by auto have "dist (p uM) (p x) \ L + dist pi_z (p uM)" using yM(3)[OF \x \ {yM..uM}\] by blast then show ?thesis using metric_space_class.dist_triangle[of m "p x" "p uM"] * by (auto simp add: metric_space_class.dist_commute) qed have P: "dist pi_z (p x) \ L" if "x \ {um..ym} \ {yM..uM}" for x proof (cases "dist m (p x) \ dist pi_z m") case True have "dist pi_z (p x) \ dist pi_z m + dist m (p x)" by (intro mono_intros) also have "... \ 2 * delta + 2 * delta" using \dist pi_z m \ 2 * delta\ True by auto finally show ?thesis using Laux \delta > 0\ by auto next case False have "dist pi_z (p x) = abs(dist pi_z m - dist (p x) m)" apply (rule dist_along_geodesic_wrt_endpoint[OF geodesic_segment_commute[OF H]]) using pH \pi_z \ H\ by auto also have "... = dist (p x) m - dist pi_z m" using False by (simp add: metric_space_class.dist_commute) finally show ?thesis using P0[OF that] by (simp add: metric_space_class.dist_commute) qed text \Auxiliary fact for later use: The distance between two points in $[um, ym]$ and $[yM, uM]$ can be controlled using the distances of their images under $f$ to $H$, thanks to the quasi-isometry property.\ have D: "dist rm rM \ lambda * (infdist (f rm) H + (L + C + 2 * delta) + infdist (f rM) H)" if "rm \ {um..ym}" "rM \ {yM..uM}" for rm rM proof - have *: "dist m (p rm) \ L + dist m pi_z" "dist m (p rM) \ L + dist m pi_z" using P0 that by force+ have "dist (p rm) (p rM) = abs(dist (p rm) m - dist (p rM) m)" apply (rule dist_along_geodesic_wrt_endpoint[OF geodesic_segment_commute[OF H]]) using pH by auto also have "... \ L + dist m pi_z" unfolding abs_le_iff using * apply (auto simp add: metric_space_class.dist_commute) by (metis diff_add_cancel le_add_same_cancel1 metric_space_class.zero_le_dist order_trans)+ finally have *: "dist (p rm) (p rM) \ L + 2 * delta" using \dist pi_z m \ 2 * delta\ by (simp add: metric_space_class.dist_commute) have "(1/lambda) * dist rm rM - C \ dist (f rm) (f rM)" apply (rule quasi_isometry_onD(2)[OF \lambda C-quasi_isometry_on {a..b} f\]) using \rm \ {um..ym}\ \ym \ {um..z}\ \um \ {a..z}\ \z \ {a..b}\ \rM \ {yM..uM}\ \yM \ {z..uM}\ \uM \ {z..b}\ by auto also have "... \ dist (f rm) (p rm) + dist (p rm) (p rM) + dist (p rM) (f rM)" by (intro mono_intros) also have "... \ infdist (f rm) H + L + 2 * delta + infdist (f rM) H" using * proj_setD(2)[OF p] by (simp add: metric_space_class.dist_commute) finally show ?thesis using \lambda \ 1\ by (simp add: algebra_simps divide_simps) qed text \Auxiliary fact for later use in the inductive argument: the distance from $f(z)$ to $pi_z$ is controlled by the distance from $f(z)$ to any intermediate geodesic between points in $f[um, ym]$ and $f[yM, uM]$, up to a constant essentially given by $L$. This is a variation around Lemma 5 in~\cite{shchur}.\ have Rec: "Gromov_product_at (f z) (f um) (f uM) \ Gromov_product_at (f z) (f rm) (f rM) + (L + 4 * delta)" if "rm \ {um..ym}" "rM \ {yM..uM}" for rm rM proof - have *: "dist (f rm) (p rm) + dist (p rm) (f z) \ dist (f rm) (f z) + 4 * deltaG(TYPE('a))" apply (rule dist_along_geodesic[of H]) using p H_def by auto have "dist (f z) pi_z \ dist (f z) (p rm) + dist (p rm) pi_z" by (intro mono_intros) also have "... \ (Gromov_product_at (f z) (f rm) (p rm) + 2 * deltaG(TYPE('a))) + L" apply (intro mono_intros) using * P \rm \ {um..ym}\ unfolding Gromov_product_at_def by (auto simp add: metric_space_class.dist_commute algebra_simps divide_simps) finally have A: "dist (f z) pi_z - L - 2 * deltaG(TYPE('a)) \ Gromov_product_at (f z) (f rm) (p rm)" by simp have *: "dist (f rM) (p rM) + dist (p rM) (f z) \ dist (f rM) (f z) + 4 * deltaG(TYPE('a))" apply (rule dist_along_geodesic[of H]) using p H_def by auto have "dist (f z) pi_z \ dist (f z) (p rM) + dist (p rM) pi_z" by (intro mono_intros) also have "... \ (Gromov_product_at (f z) (p rM) (f rM) + 2 * deltaG(TYPE('a))) + L" apply (intro mono_intros) using * P \rM \ {yM..uM}\ unfolding Gromov_product_at_def by (auto simp add: metric_space_class.dist_commute algebra_simps divide_simps) finally have B: "dist (f z) pi_z - L - 2 * deltaG(TYPE('a)) \ Gromov_product_at (f z) (p rM) (f rM)" by simp have C: "dist (f z) pi_z - L - 2 * deltaG(TYPE('a)) \ Gromov_product_at (f z) (p rm) (p rM)" proof (cases "dist (f z) (p rm) \ dist (f z) (p rM)") case True have "dist (p rm) (p rM) = abs(dist (f z) (p rm) - dist (f z) (p rM))" using proj_setD(1)[OF p] dist_along_geodesic_wrt_endpoint[OF H, of "p rm" "p rM"] by (simp add: metric_space_class.dist_commute) also have "... = dist (f z) (p rM) - dist (f z) (p rm)" using True by auto finally have *: "dist (f z) (p rm) = Gromov_product_at (f z) (p rm) (p rM)" unfolding Gromov_product_at_def by auto have "dist (f z) pi_z \ dist (f z) (p rm) + dist (p rm) pi_z" by (intro mono_intros) also have "... \ Gromov_product_at (f z) (p rm) (p rM) + L + 2 * deltaG(TYPE('a))" using * P[of rm] \rm \ {um..ym}\ apply (simp add: metric_space_class.dist_commute) using local.delta_nonneg by linarith finally show ?thesis by simp next case False have "dist (p rm) (p rM) = abs(dist (f z) (p rm) - dist (f z) (p rM))" using proj_setD(1)[OF p] dist_along_geodesic_wrt_endpoint[OF H, of "p rm" "p rM"] by (simp add: metric_space_class.dist_commute) also have "... = dist (f z) (p rm) - dist (f z) (p rM)" using False by auto finally have *: "dist (f z) (p rM) = Gromov_product_at (f z) (p rm) (p rM)" unfolding Gromov_product_at_def by auto have "dist (f z) pi_z \ dist (f z) (p rM) + dist (p rM) pi_z" by (intro mono_intros) also have "... \ Gromov_product_at (f z) (p rm) (p rM) + L + 2 * deltaG(TYPE('a))" using * P[of rM] \rM \ {yM..uM}\ apply (simp add: metric_space_class.dist_commute) using local.delta_nonneg by linarith finally show ?thesis by simp qed have "Gromov_product_at (f z) (f um) (f uM) - L - 2 * deltaG(TYPE('a)) \ Min {Gromov_product_at (f z) (f rm) (p rm), Gromov_product_at (f z) (p rm) (p rM), Gromov_product_at (f z) (p rM) (f rM)}" using A B C unfolding Dpi_z by auto also have "... \ Gromov_product_at (f z) (f rm) (f rM) + 2 * deltaG(TYPE('a))" by (intro mono_intros) finally show ?thesis using \deltaG(TYPE('a)) < delta\ by auto qed text \We have proved the basic facts we will need in the main argument. This argument starts here. It is divided in several cases.\ consider "dm \ D + 4 * C \ dM \ D + 4 * C" | "dm \ D + 4 * C \ dM \ dm" | "dM \ D + 4 * C \ dm \ dM" by linarith then show ?thesis proof (cases) text \Case 2.1 of the description before the statement: there are points in $f[um, ym]$ and in $f[yM, uM]$ which are close to $H$. Then one can conclude directly, without relying on the inductive argument, thanks to the quasi-isometry property.\ case 1 have I: "Gromov_product_at (f z) (f closestm) (f closestM) \ lambda^2 * (D + L / 2 + delta + 11/2 * C) - 6 * delta" proof (cases "dist (f closestm) (f closestM) \ 12 * delta") case True have "1/lambda * dist closestm closestM - C \ dist (f closestm) (f closestM)" using quasi_isometry_onD(2)[OF assms(2)] \closestm \ {um..ym}\ \um \ {a..z}\ \z \ {a..b}\ \ym \ {um..z}\ \closestM \ {yM..uM}\ \uM \ {z..b}\ \z \ {a..b}\ \yM \ {z..uM}\ by auto then have "dist closestm closestM \ lambda * dist (f closestm) (f closestM) + lambda * C" using \lambda \ 1\ by (auto simp add: divide_simps algebra_simps) also have "... \ lambda * (12 * delta) + lambda * C" apply (intro mono_intros True) using \lambda \ 1\ by auto finally have M: "dist closestm closestM \ lambda * (12 * delta + C)" by (auto simp add: algebra_simps) have "2 * Gromov_product_at (f z) (f closestm) (f closestM) \ dist (f closestm) (f z) + dist (f z) (f (closestM))" unfolding Gromov_product_at_def by (auto simp add: metric_space_class.dist_commute) also have "... \ (lambda * dist closestm z + C) + (lambda * dist z closestM + C)" apply (intro mono_intros quasi_isometry_onD(1)[OF assms(2)]) using \closestm \ {um..ym}\ \um \ {a..z}\ \z \ {a..b}\ \ym \ {um..z}\ \closestM \ {yM..uM}\ \uM \ {z..b}\ \z \ {a..b}\ \yM \ {z..uM}\ by auto also have "... = lambda * dist closestm closestM + 1 * 2 * C" unfolding dist_real_def using \closestm \ {um..ym}\ \um \ {a..z}\ \z \ {a..b}\ \ym \ {um..z}\ \closestM \ {yM..uM}\ \uM \ {z..b}\ \z \ {a..b}\ \yM \ {z..uM}\ by (auto simp add: algebra_simps) also have "... \ lambda * (lambda * (12 * delta + C)) + lambda^2 * 2 * C" apply (intro mono_intros M) using \lambda \ 1\ \C \ 0\ by auto also have "... = lambda^2 * (24 * delta + 3 * C) - lambda^2 * 12 * delta" by (simp add: algebra_simps power2_eq_square) also have "... \ lambda^2 * ((2 * D + L + 2 * delta) + 11 * C) - 1 * 12 * delta" apply (intro mono_intros) using Laux \lambda \ 1\ \C \ 0\ \delta > 0\ by auto finally show ?thesis by (auto simp add: divide_simps algebra_simps) next case False have "dist closestm closestM \ lambda * (dm + dM + L + 2 * delta + C)" using D[OF \closestm \ {um..ym}\ \closestM \ {yM..uM}\] dm_def dM_def by (auto simp add: algebra_simps) also have "... \ lambda * ((D + 4 * C) + (D + 4 * C) + L + 2 * delta + C)" apply (intro mono_intros) using 1 \lambda \ 1\ by auto also have "... \ lambda * (2 * D + L + 2 * delta + 9 * C)" using \lambda \ 1\ \C \ 0\ by auto finally have M: "dist closestm closestM \ lambda * (2 * D + L + 2 * delta + 9 * C)" by (auto simp add: algebra_simps divide_simps metric_space_class.dist_commute) have "dist (f closestm) (f z) + dist (f z) (f (closestM)) \ (lambda * dist closestm z + C) + (lambda * dist z closestM + C)" apply (intro mono_intros quasi_isometry_onD(1)[OF assms(2)]) using \closestm \ {um..ym}\ \um \ {a..z}\ \z \ {a..b}\ \ym \ {um..z}\ \closestM \ {yM..uM}\ \uM \ {z..b}\ \z \ {a..b}\ \yM \ {z..uM}\ by auto also have "... = lambda * dist closestm closestM + 1 * 2 * C" unfolding dist_real_def using \closestm \ {um..ym}\ \um \ {a..z}\ \z \ {a..b}\ \ym \ {um..z}\ \closestM \ {yM..uM}\ \uM \ {z..b}\ \z \ {a..b}\ \yM \ {z..uM}\ by (auto simp add: algebra_simps) also have "... \ lambda * (lambda * (2 * D + L + 2 * delta + 9 * C)) + lambda^2 * 2 * C" apply (intro mono_intros M) using \lambda \ 1\ \C \ 0\ by auto finally have "dist (f closestm) (f z) + dist (f z) (f closestM) \ lambda^2 * (2 * D + L + 2 * delta + 11 * C)" by (simp add: algebra_simps power2_eq_square) then show ?thesis unfolding Gromov_product_at_def using False by (simp add: metric_space_class.dist_commute algebra_simps divide_simps) qed have "Gromov_product_at (f z) (f um) (f uM) \ Gromov_product_at (f z) (f closestm) (f closestM) + 1 * L + 4 * delta + 0 * (1 - exp (- K * (uM - um)))" using Rec[OF \closestm \ {um..ym}\ \closestM \ {yM..uM}\] by simp also have "... \ (lambda^2 * (D + L / 2 + delta + 11/2 * C) - 6 * delta) + lambda^2 * L + 4 * delta + Kmult * (1 - exp (- K * (uM - um)))" apply (intro mono_intros I) using Laux \lambda \ 1\ \delta > 0\ \Kmult > 0\ \um \ {a..z}\ \uM \ {z..b}\ \K > 0\ by auto finally show ?thesis by (simp add: algebra_simps) text \End of the easy case 2.1\ next text \Case 2.2: $dm$ is large, i.e., all points in $f[um, ym]$ are far away from $H$. Moreover, assume that $dm \geq dM$. Then we will find a pair of points $v$ and $x$ with $um \leq v \leq x \leq ym$ satisfying the estimate~\eqref{eq:xvK}. We argue by induction: while we have not found such a pair, we can find a point $x_k$ whose projection on $V_k$, the neighborhood of size $(2^k-1) dm$ of $H$, is far enough from the projection of $um$, and such that all points in between are far enough from $V_k$ so that the corresponding projection will have good contraction properties.\ case 2 then have I: "D + 4 * C \ dm" "dM \ dm" by auto define V where "V = (\k::nat. (\g\H. cball g ((2^k - 1) * dm)))" define QC where "QC = (\k::nat. if k = 0 then 0 else 8 * delta)" have "QC k \ 0" for k unfolding QC_def using \delta > 0\ by auto have Q: "quasiconvex (0 + 8 * deltaG(TYPE('a))) (V k)" for k unfolding V_def apply (rule quasiconvex_thickening) using geodesic_segmentI[OF H] by (auto simp add: quasiconvex_of_geodesic) have "quasiconvex (QC k) (V k)" for k apply (cases "k = 0") apply (simp add: V_def QC_def quasiconvex_of_geodesic geodesic_segmentI[OF H]) apply (rule quasiconvex_mono[OF _ Q[of k]]) using \deltaG(TYPE('a)) < delta\ QC_def by auto text \Define $q(k, x)$ to be the projection of $f(x)$ on $V_k$.\ define q::"nat \ real \ 'a" where "q = (\k x. geodesic_segment_param {p x--f x} (p x) ((2^k - 1) * dm))" text \The inductive argument\ have Ind_k: "(Gromov_product_at (f z) (f um) (f uM) \ lambda^2 * (D + 3/2 * L + delta + 11/2 * C) - 2 * delta + Kmult * (1 - exp(- K * (uM - um)))) \ (\x \ {um..ym}. (\w \ {um..x}. dist (f w) (p w) \ (2^(k+1)-1) * dm) \ dist (q k um) (q k x) \ L - 4 * delta + 7 * QC k)" for k proof (induction k) text \Base case: there is a point far enough from $q 0 um$ on $H$. This is just the point $ym$, by construction.\ case 0 have *: "\x\ {um..ym}. (\w \ {um..x}. dist (f w) (p w) \ (2^(0+1)-1) * dm) \ dist (q 0 um) (q 0 x) \ L - 4 * delta + 7 * QC 0" proof (rule bexI[of _ ym], auto simp add: V_def q_def QC_def) show "um \ ym" using \ym \ {um..z}\ by auto show "L - 4 * delta \ dist (p um) (p ym)" using ym(2) apply auto using metric_space_class.zero_le_dist[of pi_z "p um"] by linarith show "\y. um \ y \ y \ ym \ dm \ dist (f y) (p y)" using dm_def closestm proj_setD(2)[OF p] by auto qed then show ?case by blast next text \The induction. The inductive assumption claims that, either the desired inequality holds, or one can construct a point with good properties. If the desired inequality holds, there is nothing left to prove. Otherwise, we can start from this point at step $k$, say $x$, and either prove the desired inequality or construct a point with the good properties at step $k+1$.\ case Suck: (Suc k) show ?case proof (cases "Gromov_product_at (f z) (f um) (f uM) \ lambda\<^sup>2 * (D + 3/2 * L + delta + 11/2 * C) - 2 * delta + Kmult * (1 - exp (- K * (uM - um)))") case True then show ?thesis by simp next case False then obtain x where x: "x \ {um..ym}" "dist (q k um) (q k x) \ L - 4 * delta + 7 * QC k" "\w. w \ {um..x} \ dist (f w) (p w) \ (2^(k+1)-1) * dm" using Suck.IH by auto text \Some auxiliary technical inequalities to be used later on.\ have aux: "(2 ^ k - 1) * dm \ (2*2^k-1) * dm" "0 \ 2 * 2 ^ k - (1::real)" "dm \ dm * 2 ^ k" apply (auto simp add: algebra_simps) apply (metis power.simps(2) two_realpow_ge_one) using \0 \ dm\ less_eq_real_def by fastforce have "L + C = (L/D) * (D + (D/L) * C)" using \L > 0\ \D > 0\ by (simp add: algebra_simps divide_simps) also have "... \ (L/D) * (D + 4 * C)" apply (intro mono_intros) using \L > 0\ \D > 0\ \C \ 0\ \D \ 4 * L\ by (auto simp add: algebra_simps divide_simps) also have "... \ (L/D) * dm" apply (intro mono_intros) using I \L > 0\ \D > 0\ by auto finally have "L + C \ (L/D) * dm" by simp moreover have "2 * delta \ (2 * delta)/D * dm" using I \C \ 0\ \delta > 0\ \D > 0\ by (auto simp add: algebra_simps divide_simps) ultimately have aux2: "L + C + 2 * delta \ ((L + 2 * delta)/D) * dm" by (auto simp add: algebra_simps divide_simps) have aux3: "(1-alpha) * D + alpha * 2^k * dm \ dm * 2^k - C/2 - QC k" proof (cases "k = 0") case True show ?thesis using I \C \ 0\ unfolding True QC_def alpha_def by auto next case False have "C/2 + QC k + (1-alpha) * D \ 2 * (1-alpha) * dm" using I \C \ 0\ unfolding QC_def alpha_def using False Laux by auto also have "... \ 2^k * (1-alpha) * dm" apply (intro mono_intros) using False alphaaux I \D > 0\ \C \ 0\ by auto finally show ?thesis by (simp add: algebra_simps) qed text \Construct a point $w$ such that its projection on $V_k$ is close to that of $um$ and therefore far away from that of $x$. This is just the intermediate value theorem (with some care as the closest point projection is not continuous).\ have "\w \ {um..x}. (dist (q k um) (q k w) \ {(9 * delta + 4 * QC k) - 4 * delta - 2 * QC k .. 9 * delta + 4 * QC k}) \ (\v \ {um..w}. dist (q k um) (q k v) \ 9 * delta + 4 * QC k)" proof (rule quasi_convex_projection_small_gaps[where ?f = f and ?G = "V k"]) show "continuous_on {um..x} f" apply (rule continuous_on_subset[OF \continuous_on {a..b} f\]) using \um \ {a..z}\ \z \ {a..b}\ \ym \ {um..z}\ \x \ {um..ym}\ by auto show "um \ x" using \x \ {um..ym}\ by auto show "quasiconvex (QC k) (V k)" by fact show "deltaG TYPE('a) < delta" by fact show "9 * delta + 4 * QC k \ {4 * delta + 2 * QC k..dist (q k um) (q k x)}" using x(2) \delta > 0\ \QC k \ 0\ Laux by auto show "q k w \ proj_set (f w) (V k)" if "w \ {um..x}" for w unfolding V_def q_def apply (rule proj_set_thickening) using aux p x(3)[OF that] by (auto simp add: metric_space_class.dist_commute) qed then obtain w where w: "w \ {um..x}" "dist (q k um) (q k w) \ {(9 * delta + 4 * QC k) - 4 * delta - 2 * QC k .. 9 * delta + 4 * QC k}" "\v. v \ {um..w} \ dist (q k um) (q k v) \ 9 * delta + 4 * QC k" by auto text \There are now two cases to be considered: either one can find a point $v$ between $um$ and $w$ which is close enough to $H$. Then this point will satisfy~\eqref{eq:xvK}, and we will be able to prove the desired inequality. Or there is no such point, and then $w$ will have the good properties at step $k+1$\ show ?thesis proof (cases "\v \ {um..w}. dist (f v) (p v) \ (2^(k+2)-1) * dm") case True text \First subcase: there is a good point $v$ between $um$ and $w$. This is the heart of the argument: we will show that the desired inequality holds.\ then obtain v where v: "v \ {um..w}" "dist (f v) (p v) \ (2^(k+2)-1) * dm" by auto text \Auxiliary basic fact to be used later on.\ have aux4: "dm * 2 ^ k \ infdist (f r) (V k)" if "r \ {v..x}" for r proof - have *: "q k r \ proj_set (f r) (V k)" unfolding q_def V_def apply (rule proj_set_thickening) using aux p[of r] x(3)[of r] that \v \ {um..w}\ \w \ {um..x}\ by (auto simp add: metric_space_class.dist_commute) have "infdist (f r) (V k) = dist (geodesic_segment_param {p r--f r} (p r) (dist (p r) (f r))) (geodesic_segment_param {p r--f r} (p r) ((2 ^ k - 1) * dm))" using proj_setD(2)[OF *] unfolding q_def by auto also have "... = abs(dist (p r) (f r) - (2 ^ k - 1) * dm)" apply (rule geodesic_segment_param(7)[where ?y = "f r"]) using x(3)[of r] \r \ {v..x}\ \v \ {um..w}\ \w \ {um..x}\ aux by (auto simp add: metric_space_class.dist_commute) also have "... = dist (f r) (p r) - (2 ^ k - 1) * dm" using x(3)[of r] \r \ {v..x}\ \v \ {um..w}\ \w \ {um..x}\ aux by (auto simp add: metric_space_class.dist_commute) finally have "dist (f r) (p r) = infdist (f r) (V k) + (2 ^ k - 1) * dm" by simp moreover have "(2^(k+1) - 1) * dm \ dist (f r) (p r)" apply (rule x(3)) using \r \ {v..x}\ \v \ {um..w}\ \w \ {um..x}\ by auto ultimately have "(2^(k+1) - 1) * dm \ infdist (f r) (V k) + (2 ^ k - 1) * dm" by simp then show ?thesis by (auto simp add: algebra_simps) qed text \Substep 1: We can control the distance from $f(v)$ to $f(closestM)$ in terms of the distance of the distance of $f(v)$ to $H$, i.e., by $2^k dm$. The same control follows for $closestM - v$ thanks to the quasi-isometry property. Then, we massage this inequality to put it in the form we will need, as an upper bound on $(x-v) \exp(-2^k dm)$.\ have "infdist (f v) H \ (2^(k+2)-1) * dm" using v proj_setD(2)[OF p[of v]] by auto have "dist v closestM \ lambda * (infdist (f v) H + (L + C + 2 * delta) + infdist (f closestM) H)" apply (rule D) using \v \ {um..w}\ \w \ {um..x}\ \x \ {um..ym}\ \ym \ {um..z}\ \um \ {a..z}\ \z \ {a..b}\ \closestM \ {yM..uM}\ \yM \ {z..uM}\ \uM \ {z..b}\ by auto also have "... \ lambda * ((2^(k+2)-1) * dm + 1 * (L + C + 2 * delta) + dM)" apply (intro mono_intros \infdist (f v) H \ (2^(k+2)-1) * dm\) using dM_def \lambda \ 1\ \L > 0\ \C \ 0\ \delta > 0\ by (auto simp add: metric_space_class.dist_commute) also have "... \ lambda * ((2^(k+2)-1) * dm + 2^k * (((L + 2 * delta)/D) * dm) + dm)" apply (intro mono_intros) using I \lambda \ 1\ \C \ 0\ \delta > 0\ \L > 0\ aux2 by auto also have "... = lambda * 2^k * (4 + (L + 2 * delta)/D) * dm" by (simp add: algebra_simps) finally have *: "dist v closestM / (lambda * (4 + (L + 2 * delta)/D)) \ 2^k * dm" using \lambda \ 1\ \L > 0\ \D > 0\ \delta > 0\ by (simp add: divide_simps, simp add: algebra_simps) text \We reformulate this control inside of an exponential, as this is the form we will use later on.\ have "exp(- (alpha * (2^k * dm) * ln 2 / (5 * delta))) \ exp(-(alpha * (dist v closestM / (lambda * (4 + (L + 2 * delta)/D))) * ln 2 / (5 * delta)))" apply (intro mono_intros *) using alphaaux \delta > 0\ by auto also have "... = exp(-K * dist v closestM)" unfolding K_def by (simp add: divide_simps) also have "... = exp(-K * (closestM - v))" unfolding dist_real_def using \v \ {um..w}\ \w \ {um..x}\ \x \ {um..ym}\ \ym \ {um..z}\ \yM \ {z..uM}\ \closestM \ {yM..uM}\ \K > 0\ by auto finally have "exp(- (alpha * (2^k * dm) * ln 2 / (5 * delta))) \ exp(-K * (closestM - v))" by simp text \Plug in $x-v$ to get the final form of this inequality.\ then have "K * (x - v) * exp(- (alpha * (2^k * dm) * ln 2 / (5 * delta))) \ K * (x - v) * exp(-K * (closestM - v))" apply (rule mult_left_mono) using \delta > 0\ \lambda \ 1\ \v \ {um..w}\ \w \ {um..x}\ \K > 0\ by auto also have "... = ((1 + K * (x - v)) - 1) * exp(- K * (closestM - v))" by (auto simp add: algebra_simps) also have "... \ (exp (K * (x - v)) - 1) * exp(-K * (closestM - v))" by (intro mono_intros, auto) also have "... = exp(-K * (closestM - x)) - exp(-K * (closestM - v))" by (simp add: algebra_simps mult_exp_exp) also have "... \ exp(-K * (closestM - x)) - exp(-K * (uM - um))" using \K > 0\ \v \ {um..w}\ \w \ {um..x}\ \x \ {um..ym}\ \ym \ {um..z}\ \yM \ {z..uM}\ \closestM \ {yM..uM}\ by auto finally have B: "(x - v) * exp(- alpha * 2^k * dm * ln 2 / (5 * delta)) \ (exp(-K * (closestM - x)) - exp(-K * (uM-um)))/K" using \K > 0\ by (auto simp add: divide_simps algebra_simps) text \End of substep 1\ text \Substep 2: The projections of $f(v)$ and $f(x)$ on the cylinder $V_k$ are well separated, by construction. This implies that $v$ and $x$ themselves are well separated, thanks to the exponential contraction property of the projection on the quasi-convex set $V_k$. This leads to a uniform lower bound for $(x-v) \exp(-2^k dm)$, which has been upper bounded in Substep 1.\ have "L - 4 * delta + 7 * QC k \ dist (q k um) (q k x)" using x by simp also have "... \ dist (q k um) (q k v) + dist (q k v) (q k x)" by (intro mono_intros) also have "... \ (9 * delta + 4 * QC k) + dist (q k v) (q k x)" using w(3)[of v] \v \ {um..w}\ by auto finally have "L - 13 * delta + 3 * QC k \ dist (q k v) (q k x)" by simp also have "... \ 3 * QC k + max (5 * deltaG(TYPE('a))) ((4 * exp(1/2 * ln 2)) * lambda * (x - v) * exp(-(dm * 2^k - C/2 - QC k) * ln 2 / (5 * delta)))" proof (cases "k = 0") text \We use different statements for the projection in the case $k = 0$ (projection on a geodesic) and $k > 0$ (projection on a quasi-convex set) as the bounds are better in the first case, which is the most important one for the final value of the constant.\ case True have "dist (q k v) (q k x) \ max (5 * deltaG(TYPE('a))) ((4 * exp(1/2 * ln 2)) * lambda * (x - v) * exp(-(dm * 2^k - C/2) * ln 2 / (5 * delta)))" proof (rule geodesic_projection_exp_contracting[where ?G = "V k" and ?f = f]) show "geodesic_segment (V k)" unfolding True V_def using geodesic_segmentI[OF H] by auto show "v \ x" using \v \ {um..w}\ \w \ {um..x}\ by auto show "q k v \ proj_set (f v) (V k)" unfolding q_def V_def apply (rule proj_set_thickening) using aux p[of v] x(3)[of v] \v \ {um..w}\ \w \ {um..x}\ by (auto simp add: metric_space_class.dist_commute) show "q k x \ proj_set (f x) (V k)" unfolding q_def V_def apply (rule proj_set_thickening) using aux p[of x] x(3)[of x] \w \ {um..x}\ by (auto simp add: metric_space_class.dist_commute) show "15/2 * delta + C/2 \ dm * 2^k" apply (rule order_trans[of _ dm]) using I \delta > 0\ \C \ 0\ Laux unfolding QC_def by auto show "deltaG TYPE('a) < delta" by fact show "\t. t \ {v..x} \ dm * 2 ^ k \ infdist (f t) (V k)" using aux4 by auto show "0 \ C" "0 \ lambda" using \C \ 0\ \lambda \ 1\ by auto show "dist (f x1) (f x2) \ lambda * dist x1 x2 + C" if "x1 \ {v..x}" "x2 \ {v..x}" for x1 x2 using quasi_isometry_onD(1)[OF assms(2)] that \v \ {um..w}\ \w \ {um..x}\ \x \ {um..ym}\ \ym \ {um..z}\ \um \ {a..z}\ \z \ {a..b}\ by auto qed then show ?thesis unfolding QC_def True by auto next case False have "dist (q k v) (q k x) \ 2 * QC k + 8 * delta + max (5 * deltaG(TYPE('a))) ((4 * exp(1/2 * ln 2)) * lambda * (x - v) * exp(-(dm * 2^k - QC k -C/2) * ln 2 / (5 * delta)))" proof (rule quasiconvex_projection_exp_contracting[where ?G = "V k" and ?f = f]) show "quasiconvex (QC k) (V k)" by fact show "v \ x" using \v \ {um..w}\ \w \ {um..x}\ by auto show "q k v \ proj_set (f v) (V k)" unfolding q_def V_def apply (rule proj_set_thickening) using aux p[of v] x(3)[of v] \v \ {um..w}\ \w \ {um..x}\ by (auto simp add: metric_space_class.dist_commute) show "q k x \ proj_set (f x) (V k)" unfolding q_def V_def apply (rule proj_set_thickening) using aux p[of x] x(3)[of x] \w \ {um..x}\ by (auto simp add: metric_space_class.dist_commute) show "15/2 * delta + QC k + C/2 \ dm * 2^k" apply (rule order_trans[of _ dm]) using I \delta > 0\ \C \ 0\ Laux unfolding QC_def by auto show "deltaG TYPE('a) < delta" by fact show "\t. t \ {v..x} \ dm * 2 ^ k \ infdist (f t) (V k)" using aux4 by auto show "0 \ C" "0 \ lambda" using \C \ 0\ \lambda \ 1\ by auto show "dist (f x1) (f x2) \ lambda * dist x1 x2 + C" if "x1 \ {v..x}" "x2 \ {v..x}" for x1 x2 using quasi_isometry_onD(1)[OF assms(2)] that \v \ {um..w}\ \w \ {um..x}\ \x \ {um..ym}\ \ym \ {um..z}\ \um \ {a..z}\ \z \ {a..b}\ by auto qed then show ?thesis unfolding QC_def using False by (auto simp add: algebra_simps) qed finally have "L - 13 * delta \ max (5 * deltaG(TYPE('a))) ((4 * exp(1/2 * ln 2)) * lambda * (x - v) * exp(-(dm * 2^k - C/2 - QC k) * ln 2 / (5 * delta)))" by auto then have "L - 13 * delta \ (4 * exp(1/2 * ln 2)) * lambda * (x - v) * exp(-(dm * 2^k - C/2 - QC k) * ln 2 / (5 * delta))" using \delta > deltaG(TYPE('a))\ Laux by auto text \We separate the exponential gain coming from the contraction into two parts, one to be spent to improve the constant, and one for the inductive argument.\ also have "... \ (4 * exp(1/2 * ln 2)) * lambda * (x - v) * exp(-((1-alpha) * D + alpha * 2^k * dm) * ln 2 / (5 * delta))" apply (intro mono_intros) using aux3 \delta > 0\ \lambda \ 1\ \v \ {um..w}\ \w \ {um..x}\ by auto also have "... = (4 * exp(1/2 * ln 2)) * lambda * (x - v) * (exp(-(1-alpha) * D * ln 2 / (5 * delta)) * exp(-alpha * 2^k * dm * ln 2 / (5 * delta)))" unfolding mult_exp_exp by (auto simp add: algebra_simps divide_simps) finally have A: "L - 13 * delta \ (4 * exp(1/2 * ln 2)) * lambda * exp(-(1-alpha) * D * ln 2 / (5 * delta)) * ((x - v) * exp(-alpha * 2^k * dm * ln 2 / (5 * delta)))" by (simp add: algebra_simps) text \This is the end of the second substep.\ text \Use the second substep to show that $x-v$ is bounded below, and therefore that $closestM - x$ (the endpoints of the new geodesic we want to consider in the inductive argument) are quantitatively closer than $uM - um$, which means that we will be able to use the inductive assumption over this new geodesic.\ also have "... \ (4 * exp(1/2 * ln 2)) * lambda * exp 0 * ((x - v) * exp 0)" apply (intro mono_intros) using \delta > 0\ \lambda \ 1\ \v \ {um..w}\ \w \ {um..x}\ alphaaux \D > 0\ \C \ 0\ I by (auto simp add: divide_simps mult_nonpos_nonneg) also have "... = (4 * exp(1/2 * ln 2)) * lambda * (x-v)" by simp also have "... \ 20 * lambda * (x - v)" apply (intro mono_intros, approximation 10) using \delta > 0\ \lambda \ 1\ \v \ {um..w}\ \w \ {um..x}\ by auto finally have "x - v \ (1/4) * delta / lambda" using \lambda \ 1\ L_def \delta > 0\ by (simp add: divide_simps algebra_simps) then have "closestM - x + (1/4) * delta / lambda \ closestM - v" by simp also have "... \ uM - um" using \closestM \ {yM..uM}\ \v \ {um..w}\ by auto also have "... \ Suc n * (1/4) * delta / lambda" by fact finally have "closestM - x \ n * (1/4) * delta / lambda" unfolding Suc_eq_plus1 by (auto simp add: algebra_simps add_divide_distrib) text \Conclusion of the proof: combine the lower bound of the second substep with the upper bound of the first substep to get a definite gain when one goes from the old geodesic to the new one. Then, apply the inductive assumption to the new one to conclude the desired inequality for the old one.\ have "L + 4 * delta = ((L + 4 * delta)/(L - 13 * delta)) * (L - 13 * delta)" using Laux \delta > 0\ by (simp add: algebra_simps divide_simps) also have "... \ ((L + 4 * delta)/(L - 13 * delta)) * ((4 * exp(1/2 * ln 2)) * lambda * exp (- (1 - alpha) * D * ln 2 / (5 * delta)) * ((x - v) * exp (- alpha * 2 ^ k * dm * ln 2 / (5 * delta))))" apply (rule mult_left_mono) using A Laux \delta > 0\ by (auto simp add: divide_simps) also have "... \ ((L + 4 * delta)/(L - 13 * delta)) * ((4 * exp(1/2 * ln 2)) * lambda * exp (- (1 - alpha) * D * ln 2 / (5 * delta)) * ((exp(-K * (closestM - x)) - exp(-K * (uM - um)))/K))" apply (intro mono_intros B) using Laux \delta > 0\ \lambda \ 1\ by (auto simp add: divide_simps) finally have C: "L + 4 * delta \ Kmult * (exp(-K * (closestM - x)) - exp(-K * (uM - um)))" unfolding Kmult_def by auto have "Gromov_product_at (f z) (f um) (f uM) \ Gromov_product_at (f z) (f x) (f closestM) + (L + 4 * delta)" apply (rule Rec) using \closestM \ {yM..uM}\ \x \ {um..ym}\ \ym \ {um..z}\ by auto also have "... \ (lambda^2 * (D + 3/2 * L + delta + 11/2 * C) - 2 * delta + Kmult * (1 - exp(- K * (closestM - x)))) + (Kmult * (exp(-K * (closestM - x)) - exp(-K * (uM-um))))" apply (intro mono_intros C Suc.IH) using \x \ {um..ym}\ \ym \ {um..z}\ \um \ {a..z}\ \closestM \ {yM..uM}\ \yM \ {z..uM}\ \uM \ {z..b}\ \closestM - x \ n * (1/4) * delta / lambda\ by auto also have "... = (lambda^2 * (D + 3/2 * L + delta + 11/2 * C) - 2 * delta + Kmult * (1 - exp(- K * (uM - um))))" unfolding K_def by (simp add: algebra_simps) finally show ?thesis by auto text \End of the first subcase, when there is a good point $v$ between $um$ and $w$.\ next case False text \Second subcase: between $um$ and $w$, all points are far away from $V_k$. We will show that this implies that $w$ is admissible for the step $k+1$.\ have "\w\{um..ym}. (\v\{um..w}. (2 ^ (Suc k + 1) - 1) * dm \ dist (f v) (p v)) \ L - 4 * delta + 7 * QC (Suc k) \ dist (q (Suc k) um) (q (Suc k) w)" proof (rule bexI[of _ w], auto) show "um \ w" "w \ ym" using \w \ {um..x}\ \x \ {um..ym}\ by auto show "(4 * 2 ^ k - 1) * dm \ dist (f x) (p x)" if "um \ x" "x \ w" for x using False \dm \ 0\ that by force have "dist (q k um) (q (k+1) um) = 2^k * dm" unfolding q_def apply (subst geodesic_segment_param(7)[where ?y = "f um"]) using x(3)[of um] \x \ {um..ym}\ aux by (auto simp add: metric_space_class.dist_commute, simp add: algebra_simps) have "dist (q k w) (q (k+1) w) = 2^k * dm" unfolding q_def apply (subst geodesic_segment_param(7)[where ?y = "f w"]) using x(3)[of w] \w \ {um..x}\ \x \ {um..ym}\ aux by (auto simp add: metric_space_class.dist_commute, simp add: algebra_simps) have i: "q k um \ proj_set (q (k+1) um) (V k)" unfolding q_def V_def apply (rule proj_set_thickening'[of _ "f um"]) using p x(3)[of um] \x \ {um..ym}\ aux by (auto simp add: algebra_simps metric_space_class.dist_commute) have j: "q k w \ proj_set (q (k+1) w) (V k)" unfolding q_def V_def apply (rule proj_set_thickening'[of _ "f w"]) using p x(3)[of w] \x \ {um..ym}\ \w \ {um..x}\ aux by (auto simp add: algebra_simps metric_space_class.dist_commute) have "5 * delta + 2 * QC k \ dist (q k um) (q k w)" using w(2) by simp also have "... \ max (5 * deltaG(TYPE('a)) + 2 * QC k) (dist (q (k + 1) um) (q (k + 1) w) - dist (q k um) (q (k + 1) um) - dist (q k w) (q (k + 1) w) + 10 * deltaG(TYPE('a)) + 4 * QC k)" by (rule proj_along_quasiconvex_contraction[OF \quasiconvex (QC k) (V k)\ i j]) finally have "5 * delta + 2 * QC k \ dist (q (k + 1) um) (q (k + 1) w) - dist (q k um) (q (k + 1) um) - dist (q k w) (q (k + 1) w) + 10 * deltaG(TYPE('a)) + 4 * QC k" using \deltaG(TYPE('a)) < delta\ by auto then have "0 \ dist (q (k + 1) um) (q (k + 1) w) + 5 * delta + 2 * QC k - dist (q k um) (q (k + 1) um) - dist (q k w) (q (k + 1) w)" using \deltaG(TYPE('a)) < delta\ by auto also have "... = dist (q (k + 1) um) (q (k + 1) w) + 5 * delta + 2 * QC k - 2^(k+1) * dm" by (simp only: \dist (q k w) (q (k+1) w) = 2^k * dm\ \dist (q k um) (q (k+1) um) = 2^k * dm\, auto) finally have *: "2^(k+1) * dm - 5 * delta - 2 * QC k \ dist (q (k+1) um) (q (k+1) w)" using \deltaG(TYPE('a)) < delta\ by auto have "L - 4 * delta + 7 * QC (k+1) \ 2 * dm - 5 * delta - 2 * QC k" unfolding QC_def L_def using \delta > 0\ Laux I \C \ 0\ by auto also have "... \ 2^(k+1) * dm - 5 * delta - 2 * QC k" using aux by (auto simp add: algebra_simps) finally show "L - 4 * delta + 7 * QC (Suc k) \ dist (q (Suc k) um) (q (Suc k) w)" using * by auto qed then show ?thesis by simp qed qed qed text \This is the end of the main induction over $k$. To conclude, choose $k$ large enough so that the second alternative in this induction is impossible. It follows that the first alternative holds, i.e., the desired inequality is true.\ have "dm > 0" using I \delta > 0\ \C \ 0\ Laux by auto have "\k. 2^k > dist (f um) (p um)/dm + 1" by (simp add: real_arch_pow) then obtain k where "2^k > dist (f um) (p um)/dm + 1" by blast then have "dist (f um) (p um) < (2^k - 1) * dm" using \dm > 0\ by (auto simp add: divide_simps algebra_simps) also have "... \ (2^(Suc k) - 1) * dm" by (intro mono_intros, auto) finally have "\((2 ^ (k + 1) - 1) * dm \ dist (f um) (p um))" by simp then show "Gromov_product_at (f z) (f um) (f uM) \ lambda\<^sup>2 * (D + 3/2 * L + delta + 11/2 * C) - 2 * delta + Kmult * (1 - exp (- K * (uM - um)))" using Ind_k[of k] by auto text \end of the case where $D + 4 * C \leq dm$ and $dM \leq dm$.\ next case 3 text \This is the exact copy of the previous case, except that the roles of the points before and after $z$ are exchanged. In a perfect world, one would use a lemma subsuming both cases, but in practice copy-paste seems to work better here as there are two many details to be changed regarding the direction of inequalities.\ then have I: "D + 4 * C \ dM" "dm \ dM" by auto define V where "V = (\k::nat. (\g\H. cball g ((2^k - 1) * dM)))" define QC where "QC = (\k::nat. if k = 0 then 0 else 8 * delta)" have "QC k \ 0" for k unfolding QC_def using \delta > 0\ by auto have Q: "quasiconvex (0 + 8 * deltaG(TYPE('a))) (V k)" for k unfolding V_def apply (rule quasiconvex_thickening) using geodesic_segmentI[OF H] by (auto simp add: quasiconvex_of_geodesic) have "quasiconvex (QC k) (V k)" for k apply (cases "k = 0") apply (simp add: V_def QC_def quasiconvex_of_geodesic geodesic_segmentI[OF H]) apply (rule quasiconvex_mono[OF _ Q[of k]]) using \deltaG(TYPE('a)) < delta\ QC_def by auto define q::"nat \ real \ 'a" where "q = (\k x. geodesic_segment_param {p x--f x} (p x) ((2^k - 1) * dM))" have Ind_k: "(Gromov_product_at (f z) (f um) (f uM) \ lambda^2 * (D + 3/2 * L + delta + 11/2 * C) - 2 * delta + Kmult * (1 - exp(- K * (uM - um)))) \ (\x \ {yM..uM}. (\y \ {x..uM}. dist (f y) (p y) \ (2^(k+1)-1) * dM) \ dist (q k uM) (q k x) \ L - 4 * delta + 7 * QC k)" for k proof (induction k) case 0 have *: "\x\ {yM..uM}. (\y \ {x..uM}. dist (f y) (p y) \ (2^(0+1)-1) * dM) \ dist (q 0 uM) (q 0 x) \ L - 4 * delta + 7 * QC 0" proof (rule bexI[of _ yM], auto simp add: V_def q_def QC_def) show "yM \ uM" using \yM \ {z..uM}\ by auto show "L -4 * delta \ dist (p uM) (p yM)" using yM(2) apply auto using metric_space_class.zero_le_dist[of pi_z "p uM"] by linarith show "\y. y \ uM \ yM \ y \ dM \ dist (f y) (p y)" using dM_def closestM proj_setD(2)[OF p] by auto qed then show ?case by blast next case Suck: (Suc k) show ?case proof (cases "Gromov_product_at (f z) (f um) (f uM) \ lambda\<^sup>2 * (D + 3/2 * L + delta + 11/2 * C) - 2 * delta + Kmult * (1 - exp (- K * (uM - um)))") case True then show ?thesis by simp next case False then obtain x where x: "x \ {yM..uM}" "dist (q k uM) (q k x) \ L - 4 * delta + 7 * QC k" "\w. w \ {x..uM} \ dist (f w) (p w) \ (2^(k+1)-1) * dM" using Suck.IH by auto have aux: "(2 ^ k - 1) * dM \ (2*2^k-1) * dM" "0 \ 2 * 2 ^ k - (1::real)" "dM \ dM * 2 ^ k" apply (auto simp add: algebra_simps) apply (metis power.simps(2) two_realpow_ge_one) using \0 \ dM\ less_eq_real_def by fastforce have "L + C = (L/D) * (D + (D/L) * C)" using \L > 0\ \D > 0\ by (simp add: algebra_simps divide_simps) also have "... \ (L/D) * (D + 4 * C)" apply (intro mono_intros) using \L > 0\ \D > 0\ \C \ 0\ \D \ 4 * L\ by (auto simp add: algebra_simps divide_simps) also have "... \ (L/D) * dM" apply (intro mono_intros) using I \L > 0\ \D > 0\ by auto finally have "L + C \ (L/D) * dM" by simp moreover have "2 * delta \ (2 * delta)/D * dM" using I \C \ 0\ \delta > 0\ \D > 0\ by (auto simp add: algebra_simps divide_simps) ultimately have aux2: "L + C + 2 * delta \ ((L + 2 * delta)/D) * dM" by (auto simp add: algebra_simps divide_simps) have aux3: "(1-alpha) * D + alpha * 2^k * dM \ dM * 2^k - C/2 - QC k" proof (cases "k = 0") case True show ?thesis using I \C \ 0\ unfolding True QC_def alpha_def by auto next case False have "C/2 + QC k + (1-alpha) * D \ 2 * (1-alpha) * dM" using I \C \ 0\ unfolding QC_def alpha_def using False Laux by auto also have "... \ 2^k * (1-alpha) * dM" apply (intro mono_intros) using False alphaaux I \D > 0\ \C \ 0\ by auto finally show ?thesis by (simp add: algebra_simps) qed have "\w \ {x..uM}. (dist (q k uM) (q k w) \ {(9 * delta + 4 * QC k) - 4 * delta - 2 * QC k .. 9 * delta + 4 * QC k}) \ (\v \ {w..uM}. dist (q k uM) (q k v) \ 9 * delta + 4 * QC k)" proof (rule quasi_convex_projection_small_gaps'[where ?f = f and ?G = "V k"]) show "continuous_on {x..uM} f" apply (rule continuous_on_subset[OF \continuous_on {a..b} f\]) using \uM \ {z..b}\ \z \ {a..b}\ \yM \ {z..uM}\ \x \ {yM..uM}\ by auto show "x \ uM" using \x \ {yM..uM}\ by auto show "quasiconvex (QC k) (V k)" by fact show "deltaG TYPE('a) < delta" by fact show "9 * delta + 4 * QC k \ {4 * delta + 2 * QC k..dist (q k x) (q k uM)}" using x(2) \delta > 0\ \QC k \ 0\ Laux by (auto simp add: metric_space_class.dist_commute) show "q k w \ proj_set (f w) (V k)" if "w \ {x..uM}" for w unfolding V_def q_def apply (rule proj_set_thickening) using aux p x(3)[OF that] by (auto simp add: metric_space_class.dist_commute) qed then obtain w where w: "w \ {x..uM}" "dist (q k uM) (q k w) \ {(9 * delta + 4 * QC k) - 4 * delta - 2 * QC k .. 9 * delta + 4 * QC k}" "\v. v \ {w..uM} \ dist (q k uM) (q k v) \ 9 * delta + 4 * QC k" by auto show ?thesis proof (cases "\v \ {w..uM}. dist (f v) (p v) \ (2^(k+2)-1) * dM") case True then obtain v where v: "v \ {w..uM}" "dist (f v) (p v) \ (2^(k+2)-1) * dM" by auto have aux4: "dM * 2 ^ k \ infdist (f r) (V k)" if "r \ {x..v}" for r proof - have *: "q k r \ proj_set (f r) (V k)" unfolding q_def V_def apply (rule proj_set_thickening) using aux p[of r] x(3)[of r] that \v \ {w..uM}\ \w \ {x..uM}\ by (auto simp add: metric_space_class.dist_commute) have "infdist (f r) (V k) = dist (geodesic_segment_param {p r--f r} (p r) (dist (p r) (f r))) (geodesic_segment_param {p r--f r} (p r) ((2 ^ k - 1) * dM))" using proj_setD(2)[OF *] unfolding q_def by auto also have "... = abs(dist (p r) (f r) - (2 ^ k - 1) * dM)" apply (rule geodesic_segment_param(7)[where ?y = "f r"]) using x(3)[of r] \r \ {x..v}\ \v \ {w..uM}\ \w \ {x..uM}\ aux by (auto simp add: metric_space_class.dist_commute) also have "... = dist (f r) (p r) - (2 ^ k - 1) * dM" using x(3)[of r] \r \ {x..v}\ \v \ {w..uM}\ \w \ {x..uM}\ aux by (auto simp add: metric_space_class.dist_commute) finally have "dist (f r) (p r) = infdist (f r) (V k) + (2 ^ k - 1) * dM" by simp moreover have "(2^(k+1) - 1) * dM \ dist (f r) (p r)" apply (rule x(3)) using \r \ {x..v}\ \v \ {w..uM}\ \w \ {x..uM}\ by auto ultimately have "(2^(k+1) - 1) * dM \ infdist (f r) (V k) + (2 ^ k - 1) * dM" by simp then show ?thesis by (auto simp add: algebra_simps) qed have "infdist (f v) H \ (2^(k+2)-1) * dM" using v proj_setD(2)[OF p[of v]] by auto have "dist closestm v \ lambda * (infdist (f closestm) H + (L + C + 2 * delta) + infdist (f v) H)" apply (rule D) using \v \ {w..uM}\ \w \ {x..uM}\ \x \ {yM..uM}\ \yM \ {z..uM}\ \uM \ {z..b}\ \z \ {a..b}\ \closestm \ {um..ym}\ \ym \ {um..z}\ \um \ {a..z}\ by auto also have "... \ lambda * (dm + 1 * (L + C + 2 * delta) + (2^(k+2)-1) * dM)" apply (intro mono_intros \infdist (f v) H \ (2^(k+2)-1) * dM\) using dm_def \lambda \ 1\ \L > 0\ \C \ 0\ \delta > 0\ by (auto simp add: metric_space_class.dist_commute) also have "... \ lambda * (dM + 2^k * (((L + 2 * delta)/D) * dM) + (2^(k+2)-1) * dM)" apply (intro mono_intros) using I \lambda \ 1\ \C \ 0\ \delta > 0\ \L > 0\ aux2 by auto also have "... = lambda * 2^k * (4 + (L + 2 * delta)/D) * dM" by (simp add: algebra_simps) finally have *: "dist closestm v / (lambda * (4 + (L + 2 * delta)/D)) \ 2^k * dM" using \lambda \ 1\ \L > 0\ \D > 0\ \delta > 0\ by (simp add: divide_simps, simp add: algebra_simps) have "exp(- (alpha * (2^k * dM) * ln 2 / (5 * delta))) \ exp(-(alpha * (dist closestm v / (lambda * (4 + (L + 2 * delta)/D))) * ln 2 / (5 * delta)))" apply (intro mono_intros *) using alphaaux \delta > 0\ by auto also have "... = exp(-K * dist closestm v)" unfolding K_def by (simp add: divide_simps) also have "... = exp(-K * (v - closestm))" unfolding dist_real_def using \v \ {w..uM}\ \w \ {x..uM}\ \x \ {yM..uM}\ \yM \ {z..uM}\ \ym \ {um..z}\ \closestm \ {um..ym}\ \K > 0\ by auto finally have "exp(- (alpha * (2^k * dM) * ln 2 / (5 * delta))) \ exp(-K * (v - closestm))" by simp then have "K * (v - x) * exp(- (alpha * (2^k * dM) * ln 2 / (5 * delta))) \ K * (v - x) * exp(-K * (v - closestm))" apply (rule mult_left_mono) using \delta > 0\ \lambda \ 1\ \v \ {w..uM}\ \w \ {x..uM}\ \K > 0\ by auto also have "... = ((1 + K * (v - x)) - 1) * exp(- K * (v - closestm))" by (auto simp add: algebra_simps) also have "... \ (exp (K * (v - x)) - 1) * exp(-K * (v - closestm))" by (intro mono_intros, auto) also have "... = exp(-K * (x - closestm)) - exp(-K * (v - closestm))" by (simp add: algebra_simps mult_exp_exp) also have "... \ exp(-K * (x - closestm)) - exp(-K * (uM - um))" using \K > 0\ \v \ {w..uM}\ \w \ {x..uM}\ \x \ {yM..uM}\ \yM \ {z..uM}\ \ym \ {um..z}\ \closestm \ {um..ym}\ by auto finally have B: "(v - x) * exp(- alpha * 2^k * dM * ln 2 / (5 * delta)) \ (exp(-K * (x - closestm)) - exp(-K * (uM - um)))/K" using \K > 0\ by (auto simp add: divide_simps algebra_simps) text \The projections of $f(v)$ and $f(x)$ on the cylinder $V_k$ are well separated, by construction. This implies that $v$ and $x$ themselves are well separated.\ have "L - 4 * delta + 7 * QC k \ dist (q k uM) (q k x)" using x by simp also have "... \ dist (q k uM) (q k v) + dist (q k v) (q k x)" by (intro mono_intros) also have "... \ (9 * delta + 4 * QC k) + dist (q k v) (q k x)" using w(3)[of v] \v \ {w..uM}\ by auto finally have "L - 13 * delta + 3 * QC k \ dist (q k x) (q k v)" by (simp add: metric_space_class.dist_commute) also have "... \ 3 * QC k + max (5 * deltaG(TYPE('a))) ((4 * exp(1/2 * ln 2)) * lambda * (v - x) * exp(-(dM * 2^k - C/2 - QC k) * ln 2 / (5 * delta)))" proof (cases "k = 0") case True have "dist (q k x) (q k v) \ max (5 * deltaG(TYPE('a))) ((4 * exp(1/2 * ln 2)) * lambda * (v - x) * exp(-(dM * 2^k - C/2) * ln 2 / (5 * delta)))" proof (rule geodesic_projection_exp_contracting[where ?G = "V k" and ?f = f]) show "geodesic_segment (V k)" unfolding V_def True using geodesic_segmentI[OF H] by auto show "x \ v" using \v \ {w..uM}\ \w \ {x..uM}\ by auto show "q k v \ proj_set (f v) (V k)" unfolding q_def V_def apply (rule proj_set_thickening) using aux p[of v] x(3)[of v] \v \ {w..uM}\ \w \ {x..uM}\ by (auto simp add: metric_space_class.dist_commute) show "q k x \ proj_set (f x) (V k)" unfolding q_def V_def apply (rule proj_set_thickening) using aux p[of x] x(3)[of x] \w \ {x..uM}\ by (auto simp add: metric_space_class.dist_commute) show "15/2 * delta + C/2 \ dM * 2^k" using I \delta > 0\ \C \ 0\ Laux unfolding QC_def True by auto show "deltaG TYPE('a) < delta" by fact show "\t. t \ {x..v} \ dM * 2 ^ k \ infdist (f t) (V k)" using aux4 by auto show "0 \ C" "0 \ lambda" using \C \ 0\ \lambda \ 1\ by auto show "dist (f x1) (f x2) \ lambda * dist x1 x2 + C" if "x1 \ {x..v}" "x2 \ {x..v}" for x1 x2 using quasi_isometry_onD(1)[OF assms(2)] that \v \ {w..uM}\ \w \ {x..uM}\ \x \ {yM..uM}\ \yM \ {z..uM}\ \uM \ {z..b}\ \z \ {a..b}\ by auto qed then show ?thesis unfolding QC_def True by auto next case False have "dist (q k x) (q k v) \ 2 * QC k + 8 * delta + max (5 * deltaG(TYPE('a))) ((4 * exp(1/2 * ln 2)) * lambda * (v - x) * exp(-(dM * 2^k - QC k - C/2) * ln 2 / (5 * delta)))" proof (rule quasiconvex_projection_exp_contracting[where ?G = "V k" and ?f = f]) show "quasiconvex (QC k) (V k)" by fact show "x \ v" using \v \ {w..uM}\ \w \ {x..uM}\ by auto show "q k v \ proj_set (f v) (V k)" unfolding q_def V_def apply (rule proj_set_thickening) using aux p[of v] x(3)[of v] \v \ {w..uM}\ \w \ {x..uM}\ by (auto simp add: metric_space_class.dist_commute) show "q k x \ proj_set (f x) (V k)" unfolding q_def V_def apply (rule proj_set_thickening) using aux p[of x] x(3)[of x] \w \ {x..uM}\ by (auto simp add: metric_space_class.dist_commute) show "15/2 * delta + QC k + C/2 \ dM * 2^k" apply (rule order_trans[of _ dM]) using I \delta > 0\ \C \ 0\ Laux unfolding QC_def by auto show "deltaG TYPE('a) < delta" by fact show "\t. t \ {x..v} \ dM * 2 ^ k \ infdist (f t) (V k)" using aux4 by auto show "0 \ C" "0 \ lambda" using \C \ 0\ \lambda \ 1\ by auto show "dist (f x1) (f x2) \ lambda * dist x1 x2 + C" if "x1 \ {x..v}" "x2 \ {x..v}" for x1 x2 using quasi_isometry_onD(1)[OF assms(2)] that \v \ {w..uM}\ \w \ {x..uM}\ \x \ {yM..uM}\ \yM \ {z..uM}\ \uM \ {z..b}\ \z \ {a..b}\ by auto qed then show ?thesis unfolding QC_def using False by (auto simp add: algebra_simps) qed finally have "L - 13 * delta \ max (5 * deltaG(TYPE('a))) ((4 * exp(1/2 * ln 2)) * lambda * (v - x) * exp(-(dM * 2^k - C/2 - QC k) * ln 2 / (5 * delta)))" by auto then have "L - 13 * delta \ (4 * exp(1/2 * ln 2)) * lambda * (v - x) * exp(-(dM * 2^k - C/2 - QC k) * ln 2 / (5 * delta))" using \delta > deltaG(TYPE('a))\ Laux by auto also have "... \ (4 * exp(1/2 * ln 2)) * lambda * (v - x) * exp(-((1-alpha) * D + alpha * 2^k * dM) * ln 2 / (5 * delta))" apply (intro mono_intros) using aux3 \delta > 0\ \lambda \ 1\ \v \ {w..uM}\ \w \ {x..uM}\ by auto also have "... = (4 * exp(1/2 * ln 2)) * lambda * (v - x) * (exp(-(1-alpha) * D * ln 2 / (5 * delta)) * exp(-alpha * 2^k * dM * ln 2 / (5 * delta)))" unfolding mult_exp_exp by (auto simp add: algebra_simps divide_simps) finally have A: "L - 13 * delta \ (4 * exp(1/2 * ln 2)) * lambda * exp(-(1-alpha) * D * ln 2 / (5 * delta)) * ((v - x) * exp(-alpha * 2^k * dM * ln 2 / (5 * delta)))" by (simp add: algebra_simps) also have "... \ (4 * exp(1/2 * ln 2)) * lambda * exp 0 * ((v - x) * exp 0)" apply (intro mono_intros) using \delta > 0\ \lambda \ 1\ \v \ {w..uM}\ \w \ {x..uM}\ alphaaux \D > 0\ \C \ 0\ I by (auto simp add: divide_simps mult_nonpos_nonneg) also have "... = (4 * exp(1/2 * ln 2)) * lambda * (v - x)" by simp also have "... \ 20 * lambda * (v - x)" apply (intro mono_intros, approximation 10) using \delta > 0\ \lambda \ 1\ \v \ {w..uM}\ \w \ {x..uM}\ by auto finally have "v - x \ (1/4) * delta / lambda" using \lambda \ 1\ L_def \delta > 0\ by (simp add: divide_simps algebra_simps) then have "x - closestm + (1/4) * delta / lambda \ v - closestm" by simp also have "... \ uM - um" using \closestm \ {um..ym}\ \v \ {w..uM}\ by auto also have "... \ Suc n * (1/4) * delta / lambda" by fact finally have "x - closestm \ n * (1/4) * delta / lambda" unfolding Suc_eq_plus1 by (auto simp add: algebra_simps add_divide_distrib) have "L + 4 * delta = ((L + 4 * delta)/(L - 13 * delta)) * (L - 13 * delta)" using Laux \delta > 0\ by (simp add: algebra_simps divide_simps) also have "... \ ((L + 4 * delta)/(L - 13 * delta)) * ((4 * exp(1/2 * ln 2)) * lambda * exp (- (1 - alpha) * D * ln 2 / (5 * delta)) * ((v - x) * exp (- alpha * 2 ^ k * dM * ln 2 / (5 * delta))))" apply (rule mult_left_mono) using A Laux \delta > 0\ by (auto simp add: divide_simps) also have "... \ ((L + 4 * delta)/(L - 13 * delta)) * ((4 * exp(1/2 * ln 2)) * lambda * exp (- (1 - alpha) * D * ln 2 / (5 * delta)) * ((exp(-K * (x - closestm)) - exp(-K * (uM - um)))/K))" apply (intro mono_intros B) using Laux \delta > 0\ \lambda \ 1\ by (auto simp add: divide_simps) finally have C: "L + 4 * delta \ Kmult * (exp(-K * (x - closestm)) - exp(-K * (uM - um)))" unfolding Kmult_def by argo have "Gromov_product_at (f z) (f um) (f uM) \ Gromov_product_at (f z) (f closestm) (f x) + (L + 4 * delta)" apply (rule Rec) using \closestm \ {um..ym}\ \x \ {yM..uM}\ \yM \ {z..uM}\ by auto also have "... \ (lambda^2 * (D + 3/2 * L + delta + 11/2 * C) - 2 * delta + Kmult * (1 - exp(- K * (x - closestm)))) + (Kmult * (exp(-K * (x - closestm)) - exp(-K * (uM-um))))" apply (intro mono_intros C Suc.IH) using \x \ {yM..uM}\ \yM \ {z..uM}\ \um \ {a..z}\ \closestm \ {um..ym}\ \ym \ {um..z}\ \uM \ {z..b}\ \x - closestm \ n * (1/4) * delta / lambda\ by auto also have "... = (lambda^2 * (D + 3/2 * L + delta + 11/2 * C) - 2 * delta + Kmult * (1 - exp(- K * (uM - um))))" unfolding K_def by (simp add: algebra_simps) finally show ?thesis by auto next case False have "\w\{yM..uM}. (\r\{w..uM}. (2 ^ (Suc k + 1) - 1) * dM \ dist (f r) (p r)) \ L - 4 * delta + 7 * QC (Suc k) \ dist (q (Suc k) uM) (q (Suc k) w)" proof (rule bexI[of _ w], auto) show "w \ uM" "yM \ w" using \w \ {x..uM}\ \x \ {yM..uM}\ by auto show "(4 * 2 ^ k - 1) * dM \ dist (f x) (p x)" if "x \ uM" "w \ x" for x using False \dM \ 0\ that by force have "dist (q k uM) (q (k+1) uM) = 2^k * dM" unfolding q_def apply (subst geodesic_segment_param(7)[where ?y = "f uM"]) using x(3)[of uM] \x \ {yM..uM}\ aux by (auto simp add: metric_space_class.dist_commute, simp add: algebra_simps) have "dist (q k w) (q (k+1) w) = 2^k * dM" unfolding q_def apply (subst geodesic_segment_param(7)[where ?y = "f w"]) using x(3)[of w] \w \ {x..uM}\ \x \ {yM..uM}\ aux by (auto simp add: metric_space_class.dist_commute, simp add: algebra_simps) have i: "q k uM \ proj_set (q (k+1) uM) (V k)" unfolding q_def V_def apply (rule proj_set_thickening'[of _ "f uM"]) using p x(3)[of uM] \x \ {yM..uM}\ aux by (auto simp add: algebra_simps metric_space_class.dist_commute) have j: "q k w \ proj_set (q (k+1) w) (V k)" unfolding q_def V_def apply (rule proj_set_thickening'[of _ "f w"]) using p x(3)[of w] \x \ {yM..uM}\ \w \ {x..uM}\ aux by (auto simp add: algebra_simps metric_space_class.dist_commute) have "5 * delta + 2 * QC k \ dist (q k uM) (q k w)" using w(2) by simp also have "... \ max (5 * deltaG(TYPE('a)) + 2 * QC k) (dist (q (k + 1) uM) (q (k + 1) w) - dist (q k uM) (q (k + 1) uM) - dist (q k w) (q (k + 1) w) + 10 * deltaG(TYPE('a)) + 4 * QC k)" by (rule proj_along_quasiconvex_contraction[OF \quasiconvex (QC k) (V k)\ i j]) finally have "5 * delta + 2 * QC k \ dist (q (k + 1) uM) (q (k + 1) w) - dist (q k uM) (q (k + 1) uM) - dist (q k w) (q (k + 1) w) + 10 * deltaG(TYPE('a)) + 4 * QC k" using \deltaG(TYPE('a)) < delta\ by auto then have "0 \ dist (q (k + 1) uM) (q (k + 1) w) + 5 * delta + 2 * QC k - dist (q k uM) (q (k + 1) uM) - dist (q k w) (q (k + 1) w)" using \deltaG(TYPE('a)) < delta\ by auto also have "... = dist (q (k + 1) uM) (q (k + 1) w) + 5 * delta + 2 * QC k - 2^(k+1) * dM" by (simp only: \dist (q k w) (q (k+1) w) = 2^k * dM\ \dist (q k uM) (q (k+1) uM) = 2^k * dM\, auto) finally have *: "2^(k+1) * dM - 5 * delta - 2 * QC k \ dist (q (k+1) uM) (q (k+1) w)" using \deltaG(TYPE('a)) < delta\ by auto have "L - 4 * delta + 7 * QC (k+1) \ 2 * dM - 5 * delta - 2 * QC k" unfolding QC_def L_def using \delta > 0\ Laux I \C \ 0\ by auto also have "... \ 2^(k+1) * dM - 5 * delta - 2 * QC k" using aux by (auto simp add: algebra_simps) finally show "L - 4 * delta + 7 * QC (Suc k) \ dist (q (Suc k) uM) (q (Suc k) w)" using * by auto qed then show ?thesis by simp qed qed qed have "dM > 0" using I \delta > 0\ \C \ 0\ Laux by auto have "\k. 2^k > dist (f uM) (p uM)/dM + 1" by (simp add: real_arch_pow) then obtain k where "2^k > dist (f uM) (p uM)/dM + 1" by blast then have "dist (f uM) (p uM) < (2^k - 1) * dM" using \dM > 0\ by (auto simp add: divide_simps algebra_simps) also have "... \ (2^(Suc k) - 1) * dM" by (intro mono_intros, auto) finally have "\((2 ^ (k + 1) - 1) * dM \ dist (f uM) (p uM))" by simp then show "Gromov_product_at (f z) (f um) (f uM) \ lambda\<^sup>2 * (D + 3/2 * L + delta + 11/2 * C) - 2 * delta + Kmult * (1 - exp (- K * (uM - um)))" using Ind_k[of k] by auto qed qed qed text \The main induction is over. To conclude, one should apply its result to the original geodesic segment joining the points $f(a)$ and $f(b)$.\ obtain n::nat where "(b - a)/((1/4) * delta / lambda) \ n" using real_arch_simple by blast then have "b - a \ n * (1/4) * delta / lambda" using \delta > 0\ \lambda \ 1\ by (auto simp add: divide_simps) have "infdist (f z) G \ Gromov_product_at (f z) (f a) (f b) + 2 * deltaG(TYPE('a))" apply (intro mono_intros) using assms by auto also have "... \ (lambda^2 * (D + 3/2 * L + delta + 11/2 * C) - 2 * delta + Kmult * (1 - exp(-K * (b - a)))) + 2 * delta" apply (intro mono_intros Main[OF _ _ \b - a \ n * (1/4) * delta / lambda\]) using assms by auto also have "... = lambda^2 * (D + 3/2 * L + delta + 11/2 * C) + Kmult * (1 - exp(-K * (b - a)))" by simp also have "... \ lambda^2 * (D + 3/2 * L + delta + 11/2 * C) + Kmult * (1 - 0)" apply (intro mono_intros) using \Kmult > 0\ by auto also have "... = lambda^2 * (11/2 * C + (3200*exp(-459/50*ln 2)/ln 2 + 83) * delta)" unfolding Kmult_def K_def L_def alpha_def D_def using \delta > 0\ \lambda \ 1\ by (simp add: algebra_simps divide_simps power2_eq_square mult_exp_exp) also have "... \ lambda^2 * (11/2 * C + 91 * delta)" apply (intro mono_intros, simp add: divide_simps, approximation 14) using \delta > 0\ by auto finally show ?thesis by (simp add: algebra_simps) qed text \Still assuming that our quasi-isometry is Lipschitz, we will improve slightly on the previous result, first going down to the hyperbolicity constant of the space, and also showing that, conversely, the geodesic is contained in a neighborhood of the quasi-geodesic. The argument for this last point goes as follows. Consider a point $x$ on the geodesic. Define two sets to be the $D$-thickenings of $[a,x]$ and $[x,b]$ respectively, where $D$ is such that any point on the quasi-geodesic is within distance $D$ of the geodesic (as given by the previous theorem). The union of these two sets covers the quasi-geodesic, and they are both closed and nonempty. By connectedness, there is a point $z$ in their intersection, $D$-close both to a point $x^-$ before $x$ and to a point $x^+$ after $x$. Then $x$ belongs to a geodesic between $x^-$ and $x^+$, which is contained in a $4\delta$-neighborhood of geodesics from $x^+$ to $z$ and from $x^-$ to $z$ by hyperbolicity. It follows that $x$ is at distance at most $D + 4\delta$ of $z$, concluding the proof.\ lemma (in Gromov_hyperbolic_space_geodesic) Morse_Gromov_theorem_aux2: fixes f::"real \ 'a" assumes "continuous_on {a..b} f" "lambda C-quasi_isometry_on {a..b} f" "geodesic_segment_between G (f a) (f b)" shows "hausdorff_distance (f`{a..b}) G \ lambda^2 * (11/2 * C + 92 * deltaG(TYPE('a)))" proof (cases "a \ b") case True have "lambda \ 1" "C \ 0" using quasi_isometry_onD[OF assms(2)] by auto have *: "infdist (f z) G \ lambda^2 * (11/2 * C + 91 * delta)" if "z \ {a..b}" "delta > deltaG(TYPE('a))" for z delta by (rule Morse_Gromov_theorem_aux1[OF assms(1) assms(2) True assms(3) that]) define D where "D = lambda^2 * (11/2 * C + 91 * deltaG(TYPE('a)))" have "D \ 0" unfolding D_def using \C \ 0\ by auto have I: "infdist (f z) G \ D" if "z \ {a..b}" for z proof - have "(infdist (f z) G/ lambda^2 - 11/2 * C)/91 \ delta" if "delta > deltaG(TYPE('a))" for delta using *[OF \z \ {a..b}\ that] \lambda \ 1\ by (auto simp add: divide_simps algebra_simps) then have "(infdist (f z) G/ lambda^2 - 11/2 * C)/91 \ deltaG(TYPE('a))" using dense_ge by blast then show ?thesis unfolding D_def using \lambda \ 1\ by (auto simp add: divide_simps algebra_simps) qed show ?thesis proof (rule hausdorff_distanceI) show "0 \ lambda\<^sup>2 * (11/2 * C + 92 * deltaG TYPE('a))" using \C \ 0\ by auto fix x assume "x \ f`{a..b}" then obtain z where z: "x = f z" "z \ {a..b}" by blast show "infdist x G \ lambda\<^sup>2 * (11/2 * C + 92 * deltaG TYPE('a))" unfolding z(1) by (rule order_trans[OF I[OF \z \ {a..b}\]], auto simp add: algebra_simps D_def) next fix x assume "x \ G" have "infdist x (f`{a..b}) \ D + 1 * deltaG TYPE('a)" proof - define p where "p = geodesic_segment_param G (f a)" then have p: "p 0 = f a" "p (dist (f a) (f b)) = f b" unfolding p_def using assms(3) by auto obtain t where t: "x = p t" "t \ {0..dist (f a) (f b)}" unfolding p_def using \x \ G\ \geodesic_segment_between G (f a) (f b)\ by (metis geodesic_segment_param(5) imageE) define Km where "Km = (\z \ p`{0..t}. cball z D)" define KM where "KM = (\z \ p`{t..dist (f a) (f b)}. cball z D)" have "f`{a..b} \ Km \ KM" proof fix x assume x: "x \ f`{a..b}" have "\z \ G. infdist x G = dist x z" apply (rule infdist_proper_attained) using geodesic_segment_topology[OF geodesic_segmentI[OF assms(3)]] by auto then obtain z where z: "z \ G" "infdist x G = dist x z" by auto obtain tz where tz: "z = p tz" "tz \ {0..dist (f a) (f b)}" unfolding p_def using \z \ G\ \geodesic_segment_between G (f a) (f b)\ by (metis geodesic_segment_param(5) imageE) have "infdist x G \ D" using I \x \ f`{a..b}\ by auto then have "dist z x \ D" using z(2) by (simp add: metric_space_class.dist_commute) then show "x \ Km \ KM" unfolding Km_def KM_def using tz by force qed then have *: "f`{a..b} = (Km \ f`{a..b}) \ (KM \ f`{a..b})" by auto have "(Km \ f`{a..b}) \ (KM \ f`{a..b}) \ {}" proof (rule connected_as_closed_union[OF _ *]) have "closed (f ` {a..b})" apply (intro compact_imp_closed compact_continuous_image) using assms(1) by auto have "closed Km" unfolding Km_def apply (intro compact_has_closed_thickening compact_continuous_image) apply (rule continuous_on_subset[of "{0..dist (f a) (f b)}" p]) unfolding p_def using assms(3) \t \ {0..dist (f a) (f b)}\ by (auto simp add: isometry_on_continuous) then show "closed (Km \ f`{a..b})" by (rule topological_space_class.closed_Int) fact have "closed KM" unfolding KM_def apply (intro compact_has_closed_thickening compact_continuous_image) apply (rule continuous_on_subset[of "{0..dist (f a) (f b)}" p]) unfolding p_def using assms(3) \t \ {0..dist (f a) (f b)}\ by (auto simp add: isometry_on_continuous) then show "closed (KM \ f`{a..b})" by (rule topological_space_class.closed_Int) fact show "connected (f`{a..b})" apply (rule connected_continuous_image) using assms(1) by auto have "f a \ Km \ f`{a..b}" using True apply auto unfolding Km_def apply auto apply (rule bexI[of _ 0]) unfolding p using \D \ 0\ t(2) by auto then show "Km \ f`{a..b} \ {}" by auto have "f b \ KM \ f`{a..b}" apply auto unfolding KM_def apply auto apply (rule bexI[of _ "dist (f a) (f b)"]) unfolding p using \D \ 0\ t(2) True by auto then show "KM \ f`{a..b} \ {}" by auto qed then obtain y where y: "y \ f`{a..b}" "y \ Km" "y \ KM" by auto obtain tm where tm: "tm \ {0..t}" "dist (p tm) y \ D" using y(2) unfolding Km_def by auto obtain tM where tM: "tM \ {t..dist (f a) (f b)}" "dist (p tM) y \ D" using y(3) unfolding KM_def by auto define H where "H = p`{tm..tM}" have *: "geodesic_segment_between H (p tm) (p tM)" unfolding H_def p_def apply (rule geodesic_segmentI2) using assms(3) \tm \ {0..t}\ \tM \ {t..dist (f a) (f b)}\ isometry_on_subset using assms(3) geodesic_segment_param(4) by (auto) fastforce have "x \ H" unfolding t(1) H_def using \tm \ {0..t}\ \tM \ {t..dist (f a) (f b)}\ by auto have "infdist x (f ` {a..b}) \ dist x y" by (rule infdist_le[OF y(1)]) also have "... \ max (dist (p tm) y) (dist (p tM) y) + deltaG(TYPE('a))" by (rule dist_le_max_dist_triangle[OF * \x \ H\]) finally show ?thesis using tm(2) tM(2) by auto qed also have "... \ D + lambda^2 * deltaG TYPE('a)" apply (intro mono_intros) using \lambda \ 1\ by auto finally show "infdist x (f ` {a..b}) \ lambda\<^sup>2 * (11/2 * C + 92 * deltaG TYPE('a))" unfolding D_def by (simp add: algebra_simps) qed next case False then have "f`{a..b} = {}" by auto then have "hausdorff_distance (f ` {a..b}) G = 0" unfolding hausdorff_distance_def by auto then show ?thesis using quasi_isometry_onD(4)[OF assms(2)] by auto qed text \The full statement of the Morse-Gromov Theorem, asserting that a quasi-geodesic is within controlled distance of a geodesic with the same endpoints. It is given in the formulation of Shchur~\cite{shchur}, with optimal control in terms of the parameters of the quasi-isometry. This statement follows readily from the previous one and from the fact that quasi-geodesics can be approximated by Lipschitz ones.\ theorem (in Gromov_hyperbolic_space_geodesic) Morse_Gromov_theorem: fixes f::"real \ 'a" assumes "lambda C-quasi_isometry_on {a..b} f" "geodesic_segment_between G (f a) (f b)" shows "hausdorff_distance (f`{a..b}) G \ 92 * lambda^2 * (C + deltaG(TYPE('a)))" proof - have C: "C \ 0" "lambda \ 1" using quasi_isometry_onD[OF assms(1)] by auto consider "dist (f a) (f b) \ 2 * C \ a \ b" | "dist (f a) (f b) \ 2 * C \ a \ b" | "b < a" by linarith then show ?thesis proof (cases) case 1 have "\d. continuous_on {a..b} d \ d a = f a \ d b = f b \ (\x\{a..b}. dist (f x) (d x) \ 4 * C) \ lambda (4 * C)-quasi_isometry_on {a..b} d \ (2 * lambda)-lipschitz_on {a..b} d \ hausdorff_distance (f`{a..b}) (d`{a..b}) \ 2 * C" apply (rule quasi_geodesic_made_lipschitz[OF assms(1)]) using 1 by auto then obtain d where d: "d a = f a" "d b = f b" "\x. x \ {a..b} \ dist (f x) (d x) \ 4 * C" "lambda (4 * C)-quasi_isometry_on {a..b} d" "(2 * lambda)-lipschitz_on {a..b} d" "hausdorff_distance (f`{a..b}) (d`{a..b}) \ 2 * C" by auto have a: "hausdorff_distance (d`{a..b}) G \ lambda^2 * ((11/2) * (4 * C) + 92 * deltaG(TYPE('a)))" apply (rule Morse_Gromov_theorem_aux2) using d assms lipschitz_on_continuous_on by auto have "hausdorff_distance (f`{a..b}) G \ hausdorff_distance (f`{a..b}) (d`{a..b}) + hausdorff_distance (d`{a..b}) G" apply (rule hausdorff_distance_triangle) using 1 apply simp by (rule quasi_isometry_on_bounded[OF d(4)], auto) also have "... \ lambda^2 * ((11/2) * (4 * C) + 92 * deltaG(TYPE('a))) + 1 * 2 * C" using a d by auto also have "... \ lambda^2 * ((11/2) * (4 * C) + 92 * deltaG(TYPE('a))) + lambda^2 * 2 * C" apply (intro mono_intros) using \lambda \ 1\ \C \ 0\ by auto also have "... = lambda^2 * (24 * C + 92 * deltaG(TYPE('a)))" by (simp add: algebra_simps divide_simps) also have "... \ lambda^2 * (92 * C + 92 * deltaG(TYPE('a)))" apply (intro mono_intros) using \lambda \ 1\ \C \ 0\ by auto finally show ?thesis by (auto simp add: algebra_simps) next case 2 have "(1/lambda) * dist a b - C \ dist (f a) (f b)" apply (rule quasi_isometry_onD[OF assms(1)]) using 2 by auto also have "... \ 2 * C" using 2 by auto finally have "dist a b \ 3 * lambda * C" using C by (auto simp add: algebra_simps divide_simps) then have *: "b - a \ 3 * lambda * C" using 2 unfolding dist_real_def by auto show ?thesis proof (rule hausdorff_distanceI2) show "0 \ 92 * lambda\<^sup>2 * (C + deltaG TYPE('a))" using C by auto fix x assume "x \ f`{a..b}" then obtain t where t: "x = f t" "t \ {a..b}" by auto have "dist x (f a) \ lambda * dist t a + C" unfolding t(1) using quasi_isometry_onD(1)[OF assms(1) t(2)] 2 by auto also have "... \ lambda * (b - a) + 1 * 1 * C + 0 * 0 * deltaG(TYPE('a))" using t(2) 2 C unfolding dist_real_def by auto also have "... \ lambda * (3 * lambda * C) + lambda^2 * (92-3) * C + lambda^2 * 92 * deltaG(TYPE('a))" apply (intro mono_intros *) using C by auto finally have *: "dist x (f a) \ 92 * lambda\<^sup>2 * (C + deltaG TYPE('a))" by (simp add: algebra_simps power2_eq_square) show "\y\G. dist x y \ 92 * lambda\<^sup>2 * (C + deltaG TYPE('a))" apply (rule bexI[of _ "f a"]) using * 2 assms(2) by auto next fix x assume "x \ G" then have "dist x (f a) \ dist (f a) (f b)" by (meson assms geodesic_segment_dist_le geodesic_segment_endpoints(1) local.some_geodesic_is_geodesic_segment(1)) also have "... \ 1 * 2 * C + lambda^2 * 0 * deltaG(TYPE('a))" using 2 by auto also have "... \ lambda^2 * 92 * C + lambda^2 * 92 * deltaG(TYPE('a))" apply (intro mono_intros) using C by auto finally have *: "dist x (f a) \ 92 * lambda\<^sup>2 * (C + deltaG TYPE('a))" by (simp add: algebra_simps) show "\y\f`{a..b}. dist x y \ 92 * lambda\<^sup>2 * (C + deltaG TYPE('a))" apply (rule bexI[of _ "f a"]) using * 2 by auto qed next case 3 then have "hausdorff_distance (f ` {a..b}) G = 0" unfolding hausdorff_distance_def by auto then show ?thesis using C by auto qed qed text \This theorem implies the same statement for two quasi-geodesics sharing their endpoints.\ theorem (in Gromov_hyperbolic_space_geodesic) Morse_Gromov_theorem2: fixes c d::"real \ 'a" assumes "lambda C-quasi_isometry_on {A..B} c" "lambda C-quasi_isometry_on {A..B} d" "c A = d A" "c B = d B" shows "hausdorff_distance (c`{A..B}) (d`{A..B}) \ 184 * lambda^2 * (C + deltaG(TYPE('a)))" proof (cases "A \ B") case False then have "hausdorff_distance (c`{A..B}) (d`{A..B}) = 0" by auto then show ?thesis using quasi_isometry_onD[OF assms(1)] delta_nonneg by auto next case True have "hausdorff_distance (c`{A..B}) {c A--c B} \ 92 * lambda^2 * (C + deltaG(TYPE('a)))" by (rule Morse_Gromov_theorem[OF assms(1)], auto) moreover have "hausdorff_distance {c A--c B} (d`{A..B}) \ 92 * lambda^2 * (C + deltaG(TYPE('a)))" unfolding \c A = d A\ \c B = d B\ apply (subst hausdorff_distance_sym) by (rule Morse_Gromov_theorem[OF assms(2)], auto) moreover have "hausdorff_distance (c`{A..B}) (d`{A..B}) \ hausdorff_distance (c`{A..B}) {c A--c B} + hausdorff_distance {c A--c B} (d`{A..B})" apply (rule hausdorff_distance_triangle) using True compact_imp_bounded[OF some_geodesic_compact] by auto ultimately show ?thesis by auto qed text \We deduce from the Morse lemma that hyperbolicity is invariant under quasi-isometry.\ text \First, we note that the image of a geodesic segment under a quasi-isometry is close to a geodesic segment in Hausdorff distance, as it is a quasi-geodesic.\ lemma geodesic_quasi_isometric_image: fixes f::"'a::metric_space \ 'b::Gromov_hyperbolic_space_geodesic" assumes "lambda C-quasi_isometry_on UNIV f" "geodesic_segment_between G x y" shows "hausdorff_distance (f`G) {f x--f y} \ 92 * lambda^2 * (C + deltaG(TYPE('b)))" proof - define c where "c = f o (geodesic_segment_param G x)" have *: "(1 * lambda) (0 * lambda + C)-quasi_isometry_on {0..dist x y} c" unfolding c_def by (rule quasi_isometry_on_compose[where Y = UNIV], auto intro!: isometry_quasi_isometry_on simp add: assms) have "hausdorff_distance (c`{0..dist x y}) {c 0--c (dist x y)} \ 92 * lambda^2 * (C + deltaG(TYPE('b)))" apply (rule Morse_Gromov_theorem) using * by auto moreover have "c`{0..dist x y} = f`G" unfolding c_def image_comp[symmetric] using assms(2) by auto moreover have "c 0 = f x" "c (dist x y) = f y" unfolding c_def using assms(2) by auto ultimately show ?thesis by auto qed text \We deduce that hyperbolicity is invariant under quasi-isometry. The proof goes as follows: we want to see that a geodesic triangle is delta-thin, i.e., a point on a side $Gxy$ is close to the union of the two other sides $Gxz$ and $Gyz$. Pull everything back by the quasi-isometry: we obtain three quasi-geodesic, each of which is close to the corresponding geodesic segment by the Morse lemma. As the geodesic triangle is thin, it follows that the quasi-geodesic triangle is also thin, i.e., a point on $f^{-1}Gxy$ is close to $f^{-1}Gxz \cup f^{-1}Gyz$ (for some explicit, albeit large, constant). Then push everything forward by $f$: as it is a quasi-isometry, it will again distort distances by a bounded amount.\ lemma Gromov_hyperbolic_invariant_under_quasi_isometry_explicit: fixes f::"'a::geodesic_space \ 'b::Gromov_hyperbolic_space_geodesic" assumes "lambda C-quasi_isometry f" shows "Gromov_hyperbolic_subset (752 * lambda^3 * (C + deltaG(TYPE('b)))) (UNIV::('a set))" proof - have C: "lambda \ 1" "C \ 0" using quasi_isometry_onD[OF assms] by auto text \The Morse lemma gives a control bounded by $K$ below. Following the proof, we deduce a bound on the thinness of triangles by an ugly constant $L$. We bound it by a more tractable (albeit still ugly) constant $M$.\ define K where "K = 92 * lambda^2 * (C + deltaG(TYPE('b)))" have HD: "hausdorff_distance (f`G) {f a--f b} \ K" if "geodesic_segment_between G a b" for G a b unfolding K_def by (rule geodesic_quasi_isometric_image[OF assms that]) define L where "L = lambda * (4 * 1 * deltaG(TYPE('b)) + 1 * 1 * C + 2 * K)" define M where "M = 188 * lambda^3 * (C + deltaG(TYPE('b)))" have "L \ lambda * (4 * lambda^2 * deltaG(TYPE('b)) + 4 * lambda^2 * C + 2 * K)" unfolding L_def apply (intro mono_intros) using C by auto also have "... = M" unfolding M_def K_def by (auto simp add: algebra_simps power2_eq_square power3_eq_cube) finally have "L \ M" by simp text \After these preliminaries, we start the real argument per se, showing that triangles are thin in the type b.\ have Thin: "infdist w (Gxz \ Gyz) \ M" if H: "geodesic_segment_between Gxy x y" "geodesic_segment_between Gxz x z" "geodesic_segment_between Gyz y z" "w \ Gxy" for w x y z::'a and Gxy Gyz Gxz proof - obtain w2 where w2: "w2 \ {f x--f y}" "infdist (f w) {f x--f y} = dist (f w) w2" using infdist_proper_attained[OF proper_of_compact, of "{f x--f y}" "f w"] by auto have "dist (f w) w2 = infdist (f w) {f x-- f y}" using w2 by simp also have "... \ hausdorff_distance (f`Gxy) {f x-- f y}" using geodesic_segment_topology(4)[OF geodesic_segmentI] H by (auto intro!: quasi_isometry_on_bounded[OF quasi_isometry_on_subset[OF assms]] infdist_le_hausdorff_distance) also have "... \ K" using HD[OF H(1)] by simp finally have *: "dist (f w) w2 \ K" by simp have "infdist w2 (f`Gxz \ f`Gyz) \ infdist w2 ({f x--f z} \ {f y--f z}) + hausdorff_distance ({f x--f z} \ {f y--f z}) (f`Gxz \ f`Gyz)" apply (rule hausdorff_distance_infdist_triangle) using geodesic_segment_topology(4)[OF geodesic_segmentI] H by (auto intro!: quasi_isometry_on_bounded[OF quasi_isometry_on_subset[OF assms]]) also have "... \ 4 * deltaG(TYPE('b)) + hausdorff_distance ({f x--f z} \ {f y--f z}) (f`Gxz \ f`Gyz)" apply (simp, rule thin_triangles[of "{f x--f z}" "f z" "f x" "{f y--f z}" "f y" "{f x--f y}" w2]) using w2 apply auto using geodesic_segment_commute some_geodesic_is_geodesic_segment(1) by blast+ also have "... \ 4 * deltaG(TYPE('b)) + max (hausdorff_distance {f x--f z} (f`Gxz)) (hausdorff_distance {f y--f z} (f`Gyz))" apply (intro mono_intros) using H by auto also have "... \ 4 * deltaG(TYPE('b)) + K" using HD[OF H(2)] HD[OF H(3)] by (auto simp add: hausdorff_distance_sym) finally have **: "infdist w2 (f`Gxz \ f`Gyz) \ 4 * deltaG(TYPE('b)) + K" by simp have "infdist (f w) (f`Gxz \ f`Gyz) \ infdist w2 (f`Gxz \ f`Gyz) + dist (f w) w2" by (rule infdist_triangle) then have A: "infdist (f w) (f`(Gxz \ Gyz)) \ 4 * deltaG(TYPE('b)) + 2 * K" using * ** by (auto simp add: image_Un) have "infdist w (Gxz \ Gyz) \ L + epsilon" if "epsilon > 0" for epsilon proof - have *: "epsilon/lambda > 0" using that C by auto have "\z \ f`(Gxz \ Gyz). dist (f w) z < 4 * deltaG(TYPE('b)) + 2 * K + epsilon/lambda" apply (rule infdist_almost_attained) using A * H(2) by auto then obtain z where z: "z \ Gxz \ Gyz" "dist (f w) (f z) < 4 * deltaG(TYPE('b)) + 2 * K + epsilon/lambda" by auto have "infdist w (Gxz \ Gyz) \ dist w z" by (auto intro!: infdist_le z(1)) also have "... \ lambda * dist (f w) (f z) + C * lambda" using quasi_isometry_onD[OF assms] by (auto simp add: algebra_simps divide_simps) also have "... \ lambda * (4 * deltaG(TYPE('b)) + 2 * K + epsilon/lambda) + C * lambda" apply (intro mono_intros) using z(2) C by auto also have "... = L + epsilon" unfolding K_def L_def using C by (auto simp add: algebra_simps) finally show ?thesis by simp qed then have "infdist w (Gxz \ Gyz) \ L" using field_le_epsilon by blast then show ?thesis using \L \ M\ by auto qed then have "Gromov_hyperbolic_subset (4 * M) (UNIV::'a set)" using thin_triangles_implies_hyperbolic[OF Thin] by auto then show ?thesis unfolding M_def by (auto simp add: algebra_simps) qed text \Most often, the precise value of the constant in the previous theorem is irrelevant, it is used in the following form.\ theorem Gromov_hyperbolic_invariant_under_quasi_isometry: assumes "quasi_isometric (UNIV::('a::geodesic_space) set) (UNIV::('b::Gromov_hyperbolic_space_geodesic) set)" shows "\delta. Gromov_hyperbolic_subset delta (UNIV::'a set)" proof - obtain C lambda f where f: "lambda C-quasi_isometry_between (UNIV::'a set) (UNIV::'b set) f" using assms unfolding quasi_isometric_def by auto show ?thesis using Gromov_hyperbolic_invariant_under_quasi_isometry_explicit[OF quasi_isometry_betweenD(1)[OF f]] by blast qed text \A central feature of hyperbolic spaces is that a path from $x$ to $y$ can not deviate too much from a geodesic from $x$ to $y$ unless it is extremely long (exponentially long in terms of the distance from $x$ to $y$). This is useful both to ensure that short paths (for instance quasi-geodesics) stay close to geodesics, see the Morse lemme below, and to ensure that paths that avoid a given large ball of radius $R$ have to be exponentially long in terms of $R$ (this is extremely useful for random walks). This proposition is the first non-trivial result on hyperbolic spaces in~\cite{bridson_haefliger} (Proposition III.H.1.6). We follow their proof. The proof is geometric, and uses the existence of geodesics and the fact that geodesic triangles are thin. In fact, the result still holds if the space is not geodesic, as it can be deduced by embedding the hyperbolic space in a geodesic hyperbolic space and using the result there.\ proposition (in Gromov_hyperbolic_space_geodesic) lipschitz_path_close_to_geodesic: fixes c::"real \ 'a" assumes "M-lipschitz_on {A..B} c" "geodesic_segment_between G (c A) (c B)" "x \ G" shows "infdist x (c`{A..B}) \ (4/ln 2) * deltaG(TYPE('a)) * max 0 (ln (B-A)) + M" proof - have "M \ 0" by (rule lipschitz_on_nonneg[OF assms(1)]) have Main: "a \ {A..B} \ b \ {A..B} \ a \ b \ b-a \ 2^(n+1) \ geodesic_segment_between H (c a) (c b) \ y \ H \ infdist y (c`{A..B}) \ 4 * deltaG(TYPE('a)) * n + M" for a b H y n proof (induction n arbitrary: a b H y) case 0 have "infdist y (c ` {A..B}) \ dist y (c b)" apply (rule infdist_le) using \b \ {A..B}\ by auto moreover have "infdist y (c ` {A..B}) \ dist y (c a)" apply (rule infdist_le) using \a \ {A..B}\ by auto ultimately have "2 * infdist y (c ` {A..B}) \ dist (c a) y + dist y (c b)" by (auto simp add: metric_space_class.dist_commute) also have "... = dist (c a) (c b)" by (rule geodesic_segment_dist[OF \geodesic_segment_between H (c a) (c b)\ \y \ H\]) also have "... \ M * abs(b - a)" using lipschitz_onD(1)[OF assms(1) \a \ {A..B}\ \b \ {A..B}\] unfolding dist_real_def by (simp add: abs_minus_commute) also have "... \ M * 2" using \a \ b\ \b - a \ 2^(0 + 1)\ \M \ 0\ mult_left_mono by auto finally show ?case by simp next case (Suc n) define m where "m = (a + b)/2" have "m \ {A..B}" using \a \ {A..B}\ \b \ {A..B}\ unfolding m_def by auto define Ha where "Ha = {c m--c a}" define Hb where "Hb = {c m--c b}" have I: "geodesic_segment_between Ha (c m) (c a)" "geodesic_segment_between Hb (c m) (c b)" unfolding Ha_def Hb_def by auto then have "Ha \ {}" "Hb \ {}" "compact Ha" "compact Hb" by (auto intro: geodesic_segment_topology) have *: "infdist y (Ha \ Hb) \ 4 * deltaG(TYPE('a))" by (rule thin_triangles[OF I \geodesic_segment_between H (c a) (c b)\ \y \ H\]) then have "infdist y Ha \ 4 * deltaG(TYPE('a)) \ infdist y Hb \ 4 * deltaG(TYPE('a))" unfolding infdist_union_min[OF \Ha \ {}\ \Hb \ {}\] by auto then show ?case proof assume H: "infdist y Ha \ 4 * deltaG TYPE('a)" obtain z where z: "z \ Ha" "infdist y Ha = dist y z" using infdist_proper_attained[OF proper_of_compact[OF \compact Ha\] \Ha \ {}\] by auto have Iz: "infdist z (c`{A..B}) \ 4 * deltaG(TYPE('a)) * n + M" proof (rule Suc.IH[OF \a \ {A..B}\ \m \ {A..B}\, of Ha]) show "a \ m" unfolding m_def using \a \ b\ by auto show "m - a \ 2^(n+1)" using \b - a \ 2^(Suc n + 1)\ \a \ b\ unfolding m_def by auto show "geodesic_segment_between Ha (c a) (c m)" by (simp add: I(1) geodesic_segment_commute) show "z \ Ha" using z by auto qed have "infdist y (c`{A..B}) \ dist y z + infdist z (c`{A..B})" by (metis add.commute infdist_triangle) also have "... \ 4 * deltaG TYPE('a) + (4 * deltaG(TYPE('a)) * n + M)" using H z Iz by (auto intro: add_mono) finally show "infdist y (c ` {A..B}) \ 4 * deltaG TYPE('a) * real (Suc n) + M" by (auto simp add: algebra_simps) next assume H: "infdist y Hb \ 4 * deltaG TYPE('a)" obtain z where z: "z \ Hb" "infdist y Hb = dist y z" using infdist_proper_attained[OF proper_of_compact[OF \compact Hb\] \Hb \ {}\] by auto have Iz: "infdist z (c`{A..B}) \ 4 * deltaG(TYPE('a)) * n + M" proof (rule Suc.IH[OF \m \ {A..B}\ \b \ {A..B}\, of Hb]) show "m \ b" unfolding m_def using \a \ b\ by auto show "b - m \ 2^(n+1)" using \b - a \ 2^(Suc n + 1)\ \a \ b\ unfolding m_def by (auto simp add: divide_simps) show "geodesic_segment_between Hb (c m) (c b)" by (simp add: I(2)) show "z \ Hb" using z by auto qed have "infdist y (c`{A..B}) \ dist y z + infdist z (c`{A..B})" by (metis add.commute infdist_triangle) also have "... \ 4 * deltaG TYPE('a) + (4 * deltaG(TYPE('a)) * n + M)" using H z Iz by (auto intro: add_mono) finally show "infdist y (c ` {A..B}) \ 4 * deltaG TYPE('a) * real (Suc n) + M" by (auto simp add: algebra_simps) qed qed consider "B-A <0" | "B-A \ 0 \ B-A \ 2" | "B-A > 2" by linarith then show ?thesis proof (cases) case 1 then have "c`{A..B} = {}" by auto then show ?thesis unfolding infdist_def using \M \ 0\ by auto next case 2 have "infdist x (c`{A..B}) \ 4 * deltaG(TYPE('a)) * real 0 + M" apply (rule Main[OF _ _ _ _ \geodesic_segment_between G (c A) (c B)\ \x \ G\]) using 2 by auto also have "... \ (4/ln 2) * deltaG(TYPE('a)) * max 0 (ln (B-A)) + M" using delta_nonneg by auto finally show ?thesis by auto next case 3 define n::nat where "n = nat(floor (log 2 (B-A)))" have "log 2 (B-A) > 0" using 3 by auto then have n: "n \ log 2 (B-A)" "log 2 (B-A) < n+1" unfolding n_def by (auto simp add: floor_less_cancel) then have *: "B-A \ 2^(n+1)" by (meson le_log_of_power linear not_less one_less_numeral_iff semiring_norm(76)) have "n \ ln (B-A) * (1/ln 2)" using n unfolding log_def by auto then have "n \ (1/ln 2) * max 0 (ln (B-A))" using 3 by (auto simp add: algebra_simps divide_simps) have "infdist x (c`{A..B}) \ 4 * deltaG(TYPE('a)) * n + M" apply (rule Main[OF _ _ _ _ \geodesic_segment_between G (c A) (c B)\ \x \ G\]) using * 3 by auto also have "... \ 4 * deltaG(TYPE('a)) * ((1/ln 2) * max 0 (ln (B-A))) + M" apply (intro mono_intros) using \n \ (1/ln 2) * max 0 (ln (B-A))\ delta_nonneg by auto finally show ?thesis by auto qed qed text \By rescaling coordinates at the origin, one obtains a variation around the previous statement.\ proposition (in Gromov_hyperbolic_space_geodesic) lipschitz_path_close_to_geodesic': fixes c::"real \ 'a" assumes "M-lipschitz_on {A..B} c" "geodesic_segment_between G (c A) (c B)" "x \ G" "a > 0" shows "infdist x (c`{A..B}) \ (4/ln 2) * deltaG(TYPE('a)) * max 0 (ln (a * (B-A))) + M/a" proof - define d where "d = c o (\t. (1/a) * t)" have *: "(M * ((1/a)* 1))-lipschitz_on {a * A..a * B} d" unfolding d_def apply (rule lipschitz_on_compose, intro lipschitz_intros) using assms by auto have "d`{a * A..a * B} = c`{A..B}" unfolding d_def image_comp[symmetric] apply (rule arg_cong[where ?f = "image c"]) using \a > 0\ by auto then have "infdist x (c`{A..B}) = infdist x (d`{a * A..a * B})" by auto also have "... \ (4/ln 2) * deltaG(TYPE('a)) * max 0 (ln ((a * B)- (a * A))) + M/a" apply (rule lipschitz_path_close_to_geodesic[OF _ _ \x \ G\]) using * assms unfolding d_def by auto finally show ?thesis by (auto simp add: algebra_simps) qed text \We can now give another proof of the Morse-Gromov Theorem, as described in~\cite{bridson_haefliger}. It is more direct than the one we have given above, but it gives a worse dependence in terms of the quasi-isometry constants. In particular, when $C = \delta = 0$, it does not recover the fact that a quasi-geodesic has to coincide with a geodesic.\ theorem (in Gromov_hyperbolic_space_geodesic) Morse_Gromov_theorem_BH_proof: fixes c::"real \ 'a" assumes "lambda C-quasi_isometry_on {A..B} c" shows "hausdorff_distance (c`{A..B}) {c A--c B} \ 72 * lambda^2 * (C + lambda + deltaG(TYPE('a))^2)" proof - have C: "C \ 0" "lambda \ 1" using quasi_isometry_onD[OF assms] by auto consider "B-A < 0" | "B-A \ 0 \ dist (c A) (c B) \ 2 * C" | "B-A \ 0 \ dist (c A) (c B) > 2 * C" by linarith then show ?thesis proof (cases) case 1 then have "c`{A..B} = {}" by auto then show ?thesis unfolding hausdorff_distance_def using delta_nonneg C by auto next case 2 have "(1/lambda) * dist A B - C \ dist (c A) (c B)" apply (rule quasi_isometry_onD[OF assms]) using 2 by auto also have "... \ 2 * C" using 2 by auto finally have "dist A B \ 3 * lambda * C" using C by (auto simp add: algebra_simps divide_simps) then have *: "B - A \ 3 * lambda * C" using 2 unfolding dist_real_def by auto show ?thesis proof (rule hausdorff_distanceI2) show "0 \ 72 * lambda^2 * (C + lambda + deltaG(TYPE('a))^2)" using C by auto fix x assume "x \ c`{A..B}" then obtain t where t: "x = c t" "t \ {A..B}" by auto have "dist x (c A) \ lambda * dist t A + C" unfolding t(1) using quasi_isometry_onD(1)[OF assms t(2), of A] 2 by auto also have "... \ lambda * (B-A) + C" using t(2) 2 C unfolding dist_real_def by auto also have "... \ 3 * lambda * lambda * C + 1 * 1 * C" using * C by auto also have "... \ 3 * lambda * lambda * C + lambda * lambda * C" apply (intro mono_intros) using C by auto also have "... = 4 * lambda * lambda * (C + 0 + 0^2)" by auto also have "... \ 72 * lambda * lambda * (C + lambda + deltaG(TYPE('a))^2)" apply (intro mono_intros) using C delta_nonneg by auto finally have *: "dist x (c A) \ 72 * lambda^2 * (C + lambda + deltaG(TYPE('a))^2)" unfolding power2_eq_square by simp show "\y\{c A--c B}. dist x y \ 72 * lambda^2 * (C + lambda + deltaG(TYPE('a))^2)" apply (rule bexI[of _ "c A"]) using * by auto next fix x assume "x \ {c A-- c B}" then have "dist x (c A) \ dist (c A) (c B)" by (meson geodesic_segment_dist_le geodesic_segment_endpoints(1) local.some_geodesic_is_geodesic_segment(1)) also have "... \ 2 * C" using 2 by auto also have "... \ 2 * 1 * 1 * (C + lambda + 0)" using 2 C unfolding dist_real_def by auto also have "... \ 72 * lambda * lambda * (C + lambda + deltaG(TYPE('a)) * deltaG(TYPE('a)))" apply (intro mono_intros) using C delta_nonneg by auto finally have *: "dist x (c A) \ 72 * lambda * lambda * (C + lambda + deltaG(TYPE('a)) * deltaG(TYPE('a)))" by simp show "\y\c`{A..B}. dist x y \ 72 * lambda^2 * (C + lambda + deltaG(TYPE('a))^2)" apply (rule bexI[of _ "c A"]) unfolding power2_eq_square using * 2 by auto qed next case 3 then obtain d where d: "continuous_on {A..B} d" "d A = c A" "d B = c B" "\x. x \ {A..B} \ dist (c x) (d x) \ 4 *C" "lambda (4 * C)-quasi_isometry_on {A..B} d" "(2 * lambda)-lipschitz_on {A..B} d" "hausdorff_distance (c`{A..B}) (d`{A..B}) \ 2 * C" using quasi_geodesic_made_lipschitz[OF assms] C(1) by fastforce have "A \ {A..B}" "B \ {A..B}" using 3 by auto text \We show that the distance of any point in the geodesic from $c(A)$ to $c(B)$ is a bounded distance away from the quasi-geodesic $d$, by considering a point $x$ where the distance $D$ is maximal and arguing around this point. Consider the point $x_m$ on the geodesic $[c(A), c(B)]$ at distance $2D$ from $x$, and the closest point $y_m$ on the image of $d$. Then the distance between $x_m$ and $y_m$ is at most $D$. Hence a point on $[x_m,y_m]$ is at distance at least $2D - D = D$ of $x$. In the same way, define $x_M$ and $y_M$ on the other side of $x$. Then the excursion from $x_m$ to $y_m$, then to $y_M$ along $d$, then to $x_M$, has length at most $D + (\lambda \cdot 6D + C) + D$ and is always at distance at least $D$ from $x$. It follows from the previous lemma that $D \leq \log(length)$, which implies a bound on $D$. This argument has to be amended if $x$ is at distance $ < 2D$ from $c(A)$ or $c(B)$. In this case, simply use $x_m = y_m = c(A)$ or $x_M = y_M = c(B)$, then everything goes through.\ have "\x \ {c A--c B}. \y \ {c A--c B}. infdist y (d`{A..B}) \ infdist x (d`{A..B})" by (rule continuous_attains_sup, auto intro: continuous_intros) then obtain x where x: "x \ {c A--c B}" "\y. y \ {c A--c B} \ infdist y (d`{A..B}) \ infdist x (d`{A..B})" by auto define D where "D = infdist x (d`{A..B})" have "D \ 0" unfolding D_def by (rule infdist_nonneg) have D_bound: "D \ 24 * lambda + 12 * C + 24 * deltaG(TYPE('a))^2" proof (cases "D \ 1") case True have "1 * 1 + 1 * 0 + 0 * 0 \ 24 * lambda + 12 * C + 24 * deltaG(TYPE('a))^2" apply (intro mono_intros) using C delta_nonneg by auto then show ?thesis using True by auto next case False then have "D \ 1" by auto have ln2mult: "2 * ln t = ln (t * t)" if "t > 0" for t::real by (simp add: that ln_mult) have "infdist (c A) (d`{A..B}) = 0" using \d A = c A\ by (metis \A \ {A..B}\ image_eqI infdist_zero) then have "x \ c A" using \D \ 1\ D_def by auto define tx where "tx = dist (c A) x" then have "tx \ {0..dist (c A) (c B)}" using \x \ {c A--c B}\ by (meson atLeastAtMost_iff geodesic_segment_dist_le some_geodesic_is_geodesic_segment(1) metric_space_class.zero_le_dist some_geodesic_endpoints(1)) have "tx > 0" using \x \ c A\ tx_def by auto have x_param: "x = geodesic_segment_param {c A--c B} (c A) tx" using \x \ {c A--c B}\ geodesic_segment_param[OF some_geodesic_is_geodesic_segment(1)] tx_def by auto define tm where "tm = max (tx - 2 * D) 0" have "tm \ {0..dist (c A) (c B)}" unfolding tm_def using \tx \ {0..dist (c A) (c B)}\ \D \ 0\ by auto define xm where "xm = geodesic_segment_param {c A--c B} (c A) tm" have "xm \ {c A--c B}" using \tm \ {0..dist (c A) (c B)}\ by (metis geodesic_segment_param(3) local.some_geodesic_is_geodesic_segment(1) xm_def) have "dist xm x = abs((max (tx - 2 * D) 0) - tx)" unfolding xm_def tm_def x_param apply (rule geodesic_segment_param[of _ _ "c B"], auto) using \tx \ {0..dist (c A) (c B)}\ \D \ 0\ by auto also have "... \ 2 * D" by (simp add: \0 \ D\ tx_def) finally have "dist xm x \ 2 * D" by auto have "\ym\d`{A..B}. infdist xm (d`{A..B}) = dist xm ym" apply (rule infdist_proper_attained) using 3 d(1) proper_of_compact compact_continuous_image by auto then obtain ym where ym: "ym \ d`{A..B}" "dist xm ym = infdist xm (d`{A..B})" by metis then obtain um where um: "um \ {A..B}" "ym = d um" by auto have "dist xm ym \ D" unfolding D_def using x ym by (simp add: \xm \ {c A--c B}\) have D1: "dist x z \ D" if "z \ {xm--ym}" for z proof (cases "tx - 2 * D < 0") case True then have "tm = 0" unfolding tm_def by auto then have "xm = c A" unfolding xm_def by (meson geodesic_segment_param(1) local.some_geodesic_is_geodesic_segment(1)) then have "infdist xm (d`{A..B}) = 0" using \d A = c A\ \A \ {A..B}\ by (metis image_eqI infdist_zero) then have "ym = xm" using ym(2) by auto then have "z = xm" using \z \ {xm--ym}\ geodesic_segment_between_x_x(3) by (metis empty_iff insert_iff some_geodesic_is_geodesic_segment(1)) then have "z \ d`{A..B}" using \ym = xm\ ym(1) by blast then show "dist x z \ D" unfolding D_def by (simp add: infdist_le) next case False then have *: "tm = tx - 2 * D" unfolding tm_def by auto have "dist xm x = abs((tx - 2 * D) - tx)" unfolding xm_def x_param * apply (rule geodesic_segment_param[of _ _ "c B"], auto) using False \tx \ {0..dist (c A) (c B)}\ \D \ 0\ by auto then have "2 * D = dist xm x" using \D \ 0\ by auto also have "... \ dist xm z + dist x z" using metric_space_class.dist_triangle2 by auto also have "... \ dist xm ym + dist x z" using \z \ {xm--ym}\ by (auto, meson geodesic_segment_dist_le some_geodesic_is_geodesic_segment(1) some_geodesic_endpoints(1)) also have "... \ D + dist x z" using \dist xm ym \ D\ by simp finally show "dist x z \ D" by auto qed define tM where "tM = min (tx + 2 * D) (dist (c A) (c B))" have "tM \ {0..dist (c A) (c B)}" unfolding tM_def using \tx \ {0..dist (c A) (c B)}\ \D \ 0\ by auto have "tm \ tM" unfolding tM_def using \D \ 0\ \tm \ {0..dist (c A) (c B)}\ \tx \ dist (c A) x\ tm_def by auto define xM where "xM = geodesic_segment_param {c A--c B} (c A) tM" have "xM \ {c A--c B}" using \tM \ {0..dist (c A) (c B)}\ by (metis geodesic_segment_param(3) local.some_geodesic_is_geodesic_segment(1) xM_def) have "dist xM x = abs((min (tx + 2 * D) (dist (c A) (c B))) - tx)" unfolding xM_def tM_def x_param apply (rule geodesic_segment_param[of _ _ "c B"], auto) using \tx \ {0..dist (c A) (c B)}\ \D \ 0\ by auto also have "... \ 2 * D" using \0 \ D\ \tx \ {0..dist (c A) (c B)}\ by auto finally have "dist xM x \ 2 * D" by auto have "\yM\d`{A..B}. infdist xM (d`{A..B}) = dist xM yM" apply (rule infdist_proper_attained) using 3 d(1) proper_of_compact compact_continuous_image by auto then obtain yM where yM: "yM \ d`{A..B}" "dist xM yM = infdist xM (d`{A..B})" by metis then obtain uM where uM: "uM \ {A..B}" "yM = d uM" by auto have "dist xM yM \ D" unfolding D_def using x yM by (simp add: \xM \ {c A--c B}\) have D3: "dist x z \ D" if "z \ {xM--yM}" for z proof (cases "tx + 2 * D > dist (c A) (c B)") case True then have "tM = dist (c A) (c B)" unfolding tM_def by auto then have "xM = c B" unfolding xM_def by (meson geodesic_segment_param(2) local.some_geodesic_is_geodesic_segment(1)) then have "infdist xM (d`{A..B}) = 0" using \d B = c B\ \B \ {A..B}\ by (metis image_eqI infdist_zero) then have "yM = xM" using yM(2) by auto then have "z = xM" using \z \ {xM--yM}\ geodesic_segment_between_x_x(3) by (metis empty_iff insert_iff some_geodesic_is_geodesic_segment(1)) then have "z \ d`{A..B}" using \yM = xM\ yM(1) by blast then show "dist x z \ D" unfolding D_def by (simp add: infdist_le) next case False then have *: "tM = tx + 2 * D" unfolding tM_def by auto have "dist xM x = abs((tx + 2 * D) - tx)" unfolding xM_def x_param * apply (rule geodesic_segment_param[of _ _ "c B"], auto) using False \tx \ {0..dist (c A) (c B)}\ \D \ 0\ by auto then have "2 * D = dist xM x" using \D \ 0\ by auto also have "... \ dist xM z + dist x z" using metric_space_class.dist_triangle2 by auto also have "... \ dist xM yM + dist x z" using \z \ {xM--yM}\ by (auto, meson geodesic_segment_dist_le local.some_geodesic_is_geodesic_segment(1) some_geodesic_endpoints(1)) also have "... \ D + dist x z" using \dist xM yM \ D\ by simp finally show "dist x z \ D" by auto qed define excursion:: "real\'a" where "excursion = (\t. if t \ {0..dist xm ym} then (geodesic_segment_param {xm--ym} xm t) else if t \ {dist xm ym..dist xm ym + abs(uM - um)} then d (um + sgn(uM-um) * (t - dist xm ym)) else geodesic_segment_param {yM--xM} yM (t - dist xm ym - abs (uM -um)))" define L where "L = dist xm ym + abs(uM - um) + dist yM xM" have E1: "excursion t = geodesic_segment_param {xm--ym} xm t" if "t \ {0..dist xm ym}" for t unfolding excursion_def using that by auto have E2: "excursion t = d (um + sgn(uM-um) * (t - dist xm ym))" if "t \ {dist xm ym..dist xm ym + abs(uM - um)}" for t unfolding excursion_def using that by (auto simp add: \ym = d um\) have E3: "excursion t = geodesic_segment_param {yM--xM} yM (t - dist xm ym - abs (uM -um))" if "t \ {dist xm ym + \uM - um\..dist xm ym + \uM - um\ + dist yM xM}" for t unfolding excursion_def using that \yM = d uM\ \ym = d um\ by (auto simp add: sgn_mult_abs) have E0: "excursion 0 = xm" unfolding excursion_def by auto have EL: "excursion L = xM" unfolding excursion_def L_def apply (auto simp add: uM(2) um(2) sgn_mult_abs) by (metis (mono_tags, hide_lams) add.left_neutral add_increasing2 add_le_same_cancel1 dist_real_def Gromov_product_e_x_x Gromov_product_nonneg metric_space_class.dist_le_zero_iff) have [simp]: "L \ 0" unfolding L_def by auto have "L > 0" proof (rule ccontr) assume "\(L > 0)" then have "L = 0" using \L \ 0\ by simp then have "xm = xM" using E0 EL by auto then have "tM = tm" unfolding xm_def xM_def using \tM \ {0..dist (c A) (c B)}\ \tm \ {0..dist (c A) (c B)}\ local.geodesic_segment_param_in_geodesic_spaces(6) by fastforce also have "... < tx" unfolding tm_def using \tx > 0\ \D \ 1\ by auto also have "... \ tM" unfolding tM_def using \D \ 0\ \tx \ {0..dist (c A) (c B)}\ by auto finally show False by simp qed have "(1/lambda) * dist um uM - (4 * C) \ dist (d um) (d uM)" by (rule quasi_isometry_onD(2)[OF \lambda (4 * C)-quasi_isometry_on {A..B} d\ \um \ {A..B}\ \uM \ {A..B}\]) also have "... \ dist ym xm + dist xm x + dist x xM + dist xM yM" unfolding um(2)[symmetric] uM(2)[symmetric] by (rule dist_triangle5) also have "... \ D + (2*D) + (2*D) + D" using \dist xm ym \ D\ \dist xm x \ 2*D\ \dist xM x \ 2*D\ \dist xM yM \ D\ by (auto simp add: metric_space_class.dist_commute intro: add_mono) finally have "(1/lambda) * dist um uM \ 6*D + 4*C" by auto then have "dist um uM \ 6*D*lambda + 4*C*lambda" using C by (auto simp add: divide_simps algebra_simps) then have "L \ D + (6*D*lambda + 4*C*lambda) + D" unfolding L_def dist_real_def using \dist xm ym \ D\ \dist xM yM \ D\ by (auto simp add: metric_space_class.dist_commute intro: add_mono) also have "... \ 8 * D * lambda + 4*C*lambda" using C \D \ 0\ by (auto intro: mono_intros) finally have L_bound: "L \ lambda * (8 * D + 4 * C)" by (auto simp add: algebra_simps) have "1 * (1 * 1 + 0) \ lambda * (8 * D + 4 * C)" using C \D \ 1\ by (intro mono_intros, auto) consider "um < uM" | "um = uM" | "um > uM" by linarith then have "((\t. um + sgn (uM - um) * (t - dist xm ym)) ` {dist xm ym..dist xm ym + \uM - um\}) \ {min um uM..max um uM}" by (cases, auto) also have "... \ {A..B}" using \um \ {A..B}\ \uM \ {A..B}\ by auto finally have middle: "((\t. um + sgn (uM - um) * (t - dist xm ym)) ` {dist xm ym..dist xm ym + \uM - um\}) \ {A..B}" by simp have "(2 * lambda)-lipschitz_on {0..L} excursion" proof (unfold L_def, rule lipschitz_on_closed_Union[of "{{0..dist xm ym}, {dist xm ym..dist xm ym + abs(uM - um)}, {dist xm ym + abs(uM - um)..dist xm ym + abs(uM - um) + dist yM xM}}" _ "\ i. i"], auto) show "lambda \ 0" using C by auto have *: "1-lipschitz_on {0..dist xm ym} (geodesic_segment_param {xm--ym} xm)" by (rule isometry_on_lipschitz, simp) have **: "1-lipschitz_on {0..dist xm ym} excursion" using lipschitz_on_transform[OF * E1] by simp show "(2 * lambda)-lipschitz_on {0..dist xm ym} excursion" apply (rule lipschitz_on_mono[OF **]) using C by auto have *: "(1*(1+0))-lipschitz_on {dist xm ym + \uM - um\..dist xm ym + \uM - um\ + dist yM xM} ((geodesic_segment_param {yM--xM} yM) o (\t. t - (dist xm ym + abs (uM -um))))" by (intro lipschitz_intros, rule isometry_on_lipschitz, auto) have **: "(1*(1+0))-lipschitz_on {dist xm ym + \uM - um\..dist xm ym + \uM - um\ + dist yM xM} excursion" apply (rule lipschitz_on_transform[OF *]) using E3 unfolding comp_def by (auto simp add: algebra_simps) show "(2 * lambda)-lipschitz_on {dist xm ym + \uM - um\..dist xm ym + \uM - um\ + dist yM xM} excursion" apply (rule lipschitz_on_mono[OF **]) using C by auto have **: "((2 * lambda) * (0 + abs(sgn (uM - um)) * (1 + 0)))-lipschitz_on {dist xm ym..dist xm ym + abs(uM - um)} (d o (\t. um + sgn(uM-um) * (t - dist xm ym)))" apply (intro lipschitz_intros, rule lipschitz_on_subset[OF _ middle]) using \(2 * lambda)-lipschitz_on {A..B} d\ by simp have ***: "(2 * lambda)-lipschitz_on {dist xm ym..dist xm ym + abs(uM - um)} (d o (\t. um + sgn(uM-um) * (t - dist xm ym)))" apply (rule lipschitz_on_mono[OF **]) using C by auto show "(2 * lambda)-lipschitz_on {dist xm ym..dist xm ym + abs(uM - um)} excursion" apply (rule lipschitz_on_transform[OF ***]) using E2 by auto qed have *: "dist x z \ D" if z: "z \ excursion`{0..L}" for z proof - obtain tz where tz: "z = excursion tz" "tz \ {0..dist xm ym + abs(uM - um) + dist yM xM}" using z L_def by auto consider "tz \ {0..dist xm ym}" | "tz \ {dist xm ym<..dist xm ym + abs(uM - um)}" | "tz \ {dist xm ym + abs(uM - um)<..dist xm ym + abs(uM - um) + dist yM xM}" using tz by force then show ?thesis proof (cases) case 1 then have "z \ {xm--ym}" unfolding tz(1) excursion_def by auto then show ?thesis using D1 by auto next case 3 then have "z \ {yM--xM}" unfolding tz(1) excursion_def using tz(2) by auto then show ?thesis using D3 by (simp add: some_geodesic_commute) next case 2 then have "z \ d`{A..B}" unfolding tz(1) excursion_def using middle by force then show ?thesis unfolding D_def by (simp add: infdist_le) qed qed text \Now comes the main point: the excursion is always at distance at least $D$ of $x$, but this distance is also bounded by the log of its length, i.e., essentially $\log D$. To have an efficient estimate, we use a rescaled version, to get rid of one term on the right hand side.\ have "1 * 1 * 1 * (1 + 0/1) \ 512 * lambda * lambda * (1+C/D)" apply (intro mono_intros) using \lambda \ 1\ \D \ 1\ \C \ 0\ by auto then have "ln (512 * lambda * lambda * (1+C/D)) \ 0" apply (subst ln_ge_zero_iff) by auto define a where "a = 64 * lambda/D" have "a > 0" unfolding a_def using \D \ 1\ \lambda \ 1\ by auto have "D \ infdist x (excursion`{0..L})" unfolding infdist_def apply auto apply (rule cInf_greatest) using * by auto also have "... \ (4/ln 2) * deltaG(TYPE('a)) * max 0 (ln (a * (L-0))) + (2 * lambda) / a" proof (rule lipschitz_path_close_to_geodesic'[of _ _ _ _ "geodesic_subsegment {c A--c B} (c A) tm tM"]) show "(2 * lambda)-lipschitz_on {0..L} excursion" by fact have *: "geodesic_subsegment {c A--c B} (c A) tm tM = geodesic_segment_param {c A--c B} (c A) ` {tm..tM} " apply (rule geodesic_subsegment(1)[of _ _ "c B"]) using \tm \ {0..dist (c A) (c B)}\ \tM \ {0..dist (c A) (c B)}\ \tm \ tM\ by auto show "x \ geodesic_subsegment {c A--c B} (c A) tm tM" unfolding * unfolding x_param tm_def tM_def using \tx \ {0..dist (c A) (c B)}\ \0 \ D\ by simp show "geodesic_segment_between (geodesic_subsegment {c A--c B} (c A) tm tM) (excursion 0) (excursion L)" unfolding E0 EL xm_def xM_def apply (rule geodesic_subsegment[of _ _ "c B"]) using \tm \ {0..dist (c A) (c B)}\ \tM \ {0..dist (c A) (c B)}\ \tm \ tM\ by auto qed (fact) also have "... = (4/ln 2) * deltaG(TYPE('a)) * max 0 (ln (a *L)) + D/32" unfolding a_def using \D \ 1\ \lambda \ 1\ by (simp add: algebra_simps) finally have "(31 * ln 2 / 128) * D \ deltaG(TYPE('a)) * max 0 (ln (a * L))" by (auto simp add: algebra_simps divide_simps) also have "... \ deltaG(TYPE('a)) * max 0 (ln ((64 * lambda/D) * (lambda * (8 * D + 4 * C))))" unfolding a_def apply (intro mono_intros) using L_bound \L > 0\ \lambda \ 1\ \D \ 1\ by auto also have "... \ deltaG(TYPE('a)) * max 0 (ln ((64 * lambda/D) * (lambda * (8 * D + 8 * C))))" apply (intro mono_intros) using L_bound \L > 0\ \lambda \ 1\ \D \ 1\ \C \ 0\ by auto also have "... = deltaG(TYPE('a)) * max 0 (ln (512 * lambda * lambda * (1+C/D)))" using \D \ 1\ by (auto simp add: algebra_simps) also have "... = deltaG(TYPE('a)) * ln (512 * lambda * lambda * (1+C/D))" using \ln (512 * lambda * lambda * (1+C/D)) \ 0\ by auto also have "... \ deltaG(TYPE('a)) * ln (512 * lambda * lambda * (1+C/1))" apply (intro mono_intros) using \lambda \ 1\ \C \ 0\ \D \ 1\ by (auto simp add: divide_simps mult_ge1_mono(1)) text \We have obtained a bound on $D$, of the form $D \leq M \delta \ln(M \lambda^2(1+C))$. This is a nice bound, but we tweak it a little bit to obtain something more manageable, without the logarithm.\ also have "... = deltaG(TYPE('a)) * (ln 512 + 2 * ln lambda + ln (1+C))" apply (subst ln2mult) using \C \ 0\ \lambda \ 1\ apply simp apply (subst ln_mult[symmetric]) apply simp using \C \ 0\ \lambda \ 1\ apply simp apply (subst ln_mult[symmetric]) using \C \ 0\ \lambda \ 1\ by auto also have "... = (deltaG(TYPE('a)) * 1) * ln 512 + 2 * (deltaG(TYPE('a)) * ln lambda) + (deltaG(TYPE('a)) * ln (1+C))" by (auto simp add: algebra_simps) text \For each term, of the form $\delta \ln c$, we bound it by $(\delta^2 + (\ln c)^2)/2$, and then bound $(\ln c)^2$ by $2c-2$. In fact, to get coefficients of the same order of magnitude on $\delta^2$ and $\lambda$, we tweak a little bit the inequality for the last two terms, using rather $uv \leq (u^2/2 + 2v^2)/2$. We also bound $\ln(32)$ by a good approximation $16/3$.\ also have "... \ (deltaG(TYPE('a))^2/2 + 1^2/2) * (25/4) + 2 * ((1/2) * deltaG(TYPE('a))^2/2 + 2 * (ln lambda)^2 / 2) + ((1/2) * deltaG(TYPE('a))^2/2 + 2 * (ln (1+C))^2 / 2)" by (intro mono_intros, auto, approximation 10) also have "... = (31/8) * deltaG(TYPE('a))^2 + 25/8 + 2 * (ln lambda)^2 + (ln (1+C))^2" by (auto simp add: algebra_simps) also have "... \ 4 * deltaG(TYPE('a))^2 + 4 + 2 * (2 * lambda - 2) + (2 * (1+C) - 2)" apply (intro mono_intros) using \C \ 0\ \lambda \ 1\ by auto also have "... \ 4 * deltaG(TYPE('a))^2 + 4 * lambda + 2 * C" by auto finally have "D \ (128 / (31 * ln 2)) * (4 * deltaG(TYPE('a))^2 + 4 * lambda + 2 * C)" by (auto simp add: divide_simps algebra_simps) also have "... \ 6 * (4 * deltaG(TYPE('a))^2 + 4 * lambda + 2 * C)" apply (intro mono_intros, approximation 10) using \lambda \ 1\ \C \ 0\ by auto also have "... \ 24 * deltaG(TYPE('a))^2 + 24 * lambda + 12 * C" using \lambda \ 1\ \C \ 0\ by auto finally show ?thesis by simp qed define D0 where "D0 = 24 * lambda + 12 * C + 24 * deltaG(TYPE('a))^2" have first_step: "infdist y (d`{A..B}) \ D0" if "y \ {c A--c B}" for y using x(2)[OF that] D_bound unfolding D0_def D_def by auto have "1 * 1 + 4 * 0 + 24 * 0 \ D0" unfolding D0_def apply (intro mono_intros) using C delta_nonneg by auto then have "D0 > 0" by simp text \This is the end of the first step, i.e., showing that $[c(A), c(B)]$ is included in the neighborhood of size $D0$ of the quasi-geodesic.\ text \Now, we start the second step: we show that the quasi-geodesic is included in the neighborhood of size $D1$ of the geodesic, where $D1 \geq D0$ is the constant defined below. The argument goes as follows. Assume that a point $y$ on the quasi-geodesic is at distance $ > D0$ of the geodesic. Consider the last point $y_m$ before $y$ which is at distance $D0$ of the geodesic, and the first point $y_M$ after $y$ likewise. On $(y_m, y_M)$, one is always at distance $ > D0$ of the geodesic. However, by the first step, the geodesic is covered by the balls of radius $D0$ centered at points on the quasi-geodesic -- and only the points before $y_m$ or after $y_M$ can be used. Let $K_m$ be the points on the geodesics that are at distance $\leq D0$ of a point on the quasi-geodesic before $y_m$, and likewise define $K_M$. These are two closed subsets of the geodesic. By connectedness, they have to intersect. This implies that some points before $y_m$ and after $y_M$ are at distance at most $2D0$. Since we are dealing with a quasi-geodesic, this gives a bound on the distance between $y_m$ and $y_M$, and therefore a bound between $y$ and the geodesic, as desired.\ define D1 where "D1 = lambda * lambda * (72 * lambda + 44 * C + 72 * deltaG(TYPE('a))^2)" have "1 * 1 * (24 * lambda + 12 * C + 24 * deltaG(TYPE('a))^2) \ lambda * lambda * (72 * lambda + 44 * C + 72 * deltaG(TYPE('a))^2)" apply (intro mono_intros) using C by auto then have "D0 \ D1" unfolding D0_def D1_def by auto have second_step: "infdist y {c A--c B} \ D1" if "y \ d`{A..B}" for y proof (cases "infdist y {c A--c B} \ D0") case True then show ?thesis using \D0 \ D1\ by auto next case False obtain ty where "ty \ {A..B}" "y = d ty" using \y \ d`{A..B}\ by auto define tm where "tm = Sup ((\t. infdist (d t) {c A--c B})-`{..D0} \ {A..ty})" have tm: "tm \ (\t. infdist (d t) {c A--c B})-`{..D0} \ {A..ty}" unfolding tm_def proof (rule closed_contains_Sup) show "closed ((\t. infdist (d t) {c A--c B})-`{..D0} \ {A..ty})" apply (rule closed_vimage_Int, auto, intro continuous_intros) apply (rule continuous_on_subset[OF d(1)]) using \ty \ {A..B}\ by auto have "A \ (\t. infdist (d t) {c A--c B})-`{..D0} \ {A..ty}" using \D0 > 0\ \ty \ {A..B}\ by (auto simp add: \d A = c A\) then show "(\t. infdist (d t) {c A--c B})-`{..D0} \ {A..ty} \ {}" by auto show "bdd_above ((\t. infdist (d t) {c A--c B}) -` {..D0} \ {A..ty})" by auto qed have *: "infdist (d t) {c A--c B} > D0" if "t \ {tm<..ty}" for t proof (rule ccontr) assume "\(infdist (d t) {c A--c B} > D0)" then have *: "t \ (\t. infdist (d t) {c A--c B})-`{..D0} \ {A..ty}" using that tm by auto have "t \ tm" unfolding tm_def apply (rule cSup_upper) using * by auto then show False using that by auto qed define tM where "tM = Inf ((\t. infdist (d t) {c A--c B})-`{..D0} \ {ty..B})" have tM: "tM \ (\t. infdist (d t) {c A--c B})-`{..D0} \ {ty..B}" unfolding tM_def proof (rule closed_contains_Inf) show "closed ((\t. infdist (d t) {c A--c B})-`{..D0} \ {ty..B})" apply (rule closed_vimage_Int, auto, intro continuous_intros) apply (rule continuous_on_subset[OF d(1)]) using \ty \ {A..B}\ by auto have "B \ (\t. infdist (d t) {c A--c B})-`{..D0} \ {ty..B}" using \D0 > 0\ \ty \ {A..B}\ by (auto simp add: \d B = c B\) then show "(\t. infdist (d t) {c A--c B})-`{..D0} \ {ty..B} \ {}" by auto show "bdd_below ((\t. infdist (d t) {c A--c B}) -` {..D0} \ {ty..B})" by auto qed have "infdist (d t) {c A--c B} > D0" if "t \ {ty..(infdist (d t) {c A--c B} > D0)" then have *: "t \ (\t. infdist (d t) {c A--c B})-`{..D0} \ {ty..B}" using that tM by auto have "t \ tM" unfolding tM_def apply (rule cInf_lower) using * by auto then show False using that by auto qed then have lower_tm_tM: "infdist (d t) {c A--c B} > D0" if "t \ {tm<.. ty", auto) define Km where "Km = (\z \ d`{A..tm}. cball z D0)" define KM where "KM = (\z \ d`{tM..B}. cball z D0)" have "{c A--c B} \ Km \ KM" proof fix x assume "x \ {c A--c B}" have "\z \ d`{A..B}. infdist x (d`{A..B}) = dist x z" apply (rule infdist_proper_attained[OF proper_of_compact], rule compact_continuous_image[OF \continuous_on {A..B} d\]) using that by auto then obtain tx where "tx \ {A..B}" "infdist x (d`{A..B}) = dist x (d tx)" by blast then have "dist x (d tx) \ D0" using first_step[OF \x \ {c A--c B}\] by auto then have "x \ cball (d tx) D0" by (auto simp add: metric_space_class.dist_commute) consider "tx \ {A..tm}" | "tx \ {tm<.. {tM..B}" using \tx \ {A..B}\ by fastforce then show "x \ Km \ KM" proof (cases) case 1 then have "x \ Km" unfolding Km_def using \x \ cball (d tx) D0\ by auto then show ?thesis by simp next case 3 then have "x \ KM" unfolding KM_def using \x \ cball (d tx) D0\ by auto then show ?thesis by simp next case 2 have "infdist (d tx) {c A--c B} \ dist (d tx) x" using \x \ {c A--c B}\ by (rule infdist_le) also have "... \ D0" using \x \ cball (d tx) D0\ by auto finally have False using lower_tm_tM[OF 2] by simp then show ?thesis by simp qed qed then have *: "{c A--c B} = (Km \ {c A--c B}) \ (KM \ {c A--c B})" by auto have "(Km \ {c A--c B}) \ (KM \ {c A--c B}) \ {}" proof (rule connected_as_closed_union[OF _ *]) have "closed Km" unfolding Km_def apply (rule compact_has_closed_thickening) apply (rule compact_continuous_image) apply (rule continuous_on_subset[OF \continuous_on {A..B} d\]) using tm \ty \ {A..B}\ by auto then show "closed (Km \ {c A--c B})" by (rule topological_space_class.closed_Int, auto) have "closed KM" unfolding KM_def apply (rule compact_has_closed_thickening) apply (rule compact_continuous_image) apply (rule continuous_on_subset[OF \continuous_on {A..B} d\]) using tM \ty \ {A..B}\ by auto then show "closed (KM \ {c A--c B})" by (rule topological_space_class.closed_Int, auto) show "connected {c A--c B}" by simp have "c A \ Km \ {c A--c B}" apply auto unfolding Km_def using tm \d A = c A\ \D0 > 0\ by (auto) (rule bexI[of _ A], auto) then show "Km \ {c A--c B} \ {}" by auto have "c B \ KM \ {c A--c B}" apply auto unfolding KM_def using tM \d B = c B\ \D0 > 0\ by (auto) (rule bexI[of _ B], auto) then show "KM \ {c A--c B} \ {}" by auto qed then obtain w where "w \ {c A--c B}" "w \ Km" "w \ KM" by auto then obtain twm twM where tw: "twm \ {A..tm}" "w \ cball (d twm) D0" "twM \ {tM..B}" "w \ cball (d twM) D0" unfolding Km_def KM_def by auto have "(1/lambda) * dist twm twM - (4*C) \ dist (d twm) (d twM)" apply (rule quasi_isometry_onD(2)[OF d(5)]) using tw tm tM by auto also have "... \ dist (d twm) w + dist w (d twM)" by (rule metric_space_class.dist_triangle) also have "... \ 2 * D0" using tw by (auto simp add: metric_space_class.dist_commute) finally have "dist twm twM \ lambda * (4*C + 2*D0)" using C by (auto simp add: divide_simps algebra_simps) then have *: "dist twm ty \ lambda * (4*C + 2*D0)" using tw tm tM dist_real_def by auto have "dist (d ty) w \ dist (d ty) (d twm) + dist (d twm) w" by (rule metric_space_class.dist_triangle) also have "... \ (lambda * dist ty twm + (4*C)) + D0" apply (intro add_mono, rule quasi_isometry_onD(1)[OF d(5)]) using tw tm tM by auto also have "... \ (lambda * (lambda * (4*C + 2*D0))) + (4*C) + D0" apply (intro mono_intros) using C * by (auto simp add: metric_space_class.dist_commute) also have "... = lambda * lambda * (4*C + 2*D0) + 1 * 1 * (4 * C) + 1 * 1 * D0" by simp also have "... \ lambda * lambda * (4*C + 2*D0) + lambda * lambda * (4 * C) + lambda * lambda * D0" apply (intro mono_intros) using C * \D0 > 0\ by auto also have "... = lambda * lambda * (8 * C + 3 * D0)" by (auto simp add: algebra_simps) also have "... = lambda * lambda * (44 * C + 72 * lambda + 72 * deltaG(TYPE('a))^2)" unfolding D0_def by auto finally have "dist y w \ D1" unfolding D1_def \y = d ty\ by (auto simp add: algebra_simps) then show "infdist y {c A--c B} \ D1" using infdist_le[OF \w \ {c A--c B}\, of y] by auto qed text \This concludes the second step.\ text \Putting the two steps together, we deduce that the Hausdorff distance between the geodesic and the quasi-geodesic is bounded by $D1$. A bound between the geodesic and the original (untamed) quasi-geodesic follows.\ have a: "hausdorff_distance (d`{A..B}) {c A--c B} \ D1" proof (rule hausdorff_distanceI) show "D1 \ 0" unfolding D1_def using C delta_nonneg by auto fix x assume "x \ d ` {A..B}" then show "infdist x {c A--c B} \ D1" using second_step by auto next fix x assume "x \ {c A--c B}" then show "infdist x (d`{A..B}) \ D1" using first_step \D0 \ D1\ by force qed have "hausdorff_distance (c`{A..B}) {c A--c B} \ hausdorff_distance (c`{A..B}) (d`{A..B}) + hausdorff_distance (d`{A..B}) {c A--c B}" apply (rule hausdorff_distance_triangle) using \A \ {A..B}\ apply blast by (rule quasi_isometry_on_bounded[OF d(5)], auto) also have "... \ D1 + 2*C" using a d by auto also have "... = lambda * lambda * (72 * lambda + 44 * C + 72 * deltaG(TYPE('a))^2) + 1 * 1 * (2 * C)" unfolding D1_def by auto also have "... \ lambda * lambda * (72 * lambda + 44 * C + 72 * deltaG(TYPE('a))^2) + lambda * lambda * (28 * C)" apply (intro mono_intros) using C delta_nonneg by auto also have "... = 72 * lambda^2 * (lambda + C + deltaG(TYPE('a))^2)" by (auto simp add: algebra_simps power2_eq_square) finally show ?thesis by (auto simp add: algebra_simps) qed qed end (*of theory Morse_Gromov_Theorem*) diff --git a/thys/KD_Tree/KDTree.thy b/thys/KD_Tree/KDTree.thy --- a/thys/KD_Tree/KDTree.thy +++ b/thys/KD_Tree/KDTree.thy @@ -1,441 +1,441 @@ (* File: KDTree.thy Author: Martin Rau, TU München *) section "Definition of the \k\-d Tree" theory KDTree imports Complex_Main begin text \ A \k\-d tree is a space-partitioning data structure for organizing points in a \k\-dimensional space. In principle the \k\-d tree is a binary tree. The leafs hold the \k\-dimensional points and the nodes contain left and right subtrees as well as a discriminator \s\ at a particular axis \a\. Every node divides the space into two parts by splitting along a hyperplane. Consider a node \n\ with associated discriminator \s\ at axis \a\. All points in the left subtree must have a value at axis \a\ that is less than or equal to \s\ and all points in the right subtree must have a value at axis \a\ that is greater than or equal to \s\. Deviations from the papers: The chosen tree representation is taken from @{cite "DBLP:journals/toms/FriedmanBF77"} with one minor adjustment. Originally the leafs hold buckets of points of size \b\. This representation fixes the bucket size to \b = 1\, a single point per Leaf. This is only a minor adjustment since the paper proves that \b = 1\ is the optimal bucket size for minimizing the runtime of the nearest neighbor algorithm @{cite "DBLP:journals/toms/FriedmanBF77"}, only simplifies building the optimized \k\-d trees @{cite "DBLP:journals/toms/FriedmanBF77"} and has little influence on the search algorithm @{cite "DBLP:journals/cacm/Bentley75"}. \ type_synonym point = "real list" type_synonym axis = nat type_synonym dimension = nat type_synonym discriminator = real datatype kdt = Leaf point | Node axis discriminator kdt kdt subsection "Definition of the \k\-d Tree Invariant and Related Functions" definition dim :: "point \ nat" where "dim p = length p" declare dim_def[simp] fun set_kdt :: "kdt \ point set" where "set_kdt (Leaf p) = {p}" | "set_kdt (Node _ _ l r) = set_kdt l \ set_kdt r" fun invar :: "dimension \ kdt \ bool" where "invar k (Leaf p) \ dim p = k" | "invar k (Node a s l r) \ (\p \ set_kdt l. p!a \ s) \ (\p \ set_kdt r. s \ p!a) \ invar k l \ invar k r \ a < k \ set_kdt l \ set_kdt r = {}" fun size_kdt :: "kdt \ nat" where "size_kdt (Leaf _) = 1" | "size_kdt (Node _ _ l r) = size_kdt l + size_kdt r" fun height :: "kdt \ nat" where "height (Leaf _) = 0" | "height (Node _ _ l r) = max (height l) (height r) + 1" fun min_height :: "kdt \ nat" where "min_height (Leaf _) = 0" | "min_height (Node _ _ l r) = min (min_height l) (min_height r) + 1" definition balanced :: "kdt \ bool" where "balanced kdt \ height kdt - min_height kdt \ 1" fun complete :: "kdt \ bool" where "complete (Leaf _) = True" | "complete (Node _ _ l r) \ complete l \ complete r \ height l = height r" lemma invar_l: "invar k (Node a s l r) \ invar k l" by simp lemma invar_r: "invar k (Node a s l r) \ invar k r" by simp lemma invar_axis_lt_d: "invar k (Node a s l r) \ a < k" by simp lemma invar_dim: "invar k kdt \ \p \ set_kdt kdt. dim p = k" by (induction kdt) auto lemma invar_l_le_a: "invar k (Node a s l r) \ \p \ set_kdt l. p!a \ s" by simp lemma invar_r_ge_a: "invar k (Node a s l r) \ \p \ set_kdt r. s \ p!a" by simp lemma invar_distinct: "invar k kdt \ kdt = Node a s l r \ set_kdt l \ set_kdt r = {}" by fastforce lemma invar_set: "set_kdt (Node a s l r) = set_kdt l \ set_kdt r" by simp subsection "Lemmas adapted from \HOL-Library.Tree\ to \k\-d Tree" lemma size_ge0[simp]: "0 < size_kdt kdt" by (induction kdt) auto lemma eq_size_1[simp]: "size_kdt kdt = 1 \ (\p. kdt = Leaf p)" apply (induction kdt) apply (auto) using size_ge0 nat_less_le apply blast+ done lemma eq_1_size[simp]: "1 = size_kdt kdt \ (\p. kdt = Leaf p)" using eq_size_1 by metis lemma neq_Leaf_iff: "(\p. kdt = Leaf p) = (\a s l r. kdt = Node a s l r)" by (cases kdt) auto lemma eq_height_0[simp]: "height kdt = 0 \ (\p. kdt = Leaf p)" by (cases kdt) auto lemma eq_0_height[simp]: "0 = height kdt \ (\p. kdt = Leaf p)" by (cases kdt) auto lemma eq_min_height_0[simp]: "min_height kdt = 0 \ (\p. kdt = Leaf p)" by (cases kdt) auto lemma eq_0_min_height[simp]: "0 = min_height kdt \ (\p. kdt = Leaf p)" by (cases kdt) auto lemma size_height: "size_kdt kdt \ 2 ^ height kdt" proof(induction kdt) case (Node a s l r) show ?case proof (cases "height l \ height r") case True have "size_kdt (Node a s l r) = size_kdt l + size_kdt r" by simp also have "\ \ 2 ^ height l + 2 ^ height r" using Node.IH by arith also have "\ \ 2 ^ height r + 2 ^ height r" using True by simp also have "\ = 2 ^ height (Node a s l r)" using True by (auto simp: max_def mult_2) finally show ?thesis . next case False have "size_kdt (Node a s l r) = size_kdt l + size_kdt r" by simp also have "\ \ 2 ^ height l + 2 ^ height r" using Node.IH by arith also have "\ \ 2 ^ height l + 2 ^ height l" using False by simp finally show ?thesis using False by (auto simp: max_def mult_2) qed qed simp lemma min_height_le_height: "min_height kdt \ height kdt" by (induction kdt) auto lemma min_height_size: "2 ^ min_height kdt \ size_kdt kdt" proof(induction kdt) case (Node a s l r) have "(2::nat) ^ min_height (Node a s l r) \ 2 ^ min_height l + 2 ^ min_height r" by (simp add: min_def) also have "\ \ size_kdt (Node a s l r)" using Node.IH by simp finally show ?case . qed simp lemma complete_iff_height: "complete kdt \ (min_height kdt = height kdt)" apply (induction kdt) apply simp apply (simp add: min_def max_def) by (metis le_antisym le_trans min_height_le_height) lemma size_if_complete: "complete kdt \ size_kdt kdt = 2 ^ height kdt" by (induction kdt) auto lemma complete_if_size_height: "size_kdt kdt = 2 ^ height kdt \ complete kdt" proof (induction "height kdt" arbitrary: kdt) case 0 thus ?case by auto next case (Suc h) hence "\p. kdt = Leaf p" by auto then obtain a s l r where [simp]: "kdt = Node a s l r" using neq_Leaf_iff by auto have 1: "height l \ h" and 2: "height r \ h" using Suc(2) by(auto) have 3: "\ height l < h" proof assume 0: "height l < h" have "size_kdt kdt = size_kdt l + size_kdt r" by simp also have "\ \ 2 ^ height l + 2 ^ height r" using size_height[of l] size_height[of r] by arith also have " \ < 2 ^ h + 2 ^ height r" using 0 by (simp) also have " \ \ 2 ^ h + 2 ^ h" using 2 by (simp) also have "\ = 2 ^ (Suc h)" by (simp) also have "\ = size_kdt kdt" using Suc(2,3) by simp finally have "size_kdt kdt < size_kdt kdt" . thus False by (simp) qed have 4: "\ height r < h" proof assume 0: "height r < h" have "size_kdt kdt = size_kdt l + size_kdt r" by simp also have "\ \ 2 ^ height l + 2 ^ height r" using size_height[of l] size_height[of r] by arith also have " \ < 2 ^ height l + 2 ^ h" using 0 by (simp) also have " \ \ 2 ^ h + 2 ^ h" using 1 by (simp) also have "\ = 2 ^ (Suc h)" by (simp) also have "\ = size_kdt kdt" using Suc(2,3) by simp finally have "size_kdt kdt < size_kdt kdt" . thus False by (simp) qed from 1 2 3 4 have *: "height l = h" "height r = h" by linarith+ hence "size_kdt l = 2 ^ height l" "size_kdt r = 2 ^ height r" using Suc(3) size_height[of l] size_height[of r] by (auto) with * Suc(1) show ?case by simp qed lemma complete_if_size_min_height: "size_kdt kdt = 2 ^ min_height kdt \ complete kdt" proof (induct "min_height kdt" arbitrary: kdt) case 0 thus ?case by auto next case (Suc h) hence "\p. kdt = Leaf p" by auto then obtain a s l r where [simp]: "kdt = Node a s l r" using neq_Leaf_iff by auto have 1: "h \ min_height l" and 2: "h \ min_height r" using Suc(2) by (auto) have 3: "\ h < min_height l" proof assume 0: "h < min_height l" have "size_kdt kdt = size_kdt l + size_kdt r" by simp also note min_height_size[of l] also(xtrans) note min_height_size[of r] also(xtrans) have "(2::nat) ^ min_height l > 2 ^ h" using 0 by (simp add: diff_less_mono) also(xtrans) have "(2::nat) ^ min_height r \ 2 ^ h" using 2 by simp also(xtrans) have "(2::nat) ^ h + 2 ^ h = 2 ^ (Suc h)" by (simp) also have "\ = size_kdt kdt" using Suc(2,3) by simp finally show False by (simp add: diff_le_mono) qed have 4: "\ h < min_height r" proof assume 0: "h < min_height r" have "size_kdt kdt = size_kdt l + size_kdt r" by simp also note min_height_size[of l] also(xtrans) note min_height_size[of r] also(xtrans) have "(2::nat) ^ min_height r > 2 ^ h" using 0 by (simp add: diff_less_mono) also(xtrans) have "(2::nat) ^ min_height l \ 2 ^ h" using 1 by simp also(xtrans) have "(2::nat) ^ h + 2 ^ h = 2 ^ (Suc h)" by (simp) also have "\ = size_kdt kdt" using Suc(2,3) by simp finally show False by (simp add: diff_le_mono) qed from 1 2 3 4 have *: "min_height l = h" "min_height r = h" by linarith+ hence "size_kdt l = 2 ^ min_height l" "size_kdt r = 2 ^ min_height r" using Suc(3) min_height_size[of l] min_height_size[of r] by (auto) with * Suc(1) show ?case by (simp add: complete_iff_height) qed lemma complete_iff_size: "complete kdt \ size_kdt kdt = 2 ^ height kdt" using complete_if_size_height size_if_complete by blast lemma size_height_if_incomplete: "\ complete kdt \ size_kdt kdt < 2 ^ height kdt" by (meson antisym_conv complete_iff_size not_le size_height) lemma min_height_size_if_incomplete: "\ complete kdt \ 2 ^ min_height kdt < size_kdt kdt" by (metis complete_if_size_min_height le_less min_height_size) lemma balanced_subtreeL: "balanced (Node a s l r) \ balanced l" by (simp add: balanced_def) lemma balanced_subtreeR: "balanced (Node a s l r) \ balanced r" by (simp add: balanced_def) lemma balanced_optimal: assumes "balanced kdt" "size_kdt kdt \ size_kdt kdt'" shows "height kdt \ height kdt'" proof (cases "complete kdt") case True have "(2::nat) ^ height kdt \ 2 ^ height kdt'" proof - have "2 ^ height kdt = size_kdt kdt" using True by (simp add: complete_iff_height size_if_complete) also have "\ \ size_kdt kdt'" using assms(2) by simp also have "\ \ 2 ^ height kdt'" by (rule size_height) finally show ?thesis . qed thus ?thesis by (simp) next case False have "(2::nat) ^ min_height kdt < 2 ^ height kdt'" proof - have "(2::nat) ^ min_height kdt < size_kdt kdt" by(rule min_height_size_if_incomplete[OF False]) also have "\ \ size_kdt kdt'" using assms(2) by simp also have "\ \ 2 ^ height kdt'" by(rule size_height) finally have "(2::nat) ^ min_height kdt < (2::nat) ^ height kdt'" . thus ?thesis . qed hence *: "min_height kdt < height kdt'" by simp have "min_height kdt + 1 = height kdt" using min_height_le_height[of kdt] assms(1) False by (simp add: complete_iff_height balanced_def) with * show ?thesis by arith qed subsection "Lemmas adapted from \HOL-Library.Tree_Real\ to \k\-d Tree" lemma size_height_log: "log 2 (size_kdt kdt) \ height kdt" by (simp add: log2_of_power_le size_height) lemma min_height_size_log: "min_height kdt \ log 2 (size_kdt kdt)" by (simp add: le_log2_of_power min_height_size) lemma size_log_if_complete: "complete kdt \ height kdt = log 2 (size_kdt kdt)" - by (simp add: log2_of_power_eq size_if_complete) + by (simp add: size_if_complete) lemma min_height_size_log_if_incomplete: "\ complete kdt \ min_height kdt < log 2 (size_kdt kdt)" by (simp add: less_log2_of_power min_height_size_if_incomplete) lemma min_height_balanced: assumes "balanced kdt" shows "min_height kdt = nat(floor(log 2 (size_kdt kdt)))" proof cases assume *: "complete kdt" hence "size_kdt kdt = 2 ^ min_height kdt" by (simp add: complete_iff_height size_if_complete) from log2_of_power_eq[OF this] show ?thesis by linarith next assume *: "\ complete kdt" hence "height kdt = min_height kdt + 1" using assms min_height_le_height[of kdt] by(auto simp add: balanced_def complete_iff_height) hence "size_kdt kdt < 2 ^ (min_height kdt + 1)" by (metis * size_height_if_incomplete) hence "log 2 (size_kdt kdt) < min_height kdt + 1" using log2_of_power_less size_ge0 by blast thus ?thesis using min_height_size_log[of kdt] by linarith qed lemma height_balanced: assumes "balanced kdt" shows "height kdt = nat(ceiling(log 2 (size_kdt kdt)))" proof cases assume *: "complete kdt" hence "size_kdt kdt = 2 ^ height kdt" by (simp add: size_if_complete) from log2_of_power_eq[OF this] show ?thesis by linarith next assume *: "\ complete kdt" hence **: "height kdt = min_height kdt + 1" using assms min_height_le_height[of kdt] by(auto simp add: balanced_def complete_iff_height) hence "size_kdt kdt \ 2 ^ (min_height kdt + 1)" by (metis size_height) from log2_of_power_le[OF this size_ge0] min_height_size_log_if_incomplete[OF *] ** show ?thesis by linarith qed lemma balanced_Node_if_wbal1: assumes "balanced l" "balanced r" "size_kdt l = size_kdt r + 1" shows "balanced (Node a s l r)" proof - from assms(3) have [simp]: "size_kdt l = size_kdt r + 1" by simp have "nat \log 2 (1 + size_kdt r)\ \ nat \log 2 (size_kdt r)\" by(rule nat_mono[OF ceiling_mono]) simp hence 1: "height(Node a s l r) = nat \log 2 (1 + size_kdt r)\ + 1" using height_balanced[OF assms(1)] height_balanced[OF assms(2)] by (simp del: nat_ceiling_le_eq add: max_def) have "nat \log 2 (1 + size_kdt r)\ \ nat \log 2 (size_kdt r)\" by(rule nat_mono[OF floor_mono]) simp hence 2: "min_height(Node a s l r) = nat \log 2 (size_kdt r)\ + 1" using min_height_balanced[OF assms(1)] min_height_balanced[OF assms(2)] by (simp) have "size_kdt r \ 1" by (simp add: Suc_leI) then obtain i where i: "2 ^ i \ size_kdt r" "size_kdt r < 2 ^ (i + 1)" using ex_power_ivl1[of 2 "size_kdt r"] by auto hence i1: "2 ^ i < size_kdt r + 1" "size_kdt r + 1 \ 2 ^ (i + 1)" by auto from 1 2 floor_log_nat_eq_if[OF i] ceiling_log_nat_eq_if[OF i1] show ?thesis by(simp add:balanced_def) qed lemma balanced_sym: "balanced (Node a s l r) \ balanced (Node a' s' r l)" by (auto simp: balanced_def) lemma balanced_Node_if_wbal2: assumes "balanced l" "balanced r" "abs(int(size_kdt l) - int(size_kdt r)) \ 1" shows "balanced (Node a s l r)" proof - have "size_kdt l = size_kdt r \ (size_kdt l = size_kdt r + 1 \ size_kdt r = size_kdt l + 1)" (is "?A \ ?B") using assms(3) by linarith thus ?thesis proof assume "?A" thus ?thesis using assms(1,2) apply(simp add: balanced_def min_def max_def) by (metis assms(1,2) balanced_optimal le_antisym le_less) next assume "?B" thus ?thesis by (meson assms(1,2) balanced_sym balanced_Node_if_wbal1) qed qed end \ No newline at end of file diff --git a/thys/Sqrt_Babylonian/Sqrt_Babylonian_Auxiliary.thy b/thys/Sqrt_Babylonian/Sqrt_Babylonian_Auxiliary.thy --- a/thys/Sqrt_Babylonian/Sqrt_Babylonian_Auxiliary.thy +++ b/thys/Sqrt_Babylonian/Sqrt_Babylonian_Auxiliary.thy @@ -1,153 +1,136 @@ (* Title: Computing Square Roots using the Babylonian Method Author: René Thiemann Maintainer: René Thiemann License: LGPL *) (* Copyright 2009-2014 René Thiemann This file is part of IsaFoR/CeTA. IsaFoR/CeTA is free software: you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation, either version 3 of the License, or (at your option) any later version. IsaFoR/CeTA is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with IsaFoR/CeTA. If not, see . *) section \Auxiliary lemmas which might be moved into the Isabelle distribution.\ theory Sqrt_Babylonian_Auxiliary imports Complex_Main begin lemma mod_div_equality_int: "(n :: int) div x * x = n - n mod x" using div_mult_mod_eq[of n x] by arith -lemma log_pow_cancel[simp]: "a > 0 \ a \ 1 \ log a (a ^ b) = b" - by (metis monoid_mult_class.mult.right_neutral log_eq_one log_nat_power) - -lemma real_of_rat_floor[simp]: "floor (real_of_rat x) = floor x" - by (metis Ratreal_def real_floor_code) - -lemma abs_of_rat[simp]: "\real_of_rat x\ = real_of_rat \x\" -proof (cases "x \ 0") - case False - define y where "y = - x" - from False have y: "y \ 0" "x = - y" by (auto simp: y_def) - thus ?thesis by (auto simp: of_rat_minus) -qed auto - -lemma real_of_rat_ceiling[simp]: "ceiling (real_of_rat x) = ceiling x" - unfolding ceiling_def by (metis of_rat_minus real_of_rat_floor) - lemma div_is_floor_divide_rat: "n div y = \rat_of_int n / rat_of_int y\" unfolding Fract_of_int_quotient[symmetric] floor_Fract by simp lemma div_is_floor_divide_real: "n div y = \real_of_int n / of_int y\" unfolding div_is_floor_divide_rat[of n y] by (metis Ratreal_def of_rat_divide of_rat_of_int_eq real_floor_code) lemma floor_div_pos_int: fixes r :: "'a :: floor_ceiling" assumes n: "n > 0" shows "\r / of_int n\ = \r\ div n" (is "?l = ?r") proof - let ?of_int = "of_int :: int \ 'a" define rhs where "rhs = \r\ div n" let ?n = "?of_int n" define m where "m = \r\ mod n" let ?m = "?of_int m" from div_mult_mod_eq[of "floor r" n] have dm: "rhs * n + m = \r\" unfolding rhs_def m_def by simp have mn: "m < n" and m0: "m \ 0" using n m_def by auto define e where "e = r - ?of_int \r\" have e0: "e \ 0" unfolding e_def by (metis diff_self eq_iff floor_diff_of_int zero_le_floor) have e1: "e < 1" unfolding e_def by (metis diff_self dual_order.refl floor_diff_of_int floor_le_zero) have "r = ?of_int \r\ + e" unfolding e_def by simp also have "\r\ = rhs * n + m" using dm by simp finally have "r = ?of_int (rhs * n + m) + e" . hence "r / ?n = ?of_int (rhs * n) / ?n + (e + ?m) / ?n" using n by (simp add: field_simps) also have "?of_int (rhs * n) / ?n = ?of_int rhs" using n by auto finally have *: "r / ?of_int n = (e + ?of_int m) / ?of_int n + ?of_int rhs" by simp have "?l = rhs + floor ((e + ?m) / ?n)" unfolding * by simp also have "floor ((e + ?m) / ?n) = 0" proof (rule floor_unique) show "?of_int 0 \ (e + ?m) / ?n" using e0 m0 n by (metis add_increasing2 divide_nonneg_pos of_int_0 of_int_0_le_iff of_int_0_less_iff) show "(e + ?m) / ?n < ?of_int 0 + 1" proof (rule ccontr) from n have n': "?n > 0" "?n \ 0" by simp_all assume "\ ?thesis" hence "(e + ?m) / ?n \ 1" by auto from mult_right_mono[OF this n'(2)] have "?n \ e + ?m" using n'(1) by simp also have "?m \ ?n - 1" using mn by (metis of_int_1 of_int_diff of_int_le_iff zle_diff1_eq) finally have "?n \ e + ?n - 1" by auto with e1 show False by arith qed qed finally show ?thesis unfolding rhs_def by simp qed lemma floor_div_neg_int: fixes r :: "'a :: floor_ceiling" assumes n: "n < 0" shows "\r / of_int n\ = \r\ div n" proof - from n have n': "- n > 0" by auto have "\r / of_int n\ = \ - r / of_int (- n)\" using n by (metis floor_of_int floor_zero less_int_code(1) minus_divide_left minus_minus nonzero_minus_divide_right of_int_minus) also have "\ = \ - r \ div (- n)" by (rule floor_div_pos_int[OF n']) also have "\ = \ r \ div n" using n by (metis ceiling_def div_minus_right) finally show ?thesis . qed lemma divide_less_floor1: "n / y < of_int (floor (n / y)) + 1" by (metis floor_less_iff less_add_one of_int_1 of_int_add) context linordered_idom begin lemma sgn_int_pow_if [simp]: "sgn x ^ p = (if even p then 1 else sgn x)" if "x \ 0" using that by (induct p) simp_all lemma compare_pow_le_iff: "p > 0 \ (x :: 'a) \ 0 \ y \ 0 \ (x ^ p \ y ^ p) = (x \ y)" by (metis eq_iff linear power_eq_imp_eq_base power_mono) lemma compare_pow_less_iff: "p > 0 \ (x :: 'a) \ 0 \ y \ 0 \ (x ^ p < y ^ p) = (x < y)" by (metis power_less_imp_less_base power_strict_mono) end lemma quotient_of_int[simp]: "quotient_of (of_int i) = (i,1)" by (metis Rat.of_int_def quotient_of_int) lemma quotient_of_nat[simp]: "quotient_of (of_nat i) = (int i,1)" by (metis Rat.of_int_def Rat.quotient_of_int of_int_of_nat_eq) lemma square_lesseq_square: "\ x y. 0 \ (x :: 'a :: linordered_field) \ 0 \ y \ (x * x \ y * y) = (x \ y)" by (metis mult_mono mult_strict_mono' not_less) lemma square_less_square: "\ x y. 0 \ (x :: 'a :: linordered_field) \ 0 \ y \ (x * x < y * y) = (x < y)" by (metis mult_mono mult_strict_mono' not_less) lemma sqrt_sqrt[simp]: "x \ 0 \ sqrt x * sqrt x = x" by (metis real_sqrt_pow2 power2_eq_square) lemma abs_lesseq_square: "abs (x :: real) \ abs y \ x * x \ y * y" using square_lesseq_square[of "abs x" "abs y"] by auto end