diff --git a/src/HOL/Decision_Procs/Approximation.thy b/src/HOL/Decision_Procs/Approximation.thy --- a/src/HOL/Decision_Procs/Approximation.thy +++ b/src/HOL/Decision_Procs/Approximation.thy @@ -1,1701 +1,1232 @@ (* Author: Johannes Hoelzl, TU Muenchen Coercions removed by Dmitriy Traytel *) theory Approximation imports Complex_Main "HOL-Library.Code_Target_Numeral" Approximation_Bounds keywords "approximate" :: diag begin section "Implement floatarith" subsection "Define syntax and semantics" datatype floatarith = Add floatarith floatarith | Minus floatarith | Mult floatarith floatarith | Inverse floatarith | Cos floatarith | Arctan floatarith | Abs floatarith | Max floatarith floatarith | Min floatarith floatarith | Pi | Sqrt floatarith | Exp floatarith | Powr floatarith floatarith | Ln floatarith | Power floatarith nat | Floor floatarith | Var nat | Num float fun interpret_floatarith :: "floatarith \ real list \ real" where "interpret_floatarith (Add a b) vs = (interpret_floatarith a vs) + (interpret_floatarith b vs)" | "interpret_floatarith (Minus a) vs = - (interpret_floatarith a vs)" | "interpret_floatarith (Mult a b) vs = (interpret_floatarith a vs) * (interpret_floatarith b vs)" | "interpret_floatarith (Inverse a) vs = inverse (interpret_floatarith a vs)" | "interpret_floatarith (Cos a) vs = cos (interpret_floatarith a vs)" | "interpret_floatarith (Arctan a) vs = arctan (interpret_floatarith a vs)" | "interpret_floatarith (Min a b) vs = min (interpret_floatarith a vs) (interpret_floatarith b vs)" | "interpret_floatarith (Max a b) vs = max (interpret_floatarith a vs) (interpret_floatarith b vs)" | "interpret_floatarith (Abs a) vs = \interpret_floatarith a vs\" | "interpret_floatarith Pi vs = pi" | "interpret_floatarith (Sqrt a) vs = sqrt (interpret_floatarith a vs)" | "interpret_floatarith (Exp a) vs = exp (interpret_floatarith a vs)" | "interpret_floatarith (Powr a b) vs = interpret_floatarith a vs powr interpret_floatarith b vs" | "interpret_floatarith (Ln a) vs = ln (interpret_floatarith a vs)" | "interpret_floatarith (Power a n) vs = (interpret_floatarith a vs)^n" | "interpret_floatarith (Floor a) vs = floor (interpret_floatarith a vs)" | "interpret_floatarith (Num f) vs = f" | "interpret_floatarith (Var n) vs = vs ! n" lemma interpret_floatarith_divide: "interpret_floatarith (Mult a (Inverse b)) vs = (interpret_floatarith a vs) / (interpret_floatarith b vs)" unfolding divide_inverse interpret_floatarith.simps .. lemma interpret_floatarith_diff: "interpret_floatarith (Add a (Minus b)) vs = (interpret_floatarith a vs) - (interpret_floatarith b vs)" unfolding interpret_floatarith.simps by simp lemma interpret_floatarith_sin: "interpret_floatarith (Cos (Add (Mult Pi (Num (Float 1 (- 1)))) (Minus a))) vs = sin (interpret_floatarith a vs)" unfolding sin_cos_eq interpret_floatarith.simps interpret_floatarith_divide interpret_floatarith_diff by auto subsection "Implement approximation function" -fun lift_bin :: "(float * float) option \ (float * float) option \ (float \ float \ float \ float \ (float * float) option) \ (float * float) option" where -"lift_bin (Some (l1, u1)) (Some (l2, u2)) f = f l1 u1 l2 u2" | +fun lift_bin :: "(float interval) option \ (float interval) option \ (float interval \ float interval \ (float interval) option) \ (float interval) option" where +"lift_bin (Some ivl1) (Some ivl2) f = f ivl1 ivl2" | "lift_bin a b f = None" -fun lift_bin' :: "(float * float) option \ (float * float) option \ (float \ float \ float \ float \ (float * float)) \ (float * float) option" where -"lift_bin' (Some (l1, u1)) (Some (l2, u2)) f = Some (f l1 u1 l2 u2)" | +fun lift_bin' :: "(float interval) option \ (float interval) option \ (float interval \ float interval \ float interval) \ (float interval) option" where +"lift_bin' (Some ivl1) (Some ivl2) f = Some (f ivl1 ivl2)" | "lift_bin' a b f = None" -fun lift_un :: "(float * float) option \ (float \ float \ ((float option) * (float option))) \ (float * float) option" where -"lift_un (Some (l1, u1)) f = (case (f l1 u1) of (Some l, Some u) \ Some (l, u) - | t \ None)" | +fun lift_un :: "float interval option \ (float interval \ float interval option) \ float interval option" where +"lift_un (Some ivl) f = f ivl" | "lift_un b f = None" -fun lift_un' :: "(float * float) option \ (float \ float \ (float * float)) \ (float * float) option" where -"lift_un' (Some (l1, u1)) f = Some (f l1 u1)" | +lemma lift_un_eq:\ \TODO\ "lift_un x f = Option.bind x f" + by (cases x) auto + +fun lift_un' :: "(float interval) option \ (float interval \ (float interval)) \ (float interval) option" where +"lift_un' (Some ivl1) f = Some (f ivl1)" | "lift_un' b f = None" -definition bounded_by :: "real list \ (float \ float) option list \ bool" where - "bounded_by xs vs \ - (\ i < length vs. case vs ! i of None \ True - | Some (l, u) \ xs ! i \ { real_of_float l .. real_of_float u })" - +definition bounded_by :: "real list \ (float interval) option list \ bool" where + "bounded_by xs vs \ (\ i < length vs. case vs ! i of None \ True | Some ivl \ xs ! i \\<^sub>r ivl)" + lemma bounded_byE: assumes "bounded_by xs vs" shows "\ i. i < length vs \ case vs ! i of None \ True - | Some (l, u) \ xs ! i \ { real_of_float l .. real_of_float u }" - using assms bounded_by_def by blast + | Some ivl \ xs ! i \\<^sub>r ivl" + using assms + by (auto simp: bounded_by_def) lemma bounded_by_update: assumes "bounded_by xs vs" - and bnd: "xs ! i \ { real_of_float l .. real_of_float u }" - shows "bounded_by xs (vs[i := Some (l,u)])" -proof - - { - fix j - let ?vs = "vs[i := Some (l,u)]" - assume "j < length ?vs" - hence [simp]: "j < length vs" by simp - have "case ?vs ! j of None \ True | Some (l, u) \ xs ! j \ { real_of_float l .. real_of_float u }" - proof (cases "?vs ! j") - case (Some b) - thus ?thesis - proof (cases "i = j") - case True - thus ?thesis using \?vs ! j = Some b\ and bnd by auto - next - case False - thus ?thesis using \bounded_by xs vs\ unfolding bounded_by_def by auto - qed - qed auto - } - thus ?thesis unfolding bounded_by_def by auto -qed + and bnd: "xs ! i \\<^sub>r ivl" + shows "bounded_by xs (vs[i := Some ivl])" + using assms + by (cases "i < length vs") (auto simp: bounded_by_def nth_list_update split: option.splits) lemma bounded_by_None: "bounded_by xs (replicate (length xs) None)" unfolding bounded_by_def by auto -fun approx approx' :: "nat \ floatarith \ (float * float) option list \ (float * float) option" where -"approx' prec a bs = (case (approx prec a bs) of Some (l, u) \ Some (float_round_down prec l, float_round_up prec u) | None \ None)" | -"approx prec (Add a b) bs = - lift_bin' (approx' prec a bs) (approx' prec b bs) - (\ l1 u1 l2 u2. (float_plus_down prec l1 l2, float_plus_up prec u1 u2))" | -"approx prec (Minus a) bs = lift_un' (approx' prec a bs) (\ l u. (-u, -l))" | -"approx prec (Mult a b) bs = - lift_bin' (approx' prec a bs) (approx' prec b bs) (bnds_mult prec)" | -"approx prec (Inverse a) bs = lift_un (approx' prec a bs) (\ l u. if (0 < l \ u < 0) then (Some (float_divl prec 1 u), Some (float_divr prec 1 l)) else (None, None))" | -"approx prec (Cos a) bs = lift_un' (approx' prec a bs) (bnds_cos prec)" | -"approx prec Pi bs = Some (lb_pi prec, ub_pi prec)" | -"approx prec (Min a b) bs = lift_bin' (approx' prec a bs) (approx' prec b bs) (\ l1 u1 l2 u2. (min l1 l2, min u1 u2))" | -"approx prec (Max a b) bs = lift_bin' (approx' prec a bs) (approx' prec b bs) (\ l1 u1 l2 u2. (max l1 l2, max u1 u2))" | -"approx prec (Abs a) bs = lift_un' (approx' prec a bs) (\l u. (if l < 0 \ 0 < u then 0 else min \l\ \u\, max \l\ \u\))" | -"approx prec (Arctan a) bs = lift_un' (approx' prec a bs) (\ l u. (lb_arctan prec l, ub_arctan prec u))" | -"approx prec (Sqrt a) bs = lift_un' (approx' prec a bs) (\ l u. (lb_sqrt prec l, ub_sqrt prec u))" | -"approx prec (Exp a) bs = lift_un' (approx' prec a bs) (\ l u. (lb_exp prec l, ub_exp prec u))" | -"approx prec (Powr a b) bs = lift_bin (approx' prec a bs) (approx' prec b bs) (bnds_powr prec)" | -"approx prec (Ln a) bs = lift_un (approx' prec a bs) (\ l u. (lb_ln prec l, ub_ln prec u))" | -"approx prec (Power a n) bs = lift_un' (approx' prec a bs) (float_power_bnds prec n)" | -"approx prec (Floor a) bs = lift_un' (approx' prec a bs) (\ l u. (floor_fl l, floor_fl u))" | -"approx prec (Num f) bs = Some (f, f)" | -"approx prec (Var i) bs = (if i < length bs then bs ! i else None)" +fun approx approx' :: "nat \ floatarith \ (float interval) option list \ (float interval) option" + where + "approx' prec a bs = (case (approx prec a bs) of Some ivl \ Some (round_interval prec ivl) | None \ None)" | + "approx prec (Add a b) bs = lift_bin' (approx' prec a bs) (approx' prec b bs) (+)" | + "approx prec (Minus a) bs = lift_un' (approx' prec a bs) uminus" | + "approx prec (Mult a b) bs = lift_bin' (approx' prec a bs) (approx' prec b bs) (mult_float_interval prec)" | + "approx prec (Inverse a) bs = lift_un (approx' prec a bs) (inverse_float_interval prec)" | + "approx prec (Cos a) bs = lift_un' (approx' prec a bs) (cos_float_interval prec)" | + "approx prec Pi bs = Some (pi_float_interval prec)" | + "approx prec (Min a b) bs = lift_bin' (approx' prec a bs) (approx' prec b bs) (min_interval)" | + "approx prec (Max a b) bs = lift_bin' (approx' prec a bs) (approx' prec b bs) (max_interval)" | + "approx prec (Abs a) bs = lift_un' (approx' prec a bs) (abs_interval)" | + "approx prec (Arctan a) bs = lift_un' (approx' prec a bs) (arctan_float_interval prec)" | + "approx prec (Sqrt a) bs = lift_un' (approx' prec a bs) (sqrt_float_interval prec)" | + "approx prec (Exp a) bs = lift_un' (approx' prec a bs) (exp_float_interval prec)" | + "approx prec (Powr a b) bs = lift_bin (approx' prec a bs) (approx' prec b bs) (powr_float_interval prec)" | + "approx prec (Ln a) bs = lift_un (approx' prec a bs) (ln_float_interval prec)" | + "approx prec (Power a n) bs = lift_un' (approx' prec a bs) (power_float_interval prec n)" | + "approx prec (Floor a) bs = lift_un' (approx' prec a bs) (floor_float_interval)" | + "approx prec (Num f) bs = Some (interval_of f)" | + "approx prec (Var i) bs = (if i < length bs then bs ! i else None)" lemma approx_approx': - assumes Pa: "\l u. Some (l, u) = approx prec a vs \ - l \ interpret_floatarith a xs \ interpret_floatarith a xs \ u" - and approx': "Some (l, u) = approx' prec a vs" - shows "l \ interpret_floatarith a xs \ interpret_floatarith a xs \ u" -proof - - obtain l' u' where S: "Some (l', u') = approx prec a vs" - using approx' unfolding approx'.simps by (cases "approx prec a vs") auto - have l': "l = float_round_down prec l'" and u': "u = float_round_up prec u'" - using approx' unfolding approx'.simps S[symmetric] by auto - show ?thesis unfolding l' u' - using order_trans[OF Pa[OF S, THEN conjunct2] float_round_up[of u']] - using order_trans[OF float_round_down[of _ l'] Pa[OF S, THEN conjunct1]] by auto -qed - -lemma lift_bin_ex: - assumes lift_bin_Some: "Some (l, u) = lift_bin a b f" - shows "\ l1 u1 l2 u2. Some (l1, u1) = a \ Some (l2, u2) = b" -proof (cases a) - case None - hence "None = lift_bin a b f" - unfolding None lift_bin.simps .. - thus ?thesis - using lift_bin_Some by auto -next - case (Some a') - show ?thesis - proof (cases b) - case None - hence "None = lift_bin a b f" - unfolding None lift_bin.simps .. - thus ?thesis using lift_bin_Some by auto - next - case (Some b') - obtain la ua where a': "a' = (la, ua)" - by (cases a') auto - obtain lb ub where b': "b' = (lb, ub)" - by (cases b') auto - thus ?thesis - unfolding \a = Some a'\ \b = Some b'\ a' b' by auto - qed -qed - -lemma lift_bin_f: - assumes lift_bin_Some: "Some (l, u) = lift_bin (g a) (g b) f" - and Pa: "\l u. Some (l, u) = g a \ P l u a" - and Pb: "\l u. Some (l, u) = g b \ P l u b" - shows "\ l1 u1 l2 u2. P l1 u1 a \ P l2 u2 b \ Some (l, u) = f l1 u1 l2 u2" -proof - - obtain l1 u1 l2 u2 - where Sa: "Some (l1, u1) = g a" - and Sb: "Some (l2, u2) = g b" - using lift_bin_ex[OF assms(1)] by auto - have lu: "Some (l, u) = f l1 u1 l2 u2" - using lift_bin_Some[unfolded Sa[symmetric] Sb[symmetric] lift_bin.simps] by auto - thus ?thesis - using Pa[OF Sa] Pb[OF Sb] by auto -qed - -lemma lift_bin: - assumes lift_bin_Some: "Some (l, u) = lift_bin (approx' prec a bs) (approx' prec b bs) f" - and Pa: "\l u. Some (l, u) = approx prec a bs \ - real_of_float l \ interpret_floatarith a xs \ interpret_floatarith a xs \ real_of_float u" (is "\l u. _ = ?g a \ ?P l u a") - and Pb: "\l u. Some (l, u) = approx prec b bs \ - real_of_float l \ interpret_floatarith b xs \ interpret_floatarith b xs \ real_of_float u" - shows "\l1 u1 l2 u2. (real_of_float l1 \ interpret_floatarith a xs \ interpret_floatarith a xs \ real_of_float u1) \ - (real_of_float l2 \ interpret_floatarith b xs \ interpret_floatarith b xs \ real_of_float u2) \ - Some (l, u) = (f l1 u1 l2 u2)" -proof - - { fix l u assume "Some (l, u) = approx' prec a bs" - with approx_approx'[of prec a bs, OF _ this] Pa - have "l \ interpret_floatarith a xs \ interpret_floatarith a xs \ u" by auto } note Pa = this - { fix l u assume "Some (l, u) = approx' prec b bs" - with approx_approx'[of prec b bs, OF _ this] Pb - have "l \ interpret_floatarith b xs \ interpret_floatarith b xs \ u" by auto } note Pb = this - - from lift_bin_f[where g="\a. approx' prec a bs" and P = ?P, OF lift_bin_Some, OF Pa Pb] - show ?thesis by auto -qed - -lemma lift_bin'_ex: - assumes lift_bin'_Some: "Some (l, u) = lift_bin' a b f" - shows "\ l1 u1 l2 u2. Some (l1, u1) = a \ Some (l2, u2) = b" -proof (cases a) - case None - hence "None = lift_bin' a b f" - unfolding None lift_bin'.simps .. - thus ?thesis - using lift_bin'_Some by auto -next - case (Some a') - show ?thesis - proof (cases b) - case None - hence "None = lift_bin' a b f" - unfolding None lift_bin'.simps .. - thus ?thesis using lift_bin'_Some by auto - next - case (Some b') - obtain la ua where a': "a' = (la, ua)" - by (cases a') auto - obtain lb ub where b': "b' = (lb, ub)" - by (cases b') auto - thus ?thesis - unfolding \a = Some a'\ \b = Some b'\ a' b' by auto - qed -qed - -lemma lift_bin'_f: - assumes lift_bin'_Some: "Some (l, u) = lift_bin' (g a) (g b) f" - and Pa: "\l u. Some (l, u) = g a \ P l u a" - and Pb: "\l u. Some (l, u) = g b \ P l u b" - shows "\ l1 u1 l2 u2. P l1 u1 a \ P l2 u2 b \ l = fst (f l1 u1 l2 u2) \ u = snd (f l1 u1 l2 u2)" -proof - - obtain l1 u1 l2 u2 - where Sa: "Some (l1, u1) = g a" - and Sb: "Some (l2, u2) = g b" - using lift_bin'_ex[OF assms(1)] by auto - have lu: "(l, u) = f l1 u1 l2 u2" - using lift_bin'_Some[unfolded Sa[symmetric] Sb[symmetric] lift_bin'.simps] by auto - have "l = fst (f l1 u1 l2 u2)" and "u = snd (f l1 u1 l2 u2)" - unfolding lu[symmetric] by auto - thus ?thesis - using Pa[OF Sa] Pb[OF Sb] by auto -qed - -lemma lift_bin': - assumes lift_bin'_Some: "Some (l, u) = lift_bin' (approx' prec a bs) (approx' prec b bs) f" - and Pa: "\l u. Some (l, u) = approx prec a bs \ - l \ interpret_floatarith a xs \ interpret_floatarith a xs \ u" (is "\l u. _ = ?g a \ ?P l u a") - and Pb: "\l u. Some (l, u) = approx prec b bs \ - l \ interpret_floatarith b xs \ interpret_floatarith b xs \ u" - shows "\l1 u1 l2 u2. (l1 \ interpret_floatarith a xs \ interpret_floatarith a xs \ u1) \ - (l2 \ interpret_floatarith b xs \ interpret_floatarith b xs \ u2) \ - l = fst (f l1 u1 l2 u2) \ u = snd (f l1 u1 l2 u2)" + assumes Pa: "\ivl. approx prec a vs = Some ivl \ interpret_floatarith a xs \\<^sub>r ivl" + and approx': "approx' prec a vs = Some ivl" + shows "interpret_floatarith a xs \\<^sub>r ivl" proof - - { fix l u assume "Some (l, u) = approx' prec a bs" - with approx_approx'[of prec a bs, OF _ this] Pa - have "l \ interpret_floatarith a xs \ interpret_floatarith a xs \ u" by auto } note Pa = this - { fix l u assume "Some (l, u) = approx' prec b bs" - with approx_approx'[of prec b bs, OF _ this] Pb - have "l \ interpret_floatarith b xs \ interpret_floatarith b xs \ u" by auto } note Pb = this - - from lift_bin'_f[where g="\a. approx' prec a bs" and P = ?P, OF lift_bin'_Some, OF Pa Pb] - show ?thesis by auto -qed - -lemma lift_un'_ex: - assumes lift_un'_Some: "Some (l, u) = lift_un' a f" - shows "\ l u. Some (l, u) = a" -proof (cases a) - case None - hence "None = lift_un' a f" - unfolding None lift_un'.simps .. - thus ?thesis - using lift_un'_Some by auto -next - case (Some a') - obtain la ua where a': "a' = (la, ua)" - by (cases a') auto - thus ?thesis - unfolding \a = Some a'\ a' by auto -qed - -lemma lift_un'_f: - assumes lift_un'_Some: "Some (l, u) = lift_un' (g a) f" - and Pa: "\l u. Some (l, u) = g a \ P l u a" - shows "\ l1 u1. P l1 u1 a \ l = fst (f l1 u1) \ u = snd (f l1 u1)" -proof - - obtain l1 u1 where Sa: "Some (l1, u1) = g a" - using lift_un'_ex[OF assms(1)] by auto - have lu: "(l, u) = f l1 u1" - using lift_un'_Some[unfolded Sa[symmetric] lift_un'.simps] by auto - have "l = fst (f l1 u1)" and "u = snd (f l1 u1)" - unfolding lu[symmetric] by auto - thus ?thesis - using Pa[OF Sa] by auto -qed - -lemma lift_un': - assumes lift_un'_Some: "Some (l, u) = lift_un' (approx' prec a bs) f" - and Pa: "\l u. Some (l, u) = approx prec a bs \ - l \ interpret_floatarith a xs \ interpret_floatarith a xs \ u" - (is "\l u. _ = ?g a \ ?P l u a") - shows "\l1 u1. (l1 \ interpret_floatarith a xs \ interpret_floatarith a xs \ u1) \ - l = fst (f l1 u1) \ u = snd (f l1 u1)" -proof - - have Pa: "l \ interpret_floatarith a xs \ interpret_floatarith a xs \ u" - if "Some (l, u) = approx' prec a bs" for l u - using approx_approx'[of prec a bs, OF _ that] Pa - by auto - from lift_un'_f[where g="\a. approx' prec a bs" and P = ?P, OF lift_un'_Some, OF Pa] - show ?thesis by auto -qed - -lemma lift_un'_bnds: - assumes bnds: "\ (x::real) lx ux. (l, u) = f lx ux \ x \ { lx .. ux } \ l \ f' x \ f' x \ u" - and lift_un'_Some: "Some (l, u) = lift_un' (approx' prec a bs) f" - and Pa: "\l u. Some (l, u) = approx prec a bs \ - l \ interpret_floatarith a xs \ interpret_floatarith a xs \ u" - shows "real_of_float l \ f' (interpret_floatarith a xs) \ f' (interpret_floatarith a xs) \ real_of_float u" -proof - - from lift_un'[OF lift_un'_Some Pa] - obtain l1 u1 where "l1 \ interpret_floatarith a xs" - and "interpret_floatarith a xs \ u1" - and "l = fst (f l1 u1)" - and "u = snd (f l1 u1)" - by blast - hence "(l, u) = f l1 u1" and "interpret_floatarith a xs \ {l1 .. u1}" - by auto - thus ?thesis - using bnds by auto -qed - -lemma lift_un_ex: - assumes lift_un_Some: "Some (l, u) = lift_un a f" - shows "\l u. Some (l, u) = a" -proof (cases a) - case None - hence "None = lift_un a f" - unfolding None lift_un.simps .. - thus ?thesis - using lift_un_Some by auto -next - case (Some a') - obtain la ua where a': "a' = (la, ua)" - by (cases a') auto - thus ?thesis - unfolding \a = Some a'\ a' by auto -qed - -lemma lift_un_f: - assumes lift_un_Some: "Some (l, u) = lift_un (g a) f" - and Pa: "\l u. Some (l, u) = g a \ P l u a" - shows "\ l1 u1. P l1 u1 a \ Some l = fst (f l1 u1) \ Some u = snd (f l1 u1)" -proof - - obtain l1 u1 where Sa: "Some (l1, u1) = g a" - using lift_un_ex[OF assms(1)] by auto - have "fst (f l1 u1) \ None \ snd (f l1 u1) \ None" - proof (rule ccontr) - assume "\ (fst (f l1 u1) \ None \ snd (f l1 u1) \ None)" - hence or: "fst (f l1 u1) = None \ snd (f l1 u1) = None" by auto - hence "lift_un (g a) f = None" - proof (cases "fst (f l1 u1) = None") - case True - then obtain b where b: "f l1 u1 = (None, b)" - by (cases "f l1 u1") auto - thus ?thesis - unfolding Sa[symmetric] lift_un.simps b by auto - next - case False - hence "snd (f l1 u1) = None" - using or by auto - with False obtain b where b: "f l1 u1 = (Some b, None)" - by (cases "f l1 u1") auto - thus ?thesis - unfolding Sa[symmetric] lift_un.simps b by auto - qed - thus False - using lift_un_Some by auto - qed - then obtain a' b' where f: "f l1 u1 = (Some a', Some b')" - by (cases "f l1 u1") auto - from lift_un_Some[unfolded Sa[symmetric] lift_un.simps f] - have "Some l = fst (f l1 u1)" and "Some u = snd (f l1 u1)" - unfolding f by auto - thus ?thesis - unfolding Sa[symmetric] lift_un.simps using Pa[OF Sa] by auto -qed - -lemma lift_un: - assumes lift_un_Some: "Some (l, u) = lift_un (approx' prec a bs) f" - and Pa: "\l u. Some (l, u) = approx prec a bs \ - l \ interpret_floatarith a xs \ interpret_floatarith a xs \ u" - (is "\l u. _ = ?g a \ ?P l u a") - shows "\l1 u1. (l1 \ interpret_floatarith a xs \ interpret_floatarith a xs \ u1) \ - Some l = fst (f l1 u1) \ Some u = snd (f l1 u1)" -proof - - have Pa: "l \ interpret_floatarith a xs \ interpret_floatarith a xs \ u" - if "Some (l, u) = approx' prec a bs" for l u - using approx_approx'[of prec a bs, OF _ that] Pa by auto - from lift_un_f[where g="\a. approx' prec a bs" and P = ?P, OF lift_un_Some, OF Pa] - show ?thesis by auto -qed - -lemma lift_un_bnds: - assumes bnds: "\(x::real) lx ux. (Some l, Some u) = f lx ux \ x \ { lx .. ux } \ l \ f' x \ f' x \ u" - and lift_un_Some: "Some (l, u) = lift_un (approx' prec a bs) f" - and Pa: "\l u. Some (l, u) = approx prec a bs \ - l \ interpret_floatarith a xs \ interpret_floatarith a xs \ u" - shows "real_of_float l \ f' (interpret_floatarith a xs) \ f' (interpret_floatarith a xs) \ real_of_float u" -proof - - from lift_un[OF lift_un_Some Pa] - obtain l1 u1 where "l1 \ interpret_floatarith a xs" - and "interpret_floatarith a xs \ u1" - and "Some l = fst (f l1 u1)" - and "Some u = snd (f l1 u1)" - by blast - hence "(Some l, Some u) = f l1 u1" and "interpret_floatarith a xs \ {l1 .. u1}" - by auto - thus ?thesis - using bnds by auto + obtain ivl' where S: "approx prec a vs = Some ivl'" and ivl_def: "ivl = round_interval prec ivl'" + using approx' unfolding approx'.simps by (cases "approx prec a vs") auto + show ?thesis + by (auto simp: ivl_def intro!: in_round_intervalI Pa[OF S]) qed lemma approx: assumes "bounded_by xs vs" - and "Some (l, u) = approx prec arith vs" (is "_ = ?g arith") - shows "l \ interpret_floatarith arith xs \ interpret_floatarith arith xs \ u" (is "?P l u arith") - using \Some (l, u) = approx prec arith vs\ -proof (induct arith arbitrary: l u) - case (Add a b) - from lift_bin'[OF Add.prems[unfolded approx.simps]] Add.hyps - obtain l1 u1 l2 u2 where "l = float_plus_down prec l1 l2" - and "u = float_plus_up prec u1 u2" "l1 \ interpret_floatarith a xs" - and "interpret_floatarith a xs \ u1" "l2 \ interpret_floatarith b xs" - and "interpret_floatarith b xs \ u2" - unfolding fst_conv snd_conv by blast - thus ?case - unfolding interpret_floatarith.simps by (auto intro!: float_plus_up_le float_plus_down_le) -next - case (Minus a) - from lift_un'[OF Minus.prems[unfolded approx.simps]] Minus.hyps - obtain l1 u1 where "l = -u1" "u = -l1" - and "l1 \ interpret_floatarith a xs" "interpret_floatarith a xs \ u1" - unfolding fst_conv snd_conv by blast - thus ?case - unfolding interpret_floatarith.simps using minus_float.rep_eq by auto -next - case (Mult a b) - from lift_bin'[OF Mult.prems[unfolded approx.simps]] Mult.hyps - obtain l1 u1 l2 u2 - where l: "l = fst (bnds_mult prec l1 u1 l2 u2)" - and u: "u = snd (bnds_mult prec l1 u1 l2 u2)" - and a: "l1 \ interpret_floatarith a xs" "interpret_floatarith a xs \ u1" - and b: "l2 \ interpret_floatarith b xs" "interpret_floatarith b xs \ u2" unfolding fst_conv snd_conv by blast - from l u have lu: "(l, u) = bnds_mult prec l1 u1 l2 u2" by simp - from bnds_mult[OF lu] a b show ?case by simp -next - case (Inverse a) - from lift_un[OF Inverse.prems[unfolded approx.simps], unfolded if_distrib[of fst] if_distrib[of snd] fst_conv snd_conv] Inverse.hyps - obtain l1 u1 where l': "Some l = (if 0 < l1 \ u1 < 0 then Some (float_divl prec 1 u1) else None)" - and u': "Some u = (if 0 < l1 \ u1 < 0 then Some (float_divr prec 1 l1) else None)" - and l1: "l1 \ interpret_floatarith a xs" - and u1: "interpret_floatarith a xs \ u1" - by blast - have either: "0 < l1 \ u1 < 0" - proof (rule ccontr) - assume P: "\ (0 < l1 \ u1 < 0)" - show False - using l' unfolding if_not_P[OF P] by auto - qed - moreover have l1_le_u1: "real_of_float l1 \ real_of_float u1" - using l1 u1 by auto - ultimately have "real_of_float l1 \ 0" and "real_of_float u1 \ 0" - by auto - - have inv: "inverse u1 \ inverse (interpret_floatarith a xs) - \ inverse (interpret_floatarith a xs) \ inverse l1" - proof (cases "0 < l1") - case True - hence "0 < real_of_float u1" and "0 < real_of_float l1" "0 < interpret_floatarith a xs" - using l1_le_u1 l1 by auto - show ?thesis - unfolding inverse_le_iff_le[OF \0 < real_of_float u1\ \0 < interpret_floatarith a xs\] - inverse_le_iff_le[OF \0 < interpret_floatarith a xs\ \0 < real_of_float l1\] - using l1 u1 by auto - next - case False - hence "u1 < 0" - using either by blast - hence "real_of_float u1 < 0" and "real_of_float l1 < 0" "interpret_floatarith a xs < 0" - using l1_le_u1 u1 by auto - show ?thesis - unfolding inverse_le_iff_le_neg[OF \real_of_float u1 < 0\ \interpret_floatarith a xs < 0\] - inverse_le_iff_le_neg[OF \interpret_floatarith a xs < 0\ \real_of_float l1 < 0\] - using l1 u1 by auto - qed - - from l' have "l = float_divl prec 1 u1" - by (cases "0 < l1 \ u1 < 0") auto - hence "l \ inverse u1" - unfolding nonzero_inverse_eq_divide[OF \real_of_float u1 \ 0\] - using float_divl[of prec 1 u1] by auto - also have "\ \ inverse (interpret_floatarith a xs)" - using inv by auto - finally have "l \ inverse (interpret_floatarith a xs)" . - moreover - from u' have "u = float_divr prec 1 l1" - by (cases "0 < l1 \ u1 < 0") auto - hence "inverse l1 \ u" - unfolding nonzero_inverse_eq_divide[OF \real_of_float l1 \ 0\] - using float_divr[of 1 l1 prec] by auto - hence "inverse (interpret_floatarith a xs) \ u" - by (rule order_trans[OF inv[THEN conjunct2]]) - ultimately show ?case - unfolding interpret_floatarith.simps using l1 u1 by auto -next - case (Abs x) - from lift_un'[OF Abs.prems[unfolded approx.simps], unfolded fst_conv snd_conv] Abs.hyps - obtain l1 u1 where l': "l = (if l1 < 0 \ 0 < u1 then 0 else min \l1\ \u1\)" - and u': "u = max \l1\ \u1\" - and l1: "l1 \ interpret_floatarith x xs" - and u1: "interpret_floatarith x xs \ u1" - by blast - thus ?case - unfolding l' u' - by (cases "l1 < 0 \ 0 < u1") (auto simp add: real_of_float_min real_of_float_max) -next - case (Min a b) - from lift_bin'[OF Min.prems[unfolded approx.simps], unfolded fst_conv snd_conv] Min.hyps - obtain l1 u1 l2 u2 where l': "l = min l1 l2" and u': "u = min u1 u2" - and l1: "l1 \ interpret_floatarith a xs" and u1: "interpret_floatarith a xs \ u1" - and l1: "l2 \ interpret_floatarith b xs" and u1: "interpret_floatarith b xs \ u2" - by blast - thus ?case - unfolding l' u' by (auto simp add: real_of_float_min) -next - case (Max a b) - from lift_bin'[OF Max.prems[unfolded approx.simps], unfolded fst_conv snd_conv] Max.hyps - obtain l1 u1 l2 u2 where l': "l = max l1 l2" and u': "u = max u1 u2" - and l1: "l1 \ interpret_floatarith a xs" and u1: "interpret_floatarith a xs \ u1" - and l1: "l2 \ interpret_floatarith b xs" and u1: "interpret_floatarith b xs \ u2" - by blast - thus ?case - unfolding l' u' by (auto simp add: real_of_float_max) -next - case (Cos a) - with lift_un'_bnds[OF bnds_cos] show ?case by auto -next - case (Arctan a) - with lift_un'_bnds[OF bnds_arctan] show ?case by auto -next - case Pi - with pi_boundaries show ?case by auto -next - case (Sqrt a) - with lift_un'_bnds[OF bnds_sqrt] show ?case by auto -next - case (Exp a) - with lift_un'_bnds[OF bnds_exp] show ?case by auto -next - case (Powr a b) - from lift_bin[OF Powr.prems[unfolded approx.simps]] Powr.hyps - obtain l1 u1 l2 u2 where lu: "Some (l, u) = bnds_powr prec l1 u1 l2 u2" - and l1: "l1 \ interpret_floatarith a xs" and u1: "interpret_floatarith a xs \ u1" - and l2: "l2 \ interpret_floatarith b xs" and u2: "interpret_floatarith b xs \ u2" - by blast - from bnds_powr[OF lu] l1 u1 l2 u2 - show ?case by simp -next - case (Ln a) - with lift_un_bnds[OF bnds_ln] show ?case by auto -next - case (Power a n) - with lift_un'_bnds[OF bnds_power] show ?case by auto -next - case (Floor a) - from lift_un'[OF Floor.prems[unfolded approx.simps] Floor.hyps] - show ?case by (auto simp: floor_fl.rep_eq floor_mono) -next - case (Num f) - thus ?case by auto -next - case (Var n) - from this[symmetric] \bounded_by xs vs\[THEN bounded_byE, of n] - show ?case by (cases "n < length vs") auto -qed + and "approx prec arith vs = Some ivl" + shows "interpret_floatarith arith xs \\<^sub>r ivl" + using \approx prec arith vs = Some ivl\ + using \bounded_by xs vs\[THEN bounded_byE] + by (induct arith arbitrary: ivl) + (force split: option.splits if_splits + intro!: plus_in_float_intervalI mult_float_intervalI uminus_in_float_intervalI + inverse_float_interval_eqI abs_in_float_intervalI + min_intervalI max_intervalI + cos_float_intervalI + arctan_float_intervalI pi_float_interval sqrt_float_intervalI exp_float_intervalI + powr_float_interval_eqI ln_float_interval_eqI power_float_intervalI floor_float_intervalI + intro: in_round_intervalI)+ datatype form = Bound floatarith floatarith floatarith form | Assign floatarith floatarith form | Less floatarith floatarith | LessEqual floatarith floatarith | AtLeastAtMost floatarith floatarith floatarith | Conj form form | Disj form form fun interpret_form :: "form \ real list \ bool" where "interpret_form (Bound x a b f) vs = (interpret_floatarith x vs \ { interpret_floatarith a vs .. interpret_floatarith b vs } \ interpret_form f vs)" | "interpret_form (Assign x a f) vs = (interpret_floatarith x vs = interpret_floatarith a vs \ interpret_form f vs)" | "interpret_form (Less a b) vs = (interpret_floatarith a vs < interpret_floatarith b vs)" | "interpret_form (LessEqual a b) vs = (interpret_floatarith a vs \ interpret_floatarith b vs)" | "interpret_form (AtLeastAtMost x a b) vs = (interpret_floatarith x vs \ { interpret_floatarith a vs .. interpret_floatarith b vs })" | "interpret_form (Conj f g) vs \ interpret_form f vs \ interpret_form g vs" | "interpret_form (Disj f g) vs \ interpret_form f vs \ interpret_form g vs" -fun approx_form' and approx_form :: "nat \ form \ (float * float) option list \ nat list \ bool" where -"approx_form' prec f 0 n l u bs ss = approx_form prec f (bs[n := Some (l, u)]) ss" | -"approx_form' prec f (Suc s) n l u bs ss = - (let m = (l + u) * Float 1 (- 1) - in (if approx_form' prec f s n l m bs ss then approx_form' prec f s n m u bs ss else False))" | +fun approx_form' and approx_form :: "nat \ form \ (float interval) option list \ nat list \ bool" where +"approx_form' prec f 0 n ivl bs ss = approx_form prec f (bs[n := Some ivl]) ss" | +"approx_form' prec f (Suc s) n ivl bs ss = + (let (ivl1, ivl2) = split_float_interval ivl + in (if approx_form' prec f s n ivl1 bs ss then approx_form' prec f s n ivl2 bs ss else False))" | "approx_form prec (Bound (Var n) a b f) bs ss = (case (approx prec a bs, approx prec b bs) - of (Some (l, _), Some (_, u)) \ approx_form' prec f (ss ! n) n l u bs ss + of (Some ivl1, Some ivl2) \ approx_form' prec f (ss ! n) n (sup ivl1 ivl2) bs ss | _ \ False)" | "approx_form prec (Assign (Var n) a f) bs ss = (case (approx prec a bs) - of (Some (l, u)) \ approx_form' prec f (ss ! n) n l u bs ss + of (Some ivl) \ approx_form' prec f (ss ! n) n ivl bs ss | _ \ False)" | "approx_form prec (Less a b) bs ss = (case (approx prec a bs, approx prec b bs) - of (Some (l, u), Some (l', u')) \ float_plus_up prec u (-l') < 0 + of (Some ivl, Some ivl') \ float_plus_up prec (upper ivl) (-lower ivl') < 0 | _ \ False)" | "approx_form prec (LessEqual a b) bs ss = (case (approx prec a bs, approx prec b bs) - of (Some (l, u), Some (l', u')) \ float_plus_up prec u (-l') \ 0 + of (Some ivl, Some ivl') \ float_plus_up prec (upper ivl) (-lower ivl') \ 0 | _ \ False)" | "approx_form prec (AtLeastAtMost x a b) bs ss = (case (approx prec x bs, approx prec a bs, approx prec b bs) - of (Some (lx, ux), Some (l, u), Some (l', u')) \ float_plus_up prec u (-lx) \ 0 \ float_plus_up prec ux (-l') \ 0 + of (Some ivlx, Some ivl, Some ivl') \ + float_plus_up prec (upper ivl) (-lower ivlx) \ 0 \ + float_plus_up prec (upper ivlx) (-lower ivl') \ 0 | _ \ False)" | "approx_form prec (Conj a b) bs ss \ approx_form prec a bs ss \ approx_form prec b bs ss" | "approx_form prec (Disj a b) bs ss \ approx_form prec a bs ss \ approx_form prec b bs ss" | "approx_form _ _ _ _ = False" lemma lazy_conj: "(if A then B else False) = (A \ B)" by simp lemma approx_form_approx_form': - assumes "approx_form' prec f s n l u bs ss" - and "(x::real) \ { l .. u }" - obtains l' u' where "x \ { l' .. u' }" - and "approx_form prec f (bs[n := Some (l', u')]) ss" -using assms proof (induct s arbitrary: l u) + assumes "approx_form' prec f s n ivl bs ss" + and "(x::real) \\<^sub>r ivl" + obtains ivl' where "x \\<^sub>r ivl'" + and "approx_form prec f (bs[n := Some ivl']) ss" +using assms proof (induct s arbitrary: ivl) case 0 - from this(1)[of l u] this(2,3) + from this(1)[of ivl] this(2,3) show thesis by auto next case (Suc s) - - let ?m = "(l + u) * Float 1 (- 1)" - have "real_of_float l \ ?m" and "?m \ real_of_float u" - unfolding less_eq_float_def using Suc.prems by auto - - with \x \ { l .. u }\ - have "x \ { l .. ?m} \ x \ { ?m .. u }" by auto - thus thesis - proof (rule disjE) - assume *: "x \ { l .. ?m }" - with Suc.hyps[OF _ _ *] Suc.prems - show thesis by (simp add: Let_def lazy_conj) + then obtain ivl1 ivl2 where ivl_split: "split_float_interval ivl = (ivl1, ivl2)" + by (fastforce dest: split_float_intervalD) + from split_float_interval_realD[OF this \x \\<^sub>r ivl\] + consider "x \\<^sub>r ivl1" | "x \\<^sub>r ivl2" + by atomize_elim + then show thesis + proof cases + case *: 1 + from Suc.hyps[OF _ _ *] Suc.prems + show ?thesis + by (simp add: lazy_conj ivl_split split: prod.splits) next - assume *: "x \ { ?m .. u }" - with Suc.hyps[OF _ _ *] Suc.prems - show thesis by (simp add: Let_def lazy_conj) + case *: 2 + from Suc.hyps[OF _ _ *] Suc.prems + show ?thesis + by (simp add: lazy_conj ivl_split split: prod.splits) qed qed lemma approx_form_aux: assumes "approx_form prec f vs ss" and "bounded_by xs vs" shows "interpret_form f xs" using assms proof (induct f arbitrary: vs) case (Bound x a b f) then obtain n where x_eq: "x = Var n" by (cases x) auto - with Bound.prems obtain l u' l' u - where l_eq: "Some (l, u') = approx prec a vs" - and u_eq: "Some (l', u) = approx prec b vs" - and approx_form': "approx_form' prec f (ss ! n) n l u vs ss" + with Bound.prems obtain ivl1 ivl2 + where l_eq: "approx prec a vs = Some ivl1" + and u_eq: "approx prec b vs = Some ivl2" + and approx_form': "approx_form' prec f (ss ! n) n (sup ivl1 ivl2) vs ss" by (cases "approx prec a vs", simp) (cases "approx prec b vs", auto) have "interpret_form f xs" if "xs ! n \ { interpret_floatarith a xs .. interpret_floatarith b xs }" proof - from approx[OF Bound.prems(2) l_eq] and approx[OF Bound.prems(2) u_eq] that - have "xs ! n \ { l .. u}" by auto + have "xs ! n \\<^sub>r (sup ivl1 ivl2)" + by (auto simp: set_of_eq sup_float_def max_def inf_float_def min_def) from approx_form_approx_form'[OF approx_form' this] - obtain lx ux where bnds: "xs ! n \ { lx .. ux }" - and approx_form: "approx_form prec f (vs[n := Some (lx, ux)]) ss" . + obtain ivlx where bnds: "xs ! n \\<^sub>r ivlx" + and approx_form: "approx_form prec f (vs[n := Some ivlx]) ss" . - from \bounded_by xs vs\ bnds have "bounded_by xs (vs[n := Some (lx, ux)])" + from \bounded_by xs vs\ bnds have "bounded_by xs (vs[n := Some ivlx])" by (rule bounded_by_update) with Bound.hyps[OF approx_form] show ?thesis by blast qed thus ?case using interpret_form.simps x_eq and interpret_floatarith.simps by simp next case (Assign x a f) then obtain n where x_eq: "x = Var n" by (cases x) auto - with Assign.prems obtain l u - where bnd_eq: "Some (l, u) = approx prec a vs" + with Assign.prems obtain ivl + where bnd_eq: "approx prec a vs = Some ivl" and x_eq: "x = Var n" - and approx_form': "approx_form' prec f (ss ! n) n l u vs ss" + and approx_form': "approx_form' prec f (ss ! n) n ivl vs ss" by (cases "approx prec a vs") auto have "interpret_form f xs" if bnds: "xs ! n = interpret_floatarith a xs" proof - from approx[OF Assign.prems(2) bnd_eq] bnds - have "xs ! n \ { l .. u}" by auto + have "xs ! n \\<^sub>r ivl" by auto from approx_form_approx_form'[OF approx_form' this] - obtain lx ux where bnds: "xs ! n \ { lx .. ux }" - and approx_form: "approx_form prec f (vs[n := Some (lx, ux)]) ss" . + obtain ivlx where bnds: "xs ! n \\<^sub>r ivlx" + and approx_form: "approx_form prec f (vs[n := Some ivlx]) ss" . - from \bounded_by xs vs\ bnds have "bounded_by xs (vs[n := Some (lx, ux)])" + from \bounded_by xs vs\ bnds have "bounded_by xs (vs[n := Some ivlx])" by (rule bounded_by_update) with Assign.hyps[OF approx_form] show ?thesis by blast qed thus ?case using interpret_form.simps x_eq and interpret_floatarith.simps by simp next case (Less a b) - then obtain l u l' u' - where l_eq: "Some (l, u) = approx prec a vs" - and u_eq: "Some (l', u') = approx prec b vs" - and inequality: "real_of_float (float_plus_up prec u (-l')) < 0" + then obtain ivl ivl' + where l_eq: "approx prec a vs = Some ivl" + and u_eq: "approx prec b vs = Some ivl'" + and inequality: "real_of_float (float_plus_up prec (upper ivl) (-lower ivl')) < 0" by (cases "approx prec a vs", auto, cases "approx prec b vs", auto) from le_less_trans[OF float_plus_up inequality] approx[OF Less.prems(2) l_eq] approx[OF Less.prems(2) u_eq] - show ?case by auto + show ?case by (auto simp: set_of_eq) next case (LessEqual a b) - then obtain l u l' u' - where l_eq: "Some (l, u) = approx prec a vs" - and u_eq: "Some (l', u') = approx prec b vs" - and inequality: "real_of_float (float_plus_up prec u (-l')) \ 0" + then obtain ivl ivl' + where l_eq: "approx prec a vs = Some ivl" + and u_eq: "approx prec b vs = Some ivl'" + and inequality: "real_of_float (float_plus_up prec (upper ivl) (-lower ivl')) \ 0" by (cases "approx prec a vs", auto, cases "approx prec b vs", auto) from order_trans[OF float_plus_up inequality] approx[OF LessEqual.prems(2) l_eq] approx[OF LessEqual.prems(2) u_eq] - show ?case by auto + show ?case by (auto simp: set_of_eq) next case (AtLeastAtMost x a b) - then obtain lx ux l u l' u' - where x_eq: "Some (lx, ux) = approx prec x vs" - and l_eq: "Some (l, u) = approx prec a vs" - and u_eq: "Some (l', u') = approx prec b vs" - and inequality: "real_of_float (float_plus_up prec u (-lx)) \ 0" "real_of_float (float_plus_up prec ux (-l')) \ 0" + then obtain ivlx ivl ivl' + where x_eq: "approx prec x vs = Some ivlx" + and l_eq: "approx prec a vs = Some ivl" + and u_eq: "approx prec b vs = Some ivl'" + and inequality: "real_of_float (float_plus_up prec (upper ivl) (-lower ivlx)) \ 0" + "real_of_float (float_plus_up prec (upper ivlx) (-lower ivl')) \ 0" by (cases "approx prec x vs", auto, cases "approx prec a vs", auto, cases "approx prec b vs", auto) from order_trans[OF float_plus_up inequality(1)] order_trans[OF float_plus_up inequality(2)] approx[OF AtLeastAtMost.prems(2) l_eq] approx[OF AtLeastAtMost.prems(2) u_eq] approx[OF AtLeastAtMost.prems(2) x_eq] - show ?case by auto + show ?case by (auto simp: set_of_eq) qed auto lemma approx_form: assumes "n = length xs" and "approx_form prec f (replicate n None) ss" shows "interpret_form f xs" using approx_form_aux[OF _ bounded_by_None] assms by auto subsection \Implementing Taylor series expansion\ fun isDERIV :: "nat \ floatarith \ real list \ bool" where "isDERIV x (Add a b) vs = (isDERIV x a vs \ isDERIV x b vs)" | "isDERIV x (Mult a b) vs = (isDERIV x a vs \ isDERIV x b vs)" | "isDERIV x (Minus a) vs = isDERIV x a vs" | "isDERIV x (Inverse a) vs = (isDERIV x a vs \ interpret_floatarith a vs \ 0)" | "isDERIV x (Cos a) vs = isDERIV x a vs" | "isDERIV x (Arctan a) vs = isDERIV x a vs" | "isDERIV x (Min a b) vs = False" | "isDERIV x (Max a b) vs = False" | "isDERIV x (Abs a) vs = False" | "isDERIV x Pi vs = True" | "isDERIV x (Sqrt a) vs = (isDERIV x a vs \ interpret_floatarith a vs > 0)" | "isDERIV x (Exp a) vs = isDERIV x a vs" | "isDERIV x (Powr a b) vs = (isDERIV x a vs \ isDERIV x b vs \ interpret_floatarith a vs > 0)" | "isDERIV x (Ln a) vs = (isDERIV x a vs \ interpret_floatarith a vs > 0)" | "isDERIV x (Floor a) vs = (isDERIV x a vs \ interpret_floatarith a vs \ \)" | "isDERIV x (Power a 0) vs = True" | "isDERIV x (Power a (Suc n)) vs = isDERIV x a vs" | "isDERIV x (Num f) vs = True" | "isDERIV x (Var n) vs = True" fun DERIV_floatarith :: "nat \ floatarith \ floatarith" where "DERIV_floatarith x (Add a b) = Add (DERIV_floatarith x a) (DERIV_floatarith x b)" | "DERIV_floatarith x (Mult a b) = Add (Mult a (DERIV_floatarith x b)) (Mult (DERIV_floatarith x a) b)" | "DERIV_floatarith x (Minus a) = Minus (DERIV_floatarith x a)" | "DERIV_floatarith x (Inverse a) = Minus (Mult (DERIV_floatarith x a) (Inverse (Power a 2)))" | "DERIV_floatarith x (Cos a) = Minus (Mult (Cos (Add (Mult Pi (Num (Float 1 (- 1)))) (Minus a))) (DERIV_floatarith x a))" | "DERIV_floatarith x (Arctan a) = Mult (Inverse (Add (Num 1) (Power a 2))) (DERIV_floatarith x a)" | "DERIV_floatarith x (Min a b) = Num 0" | "DERIV_floatarith x (Max a b) = Num 0" | "DERIV_floatarith x (Abs a) = Num 0" | "DERIV_floatarith x Pi = Num 0" | "DERIV_floatarith x (Sqrt a) = (Mult (Inverse (Mult (Sqrt a) (Num 2))) (DERIV_floatarith x a))" | "DERIV_floatarith x (Exp a) = Mult (Exp a) (DERIV_floatarith x a)" | "DERIV_floatarith x (Powr a b) = Mult (Powr a b) (Add (Mult (DERIV_floatarith x b) (Ln a)) (Mult (Mult (DERIV_floatarith x a) b) (Inverse a)))" | "DERIV_floatarith x (Ln a) = Mult (Inverse a) (DERIV_floatarith x a)" | "DERIV_floatarith x (Power a 0) = Num 0" | "DERIV_floatarith x (Power a (Suc n)) = Mult (Num (Float (int (Suc n)) 0)) (Mult (Power a n) (DERIV_floatarith x a))" | "DERIV_floatarith x (Floor a) = Num 0" | "DERIV_floatarith x (Num f) = Num 0" | "DERIV_floatarith x (Var n) = (if x = n then Num 1 else Num 0)" lemma has_real_derivative_powr': fixes f g :: "real \ real" assumes "(f has_real_derivative f') (at x)" assumes "(g has_real_derivative g') (at x)" assumes "f x > 0" defines "h \ \x. f x powr g x * (g' * ln (f x) + f' * g x / f x)" shows "((\x. f x powr g x) has_real_derivative h x) (at x)" proof (subst DERIV_cong_ev[OF refl _ refl]) from assms have "isCont f x" by (simp add: DERIV_continuous) hence "f \x\ f x" by (simp add: continuous_at) with \f x > 0\ have "eventually (\x. f x > 0) (nhds x)" by (auto simp: tendsto_at_iff_tendsto_nhds dest: order_tendstoD) thus "eventually (\x. f x powr g x = exp (g x * ln (f x))) (nhds x)" by eventually_elim (simp add: powr_def) next from assms show "((\x. exp (g x * ln (f x))) has_real_derivative h x) (at x)" by (auto intro!: derivative_eq_intros simp: h_def powr_def) qed lemma DERIV_floatarith: assumes "n < length vs" assumes isDERIV: "isDERIV n f (vs[n := x])" shows "DERIV (\ x'. interpret_floatarith f (vs[n := x'])) x :> interpret_floatarith (DERIV_floatarith n f) (vs[n := x])" (is "DERIV (?i f) x :> _") using isDERIV proof (induct f arbitrary: x) case (Inverse a) thus ?case by (auto intro!: derivative_eq_intros simp add: algebra_simps power2_eq_square) next case (Cos a) thus ?case by (auto intro!: derivative_eq_intros simp del: interpret_floatarith.simps(5) simp add: interpret_floatarith_sin interpret_floatarith.simps(5)[of a]) next case (Power a n) thus ?case by (cases n) (auto intro!: derivative_eq_intros simp del: power_Suc) next case (Floor a) thus ?case by (auto intro!: derivative_eq_intros DERIV_isCont floor_has_real_derivative) next case (Ln a) thus ?case by (auto intro!: derivative_eq_intros simp add: divide_inverse) next case (Var i) thus ?case using \n < length vs\ by auto next case (Powr a b) note [derivative_intros] = has_real_derivative_powr' from Powr show ?case by (auto intro!: derivative_eq_intros simp: field_simps) qed (auto intro!: derivative_eq_intros) declare approx.simps[simp del] -fun isDERIV_approx :: "nat \ nat \ floatarith \ (float * float) option list \ bool" where +fun isDERIV_approx :: "nat \ nat \ floatarith \ (float interval) option list \ bool" where "isDERIV_approx prec x (Add a b) vs = (isDERIV_approx prec x a vs \ isDERIV_approx prec x b vs)" | "isDERIV_approx prec x (Mult a b) vs = (isDERIV_approx prec x a vs \ isDERIV_approx prec x b vs)" | "isDERIV_approx prec x (Minus a) vs = isDERIV_approx prec x a vs" | "isDERIV_approx prec x (Inverse a) vs = - (isDERIV_approx prec x a vs \ (case approx prec a vs of Some (l, u) \ 0 < l \ u < 0 | None \ False))" | + (isDERIV_approx prec x a vs \ (case approx prec a vs of Some ivl \ 0 < lower ivl \ upper ivl < 0 | None \ False))" | "isDERIV_approx prec x (Cos a) vs = isDERIV_approx prec x a vs" | "isDERIV_approx prec x (Arctan a) vs = isDERIV_approx prec x a vs" | "isDERIV_approx prec x (Min a b) vs = False" | "isDERIV_approx prec x (Max a b) vs = False" | "isDERIV_approx prec x (Abs a) vs = False" | "isDERIV_approx prec x Pi vs = True" | "isDERIV_approx prec x (Sqrt a) vs = - (isDERIV_approx prec x a vs \ (case approx prec a vs of Some (l, u) \ 0 < l | None \ False))" | + (isDERIV_approx prec x a vs \ (case approx prec a vs of Some ivl \ 0 < lower ivl | None \ False))" | "isDERIV_approx prec x (Exp a) vs = isDERIV_approx prec x a vs" | "isDERIV_approx prec x (Powr a b) vs = - (isDERIV_approx prec x a vs \ isDERIV_approx prec x b vs \ (case approx prec a vs of Some (l, u) \ 0 < l | None \ False))" | + (isDERIV_approx prec x a vs \ isDERIV_approx prec x b vs \ (case approx prec a vs of Some ivl \ 0 < lower ivl | None \ False))" | "isDERIV_approx prec x (Ln a) vs = - (isDERIV_approx prec x a vs \ (case approx prec a vs of Some (l, u) \ 0 < l | None \ False))" | + (isDERIV_approx prec x a vs \ (case approx prec a vs of Some ivl \ 0 < lower ivl | None \ False))" | "isDERIV_approx prec x (Power a 0) vs = True" | "isDERIV_approx prec x (Floor a) vs = - (isDERIV_approx prec x a vs \ (case approx prec a vs of Some (l, u) \ l > floor u \ u < ceiling l | None \ False))" | + (isDERIV_approx prec x a vs \ (case approx prec a vs of Some ivl \ lower ivl > floor (upper ivl) \ upper ivl < ceiling (lower ivl) | None \ False))" | "isDERIV_approx prec x (Power a (Suc n)) vs = isDERIV_approx prec x a vs" | "isDERIV_approx prec x (Num f) vs = True" | "isDERIV_approx prec x (Var n) vs = True" lemma isDERIV_approx: assumes "bounded_by xs vs" and isDERIV_approx: "isDERIV_approx prec x f vs" shows "isDERIV x f xs" using isDERIV_approx proof (induct f) case (Inverse a) - then obtain l u where approx_Some: "Some (l, u) = approx prec a vs" - and *: "0 < l \ u < 0" + then obtain ivl where approx_Some: "approx prec a vs = Some ivl" + and *: "0 < lower ivl \ upper ivl < 0" by (cases "approx prec a vs") auto with approx[OF \bounded_by xs vs\ approx_Some] - have "interpret_floatarith a xs \ 0" by auto + have "interpret_floatarith a xs \ 0" by (auto simp: set_of_eq) thus ?case using Inverse by auto next case (Ln a) - then obtain l u where approx_Some: "Some (l, u) = approx prec a vs" - and *: "0 < l" + then obtain ivl where approx_Some: "approx prec a vs = Some ivl" + and *: "0 < lower ivl" by (cases "approx prec a vs") auto with approx[OF \bounded_by xs vs\ approx_Some] - have "0 < interpret_floatarith a xs" by auto + have "0 < interpret_floatarith a xs" by (auto simp: set_of_eq) thus ?case using Ln by auto next case (Sqrt a) - then obtain l u where approx_Some: "Some (l, u) = approx prec a vs" - and *: "0 < l" + then obtain ivl where approx_Some: "approx prec a vs = Some ivl" + and *: "0 < lower ivl" by (cases "approx prec a vs") auto with approx[OF \bounded_by xs vs\ approx_Some] - have "0 < interpret_floatarith a xs" by auto + have "0 < interpret_floatarith a xs" by (auto simp: set_of_eq) thus ?case using Sqrt by auto next case (Power a n) thus ?case by (cases n) auto next case (Powr a b) - from Powr obtain l1 u1 where a: "Some (l1, u1) = approx prec a vs" and pos: "0 < l1" + from Powr obtain ivl1 where a: "approx prec a vs = Some ivl1" + and pos: "0 < lower ivl1" by (cases "approx prec a vs") auto with approx[OF \bounded_by xs vs\ a] - have "0 < interpret_floatarith a xs" by auto + have "0 < interpret_floatarith a xs" by (auto simp: set_of_eq) with Powr show ?case by auto next case (Floor a) - then obtain l u where approx_Some: "Some (l, u) = approx prec a vs" - and "real_of_int \real_of_float u\ < real_of_float l" "real_of_float u < real_of_int \real_of_float l\" + then obtain ivl where approx_Some: "approx prec a vs = Some ivl" + and "real_of_int \real_of_float (upper ivl)\ < real_of_float (lower ivl)" + "real_of_float (upper ivl) < real_of_int \real_of_float (lower ivl)\" and "isDERIV x a xs" by (cases "approx prec a vs") auto with approx[OF \bounded_by xs vs\ approx_Some] le_floor_iff show ?case - by (force elim!: Ints_cases) + by (force elim!: Ints_cases simp: set_of_eq) qed auto lemma bounded_by_update_var: assumes "bounded_by xs vs" - and "vs ! i = Some (l, u)" - and bnd: "x \ { real_of_float l .. real_of_float u }" + and "vs ! i = Some ivl" + and bnd: "x \\<^sub>r ivl" shows "bounded_by (xs[i := x]) vs" -proof (cases "i < length xs") - case False - thus ?thesis - using \bounded_by xs vs\ by auto -next - case True - let ?xs = "xs[i := x]" - from True have "i < length ?xs" by auto - have "case vs ! j of None \ True | Some (l, u) \ ?xs ! j \ {real_of_float l .. real_of_float u}" - if "j < length vs" for j - proof (cases "vs ! j") - case None - then show ?thesis by simp - next - case (Some b) - thus ?thesis - proof (cases "i = j") - case True - thus ?thesis using \vs ! i = Some (l, u)\ Some and bnd \i < length ?xs\ - by auto - next - case False - thus ?thesis - using \bounded_by xs vs\[THEN bounded_byE, OF \j < length vs\] Some by auto - qed - qed - thus ?thesis - unfolding bounded_by_def by auto -qed + using assms + using nth_list_update + by (cases "i < length xs") + (force simp: bounded_by_def split: option.splits)+ lemma isDERIV_approx': assumes "bounded_by xs vs" - and vs_x: "vs ! x = Some (l, u)" - and X_in: "X \ {real_of_float l .. real_of_float u}" + and vs_x: "vs ! x = Some ivl" + and X_in: "X \\<^sub>r ivl" and approx: "isDERIV_approx prec x f vs" shows "isDERIV x f (xs[x := X])" proof - from bounded_by_update_var[OF \bounded_by xs vs\ vs_x X_in] approx show ?thesis by (rule isDERIV_approx) qed lemma DERIV_approx: assumes "n < length xs" and bnd: "bounded_by xs vs" and isD: "isDERIV_approx prec n f vs" - and app: "Some (l, u) = approx prec (DERIV_floatarith n f) vs" (is "_ = approx _ ?D _") - shows "\(x::real). l \ x \ x \ u \ + and app: "approx prec (DERIV_floatarith n f) vs = Some ivl" (is "approx _ ?D _ = _") + shows "\(x::real). x \\<^sub>r ivl \ DERIV (\ x. interpret_floatarith f (xs[n := x])) (xs!n) :> x" - (is "\ x. _ \ _ \ DERIV (?i f) _ :> _") -proof (rule exI[of _ "?i ?D (xs!n)"], rule conjI[OF _ conjI]) + (is "\ x. _ \ DERIV (?i f) _ :> _") +proof (rule exI[of _ "?i ?D (xs!n)"], rule conjI) let "?i f" = "\x. interpret_floatarith f (xs[n := x])" from approx[OF bnd app] - show "l \ ?i ?D (xs!n)" and "?i ?D (xs!n) \ u" + show "?i ?D (xs!n) \\<^sub>r ivl" using \n < length xs\ by auto from DERIV_floatarith[OF \n < length xs\, of f "xs!n"] isDERIV_approx[OF bnd isD] show "DERIV (?i f) (xs!n) :> (?i ?D (xs!n))" by simp qed lemma lift_bin_aux: - assumes lift_bin_Some: "Some (l, u) = lift_bin a b f" - obtains l1 u1 l2 u2 - where "a = Some (l1, u1)" - and "b = Some (l2, u2)" - and "f l1 u1 l2 u2 = Some (l, u)" + assumes lift_bin_Some: "lift_bin a b f = Some ivl" + obtains ivl1 ivl2 + where "a = Some ivl1" + and "b = Some ivl2" + and "f ivl1 ivl2 = Some ivl" using assms by (cases a, simp, cases b, simp, auto) fun approx_tse where "approx_tse prec n 0 c k f bs = approx prec f bs" | "approx_tse prec n (Suc s) c k f bs = (if isDERIV_approx prec n f bs then - lift_bin (approx prec f (bs[n := Some (c,c)])) + lift_bin (approx prec f (bs[n := Some (interval_of c)])) (approx_tse prec n s c (Suc k) (DERIV_floatarith n f) bs) - (\ l1 u1 l2 u2. approx prec + (\ivl1 ivl2. approx prec (Add (Var 0) (Mult (Inverse (Num (Float (int k) 0))) (Mult (Add (Var (Suc (Suc 0))) (Minus (Num c))) - (Var (Suc 0))))) [Some (l1, u1), Some (l2, u2), bs!n]) + (Var (Suc 0))))) [Some ivl1, Some ivl2, bs!n]) else approx prec f bs)" lemma bounded_by_Cons: assumes bnd: "bounded_by xs vs" - and x: "x \ { real_of_float l .. real_of_float u }" - shows "bounded_by (x#xs) ((Some (l, u))#vs)" + and x: "x \\<^sub>r ivl" + shows "bounded_by (x#xs) ((Some ivl)#vs)" proof - - have "case ((Some (l,u))#vs) ! i of Some (l, u) \ (x#xs)!i \ { real_of_float l .. real_of_float u } | None \ True" - if *: "i < length ((Some (l, u))#vs)" for i + have "case ((Some ivl)#vs) ! i of Some ivl \ (x#xs)!i \\<^sub>r ivl | None \ True" + if *: "i < length ((Some ivl)#vs)" for i proof (cases i) case 0 with x show ?thesis by auto next case (Suc i) with * have "i < length vs" by auto from bnd[THEN bounded_byE, OF this] show ?thesis unfolding Suc nth_Cons_Suc . qed thus ?thesis by (auto simp add: bounded_by_def) qed lemma approx_tse_generic: assumes "bounded_by xs vs" and bnd_c: "bounded_by (xs[x := c]) vs" and "x < length vs" and "x < length xs" - and bnd_x: "vs ! x = Some (lx, ux)" - and ate: "Some (l, u) = approx_tse prec x s c k f vs" - shows "\ n. (\ m < n. \ (z::real) \ {lx .. ux}. + and bnd_x: "vs ! x = Some ivlx" + and ate: "approx_tse prec x s c k f vs = Some ivl" + shows "\ n. (\ m < n. \ (z::real) \ set_of (real_interval ivlx). DERIV (\ y. interpret_floatarith ((DERIV_floatarith x ^^ m) f) (xs[x := y])) z :> (interpret_floatarith ((DERIV_floatarith x ^^ (Suc m)) f) (xs[x := z]))) - \ (\ (t::real) \ {lx .. ux}. (\ i = 0.. j \ {k.. (\ (t::real) \ set_of (real_interval ivlx). (\ i = 0.. j \ {k.. j \ {k.. {l .. u})" (is "\ n. ?taylor f k l u n") + (xs!x - c)^n \\<^sub>r ivl)" (is "\ n. ?taylor f k ivl n") using ate -proof (induct s arbitrary: k f l u) +proof (induct s arbitrary: k f ivl) case 0 { - fix t::real assume "t \ {lx .. ux}" + fix t::real assume "t \\<^sub>r ivlx" note bounded_by_update_var[OF \bounded_by xs vs\ bnd_x this] from approx[OF this 0[unfolded approx_tse.simps]] - have "(interpret_floatarith f (xs[x := t])) \ {l .. u}" + have "(interpret_floatarith f (xs[x := t])) \\<^sub>r ivl" by (auto simp add: algebra_simps) } thus ?case by (auto intro!: exI[of _ 0]) next case (Suc s) show ?case proof (cases "isDERIV_approx prec x f vs") case False note ap = Suc.prems[unfolded approx_tse.simps if_not_P[OF False]] { - fix t::real assume "t \ {lx .. ux}" + fix t::real assume "t \\<^sub>r ivlx" note bounded_by_update_var[OF \bounded_by xs vs\ bnd_x this] from approx[OF this ap] - have "(interpret_floatarith f (xs[x := t])) \ {l .. u}" + have "(interpret_floatarith f (xs[x := t])) \\<^sub>r ivl" by (auto simp add: algebra_simps) } thus ?thesis by (auto intro!: exI[of _ 0]) next case True with Suc.prems - obtain l1 u1 l2 u2 - where a: "Some (l1, u1) = approx prec f (vs[x := Some (c,c)])" - and ate: "Some (l2, u2) = approx_tse prec x s c (Suc k) (DERIV_floatarith x f) vs" - and final: "Some (l, u) = approx prec + obtain ivl1 ivl2 + where a: "approx prec f (vs[x := Some (interval_of c)]) = Some ivl1" + and ate: "approx_tse prec x s c (Suc k) (DERIV_floatarith x f) vs = Some ivl2" + and final: "approx prec (Add (Var 0) (Mult (Inverse (Num (Float (int k) 0))) (Mult (Add (Var (Suc (Suc 0))) (Minus (Num c))) - (Var (Suc 0))))) [Some (l1, u1), Some (l2, u2), vs!x]" + (Var (Suc 0))))) [Some ivl1, Some ivl2, vs!x] = Some ivl" by (auto elim!: lift_bin_aux) from bnd_c \x < length xs\ - have bnd: "bounded_by (xs[x:=c]) (vs[x:= Some (c,c)])" + have bnd: "bounded_by (xs[x:=c]) (vs[x:= Some (interval_of c)])" by (auto intro!: bounded_by_update) from approx[OF this a] - have f_c: "interpret_floatarith ((DERIV_floatarith x ^^ 0) f) (xs[x := c]) \ { l1 .. u1 }" + have f_c: "interpret_floatarith ((DERIV_floatarith x ^^ 0) f) (xs[x := c]) \\<^sub>r ivl1" (is "?f 0 (real_of_float c) \ _") by auto have funpow_Suc[symmetric]: "(f ^^ Suc n) x = (f ^^ n) (f x)" for f :: "'a \ 'a" and n :: nat and x :: 'a by (induct n) auto from Suc.hyps[OF ate, unfolded this] obtain n - where DERIV_hyp: "\m z. \ m < n ; (z::real) \ { lx .. ux } \ \ + where DERIV_hyp: "\m z. \ m < n ; (z::real) \\<^sub>r ivlx \ \ DERIV (?f (Suc m)) z :> ?f (Suc (Suc m)) z" - and hyp: "\t \ {real_of_float lx .. real_of_float ux}. + and hyp: "\t \ set_of (real_interval ivlx). (\ i = 0.. j \ {Suc k.. j \ {Suc k.. {l2 .. u2}" + inverse (real (\ j \ {Suc k..\<^sub>r ivl2" (is "\ t \ _. ?X (Suc k) f n t \ _") by blast have DERIV: "DERIV (?f m) z :> ?f (Suc m) z" - if "m < Suc n" and bnd_z: "z \ { lx .. ux }" for m and z::real + if "m < Suc n" and bnd_z: "z \\<^sub>r ivlx" for m and z::real proof (cases m) case 0 with DERIV_floatarith[OF \x < length xs\ isDERIV_approx'[OF \bounded_by xs vs\ bnd_x bnd_z True]] show ?thesis by simp next case (Suc m') hence "m' < n" using \m < Suc n\ by auto from DERIV_hyp[OF this bnd_z] show ?thesis using Suc by simp qed have "\k i. k < i \ {k ..< i} = insert k {Suc k ..< i}" by auto hence prod_head_Suc: "\k i. \{k ..< k + Suc i} = k * \{Suc k ..< Suc k + i}" by auto have sum_move0: "\k F. sum F {0.. k. F (Suc k)) {0.. {lx .. ux}" + fix t::real assume t: "t \\<^sub>r ivlx" hence "bounded_by [xs!x] [vs!x]" using \bounded_by xs vs\[THEN bounded_byE, OF \x < length vs\] by (cases "vs!x", auto simp add: bounded_by_def) with hyp[THEN bspec, OF t] f_c - have "bounded_by [?f 0 c, ?X (Suc k) f n t, xs!x] [Some (l1, u1), Some (l2, u2), vs!x]" + have "bounded_by [?f 0 c, ?X (Suc k) f n t, xs!x] [Some ivl1, Some ivl2, vs!x]" by (auto intro!: bounded_by_Cons) from approx[OF this final, unfolded atLeastAtMost_iff[symmetric]] - have "?X (Suc k) f n t * (xs!x - real_of_float c) * inverse k + ?f 0 c \ {l .. u}" + have "?X (Suc k) f n t * (xs!x - real_of_float c) * inverse k + ?f 0 c \\<^sub>r ivl" by (auto simp add: algebra_simps) also have "?X (Suc k) f n t * (xs!x - real_of_float c) * inverse (real k) + ?f 0 c = (\ i = 0.. j \ {k.. j \ {k.. {l .. u}" . + finally have "?T \\<^sub>r ivl" . } thus ?thesis using DERIV by blast qed qed lemma prod_fact: "real (\ {1..<1 + k}) = fact (k :: nat)" by (simp add: fact_prod atLeastLessThanSuc_atLeastAtMost) lemma approx_tse: assumes "bounded_by xs vs" - and bnd_x: "vs ! x = Some (lx, ux)" - and bnd_c: "real_of_float c \ {lx .. ux}" + and bnd_x: "vs ! x = Some ivlx" + and bnd_c: "real_of_float c \\<^sub>r ivlx" and "x < length vs" and "x < length xs" - and ate: "Some (l, u) = approx_tse prec x s c 1 f vs" - shows "interpret_floatarith f xs \ {l .. u}" + and ate: "approx_tse prec x s c 1 f vs = Some ivl" + shows "interpret_floatarith f xs \\<^sub>r ivl" proof - define F where [abs_def]: "F n z = interpret_floatarith ((DERIV_floatarith x ^^ n) f) (xs[x := z])" for n z hence F0: "F 0 = (\ z. interpret_floatarith f (xs[x := z]))" by auto hence "bounded_by (xs[x := c]) vs" and "x < length vs" "x < length xs" using \bounded_by xs vs\ bnd_x bnd_c \x < length vs\ \x < length xs\ by (auto intro!: bounded_by_update_var) from approx_tse_generic[OF \bounded_by xs vs\ this bnd_x ate] obtain n - where DERIV: "\ m z. m < n \ real_of_float lx \ z \ z \ real_of_float ux \ DERIV (F m) z :> F (Suc m) z" - and hyp: "\ (t::real). t \ {lx .. ux} \ + where DERIV: "\ m z. m < n \ real_of_float (lower ivlx) \ z \ z \ real_of_float (upper ivlx) \ DERIV (F m) z :> F (Suc m) z" + and hyp: "\ (t::real). t \\<^sub>r ivlx \ (\ j = 0.. {l .. u}" (is "\ t. _ \ ?taylor t \ _") + \\<^sub>r ivl" (is "\ t. _ \ ?taylor t \ _") unfolding F_def atLeastAtMost_iff[symmetric] prod_fact - by blast + by (auto simp: set_of_eq Ball_def) - have bnd_xs: "xs ! x \ { lx .. ux }" + have bnd_xs: "xs ! x \\<^sub>r ivlx" using \bounded_by xs vs\[THEN bounded_byE, OF \x < length vs\] bnd_x by auto show ?thesis proof (cases n) case 0 thus ?thesis using hyp[OF bnd_xs] unfolding F_def by auto next case (Suc n') show ?thesis proof (cases "xs ! x = c") case True from True[symmetric] hyp[OF bnd_xs] Suc show ?thesis unfolding F_def Suc sum.atLeast_Suc_lessThan[OF zero_less_Suc] sum.shift_bounds_Suc_ivl by auto next case False - have "lx \ real_of_float c" "real_of_float c \ ux" "lx \ xs!x" "xs!x \ ux" - using Suc bnd_c \bounded_by xs vs\[THEN bounded_byE, OF \x < length vs\] bnd_x by auto + have "lower ivlx \ real_of_float c" "real_of_float c \ upper ivlx" "lower ivlx \ xs!x" "xs!x \ upper ivlx" + using Suc bnd_c \bounded_by xs vs\[THEN bounded_byE, OF \x < length vs\] bnd_x + by (auto simp: set_of_eq) from Taylor[OF zero_less_Suc, of F, OF F0 DERIV[unfolded Suc] this False] obtain t::real where t_bnd: "if xs ! x < c then xs ! x < t \ t < c else c < t \ t < xs ! x" and fl_eq: "interpret_floatarith f (xs[x := xs ! x]) = (\m = 0.. {lx .. ux}" - by (cases "xs ! x < c") auto + from t_bnd bnd_xs bnd_c have *: "t \\<^sub>r ivlx" + by (cases "xs ! x < c") (auto simp: set_of_eq) have "interpret_floatarith f (xs[x := xs ! x]) = ?taylor t" unfolding fl_eq Suc by (auto simp add: algebra_simps divide_inverse) - also have "\ \ {l .. u}" + also have "\ \\<^sub>r ivl" using * by (rule hyp) finally show ?thesis by simp qed qed qed fun approx_tse_form' where -"approx_tse_form' prec t f 0 l u cmp = - (case approx_tse prec 0 t ((l + u) * Float 1 (- 1)) 1 f [Some (l, u)] - of Some (l, u) \ cmp l u | None \ False)" | -"approx_tse_form' prec t f (Suc s) l u cmp = - (let m = (l + u) * Float 1 (- 1) - in (if approx_tse_form' prec t f s l m cmp then - approx_tse_form' prec t f s m u cmp else False))" +"approx_tse_form' prec t f 0 ivl cmp = + (case approx_tse prec 0 t (mid ivl) 1 f [Some ivl] + of Some ivl \ cmp ivl | None \ False)" | +"approx_tse_form' prec t f (Suc s) ivl cmp = + (let (ivl1, ivl2) = split_float_interval ivl + in (if approx_tse_form' prec t f s ivl1 cmp then + approx_tse_form' prec t f s ivl2 cmp else False))" lemma approx_tse_form': fixes x :: real - assumes "approx_tse_form' prec t f s l u cmp" - and "x \ {l .. u}" - shows "\l' u' ly uy. x \ {l' .. u'} \ real_of_float l \ l' \ u' \ real_of_float u \ cmp ly uy \ - approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 f [Some (l', u')] = Some (ly, uy)" + assumes "approx_tse_form' prec t f s ivl cmp" + and "x \\<^sub>r ivl" + shows "\ivl' ivly. x \\<^sub>r ivl' \ ivl' \ ivl \ cmp ivly \ + approx_tse prec 0 t (mid ivl') 1 f [Some ivl'] = Some ivly" using assms -proof (induct s arbitrary: l u) +proof (induct s arbitrary: ivl) case 0 - then obtain ly uy - where *: "approx_tse prec 0 t ((l + u) * Float 1 (- 1)) 1 f [Some (l, u)] = Some (ly, uy)" - and **: "cmp ly uy" by (auto elim!: case_optionE) + then obtain ivly + where *: "approx_tse prec 0 t (mid ivl) 1 f [Some ivl] = Some ivly" + and **: "cmp ivly" by (auto elim!: case_optionE) with 0 show ?case by auto next case (Suc s) - let ?m = "(l + u) * Float 1 (- 1)" + let ?ivl1 = "fst (split_float_interval ivl)" + let ?ivl2 = "snd (split_float_interval ivl)" from Suc.prems - have l: "approx_tse_form' prec t f s l ?m cmp" - and u: "approx_tse_form' prec t f s ?m u cmp" + have l: "approx_tse_form' prec t f s ?ivl1 cmp" + and u: "approx_tse_form' prec t f s ?ivl2 cmp" by (auto simp add: Let_def lazy_conj) - have m_l: "real_of_float l \ ?m" and m_u: "?m \ real_of_float u" - unfolding less_eq_float_def using Suc.prems by auto - with \x \ { l .. u }\ consider "x \ { l .. ?m}" | "x \ {?m .. u}" - by atomize_elim auto - thus ?case + have subintervals: "?ivl1 \ ivl" "?ivl2 \ ivl" + using mid_le + by (auto simp: less_eq_interval_def split_float_interval_bounds) + + from split_float_interval_realD[OF _ \x \\<^sub>r ivl\] consider "x \\<^sub>r ?ivl1" | "x \\<^sub>r ?ivl2" + by (auto simp: prod_eq_iff) + then show ?case proof cases case 1 from Suc.hyps[OF l this] - obtain l' u' ly uy where - "x \ {l' .. u'} \ real_of_float l \ l' \ real_of_float u' \ ?m \ cmp ly uy \ - approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 f [Some (l', u')] = Some (ly, uy)" + obtain ivl' ivly where + "x \\<^sub>r ivl' \ ivl' \ fst (split_float_interval ivl) \ cmp ivly \ + approx_tse prec 0 t (mid ivl') 1 f [Some ivl'] = Some ivly" by blast - with m_u show ?thesis - by (auto intro!: exI) + then show ?thesis + using subintervals + by (auto) next case 2 from Suc.hyps[OF u this] - obtain l' u' ly uy where - "x \ { l' .. u' } \ ?m \ real_of_float l' \ u' \ real_of_float u \ cmp ly uy \ - approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 f [Some (l', u')] = Some (ly, uy)" + obtain ivl' ivly where + "x \\<^sub>r ivl' \ ivl' \ snd (split_float_interval ivl) \ cmp ivly \ + approx_tse prec 0 t (mid ivl') 1 f [Some ivl'] = Some ivly" by blast - with m_u show ?thesis - by (auto intro!: exI) + then show ?thesis + using subintervals + by (auto) qed qed lemma approx_tse_form'_less: fixes x :: real - assumes tse: "approx_tse_form' prec t (Add a (Minus b)) s l u (\ l u. 0 < l)" - and x: "x \ {l .. u}" + assumes tse: "approx_tse_form' prec t (Add a (Minus b)) s ivl (\ ivl. 0 < lower ivl)" + and x: "x \\<^sub>r ivl" shows "interpret_floatarith b [x] < interpret_floatarith a [x]" proof - from approx_tse_form'[OF tse x] - obtain l' u' ly uy - where x': "x \ {l' .. u'}" - and "real_of_float l \ real_of_float l'" - and "real_of_float u' \ real_of_float u" and "0 < ly" - and tse: "approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 (Add a (Minus b)) [Some (l', u')] = Some (ly, uy)" + obtain ivl' ivly + where x': "x \\<^sub>r ivl'" + and "ivl' \ ivl" and "0 < lower ivly" + and tse: "approx_tse prec 0 t (mid ivl') 1 (Add a (Minus b)) [Some ivl'] = Some ivly" by blast - hence "bounded_by [x] [Some (l', u')]" + hence "bounded_by [x] [Some ivl']" by (auto simp add: bounded_by_def) - from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x' - have "ly \ interpret_floatarith a [x] - interpret_floatarith b [x]" - by auto - from order_less_le_trans[OF _ this, of 0] \0 < ly\ show ?thesis + from approx_tse[OF this _ _ _ _ tse, of ivl'] x' mid_le + have "lower ivly \ interpret_floatarith a [x] - interpret_floatarith b [x]" + by (auto simp: set_of_eq) + from order_less_le_trans[OF _ this, of 0] \0 < lower ivly\ show ?thesis by auto qed lemma approx_tse_form'_le: fixes x :: real - assumes tse: "approx_tse_form' prec t (Add a (Minus b)) s l u (\ l u. 0 \ l)" - and x: "x \ {l .. u}" + assumes tse: "approx_tse_form' prec t (Add a (Minus b)) s ivl (\ ivl. 0 \ lower ivl)" + and x: "x \\<^sub>r ivl" shows "interpret_floatarith b [x] \ interpret_floatarith a [x]" proof - from approx_tse_form'[OF tse x] - obtain l' u' ly uy - where x': "x \ {l' .. u'}" - and "l \ real_of_float l'" - and "real_of_float u' \ u" and "0 \ ly" - and tse: "approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 (Add a (Minus b)) [Some (l', u')] = Some (ly, uy)" + obtain ivl' ivly + where x': "x \\<^sub>r ivl'" + and "ivl' \ ivl" and "0 \ lower ivly" + and tse: "approx_tse prec 0 t (mid ivl') 1 (Add a (Minus b)) [Some (ivl')] = Some ivly" by blast - hence "bounded_by [x] [Some (l', u')]" by (auto simp add: bounded_by_def) - from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x' - have "ly \ interpret_floatarith a [x] - interpret_floatarith b [x]" - by auto - from order_trans[OF _ this, of 0] \0 \ ly\ show ?thesis + hence "bounded_by [x] [Some ivl']" by (auto simp add: bounded_by_def) + from approx_tse[OF this _ _ _ _ tse, of ivl'] x' mid_le + have "lower ivly \ interpret_floatarith a [x] - interpret_floatarith b [x]" + by (auto simp: set_of_eq) + from order_trans[OF _ this, of 0] \0 \ lower ivly\ show ?thesis by auto qed fun approx_tse_concl where -"approx_tse_concl prec t (Less lf rt) s l u l' u' \ - approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\ l u. 0 < l)" | -"approx_tse_concl prec t (LessEqual lf rt) s l u l' u' \ - approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\ l u. 0 \ l)" | -"approx_tse_concl prec t (AtLeastAtMost x lf rt) s l u l' u' \ - (if approx_tse_form' prec t (Add x (Minus lf)) s l u' (\ l u. 0 \ l) then - approx_tse_form' prec t (Add rt (Minus x)) s l u' (\ l u. 0 \ l) else False)" | -"approx_tse_concl prec t (Conj f g) s l u l' u' \ - approx_tse_concl prec t f s l u l' u' \ approx_tse_concl prec t g s l u l' u'" | -"approx_tse_concl prec t (Disj f g) s l u l' u' \ - approx_tse_concl prec t f s l u l' u' \ approx_tse_concl prec t g s l u l' u'" | -"approx_tse_concl _ _ _ _ _ _ _ _ \ False" +"approx_tse_concl prec t (Less lf rt) s ivl ivl' \ + approx_tse_form' prec t (Add rt (Minus lf)) s (sup ivl ivl') (\ ivl. 0 < lower ivl)" | +"approx_tse_concl prec t (LessEqual lf rt) s ivl ivl' \ + approx_tse_form' prec t (Add rt (Minus lf)) s (sup ivl ivl') (\ ivl. 0 \ lower ivl)" | +"approx_tse_concl prec t (AtLeastAtMost x lf rt) s ivl ivl' \ + (if approx_tse_form' prec t (Add x (Minus lf)) s (sup ivl ivl') (\ ivl. 0 \ lower ivl) then + approx_tse_form' prec t (Add rt (Minus x)) s (sup ivl ivl') (\ ivl. 0 \ lower ivl) else False)" | +"approx_tse_concl prec t (Conj f g) s ivl ivl' \ + approx_tse_concl prec t f s ivl ivl' \ approx_tse_concl prec t g s ivl ivl'" | +"approx_tse_concl prec t (Disj f g) s ivl ivl' \ + approx_tse_concl prec t f s ivl ivl' \ approx_tse_concl prec t g s ivl ivl'" | +"approx_tse_concl _ _ _ _ _ _ \ False" definition "approx_tse_form prec t s f = (case f of Bound x a b f \ x = Var 0 \ (case (approx prec a [None], approx prec b [None]) of - (Some (l, u), Some (l', u')) \ approx_tse_concl prec t f s l u l' u' + (Some ivl, Some ivl') \ approx_tse_concl prec t f s ivl ivl' | _ \ False) | _ \ False)" lemma approx_tse_form: assumes "approx_tse_form prec t s f" shows "interpret_form f [x]" proof (cases f) case f_def: (Bound i a b f') - with assms obtain l u l' u' - where a: "approx prec a [None] = Some (l, u)" - and b: "approx prec b [None] = Some (l', u')" + with assms obtain ivl ivl' + where a: "approx prec a [None] = Some ivl" + and b: "approx prec b [None] = Some ivl'" unfolding approx_tse_form_def by (auto elim!: case_optionE) from f_def assms have "i = Var 0" unfolding approx_tse_form_def by auto hence i: "interpret_floatarith i [x] = x" by auto { let ?f = "\z. interpret_floatarith z [x]" assume "?f i \ { ?f a .. ?f b }" - with approx[OF _ a[symmetric], of "[x]"] approx[OF _ b[symmetric], of "[x]"] - have bnd: "x \ { l .. u'}" unfolding bounded_by_def i by auto + with approx[OF _ a, of "[x]"] approx[OF _ b, of "[x]"] + have bnd: "x \\<^sub>r sup ivl ivl'" unfolding bounded_by_def i + by (auto simp: set_of_eq sup_float_def inf_float_def min_def max_def) have "interpret_form f' [x]" using assms[unfolded f_def] proof (induct f') case (Less lf rt) with a b - have "approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\ l u. 0 < l)" + have "approx_tse_form' prec t (Add rt (Minus lf)) s (sup ivl ivl') (\ ivl. 0 < lower ivl)" unfolding approx_tse_form_def by auto from approx_tse_form'_less[OF this bnd] show ?case using Less by auto next case (LessEqual lf rt) with f_def a b assms - have "approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\ l u. 0 \ l)" + have "approx_tse_form' prec t (Add rt (Minus lf)) s (sup ivl ivl') (\ ivl. 0 \ lower ivl)" unfolding approx_tse_form_def by auto from approx_tse_form'_le[OF this bnd] show ?case using LessEqual by auto next case (AtLeastAtMost x lf rt) with f_def a b assms - have "approx_tse_form' prec t (Add rt (Minus x)) s l u' (\ l u. 0 \ l)" - and "approx_tse_form' prec t (Add x (Minus lf)) s l u' (\ l u. 0 \ l)" + have "approx_tse_form' prec t (Add rt (Minus x)) s (sup ivl ivl') (\ ivl. 0 \ lower ivl)" + and "approx_tse_form' prec t (Add x (Minus lf)) s (sup ivl ivl') (\ ivl. 0 \ lower ivl)" unfolding approx_tse_form_def lazy_conj by (auto split: if_split_asm) from approx_tse_form'_le[OF this(1) bnd] approx_tse_form'_le[OF this(2) bnd] show ?case using AtLeastAtMost by auto qed (auto simp: f_def approx_tse_form_def elim!: case_optionE) } thus ?thesis unfolding f_def by auto qed (insert assms, auto simp add: approx_tse_form_def) text \\<^term>\approx_form_eval\ is only used for the {\tt value}-command.\ -fun approx_form_eval :: "nat \ form \ (float * float) option list \ (float * float) option list" where +fun approx_form_eval :: "nat \ form \ (float interval) option list \ (float interval) option list" where "approx_form_eval prec (Bound (Var n) a b f) bs = (case (approx prec a bs, approx prec b bs) - of (Some (l, _), Some (_, u)) \ approx_form_eval prec f (bs[n := Some (l, u)]) + of (Some ivl1, Some ivl2) \ approx_form_eval prec f (bs[n := Some (sup ivl1 ivl2)]) | _ \ bs)" | "approx_form_eval prec (Assign (Var n) a f) bs = (case (approx prec a bs) - of (Some (l, u)) \ approx_form_eval prec f (bs[n := Some (l, u)]) + of (Some ivl) \ approx_form_eval prec f (bs[n := Some ivl]) | _ \ bs)" | "approx_form_eval prec (Less a b) bs = bs @ [approx prec a bs, approx prec b bs]" | "approx_form_eval prec (LessEqual a b) bs = bs @ [approx prec a bs, approx prec b bs]" | "approx_form_eval prec (AtLeastAtMost x a b) bs = bs @ [approx prec x bs, approx prec a bs, approx prec b bs]" | "approx_form_eval _ _ bs = bs" subsection \Implement proof method \texttt{approximation}\ ML \ val _ = \ \Trusting the oracle \@{oracle_name "holds_by_evaluation"} signature APPROXIMATION_COMPUTATION = sig val approx_bool: Proof.context -> term -> term val approx_arith: Proof.context -> term -> term val approx_form_eval: Proof.context -> term -> term val approx_conv: Proof.context -> conv end structure Approximation_Computation : APPROXIMATION_COMPUTATION = struct fun approx_conv ctxt ct = @{computation_check terms: Trueprop "0 :: nat" "1 :: nat" "2 :: nat" "3 :: nat" Suc "(+)::nat\nat\nat" "(-)::nat\nat\nat" "(*)::nat\nat\nat" "0 :: int" "1 :: int" "2 :: int" "3 :: int" "-1 :: int" "(+)::int\int\int" "(-)::int\int\int" "(*)::int\int\int" "uminus::int\int" - "replicate :: _ \ (float \ float) option \ _" + "replicate :: _ \ (float interval) option \ _" + "interval_of::float\float interval" approx' approx_form approx_tse_form approx_form_eval datatypes: bool int nat integer "nat list" - "(float \ float) option list" + float + "(float interval) option list" floatarith form } ctxt ct fun term_of_bool true = \<^term>\True\ | term_of_bool false = \<^term>\False\; val mk_int = HOLogic.mk_number \<^typ>\int\ o @{code integer_of_int}; fun term_of_float (@{code Float} (k, l)) = \<^term>\Float\ $ mk_int k $ mk_int l; - fun term_of_float_float_option NONE = \<^term>\None :: (float \ float) option\ - | term_of_float_float_option (SOME ff) = \<^term>\Some :: float \ float \ _\ - $ HOLogic.mk_prod (apply2 term_of_float ff); + fun term_of_float_interval x = @{term "Interval::_\float interval"} $ + HOLogic.mk_prod + (apply2 term_of_float (@{code lowerF} x, @{code upperF} x)) - val term_of_float_float_option_list = - HOLogic.mk_list \<^typ>\(float \ float) option\ o map term_of_float_float_option; + fun term_of_float_interval_option NONE = \<^term>\None :: (float interval) option\ + | term_of_float_interval_option (SOME ff) = \<^term>\Some :: float interval \ _\ + $ (term_of_float_interval ff); + + val term_of_float_interval_option_list = + HOLogic.mk_list \<^typ>\(float interval) option\ o map term_of_float_interval_option; val approx_bool = @{computation bool} (fn _ => fn x => case x of SOME b => term_of_bool b | NONE => error "Computation approx_bool failed.") - val approx_arith = @{computation "(float \ float) option"} - (fn _ => fn x => case x of SOME ffo => term_of_float_float_option ffo + val approx_arith = @{computation "float interval option"} + (fn _ => fn x => case x of SOME ffo => term_of_float_interval_option ffo | NONE => error "Computation approx_arith failed.") - val approx_form_eval = @{computation "(float \ float) option list"} - (fn _ => fn x => case x of SOME ffos => term_of_float_float_option_list ffos + val approx_form_eval = @{computation "float interval option list"} + (fn _ => fn x => case x of SOME ffos => term_of_float_interval_option_list ffos | NONE => error "Computation approx_form_eval failed.") end \ lemma intervalE: "a \ x \ x \ b \ \ x \ { a .. b } \ P\ \ P" by auto lemma meta_eqE: "x \ a \ \ x = a \ P\ \ P" by auto named_theorems approximation_preproc lemma approximation_preproc_floatarith[approximation_preproc]: "0 = real_of_float 0" "1 = real_of_float 1" "0 = Float 0 0" "1 = Float 1 0" "numeral a = Float (numeral a) 0" "numeral a = real_of_float (numeral a)" "x - y = x + - y" "x / y = x * inverse y" "ceiling x = - floor (- x)" "log x y = ln y * inverse (ln x)" "sin x = cos (pi / 2 - x)" "tan x = sin x / cos x" by (simp_all add: inverse_eq_divide ceiling_def log_def sin_cos_eq tan_def real_of_float_eq) lemma approximation_preproc_int[approximation_preproc]: "real_of_int 0 = real_of_float 0" "real_of_int 1 = real_of_float 1" "real_of_int (i + j) = real_of_int i + real_of_int j" "real_of_int (- i) = - real_of_int i" "real_of_int (i - j) = real_of_int i - real_of_int j" "real_of_int (i * j) = real_of_int i * real_of_int j" "real_of_int (i div j) = real_of_int (floor (real_of_int i / real_of_int j))" "real_of_int (min i j) = min (real_of_int i) (real_of_int j)" "real_of_int (max i j) = max (real_of_int i) (real_of_int j)" "real_of_int (abs i) = abs (real_of_int i)" "real_of_int (i ^ n) = (real_of_int i) ^ n" "real_of_int (numeral a) = real_of_float (numeral a)" "i mod j = i - i div j * j" "i = j \ real_of_int i = real_of_int j" "i \ j \ real_of_int i \ real_of_int j" "i < j \ real_of_int i < real_of_int j" "i \ {j .. k} \ real_of_int i \ {real_of_int j .. real_of_int k}" by (simp_all add: floor_divide_of_int_eq minus_div_mult_eq_mod [symmetric]) lemma approximation_preproc_nat[approximation_preproc]: "real 0 = real_of_float 0" "real 1 = real_of_float 1" "real (i + j) = real i + real j" "real (i - j) = max (real i - real j) 0" "real (i * j) = real i * real j" "real (i div j) = real_of_int (floor (real i / real j))" "real (min i j) = min (real i) (real j)" "real (max i j) = max (real i) (real j)" "real (i ^ n) = (real i) ^ n" "real (numeral a) = real_of_float (numeral a)" "i mod j = i - i div j * j" "n = m \ real n = real m" "n \ m \ real n \ real m" "n < m \ real n < real m" "n \ {m .. l} \ real n \ {real m .. real l}" by (simp_all add: real_div_nat_eq_floor_of_divide minus_div_mult_eq_mod [symmetric]) ML_file \approximation.ML\ method_setup approximation = \ let val free = Args.context -- Args.term >> (fn (_, Free (n, _)) => n | (ctxt, t) => error ("Bad free variable: " ^ Syntax.string_of_term ctxt t)); in Scan.lift Parse.nat -- Scan.optional (Scan.lift (Args.$$$ "splitting" |-- Args.colon) |-- Parse.and_list' (free --| Scan.lift (Args.$$$ "=") -- Scan.lift Parse.nat)) [] -- Scan.option (Scan.lift (Args.$$$ "taylor" |-- Args.colon) |-- (free |-- Scan.lift (Args.$$$ "=") |-- Scan.lift Parse.nat)) >> (fn ((prec, splitting), taylor) => fn ctxt => SIMPLE_METHOD' (Approximation.approximation_tac prec splitting taylor ctxt)) end \ "real number approximation" section "Quickcheck Generator" lemma approximation_preproc_push_neg[approximation_preproc]: fixes a b::real shows "\ (a < b) \ b \ a" "\ (a \ b) \ b < a" "\ (a = b) \ b < a \ a < b" "\ (p \ q) \ \ p \ \ q" "\ (p \ q) \ \ p \ \ q" "\ \ q \ q" by auto ML_file \approximation_generator.ML\ setup "Approximation_Generator.setup" section "Avoid pollution of name space" bundle floatarith_notation begin notation Add ("Add") notation Minus ("Minus") notation Mult ("Mult") notation Inverse ("Inverse") notation Cos ("Cos") notation Arctan ("Arctan") notation Abs ("Abs") notation Max ("Max") notation Min ("Min") notation Pi ("Pi") notation Sqrt ("Sqrt") notation Exp ("Exp") notation Powr ("Powr") notation Ln ("Ln") notation Power ("Power") notation Floor ("Floor") notation Var ("Var") notation Num ("Num") end bundle no_floatarith_notation begin no_notation Add ("Add") no_notation Minus ("Minus") no_notation Mult ("Mult") no_notation Inverse ("Inverse") no_notation Cos ("Cos") no_notation Arctan ("Arctan") no_notation Abs ("Abs") no_notation Max ("Max") no_notation Min ("Min") no_notation Pi ("Pi") no_notation Sqrt ("Sqrt") no_notation Exp ("Exp") no_notation Powr ("Powr") no_notation Ln ("Ln") no_notation Power ("Power") no_notation Floor ("Floor") no_notation Var ("Var") no_notation Num ("Num") end hide_const (open) Add Minus Mult Inverse Cos Arctan Abs Max Min Pi Sqrt Exp Powr Ln Power Floor Var Num end diff --git a/src/HOL/Decision_Procs/Approximation_Bounds.thy b/src/HOL/Decision_Procs/Approximation_Bounds.thy --- a/src/HOL/Decision_Procs/Approximation_Bounds.thy +++ b/src/HOL/Decision_Procs/Approximation_Bounds.thy @@ -1,3036 +1,3047 @@ (* Author: Johannes Hoelzl, TU Muenchen Coercions removed by Dmitriy Traytel This file contains only general material about computing lower/upper bounds on real functions. Approximation.thy contains the actual approximation algorithm and the approximation oracle. This is in order to make a clear separation between "morally immaculate" material about upper/lower bounds and the trusted oracle/reflection. *) theory Approximation_Bounds imports Complex_Main "HOL-Library.Interval_Float" Dense_Linear_Order begin declare powr_neg_one [simp] declare powr_neg_numeral [simp] context includes interval.lifting begin section "Horner Scheme" subsection \Define auxiliary helper \horner\ function\ primrec horner :: "(nat \ nat) \ (nat \ nat \ nat) \ nat \ nat \ nat \ real \ real" where "horner F G 0 i k x = 0" | "horner F G (Suc n) i k x = 1 / k - x * horner F G n (F i) (G i k) x" lemma horner_schema': fixes x :: real and a :: "nat \ real" shows "a 0 - x * (\ i=0.. i=0..i. - (x * ((-1)^i * a (Suc i) * x ^ i)) = (-1)^(Suc i) * a (Suc i) * x ^ (Suc i)" by auto show ?thesis unfolding sum_distrib_left shift_pow uminus_add_conv_diff [symmetric] sum_negf[symmetric] sum.atLeast_Suc_lessThan[OF zero_less_Suc] sum.reindex[OF inj_Suc, unfolded comp_def, symmetric, of "\ n. (-1)^n *a n * x^n"] by auto qed lemma horner_schema: fixes f :: "nat \ nat" and G :: "nat \ nat \ nat" and F :: "nat \ nat" assumes f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" shows "horner F G n ((F ^^ j') s) (f j') x = (\ j = 0..< n. (- 1) ^ j * (1 / (f (j' + j))) * x ^ j)" proof (induct n arbitrary: j') case 0 then show ?case by auto next case (Suc n) show ?case unfolding horner.simps Suc[where j'="Suc j'", unfolded funpow.simps comp_def f_Suc] using horner_schema'[of "\ j. 1 / (f (j' + j))"] by auto qed lemma horner_bounds': fixes lb :: "nat \ nat \ nat \ float \ float" and ub :: "nat \ nat \ nat \ float \ float" assumes "0 \ real_of_float x" and f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" and lb_0: "\ i k x. lb 0 i k x = 0" and lb_Suc: "\ n i k x. lb (Suc n) i k x = float_plus_down prec (lapprox_rat prec 1 k) (- float_round_up prec (x * (ub n (F i) (G i k) x)))" and ub_0: "\ i k x. ub 0 i k x = 0" and ub_Suc: "\ n i k x. ub (Suc n) i k x = float_plus_up prec (rapprox_rat prec 1 k) (- float_round_down prec (x * (lb n (F i) (G i k) x)))" shows "(lb n ((F ^^ j') s) (f j') x) \ horner F G n ((F ^^ j') s) (f j') x \ horner F G n ((F ^^ j') s) (f j') x \ (ub n ((F ^^ j') s) (f j') x)" (is "?lb n j' \ ?horner n j' \ ?horner n j' \ ?ub n j'") proof (induct n arbitrary: j') case 0 thus ?case unfolding lb_0 ub_0 horner.simps by auto next case (Suc n) thus ?case using lapprox_rat[of prec 1 "f j'"] using rapprox_rat[of 1 "f j'" prec] Suc[where j'="Suc j'"] \0 \ real_of_float x\ by (auto intro!: add_mono mult_left_mono float_round_down_le float_round_up_le order_trans[OF add_mono[OF _ float_plus_down_le]] order_trans[OF _ add_mono[OF _ float_plus_up_le]] simp add: lb_Suc ub_Suc field_simps f_Suc) qed subsection "Theorems for floating point functions implementing the horner scheme" text \ Here \<^term_type>\f :: nat \ nat\ is the sequence defining the Taylor series, the coefficients are all alternating and reciprocs. We use \<^term>\G\ and \<^term>\F\ to describe the computation of \<^term>\f\. \ lemma horner_bounds: fixes F :: "nat \ nat" and G :: "nat \ nat \ nat" assumes "0 \ real_of_float x" and f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" and lb_0: "\ i k x. lb 0 i k x = 0" and lb_Suc: "\ n i k x. lb (Suc n) i k x = float_plus_down prec (lapprox_rat prec 1 k) (- float_round_up prec (x * (ub n (F i) (G i k) x)))" and ub_0: "\ i k x. ub 0 i k x = 0" and ub_Suc: "\ n i k x. ub (Suc n) i k x = float_plus_up prec (rapprox_rat prec 1 k) (- float_round_down prec (x * (lb n (F i) (G i k) x)))" shows "(lb n ((F ^^ j') s) (f j') x) \ (\j=0..j=0.. (ub n ((F ^^ j') s) (f j') x)" (is "?ub") proof - have "?lb \ ?ub" using horner_bounds'[where lb=lb, OF \0 \ real_of_float x\ f_Suc lb_0 lb_Suc ub_0 ub_Suc] unfolding horner_schema[where f=f, OF f_Suc] by simp thus "?lb" and "?ub" by auto qed lemma horner_bounds_nonpos: fixes F :: "nat \ nat" and G :: "nat \ nat \ nat" assumes "real_of_float x \ 0" and f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" and lb_0: "\ i k x. lb 0 i k x = 0" and lb_Suc: "\ n i k x. lb (Suc n) i k x = float_plus_down prec (lapprox_rat prec 1 k) (float_round_down prec (x * (ub n (F i) (G i k) x)))" and ub_0: "\ i k x. ub 0 i k x = 0" and ub_Suc: "\ n i k x. ub (Suc n) i k x = float_plus_up prec (rapprox_rat prec 1 k) (float_round_up prec (x * (lb n (F i) (G i k) x)))" shows "(lb n ((F ^^ j') s) (f j') x) \ (\j=0..j=0.. (ub n ((F ^^ j') s) (f j') x)" (is "?ub") proof - have diff_mult_minus: "x - y * z = x + - y * z" for x y z :: float by simp have sum_eq: "(\j=0..j = 0.. real_of_float (-x)" using assms by auto from horner_bounds[where G=G and F=F and f=f and s=s and prec=prec and lb="\ n i k x. lb n i k (-x)" and ub="\ n i k x. ub n i k (-x)", unfolded lb_Suc ub_Suc diff_mult_minus, OF this f_Suc lb_0 _ ub_0 _] show "?lb" and "?ub" unfolding minus_minus sum_eq by (auto simp: minus_float_round_up_eq minus_float_round_down_eq) qed subsection \Selectors for next even or odd number\ text \ The horner scheme computes alternating series. To get the upper and lower bounds we need to guarantee to access a even or odd member. To do this we use \<^term>\get_odd\ and \<^term>\get_even\. \ definition get_odd :: "nat \ nat" where "get_odd n = (if odd n then n else (Suc n))" definition get_even :: "nat \ nat" where "get_even n = (if even n then n else (Suc n))" lemma get_odd[simp]: "odd (get_odd n)" unfolding get_odd_def by (cases "odd n") auto lemma get_even[simp]: "even (get_even n)" unfolding get_even_def by (cases "even n") auto lemma get_odd_ex: "\ k. Suc k = get_odd n \ odd (Suc k)" by (auto simp: get_odd_def odd_pos intro!: exI[of _ "n - 1"]) lemma get_even_double: "\i. get_even n = 2 * i" using get_even by (blast elim: evenE) lemma get_odd_double: "\i. get_odd n = 2 * i + 1" using get_odd by (blast elim: oddE) section "Power function" definition float_power_bnds :: "nat \ nat \ float \ float \ float * float" where "float_power_bnds prec n l u = (if 0 < l then (power_down_fl prec l n, power_up_fl prec u n) else if odd n then (- power_up_fl prec \l\ n, if u < 0 then - power_down_fl prec \u\ n else power_up_fl prec u n) else if u < 0 then (power_down_fl prec \u\ n, power_up_fl prec \l\ n) else (0, power_up_fl prec (max \l\ \u\) n))" lemma le_minus_power_downI: "0 \ x \ x ^ n \ - a \ a \ - power_down prec x n" by (subst le_minus_iff) (auto intro: power_down_le power_mono_odd) lemma float_power_bnds: "(l1, u1) = float_power_bnds prec n l u \ x \ {l .. u} \ (x::real) ^ n \ {l1..u1}" by (auto simp: float_power_bnds_def max_def real_power_up_fl real_power_down_fl minus_le_iff split: if_split_asm intro!: power_up_le power_down_le le_minus_power_downI intro: power_mono_odd power_mono power_mono_even zero_le_even_power) lemma bnds_power: "\(x::real) l u. (l1, u1) = float_power_bnds prec n l u \ x \ {l .. u} \ l1 \ x ^ n \ x ^ n \ u1" using float_power_bnds by auto lift_definition power_float_interval :: "nat \ nat \ float interval \ float interval" is "\p n (l, u). float_power_bnds p n l u" using float_power_bnds by (auto simp: bnds_power dest!: float_power_bnds[OF sym]) lemma lower_power_float_interval: "lower (power_float_interval p n x) = fst (float_power_bnds p n (lower x) (upper x))" by transfer auto lemma upper_power_float_interval: "upper (power_float_interval p n x) = snd (float_power_bnds p n (lower x) (upper x))" by transfer auto lemma power_float_intervalI: "x \\<^sub>r X \ x ^ n \\<^sub>r power_float_interval p n X" using float_power_bnds[OF prod.collapse] by (auto simp: set_of_eq lower_power_float_interval upper_power_float_interval) lemma power_float_interval_mono: "set_of (power_float_interval prec n A) \ set_of (power_float_interval prec n B)" if "set_of A \ set_of B" proof - define la where "la = real_of_float (lower A)" define ua where "ua = real_of_float (upper A)" define lb where "lb = real_of_float (lower B)" define ub where "ub = real_of_float (upper B)" have ineqs: "lb \ la" "la \ ua" "ua \ ub" "lb \ ub" using that lower_le_upper[of A] lower_le_upper[of B] by (auto simp: la_def ua_def lb_def ub_def set_of_eq) show ?thesis using ineqs by (simp add: set_of_subset_iff float_power_bnds_def max_def power_down_fl.rep_eq power_up_fl.rep_eq lower_power_float_interval upper_power_float_interval la_def[symmetric] ua_def[symmetric] lb_def[symmetric] ub_def[symmetric]) (auto intro!: power_down_mono power_up_mono intro: order_trans[where y=0]) qed section \Approximation utility functions\ definition bnds_mult :: "nat \ float \ float \ float \ float \ float \ float" where "bnds_mult prec a1 a2 b1 b2 = (float_plus_down prec (nprt a1 * pprt b2) (float_plus_down prec (nprt a2 * nprt b2) (float_plus_down prec (pprt a1 * pprt b1) (pprt a2 * nprt b1))), float_plus_up prec (pprt a2 * pprt b2) (float_plus_up prec (pprt a1 * nprt b2) (float_plus_up prec (nprt a2 * pprt b1) (nprt a1 * nprt b1))))" lemma bnds_mult: fixes prec :: nat and a1 aa2 b1 b2 :: float assumes "(l, u) = bnds_mult prec a1 a2 b1 b2" assumes "a \ {real_of_float a1..real_of_float a2}" assumes "b \ {real_of_float b1..real_of_float b2}" shows "a * b \ {real_of_float l..real_of_float u}" proof - from assms have "real_of_float l \ a * b" by (intro order.trans[OF _ mult_ge_prts[of a1 a a2 b1 b b2]]) (auto simp: bnds_mult_def intro!: float_plus_down_le) moreover from assms have "real_of_float u \ a * b" by (intro order.trans[OF mult_le_prts[of a1 a a2 b1 b b2]]) (auto simp: bnds_mult_def intro!: float_plus_up_le) ultimately show ?thesis by simp qed lift_definition mult_float_interval::"nat \ float interval \ float interval \ float interval" is "\prec. \(a1, a2). \(b1, b2). bnds_mult prec a1 a2 b1 b2" by (auto dest!: bnds_mult[OF sym]) lemma lower_mult_float_interval: "lower (mult_float_interval p x y) = fst (bnds_mult p (lower x) (upper x) (lower y) (upper y))" by transfer auto lemma upper_mult_float_interval: "upper (mult_float_interval p x y) = snd (bnds_mult p (lower x) (upper x) (lower y) (upper y))" by transfer auto lemma mult_float_interval: "set_of (real_interval A) * set_of (real_interval B) \ set_of (real_interval (mult_float_interval prec A B))" proof - let ?bm = "bnds_mult prec (lower A) (upper A) (lower B) (upper B)" show ?thesis using bnds_mult[of "fst ?bm" "snd ?bm", simplified, OF refl] by (auto simp: set_of_eq set_times_def upper_mult_float_interval lower_mult_float_interval) qed lemma mult_float_intervalI: "x * y \\<^sub>r mult_float_interval prec A B" if "x \\<^sub>i real_interval A" "y \\<^sub>i real_interval B" using mult_float_interval[of A B] that by (auto simp: ) lemma mult_float_interval_mono: "set_of (mult_float_interval prec A B) \ set_of (mult_float_interval prec X Y)" if "set_of A \ set_of X" "set_of B \ set_of Y" using that apply transfer unfolding bnds_mult_def atLeastatMost_subset_iff float_plus_down.rep_eq float_plus_up.rep_eq by (auto simp: float_plus_down.rep_eq float_plus_up.rep_eq mult_float_mono1 mult_float_mono2) definition map_bnds :: "(nat \ float \ float) \ (nat \ float \ float) \ nat \ (float \ float) \ (float \ float)" where "map_bnds lb ub prec = (\(l,u). (lb prec l, ub prec u))" lemma map_bnds: assumes "(lf, uf) = map_bnds lb ub prec (l, u)" assumes "mono f" assumes "x \ {real_of_float l..real_of_float u}" assumes "real_of_float (lb prec l) \ f (real_of_float l)" assumes "real_of_float (ub prec u) \ f (real_of_float u)" shows "f x \ {real_of_float lf..real_of_float uf}" proof - from assms have "real_of_float lf = real_of_float (lb prec l)" by (simp add: map_bnds_def) also have "real_of_float (lb prec l) \ f (real_of_float l)" by fact also from assms have "\ \ f x" by (intro monoD[OF \mono f\]) auto finally have lf: "real_of_float lf \ f x" . from assms have "f x \ f (real_of_float u)" by (intro monoD[OF \mono f\]) auto also have "\ \ real_of_float (ub prec u)" by fact also from assms have "\ = real_of_float uf" by (simp add: map_bnds_def) finally have uf: "f x \ real_of_float uf" . from lf uf show ?thesis by simp qed section "Square root" text \ The square root computation is implemented as newton iteration. As first first step we use the nearest power of two greater than the square root. \ fun sqrt_iteration :: "nat \ nat \ float \ float" where "sqrt_iteration prec 0 x = Float 1 ((bitlen \mantissa x\ + exponent x) div 2 + 1)" | "sqrt_iteration prec (Suc m) x = (let y = sqrt_iteration prec m x in Float 1 (- 1) * float_plus_up prec y (float_divr prec x y))" lemma compute_sqrt_iteration_base[code]: shows "sqrt_iteration prec n (Float m e) = (if n = 0 then Float 1 ((if m = 0 then 0 else bitlen \m\ + e) div 2 + 1) else (let y = sqrt_iteration prec (n - 1) (Float m e) in Float 1 (- 1) * float_plus_up prec y (float_divr prec (Float m e) y)))" using bitlen_Float by (cases n) simp_all function ub_sqrt lb_sqrt :: "nat \ float \ float" where "ub_sqrt prec x = (if 0 < x then (sqrt_iteration prec prec x) else if x < 0 then - lb_sqrt prec (- x) else 0)" | "lb_sqrt prec x = (if 0 < x then (float_divl prec x (sqrt_iteration prec prec x)) else if x < 0 then - ub_sqrt prec (- x) else 0)" by pat_completeness auto termination by (relation "measure (\ v. let (prec, x) = case_sum id id v in (if x < 0 then 1 else 0))", auto) declare lb_sqrt.simps[simp del] declare ub_sqrt.simps[simp del] lemma sqrt_ub_pos_pos_1: assumes "sqrt x < b" and "0 < b" and "0 < x" shows "sqrt x < (b + x / b)/2" proof - from assms have "0 < (b - sqrt x)\<^sup>2 " by simp also have "\ = b\<^sup>2 - 2 * b * sqrt x + (sqrt x)\<^sup>2" by algebra also have "\ = b\<^sup>2 - 2 * b * sqrt x + x" using assms by simp finally have "0 < b\<^sup>2 - 2 * b * sqrt x + x" . hence "0 < b / 2 - sqrt x + x / (2 * b)" using assms by (simp add: field_simps power2_eq_square) thus ?thesis by (simp add: field_simps) qed lemma sqrt_iteration_bound: assumes "0 < real_of_float x" shows "sqrt x < sqrt_iteration prec n x" proof (induct n) case 0 show ?case proof (cases x) case (Float m e) hence "0 < m" using assms by (auto simp: algebra_split_simps) hence "0 < sqrt m" by auto have int_nat_bl: "(nat (bitlen m)) = bitlen m" using bitlen_nonneg by auto have "x = (m / 2^nat (bitlen m)) * 2 powr (e + (nat (bitlen m)))" unfolding Float by (auto simp: powr_realpow[symmetric] field_simps powr_add) also have "\ < 1 * 2 powr (e + nat (bitlen m))" proof (rule mult_strict_right_mono, auto) show "m < 2^nat (bitlen m)" using bitlen_bounds[OF \0 < m\, THEN conjunct2] unfolding of_int_less_iff[of m, symmetric] by auto qed finally have "sqrt x < sqrt (2 powr (e + bitlen m))" unfolding int_nat_bl by auto also have "\ \ 2 powr ((e + bitlen m) div 2 + 1)" proof - let ?E = "e + bitlen m" have E_mod_pow: "2 powr (?E mod 2) < 4" proof (cases "?E mod 2 = 1") case True thus ?thesis by auto next case False have "0 \ ?E mod 2" by auto have "?E mod 2 < 2" by auto from this[THEN zless_imp_add1_zle] have "?E mod 2 \ 0" using False by auto from xt1(5)[OF \0 \ ?E mod 2\ this] show ?thesis by auto qed hence "sqrt (2 powr (?E mod 2)) < sqrt (2 * 2)" by (intro real_sqrt_less_mono) auto hence E_mod_pow: "sqrt (2 powr (?E mod 2)) < 2" by auto have E_eq: "2 powr ?E = 2 powr (?E div 2 + ?E div 2 + ?E mod 2)" by auto have "sqrt (2 powr ?E) = sqrt (2 powr (?E div 2) * 2 powr (?E div 2) * 2 powr (?E mod 2))" unfolding E_eq unfolding powr_add[symmetric] by (metis of_int_add) also have "\ = 2 powr (?E div 2) * sqrt (2 powr (?E mod 2))" unfolding real_sqrt_mult[of _ "2 powr (?E mod 2)"] real_sqrt_abs2 by auto also have "\ < 2 powr (?E div 2) * 2 powr 1" by (rule mult_strict_left_mono) (auto intro: E_mod_pow) also have "\ = 2 powr (?E div 2 + 1)" unfolding add.commute[of _ 1] powr_add[symmetric] by simp finally show ?thesis by auto qed finally show ?thesis using \0 < m\ unfolding Float by (subst compute_sqrt_iteration_base) (simp add: ac_simps) qed next case (Suc n) let ?b = "sqrt_iteration prec n x" have "0 < sqrt x" using \0 < real_of_float x\ by auto also have "\ < real_of_float ?b" using Suc . finally have "sqrt x < (?b + x / ?b)/2" using sqrt_ub_pos_pos_1[OF Suc _ \0 < real_of_float x\] by auto also have "\ \ (?b + (float_divr prec x ?b))/2" by (rule divide_right_mono, auto simp add: float_divr) also have "\ = (Float 1 (- 1)) * (?b + (float_divr prec x ?b))" by simp also have "\ \ (Float 1 (- 1)) * (float_plus_up prec ?b (float_divr prec x ?b))" by (auto simp add: algebra_simps float_plus_up_le) finally show ?case unfolding sqrt_iteration.simps Let_def distrib_left . qed lemma sqrt_iteration_lower_bound: assumes "0 < real_of_float x" shows "0 < real_of_float (sqrt_iteration prec n x)" (is "0 < ?sqrt") proof - have "0 < sqrt x" using assms by auto also have "\ < ?sqrt" using sqrt_iteration_bound[OF assms] . finally show ?thesis . qed lemma lb_sqrt_lower_bound: assumes "0 \ real_of_float x" shows "0 \ real_of_float (lb_sqrt prec x)" proof (cases "0 < x") case True hence "0 < real_of_float x" and "0 \ x" using \0 \ real_of_float x\ by auto hence "0 < sqrt_iteration prec prec x" using sqrt_iteration_lower_bound by auto hence "0 \ real_of_float (float_divl prec x (sqrt_iteration prec prec x))" using float_divl_lower_bound[OF \0 \ x\] unfolding less_eq_float_def by auto thus ?thesis unfolding lb_sqrt.simps using True by auto next case False with \0 \ real_of_float x\ have "real_of_float x = 0" by auto thus ?thesis unfolding lb_sqrt.simps by auto qed lemma bnds_sqrt': "sqrt x \ {(lb_sqrt prec x) .. (ub_sqrt prec x)}" proof - have lb: "lb_sqrt prec x \ sqrt x" if "0 < x" for x :: float proof - from that have "0 < real_of_float x" and "0 \ real_of_float x" by auto hence sqrt_gt0: "0 < sqrt x" by auto hence sqrt_ub: "sqrt x < sqrt_iteration prec prec x" using sqrt_iteration_bound by auto have "(float_divl prec x (sqrt_iteration prec prec x)) \ x / (sqrt_iteration prec prec x)" by (rule float_divl) also have "\ < x / sqrt x" by (rule divide_strict_left_mono[OF sqrt_ub \0 < real_of_float x\ mult_pos_pos[OF order_less_trans[OF sqrt_gt0 sqrt_ub] sqrt_gt0]]) also have "\ = sqrt x" unfolding inverse_eq_iff_eq[of _ "sqrt x", symmetric] sqrt_divide_self_eq[OF \0 \ real_of_float x\, symmetric] by auto finally show ?thesis unfolding lb_sqrt.simps if_P[OF \0 < x\] by auto qed have ub: "sqrt x \ ub_sqrt prec x" if "0 < x" for x :: float proof - from that have "0 < real_of_float x" by auto hence "0 < sqrt x" by auto hence "sqrt x < sqrt_iteration prec prec x" using sqrt_iteration_bound by auto then show ?thesis unfolding ub_sqrt.simps if_P[OF \0 < x\] by auto qed show ?thesis using lb[of "-x"] ub[of "-x"] lb[of x] ub[of x] by (auto simp add: lb_sqrt.simps ub_sqrt.simps real_sqrt_minus) qed lemma bnds_sqrt: "\(x::real) lx ux. (l, u) = (lb_sqrt prec lx, ub_sqrt prec ux) \ x \ {lx .. ux} \ l \ sqrt x \ sqrt x \ u" proof ((rule allI) +, rule impI, erule conjE, rule conjI) fix x :: real fix lx ux assume "(l, u) = (lb_sqrt prec lx, ub_sqrt prec ux)" and x: "x \ {lx .. ux}" hence l: "l = lb_sqrt prec lx " and u: "u = ub_sqrt prec ux" by auto have "sqrt lx \ sqrt x" using x by auto from order_trans[OF _ this] show "l \ sqrt x" unfolding l using bnds_sqrt'[of lx prec] by auto have "sqrt x \ sqrt ux" using x by auto from order_trans[OF this] show "sqrt x \ u" unfolding u using bnds_sqrt'[of ux prec] by auto qed lift_definition sqrt_float_interval::"nat \ float interval \ float interval" is "\prec. \(lx, ux). (lb_sqrt prec lx, ub_sqrt prec ux)" using bnds_sqrt' by auto (meson order_trans real_sqrt_le_iff) lemma lower_float_interval: "lower (sqrt_float_interval prec X) = lb_sqrt prec (lower X)" by transfer auto lemma upper_float_interval: "upper (sqrt_float_interval prec X) = ub_sqrt prec (upper X)" by transfer auto lemma sqrt_float_interval: "sqrt ` set_of (real_interval X) \ set_of (real_interval (sqrt_float_interval prec X))" using bnds_sqrt by (auto simp: set_of_eq lower_float_interval upper_float_interval) lemma sqrt_float_intervalI: "sqrt x \\<^sub>r sqrt_float_interval p X" if "x \\<^sub>r X" using sqrt_float_interval[of X p] that by auto section "Arcus tangens and \" subsection "Compute arcus tangens series" text \ As first step we implement the computation of the arcus tangens series. This is only valid in the range \<^term>\{-1 :: real .. 1}\. This is used to compute \ and then the entire arcus tangens. \ fun ub_arctan_horner :: "nat \ nat \ nat \ float \ float" and lb_arctan_horner :: "nat \ nat \ nat \ float \ float" where "ub_arctan_horner prec 0 k x = 0" | "ub_arctan_horner prec (Suc n) k x = float_plus_up prec (rapprox_rat prec 1 k) (- float_round_down prec (x * (lb_arctan_horner prec n (k + 2) x)))" | "lb_arctan_horner prec 0 k x = 0" | "lb_arctan_horner prec (Suc n) k x = float_plus_down prec (lapprox_rat prec 1 k) (- float_round_up prec (x * (ub_arctan_horner prec n (k + 2) x)))" lemma arctan_0_1_bounds': assumes "0 \ real_of_float y" "real_of_float y \ 1" and "even n" shows "arctan (sqrt y) \ {(sqrt y * lb_arctan_horner prec n 1 y) .. (sqrt y * ub_arctan_horner prec (Suc n) 1 y)}" proof - let ?c = "\i. (- 1) ^ i * (1 / (i * 2 + (1::nat)) * sqrt y ^ (i * 2 + 1))" let ?S = "\n. \ i=0.. sqrt y" using assms by auto have "sqrt y \ 1" using assms by auto from \even n\ obtain m where "2 * m = n" by (blast elim: evenE) have "arctan (sqrt y) \ { ?S n .. ?S (Suc n) }" proof (cases "sqrt y = 0") case True then show ?thesis by simp next case False hence "0 < sqrt y" using \0 \ sqrt y\ by auto hence prem: "0 < 1 / (0 * 2 + (1::nat)) * sqrt y ^ (0 * 2 + 1)" by auto have "\ sqrt y \ \ 1" using \0 \ sqrt y\ \sqrt y \ 1\ by auto from mp[OF summable_Leibniz(2)[OF zeroseq_arctan_series[OF this] monoseq_arctan_series[OF this]] prem, THEN spec, of m, unfolded \2 * m = n\] show ?thesis unfolding arctan_series[OF \\ sqrt y \ \ 1\] Suc_eq_plus1 atLeast0LessThan . qed note arctan_bounds = this[unfolded atLeastAtMost_iff] have F: "\n. 2 * Suc n + 1 = 2 * n + 1 + 2" by auto note bounds = horner_bounds[where s=1 and f="\i. 2 * i + 1" and j'=0 and lb="\n i k x. lb_arctan_horner prec n k x" and ub="\n i k x. ub_arctan_horner prec n k x", OF \0 \ real_of_float y\ F lb_arctan_horner.simps ub_arctan_horner.simps] have "(sqrt y * lb_arctan_horner prec n 1 y) \ arctan (sqrt y)" proof - have "(sqrt y * lb_arctan_horner prec n 1 y) \ ?S n" using bounds(1) \0 \ sqrt y\ apply (simp only: power_add power_one_right mult.assoc[symmetric] sum_distrib_right[symmetric]) apply (simp only: mult.commute[where 'a=real] mult.commute[of _ "2::nat"] power_mult) apply (auto intro!: mult_left_mono) done also have "\ \ arctan (sqrt y)" using arctan_bounds .. finally show ?thesis . qed moreover have "arctan (sqrt y) \ (sqrt y * ub_arctan_horner prec (Suc n) 1 y)" proof - have "arctan (sqrt y) \ ?S (Suc n)" using arctan_bounds .. also have "\ \ (sqrt y * ub_arctan_horner prec (Suc n) 1 y)" using bounds(2)[of "Suc n"] \0 \ sqrt y\ apply (simp only: power_add power_one_right mult.assoc[symmetric] sum_distrib_right[symmetric]) apply (simp only: mult.commute[where 'a=real] mult.commute[of _ "2::nat"] power_mult) apply (auto intro!: mult_left_mono) done finally show ?thesis . qed ultimately show ?thesis by auto qed lemma arctan_0_1_bounds: assumes "0 \ real_of_float y" "real_of_float y \ 1" shows "arctan (sqrt y) \ {(sqrt y * lb_arctan_horner prec (get_even n) 1 y) .. (sqrt y * ub_arctan_horner prec (get_odd n) 1 y)}" using arctan_0_1_bounds'[OF assms, of n prec] arctan_0_1_bounds'[OF assms, of "n + 1" prec] arctan_0_1_bounds'[OF assms, of "n - 1" prec] by (auto simp: get_even_def get_odd_def odd_pos simp del: ub_arctan_horner.simps lb_arctan_horner.simps) lemma arctan_lower_bound: assumes "0 \ x" shows "x / (1 + x\<^sup>2) \ arctan x" (is "?l x \ _") proof - have "?l x - arctan x \ ?l 0 - arctan 0" using assms by (intro DERIV_nonpos_imp_nonincreasing[where f="\x. ?l x - arctan x"]) (auto intro!: derivative_eq_intros simp: add_nonneg_eq_0_iff field_simps) thus ?thesis by simp qed lemma arctan_divide_mono: "0 < x \ x \ y \ arctan y / y \ arctan x / x" by (rule DERIV_nonpos_imp_nonincreasing[where f="\x. arctan x / x"]) (auto intro!: derivative_eq_intros divide_nonpos_nonneg simp: inverse_eq_divide arctan_lower_bound) lemma arctan_mult_mono: "0 \ x \ x \ y \ x * arctan y \ y * arctan x" using arctan_divide_mono[of x y] by (cases "x = 0") (simp_all add: field_simps) lemma arctan_mult_le: assumes "0 \ x" "x \ y" "y * z \ arctan y" shows "x * z \ arctan x" proof (cases "x = 0") case True then show ?thesis by simp next case False with assms have "z \ arctan y / y" by (simp add: field_simps) also have "\ \ arctan x / x" using assms \x \ 0\ by (auto intro!: arctan_divide_mono) finally show ?thesis using assms \x \ 0\ by (simp add: field_simps) qed lemma arctan_le_mult: assumes "0 < x" "x \ y" "arctan x \ x * z" shows "arctan y \ y * z" proof - from assms have "arctan y / y \ arctan x / x" by (auto intro!: arctan_divide_mono) also have "\ \ z" using assms by (auto simp: field_simps) finally show ?thesis using assms by (simp add: field_simps) qed lemma arctan_0_1_bounds_le: assumes "0 \ x" "x \ 1" "0 < real_of_float xl" "real_of_float xl \ x * x" "x * x \ real_of_float xu" "real_of_float xu \ 1" shows "arctan x \ {x * lb_arctan_horner p1 (get_even n) 1 xu .. x * ub_arctan_horner p2 (get_odd n) 1 xl}" proof - from assms have "real_of_float xl \ 1" "sqrt (real_of_float xl) \ x" "x \ sqrt (real_of_float xu)" "0 \ real_of_float xu" "0 \ real_of_float xl" "0 < sqrt (real_of_float xl)" by (auto intro!: real_le_rsqrt real_le_lsqrt simp: power2_eq_square) from arctan_0_1_bounds[OF \0 \ real_of_float xu\ \real_of_float xu \ 1\] have "sqrt (real_of_float xu) * real_of_float (lb_arctan_horner p1 (get_even n) 1 xu) \ arctan (sqrt (real_of_float xu))" by simp from arctan_mult_le[OF \0 \ x\ \x \ sqrt _\ this] have "x * real_of_float (lb_arctan_horner p1 (get_even n) 1 xu) \ arctan x" . moreover from arctan_0_1_bounds[OF \0 \ real_of_float xl\ \real_of_float xl \ 1\] have "arctan (sqrt (real_of_float xl)) \ sqrt (real_of_float xl) * real_of_float (ub_arctan_horner p2 (get_odd n) 1 xl)" by simp from arctan_le_mult[OF \0 < sqrt xl\ \sqrt xl \ x\ this] have "arctan x \ x * real_of_float (ub_arctan_horner p2 (get_odd n) 1 xl)" . ultimately show ?thesis by simp qed lemma arctan_0_1_bounds_round: assumes "0 \ real_of_float x" "real_of_float x \ 1" shows "arctan x \ {real_of_float x * lb_arctan_horner p1 (get_even n) 1 (float_round_up (Suc p2) (x * x)) .. real_of_float x * ub_arctan_horner p3 (get_odd n) 1 (float_round_down (Suc p4) (x * x))}" using assms apply (cases "x > 0") apply (intro arctan_0_1_bounds_le) apply (auto simp: float_round_down.rep_eq float_round_up.rep_eq intro!: truncate_up_le1 mult_le_one truncate_down_le truncate_up_le truncate_down_pos mult_pos_pos) done subsection "Compute \" definition ub_pi :: "nat \ float" where "ub_pi prec = (let A = rapprox_rat prec 1 5 ; B = lapprox_rat prec 1 239 in ((Float 1 2) * float_plus_up prec ((Float 1 2) * float_round_up prec (A * (ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (A * A))))) (- float_round_down prec (B * (lb_arctan_horner prec (get_even (prec div 14 + 1)) 1 (float_round_up (Suc prec) (B * B)))))))" definition lb_pi :: "nat \ float" where "lb_pi prec = (let A = lapprox_rat prec 1 5 ; B = rapprox_rat prec 1 239 in ((Float 1 2) * float_plus_down prec ((Float 1 2) * float_round_down prec (A * (lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (float_round_up (Suc prec) (A * A))))) (- float_round_up prec (B * (ub_arctan_horner prec (get_odd (prec div 14 + 1)) 1 (float_round_down (Suc prec) (B * B)))))))" lemma pi_boundaries: "pi \ {(lb_pi n) .. (ub_pi n)}" proof - have machin_pi: "pi = 4 * (4 * arctan (1 / 5) - arctan (1 / 239))" unfolding machin[symmetric] by auto { fix prec n :: nat fix k :: int assume "1 < k" hence "0 \ k" and "0 < k" and "1 \ k" by auto let ?k = "rapprox_rat prec 1 k" let ?kl = "float_round_down (Suc prec) (?k * ?k)" have "1 div k = 0" using div_pos_pos_trivial[OF _ \1 < k\] by auto have "0 \ real_of_float ?k" by (rule order_trans[OF _ rapprox_rat]) (auto simp add: \0 \ k\) have "real_of_float ?k \ 1" by (auto simp add: \0 < k\ \1 \ k\ less_imp_le intro!: mult_le_one order_trans[OF _ rapprox_rat] rapprox_rat_le1) have "1 / k \ ?k" using rapprox_rat[where x=1 and y=k] by auto hence "arctan (1 / k) \ arctan ?k" by (rule arctan_monotone') also have "\ \ (?k * ub_arctan_horner prec (get_odd n) 1 ?kl)" using arctan_0_1_bounds_round[OF \0 \ real_of_float ?k\ \real_of_float ?k \ 1\] by auto finally have "arctan (1 / k) \ ?k * ub_arctan_horner prec (get_odd n) 1 ?kl" . } note ub_arctan = this { fix prec n :: nat fix k :: int assume "1 < k" hence "0 \ k" and "0 < k" by auto let ?k = "lapprox_rat prec 1 k" let ?ku = "float_round_up (Suc prec) (?k * ?k)" have "1 div k = 0" using div_pos_pos_trivial[OF _ \1 < k\] by auto have "1 / k \ 1" using \1 < k\ by auto have "0 \ real_of_float ?k" using lapprox_rat_nonneg[where x=1 and y=k, OF zero_le_one \0 \ k\] by (auto simp add: \1 div k = 0\) have "0 \ real_of_float (?k * ?k)" by simp have "real_of_float ?k \ 1" using lapprox_rat by (rule order_trans, auto simp add: \1 / k \ 1\) hence "real_of_float (?k * ?k) \ 1" using \0 \ real_of_float ?k\ by (auto intro!: mult_le_one) have "?k \ 1 / k" using lapprox_rat[where x=1 and y=k] by auto have "?k * lb_arctan_horner prec (get_even n) 1 ?ku \ arctan ?k" using arctan_0_1_bounds_round[OF \0 \ real_of_float ?k\ \real_of_float ?k \ 1\] by auto also have "\ \ arctan (1 / k)" using \?k \ 1 / k\ by (rule arctan_monotone') finally have "?k * lb_arctan_horner prec (get_even n) 1 ?ku \ arctan (1 / k)" . } note lb_arctan = this have "pi \ ub_pi n " unfolding ub_pi_def machin_pi Let_def times_float.rep_eq Float_num using lb_arctan[of 239] ub_arctan[of 5] powr_realpow[of 2 2] by (intro mult_left_mono float_plus_up_le float_plus_down_le) (auto intro!: mult_left_mono float_round_down_le float_round_up_le diff_mono) moreover have "lb_pi n \ pi" unfolding lb_pi_def machin_pi Let_def times_float.rep_eq Float_num using lb_arctan[of 5] ub_arctan[of 239] by (intro mult_left_mono float_plus_up_le float_plus_down_le) (auto intro!: mult_left_mono float_round_down_le float_round_up_le diff_mono) ultimately show ?thesis by auto qed lift_definition pi_float_interval::"nat \ float interval" is "\prec. (lb_pi prec, ub_pi prec)" using pi_boundaries by (auto intro: order_trans) lemma lower_pi_float_interval: "lower (pi_float_interval prec) = lb_pi prec" by transfer auto lemma upper_pi_float_interval: "upper (pi_float_interval prec) = ub_pi prec" by transfer auto lemma pi_float_interval: "pi \ set_of (real_interval (pi_float_interval prec))" using pi_boundaries by (auto simp: set_of_eq lower_pi_float_interval upper_pi_float_interval) subsection "Compute arcus tangens in the entire domain" function lb_arctan :: "nat \ float \ float" and ub_arctan :: "nat \ float \ float" where "lb_arctan prec x = (let ub_horner = \ x. float_round_up prec (x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x))); lb_horner = \ x. float_round_down prec (x * lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (float_round_up (Suc prec) (x * x))) in if x < 0 then - ub_arctan prec (-x) else if x \ Float 1 (- 1) then lb_horner x else if x \ Float 1 1 then Float 1 1 * lb_horner (float_divl prec x (float_plus_up prec 1 (ub_sqrt prec (float_plus_up prec 1 (float_round_up prec (x * x)))))) else let inv = float_divr prec 1 x in if inv > 1 then 0 else float_plus_down prec (lb_pi prec * Float 1 (- 1)) ( - ub_horner inv))" | "ub_arctan prec x = (let lb_horner = \ x. float_round_down prec (x * lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (float_round_up (Suc prec) (x * x))) ; ub_horner = \ x. float_round_up prec (x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x))) in if x < 0 then - lb_arctan prec (-x) else if x \ Float 1 (- 1) then ub_horner x else if x \ Float 1 1 then let y = float_divr prec x (float_plus_down (Suc prec) 1 (lb_sqrt prec (float_plus_down prec 1 (float_round_down prec (x * x))))) in if y > 1 then ub_pi prec * Float 1 (- 1) else Float 1 1 * ub_horner y else float_plus_up prec (ub_pi prec * Float 1 (- 1)) ( - lb_horner (float_divl prec 1 x)))" by pat_completeness auto termination by (relation "measure (\ v. let (prec, x) = case_sum id id v in (if x < 0 then 1 else 0))", auto) declare ub_arctan_horner.simps[simp del] declare lb_arctan_horner.simps[simp del] lemma lb_arctan_bound': assumes "0 \ real_of_float x" shows "lb_arctan prec x \ arctan x" proof - have "\ x < 0" and "0 \ x" using \0 \ real_of_float x\ by (auto intro!: truncate_up_le ) let "?ub_horner x" = "x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x))" and "?lb_horner x" = "x * lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (float_round_up (Suc prec) (x * x))" show ?thesis proof (cases "x \ Float 1 (- 1)") case True hence "real_of_float x \ 1" by simp from arctan_0_1_bounds_round[OF \0 \ real_of_float x\ \real_of_float x \ 1\] show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_P[OF True] using \0 \ x\ by (auto intro!: float_round_down_le) next case False hence "0 < real_of_float x" by auto let ?R = "1 + sqrt (1 + real_of_float x * real_of_float x)" let ?sxx = "float_plus_up prec 1 (float_round_up prec (x * x))" let ?fR = "float_plus_up prec 1 (ub_sqrt prec ?sxx)" let ?DIV = "float_divl prec x ?fR" have divisor_gt0: "0 < ?R" by (auto intro: add_pos_nonneg) have "sqrt (1 + x*x) \ sqrt ?sxx" by (auto simp: float_plus_up.rep_eq plus_up_def float_round_up.rep_eq intro!: truncate_up_le) also have "\ \ ub_sqrt prec ?sxx" using bnds_sqrt'[of ?sxx prec] by auto finally have "sqrt (1 + x*x) \ ub_sqrt prec ?sxx" . hence "?R \ ?fR" by (auto simp: float_plus_up.rep_eq plus_up_def intro!: truncate_up_le) hence "0 < ?fR" and "0 < real_of_float ?fR" using \0 < ?R\ by auto have monotone: "?DIV \ x / ?R" proof - have "?DIV \ real_of_float x / ?fR" by (rule float_divl) also have "\ \ x / ?R" by (rule divide_left_mono[OF \?R \ ?fR\ \0 \ real_of_float x\ mult_pos_pos[OF order_less_le_trans[OF divisor_gt0 \?R \ real_of_float ?fR\] divisor_gt0]]) finally show ?thesis . qed show ?thesis proof (cases "x \ Float 1 1") case True have "x \ sqrt (1 + x * x)" using real_sqrt_sum_squares_ge2[where x=1, unfolded numeral_2_eq_2] by auto also note \\ \ (ub_sqrt prec ?sxx)\ finally have "real_of_float x \ ?fR" by (auto simp: float_plus_up.rep_eq plus_up_def intro!: truncate_up_le) moreover have "?DIV \ real_of_float x / ?fR" by (rule float_divl) ultimately have "real_of_float ?DIV \ 1" unfolding divide_le_eq_1_pos[OF \0 < real_of_float ?fR\, symmetric] by auto have "0 \ real_of_float ?DIV" using float_divl_lower_bound[OF \0 \ x\] \0 < ?fR\ unfolding less_eq_float_def by auto from arctan_0_1_bounds_round[OF \0 \ real_of_float (?DIV)\ \real_of_float (?DIV) \ 1\] have "Float 1 1 * ?lb_horner ?DIV \ 2 * arctan ?DIV" by simp also have "\ \ 2 * arctan (x / ?R)" using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono arctan_monotone') also have "2 * arctan (x / ?R) = arctan x" using arctan_half[symmetric] unfolding numeral_2_eq_2 power_Suc2 power_0 mult_1_left . finally show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_not_P[OF \\ x \ Float 1 (- 1)\] if_P[OF True] by (auto simp: float_round_down.rep_eq intro!: order_trans[OF mult_left_mono[OF truncate_down]]) next case False hence "2 < real_of_float x" by auto hence "1 \ real_of_float x" by auto let "?invx" = "float_divr prec 1 x" have "0 \ arctan x" using arctan_monotone'[OF \0 \ real_of_float x\] using arctan_tan[of 0, unfolded tan_zero] by auto show ?thesis proof (cases "1 < ?invx") case True show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_not_P[OF \\ x \ Float 1 (- 1)\] if_not_P[OF False] if_P[OF True] using \0 \ arctan x\ by auto next case False hence "real_of_float ?invx \ 1" by auto have "0 \ real_of_float ?invx" by (rule order_trans[OF _ float_divr]) (auto simp add: \0 \ real_of_float x\) have "1 / x \ 0" and "0 < 1 / x" using \0 < real_of_float x\ by auto have "arctan (1 / x) \ arctan ?invx" unfolding one_float.rep_eq[symmetric] by (rule arctan_monotone', rule float_divr) also have "\ \ ?ub_horner ?invx" using arctan_0_1_bounds_round[OF \0 \ real_of_float ?invx\ \real_of_float ?invx \ 1\] by (auto intro!: float_round_up_le) also note float_round_up finally have "pi / 2 - float_round_up prec (?ub_horner ?invx) \ arctan x" using \0 \ arctan x\ arctan_inverse[OF \1 / x \ 0\] unfolding sgn_pos[OF \0 < 1 / real_of_float x\] le_diff_eq by auto moreover have "lb_pi prec * Float 1 (- 1) \ pi / 2" unfolding Float_num times_divide_eq_right mult_1_left using pi_boundaries by simp ultimately show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_not_P[OF \\ x \ Float 1 (- 1)\] if_not_P[OF \\ x \ Float 1 1\] if_not_P[OF False] by (auto intro!: float_plus_down_le) qed qed qed qed lemma ub_arctan_bound': assumes "0 \ real_of_float x" shows "arctan x \ ub_arctan prec x" proof - have "\ x < 0" and "0 \ x" using \0 \ real_of_float x\ by auto let "?ub_horner x" = "float_round_up prec (x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x)))" let "?lb_horner x" = "float_round_down prec (x * lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (float_round_up (Suc prec) (x * x)))" show ?thesis proof (cases "x \ Float 1 (- 1)") case True hence "real_of_float x \ 1" by auto show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_P[OF True] using arctan_0_1_bounds_round[OF \0 \ real_of_float x\ \real_of_float x \ 1\] by (auto intro!: float_round_up_le) next case False hence "0 < real_of_float x" by auto let ?R = "1 + sqrt (1 + real_of_float x * real_of_float x)" let ?sxx = "float_plus_down prec 1 (float_round_down prec (x * x))" let ?fR = "float_plus_down (Suc prec) 1 (lb_sqrt prec ?sxx)" let ?DIV = "float_divr prec x ?fR" have sqr_ge0: "0 \ 1 + real_of_float x * real_of_float x" using sum_power2_ge_zero[of 1 "real_of_float x", unfolded numeral_2_eq_2] by auto hence "0 \ real_of_float (1 + x*x)" by auto hence divisor_gt0: "0 < ?R" by (auto intro: add_pos_nonneg) have "lb_sqrt prec ?sxx \ sqrt ?sxx" using bnds_sqrt'[of ?sxx] by auto also have "\ \ sqrt (1 + x*x)" by (auto simp: float_plus_down.rep_eq plus_down_def float_round_down.rep_eq truncate_down_le) finally have "lb_sqrt prec ?sxx \ sqrt (1 + x*x)" . hence "?fR \ ?R" by (auto simp: float_plus_down.rep_eq plus_down_def truncate_down_le) have "0 < real_of_float ?fR" by (auto simp: float_plus_down.rep_eq plus_down_def float_round_down.rep_eq intro!: truncate_down_ge1 lb_sqrt_lower_bound order_less_le_trans[OF zero_less_one] truncate_down_nonneg add_nonneg_nonneg) have monotone: "x / ?R \ (float_divr prec x ?fR)" proof - from divide_left_mono[OF \?fR \ ?R\ \0 \ real_of_float x\ mult_pos_pos[OF divisor_gt0 \0 < real_of_float ?fR\]] have "x / ?R \ x / ?fR" . also have "\ \ ?DIV" by (rule float_divr) finally show ?thesis . qed show ?thesis proof (cases "x \ Float 1 1") case True show ?thesis proof (cases "?DIV > 1") case True have "pi / 2 \ ub_pi prec * Float 1 (- 1)" unfolding Float_num times_divide_eq_right mult_1_left using pi_boundaries by auto from order_less_le_trans[OF arctan_ubound this, THEN less_imp_le] show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_not_P[OF \\ x \ Float 1 (- 1)\] if_P[OF \x \ Float 1 1\] if_P[OF True] . next case False hence "real_of_float ?DIV \ 1" by auto have "0 \ x / ?R" using \0 \ real_of_float x\ \0 < ?R\ unfolding zero_le_divide_iff by auto hence "0 \ real_of_float ?DIV" using monotone by (rule order_trans) have "arctan x = 2 * arctan (x / ?R)" using arctan_half unfolding numeral_2_eq_2 power_Suc2 power_0 mult_1_left . also have "\ \ 2 * arctan (?DIV)" using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono) also have "\ \ (Float 1 1 * ?ub_horner ?DIV)" unfolding Float_num using arctan_0_1_bounds_round[OF \0 \ real_of_float ?DIV\ \real_of_float ?DIV \ 1\] by (auto intro!: float_round_up_le) finally show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_not_P[OF \\ x \ Float 1 (- 1)\] if_P[OF \x \ Float 1 1\] if_not_P[OF False] . qed next case False hence "2 < real_of_float x" by auto hence "1 \ real_of_float x" by auto hence "0 < real_of_float x" by auto hence "0 < x" by auto let "?invx" = "float_divl prec 1 x" have "0 \ arctan x" using arctan_monotone'[OF \0 \ real_of_float x\] and arctan_tan[of 0, unfolded tan_zero] by auto have "real_of_float ?invx \ 1" unfolding less_float_def by (rule order_trans[OF float_divl]) (auto simp add: \1 \ real_of_float x\ divide_le_eq_1_pos[OF \0 < real_of_float x\]) have "0 \ real_of_float ?invx" using \0 < x\ by (intro float_divl_lower_bound) auto have "1 / x \ 0" and "0 < 1 / x" using \0 < real_of_float x\ by auto have "(?lb_horner ?invx) \ arctan (?invx)" using arctan_0_1_bounds_round[OF \0 \ real_of_float ?invx\ \real_of_float ?invx \ 1\] by (auto intro!: float_round_down_le) also have "\ \ arctan (1 / x)" unfolding one_float.rep_eq[symmetric] by (rule arctan_monotone') (rule float_divl) finally have "arctan x \ pi / 2 - (?lb_horner ?invx)" using \0 \ arctan x\ arctan_inverse[OF \1 / x \ 0\] unfolding sgn_pos[OF \0 < 1 / x\] le_diff_eq by auto moreover have "pi / 2 \ ub_pi prec * Float 1 (- 1)" unfolding Float_num times_divide_eq_right mult_1_right using pi_boundaries by auto ultimately show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_not_P[OF \\ x \ Float 1 (- 1)\] if_not_P[OF False] by (auto intro!: float_round_up_le float_plus_up_le) qed qed qed lemma arctan_boundaries: "arctan x \ {(lb_arctan prec x) .. (ub_arctan prec x)}" proof (cases "0 \ x") case True hence "0 \ real_of_float x" by auto show ?thesis using ub_arctan_bound'[OF \0 \ real_of_float x\] lb_arctan_bound'[OF \0 \ real_of_float x\] unfolding atLeastAtMost_iff by auto next case False let ?mx = "-x" from False have "x < 0" and "0 \ real_of_float ?mx" by auto hence bounds: "lb_arctan prec ?mx \ arctan ?mx \ arctan ?mx \ ub_arctan prec ?mx" using ub_arctan_bound'[OF \0 \ real_of_float ?mx\] lb_arctan_bound'[OF \0 \ real_of_float ?mx\] by auto show ?thesis unfolding minus_float.rep_eq arctan_minus lb_arctan.simps[where x=x] ub_arctan.simps[where x=x] Let_def if_P[OF \x < 0\] unfolding atLeastAtMost_iff using bounds[unfolded minus_float.rep_eq arctan_minus] by (simp add: arctan_minus) qed lemma bnds_arctan: "\ (x::real) lx ux. (l, u) = (lb_arctan prec lx, ub_arctan prec ux) \ x \ {lx .. ux} \ l \ arctan x \ arctan x \ u" proof (rule allI, rule allI, rule allI, rule impI) fix x :: real fix lx ux assume "(l, u) = (lb_arctan prec lx, ub_arctan prec ux) \ x \ {lx .. ux}" hence l: "lb_arctan prec lx = l " and u: "ub_arctan prec ux = u" and x: "x \ {lx .. ux}" by auto show "l \ arctan x \ arctan x \ u" proof show "l \ arctan x" proof - from arctan_boundaries[of lx prec, unfolded l] have "l \ arctan lx" by (auto simp del: lb_arctan.simps) also have "\ \ arctan x" using x by (auto intro: arctan_monotone') finally show ?thesis . qed show "arctan x \ u" proof - have "arctan x \ arctan ux" using x by (auto intro: arctan_monotone') also have "\ \ u" using arctan_boundaries[of ux prec, unfolded u] by (auto simp del: ub_arctan.simps) finally show ?thesis . qed qed qed lemmas [simp del] = lb_arctan.simps ub_arctan.simps lemma lb_arctan: "arctan (real_of_float x) \ y \ real_of_float (lb_arctan prec x) \ y" and ub_arctan: "y \ arctan x \ y \ ub_arctan prec x" for x::float and y::real using arctan_boundaries[of x prec] by auto lift_definition arctan_float_interval :: "nat \ float interval \ float interval" is "\prec. \(lx, ux). (lb_arctan prec lx, ub_arctan prec ux)" by (auto intro!: lb_arctan ub_arctan arctan_monotone') lemma lower_arctan_float_interval: "lower (arctan_float_interval p x) = lb_arctan p (lower x)" by transfer auto lemma upper_arctan_float_interval: "upper (arctan_float_interval p x) = ub_arctan p (upper x)" by transfer auto lemma arctan_float_interval: "arctan ` set_of (real_interval x) \ set_of (real_interval (arctan_float_interval p x))" by (auto simp: set_of_eq lower_arctan_float_interval upper_arctan_float_interval intro!: lb_arctan ub_arctan arctan_monotone') lemma arctan_float_intervalI: "arctan x \\<^sub>r arctan_float_interval p X" if "x \\<^sub>r X" using arctan_float_interval[of X p] that by auto section "Sinus and Cosinus" subsection "Compute the cosinus and sinus series" fun ub_sin_cos_aux :: "nat \ nat \ nat \ nat \ float \ float" and lb_sin_cos_aux :: "nat \ nat \ nat \ nat \ float \ float" where "ub_sin_cos_aux prec 0 i k x = 0" | "ub_sin_cos_aux prec (Suc n) i k x = float_plus_up prec (rapprox_rat prec 1 k) (- float_round_down prec (x * (lb_sin_cos_aux prec n (i + 2) (k * i * (i + 1)) x)))" | "lb_sin_cos_aux prec 0 i k x = 0" | "lb_sin_cos_aux prec (Suc n) i k x = float_plus_down prec (lapprox_rat prec 1 k) (- float_round_up prec (x * (ub_sin_cos_aux prec n (i + 2) (k * i * (i + 1)) x)))" lemma cos_aux: shows "(lb_sin_cos_aux prec n 1 1 (x * x)) \ (\ i=0.. i=0.. (ub_sin_cos_aux prec n 1 1 (x * x))" (is "?ub") proof - have "0 \ real_of_float (x * x)" by auto let "?f n" = "fact (2 * n) :: nat" have f_eq: "?f (Suc n) = ?f n * ((\i. i + 2) ^^ n) 1 * (((\i. i + 2) ^^ n) 1 + 1)" for n proof - have "\m. ((\i. i + 2) ^^ n) m = m + 2 * n" by (induct n) auto then show ?thesis by auto qed from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0, OF \0 \ real_of_float (x * x)\ f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps] show ?lb and ?ub by (auto simp add: power_mult power2_eq_square[of "real_of_float x"]) qed lemma lb_sin_cos_aux_zero_le_one: "lb_sin_cos_aux prec n i j 0 \ 1" by (cases j n rule: nat.exhaust[case_product nat.exhaust]) (auto intro!: float_plus_down_le order_trans[OF lapprox_rat]) lemma one_le_ub_sin_cos_aux: "odd n \ 1 \ ub_sin_cos_aux prec n i (Suc 0) 0" by (cases n) (auto intro!: float_plus_up_le order_trans[OF _ rapprox_rat]) lemma cos_boundaries: assumes "0 \ real_of_float x" and "x \ pi / 2" shows "cos x \ {(lb_sin_cos_aux prec (get_even n) 1 1 (x * x)) .. (ub_sin_cos_aux prec (get_odd n) 1 1 (x * x))}" proof (cases "real_of_float x = 0") case False hence "real_of_float x \ 0" by auto hence "0 < x" and "0 < real_of_float x" using \0 \ real_of_float x\ by auto have "0 < x * x" using \0 < x\ by simp have morph_to_if_power: "(\ i=0.. i = 0 ..< 2 * n. (if even(i) then ((- 1) ^ (i div 2))/((fact i)) else 0) * x ^ i)" (is "?sum = ?ifsum") for x n proof - have "?sum = ?sum + (\ j = 0 ..< n. 0)" by auto also have "\ = (\ j = 0 ..< n. (- 1) ^ ((2 * j) div 2) / ((fact (2 * j))) * x ^(2 * j)) + (\ j = 0 ..< n. 0)" by auto also have "\ = (\ i = 0 ..< 2 * n. if even i then (- 1) ^ (i div 2) / ((fact i)) * x ^ i else 0)" unfolding sum_split_even_odd atLeast0LessThan .. also have "\ = (\ i = 0 ..< 2 * n. (if even i then (- 1) ^ (i div 2) / ((fact i)) else 0) * x ^ i)" by (rule sum.cong) auto finally show ?thesis . qed { fix n :: nat assume "0 < n" hence "0 < 2 * n" by auto obtain t where "0 < t" and "t < real_of_float x" and cos_eq: "cos x = (\ i = 0 ..< 2 * n. (if even(i) then ((- 1) ^ (i div 2))/((fact i)) else 0) * (real_of_float x) ^ i) + (cos (t + 1/2 * (2 * n) * pi) / (fact (2*n))) * (real_of_float x)^(2*n)" (is "_ = ?SUM + ?rest / ?fact * ?pow") using Maclaurin_cos_expansion2[OF \0 < real_of_float x\ \0 < 2 * n\] unfolding cos_coeff_def atLeast0LessThan by auto have "cos t * (- 1) ^ n = cos t * cos (n * pi) + sin t * sin (n * pi)" by auto also have "\ = cos (t + n * pi)" by (simp add: cos_add) also have "\ = ?rest" by auto finally have "cos t * (- 1) ^ n = ?rest" . moreover have "t \ pi / 2" using \t < real_of_float x\ and \x \ pi / 2\ by auto hence "0 \ cos t" using \0 < t\ and cos_ge_zero by auto ultimately have even: "even n \ 0 \ ?rest" and odd: "odd n \ 0 \ - ?rest " by auto have "0 < ?fact" by auto have "0 < ?pow" using \0 < real_of_float x\ by auto { assume "even n" have "(lb_sin_cos_aux prec n 1 1 (x * x)) \ ?SUM" unfolding morph_to_if_power[symmetric] using cos_aux by auto also have "\ \ cos x" proof - from even[OF \even n\] \0 < ?fact\ \0 < ?pow\ have "0 \ (?rest / ?fact) * ?pow" by simp thus ?thesis unfolding cos_eq by auto qed finally have "(lb_sin_cos_aux prec n 1 1 (x * x)) \ cos x" . } note lb = this { assume "odd n" have "cos x \ ?SUM" proof - from \0 < ?fact\ and \0 < ?pow\ and odd[OF \odd n\] have "0 \ (- ?rest) / ?fact * ?pow" by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le) thus ?thesis unfolding cos_eq by auto qed also have "\ \ (ub_sin_cos_aux prec n 1 1 (x * x))" unfolding morph_to_if_power[symmetric] using cos_aux by auto finally have "cos x \ (ub_sin_cos_aux prec n 1 1 (x * x))" . } note ub = this and lb } note ub = this(1) and lb = this(2) have "cos x \ (ub_sin_cos_aux prec (get_odd n) 1 1 (x * x))" using ub[OF odd_pos[OF get_odd] get_odd] . moreover have "(lb_sin_cos_aux prec (get_even n) 1 1 (x * x)) \ cos x" proof (cases "0 < get_even n") case True show ?thesis using lb[OF True get_even] . next case False hence "get_even n = 0" by auto have "- (pi / 2) \ x" by (rule order_trans[OF _ \0 < real_of_float x\[THEN less_imp_le]]) auto with \x \ pi / 2\ show ?thesis unfolding \get_even n = 0\ lb_sin_cos_aux.simps minus_float.rep_eq zero_float.rep_eq using cos_ge_zero by auto qed ultimately show ?thesis by auto next case True hence "x = 0" by (simp add: real_of_float_eq) thus ?thesis using lb_sin_cos_aux_zero_le_one one_le_ub_sin_cos_aux by simp qed lemma sin_aux: assumes "0 \ real_of_float x" shows "(x * lb_sin_cos_aux prec n 2 1 (x * x)) \ (\ i=0.. i=0.. (x * ub_sin_cos_aux prec n 2 1 (x * x))" (is "?ub") proof - have "0 \ real_of_float (x * x)" by auto let "?f n" = "fact (2 * n + 1) :: nat" have f_eq: "?f (Suc n) = ?f n * ((\i. i + 2) ^^ n) 2 * (((\i. i + 2) ^^ n) 2 + 1)" for n proof - have F: "\m. ((\i. i + 2) ^^ n) m = m + 2 * n" by (induct n) auto show ?thesis unfolding F by auto qed from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0, OF \0 \ real_of_float (x * x)\ f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps] show "?lb" and "?ub" using \0 \ real_of_float x\ apply (simp_all only: power_add power_one_right mult.assoc[symmetric] sum_distrib_right[symmetric]) apply (simp_all only: mult.commute[where 'a=real] of_nat_fact) apply (auto intro!: mult_left_mono simp add: power_mult power2_eq_square[of "real_of_float x"]) done qed lemma sin_boundaries: assumes "0 \ real_of_float x" and "x \ pi / 2" shows "sin x \ {(x * lb_sin_cos_aux prec (get_even n) 2 1 (x * x)) .. (x * ub_sin_cos_aux prec (get_odd n) 2 1 (x * x))}" proof (cases "real_of_float x = 0") case False hence "real_of_float x \ 0" by auto hence "0 < x" and "0 < real_of_float x" using \0 \ real_of_float x\ by auto have "0 < x * x" using \0 < x\ by simp have sum_morph: "(\j = 0 ..< n. (- 1) ^ (((2 * j + 1) - Suc 0) div 2) / ((fact (2 * j + 1))) * x ^(2 * j + 1)) = (\ i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * x ^ i)" (is "?SUM = _") for x :: real and n proof - have pow: "!!i. x ^ (2 * i + 1) = x * x ^ (2 * i)" by auto have "?SUM = (\ j = 0 ..< n. 0) + ?SUM" by auto also have "\ = (\ i = 0 ..< 2 * n. if even i then 0 else (- 1) ^ ((i - Suc 0) div 2) / ((fact i)) * x ^ i)" unfolding sum_split_even_odd atLeast0LessThan .. also have "\ = (\ i = 0 ..< 2 * n. (if even i then 0 else (- 1) ^ ((i - Suc 0) div 2) / ((fact i))) * x ^ i)" by (rule sum.cong) auto finally show ?thesis . qed { fix n :: nat assume "0 < n" hence "0 < 2 * n + 1" by auto obtain t where "0 < t" and "t < real_of_float x" and sin_eq: "sin x = (\ i = 0 ..< 2 * n + 1. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real_of_float x) ^ i) + (sin (t + 1/2 * (2 * n + 1) * pi) / (fact (2*n + 1))) * (real_of_float x)^(2*n + 1)" (is "_ = ?SUM + ?rest / ?fact * ?pow") using Maclaurin_sin_expansion3[OF \0 < 2 * n + 1\ \0 < real_of_float x\] unfolding sin_coeff_def atLeast0LessThan by auto have "?rest = cos t * (- 1) ^ n" unfolding sin_add cos_add of_nat_add distrib_right distrib_left by auto moreover have "t \ pi / 2" using \t < real_of_float x\ and \x \ pi / 2\ by auto hence "0 \ cos t" using \0 < t\ and cos_ge_zero by auto ultimately have even: "even n \ 0 \ ?rest" and odd: "odd n \ 0 \ - ?rest" by auto have "0 < ?fact" by (simp del: fact_Suc) have "0 < ?pow" using \0 < real_of_float x\ by (rule zero_less_power) { assume "even n" have "(x * lb_sin_cos_aux prec n 2 1 (x * x)) \ (\ i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real_of_float x) ^ i)" using sin_aux[OF \0 \ real_of_float x\] unfolding sum_morph[symmetric] by auto also have "\ \ ?SUM" by auto also have "\ \ sin x" proof - from even[OF \even n\] \0 < ?fact\ \0 < ?pow\ have "0 \ (?rest / ?fact) * ?pow" by simp thus ?thesis unfolding sin_eq by auto qed finally have "(x * lb_sin_cos_aux prec n 2 1 (x * x)) \ sin x" . } note lb = this { assume "odd n" have "sin x \ ?SUM" proof - from \0 < ?fact\ and \0 < ?pow\ and odd[OF \odd n\] have "0 \ (- ?rest) / ?fact * ?pow" by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le) thus ?thesis unfolding sin_eq by auto qed also have "\ \ (\ i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real_of_float x) ^ i)" by auto also have "\ \ (x * ub_sin_cos_aux prec n 2 1 (x * x))" using sin_aux[OF \0 \ real_of_float x\] unfolding sum_morph[symmetric] by auto finally have "sin x \ (x * ub_sin_cos_aux prec n 2 1 (x * x))" . } note ub = this and lb } note ub = this(1) and lb = this(2) have "sin x \ (x * ub_sin_cos_aux prec (get_odd n) 2 1 (x * x))" using ub[OF odd_pos[OF get_odd] get_odd] . moreover have "(x * lb_sin_cos_aux prec (get_even n) 2 1 (x * x)) \ sin x" proof (cases "0 < get_even n") case True show ?thesis using lb[OF True get_even] . next case False hence "get_even n = 0" by auto with \x \ pi / 2\ \0 \ real_of_float x\ show ?thesis unfolding \get_even n = 0\ ub_sin_cos_aux.simps minus_float.rep_eq using sin_ge_zero by auto qed ultimately show ?thesis by auto next case True show ?thesis proof (cases "n = 0") case True thus ?thesis unfolding \n = 0\ get_even_def get_odd_def using \real_of_float x = 0\ lapprox_rat[where x="-1" and y=1] by auto next case False with not0_implies_Suc obtain m where "n = Suc m" by blast thus ?thesis unfolding \n = Suc m\ get_even_def get_odd_def using \real_of_float x = 0\ rapprox_rat[where x=1 and y=1] lapprox_rat[where x=1 and y=1] by (cases "even (Suc m)") auto qed qed subsection "Compute the cosinus in the entire domain" definition lb_cos :: "nat \ float \ float" where "lb_cos prec x = (let horner = \ x. lb_sin_cos_aux prec (get_even (prec div 4 + 1)) 1 1 (x * x) ; half = \ x. if x < 0 then - 1 else float_plus_down prec (Float 1 1 * x * x) (- 1) in if x < Float 1 (- 1) then horner x else if x < 1 then half (horner (x * Float 1 (- 1))) else half (half (horner (x * Float 1 (- 2)))))" definition ub_cos :: "nat \ float \ float" where "ub_cos prec x = (let horner = \ x. ub_sin_cos_aux prec (get_odd (prec div 4 + 1)) 1 1 (x * x) ; half = \ x. float_plus_up prec (Float 1 1 * x * x) (- 1) in if x < Float 1 (- 1) then horner x else if x < 1 then half (horner (x * Float 1 (- 1))) else half (half (horner (x * Float 1 (- 2)))))" lemma lb_cos: assumes "0 \ real_of_float x" and "x \ pi" shows "cos x \ {(lb_cos prec x) .. (ub_cos prec x)}" (is "?cos x \ {(?lb x) .. (?ub x) }") proof - have x_half[symmetric]: "cos x = 2 * cos (x / 2) * cos (x / 2) - 1" for x :: real proof - have "cos x = cos (x / 2 + x / 2)" by auto also have "\ = cos (x / 2) * cos (x / 2) + sin (x / 2) * sin (x / 2) - sin (x / 2) * sin (x / 2) + cos (x / 2) * cos (x / 2) - 1" unfolding cos_add by auto also have "\ = 2 * cos (x / 2) * cos (x / 2) - 1" by algebra finally show ?thesis . qed have "\ x < 0" using \0 \ real_of_float x\ by auto let "?ub_horner x" = "ub_sin_cos_aux prec (get_odd (prec div 4 + 1)) 1 1 (x * x)" let "?lb_horner x" = "lb_sin_cos_aux prec (get_even (prec div 4 + 1)) 1 1 (x * x)" let "?ub_half x" = "float_plus_up prec (Float 1 1 * x * x) (- 1)" let "?lb_half x" = "if x < 0 then - 1 else float_plus_down prec (Float 1 1 * x * x) (- 1)" show ?thesis proof (cases "x < Float 1 (- 1)") case True hence "x \ pi / 2" using pi_ge_two by auto show ?thesis unfolding lb_cos_def[where x=x] ub_cos_def[where x=x] if_not_P[OF \\ x < 0\] if_P[OF \x < Float 1 (- 1)\] Let_def using cos_boundaries[OF \0 \ real_of_float x\ \x \ pi / 2\] . next case False { fix y x :: float let ?x2 = "(x * Float 1 (- 1))" assume "y \ cos ?x2" and "-pi \ x" and "x \ pi" hence "- (pi / 2) \ ?x2" and "?x2 \ pi / 2" using pi_ge_two unfolding Float_num by auto hence "0 \ cos ?x2" by (rule cos_ge_zero) have "(?lb_half y) \ cos x" proof (cases "y < 0") case True show ?thesis using cos_ge_minus_one unfolding if_P[OF True] by auto next case False hence "0 \ real_of_float y" by auto from mult_mono[OF \y \ cos ?x2\ \y \ cos ?x2\ \0 \ cos ?x2\ this] have "real_of_float y * real_of_float y \ cos ?x2 * cos ?x2" . hence "2 * real_of_float y * real_of_float y \ 2 * cos ?x2 * cos ?x2" by auto hence "2 * real_of_float y * real_of_float y - 1 \ 2 * cos (x / 2) * cos (x / 2) - 1" unfolding Float_num by auto thus ?thesis unfolding if_not_P[OF False] x_half Float_num by (auto intro!: float_plus_down_le) qed } note lb_half = this { fix y x :: float let ?x2 = "(x * Float 1 (- 1))" assume ub: "cos ?x2 \ y" and "- pi \ x" and "x \ pi" hence "- (pi / 2) \ ?x2" and "?x2 \ pi / 2" using pi_ge_two unfolding Float_num by auto hence "0 \ cos ?x2" by (rule cos_ge_zero) have "cos x \ (?ub_half y)" proof - have "0 \ real_of_float y" using \0 \ cos ?x2\ ub by (rule order_trans) from mult_mono[OF ub ub this \0 \ cos ?x2\] have "cos ?x2 * cos ?x2 \ real_of_float y * real_of_float y" . hence "2 * cos ?x2 * cos ?x2 \ 2 * real_of_float y * real_of_float y" by auto hence "2 * cos (x / 2) * cos (x / 2) - 1 \ 2 * real_of_float y * real_of_float y - 1" unfolding Float_num by auto thus ?thesis unfolding x_half Float_num by (auto intro!: float_plus_up_le) qed } note ub_half = this let ?x2 = "x * Float 1 (- 1)" let ?x4 = "x * Float 1 (- 1) * Float 1 (- 1)" have "-pi \ x" using pi_ge_zero[THEN le_imp_neg_le, unfolded minus_zero] \0 \ real_of_float x\ by (rule order_trans) show ?thesis proof (cases "x < 1") case True hence "real_of_float x \ 1" by auto have "0 \ real_of_float ?x2" and "?x2 \ pi / 2" using pi_ge_two \0 \ real_of_float x\ using assms by auto from cos_boundaries[OF this] have lb: "(?lb_horner ?x2) \ ?cos ?x2" and ub: "?cos ?x2 \ (?ub_horner ?x2)" by auto have "(?lb x) \ ?cos x" proof - from lb_half[OF lb \-pi \ x\ \x \ pi\] show ?thesis unfolding lb_cos_def[where x=x] Let_def using \\ x < 0\ \\ x < Float 1 (- 1)\ \x < 1\ by auto qed moreover have "?cos x \ (?ub x)" proof - from ub_half[OF ub \-pi \ x\ \x \ pi\] show ?thesis unfolding ub_cos_def[where x=x] Let_def using \\ x < 0\ \\ x < Float 1 (- 1)\ \x < 1\ by auto qed ultimately show ?thesis by auto next case False have "0 \ real_of_float ?x4" and "?x4 \ pi / 2" using pi_ge_two \0 \ real_of_float x\ \x \ pi\ unfolding Float_num by auto from cos_boundaries[OF this] have lb: "(?lb_horner ?x4) \ ?cos ?x4" and ub: "?cos ?x4 \ (?ub_horner ?x4)" by auto have eq_4: "?x2 * Float 1 (- 1) = x * Float 1 (- 2)" by (auto simp: real_of_float_eq) have "(?lb x) \ ?cos x" proof - have "-pi \ ?x2" and "?x2 \ pi" using pi_ge_two \0 \ real_of_float x\ \x \ pi\ by auto from lb_half[OF lb_half[OF lb this] \-pi \ x\ \x \ pi\, unfolded eq_4] show ?thesis unfolding lb_cos_def[where x=x] if_not_P[OF \\ x < 0\] if_not_P[OF \\ x < Float 1 (- 1)\] if_not_P[OF \\ x < 1\] Let_def . qed moreover have "?cos x \ (?ub x)" proof - have "-pi \ ?x2" and "?x2 \ pi" using pi_ge_two \0 \ real_of_float x\ \ x \ pi\ by auto from ub_half[OF ub_half[OF ub this] \-pi \ x\ \x \ pi\, unfolded eq_4] show ?thesis unfolding ub_cos_def[where x=x] if_not_P[OF \\ x < 0\] if_not_P[OF \\ x < Float 1 (- 1)\] if_not_P[OF \\ x < 1\] Let_def . qed ultimately show ?thesis by auto qed qed qed lemma lb_cos_minus: assumes "-pi \ x" and "real_of_float x \ 0" shows "cos (real_of_float(-x)) \ {(lb_cos prec (-x)) .. (ub_cos prec (-x))}" proof - have "0 \ real_of_float (-x)" and "(-x) \ pi" using \-pi \ x\ \real_of_float x \ 0\ by auto from lb_cos[OF this] show ?thesis . qed definition bnds_cos :: "nat \ float \ float \ float * float" where "bnds_cos prec lx ux = (let lpi = float_round_down prec (lb_pi prec) ; upi = float_round_up prec (ub_pi prec) ; k = floor_fl (float_divr prec (lx + lpi) (2 * lpi)) ; lx = float_plus_down prec lx (- k * 2 * (if k < 0 then lpi else upi)) ; ux = float_plus_up prec ux (- k * 2 * (if k < 0 then upi else lpi)) in if - lpi \ lx \ ux \ 0 then (lb_cos prec (-lx), ub_cos prec (-ux)) else if 0 \ lx \ ux \ lpi then (lb_cos prec ux, ub_cos prec lx) else if - lpi \ lx \ ux \ lpi then (min (lb_cos prec (-lx)) (lb_cos prec ux), Float 1 0) else if 0 \ lx \ ux \ 2 * lpi then (Float (- 1) 0, max (ub_cos prec lx) (ub_cos prec (- (ux - 2 * lpi)))) else if -2 * lpi \ lx \ ux \ 0 then (Float (- 1) 0, max (ub_cos prec (lx + 2 * lpi)) (ub_cos prec (-ux))) else (Float (- 1) 0, Float 1 0))" lemma floor_int: obtains k :: int where "real_of_int k = (floor_fl f)" by (simp add: floor_fl_def) lemma cos_periodic_nat[simp]: fixes n :: nat shows "cos (x + n * (2 * pi)) = cos x" proof (induct n arbitrary: x) case 0 then show ?case by simp next case (Suc n) have split_pi_off: "x + (Suc n) * (2 * pi) = (x + n * (2 * pi)) + 2 * pi" unfolding Suc_eq_plus1 of_nat_add of_int_1 distrib_right by auto show ?case unfolding split_pi_off using Suc by auto qed lemma cos_periodic_int[simp]: fixes i :: int shows "cos (x + i * (2 * pi)) = cos x" proof (cases "0 \ i") case True hence i_nat: "real_of_int i = nat i" by auto show ?thesis unfolding i_nat by auto next case False hence i_nat: "i = - real (nat (-i))" by auto have "cos x = cos (x + i * (2 * pi) - i * (2 * pi))" by auto also have "\ = cos (x + i * (2 * pi))" unfolding i_nat mult_minus_left diff_minus_eq_add by (rule cos_periodic_nat) finally show ?thesis by auto qed lemma bnds_cos: "\(x::real) lx ux. (l, u) = bnds_cos prec lx ux \ x \ {lx .. ux} \ l \ cos x \ cos x \ u" proof (rule allI | rule impI | erule conjE)+ fix x :: real fix lx ux assume bnds: "(l, u) = bnds_cos prec lx ux" and x: "x \ {lx .. ux}" let ?lpi = "float_round_down prec (lb_pi prec)" let ?upi = "float_round_up prec (ub_pi prec)" let ?k = "floor_fl (float_divr prec (lx + ?lpi) (2 * ?lpi))" let ?lx2 = "(- ?k * 2 * (if ?k < 0 then ?lpi else ?upi))" let ?ux2 = "(- ?k * 2 * (if ?k < 0 then ?upi else ?lpi))" let ?lx = "float_plus_down prec lx ?lx2" let ?ux = "float_plus_up prec ux ?ux2" obtain k :: int where k: "k = real_of_float ?k" by (rule floor_int) have upi: "pi \ ?upi" and lpi: "?lpi \ pi" using float_round_up[of "ub_pi prec" prec] pi_boundaries[of prec] float_round_down[of prec "lb_pi prec"] by auto hence "lx + ?lx2 \ x - k * (2 * pi) \ x - k * (2 * pi) \ ux + ?ux2" using x by (cases "k = 0") (auto intro!: add_mono simp add: k [symmetric] uminus_add_conv_diff [symmetric] simp del: uminus_add_conv_diff) hence "?lx \ x - k * (2 * pi) \ x - k * (2 * pi) \ ?ux" by (auto intro!: float_plus_down_le float_plus_up_le) note lx = this[THEN conjunct1] and ux = this[THEN conjunct2] hence lx_less_ux: "?lx \ real_of_float ?ux" by (rule order_trans) { assume "- ?lpi \ ?lx" and x_le_0: "x - k * (2 * pi) \ 0" with lpi[THEN le_imp_neg_le] lx have pi_lx: "- pi \ ?lx" and lx_0: "real_of_float ?lx \ 0" by simp_all have "(lb_cos prec (- ?lx)) \ cos (real_of_float (- ?lx))" using lb_cos_minus[OF pi_lx lx_0] by simp also have "\ \ cos (x + (-k) * (2 * pi))" using cos_monotone_minus_pi_0'[OF pi_lx lx x_le_0] by (simp only: uminus_float.rep_eq of_int_minus cos_minus mult_minus_left) simp finally have "(lb_cos prec (- ?lx)) \ cos x" unfolding cos_periodic_int . } note negative_lx = this { assume "0 \ ?lx" and pi_x: "x - k * (2 * pi) \ pi" with lx have pi_lx: "?lx \ pi" and lx_0: "0 \ real_of_float ?lx" by auto have "cos (x + (-k) * (2 * pi)) \ cos ?lx" using cos_monotone_0_pi_le[OF lx_0 lx pi_x] by (simp only: of_int_minus cos_minus mult_minus_left) simp also have "\ \ (ub_cos prec ?lx)" using lb_cos[OF lx_0 pi_lx] by simp finally have "cos x \ (ub_cos prec ?lx)" unfolding cos_periodic_int . } note positive_lx = this { assume pi_x: "- pi \ x - k * (2 * pi)" and "?ux \ 0" with ux have pi_ux: "- pi \ ?ux" and ux_0: "real_of_float ?ux \ 0" by simp_all have "cos (x + (-k) * (2 * pi)) \ cos (real_of_float (- ?ux))" using cos_monotone_minus_pi_0'[OF pi_x ux ux_0] by (simp only: uminus_float.rep_eq of_int_minus cos_minus mult_minus_left) simp also have "\ \ (ub_cos prec (- ?ux))" using lb_cos_minus[OF pi_ux ux_0, of prec] by simp finally have "cos x \ (ub_cos prec (- ?ux))" unfolding cos_periodic_int . } note negative_ux = this { assume "?ux \ ?lpi" and x_ge_0: "0 \ x - k * (2 * pi)" with lpi ux have pi_ux: "?ux \ pi" and ux_0: "0 \ real_of_float ?ux" by simp_all have "(lb_cos prec ?ux) \ cos ?ux" using lb_cos[OF ux_0 pi_ux] by simp also have "\ \ cos (x + (-k) * (2 * pi))" using cos_monotone_0_pi_le[OF x_ge_0 ux pi_ux] by (simp only: of_int_minus cos_minus mult_minus_left) simp finally have "(lb_cos prec ?ux) \ cos x" unfolding cos_periodic_int . } note positive_ux = this show "l \ cos x \ cos x \ u" proof (cases "- ?lpi \ ?lx \ ?ux \ 0") case True with bnds have l: "l = lb_cos prec (-?lx)" and u: "u = ub_cos prec (-?ux)" by (auto simp add: bnds_cos_def Let_def) from True lpi[THEN le_imp_neg_le] lx ux have "- pi \ x - k * (2 * pi)" and "x - k * (2 * pi) \ 0" by auto with True negative_ux negative_lx show ?thesis unfolding l u by simp next case 1: False show ?thesis proof (cases "0 \ ?lx \ ?ux \ ?lpi") case True with bnds 1 have l: "l = lb_cos prec ?ux" and u: "u = ub_cos prec ?lx" by (auto simp add: bnds_cos_def Let_def) from True lpi lx ux have "0 \ x - k * (2 * pi)" and "x - k * (2 * pi) \ pi" by auto with True positive_ux positive_lx show ?thesis unfolding l u by simp next case 2: False show ?thesis proof (cases "- ?lpi \ ?lx \ ?ux \ ?lpi") case Cond: True with bnds 1 2 have l: "l = min (lb_cos prec (-?lx)) (lb_cos prec ?ux)" and u: "u = Float 1 0" by (auto simp add: bnds_cos_def Let_def) show ?thesis unfolding u l using negative_lx positive_ux Cond by (cases "x - k * (2 * pi) < 0") (auto simp add: real_of_float_min) next case 3: False show ?thesis proof (cases "0 \ ?lx \ ?ux \ 2 * ?lpi") case Cond: True with bnds 1 2 3 have l: "l = Float (- 1) 0" and u: "u = max (ub_cos prec ?lx) (ub_cos prec (- (?ux - 2 * ?lpi)))" by (auto simp add: bnds_cos_def Let_def) have "cos x \ real_of_float u" proof (cases "x - k * (2 * pi) < pi") case True hence "x - k * (2 * pi) \ pi" by simp from positive_lx[OF Cond[THEN conjunct1] this] show ?thesis unfolding u by (simp add: real_of_float_max) next case False hence "pi \ x - k * (2 * pi)" by simp hence pi_x: "- pi \ x - k * (2 * pi) - 2 * pi" by simp have "?ux \ 2 * pi" using Cond lpi by auto hence "x - k * (2 * pi) - 2 * pi \ 0" using ux by simp have ux_0: "real_of_float (?ux - 2 * ?lpi) \ 0" using Cond by auto from 2 and Cond have "\ ?ux \ ?lpi" by auto hence "- ?lpi \ ?ux - 2 * ?lpi" by auto hence pi_ux: "- pi \ (?ux - 2 * ?lpi)" using lpi[THEN le_imp_neg_le] by auto have x_le_ux: "x - k * (2 * pi) - 2 * pi \ (?ux - 2 * ?lpi)" using ux lpi by auto have "cos x = cos (x + (-k) * (2 * pi) + (-1::int) * (2 * pi))" unfolding cos_periodic_int .. also have "\ \ cos ((?ux - 2 * ?lpi))" using cos_monotone_minus_pi_0'[OF pi_x x_le_ux ux_0] by (simp only: minus_float.rep_eq of_int_minus of_int_1 mult_minus_left mult_1_left) simp also have "\ = cos ((- (?ux - 2 * ?lpi)))" unfolding uminus_float.rep_eq cos_minus .. also have "\ \ (ub_cos prec (- (?ux - 2 * ?lpi)))" using lb_cos_minus[OF pi_ux ux_0] by simp finally show ?thesis unfolding u by (simp add: real_of_float_max) qed thus ?thesis unfolding l by auto next case 4: False show ?thesis proof (cases "-2 * ?lpi \ ?lx \ ?ux \ 0") case Cond: True with bnds 1 2 3 4 have l: "l = Float (- 1) 0" and u: "u = max (ub_cos prec (?lx + 2 * ?lpi)) (ub_cos prec (-?ux))" by (auto simp add: bnds_cos_def Let_def) have "cos x \ u" proof (cases "-pi < x - k * (2 * pi)") case True hence "-pi \ x - k * (2 * pi)" by simp from negative_ux[OF this Cond[THEN conjunct2]] show ?thesis unfolding u by (simp add: real_of_float_max) next case False hence "x - k * (2 * pi) \ -pi" by simp hence pi_x: "x - k * (2 * pi) + 2 * pi \ pi" by simp have "-2 * pi \ ?lx" using Cond lpi by auto hence "0 \ x - k * (2 * pi) + 2 * pi" using lx by simp have lx_0: "0 \ real_of_float (?lx + 2 * ?lpi)" using Cond lpi by auto from 1 and Cond have "\ -?lpi \ ?lx" by auto hence "?lx + 2 * ?lpi \ ?lpi" by auto hence pi_lx: "(?lx + 2 * ?lpi) \ pi" using lpi[THEN le_imp_neg_le] by auto have lx_le_x: "(?lx + 2 * ?lpi) \ x - k * (2 * pi) + 2 * pi" using lx lpi by auto have "cos x = cos (x + (-k) * (2 * pi) + (1 :: int) * (2 * pi))" unfolding cos_periodic_int .. also have "\ \ cos ((?lx + 2 * ?lpi))" using cos_monotone_0_pi_le[OF lx_0 lx_le_x pi_x] by (simp only: minus_float.rep_eq of_int_minus of_int_1 mult_minus_left mult_1_left) simp also have "\ \ (ub_cos prec (?lx + 2 * ?lpi))" using lb_cos[OF lx_0 pi_lx] by simp finally show ?thesis unfolding u by (simp add: real_of_float_max) qed thus ?thesis unfolding l by auto next case False with bnds 1 2 3 4 show ?thesis by (auto simp add: bnds_cos_def Let_def) qed qed qed qed qed qed lemma bnds_cos_lower: "\x. real_of_float xl \ x \ x \ real_of_float xu \ cos x \ y \ real_of_float (fst (bnds_cos prec xl xu)) \ y" and bnds_cos_upper: "\x. real_of_float xl \ x \ x \ real_of_float xu \ y \ cos x \ y \ real_of_float (snd (bnds_cos prec xl xu))" for xl xu::float and y::real using bnds_cos[of "fst (bnds_cos prec xl xu)" "snd (bnds_cos prec xl xu)" prec] by force+ lift_definition cos_float_interval :: "nat \ float interval \ float interval" is "\prec. \(lx, ux). bnds_cos prec lx ux" using bnds_cos by auto (metis (full_types) order_refl order_trans) lemma lower_cos_float_interval: "lower (cos_float_interval p x) = fst (bnds_cos p (lower x) (upper x))" by transfer auto lemma upper_cos_float_interval: "upper (cos_float_interval p x) = snd (bnds_cos p (lower x) (upper x))" by transfer auto lemma cos_float_interval: "cos ` set_of (real_interval x) \ set_of (real_interval (cos_float_interval p x))" by (auto simp: set_of_eq bnds_cos_lower bnds_cos_upper lower_cos_float_interval upper_cos_float_interval) lemma cos_float_intervalI: "cos x \\<^sub>r cos_float_interval p X" if "x \\<^sub>r X" using cos_float_interval[of X p] that by auto section "Exponential function" subsection "Compute the series of the exponential function" fun ub_exp_horner :: "nat \ nat \ nat \ nat \ float \ float" and lb_exp_horner :: "nat \ nat \ nat \ nat \ float \ float" where "ub_exp_horner prec 0 i k x = 0" | "ub_exp_horner prec (Suc n) i k x = float_plus_up prec (rapprox_rat prec 1 (int k)) (float_round_up prec (x * lb_exp_horner prec n (i + 1) (k * i) x))" | "lb_exp_horner prec 0 i k x = 0" | "lb_exp_horner prec (Suc n) i k x = float_plus_down prec (lapprox_rat prec 1 (int k)) (float_round_down prec (x * ub_exp_horner prec n (i + 1) (k * i) x))" lemma bnds_exp_horner: assumes "real_of_float x \ 0" shows "exp x \ {lb_exp_horner prec (get_even n) 1 1 x .. ub_exp_horner prec (get_odd n) 1 1 x}" proof - have f_eq: "fact (Suc n) = fact n * ((\i::nat. i + 1) ^^ n) 1" for n proof - have F: "\ m. ((\i. i + 1) ^^ n) m = n + m" by (induct n) auto show ?thesis unfolding F by auto qed note bounds = horner_bounds_nonpos[where f="fact" and lb="lb_exp_horner prec" and ub="ub_exp_horner prec" and j'=0 and s=1, OF assms f_eq lb_exp_horner.simps ub_exp_horner.simps] have "lb_exp_horner prec (get_even n) 1 1 x \ exp x" proof - have "lb_exp_horner prec (get_even n) 1 1 x \ (\j = 0.. \ exp x" proof - obtain t where "\t\ \ \real_of_float x\" and "exp x = (\m = 0.. exp t / (fact (get_even n)) * (real_of_float x) ^ (get_even n)" by (auto simp: zero_le_even_power) ultimately show ?thesis using get_odd exp_gt_zero by auto qed finally show ?thesis . qed moreover have "exp x \ ub_exp_horner prec (get_odd n) 1 1 x" proof - have x_less_zero: "real_of_float x ^ get_odd n \ 0" proof (cases "real_of_float x = 0") case True have "(get_odd n) \ 0" using get_odd[THEN odd_pos] by auto thus ?thesis unfolding True power_0_left by auto next case False hence "real_of_float x < 0" using \real_of_float x \ 0\ by auto show ?thesis by (rule less_imp_le, auto simp add: \real_of_float x < 0\) qed obtain t where "\t\ \ \real_of_float x\" and "exp x = (\m = 0.. 0" by (auto intro!: mult_nonneg_nonpos divide_nonpos_pos simp add: x_less_zero) ultimately have "exp x \ (\j = 0.. \ ub_exp_horner prec (get_odd n) 1 1 x" using bounds(2) by auto finally show ?thesis . qed ultimately show ?thesis by auto qed lemma ub_exp_horner_nonneg: "real_of_float x \ 0 \ 0 \ real_of_float (ub_exp_horner prec (get_odd n) (Suc 0) (Suc 0) x)" using bnds_exp_horner[of x prec n] by (intro order_trans[OF exp_ge_zero]) auto subsection "Compute the exponential function on the entire domain" function ub_exp :: "nat \ float \ float" and lb_exp :: "nat \ float \ float" where "lb_exp prec x = (if 0 < x then float_divl prec 1 (ub_exp prec (-x)) else let horner = (\ x. let y = lb_exp_horner prec (get_even (prec + 2)) 1 1 x in if y \ 0 then Float 1 (- 2) else y) in if x < - 1 then power_down_fl prec (horner (float_divl prec x (- floor_fl x))) (nat (- int_floor_fl x)) else horner x)" | "ub_exp prec x = (if 0 < x then float_divr prec 1 (lb_exp prec (-x)) else if x < - 1 then power_up_fl prec (ub_exp_horner prec (get_odd (prec + 2)) 1 1 (float_divr prec x (- floor_fl x))) (nat (- int_floor_fl x)) else ub_exp_horner prec (get_odd (prec + 2)) 1 1 x)" by pat_completeness auto termination by (relation "measure (\ v. let (prec, x) = case_sum id id v in (if 0 < x then 1 else 0))") auto lemma exp_m1_ge_quarter: "(1 / 4 :: real) \ exp (- 1)" proof - have eq4: "4 = Suc (Suc (Suc (Suc 0)))" by auto have "1 / 4 = (Float 1 (- 2))" unfolding Float_num by auto also have "\ \ lb_exp_horner 3 (get_even 3) 1 1 (- 1)" by (subst less_eq_float.rep_eq [symmetric]) code_simp also have "\ \ exp (- 1 :: float)" using bnds_exp_horner[where x="- 1"] by auto finally show ?thesis by simp qed lemma lb_exp_pos: assumes "\ 0 < x" shows "0 < lb_exp prec x" proof - let "?lb_horner x" = "lb_exp_horner prec (get_even (prec + 2)) 1 1 x" let "?horner x" = "let y = ?lb_horner x in if y \ 0 then Float 1 (- 2) else y" have pos_horner: "0 < ?horner x" for x unfolding Let_def by (cases "?lb_horner x \ 0") auto moreover have "0 < real_of_float ((?horner x) ^ num)" for x :: float and num :: nat proof - have "0 < real_of_float (?horner x) ^ num" using \0 < ?horner x\ by simp also have "\ = (?horner x) ^ num" by auto finally show ?thesis . qed ultimately show ?thesis unfolding lb_exp.simps if_not_P[OF \\ 0 < x\] Let_def by (cases "floor_fl x", cases "x < - 1") (auto simp: real_power_up_fl real_power_down_fl intro!: power_up_less power_down_pos) qed lemma exp_boundaries': assumes "x \ 0" shows "exp x \ { (lb_exp prec x) .. (ub_exp prec x)}" proof - let "?lb_exp_horner x" = "lb_exp_horner prec (get_even (prec + 2)) 1 1 x" let "?ub_exp_horner x" = "ub_exp_horner prec (get_odd (prec + 2)) 1 1 x" have "real_of_float x \ 0" and "\ x > 0" using \x \ 0\ by auto show ?thesis proof (cases "x < - 1") case False hence "- 1 \ real_of_float x" by auto show ?thesis proof (cases "?lb_exp_horner x \ 0") case True from \\ x < - 1\ have "- 1 \ real_of_float x" by auto hence "exp (- 1) \ exp x" unfolding exp_le_cancel_iff . from order_trans[OF exp_m1_ge_quarter this] have "Float 1 (- 2) \ exp x" unfolding Float_num . with True show ?thesis using bnds_exp_horner \real_of_float x \ 0\ \\ x > 0\ \\ x < - 1\ by auto next case False thus ?thesis using bnds_exp_horner \real_of_float x \ 0\ \\ x > 0\ \\ x < - 1\ by (auto simp add: Let_def) qed next case True let ?num = "nat (- int_floor_fl x)" have "real_of_int (int_floor_fl x) < - 1" using int_floor_fl[of x] \x < - 1\ by simp hence "real_of_int (int_floor_fl x) < 0" by simp hence "int_floor_fl x < 0" by auto hence "1 \ - int_floor_fl x" by auto hence "0 < nat (- int_floor_fl x)" by auto hence "0 < ?num" by auto hence "real ?num \ 0" by auto have num_eq: "real ?num = - int_floor_fl x" using \0 < nat (- int_floor_fl x)\ by auto have "0 < - int_floor_fl x" using \0 < ?num\[unfolded of_nat_less_iff[symmetric]] by simp hence "real_of_int (int_floor_fl x) < 0" unfolding less_float_def by auto have fl_eq: "real_of_int (- int_floor_fl x) = real_of_float (- floor_fl x)" by (simp add: floor_fl_def int_floor_fl_def) from \0 < - int_floor_fl x\ have "0 \ real_of_float (- floor_fl x)" by (simp add: floor_fl_def int_floor_fl_def) from \real_of_int (int_floor_fl x) < 0\ have "real_of_float (floor_fl x) < 0" by (simp add: floor_fl_def int_floor_fl_def) have "exp x \ ub_exp prec x" proof - have div_less_zero: "real_of_float (float_divr prec x (- floor_fl x)) \ 0" using float_divr_nonpos_pos_upper_bound[OF \real_of_float x \ 0\ \0 \ real_of_float (- floor_fl x)\] unfolding less_eq_float_def zero_float.rep_eq . have "exp x = exp (?num * (x / ?num))" using \real ?num \ 0\ by auto also have "\ = exp (x / ?num) ^ ?num" unfolding exp_of_nat_mult .. also have "\ \ exp (float_divr prec x (- floor_fl x)) ^ ?num" unfolding num_eq fl_eq by (rule power_mono, rule exp_le_cancel_iff[THEN iffD2], rule float_divr) auto also have "\ \ (?ub_exp_horner (float_divr prec x (- floor_fl x))) ^ ?num" unfolding real_of_float_power by (rule power_mono, rule bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct2], auto) also have "\ \ real_of_float (power_up_fl prec (?ub_exp_horner (float_divr prec x (- floor_fl x))) ?num)" by (auto simp add: real_power_up_fl intro!: power_up ub_exp_horner_nonneg div_less_zero) finally show ?thesis unfolding ub_exp.simps if_not_P[OF \\ 0 < x\] if_P[OF \x < - 1\] floor_fl_def Let_def . qed moreover have "lb_exp prec x \ exp x" proof - let ?divl = "float_divl prec x (- floor_fl x)" let ?horner = "?lb_exp_horner ?divl" show ?thesis proof (cases "?horner \ 0") case False hence "0 \ real_of_float ?horner" by auto have div_less_zero: "real_of_float (float_divl prec x (- floor_fl x)) \ 0" using \real_of_float (floor_fl x) < 0\ \real_of_float x \ 0\ by (auto intro!: order_trans[OF float_divl] divide_nonpos_neg) have "(?lb_exp_horner (float_divl prec x (- floor_fl x))) ^ ?num \ exp (float_divl prec x (- floor_fl x)) ^ ?num" using \0 \ real_of_float ?horner\[unfolded floor_fl_def[symmetric]] bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct1] by (auto intro!: power_mono) also have "\ \ exp (x / ?num) ^ ?num" unfolding num_eq fl_eq using float_divl by (auto intro!: power_mono simp del: uminus_float.rep_eq) also have "\ = exp (?num * (x / ?num))" unfolding exp_of_nat_mult .. also have "\ = exp x" using \real ?num \ 0\ by auto finally show ?thesis using False unfolding lb_exp.simps if_not_P[OF \\ 0 < x\] if_P[OF \x < - 1\] int_floor_fl_def Let_def if_not_P[OF False] by (auto simp: real_power_down_fl intro!: power_down_le) next case True have "power_down_fl prec (Float 1 (- 2)) ?num \ (Float 1 (- 2)) ^ ?num" by (metis Float_le_zero_iff less_imp_le linorder_not_less not_numeral_le_zero numeral_One power_down_fl) then have "power_down_fl prec (Float 1 (- 2)) ?num \ real_of_float (Float 1 (- 2)) ^ ?num" by simp also have "real_of_float (floor_fl x) \ 0" and "real_of_float (floor_fl x) \ 0" using \real_of_float (floor_fl x) < 0\ by auto from divide_right_mono_neg[OF floor_fl[of x] \real_of_float (floor_fl x) \ 0\, unfolded divide_self[OF \real_of_float (floor_fl x) \ 0\]] have "- 1 \ x / (- floor_fl x)" unfolding minus_float.rep_eq by auto from order_trans[OF exp_m1_ge_quarter this[unfolded exp_le_cancel_iff[where x="- 1", symmetric]]] have "Float 1 (- 2) \ exp (x / (- floor_fl x))" unfolding Float_num . hence "real_of_float (Float 1 (- 2)) ^ ?num \ exp (x / (- floor_fl x)) ^ ?num" by (metis Float_num(5) power_mono zero_le_divide_1_iff zero_le_numeral) also have "\ = exp x" unfolding num_eq fl_eq exp_of_nat_mult[symmetric] using \real_of_float (floor_fl x) \ 0\ by auto finally show ?thesis unfolding lb_exp.simps if_not_P[OF \\ 0 < x\] if_P[OF \x < - 1\] int_floor_fl_def Let_def if_P[OF True] real_of_float_power . qed qed ultimately show ?thesis by auto qed qed lemma exp_boundaries: "exp x \ { lb_exp prec x .. ub_exp prec x }" proof - show ?thesis proof (cases "0 < x") case False hence "x \ 0" by auto from exp_boundaries'[OF this] show ?thesis . next case True hence "-x \ 0" by auto have "lb_exp prec x \ exp x" proof - from exp_boundaries'[OF \-x \ 0\] have ub_exp: "exp (- real_of_float x) \ ub_exp prec (-x)" unfolding atLeastAtMost_iff minus_float.rep_eq by auto have "float_divl prec 1 (ub_exp prec (-x)) \ 1 / ub_exp prec (-x)" using float_divl[where x=1] by auto also have "\ \ exp x" using ub_exp[unfolded inverse_le_iff_le[OF order_less_le_trans[OF exp_gt_zero ub_exp] exp_gt_zero, symmetric]] unfolding exp_minus nonzero_inverse_inverse_eq[OF exp_not_eq_zero] inverse_eq_divide by auto finally show ?thesis unfolding lb_exp.simps if_P[OF True] . qed moreover have "exp x \ ub_exp prec x" proof - have "\ 0 < -x" using \0 < x\ by auto from exp_boundaries'[OF \-x \ 0\] have lb_exp: "lb_exp prec (-x) \ exp (- real_of_float x)" unfolding atLeastAtMost_iff minus_float.rep_eq by auto have "exp x \ (1 :: float) / lb_exp prec (-x)" using lb_exp lb_exp_pos[OF \\ 0 < -x\, of prec] by (simp del: lb_exp.simps add: exp_minus field_simps) also have "\ \ float_divr prec 1 (lb_exp prec (-x))" using float_divr . finally show ?thesis unfolding ub_exp.simps if_P[OF True] . qed ultimately show ?thesis by auto qed qed lemma bnds_exp: "\(x::real) lx ux. (l, u) = (lb_exp prec lx, ub_exp prec ux) \ x \ {lx .. ux} \ l \ exp x \ exp x \ u" proof (rule allI, rule allI, rule allI, rule impI) fix x :: real and lx ux assume "(l, u) = (lb_exp prec lx, ub_exp prec ux) \ x \ {lx .. ux}" hence l: "lb_exp prec lx = l " and u: "ub_exp prec ux = u" and x: "x \ {lx .. ux}" by auto show "l \ exp x \ exp x \ u" proof show "l \ exp x" proof - from exp_boundaries[of lx prec, unfolded l] have "l \ exp lx" by (auto simp del: lb_exp.simps) also have "\ \ exp x" using x by auto finally show ?thesis . qed show "exp x \ u" proof - have "exp x \ exp ux" using x by auto also have "\ \ u" using exp_boundaries[of ux prec, unfolded u] by (auto simp del: ub_exp.simps) finally show ?thesis . qed qed qed lemmas [simp del] = lb_exp.simps ub_exp.simps lemma lb_exp: "exp x \ y \ lb_exp prec x \ y" and ub_exp: "y \ exp x \ y \ ub_exp prec x" for x::float and y::real using exp_boundaries[of x prec] by auto lift_definition exp_float_interval :: "nat \ float interval \ float interval" is "\prec. \(lx, ux). (lb_exp prec lx, ub_exp prec ux)" by (auto simp: lb_exp ub_exp) lemma lower_exp_float_interval: "lower (exp_float_interval p x) = lb_exp p (lower x)" by transfer auto lemma upper_exp_float_interval: "upper (exp_float_interval p x) = ub_exp p (upper x)" by transfer auto lemma exp_float_interval: "exp ` set_of (real_interval x) \ set_of (real_interval (exp_float_interval p x))" using exp_boundaries by (auto simp: set_of_eq lower_exp_float_interval upper_exp_float_interval lb_exp ub_exp) lemma exp_float_intervalI: "exp x \\<^sub>r exp_float_interval p X" if "x \\<^sub>r X" using exp_float_interval[of X p] that by auto section "Logarithm" subsection "Compute the logarithm series" fun ub_ln_horner :: "nat \ nat \ nat \ float \ float" and lb_ln_horner :: "nat \ nat \ nat \ float \ float" where "ub_ln_horner prec 0 i x = 0" | "ub_ln_horner prec (Suc n) i x = float_plus_up prec (rapprox_rat prec 1 (int i)) (- float_round_down prec (x * lb_ln_horner prec n (Suc i) x))" | "lb_ln_horner prec 0 i x = 0" | "lb_ln_horner prec (Suc n) i x = float_plus_down prec (lapprox_rat prec 1 (int i)) (- float_round_up prec (x * ub_ln_horner prec n (Suc i) x))" lemma ln_bounds: assumes "0 \ x" and "x < 1" shows "(\i=0..<2*n. (- 1) ^ i * (1 / real (i + 1)) * x ^ (Suc i)) \ ln (x + 1)" (is "?lb") and "ln (x + 1) \ (\i=0..<2*n + 1. (- 1) ^ i * (1 / real (i + 1)) * x ^ (Suc i))" (is "?ub") proof - let "?a n" = "(1/real (n +1)) * x ^ (Suc n)" have ln_eq: "(\ i. (- 1) ^ i * ?a i) = ln (x + 1)" using ln_series[of "x + 1"] \0 \ x\ \x < 1\ by auto have "norm x < 1" using assms by auto have "?a \ 0" unfolding Suc_eq_plus1[symmetric] inverse_eq_divide[symmetric] using tendsto_mult[OF LIMSEQ_inverse_real_of_nat LIMSEQ_Suc[OF LIMSEQ_power_zero[OF \norm x < 1\]]] by auto have "0 \ ?a n" for n by (rule mult_nonneg_nonneg) (auto simp: \0 \ x\) have "?a (Suc n) \ ?a n" for n unfolding inverse_eq_divide[symmetric] proof (rule mult_mono) show "0 \ x ^ Suc (Suc n)" by (auto simp add: \0 \ x\) have "x ^ Suc (Suc n) \ x ^ Suc n * 1" unfolding power_Suc2 mult.assoc[symmetric] by (rule mult_left_mono, fact less_imp_le[OF \x < 1\]) (auto simp: \0 \ x\) thus "x ^ Suc (Suc n) \ x ^ Suc n" by auto qed auto from summable_Leibniz'(2,4)[OF \?a \ 0\ \\n. 0 \ ?a n\, OF \\n. ?a (Suc n) \ ?a n\, unfolded ln_eq] show ?lb and ?ub unfolding atLeast0LessThan by auto qed lemma ln_float_bounds: assumes "0 \ real_of_float x" and "real_of_float x < 1" shows "x * lb_ln_horner prec (get_even n) 1 x \ ln (x + 1)" (is "?lb \ ?ln") and "ln (x + 1) \ x * ub_ln_horner prec (get_odd n) 1 x" (is "?ln \ ?ub") proof - obtain ev where ev: "get_even n = 2 * ev" using get_even_double .. obtain od where od: "get_odd n = 2 * od + 1" using get_odd_double .. let "?s n" = "(- 1) ^ n * (1 / real (1 + n)) * (real_of_float x)^(Suc n)" have "?lb \ sum ?s {0 ..< 2 * ev}" unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq sum_distrib_right[symmetric] unfolding mult.commute[of "real_of_float x"] ev using horner_bounds(1)[where G="\ i k. Suc k" and F="\x. x" and f="\x. x" and lb="\n i k x. lb_ln_horner prec n k x" and ub="\n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*ev", OF \0 \ real_of_float x\ refl lb_ln_horner.simps ub_ln_horner.simps] \0 \ real_of_float x\ unfolding real_of_float_power by (rule mult_right_mono) also have "\ \ ?ln" using ln_bounds(1)[OF \0 \ real_of_float x\ \real_of_float x < 1\] by auto finally show "?lb \ ?ln" . have "?ln \ sum ?s {0 ..< 2 * od + 1}" using ln_bounds(2)[OF \0 \ real_of_float x\ \real_of_float x < 1\] by auto also have "\ \ ?ub" unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq sum_distrib_right[symmetric] unfolding mult.commute[of "real_of_float x"] od using horner_bounds(2)[where G="\ i k. Suc k" and F="\x. x" and f="\x. x" and lb="\n i k x. lb_ln_horner prec n k x" and ub="\n i k x. ub_ln_horner prec n k x" and j'=1 and n="2 * od + 1", OF \0 \ real_of_float x\ refl lb_ln_horner.simps ub_ln_horner.simps] \0 \ real_of_float x\ unfolding real_of_float_power by (rule mult_right_mono) finally show "?ln \ ?ub" . qed lemma ln_add: fixes x :: real assumes "0 < x" and "0 < y" shows "ln (x + y) = ln x + ln (1 + y / x)" proof - have "x \ 0" using assms by auto have "x + y = x * (1 + y / x)" unfolding distrib_left times_divide_eq_right nonzero_mult_div_cancel_left[OF \x \ 0\] by auto moreover have "0 < y / x" using assms by auto hence "0 < 1 + y / x" by auto ultimately show ?thesis using ln_mult assms by auto qed subsection "Compute the logarithm of 2" definition ub_ln2 where "ub_ln2 prec = (let third = rapprox_rat (max prec 1) 1 3 in float_plus_up prec ((Float 1 (- 1) * ub_ln_horner prec (get_odd prec) 1 (Float 1 (- 1)))) (float_round_up prec (third * ub_ln_horner prec (get_odd prec) 1 third)))" definition lb_ln2 where "lb_ln2 prec = (let third = lapprox_rat prec 1 3 in float_plus_down prec ((Float 1 (- 1) * lb_ln_horner prec (get_even prec) 1 (Float 1 (- 1)))) (float_round_down prec (third * lb_ln_horner prec (get_even prec) 1 third)))" lemma ub_ln2: "ln 2 \ ub_ln2 prec" (is "?ub_ln2") and lb_ln2: "lb_ln2 prec \ ln 2" (is "?lb_ln2") proof - let ?uthird = "rapprox_rat (max prec 1) 1 3" let ?lthird = "lapprox_rat prec 1 3" have ln2_sum: "ln 2 = ln (1/2 + 1) + ln (1 / 3 + 1::real)" using ln_add[of "3 / 2" "1 / 2"] by auto have lb3: "?lthird \ 1 / 3" using lapprox_rat[of prec 1 3] by auto hence lb3_ub: "real_of_float ?lthird < 1" by auto have lb3_lb: "0 \ real_of_float ?lthird" using lapprox_rat_nonneg[of 1 3] by auto have ub3: "1 / 3 \ ?uthird" using rapprox_rat[of 1 3] by auto hence ub3_lb: "0 \ real_of_float ?uthird" by auto have lb2: "0 \ real_of_float (Float 1 (- 1))" and ub2: "real_of_float (Float 1 (- 1)) < 1" unfolding Float_num by auto have "0 \ (1::int)" and "0 < (3::int)" by auto have ub3_ub: "real_of_float ?uthird < 1" by (simp add: Float.compute_rapprox_rat Float.compute_lapprox_rat rapprox_posrat_less1) have third_gt0: "(0 :: real) < 1 / 3 + 1" by auto have uthird_gt0: "0 < real_of_float ?uthird + 1" using ub3_lb by auto have lthird_gt0: "0 < real_of_float ?lthird + 1" using lb3_lb by auto show ?ub_ln2 unfolding ub_ln2_def Let_def ln2_sum Float_num(4)[symmetric] proof (rule float_plus_up_le, rule add_mono, fact ln_float_bounds(2)[OF lb2 ub2]) have "ln (1 / 3 + 1) \ ln (real_of_float ?uthird + 1)" unfolding ln_le_cancel_iff[OF third_gt0 uthird_gt0] using ub3 by auto also have "\ \ ?uthird * ub_ln_horner prec (get_odd prec) 1 ?uthird" using ln_float_bounds(2)[OF ub3_lb ub3_ub] . also note float_round_up finally show "ln (1 / 3 + 1) \ float_round_up prec (?uthird * ub_ln_horner prec (get_odd prec) 1 ?uthird)" . qed show ?lb_ln2 unfolding lb_ln2_def Let_def ln2_sum Float_num(4)[symmetric] proof (rule float_plus_down_le, rule add_mono, fact ln_float_bounds(1)[OF lb2 ub2]) have "?lthird * lb_ln_horner prec (get_even prec) 1 ?lthird \ ln (real_of_float ?lthird + 1)" using ln_float_bounds(1)[OF lb3_lb lb3_ub] . note float_round_down_le[OF this] also have "\ \ ln (1 / 3 + 1)" unfolding ln_le_cancel_iff[OF lthird_gt0 third_gt0] using lb3 by auto finally show "float_round_down prec (?lthird * lb_ln_horner prec (get_even prec) 1 ?lthird) \ ln (1 / 3 + 1)" . qed qed subsection "Compute the logarithm in the entire domain" function ub_ln :: "nat \ float \ float option" and lb_ln :: "nat \ float \ float option" where "ub_ln prec x = (if x \ 0 then None else if x < 1 then Some (- the (lb_ln prec (float_divl (max prec 1) 1 x))) else let horner = \x. float_round_up prec (x * ub_ln_horner prec (get_odd prec) 1 x) in if x \ Float 3 (- 1) then Some (horner (x - 1)) else if x < Float 1 1 then Some (float_round_up prec (horner (Float 1 (- 1)) + horner (x * rapprox_rat prec 2 3 - 1))) else let l = bitlen (mantissa x) - 1 in Some (float_plus_up prec (float_round_up prec (ub_ln2 prec * (Float (exponent x + l) 0))) (horner (Float (mantissa x) (- l) - 1))))" | "lb_ln prec x = (if x \ 0 then None else if x < 1 then Some (- the (ub_ln prec (float_divr prec 1 x))) else let horner = \x. float_round_down prec (x * lb_ln_horner prec (get_even prec) 1 x) in if x \ Float 3 (- 1) then Some (horner (x - 1)) else if x < Float 1 1 then Some (float_round_down prec (horner (Float 1 (- 1)) + horner (max (x * lapprox_rat prec 2 3 - 1) 0))) else let l = bitlen (mantissa x) - 1 in Some (float_plus_down prec (float_round_down prec (lb_ln2 prec * (Float (exponent x + l) 0))) (horner (Float (mantissa x) (- l) - 1))))" by pat_completeness auto termination proof (relation "measure (\ v. let (prec, x) = case_sum id id v in (if x < 1 then 1 else 0))", auto) fix prec and x :: float assume "\ real_of_float x \ 0" and "real_of_float x < 1" and "real_of_float (float_divl (max prec (Suc 0)) 1 x) < 1" hence "0 < real_of_float x" "1 \ max prec (Suc 0)" "real_of_float x < 1" by auto from float_divl_pos_less1_bound[OF \0 < real_of_float x\ \real_of_float x < 1\[THEN less_imp_le] \1 \ max prec (Suc 0)\] show False using \real_of_float (float_divl (max prec (Suc 0)) 1 x) < 1\ by auto next fix prec x assume "\ real_of_float x \ 0" and "real_of_float x < 1" and "real_of_float (float_divr prec 1 x) < 1" hence "0 < x" by auto from float_divr_pos_less1_lower_bound[OF \0 < x\, of prec] \real_of_float x < 1\ show False using \real_of_float (float_divr prec 1 x) < 1\ by auto qed lemmas float_pos_eq_mantissa_pos = mantissa_pos_iff[symmetric] lemma Float_pos_eq_mantissa_pos: "Float m e > 0 \ m > 0" using powr_gt_zero[of 2 "e"] by (auto simp add: zero_less_mult_iff zero_float_def simp del: powr_gt_zero dest: less_zeroE) lemma Float_representation_aux: fixes m e defines [THEN meta_eq_to_obj_eq]: "x \ Float m e" assumes "x > 0" shows "Float (exponent x + (bitlen (mantissa x) - 1)) 0 = Float (e + (bitlen m - 1)) 0" (is ?th1) and "Float (mantissa x) (- (bitlen (mantissa x) - 1)) = Float m ( - (bitlen m - 1))" (is ?th2) proof - from assms have mantissa_pos: "m > 0" "mantissa x > 0" using Float_pos_eq_mantissa_pos[of m e] float_pos_eq_mantissa_pos[of x] by simp_all thus ?th1 using bitlen_Float[of m e] assms by (auto simp add: zero_less_mult_iff intro!: arg_cong2[where f=Float]) have "x \ 0" unfolding zero_float_def[symmetric] using \0 < x\ by auto from denormalize_shift[OF x_def this] obtain i where i: "m = mantissa x * 2 ^ i" "e = exponent x - int i" . have "2 powr (1 - (real_of_int (bitlen (mantissa x)) + real_of_int i)) = 2 powr (1 - (real_of_int (bitlen (mantissa x)))) * inverse (2 powr (real i))" by (simp add: powr_minus[symmetric] powr_add[symmetric] field_simps) hence "real_of_int (mantissa x) * 2 powr (1 - real_of_int (bitlen (mantissa x))) = (real_of_int (mantissa x) * 2 ^ i) * 2 powr (1 - real_of_int (bitlen (mantissa x * 2 ^ i)))" using \mantissa x > 0\ by (simp add: powr_realpow) then show ?th2 unfolding i by (auto simp: real_of_float_eq) qed lemma compute_ln[code]: fixes m e defines "x \ Float m e" shows "ub_ln prec x = (if x \ 0 then None else if x < 1 then Some (- the (lb_ln prec (float_divl (max prec 1) 1 x))) else let horner = \x. float_round_up prec (x * ub_ln_horner prec (get_odd prec) 1 x) in if x \ Float 3 (- 1) then Some (horner (x - 1)) else if x < Float 1 1 then Some (float_round_up prec (horner (Float 1 (- 1)) + horner (x * rapprox_rat prec 2 3 - 1))) else let l = bitlen m - 1 in Some (float_plus_up prec (float_round_up prec (ub_ln2 prec * (Float (e + l) 0))) (horner (Float m (- l) - 1))))" (is ?th1) and "lb_ln prec x = (if x \ 0 then None else if x < 1 then Some (- the (ub_ln prec (float_divr prec 1 x))) else let horner = \x. float_round_down prec (x * lb_ln_horner prec (get_even prec) 1 x) in if x \ Float 3 (- 1) then Some (horner (x - 1)) else if x < Float 1 1 then Some (float_round_down prec (horner (Float 1 (- 1)) + horner (max (x * lapprox_rat prec 2 3 - 1) 0))) else let l = bitlen m - 1 in Some (float_plus_down prec (float_round_down prec (lb_ln2 prec * (Float (e + l) 0))) (horner (Float m (- l) - 1))))" (is ?th2) proof - from assms Float_pos_eq_mantissa_pos have "x > 0 \ m > 0" by simp thus ?th1 ?th2 using Float_representation_aux[of m e] unfolding x_def[symmetric] by (auto dest: not_le_imp_less) qed lemma ln_shifted_float: assumes "0 < m" shows "ln (Float m e) = ln 2 * (e + (bitlen m - 1)) + ln (Float m (- (bitlen m - 1)))" proof - let ?B = "2^nat (bitlen m - 1)" define bl where "bl = bitlen m - 1" have "0 < real_of_int m" and "\X. (0 :: real) < 2^X" and "0 < (2 :: real)" and "m \ 0" using assms by auto hence "0 \ bl" by (simp add: bitlen_alt_def bl_def) show ?thesis proof (cases "0 \ e") case True thus ?thesis unfolding bl_def[symmetric] using \0 < real_of_int m\ \0 \ bl\ apply (simp add: ln_mult) apply (cases "e=0") apply (cases "bl = 0", simp_all add: powr_minus ln_inverse ln_powr) apply (cases "bl = 0", simp_all add: powr_minus ln_inverse ln_powr field_simps) done next case False hence "0 < -e" by auto have lne: "ln (2 powr real_of_int e) = ln (inverse (2 powr - e))" by (simp add: powr_minus) hence pow_gt0: "(0::real) < 2^nat (-e)" by auto hence inv_gt0: "(0::real) < inverse (2^nat (-e))" by auto show ?thesis using False unfolding bl_def[symmetric] using \0 < real_of_int m\ \0 \ bl\ by (auto simp add: lne ln_mult ln_powr ln_div field_simps) qed qed lemma ub_ln_lb_ln_bounds': assumes "1 \ x" shows "the (lb_ln prec x) \ ln x \ ln x \ the (ub_ln prec x)" (is "?lb \ ?ln \ ?ln \ ?ub") proof (cases "x < Float 1 1") case True hence "real_of_float (x - 1) < 1" and "real_of_float x < 2" by auto have "\ x \ 0" and "\ x < 1" using \1 \ x\ by auto hence "0 \ real_of_float (x - 1)" using \1 \ x\ by auto have [simp]: "(Float 3 (- 1)) = 3 / 2" by simp show ?thesis proof (cases "x \ Float 3 (- 1)") case True show ?thesis unfolding lb_ln.simps unfolding ub_ln.simps Let_def using ln_float_bounds[OF \0 \ real_of_float (x - 1)\ \real_of_float (x - 1) < 1\, of prec] \\ x \ 0\ \\ x < 1\ True by (auto intro!: float_round_down_le float_round_up_le) next case False hence *: "3 / 2 < x" by auto with ln_add[of "3 / 2" "x - 3 / 2"] have add: "ln x = ln (3 / 2) + ln (real_of_float x * 2 / 3)" by (auto simp add: algebra_simps diff_divide_distrib) let "?ub_horner x" = "float_round_up prec (x * ub_ln_horner prec (get_odd prec) 1 x)" let "?lb_horner x" = "float_round_down prec (x * lb_ln_horner prec (get_even prec) 1 x)" { have up: "real_of_float (rapprox_rat prec 2 3) \ 1" by (rule rapprox_rat_le1) simp_all have low: "2 / 3 \ rapprox_rat prec 2 3" by (rule order_trans[OF _ rapprox_rat]) simp from mult_less_le_imp_less[OF * low] * have pos: "0 < real_of_float (x * rapprox_rat prec 2 3 - 1)" by auto have "ln (real_of_float x * 2/3) \ ln (real_of_float (x * rapprox_rat prec 2 3 - 1) + 1)" proof (rule ln_le_cancel_iff[symmetric, THEN iffD1]) show "real_of_float x * 2 / 3 \ real_of_float (x * rapprox_rat prec 2 3 - 1) + 1" using * low by auto show "0 < real_of_float x * 2 / 3" using * by simp show "0 < real_of_float (x * rapprox_rat prec 2 3 - 1) + 1" using pos by auto qed also have "\ \ ?ub_horner (x * rapprox_rat prec 2 3 - 1)" proof (rule float_round_up_le, rule ln_float_bounds(2)) from mult_less_le_imp_less[OF \real_of_float x < 2\ up] low * show "real_of_float (x * rapprox_rat prec 2 3 - 1) < 1" by auto show "0 \ real_of_float (x * rapprox_rat prec 2 3 - 1)" using pos by auto qed finally have "ln x \ ?ub_horner (Float 1 (-1)) + ?ub_horner ((x * rapprox_rat prec 2 3 - 1))" using ln_float_bounds(2)[of "Float 1 (- 1)" prec prec] add by (auto intro!: add_mono float_round_up_le) note float_round_up_le[OF this, of prec] } moreover { let ?max = "max (x * lapprox_rat prec 2 3 - 1) 0" have up: "lapprox_rat prec 2 3 \ 2/3" by (rule order_trans[OF lapprox_rat], simp) have low: "0 \ real_of_float (lapprox_rat prec 2 3)" using lapprox_rat_nonneg[of 2 3 prec] by simp have "?lb_horner ?max \ ln (real_of_float ?max + 1)" proof (rule float_round_down_le, rule ln_float_bounds(1)) from mult_less_le_imp_less[OF \real_of_float x < 2\ up] * low show "real_of_float ?max < 1" by (cases "real_of_float (lapprox_rat prec 2 3) = 0", auto simp add: real_of_float_max) show "0 \ real_of_float ?max" by (auto simp add: real_of_float_max) qed also have "\ \ ln (real_of_float x * 2/3)" proof (rule ln_le_cancel_iff[symmetric, THEN iffD1]) show "0 < real_of_float ?max + 1" by (auto simp add: real_of_float_max) show "0 < real_of_float x * 2/3" using * by auto show "real_of_float ?max + 1 \ real_of_float x * 2/3" using * up by (cases "0 < real_of_float x * real_of_float (lapprox_posrat prec 2 3) - 1", auto simp add: max_def) qed finally have "?lb_horner (Float 1 (- 1)) + ?lb_horner ?max \ ln x" using ln_float_bounds(1)[of "Float 1 (- 1)" prec prec] add by (auto intro!: add_mono float_round_down_le) note float_round_down_le[OF this, of prec] } ultimately show ?thesis unfolding lb_ln.simps unfolding ub_ln.simps Let_def using \\ x \ 0\ \\ x < 1\ True False by auto qed next case False hence "\ x \ 0" and "\ x < 1" "0 < x" "\ x \ Float 3 (- 1)" using \1 \ x\ by auto show ?thesis proof - define m where "m = mantissa x" define e where "e = exponent x" from Float_mantissa_exponent[of x] have Float: "x = Float m e" by (simp add: m_def e_def) let ?s = "Float (e + (bitlen m - 1)) 0" let ?x = "Float m (- (bitlen m - 1))" have "0 < m" and "m \ 0" using \0 < x\ Float powr_gt_zero[of 2 e] by (auto simp add: zero_less_mult_iff) define bl where "bl = bitlen m - 1" then have bitlen: "bitlen m = bl + 1" by simp hence "bl \ 0" using \m > 0\ by (auto simp add: bitlen_alt_def) have "1 \ Float m e" using \1 \ x\ Float unfolding less_eq_float_def by auto from bitlen_div[OF \0 < m\] float_gt1_scale[OF \1 \ Float m e\] \bl \ 0\ have x_bnds: "0 \ real_of_float (?x - 1)" "real_of_float (?x - 1) < 1" using abs_real_le_2_powr_bitlen [of m] \m > 0\ by (simp_all add: bitlen powr_realpow [symmetric] powr_minus powr_add field_simps) { have "float_round_down prec (lb_ln2 prec * ?s) \ ln 2 * (e + (bitlen m - 1))" (is "real_of_float ?lb2 \ _") apply (rule float_round_down_le) unfolding nat_0 power_0 mult_1_right times_float.rep_eq using lb_ln2[of prec] proof (rule mult_mono) from float_gt1_scale[OF \1 \ Float m e\] show "0 \ real_of_float (Float (e + (bitlen m - 1)) 0)" by simp qed auto moreover from ln_float_bounds(1)[OF x_bnds] have "float_round_down prec ((?x - 1) * lb_ln_horner prec (get_even prec) 1 (?x - 1)) \ ln ?x" (is "real_of_float ?lb_horner \ _") by (auto intro!: float_round_down_le) ultimately have "float_plus_down prec ?lb2 ?lb_horner \ ln x" unfolding Float ln_shifted_float[OF \0 < m\, of e] by (auto intro!: float_plus_down_le) } moreover { from ln_float_bounds(2)[OF x_bnds] have "ln ?x \ float_round_up prec ((?x - 1) * ub_ln_horner prec (get_odd prec) 1 (?x - 1))" (is "_ \ real_of_float ?ub_horner") by (auto intro!: float_round_up_le) moreover have "ln 2 * (e + (bitlen m - 1)) \ float_round_up prec (ub_ln2 prec * ?s)" (is "_ \ real_of_float ?ub2") apply (rule float_round_up_le) unfolding nat_0 power_0 mult_1_right times_float.rep_eq using ub_ln2[of prec] proof (rule mult_mono) from float_gt1_scale[OF \1 \ Float m e\] show "0 \ real_of_int (e + (bitlen m - 1))" by auto have "0 \ ln (2 :: real)" by simp thus "0 \ real_of_float (ub_ln2 prec)" using ub_ln2[of prec] by arith qed auto ultimately have "ln x \ float_plus_up prec ?ub2 ?ub_horner" unfolding Float ln_shifted_float[OF \0 < m\, of e] by (auto intro!: float_plus_up_le) } ultimately show ?thesis unfolding lb_ln.simps unfolding ub_ln.simps unfolding if_not_P[OF \\ x \ 0\] if_not_P[OF \\ x < 1\] if_not_P[OF False] if_not_P[OF \\ x \ Float 3 (- 1)\] Let_def unfolding plus_float.rep_eq e_def[symmetric] m_def[symmetric] by simp qed qed lemma ub_ln_lb_ln_bounds: assumes "0 < x" shows "the (lb_ln prec x) \ ln x \ ln x \ the (ub_ln prec x)" (is "?lb \ ?ln \ ?ln \ ?ub") proof (cases "x < 1") case False hence "1 \ x" unfolding less_float_def less_eq_float_def by auto show ?thesis using ub_ln_lb_ln_bounds'[OF \1 \ x\] . next case True have "\ x \ 0" using \0 < x\ by auto from True have "real_of_float x \ 1" "x \ 1" by simp_all have "0 < real_of_float x" and "real_of_float x \ 0" using \0 < x\ by auto hence A: "0 < 1 / real_of_float x" by auto { let ?divl = "float_divl (max prec 1) 1 x" have A': "1 \ ?divl" using float_divl_pos_less1_bound[OF \0 < real_of_float x\ \real_of_float x \ 1\] by auto hence B: "0 < real_of_float ?divl" by auto have "ln ?divl \ ln (1 / x)" unfolding ln_le_cancel_iff[OF B A] using float_divl[of _ 1 x] by auto hence "ln x \ - ln ?divl" unfolding nonzero_inverse_eq_divide[OF \real_of_float x \ 0\, symmetric] ln_inverse[OF \0 < real_of_float x\] by auto from this ub_ln_lb_ln_bounds'[OF A', THEN conjunct1, THEN le_imp_neg_le] have "?ln \ - the (lb_ln prec ?divl)" unfolding uminus_float.rep_eq by (rule order_trans) } moreover { let ?divr = "float_divr prec 1 x" have A': "1 \ ?divr" using float_divr_pos_less1_lower_bound[OF \0 < x\ \x \ 1\] unfolding less_eq_float_def less_float_def by auto hence B: "0 < real_of_float ?divr" by auto have "ln (1 / x) \ ln ?divr" unfolding ln_le_cancel_iff[OF A B] using float_divr[of 1 x] by auto hence "- ln ?divr \ ln x" unfolding nonzero_inverse_eq_divide[OF \real_of_float x \ 0\, symmetric] ln_inverse[OF \0 < real_of_float x\] by auto from ub_ln_lb_ln_bounds'[OF A', THEN conjunct2, THEN le_imp_neg_le] this have "- the (ub_ln prec ?divr) \ ?ln" unfolding uminus_float.rep_eq by (rule order_trans) } ultimately show ?thesis unfolding lb_ln.simps[where x=x] ub_ln.simps[where x=x] unfolding if_not_P[OF \\ x \ 0\] if_P[OF True] by auto qed lemma lb_ln: assumes "Some y = lb_ln prec x" shows "y \ ln x" and "0 < real_of_float x" proof - have "0 < x" proof (rule ccontr) assume "\ 0 < x" hence "x \ 0" unfolding less_eq_float_def less_float_def by auto thus False using assms by auto qed thus "0 < real_of_float x" by auto have "the (lb_ln prec x) \ ln x" using ub_ln_lb_ln_bounds[OF \0 < x\] .. thus "y \ ln x" unfolding assms[symmetric] by auto qed lemma ub_ln: assumes "Some y = ub_ln prec x" shows "ln x \ y" and "0 < real_of_float x" proof - have "0 < x" proof (rule ccontr) assume "\ 0 < x" hence "x \ 0" by auto thus False using assms by auto qed thus "0 < real_of_float x" by auto have "ln x \ the (ub_ln prec x)" using ub_ln_lb_ln_bounds[OF \0 < x\] .. thus "ln x \ y" unfolding assms[symmetric] by auto qed lemma bnds_ln: "\(x::real) lx ux. (Some l, Some u) = (lb_ln prec lx, ub_ln prec ux) \ x \ {lx .. ux} \ l \ ln x \ ln x \ u" proof (rule allI, rule allI, rule allI, rule impI) fix x :: real fix lx ux assume "(Some l, Some u) = (lb_ln prec lx, ub_ln prec ux) \ x \ {lx .. ux}" hence l: "Some l = lb_ln prec lx " and u: "Some u = ub_ln prec ux" and x: "x \ {lx .. ux}" by auto have "ln ux \ u" and "0 < real_of_float ux" using ub_ln u by auto have "l \ ln lx" and "0 < real_of_float lx" and "0 < x" using lb_ln[OF l] x by auto from ln_le_cancel_iff[OF \0 < real_of_float lx\ \0 < x\] \l \ ln lx\ have "l \ ln x" using x unfolding atLeastAtMost_iff by auto moreover from ln_le_cancel_iff[OF \0 < x\ \0 < real_of_float ux\] \ln ux \ real_of_float u\ have "ln x \ u" using x unfolding atLeastAtMost_iff by auto ultimately show "l \ ln x \ ln x \ u" .. qed lemmas [simp del] = lb_ln.simps ub_ln.simps lemma lb_lnD: "y \ ln x \ 0 < real_of_float x" if "lb_ln prec x = Some y" using lb_ln[OF that[symmetric]] by auto lemma ub_lnD: "ln x \ y\ 0 < real_of_float x" if "ub_ln prec x = Some y" using ub_ln[OF that[symmetric]] by auto lift_definition(code_dt) ln_float_interval :: "nat \ float interval \ float interval option" is "\prec. \(lx, ux). Option.bind (lb_ln prec lx) (\l. Option.bind (ub_ln prec ux) (\u. Some (l, u)))" by (auto simp: pred_option_def bind_eq_Some_conv ln_le_cancel_iff[symmetric] simp del: ln_le_cancel_iff dest!: lb_lnD ub_lnD) lemma ln_float_interval_eq_Some_conv: "ln_float_interval p x = Some y \ lb_ln p (lower x) = Some (lower y) \ ub_ln p (upper x) = Some (upper y)" by transfer (auto simp: bind_eq_Some_conv) lemma ln_float_interval: "ln ` set_of (real_interval x) \ set_of (real_interval y)" if "ln_float_interval p x = Some y" using that lb_ln[of "lower y" p "lower x"] ub_ln[of "lower y" p "lower x"] apply (auto simp add: set_of_eq ln_float_interval_eq_Some_conv ln_le_cancel_iff) apply (meson less_le_trans ln_less_cancel_iff not_le) by (meson less_le_trans ln_less_cancel_iff not_le ub_lnD) lemma ln_float_intervalI: "ln x \ set_of' (ln_float_interval p X)" if "x \\<^sub>r X" using ln_float_interval[of p X] that by (auto simp: set_of'_def split: option.splits) +lemma ln_float_interval_eqI: + "ln x \\<^sub>r IVL" if "ln_float_interval p X = Some IVL" "x \\<^sub>r X" + using ln_float_intervalI[of x X p] that + by (auto simp: set_of'_def split: option.splits) + section \Real power function\ definition bnds_powr :: "nat \ float \ float \ float \ float \ (float \ float) option" where "bnds_powr prec l1 u1 l2 u2 = ( if l1 = 0 \ u1 = 0 then Some (0, 0) else if l1 = 0 \ l2 \ 1 then let uln = the (ub_ln prec u1) in Some (0, ub_exp prec (float_round_up prec (uln * (if uln \ 0 then u2 else l2)))) else if l1 \ 0 then None else Some (map_bnds lb_exp ub_exp prec (bnds_mult prec (the (lb_ln prec l1)) (the (ub_ln prec u1)) l2 u2)))" lemma mono_exp_real: "mono (exp :: real \ real)" by (auto simp: mono_def) lemma ub_exp_nonneg: "real_of_float (ub_exp prec x) \ 0" proof - have "0 \ exp (real_of_float x)" by simp also from exp_boundaries[of x prec] have "\ \ real_of_float (ub_exp prec x)" by simp finally show ?thesis . qed lemma bnds_powr: assumes lu: "Some (l, u) = bnds_powr prec l1 u1 l2 u2" assumes x: "x \ {real_of_float l1..real_of_float u1}" assumes y: "y \ {real_of_float l2..real_of_float u2}" shows "x powr y \ {real_of_float l..real_of_float u}" proof - consider "l1 = 0" "u1 = 0" | "l1 = 0" "u1 \ 0" "l2 \ 1" | "l1 \ 0" "\(l1 = 0 \ (u1 = 0 \ l2 \ 1))" | "l1 > 0" by force thus ?thesis proof cases assume "l1 = 0" "u1 = 0" with x lu show ?thesis by (auto simp: bnds_powr_def) next assume A: "l1 = 0" "u1 \ 0" "l2 \ 1" define uln where "uln = the (ub_ln prec u1)" show ?thesis proof (cases "x = 0") case False with A x y have "x powr y = exp (ln x * y)" by (simp add: powr_def) also { from A x False have "ln x \ ln (real_of_float u1)" by simp also from ub_ln_lb_ln_bounds[of u1 prec] A y x False have "ln (real_of_float u1) \ real_of_float uln" by (simp add: uln_def del: lb_ln.simps) also from A x y have "\ * y \ real_of_float uln * (if uln \ 0 then u2 else l2)" by (auto intro: mult_left_mono mult_left_mono_neg) also have "\ \ real_of_float (float_round_up prec (uln * (if uln \ 0 then u2 else l2)))" by (simp add: float_round_up_le) finally have "ln x * y \ \" using A y by - simp } also have "exp (real_of_float (float_round_up prec (uln * (if uln \ 0 then u2 else l2)))) \ real_of_float (ub_exp prec (float_round_up prec (uln * (if uln \ 0 then u2 else l2))))" using exp_boundaries by simp finally show ?thesis using A x y lu by (simp add: bnds_powr_def uln_def Let_def del: lb_ln.simps ub_ln.simps) qed (insert x y lu A, simp_all add: bnds_powr_def Let_def ub_exp_nonneg del: lb_ln.simps ub_ln.simps) next assume "l1 \ 0" "\(l1 = 0 \ (u1 = 0 \ l2 \ 1))" with lu show ?thesis by (simp add: bnds_powr_def split: if_split_asm) next assume l1: "l1 > 0" obtain lm um where lmum: "(lm, um) = bnds_mult prec (the (lb_ln prec l1)) (the (ub_ln prec u1)) l2 u2" by (cases "bnds_mult prec (the (lb_ln prec l1)) (the (ub_ln prec u1)) l2 u2") simp with l1 have "(l, u) = map_bnds lb_exp ub_exp prec (lm, um)" using lu by (simp add: bnds_powr_def del: lb_ln.simps ub_ln.simps split: if_split_asm) hence "exp (ln x * y) \ {real_of_float l..real_of_float u}" proof (rule map_bnds[OF _ mono_exp_real], goal_cases) case 1 let ?lln = "the (lb_ln prec l1)" and ?uln = "the (ub_ln prec u1)" from ub_ln_lb_ln_bounds[of l1 prec] ub_ln_lb_ln_bounds[of u1 prec] x l1 have "real_of_float ?lln \ ln (real_of_float l1) \ ln (real_of_float u1) \ real_of_float ?uln" by (auto simp del: lb_ln.simps ub_ln.simps) moreover from l1 x have "ln (real_of_float l1) \ ln x \ ln x \ ln (real_of_float u1)" by auto ultimately have ln: "real_of_float ?lln \ ln x \ ln x \ real_of_float ?uln" by simp from lmum show ?case by (rule bnds_mult) (insert y ln, simp_all) qed (insert exp_boundaries[of lm prec] exp_boundaries[of um prec], simp_all) with x l1 show ?thesis by (simp add: powr_def mult_ac) qed qed lift_definition(code_dt) powr_float_interval :: "nat \ float interval \ float interval \ float interval option" is "\prec. \(l1, u1). \(l2, u2). bnds_powr prec l1 u1 l2 u2" by (auto simp: pred_option_def dest!: bnds_powr[OF sym]) lemma powr_float_interval: "{x powr y | x y. x \ set_of (real_interval X) \ y \ set_of (real_interval Y)} \ set_of (real_interval R)" if "powr_float_interval prec X Y = Some R" using that by transfer (auto dest!: bnds_powr[OF sym]) lemma powr_float_intervalI: "x powr y \ set_of' (powr_float_interval p X Y)" if "x \\<^sub>r X" "y \\<^sub>r Y" using powr_float_interval[of p X Y] that by (auto simp: set_of'_def split: option.splits) +lemma powr_float_interval_eqI: + "x powr y \\<^sub>r IVL" + if "powr_float_interval p X Y = Some IVL" "x \\<^sub>r X" "y \\<^sub>r Y" + using powr_float_intervalI[of x X y Y p] that + by (auto simp: set_of'_def split: option.splits) + end end \ No newline at end of file diff --git a/src/HOL/Decision_Procs/approximation.ML b/src/HOL/Decision_Procs/approximation.ML --- a/src/HOL/Decision_Procs/approximation.ML +++ b/src/HOL/Decision_Procs/approximation.ML @@ -1,292 +1,294 @@ (* Title: HOL/Decision_Procs/approximation.ML Author: Johannes Hoelzl, TU Muenchen *) signature APPROXIMATION = sig val reify_form: Proof.context -> term -> term val approx: int -> Proof.context -> term -> term val approximate: Proof.context -> term -> term val approximation_tac : int -> (string * int) list -> int option -> Proof.context -> int -> tactic end -structure Approximation: APPROXIMATION = +structure Approximation = struct fun reorder_bounds_tac ctxt prems i = let fun variable_of_bound (Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\Set.member\, _) $ Free (name, _) $ _)) = name | variable_of_bound (Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ Free (name, _) $ _)) = name | variable_of_bound t = raise TERM ("variable_of_bound", [t]) val variable_bounds = map (`(variable_of_bound o Thm.prop_of)) prems fun add_deps (name, bnds) = Graph.add_deps_acyclic (name, remove (op =) name (Term.add_free_names (Thm.prop_of bnds) [])) val order = Graph.empty |> fold Graph.new_node variable_bounds |> fold add_deps variable_bounds |> Graph.strong_conn |> map the_single |> rev |> map_filter (AList.lookup (op =) variable_bounds) fun prepend_prem th tac = tac THEN resolve_tac ctxt [th RSN (2, @{thm mp})] i in fold prepend_prem order all_tac end fun approximate ctxt t = case fastype_of t of \<^typ>\bool\ => Approximation_Computation.approx_bool ctxt t - | \<^typ>\(float \ float) option\ => + | \<^typ>\(float interval) option\ => Approximation_Computation.approx_arith ctxt t - | \<^typ>\(float \ float) option list\ => + | \<^typ>\(float interval) option list\ => Approximation_Computation.approx_form_eval ctxt t | _ => error ("Bad term: " ^ Syntax.string_of_term ctxt t); fun rewrite_interpret_form_tac ctxt prec splitting taylor i st = let fun lookup_splitting (Free (name, _)) = (case AList.lookup (op =) splitting name of SOME s => HOLogic.mk_number \<^typ>\nat\ s | NONE => \<^term>\0 :: nat\) | lookup_splitting t = raise TERM ("lookup_splitting", [t]) val vs = nth (Thm.prems_of st) (i - 1) |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |> Term.strip_comb |> snd |> List.last |> HOLogic.dest_list val p = prec |> HOLogic.mk_number \<^typ>\nat\ |> Thm.cterm_of ctxt in case taylor of NONE => let val n = vs |> length |> HOLogic.mk_number \<^typ>\nat\ |> Thm.cterm_of ctxt val s = vs |> map lookup_splitting |> HOLogic.mk_list \<^typ>\nat\ |> Thm.cterm_of ctxt in (resolve_tac ctxt [Thm.instantiate ([], [((("n", 0), \<^typ>\nat\), n), ((("prec", 0), \<^typ>\nat\), p), ((("ss", 0), \<^typ>\nat list\), s)]) @{thm approx_form}] i THEN simp_tac (put_simpset (simpset_of \<^context>) ctxt) i) st end | SOME t => if length vs <> 1 then raise (TERM ("More than one variable used for taylor series expansion", [Thm.prop_of st])) else let val t = t |> HOLogic.mk_number \<^typ>\nat\ |> Thm.cterm_of ctxt val s = vs |> map lookup_splitting |> hd |> Thm.cterm_of ctxt in resolve_tac ctxt [Thm.instantiate ([], [((("s", 0), \<^typ>\nat\), s), ((("t", 0), \<^typ>\nat\), t), ((("prec", 0), \<^typ>\nat\), p)]) @{thm approx_tse_form}] i st end end fun calculated_subterms (\<^const>\Trueprop\ $ t) = calculated_subterms t | calculated_subterms (\<^const>\HOL.implies\ $ _ $ t) = calculated_subterms t | calculated_subterms (\<^term>\(\) :: real \ real \ bool\ $ t1 $ t2) = [t1, t2] | calculated_subterms (\<^term>\(<) :: real \ real \ bool\ $ t1 $ t2) = [t1, t2] | calculated_subterms (\<^term>\(\) :: real \ real set \ bool\ $ t1 $ (\<^term>\atLeastAtMost :: real \ real \ real set\ $ t2 $ t3)) = [t1, t2, t3] | calculated_subterms t = raise TERM ("calculated_subterms", [t]) fun dest_interpret_form (\<^const>\interpret_form\ $ b $ xs) = (b, xs) | dest_interpret_form t = raise TERM ("dest_interpret_form", [t]) fun dest_interpret (\<^const>\interpret_floatarith\ $ b $ xs) = (b, xs) | dest_interpret t = raise TERM ("dest_interpret", [t]) fun dest_interpret_env (\<^const>\interpret_form\ $ _ $ xs) = xs | dest_interpret_env (\<^const>\interpret_floatarith\ $ _ $ xs) = xs | dest_interpret_env t = raise TERM ("dest_interpret_env", [t]) fun dest_float (\<^const>\Float\ $ m $ e) = (snd (HOLogic.dest_number m), snd (HOLogic.dest_number e)) | dest_float t = raise TERM ("dest_float", [t]) -fun dest_ivl (Const (\<^const_name>\Some\, _) $ - (Const (\<^const_name>\Pair\, _) $ u $ l)) = SOME (dest_float u, dest_float l) +fun dest_ivl + (Const (\<^const_name>\Some\, _) $ + (Const (\<^const_name>\Interval\, _) $ ((Const (\<^const_name>\Pair\, _) $ u $ l)))) = + SOME (dest_float u, dest_float l) | dest_ivl (Const (\<^const_name>\None\, _)) = NONE | dest_ivl t = raise TERM ("dest_result", [t]) fun mk_approx' prec t = (\<^const>\approx'\ $ HOLogic.mk_number \<^typ>\nat\ prec - $ t $ \<^term>\[] :: (float * float) option list\) + $ t $ \<^term>\[] :: (float interval) option list\) fun mk_approx_form_eval prec t xs = (\<^const>\approx_form_eval\ $ HOLogic.mk_number \<^typ>\nat\ prec $ t $ xs) fun float2_float10 prec round_down (m, e) = ( let val (m, e) = (if e < 0 then (m,e) else (m * Integer.pow e 2, 0)) fun frac _ _ 0 digits cnt = (digits, cnt, 0) | frac _ 0 r digits cnt = (digits, cnt, r) | frac c p r digits cnt = (let val (d, r) = Integer.div_mod (r * 10) (Integer.pow (~e) 2) in frac (c orelse d <> 0) (if d <> 0 orelse c then p - 1 else p) r (digits * 10 + d) (cnt + 1) end) val sgn = Int.sign m val m = abs m val round_down = (sgn = 1 andalso round_down) orelse (sgn = ~1 andalso not round_down) val (x, r) = Integer.div_mod m (Integer.pow (~e) 2) val p = ((if x = 0 then prec else prec - (IntInf.log2 x + 1)) * 3) div 10 + 1 val (digits, e10, r) = if p > 0 then frac (x <> 0) p r 0 0 else (0,0,0) val digits = if round_down orelse r = 0 then digits else digits + 1 in (sgn * (digits + x * (Integer.pow e10 10)), ~e10) end) fun mk_result prec (SOME (l, u)) = (let fun mk_float10 rnd x = (let val (m, e) = float2_float10 prec rnd x in if e = 0 then HOLogic.mk_number \<^typ>\real\ m else if e = 1 then \<^term>\divide :: real \ real \ real\ $ HOLogic.mk_number \<^typ>\real\ m $ \<^term>\10\ else \<^term>\divide :: real \ real \ real\ $ HOLogic.mk_number \<^typ>\real\ m $ (\<^term>\power 10 :: nat \ real\ $ HOLogic.mk_number \<^typ>\nat\ (~e)) end) in \<^term>\atLeastAtMost :: real \ real \ real set\ $ mk_float10 true l $ mk_float10 false u end) | mk_result _ NONE = \<^term>\UNIV :: real set\ fun realify t = let val t = Logic.varify_global t val m = map (fn (name, _) => (name, \<^typ>\real\)) (Term.add_tvars t []) val t = Term.subst_TVars m t in t end fun apply_tactic ctxt term tactic = Thm.cterm_of ctxt term |> Goal.init |> SINGLE tactic |> the |> Thm.prems_of |> hd fun preproc_form_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps (Named_Theorems.get ctxt \<^named_theorems>\approximation_preproc\)) fun reify_form_conv ctxt ct = let val thm = Reification.conv ctxt @{thms interpret_form.simps interpret_floatarith.simps} ct handle ERROR msg => cat_error ("Reification failed: " ^ msg) ("Approximation does not support " ^ quote (Syntax.string_of_term ctxt (Thm.term_of ct))) fun check_env (Free _) = () | check_env (Var _) = () | check_env t = cat_error "Term not supported by approximation:" (Syntax.string_of_term ctxt t) val _ = Thm.rhs_of thm |> Thm.term_of |> dest_interpret_env |> HOLogic.dest_list |> map check_env in thm end fun reify_form_tac ctxt i = CONVERSION (Conv.arg_conv (reify_form_conv ctxt)) i fun prepare_form_tac ctxt i = REPEAT (FIRST' [eresolve_tac ctxt @{thms intervalE}, eresolve_tac ctxt @{thms meta_eqE}, resolve_tac ctxt @{thms impI}] i) THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => reorder_bounds_tac ctxt' prems i) ctxt i THEN DETERM (TRY (filter_prems_tac ctxt (K false) i)) THEN CONVERSION (Conv.arg_conv (preproc_form_conv ctxt)) i fun prepare_form ctxt term = apply_tactic ctxt term (prepare_form_tac ctxt 1) fun apply_reify_form ctxt t = apply_tactic ctxt t (reify_form_tac ctxt 1) fun reify_form ctxt t = HOLogic.mk_Trueprop t |> prepare_form ctxt |> apply_reify_form ctxt |> HOLogic.dest_Trueprop fun approx_form prec ctxt t = realify t |> prepare_form ctxt |> (fn arith_term => apply_reify_form ctxt arith_term |> HOLogic.dest_Trueprop |> dest_interpret_form |> (fn (data, xs) => - mk_approx_form_eval prec data (HOLogic.mk_list \<^typ>\(float * float) option\ - (map (fn _ => \<^term>\None :: (float * float) option\) (HOLogic.dest_list xs))) + mk_approx_form_eval prec data (HOLogic.mk_list \<^typ>\(float interval) option\ + (map (fn _ => \<^term>\None :: (float interval) option\) (HOLogic.dest_list xs))) |> approximate ctxt |> HOLogic.dest_list |> curry ListPair.zip (HOLogic.dest_list xs @ calculated_subterms arith_term) |> map (fn (elem, s) => \<^term>\(\) :: real \ real set \ bool\ $ elem $ mk_result prec (dest_ivl s)) |> foldr1 HOLogic.mk_conj)) fun approx_arith prec ctxt t = realify t |> Thm.cterm_of ctxt |> (preproc_form_conv ctxt then_conv reify_form_conv ctxt) |> Thm.prop_of |> Logic.dest_equals |> snd |> dest_interpret |> fst |> mk_approx' prec |> approximate ctxt |> dest_ivl |> mk_result prec fun approx prec ctxt t = if type_of t = \<^typ>\prop\ then approx_form prec ctxt t else if type_of t = \<^typ>\bool\ then approx_form prec ctxt (\<^const>\Trueprop\ $ t) else approx_arith prec ctxt t fun approximate_cmd modes raw_t state = let val ctxt = Toplevel.context_of state; val t = Syntax.read_term ctxt raw_t; val t' = approx 30 ctxt t; val ty' = Term.type_of t'; val ctxt' = Proof_Context.augment t' ctxt; in Print_Mode.with_modes modes (fn () => Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk, Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) () end |> Pretty.writeln; val opt_modes = Scan.optional (\<^keyword>\(\ |-- Parse.!!! (Scan.repeat1 Parse.name --| \<^keyword>\)\)) []; val _ = Outer_Syntax.command \<^command_keyword>\approximate\ "print approximation of term" (opt_modes -- Parse.term >> (fn (modes, t) => Toplevel.keep (approximate_cmd modes t))); fun approximation_tac prec splitting taylor ctxt = prepare_form_tac ctxt THEN' reify_form_tac ctxt THEN' rewrite_interpret_form_tac ctxt prec splitting taylor THEN' CONVERSION (Approximation_Computation.approx_conv ctxt) THEN' resolve_tac ctxt [TrueI] end; diff --git a/src/HOL/Decision_Procs/approximation_generator.ML b/src/HOL/Decision_Procs/approximation_generator.ML --- a/src/HOL/Decision_Procs/approximation_generator.ML +++ b/src/HOL/Decision_Procs/approximation_generator.ML @@ -1,218 +1,218 @@ (* Title: HOL/Decision_Procs/approximation_generator.ML Author: Fabian Immler, TU Muenchen *) signature APPROXIMATION_GENERATOR = sig val custom_seed: int Config.T val precision: int Config.T val epsilon: real Config.T val approximation_generator: Proof.context -> (term * term list) list -> bool -> int list -> (bool * term list) option * Quickcheck.report option val setup: theory -> theory end; structure Approximation_Generator : APPROXIMATION_GENERATOR = struct val custom_seed = Attrib.setup_config_int \<^binding>\quickcheck_approximation_custom_seed\ (K ~1) val precision = Attrib.setup_config_int \<^binding>\quickcheck_approximation_precision\ (K 30) val epsilon = Attrib.setup_config_real \<^binding>\quickcheck_approximation_epsilon\ (K 0.0) val random_float = @{code "random_class.random::_ \ _ \ (float \ (unit \ term)) \ _"} fun nat_of_term t = (HOLogic.dest_nat t handle TERM _ => snd (HOLogic.dest_number t) handle TERM _ => raise TERM ("nat_of_term", [t])); fun int_of_term t = snd (HOLogic.dest_number t) handle TERM _ => raise TERM ("int_of_term", [t]); fun real_of_man_exp m e = Real.fromManExp {man = Real.fromInt m, exp = e} fun mapprox_float (\<^term>\Float\ $ m $ e) = real_of_man_exp (int_of_term m) (int_of_term e) | mapprox_float t = Real.fromInt (snd (HOLogic.dest_number t)) handle TERM _ => raise TERM ("mapprox_float", [t]); (* TODO: define using compiled terms? *) fun mapprox_floatarith (\<^term>\Add\ $ a $ b) xs = mapprox_floatarith a xs + mapprox_floatarith b xs | mapprox_floatarith (\<^term>\Minus\ $ a) xs = ~ (mapprox_floatarith a xs) | mapprox_floatarith (\<^term>\Mult\ $ a $ b) xs = mapprox_floatarith a xs * mapprox_floatarith b xs | mapprox_floatarith (\<^term>\Inverse\ $ a) xs = 1.0 / mapprox_floatarith a xs | mapprox_floatarith (\<^term>\Cos\ $ a) xs = Math.cos (mapprox_floatarith a xs) | mapprox_floatarith (\<^term>\Arctan\ $ a) xs = Math.atan (mapprox_floatarith a xs) | mapprox_floatarith (\<^term>\Abs\ $ a) xs = abs (mapprox_floatarith a xs) | mapprox_floatarith (\<^term>\Max\ $ a $ b) xs = Real.max (mapprox_floatarith a xs, mapprox_floatarith b xs) | mapprox_floatarith (\<^term>\Min\ $ a $ b) xs = Real.min (mapprox_floatarith a xs, mapprox_floatarith b xs) | mapprox_floatarith \<^term>\Pi\ _ = Math.pi | mapprox_floatarith (\<^term>\Sqrt\ $ a) xs = Math.sqrt (mapprox_floatarith a xs) | mapprox_floatarith (\<^term>\Exp\ $ a) xs = Math.exp (mapprox_floatarith a xs) | mapprox_floatarith (\<^term>\Powr\ $ a $ b) xs = Math.pow (mapprox_floatarith a xs, mapprox_floatarith b xs) | mapprox_floatarith (\<^term>\Ln\ $ a) xs = Math.ln (mapprox_floatarith a xs) | mapprox_floatarith (\<^term>\Power\ $ a $ n) xs = Math.pow (mapprox_floatarith a xs, Real.fromInt (nat_of_term n)) | mapprox_floatarith (\<^term>\Floor\ $ a) xs = Real.fromInt (floor (mapprox_floatarith a xs)) | mapprox_floatarith (\<^term>\Var\ $ n) xs = nth xs (nat_of_term n) | mapprox_floatarith (\<^term>\Num\ $ m) _ = mapprox_float m | mapprox_floatarith t _ = raise TERM ("mapprox_floatarith", [t]) fun mapprox_atLeastAtMost eps x a b xs = let val x' = mapprox_floatarith x xs in mapprox_floatarith a xs + eps <= x' andalso x' + eps <= mapprox_floatarith b xs end fun mapprox_form eps (\<^term>\Bound\ $ x $ a $ b $ f) xs = (not (mapprox_atLeastAtMost eps x a b xs)) orelse mapprox_form eps f xs | mapprox_form eps (\<^term>\Assign\ $ x $ a $ f) xs = (Real.!= (mapprox_floatarith x xs, mapprox_floatarith a xs)) orelse mapprox_form eps f xs | mapprox_form eps (\<^term>\Less\ $ a $ b) xs = mapprox_floatarith a xs + eps < mapprox_floatarith b xs | mapprox_form eps (\<^term>\LessEqual\ $ a $ b) xs = mapprox_floatarith a xs + eps <= mapprox_floatarith b xs | mapprox_form eps (\<^term>\AtLeastAtMost\ $ x $ a $ b) xs = mapprox_atLeastAtMost eps x a b xs | mapprox_form eps (\<^term>\Conj\ $ f $ g) xs = mapprox_form eps f xs andalso mapprox_form eps g xs | mapprox_form eps (\<^term>\Disj\ $ f $ g) xs = mapprox_form eps f xs orelse mapprox_form eps g xs | mapprox_form _ t _ = raise TERM ("mapprox_form", [t]) fun dest_interpret_form (\<^const>\interpret_form\ $ b $ xs) = (b, xs) | dest_interpret_form t = raise TERM ("dest_interpret_form", [t]) fun optionT t = Type (\<^type_name>\option\, [t]) fun mk_Some t = Const (\<^const_name>\Some\, t --> optionT t) fun random_float_list size xs seed = fold (K (apsnd (random_float size) #-> (fn c => apfst (fn b => b::c)))) xs ([],seed) fun real_of_Float (@{code Float} (m, e)) = real_of_man_exp (@{code integer_of_int} m) (@{code integer_of_int} e) fun term_of_int i = (HOLogic.mk_number @{typ int} (@{code integer_of_int} i)) fun term_of_Float (@{code Float} (m, e)) = @{term Float} $ term_of_int m $ term_of_int e fun is_True \<^term>\True\ = true | is_True _ = false val postproc_form_eqs = @{lemma "real_of_float (Float 0 a) = 0" "real_of_float (Float (numeral m) 0) = numeral m" "real_of_float (Float 1 0) = 1" "real_of_float (Float (- 1) 0) = - 1" "real_of_float (Float 1 (numeral e)) = 2 ^ numeral e" "real_of_float (Float 1 (- numeral e)) = 1 / 2 ^ numeral e" "real_of_float (Float a 1) = a * 2" "real_of_float (Float a (-1)) = a / 2" "real_of_float (Float (- a) b) = - real_of_float (Float a b)" "real_of_float (Float (numeral m) (numeral e)) = numeral m * 2 ^ (numeral e)" "real_of_float (Float (numeral m) (- numeral e)) = numeral m / 2 ^ (numeral e)" "- (c * d::real) = -c * d" "- (c / d::real) = -c / d" "- (0::real) = 0" "int_of_integer (numeral k) = numeral k" "int_of_integer (- numeral k) = - numeral k" "int_of_integer 0 = 0" "int_of_integer 1 = 1" "int_of_integer (- 1) = - 1" by auto } fun rewrite_with ctxt thms = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps thms) fun conv_term ctxt conv r = Thm.cterm_of ctxt r |> conv |> Thm.prop_of |> Logic.dest_equals |> snd fun approx_random ctxt prec eps frees e xs genuine_only size seed = let val (rs, seed') = random_float_list size xs seed fun mk_approx_form e ts = \<^const>\approx_form\ $ HOLogic.mk_number \<^typ>\nat\ prec $ e $ - (HOLogic.mk_list \<^typ>\(float * float) option\ - (map (fn t => mk_Some \<^typ>\float * float\ $ HOLogic.mk_prod (t, t)) ts)) $ + (HOLogic.mk_list \<^typ>\(float interval) option\ + (map (fn t => mk_Some \<^typ>\float interval\ $ (\<^term>\interval_of::float\_\ $ t)) ts)) $ \<^term>\[] :: nat list\ in (if mapprox_form eps e (map (real_of_Float o fst) rs) handle General.Overflow => false | General.Domain => false | General.Div => false | IEEEReal.Unordered => false then let val ts = map (fn x => term_of_Float (fst x)) rs val ts' = map (AList.lookup op = (map dest_Free xs ~~ ts) #> the_default Term.dummy #> curry op $ \<^term>\real_of_float::float\_\ #> conv_term ctxt (rewrite_with ctxt postproc_form_eqs)) frees in if Approximation.approximate ctxt (mk_approx_form e ts) |> is_True then SOME (true, ts') else (if genuine_only then NONE else SOME (false, ts')) end else NONE, seed') end val preproc_form_eqs = @{lemma "(a::real) \ {b .. c} \ b \ a \ a \ c" "a = b \ a \ b \ b \ a" "(p \ q) \ \p \ q" "(p \ q) \ (p \ q) \ (q \ p)" by auto} fun reify_goal ctxt t = HOLogic.mk_not t |> conv_term ctxt (rewrite_with ctxt preproc_form_eqs) |> Approximation.reify_form ctxt |> dest_interpret_form ||> HOLogic.dest_list fun approximation_generator_raw ctxt t = let val iterations = Config.get ctxt Quickcheck.iterations val prec = Config.get ctxt precision val eps = Config.get ctxt epsilon val cs = Config.get ctxt custom_seed val seed = (Code_Numeral.natural_of_integer (cs + 1), Code_Numeral.natural_of_integer 1) val run = if cs < 0 then (fn f => fn seed => (Random_Engine.run f, seed)) else (fn f => fn seed => f seed) val frees = Term.add_frees t [] val (e, xs) = reify_goal ctxt t fun single_tester b s = approx_random ctxt prec eps frees e xs b s |> run fun iterate _ _ 0 _ = NONE | iterate genuine_only size j seed = case single_tester genuine_only size seed of (NONE, seed') => iterate genuine_only size (j - 1) seed' | (SOME q, _) => SOME q in fn genuine_only => fn size => (iterate genuine_only size iterations seed, NONE) end fun approximation_generator ctxt [(t, _)] = (fn genuine_only => fn [_, size] => approximation_generator_raw ctxt t genuine_only (Code_Numeral.natural_of_integer size)) | approximation_generator _ _ = error "Quickcheck-approximation does not support type variables (or finite instantiations)" val test_goals = Quickcheck_Common.generator_test_goal_terms ("approximation", (fn _ => fn _ => false, approximation_generator)) val active = Attrib.setup_config_bool \<^binding>\quickcheck_approximation_active\ (K false) val setup = Context.theory_map (Quickcheck.add_tester ("approximation", (active, test_goals))) end diff --git a/src/HOL/Library/Interval.thy b/src/HOL/Library/Interval.thy --- a/src/HOL/Library/Interval.thy +++ b/src/HOL/Library/Interval.thy @@ -1,829 +1,838 @@ (* Title: Interval Author: Christoph Traut, TU Muenchen Fabian Immler, TU Muenchen *) section \Interval Type\ theory Interval imports Complex_Main Lattice_Algebras Set_Algebras begin text \A type of non-empty, closed intervals.\ typedef (overloaded) 'a interval = "{(a::'a::preorder, b). a \ b}" morphisms bounds_of_interval Interval by auto setup_lifting type_definition_interval lift_definition lower::"('a::preorder) interval \ 'a" is fst . lift_definition upper::"('a::preorder) interval \ 'a" is snd . lemma interval_eq_iff: "a = b \ lower a = lower b \ upper a = upper b" by transfer auto lemma interval_eqI: "lower a = lower b \ upper a = upper b \ a = b" by (auto simp: interval_eq_iff) lemma lower_le_upper[simp]: "lower i \ upper i" by transfer auto lift_definition set_of :: "'a::preorder interval \ 'a set" is "\x. {fst x .. snd x}" . lemma set_of_eq: "set_of x = {lower x .. upper x}" by transfer simp context notes [[typedef_overloaded]] begin lift_definition(code_dt) Interval'::"'a::preorder \ 'a::preorder \ 'a interval option" is "\a b. if a \ b then Some (a, b) else None" by auto end instantiation "interval" :: ("{preorder,equal}") equal begin definition "equal_class.equal a b \ (lower a = lower b) \ (upper a = upper b)" instance proof qed (simp add: equal_interval_def interval_eq_iff) end instantiation interval :: ("preorder") ord begin definition less_eq_interval :: "'a interval \ 'a interval \ bool" where "less_eq_interval a b \ lower b \ lower a \ upper a \ upper b" definition less_interval :: "'a interval \ 'a interval \ bool" where "less_interval x y = (x \ y \ \ y \ x)" instance proof qed end instantiation interval :: ("lattice") semilattice_sup begin lift_definition sup_interval :: "'a interval \ 'a interval \ 'a interval" is "\(a, b) (c, d). (inf a c, sup b d)" by (auto simp: le_infI1 le_supI1) lemma lower_sup[simp]: "lower (sup A B) = inf (lower A) (lower B)" by transfer auto lemma upper_sup[simp]: "upper (sup A B) = sup (upper A) (upper B)" by transfer auto instance proof qed (auto simp: less_eq_interval_def less_interval_def interval_eq_iff) end lemma set_of_interval_union: "set_of A \ set_of B \ set_of (sup A B)" for A::"'a::lattice interval" by (auto simp: set_of_eq) lemma interval_union_commute: "sup A B = sup B A" for A::"'a::lattice interval" by (auto simp add: interval_eq_iff inf.commute sup.commute) lemma interval_union_mono1: "set_of a \ set_of (sup a A)" for A :: "'a::lattice interval" using set_of_interval_union by blast lemma interval_union_mono2: "set_of A \ set_of (sup a A)" for A :: "'a::lattice interval" using set_of_interval_union by blast lift_definition interval_of :: "'a::preorder \ 'a interval" is "\x. (x, x)" by auto lemma lower_interval_of[simp]: "lower (interval_of a) = a" by transfer auto lemma upper_interval_of[simp]: "upper (interval_of a) = a" by transfer auto definition width :: "'a::{preorder,minus} interval \ 'a" where "width i = upper i - lower i" instantiation "interval" :: ("ordered_ab_semigroup_add") ab_semigroup_add begin lift_definition plus_interval::"'a interval \ 'a interval \ 'a interval" is "\(a, b). \(c, d). (a + c, b + d)" by (auto intro!: add_mono) lemma lower_plus[simp]: "lower (plus A B) = plus (lower A) (lower B)" by transfer auto lemma upper_plus[simp]: "upper (plus A B) = plus (upper A) (upper B)" by transfer auto instance proof qed (auto simp: interval_eq_iff less_eq_interval_def ac_simps) end instance "interval" :: ("{ordered_ab_semigroup_add, lattice}") ordered_ab_semigroup_add proof qed (auto simp: less_eq_interval_def intro!: add_mono) instantiation "interval" :: ("{preorder,zero}") zero begin lift_definition zero_interval::"'a interval" is "(0, 0)" by auto lemma lower_zero[simp]: "lower 0 = 0" by transfer auto lemma upper_zero[simp]: "upper 0 = 0" by transfer auto instance proof qed end instance "interval" :: ("{ordered_comm_monoid_add}") comm_monoid_add proof qed (auto simp: interval_eq_iff) instance "interval" :: ("{ordered_comm_monoid_add,lattice}") ordered_comm_monoid_add .. instantiation "interval" :: ("{ordered_ab_group_add}") uminus begin lift_definition uminus_interval::"'a interval \ 'a interval" is "\(a, b). (-b, -a)" by auto lemma lower_uminus[simp]: "lower (- A) = - upper A" by transfer auto lemma upper_uminus[simp]: "upper (- A) = - lower A" by transfer auto instance .. end instantiation "interval" :: ("{ordered_ab_group_add}") minus begin definition minus_interval::"'a interval \ 'a interval \ 'a interval" where "minus_interval a b = a + - b" lemma lower_minus[simp]: "lower (minus A B) = minus (lower A) (upper B)" by (auto simp: minus_interval_def) lemma upper_minus[simp]: "upper (minus A B) = minus (upper A) (lower B)" by (auto simp: minus_interval_def) instance .. end instantiation "interval" :: (linordered_semiring) times begin lift_definition times_interval :: "'a interval \ 'a interval \ 'a interval" is "\(a1, a2). \(b1, b2). (let x1 = a1 * b1; x2 = a1 * b2; x3 = a2 * b1; x4 = a2 * b2 in (min x1 (min x2 (min x3 x4)), max x1 (max x2 (max x3 x4))))" by (auto simp: Let_def intro!: min.coboundedI1 max.coboundedI1) lemma lower_times: "lower (times A B) = Min {lower A * lower B, lower A * upper B, upper A * lower B, upper A * upper B}" by transfer (auto simp: Let_def) lemma upper_times: "upper (times A B) = Max {lower A * lower B, lower A * upper B, upper A * lower B, upper A * upper B}" by transfer (auto simp: Let_def) instance .. end lemma interval_eq_set_of_iff: "X = Y \ set_of X = set_of Y" for X Y::"'a::order interval" by (auto simp: set_of_eq interval_eq_iff) subsection \Membership\ abbreviation (in preorder) in_interval ("(_/ \\<^sub>i _)" [51, 51] 50) where "in_interval x X \ x \ set_of X" lemma in_interval_to_interval[intro!]: "a \\<^sub>i interval_of a" by (auto simp: set_of_eq) lemma plus_in_intervalI: fixes x y :: "'a :: ordered_ab_semigroup_add" shows "x \\<^sub>i X \ y \\<^sub>i Y \ x + y \\<^sub>i X + Y" by (simp add: add_mono_thms_linordered_semiring(1) set_of_eq) lemma connected_set_of[intro, simp]: "connected (set_of X)" for X::"'a::linear_continuum_topology interval" by (auto simp: set_of_eq ) lemma ex_sum_in_interval_lemma: "\xa\{la .. ua}. \xb\{lb .. ub}. x = xa + xb" if "la \ ua" "lb \ ub" "la + lb \ x" "x \ ua + ub" "ua - la \ ub - lb" for la b c d::"'a::linordered_ab_group_add" proof - define wa where "wa = ua - la" define wb where "wb = ub - lb" define w where "w = wa + wb" define d where "d = x - la - lb" define da where "da = max 0 (min wa (d - wa))" define db where "db = d - da" from that have nonneg: "0 \ wa" "0 \ wb" "0 \ w" "0 \ d" "d \ w" by (auto simp add: wa_def wb_def w_def d_def add.commute le_diff_eq) have "0 \ db" by (auto simp: da_def nonneg db_def intro!: min.coboundedI2) have "x = (la + da) + (lb + db)" by (simp add: da_def db_def d_def) moreover have "x - la - ub \ da" using that unfolding da_def by (intro max.coboundedI2) (auto simp: wa_def d_def diff_le_eq diff_add_eq) then have "db \ wb" by (auto simp: db_def d_def wb_def algebra_simps) with \0 \ db\ that nonneg have "lb + db \ {lb..ub}" by (auto simp: wb_def algebra_simps) moreover have "da \ wa" by (auto simp: da_def nonneg) then have "la + da \ {la..ua}" by (auto simp: da_def wa_def algebra_simps) ultimately show ?thesis by force qed lemma ex_sum_in_interval: "\xa\la. xa \ ua \ (\xb\lb. xb \ ub \ x = xa + xb)" if a: "la \ ua" and b: "lb \ ub" and x: "la + lb \ x" "x \ ua + ub" for la b c d::"'a::linordered_ab_group_add" proof - from linear consider "ua - la \ ub - lb" | "ub - lb \ ua - la" by blast then show ?thesis proof cases case 1 from ex_sum_in_interval_lemma[OF that 1] show ?thesis by auto next case 2 from x have "lb + la \ x" "x \ ub + ua" by (simp_all add: ac_simps) from ex_sum_in_interval_lemma[OF b a this 2] show ?thesis by auto qed qed lemma Icc_plus_Icc: "{a .. b} + {c .. d} = {a + c .. b + d}" if "a \ b" "c \ d" for a b c d::"'a::linordered_ab_group_add" using ex_sum_in_interval[OF that] by (auto intro: add_mono simp: atLeastAtMost_iff Bex_def set_plus_def) lemma set_of_plus: fixes A :: "'a::linordered_ab_group_add interval" shows "set_of (A + B) = set_of A + set_of B" using Icc_plus_Icc[of "lower A" "upper A" "lower B" "upper B"] by (auto simp: set_of_eq) lemma plus_in_intervalE: fixes xy :: "'a :: linordered_ab_group_add" assumes "xy \\<^sub>i X + Y" obtains x y where "xy = x + y" "x \\<^sub>i X" "y \\<^sub>i Y" using assms unfolding set_of_plus set_plus_def by auto lemma set_of_uminus: "set_of (-X) = {- x | x. x \ set_of X}" for X :: "'a :: ordered_ab_group_add interval" by (auto simp: set_of_eq simp: le_minus_iff minus_le_iff intro!: exI[where x="-x" for x]) lemma uminus_in_intervalI: fixes x :: "'a :: ordered_ab_group_add" shows "x \\<^sub>i X \ -x \\<^sub>i -X" by (auto simp: set_of_uminus) lemma uminus_in_intervalD: fixes x :: "'a :: ordered_ab_group_add" shows "x \\<^sub>i - X \ - x \\<^sub>i X" by (auto simp: set_of_uminus) lemma minus_in_intervalI: fixes x y :: "'a :: ordered_ab_group_add" shows "x \\<^sub>i X \ y \\<^sub>i Y \ x - y \\<^sub>i X - Y" by (metis diff_conv_add_uminus minus_interval_def plus_in_intervalI uminus_in_intervalI) lemma set_of_minus: "set_of (X - Y) = {x - y | x y . x \ set_of X \ y \ set_of Y}" for X Y :: "'a :: linordered_ab_group_add interval" unfolding minus_interval_def set_of_plus set_of_uminus set_plus_def by force lemma times_in_intervalI: fixes x y::"'a::linordered_ring" assumes "x \\<^sub>i X" "y \\<^sub>i Y" shows "x * y \\<^sub>i X * Y" proof - define X1 where "X1 \ lower X" define X2 where "X2 \ upper X" define Y1 where "Y1 \ lower Y" define Y2 where "Y2 \ upper Y" from assms have assms: "X1 \ x" "x \ X2" "Y1 \ y" "y \ Y2" by (auto simp: X1_def X2_def Y1_def Y2_def set_of_eq) have "(X1 * Y1 \ x * y \ X1 * Y2 \ x * y \ X2 * Y1 \ x * y \ X2 * Y2 \ x * y) \ (X1 * Y1 \ x * y \ X1 * Y2 \ x * y \ X2 * Y1 \ x * y \ X2 * Y2 \ x * y)" proof (cases x "0::'a" rule: linorder_cases) case x0: less show ?thesis proof (cases "y < 0") case y0: True from y0 x0 assms have "x * y \ X1 * y" by (intro mult_right_mono_neg, auto) also from x0 y0 assms have "X1 * y \ X1 * Y1" by (intro mult_left_mono_neg, auto) finally have 1: "x * y \ X1 * Y1". show ?thesis proof(cases "X2 \ 0") case True with assms have "X2 * Y2 \ X2 * y" by (auto intro: mult_left_mono_neg) also from assms y0 have "... \ x * y" by (auto intro: mult_right_mono_neg) finally have "X2 * Y2 \ x * y". with 1 show ?thesis by auto next case False with assms have "X2 * Y1 \ X2 * y" by (auto intro: mult_left_mono) also from assms y0 have "... \ x * y" by (auto intro: mult_right_mono_neg) finally have "X2 * Y1 \ x * y". with 1 show ?thesis by auto qed next case False then have y0: "y \ 0" by auto from x0 y0 assms have "X1 * Y2 \ x * Y2" by (intro mult_right_mono, auto) also from y0 x0 assms have "... \ x * y" by (intro mult_left_mono_neg, auto) finally have 1: "X1 * Y2 \ x * y". show ?thesis proof(cases "X2 \ 0") case X2: True from assms y0 have "x * y \ X2 * y" by (intro mult_right_mono) also from assms X2 have "... \ X2 * Y1" by (auto intro: mult_left_mono_neg) finally have "x * y \ X2 * Y1". with 1 show ?thesis by auto next case X2: False from assms y0 have "x * y \ X2 * y" by (intro mult_right_mono) also from assms X2 have "... \ X2 * Y2" by (auto intro: mult_left_mono) finally have "x * y \ X2 * Y2". with 1 show ?thesis by auto qed qed next case [simp]: equal with assms show ?thesis by (cases "Y2 \ 0", auto intro:mult_sign_intros) next case x0: greater show ?thesis proof (cases "y < 0") case y0: True from x0 y0 assms have "X2 * Y1 \ X2 * y" by (intro mult_left_mono, auto) also from y0 x0 assms have "X2 * y \ x * y" by (intro mult_right_mono_neg, auto) finally have 1: "X2 * Y1 \ x * y". show ?thesis proof(cases "Y2 \ 0") case Y2: True from x0 assms have "x * y \ x * Y2" by (auto intro: mult_left_mono) also from assms Y2 have "... \ X1 * Y2" by (auto intro: mult_right_mono_neg) finally have "x * y \ X1 * Y2". with 1 show ?thesis by auto next case Y2: False from x0 assms have "x * y \ x * Y2" by (auto intro: mult_left_mono) also from assms Y2 have "... \ X2 * Y2" by (auto intro: mult_right_mono) finally have "x * y \ X2 * Y2". with 1 show ?thesis by auto qed next case y0: False from x0 y0 assms have "x * y \ X2 * y" by (intro mult_right_mono, auto) also from y0 x0 assms have "... \ X2 * Y2" by (intro mult_left_mono, auto) finally have 1: "x * y \ X2 * Y2". show ?thesis proof(cases "X1 \ 0") case True with assms have "X1 * Y2 \ X1 * y" by (auto intro: mult_left_mono_neg) also from assms y0 have "... \ x * y" by (auto intro: mult_right_mono) finally have "X1 * Y2 \ x * y". with 1 show ?thesis by auto next case False with assms have "X1 * Y1 \ X1 * y" by (auto intro: mult_left_mono) also from assms y0 have "... \ x * y" by (auto intro: mult_right_mono) finally have "X1 * Y1 \ x * y". with 1 show ?thesis by auto qed qed qed hence min:"min (X1 * Y1) (min (X1 * Y2) (min (X2 * Y1) (X2 * Y2))) \ x * y" and max:"x * y \ max (X1 * Y1) (max (X1 * Y2) (max (X2 * Y1) (X2 * Y2)))" by (auto simp:min_le_iff_disj le_max_iff_disj) show ?thesis using min max by (auto simp: Let_def X1_def X2_def Y1_def Y2_def set_of_eq lower_times upper_times) qed lemma times_in_intervalE: fixes xy :: "'a :: {linordered_semiring, real_normed_algebra, linear_continuum_topology}" \ \TODO: linear continuum topology is pretty strong\ assumes "xy \\<^sub>i X * Y" obtains x y where "xy = x * y" "x \\<^sub>i X" "y \\<^sub>i Y" proof - let ?mult = "\(x, y). x * y" let ?XY = "set_of X \ set_of Y" have cont: "continuous_on ?XY ?mult" by (auto intro!: tendsto_eq_intros simp: continuous_on_def split_beta') have conn: "connected (?mult ` ?XY)" by (rule connected_continuous_image[OF cont]) auto have "lower (X * Y) \ ?mult ` ?XY" "upper (X * Y) \ ?mult ` ?XY" by (auto simp: set_of_eq lower_times upper_times min_def max_def split: if_splits) from connectedD_interval[OF conn this, of xy] assms obtain x y where "xy = x * y" "x \\<^sub>i X" "y \\<^sub>i Y" by (auto simp: set_of_eq) then show ?thesis .. qed lemma set_of_times: "set_of (X * Y) = {x * y | x y. x \ set_of X \ y \ set_of Y}" for X Y::"'a :: {linordered_ring, real_normed_algebra, linear_continuum_topology} interval" by (auto intro!: times_in_intervalI elim!: times_in_intervalE) instance "interval" :: (linordered_idom) cancel_semigroup_add proof qed (auto simp: interval_eq_iff) lemma interval_mul_commute: "A * B = B * A" for A B:: "'a::linordered_idom interval" by (simp add: interval_eq_iff lower_times upper_times ac_simps) lemma interval_times_zero_right[simp]: "A * 0 = 0" for A :: "'a::linordered_ring interval" by (simp add: interval_eq_iff lower_times upper_times ac_simps) lemma interval_times_zero_left[simp]: "0 * A = 0" for A :: "'a::linordered_ring interval" by (simp add: interval_eq_iff lower_times upper_times ac_simps) instantiation "interval" :: ("{preorder,one}") one begin lift_definition one_interval::"'a interval" is "(1, 1)" by auto lemma lower_one[simp]: "lower 1 = 1" by transfer auto lemma upper_one[simp]: "upper 1 = 1" by transfer auto instance proof qed end instance interval :: ("{one, preorder, linordered_semiring}") power proof qed lemma set_of_one[simp]: "set_of (1::'a::{one, order} interval) = {1}" by (auto simp: set_of_eq) instance "interval" :: ("{linordered_idom,linordered_ring, real_normed_algebra, linear_continuum_topology}") monoid_mult apply standard unfolding interval_eq_set_of_iff set_of_times subgoal by (auto simp: interval_eq_set_of_iff set_of_times; metis mult.assoc) by auto lemma one_times_ivl_left[simp]: "1 * A = A" for A :: "'a::linordered_idom interval" by (simp add: interval_eq_iff lower_times upper_times ac_simps min_def max_def) lemma one_times_ivl_right[simp]: "A * 1 = A" for A :: "'a::linordered_idom interval" by (metis interval_mul_commute one_times_ivl_left) lemma set_of_power_mono: "a^n \ set_of (A^n)" if "a \ set_of A" for a :: "'a::linordered_idom" using that by (induction n) (auto intro!: times_in_intervalI) lemma set_of_add_cong: "set_of (A + B) = set_of (A' + B')" if "set_of A = set_of A'" "set_of B = set_of B'" for A :: "'a::linordered_ab_group_add interval" unfolding set_of_plus that .. lemma set_of_add_inc_left: "set_of (A + B) \ set_of (A' + B)" if "set_of A \ set_of A'" for A :: "'a::linordered_ab_group_add interval" unfolding set_of_plus using that by (auto simp: set_plus_def) lemma set_of_add_inc_right: "set_of (A + B) \ set_of (A + B')" if "set_of B \ set_of B'" for A :: "'a::linordered_ab_group_add interval" using set_of_add_inc_left[OF that] by (simp add: add.commute) lemma set_of_add_inc: "set_of (A + B) \ set_of (A' + B')" if "set_of A \ set_of A'" "set_of B \ set_of B'" for A :: "'a::linordered_ab_group_add interval" using set_of_add_inc_left[OF that(1)] set_of_add_inc_right[OF that(2)] by auto lemma set_of_neg_inc: "set_of (-A) \ set_of (-A')" if "set_of A \ set_of A'" for A :: "'a::ordered_ab_group_add interval" using that unfolding set_of_uminus by auto lemma set_of_sub_inc_left: "set_of (A - B) \ set_of (A' - B)" if "set_of A \ set_of A'" for A :: "'a::linordered_ab_group_add interval" using that unfolding set_of_minus by auto lemma set_of_sub_inc_right: "set_of (A - B) \ set_of (A - B')" if "set_of B \ set_of B'" for A :: "'a::linordered_ab_group_add interval" using that unfolding set_of_minus by auto lemma set_of_sub_inc: "set_of (A - B) \ set_of (A' - B')" if "set_of A \ set_of A'" "set_of B \ set_of B'" for A :: "'a::linordered_idom interval" using set_of_sub_inc_left[OF that(1)] set_of_sub_inc_right[OF that(2)] by auto lemma set_of_mul_inc_right: "set_of (A * B) \ set_of (A * B')" if "set_of B \ set_of B'" for A :: "'a::linordered_ring interval" using that apply transfer apply (clarsimp simp add: Let_def) apply (intro conjI) apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) done lemma set_of_distrib_left: "set_of (B * (A1 + A2)) \ set_of (B * A1 + B * A2)" for A1 :: "'a::linordered_ring interval" apply transfer apply (clarsimp simp: Let_def distrib_left distrib_right) apply (intro conjI) apply (metis add_mono min.cobounded1 min.left_commute) apply (metis add_mono min.cobounded1 min.left_commute) apply (metis add_mono min.cobounded1 min.left_commute) apply (metis add_mono min.assoc min.cobounded2) apply (meson add_mono order.trans max.cobounded1 max.cobounded2) apply (meson add_mono order.trans max.cobounded1 max.cobounded2) apply (meson add_mono order.trans max.cobounded1 max.cobounded2) apply (meson add_mono order.trans max.cobounded1 max.cobounded2) done lemma set_of_distrib_right: "set_of ((A1 + A2) * B) \ set_of (A1 * B + A2 * B)" for A1 A2 B :: "'a::{linordered_ring, real_normed_algebra, linear_continuum_topology} interval" unfolding set_of_times set_of_plus set_plus_def apply clarsimp subgoal for b a1 a2 apply (rule exI[where x="a1 * b"]) apply (rule conjI) subgoal by force subgoal apply (rule exI[where x="a2 * b"]) apply (rule conjI) subgoal by force subgoal by (simp add: algebra_simps) done done done lemma set_of_mul_inc_left: "set_of (A * B) \ set_of (A' * B)" if "set_of A \ set_of A'" for A :: "'a::{linordered_ring, real_normed_algebra, linear_continuum_topology} interval" using that unfolding set_of_times by auto lemma set_of_mul_inc: "set_of (A * B) \ set_of (A' * B')" if "set_of A \ set_of A'" "set_of B \ set_of B'" for A :: "'a::{linordered_ring, real_normed_algebra, linear_continuum_topology} interval" using that unfolding set_of_times by auto lemma set_of_pow_inc: "set_of (A^n) \ set_of (A'^n)" if "set_of A \ set_of A'" for A :: "'a::{linordered_idom, real_normed_algebra, linear_continuum_topology} interval" using that by (induction n, simp_all add: set_of_mul_inc) lemma set_of_distrib_right_left: "set_of ((A1 + A2) * (B1 + B2)) \ set_of (A1 * B1 + A1 * B2 + A2 * B1 + A2 * B2)" for A1 :: "'a::{linordered_idom, real_normed_algebra, linear_continuum_topology} interval" proof- have "set_of ((A1 + A2) * (B1 + B2)) \ set_of (A1 * (B1 + B2) + A2 * (B1 + B2))" by (rule set_of_distrib_right) also have "... \ set_of ((A1 * B1 + A1 * B2) + A2 * (B1 + B2))" by (rule set_of_add_inc_left[OF set_of_distrib_left]) also have "... \ set_of ((A1 * B1 + A1 * B2) + (A2 * B1 + A2 * B2))" by (rule set_of_add_inc_right[OF set_of_distrib_left]) finally show ?thesis by (simp add: add.assoc) qed lemma mult_bounds_enclose_zero1: "min (la * lb) (min (la * ub) (min (lb * ua) (ua * ub))) \ 0" "0 \ max (la * lb) (max (la * ub) (max (lb * ua) (ua * ub)))" if "la \ 0" "0 \ ua" for la lb ua ub:: "'a::linordered_idom" subgoal by (metis (no_types, hide_lams) that eq_iff min_le_iff_disj mult_zero_left mult_zero_right zero_le_mult_iff) subgoal by (metis that le_max_iff_disj mult_zero_right order_refl zero_le_mult_iff) done lemma mult_bounds_enclose_zero2: "min (la * lb) (min (la * ub) (min (lb * ua) (ua * ub))) \ 0" "0 \ max (la * lb) (max (la * ub) (max (lb * ua) (ua * ub)))" if "lb \ 0" "0 \ ub" for la lb ua ub:: "'a::linordered_idom" using mult_bounds_enclose_zero1[OF that, of la ua] by (simp_all add: ac_simps) lemma set_of_mul_contains_zero: "0 \ set_of (A * B)" if "0 \ set_of A \ 0 \ set_of B" for A :: "'a::linordered_idom interval" using that by (auto simp: set_of_eq lower_times upper_times algebra_simps mult_le_0_iff mult_bounds_enclose_zero1 mult_bounds_enclose_zero2) instance "interval" :: (linordered_semiring) mult_zero apply standard subgoal by transfer auto subgoal by transfer auto done lift_definition min_interval::"'a::linorder interval \ 'a interval \ 'a interval" is "\(l1, u1). \(l2, u2). (min l1 l2, min u1 u2)" by (auto simp: min_def) lemma lower_min_interval[simp]: "lower (min_interval x y) = min (lower x) (lower y)" by transfer auto lemma upper_min_interval[simp]: "upper (min_interval x y) = min (upper x) (upper y)" by transfer auto lemma min_intervalI: "a \\<^sub>i A \ b \\<^sub>i B \ min a b \\<^sub>i min_interval A B" by (auto simp: set_of_eq min_def) lift_definition max_interval::"'a::linorder interval \ 'a interval \ 'a interval" is "\(l1, u1). \(l2, u2). (max l1 l2, max u1 u2)" by (auto simp: max_def) lemma lower_max_interval[simp]: "lower (max_interval x y) = max (lower x) (lower y)" by transfer auto lemma upper_max_interval[simp]: "upper (max_interval x y) = max (upper x) (upper y)" by transfer auto lemma max_intervalI: "a \\<^sub>i A \ b \\<^sub>i B \ max a b \\<^sub>i max_interval A B" by (auto simp: set_of_eq max_def) lift_definition abs_interval::"'a::linordered_idom interval \ 'a interval" is "(\(l,u). (if l < 0 \ 0 < u then 0 else min \l\ \u\, max \l\ \u\))" by auto lemma lower_abs_interval[simp]: "lower (abs_interval x) = (if lower x < 0 \ 0 < upper x then 0 else min \lower x\ \upper x\)" by transfer auto lemma upper_abs_interval[simp]: "upper (abs_interval x) = max \lower x\ \upper x\" by transfer auto lemma in_abs_intervalI1: "lx < 0 \ 0 < ux \ 0 \ xa \ xa \ max (- lx) (ux) \ xa \ abs ` {lx..ux}" for xa::"'a::linordered_idom" by (metis abs_minus_cancel abs_of_nonneg atLeastAtMost_iff image_eqI le_less le_max_iff_disj le_minus_iff neg_le_0_iff_le order_trans) lemma in_abs_intervalI2: "min (\lx\) \ux\ \ xa \ xa \ max \lx\ \ux\ \ lx \ ux \ 0 \ lx \ ux \ 0 \ xa \ abs ` {lx..ux}" for xa::"'a::linordered_idom" by (force intro: image_eqI[where x="-xa"] image_eqI[where x="xa"]) lemma set_of_abs_interval: "set_of (abs_interval x) = abs ` set_of x" by (auto simp: set_of_eq not_less intro: in_abs_intervalI1 in_abs_intervalI2 cong del: image_cong_simp) fun split_domain :: "('a::preorder interval \ 'a interval list) \ 'a interval list \ 'a interval list list" where "split_domain split [] = [[]]" | "split_domain split (I#Is) = ( let S = split I; D = split_domain split Is in concat (map (\d. map (\s. s # d) S) D) )" context notes [[typedef_overloaded]] begin lift_definition(code_dt) split_interval::"'a::linorder interval \ 'a \ ('a interval \ 'a interval)" is "\(l, u) x. ((min l x, max l x), (min u x, max u x))" by (auto simp: min_def) end lemma split_domain_nonempty: assumes "\I. split I \ []" shows "split_domain split I \ []" using last_in_set assms by (induction I, auto) +lemma lower_split_interval1: "lower (fst (split_interval X m)) = min (lower X) m" + and lower_split_interval2: "lower (snd (split_interval X m)) = min (upper X) m" + and upper_split_interval1: "upper (fst (split_interval X m)) = max (lower X) m" + and upper_split_interval2: "upper (snd (split_interval X m)) = max (upper X) m" + subgoal by transfer auto + subgoal by transfer (auto simp: min.commute) + subgoal by transfer (auto simp: ) + subgoal by transfer (auto simp: ) + done lemma split_intervalD: "split_interval X x = (A, B) \ set_of X \ set_of A \ set_of B" unfolding set_of_eq by transfer (auto simp: min_def max_def split: if_splits) instantiation interval :: ("{topological_space, preorder}") topological_space begin definition open_interval_def[code del]: "open (X::'a interval set) = (\x\X. \A B. open A \ open B \ lower x \ A \ upper x \ B \ Interval ` (A \ B) \ X)" instance proof show "open (UNIV :: ('a interval) set)" unfolding open_interval_def by auto next fix S T :: "('a interval) set" assume "open S" "open T" show "open (S \ T)" unfolding open_interval_def proof (safe) fix x assume "x \ S" "x \ T" from \x \ S\ \open S\ obtain Sl Su where S: "open Sl" "open Su" "lower x \ Sl" "upper x \ Su" "Interval ` (Sl \ Su) \ S" by (auto simp: open_interval_def) from \x \ T\ \open T\ obtain Tl Tu where T: "open Tl" "open Tu" "lower x \ Tl" "upper x \ Tu" "Interval ` (Tl \ Tu) \ T" by (auto simp: open_interval_def) let ?L = "Sl \ Tl" and ?U = "Su \ Tu" have "open ?L \ open ?U \ lower x \ ?L \ upper x \ ?U \ Interval ` (?L \ ?U) \ S \ T" using S T by (auto simp add: open_Int) then show "\A B. open A \ open B \ lower x \ A \ upper x \ B \ Interval ` (A \ B) \ S \ T" by fast qed qed (unfold open_interval_def, fast) end subsection \Quickcheck\ lift_definition Ivl::"'a \ 'a::preorder \ 'a interval" is "\a b. (min a b, b)" by (auto simp: min_def) instantiation interval :: ("{exhaustive,preorder}") exhaustive begin definition exhaustive_interval::"('a interval \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "exhaustive_interval f d = Quickcheck_Exhaustive.exhaustive (\x. Quickcheck_Exhaustive.exhaustive (\y. f (Ivl x y)) d) d" instance .. end definition (in term_syntax) [code_unfold]: "valtermify_interval x y = Code_Evaluation.valtermify (Ivl::'a::{preorder,typerep}\_) {\} x {\} y" instantiation interval :: ("{full_exhaustive,preorder,typerep}") full_exhaustive begin definition full_exhaustive_interval:: "('a interval \ (unit \ term) \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "full_exhaustive_interval f d = Quickcheck_Exhaustive.full_exhaustive (\x. Quickcheck_Exhaustive.full_exhaustive (\y. f (valtermify_interval x y)) d) d" instance .. end instantiation interval :: ("{random,preorder,typerep}") random begin definition random_interval :: "natural \ natural \ natural \ ('a interval \ (unit \ term)) \ natural \ natural" where "random_interval i = scomp (Quickcheck_Random.random i) (\man. scomp (Quickcheck_Random.random i) (\exp. Pair (valtermify_interval man exp)))" instance .. end lifting_update interval.lifting lifting_forget interval.lifting end diff --git a/src/HOL/Library/Interval_Float.thy b/src/HOL/Library/Interval_Float.thy --- a/src/HOL/Library/Interval_Float.thy +++ b/src/HOL/Library/Interval_Float.thy @@ -1,289 +1,354 @@ section \Approximate Operations on Intervals of Floating Point Numbers\ theory Interval_Float imports Interval Float begin -definition "split_float_interval x = split_interval x ((lower x + upper x) * Float 1 (-1))" - -lemma split_float_intervalD: "split_float_interval X = (A, B) \ set_of X \ set_of A \ set_of B" - by (auto dest!: split_intervalD simp: split_float_interval_def) - -lemmas float_round_down_le[intro] = order_trans[OF float_round_down] - and float_round_up_ge[intro] = order_trans[OF _ float_round_up] - definition mid :: "float interval \ float" where "mid i = (lower i + upper i) * Float 1 (-1)" lemma mid_in_interval: "mid i \\<^sub>i i" using lower_le_upper[of i] by (auto simp: mid_def set_of_eq powr_minus) +lemma mid_le: "lower i \ mid i" "mid i \ upper i" + using mid_in_interval + by (auto simp: set_of_eq) + definition centered :: "float interval \ float interval" where "centered i = i - interval_of (mid i)" +definition "split_float_interval x = split_interval x ((lower x + upper x) * Float 1 (-1))" + +lemma split_float_intervalD: "split_float_interval X = (A, B) \ set_of X \ set_of A \ set_of B" + by (auto dest!: split_intervalD simp: split_float_interval_def) + +lemma split_float_interval_bounds: + shows + lower_split_float_interval1: "lower (fst (split_float_interval X)) = lower X" + and lower_split_float_interval2: "lower (snd (split_float_interval X)) = mid X" + and upper_split_float_interval1: "upper (fst (split_float_interval X)) = mid X" + and upper_split_float_interval2: "upper (snd (split_float_interval X)) = upper X" + using mid_le[of X] + by (auto simp: split_float_interval_def mid_def[symmetric] min_def max_def real_of_float_eq + lower_split_interval1 lower_split_interval2 + upper_split_interval1 upper_split_interval2) + +lemmas float_round_down_le[intro] = order_trans[OF float_round_down] + and float_round_up_ge[intro] = order_trans[OF _ float_round_up] + text \TODO: many of the lemmas should move to theories Float or Approximation (the latter should be based on type @{type interval}.\ subsection "Intervals with Floating Point Bounds" context includes interval.lifting begin lift_definition round_interval :: "nat \ float interval \ float interval" is "\p. \(l, u). (float_round_down p l, float_round_up p u)" by (auto simp: intro!: float_round_down_le float_round_up_le) lemma lower_round_ivl[simp]: "lower (round_interval p x) = float_round_down p (lower x)" by transfer auto lemma upper_round_ivl[simp]: "upper (round_interval p x) = float_round_up p (upper x)" by transfer auto lemma round_ivl_correct: "set_of A \ set_of (round_interval prec A)" by (auto simp: set_of_eq float_round_down_le float_round_up_le) lift_definition truncate_ivl :: "nat \ real interval \ real interval" is "\p. \(l, u). (truncate_down p l, truncate_up p u)" by (auto intro!: truncate_down_le truncate_up_le) lemma lower_truncate_ivl[simp]: "lower (truncate_ivl p x) = truncate_down p (lower x)" by transfer auto lemma upper_truncate_ivl[simp]: "upper (truncate_ivl p x) = truncate_up p (upper x)" by transfer auto lemma truncate_ivl_correct: "set_of A \ set_of (truncate_ivl prec A)" by (auto simp: set_of_eq intro!: truncate_down_le truncate_up_le) lift_definition real_interval::"float interval \ real interval" is "\(l, u). (real_of_float l, real_of_float u)" by auto lemma lower_real_interval[simp]: "lower (real_interval x) = lower x" by transfer auto lemma upper_real_interval[simp]: "upper (real_interval x) = upper x" by transfer auto definition "set_of' x = (case x of None \ UNIV | Some i \ set_of (real_interval i))" lemma real_interval_min_interval[simp]: "real_interval (min_interval a b) = min_interval (real_interval a) (real_interval b)" by (auto simp: interval_eq_set_of_iff set_of_eq real_of_float_min) lemma real_interval_max_interval[simp]: "real_interval (max_interval a b) = max_interval (real_interval a) (real_interval b)" by (auto simp: interval_eq_set_of_iff set_of_eq real_of_float_max) lemma in_intervalI: "x \\<^sub>i X" if "lower X \ x" "x \ upper X" using that by (auto simp: set_of_eq) abbreviation in_real_interval ("(_/ \\<^sub>r _)" [51, 51] 50) where "x \\<^sub>r X \ x \\<^sub>i real_interval X" lemma in_real_intervalI: "x \\<^sub>r X" if "lower X \ x" "x \ upper X" for x::real and X::"float interval" using that by (intro in_intervalI) auto +subsection \intros for \real_interval\\ + +lemma in_round_intervalI: "x \\<^sub>r A \ x \\<^sub>r (round_interval prec A)" + by (auto simp: set_of_eq float_round_down_le float_round_up_le) + +lemma plus_in_float_intervalI: "a + b \\<^sub>r A + B" if "a \\<^sub>r A" "b \\<^sub>r B" + using that + by (auto simp: set_of_eq) + +lemma minus_in_float_intervalI: "a - b \\<^sub>r A - B" if "a \\<^sub>r A" "b \\<^sub>r B" + using that + by (auto simp: set_of_eq) + +lemma uminus_in_float_intervalI: "-a \\<^sub>r -A" if "a \\<^sub>r A" + using that + by (auto simp: set_of_eq) + +lemma real_interval_times: "real_interval (A * B) = real_interval A * real_interval B" + by (auto simp: interval_eq_iff lower_times upper_times min_def max_def) + +lemma times_in_float_intervalI: "a * b \\<^sub>r A * B" if "a \\<^sub>r A" "b \\<^sub>r B" + using times_in_intervalI[OF that] + by (auto simp: real_interval_times) + +lemma real_interval_abs: "real_interval (abs_interval A) = abs_interval (real_interval A)" + by (auto simp: interval_eq_iff min_def max_def) + +lemma abs_in_float_intervalI: "abs a \\<^sub>r abs_interval A" if "a \\<^sub>r A" + by (auto simp: set_of_abs_interval real_interval_abs intro!: imageI that) + +lemma interval_of[intro,simp]: "x \\<^sub>r interval_of x" + by (auto simp: set_of_eq) + +lemma split_float_interval_realD: "split_float_interval X = (A, B) \ x \\<^sub>r X \ x \\<^sub>r A \ x \\<^sub>r B" + by (auto simp: set_of_eq prod_eq_iff split_float_interval_bounds) + + +subsection \bounds for lists\ + lemma lower_Interval: "lower (Interval x) = fst x" and upper_Interval: "upper (Interval x) = snd x" if "fst x \ snd x" using that by (auto simp: lower_def upper_def Interval_inverse split_beta') definition all_in_i :: "'a::preorder list \ 'a interval list \ bool" (infix "(all'_in\<^sub>i)" 50) where "x all_in\<^sub>i I = (length x = length I \ (\i < length I. x ! i \\<^sub>i I ! i))" definition all_in :: "real list \ float interval list \ bool" (infix "(all'_in)" 50) where "x all_in I = (length x = length I \ (\i < length I. x ! i \\<^sub>r I ! i))" definition all_subset :: "'a::order interval list \ 'a interval list \ bool" (infix "(all'_subset)" 50) where "I all_subset J = (length I = length J \ (\i < length I. set_of (I!i) \ set_of (J!i)))" lemmas [simp] = all_in_def all_subset_def lemma all_subsetD: assumes "I all_subset J" assumes "x all_in I" shows "x all_in J" using assms by (auto simp: set_of_eq; fastforce) lemma round_interval_mono: "set_of (round_interval prec X) \ set_of (round_interval prec Y)" if "set_of X \ set_of Y" using that by transfer (auto simp: float_round_down.rep_eq float_round_up.rep_eq truncate_down_mono truncate_up_mono) lemma Ivl_simps[simp]: "lower (Ivl a b) = min a b" "upper (Ivl a b) = b" subgoal by transfer simp subgoal by transfer simp done lemma set_of_subset_iff: "set_of X \ set_of Y \ lower Y \ lower X \ upper X \ upper Y" for X Y::"'a::linorder interval" by (auto simp: set_of_eq subset_iff) lemma bounds_of_interval_eq_lower_upper: "bounds_of_interval ivl = (lower ivl, upper ivl)" if "lower ivl \ upper ivl" using that by (auto simp: lower.rep_eq upper.rep_eq) lemma real_interval_Ivl: "real_interval (Ivl a b) = Ivl a b" by transfer (auto simp: min_def) lemma set_of_mul_contains_real_zero: "0 \\<^sub>r (A * B)" if "0 \\<^sub>r A \ 0 \\<^sub>r B" using that set_of_mul_contains_zero[of A B] by (auto simp: set_of_eq) fun subdivide_interval :: "nat \ float interval \ float interval list" where "subdivide_interval 0 I = [I]" | "subdivide_interval (Suc n) I = ( let m = mid I in (subdivide_interval n (Ivl (lower I) m)) @ (subdivide_interval n (Ivl m (upper I))) )" lemma subdivide_interval_length: shows "length (subdivide_interval n I) = 2^n" by(induction n arbitrary: I, simp_all add: Let_def) lemma lower_le_mid: "lower x \ mid x" "real_of_float (lower x) \ mid x" and mid_le_upper: "mid x \ upper x" "real_of_float (mid x) \ upper x" unfolding mid_def subgoal by transfer (auto simp: powr_neg_one) subgoal by transfer (auto simp: powr_neg_one) subgoal by transfer (auto simp: powr_neg_one) subgoal by transfer (auto simp: powr_neg_one) done lemma subdivide_interval_correct: "list_ex (\i. x \\<^sub>r i) (subdivide_interval n I)" if "x \\<^sub>r I" for x::real using that proof(induction n arbitrary: x I) case 0 then show ?case by simp next case (Suc n) from \x \\<^sub>r I\ consider "x \\<^sub>r Ivl (lower I) (mid I)" | "x \\<^sub>r Ivl (mid I) (upper I)" by (cases "x \ real_of_float (mid I)") (auto simp: set_of_eq min_def lower_le_mid mid_le_upper) from this[case_names lower upper] show ?case by cases (use Suc.IH in \auto simp: Let_def\) qed fun interval_list_union :: "'a::lattice interval list \ 'a interval" where "interval_list_union [] = undefined" | "interval_list_union [I] = I" | "interval_list_union (I#Is) = sup I (interval_list_union Is)" lemma interval_list_union_correct: assumes "S \ []" assumes "i < length S" shows "set_of (S!i) \ set_of (interval_list_union S)" using assms proof(induction S arbitrary: i) case (Cons a S i) thus ?case proof(cases S) fix b S' assume "S = b # S'" hence "S \ []" by simp show ?thesis proof(cases i) case 0 show ?thesis apply(cases S) using interval_union_mono1 by (auto simp add: 0) next case (Suc i_prev) hence "i_prev < length S" using Cons(3) by simp from Cons(1)[OF \S \ []\ this] Cons(1) have "set_of ((a # S) ! i) \ set_of (interval_list_union S)" by (simp add: \i = Suc i_prev\) also have "... \ set_of (interval_list_union (a # S))" using \S \ []\ apply(cases S) using interval_union_mono2 by auto finally show ?thesis . qed qed simp qed simp lemma split_domain_correct: fixes x :: "real list" assumes "x all_in I" assumes split_correct: "\x a I. x \\<^sub>r I \ list_ex (\i::float interval. x \\<^sub>r i) (split I)" shows "list_ex (\s. x all_in s) (split_domain split I)" using assms(1) proof(induction I arbitrary: x) case (Cons I Is x) have "x \ []" using Cons(2) by auto obtain x' xs where x_decomp: "x = x' # xs" using \x \ []\ list.exhaust by auto hence "x' \\<^sub>r I" "xs all_in Is" using Cons(2) by auto show ?case using Cons(1)[OF \xs all_in Is\] split_correct[OF \x' \\<^sub>r I\] apply (auto simp add: list_ex_iff set_of_eq) by (smt length_Cons less_Suc_eq_0_disj nth_Cons_0 nth_Cons_Suc x_decomp) qed simp lift_definition(code_dt) inverse_float_interval::"nat \ float interval \ float interval option" is "\prec (l, u). if (0 < l \ u < 0) then Some (float_divl prec 1 u, float_divr prec 1 l) else None" by (auto intro!: order_trans[OF float_divl] order_trans[OF _ float_divr] simp: divide_simps) lemma inverse_float_interval_eq_Some_conv: defines "one \ (1::float)" shows "inverse_float_interval p X = Some R \ (lower X > 0 \ upper X < 0) \ lower R = float_divl p one (upper X) \ upper R = float_divr p one (lower X)" by clarsimp (transfer fixing: one, force simp: one_def split: if_splits) lemma inverse_float_interval: "inverse ` set_of (real_interval X) \ set_of (real_interval Y)" if "inverse_float_interval p X = Some Y" using that apply (clarsimp simp: set_of_eq inverse_float_interval_eq_Some_conv) by (intro order_trans[OF float_divl] order_trans[OF _ float_divr] conjI) (auto simp: divide_simps) lemma inverse_float_intervalI: "x \\<^sub>r X \ inverse x \ set_of' (inverse_float_interval p X)" using inverse_float_interval[of p X] by (auto simp: set_of'_def split: option.splits) +lemma inverse_float_interval_eqI: "inverse_float_interval p X = Some IVL \ x \\<^sub>r X \ inverse x \\<^sub>r IVL" + using inverse_float_intervalI[of x X p] + by (auto simp: set_of'_def) + lemma real_interval_abs_interval[simp]: "real_interval (abs_interval x) = abs_interval (real_interval x)" by (auto simp: interval_eq_set_of_iff set_of_eq real_of_float_max real_of_float_min) lift_definition floor_float_interval::"float interval \ float interval" is "\(l, u). (floor_fl l, floor_fl u)" by (auto intro!: floor_mono simp: floor_fl.rep_eq) lemma lower_floor_float_interval[simp]: "lower (floor_float_interval x) = floor_fl (lower x)" by transfer auto lemma upper_floor_float_interval[simp]: "upper (floor_float_interval x) = floor_fl (upper x)" by transfer auto lemma floor_float_intervalI: "\x\ \\<^sub>r floor_float_interval X" if "x \\<^sub>r X" using that by (auto simp: set_of_eq floor_fl_def floor_mono) end + +subsection \constants for code generation\ + +definition lowerF::"float interval \ float" where "lowerF = lower" +definition upperF::"float interval \ float" where "upperF = upper" + + end \ No newline at end of file