# HG changeset patch # User blanchet # Date 1334850542 -7200 # Node ID 075b98ed1cabf5b63d9d3822d277de5dadf6d497 # Parent c0fe12591c93f8866cf464ace4f003e2f829b2ca# Parent b716b16ab2ac498e6461bd3fbc2bb7b9fcc31f1d merged diff -r c0fe12591c93 -r 075b98ed1cab Admin/contributed_components --- a/Admin/contributed_components Thu Apr 19 11:14:57 2012 +0200 +++ b/Admin/contributed_components Thu Apr 19 17:49:02 2012 +0200 @@ -3,8 +3,8 @@ contrib/e-1.4 contrib/hol-light-bundle-0.5-126 contrib/kodkodi-1.2.16 -contrib_devel/spass-3.8ds +contrib/spass-3.8ds contrib/scala-2.9.2 contrib/vampire-1.0 contrib/yices-1.0.28 -contrib_devel/z3-3.2 +contrib/z3-3.2 diff -r c0fe12591c93 -r 075b98ed1cab Admin/isatest/settings/mac-poly-M4 --- a/Admin/isatest/settings/mac-poly-M4 Thu Apr 19 11:14:57 2012 +0200 +++ b/Admin/isatest/settings/mac-poly-M4 Thu Apr 19 17:49:02 2012 +0200 @@ -1,10 +1,10 @@ # -*- shell-script -*- :mode=shellscript: - POLYML_HOME="/home/polyml/polyml-svn" - ML_SYSTEM="polyml-5.4.2" + POLYML_HOME="/home/polyml/polyml-5.4.1" + ML_SYSTEM="polyml-5.4.1" ML_PLATFORM="x86-darwin" ML_HOME="$POLYML_HOME/$ML_PLATFORM" - ML_OPTIONS="-H 500 --gcthreads 4 --gcshare 0" + ML_OPTIONS="-H 1000" ISABELLE_HOME_USER=~/isabelle-mac-poly-M4 diff -r c0fe12591c93 -r 075b98ed1cab Admin/isatest/settings/mac-poly-M8 --- a/Admin/isatest/settings/mac-poly-M8 Thu Apr 19 11:14:57 2012 +0200 +++ b/Admin/isatest/settings/mac-poly-M8 Thu Apr 19 17:49:02 2012 +0200 @@ -1,10 +1,10 @@ # -*- shell-script -*- :mode=shellscript: - POLYML_HOME="/home/polyml/polyml-svn" - ML_SYSTEM="polyml-5.4.2" + POLYML_HOME="/home/polyml/polyml-5.4.1" + ML_SYSTEM="polyml-5.4.1" ML_PLATFORM="x86-darwin" ML_HOME="$POLYML_HOME/$ML_PLATFORM" - ML_OPTIONS="-H 500 --gcthreads 8 --gcshare 0" + ML_OPTIONS="-H 1000" ISABELLE_HOME_USER=~/isabelle-mac-poly-M8 diff -r c0fe12591c93 -r 075b98ed1cab Admin/isatest/settings/mac-poly64-M4 --- a/Admin/isatest/settings/mac-poly64-M4 Thu Apr 19 11:14:57 2012 +0200 +++ b/Admin/isatest/settings/mac-poly64-M4 Thu Apr 19 17:49:02 2012 +0200 @@ -1,10 +1,10 @@ # -*- shell-script -*- :mode=shellscript: - POLYML_HOME="/home/polyml/polyml-svn" - ML_SYSTEM="polyml-5.4.2" + POLYML_HOME="/home/polyml/polyml-5.4.1" + ML_SYSTEM="polyml-5.4.1" ML_PLATFORM="x86_64-darwin" ML_HOME="$POLYML_HOME/$ML_PLATFORM" - ML_OPTIONS="-H 1000 --gcthreads 4 --gcshare 0" + ML_OPTIONS="-H 2000 --gcthreads 4" ISABELLE_HOME_USER=~/isabelle-mac-poly64-M4 diff -r c0fe12591c93 -r 075b98ed1cab Admin/isatest/settings/mac-poly64-M8 --- a/Admin/isatest/settings/mac-poly64-M8 Thu Apr 19 11:14:57 2012 +0200 +++ b/Admin/isatest/settings/mac-poly64-M8 Thu Apr 19 17:49:02 2012 +0200 @@ -1,10 +1,10 @@ # -*- shell-script -*- :mode=shellscript: - POLYML_HOME="/home/polyml/polyml-svn" - ML_SYSTEM="polyml-5.4.2" + POLYML_HOME="/home/polyml/polyml-5.4.1" + ML_SYSTEM="polyml-5.4.1" ML_PLATFORM="x86_64-darwin" ML_HOME="$POLYML_HOME/$ML_PLATFORM" - ML_OPTIONS="-H 1000 --gcthreads 8 --gcshare 0" + ML_OPTIONS="-H 2000 --gcthreads 8" ISABELLE_HOME_USER=~/isabelle-mac-poly64-M8 diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/Archimedean_Field.thy --- a/src/HOL/Archimedean_Field.thy Thu Apr 19 11:14:57 2012 +0200 +++ b/src/HOL/Archimedean_Field.thy Thu Apr 19 17:49:02 2012 +0200 @@ -446,6 +446,17 @@ lemma ceiling_diff_one [simp]: "ceiling (x - 1) = ceiling x - 1" using ceiling_diff_of_int [of x 1] by simp +lemma ceiling_diff_floor_le_1: "ceiling x - floor x \ 1" +proof - + have "of_int \x\ - 1 < x" + using ceiling_correct[of x] by simp + also have "x < of_int \x\ + 1" + using floor_correct[of x] by simp_all + finally have "of_int (\x\ - \x\) < (of_int 2::'a)" + by simp + then show ?thesis + unfolding of_int_less_iff by simp +qed subsection {* Negation *} diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/Decision_Procs/Approximation.thy --- a/src/HOL/Decision_Procs/Approximation.thy Thu Apr 19 11:14:57 2012 +0200 +++ b/src/HOL/Decision_Procs/Approximation.thy Thu Apr 19 17:49:02 2012 +0200 @@ -12,6 +12,9 @@ "~~/src/HOL/Library/Efficient_Nat" begin +declare powr_numeral[simp] +declare powr_neg_numeral[simp] + section "Horner Scheme" subsection {* Define auxiliary helper @{text horner} function *} @@ -54,23 +57,9 @@ case 0 thus ?case unfolding lb_0 ub_0 horner.simps by auto next case (Suc n) - have "?lb (Suc n) j' \ ?horner (Suc n) j'" unfolding lb_Suc ub_Suc horner.simps real_of_float_sub diff_minus - proof (rule add_mono) - show "(lapprox_rat prec 1 (f j')) \ 1 / (f j')" using lapprox_rat[of prec 1 "f j'"] by auto - from Suc[where j'="Suc j'", unfolded funpow.simps comp_def f_Suc, THEN conjunct2] `0 \ real x` - show "- real (x * ub n (F ((F ^^ j') s)) (G ((F ^^ j') s) (f j')) x) \ - - (x * horner F G n (F ((F ^^ j') s)) (G ((F ^^ j') s) (f j')) x)" - unfolding real_of_float_mult neg_le_iff_le by (rule mult_left_mono) - qed - moreover have "?horner (Suc n) j' \ ?ub (Suc n) j'" unfolding ub_Suc horner.simps real_of_float_sub diff_minus - proof (rule add_mono) - show "1 / (f j') \ (rapprox_rat prec 1 (f j'))" using rapprox_rat[of 1 "f j'" prec] by auto - from Suc[where j'="Suc j'", unfolded funpow.simps comp_def f_Suc, THEN conjunct1] `0 \ real x` - show "- (x * horner F G n (F ((F ^^ j') s)) (G ((F ^^ j') s) (f j')) x) \ - - real (x * lb n (F ((F ^^ j') s)) (G ((F ^^ j') s) (f j')) x)" - unfolding real_of_float_mult neg_le_iff_le by (rule mult_left_mono) - qed - ultimately show ?case by blast + 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 x` + by (auto intro!: add_mono mult_left_mono simp add: lb_Suc ub_Suc field_simps f_Suc) qed subsection "Theorems for floating point functions implementing the horner scheme" @@ -106,24 +95,10 @@ shows "(lb n ((F ^^ j') s) (f j') x) \ (\j=0..j=0.. (ub n ((F ^^ j') s) (f j') x)" (is "?ub") proof - - { fix x y z :: float have "x - y * z = x + - y * z" - by (cases x, cases y, cases z, simp add: plus_float.simps minus_float_def times_float.simps algebra_simps) - } note diff_mult_minus = this - - { fix x :: float have "- (- x) = x" by (cases x) auto } note minus_minus = this - - have move_minus: "(-x) = -1 * real x" by auto (* coercion "inside" is necessary *) - + { fix x y z :: float have "x - y * z = x + - y * z" by simp } note diff_mult_minus = this have sum_eq: "(\j=0..j = 0.. {0 ..< n}" - show "1 / (f (j' + j)) * real x ^ j = -1 ^ j * (1 / (f (j' + j))) * real (- x) ^ j" - unfolding move_minus power_mult_distrib mult_assoc[symmetric] - unfolding mult_commute unfolding mult_assoc[of "-1 ^ j", symmetric] power_mult_distrib[symmetric] - by auto - qed - + by (auto simp add: field_simps power_mult_distrib[symmetric]) have "0 \ real (-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, @@ -176,37 +151,39 @@ case True show ?thesis proof (cases "0 < l") - case True hence "odd n \ 0 < l" and "0 \ real l" unfolding less_float_def by auto - have u1: "u1 = u ^ n" and l1: "l1 = l ^ n" using assms unfolding float_power_bnds_def if_P[OF `odd n \ 0 < l`] by auto - have "real l ^ n \ x ^ n" and "x ^ n \ real u ^ n " using `0 \ real l` and assms unfolding atLeastAtMost_iff using power_mono[of l x] power_mono[of x u] by auto - thus ?thesis using assms `0 < l` unfolding atLeastAtMost_iff l1 u1 float_power less_float_def by auto + case True hence "odd n \ 0 < l" and "0 \ real l" by auto + have u1: "u1 = u ^ n" and l1: "l1 = l ^ n" using assms + unfolding float_power_bnds_def if_P[OF `odd n \ 0 < l`] by auto + have "real l ^ n \ x ^ n" and "x ^ n \ real u ^ n " using `0 \ real l` assms + by (auto simp: power_mono) + thus ?thesis using assms `0 < l` unfolding l1 u1 by auto next case False hence P: "\ (odd n \ 0 < l)" using `even n` by auto show ?thesis proof (cases "u < 0") - case True hence "0 \ - real u" and "- real u \ - x" and "0 \ - x" and "-x \ - real l" using assms unfolding less_float_def by auto + case True hence "0 \ - real u" and "- real u \ - x" and "0 \ - x" and "-x \ - real l" using assms by auto hence "real u ^ n \ x ^ n" and "x ^ n \ real l ^ n" using power_mono[of "-x" "-real l" n] power_mono[of "-real u" "-x" n] unfolding power_minus_even[OF `even n`] by auto moreover have u1: "u1 = l ^ n" and l1: "l1 = u ^ n" using assms unfolding float_power_bnds_def if_not_P[OF P] if_P[OF True] by auto - ultimately show ?thesis using float_power by auto + ultimately show ?thesis by auto next case False have "\x\ \ real (max (-l) u)" proof (cases "-l \ u") - case True thus ?thesis unfolding max_def if_P[OF True] using assms unfolding le_float_def by auto + case True thus ?thesis unfolding max_def if_P[OF True] using assms by auto next - case False thus ?thesis unfolding max_def if_not_P[OF False] using assms unfolding le_float_def by auto + case False thus ?thesis unfolding max_def if_not_P[OF False] using assms by auto qed hence x_abs: "\x\ \ \real (max (-l) u)\" by auto have u1: "u1 = (max (-l) u) ^ n" and l1: "l1 = 0" using assms unfolding float_power_bnds_def if_not_P[OF P] if_not_P[OF False] by auto - show ?thesis unfolding atLeastAtMost_iff l1 u1 float_power using zero_le_even_power[OF `even n`] power_mono_even[OF `even n` x_abs] by auto + show ?thesis unfolding atLeastAtMost_iff l1 u1 using zero_le_even_power[OF `even n`] power_mono_even[OF `even n` x_abs] by auto qed qed next case False hence "odd n \ 0 < l" by auto have u1: "u1 = u ^ n" and l1: "l1 = l ^ n" using assms unfolding float_power_bnds_def if_P[OF `odd n \ 0 < l`] by auto have "real l ^ n \ x ^ n" and "x ^ n \ real u ^ n " using assms unfolding atLeastAtMost_iff using power_mono_odd[OF False] by auto - thus ?thesis unfolding atLeastAtMost_iff l1 u1 float_power less_float_def by auto + thus ?thesis unfolding atLeastAtMost_iff l1 u1 less_float_def by auto qed lemma bnds_power: "\ (x::real) l u. (l1, u1) = float_power_bnds n l u \ x \ {l .. u} \ l1 \ x ^ n \ x ^ n \ u1" @@ -222,10 +199,17 @@ *} fun sqrt_iteration :: "nat \ nat \ float \ float" where -"sqrt_iteration prec 0 (Float m e) = Float 1 ((e + bitlen m) div 2 + 1)" | +"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 * (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 * (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) @@ -234,7 +218,7 @@ else if x < 0 then - ub_sqrt prec (- x) else 0)" by pat_completeness auto -termination by (relation "measure (\ v. let (prec, x) = sum_case id id v in (if x < 0 then 1 else 0))", auto simp add: less_float_def) +termination by (relation "measure (\ v. let (prec, x) = sum_case id id v in (if x < 0 then 1 else 0))", auto) declare lb_sqrt.simps[simp del] declare ub_sqrt.simps[simp del] @@ -259,23 +243,23 @@ show ?case proof (cases x) case (Float m e) - hence "0 < m" using float_pos_m_pos[unfolded less_float_def] assms by auto + hence "0 < m" using assms powr_gt_zero[of 2 e] by (auto simp: sign_simps) hence "0 < sqrt m" by auto - have int_nat_bl: "(nat (bitlen m)) = bitlen m" using bitlen_ge0 by auto - - have "x = (m / 2^nat (bitlen m)) * pow2 (e + (nat (bitlen m)))" - unfolding pow2_add pow2_int Float real_of_float_simp by auto - also have "\ < 1 * pow2 (e + nat (bitlen m))" + 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 "real m < 2^nat (bitlen m)" using bitlen_bounds[OF `0 < m`, THEN conjunct2] unfolding real_of_int_less_iff[of m, symmetric] by auto qed - finally have "sqrt x < sqrt (pow2 (e + bitlen m))" unfolding int_nat_bl by auto - also have "\ \ pow2 ((e + bitlen m) div 2 + 1)" + 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: "pow2 (?E mod 2) < 4" + have E_mod_pow: "2 powr (?E mod 2) < 4" proof (cases "?E mod 2 = 1") case True thus ?thesis by auto next @@ -287,21 +271,23 @@ from xt1(5)[OF `0 \ ?E mod 2` this] show ?thesis by auto qed - hence "sqrt (pow2 (?E mod 2)) < sqrt (2 * 2)" by auto - hence E_mod_pow: "sqrt (pow2 (?E mod 2)) < 2" unfolding real_sqrt_abs2 by auto - - have E_eq: "pow2 ?E = pow2 (?E div 2 + ?E div 2 + ?E mod 2)" by auto - have "sqrt (pow2 ?E) = sqrt (pow2 (?E div 2) * pow2 (?E div 2) * pow2 (?E mod 2))" - unfolding E_eq unfolding pow2_add .. - also have "\ = pow2 (?E div 2) * sqrt (pow2 (?E mod 2))" - unfolding real_sqrt_mult[of _ "pow2 (?E mod 2)"] real_sqrt_abs2 by auto - also have "\ < pow2 (?E div 2) * 2" + hence "sqrt (2 powr (?E mod 2)) < sqrt (2 * 2)" by auto + hence E_mod_pow: "sqrt (2 powr (?E mod 2)) < 2" unfolding real_sqrt_abs2 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 (simp add: int_of_reals del: real_of_ints) + 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 "\ = pow2 (?E div 2 + 1)" unfolding add_commute[of _ 1] pow2_add1 by auto + 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 - unfolding Float sqrt_iteration.simps real_of_float_simp by auto + finally show ?thesis using `0 < m` + unfolding Float + by (subst compute_sqrt_iteration_base) (simp add: ac_simps) qed next case (Suc n) @@ -310,8 +296,8 @@ also have "\ < real ?b" using Suc . finally have "sqrt x < (?b + x / ?b)/2" using sqrt_ub_pos_pos_1[OF Suc _ `0 < real 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 auto - finally show ?case unfolding sqrt_iteration.simps Let_def real_of_float_mult real_of_float_add right_distrib . + also have "\ = (Float 1 -1) * (?b + (float_divr prec x ?b))" by simp + finally show ?case unfolding sqrt_iteration.simps Let_def right_distrib . qed lemma sqrt_iteration_lower_bound: assumes "0 < real x" @@ -325,20 +311,20 @@ lemma lb_sqrt_lower_bound: assumes "0 \ real x" shows "0 \ real (lb_sqrt prec x)" proof (cases "0 < x") - case True hence "0 < real x" and "0 \ x" using `0 \ real x` unfolding less_float_def le_float_def by auto - hence "0 < sqrt_iteration prec prec x" unfolding less_float_def using sqrt_iteration_lower_bound by auto - hence "0 \ real (float_divl prec x (sqrt_iteration prec prec x))" using float_divl_lower_bound[OF `0 \ x`] unfolding le_float_def by auto + case True hence "0 < real x" and "0 \ x" using `0 \ real x` by auto + hence "0 < sqrt_iteration prec prec x" using sqrt_iteration_lower_bound by auto + hence "0 \ real (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 x` have "real x = 0" unfolding less_float_def by auto - thus ?thesis unfolding lb_sqrt.simps less_float_def by auto + case False with `0 \ real x` have "real x = 0" by auto + thus ?thesis unfolding lb_sqrt.simps by auto qed lemma bnds_sqrt': shows "sqrt x \ {(lb_sqrt prec x) .. (ub_sqrt prec x) }" proof - { fix x :: float assume "0 < x" - hence "0 < real x" and "0 \ real x" unfolding less_float_def by auto + hence "0 < real x" and "0 \ real 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 @@ -355,7 +341,7 @@ note lb = this { fix x :: float assume "0 < x" - hence "0 < real x" unfolding less_float_def by auto + hence "0 < real x" by auto hence "0 < sqrt x" by auto hence "sqrt x < sqrt_iteration prec prec x" using sqrt_iteration_bound by auto @@ -369,10 +355,10 @@ next case False show ?thesis proof (cases "real x = 0") case True thus ?thesis - by (auto simp add: less_float_def lb_sqrt.simps ub_sqrt.simps) + by (auto simp add: lb_sqrt.simps ub_sqrt.simps) next case False with `\ 0 < x` have "x < 0" and "0 < -x" - by (auto simp add: less_float_def) + by auto with `\ 0 < x` show ?thesis using lb[OF `0 < -x`] ub[OF `0 < -x`] @@ -446,7 +432,7 @@ { have "(x * lb_arctan_horner prec n 1 (x*x)) \ ?S n" using bounds(1) `0 \ real x` - unfolding real_of_float_mult power_add power_one_right mult_assoc[symmetric] setsum_left_distrib[symmetric] + unfolding power_add power_one_right mult_assoc[symmetric] setsum_left_distrib[symmetric] unfolding mult_commute[where 'a=real] mult_commute[of _ "2::nat"] power_mult power2_eq_square[of "real x"] by (auto intro!: mult_left_mono) also have "\ \ arctan x" using arctan_bounds .. @@ -455,7 +441,7 @@ { have "arctan x \ ?S (Suc n)" using arctan_bounds .. also have "\ \ (x * ub_arctan_horner prec (Suc n) 1 (x*x))" using bounds(2)[of "Suc n"] `0 \ real x` - unfolding real_of_float_mult power_add power_one_right mult_assoc[symmetric] setsum_left_distrib[symmetric] + unfolding power_add power_one_right mult_assoc[symmetric] setsum_left_distrib[symmetric] unfolding mult_commute[where 'a=real] mult_commute[of _ "2::nat"] power_mult power2_eq_square[of "real x"] by (auto intro!: mult_left_mono) finally have "arctan x \ (x * ub_arctan_horner prec (Suc n) 1 (x*x))" . } @@ -512,8 +498,8 @@ have "1 div k = 0" using div_pos_pos_trivial[OF _ `1 < k`] by auto have "0 \ real ?k" by (rule order_trans[OF _ rapprox_rat], auto simp add: `0 \ k`) - have "real ?k \ 1" unfolding rapprox_rat.simps(2)[OF zero_le_one `0 < k`] - by (rule rapprox_posrat_le1, auto simp add: `0 < k` `1 \ k`) + have "real ?k \ 1" + by (rule rapprox_rat_le1, auto simp add: `0 < k` `1 \ k`) 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') @@ -526,8 +512,7 @@ let ?k = "lapprox_rat prec 1 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 "\n. 0 \ real ?k" using lapprox_rat_bottom[where x=1 and y=k, OF zero_le_one `0 < k`] by (auto simp add: `1 div k = 0`) + have "\n. 0 \ real ?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 "\n. real ?k \ 1" using lapprox_rat by (rule order_trans, auto simp add: `1 / k \ 1`) have "?k \ 1 / k" using lapprox_rat[where x=1 and y=k] by auto @@ -539,14 +524,14 @@ } note lb_arctan = this have "pi \ ub_pi n" - unfolding ub_pi_def machin_pi Let_def real_of_float_mult real_of_float_sub unfolding Float_num - using lb_arctan[of 239] ub_arctan[of 5] - by (auto intro!: mult_left_mono add_mono simp add: diff_minus simp del: lapprox_rat.simps rapprox_rat.simps) + unfolding ub_pi_def machin_pi Let_def unfolding Float_num + using lb_arctan[of 239] ub_arctan[of 5] powr_realpow[of 2 2] + by (auto intro!: mult_left_mono add_mono simp add: diff_minus) moreover have "lb_pi n \ pi" - unfolding lb_pi_def machin_pi Let_def real_of_float_mult real_of_float_sub Float_num - using lb_arctan[of 5] ub_arctan[of 239] - by (auto intro!: mult_left_mono add_mono simp add: diff_minus simp del: lapprox_rat.simps rapprox_rat.simps) + unfolding lb_pi_def machin_pi Let_def Float_num + using lb_arctan[of 5] ub_arctan[of 239] powr_realpow[of 2 2] + by (auto intro!: mult_left_mono add_mono simp add: diff_minus) ultimately show ?thesis by auto qed @@ -571,7 +556,7 @@ else Float 1 1 * ub_horner y else 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) = sum_case id id v in (if x < 0 then 1 else 0))", auto simp add: less_float_def) +termination by (relation "measure (\ v. let (prec, x) = sum_case 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] @@ -579,17 +564,17 @@ lemma lb_arctan_bound': assumes "0 \ real x" shows "lb_arctan prec x \ arctan x" proof - - have "\ x < 0" and "0 \ x" unfolding less_float_def le_float_def using `0 \ real x` by auto + have "\ x < 0" and "0 \ x" using `0 \ real x` by auto let "?ub_horner x" = "x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (x * x)" and "?lb_horner x" = "x * lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (x * x)" show ?thesis proof (cases "x \ Float 1 -1") - case True hence "real x \ 1" unfolding le_float_def Float_num by auto + case True hence "real x \ 1" by auto show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF `\ x < 0`] if_P[OF True] using arctan_0_1_bounds[OF `0 \ real x` `real x \ 1`] by auto next - case False hence "0 < real x" unfolding le_float_def Float_num by auto + case False hence "0 < real x" by auto let ?R = "1 + sqrt (1 + real x * real x)" let ?fR = "1 + ub_sqrt prec (1 + x * x)" let ?DIV = "float_divl prec x ?fR" @@ -601,7 +586,7 @@ using bnds_sqrt'[of "1 + x * x"] by auto hence "?R \ ?fR" by auto - hence "0 < ?fR" and "0 < real ?fR" unfolding less_float_def using `0 < ?R` by auto + hence "0 < ?fR" and "0 < real ?fR" using `0 < ?R` by auto have monotone: "(float_divl prec x ?fR) \ x / ?R" proof - @@ -621,9 +606,9 @@ moreover have "?DIV \ real x / ?fR" by (rule float_divl) ultimately have "real ?DIV \ 1" unfolding divide_le_eq_1_pos[OF `0 < real ?fR`, symmetric] by auto - have "0 \ real ?DIV" using float_divl_lower_bound[OF `0 \ x` `0 < ?fR`] unfolding le_float_def by auto - - have "(Float 1 1 * ?lb_horner ?DIV) \ 2 * arctan (float_divl prec x ?fR)" unfolding real_of_float_mult[of "Float 1 1"] Float_num + have "0 \ real ?DIV" using float_divl_lower_bound[OF `0 \ x` `0 < ?fR`] unfolding less_eq_float_def by auto + + have "(Float 1 1 * ?lb_horner ?DIV) \ 2 * arctan (float_divl prec x ?fR)" using arctan_0_1_bounds[OF `0 \ real ?DIV` `real ?DIV \ 1`] by auto also have "\ \ 2 * arctan (x / ?R)" using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono) @@ -631,7 +616,7 @@ 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] . next case False - hence "2 < real x" unfolding le_float_def Float_num by auto + hence "2 < real x" by auto hence "1 \ real x" by auto let "?invx" = "float_divr prec 1 x" @@ -644,18 +629,19 @@ using `0 \ arctan x` by auto next case False - hence "real ?invx \ 1" unfolding less_float_def by auto + hence "real ?invx \ 1" by auto have "0 \ real ?invx" by (rule order_trans[OF _ float_divr], auto simp add: `0 \ real x`) have "1 / x \ 0" and "0 < 1 / x" using `0 < real x` by auto - have "arctan (1 / x) \ arctan ?invx" unfolding real_of_float_1[symmetric] by (rule arctan_monotone', rule float_divr) + 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[OF `0 \ real ?invx` `real ?invx \ 1`] by auto finally have "pi / 2 - (?ub_horner ?invx) \ arctan x" using `0 \ arctan x` arctan_inverse[OF `1 / x \ 0`] unfolding real_sgn_pos[OF `0 < 1 / real x`] le_diff_eq by auto moreover - have "lb_pi prec * Float 1 -1 \ pi / 2" unfolding real_of_float_mult Float_num times_divide_eq_right mult_1_left using pi_boundaries by auto + 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 @@ -667,18 +653,18 @@ lemma ub_arctan_bound': assumes "0 \ real x" shows "arctan x \ ub_arctan prec x" proof - - have "\ x < 0" and "0 \ x" unfolding less_float_def le_float_def using `0 \ real x` by auto + have "\ x < 0" and "0 \ x" using `0 \ real x` by auto let "?ub_horner x" = "x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (x * x)" and "?lb_horner x" = "x * lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (x * x)" show ?thesis proof (cases "x \ Float 1 -1") - case True hence "real x \ 1" unfolding le_float_def Float_num by auto + case True hence "real 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[OF `0 \ real x` `real x \ 1`] by auto next - case False hence "0 < real x" unfolding le_float_def Float_num by auto + case False hence "0 < real x" by auto let ?R = "1 + sqrt (1 + real x * real x)" let ?fR = "1 + lb_sqrt prec (1 + x * x)" let ?DIV = "float_divr prec x ?fR" @@ -691,7 +677,7 @@ have "lb_sqrt prec (1 + x * x) \ sqrt (1 + x * x)" using bnds_sqrt'[of "1 + x * x"] by auto hence "?fR \ ?R" by auto - have "0 < real ?fR" unfolding real_of_float_add real_of_float_1 by (rule order_less_le_trans[OF zero_less_one], auto simp add: lb_sqrt_lower_bound[OF `0 \ real (1 + x*x)`]) + have "0 < real ?fR" by (rule order_less_le_trans[OF zero_less_one], auto simp add: lb_sqrt_lower_bound[OF `0 \ real (1 + x*x)`]) have monotone: "x / ?R \ (float_divr prec x ?fR)" proof - @@ -707,12 +693,12 @@ show ?thesis proof (cases "?DIV > 1") case True - have "pi / 2 \ ub_pi prec * Float 1 -1" unfolding real_of_float_mult Float_num times_divide_eq_right mult_1_left using pi_boundaries by auto + 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 ?DIV \ 1" unfolding less_float_def by auto + hence "real ?DIV \ 1" by auto have "0 \ x / ?R" using `0 \ real x` `0 < ?R` unfolding zero_le_divide_iff by auto hence "0 \ real ?DIV" using monotone by (rule order_trans) @@ -720,32 +706,32 @@ 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 real_of_float_mult[of "Float 1 1"] Float_num + also have "\ \ (Float 1 1 * ?ub_horner ?DIV)" unfolding Float_num using arctan_0_1_bounds[OF `0 \ real ?DIV` `real ?DIV \ 1`] by auto 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 x" unfolding le_float_def Float_num by auto + hence "2 < real x" by auto hence "1 \ real x" by auto hence "0 < real x" by auto - hence "0 < x" unfolding less_float_def by auto + hence "0 < x" by auto let "?invx" = "float_divl prec 1 x" have "0 \ arctan x" using arctan_monotone'[OF `0 \ real x`] using arctan_tan[of 0, unfolded tan_zero] by auto have "real ?invx \ 1" unfolding less_float_def by (rule order_trans[OF float_divl], auto simp add: `1 \ real x` divide_le_eq_1_pos[OF `0 < real x`]) - have "0 \ real ?invx" unfolding real_of_float_0[symmetric] by (rule float_divl_lower_bound[unfolded le_float_def], auto simp add: `0 < x`) + have "0 \ real ?invx" using `0 < x` by (intro float_divl_lower_bound) auto have "1 / x \ 0" and "0 < 1 / x" using `0 < real x` by auto have "(?lb_horner ?invx) \ arctan (?invx)" using arctan_0_1_bounds[OF `0 \ real ?invx` `real ?invx \ 1`] by auto - also have "\ \ arctan (1 / x)" unfolding real_of_float_1[symmetric] by (rule arctan_monotone', rule float_divl) + 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 real_sgn_pos[OF `0 < 1 / x`] le_diff_eq by auto moreover - have "pi / 2 \ ub_pi prec * Float 1 -1" unfolding real_of_float_mult Float_num times_divide_eq_right mult_1_right using pi_boundaries by auto + 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 @@ -756,15 +742,16 @@ lemma arctan_boundaries: "arctan x \ {(lb_arctan prec x) .. (ub_arctan prec x)}" proof (cases "0 \ x") - case True hence "0 \ real x" unfolding le_float_def by auto + case True hence "0 \ real x" by auto show ?thesis using ub_arctan_bound'[OF `0 \ real x`] lb_arctan_bound'[OF `0 \ real x`] unfolding atLeastAtMost_iff by auto next let ?mx = "-x" - case False hence "x < 0" and "0 \ real ?mx" unfolding le_float_def less_float_def by auto + case False hence "x < 0" and "0 \ real ?mx" by auto hence bounds: "lb_arctan prec ?mx \ arctan ?mx \ arctan ?mx \ ub_arctan prec ?mx" using ub_arctan_bound'[OF `0 \ real ?mx`] lb_arctan_bound'[OF `0 \ real ?mx`] by auto - show ?thesis unfolding real_of_float_minus 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 real_of_float_minus arctan_minus] 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" @@ -796,11 +783,12 @@ | "lb_sin_cos_aux prec 0 i k x = 0" | "lb_sin_cos_aux prec (Suc n) i k x = (lapprox_rat prec 1 k) - 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 (x * x)" unfolding real_of_float_mult by auto + have "0 \ real (x * x)" by auto let "?f n" = "fact (2 * n)" { fix n @@ -817,8 +805,8 @@ 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 x = 0") case False hence "real x \ 0" by auto - hence "0 < x" and "0 < real x" using `0 \ real x` unfolding less_float_def by auto - have "0 < x * x" using `0 < x` unfolding less_float_def real_of_float_mult real_of_float_0 + hence "0 < x" and "0 < real x" using `0 \ real x` by auto + have "0 < x * x" using `0 < x` using mult_pos_pos[where a="real x" and b="real x"] by auto { fix x n have "(\ i=0.. x" by (rule order_trans[OF _ `0 < real x`[THEN less_imp_le]], auto) with `x \ pi / 2` - show ?thesis unfolding `get_even n = 0` lb_sin_cos_aux.simps real_of_float_minus real_of_float_0 using cos_ge_zero by auto + 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 @@ -901,7 +889,9 @@ show ?thesis proof (cases "n = 0") case True - thus ?thesis unfolding `n = 0` get_even_def get_odd_def using `real x = 0` lapprox_rat[where x="-1" and y=1] by auto + thus ?thesis unfolding `n = 0` get_even_def get_odd_def + using `real x = 0` lapprox_rat[where x="-1" and y=1] + by (auto simp: compute_lapprox_rat compute_rapprox_rat) 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 x = 0` rapprox_rat[where x=1 and y=1] lapprox_rat[where x=1 and y=1] by (cases "even (Suc m)", auto) @@ -912,7 +902,7 @@ 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 (x * x)" unfolding real_of_float_mult by auto + have "0 \ real (x * x)" by auto let "?f n" = "fact (2 * n + 1)" { fix n @@ -922,7 +912,7 @@ from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0, OF `0 \ real (x * x)` f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps] - show "?lb" and "?ub" using `0 \ real x` unfolding real_of_float_mult + show "?lb" and "?ub" using `0 \ real x` unfolding power_add power_one_right mult_assoc[symmetric] setsum_left_distrib[symmetric] unfolding mult_commute[where 'a=real] by (auto intro!: mult_left_mono simp add: power_mult power2_eq_square[of "real x"]) @@ -932,8 +922,8 @@ 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 x = 0") case False hence "real x \ 0" by auto - hence "0 < x" and "0 < real x" using `0 \ real x` unfolding less_float_def by auto - have "0 < x * x" using `0 < x` unfolding less_float_def real_of_float_mult real_of_float_0 + hence "0 < x" and "0 < real x" using `0 \ real x` by auto + have "0 < x * x" using `0 < x` using mult_pos_pos[where a="real x" and b="real x"] by auto { fix x n have "(\ j = 0 ..< n. -1 ^ (((2 * j + 1) - Suc 0) div 2) / (real (fact (2 * j + 1))) * x ^(2 * j + 1)) @@ -1006,7 +996,7 @@ case False hence "get_even n = 0" by auto with `x \ pi / 2` `0 \ real x` - show ?thesis unfolding `get_even n = 0` ub_sin_cos_aux.simps real_of_float_minus real_of_float_0 using sin_ge_zero by auto + 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 @@ -1050,7 +1040,7 @@ finally have "cos x = 2 * cos (x / 2) * cos (x / 2) - 1" . } note x_half = this[symmetric] - have "\ x < 0" using `0 \ real x` unfolding less_float_def by auto + have "\ x < 0" using `0 \ real 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 1 1 * x * x - 1" @@ -1058,14 +1048,14 @@ show ?thesis proof (cases "x < Float 1 -1") - case True hence "x \ pi / 2" unfolding less_float_def using pi_ge_two by auto + 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 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 real_of_float_mult Float_num by auto + 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" @@ -1073,18 +1063,18 @@ case True show ?thesis using cos_ge_minus_one unfolding if_P[OF True] by auto next case False - hence "0 \ real y" unfolding less_float_def by auto + hence "0 \ real y" by auto from mult_mono[OF `y \ cos ?x2` `y \ cos ?x2` `0 \ cos ?x2` this] have "real y * real y \ cos ?x2 * cos ?x2" . hence "2 * real y * real y \ 2 * cos ?x2 * cos ?x2" by auto - hence "2 * real y * real y - 1 \ 2 * cos (x / 2) * cos (x / 2) - 1" unfolding Float_num real_of_float_mult by auto - thus ?thesis unfolding if_not_P[OF False] x_half Float_num real_of_float_mult real_of_float_sub by auto + hence "2 * real y * real 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 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 real_of_float_mult Float_num by auto + 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)" @@ -1093,8 +1083,8 @@ from mult_mono[OF ub ub this `0 \ cos ?x2`] have "cos ?x2 * cos ?x2 \ real y * real y" . hence "2 * cos ?x2 * cos ?x2 \ 2 * real y * real y" by auto - hence "2 * cos (x / 2) * cos (x / 2) - 1 \ 2 * real y * real y - 1" unfolding Float_num real_of_float_mult by auto - thus ?thesis unfolding x_half real_of_float_mult Float_num real_of_float_sub by auto + hence "2 * cos (x / 2) * cos (x / 2) - 1 \ 2 * real y * real y - 1" unfolding Float_num by auto + thus ?thesis unfolding x_half Float_num by auto qed } note ub_half = this @@ -1105,8 +1095,8 @@ show ?thesis proof (cases "x < 1") - case True hence "real x \ 1" unfolding less_float_def by auto - have "0 \ real ?x2" and "?x2 \ pi / 2" using pi_ge_two `0 \ real x` unfolding real_of_float_mult Float_num using assms by auto + case True hence "real x \ 1" by auto + have "0 \ real ?x2" and "?x2 \ pi / 2" using pi_ge_two `0 \ real 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 @@ -1123,21 +1113,21 @@ ultimately show ?thesis by auto next case False - have "0 \ real ?x4" and "?x4 \ pi / 2" using pi_ge_two `0 \ real x` `x \ pi` unfolding real_of_float_mult Float_num by auto + have "0 \ real ?x4" and "?x4 \ pi / 2" using pi_ge_two `0 \ real 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 (cases x, auto simp add: times_float.simps) + have eq_4: "?x2 * Float 1 -1 = x * Float 1 -2" by transfer simp have "(?lb x) \ ?cos x" proof - - have "-pi \ ?x2" and "?x2 \ pi" unfolding real_of_float_mult Float_num using pi_ge_two `0 \ real x` `x \ pi` by auto + have "-pi \ ?x2" and "?x2 \ pi" using pi_ge_two `0 \ real 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" unfolding real_of_float_mult Float_num using pi_ge_two `0 \ real x` ` x \ pi` by auto + have "-pi \ ?x2" and "?x2 \ pi" using pi_ge_two `0 \ real 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 @@ -1155,8 +1145,8 @@ definition bnds_cos :: "nat \ float \ float \ float * float" where "bnds_cos prec lx ux = (let - lpi = round_down prec (lb_pi prec) ; - upi = round_up prec (ub_pi prec) ; + 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 = lx - k * 2 * (if k < 0 then lpi else upi) ; ux = ux - k * 2 * (if k < 0 then upi else lpi) @@ -1169,14 +1159,7 @@ lemma floor_int: obtains k :: int where "real k = (floor_fl f)" -proof - - assume *: "\ k :: int. real k = (floor_fl f) \ thesis" - obtain m e where fl: "Float m e = floor_fl f" by (cases "floor_fl f", auto) - from floor_pos_exp[OF this] - have "real (m* 2^(nat e)) = (floor_fl f)" - by (auto simp add: fl[symmetric] real_of_float_def pow2_def) - from *[OF this] show thesis by blast -qed + 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) @@ -1203,8 +1186,8 @@ fix x :: real fix lx ux assume bnds: "(l, u) = bnds_cos prec lx ux" and x: "x \ {lx .. ux}" - let ?lpi = "round_down prec (lb_pi prec)" - let ?upi = "round_up prec (ub_pi prec)" + 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 ?lx = "lx - ?k * 2 * (if ?k < 0 then ?lpi else ?upi)" let ?ux = "ux - ?k * 2 * (if ?k < 0 then ?upi else ?lpi)" @@ -1212,24 +1195,27 @@ obtain k :: int where k: "k = real ?k" using floor_int . have upi: "pi \ ?upi" and lpi: "?lpi \ pi" - using round_up[of "ub_pi prec" prec] pi_boundaries[of prec] - round_down[of prec "lb_pi prec"] by auto + 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 \ x - k * (2 * pi) \ x - k * (2 * pi) \ ?ux" - using x by (cases "k = 0") (auto intro!: add_mono - simp add: diff_minus k[symmetric] less_float_def) + using x unfolding k[symmetric] + by (cases "k = 0") + (auto intro!: add_mono + simp add: diff_minus k[symmetric] + simp del: float_of_numeral) note lx = this[THEN conjunct1] and ux = this[THEN conjunct2] hence lx_less_ux: "?lx \ real ?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 ?lx \ 0" - by (simp_all add: le_float_def) + by simp_all have "(lb_cos prec (- ?lx)) \ cos (real (- ?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: real_of_float_minus real_of_int_minus + by (simp only: uminus_float.rep_eq real_of_int_minus cos_minus diff_minus mult_minus_left) finally have "(lb_cos prec (- ?lx)) \ cos x" unfolding cos_periodic_int . } @@ -1238,11 +1224,11 @@ { assume "0 \ ?lx" and pi_x: "x - k * (2 * pi) \ pi" with lx have pi_lx: "?lx \ pi" and lx_0: "0 \ real ?lx" - by (auto simp add: le_float_def) + by auto have "cos (x + (-k) * (2 * pi)) \ cos ?lx" using cos_monotone_0_pi'[OF lx_0 lx pi_x] - by (simp only: real_of_float_minus real_of_int_minus + by (simp only: real_of_int_minus cos_minus diff_minus mult_minus_left) also have "\ \ (ub_cos prec ?lx)" using lb_cos[OF lx_0 pi_lx] by simp @@ -1253,11 +1239,11 @@ { assume pi_x: "- pi \ x - k * (2 * pi)" and "?ux \ 0" with ux have pi_ux: "- pi \ ?ux" and ux_0: "real ?ux \ 0" - by (simp_all add: le_float_def) + by simp_all have "cos (x + (-k) * (2 * pi)) \ cos (real (- ?ux))" using cos_monotone_minus_pi_0'[OF pi_x ux ux_0] - by (simp only: real_of_float_minus real_of_int_minus + by (simp only: uminus_float.rep_eq real_of_int_minus cos_minus diff_minus mult_minus_left) also have "\ \ (ub_cos prec (- ?ux))" using lb_cos_minus[OF pi_ux ux_0, of prec] by simp @@ -1268,13 +1254,13 @@ { 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 ?ux" - by (simp_all add: le_float_def) + 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'[OF x_ge_0 ux pi_ux] - by (simp only: real_of_float_minus real_of_int_minus + by (simp only: real_of_int_minus cos_minus diff_minus mult_minus_left) finally have "(lb_cos prec ?ux) \ cos x" unfolding cos_periodic_int . } @@ -1290,7 +1276,7 @@ from True lpi[THEN le_imp_neg_le] lx ux have "- pi \ x - k * (2 * pi)" and "x - k * (2 * pi) \ 0" - by (auto simp add: le_float_def) + by auto with True negative_ux negative_lx show ?thesis unfolding l u by simp next case False note 1 = this show ?thesis @@ -1303,7 +1289,7 @@ from True lpi lx ux have "0 \ x - k * (2 * pi)" and "x - k * (2 * pi) \ pi" - by (auto simp add: le_float_def) + by auto with True positive_ux positive_lx show ?thesis unfolding l u by simp next case False note 2 = this show ?thesis @@ -1314,7 +1300,8 @@ 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", simp_all add: real_of_float_min) + by (cases "x - k * (2 * pi) < 0") (auto simp add: real_of_float_min) + next case False note 3 = this show ?thesis proof (cases "0 \ ?lx \ ?ux \ 2 * ?lpi") case True note Cond = this with bnds 1 2 3 @@ -1331,28 +1318,27 @@ 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 simp add: le_float_def) + 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 (?ux - 2 * ?lpi) \ 0" - using Cond by (auto simp add: le_float_def) + using Cond by auto from 2 and Cond have "\ ?ux \ ?lpi" by auto - hence "- ?lpi \ ?ux - 2 * ?lpi" by (auto simp add: le_float_def) + hence "- ?lpi \ ?ux - 2 * ?lpi" by auto hence pi_ux: "- pi \ (?ux - 2 * ?lpi)" - using lpi[THEN le_imp_neg_le] by (auto simp add: le_float_def) + 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: real_of_float_minus real_of_int_minus real_of_one - minus_one [symmetric] diff_minus mult_minus_left mult_1_left) + by (simp only: minus_float.rep_eq real_of_int_minus real_of_one minus_one[symmetric] + diff_minus mult_minus_left mult_1_left) also have "\ = cos ((- (?ux - 2 * ?lpi)))" - unfolding real_of_float_minus cos_minus .. + 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) @@ -1374,17 +1360,17 @@ 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 simp add: le_float_def) + 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 (?lx + 2 * ?lpi)" - using Cond lpi by (auto simp add: le_float_def) + using Cond lpi by auto from 1 and Cond have "\ -?lpi \ ?lx" by auto - hence "?lx + 2 * ?lpi \ ?lpi" by (auto simp add: le_float_def) + hence "?lx + 2 * ?lpi \ ?lpi" by auto hence pi_lx: "(?lx + 2 * ?lpi) \ pi" - using lpi[THEN le_imp_neg_le] by (auto simp add: le_float_def) + 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 @@ -1393,8 +1379,8 @@ unfolding cos_periodic_int .. also have "\ \ cos ((?lx + 2 * ?lpi))" using cos_monotone_0_pi'[OF lx_0 lx_le_x pi_x] - by (simp only: real_of_float_minus real_of_int_minus real_of_one - minus_one [symmetric] diff_minus mult_minus_left mult_1_left) + by (simp only: minus_float.rep_eq real_of_int_minus real_of_one + minus_one[symmetric] diff_minus mult_minus_left mult_1_left) 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) @@ -1467,14 +1453,13 @@ "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 (case floor_fl x of (Float m e) \ (horner (float_divl prec x (- Float m e))) ^ (nat (-m) * 2 ^ nat e)) + in if x < - 1 then (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 (case floor_fl x of (Float m e) \ - (ub_exp_horner prec (get_odd (prec + 2)) 1 1 (float_divr prec x (- Float m e))) ^ (nat (-m) * 2 ^ nat e)) + else if x < - 1 then 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) = sum_case id id v in (if 0 < x then 1 else 0))", auto simp add: less_float_def) +termination by (relation "measure (\ v. let (prec, x) = sum_case id id v in (if 0 < x then 1 else 0))", auto) lemma exp_m1_ge_quarter: "(1 / 4 :: real) \ exp (- 1)" proof - @@ -1483,24 +1468,24 @@ have "1 / 4 = (Float 1 -2)" unfolding Float_num by auto also have "\ \ lb_exp_horner 1 (get_even 4) 1 1 (- 1)" unfolding get_even_def eq4 - by (auto simp add: lapprox_posrat_def rapprox_posrat_def normfloat.simps) + by (auto simp add: compute_lapprox_rat compute_rapprox_rat compute_lapprox_posrat compute_rapprox_posrat rat_precision_def compute_bitlen) also have "\ \ exp (- 1 :: float)" using bnds_exp_horner[where x="- 1"] by auto - finally show ?thesis unfolding real_of_float_minus real_of_float_1 . + 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: "\ x. 0 < ?horner x" unfolding Let_def by (cases "?lb_horner x \ 0", auto simp add: le_float_def less_float_def) + have pos_horner: "\ x. 0 < ?horner x" unfolding Let_def by (cases "?lb_horner x \ 0", auto) moreover { fix x :: float fix num :: nat - have "0 < real (?horner x) ^ num" using `0 < ?horner x`[unfolded less_float_def real_of_float_0] by (rule zero_less_power) - also have "\ = (?horner x) ^ num" using float_power by auto + have "0 < real (?horner x) ^ num" using `0 < ?horner x` by simp + also have "\ = (?horner x) ^ num" by auto finally have "0 < real ((?horner x) ^ num)" . } 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 add: float_power le_float_def less_float_def) + by (cases "floor_fl x", cases "x < - 1", auto) qed lemma exp_boundaries': assumes "x \ 0" @@ -1509,13 +1494,13 @@ 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 x \ 0" and "\ x > 0" using `x \ 0` unfolding le_float_def less_float_def by auto + have "real x \ 0" and "\ x > 0" using `x \ 0` by auto show ?thesis proof (cases "x < - 1") - case False hence "- 1 \ real x" unfolding less_float_def by auto + case False hence "- 1 \ real x" by auto show ?thesis proof (cases "?lb_exp_horner x \ 0") - from `\ x < - 1` have "- 1 \ real x" unfolding less_float_def by auto + from `\ x < - 1` have "- 1 \ real 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 . @@ -1527,74 +1512,74 @@ next case True - obtain m e where Float_floor: "floor_fl x = Float m e" by (cases "floor_fl x", auto) - let ?num = "nat (- m) * 2 ^ nat e" - - have "real (floor_fl x) < - 1" using floor_fl `x < - 1` unfolding le_float_def less_float_def real_of_float_minus real_of_float_1 by (rule order_le_less_trans) - hence "real (floor_fl x) < 0" unfolding Float_floor real_of_float_simp using zero_less_pow2[of xe] by auto - hence "m < 0" - unfolding less_float_def real_of_float_0 Float_floor real_of_float_simp - unfolding pos_prod_lt[OF zero_less_pow2[of e], unfolded mult_commute] by auto - hence "1 \ - m" by auto - hence "0 < nat (- m)" by auto - moreover - have "0 \ e" using floor_pos_exp Float_floor[symmetric] by auto - hence "(0::nat) < 2 ^ nat e" by auto - ultimately have "0 < ?num" by auto + let ?num = "nat (- int_floor_fl x)" + + have "real (int_floor_fl x) < - 1" using int_floor_fl[of x] `x < - 1` + by simp + hence "real (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 e_nat: "(nat e) = e" using `0 \ e` by auto - have num_eq: "real ?num = - floor_fl x" using `0 < nat (- m)` - unfolding Float_floor real_of_float_minus real_of_float_simp real_of_nat_mult pow2_int[of "nat e", unfolded e_nat] real_of_nat_power by auto - have "0 < - floor_fl x" using `0 < ?num`[unfolded real_of_nat_less_iff[symmetric]] unfolding less_float_def num_eq[symmetric] real_of_float_0 real_of_nat_zero . - hence "real (floor_fl x) < 0" unfolding less_float_def 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 real_of_nat_less_iff[symmetric]] by simp + hence "real (int_floor_fl x) < 0" unfolding less_float_def by auto + have fl_eq: "real (- int_floor_fl x) = real (- floor_fl x)" + by (simp add: floor_fl_def int_floor_fl_def) + from `0 < - int_floor_fl x` have "0 < real (- floor_fl x)" + by (simp add: floor_fl_def int_floor_fl_def) + from `real (int_floor_fl x) < 0` have "real (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 (float_divr prec x (- floor_fl x)) \ 0" - using float_divr_nonpos_pos_upper_bound[OF `x \ 0` `0 < - floor_fl x`] unfolding le_float_def real_of_float_0 . + using float_divr_nonpos_pos_upper_bound[OF `real x \ 0` `0 < real (- 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_real_of_nat_mult .. - also have "\ \ exp (float_divr prec x (- floor_fl x)) ^ ?num" unfolding num_eq + 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 float_power + 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) - finally show ?thesis unfolding ub_exp.simps if_not_P[OF `\ 0 < x`] if_P[OF `x < - 1`] float.cases Float_floor Let_def . + 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 (- Float m e)" + 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 ?horner" unfolding le_float_def by auto + case False hence "0 \ real ?horner" by auto have div_less_zero: "real (float_divl prec x (- floor_fl x)) \ 0" using `real (floor_fl x) < 0` `real 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" unfolding float_power - using `0 \ real ?horner`[unfolded Float_floor[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 - using float_divl by (auto intro!: power_mono simp del: real_of_float_minus) + exp (float_divl prec x (- floor_fl x)) ^ ?num" + using `0 \ real ?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_real_of_nat_mult .. also have "\ = exp x" using `real ?num \ 0` by auto finally show ?thesis - unfolding lb_exp.simps if_not_P[OF `\ 0 < x`] if_P[OF `x < - 1`] float.cases Float_floor Let_def if_not_P[OF False] by auto + 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 next case True have "real (floor_fl x) \ 0" and "real (floor_fl x) \ 0" using `real (floor_fl x) < 0` by auto from divide_right_mono_neg[OF floor_fl[of x] `real (floor_fl x) \ 0`, unfolded divide_self[OF `real (floor_fl x) \ 0`]] - have "- 1 \ x / (- floor_fl x)" unfolding real_of_float_minus by auto + 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 (Float 1 -2) ^ ?num \ exp (x / (- floor_fl x)) ^ ?num" by (auto intro!: power_mono) - also have "\ = exp x" unfolding num_eq exp_real_of_nat_mult[symmetric] using `real (floor_fl x) \ 0` by auto + also have "\ = exp x" unfolding num_eq fl_eq exp_real_of_nat_mult[symmetric] using `real (floor_fl x) \ 0` by auto finally show ?thesis - unfolding lb_exp.simps if_not_P[OF `\ 0 < x`] if_P[OF `x < - 1`] float.cases Float_floor Let_def if_P[OF True] float_power . + 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 @@ -1605,15 +1590,15 @@ proof - show ?thesis proof (cases "0 < x") - case False hence "x \ 0" unfolding less_float_def le_float_def by auto + case False hence "x \ 0" by auto from exp_boundaries'[OF this] show ?thesis . next - case True hence "-x \ 0" unfolding less_float_def le_float_def by auto + 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 x) \ ub_exp prec (-x)" unfolding atLeastAtMost_iff real_of_float_minus by auto + have ub_exp: "exp (- real 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" @@ -1624,15 +1609,14 @@ moreover have "exp x \ ub_exp prec x" proof - - have "\ 0 < -x" using `0 < x` unfolding less_float_def by auto + have "\ 0 < -x" using `0 < x` by auto from exp_boundaries'[OF `-x \ 0`] - have lb_exp: "lb_exp prec (-x) \ exp (- real x)" unfolding atLeastAtMost_iff real_of_float_minus by auto + have lb_exp: "lb_exp prec (-x) \ exp (- real x)" unfolding atLeastAtMost_iff minus_float.rep_eq by auto have "exp x \ (1 :: float) / lb_exp prec (-x)" - using lb_exp[unfolded inverse_le_iff_le[OF exp_gt_zero lb_exp_pos[OF `\ 0 < -x`, unfolded less_float_def real_of_float_0], - symmetric]] - unfolding exp_minus nonzero_inverse_inverse_eq[OF exp_not_eq_zero] inverse_eq_divide real_of_float_1 by auto + using lb_exp lb_exp_pos[OF `\ 0 < -x`, of prec] + by (simp del: lb_exp.simps add: exp_minus inverse_eq_divide 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 @@ -1703,7 +1687,7 @@ let "?s n" = "-1^n * (1 / real (1 + n)) * (real x)^(Suc n)" - have "?lb \ setsum ?s {0 ..< 2 * ev}" unfolding power_Suc2 mult_assoc[symmetric] real_of_float_mult setsum_left_distrib[symmetric] unfolding mult_commute[of "real x"] ev + have "?lb \ setsum ?s {0 ..< 2 * ev}" unfolding power_Suc2 mult_assoc[symmetric] times_float.rep_eq setsum_left_distrib[symmetric] unfolding mult_commute[of "real 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 x` refl lb_ln_horner.simps ub_ln_horner.simps] `0 \ real x` by (rule mult_right_mono) @@ -1711,7 +1695,7 @@ finally show "?lb \ ?ln" . have "?ln \ setsum ?s {0 ..< 2 * od + 1}" using ln_bounds(2)[OF `0 \ real x` `real x < 1`] by auto - also have "\ \ ?ub" unfolding power_Suc2 mult_assoc[symmetric] real_of_float_mult setsum_left_distrib[symmetric] unfolding mult_commute[of "real x"] od + also have "\ \ ?ub" unfolding power_Suc2 mult_assoc[symmetric] times_float.rep_eq setsum_left_distrib[symmetric] unfolding mult_commute[of "real 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 x` refl lb_ln_horner.simps ub_ln_horner.simps] `0 \ real x` by (rule mult_right_mono) @@ -1747,28 +1731,27 @@ 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 ?lthird < 1" by auto - have lb3_lb: "0 \ real ?lthird" using lapprox_rat_bottom[of 1 3] by auto + have lb3_lb: "0 \ real ?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 ?uthird" by auto have lb2: "0 \ real (Float 1 -1)" and ub2: "real (Float 1 -1) < 1" unfolding Float_num by auto have "0 \ (1::int)" and "0 < (3::int)" by auto - have ub3_ub: "real ?uthird < 1" unfolding rapprox_rat.simps(2)[OF `0 \ 1` `0 < 3`] - by (rule rapprox_posrat_less1, auto) + have ub3_ub: "real ?uthird < 1" by (simp add: compute_rapprox_rat rapprox_posrat_less1) have third_gt0: "(0 :: real) < 1 / 3 + 1" by auto have uthird_gt0: "0 < real ?uthird + 1" using ub3_lb by auto have lthird_gt0: "0 < real ?lthird + 1" using lb3_lb by auto - show ?ub_ln2 unfolding ub_ln2_def Let_def real_of_float_add ln2_sum Float_num(4)[symmetric] + show ?ub_ln2 unfolding ub_ln2_def Let_def plus_float.rep_eq ln2_sum Float_num(4)[symmetric] proof (rule add_mono, fact ln_float_bounds(2)[OF lb2 ub2]) have "ln (1 / 3 + 1) \ ln (real ?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] . finally show "ln (1 / 3 + 1) \ ?uthird * ub_ln_horner prec (get_odd prec) 1 ?uthird" . qed - show ?lb_ln2 unfolding lb_ln2_def Let_def real_of_float_add ln2_sum Float_num(4)[symmetric] + show ?lb_ln2 unfolding lb_ln2_def Let_def plus_float.rep_eq ln2_sum Float_num(4)[symmetric] proof (rule add_mono, fact ln_float_bounds(1)[OF lb2 ub2]) have "?lthird * lb_ln_horner prec (get_even prec) 1 ?lthird \ ln (real ?lthird + 1)" using ln_float_bounds(1)[OF lb3_lb lb3_ub] . @@ -1786,7 +1769,7 @@ if x \ Float 3 -1 then Some (horner (x - 1)) else if x < Float 1 1 then Some (horner (Float 1 -1) + horner (x * rapprox_rat prec 2 3 - 1)) else let l = bitlen (mantissa x) - 1 in - Some (ub_ln2 prec * (Float (scale x + l) 0) + horner (Float (mantissa x) (- l) - 1)))" | + Some (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. x * lb_ln_horner prec (get_even prec) 1 x in @@ -1794,41 +1777,112 @@ else if x < Float 1 1 then Some (horner (Float 1 -1) + horner (max (x * lapprox_rat prec 2 3 - 1) 0)) else let l = bitlen (mantissa x) - 1 in - Some (lb_ln2 prec * (Float (scale x + l) 0) + horner (Float (mantissa x) (- l) - 1)))" + Some (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) = sum_case id id v in (if x < 1 then 1 else 0))", auto) - fix prec x assume "\ x \ 0" and "x < 1" and "float_divl (max prec (Suc 0)) 1 x < 1" - hence "0 < x" and "0 < max prec (Suc 0)" unfolding less_float_def le_float_def by auto - from float_divl_pos_less1_bound[OF `0 < x` `x < 1` `0 < max prec (Suc 0)`] - show False using `float_divl (max prec (Suc 0)) 1 x < 1` unfolding less_float_def le_float_def by auto + fix prec and x :: float assume "\ real x \ 0" and "real x < 1" and "real (float_divl (max prec (Suc 0)) 1 x) < 1" + hence "0 < real x" "1 \ max prec (Suc 0)" "real x < 1" by auto + from float_divl_pos_less1_bound[OF `0 < real x` `real x < 1` `1 \ max prec (Suc 0)`] + show False using `real (float_divl (max prec (Suc 0)) 1 x) < 1` by auto next - fix prec x assume "\ x \ 0" and "x < 1" and "float_divr prec 1 x < 1" - hence "0 < x" unfolding less_float_def le_float_def by auto - from float_divr_pos_less1_lower_bound[OF `0 < x` `x < 1`, of prec] - show False using `float_divr prec 1 x < 1` unfolding less_float_def le_float_def by auto + fix prec x assume "\ real x \ 0" and "real x < 1" and "real (float_divr prec 1 x) < 1" + hence "0 < x" by auto + from float_divr_pos_less1_lower_bound[OF `0 < x`, of prec] `real x < 1` + show False using `real (float_divr prec 1 x) < 1` by auto +qed + +lemma float_pos_eq_mantissa_pos: "x > 0 \ mantissa x > 0" + apply (subst Float_mantissa_exponent[of x, symmetric]) + apply (auto simp add: zero_less_mult_iff zero_float_def powr_gt_zero[of 2 "exponent x"] dest: less_zeroE) + using powr_gt_zero[of 2 "exponent x"] + apply simp + done + +lemma Float_pos_eq_mantissa_pos: "Float m e > 0 \ m > 0" + apply (auto simp add: zero_less_mult_iff zero_float_def powr_gt_zero[of 2 "exponent x"] dest: less_zeroE) + using powr_gt_zero[of 2 "e"] + apply simp + done + +lemma Float_representation_aux: + fixes m e + defines "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 \ float_of 0" + unfolding zero_float_def[symmetric] using `0 < x` by auto + from denormalize_shift[OF assms(1) this] guess i . note i = this + + have "2 powr (1 - (real (bitlen (mantissa x)) + real i)) = + 2 powr (1 - (real (bitlen (mantissa x)))) * inverse (2 powr (real i))" + by (simp add: powr_minus[symmetric] powr_add[symmetric] field_simps) + hence "real (mantissa x) * 2 powr (1 - real (bitlen (mantissa x))) = + (real (mantissa x) * 2 ^ i) * 2 powr (1 - real (bitlen (mantissa x * 2 ^ i)))" + using `mantissa x > 0` by (simp add: powr_realpow) + then show ?th2 + unfolding i by transfer auto +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. 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 (horner (Float 1 -1) + horner (x * rapprox_rat prec 2 3 - 1)) + else let l = bitlen m - 1 in + Some (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. 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 (horner (Float 1 -1) + + horner (max (x * lapprox_rat prec 2 3 - 1) 0)) + else let l = bitlen m - 1 in + Some (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_leE) 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)" + def bl \ "bitlen m - 1" have "0 < real m" and "\X. (0 :: real) < 2^X" and "0 < (2 :: real)" and "m \ 0" using assms by auto - hence "0 \ bitlen m - 1" using bitlen_ge1[OF `m \ 0`] by auto + hence "0 \ bl" by (simp add: bitlen_def bl_def) show ?thesis proof (cases "0 \ e") - case True - show ?thesis unfolding normalized_float[OF `m \ 0`] - unfolding ln_div[OF `0 < real m` `0 < ?B`] real_of_int_add ln_realpow[OF `0 < 2`] - unfolding real_of_float_ge0_exp[OF True] ln_mult[OF `0 < real m` `0 < 2^nat e`] - ln_realpow[OF `0 < 2`] algebra_simps using `0 \ bitlen m - 1` True by auto + case True + thus ?thesis + unfolding bl_def[symmetric] using `0 < real 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 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 unfolding normalized_float[OF `m \ 0`] - unfolding ln_div[OF `0 < real m` `0 < ?B`] real_of_int_add ln_realpow[OF `0 < 2`] - unfolding real_of_float_nge0_exp[OF False] ln_mult[OF `0 < real m` inv_gt0] ln_inverse[OF pow_gt0] - ln_realpow[OF `0 < 2`] algebra_simps using `0 \ bitlen m - 1` False by auto + show ?thesis using False unfolding bl_def[symmetric] using `0 < real m` `0 \ bl` + apply (simp add: ln_mult lne) + apply (cases "e=0") + apply (cases "bl = 0", simp_all add: powr_minus ln_inverse ln_powr) + apply (simp add: ln_inverse lne) + apply (cases "bl = 0", simp_all add: ln_inverse ln_powr field_simps) + done qed qed @@ -1837,11 +1891,11 @@ (is "?lb \ ?ln \ ?ln \ ?ub") proof (cases "x < Float 1 1") case True - hence "real (x - 1) < 1" and "real x < 2" unfolding less_float_def Float_num by auto - have "\ x \ 0" and "\ x < 1" using `1 \ x` unfolding less_float_def le_float_def by auto - hence "0 \ real (x - 1)" using `1 \ x` unfolding less_float_def Float_num by auto - - have [simp]: "(Float 3 -1) = 3 / 2" by (simp add: real_of_float_def pow2_def) + hence "real (x - 1) < 1" and "real x < 2" by auto + have "\ x \ 0" and "\ x < 1" using `1 \ x` by auto + hence "0 \ real (x - 1)" using `1 \ x` by auto + + have [simp]: "(Float 3 -1) = 3 / 2" by simp show ?thesis proof (cases "x \ Float 3 -1") @@ -1850,7 +1904,7 @@ using ln_float_bounds[OF `0 \ real (x - 1)` `real (x - 1) < 1`, of prec] `\ x \ 0` `\ x < 1` True by auto next - case False hence *: "3 / 2 < x" by (auto simp add: le_float_def) + 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 x * 2 / 3)" @@ -1891,7 +1945,7 @@ by (rule order_trans[OF lapprox_rat], simp) have low: "0 \ real (lapprox_rat prec 2 3)" - using lapprox_rat_bottom[of 2 3 prec] by simp + using lapprox_rat_nonneg[of 2 3 prec] by simp have "?lb_horner ?max \ ln (real ?max + 1)" @@ -1907,7 +1961,7 @@ show "0 < real x * 2/3" using * by auto show "real ?max + 1 \ real x * 2/3" using * up by (cases "0 < real x * real (lapprox_posrat prec 2 3) - 1", - auto simp add: real_of_float_max min_max.sup_absorb1) + auto simp add: max_def) qed finally have "?lb_horner (Float 1 -1) + ?lb_horner ?max \ ln x" @@ -1919,55 +1973,60 @@ next case False hence "\ x \ 0" and "\ x < 1" "0 < x" "\ x \ Float 3 -1" - using `1 \ x` unfolding less_float_def le_float_def real_of_float_simp pow2_def - by auto + using `1 \ x` by auto show ?thesis - proof (cases x) - case (Float m e) + proof - + def m \ "mantissa x" + def 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 float_pos_m_pos `0 < x` Float by auto + have "0 < m" and "m \ 0" using `0 < x` Float powr_gt_zero[of 2 e] + by (auto simp: zero_less_mult_iff) + def bl \ "bitlen m - 1" hence "bl \ 0" using `m > 0` by (simp add: bitlen_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 (?x - 1)" "real (?x - 1) < 1" + unfolding bl_def[symmetric] + by (auto simp: powr_realpow[symmetric] field_simps inverse_eq_divide) + (auto simp : powr_minus field_simps inverse_eq_divide) { have "lb_ln2 prec * ?s \ ln 2 * (e + (bitlen m - 1))" (is "?lb2 \ _") - unfolding real_of_float_mult real_of_float_ge0_exp[OF order_refl] nat_0 power_0 mult_1_right + unfolding nat_0 power_0 mult_1_right times_float.rep_eq using lb_ln2[of prec] - proof (rule mult_right_mono) - have "1 \ Float m e" using `1 \ x` Float unfolding le_float_def by auto - from float_gt1_scale[OF this] - show "0 \ real (e + (bitlen m - 1))" by auto - qed + proof (rule mult_mono) + from float_gt1_scale[OF `1 \ Float m e`] + show "0 \ real (Float (e + (bitlen m - 1)) 0)" by simp + qed auto moreover - from bitlen_div[OF `0 < m`, unfolded normalized_float[OF `m \ 0`, symmetric]] - have "0 \ real (?x - 1)" and "real (?x - 1) < 1" by auto - from ln_float_bounds(1)[OF this] + from ln_float_bounds(1)[OF x_bnds] have "(?x - 1) * lb_ln_horner prec (get_even prec) 1 (?x - 1) \ ln ?x" (is "?lb_horner \ _") by auto ultimately have "?lb2 + ?lb_horner \ ln x" unfolding Float ln_shifted_float[OF `0 < m`, of e] by auto } moreover { - from bitlen_div[OF `0 < m`, unfolded normalized_float[OF `m \ 0`, symmetric]] - have "0 \ real (?x - 1)" and "real (?x - 1) < 1" by auto - from ln_float_bounds(2)[OF this] + from ln_float_bounds(2)[OF x_bnds] have "ln ?x \ (?x - 1) * ub_ln_horner prec (get_odd prec) 1 (?x - 1)" (is "_ \ ?ub_horner") by auto moreover have "ln 2 * (e + (bitlen m - 1)) \ ub_ln2 prec * ?s" (is "_ \ ?ub2") - unfolding real_of_float_mult real_of_float_ge0_exp[OF order_refl] nat_0 power_0 mult_1_right + unfolding nat_0 power_0 mult_1_right times_float.rep_eq using ub_ln2[of prec] - proof (rule mult_right_mono) - have "1 \ Float m e" using `1 \ x` Float unfolding le_float_def by auto - from float_gt1_scale[OF this] + proof (rule mult_mono) + from float_gt1_scale[OF `1 \ Float m e`] show "0 \ real (e + (bitlen m - 1))" by auto - qed + next + have "0 \ ln 2" by simp + thus "0 \ real (ub_ln2 prec)" using ub_ln2[of prec] by arith + qed auto ultimately have "ln x \ ?ub2 + ?ub_horner" unfolding Float ln_shifted_float[OF `0 < m`, of e] by auto } 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 scale.simps[of m e, unfolded Float[symmetric]] mantissa.simps[of m e, unfolded Float[symmetric]] real_of_float_add - by auto + unfolding plus_float.rep_eq e_def[symmetric] m_def[symmetric] by simp qed qed @@ -1975,33 +2034,33 @@ 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 le_float_def by auto + 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` unfolding less_float_def le_float_def by auto - - have "0 < real x" and "real x \ 0" using `0 < x` unfolding less_float_def by auto + case True have "\ x \ 0" using `0 < x` by auto + from True have "real x < 1" by simp + have "0 < real x" and "real x \ 0" using `0 < x` by auto hence A: "0 < 1 / real x" by auto { let ?divl = "float_divl (max prec 1) 1 x" - have A': "1 \ ?divl" using float_divl_pos_less1_bound[OF `0 < x` `x < 1`] unfolding le_float_def less_float_def by auto - hence B: "0 < real ?divl" unfolding le_float_def by auto + have A': "1 \ ?divl" using float_divl_pos_less1_bound[OF `0 < real x` `real x < 1`] by auto + hence B: "0 < real ?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 x \ 0`, symmetric] ln_inverse[OF `0 < real 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 real_of_float_minus by (rule order_trans) + 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 le_float_def less_float_def by auto - hence B: "0 < real ?divr" unfolding le_float_def by auto + 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 ?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 x \ 0`, symmetric] ln_inverse[OF `0 < real 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 real_of_float_minus by (rule order_trans) + 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 @@ -2012,10 +2071,10 @@ proof - have "0 < x" proof (rule ccontr) - assume "\ 0 < x" hence "x \ 0" unfolding le_float_def less_float_def by auto + 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 x" unfolding less_float_def by auto + thus "0 < real 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 @@ -2025,10 +2084,10 @@ proof - have "0 < x" proof (rule ccontr) - assume "\ 0 < x" hence "x \ 0" unfolding le_float_def less_float_def by auto + assume "\ 0 < x" hence "x \ 0" by auto thus False using assms by auto qed - thus "0 < real x" unfolding less_float_def by auto + thus "0 < real 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 @@ -2174,12 +2233,12 @@ 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 (round_down prec l, round_up prec u) | None \ None)" | +"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. (l1 + l2, 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) - (\ a1 a2 b1 b2. (float_nprt a1 * float_pprt b2 + float_nprt a2 * float_nprt b2 + float_pprt a1 * float_pprt b1 + float_pprt a2 * float_nprt b1, - float_pprt a2 * float_pprt b2 + float_pprt a1 * float_nprt b2 + float_nprt a2 * float_pprt b1 + float_nprt a1 * float_nprt b1))" | + (\ a1 a2 b1 b2. (nprt a1 * pprt b2 + nprt a2 * nprt b2 + pprt a1 * pprt b1 + pprt a2 * nprt b1, + pprt a2 * pprt b2 + pprt a1 * nprt b2 + nprt a2 * pprt b1 + nprt a1 * nprt b1))" | "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)" | @@ -2233,11 +2292,11 @@ 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 = round_down prec l'" and u': "u = round_up prec u'" + 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] round_up[of u']] - using order_trans[OF round_down[of _ l'] Pa[OF S, THEN conjunct1]] by auto + 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': @@ -2389,16 +2448,16 @@ from lift_un'[OF Minus.prems[unfolded approx.simps]] Minus.hyps obtain l1 u1 where "l = -u1" and "u = -l1" "l1 \ interpret_floatarith a xs" and "interpret_floatarith a xs \ u1" unfolding fst_conv snd_conv by blast - thus ?case unfolding interpret_floatarith.simps using real_of_float_minus by auto + 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 = float_nprt l1 * float_pprt u2 + float_nprt u1 * float_nprt u2 + float_pprt l1 * float_pprt l2 + float_pprt u1 * float_nprt l2" - and u: "u = float_pprt u1 * float_pprt u2 + float_pprt l1 * float_nprt u2 + float_nprt u1 * float_pprt l2 + float_nprt l1 * float_nprt l2" + where l: "l = nprt l1 * pprt u2 + nprt u1 * nprt u2 + pprt l1 * pprt l2 + pprt u1 * nprt l2" + and u: "u = pprt u1 * pprt u2 + pprt l1 * nprt u2 + nprt u1 * pprt l2 + nprt l1 * nprt l2" and "l1 \ interpret_floatarith a xs" and "interpret_floatarith a xs \ u1" and "l2 \ interpret_floatarith b xs" and "interpret_floatarith b xs \ u2" unfolding fst_conv snd_conv by blast - thus ?case unfolding interpret_floatarith.simps l u real_of_float_add real_of_float_mult real_of_float_nprt real_of_float_pprt + thus ?case unfolding interpret_floatarith.simps l u using mult_le_prts mult_ge_prts by auto next case (Inverse a) @@ -2408,13 +2467,13 @@ 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 l1 \ real u1" using l1 u1 by auto - ultimately have "real l1 \ 0" and "real u1 \ 0" unfolding less_float_def by auto + ultimately have "real l1 \ 0" and "real 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 u1" and "0 < real l1" "0 < interpret_floatarith a xs" - unfolding less_float_def using l1_le_u1 l1 by auto + using l1_le_u1 l1 by auto show ?thesis unfolding inverse_le_iff_le[OF `0 < real u1` `0 < interpret_floatarith a xs`] inverse_le_iff_le[OF `0 < interpret_floatarith a xs` `0 < real l1`] @@ -2422,7 +2481,7 @@ next case False hence "u1 < 0" using either by blast hence "real u1 < 0" and "real l1 < 0" "interpret_floatarith a xs < 0" - unfolding less_float_def using l1_le_u1 u1 by auto + using l1_le_u1 u1 by auto show ?thesis unfolding inverse_le_iff_le_neg[OF `real u1 < 0` `interpret_floatarith a xs < 0`] inverse_le_iff_le_neg[OF `interpret_floatarith a xs < 0` `real l1 < 0`] @@ -2443,7 +2502,7 @@ 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 real_of_float_abs less_float_def) + 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 @@ -2527,7 +2586,7 @@ let ?m = "(l + u) * Float 1 -1" have "real l \ ?m" and "?m \ real u" - unfolding le_float_def using Suc.prems by auto + unfolding less_eq_float_def using Suc.prems by auto with `x \ { l .. u }` have "x \ { l .. ?m} \ x \ { ?m .. u }" by auto @@ -2576,7 +2635,7 @@ then obtain n where x_eq: "x = Var n" by (cases x) auto - with Assign.prems obtain l u' l' u + with Assign.prems obtain l u where bnd_eq: "Some (l, u) = approx prec a vs" and x_eq: "x = Var n" and approx_form': "approx_form' prec f (ss ! n) n l u vs ss" @@ -2602,7 +2661,7 @@ and inequality: "u < l'" by (cases "approx prec a vs", auto, cases "approx prec b vs", auto) - from inequality[unfolded less_float_def] approx[OF Less.prems(2) l_eq] approx[OF Less.prems(2) u_eq] + from inequality approx[OF Less.prems(2) l_eq] approx[OF Less.prems(2) u_eq] show ?case by auto next case (LessEqual a b) @@ -2612,7 +2671,7 @@ and inequality: "u \ l'" by (cases "approx prec a vs", auto, cases "approx prec b vs", auto) - from inequality[unfolded le_float_def] approx[OF LessEqual.prems(2) l_eq] approx[OF LessEqual.prems(2) u_eq] + from inequality approx[OF LessEqual.prems(2) l_eq] approx[OF LessEqual.prems(2) u_eq] show ?case by auto next case (AtLeastAtMost x a b) @@ -2624,7 +2683,7 @@ by (cases "approx prec x vs", auto, cases "approx prec a vs", auto, cases "approx prec b vs", auto, blast) - from inequality[unfolded le_float_def] approx[OF AtLeastAtMost.prems(2) l_eq] approx[OF AtLeastAtMost.prems(2) u_eq] approx[OF AtLeastAtMost.prems(2) x_eq] + from inequality 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 qed @@ -2685,12 +2744,13 @@ by (auto intro!: DERIV_intros simp add: algebra_simps power2_eq_square) next case (Cos a) thus ?case - by (auto intro!: DERIV_intros + apply (auto intro!: DERIV_intros simp del: interpret_floatarith.simps(5) simp add: interpret_floatarith_sin interpret_floatarith.simps(5)[of a]) + done next case (Power a n) thus ?case by (cases n, auto intro!: DERIV_intros - simp del: power_Suc simp add: real_eq_of_nat) + simp del: power_Suc) next case (Ln a) thus ?case by (auto intro!: DERIV_intros simp add: divide_inverse) next case (Var i) thus ?case using `n < length vs` by auto @@ -2730,7 +2790,7 @@ and *: "0 < l \ u < 0" by (cases "approx prec a vs", auto) with approx[OF `bounded_by xs vs` approx_Some] - have "interpret_floatarith a xs \ 0" unfolding less_float_def by auto + have "interpret_floatarith a xs \ 0" by auto thus ?case using Inverse by auto next case (Ln a) @@ -2738,7 +2798,7 @@ and *: "0 < l" by (cases "approx prec a vs", auto) with approx[OF `bounded_by xs vs` approx_Some] - have "0 < interpret_floatarith a xs" unfolding less_float_def by auto + have "0 < interpret_floatarith a xs" by auto thus ?case using Ln by auto next case (Sqrt a) @@ -2746,7 +2806,7 @@ and *: "0 < l" by (cases "approx prec a vs", auto) with approx[OF `bounded_by xs vs` approx_Some] - have "0 < interpret_floatarith a xs" unfolding less_float_def by auto + have "0 < interpret_floatarith a xs" by auto thus ?case using Sqrt by auto next case (Power a n) thus ?case by (cases n, auto) @@ -3056,7 +3116,7 @@ by (auto simp add: Let_def lazy_conj) have m_l: "real l \ ?m" and m_u: "?m \ real u" - unfolding le_float_def using Suc.prems by auto + unfolding less_eq_float_def using Suc.prems by auto with `x \ { l .. u }` have "x \ { l .. ?m} \ x \ { ?m .. u }" by auto @@ -3096,7 +3156,7 @@ from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x' have "ly \ interpret_floatarith a [x] - interpret_floatarith b [x]" by (auto simp add: diff_minus) - from order_less_le_trans[OF `0 < ly`[unfolded less_float_def] this] + from order_less_le_trans[OF _ this, of 0] `0 < ly` show ?thesis by auto qed @@ -3118,7 +3178,7 @@ from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x' have "ly \ interpret_floatarith a [x] - interpret_floatarith b [x]" by (auto simp add: diff_minus) - from order_trans[OF `0 \ ly`[unfolded le_float_def] this] + from order_trans[OF _ this, of 0] `0 \ ly` show ?thesis by auto qed diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/Decision_Procs/ex/Approximation_Ex.thy --- a/src/HOL/Decision_Procs/ex/Approximation_Ex.thy Thu Apr 19 11:14:57 2012 +0200 +++ b/src/HOL/Decision_Procs/ex/Approximation_Ex.thy Thu Apr 19 17:49:02 2012 +0200 @@ -36,19 +36,19 @@ section "Compute some transcendental values" lemma "\ ln 2 - 544531980202654583340825686620847 / 785593587443817081832229725798400 \ < inverse (2^51) " - by (approximation 80) + by (approximation 60) lemma "\ exp 1.626 - 5.083499996273 \ < (inverse 10 ^ 10 :: real)" - by (approximation 80) + by (approximation 40) lemma "\ sqrt 2 - 1.4142135623730951 \ < inverse 10 ^ 16" - by (approximation 80) + by (approximation 60) lemma "\ pi - 3.1415926535897932385 \ < inverse 10 ^ 18" - by (approximation 80) + by (approximation 70) lemma "\ sin 100 + 0.50636564110975879 \ < inverse 10 ^ 17" - by (approximation 80) + by (approximation 70) section "Use variable ranges" @@ -70,7 +70,7 @@ lemma defines "g \ 9.80665" and "v \ 128.61" and "d \ pi / 180" shows "g / v * tan (35 * d) \ { 3 * d .. 3.1 * d }" - using assms by (approximation 80) + using assms by (approximation 20) lemma "x \ { 0 .. 1 :: real } \ x ^ 2 \ x" by (approximation 30 splitting: x=1 taylor: x = 3) diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/Fun.thy --- a/src/HOL/Fun.thy Thu Apr 19 11:14:57 2012 +0200 +++ b/src/HOL/Fun.thy Thu Apr 19 17:49:02 2012 +0200 @@ -24,11 +24,11 @@ lemma id_apply [simp]: "id x = x" by (simp add: id_def) -lemma image_id [simp]: "id ` Y = Y" - by (simp add: id_def) +lemma image_id [simp]: "image id = id" + by (simp add: id_def fun_eq_iff) -lemma vimage_id [simp]: "id -` A = A" - by (simp add: id_def) +lemma vimage_id [simp]: "vimage id = id" + by (simp add: id_def fun_eq_iff) subsection {* The Composition Operator @{text "f \ g"} *} diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_Int0.thy --- a/src/HOL/IMP/Abs_Int0.thy Thu Apr 19 11:14:57 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,411 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory Abs_Int0 -imports Abs_State -begin - -subsection "Computable Abstract Interpretation" - -text{* Abstract interpretation over type @{text st} instead of -functions. *} - -context Gamma -begin - -fun aval' :: "aexp \ 'av st \ 'av" where -"aval' (N n) S = num' n" | -"aval' (V x) S = lookup S x" | -"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" - -lemma aval'_sound: "s : \\<^isub>f S \ aval a s : \(aval' a S)" -by (induction a) (auto simp: gamma_num' gamma_plus' \_st_def lookup_def) - -end - -text{* The for-clause (here and elsewhere) only serves the purpose of fixing -the name of the type parameter @{typ 'av} which would otherwise be renamed to -@{typ 'a}. *} - -locale Abs_Int = Gamma where \=\ for \ :: "'av::SL_top \ val set" -begin - -fun step' :: "'av st option \ 'av st option acom \ 'av st option acom" where -"step' S (SKIP {P}) = (SKIP {S})" | -"step' S (x ::= e {P}) = - x ::= e {case S of None \ None | Some S \ Some(update S x (aval' e S))}" | -"step' S (c1; c2) = step' S c1; step' (post c1) c2" | -"step' S (IF b THEN c1 ELSE c2 {P}) = - (let c1' = step' S c1; c2' = step' S c2 - in IF b THEN c1' ELSE c2' {post c1 \ post c2})" | -"step' S ({Inv} WHILE b DO c {P}) = - {S \ post c} WHILE b DO step' Inv c {Inv}" - -definition AI :: "com \ 'av st option acom option" where -"AI = lpfp\<^isub>c (step' \)" - - -lemma strip_step'[simp]: "strip(step' S c) = strip c" -by(induct c arbitrary: S) (simp_all add: Let_def) - - -text{* Soundness: *} - -lemma in_gamma_update: - "\ s : \\<^isub>f S; i : \ a \ \ s(x := i) : \\<^isub>f(update S x a)" -by(simp add: \_st_def lookup_update) - -text{* The soundness proofs are textually identical to the ones for the step -function operating on states as functions. *} - -lemma step_preserves_le: - "\ S \ \\<^isub>o S'; c \ \\<^isub>c c' \ \ step S c \ \\<^isub>c (step' S' c')" -proof(induction c arbitrary: c' S S') - case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) -next - case Assign thus ?case - by (fastforce simp: Assign_le map_acom_Assign intro: aval'_sound in_gamma_update - split: option.splits del:subsetD) -next - case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi) - by (metis le_post post_map_acom) -next - case (If b c1 c2 P) - then obtain c1' c2' P' where - "c' = IF b THEN c1' ELSE c2' {P'}" - "P \ \\<^isub>o P'" "c1 \ \\<^isub>c c1'" "c2 \ \\<^isub>c c2'" - by (fastforce simp: If_le map_acom_If) - moreover have "post c1 \ \\<^isub>o(post c1' \ post c2')" - by (metis (no_types) `c1 \ \\<^isub>c c1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom) - moreover have "post c2 \ \\<^isub>o(post c1' \ post c2')" - by (metis (no_types) `c2 \ \\<^isub>c c2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom) - ultimately show ?case using `S \ \\<^isub>o S'` by (simp add: If.IH subset_iff) -next - case (While I b c1 P) - then obtain c1' I' P' where - "c' = {I'} WHILE b DO c1' {P'}" - "I \ \\<^isub>o I'" "P \ \\<^isub>o P'" "c1 \ \\<^isub>c c1'" - by (fastforce simp: map_acom_While While_le) - moreover have "S \ post c1 \ \\<^isub>o (S' \ post c1')" - using `S \ \\<^isub>o S'` le_post[OF `c1 \ \\<^isub>c c1'`, simplified] - by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans) - ultimately show ?case by (simp add: While.IH subset_iff) -qed - -lemma AI_sound: "AI c = Some c' \ CS c \ \\<^isub>c c'" -proof(simp add: CS_def AI_def) - assume 1: "lpfp\<^isub>c (step' \) c = Some c'" - have 2: "step' \ c' \ c'" by(rule lpfpc_pfp[OF 1]) - have 3: "strip (\\<^isub>c (step' \ c')) = c" - by(simp add: strip_lpfpc[OF _ 1]) - have "lfp (step UNIV) c \ \\<^isub>c (step' \ c')" - proof(rule lfp_lowerbound[simplified,OF 3]) - show "step UNIV (\\<^isub>c (step' \ c')) \ \\<^isub>c (step' \ c')" - proof(rule step_preserves_le[OF _ _]) - show "UNIV \ \\<^isub>o \" by simp - show "\\<^isub>c (step' \ c') \ \\<^isub>c c'" by(rule mono_gamma_c[OF 2]) - qed - qed - from this 2 show "lfp (step UNIV) c \ \\<^isub>c c'" - by (blast intro: mono_gamma_c order_trans) -qed - -end - - -subsubsection "Monotonicity" - -locale Abs_Int_mono = Abs_Int + -assumes mono_plus': "a1 \ b1 \ a2 \ b2 \ plus' a1 a2 \ plus' b1 b2" -begin - -lemma mono_aval': "S \ S' \ aval' e S \ aval' e S'" -by(induction e) (auto simp: le_st_def lookup_def mono_plus') - -lemma mono_update: "a \ a' \ S \ S' \ update S x a \ update S' x a'" -by(auto simp add: le_st_def lookup_def update_def) - -lemma mono_step': "S \ S' \ c \ c' \ step' S c \ step' S' c'" -apply(induction c c' arbitrary: S S' rule: le_acom.induct) -apply (auto simp: Let_def mono_update mono_aval' mono_post le_join_disj - split: option.split) -done - -end - - -subsubsection "Ascending Chain Condition" - -hide_const (open) acc - -abbreviation "strict r == r \ -(r^-1)" -abbreviation "acc r == wf((strict r)^-1)" - -lemma strict_inv_image: "strict(inv_image r f) = inv_image (strict r) f" -by(auto simp: inv_image_def) - -lemma acc_inv_image: - "acc r \ acc (inv_image r f)" -by (metis converse_inv_image strict_inv_image wf_inv_image) - -text{* ACC for option type: *} - -lemma acc_option: assumes "acc {(x,y::'a::preord). x \ y}" -shows "acc {(x,y::'a::preord option). x \ y}" -proof(auto simp: wf_eq_minimal) - fix xo :: "'a option" and Qo assume "xo : Qo" - let ?Q = "{x. Some x \ Qo}" - show "\yo\Qo. \zo. yo \ zo \ ~ zo \ yo \ zo \ Qo" (is "\zo\Qo. ?P zo") - proof cases - assume "?Q = {}" - hence "?P None" by auto - moreover have "None \ Qo" using `?Q = {}` `xo : Qo` - by auto (metis not_Some_eq) - ultimately show ?thesis by blast - next - assume "?Q \ {}" - with assms show ?thesis - apply(auto simp: wf_eq_minimal) - apply(erule_tac x="?Q" in allE) - apply auto - apply(rule_tac x = "Some z" in bexI) - by auto - qed -qed - -text{* ACC for abstract states, via measure functions. *} - -(*FIXME mv*) -lemma setsum_strict_mono1: -fixes f :: "'a \ 'b::{comm_monoid_add, ordered_cancel_ab_semigroup_add}" -assumes "finite A" and "ALL x:A. f x \ g x" and "EX a:A. f a < g a" -shows "setsum f A < setsum g A" -proof- - from assms(3) obtain a where a: "a:A" "f a < g a" by blast - have "setsum f A = setsum f ((A-{a}) \ {a})" - by(simp add:insert_absorb[OF `a:A`]) - also have "\ = setsum f (A-{a}) + setsum f {a}" - using `finite A` by(subst setsum_Un_disjoint) auto - also have "setsum f (A-{a}) \ setsum g (A-{a})" - by(rule setsum_mono)(simp add: assms(2)) - also have "setsum f {a} < setsum g {a}" using a by simp - also have "setsum g (A - {a}) + setsum g {a} = setsum g((A-{a}) \ {a})" - using `finite A` by(subst setsum_Un_disjoint[symmetric]) auto - also have "\ = setsum g A" by(simp add:insert_absorb[OF `a:A`]) - finally show ?thesis by (metis add_right_mono add_strict_left_mono) -qed - -lemma measure_st: assumes "(strict{(x,y::'a::SL_top). x \ y})^-1 <= measure m" -and "\x y::'a::SL_top. x \ y \ y \ x \ m x = m y" -shows "(strict{(S,S'::'a::SL_top st). S \ S'})^-1 \ - measure(%fd. \x| x\set(dom fd) \ ~ \ \ fun fd x. m(fun fd x)+1)" -proof- - { fix S S' :: "'a st" assume "S \ S'" "~ S' \ S" - let ?X = "set(dom S)" let ?Y = "set(dom S')" - let ?f = "fun S" let ?g = "fun S'" - let ?X' = "{x:?X. ~ \ \ ?f x}" let ?Y' = "{y:?Y. ~ \ \ ?g y}" - from `S \ S'` have "ALL y:?Y'\?X. ?f y \ ?g y" - by(auto simp: le_st_def lookup_def) - hence 1: "ALL y:?Y'\?X. m(?g y)+1 \ m(?f y)+1" - using assms(1,2) by(fastforce) - from `~ S' \ S` obtain u where u: "u : ?X" "~ lookup S' u \ ?f u" - by(auto simp: le_st_def) - hence "u : ?X'" by simp (metis preord_class.le_trans top) - have "?Y'-?X = {}" using `S \ S'` by(fastforce simp: le_st_def lookup_def) - have "?Y'\?X <= ?X'" apply auto - apply (metis `S \ S'` le_st_def lookup_def preord_class.le_trans) - done - have "(\y\?Y'. m(?g y)+1) = (\y\(?Y'-?X) \ (?Y'\?X). m(?g y)+1)" - by (metis Un_Diff_Int) - also have "\ = (\y\?Y'\?X. m(?g y)+1)" - using `?Y'-?X = {}` by (metis Un_empty_left) - also have "\ < (\x\?X'. m(?f x)+1)" - proof cases - assume "u \ ?Y'" - hence "m(?g u) < m(?f u)" using assms(1) `S \ S'` u - by (fastforce simp: le_st_def lookup_def) - have "(\y\?Y'\?X. m(?g y)+1) < (\y\?Y'\?X. m(?f y)+1)" - using `u:?X` `u:?Y'` `m(?g u) < m(?f u)` - by(fastforce intro!: setsum_strict_mono1[OF _ 1]) - also have "\ \ (\y\?X'. m(?f y)+1)" - by(simp add: setsum_mono3[OF _ `?Y'\?X <= ?X'`]) - finally show ?thesis . - next - assume "u \ ?Y'" - with `?Y'\?X <= ?X'` have "?Y'\?X - {u} <= ?X' - {u}" by blast - have "(\y\?Y'\?X. m(?g y)+1) = (\y\?Y'\?X - {u}. m(?g y)+1)" - proof- - have "?Y'\?X = ?Y'\?X - {u}" using `u \ ?Y'` by auto - thus ?thesis by metis - qed - also have "\ < (\y\?Y'\?X-{u}. m(?g y)+1) + (\y\{u}. m(?f y)+1)" by simp - also have "(\y\?Y'\?X-{u}. m(?g y)+1) \ (\y\?Y'\?X-{u}. m(?f y)+1)" - using 1 by(blast intro: setsum_mono) - also have "\ \ (\y\?X'-{u}. m(?f y)+1)" - by(simp add: setsum_mono3[OF _ `?Y'\?X-{u} <= ?X'-{u}`]) - also have "\ + (\y\{u}. m(?f y)+1)= (\y\(?X'-{u}) \ {u}. m(?f y)+1)" - using `u:?X'` by(subst setsum_Un_disjoint[symmetric]) auto - also have "\ = (\x\?X'. m(?f x)+1)" - using `u : ?X'` by(simp add:insert_absorb) - finally show ?thesis by (blast intro: add_right_mono) - qed - finally have "(\y\?Y'. m(?g y)+1) < (\x\?X'. m(?f x)+1)" . - } thus ?thesis by(auto simp add: measure_def inv_image_def) -qed - -text{* ACC for acom. First the ordering on acom is related to an ordering on -lists of annotations. *} - -(* FIXME mv and add [simp] *) -lemma listrel_Cons_iff: - "(x#xs, y#ys) : listrel r \ (x,y) \ r \ (xs,ys) \ listrel r" -by (blast intro:listrel.Cons) - -lemma listrel_app: "(xs1,ys1) : listrel r \ (xs2,ys2) : listrel r - \ (xs1@xs2, ys1@ys2) : listrel r" -by(auto simp add: listrel_iff_zip) - -lemma listrel_app_same_size: "size xs1 = size ys1 \ size xs2 = size ys2 \ - (xs1@xs2, ys1@ys2) : listrel r \ - (xs1,ys1) : listrel r \ (xs2,ys2) : listrel r" -by(auto simp add: listrel_iff_zip) - -lemma listrel_converse: "listrel(r^-1) = (listrel r)^-1" -proof- - { fix xs ys - have "(xs,ys) : listrel(r^-1) \ (ys,xs) : listrel r" - apply(induct xs arbitrary: ys) - apply (fastforce simp: listrel.Nil) - apply (fastforce simp: listrel_Cons_iff) - done - } thus ?thesis by auto -qed - -(* It would be nice to get rid of refl & trans and build them into the proof *) -lemma acc_listrel: fixes r :: "('a*'a)set" assumes "refl r" and "trans r" -and "acc r" shows "acc (listrel r - {([],[])})" -proof- - have refl: "!!x. (x,x) : r" using `refl r` unfolding refl_on_def by blast - have trans: "!!x y z. (x,y) : r \ (y,z) : r \ (x,z) : r" - using `trans r` unfolding trans_def by blast - from assms(3) obtain mx :: "'a set \ 'a" where - mx: "!!S x. x:S \ mx S : S \ (\y. (mx S,y) : strict r \ y \ S)" - by(simp add: wf_eq_minimal) metis - let ?R = "listrel r - {([], [])}" - { fix Q and xs :: "'a list" - have "xs \ Q \ \ys. ys\Q \ (\zs. (ys, zs) \ strict ?R \ zs \ Q)" - (is "_ \ \ys. ?P Q ys") - proof(induction xs arbitrary: Q rule: length_induct) - case (1 xs) - { have "!!ys Q. size ys < size xs \ ys : Q \ EX ms. ?P Q ms" - using "1.IH" by blast - } note IH = this - show ?case - proof(cases xs) - case Nil with `xs : Q` have "?P Q []" by auto - thus ?thesis by blast - next - case (Cons x ys) - let ?Q1 = "{a. \bs. size bs = size ys \ a#bs : Q}" - have "x : ?Q1" using `xs : Q` Cons by auto - from mx[OF this] obtain m1 where - 1: "m1 \ ?Q1 \ (\y. (m1,y) \ strict r \ y \ ?Q1)" by blast - then obtain ms1 where "size ms1 = size ys" "m1#ms1 : Q" by blast+ - hence "size ms1 < size xs" using Cons by auto - let ?Q2 = "{bs. \m1'. (m1',m1):r \ (m1,m1'):r \ m1'#bs : Q \ size bs = size ms1}" - have "ms1 : ?Q2" using `m1#ms1 : Q` by(blast intro: refl) - from IH[OF `size ms1 < size xs` this] - obtain ms where 2: "?P ?Q2 ms" by auto - then obtain m1' where m1': "(m1',m1) : r \ (m1,m1') : r \ m1'#ms : Q" - by blast - hence "\ab. (m1'#ms,ab) : strict ?R \ ab \ Q" using 1 2 - apply (auto simp: listrel_Cons_iff) - apply (metis `length ms1 = length ys` listrel_eq_len trans) - by (metis `length ms1 = length ys` listrel_eq_len trans) - with m1' show ?thesis by blast - qed - qed - } - thus ?thesis unfolding wf_eq_minimal by (metis converse_iff) -qed - - -fun annos :: "'a acom \ 'a list" where -"annos (SKIP {a}) = [a]" | -"annos (x::=e {a}) = [a]" | -"annos (c1;c2) = annos c1 @ annos c2" | -"annos (IF b THEN c1 ELSE c2 {a}) = a # annos c1 @ annos c2" | -"annos ({i} WHILE b DO c {a}) = i # a # annos c" - -lemma size_annos_same: "strip c1 = strip c2 \ size(annos c1) = size(annos c2)" -apply(induct c2 arbitrary: c1) -apply (auto simp: strip_eq_SKIP strip_eq_Assign strip_eq_Semi strip_eq_If strip_eq_While) -done - -lemmas size_annos_same2 = eqTrueI[OF size_annos_same] - -lemma set_annos_anno: "set (annos (anno a c)) = {a}" -by(induction c)(auto) - -lemma le_iff_le_annos: "c1 \ c2 \ - (annos c1, annos c2) : listrel{(x,y). x \ y} \ strip c1 = strip c2" -apply(induct c1 c2 rule: le_acom.induct) -apply (auto simp: listrel.Nil listrel_Cons_iff listrel_app size_annos_same2) -apply (metis listrel_app_same_size size_annos_same)+ -done - -lemma le_acom_subset_same_annos: - "(strict{(c,c'::'a::preord acom). c \ c'})^-1 \ - (strict(inv_image (listrel{(a,a'::'a). a \ a'} - {([],[])}) annos))^-1" -by(auto simp: le_iff_le_annos) - -lemma acc_acom: "acc {(a,a'::'a::preord). a \ a'} \ - acc {(c,c'::'a acom). c \ c'}" -apply(rule wf_subset[OF _ le_acom_subset_same_annos]) -apply(rule acc_inv_image[OF acc_listrel]) -apply(auto simp: refl_on_def trans_def intro: le_trans) -done - -text{* Termination of the fixed-point finders, assuming monotone functions: *} - -lemma pfp_termination: -fixes x0 :: "'a::preord" -assumes mono: "\x y. x \ y \ f x \ f y" and "acc {(x::'a,y). x \ y}" -and "x0 \ f x0" shows "EX x. pfp f x0 = Some x" -proof(simp add: pfp_def, rule wf_while_option_Some[where P = "%x. x \ f x"]) - show "wf {(x, s). (s \ f s \ \ f s \ s) \ x = f s}" - by(rule wf_subset[OF assms(2)]) auto -next - show "x0 \ f x0" by(rule assms) -next - fix x assume "x \ f x" thus "f x \ f(f x)" by(rule mono) -qed - -lemma lpfpc_termination: - fixes f :: "(('a::SL_top)option acom \ 'a option acom)" - assumes "acc {(x::'a,y). x \ y}" and "\x y. x \ y \ f x \ f y" - and "\c. strip(f c) = strip c" - shows "\c'. lpfp\<^isub>c f c = Some c'" -unfolding lpfp\<^isub>c_def -apply(rule pfp_termination) - apply(erule assms(2)) - apply(rule acc_acom[OF acc_option[OF assms(1)]]) -apply(simp add: bot_acom assms(3)) -done - -context Abs_Int_mono -begin - -lemma AI_Some_measure: -assumes "(strict{(x,y::'a). x \ y})^-1 <= measure m" -and "\x y::'a. x \ y \ y \ x \ m x = m y" -shows "\c'. AI c = Some c'" -unfolding AI_def -apply(rule lpfpc_termination) -apply(rule wf_subset[OF wf_measure measure_st[OF assms]]) -apply(erule mono_step'[OF le_refl]) -apply(rule strip_step') -done - -end - -end diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_Int0_const.thy --- a/src/HOL/IMP/Abs_Int0_const.thy Thu Apr 19 11:14:57 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,139 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory Abs_Int0_const -imports Abs_Int0 Abs_Int_Tests -begin - -subsection "Constant Propagation" - -datatype const = Const val | Any - -fun \_const where -"\_const (Const n) = {n}" | -"\_const (Any) = UNIV" - -fun plus_const where -"plus_const (Const m) (Const n) = Const(m+n)" | -"plus_const _ _ = Any" - -lemma plus_const_cases: "plus_const a1 a2 = - (case (a1,a2) of (Const m, Const n) \ Const(m+n) | _ \ Any)" -by(auto split: prod.split const.split) - -instantiation const :: SL_top -begin - -fun le_const where -"_ \ Any = True" | -"Const n \ Const m = (n=m)" | -"Any \ Const _ = False" - -fun join_const where -"Const m \ Const n = (if n=m then Const m else Any)" | -"_ \ _ = Any" - -definition "\ = Any" - -instance -proof - case goal1 thus ?case by (cases x) simp_all -next - case goal2 thus ?case by(cases z, cases y, cases x, simp_all) -next - case goal3 thus ?case by(cases x, cases y, simp_all) -next - case goal4 thus ?case by(cases y, cases x, simp_all) -next - case goal5 thus ?case by(cases z, cases y, cases x, simp_all) -next - case goal6 thus ?case by(simp add: Top_const_def) -qed - -end - - -interpretation Val_abs -where \ = \_const and num' = Const and plus' = plus_const -proof - case goal1 thus ?case - by(cases a, cases b, simp, simp, cases b, simp, simp) -next - case goal2 show ?case by(simp add: Top_const_def) -next - case goal3 show ?case by simp -next - case goal4 thus ?case - by(auto simp: plus_const_cases split: const.split) -qed - -interpretation Abs_Int -where \ = \_const and num' = Const and plus' = plus_const -defines AI_const is AI and step_const is step' and aval'_const is aval' -.. - - -subsubsection "Tests" - -value "show_acom (((step_const \)^^0) (\\<^sub>c test1_const))" -value "show_acom (((step_const \)^^1) (\\<^sub>c test1_const))" -value "show_acom (((step_const \)^^2) (\\<^sub>c test1_const))" -value "show_acom (((step_const \)^^3) (\\<^sub>c test1_const))" -value "show_acom_opt (AI_const test1_const)" - -value "show_acom_opt (AI_const test2_const)" -value "show_acom_opt (AI_const test3_const)" - -value "show_acom (((step_const \)^^0) (\\<^sub>c test4_const))" -value "show_acom (((step_const \)^^1) (\\<^sub>c test4_const))" -value "show_acom (((step_const \)^^2) (\\<^sub>c test4_const))" -value "show_acom (((step_const \)^^3) (\\<^sub>c test4_const))" -value "show_acom_opt (AI_const test4_const)" - -value "show_acom (((step_const \)^^0) (\\<^sub>c test5_const))" -value "show_acom (((step_const \)^^1) (\\<^sub>c test5_const))" -value "show_acom (((step_const \)^^2) (\\<^sub>c test5_const))" -value "show_acom (((step_const \)^^3) (\\<^sub>c test5_const))" -value "show_acom (((step_const \)^^4) (\\<^sub>c test5_const))" -value "show_acom (((step_const \)^^5) (\\<^sub>c test5_const))" -value "show_acom_opt (AI_const test5_const)" - -value "show_acom (((step_const \)^^0) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^1) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^2) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^3) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^4) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^5) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^6) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^7) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^8) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^9) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^10) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^11) (\\<^sub>c test6_const))" -value "show_acom_opt (AI_const test6_const)" - - -text{* Monotonicity: *} - -interpretation Abs_Int_mono -where \ = \_const and num' = Const and plus' = plus_const -proof - case goal1 thus ?case - by(auto simp: plus_const_cases split: const.split) -qed - -text{* Termination: *} - -definition "m_const x = (case x of Const _ \ 1 | Any \ 0)" - -lemma measure_const: - "(strict{(x::const,y). x \ y})^-1 \ measure m_const" -by(auto simp: m_const_def split: const.splits) - -lemma measure_const_eq: - "\ x y::const. x \ y \ y \ x \ m_const x = m_const y" -by(auto simp: m_const_def split: const.splits) - -lemma "EX c'. AI_const c = Some c'" -by(rule AI_Some_measure[OF measure_const measure_const_eq]) - -end diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_Int0_fun.thy --- a/src/HOL/IMP/Abs_Int0_fun.thy Thu Apr 19 11:14:57 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,397 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory Abs_Int0_fun -imports "~~/src/HOL/ex/Interpretation_with_Defs" - "~~/src/HOL/Library/While_Combinator" - Collecting -begin - -subsection "Orderings" - -class preord = -fixes le :: "'a \ 'a \ bool" (infix "\" 50) -assumes le_refl[simp]: "x \ x" -and le_trans: "x \ y \ y \ z \ x \ z" -begin - -definition mono where "mono f = (\x y. x \ y \ f x \ f y)" - -lemma monoD: "mono f \ x \ y \ f x \ f y" by(simp add: mono_def) - -lemma mono_comp: "mono f \ mono g \ mono (g o f)" -by(simp add: mono_def) - -declare le_trans[trans] - -end - -text{* Note: no antisymmetry. Allows implementations where some abstract -element is implemented by two different values @{prop "x \ y"} -such that @{prop"x \ y"} and @{prop"y \ x"}. Antisymmetry is not -needed because we never compare elements for equality but only for @{text"\"}. -*} - -class SL_top = preord + -fixes join :: "'a \ 'a \ 'a" (infixl "\" 65) -fixes Top :: "'a" ("\") -assumes join_ge1 [simp]: "x \ x \ y" -and join_ge2 [simp]: "y \ x \ y" -and join_least: "x \ z \ y \ z \ x \ y \ z" -and top[simp]: "x \ \" -begin - -lemma join_le_iff[simp]: "x \ y \ z \ x \ z \ y \ z" -by (metis join_ge1 join_ge2 join_least le_trans) - -lemma le_join_disj: "x \ y \ x \ z \ x \ y \ z" -by (metis join_ge1 join_ge2 le_trans) - -end - -instantiation "fun" :: (type, SL_top) SL_top -begin - -definition "f \ g = (ALL x. f x \ g x)" -definition "f \ g = (\x. f x \ g x)" -definition "\ = (\x. \)" - -lemma join_apply[simp]: "(f \ g) x = f x \ g x" -by (simp add: join_fun_def) - -instance -proof - case goal2 thus ?case by (metis le_fun_def preord_class.le_trans) -qed (simp_all add: le_fun_def Top_fun_def) - -end - - -instantiation acom :: (preord) preord -begin - -fun le_acom :: "('a::preord)acom \ 'a acom \ bool" where -"le_acom (SKIP {S}) (SKIP {S'}) = (S \ S')" | -"le_acom (x ::= e {S}) (x' ::= e' {S'}) = (x=x' \ e=e' \ S \ S')" | -"le_acom (c1;c2) (c1';c2') = (le_acom c1 c1' \ le_acom c2 c2')" | -"le_acom (IF b THEN c1 ELSE c2 {S}) (IF b' THEN c1' ELSE c2' {S'}) = - (b=b' \ le_acom c1 c1' \ le_acom c2 c2' \ S \ S')" | -"le_acom ({Inv} WHILE b DO c {P}) ({Inv'} WHILE b' DO c' {P'}) = - (b=b' \ le_acom c c' \ Inv \ Inv' \ P \ P')" | -"le_acom _ _ = False" - -lemma [simp]: "SKIP {S} \ c \ (\S'. c = SKIP {S'} \ S \ S')" -by (cases c) auto - -lemma [simp]: "x ::= e {S} \ c \ (\S'. c = x ::= e {S'} \ S \ S')" -by (cases c) auto - -lemma [simp]: "c1;c2 \ c \ (\c1' c2'. c = c1';c2' \ c1 \ c1' \ c2 \ c2')" -by (cases c) auto - -lemma [simp]: "IF b THEN c1 ELSE c2 {S} \ c \ - (\c1' c2' S'. c = IF b THEN c1' ELSE c2' {S'} \ c1 \ c1' \ c2 \ c2' \ S \ S')" -by (cases c) auto - -lemma [simp]: "{Inv} WHILE b DO c {P} \ w \ - (\Inv' c' P'. w = {Inv'} WHILE b DO c' {P'} \ c \ c' \ Inv \ Inv' \ P \ P')" -by (cases w) auto - -instance -proof - case goal1 thus ?case by (induct x) auto -next - case goal2 thus ?case - apply(induct x y arbitrary: z rule: le_acom.induct) - apply (auto intro: le_trans) - done -qed - -end - - -subsubsection "Lifting" - -instantiation option :: (preord)preord -begin - -fun le_option where -"Some x \ Some y = (x \ y)" | -"None \ y = True" | -"Some _ \ None = False" - -lemma [simp]: "(x \ None) = (x = None)" -by (cases x) simp_all - -lemma [simp]: "(Some x \ u) = (\y. u = Some y & x \ y)" -by (cases u) auto - -instance proof - case goal1 show ?case by(cases x, simp_all) -next - case goal2 thus ?case - by(cases z, simp, cases y, simp, cases x, auto intro: le_trans) -qed - -end - -instantiation option :: (SL_top)SL_top -begin - -fun join_option where -"Some x \ Some y = Some(x \ y)" | -"None \ y = y" | -"x \ None = x" - -lemma join_None2[simp]: "x \ None = x" -by (cases x) simp_all - -definition "\ = Some \" - -instance proof - case goal1 thus ?case by(cases x, simp, cases y, simp_all) -next - case goal2 thus ?case by(cases y, simp, cases x, simp_all) -next - case goal3 thus ?case by(cases z, simp, cases y, simp, cases x, simp_all) -next - case goal4 thus ?case by(cases x, simp_all add: Top_option_def) -qed - -end - -definition bot_acom :: "com \ ('a::SL_top)option acom" ("\\<^sub>c") where -"\\<^sub>c = anno None" - -lemma strip_bot_acom[simp]: "strip(\\<^sub>c c) = c" -by(simp add: bot_acom_def) - -lemma bot_acom[rule_format]: "strip c' = c \ \\<^sub>c c \ c'" -apply(induct c arbitrary: c') -apply (simp_all add: bot_acom_def) - apply(induct_tac c') - apply simp_all - apply(induct_tac c') - apply simp_all - apply(induct_tac c') - apply simp_all - apply(induct_tac c') - apply simp_all - apply(induct_tac c') - apply simp_all -done - - -subsubsection "Post-fixed point iteration" - -definition - pfp :: "(('a::preord) \ 'a) \ 'a \ 'a option" where -"pfp f = while_option (\x. \ f x \ x) f" - -lemma pfp_pfp: assumes "pfp f x0 = Some x" shows "f x \ x" -using while_option_stop[OF assms[simplified pfp_def]] by simp - -lemma pfp_least: -assumes mono: "\x y. x \ y \ f x \ f y" -and "f p \ p" and "x0 \ p" and "pfp f x0 = Some x" shows "x \ p" -proof- - { fix x assume "x \ p" - hence "f x \ f p" by(rule mono) - from this `f p \ p` have "f x \ p" by(rule le_trans) - } - thus "x \ p" using assms(2-) while_option_rule[where P = "%x. x \ p"] - unfolding pfp_def by blast -qed - -definition - lpfp\<^isub>c :: "(('a::SL_top)option acom \ 'a option acom) \ com \ 'a option acom option" where -"lpfp\<^isub>c f c = pfp f (\\<^sub>c c)" - -lemma lpfpc_pfp: "lpfp\<^isub>c f c0 = Some c \ f c \ c" -by(simp add: pfp_pfp lpfp\<^isub>c_def) - -lemma strip_pfp: -assumes "\x. g(f x) = g x" and "pfp f x0 = Some x" shows "g x = g x0" -using assms while_option_rule[where P = "%x. g x = g x0" and c = f] -unfolding pfp_def by metis - -lemma strip_lpfpc: assumes "\c. strip(f c) = strip c" and "lpfp\<^isub>c f c = Some c'" -shows "strip c' = c" -using assms(1) strip_pfp[OF _ assms(2)[simplified lpfp\<^isub>c_def]] -by(metis strip_bot_acom) - -lemma lpfpc_least: -assumes mono: "\x y. x \ y \ f x \ f y" -and "strip p = c0" and "f p \ p" and lp: "lpfp\<^isub>c f c0 = Some c" shows "c \ p" -using pfp_least[OF _ _ bot_acom[OF `strip p = c0`] lp[simplified lpfp\<^isub>c_def]] - mono `f p \ p` -by blast - - -subsection "Abstract Interpretation" - -definition \_fun :: "('a \ 'b set) \ ('c \ 'a) \ ('c \ 'b)set" where -"\_fun \ F = {f. \x. f x \ \(F x)}" - -fun \_option :: "('a \ 'b set) \ 'a option \ 'b set" where -"\_option \ None = {}" | -"\_option \ (Some a) = \ a" - -text{* The interface for abstract values: *} - -locale Val_abs = -fixes \ :: "'av::SL_top \ val set" - assumes mono_gamma: "a \ b \ \ a \ \ b" - and gamma_Top[simp]: "\ \ = UNIV" -fixes num' :: "val \ 'av" -and plus' :: "'av \ 'av \ 'av" - assumes gamma_num': "n : \(num' n)" - and gamma_plus': - "n1 : \ a1 \ n2 : \ a2 \ n1+n2 : \(plus' a1 a2)" - -type_synonym 'av st = "(vname \ 'av)" - -locale Abs_Int_Fun = Val_abs \ for \ :: "'av::SL_top \ val set" -begin - -fun aval' :: "aexp \ 'av st \ 'av" where -"aval' (N n) S = num' n" | -"aval' (V x) S = S x" | -"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" - -fun step' :: "'av st option \ 'av st option acom \ 'av st option acom" - where -"step' S (SKIP {P}) = (SKIP {S})" | -"step' S (x ::= e {P}) = - x ::= e {case S of None \ None | Some S \ Some(S(x := aval' e S))}" | -"step' S (c1; c2) = step' S c1; step' (post c1) c2" | -"step' S (IF b THEN c1 ELSE c2 {P}) = - IF b THEN step' S c1 ELSE step' S c2 {post c1 \ post c2}" | -"step' S ({Inv} WHILE b DO c {P}) = - {S \ post c} WHILE b DO (step' Inv c) {Inv}" - -definition AI :: "com \ 'av st option acom option" where -"AI = lpfp\<^isub>c (step' \)" - - -lemma strip_step'[simp]: "strip(step' S c) = strip c" -by(induct c arbitrary: S) (simp_all add: Let_def) - - -abbreviation \\<^isub>f :: "'av st \ state set" -where "\\<^isub>f == \_fun \" - -abbreviation \\<^isub>o :: "'av st option \ state set" -where "\\<^isub>o == \_option \\<^isub>f" - -abbreviation \\<^isub>c :: "'av st option acom \ state set acom" -where "\\<^isub>c == map_acom \\<^isub>o" - -lemma gamma_f_Top[simp]: "\\<^isub>f Top = UNIV" -by(simp add: Top_fun_def \_fun_def) - -lemma gamma_o_Top[simp]: "\\<^isub>o Top = UNIV" -by (simp add: Top_option_def) - -(* FIXME (maybe also le \ sqle?) *) - -lemma mono_gamma_f: "f \ g \ \\<^isub>f f \ \\<^isub>f g" -by(auto simp: le_fun_def \_fun_def dest: mono_gamma) - -lemma mono_gamma_o: - "sa \ sa' \ \\<^isub>o sa \ \\<^isub>o sa'" -by(induction sa sa' rule: le_option.induct)(simp_all add: mono_gamma_f) - -lemma mono_gamma_c: "ca \ ca' \ \\<^isub>c ca \ \\<^isub>c ca'" -by (induction ca ca' rule: le_acom.induct) (simp_all add:mono_gamma_o) - -text{* Soundness: *} - -lemma aval'_sound: "s : \\<^isub>f S \ aval a s : \(aval' a S)" -by (induct a) (auto simp: gamma_num' gamma_plus' \_fun_def) - -lemma in_gamma_update: - "\ s : \\<^isub>f S; i : \ a \ \ s(x := i) : \\<^isub>f(S(x := a))" -by(simp add: \_fun_def) - -lemma step_preserves_le: - "\ S \ \\<^isub>o S'; c \ \\<^isub>c c' \ \ step S c \ \\<^isub>c (step' S' c')" -proof(induction c arbitrary: c' S S') - case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) -next - case Assign thus ?case - by (fastforce simp: Assign_le map_acom_Assign intro: aval'_sound in_gamma_update - split: option.splits del:subsetD) -next - case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi) - by (metis le_post post_map_acom) -next - case (If b c1 c2 P) - then obtain c1' c2' P' where - "c' = IF b THEN c1' ELSE c2' {P'}" - "P \ \\<^isub>o P'" "c1 \ \\<^isub>c c1'" "c2 \ \\<^isub>c c2'" - by (fastforce simp: If_le map_acom_If) - moreover have "post c1 \ \\<^isub>o(post c1' \ post c2')" - by (metis (no_types) `c1 \ \\<^isub>c c1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom) - moreover have "post c2 \ \\<^isub>o(post c1' \ post c2')" - by (metis (no_types) `c2 \ \\<^isub>c c2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom) - ultimately show ?case using `S \ \\<^isub>o S'` by (simp add: If.IH subset_iff) -next - case (While I b c1 P) - then obtain c1' I' P' where - "c' = {I'} WHILE b DO c1' {P'}" - "I \ \\<^isub>o I'" "P \ \\<^isub>o P'" "c1 \ \\<^isub>c c1'" - by (fastforce simp: map_acom_While While_le) - moreover have "S \ post c1 \ \\<^isub>o (S' \ post c1')" - using `S \ \\<^isub>o S'` le_post[OF `c1 \ \\<^isub>c c1'`, simplified] - by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans) - ultimately show ?case by (simp add: While.IH subset_iff) -qed - -lemma AI_sound: "AI c = Some c' \ CS c \ \\<^isub>c c'" -proof(simp add: CS_def AI_def) - assume 1: "lpfp\<^isub>c (step' \) c = Some c'" - have 2: "step' \ c' \ c'" by(rule lpfpc_pfp[OF 1]) - have 3: "strip (\\<^isub>c (step' \ c')) = c" - by(simp add: strip_lpfpc[OF _ 1]) - have "lfp (step UNIV) c \ \\<^isub>c (step' \ c')" - proof(rule lfp_lowerbound[simplified,OF 3]) - show "step UNIV (\\<^isub>c (step' \ c')) \ \\<^isub>c (step' \ c')" - proof(rule step_preserves_le[OF _ _]) - show "UNIV \ \\<^isub>o \" by simp - show "\\<^isub>c (step' \ c') \ \\<^isub>c c'" by(rule mono_gamma_c[OF 2]) - qed - qed - with 2 show "lfp (step UNIV) c \ \\<^isub>c c'" - by (blast intro: mono_gamma_c order_trans) -qed - -end - - -subsubsection "Monotonicity" - -lemma mono_post: "c \ c' \ post c \ post c'" -by(induction c c' rule: le_acom.induct) (auto) - -locale Abs_Int_Fun_mono = Abs_Int_Fun + -assumes mono_plus': "a1 \ b1 \ a2 \ b2 \ plus' a1 a2 \ plus' b1 b2" -begin - -lemma mono_aval': "S \ S' \ aval' e S \ aval' e S'" -by(induction e)(auto simp: le_fun_def mono_plus') - -lemma mono_update: "a \ a' \ S \ S' \ S(x := a) \ S'(x := a')" -by(simp add: le_fun_def) - -lemma mono_step': "S \ S' \ c \ c' \ step' S c \ step' S' c'" -apply(induction c c' arbitrary: S S' rule: le_acom.induct) -apply (auto simp: Let_def mono_update mono_aval' mono_post le_join_disj - split: option.split) -done - -end - -text{* Problem: not executable because of the comparison of abstract states, -i.e. functions, in the post-fixedpoint computation. *} - -end diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_Int0_parity.thy --- a/src/HOL/IMP/Abs_Int0_parity.thy Thu Apr 19 11:14:57 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,166 +0,0 @@ -theory Abs_Int0_parity -imports Abs_Int0 -begin - -subsection "Parity Analysis" - -datatype parity = Even | Odd | Either - -text{* Instantiation of class @{class preord} with type @{typ parity}: *} - -instantiation parity :: preord -begin - -text{* First the definition of the interface function @{text"\"}. Note that -the header of the definition must refer to the ascii name @{const le} of the -constants as @{text le_parity} and the definition is named @{text -le_parity_def}. Inside the definition the symbolic names can be used. *} - -definition le_parity where -"x \ y = (y = Either \ x=y)" - -text{* Now the instance proof, i.e.\ the proof that the definition fulfills -the axioms (assumptions) of the class. The initial proof-step generates the -necessary proof obligations. *} - -instance -proof - fix x::parity show "x \ x" by(auto simp: le_parity_def) -next - fix x y z :: parity assume "x \ y" "y \ z" thus "x \ z" - by(auto simp: le_parity_def) -qed - -end - -text{* Instantiation of class @{class SL_top} with type @{typ parity}: *} - -instantiation parity :: SL_top -begin - - -definition join_parity where -"x \ y = (if x \ y then y else if y \ x then x else Either)" - -definition Top_parity where -"\ = Either" - -text{* Now the instance proof. This time we take a lazy shortcut: we do not -write out the proof obligations but use the @{text goali} primitive to refer -to the assumptions of subgoal i and @{text "case?"} to refer to the -conclusion of subgoal i. The class axioms are presented in the same order as -in the class definition. *} - -instance -proof - case goal1 (*join1*) show ?case by(auto simp: le_parity_def join_parity_def) -next - case goal2 (*join2*) show ?case by(auto simp: le_parity_def join_parity_def) -next - case goal3 (*join least*) thus ?case by(auto simp: le_parity_def join_parity_def) -next - case goal4 (*Top*) show ?case by(auto simp: le_parity_def Top_parity_def) -qed - -end - - -text{* Now we define the functions used for instantiating the abstract -interpretation locales. Note that the Isabelle terminology is -\emph{interpretation}, not \emph{instantiation} of locales, but we use -instantiation to avoid confusion with abstract interpretation. *} - -fun \_parity :: "parity \ val set" where -"\_parity Even = {i. i mod 2 = 0}" | -"\_parity Odd = {i. i mod 2 = 1}" | -"\_parity Either = UNIV" - -fun num_parity :: "val \ parity" where -"num_parity i = (if i mod 2 = 0 then Even else Odd)" - -fun plus_parity :: "parity \ parity \ parity" where -"plus_parity Even Even = Even" | -"plus_parity Odd Odd = Even" | -"plus_parity Even Odd = Odd" | -"plus_parity Odd Even = Odd" | -"plus_parity Either y = Either" | -"plus_parity x Either = Either" - -text{* First we instantiate the abstract value interface and prove that the -functions on type @{typ parity} have all the necessary properties: *} - -interpretation Val_abs -where \ = \_parity and num' = num_parity and plus' = plus_parity -proof txt{* of the locale axioms *} - fix a b :: parity - assume "a \ b" thus "\_parity a \ \_parity b" - by(auto simp: le_parity_def) -next txt{* The rest in the lazy, implicit way *} - case goal2 show ?case by(auto simp: Top_parity_def) -next - case goal3 show ?case by auto -next - txt{* Warning: this subproof refers to the names @{text a1} and @{text a2} - from the statement of the axiom. *} - case goal4 thus ?case - proof(cases a1 a2 rule: parity.exhaust[case_product parity.exhaust]) - qed (auto simp add:mod_add_eq) -qed - -text{* Instantiating the abstract interpretation locale requires no more -proofs (they happened in the instatiation above) but delivers the -instantiated abstract interpreter which we call AI: *} - -interpretation Abs_Int -where \ = \_parity and num' = num_parity and plus' = plus_parity -defines aval_parity is aval' and step_parity is step' and AI_parity is AI -.. - - -subsubsection "Tests" - -definition "test1_parity = - ''x'' ::= N 1; - WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 2)" - -value "show_acom_opt (AI_parity test1_parity)" - -definition "test2_parity = - ''x'' ::= N 1; - WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 3)" - -value "show_acom ((step_parity \ ^^1) (anno None test2_parity))" -value "show_acom ((step_parity \ ^^2) (anno None test2_parity))" -value "show_acom ((step_parity \ ^^3) (anno None test2_parity))" -value "show_acom ((step_parity \ ^^4) (anno None test2_parity))" -value "show_acom ((step_parity \ ^^5) (anno None test2_parity))" -value "show_acom_opt (AI_parity test2_parity)" - - -subsubsection "Termination" - -interpretation Abs_Int_mono -where \ = \_parity and num' = num_parity and plus' = plus_parity -proof - case goal1 thus ?case - proof(cases a1 a2 b1 b2 - rule: parity.exhaust[case_product parity.exhaust[case_product parity.exhaust[case_product parity.exhaust]]]) (* FIXME - UGLY! *) - qed (auto simp add:le_parity_def) -qed - - -definition m_parity :: "parity \ nat" where -"m_parity x = (if x=Either then 0 else 1)" - -lemma measure_parity: - "(strict{(x::parity,y). x \ y})^-1 \ measure m_parity" -by(auto simp add: m_parity_def le_parity_def) - -lemma measure_parity_eq: - "\x y::parity. x \ y \ y \ x \ m_parity x = m_parity y" -by(auto simp add: m_parity_def le_parity_def) - -lemma AI_parity_Some: "\c'. AI_parity c = Some c'" -by(rule AI_Some_measure[OF measure_parity measure_parity_eq]) - -end diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_Int1.thy --- a/src/HOL/IMP/Abs_Int1.thy Thu Apr 19 11:14:57 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,358 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory Abs_Int1 -imports Abs_Int0 Vars -begin - -instantiation prod :: (preord,preord) preord -begin - -definition "le_prod p1 p2 = (fst p1 \ fst p2 \ snd p1 \ snd p2)" - -instance -proof - case goal1 show ?case by(simp add: le_prod_def) -next - case goal2 thus ?case unfolding le_prod_def by(metis le_trans) -qed - -end - -instantiation com :: vars -begin - -fun vars_com :: "com \ vname set" where -"vars com.SKIP = {}" | -"vars (x::=e) = {x} \ vars e" | -"vars (c1;c2) = vars c1 \ vars c2" | -"vars (IF b THEN c1 ELSE c2) = vars b \ vars c1 \ vars c2" | -"vars (WHILE b DO c) = vars b \ vars c" - -instance .. - -end - -lemma finite_avars: "finite(vars(a::aexp))" -by(induction a) simp_all - -lemma finite_bvars: "finite(vars(b::bexp))" -by(induction b) (simp_all add: finite_avars) - -lemma finite_cvars: "finite(vars(c::com))" -by(induction c) (simp_all add: finite_avars finite_bvars) - - -subsection "Backward Analysis of Expressions" - -hide_const bot - -class L_top_bot = SL_top + -fixes meet :: "'a \ 'a \ 'a" (infixl "\" 65) -and bot :: "'a" ("\") -assumes meet_le1 [simp]: "x \ y \ x" -and meet_le2 [simp]: "x \ y \ y" -and meet_greatest: "x \ y \ x \ z \ x \ y \ z" -assumes bot[simp]: "\ \ x" -begin - -lemma mono_meet: "x \ x' \ y \ y' \ x \ y \ x' \ y'" -by (metis meet_greatest meet_le1 meet_le2 le_trans) - -end - -locale Val_abs1_gamma = - Gamma where \ = \ for \ :: "'av::L_top_bot \ val set" + -assumes inter_gamma_subset_gamma_meet: - "\ a1 \ \ a2 \ \(a1 \ a2)" -and gamma_Bot[simp]: "\ \ = {}" -begin - -lemma in_gamma_meet: "x : \ a1 \ x : \ a2 \ x : \(a1 \ a2)" -by (metis IntI inter_gamma_subset_gamma_meet set_mp) - -lemma gamma_meet[simp]: "\(a1 \ a2) = \ a1 \ \ a2" -by (metis equalityI inter_gamma_subset_gamma_meet le_inf_iff mono_gamma meet_le1 meet_le2) - -end - - -locale Val_abs1 = - Val_abs1_gamma where \ = \ - for \ :: "'av::L_top_bot \ val set" + -fixes test_num' :: "val \ 'av \ bool" -and filter_plus' :: "'av \ 'av \ 'av \ 'av * 'av" -and filter_less' :: "bool \ 'av \ 'av \ 'av * 'av" -assumes test_num': "test_num' n a = (n : \ a)" -and filter_plus': "filter_plus' a a1 a2 = (b1,b2) \ - n1 : \ a1 \ n2 : \ a2 \ n1+n2 : \ a \ n1 : \ b1 \ n2 : \ b2" -and filter_less': "filter_less' (n1 - n1 : \ a1 \ n2 : \ a2 \ n1 : \ b1 \ n2 : \ b2" - - -locale Abs_Int1 = - Val_abs1 where \ = \ for \ :: "'av::L_top_bot \ val set" -begin - -lemma in_gamma_join_UpI: "s : \\<^isub>o S1 \ s : \\<^isub>o S2 \ s : \\<^isub>o(S1 \ S2)" -by (metis (no_types) join_ge1 join_ge2 mono_gamma_o set_rev_mp) - -fun aval'' :: "aexp \ 'av st option \ 'av" where -"aval'' e None = \" | -"aval'' e (Some sa) = aval' e sa" - -lemma aval''_sound: "s : \\<^isub>o S \ aval a s : \(aval'' a S)" -by(cases S)(simp add: aval'_sound)+ - -subsubsection "Backward analysis" - -fun afilter :: "aexp \ 'av \ 'av st option \ 'av st option" where -"afilter (N n) a S = (if test_num' n a then S else None)" | -"afilter (V x) a S = (case S of None \ None | Some S \ - let a' = lookup S x \ a in - if a' \ \ then None else Some(update S x a'))" | -"afilter (Plus e1 e2) a S = - (let (a1,a2) = filter_plus' a (aval'' e1 S) (aval'' e2 S) - in afilter e1 a1 (afilter e2 a2 S))" - -text{* The test for @{const bot} in the @{const V}-case is important: @{const -bot} indicates that a variable has no possible values, i.e.\ that the current -program point is unreachable. But then the abstract state should collapse to -@{const None}. Put differently, we maintain the invariant that in an abstract -state of the form @{term"Some s"}, all variables are mapped to non-@{const -bot} values. Otherwise the (pointwise) join of two abstract states, one of -which contains @{const bot} values, may produce too large a result, thus -making the analysis less precise. *} - - -fun bfilter :: "bexp \ bool \ 'av st option \ 'av st option" where -"bfilter (Bc v) res S = (if v=res then S else None)" | -"bfilter (Not b) res S = bfilter b (\ res) S" | -"bfilter (And b1 b2) res S = - (if res then bfilter b1 True (bfilter b2 True S) - else bfilter b1 False S \ bfilter b2 False S)" | -"bfilter (Less e1 e2) res S = - (let (res1,res2) = filter_less' res (aval'' e1 S) (aval'' e2 S) - in afilter e1 res1 (afilter e2 res2 S))" - -lemma afilter_sound: "s : \\<^isub>o S \ aval e s : \ a \ s : \\<^isub>o (afilter e a S)" -proof(induction e arbitrary: a S) - case N thus ?case by simp (metis test_num') -next - case (V x) - obtain S' where "S = Some S'" and "s : \\<^isub>f S'" using `s : \\<^isub>o S` - by(auto simp: in_gamma_option_iff) - moreover hence "s x : \ (lookup S' x)" by(simp add: \_st_def) - moreover have "s x : \ a" using V by simp - ultimately show ?case using V(1) - by(simp add: lookup_update Let_def \_st_def) - (metis mono_gamma emptyE in_gamma_meet gamma_Bot subset_empty) -next - case (Plus e1 e2) thus ?case - using filter_plus'[OF _ aval''_sound[OF Plus(3)] aval''_sound[OF Plus(3)]] - by (auto split: prod.split) -qed - -lemma bfilter_sound: "s : \\<^isub>o S \ bv = bval b s \ s : \\<^isub>o(bfilter b bv S)" -proof(induction b arbitrary: S bv) - case Bc thus ?case by simp -next - case (Not b) thus ?case by simp -next - case (And b1 b2) thus ?case by(fastforce simp: in_gamma_join_UpI) -next - case (Less e1 e2) thus ?case - by (auto split: prod.split) - (metis afilter_sound filter_less' aval''_sound Less) -qed - - -fun step' :: "'av st option \ 'av st option acom \ 'av st option acom" - where -"step' S (SKIP {P}) = (SKIP {S})" | -"step' S (x ::= e {P}) = - x ::= e {case S of None \ None | Some S \ Some(update S x (aval' e S))}" | -"step' S (c1; c2) = step' S c1; step' (post c1) c2" | -"step' S (IF b THEN c1 ELSE c2 {P}) = - (let c1' = step' (bfilter b True S) c1; c2' = step' (bfilter b False S) c2 - in IF b THEN c1' ELSE c2' {post c1 \ post c2})" | -"step' S ({Inv} WHILE b DO c {P}) = - {S \ post c} - WHILE b DO step' (bfilter b True Inv) c - {bfilter b False Inv}" - -definition AI :: "com \ 'av st option acom option" where -"AI = lpfp\<^isub>c (step' \)" - -lemma strip_step'[simp]: "strip(step' S c) = strip c" -by(induct c arbitrary: S) (simp_all add: Let_def) - - -subsubsection "Soundness" - -lemma in_gamma_update: - "\ s : \\<^isub>f S; i : \ a \ \ s(x := i) : \\<^isub>f(update S x a)" -by(simp add: \_st_def lookup_update) - -lemma step_preserves_le: - "\ S \ \\<^isub>o S'; cs \ \\<^isub>c ca \ \ step S cs \ \\<^isub>c (step' S' ca)" -proof(induction cs arbitrary: ca S S') - case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) -next - case Assign thus ?case - by (fastforce simp: Assign_le map_acom_Assign intro: aval'_sound in_gamma_update - split: option.splits del:subsetD) -next - case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi) - by (metis le_post post_map_acom) -next - case (If b cs1 cs2 P) - then obtain ca1 ca2 Pa where - "ca= IF b THEN ca1 ELSE ca2 {Pa}" - "P \ \\<^isub>o Pa" "cs1 \ \\<^isub>c ca1" "cs2 \ \\<^isub>c ca2" - by (fastforce simp: If_le map_acom_If) - moreover have "post cs1 \ \\<^isub>o(post ca1 \ post ca2)" - by (metis (no_types) `cs1 \ \\<^isub>c ca1` join_ge1 le_post mono_gamma_o order_trans post_map_acom) - moreover have "post cs2 \ \\<^isub>o(post ca1 \ post ca2)" - by (metis (no_types) `cs2 \ \\<^isub>c ca2` join_ge2 le_post mono_gamma_o order_trans post_map_acom) - ultimately show ?case using `S \ \\<^isub>o S'` - by (simp add: If.IH subset_iff bfilter_sound) -next - case (While I b cs1 P) - then obtain ca1 Ia Pa where - "ca = {Ia} WHILE b DO ca1 {Pa}" - "I \ \\<^isub>o Ia" "P \ \\<^isub>o Pa" "cs1 \ \\<^isub>c ca1" - by (fastforce simp: map_acom_While While_le) - moreover have "S \ post cs1 \ \\<^isub>o (S' \ post ca1)" - using `S \ \\<^isub>o S'` le_post[OF `cs1 \ \\<^isub>c ca1`, simplified] - by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans) - ultimately show ?case by (simp add: While.IH subset_iff bfilter_sound) -qed - -lemma AI_sound: "AI c = Some c' \ CS c \ \\<^isub>c c'" -proof(simp add: CS_def AI_def) - assume 1: "lpfp\<^isub>c (step' \) c = Some c'" - have 2: "step' \ c' \ c'" by(rule lpfpc_pfp[OF 1]) - have 3: "strip (\\<^isub>c (step' \ c')) = c" - by(simp add: strip_lpfpc[OF _ 1]) - have "lfp (step UNIV) c \ \\<^isub>c (step' \ c')" - proof(rule lfp_lowerbound[simplified,OF 3]) - show "step UNIV (\\<^isub>c (step' \ c')) \ \\<^isub>c (step' \ c')" - proof(rule step_preserves_le[OF _ _]) - show "UNIV \ \\<^isub>o \" by simp - show "\\<^isub>c (step' \ c') \ \\<^isub>c c'" by(rule mono_gamma_c[OF 2]) - qed - qed - from this 2 show "lfp (step UNIV) c \ \\<^isub>c c'" - by (blast intro: mono_gamma_c order_trans) -qed - - -subsubsection "Commands over a set of variables" - -text{* Key invariant: the domains of all abstract states are subsets of the -set of variables of the program. *} - -definition "domo S = (case S of None \ {} | Some S' \ set(dom S'))" - -definition Com :: "vname set \ 'a st option acom set" where -"Com X = {c. \S \ set(annos c). domo S \ X}" - -lemma domo_Top[simp]: "domo \ = {}" -by(simp add: domo_def Top_st_def Top_option_def) - -lemma bot_acom_Com[simp]: "\\<^sub>c c \ Com X" -by(simp add: bot_acom_def Com_def domo_def set_annos_anno) - -lemma post_in_annos: "post c : set(annos c)" -by(induction c) simp_all - -lemma domo_join: "domo (S \ T) \ domo S \ domo T" -by(auto simp: domo_def join_st_def split: option.split) - -lemma domo_afilter: "vars a \ X \ domo S \ X \ domo(afilter a i S) \ X" -apply(induction a arbitrary: i S) -apply(simp add: domo_def) -apply(simp add: domo_def Let_def update_def lookup_def split: option.splits) -apply blast -apply(simp split: prod.split) -done - -lemma domo_bfilter: "vars b \ X \ domo S \ X \ domo(bfilter b bv S) \ X" -apply(induction b arbitrary: bv S) -apply(simp add: domo_def) -apply(simp) -apply(simp) -apply rule -apply (metis le_sup_iff subset_trans[OF domo_join]) -apply(simp split: prod.split) -by (metis domo_afilter) - -lemma step'_Com: - "domo S \ X \ vars(strip c) \ X \ c : Com X \ step' S c : Com X" -apply(induction c arbitrary: S) -apply(simp add: Com_def) -apply(simp add: Com_def domo_def update_def split: option.splits) -apply(simp (no_asm_use) add: Com_def ball_Un) -apply (metis post_in_annos) -apply(simp (no_asm_use) add: Com_def ball_Un) -apply rule -apply (metis Un_assoc domo_join order_trans post_in_annos subset_Un_eq) -apply (metis domo_bfilter) -apply(simp (no_asm_use) add: Com_def) -apply rule -apply (metis domo_join le_sup_iff post_in_annos subset_trans) -apply rule -apply (metis domo_bfilter) -by (metis domo_bfilter) - -end - - -subsubsection "Monotonicity" - -locale Abs_Int1_mono = Abs_Int1 + -assumes mono_plus': "a1 \ b1 \ a2 \ b2 \ plus' a1 a2 \ plus' b1 b2" -and mono_filter_plus': "a1 \ b1 \ a2 \ b2 \ r \ r' \ - filter_plus' r a1 a2 \ filter_plus' r' b1 b2" -and mono_filter_less': "a1 \ b1 \ a2 \ b2 \ - filter_less' bv a1 a2 \ filter_less' bv b1 b2" -begin - -lemma mono_aval': "S \ S' \ aval' e S \ aval' e S'" -by(induction e) (auto simp: le_st_def lookup_def mono_plus') - -lemma mono_aval'': "S \ S' \ aval'' e S \ aval'' e S'" -apply(cases S) - apply simp -apply(cases S') - apply simp -by (simp add: mono_aval') - -lemma mono_afilter: "r \ r' \ S \ S' \ afilter e r S \ afilter e r' S'" -apply(induction e arbitrary: r r' S S') -apply(auto simp: test_num' Let_def split: option.splits prod.splits) -apply (metis mono_gamma subsetD) -apply(drule_tac x = "list" in mono_lookup) -apply (metis mono_meet le_trans) -apply (metis mono_meet mono_lookup mono_update) -apply(metis mono_aval'' mono_filter_plus'[simplified le_prod_def] fst_conv snd_conv) -done - -lemma mono_bfilter: "S \ S' \ bfilter b r S \ bfilter b r S'" -apply(induction b arbitrary: r S S') -apply(auto simp: le_trans[OF _ join_ge1] le_trans[OF _ join_ge2] split: prod.splits) -apply(metis mono_aval'' mono_afilter mono_filter_less'[simplified le_prod_def] fst_conv snd_conv) -done - -lemma mono_step': "S \ S' \ c \ c' \ step' S c \ step' S' c'" -apply(induction c c' arbitrary: S S' rule: le_acom.induct) -apply (auto simp: mono_post mono_bfilter mono_update mono_aval' Let_def le_join_disj - split: option.split) -done - -lemma mono_step'2: "mono (step' S)" -by(simp add: mono_def mono_step'[OF le_refl]) - -end - -end diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_Int1_ivl.thy --- a/src/HOL/IMP/Abs_Int1_ivl.thy Thu Apr 19 11:14:57 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,285 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory Abs_Int1_ivl -imports Abs_Int1 Abs_Int_Tests -begin - -subsection "Interval Analysis" - -datatype ivl = I "int option" "int option" - -definition "\_ivl i = (case i of - I (Some l) (Some h) \ {l..h} | - I (Some l) None \ {l..} | - I None (Some h) \ {..h} | - I None None \ UNIV)" - -abbreviation I_Some_Some :: "int \ int \ ivl" ("{_\_}") where -"{lo\hi} == I (Some lo) (Some hi)" -abbreviation I_Some_None :: "int \ ivl" ("{_\}") where -"{lo\} == I (Some lo) None" -abbreviation I_None_Some :: "int \ ivl" ("{\_}") where -"{\hi} == I None (Some hi)" -abbreviation I_None_None :: "ivl" ("{\}") where -"{\} == I None None" - -definition "num_ivl n = {n\n}" - -fun in_ivl :: "int \ ivl \ bool" where -"in_ivl k (I (Some l) (Some h)) \ l \ k \ k \ h" | -"in_ivl k (I (Some l) None) \ l \ k" | -"in_ivl k (I None (Some h)) \ k \ h" | -"in_ivl k (I None None) \ True" - -instantiation option :: (plus)plus -begin - -fun plus_option where -"Some x + Some y = Some(x+y)" | -"_ + _ = None" - -instance .. - -end - -definition empty where "empty = {1\0}" - -fun is_empty where -"is_empty {l\h} = (h (case h of Some h \ h False) | None \ False)" -by(auto split:option.split) - -lemma [simp]: "is_empty i \ \_ivl i = {}" -by(auto simp add: \_ivl_def split: ivl.split option.split) - -definition "plus_ivl i1 i2 = (if is_empty i1 | is_empty i2 then empty else - case (i1,i2) of (I l1 h1, I l2 h2) \ I (l1+l2) (h1+h2))" - -instantiation ivl :: SL_top -begin - -definition le_option :: "bool \ int option \ int option \ bool" where -"le_option pos x y = - (case x of (Some i) \ (case y of Some j \ i\j | None \ pos) - | None \ (case y of Some j \ \pos | None \ True))" - -fun le_aux where -"le_aux (I l1 h1) (I l2 h2) = (le_option False l2 l1 & le_option True h1 h2)" - -definition le_ivl where -"i1 \ i2 = - (if is_empty i1 then True else - if is_empty i2 then False else le_aux i1 i2)" - -definition min_option :: "bool \ int option \ int option \ int option" where -"min_option pos o1 o2 = (if le_option pos o1 o2 then o1 else o2)" - -definition max_option :: "bool \ int option \ int option \ int option" where -"max_option pos o1 o2 = (if le_option pos o1 o2 then o2 else o1)" - -definition "i1 \ i2 = - (if is_empty i1 then i2 else if is_empty i2 then i1 - else case (i1,i2) of (I l1 h1, I l2 h2) \ - I (min_option False l1 l2) (max_option True h1 h2))" - -definition "\ = {\}" - -instance -proof - case goal1 thus ?case - by(cases x, simp add: le_ivl_def le_option_def split: option.split) -next - case goal2 thus ?case - by(cases x, cases y, cases z, auto simp: le_ivl_def le_option_def split: option.splits if_splits) -next - case goal3 thus ?case - by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits) -next - case goal4 thus ?case - by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits) -next - case goal5 thus ?case - by(cases x, cases y, cases z, auto simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits if_splits) -next - case goal6 thus ?case - by(cases x, simp add: Top_ivl_def le_ivl_def le_option_def split: option.split) -qed - -end - - -instantiation ivl :: L_top_bot -begin - -definition "i1 \ i2 = (if is_empty i1 \ is_empty i2 then empty else - case (i1,i2) of (I l1 h1, I l2 h2) \ - I (max_option False l1 l2) (min_option True h1 h2))" - -definition "\ = empty" - -instance -proof - case goal1 thus ?case - by (simp add:meet_ivl_def empty_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits) -next - case goal2 thus ?case - by (simp add: empty_def meet_ivl_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits) -next - case goal3 thus ?case - by (cases x, cases y, cases z, auto simp add: le_ivl_def meet_ivl_def empty_def le_option_def max_option_def min_option_def split: option.splits if_splits) -next - case goal4 show ?case by(cases x, simp add: bot_ivl_def empty_def le_ivl_def) -qed - -end - -instantiation option :: (minus)minus -begin - -fun minus_option where -"Some x - Some y = Some(x-y)" | -"_ - _ = None" - -instance .. - -end - -definition "minus_ivl i1 i2 = (if is_empty i1 | is_empty i2 then empty else - case (i1,i2) of (I l1 h1, I l2 h2) \ I (l1-h2) (h1-l2))" - -lemma gamma_minus_ivl: - "n1 : \_ivl i1 \ n2 : \_ivl i2 \ n1-n2 : \_ivl(minus_ivl i1 i2)" -by(auto simp add: minus_ivl_def \_ivl_def split: ivl.splits option.splits) - -definition "filter_plus_ivl i i1 i2 = ((*if is_empty i then empty else*) - i1 \ minus_ivl i i2, i2 \ minus_ivl i i1)" - -fun filter_less_ivl :: "bool \ ivl \ ivl \ ivl * ivl" where -"filter_less_ivl res (I l1 h1) (I l2 h2) = - (if is_empty(I l1 h1) \ is_empty(I l2 h2) then (empty, empty) else - if res - then (I l1 (min_option True h1 (h2 - Some 1)), - I (max_option False (l1 + Some 1) l2) h2) - else (I (max_option False l1 l2) h1, I l2 (min_option True h1 h2)))" - -interpretation Val_abs -where \ = \_ivl and num' = num_ivl and plus' = plus_ivl -proof - case goal1 thus ?case - by(auto simp: \_ivl_def le_ivl_def le_option_def split: ivl.split option.split if_splits) -next - case goal2 show ?case by(simp add: \_ivl_def Top_ivl_def) -next - case goal3 thus ?case by(simp add: \_ivl_def num_ivl_def) -next - case goal4 thus ?case - by(auto simp add: \_ivl_def plus_ivl_def split: ivl.split option.splits) -qed - -interpretation Val_abs1_gamma -where \ = \_ivl and num' = num_ivl and plus' = plus_ivl -defines aval_ivl is aval' -proof - case goal1 thus ?case - by(auto simp add: \_ivl_def meet_ivl_def empty_def min_option_def max_option_def split: ivl.split option.split) -next - case goal2 show ?case by(auto simp add: bot_ivl_def \_ivl_def empty_def) -qed - -lemma mono_minus_ivl: - "i1 \ i1' \ i2 \ i2' \ minus_ivl i1 i2 \ minus_ivl i1' i2'" -apply(auto simp add: minus_ivl_def empty_def le_ivl_def le_option_def split: ivl.splits) - apply(simp split: option.splits) - apply(simp split: option.splits) -apply(simp split: option.splits) -done - - -interpretation Val_abs1 -where \ = \_ivl and num' = num_ivl and plus' = plus_ivl -and test_num' = in_ivl -and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl -proof - case goal1 thus ?case - by (simp add: \_ivl_def split: ivl.split option.split) -next - case goal2 thus ?case - by(auto simp add: filter_plus_ivl_def) - (metis gamma_minus_ivl add_diff_cancel add_commute)+ -next - case goal3 thus ?case - by(cases a1, cases a2, - auto simp: \_ivl_def min_option_def max_option_def le_option_def split: if_splits option.splits) -qed - -interpretation Abs_Int1 -where \ = \_ivl and num' = num_ivl and plus' = plus_ivl -and test_num' = in_ivl -and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl -defines afilter_ivl is afilter -and bfilter_ivl is bfilter -and step_ivl is step' -and AI_ivl is AI -and aval_ivl' is aval'' -.. - - -text{* Monotonicity: *} - -interpretation Abs_Int1_mono -where \ = \_ivl and num' = num_ivl and plus' = plus_ivl -and test_num' = in_ivl -and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl -proof - case goal1 thus ?case - by(auto simp: plus_ivl_def le_ivl_def le_option_def empty_def split: if_splits ivl.splits option.splits) -next - case goal2 thus ?case - by(auto simp: filter_plus_ivl_def le_prod_def mono_meet mono_minus_ivl) -next - case goal3 thus ?case - apply(cases a1, cases b1, cases a2, cases b2, auto simp: le_prod_def) - by(auto simp add: empty_def le_ivl_def le_option_def min_option_def max_option_def split: option.splits) -qed - - -subsubsection "Tests" - -value "show_acom_opt (AI_ivl test1_ivl)" - -text{* Better than @{text AI_const}: *} -value "show_acom_opt (AI_ivl test3_const)" -value "show_acom_opt (AI_ivl test4_const)" -value "show_acom_opt (AI_ivl test6_const)" - -value "show_acom_opt (AI_ivl test2_ivl)" -value "show_acom (((step_ivl \)^^0) (\\<^sub>c test2_ivl))" -value "show_acom (((step_ivl \)^^1) (\\<^sub>c test2_ivl))" -value "show_acom (((step_ivl \)^^2) (\\<^sub>c test2_ivl))" - -text{* Fixed point reached in 2 steps. Not so if the start value of x is known: *} - -value "show_acom_opt (AI_ivl test3_ivl)" -value "show_acom (((step_ivl \)^^0) (\\<^sub>c test3_ivl))" -value "show_acom (((step_ivl \)^^1) (\\<^sub>c test3_ivl))" -value "show_acom (((step_ivl \)^^2) (\\<^sub>c test3_ivl))" -value "show_acom (((step_ivl \)^^3) (\\<^sub>c test3_ivl))" -value "show_acom (((step_ivl \)^^4) (\\<^sub>c test3_ivl))" - -text{* Takes as many iterations as the actual execution. Would diverge if -loop did not terminate. Worse still, as the following example shows: even if -the actual execution terminates, the analysis may not. The value of y keeps -decreasing as the analysis is iterated, no matter how long: *} - -value "show_acom (((step_ivl \)^^50) (\\<^sub>c test4_ivl))" - -text{* Relationships between variables are NOT captured: *} -value "show_acom_opt (AI_ivl test5_ivl)" - -text{* Again, the analysis would not terminate: *} -value "show_acom (((step_ivl \)^^50) (\\<^sub>c test6_ivl))" - -end diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_Int2.thy --- a/src/HOL/IMP/Abs_Int2.thy Thu Apr 19 11:14:57 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,683 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory Abs_Int2 -imports Abs_Int1_ivl -begin - -subsection "Widening and Narrowing" - -class WN = SL_top + -fixes widen :: "'a \ 'a \ 'a" (infix "\" 65) -assumes widen1: "x \ x \ y" -assumes widen2: "y \ x \ y" -fixes narrow :: "'a \ 'a \ 'a" (infix "\" 65) -assumes narrow1: "y \ x \ y \ x \ y" -assumes narrow2: "y \ x \ x \ y \ x" - - -subsubsection "Intervals" - -instantiation ivl :: WN -begin - -definition "widen_ivl ivl1 ivl2 = - ((*if is_empty ivl1 then ivl2 else - if is_empty ivl2 then ivl1 else*) - case (ivl1,ivl2) of (I l1 h1, I l2 h2) \ - I (if le_option False l2 l1 \ l2 \ l1 then None else l1) - (if le_option True h1 h2 \ h1 \ h2 then None else h1))" - -definition "narrow_ivl ivl1 ivl2 = - ((*if is_empty ivl1 \ is_empty ivl2 then empty else*) - case (ivl1,ivl2) of (I l1 h1, I l2 h2) \ - I (if l1 = None then l2 else l1) - (if h1 = None then h2 else h1))" - -instance -proof qed - (auto simp add: widen_ivl_def narrow_ivl_def le_option_def le_ivl_def empty_def split: ivl.split option.split if_splits) - -end - - -subsubsection "Abstract State" - -instantiation st :: (WN)WN -begin - -definition "widen_st F1 F2 = - FunDom (\x. fun F1 x \ fun F2 x) (inter_list (dom F1) (dom F2))" - -definition "narrow_st F1 F2 = - FunDom (\x. fun F1 x \ fun F2 x) (inter_list (dom F1) (dom F2))" - -instance -proof - case goal1 thus ?case - by(simp add: widen_st_def le_st_def lookup_def widen1) -next - case goal2 thus ?case - by(simp add: widen_st_def le_st_def lookup_def widen2) -next - case goal3 thus ?case - by(auto simp: narrow_st_def le_st_def lookup_def narrow1) -next - case goal4 thus ?case - by(auto simp: narrow_st_def le_st_def lookup_def narrow2) -qed - -end - - -subsubsection "Option" - -instantiation option :: (WN)WN -begin - -fun widen_option where -"None \ x = x" | -"x \ None = x" | -"(Some x) \ (Some y) = Some(x \ y)" - -fun narrow_option where -"None \ x = None" | -"x \ None = None" | -"(Some x) \ (Some y) = Some(x \ y)" - -instance -proof - case goal1 show ?case - by(induct x y rule: widen_option.induct) (simp_all add: widen1) -next - case goal2 show ?case - by(induct x y rule: widen_option.induct) (simp_all add: widen2) -next - case goal3 thus ?case - by(induct x y rule: narrow_option.induct) (simp_all add: narrow1) -next - case goal4 thus ?case - by(induct x y rule: narrow_option.induct) (simp_all add: narrow2) -qed - -end - - -subsubsection "Annotated commands" - -fun map2_acom :: "('a \ 'a \ 'a) \ 'a acom \ 'a acom \ 'a acom" where -"map2_acom f (SKIP {a1}) (SKIP {a2}) = (SKIP {f a1 a2})" | -"map2_acom f (x ::= e {a1}) (x' ::= e' {a2}) = (x ::= e {f a1 a2})" | -"map2_acom f (c1;c2) (c1';c2') = (map2_acom f c1 c1'; map2_acom f c2 c2')" | -"map2_acom f (IF b THEN c1 ELSE c2 {a1}) (IF b' THEN c1' ELSE c2' {a2}) = - (IF b THEN map2_acom f c1 c1' ELSE map2_acom f c2 c2' {f a1 a2})" | -"map2_acom f ({a1} WHILE b DO c {a2}) ({a3} WHILE b' DO c' {a4}) = - ({f a1 a3} WHILE b DO map2_acom f c c' {f a2 a4})" - -abbreviation widen_acom :: "('a::WN)acom \ 'a acom \ 'a acom" (infix "\\<^sub>c" 65) -where "widen_acom == map2_acom (op \)" - -abbreviation narrow_acom :: "('a::WN)acom \ 'a acom \ 'a acom" (infix "\\<^sub>c" 65) -where "narrow_acom == map2_acom (op \)" - -lemma widen1_acom: "strip c = strip c' \ c \ c \\<^sub>c c'" -by(induct c c' rule: le_acom.induct)(simp_all add: widen1) - -lemma widen2_acom: "strip c = strip c' \ c' \ c \\<^sub>c c'" -by(induct c c' rule: le_acom.induct)(simp_all add: widen2) - -lemma narrow1_acom: "y \ x \ y \ x \\<^sub>c y" -by(induct y x rule: le_acom.induct) (simp_all add: narrow1) - -lemma narrow2_acom: "y \ x \ x \\<^sub>c y \ x" -by(induct y x rule: le_acom.induct) (simp_all add: narrow2) - - -subsubsection "Post-fixed point computation" - -definition iter_widen :: "('a acom \ 'a acom) \ 'a acom \ ('a::WN)acom option" -where "iter_widen f = while_option (\c. \ f c \ c) (\c. c \\<^sub>c f c)" - -definition iter_narrow :: "('a acom \ 'a acom) \ 'a acom \ 'a::WN acom option" -where "iter_narrow f = while_option (\c. \ c \ c \\<^sub>c f c) (\c. c \\<^sub>c f c)" - -definition pfp_wn :: - "(('a::WN)option acom \ 'a option acom) \ com \ 'a option acom option" -where "pfp_wn f c = (case iter_widen f (\\<^sub>c c) of None \ None - | Some c' \ iter_narrow f c')" - -lemma strip_map2_acom: - "strip c1 = strip c2 \ strip(map2_acom f c1 c2) = strip c1" -by(induct f c1 c2 rule: map2_acom.induct) simp_all - -lemma iter_widen_pfp: "iter_widen f c = Some c' \ f c' \ c'" -by(auto simp add: iter_widen_def dest: while_option_stop) - -lemma strip_while: fixes f :: "'a acom \ 'a acom" -assumes "\c. strip (f c) = strip c" and "while_option P f c = Some c'" -shows "strip c' = strip c" -using while_option_rule[where P = "\c'. strip c' = strip c", OF _ assms(2)] -by (metis assms(1)) - -lemma strip_iter_widen: fixes f :: "'a::WN acom \ 'a acom" -assumes "\c. strip (f c) = strip c" and "iter_widen f c = Some c'" -shows "strip c' = strip c" -proof- - have "\c. strip(c \\<^sub>c f c) = strip c" by (metis assms(1) strip_map2_acom) - from strip_while[OF this] assms(2) show ?thesis by(simp add: iter_widen_def) -qed - -lemma iter_narrow_pfp: assumes "mono f" and "f c0 \ c0" -and "iter_narrow f c0 = Some c" -shows "f c \ c \ c \ c0" (is "?P c") -proof- - { fix c assume "?P c" - note 1 = conjunct1[OF this] and 2 = conjunct2[OF this] - let ?c' = "c \\<^sub>c f c" - have "?P ?c'" - proof - have "f ?c' \ f c" by(rule monoD[OF `mono f` narrow2_acom[OF 1]]) - also have "\ \ ?c'" by(rule narrow1_acom[OF 1]) - finally show "f ?c' \ ?c'" . - have "?c' \ c" by (rule narrow2_acom[OF 1]) - also have "c \ c0" by(rule 2) - finally show "?c' \ c0" . - qed - } - with while_option_rule[where P = ?P, OF _ assms(3)[simplified iter_narrow_def]] - assms(2) le_refl - show ?thesis by blast -qed - -lemma pfp_wn_pfp: - "\ mono f; pfp_wn f c = Some c' \ \ f c' \ c'" -unfolding pfp_wn_def -by (auto dest: iter_widen_pfp iter_narrow_pfp split: option.splits) - -lemma strip_pfp_wn: - "\ \c. strip(f c) = strip c; pfp_wn f c = Some c' \ \ strip c' = c" -apply(auto simp add: pfp_wn_def iter_narrow_def split: option.splits) -by (metis (no_types) strip_map2_acom strip_while strip_bot_acom strip_iter_widen) - -locale Abs_Int2 = Abs_Int1_mono -where \=\ for \ :: "'av::{WN,L_top_bot} \ val set" -begin - -definition AI_wn :: "com \ 'av st option acom option" where -"AI_wn = pfp_wn (step' \)" - -lemma AI_wn_sound: "AI_wn c = Some c' \ CS c \ \\<^isub>c c'" -proof(simp add: CS_def AI_wn_def) - assume 1: "pfp_wn (step' \) c = Some c'" - from pfp_wn_pfp[OF mono_step'2 1] - have 2: "step' \ c' \ c'" . - have 3: "strip (\\<^isub>c (step' \ c')) = c" by(simp add: strip_pfp_wn[OF _ 1]) - have "lfp (step UNIV) c \ \\<^isub>c (step' \ c')" - proof(rule lfp_lowerbound[simplified,OF 3]) - show "step UNIV (\\<^isub>c (step' \ c')) \ \\<^isub>c (step' \ c')" - proof(rule step_preserves_le[OF _ _]) - show "UNIV \ \\<^isub>o \" by simp - show "\\<^isub>c (step' \ c') \ \\<^isub>c c'" by(rule mono_gamma_c[OF 2]) - qed - qed - from this 2 show "lfp (step UNIV) c \ \\<^isub>c c'" - by (blast intro: mono_gamma_c order_trans) -qed - -end - -interpretation Abs_Int2 -where \ = \_ivl and num' = num_ivl and plus' = plus_ivl -and test_num' = in_ivl -and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl -defines AI_ivl' is AI_wn -.. - - -subsubsection "Tests" - -definition "step_up_ivl n = ((\c. c \\<^sub>c step_ivl \ c)^^n)" -definition "step_down_ivl n = ((\c. c \\<^sub>c step_ivl \ c)^^n)" - -text{* For @{const test3_ivl}, @{const AI_ivl} needed as many iterations as -the loop took to execute. In contrast, @{const AI_ivl'} converges in a -constant number of steps: *} - -value "show_acom (step_up_ivl 1 (\\<^sub>c test3_ivl))" -value "show_acom (step_up_ivl 2 (\\<^sub>c test3_ivl))" -value "show_acom (step_up_ivl 3 (\\<^sub>c test3_ivl))" -value "show_acom (step_up_ivl 4 (\\<^sub>c test3_ivl))" -value "show_acom (step_up_ivl 5 (\\<^sub>c test3_ivl))" -value "show_acom (step_down_ivl 1 (step_up_ivl 5 (\\<^sub>c test3_ivl)))" -value "show_acom (step_down_ivl 2 (step_up_ivl 5 (\\<^sub>c test3_ivl)))" -value "show_acom (step_down_ivl 3 (step_up_ivl 5 (\\<^sub>c test3_ivl)))" - -text{* Now all the analyses terminate: *} - -value "show_acom_opt (AI_ivl' test4_ivl)" -value "show_acom_opt (AI_ivl' test5_ivl)" -value "show_acom_opt (AI_ivl' test6_ivl)" - - -subsubsection "Termination: Intervals" - -definition m_ivl :: "ivl \ nat" where -"m_ivl ivl = (case ivl of I l h \ - (case l of None \ 0 | Some _ \ 1) + (case h of None \ 0 | Some _ \ 1))" - -lemma m_ivl_height: "m_ivl ivl \ 2" -by(simp add: m_ivl_def split: ivl.split option.split) - -lemma m_ivl_anti_mono: "(y::ivl) \ x \ m_ivl x \ m_ivl y" -by(auto simp: m_ivl_def le_option_def le_ivl_def - split: ivl.split option.split if_splits) - -lemma m_ivl_widen: - "~ y \ x \ m_ivl(x \ y) < m_ivl x" -by(auto simp: m_ivl_def widen_ivl_def le_option_def le_ivl_def - split: ivl.splits option.splits if_splits) - -lemma Top_less_ivl: "\ \ x \ m_ivl x = 0" -by(auto simp: m_ivl_def le_option_def le_ivl_def empty_def Top_ivl_def - split: ivl.split option.split if_splits) - - -definition n_ivl :: "ivl \ nat" where -"n_ivl ivl = 2 - m_ivl ivl" - -lemma n_ivl_mono: "(x::ivl) \ y \ n_ivl x \ n_ivl y" -unfolding n_ivl_def by (metis diff_le_mono2 m_ivl_anti_mono) - -lemma n_ivl_narrow: - "~ x \ x \ y \ n_ivl(x \ y) < n_ivl x" -by(auto simp: n_ivl_def m_ivl_def narrow_ivl_def le_option_def le_ivl_def - split: ivl.splits option.splits if_splits) - - -subsubsection "Termination: Abstract State" - -definition "m_st m st = (\x\set(dom st). m(fun st x))" - -lemma m_st_height: assumes "finite X" and "set (dom S) \ X" -shows "m_st m_ivl S \ 2 * card X" -proof(auto simp: m_st_def) - have "(\x\set(dom S). m_ivl (fun S x)) \ (\x\set(dom S). 2)" (is "?L \ _") - by(rule setsum_mono)(simp add:m_ivl_height) - also have "\ \ (\x\X. 2)" - by(rule setsum_mono3[OF assms]) simp - also have "\ = 2 * card X" by(simp add: setsum_constant) - finally show "?L \ \" . -qed - -lemma m_st_anti_mono: - "S1 \ S2 \ m_st m_ivl S2 \ m_st m_ivl S1" -proof(auto simp: m_st_def le_st_def lookup_def split: if_splits) - let ?X = "set(dom S1)" let ?Y = "set(dom S2)" - let ?f = "fun S1" let ?g = "fun S2" - assume asm: "\x\?Y. (x \ ?X \ ?f x \ ?g x) \ (x \ ?X \ \ \ ?g x)" - hence 1: "\y\?Y\?X. m_ivl(?g y) \ m_ivl(?f y)" by(simp add: m_ivl_anti_mono) - have 0: "\x\?Y-?X. m_ivl(?g x) = 0" using asm by (auto simp: Top_less_ivl) - have "(\y\?Y. m_ivl(?g y)) = (\y\(?Y-?X) \ (?Y\?X). m_ivl(?g y))" - by (metis Un_Diff_Int) - also have "\ = (\y\?Y-?X. m_ivl(?g y)) + (\y\?Y\?X. m_ivl(?g y))" - by(subst setsum_Un_disjoint) auto - also have "(\y\?Y-?X. m_ivl(?g y)) = 0" using 0 by simp - also have "0 + (\y\?Y\?X. m_ivl(?g y)) = (\y\?Y\?X. m_ivl(?g y))" by simp - also have "\ \ (\y\?Y\?X. m_ivl(?f y))" - by(rule setsum_mono)(simp add: 1) - also have "\ \ (\y\?X. m_ivl(?f y))" - by(simp add: setsum_mono3[of "?X" "?Y Int ?X", OF _ Int_lower2]) - finally show "(\y\?Y. m_ivl(?g y)) \ (\x\?X. m_ivl(?f x))" - by (metis add_less_cancel_left) -qed - -lemma m_st_widen: -assumes "\ S2 \ S1" shows "m_st m_ivl (S1 \ S2) < m_st m_ivl S1" -proof- - { let ?X = "set(dom S1)" let ?Y = "set(dom S2)" - let ?f = "fun S1" let ?g = "fun S2" - fix x assume "x \ ?X" "\ lookup S2 x \ ?f x" - have "(\x\?X\?Y. m_ivl(?f x \ ?g x)) < (\x\?X. m_ivl(?f x))" (is "?L < ?R") - proof cases - assume "x : ?Y" - have "?L < (\x\?X\?Y. m_ivl(?f x))" - proof(rule setsum_strict_mono1, simp) - show "\x\?X\?Y. m_ivl(?f x \ ?g x) \ m_ivl (?f x)" - by (metis m_ivl_anti_mono widen1) - next - show "\x\?X\?Y. m_ivl(?f x \ ?g x) < m_ivl(?f x)" - using `x:?X` `x:?Y` `\ lookup S2 x \ ?f x` - by (metis IntI m_ivl_widen lookup_def) - qed - also have "\ \ ?R" by(simp add: setsum_mono3[OF _ Int_lower1]) - finally show ?thesis . - next - assume "x ~: ?Y" - have "?L \ (\x\?X\?Y. m_ivl(?f x))" - proof(rule setsum_mono, simp) - fix x assume "x:?X \ x:?Y" show "m_ivl(?f x \ ?g x) \ m_ivl (?f x)" - by (metis m_ivl_anti_mono widen1) - qed - also have "\ < m_ivl(?f x) + \" - using m_ivl_widen[OF `\ lookup S2 x \ ?f x`] - by (metis Nat.le_refl add_strict_increasing gr0I not_less0) - also have "\ = (\y\insert x (?X\?Y). m_ivl(?f y))" - using `x ~: ?Y` by simp - also have "\ \ (\x\?X. m_ivl(?f x))" - by(rule setsum_mono3)(insert `x:?X`, auto) - finally show ?thesis . - qed - } with assms show ?thesis - by(auto simp: le_st_def widen_st_def m_st_def Int_def) -qed - -definition "n_st m X st = (\x\X. m(lookup st x))" - -lemma n_st_mono: assumes "set(dom S1) \ X" "set(dom S2) \ X" "S1 \ S2" -shows "n_st n_ivl X S1 \ n_st n_ivl X S2" -proof- - have "(\x\X. n_ivl(lookup S1 x)) \ (\x\X. n_ivl(lookup S2 x))" - apply(rule setsum_mono) using assms - by(auto simp: le_st_def lookup_def n_ivl_mono split: if_splits) - thus ?thesis by(simp add: n_st_def) -qed - -lemma n_st_narrow: -assumes "finite X" and "set(dom S1) \ X" "set(dom S2) \ X" -and "S2 \ S1" "\ S1 \ S1 \ S2" -shows "n_st n_ivl X (S1 \ S2) < n_st n_ivl X S1" -proof- - have 1: "\x\X. n_ivl (lookup (S1 \ S2) x) \ n_ivl (lookup S1 x)" - using assms(2-4) - by(auto simp: le_st_def narrow_st_def lookup_def n_ivl_mono narrow2 - split: if_splits) - have 2: "\x\X. n_ivl (lookup (S1 \ S2) x) < n_ivl (lookup S1 x)" - using assms(2-5) - by(auto simp: le_st_def narrow_st_def lookup_def intro: n_ivl_narrow - split: if_splits) - have "(\x\X. n_ivl(lookup (S1 \ S2) x)) < (\x\X. n_ivl(lookup S1 x))" - apply(rule setsum_strict_mono1[OF `finite X`]) using 1 2 by blast+ - thus ?thesis by(simp add: n_st_def) -qed - - -subsubsection "Termination: Option" - -definition "m_o m n opt = (case opt of None \ n+1 | Some x \ m x)" - -lemma m_o_anti_mono: "finite X \ domo S2 \ X \ S1 \ S2 \ - m_o (m_st m_ivl) (2 * card X) S2 \ m_o (m_st m_ivl) (2 * card X) S1" -apply(induction S1 S2 rule: le_option.induct) -apply(auto simp: domo_def m_o_def m_st_anti_mono le_SucI m_st_height - split: option.splits) -done - -lemma m_o_widen: "\ finite X; domo S2 \ X; \ S2 \ S1 \ \ - m_o (m_st m_ivl) (2 * card X) (S1 \ S2) < m_o (m_st m_ivl) (2 * card X) S1" -by(auto simp: m_o_def domo_def m_st_height less_Suc_eq_le m_st_widen - split: option.split) - -definition "n_o n opt = (case opt of None \ 0 | Some x \ n x + 1)" - -lemma n_o_mono: "domo S1 \ X \ domo S2 \ X \ S1 \ S2 \ - n_o (n_st n_ivl X) S1 \ n_o (n_st n_ivl X) S2" -apply(induction S1 S2 rule: le_option.induct) -apply(auto simp: domo_def n_o_def n_st_mono - split: option.splits) -done - -lemma n_o_narrow: - "\ finite X; domo S1 \ X; domo S2 \ X; S2 \ S1; \ S1 \ S1 \ S2 \ - \ n_o (n_st n_ivl X) (S1 \ S2) < n_o (n_st n_ivl X) S1" -apply(induction S1 S2 rule: narrow_option.induct) -apply(auto simp: n_o_def domo_def n_st_narrow) -done - -lemma domo_widen_subset: "domo (S1 \ S2) \ domo S1 \ domo S2" -apply(induction S1 S2 rule: widen_option.induct) -apply (auto simp: domo_def widen_st_def) -done - -lemma domo_narrow_subset: "domo (S1 \ S2) \ domo S1 \ domo S2" -apply(induction S1 S2 rule: narrow_option.induct) -apply (auto simp: domo_def narrow_st_def) -done - -subsubsection "Termination: Commands" - -lemma strip_widen_acom[simp]: - "strip c' = strip (c::'a::WN acom) \ strip (c \\<^sub>c c') = strip c" -by(induction "widen::'a\'a\'a" c c' rule: map2_acom.induct) simp_all - -lemma strip_narrow_acom[simp]: - "strip c' = strip (c::'a::WN acom) \ strip (c \\<^sub>c c') = strip c" -by(induction "narrow::'a\'a\'a" c c' rule: map2_acom.induct) simp_all - -lemma annos_widen_acom[simp]: "strip c1 = strip (c2::'a::WN acom) \ - annos(c1 \\<^sub>c c2) = map (%(x,y).x\y) (zip (annos c1) (annos(c2::'a::WN acom)))" -by(induction "widen::'a\'a\'a" c1 c2 rule: map2_acom.induct) - (simp_all add: size_annos_same2) - -lemma annos_narrow_acom[simp]: "strip c1 = strip (c2::'a::WN acom) \ - annos(c1 \\<^sub>c c2) = map (%(x,y).x\y) (zip (annos c1) (annos(c2::'a::WN acom)))" -by(induction "narrow::'a\'a\'a" c1 c2 rule: map2_acom.induct) - (simp_all add: size_annos_same2) - -lemma widen_acom_Com[simp]: "strip c2 = strip c1 \ - c1 : Com X \ c2 : Com X \ (c1 \\<^sub>c c2) : Com X" -apply(auto simp add: Com_def) -apply(rename_tac S S' x) -apply(erule in_set_zipE) -apply(auto simp: domo_def split: option.splits) -apply(case_tac S) -apply(case_tac S') -apply simp -apply fastforce -apply(case_tac S') -apply fastforce -apply (fastforce simp: widen_st_def) -done - -lemma narrow_acom_Com[simp]: "strip c2 = strip c1 \ - c1 : Com X \ c2 : Com X \ (c1 \\<^sub>c c2) : Com X" -apply(auto simp add: Com_def) -apply(rename_tac S S' x) -apply(erule in_set_zipE) -apply(auto simp: domo_def split: option.splits) -apply(case_tac S) -apply(case_tac S') -apply simp -apply fastforce -apply(case_tac S') -apply fastforce -apply (fastforce simp: narrow_st_def) -done - -definition "m_c m c = (let as = annos c in \i=0.. {(c, c \\<^sub>c c') |c c'\ivl st option acom. - strip c' = strip c \ c : Com X \ c' : Com X \ \ c' \ c}\ - \ measure(m_c(m_o (m_st m_ivl) (2*card(X))))" -apply(auto simp: m_c_def Let_def Com_def) -apply(subgoal_tac "length(annos c') = length(annos c)") -prefer 2 apply (simp add: size_annos_same2) -apply (auto) -apply(rule setsum_strict_mono1) -apply simp -apply (clarsimp) -apply(erule m_o_anti_mono) -apply(rule subset_trans[OF domo_widen_subset]) -apply fastforce -apply(rule widen1) -apply(auto simp: le_iff_le_annos listrel_iff_nth) -apply(rule_tac x=n in bexI) -prefer 2 apply simp -apply(erule m_o_widen) -apply (simp)+ -done - -lemma measure_n_c: "finite X \ {(c, c \\<^sub>c c') |c c'. - strip c = strip c' \ c \ Com X \ c' \ Com X \ c' \ c \ \ c \ c \\<^sub>c c'}\ - \ measure(m_c(n_o (n_st n_ivl X)))" -apply(auto simp: m_c_def Let_def Com_def) -apply(subgoal_tac "length(annos c') = length(annos c)") -prefer 2 apply (simp add: size_annos_same2) -apply (auto) -apply(rule setsum_strict_mono1) -apply simp -apply (clarsimp) -apply(rule n_o_mono) -using domo_narrow_subset apply fastforce -apply fastforce -apply(rule narrow2) -apply(fastforce simp: le_iff_le_annos listrel_iff_nth) -apply(auto simp: le_iff_le_annos listrel_iff_nth strip_narrow_acom) -apply(rule_tac x=n in bexI) -prefer 2 apply simp -apply(erule n_o_narrow) -apply (simp)+ -done - - -subsubsection "Termination: Post-Fixed Point Iterations" - -lemma iter_widen_termination: -fixes c0 :: "'a::WN acom" -assumes P_f: "\c. P c \ P(f c)" -assumes P_widen: "\c c'. P c \ P c' \ P(c \\<^sub>c c')" -and "wf({(c::'a acom,c \\<^sub>c c')|c c'. P c \ P c' \ ~ c' \ c}^-1)" -and "P c0" and "c0 \ f c0" shows "EX c. iter_widen f c0 = Some c" -proof(simp add: iter_widen_def, rule wf_while_option_Some[where P = "P"]) - show "wf {(cc', c). (P c \ \ f c \ c) \ cc' = c \\<^sub>c f c}" - apply(rule wf_subset[OF assms(3)]) by(blast intro: P_f) -next - show "P c0" by(rule `P c0`) -next - fix c assume "P c" thus "P (c \\<^sub>c f c)" by(simp add: P_f P_widen) -qed - -lemma iter_narrow_termination: -assumes P_f: "\c. P c \ P(c \\<^sub>c f c)" -and wf: "wf({(c, c \\<^sub>c f c)|c c'. P c \ ~ c \ c \\<^sub>c f c}^-1)" -and "P c0" shows "EX c. iter_narrow f c0 = Some c" -proof(simp add: iter_narrow_def, rule wf_while_option_Some[where P = "P"]) - show "wf {(c', c). (P c \ \ c \ c \\<^sub>c f c) \ c' = c \\<^sub>c f c}" - apply(rule wf_subset[OF wf]) by(blast intro: P_f) -next - show "P c0" by(rule `P c0`) -next - fix c assume "P c" thus "P (c \\<^sub>c f c)" by(simp add: P_f) -qed - -lemma iter_winden_step_ivl_termination: - "EX c. iter_widen (step_ivl \) (\\<^sub>c c0) = Some c" -apply(rule iter_widen_termination[where - P = "%c. strip c = c0 \ c : Com(vars c0)"]) -apply (simp_all add: step'_Com bot_acom) -apply(rule wf_subset) -apply(rule wf_measure) -apply(rule subset_trans) -prefer 2 -apply(rule measure_m_c[where X = "vars c0", OF finite_cvars]) -apply blast -done - -lemma iter_narrow_step_ivl_termination: - "c0 \ Com (vars(strip c0)) \ step_ivl \ c0 \ c0 \ - EX c. iter_narrow (step_ivl \) c0 = Some c" -apply(rule iter_narrow_termination[where - P = "%c. strip c = strip c0 \ c : Com(vars(strip c0)) \ step_ivl \ c \ c"]) -apply (simp_all add: step'_Com) -apply(clarify) -apply(frule narrow2_acom, drule mono_step'[OF le_refl], erule le_trans[OF _ narrow1_acom]) -apply assumption -apply(rule wf_subset) -apply(rule wf_measure) -apply(rule subset_trans) -prefer 2 -apply(rule measure_n_c[where X = "vars(strip c0)", OF finite_cvars]) -apply auto -by (metis bot_least domo_Top order_refl step'_Com strip_step') - -(* FIXME: simplify type system: Combine Com(X) and vars <= X?? *) -lemma while_Com: -fixes c :: "'a st option acom" -assumes "while_option P f c = Some c'" -and "!!c. strip(f c) = strip c" -and "\c::'a st option acom. c : Com(X) \ vars(strip c) \ X \ f c : Com(X)" -and "c : Com(X)" and "vars(strip c) \ X" shows "c' : Com(X)" -using while_option_rule[where P = "\c'. c' : Com(X) \ vars(strip c') \ X", OF _ assms(1)] -by(simp add: assms(2-)) - -lemma iter_widen_Com: fixes f :: "'a::WN st option acom \ 'a st option acom" -assumes "iter_widen f c = Some c'" -and "\c. c : Com(X) \ vars(strip c) \ X \ f c : Com(X)" -and "!!c. strip(f c) = strip c" -and "c : Com(X)" and "vars (strip c) \ X" shows "c' : Com(X)" -proof- - have "\c. c : Com(X) \ vars(strip c) \ X \ c \\<^sub>c f c : Com(X)" - by (metis (full_types) widen_acom_Com assms(2,3)) - from while_Com[OF assms(1)[simplified iter_widen_def] _ this assms(4,5)] - show ?thesis using assms(3) by(simp) -qed - - -context Abs_Int2 -begin - -lemma iter_widen_step'_Com: - "iter_widen (step' \) c = Some c' \ vars(strip c) \ X \ c : Com(X) - \ c' : Com(X)" -apply(subgoal_tac "strip c'= strip c") -prefer 2 apply (metis strip_iter_widen strip_step') -apply(drule iter_widen_Com) -prefer 3 apply assumption -prefer 3 apply assumption -apply (auto simp: step'_Com) -done - -end - -theorem AI_ivl'_termination: - "EX c'. AI_ivl' c = Some c'" -apply(auto simp: AI_wn_def pfp_wn_def iter_winden_step_ivl_termination split: option.split) -apply(rule iter_narrow_step_ivl_termination) -apply (metis bot_acom_Com iter_widen_step'_Com[OF _ subset_refl] strip_iter_widen strip_step') -apply(erule iter_widen_pfp) -done - -end - - -(* interesting(?) relic -lemma widen_assoc: - "~ (y::ivl) \ x \ ~ z \ x \ y \ ((x::ivl) \ y) \ z = x \ (y \ z)" -apply(cases x) -apply(cases y) -apply(cases z) -apply(rename_tac x1 x2 y1 y2 z1 z2) -apply(simp add: le_ivl_def) -apply(case_tac x1) -apply(case_tac x2) -apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits) -apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits) -apply(case_tac x2) -apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits) -apply(case_tac y1) -apply(case_tac y2) -apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits) -apply(case_tac z1) -apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits ivl.splits)[1] -apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits ivl.splits)[1] -apply(case_tac y2) -apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits ivl.splits)[1] -apply(case_tac z1) -apply(auto simp add:le_option_def widen_ivl_def split: if_splits ivl.splits option.splits)[1] -apply(case_tac z2) -apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits)[1] -apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits)[1] -done - -lemma widen_step_trans: - "~ (y::ivl) \ x \ ~ z \ x \ y \ EX u. (x \ y) \ z = x \ u \ ~ u \ x" -by (metis widen_assoc preord_class.le_trans widen1) -*) diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_Int_ITP/Abs_Int0_ITP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int0_ITP.thy Thu Apr 19 17:49:02 2012 +0200 @@ -0,0 +1,397 @@ +(* Author: Tobias Nipkow *) + +theory Abs_Int0_ITP +imports "~~/src/HOL/ex/Interpretation_with_Defs" + "~~/src/HOL/Library/While_Combinator" + Collecting +begin + +subsection "Orderings" + +class preord = +fixes le :: "'a \ 'a \ bool" (infix "\" 50) +assumes le_refl[simp]: "x \ x" +and le_trans: "x \ y \ y \ z \ x \ z" +begin + +definition mono where "mono f = (\x y. x \ y \ f x \ f y)" + +lemma monoD: "mono f \ x \ y \ f x \ f y" by(simp add: mono_def) + +lemma mono_comp: "mono f \ mono g \ mono (g o f)" +by(simp add: mono_def) + +declare le_trans[trans] + +end + +text{* Note: no antisymmetry. Allows implementations where some abstract +element is implemented by two different values @{prop "x \ y"} +such that @{prop"x \ y"} and @{prop"y \ x"}. Antisymmetry is not +needed because we never compare elements for equality but only for @{text"\"}. +*} + +class SL_top = preord + +fixes join :: "'a \ 'a \ 'a" (infixl "\" 65) +fixes Top :: "'a" ("\") +assumes join_ge1 [simp]: "x \ x \ y" +and join_ge2 [simp]: "y \ x \ y" +and join_least: "x \ z \ y \ z \ x \ y \ z" +and top[simp]: "x \ \" +begin + +lemma join_le_iff[simp]: "x \ y \ z \ x \ z \ y \ z" +by (metis join_ge1 join_ge2 join_least le_trans) + +lemma le_join_disj: "x \ y \ x \ z \ x \ y \ z" +by (metis join_ge1 join_ge2 le_trans) + +end + +instantiation "fun" :: (type, SL_top) SL_top +begin + +definition "f \ g = (ALL x. f x \ g x)" +definition "f \ g = (\x. f x \ g x)" +definition "\ = (\x. \)" + +lemma join_apply[simp]: "(f \ g) x = f x \ g x" +by (simp add: join_fun_def) + +instance +proof + case goal2 thus ?case by (metis le_fun_def preord_class.le_trans) +qed (simp_all add: le_fun_def Top_fun_def) + +end + + +instantiation acom :: (preord) preord +begin + +fun le_acom :: "('a::preord)acom \ 'a acom \ bool" where +"le_acom (SKIP {S}) (SKIP {S'}) = (S \ S')" | +"le_acom (x ::= e {S}) (x' ::= e' {S'}) = (x=x' \ e=e' \ S \ S')" | +"le_acom (c1;c2) (c1';c2') = (le_acom c1 c1' \ le_acom c2 c2')" | +"le_acom (IF b THEN c1 ELSE c2 {S}) (IF b' THEN c1' ELSE c2' {S'}) = + (b=b' \ le_acom c1 c1' \ le_acom c2 c2' \ S \ S')" | +"le_acom ({Inv} WHILE b DO c {P}) ({Inv'} WHILE b' DO c' {P'}) = + (b=b' \ le_acom c c' \ Inv \ Inv' \ P \ P')" | +"le_acom _ _ = False" + +lemma [simp]: "SKIP {S} \ c \ (\S'. c = SKIP {S'} \ S \ S')" +by (cases c) auto + +lemma [simp]: "x ::= e {S} \ c \ (\S'. c = x ::= e {S'} \ S \ S')" +by (cases c) auto + +lemma [simp]: "c1;c2 \ c \ (\c1' c2'. c = c1';c2' \ c1 \ c1' \ c2 \ c2')" +by (cases c) auto + +lemma [simp]: "IF b THEN c1 ELSE c2 {S} \ c \ + (\c1' c2' S'. c = IF b THEN c1' ELSE c2' {S'} \ c1 \ c1' \ c2 \ c2' \ S \ S')" +by (cases c) auto + +lemma [simp]: "{Inv} WHILE b DO c {P} \ w \ + (\Inv' c' P'. w = {Inv'} WHILE b DO c' {P'} \ c \ c' \ Inv \ Inv' \ P \ P')" +by (cases w) auto + +instance +proof + case goal1 thus ?case by (induct x) auto +next + case goal2 thus ?case + apply(induct x y arbitrary: z rule: le_acom.induct) + apply (auto intro: le_trans) + done +qed + +end + + +subsubsection "Lifting" + +instantiation option :: (preord)preord +begin + +fun le_option where +"Some x \ Some y = (x \ y)" | +"None \ y = True" | +"Some _ \ None = False" + +lemma [simp]: "(x \ None) = (x = None)" +by (cases x) simp_all + +lemma [simp]: "(Some x \ u) = (\y. u = Some y & x \ y)" +by (cases u) auto + +instance proof + case goal1 show ?case by(cases x, simp_all) +next + case goal2 thus ?case + by(cases z, simp, cases y, simp, cases x, auto intro: le_trans) +qed + +end + +instantiation option :: (SL_top)SL_top +begin + +fun join_option where +"Some x \ Some y = Some(x \ y)" | +"None \ y = y" | +"x \ None = x" + +lemma join_None2[simp]: "x \ None = x" +by (cases x) simp_all + +definition "\ = Some \" + +instance proof + case goal1 thus ?case by(cases x, simp, cases y, simp_all) +next + case goal2 thus ?case by(cases y, simp, cases x, simp_all) +next + case goal3 thus ?case by(cases z, simp, cases y, simp, cases x, simp_all) +next + case goal4 thus ?case by(cases x, simp_all add: Top_option_def) +qed + +end + +definition bot_acom :: "com \ ('a::SL_top)option acom" ("\\<^sub>c") where +"\\<^sub>c = anno None" + +lemma strip_bot_acom[simp]: "strip(\\<^sub>c c) = c" +by(simp add: bot_acom_def) + +lemma bot_acom[rule_format]: "strip c' = c \ \\<^sub>c c \ c'" +apply(induct c arbitrary: c') +apply (simp_all add: bot_acom_def) + apply(induct_tac c') + apply simp_all + apply(induct_tac c') + apply simp_all + apply(induct_tac c') + apply simp_all + apply(induct_tac c') + apply simp_all + apply(induct_tac c') + apply simp_all +done + + +subsubsection "Post-fixed point iteration" + +definition + pfp :: "(('a::preord) \ 'a) \ 'a \ 'a option" where +"pfp f = while_option (\x. \ f x \ x) f" + +lemma pfp_pfp: assumes "pfp f x0 = Some x" shows "f x \ x" +using while_option_stop[OF assms[simplified pfp_def]] by simp + +lemma pfp_least: +assumes mono: "\x y. x \ y \ f x \ f y" +and "f p \ p" and "x0 \ p" and "pfp f x0 = Some x" shows "x \ p" +proof- + { fix x assume "x \ p" + hence "f x \ f p" by(rule mono) + from this `f p \ p` have "f x \ p" by(rule le_trans) + } + thus "x \ p" using assms(2-) while_option_rule[where P = "%x. x \ p"] + unfolding pfp_def by blast +qed + +definition + lpfp\<^isub>c :: "(('a::SL_top)option acom \ 'a option acom) \ com \ 'a option acom option" where +"lpfp\<^isub>c f c = pfp f (\\<^sub>c c)" + +lemma lpfpc_pfp: "lpfp\<^isub>c f c0 = Some c \ f c \ c" +by(simp add: pfp_pfp lpfp\<^isub>c_def) + +lemma strip_pfp: +assumes "\x. g(f x) = g x" and "pfp f x0 = Some x" shows "g x = g x0" +using assms while_option_rule[where P = "%x. g x = g x0" and c = f] +unfolding pfp_def by metis + +lemma strip_lpfpc: assumes "\c. strip(f c) = strip c" and "lpfp\<^isub>c f c = Some c'" +shows "strip c' = c" +using assms(1) strip_pfp[OF _ assms(2)[simplified lpfp\<^isub>c_def]] +by(metis strip_bot_acom) + +lemma lpfpc_least: +assumes mono: "\x y. x \ y \ f x \ f y" +and "strip p = c0" and "f p \ p" and lp: "lpfp\<^isub>c f c0 = Some c" shows "c \ p" +using pfp_least[OF _ _ bot_acom[OF `strip p = c0`] lp[simplified lpfp\<^isub>c_def]] + mono `f p \ p` +by blast + + +subsection "Abstract Interpretation" + +definition \_fun :: "('a \ 'b set) \ ('c \ 'a) \ ('c \ 'b)set" where +"\_fun \ F = {f. \x. f x \ \(F x)}" + +fun \_option :: "('a \ 'b set) \ 'a option \ 'b set" where +"\_option \ None = {}" | +"\_option \ (Some a) = \ a" + +text{* The interface for abstract values: *} + +locale Val_abs = +fixes \ :: "'av::SL_top \ val set" + assumes mono_gamma: "a \ b \ \ a \ \ b" + and gamma_Top[simp]: "\ \ = UNIV" +fixes num' :: "val \ 'av" +and plus' :: "'av \ 'av \ 'av" + assumes gamma_num': "n : \(num' n)" + and gamma_plus': + "n1 : \ a1 \ n2 : \ a2 \ n1+n2 : \(plus' a1 a2)" + +type_synonym 'av st = "(vname \ 'av)" + +locale Abs_Int_Fun = Val_abs \ for \ :: "'av::SL_top \ val set" +begin + +fun aval' :: "aexp \ 'av st \ 'av" where +"aval' (N n) S = num' n" | +"aval' (V x) S = S x" | +"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" + +fun step' :: "'av st option \ 'av st option acom \ 'av st option acom" + where +"step' S (SKIP {P}) = (SKIP {S})" | +"step' S (x ::= e {P}) = + x ::= e {case S of None \ None | Some S \ Some(S(x := aval' e S))}" | +"step' S (c1; c2) = step' S c1; step' (post c1) c2" | +"step' S (IF b THEN c1 ELSE c2 {P}) = + IF b THEN step' S c1 ELSE step' S c2 {post c1 \ post c2}" | +"step' S ({Inv} WHILE b DO c {P}) = + {S \ post c} WHILE b DO (step' Inv c) {Inv}" + +definition AI :: "com \ 'av st option acom option" where +"AI = lpfp\<^isub>c (step' \)" + + +lemma strip_step'[simp]: "strip(step' S c) = strip c" +by(induct c arbitrary: S) (simp_all add: Let_def) + + +abbreviation \\<^isub>f :: "'av st \ state set" +where "\\<^isub>f == \_fun \" + +abbreviation \\<^isub>o :: "'av st option \ state set" +where "\\<^isub>o == \_option \\<^isub>f" + +abbreviation \\<^isub>c :: "'av st option acom \ state set acom" +where "\\<^isub>c == map_acom \\<^isub>o" + +lemma gamma_f_Top[simp]: "\\<^isub>f Top = UNIV" +by(simp add: Top_fun_def \_fun_def) + +lemma gamma_o_Top[simp]: "\\<^isub>o Top = UNIV" +by (simp add: Top_option_def) + +(* FIXME (maybe also le \ sqle?) *) + +lemma mono_gamma_f: "f \ g \ \\<^isub>f f \ \\<^isub>f g" +by(auto simp: le_fun_def \_fun_def dest: mono_gamma) + +lemma mono_gamma_o: + "sa \ sa' \ \\<^isub>o sa \ \\<^isub>o sa'" +by(induction sa sa' rule: le_option.induct)(simp_all add: mono_gamma_f) + +lemma mono_gamma_c: "ca \ ca' \ \\<^isub>c ca \ \\<^isub>c ca'" +by (induction ca ca' rule: le_acom.induct) (simp_all add:mono_gamma_o) + +text{* Soundness: *} + +lemma aval'_sound: "s : \\<^isub>f S \ aval a s : \(aval' a S)" +by (induct a) (auto simp: gamma_num' gamma_plus' \_fun_def) + +lemma in_gamma_update: + "\ s : \\<^isub>f S; i : \ a \ \ s(x := i) : \\<^isub>f(S(x := a))" +by(simp add: \_fun_def) + +lemma step_preserves_le: + "\ S \ \\<^isub>o S'; c \ \\<^isub>c c' \ \ step S c \ \\<^isub>c (step' S' c')" +proof(induction c arbitrary: c' S S') + case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) +next + case Assign thus ?case + by (fastforce simp: Assign_le map_acom_Assign intro: aval'_sound in_gamma_update + split: option.splits del:subsetD) +next + case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi) + by (metis le_post post_map_acom) +next + case (If b c1 c2 P) + then obtain c1' c2' P' where + "c' = IF b THEN c1' ELSE c2' {P'}" + "P \ \\<^isub>o P'" "c1 \ \\<^isub>c c1'" "c2 \ \\<^isub>c c2'" + by (fastforce simp: If_le map_acom_If) + moreover have "post c1 \ \\<^isub>o(post c1' \ post c2')" + by (metis (no_types) `c1 \ \\<^isub>c c1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom) + moreover have "post c2 \ \\<^isub>o(post c1' \ post c2')" + by (metis (no_types) `c2 \ \\<^isub>c c2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom) + ultimately show ?case using `S \ \\<^isub>o S'` by (simp add: If.IH subset_iff) +next + case (While I b c1 P) + then obtain c1' I' P' where + "c' = {I'} WHILE b DO c1' {P'}" + "I \ \\<^isub>o I'" "P \ \\<^isub>o P'" "c1 \ \\<^isub>c c1'" + by (fastforce simp: map_acom_While While_le) + moreover have "S \ post c1 \ \\<^isub>o (S' \ post c1')" + using `S \ \\<^isub>o S'` le_post[OF `c1 \ \\<^isub>c c1'`, simplified] + by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans) + ultimately show ?case by (simp add: While.IH subset_iff) +qed + +lemma AI_sound: "AI c = Some c' \ CS c \ \\<^isub>c c'" +proof(simp add: CS_def AI_def) + assume 1: "lpfp\<^isub>c (step' \) c = Some c'" + have 2: "step' \ c' \ c'" by(rule lpfpc_pfp[OF 1]) + have 3: "strip (\\<^isub>c (step' \ c')) = c" + by(simp add: strip_lpfpc[OF _ 1]) + have "lfp (step UNIV) c \ \\<^isub>c (step' \ c')" + proof(rule lfp_lowerbound[simplified,OF 3]) + show "step UNIV (\\<^isub>c (step' \ c')) \ \\<^isub>c (step' \ c')" + proof(rule step_preserves_le[OF _ _]) + show "UNIV \ \\<^isub>o \" by simp + show "\\<^isub>c (step' \ c') \ \\<^isub>c c'" by(rule mono_gamma_c[OF 2]) + qed + qed + with 2 show "lfp (step UNIV) c \ \\<^isub>c c'" + by (blast intro: mono_gamma_c order_trans) +qed + +end + + +subsubsection "Monotonicity" + +lemma mono_post: "c \ c' \ post c \ post c'" +by(induction c c' rule: le_acom.induct) (auto) + +locale Abs_Int_Fun_mono = Abs_Int_Fun + +assumes mono_plus': "a1 \ b1 \ a2 \ b2 \ plus' a1 a2 \ plus' b1 b2" +begin + +lemma mono_aval': "S \ S' \ aval' e S \ aval' e S'" +by(induction e)(auto simp: le_fun_def mono_plus') + +lemma mono_update: "a \ a' \ S \ S' \ S(x := a) \ S'(x := a')" +by(simp add: le_fun_def) + +lemma mono_step': "S \ S' \ c \ c' \ step' S c \ step' S' c'" +apply(induction c c' arbitrary: S S' rule: le_acom.induct) +apply (auto simp: Let_def mono_update mono_aval' mono_post le_join_disj + split: option.split) +done + +end + +text{* Problem: not executable because of the comparison of abstract states, +i.e. functions, in the post-fixedpoint computation. *} + +end diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy Thu Apr 19 17:49:02 2012 +0200 @@ -0,0 +1,411 @@ +(* Author: Tobias Nipkow *) + +theory Abs_Int1_ITP +imports Abs_State_ITP +begin + +subsection "Computable Abstract Interpretation" + +text{* Abstract interpretation over type @{text st} instead of +functions. *} + +context Gamma +begin + +fun aval' :: "aexp \ 'av st \ 'av" where +"aval' (N n) S = num' n" | +"aval' (V x) S = lookup S x" | +"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" + +lemma aval'_sound: "s : \\<^isub>f S \ aval a s : \(aval' a S)" +by (induction a) (auto simp: gamma_num' gamma_plus' \_st_def lookup_def) + +end + +text{* The for-clause (here and elsewhere) only serves the purpose of fixing +the name of the type parameter @{typ 'av} which would otherwise be renamed to +@{typ 'a}. *} + +locale Abs_Int = Gamma where \=\ for \ :: "'av::SL_top \ val set" +begin + +fun step' :: "'av st option \ 'av st option acom \ 'av st option acom" where +"step' S (SKIP {P}) = (SKIP {S})" | +"step' S (x ::= e {P}) = + x ::= e {case S of None \ None | Some S \ Some(update S x (aval' e S))}" | +"step' S (c1; c2) = step' S c1; step' (post c1) c2" | +"step' S (IF b THEN c1 ELSE c2 {P}) = + (let c1' = step' S c1; c2' = step' S c2 + in IF b THEN c1' ELSE c2' {post c1 \ post c2})" | +"step' S ({Inv} WHILE b DO c {P}) = + {S \ post c} WHILE b DO step' Inv c {Inv}" + +definition AI :: "com \ 'av st option acom option" where +"AI = lpfp\<^isub>c (step' \)" + + +lemma strip_step'[simp]: "strip(step' S c) = strip c" +by(induct c arbitrary: S) (simp_all add: Let_def) + + +text{* Soundness: *} + +lemma in_gamma_update: + "\ s : \\<^isub>f S; i : \ a \ \ s(x := i) : \\<^isub>f(update S x a)" +by(simp add: \_st_def lookup_update) + +text{* The soundness proofs are textually identical to the ones for the step +function operating on states as functions. *} + +lemma step_preserves_le: + "\ S \ \\<^isub>o S'; c \ \\<^isub>c c' \ \ step S c \ \\<^isub>c (step' S' c')" +proof(induction c arbitrary: c' S S') + case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) +next + case Assign thus ?case + by (fastforce simp: Assign_le map_acom_Assign intro: aval'_sound in_gamma_update + split: option.splits del:subsetD) +next + case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi) + by (metis le_post post_map_acom) +next + case (If b c1 c2 P) + then obtain c1' c2' P' where + "c' = IF b THEN c1' ELSE c2' {P'}" + "P \ \\<^isub>o P'" "c1 \ \\<^isub>c c1'" "c2 \ \\<^isub>c c2'" + by (fastforce simp: If_le map_acom_If) + moreover have "post c1 \ \\<^isub>o(post c1' \ post c2')" + by (metis (no_types) `c1 \ \\<^isub>c c1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom) + moreover have "post c2 \ \\<^isub>o(post c1' \ post c2')" + by (metis (no_types) `c2 \ \\<^isub>c c2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom) + ultimately show ?case using `S \ \\<^isub>o S'` by (simp add: If.IH subset_iff) +next + case (While I b c1 P) + then obtain c1' I' P' where + "c' = {I'} WHILE b DO c1' {P'}" + "I \ \\<^isub>o I'" "P \ \\<^isub>o P'" "c1 \ \\<^isub>c c1'" + by (fastforce simp: map_acom_While While_le) + moreover have "S \ post c1 \ \\<^isub>o (S' \ post c1')" + using `S \ \\<^isub>o S'` le_post[OF `c1 \ \\<^isub>c c1'`, simplified] + by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans) + ultimately show ?case by (simp add: While.IH subset_iff) +qed + +lemma AI_sound: "AI c = Some c' \ CS c \ \\<^isub>c c'" +proof(simp add: CS_def AI_def) + assume 1: "lpfp\<^isub>c (step' \) c = Some c'" + have 2: "step' \ c' \ c'" by(rule lpfpc_pfp[OF 1]) + have 3: "strip (\\<^isub>c (step' \ c')) = c" + by(simp add: strip_lpfpc[OF _ 1]) + have "lfp (step UNIV) c \ \\<^isub>c (step' \ c')" + proof(rule lfp_lowerbound[simplified,OF 3]) + show "step UNIV (\\<^isub>c (step' \ c')) \ \\<^isub>c (step' \ c')" + proof(rule step_preserves_le[OF _ _]) + show "UNIV \ \\<^isub>o \" by simp + show "\\<^isub>c (step' \ c') \ \\<^isub>c c'" by(rule mono_gamma_c[OF 2]) + qed + qed + from this 2 show "lfp (step UNIV) c \ \\<^isub>c c'" + by (blast intro: mono_gamma_c order_trans) +qed + +end + + +subsubsection "Monotonicity" + +locale Abs_Int_mono = Abs_Int + +assumes mono_plus': "a1 \ b1 \ a2 \ b2 \ plus' a1 a2 \ plus' b1 b2" +begin + +lemma mono_aval': "S \ S' \ aval' e S \ aval' e S'" +by(induction e) (auto simp: le_st_def lookup_def mono_plus') + +lemma mono_update: "a \ a' \ S \ S' \ update S x a \ update S' x a'" +by(auto simp add: le_st_def lookup_def update_def) + +lemma mono_step': "S \ S' \ c \ c' \ step' S c \ step' S' c'" +apply(induction c c' arbitrary: S S' rule: le_acom.induct) +apply (auto simp: Let_def mono_update mono_aval' mono_post le_join_disj + split: option.split) +done + +end + + +subsubsection "Ascending Chain Condition" + +hide_const (open) acc + +abbreviation "strict r == r \ -(r^-1)" +abbreviation "acc r == wf((strict r)^-1)" + +lemma strict_inv_image: "strict(inv_image r f) = inv_image (strict r) f" +by(auto simp: inv_image_def) + +lemma acc_inv_image: + "acc r \ acc (inv_image r f)" +by (metis converse_inv_image strict_inv_image wf_inv_image) + +text{* ACC for option type: *} + +lemma acc_option: assumes "acc {(x,y::'a::preord). x \ y}" +shows "acc {(x,y::'a::preord option). x \ y}" +proof(auto simp: wf_eq_minimal) + fix xo :: "'a option" and Qo assume "xo : Qo" + let ?Q = "{x. Some x \ Qo}" + show "\yo\Qo. \zo. yo \ zo \ ~ zo \ yo \ zo \ Qo" (is "\zo\Qo. ?P zo") + proof cases + assume "?Q = {}" + hence "?P None" by auto + moreover have "None \ Qo" using `?Q = {}` `xo : Qo` + by auto (metis not_Some_eq) + ultimately show ?thesis by blast + next + assume "?Q \ {}" + with assms show ?thesis + apply(auto simp: wf_eq_minimal) + apply(erule_tac x="?Q" in allE) + apply auto + apply(rule_tac x = "Some z" in bexI) + by auto + qed +qed + +text{* ACC for abstract states, via measure functions. *} + +(*FIXME mv*) +lemma setsum_strict_mono1: +fixes f :: "'a \ 'b::{comm_monoid_add, ordered_cancel_ab_semigroup_add}" +assumes "finite A" and "ALL x:A. f x \ g x" and "EX a:A. f a < g a" +shows "setsum f A < setsum g A" +proof- + from assms(3) obtain a where a: "a:A" "f a < g a" by blast + have "setsum f A = setsum f ((A-{a}) \ {a})" + by(simp add:insert_absorb[OF `a:A`]) + also have "\ = setsum f (A-{a}) + setsum f {a}" + using `finite A` by(subst setsum_Un_disjoint) auto + also have "setsum f (A-{a}) \ setsum g (A-{a})" + by(rule setsum_mono)(simp add: assms(2)) + also have "setsum f {a} < setsum g {a}" using a by simp + also have "setsum g (A - {a}) + setsum g {a} = setsum g((A-{a}) \ {a})" + using `finite A` by(subst setsum_Un_disjoint[symmetric]) auto + also have "\ = setsum g A" by(simp add:insert_absorb[OF `a:A`]) + finally show ?thesis by (metis add_right_mono add_strict_left_mono) +qed + +lemma measure_st: assumes "(strict{(x,y::'a::SL_top). x \ y})^-1 <= measure m" +and "\x y::'a::SL_top. x \ y \ y \ x \ m x = m y" +shows "(strict{(S,S'::'a::SL_top st). S \ S'})^-1 \ + measure(%fd. \x| x\set(dom fd) \ ~ \ \ fun fd x. m(fun fd x)+1)" +proof- + { fix S S' :: "'a st" assume "S \ S'" "~ S' \ S" + let ?X = "set(dom S)" let ?Y = "set(dom S')" + let ?f = "fun S" let ?g = "fun S'" + let ?X' = "{x:?X. ~ \ \ ?f x}" let ?Y' = "{y:?Y. ~ \ \ ?g y}" + from `S \ S'` have "ALL y:?Y'\?X. ?f y \ ?g y" + by(auto simp: le_st_def lookup_def) + hence 1: "ALL y:?Y'\?X. m(?g y)+1 \ m(?f y)+1" + using assms(1,2) by(fastforce) + from `~ S' \ S` obtain u where u: "u : ?X" "~ lookup S' u \ ?f u" + by(auto simp: le_st_def) + hence "u : ?X'" by simp (metis preord_class.le_trans top) + have "?Y'-?X = {}" using `S \ S'` by(fastforce simp: le_st_def lookup_def) + have "?Y'\?X <= ?X'" apply auto + apply (metis `S \ S'` le_st_def lookup_def preord_class.le_trans) + done + have "(\y\?Y'. m(?g y)+1) = (\y\(?Y'-?X) \ (?Y'\?X). m(?g y)+1)" + by (metis Un_Diff_Int) + also have "\ = (\y\?Y'\?X. m(?g y)+1)" + using `?Y'-?X = {}` by (metis Un_empty_left) + also have "\ < (\x\?X'. m(?f x)+1)" + proof cases + assume "u \ ?Y'" + hence "m(?g u) < m(?f u)" using assms(1) `S \ S'` u + by (fastforce simp: le_st_def lookup_def) + have "(\y\?Y'\?X. m(?g y)+1) < (\y\?Y'\?X. m(?f y)+1)" + using `u:?X` `u:?Y'` `m(?g u) < m(?f u)` + by(fastforce intro!: setsum_strict_mono1[OF _ 1]) + also have "\ \ (\y\?X'. m(?f y)+1)" + by(simp add: setsum_mono3[OF _ `?Y'\?X <= ?X'`]) + finally show ?thesis . + next + assume "u \ ?Y'" + with `?Y'\?X <= ?X'` have "?Y'\?X - {u} <= ?X' - {u}" by blast + have "(\y\?Y'\?X. m(?g y)+1) = (\y\?Y'\?X - {u}. m(?g y)+1)" + proof- + have "?Y'\?X = ?Y'\?X - {u}" using `u \ ?Y'` by auto + thus ?thesis by metis + qed + also have "\ < (\y\?Y'\?X-{u}. m(?g y)+1) + (\y\{u}. m(?f y)+1)" by simp + also have "(\y\?Y'\?X-{u}. m(?g y)+1) \ (\y\?Y'\?X-{u}. m(?f y)+1)" + using 1 by(blast intro: setsum_mono) + also have "\ \ (\y\?X'-{u}. m(?f y)+1)" + by(simp add: setsum_mono3[OF _ `?Y'\?X-{u} <= ?X'-{u}`]) + also have "\ + (\y\{u}. m(?f y)+1)= (\y\(?X'-{u}) \ {u}. m(?f y)+1)" + using `u:?X'` by(subst setsum_Un_disjoint[symmetric]) auto + also have "\ = (\x\?X'. m(?f x)+1)" + using `u : ?X'` by(simp add:insert_absorb) + finally show ?thesis by (blast intro: add_right_mono) + qed + finally have "(\y\?Y'. m(?g y)+1) < (\x\?X'. m(?f x)+1)" . + } thus ?thesis by(auto simp add: measure_def inv_image_def) +qed + +text{* ACC for acom. First the ordering on acom is related to an ordering on +lists of annotations. *} + +(* FIXME mv and add [simp] *) +lemma listrel_Cons_iff: + "(x#xs, y#ys) : listrel r \ (x,y) \ r \ (xs,ys) \ listrel r" +by (blast intro:listrel.Cons) + +lemma listrel_app: "(xs1,ys1) : listrel r \ (xs2,ys2) : listrel r + \ (xs1@xs2, ys1@ys2) : listrel r" +by(auto simp add: listrel_iff_zip) + +lemma listrel_app_same_size: "size xs1 = size ys1 \ size xs2 = size ys2 \ + (xs1@xs2, ys1@ys2) : listrel r \ + (xs1,ys1) : listrel r \ (xs2,ys2) : listrel r" +by(auto simp add: listrel_iff_zip) + +lemma listrel_converse: "listrel(r^-1) = (listrel r)^-1" +proof- + { fix xs ys + have "(xs,ys) : listrel(r^-1) \ (ys,xs) : listrel r" + apply(induct xs arbitrary: ys) + apply (fastforce simp: listrel.Nil) + apply (fastforce simp: listrel_Cons_iff) + done + } thus ?thesis by auto +qed + +(* It would be nice to get rid of refl & trans and build them into the proof *) +lemma acc_listrel: fixes r :: "('a*'a)set" assumes "refl r" and "trans r" +and "acc r" shows "acc (listrel r - {([],[])})" +proof- + have refl: "!!x. (x,x) : r" using `refl r` unfolding refl_on_def by blast + have trans: "!!x y z. (x,y) : r \ (y,z) : r \ (x,z) : r" + using `trans r` unfolding trans_def by blast + from assms(3) obtain mx :: "'a set \ 'a" where + mx: "!!S x. x:S \ mx S : S \ (\y. (mx S,y) : strict r \ y \ S)" + by(simp add: wf_eq_minimal) metis + let ?R = "listrel r - {([], [])}" + { fix Q and xs :: "'a list" + have "xs \ Q \ \ys. ys\Q \ (\zs. (ys, zs) \ strict ?R \ zs \ Q)" + (is "_ \ \ys. ?P Q ys") + proof(induction xs arbitrary: Q rule: length_induct) + case (1 xs) + { have "!!ys Q. size ys < size xs \ ys : Q \ EX ms. ?P Q ms" + using "1.IH" by blast + } note IH = this + show ?case + proof(cases xs) + case Nil with `xs : Q` have "?P Q []" by auto + thus ?thesis by blast + next + case (Cons x ys) + let ?Q1 = "{a. \bs. size bs = size ys \ a#bs : Q}" + have "x : ?Q1" using `xs : Q` Cons by auto + from mx[OF this] obtain m1 where + 1: "m1 \ ?Q1 \ (\y. (m1,y) \ strict r \ y \ ?Q1)" by blast + then obtain ms1 where "size ms1 = size ys" "m1#ms1 : Q" by blast+ + hence "size ms1 < size xs" using Cons by auto + let ?Q2 = "{bs. \m1'. (m1',m1):r \ (m1,m1'):r \ m1'#bs : Q \ size bs = size ms1}" + have "ms1 : ?Q2" using `m1#ms1 : Q` by(blast intro: refl) + from IH[OF `size ms1 < size xs` this] + obtain ms where 2: "?P ?Q2 ms" by auto + then obtain m1' where m1': "(m1',m1) : r \ (m1,m1') : r \ m1'#ms : Q" + by blast + hence "\ab. (m1'#ms,ab) : strict ?R \ ab \ Q" using 1 2 + apply (auto simp: listrel_Cons_iff) + apply (metis `length ms1 = length ys` listrel_eq_len trans) + by (metis `length ms1 = length ys` listrel_eq_len trans) + with m1' show ?thesis by blast + qed + qed + } + thus ?thesis unfolding wf_eq_minimal by (metis converse_iff) +qed + + +fun annos :: "'a acom \ 'a list" where +"annos (SKIP {a}) = [a]" | +"annos (x::=e {a}) = [a]" | +"annos (c1;c2) = annos c1 @ annos c2" | +"annos (IF b THEN c1 ELSE c2 {a}) = a # annos c1 @ annos c2" | +"annos ({i} WHILE b DO c {a}) = i # a # annos c" + +lemma size_annos_same: "strip c1 = strip c2 \ size(annos c1) = size(annos c2)" +apply(induct c2 arbitrary: c1) +apply (auto simp: strip_eq_SKIP strip_eq_Assign strip_eq_Semi strip_eq_If strip_eq_While) +done + +lemmas size_annos_same2 = eqTrueI[OF size_annos_same] + +lemma set_annos_anno: "set (annos (anno a c)) = {a}" +by(induction c)(auto) + +lemma le_iff_le_annos: "c1 \ c2 \ + (annos c1, annos c2) : listrel{(x,y). x \ y} \ strip c1 = strip c2" +apply(induct c1 c2 rule: le_acom.induct) +apply (auto simp: listrel.Nil listrel_Cons_iff listrel_app size_annos_same2) +apply (metis listrel_app_same_size size_annos_same)+ +done + +lemma le_acom_subset_same_annos: + "(strict{(c,c'::'a::preord acom). c \ c'})^-1 \ + (strict(inv_image (listrel{(a,a'::'a). a \ a'} - {([],[])}) annos))^-1" +by(auto simp: le_iff_le_annos) + +lemma acc_acom: "acc {(a,a'::'a::preord). a \ a'} \ + acc {(c,c'::'a acom). c \ c'}" +apply(rule wf_subset[OF _ le_acom_subset_same_annos]) +apply(rule acc_inv_image[OF acc_listrel]) +apply(auto simp: refl_on_def trans_def intro: le_trans) +done + +text{* Termination of the fixed-point finders, assuming monotone functions: *} + +lemma pfp_termination: +fixes x0 :: "'a::preord" +assumes mono: "\x y. x \ y \ f x \ f y" and "acc {(x::'a,y). x \ y}" +and "x0 \ f x0" shows "EX x. pfp f x0 = Some x" +proof(simp add: pfp_def, rule wf_while_option_Some[where P = "%x. x \ f x"]) + show "wf {(x, s). (s \ f s \ \ f s \ s) \ x = f s}" + by(rule wf_subset[OF assms(2)]) auto +next + show "x0 \ f x0" by(rule assms) +next + fix x assume "x \ f x" thus "f x \ f(f x)" by(rule mono) +qed + +lemma lpfpc_termination: + fixes f :: "(('a::SL_top)option acom \ 'a option acom)" + assumes "acc {(x::'a,y). x \ y}" and "\x y. x \ y \ f x \ f y" + and "\c. strip(f c) = strip c" + shows "\c'. lpfp\<^isub>c f c = Some c'" +unfolding lpfp\<^isub>c_def +apply(rule pfp_termination) + apply(erule assms(2)) + apply(rule acc_acom[OF acc_option[OF assms(1)]]) +apply(simp add: bot_acom assms(3)) +done + +context Abs_Int_mono +begin + +lemma AI_Some_measure: +assumes "(strict{(x,y::'a). x \ y})^-1 <= measure m" +and "\x y::'a. x \ y \ y \ x \ m x = m y" +shows "\c'. AI c = Some c'" +unfolding AI_def +apply(rule lpfpc_termination) +apply(rule wf_subset[OF wf_measure measure_st[OF assms]]) +apply(erule mono_step'[OF le_refl]) +apply(rule strip_step') +done + +end + +end diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_Int_ITP/Abs_Int1_const_ITP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_const_ITP.thy Thu Apr 19 17:49:02 2012 +0200 @@ -0,0 +1,139 @@ +(* Author: Tobias Nipkow *) + +theory Abs_Int1_const_ITP +imports Abs_Int1_ITP Abs_Int_Tests +begin + +subsection "Constant Propagation" + +datatype const = Const val | Any + +fun \_const where +"\_const (Const n) = {n}" | +"\_const (Any) = UNIV" + +fun plus_const where +"plus_const (Const m) (Const n) = Const(m+n)" | +"plus_const _ _ = Any" + +lemma plus_const_cases: "plus_const a1 a2 = + (case (a1,a2) of (Const m, Const n) \ Const(m+n) | _ \ Any)" +by(auto split: prod.split const.split) + +instantiation const :: SL_top +begin + +fun le_const where +"_ \ Any = True" | +"Const n \ Const m = (n=m)" | +"Any \ Const _ = False" + +fun join_const where +"Const m \ Const n = (if n=m then Const m else Any)" | +"_ \ _ = Any" + +definition "\ = Any" + +instance +proof + case goal1 thus ?case by (cases x) simp_all +next + case goal2 thus ?case by(cases z, cases y, cases x, simp_all) +next + case goal3 thus ?case by(cases x, cases y, simp_all) +next + case goal4 thus ?case by(cases y, cases x, simp_all) +next + case goal5 thus ?case by(cases z, cases y, cases x, simp_all) +next + case goal6 thus ?case by(simp add: Top_const_def) +qed + +end + + +interpretation Val_abs +where \ = \_const and num' = Const and plus' = plus_const +proof + case goal1 thus ?case + by(cases a, cases b, simp, simp, cases b, simp, simp) +next + case goal2 show ?case by(simp add: Top_const_def) +next + case goal3 show ?case by simp +next + case goal4 thus ?case + by(auto simp: plus_const_cases split: const.split) +qed + +interpretation Abs_Int +where \ = \_const and num' = Const and plus' = plus_const +defines AI_const is AI and step_const is step' and aval'_const is aval' +.. + + +subsubsection "Tests" + +value "show_acom (((step_const \)^^0) (\\<^sub>c test1_const))" +value "show_acom (((step_const \)^^1) (\\<^sub>c test1_const))" +value "show_acom (((step_const \)^^2) (\\<^sub>c test1_const))" +value "show_acom (((step_const \)^^3) (\\<^sub>c test1_const))" +value "show_acom_opt (AI_const test1_const)" + +value "show_acom_opt (AI_const test2_const)" +value "show_acom_opt (AI_const test3_const)" + +value "show_acom (((step_const \)^^0) (\\<^sub>c test4_const))" +value "show_acom (((step_const \)^^1) (\\<^sub>c test4_const))" +value "show_acom (((step_const \)^^2) (\\<^sub>c test4_const))" +value "show_acom (((step_const \)^^3) (\\<^sub>c test4_const))" +value "show_acom_opt (AI_const test4_const)" + +value "show_acom (((step_const \)^^0) (\\<^sub>c test5_const))" +value "show_acom (((step_const \)^^1) (\\<^sub>c test5_const))" +value "show_acom (((step_const \)^^2) (\\<^sub>c test5_const))" +value "show_acom (((step_const \)^^3) (\\<^sub>c test5_const))" +value "show_acom (((step_const \)^^4) (\\<^sub>c test5_const))" +value "show_acom (((step_const \)^^5) (\\<^sub>c test5_const))" +value "show_acom_opt (AI_const test5_const)" + +value "show_acom (((step_const \)^^0) (\\<^sub>c test6_const))" +value "show_acom (((step_const \)^^1) (\\<^sub>c test6_const))" +value "show_acom (((step_const \)^^2) (\\<^sub>c test6_const))" +value "show_acom (((step_const \)^^3) (\\<^sub>c test6_const))" +value "show_acom (((step_const \)^^4) (\\<^sub>c test6_const))" +value "show_acom (((step_const \)^^5) (\\<^sub>c test6_const))" +value "show_acom (((step_const \)^^6) (\\<^sub>c test6_const))" +value "show_acom (((step_const \)^^7) (\\<^sub>c test6_const))" +value "show_acom (((step_const \)^^8) (\\<^sub>c test6_const))" +value "show_acom (((step_const \)^^9) (\\<^sub>c test6_const))" +value "show_acom (((step_const \)^^10) (\\<^sub>c test6_const))" +value "show_acom (((step_const \)^^11) (\\<^sub>c test6_const))" +value "show_acom_opt (AI_const test6_const)" + + +text{* Monotonicity: *} + +interpretation Abs_Int_mono +where \ = \_const and num' = Const and plus' = plus_const +proof + case goal1 thus ?case + by(auto simp: plus_const_cases split: const.split) +qed + +text{* Termination: *} + +definition "m_const x = (case x of Const _ \ 1 | Any \ 0)" + +lemma measure_const: + "(strict{(x::const,y). x \ y})^-1 \ measure m_const" +by(auto simp: m_const_def split: const.splits) + +lemma measure_const_eq: + "\ x y::const. x \ y \ y \ x \ m_const x = m_const y" +by(auto simp: m_const_def split: const.splits) + +lemma "EX c'. AI_const c = Some c'" +by(rule AI_Some_measure[OF measure_const measure_const_eq]) + +end diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_Int_ITP/Abs_Int1_parity_ITP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_parity_ITP.thy Thu Apr 19 17:49:02 2012 +0200 @@ -0,0 +1,168 @@ +(* Author: Tobias Nipkow *) + +theory Abs_Int1_parity_ITP +imports Abs_Int1_ITP +begin + +subsection "Parity Analysis" + +datatype parity = Even | Odd | Either + +text{* Instantiation of class @{class preord} with type @{typ parity}: *} + +instantiation parity :: preord +begin + +text{* First the definition of the interface function @{text"\"}. Note that +the header of the definition must refer to the ascii name @{const le} of the +constants as @{text le_parity} and the definition is named @{text +le_parity_def}. Inside the definition the symbolic names can be used. *} + +definition le_parity where +"x \ y = (y = Either \ x=y)" + +text{* Now the instance proof, i.e.\ the proof that the definition fulfills +the axioms (assumptions) of the class. The initial proof-step generates the +necessary proof obligations. *} + +instance +proof + fix x::parity show "x \ x" by(auto simp: le_parity_def) +next + fix x y z :: parity assume "x \ y" "y \ z" thus "x \ z" + by(auto simp: le_parity_def) +qed + +end + +text{* Instantiation of class @{class SL_top} with type @{typ parity}: *} + +instantiation parity :: SL_top +begin + + +definition join_parity where +"x \ y = (if x \ y then y else if y \ x then x else Either)" + +definition Top_parity where +"\ = Either" + +text{* Now the instance proof. This time we take a lazy shortcut: we do not +write out the proof obligations but use the @{text goali} primitive to refer +to the assumptions of subgoal i and @{text "case?"} to refer to the +conclusion of subgoal i. The class axioms are presented in the same order as +in the class definition. *} + +instance +proof + case goal1 (*join1*) show ?case by(auto simp: le_parity_def join_parity_def) +next + case goal2 (*join2*) show ?case by(auto simp: le_parity_def join_parity_def) +next + case goal3 (*join least*) thus ?case by(auto simp: le_parity_def join_parity_def) +next + case goal4 (*Top*) show ?case by(auto simp: le_parity_def Top_parity_def) +qed + +end + + +text{* Now we define the functions used for instantiating the abstract +interpretation locales. Note that the Isabelle terminology is +\emph{interpretation}, not \emph{instantiation} of locales, but we use +instantiation to avoid confusion with abstract interpretation. *} + +fun \_parity :: "parity \ val set" where +"\_parity Even = {i. i mod 2 = 0}" | +"\_parity Odd = {i. i mod 2 = 1}" | +"\_parity Either = UNIV" + +fun num_parity :: "val \ parity" where +"num_parity i = (if i mod 2 = 0 then Even else Odd)" + +fun plus_parity :: "parity \ parity \ parity" where +"plus_parity Even Even = Even" | +"plus_parity Odd Odd = Even" | +"plus_parity Even Odd = Odd" | +"plus_parity Odd Even = Odd" | +"plus_parity Either y = Either" | +"plus_parity x Either = Either" + +text{* First we instantiate the abstract value interface and prove that the +functions on type @{typ parity} have all the necessary properties: *} + +interpretation Val_abs +where \ = \_parity and num' = num_parity and plus' = plus_parity +proof txt{* of the locale axioms *} + fix a b :: parity + assume "a \ b" thus "\_parity a \ \_parity b" + by(auto simp: le_parity_def) +next txt{* The rest in the lazy, implicit way *} + case goal2 show ?case by(auto simp: Top_parity_def) +next + case goal3 show ?case by auto +next + txt{* Warning: this subproof refers to the names @{text a1} and @{text a2} + from the statement of the axiom. *} + case goal4 thus ?case + proof(cases a1 a2 rule: parity.exhaust[case_product parity.exhaust]) + qed (auto simp add:mod_add_eq) +qed + +text{* Instantiating the abstract interpretation locale requires no more +proofs (they happened in the instatiation above) but delivers the +instantiated abstract interpreter which we call AI: *} + +interpretation Abs_Int +where \ = \_parity and num' = num_parity and plus' = plus_parity +defines aval_parity is aval' and step_parity is step' and AI_parity is AI +.. + + +subsubsection "Tests" + +definition "test1_parity = + ''x'' ::= N 1; + WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 2)" + +value "show_acom_opt (AI_parity test1_parity)" + +definition "test2_parity = + ''x'' ::= N 1; + WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 3)" + +value "show_acom ((step_parity \ ^^1) (anno None test2_parity))" +value "show_acom ((step_parity \ ^^2) (anno None test2_parity))" +value "show_acom ((step_parity \ ^^3) (anno None test2_parity))" +value "show_acom ((step_parity \ ^^4) (anno None test2_parity))" +value "show_acom ((step_parity \ ^^5) (anno None test2_parity))" +value "show_acom_opt (AI_parity test2_parity)" + + +subsubsection "Termination" + +interpretation Abs_Int_mono +where \ = \_parity and num' = num_parity and plus' = plus_parity +proof + case goal1 thus ?case + proof(cases a1 a2 b1 b2 + rule: parity.exhaust[case_product parity.exhaust[case_product parity.exhaust[case_product parity.exhaust]]]) (* FIXME - UGLY! *) + qed (auto simp add:le_parity_def) +qed + + +definition m_parity :: "parity \ nat" where +"m_parity x = (if x=Either then 0 else 1)" + +lemma measure_parity: + "(strict{(x::parity,y). x \ y})^-1 \ measure m_parity" +by(auto simp add: m_parity_def le_parity_def) + +lemma measure_parity_eq: + "\x y::parity. x \ y \ y \ x \ m_parity x = m_parity y" +by(auto simp add: m_parity_def le_parity_def) + +lemma AI_parity_Some: "\c'. AI_parity c = Some c'" +by(rule AI_Some_measure[OF measure_parity measure_parity_eq]) + +end diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ITP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ITP.thy Thu Apr 19 17:49:02 2012 +0200 @@ -0,0 +1,358 @@ +(* Author: Tobias Nipkow *) + +theory Abs_Int2_ITP +imports Abs_Int1_ITP Vars +begin + +instantiation prod :: (preord,preord) preord +begin + +definition "le_prod p1 p2 = (fst p1 \ fst p2 \ snd p1 \ snd p2)" + +instance +proof + case goal1 show ?case by(simp add: le_prod_def) +next + case goal2 thus ?case unfolding le_prod_def by(metis le_trans) +qed + +end + +instantiation com :: vars +begin + +fun vars_com :: "com \ vname set" where +"vars com.SKIP = {}" | +"vars (x::=e) = {x} \ vars e" | +"vars (c1;c2) = vars c1 \ vars c2" | +"vars (IF b THEN c1 ELSE c2) = vars b \ vars c1 \ vars c2" | +"vars (WHILE b DO c) = vars b \ vars c" + +instance .. + +end + +lemma finite_avars: "finite(vars(a::aexp))" +by(induction a) simp_all + +lemma finite_bvars: "finite(vars(b::bexp))" +by(induction b) (simp_all add: finite_avars) + +lemma finite_cvars: "finite(vars(c::com))" +by(induction c) (simp_all add: finite_avars finite_bvars) + + +subsection "Backward Analysis of Expressions" + +hide_const bot + +class L_top_bot = SL_top + +fixes meet :: "'a \ 'a \ 'a" (infixl "\" 65) +and bot :: "'a" ("\") +assumes meet_le1 [simp]: "x \ y \ x" +and meet_le2 [simp]: "x \ y \ y" +and meet_greatest: "x \ y \ x \ z \ x \ y \ z" +assumes bot[simp]: "\ \ x" +begin + +lemma mono_meet: "x \ x' \ y \ y' \ x \ y \ x' \ y'" +by (metis meet_greatest meet_le1 meet_le2 le_trans) + +end + +locale Val_abs1_gamma = + Gamma where \ = \ for \ :: "'av::L_top_bot \ val set" + +assumes inter_gamma_subset_gamma_meet: + "\ a1 \ \ a2 \ \(a1 \ a2)" +and gamma_Bot[simp]: "\ \ = {}" +begin + +lemma in_gamma_meet: "x : \ a1 \ x : \ a2 \ x : \(a1 \ a2)" +by (metis IntI inter_gamma_subset_gamma_meet set_mp) + +lemma gamma_meet[simp]: "\(a1 \ a2) = \ a1 \ \ a2" +by (metis equalityI inter_gamma_subset_gamma_meet le_inf_iff mono_gamma meet_le1 meet_le2) + +end + + +locale Val_abs1 = + Val_abs1_gamma where \ = \ + for \ :: "'av::L_top_bot \ val set" + +fixes test_num' :: "val \ 'av \ bool" +and filter_plus' :: "'av \ 'av \ 'av \ 'av * 'av" +and filter_less' :: "bool \ 'av \ 'av \ 'av * 'av" +assumes test_num': "test_num' n a = (n : \ a)" +and filter_plus': "filter_plus' a a1 a2 = (b1,b2) \ + n1 : \ a1 \ n2 : \ a2 \ n1+n2 : \ a \ n1 : \ b1 \ n2 : \ b2" +and filter_less': "filter_less' (n1 + n1 : \ a1 \ n2 : \ a2 \ n1 : \ b1 \ n2 : \ b2" + + +locale Abs_Int1 = + Val_abs1 where \ = \ for \ :: "'av::L_top_bot \ val set" +begin + +lemma in_gamma_join_UpI: "s : \\<^isub>o S1 \ s : \\<^isub>o S2 \ s : \\<^isub>o(S1 \ S2)" +by (metis (no_types) join_ge1 join_ge2 mono_gamma_o set_rev_mp) + +fun aval'' :: "aexp \ 'av st option \ 'av" where +"aval'' e None = \" | +"aval'' e (Some sa) = aval' e sa" + +lemma aval''_sound: "s : \\<^isub>o S \ aval a s : \(aval'' a S)" +by(cases S)(simp add: aval'_sound)+ + +subsubsection "Backward analysis" + +fun afilter :: "aexp \ 'av \ 'av st option \ 'av st option" where +"afilter (N n) a S = (if test_num' n a then S else None)" | +"afilter (V x) a S = (case S of None \ None | Some S \ + let a' = lookup S x \ a in + if a' \ \ then None else Some(update S x a'))" | +"afilter (Plus e1 e2) a S = + (let (a1,a2) = filter_plus' a (aval'' e1 S) (aval'' e2 S) + in afilter e1 a1 (afilter e2 a2 S))" + +text{* The test for @{const bot} in the @{const V}-case is important: @{const +bot} indicates that a variable has no possible values, i.e.\ that the current +program point is unreachable. But then the abstract state should collapse to +@{const None}. Put differently, we maintain the invariant that in an abstract +state of the form @{term"Some s"}, all variables are mapped to non-@{const +bot} values. Otherwise the (pointwise) join of two abstract states, one of +which contains @{const bot} values, may produce too large a result, thus +making the analysis less precise. *} + + +fun bfilter :: "bexp \ bool \ 'av st option \ 'av st option" where +"bfilter (Bc v) res S = (if v=res then S else None)" | +"bfilter (Not b) res S = bfilter b (\ res) S" | +"bfilter (And b1 b2) res S = + (if res then bfilter b1 True (bfilter b2 True S) + else bfilter b1 False S \ bfilter b2 False S)" | +"bfilter (Less e1 e2) res S = + (let (res1,res2) = filter_less' res (aval'' e1 S) (aval'' e2 S) + in afilter e1 res1 (afilter e2 res2 S))" + +lemma afilter_sound: "s : \\<^isub>o S \ aval e s : \ a \ s : \\<^isub>o (afilter e a S)" +proof(induction e arbitrary: a S) + case N thus ?case by simp (metis test_num') +next + case (V x) + obtain S' where "S = Some S'" and "s : \\<^isub>f S'" using `s : \\<^isub>o S` + by(auto simp: in_gamma_option_iff) + moreover hence "s x : \ (lookup S' x)" by(simp add: \_st_def) + moreover have "s x : \ a" using V by simp + ultimately show ?case using V(1) + by(simp add: lookup_update Let_def \_st_def) + (metis mono_gamma emptyE in_gamma_meet gamma_Bot subset_empty) +next + case (Plus e1 e2) thus ?case + using filter_plus'[OF _ aval''_sound[OF Plus(3)] aval''_sound[OF Plus(3)]] + by (auto split: prod.split) +qed + +lemma bfilter_sound: "s : \\<^isub>o S \ bv = bval b s \ s : \\<^isub>o(bfilter b bv S)" +proof(induction b arbitrary: S bv) + case Bc thus ?case by simp +next + case (Not b) thus ?case by simp +next + case (And b1 b2) thus ?case by(fastforce simp: in_gamma_join_UpI) +next + case (Less e1 e2) thus ?case + by (auto split: prod.split) + (metis afilter_sound filter_less' aval''_sound Less) +qed + + +fun step' :: "'av st option \ 'av st option acom \ 'av st option acom" + where +"step' S (SKIP {P}) = (SKIP {S})" | +"step' S (x ::= e {P}) = + x ::= e {case S of None \ None | Some S \ Some(update S x (aval' e S))}" | +"step' S (c1; c2) = step' S c1; step' (post c1) c2" | +"step' S (IF b THEN c1 ELSE c2 {P}) = + (let c1' = step' (bfilter b True S) c1; c2' = step' (bfilter b False S) c2 + in IF b THEN c1' ELSE c2' {post c1 \ post c2})" | +"step' S ({Inv} WHILE b DO c {P}) = + {S \ post c} + WHILE b DO step' (bfilter b True Inv) c + {bfilter b False Inv}" + +definition AI :: "com \ 'av st option acom option" where +"AI = lpfp\<^isub>c (step' \)" + +lemma strip_step'[simp]: "strip(step' S c) = strip c" +by(induct c arbitrary: S) (simp_all add: Let_def) + + +subsubsection "Soundness" + +lemma in_gamma_update: + "\ s : \\<^isub>f S; i : \ a \ \ s(x := i) : \\<^isub>f(update S x a)" +by(simp add: \_st_def lookup_update) + +lemma step_preserves_le: + "\ S \ \\<^isub>o S'; cs \ \\<^isub>c ca \ \ step S cs \ \\<^isub>c (step' S' ca)" +proof(induction cs arbitrary: ca S S') + case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) +next + case Assign thus ?case + by (fastforce simp: Assign_le map_acom_Assign intro: aval'_sound in_gamma_update + split: option.splits del:subsetD) +next + case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi) + by (metis le_post post_map_acom) +next + case (If b cs1 cs2 P) + then obtain ca1 ca2 Pa where + "ca= IF b THEN ca1 ELSE ca2 {Pa}" + "P \ \\<^isub>o Pa" "cs1 \ \\<^isub>c ca1" "cs2 \ \\<^isub>c ca2" + by (fastforce simp: If_le map_acom_If) + moreover have "post cs1 \ \\<^isub>o(post ca1 \ post ca2)" + by (metis (no_types) `cs1 \ \\<^isub>c ca1` join_ge1 le_post mono_gamma_o order_trans post_map_acom) + moreover have "post cs2 \ \\<^isub>o(post ca1 \ post ca2)" + by (metis (no_types) `cs2 \ \\<^isub>c ca2` join_ge2 le_post mono_gamma_o order_trans post_map_acom) + ultimately show ?case using `S \ \\<^isub>o S'` + by (simp add: If.IH subset_iff bfilter_sound) +next + case (While I b cs1 P) + then obtain ca1 Ia Pa where + "ca = {Ia} WHILE b DO ca1 {Pa}" + "I \ \\<^isub>o Ia" "P \ \\<^isub>o Pa" "cs1 \ \\<^isub>c ca1" + by (fastforce simp: map_acom_While While_le) + moreover have "S \ post cs1 \ \\<^isub>o (S' \ post ca1)" + using `S \ \\<^isub>o S'` le_post[OF `cs1 \ \\<^isub>c ca1`, simplified] + by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans) + ultimately show ?case by (simp add: While.IH subset_iff bfilter_sound) +qed + +lemma AI_sound: "AI c = Some c' \ CS c \ \\<^isub>c c'" +proof(simp add: CS_def AI_def) + assume 1: "lpfp\<^isub>c (step' \) c = Some c'" + have 2: "step' \ c' \ c'" by(rule lpfpc_pfp[OF 1]) + have 3: "strip (\\<^isub>c (step' \ c')) = c" + by(simp add: strip_lpfpc[OF _ 1]) + have "lfp (step UNIV) c \ \\<^isub>c (step' \ c')" + proof(rule lfp_lowerbound[simplified,OF 3]) + show "step UNIV (\\<^isub>c (step' \ c')) \ \\<^isub>c (step' \ c')" + proof(rule step_preserves_le[OF _ _]) + show "UNIV \ \\<^isub>o \" by simp + show "\\<^isub>c (step' \ c') \ \\<^isub>c c'" by(rule mono_gamma_c[OF 2]) + qed + qed + from this 2 show "lfp (step UNIV) c \ \\<^isub>c c'" + by (blast intro: mono_gamma_c order_trans) +qed + + +subsubsection "Commands over a set of variables" + +text{* Key invariant: the domains of all abstract states are subsets of the +set of variables of the program. *} + +definition "domo S = (case S of None \ {} | Some S' \ set(dom S'))" + +definition Com :: "vname set \ 'a st option acom set" where +"Com X = {c. \S \ set(annos c). domo S \ X}" + +lemma domo_Top[simp]: "domo \ = {}" +by(simp add: domo_def Top_st_def Top_option_def) + +lemma bot_acom_Com[simp]: "\\<^sub>c c \ Com X" +by(simp add: bot_acom_def Com_def domo_def set_annos_anno) + +lemma post_in_annos: "post c : set(annos c)" +by(induction c) simp_all + +lemma domo_join: "domo (S \ T) \ domo S \ domo T" +by(auto simp: domo_def join_st_def split: option.split) + +lemma domo_afilter: "vars a \ X \ domo S \ X \ domo(afilter a i S) \ X" +apply(induction a arbitrary: i S) +apply(simp add: domo_def) +apply(simp add: domo_def Let_def update_def lookup_def split: option.splits) +apply blast +apply(simp split: prod.split) +done + +lemma domo_bfilter: "vars b \ X \ domo S \ X \ domo(bfilter b bv S) \ X" +apply(induction b arbitrary: bv S) +apply(simp add: domo_def) +apply(simp) +apply(simp) +apply rule +apply (metis le_sup_iff subset_trans[OF domo_join]) +apply(simp split: prod.split) +by (metis domo_afilter) + +lemma step'_Com: + "domo S \ X \ vars(strip c) \ X \ c : Com X \ step' S c : Com X" +apply(induction c arbitrary: S) +apply(simp add: Com_def) +apply(simp add: Com_def domo_def update_def split: option.splits) +apply(simp (no_asm_use) add: Com_def ball_Un) +apply (metis post_in_annos) +apply(simp (no_asm_use) add: Com_def ball_Un) +apply rule +apply (metis Un_assoc domo_join order_trans post_in_annos subset_Un_eq) +apply (metis domo_bfilter) +apply(simp (no_asm_use) add: Com_def) +apply rule +apply (metis domo_join le_sup_iff post_in_annos subset_trans) +apply rule +apply (metis domo_bfilter) +by (metis domo_bfilter) + +end + + +subsubsection "Monotonicity" + +locale Abs_Int1_mono = Abs_Int1 + +assumes mono_plus': "a1 \ b1 \ a2 \ b2 \ plus' a1 a2 \ plus' b1 b2" +and mono_filter_plus': "a1 \ b1 \ a2 \ b2 \ r \ r' \ + filter_plus' r a1 a2 \ filter_plus' r' b1 b2" +and mono_filter_less': "a1 \ b1 \ a2 \ b2 \ + filter_less' bv a1 a2 \ filter_less' bv b1 b2" +begin + +lemma mono_aval': "S \ S' \ aval' e S \ aval' e S'" +by(induction e) (auto simp: le_st_def lookup_def mono_plus') + +lemma mono_aval'': "S \ S' \ aval'' e S \ aval'' e S'" +apply(cases S) + apply simp +apply(cases S') + apply simp +by (simp add: mono_aval') + +lemma mono_afilter: "r \ r' \ S \ S' \ afilter e r S \ afilter e r' S'" +apply(induction e arbitrary: r r' S S') +apply(auto simp: test_num' Let_def split: option.splits prod.splits) +apply (metis mono_gamma subsetD) +apply(drule_tac x = "list" in mono_lookup) +apply (metis mono_meet le_trans) +apply (metis mono_meet mono_lookup mono_update) +apply(metis mono_aval'' mono_filter_plus'[simplified le_prod_def] fst_conv snd_conv) +done + +lemma mono_bfilter: "S \ S' \ bfilter b r S \ bfilter b r S'" +apply(induction b arbitrary: r S S') +apply(auto simp: le_trans[OF _ join_ge1] le_trans[OF _ join_ge2] split: prod.splits) +apply(metis mono_aval'' mono_afilter mono_filter_less'[simplified le_prod_def] fst_conv snd_conv) +done + +lemma mono_step': "S \ S' \ c \ c' \ step' S c \ step' S' c'" +apply(induction c c' arbitrary: S S' rule: le_acom.induct) +apply (auto simp: mono_post mono_bfilter mono_update mono_aval' Let_def le_join_disj + split: option.split) +done + +lemma mono_step'2: "mono (step' S)" +by(simp add: mono_def mono_step'[OF le_refl]) + +end + +end diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ivl_ITP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ivl_ITP.thy Thu Apr 19 17:49:02 2012 +0200 @@ -0,0 +1,285 @@ +(* Author: Tobias Nipkow *) + +theory Abs_Int2_ivl_ITP +imports Abs_Int2_ITP Abs_Int_Tests +begin + +subsection "Interval Analysis" + +datatype ivl = I "int option" "int option" + +definition "\_ivl i = (case i of + I (Some l) (Some h) \ {l..h} | + I (Some l) None \ {l..} | + I None (Some h) \ {..h} | + I None None \ UNIV)" + +abbreviation I_Some_Some :: "int \ int \ ivl" ("{_\_}") where +"{lo\hi} == I (Some lo) (Some hi)" +abbreviation I_Some_None :: "int \ ivl" ("{_\}") where +"{lo\} == I (Some lo) None" +abbreviation I_None_Some :: "int \ ivl" ("{\_}") where +"{\hi} == I None (Some hi)" +abbreviation I_None_None :: "ivl" ("{\}") where +"{\} == I None None" + +definition "num_ivl n = {n\n}" + +fun in_ivl :: "int \ ivl \ bool" where +"in_ivl k (I (Some l) (Some h)) \ l \ k \ k \ h" | +"in_ivl k (I (Some l) None) \ l \ k" | +"in_ivl k (I None (Some h)) \ k \ h" | +"in_ivl k (I None None) \ True" + +instantiation option :: (plus)plus +begin + +fun plus_option where +"Some x + Some y = Some(x+y)" | +"_ + _ = None" + +instance .. + +end + +definition empty where "empty = {1\0}" + +fun is_empty where +"is_empty {l\h} = (h (case h of Some h \ h False) | None \ False)" +by(auto split:option.split) + +lemma [simp]: "is_empty i \ \_ivl i = {}" +by(auto simp add: \_ivl_def split: ivl.split option.split) + +definition "plus_ivl i1 i2 = (if is_empty i1 | is_empty i2 then empty else + case (i1,i2) of (I l1 h1, I l2 h2) \ I (l1+l2) (h1+h2))" + +instantiation ivl :: SL_top +begin + +definition le_option :: "bool \ int option \ int option \ bool" where +"le_option pos x y = + (case x of (Some i) \ (case y of Some j \ i\j | None \ pos) + | None \ (case y of Some j \ \pos | None \ True))" + +fun le_aux where +"le_aux (I l1 h1) (I l2 h2) = (le_option False l2 l1 & le_option True h1 h2)" + +definition le_ivl where +"i1 \ i2 = + (if is_empty i1 then True else + if is_empty i2 then False else le_aux i1 i2)" + +definition min_option :: "bool \ int option \ int option \ int option" where +"min_option pos o1 o2 = (if le_option pos o1 o2 then o1 else o2)" + +definition max_option :: "bool \ int option \ int option \ int option" where +"max_option pos o1 o2 = (if le_option pos o1 o2 then o2 else o1)" + +definition "i1 \ i2 = + (if is_empty i1 then i2 else if is_empty i2 then i1 + else case (i1,i2) of (I l1 h1, I l2 h2) \ + I (min_option False l1 l2) (max_option True h1 h2))" + +definition "\ = {\}" + +instance +proof + case goal1 thus ?case + by(cases x, simp add: le_ivl_def le_option_def split: option.split) +next + case goal2 thus ?case + by(cases x, cases y, cases z, auto simp: le_ivl_def le_option_def split: option.splits if_splits) +next + case goal3 thus ?case + by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits) +next + case goal4 thus ?case + by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits) +next + case goal5 thus ?case + by(cases x, cases y, cases z, auto simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits if_splits) +next + case goal6 thus ?case + by(cases x, simp add: Top_ivl_def le_ivl_def le_option_def split: option.split) +qed + +end + + +instantiation ivl :: L_top_bot +begin + +definition "i1 \ i2 = (if is_empty i1 \ is_empty i2 then empty else + case (i1,i2) of (I l1 h1, I l2 h2) \ + I (max_option False l1 l2) (min_option True h1 h2))" + +definition "\ = empty" + +instance +proof + case goal1 thus ?case + by (simp add:meet_ivl_def empty_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits) +next + case goal2 thus ?case + by (simp add: empty_def meet_ivl_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits) +next + case goal3 thus ?case + by (cases x, cases y, cases z, auto simp add: le_ivl_def meet_ivl_def empty_def le_option_def max_option_def min_option_def split: option.splits if_splits) +next + case goal4 show ?case by(cases x, simp add: bot_ivl_def empty_def le_ivl_def) +qed + +end + +instantiation option :: (minus)minus +begin + +fun minus_option where +"Some x - Some y = Some(x-y)" | +"_ - _ = None" + +instance .. + +end + +definition "minus_ivl i1 i2 = (if is_empty i1 | is_empty i2 then empty else + case (i1,i2) of (I l1 h1, I l2 h2) \ I (l1-h2) (h1-l2))" + +lemma gamma_minus_ivl: + "n1 : \_ivl i1 \ n2 : \_ivl i2 \ n1-n2 : \_ivl(minus_ivl i1 i2)" +by(auto simp add: minus_ivl_def \_ivl_def split: ivl.splits option.splits) + +definition "filter_plus_ivl i i1 i2 = ((*if is_empty i then empty else*) + i1 \ minus_ivl i i2, i2 \ minus_ivl i i1)" + +fun filter_less_ivl :: "bool \ ivl \ ivl \ ivl * ivl" where +"filter_less_ivl res (I l1 h1) (I l2 h2) = + (if is_empty(I l1 h1) \ is_empty(I l2 h2) then (empty, empty) else + if res + then (I l1 (min_option True h1 (h2 - Some 1)), + I (max_option False (l1 + Some 1) l2) h2) + else (I (max_option False l1 l2) h1, I l2 (min_option True h1 h2)))" + +interpretation Val_abs +where \ = \_ivl and num' = num_ivl and plus' = plus_ivl +proof + case goal1 thus ?case + by(auto simp: \_ivl_def le_ivl_def le_option_def split: ivl.split option.split if_splits) +next + case goal2 show ?case by(simp add: \_ivl_def Top_ivl_def) +next + case goal3 thus ?case by(simp add: \_ivl_def num_ivl_def) +next + case goal4 thus ?case + by(auto simp add: \_ivl_def plus_ivl_def split: ivl.split option.splits) +qed + +interpretation Val_abs1_gamma +where \ = \_ivl and num' = num_ivl and plus' = plus_ivl +defines aval_ivl is aval' +proof + case goal1 thus ?case + by(auto simp add: \_ivl_def meet_ivl_def empty_def min_option_def max_option_def split: ivl.split option.split) +next + case goal2 show ?case by(auto simp add: bot_ivl_def \_ivl_def empty_def) +qed + +lemma mono_minus_ivl: + "i1 \ i1' \ i2 \ i2' \ minus_ivl i1 i2 \ minus_ivl i1' i2'" +apply(auto simp add: minus_ivl_def empty_def le_ivl_def le_option_def split: ivl.splits) + apply(simp split: option.splits) + apply(simp split: option.splits) +apply(simp split: option.splits) +done + + +interpretation Val_abs1 +where \ = \_ivl and num' = num_ivl and plus' = plus_ivl +and test_num' = in_ivl +and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl +proof + case goal1 thus ?case + by (simp add: \_ivl_def split: ivl.split option.split) +next + case goal2 thus ?case + by(auto simp add: filter_plus_ivl_def) + (metis gamma_minus_ivl add_diff_cancel add_commute)+ +next + case goal3 thus ?case + by(cases a1, cases a2, + auto simp: \_ivl_def min_option_def max_option_def le_option_def split: if_splits option.splits) +qed + +interpretation Abs_Int1 +where \ = \_ivl and num' = num_ivl and plus' = plus_ivl +and test_num' = in_ivl +and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl +defines afilter_ivl is afilter +and bfilter_ivl is bfilter +and step_ivl is step' +and AI_ivl is AI +and aval_ivl' is aval'' +.. + + +text{* Monotonicity: *} + +interpretation Abs_Int1_mono +where \ = \_ivl and num' = num_ivl and plus' = plus_ivl +and test_num' = in_ivl +and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl +proof + case goal1 thus ?case + by(auto simp: plus_ivl_def le_ivl_def le_option_def empty_def split: if_splits ivl.splits option.splits) +next + case goal2 thus ?case + by(auto simp: filter_plus_ivl_def le_prod_def mono_meet mono_minus_ivl) +next + case goal3 thus ?case + apply(cases a1, cases b1, cases a2, cases b2, auto simp: le_prod_def) + by(auto simp add: empty_def le_ivl_def le_option_def min_option_def max_option_def split: option.splits) +qed + + +subsubsection "Tests" + +value "show_acom_opt (AI_ivl test1_ivl)" + +text{* Better than @{text AI_const}: *} +value "show_acom_opt (AI_ivl test3_const)" +value "show_acom_opt (AI_ivl test4_const)" +value "show_acom_opt (AI_ivl test6_const)" + +value "show_acom_opt (AI_ivl test2_ivl)" +value "show_acom (((step_ivl \)^^0) (\\<^sub>c test2_ivl))" +value "show_acom (((step_ivl \)^^1) (\\<^sub>c test2_ivl))" +value "show_acom (((step_ivl \)^^2) (\\<^sub>c test2_ivl))" + +text{* Fixed point reached in 2 steps. Not so if the start value of x is known: *} + +value "show_acom_opt (AI_ivl test3_ivl)" +value "show_acom (((step_ivl \)^^0) (\\<^sub>c test3_ivl))" +value "show_acom (((step_ivl \)^^1) (\\<^sub>c test3_ivl))" +value "show_acom (((step_ivl \)^^2) (\\<^sub>c test3_ivl))" +value "show_acom (((step_ivl \)^^3) (\\<^sub>c test3_ivl))" +value "show_acom (((step_ivl \)^^4) (\\<^sub>c test3_ivl))" + +text{* Takes as many iterations as the actual execution. Would diverge if +loop did not terminate. Worse still, as the following example shows: even if +the actual execution terminates, the analysis may not. The value of y keeps +decreasing as the analysis is iterated, no matter how long: *} + +value "show_acom (((step_ivl \)^^50) (\\<^sub>c test4_ivl))" + +text{* Relationships between variables are NOT captured: *} +value "show_acom_opt (AI_ivl test5_ivl)" + +text{* Again, the analysis would not terminate: *} +value "show_acom (((step_ivl \)^^50) (\\<^sub>c test6_ivl))" + +end diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy Thu Apr 19 17:49:02 2012 +0200 @@ -0,0 +1,683 @@ +(* Author: Tobias Nipkow *) + +theory Abs_Int3_ITP +imports Abs_Int2_ivl_ITP +begin + +subsection "Widening and Narrowing" + +class WN = SL_top + +fixes widen :: "'a \ 'a \ 'a" (infix "\" 65) +assumes widen1: "x \ x \ y" +assumes widen2: "y \ x \ y" +fixes narrow :: "'a \ 'a \ 'a" (infix "\" 65) +assumes narrow1: "y \ x \ y \ x \ y" +assumes narrow2: "y \ x \ x \ y \ x" + + +subsubsection "Intervals" + +instantiation ivl :: WN +begin + +definition "widen_ivl ivl1 ivl2 = + ((*if is_empty ivl1 then ivl2 else + if is_empty ivl2 then ivl1 else*) + case (ivl1,ivl2) of (I l1 h1, I l2 h2) \ + I (if le_option False l2 l1 \ l2 \ l1 then None else l1) + (if le_option True h1 h2 \ h1 \ h2 then None else h1))" + +definition "narrow_ivl ivl1 ivl2 = + ((*if is_empty ivl1 \ is_empty ivl2 then empty else*) + case (ivl1,ivl2) of (I l1 h1, I l2 h2) \ + I (if l1 = None then l2 else l1) + (if h1 = None then h2 else h1))" + +instance +proof qed + (auto simp add: widen_ivl_def narrow_ivl_def le_option_def le_ivl_def empty_def split: ivl.split option.split if_splits) + +end + + +subsubsection "Abstract State" + +instantiation st :: (WN)WN +begin + +definition "widen_st F1 F2 = + FunDom (\x. fun F1 x \ fun F2 x) (inter_list (dom F1) (dom F2))" + +definition "narrow_st F1 F2 = + FunDom (\x. fun F1 x \ fun F2 x) (inter_list (dom F1) (dom F2))" + +instance +proof + case goal1 thus ?case + by(simp add: widen_st_def le_st_def lookup_def widen1) +next + case goal2 thus ?case + by(simp add: widen_st_def le_st_def lookup_def widen2) +next + case goal3 thus ?case + by(auto simp: narrow_st_def le_st_def lookup_def narrow1) +next + case goal4 thus ?case + by(auto simp: narrow_st_def le_st_def lookup_def narrow2) +qed + +end + + +subsubsection "Option" + +instantiation option :: (WN)WN +begin + +fun widen_option where +"None \ x = x" | +"x \ None = x" | +"(Some x) \ (Some y) = Some(x \ y)" + +fun narrow_option where +"None \ x = None" | +"x \ None = None" | +"(Some x) \ (Some y) = Some(x \ y)" + +instance +proof + case goal1 show ?case + by(induct x y rule: widen_option.induct) (simp_all add: widen1) +next + case goal2 show ?case + by(induct x y rule: widen_option.induct) (simp_all add: widen2) +next + case goal3 thus ?case + by(induct x y rule: narrow_option.induct) (simp_all add: narrow1) +next + case goal4 thus ?case + by(induct x y rule: narrow_option.induct) (simp_all add: narrow2) +qed + +end + + +subsubsection "Annotated commands" + +fun map2_acom :: "('a \ 'a \ 'a) \ 'a acom \ 'a acom \ 'a acom" where +"map2_acom f (SKIP {a1}) (SKIP {a2}) = (SKIP {f a1 a2})" | +"map2_acom f (x ::= e {a1}) (x' ::= e' {a2}) = (x ::= e {f a1 a2})" | +"map2_acom f (c1;c2) (c1';c2') = (map2_acom f c1 c1'; map2_acom f c2 c2')" | +"map2_acom f (IF b THEN c1 ELSE c2 {a1}) (IF b' THEN c1' ELSE c2' {a2}) = + (IF b THEN map2_acom f c1 c1' ELSE map2_acom f c2 c2' {f a1 a2})" | +"map2_acom f ({a1} WHILE b DO c {a2}) ({a3} WHILE b' DO c' {a4}) = + ({f a1 a3} WHILE b DO map2_acom f c c' {f a2 a4})" + +abbreviation widen_acom :: "('a::WN)acom \ 'a acom \ 'a acom" (infix "\\<^sub>c" 65) +where "widen_acom == map2_acom (op \)" + +abbreviation narrow_acom :: "('a::WN)acom \ 'a acom \ 'a acom" (infix "\\<^sub>c" 65) +where "narrow_acom == map2_acom (op \)" + +lemma widen1_acom: "strip c = strip c' \ c \ c \\<^sub>c c'" +by(induct c c' rule: le_acom.induct)(simp_all add: widen1) + +lemma widen2_acom: "strip c = strip c' \ c' \ c \\<^sub>c c'" +by(induct c c' rule: le_acom.induct)(simp_all add: widen2) + +lemma narrow1_acom: "y \ x \ y \ x \\<^sub>c y" +by(induct y x rule: le_acom.induct) (simp_all add: narrow1) + +lemma narrow2_acom: "y \ x \ x \\<^sub>c y \ x" +by(induct y x rule: le_acom.induct) (simp_all add: narrow2) + + +subsubsection "Post-fixed point computation" + +definition iter_widen :: "('a acom \ 'a acom) \ 'a acom \ ('a::WN)acom option" +where "iter_widen f = while_option (\c. \ f c \ c) (\c. c \\<^sub>c f c)" + +definition iter_narrow :: "('a acom \ 'a acom) \ 'a acom \ 'a::WN acom option" +where "iter_narrow f = while_option (\c. \ c \ c \\<^sub>c f c) (\c. c \\<^sub>c f c)" + +definition pfp_wn :: + "(('a::WN)option acom \ 'a option acom) \ com \ 'a option acom option" +where "pfp_wn f c = (case iter_widen f (\\<^sub>c c) of None \ None + | Some c' \ iter_narrow f c')" + +lemma strip_map2_acom: + "strip c1 = strip c2 \ strip(map2_acom f c1 c2) = strip c1" +by(induct f c1 c2 rule: map2_acom.induct) simp_all + +lemma iter_widen_pfp: "iter_widen f c = Some c' \ f c' \ c'" +by(auto simp add: iter_widen_def dest: while_option_stop) + +lemma strip_while: fixes f :: "'a acom \ 'a acom" +assumes "\c. strip (f c) = strip c" and "while_option P f c = Some c'" +shows "strip c' = strip c" +using while_option_rule[where P = "\c'. strip c' = strip c", OF _ assms(2)] +by (metis assms(1)) + +lemma strip_iter_widen: fixes f :: "'a::WN acom \ 'a acom" +assumes "\c. strip (f c) = strip c" and "iter_widen f c = Some c'" +shows "strip c' = strip c" +proof- + have "\c. strip(c \\<^sub>c f c) = strip c" by (metis assms(1) strip_map2_acom) + from strip_while[OF this] assms(2) show ?thesis by(simp add: iter_widen_def) +qed + +lemma iter_narrow_pfp: assumes "mono f" and "f c0 \ c0" +and "iter_narrow f c0 = Some c" +shows "f c \ c \ c \ c0" (is "?P c") +proof- + { fix c assume "?P c" + note 1 = conjunct1[OF this] and 2 = conjunct2[OF this] + let ?c' = "c \\<^sub>c f c" + have "?P ?c'" + proof + have "f ?c' \ f c" by(rule monoD[OF `mono f` narrow2_acom[OF 1]]) + also have "\ \ ?c'" by(rule narrow1_acom[OF 1]) + finally show "f ?c' \ ?c'" . + have "?c' \ c" by (rule narrow2_acom[OF 1]) + also have "c \ c0" by(rule 2) + finally show "?c' \ c0" . + qed + } + with while_option_rule[where P = ?P, OF _ assms(3)[simplified iter_narrow_def]] + assms(2) le_refl + show ?thesis by blast +qed + +lemma pfp_wn_pfp: + "\ mono f; pfp_wn f c = Some c' \ \ f c' \ c'" +unfolding pfp_wn_def +by (auto dest: iter_widen_pfp iter_narrow_pfp split: option.splits) + +lemma strip_pfp_wn: + "\ \c. strip(f c) = strip c; pfp_wn f c = Some c' \ \ strip c' = c" +apply(auto simp add: pfp_wn_def iter_narrow_def split: option.splits) +by (metis (no_types) strip_map2_acom strip_while strip_bot_acom strip_iter_widen) + +locale Abs_Int2 = Abs_Int1_mono +where \=\ for \ :: "'av::{WN,L_top_bot} \ val set" +begin + +definition AI_wn :: "com \ 'av st option acom option" where +"AI_wn = pfp_wn (step' \)" + +lemma AI_wn_sound: "AI_wn c = Some c' \ CS c \ \\<^isub>c c'" +proof(simp add: CS_def AI_wn_def) + assume 1: "pfp_wn (step' \) c = Some c'" + from pfp_wn_pfp[OF mono_step'2 1] + have 2: "step' \ c' \ c'" . + have 3: "strip (\\<^isub>c (step' \ c')) = c" by(simp add: strip_pfp_wn[OF _ 1]) + have "lfp (step UNIV) c \ \\<^isub>c (step' \ c')" + proof(rule lfp_lowerbound[simplified,OF 3]) + show "step UNIV (\\<^isub>c (step' \ c')) \ \\<^isub>c (step' \ c')" + proof(rule step_preserves_le[OF _ _]) + show "UNIV \ \\<^isub>o \" by simp + show "\\<^isub>c (step' \ c') \ \\<^isub>c c'" by(rule mono_gamma_c[OF 2]) + qed + qed + from this 2 show "lfp (step UNIV) c \ \\<^isub>c c'" + by (blast intro: mono_gamma_c order_trans) +qed + +end + +interpretation Abs_Int2 +where \ = \_ivl and num' = num_ivl and plus' = plus_ivl +and test_num' = in_ivl +and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl +defines AI_ivl' is AI_wn +.. + + +subsubsection "Tests" + +definition "step_up_ivl n = ((\c. c \\<^sub>c step_ivl \ c)^^n)" +definition "step_down_ivl n = ((\c. c \\<^sub>c step_ivl \ c)^^n)" + +text{* For @{const test3_ivl}, @{const AI_ivl} needed as many iterations as +the loop took to execute. In contrast, @{const AI_ivl'} converges in a +constant number of steps: *} + +value "show_acom (step_up_ivl 1 (\\<^sub>c test3_ivl))" +value "show_acom (step_up_ivl 2 (\\<^sub>c test3_ivl))" +value "show_acom (step_up_ivl 3 (\\<^sub>c test3_ivl))" +value "show_acom (step_up_ivl 4 (\\<^sub>c test3_ivl))" +value "show_acom (step_up_ivl 5 (\\<^sub>c test3_ivl))" +value "show_acom (step_down_ivl 1 (step_up_ivl 5 (\\<^sub>c test3_ivl)))" +value "show_acom (step_down_ivl 2 (step_up_ivl 5 (\\<^sub>c test3_ivl)))" +value "show_acom (step_down_ivl 3 (step_up_ivl 5 (\\<^sub>c test3_ivl)))" + +text{* Now all the analyses terminate: *} + +value "show_acom_opt (AI_ivl' test4_ivl)" +value "show_acom_opt (AI_ivl' test5_ivl)" +value "show_acom_opt (AI_ivl' test6_ivl)" + + +subsubsection "Termination: Intervals" + +definition m_ivl :: "ivl \ nat" where +"m_ivl ivl = (case ivl of I l h \ + (case l of None \ 0 | Some _ \ 1) + (case h of None \ 0 | Some _ \ 1))" + +lemma m_ivl_height: "m_ivl ivl \ 2" +by(simp add: m_ivl_def split: ivl.split option.split) + +lemma m_ivl_anti_mono: "(y::ivl) \ x \ m_ivl x \ m_ivl y" +by(auto simp: m_ivl_def le_option_def le_ivl_def + split: ivl.split option.split if_splits) + +lemma m_ivl_widen: + "~ y \ x \ m_ivl(x \ y) < m_ivl x" +by(auto simp: m_ivl_def widen_ivl_def le_option_def le_ivl_def + split: ivl.splits option.splits if_splits) + +lemma Top_less_ivl: "\ \ x \ m_ivl x = 0" +by(auto simp: m_ivl_def le_option_def le_ivl_def empty_def Top_ivl_def + split: ivl.split option.split if_splits) + + +definition n_ivl :: "ivl \ nat" where +"n_ivl ivl = 2 - m_ivl ivl" + +lemma n_ivl_mono: "(x::ivl) \ y \ n_ivl x \ n_ivl y" +unfolding n_ivl_def by (metis diff_le_mono2 m_ivl_anti_mono) + +lemma n_ivl_narrow: + "~ x \ x \ y \ n_ivl(x \ y) < n_ivl x" +by(auto simp: n_ivl_def m_ivl_def narrow_ivl_def le_option_def le_ivl_def + split: ivl.splits option.splits if_splits) + + +subsubsection "Termination: Abstract State" + +definition "m_st m st = (\x\set(dom st). m(fun st x))" + +lemma m_st_height: assumes "finite X" and "set (dom S) \ X" +shows "m_st m_ivl S \ 2 * card X" +proof(auto simp: m_st_def) + have "(\x\set(dom S). m_ivl (fun S x)) \ (\x\set(dom S). 2)" (is "?L \ _") + by(rule setsum_mono)(simp add:m_ivl_height) + also have "\ \ (\x\X. 2)" + by(rule setsum_mono3[OF assms]) simp + also have "\ = 2 * card X" by(simp add: setsum_constant) + finally show "?L \ \" . +qed + +lemma m_st_anti_mono: + "S1 \ S2 \ m_st m_ivl S2 \ m_st m_ivl S1" +proof(auto simp: m_st_def le_st_def lookup_def split: if_splits) + let ?X = "set(dom S1)" let ?Y = "set(dom S2)" + let ?f = "fun S1" let ?g = "fun S2" + assume asm: "\x\?Y. (x \ ?X \ ?f x \ ?g x) \ (x \ ?X \ \ \ ?g x)" + hence 1: "\y\?Y\?X. m_ivl(?g y) \ m_ivl(?f y)" by(simp add: m_ivl_anti_mono) + have 0: "\x\?Y-?X. m_ivl(?g x) = 0" using asm by (auto simp: Top_less_ivl) + have "(\y\?Y. m_ivl(?g y)) = (\y\(?Y-?X) \ (?Y\?X). m_ivl(?g y))" + by (metis Un_Diff_Int) + also have "\ = (\y\?Y-?X. m_ivl(?g y)) + (\y\?Y\?X. m_ivl(?g y))" + by(subst setsum_Un_disjoint) auto + also have "(\y\?Y-?X. m_ivl(?g y)) = 0" using 0 by simp + also have "0 + (\y\?Y\?X. m_ivl(?g y)) = (\y\?Y\?X. m_ivl(?g y))" by simp + also have "\ \ (\y\?Y\?X. m_ivl(?f y))" + by(rule setsum_mono)(simp add: 1) + also have "\ \ (\y\?X. m_ivl(?f y))" + by(simp add: setsum_mono3[of "?X" "?Y Int ?X", OF _ Int_lower2]) + finally show "(\y\?Y. m_ivl(?g y)) \ (\x\?X. m_ivl(?f x))" + by (metis add_less_cancel_left) +qed + +lemma m_st_widen: +assumes "\ S2 \ S1" shows "m_st m_ivl (S1 \ S2) < m_st m_ivl S1" +proof- + { let ?X = "set(dom S1)" let ?Y = "set(dom S2)" + let ?f = "fun S1" let ?g = "fun S2" + fix x assume "x \ ?X" "\ lookup S2 x \ ?f x" + have "(\x\?X\?Y. m_ivl(?f x \ ?g x)) < (\x\?X. m_ivl(?f x))" (is "?L < ?R") + proof cases + assume "x : ?Y" + have "?L < (\x\?X\?Y. m_ivl(?f x))" + proof(rule setsum_strict_mono1, simp) + show "\x\?X\?Y. m_ivl(?f x \ ?g x) \ m_ivl (?f x)" + by (metis m_ivl_anti_mono widen1) + next + show "\x\?X\?Y. m_ivl(?f x \ ?g x) < m_ivl(?f x)" + using `x:?X` `x:?Y` `\ lookup S2 x \ ?f x` + by (metis IntI m_ivl_widen lookup_def) + qed + also have "\ \ ?R" by(simp add: setsum_mono3[OF _ Int_lower1]) + finally show ?thesis . + next + assume "x ~: ?Y" + have "?L \ (\x\?X\?Y. m_ivl(?f x))" + proof(rule setsum_mono, simp) + fix x assume "x:?X \ x:?Y" show "m_ivl(?f x \ ?g x) \ m_ivl (?f x)" + by (metis m_ivl_anti_mono widen1) + qed + also have "\ < m_ivl(?f x) + \" + using m_ivl_widen[OF `\ lookup S2 x \ ?f x`] + by (metis Nat.le_refl add_strict_increasing gr0I not_less0) + also have "\ = (\y\insert x (?X\?Y). m_ivl(?f y))" + using `x ~: ?Y` by simp + also have "\ \ (\x\?X. m_ivl(?f x))" + by(rule setsum_mono3)(insert `x:?X`, auto) + finally show ?thesis . + qed + } with assms show ?thesis + by(auto simp: le_st_def widen_st_def m_st_def Int_def) +qed + +definition "n_st m X st = (\x\X. m(lookup st x))" + +lemma n_st_mono: assumes "set(dom S1) \ X" "set(dom S2) \ X" "S1 \ S2" +shows "n_st n_ivl X S1 \ n_st n_ivl X S2" +proof- + have "(\x\X. n_ivl(lookup S1 x)) \ (\x\X. n_ivl(lookup S2 x))" + apply(rule setsum_mono) using assms + by(auto simp: le_st_def lookup_def n_ivl_mono split: if_splits) + thus ?thesis by(simp add: n_st_def) +qed + +lemma n_st_narrow: +assumes "finite X" and "set(dom S1) \ X" "set(dom S2) \ X" +and "S2 \ S1" "\ S1 \ S1 \ S2" +shows "n_st n_ivl X (S1 \ S2) < n_st n_ivl X S1" +proof- + have 1: "\x\X. n_ivl (lookup (S1 \ S2) x) \ n_ivl (lookup S1 x)" + using assms(2-4) + by(auto simp: le_st_def narrow_st_def lookup_def n_ivl_mono narrow2 + split: if_splits) + have 2: "\x\X. n_ivl (lookup (S1 \ S2) x) < n_ivl (lookup S1 x)" + using assms(2-5) + by(auto simp: le_st_def narrow_st_def lookup_def intro: n_ivl_narrow + split: if_splits) + have "(\x\X. n_ivl(lookup (S1 \ S2) x)) < (\x\X. n_ivl(lookup S1 x))" + apply(rule setsum_strict_mono1[OF `finite X`]) using 1 2 by blast+ + thus ?thesis by(simp add: n_st_def) +qed + + +subsubsection "Termination: Option" + +definition "m_o m n opt = (case opt of None \ n+1 | Some x \ m x)" + +lemma m_o_anti_mono: "finite X \ domo S2 \ X \ S1 \ S2 \ + m_o (m_st m_ivl) (2 * card X) S2 \ m_o (m_st m_ivl) (2 * card X) S1" +apply(induction S1 S2 rule: le_option.induct) +apply(auto simp: domo_def m_o_def m_st_anti_mono le_SucI m_st_height + split: option.splits) +done + +lemma m_o_widen: "\ finite X; domo S2 \ X; \ S2 \ S1 \ \ + m_o (m_st m_ivl) (2 * card X) (S1 \ S2) < m_o (m_st m_ivl) (2 * card X) S1" +by(auto simp: m_o_def domo_def m_st_height less_Suc_eq_le m_st_widen + split: option.split) + +definition "n_o n opt = (case opt of None \ 0 | Some x \ n x + 1)" + +lemma n_o_mono: "domo S1 \ X \ domo S2 \ X \ S1 \ S2 \ + n_o (n_st n_ivl X) S1 \ n_o (n_st n_ivl X) S2" +apply(induction S1 S2 rule: le_option.induct) +apply(auto simp: domo_def n_o_def n_st_mono + split: option.splits) +done + +lemma n_o_narrow: + "\ finite X; domo S1 \ X; domo S2 \ X; S2 \ S1; \ S1 \ S1 \ S2 \ + \ n_o (n_st n_ivl X) (S1 \ S2) < n_o (n_st n_ivl X) S1" +apply(induction S1 S2 rule: narrow_option.induct) +apply(auto simp: n_o_def domo_def n_st_narrow) +done + +lemma domo_widen_subset: "domo (S1 \ S2) \ domo S1 \ domo S2" +apply(induction S1 S2 rule: widen_option.induct) +apply (auto simp: domo_def widen_st_def) +done + +lemma domo_narrow_subset: "domo (S1 \ S2) \ domo S1 \ domo S2" +apply(induction S1 S2 rule: narrow_option.induct) +apply (auto simp: domo_def narrow_st_def) +done + +subsubsection "Termination: Commands" + +lemma strip_widen_acom[simp]: + "strip c' = strip (c::'a::WN acom) \ strip (c \\<^sub>c c') = strip c" +by(induction "widen::'a\'a\'a" c c' rule: map2_acom.induct) simp_all + +lemma strip_narrow_acom[simp]: + "strip c' = strip (c::'a::WN acom) \ strip (c \\<^sub>c c') = strip c" +by(induction "narrow::'a\'a\'a" c c' rule: map2_acom.induct) simp_all + +lemma annos_widen_acom[simp]: "strip c1 = strip (c2::'a::WN acom) \ + annos(c1 \\<^sub>c c2) = map (%(x,y).x\y) (zip (annos c1) (annos(c2::'a::WN acom)))" +by(induction "widen::'a\'a\'a" c1 c2 rule: map2_acom.induct) + (simp_all add: size_annos_same2) + +lemma annos_narrow_acom[simp]: "strip c1 = strip (c2::'a::WN acom) \ + annos(c1 \\<^sub>c c2) = map (%(x,y).x\y) (zip (annos c1) (annos(c2::'a::WN acom)))" +by(induction "narrow::'a\'a\'a" c1 c2 rule: map2_acom.induct) + (simp_all add: size_annos_same2) + +lemma widen_acom_Com[simp]: "strip c2 = strip c1 \ + c1 : Com X \ c2 : Com X \ (c1 \\<^sub>c c2) : Com X" +apply(auto simp add: Com_def) +apply(rename_tac S S' x) +apply(erule in_set_zipE) +apply(auto simp: domo_def split: option.splits) +apply(case_tac S) +apply(case_tac S') +apply simp +apply fastforce +apply(case_tac S') +apply fastforce +apply (fastforce simp: widen_st_def) +done + +lemma narrow_acom_Com[simp]: "strip c2 = strip c1 \ + c1 : Com X \ c2 : Com X \ (c1 \\<^sub>c c2) : Com X" +apply(auto simp add: Com_def) +apply(rename_tac S S' x) +apply(erule in_set_zipE) +apply(auto simp: domo_def split: option.splits) +apply(case_tac S) +apply(case_tac S') +apply simp +apply fastforce +apply(case_tac S') +apply fastforce +apply (fastforce simp: narrow_st_def) +done + +definition "m_c m c = (let as = annos c in \i=0.. {(c, c \\<^sub>c c') |c c'\ivl st option acom. + strip c' = strip c \ c : Com X \ c' : Com X \ \ c' \ c}\ + \ measure(m_c(m_o (m_st m_ivl) (2*card(X))))" +apply(auto simp: m_c_def Let_def Com_def) +apply(subgoal_tac "length(annos c') = length(annos c)") +prefer 2 apply (simp add: size_annos_same2) +apply (auto) +apply(rule setsum_strict_mono1) +apply simp +apply (clarsimp) +apply(erule m_o_anti_mono) +apply(rule subset_trans[OF domo_widen_subset]) +apply fastforce +apply(rule widen1) +apply(auto simp: le_iff_le_annos listrel_iff_nth) +apply(rule_tac x=n in bexI) +prefer 2 apply simp +apply(erule m_o_widen) +apply (simp)+ +done + +lemma measure_n_c: "finite X \ {(c, c \\<^sub>c c') |c c'. + strip c = strip c' \ c \ Com X \ c' \ Com X \ c' \ c \ \ c \ c \\<^sub>c c'}\ + \ measure(m_c(n_o (n_st n_ivl X)))" +apply(auto simp: m_c_def Let_def Com_def) +apply(subgoal_tac "length(annos c') = length(annos c)") +prefer 2 apply (simp add: size_annos_same2) +apply (auto) +apply(rule setsum_strict_mono1) +apply simp +apply (clarsimp) +apply(rule n_o_mono) +using domo_narrow_subset apply fastforce +apply fastforce +apply(rule narrow2) +apply(fastforce simp: le_iff_le_annos listrel_iff_nth) +apply(auto simp: le_iff_le_annos listrel_iff_nth strip_narrow_acom) +apply(rule_tac x=n in bexI) +prefer 2 apply simp +apply(erule n_o_narrow) +apply (simp)+ +done + + +subsubsection "Termination: Post-Fixed Point Iterations" + +lemma iter_widen_termination: +fixes c0 :: "'a::WN acom" +assumes P_f: "\c. P c \ P(f c)" +assumes P_widen: "\c c'. P c \ P c' \ P(c \\<^sub>c c')" +and "wf({(c::'a acom,c \\<^sub>c c')|c c'. P c \ P c' \ ~ c' \ c}^-1)" +and "P c0" and "c0 \ f c0" shows "EX c. iter_widen f c0 = Some c" +proof(simp add: iter_widen_def, rule wf_while_option_Some[where P = "P"]) + show "wf {(cc', c). (P c \ \ f c \ c) \ cc' = c \\<^sub>c f c}" + apply(rule wf_subset[OF assms(3)]) by(blast intro: P_f) +next + show "P c0" by(rule `P c0`) +next + fix c assume "P c" thus "P (c \\<^sub>c f c)" by(simp add: P_f P_widen) +qed + +lemma iter_narrow_termination: +assumes P_f: "\c. P c \ P(c \\<^sub>c f c)" +and wf: "wf({(c, c \\<^sub>c f c)|c c'. P c \ ~ c \ c \\<^sub>c f c}^-1)" +and "P c0" shows "EX c. iter_narrow f c0 = Some c" +proof(simp add: iter_narrow_def, rule wf_while_option_Some[where P = "P"]) + show "wf {(c', c). (P c \ \ c \ c \\<^sub>c f c) \ c' = c \\<^sub>c f c}" + apply(rule wf_subset[OF wf]) by(blast intro: P_f) +next + show "P c0" by(rule `P c0`) +next + fix c assume "P c" thus "P (c \\<^sub>c f c)" by(simp add: P_f) +qed + +lemma iter_winden_step_ivl_termination: + "EX c. iter_widen (step_ivl \) (\\<^sub>c c0) = Some c" +apply(rule iter_widen_termination[where + P = "%c. strip c = c0 \ c : Com(vars c0)"]) +apply (simp_all add: step'_Com bot_acom) +apply(rule wf_subset) +apply(rule wf_measure) +apply(rule subset_trans) +prefer 2 +apply(rule measure_m_c[where X = "vars c0", OF finite_cvars]) +apply blast +done + +lemma iter_narrow_step_ivl_termination: + "c0 \ Com (vars(strip c0)) \ step_ivl \ c0 \ c0 \ + EX c. iter_narrow (step_ivl \) c0 = Some c" +apply(rule iter_narrow_termination[where + P = "%c. strip c = strip c0 \ c : Com(vars(strip c0)) \ step_ivl \ c \ c"]) +apply (simp_all add: step'_Com) +apply(clarify) +apply(frule narrow2_acom, drule mono_step'[OF le_refl], erule le_trans[OF _ narrow1_acom]) +apply assumption +apply(rule wf_subset) +apply(rule wf_measure) +apply(rule subset_trans) +prefer 2 +apply(rule measure_n_c[where X = "vars(strip c0)", OF finite_cvars]) +apply auto +by (metis bot_least domo_Top order_refl step'_Com strip_step') + +(* FIXME: simplify type system: Combine Com(X) and vars <= X?? *) +lemma while_Com: +fixes c :: "'a st option acom" +assumes "while_option P f c = Some c'" +and "!!c. strip(f c) = strip c" +and "\c::'a st option acom. c : Com(X) \ vars(strip c) \ X \ f c : Com(X)" +and "c : Com(X)" and "vars(strip c) \ X" shows "c' : Com(X)" +using while_option_rule[where P = "\c'. c' : Com(X) \ vars(strip c') \ X", OF _ assms(1)] +by(simp add: assms(2-)) + +lemma iter_widen_Com: fixes f :: "'a::WN st option acom \ 'a st option acom" +assumes "iter_widen f c = Some c'" +and "\c. c : Com(X) \ vars(strip c) \ X \ f c : Com(X)" +and "!!c. strip(f c) = strip c" +and "c : Com(X)" and "vars (strip c) \ X" shows "c' : Com(X)" +proof- + have "\c. c : Com(X) \ vars(strip c) \ X \ c \\<^sub>c f c : Com(X)" + by (metis (full_types) widen_acom_Com assms(2,3)) + from while_Com[OF assms(1)[simplified iter_widen_def] _ this assms(4,5)] + show ?thesis using assms(3) by(simp) +qed + + +context Abs_Int2 +begin + +lemma iter_widen_step'_Com: + "iter_widen (step' \) c = Some c' \ vars(strip c) \ X \ c : Com(X) + \ c' : Com(X)" +apply(subgoal_tac "strip c'= strip c") +prefer 2 apply (metis strip_iter_widen strip_step') +apply(drule iter_widen_Com) +prefer 3 apply assumption +prefer 3 apply assumption +apply (auto simp: step'_Com) +done + +end + +theorem AI_ivl'_termination: + "EX c'. AI_ivl' c = Some c'" +apply(auto simp: AI_wn_def pfp_wn_def iter_winden_step_ivl_termination split: option.split) +apply(rule iter_narrow_step_ivl_termination) +apply (metis bot_acom_Com iter_widen_step'_Com[OF _ subset_refl] strip_iter_widen strip_step') +apply(erule iter_widen_pfp) +done + +end + + +(* interesting(?) relic +lemma widen_assoc: + "~ (y::ivl) \ x \ ~ z \ x \ y \ ((x::ivl) \ y) \ z = x \ (y \ z)" +apply(cases x) +apply(cases y) +apply(cases z) +apply(rename_tac x1 x2 y1 y2 z1 z2) +apply(simp add: le_ivl_def) +apply(case_tac x1) +apply(case_tac x2) +apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits) +apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits) +apply(case_tac x2) +apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits) +apply(case_tac y1) +apply(case_tac y2) +apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits) +apply(case_tac z1) +apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits ivl.splits)[1] +apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits ivl.splits)[1] +apply(case_tac y2) +apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits ivl.splits)[1] +apply(case_tac z1) +apply(auto simp add:le_option_def widen_ivl_def split: if_splits ivl.splits option.splits)[1] +apply(case_tac z2) +apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits)[1] +apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits)[1] +done + +lemma widen_step_trans: + "~ (y::ivl) \ x \ ~ z \ x \ y \ EX u. (x \ y) \ z = x \ u \ ~ u \ x" +by (metis widen_assoc preord_class.le_trans widen1) +*) diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_Int_ITP/Abs_State_ITP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/Abs_Int_ITP/Abs_State_ITP.thy Thu Apr 19 17:49:02 2012 +0200 @@ -0,0 +1,98 @@ +(* Author: Tobias Nipkow *) + +theory Abs_State_ITP +imports Abs_Int0_ITP + "~~/src/HOL/Library/Char_ord" "~~/src/HOL/Library/List_lexord" + (* Library import merely to allow string lists to be sorted for output *) +begin + +subsection "Abstract State with Computable Ordering" + +text{* A concrete type of state with computable @{text"\"}: *} + +datatype 'a st = FunDom "vname \ 'a" "vname list" + +fun "fun" where "fun (FunDom f xs) = f" +fun dom where "dom (FunDom f xs) = xs" + +definition [simp]: "inter_list xs ys = [x\xs. x \ set ys]" + +definition "show_st S = [(x,fun S x). x \ sort(dom S)]" + +definition "show_acom = map_acom (Option.map show_st)" +definition "show_acom_opt = Option.map show_acom" + +definition "lookup F x = (if x : set(dom F) then fun F x else \)" + +definition "update F x y = + FunDom ((fun F)(x:=y)) (if x \ set(dom F) then dom F else x # dom F)" + +lemma lookup_update: "lookup (update S x y) = (lookup S)(x:=y)" +by(rule ext)(auto simp: lookup_def update_def) + +definition "\_st \ F = {f. \x. f x \ \(lookup F x)}" + +instantiation st :: (SL_top) SL_top +begin + +definition "le_st F G = (ALL x : set(dom G). lookup F x \ fun G x)" + +definition +"join_st F G = + FunDom (\x. fun F x \ fun G x) (inter_list (dom F) (dom G))" + +definition "\ = FunDom (\x. \) []" + +instance +proof + case goal2 thus ?case + apply(auto simp: le_st_def) + by (metis lookup_def preord_class.le_trans top) +qed (auto simp: le_st_def lookup_def join_st_def Top_st_def) + +end + +lemma mono_lookup: "F \ F' \ lookup F x \ lookup F' x" +by(auto simp add: lookup_def le_st_def) + +lemma mono_update: "a \ a' \ S \ S' \ update S x a \ update S' x a'" +by(auto simp add: le_st_def lookup_def update_def) + +locale Gamma = Val_abs where \=\ for \ :: "'av::SL_top \ val set" +begin + +abbreviation \\<^isub>f :: "'av st \ state set" +where "\\<^isub>f == \_st \" + +abbreviation \\<^isub>o :: "'av st option \ state set" +where "\\<^isub>o == \_option \\<^isub>f" + +abbreviation \\<^isub>c :: "'av st option acom \ state set acom" +where "\\<^isub>c == map_acom \\<^isub>o" + +lemma gamma_f_Top[simp]: "\\<^isub>f Top = UNIV" +by(auto simp: Top_st_def \_st_def lookup_def) + +lemma gamma_o_Top[simp]: "\\<^isub>o Top = UNIV" +by (simp add: Top_option_def) + +(* FIXME (maybe also le \ sqle?) *) + +lemma mono_gamma_f: "f \ g \ \\<^isub>f f \ \\<^isub>f g" +apply(simp add:\_st_def subset_iff lookup_def le_st_def split: if_splits) +by (metis UNIV_I mono_gamma gamma_Top subsetD) + +lemma mono_gamma_o: + "sa \ sa' \ \\<^isub>o sa \ \\<^isub>o sa'" +by(induction sa sa' rule: le_option.induct)(simp_all add: mono_gamma_f) + +lemma mono_gamma_c: "ca \ ca' \ \\<^isub>c ca \ \\<^isub>c ca'" +by (induction ca ca' rule: le_acom.induct) (simp_all add:mono_gamma_o) + +lemma in_gamma_option_iff: + "x : \_option r u \ (\u'. u = Some u' \ x : r u')" +by (cases u) auto + +end + +end diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/Abs_State.thy --- a/src/HOL/IMP/Abs_State.thy Thu Apr 19 11:14:57 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,98 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory Abs_State -imports Abs_Int0_fun - "~~/src/HOL/Library/Char_ord" "~~/src/HOL/Library/List_lexord" - (* Library import merely to allow string lists to be sorted for output *) -begin - -subsection "Abstract State with Computable Ordering" - -text{* A concrete type of state with computable @{text"\"}: *} - -datatype 'a st = FunDom "vname \ 'a" "vname list" - -fun "fun" where "fun (FunDom f xs) = f" -fun dom where "dom (FunDom f xs) = xs" - -definition [simp]: "inter_list xs ys = [x\xs. x \ set ys]" - -definition "show_st S = [(x,fun S x). x \ sort(dom S)]" - -definition "show_acom = map_acom (Option.map show_st)" -definition "show_acom_opt = Option.map show_acom" - -definition "lookup F x = (if x : set(dom F) then fun F x else \)" - -definition "update F x y = - FunDom ((fun F)(x:=y)) (if x \ set(dom F) then dom F else x # dom F)" - -lemma lookup_update: "lookup (update S x y) = (lookup S)(x:=y)" -by(rule ext)(auto simp: lookup_def update_def) - -definition "\_st \ F = {f. \x. f x \ \(lookup F x)}" - -instantiation st :: (SL_top) SL_top -begin - -definition "le_st F G = (ALL x : set(dom G). lookup F x \ fun G x)" - -definition -"join_st F G = - FunDom (\x. fun F x \ fun G x) (inter_list (dom F) (dom G))" - -definition "\ = FunDom (\x. \) []" - -instance -proof - case goal2 thus ?case - apply(auto simp: le_st_def) - by (metis lookup_def preord_class.le_trans top) -qed (auto simp: le_st_def lookup_def join_st_def Top_st_def) - -end - -lemma mono_lookup: "F \ F' \ lookup F x \ lookup F' x" -by(auto simp add: lookup_def le_st_def) - -lemma mono_update: "a \ a' \ S \ S' \ update S x a \ update S' x a'" -by(auto simp add: le_st_def lookup_def update_def) - -locale Gamma = Val_abs where \=\ for \ :: "'av::SL_top \ val set" -begin - -abbreviation \\<^isub>f :: "'av st \ state set" -where "\\<^isub>f == \_st \" - -abbreviation \\<^isub>o :: "'av st option \ state set" -where "\\<^isub>o == \_option \\<^isub>f" - -abbreviation \\<^isub>c :: "'av st option acom \ state set acom" -where "\\<^isub>c == map_acom \\<^isub>o" - -lemma gamma_f_Top[simp]: "\\<^isub>f Top = UNIV" -by(auto simp: Top_st_def \_st_def lookup_def) - -lemma gamma_o_Top[simp]: "\\<^isub>o Top = UNIV" -by (simp add: Top_option_def) - -(* FIXME (maybe also le \ sqle?) *) - -lemma mono_gamma_f: "f \ g \ \\<^isub>f f \ \\<^isub>f g" -apply(simp add:\_st_def subset_iff lookup_def le_st_def split: if_splits) -by (metis UNIV_I mono_gamma gamma_Top subsetD) - -lemma mono_gamma_o: - "sa \ sa' \ \\<^isub>o sa \ \\<^isub>o sa'" -by(induction sa sa' rule: le_option.induct)(simp_all add: mono_gamma_f) - -lemma mono_gamma_c: "ca \ ca' \ \\<^isub>c ca \ \\<^isub>c ca'" -by (induction ca ca' rule: le_acom.induct) (simp_all add:mono_gamma_o) - -lemma in_gamma_option_iff: - "x : \_option r u \ (\u'. u = Some u' \ x : r u')" -by (cases u) auto - -end - -end diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/ROOT.ML --- a/src/HOL/IMP/ROOT.ML Thu Apr 19 11:14:57 2012 +0200 +++ b/src/HOL/IMP/ROOT.ML Thu Apr 19 17:49:02 2012 +0200 @@ -1,5 +1,6 @@ no_document use_thys ["~~/src/HOL/ex/Interpretation_with_Defs", + "~~/src/HOL/Library/While_Combinator", "~~/src/HOL/Library/Char_ord", "~~/src/HOL/Library/List_lexord" ]; @@ -21,10 +22,10 @@ "HoareT", "Collecting1", "Collecting_list", - "Abs_Int0", - "Abs_Int0_parity", - "Abs_Int0_const", - "Abs_Int2", + "Abs_Int_Tests", + "Abs_Int_ITP/Abs_Int1_parity_ITP", + "Abs_Int_ITP/Abs_Int1_const_ITP", + "Abs_Int_ITP/Abs_Int3_ITP", "Abs_Int_Den/Abs_Int_den2", "Procs_Dyn_Vars_Dyn", "Procs_Stat_Vars_Dyn", diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IMP/document/root.tex --- a/src/HOL/IMP/document/root.tex Thu Apr 19 11:14:57 2012 +0200 +++ b/src/HOL/IMP/document/root.tex Thu Apr 19 17:49:02 2012 +0200 @@ -54,6 +54,7 @@ \author{TN \& GK} \maketitle +\setcounter{tocdepth}{2} \tableofcontents \newpage diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Thu Apr 19 11:14:57 2012 +0200 +++ b/src/HOL/IsaMakefile Thu Apr 19 17:49:02 2012 +0200 @@ -525,9 +525,12 @@ HOL-IMP: HOL $(OUT)/HOL-IMP $(OUT)/HOL-IMP: $(OUT)/HOL \ - IMP/ACom.thy IMP/Abs_Int0_fun.thy IMP/Abs_State.thy IMP/Abs_Int0.thy \ - IMP/Abs_Int0_const.thy IMP/Abs_Int1.thy IMP/Abs_Int1_ivl.thy \ - IMP/Abs_Int2.thy IMP/Abs_Int_Den/Abs_Int_den0_fun.thy \ + IMP/ACom.thy \ + IMP/Abs_Int_ITP/Abs_Int0_ITP.thy IMP/Abs_Int_ITP/Abs_State_ITP.thy \ + IMP/Abs_Int_ITP/Abs_Int1_ITP.thy IMP/Abs_Int_ITP/Abs_Int1_const_ITP.thy \ + IMP/Abs_Int_ITP/Abs_Int1_parity_ITP.thy IMP/Abs_Int_ITP/Abs_Int2_ITP.thy \ + IMP/Abs_Int_ITP/Abs_Int2_ivl_ITP.thy IMP/Abs_Int_ITP/Abs_Int3_ITP.thy \ + IMP/Abs_Int_Den/Abs_Int_den0_fun.thy \ IMP/Abs_Int_Den/Abs_State_den.thy IMP/Abs_Int_Den/Abs_Int_den0.thy \ IMP/Abs_Int_Den/Abs_Int_den0_const.thy IMP/Abs_Int_Den/Abs_Int_den1.thy \ IMP/Abs_Int_Den/Abs_Int_den1_ivl.thy IMP/Abs_Int_Den/Abs_Int_den2.thy \ diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/Library/Float.thy --- a/src/HOL/Library/Float.thy Thu Apr 19 11:14:57 2012 +0200 +++ b/src/HOL/Library/Float.thy Thu Apr 19 17:49:02 2012 +0200 @@ -1,509 +1,806 @@ -(* Title: HOL/Library/Float.thy - Author: Steven Obua 2008 - Author: Johannes Hoelzl, TU Muenchen 2008 / 2009 -*) - header {* Floating-Point Numbers *} theory Float -imports Complex_Main Lattice_Algebras +imports Complex_Main "~~/src/HOL/Library/Lattice_Algebras" begin -definition pow2 :: "int \ real" where - [simp]: "pow2 a = (if (0 <= a) then (2^(nat a)) else (inverse (2^(nat (-a)))))" - -datatype float = Float int int - -primrec of_float :: "float \ real" where - "of_float (Float a b) = real a * pow2 b" - -defs (overloaded) - real_of_float_def [code_unfold]: "real == of_float" - -declare [[coercion "% x . Float x 0"]] -declare [[coercion "real::float\real"]] - -primrec mantissa :: "float \ int" where - "mantissa (Float a b) = a" - -primrec scale :: "float \ int" where - "scale (Float a b) = b" - -instantiation float :: zero -begin -definition zero_float where "0 = Float 0 0" -instance .. -end - -instantiation float :: one -begin -definition one_float where "1 = Float 1 0" -instance .. -end - -lemma real_of_float_simp[simp]: "real (Float a b) = real a * pow2 b" - unfolding real_of_float_def using of_float.simps . - -lemma real_of_float_neg_exp: "e < 0 \ real (Float m e) = real m * inverse (2^nat (-e))" by auto -lemma real_of_float_nge0_exp: "\ 0 \ e \ real (Float m e) = real m * inverse (2^nat (-e))" by auto -lemma real_of_float_ge0_exp: "0 \ e \ real (Float m e) = real m * (2^nat e)" by auto - -lemma Float_num[simp]: shows - "real (Float 1 0) = 1" and "real (Float 1 1) = 2" and "real (Float 1 2) = 4" and - "real (Float 1 -1) = 1/2" and "real (Float 1 -2) = 1/4" and "real (Float 1 -3) = 1/8" and - "real (Float -1 0) = -1" and "real (Float (numeral n) 0) = numeral n" +typedef float = "{m * 2 powr e | (m :: int) (e :: int). True }" + morphisms real_of_float float_of by auto -lemma float_number_of_int[simp]: "real (Float n 0) = real n" - by simp +defs (overloaded) + real_of_float_def[code_unfold]: "real \ real_of_float" + +lemma type_definition_float': "type_definition real float_of float" + using type_definition_float unfolding real_of_float_def . -lemma pow2_0[simp]: "pow2 0 = 1" by simp -lemma pow2_1[simp]: "pow2 1 = 2" by simp -lemma pow2_neg: "pow2 x = inverse (pow2 (-x))" by simp +setup_lifting (no_code) type_definition_float' -lemma pow2_powr: "pow2 a = 2 powr a" - by (simp add: powr_realpow[symmetric] powr_minus) +lemmas float_of_inject[simp] -declare pow2_def[simp del] +declare [[coercion "real :: float \ real"]] -lemma pow2_add: "pow2 (a+b) = (pow2 a) * (pow2 b)" - by (simp add: pow2_powr powr_add) +lemma real_of_float_eq: + fixes f1 f2 :: float shows "f1 = f2 \ real f1 = real f2" + unfolding real_of_float_def real_of_float_inject .. + +lemma float_of_real[simp]: "float_of (real x) = x" + unfolding real_of_float_def by (rule real_of_float_inverse) -lemma pow2_diff: "pow2 (a - b) = pow2 a / pow2 b" - by (simp add: pow2_powr powr_divide2) - -lemma pow2_add1: "pow2 (1 + a) = 2 * (pow2 a)" - by (simp add: pow2_add) +lemma real_float[simp]: "x \ float \ real (float_of x) = x" + unfolding real_of_float_def by (rule float_of_inverse) -lemma float_components[simp]: "Float (mantissa f) (scale f) = f" by (cases f) auto +subsection {* Real operations preserving the representation as floating point number *} -lemma float_split: "\ a b. x = Float a b" by (cases x) auto +lemma floatI: fixes m e :: int shows "m * 2 powr e = x \ x \ float" + by (auto simp: float_def) -lemma float_split2: "(\ a b. x \ Float a b) = False" by (auto simp add: float_split) - -lemma float_zero[simp]: "real (Float 0 e) = 0" by simp - -lemma abs_div_2_less: "a \ 0 \ a \ -1 \ abs((a::int) div 2) < abs a" -by arith +lemma zero_float[simp]: "0 \ float" by (auto simp: float_def) +lemma one_float[simp]: "1 \ float" by (intro floatI[of 1 0]) simp +lemma numeral_float[simp]: "numeral i \ float" by (intro floatI[of "numeral i" 0]) simp +lemma neg_numeral_float[simp]: "neg_numeral i \ float" by (intro floatI[of "neg_numeral i" 0]) simp +lemma real_of_int_float[simp]: "real (x :: int) \ float" by (intro floatI[of x 0]) simp +lemma real_of_nat_float[simp]: "real (x :: nat) \ float" by (intro floatI[of x 0]) simp +lemma two_powr_int_float[simp]: "2 powr (real (i::int)) \ float" by (intro floatI[of 1 i]) simp +lemma two_powr_nat_float[simp]: "2 powr (real (i::nat)) \ float" by (intro floatI[of 1 i]) simp +lemma two_powr_minus_int_float[simp]: "2 powr - (real (i::int)) \ float" by (intro floatI[of 1 "-i"]) simp +lemma two_powr_minus_nat_float[simp]: "2 powr - (real (i::nat)) \ float" by (intro floatI[of 1 "-i"]) simp +lemma two_powr_numeral_float[simp]: "2 powr numeral i \ float" by (intro floatI[of 1 "numeral i"]) simp +lemma two_powr_neg_numeral_float[simp]: "2 powr neg_numeral i \ float" by (intro floatI[of 1 "neg_numeral i"]) simp +lemma two_pow_float[simp]: "2 ^ n \ float" by (intro floatI[of 1 "n"]) (simp add: powr_realpow) +lemma real_of_float_float[simp]: "real (f::float) \ float" by (cases f) simp -function normfloat :: "float \ float" where - "normfloat (Float a b) = - (if a \ 0 \ even a then normfloat (Float (a div 2) (b+1)) - else if a=0 then Float 0 0 else Float a b)" -by pat_completeness auto -termination by (relation "measure (nat o abs o mantissa)") (auto intro: abs_div_2_less) -declare normfloat.simps[simp del] - -theorem normfloat[symmetric, simp]: "real f = real (normfloat f)" -proof (induct f rule: normfloat.induct) - case (1 a b) - have real2: "2 = real (2::int)" - by auto - show ?case - apply (subst normfloat.simps) - apply auto - apply (subst 1[symmetric]) - apply (auto simp add: pow2_add even_def) - done +lemma plus_float[simp]: "r \ float \ p \ float \ r + p \ float" + unfolding float_def +proof (safe, simp) + fix e1 m1 e2 m2 :: int + { fix e1 m1 e2 m2 :: int assume "e1 \ e2" + then have "m1 * 2 powr e1 + m2 * 2 powr e2 = (m1 + m2 * 2 ^ nat (e2 - e1)) * 2 powr e1" + by (simp add: powr_realpow[symmetric] powr_divide2[symmetric] field_simps) + then have "\(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e" + by blast } + note * = this + show "\(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e" + proof (cases e1 e2 rule: linorder_le_cases) + assume "e2 \ e1" from *[OF this, of m2 m1] show ?thesis by (simp add: ac_simps) + qed (rule *) qed -lemma pow2_neq_zero[simp]: "pow2 x \ 0" - by (auto simp add: pow2_def) +lemma uminus_float[simp]: "x \ float \ -x \ float" + apply (auto simp: float_def) + apply (rule_tac x="-x" in exI) + apply (rule_tac x="xa" in exI) + apply (simp add: field_simps) + done -lemma pow2_int: "pow2 (int c) = 2^c" - by (simp add: pow2_def) +lemma times_float[simp]: "x \ float \ y \ float \ x * y \ float" + apply (auto simp: float_def) + apply (rule_tac x="x * xa" in exI) + apply (rule_tac x="xb + xc" in exI) + apply (simp add: powr_add) + done -lemma zero_less_pow2[simp]: "0 < pow2 x" - by (simp add: pow2_powr) +lemma minus_float[simp]: "x \ float \ y \ float \ x - y \ float" + unfolding ab_diff_minus by (intro uminus_float plus_float) + +lemma abs_float[simp]: "x \ float \ abs x \ float" + by (cases x rule: linorder_cases[of 0]) auto + +lemma sgn_of_float[simp]: "x \ float \ sgn x \ float" + by (cases x rule: linorder_cases[of 0]) (auto intro!: uminus_float) -lemma normfloat_imp_odd_or_zero: - "normfloat f = Float a b \ odd a \ (a = 0 \ b = 0)" -proof (induct f rule: normfloat.induct) - case (1 u v) - from 1 have ab: "normfloat (Float u v) = Float a b" by auto - { - assume eu: "even u" - assume z: "u \ 0" - have "normfloat (Float u v) = normfloat (Float (u div 2) (v + 1))" - apply (subst normfloat.simps) - by (simp add: eu z) - with ab have "normfloat (Float (u div 2) (v + 1)) = Float a b" by simp - with 1 eu z have ?case by auto - } - note case1 = this - { - assume "odd u \ u = 0" - then have ou: "\ (u \ 0 \ even u)" by auto - have "normfloat (Float u v) = (if u = 0 then Float 0 0 else Float u v)" - apply (subst normfloat.simps) - apply (simp add: ou) - done - with ab have "Float a b = (if u = 0 then Float 0 0 else Float u v)" by auto - then have ?case - apply (case_tac "u=0") - apply (auto) - by (insert ou, auto) - } - note case2 = this - show ?case - apply (case_tac "odd u \ u = 0") - apply (rule case2) - apply simp - apply (rule case1) - apply auto - done +lemma div_power_2_float[simp]: "x \ float \ x / 2^d \ float" + apply (auto simp add: float_def) + apply (rule_tac x="x" in exI) + apply (rule_tac x="xa - d" in exI) + apply (simp add: powr_realpow[symmetric] field_simps powr_add[symmetric]) + done + +lemma div_power_2_int_float[simp]: "x \ float \ x / (2::int)^d \ float" + apply (auto simp add: float_def) + apply (rule_tac x="x" in exI) + apply (rule_tac x="xa - d" in exI) + apply (simp add: powr_realpow[symmetric] field_simps powr_add[symmetric]) + done + +lemma div_numeral_Bit0_float[simp]: + assumes x: "x / numeral n \ float" shows "x / (numeral (Num.Bit0 n)) \ float" +proof - + have "(x / numeral n) / 2^1 \ float" + by (intro x div_power_2_float) + also have "(x / numeral n) / 2^1 = x / (numeral (Num.Bit0 n))" + by (induct n) auto + finally show ?thesis . +qed + +lemma div_neg_numeral_Bit0_float[simp]: + assumes x: "x / numeral n \ float" shows "x / (neg_numeral (Num.Bit0 n)) \ float" +proof - + have "- (x / numeral (Num.Bit0 n)) \ float" using x by simp + also have "- (x / numeral (Num.Bit0 n)) = x / neg_numeral (Num.Bit0 n)" + unfolding neg_numeral_def by (simp del: minus_numeral) + finally show ?thesis . qed -lemma float_eq_odd_helper: - assumes odd: "odd a'" - and floateq: "real (Float a b) = real (Float a' b')" - shows "b \ b'" -proof - - from odd have "a' \ 0" by auto - with floateq have a': "real a' = real a * pow2 (b - b')" - by (simp add: pow2_diff field_simps) +lift_definition Float :: "int \ int \ float" is "\(m::int) (e::int). m * 2 powr e" by simp +declare Float.rep_eq[simp] + +code_datatype Float + +subsection {* Arithmetic operations on floating point numbers *} + +instantiation float :: "{ring_1, linorder, linordered_ring, linordered_idom, numeral, equal}" +begin - { - assume bcmp: "b > b'" - then obtain c :: nat where "b - b' = int c + 1" - by atomize_elim arith - with a' have "real a' = real (a * 2^c * 2)" - by (simp add: pow2_def nat_add_distrib) - with odd have False - unfolding real_of_int_inject by simp - } - then show ?thesis by arith -qed +lift_definition zero_float :: float is 0 by simp +declare zero_float.rep_eq[simp] +lift_definition one_float :: float is 1 by simp +declare one_float.rep_eq[simp] +lift_definition plus_float :: "float \ float \ float" is "op +" by simp +declare plus_float.rep_eq[simp] +lift_definition times_float :: "float \ float \ float" is "op *" by simp +declare times_float.rep_eq[simp] +lift_definition minus_float :: "float \ float \ float" is "op -" by simp +declare minus_float.rep_eq[simp] +lift_definition uminus_float :: "float \ float" is "uminus" by simp +declare uminus_float.rep_eq[simp] + +lift_definition abs_float :: "float \ float" is abs by simp +declare abs_float.rep_eq[simp] +lift_definition sgn_float :: "float \ float" is sgn by simp +declare sgn_float.rep_eq[simp] + +lift_definition equal_float :: "float \ float \ bool" is "op = :: real \ real \ bool" .. -lemma float_eq_odd: - assumes odd1: "odd a" - and odd2: "odd a'" - and floateq: "real (Float a b) = real (Float a' b')" - shows "a = a' \ b = b'" -proof - - from - float_eq_odd_helper[OF odd2 floateq] - float_eq_odd_helper[OF odd1 floateq[symmetric]] - have beq: "b = b'" by arith - with floateq show ?thesis by auto -qed +lift_definition less_eq_float :: "float \ float \ bool" is "op \" .. +declare less_eq_float.rep_eq[simp] +lift_definition less_float :: "float \ float \ bool" is "op <" .. +declare less_float.rep_eq[simp] + +instance + proof qed (transfer, fastforce simp add: field_simps intro: mult_left_mono mult_right_mono)+ +end + +lemma real_of_float_power[simp]: fixes f::float shows "real (f^n) = real f^n" + by (induct n) simp_all + +lemma fixes x y::float + shows real_of_float_min: "real (min x y) = min (real x) (real y)" + and real_of_float_max: "real (max x y) = max (real x) (real y)" + by (simp_all add: min_def max_def) -theorem normfloat_unique: - assumes real_of_float_eq: "real f = real g" - shows "normfloat f = normfloat g" -proof - - from float_split[of "normfloat f"] obtain a b where normf:"normfloat f = Float a b" by auto - from float_split[of "normfloat g"] obtain a' b' where normg:"normfloat g = Float a' b'" by auto - have "real (normfloat f) = real (normfloat g)" - by (simp add: real_of_float_eq) - then have float_eq: "real (Float a b) = real (Float a' b')" - by (simp add: normf normg) - have ab: "odd a \ (a = 0 \ b = 0)" by (rule normfloat_imp_odd_or_zero[OF normf]) - have ab': "odd a' \ (a' = 0 \ b' = 0)" by (rule normfloat_imp_odd_or_zero[OF normg]) - { - assume odd: "odd a" - then have "a \ 0" by (simp add: even_def) arith - with float_eq have "a' \ 0" by auto - with ab' have "odd a'" by simp - from odd this float_eq have "a = a' \ b = b'" by (rule float_eq_odd) - } - note odd_case = this - { - assume even: "even a" - with ab have a0: "a = 0" by simp - with float_eq have a0': "a' = 0" by auto - from a0 a0' ab ab' have "a = a' \ b = b'" by auto - } - note even_case = this - from odd_case even_case show ?thesis - apply (simp add: normf normg) - apply (case_tac "even a") - apply auto +instance float :: dense_linorder +proof + fix a b :: float + show "\c. a < c" + apply (intro exI[of _ "a + 1"]) + apply transfer + apply simp + done + show "\c. c < a" + apply (intro exI[of _ "a - 1"]) + apply transfer + apply simp + done + assume "a < b" + then show "\c. a < c \ c < b" + apply (intro exI[of _ "(a + b) * Float 1 -1"]) + apply transfer + apply (simp add: powr_neg_numeral) done qed -instantiation float :: plus -begin -fun plus_float where -[simp del]: "(Float a_m a_e) + (Float b_m b_e) = - (if a_e \ b_e then Float (a_m + b_m * 2^(nat(b_e - a_e))) a_e - else Float (a_m * 2^(nat (a_e - b_e)) + b_m) b_e)" -instance .. -end - -instantiation float :: uminus +instantiation float :: lattice_ab_group_add begin -primrec uminus_float where [simp del]: "uminus_float (Float m e) = Float (-m) e" -instance .. -end + +definition inf_float::"float\float\float" +where "inf_float a b = min a b" -instantiation float :: minus -begin -definition minus_float where "(z::float) - w = z + (- w)" -instance .. -end +definition sup_float::"float\float\float" +where "sup_float a b = max a b" -instantiation float :: times -begin -fun times_float where [simp del]: "(Float a_m a_e) * (Float b_m b_e) = Float (a_m * b_m) (a_e + b_e)" -instance .. +instance + by default + (transfer, simp_all add: inf_float_def sup_float_def real_of_float_min real_of_float_max)+ end -primrec float_pprt :: "float \ float" where - "float_pprt (Float a e) = (if 0 <= a then (Float a e) else 0)" - -primrec float_nprt :: "float \ float" where - "float_nprt (Float a e) = (if 0 <= a then 0 else (Float a e))" - -instantiation float :: ord -begin -definition le_float_def: "z \ (w :: float) \ real z \ real w" -definition less_float_def: "z < (w :: float) \ real z < real w" -instance .. -end - -lemma real_of_float_add[simp]: "real (a + b) = real a + real (b :: float)" - by (cases a, cases b) (simp add: algebra_simps plus_float.simps, - auto simp add: pow2_int[symmetric] pow2_add[symmetric]) - -lemma real_of_float_minus[simp]: "real (- a) = - real (a :: float)" - by (cases a) simp - -lemma real_of_float_sub[simp]: "real (a - b) = real a - real (b :: float)" - by (cases a, cases b) (simp add: minus_float_def) - -lemma real_of_float_mult[simp]: "real (a*b) = real a * real (b :: float)" - by (cases a, cases b) (simp add: times_float.simps pow2_add) - -lemma real_of_float_0[simp]: "real (0 :: float) = 0" - by (auto simp add: zero_float_def) - -lemma real_of_float_1[simp]: "real (1 :: float) = 1" - by (auto simp add: one_float_def) - -lemma zero_le_float: - "(0 <= real (Float a b)) = (0 <= a)" - apply auto - apply (auto simp add: zero_le_mult_iff) - apply (insert zero_less_pow2[of b]) - apply (simp_all) +lemma float_numeral[simp]: "real (numeral x :: float) = numeral x" + apply (induct x) + apply simp + apply (simp_all only: numeral_Bit0 numeral_Bit1 real_of_float_eq real_float + plus_float.rep_eq one_float.rep_eq plus_float numeral_float one_float) done -lemma float_le_zero: - "(real (Float a b) <= 0) = (a <= 0)" - apply auto - apply (auto simp add: mult_le_0_iff) - apply (insert zero_less_pow2[of b]) - apply auto - done +lemma transfer_numeral [transfer_rule]: + "fun_rel (op =) cr_float (numeral :: _ \ real) (numeral :: _ \ float)" + unfolding fun_rel_def cr_float_def by (simp add: real_of_float_def[symmetric]) + +lemma float_neg_numeral[simp]: "real (neg_numeral x :: float) = neg_numeral x" + by (simp add: minus_numeral[symmetric] del: minus_numeral) + +lemma transfer_neg_numeral [transfer_rule]: + "fun_rel (op =) cr_float (neg_numeral :: _ \ real) (neg_numeral :: _ \ float)" + unfolding fun_rel_def cr_float_def by (simp add: real_of_float_def[symmetric]) + +lemma + shows float_of_numeral[simp]: "numeral k = float_of (numeral k)" + and float_of_neg_numeral[simp]: "neg_numeral k = float_of (neg_numeral k)" + unfolding real_of_float_eq by simp_all + +subsection {* Represent floats as unique mantissa and exponent *} -lemma zero_less_float: - "(0 < real (Float a b)) = (0 < a)" - apply auto - apply (auto simp add: zero_less_mult_iff) - apply (insert zero_less_pow2[of b]) - apply (simp_all) - done +lemma int_induct_abs[case_names less]: + fixes j :: int + assumes H: "\n. (\i. \i\ < \n\ \ P i) \ P n" + shows "P j" +proof (induct "nat \j\" arbitrary: j rule: less_induct) + case less show ?case by (rule H[OF less]) simp +qed + +lemma int_cancel_factors: + fixes n :: int assumes "1 < r" shows "n = 0 \ (\k i. n = k * r ^ i \ \ r dvd k)" +proof (induct n rule: int_induct_abs) + case (less n) + { fix m assume n: "n \ 0" "n = m * r" + then have "\m \ < \n\" + by (metis abs_dvd_iff abs_ge_self assms comm_semiring_1_class.normalizing_semiring_rules(7) + dvd_imp_le_int dvd_refl dvd_triv_right linorder_neq_iff linorder_not_le + mult_eq_0_iff zdvd_mult_cancel1) + from less[OF this] n have "\k i. n = k * r ^ Suc i \ \ r dvd k" by auto } + then show ?case + by (metis comm_semiring_1_class.normalizing_semiring_rules(12,7) dvdE power_0) +qed -lemma float_less_zero: - "(real (Float a b) < 0) = (a < 0)" - apply auto - apply (auto simp add: mult_less_0_iff) - apply (insert zero_less_pow2[of b]) - apply (simp_all) - done - -declare real_of_float_simp[simp del] - -lemma real_of_float_pprt[simp]: "real (float_pprt a) = pprt (real a)" - by (cases a) (auto simp add: zero_le_float float_le_zero) +lemma mult_powr_eq_mult_powr_iff_asym: + fixes m1 m2 e1 e2 :: int + assumes m1: "\ 2 dvd m1" and "e1 \ e2" + shows "m1 * 2 powr e1 = m2 * 2 powr e2 \ m1 = m2 \ e1 = e2" +proof + have "m1 \ 0" using m1 unfolding dvd_def by auto + assume eq: "m1 * 2 powr e1 = m2 * 2 powr e2" + with `e1 \ e2` have "m1 = m2 * 2 powr nat (e2 - e1)" + by (simp add: powr_divide2[symmetric] field_simps) + also have "\ = m2 * 2^nat (e2 - e1)" + by (simp add: powr_realpow) + finally have m1_eq: "m1 = m2 * 2^nat (e2 - e1)" + unfolding real_of_int_inject . + with m1 have "m1 = m2" + by (cases "nat (e2 - e1)") (auto simp add: dvd_def) + then show "m1 = m2 \ e1 = e2" + using eq `m1 \ 0` by (simp add: powr_inj) +qed simp -lemma real_of_float_nprt[simp]: "real (float_nprt a) = nprt (real a)" - by (cases a) (auto simp add: zero_le_float float_le_zero) +lemma mult_powr_eq_mult_powr_iff: + fixes m1 m2 e1 e2 :: int + shows "\ 2 dvd m1 \ \ 2 dvd m2 \ m1 * 2 powr e1 = m2 * 2 powr e2 \ m1 = m2 \ e1 = e2" + using mult_powr_eq_mult_powr_iff_asym[of m1 e1 e2 m2] + using mult_powr_eq_mult_powr_iff_asym[of m2 e2 e1 m1] + by (cases e1 e2 rule: linorder_le_cases) auto -instance float :: ab_semigroup_add -proof (intro_classes) - fix a b c :: float - show "a + b + c = a + (b + c)" - by (cases a, cases b, cases c) - (auto simp add: algebra_simps power_add[symmetric] plus_float.simps) -next - fix a b :: float - show "a + b = b + a" - by (cases a, cases b) (simp add: plus_float.simps) +lemma floatE_normed: + assumes x: "x \ float" + obtains (zero) "x = 0" + | (powr) m e :: int where "x = m * 2 powr e" "\ 2 dvd m" "x \ 0" +proof atomize_elim + { assume "x \ 0" + from x obtain m e :: int where x: "x = m * 2 powr e" by (auto simp: float_def) + with `x \ 0` int_cancel_factors[of 2 m] obtain k i where "m = k * 2 ^ i" "\ 2 dvd k" + by auto + with `\ 2 dvd k` x have "\(m::int) (e::int). x = m * 2 powr e \ \ (2::int) dvd m" + by (rule_tac exI[of _ "k"], rule_tac exI[of _ "e + int i"]) + (simp add: powr_add powr_realpow) } + then show "x = 0 \ (\(m::int) (e::int). x = m * 2 powr e \ \ (2::int) dvd m \ x \ 0)" + by blast +qed + +lemma float_normed_cases: + fixes f :: float + obtains (zero) "f = 0" + | (powr) m e :: int where "real f = m * 2 powr e" "\ 2 dvd m" "f \ 0" +proof (atomize_elim, induct f) + case (float_of y) then show ?case + by (cases rule: floatE_normed) (auto simp: zero_float_def) qed -instance float :: numeral .. +definition mantissa :: "float \ int" where + "mantissa f = fst (SOME p::int \ int. (f = 0 \ fst p = 0 \ snd p = 0) + \ (f \ 0 \ real f = real (fst p) * 2 powr real (snd p) \ \ 2 dvd fst p))" + +definition exponent :: "float \ int" where + "exponent f = snd (SOME p::int \ int. (f = 0 \ fst p = 0 \ snd p = 0) + \ (f \ 0 \ real f = real (fst p) * 2 powr real (snd p) \ \ 2 dvd fst p))" -lemma Float_add_same_scale: "Float x e + Float y e = Float (x + y) e" - by (simp add: plus_float.simps) +lemma + shows exponent_0[simp]: "exponent (float_of 0) = 0" (is ?E) + and mantissa_0[simp]: "mantissa (float_of 0) = 0" (is ?M) +proof - + have "\p::int \ int. fst p = 0 \ snd p = 0 \ p = (0, 0)" by auto + then show ?E ?M + by (auto simp add: mantissa_def exponent_def zero_float_def) +qed -(* FIXME: define other constant for code_unfold_post *) -lemma numeral_float_Float (*[code_unfold_post]*): - "numeral k = Float (numeral k) 0" - by (induct k, simp_all only: numeral.simps one_float_def - Float_add_same_scale) +lemma + shows mantissa_exponent: "real f = mantissa f * 2 powr exponent f" (is ?E) + and mantissa_not_dvd: "f \ (float_of 0) \ \ 2 dvd mantissa f" (is "_ \ ?D") +proof cases + assume [simp]: "f \ (float_of 0)" + have "f = mantissa f * 2 powr exponent f \ \ 2 dvd mantissa f" + proof (cases f rule: float_normed_cases) + case (powr m e) + then have "\p::int \ int. (f = 0 \ fst p = 0 \ snd p = 0) + \ (f \ 0 \ real f = real (fst p) * 2 powr real (snd p) \ \ 2 dvd fst p)" + by auto + then show ?thesis + unfolding exponent_def mantissa_def + by (rule someI2_ex) (simp add: zero_float_def) + qed (simp add: zero_float_def) + then show ?E ?D by auto +qed simp -lemma float_number_of[simp]: "real (numeral x :: float) = numeral x" - by (simp only: numeral_float_Float Float_num) +lemma mantissa_noteq_0: "f \ float_of 0 \ mantissa f \ 0" + using mantissa_not_dvd[of f] by auto - -instance float :: comm_monoid_mult -proof (intro_classes) - fix a b c :: float - show "a * b * c = a * (b * c)" - by (cases a, cases b, cases c) (simp add: times_float.simps) +lemma + fixes m e :: int + defines "f \ float_of (m * 2 powr e)" + assumes dvd: "\ 2 dvd m" + shows mantissa_float: "mantissa f = m" (is "?M") + and exponent_float: "m \ 0 \ exponent f = e" (is "_ \ ?E") +proof cases + assume "m = 0" with dvd show "mantissa f = m" by auto next - fix a b :: float - show "a * b = b * a" - by (cases a, cases b) (simp add: times_float.simps) -next - fix a :: float - show "1 * a = a" - by (cases a) (simp add: times_float.simps one_float_def) + assume "m \ 0" + then have f_not_0: "f \ float_of 0" by (simp add: f_def) + from mantissa_exponent[of f] + have "m * 2 powr e = mantissa f * 2 powr exponent f" + by (auto simp add: f_def) + then show "?M" "?E" + using mantissa_not_dvd[OF f_not_0] dvd + by (auto simp: mult_powr_eq_mult_powr_iff) +qed + +subsection {* Compute arithmetic operations *} + +lemma real_of_float_Float[code]: "real_of_float (Float m e) = + (if e \ 0 then m * 2 ^ nat e else m * inverse (2 ^ nat (- e)))" +by (auto simp add: powr_realpow[symmetric] powr_minus real_of_float_def[symmetric]) + +lemma Float_mantissa_exponent: "Float (mantissa f) (exponent f) = f" + unfolding real_of_float_eq mantissa_exponent[of f] by simp + +lemma Float_cases[case_names Float, cases type: float]: + fixes f :: float + obtains (Float) m e :: int where "f = Float m e" + using Float_mantissa_exponent[symmetric] + by (atomize_elim) auto + +lemma denormalize_shift: + assumes f_def: "f \ Float m e" and not_0: "f \ float_of 0" + obtains i where "m = mantissa f * 2 ^ i" "e = exponent f - i" +proof + from mantissa_exponent[of f] f_def + have "m * 2 powr e = mantissa f * 2 powr exponent f" + by simp + then have eq: "m = mantissa f * 2 powr (exponent f - e)" + by (simp add: powr_divide2[symmetric] field_simps) + moreover + have "e \ exponent f" + proof (rule ccontr) + assume "\ e \ exponent f" + then have pos: "exponent f < e" by simp + then have "2 powr (exponent f - e) = 2 powr - real (e - exponent f)" + by simp + also have "\ = 1 / 2^nat (e - exponent f)" + using pos by (simp add: powr_realpow[symmetric] powr_divide2[symmetric]) + finally have "m * 2^nat (e - exponent f) = real (mantissa f)" + using eq by simp + then have "mantissa f = m * 2^nat (e - exponent f)" + unfolding real_of_int_inject by simp + with `exponent f < e` have "2 dvd mantissa f" + apply (intro dvdI[where k="m * 2^(nat (e-exponent f)) div 2"]) + apply (cases "nat (e - exponent f)") + apply auto + done + then show False using mantissa_not_dvd[OF not_0] by simp + qed + ultimately have "real m = mantissa f * 2^nat (exponent f - e)" + by (simp add: powr_realpow[symmetric]) + with `e \ exponent f` + show "m = mantissa f * 2 ^ nat (exponent f - e)" "e = exponent f - nat (exponent f - e)" + unfolding real_of_int_inject by auto qed -(* Floats do NOT form a cancel_semigroup_add: *) -lemma "0 + Float 0 1 = 0 + Float 0 2" - by (simp add: plus_float.simps zero_float_def) +lemma compute_zero[code_unfold, code]: "0 = Float 0 0" + by transfer simp + +lemma compute_one[code_unfold, code]: "1 = Float 1 0" + by transfer simp + +definition normfloat :: "float \ float" where + [simp]: "normfloat x = x" + +lemma compute_normfloat[code]: "normfloat (Float m e) = + (if m mod 2 = 0 \ m \ 0 then normfloat (Float (m div 2) (e + 1)) + else if m = 0 then 0 else Float m e)" + unfolding normfloat_def + by transfer (auto simp add: powr_add zmod_eq_0_iff) + +lemma compute_float_numeral[code_abbrev]: "Float (numeral k) 0 = numeral k" + by transfer simp + +lemma compute_float_neg_numeral[code_abbrev]: "Float (neg_numeral k) 0 = neg_numeral k" + by transfer simp + +lemma compute_float_uminus[code]: "- Float m1 e1 = Float (- m1) e1" + by transfer simp + +lemma compute_float_times[code]: "Float m1 e1 * Float m2 e2 = Float (m1 * m2) (e1 + e2)" + by transfer (simp add: field_simps powr_add) + +lemma compute_float_plus[code]: "Float m1 e1 + Float m2 e2 = + (if e1 \ e2 then Float (m1 + m2 * 2^nat (e2 - e1)) e1 + else Float (m2 + m1 * 2^nat (e1 - e2)) e2)" + by transfer (simp add: field_simps powr_realpow[symmetric] powr_divide2[symmetric]) + +lemma compute_float_minus[code]: fixes f g::float shows "f - g = f + (-g)" + by simp + +lemma compute_float_sgn[code]: "sgn (Float m1 e1) = (if 0 < m1 then 1 else if m1 < 0 then -1 else 0)" + by transfer (simp add: sgn_times) + +lift_definition is_float_pos :: "float \ bool" is "op < 0 :: real \ bool" .. + +lemma compute_is_float_pos[code]: "is_float_pos (Float m e) \ 0 < m" + by transfer (auto simp add: zero_less_mult_iff not_le[symmetric, of _ 0]) + +lemma compute_float_less[code]: "a < b \ is_float_pos (b - a)" + by transfer (simp add: field_simps) + +lift_definition is_float_nonneg :: "float \ bool" is "op \ 0 :: real \ bool" .. + +lemma compute_is_float_nonneg[code]: "is_float_nonneg (Float m e) \ 0 \ m" + by transfer (auto simp add: zero_le_mult_iff not_less[symmetric, of _ 0]) -instance float :: comm_semiring -proof (intro_classes) - fix a b c :: float - show "(a + b) * c = a * c + b * c" - by (cases a, cases b, cases c) (simp add: plus_float.simps times_float.simps algebra_simps) -qed +lemma compute_float_le[code]: "a \ b \ is_float_nonneg (b - a)" + by transfer (simp add: field_simps) + +lift_definition is_float_zero :: "float \ bool" is "op = 0 :: real \ bool" by simp + +lemma compute_is_float_zero[code]: "is_float_zero (Float m e) \ 0 = m" + by transfer (auto simp add: is_float_zero_def) + +lemma compute_float_abs[code]: "abs (Float m e) = Float (abs m) e" + by transfer (simp add: abs_mult) + +lemma compute_float_eq[code]: "equal_class.equal f g = is_float_zero (f - g)" + by transfer simp + +subsection {* Rounding Real numbers *} + +definition round_down :: "int \ real \ real" where + "round_down prec x = floor (x * 2 powr prec) * 2 powr -prec" + +definition round_up :: "int \ real \ real" where + "round_up prec x = ceiling (x * 2 powr prec) * 2 powr -prec" + +lemma round_down_float[simp]: "round_down prec x \ float" + unfolding round_down_def + by (auto intro!: times_float simp: real_of_int_minus[symmetric] simp del: real_of_int_minus) -(* Floats do NOT form an order, because "(x < y) = (x <= y & x <> y)" does NOT hold *) +lemma round_up_float[simp]: "round_up prec x \ float" + unfolding round_up_def + by (auto intro!: times_float simp: real_of_int_minus[symmetric] simp del: real_of_int_minus) + +lemma round_up: "x \ round_up prec x" + by (simp add: powr_minus_divide le_divide_eq round_up_def) + +lemma round_down: "round_down prec x \ x" + by (simp add: powr_minus_divide divide_le_eq round_down_def) + +lemma round_up_0[simp]: "round_up p 0 = 0" + unfolding round_up_def by simp -instance float :: zero_neq_one -proof (intro_classes) - show "(0::float) \ 1" - by (simp add: zero_float_def one_float_def) +lemma round_down_0[simp]: "round_down p 0 = 0" + unfolding round_down_def by simp + +lemma round_up_diff_round_down: + "round_up prec x - round_down prec x \ 2 powr -prec" +proof - + have "round_up prec x - round_down prec x = + (ceiling (x * 2 powr prec) - floor (x * 2 powr prec)) * 2 powr -prec" + by (simp add: round_up_def round_down_def field_simps) + also have "\ \ 1 * 2 powr -prec" + by (rule mult_mono) + (auto simp del: real_of_int_diff + simp: real_of_int_diff[symmetric] real_of_int_le_one_cancel_iff ceiling_diff_floor_le_1) + finally show ?thesis by simp qed -lemma float_le_simp: "((x::float) \ y) = (0 \ y - x)" - by (auto simp add: le_float_def) - -lemma float_less_simp: "((x::float) < y) = (0 < y - x)" - by (auto simp add: less_float_def) +lemma round_down_shift: "round_down p (x * 2 powr k) = 2 powr k * round_down (p + k) x" + unfolding round_down_def + by (simp add: powr_add powr_mult field_simps powr_divide2[symmetric]) + (simp add: powr_add[symmetric]) -lemma real_of_float_min: "real (min x y :: float) = min (real x) (real y)" unfolding min_def le_float_def by auto -lemma real_of_float_max: "real (max a b :: float) = max (real a) (real b)" unfolding max_def le_float_def by auto +lemma round_up_shift: "round_up p (x * 2 powr k) = 2 powr k * round_up (p + k) x" + unfolding round_up_def + by (simp add: powr_add powr_mult field_simps powr_divide2[symmetric]) + (simp add: powr_add[symmetric]) -lemma float_power: "real (x ^ n :: float) = real x ^ n" - by (induct n) simp_all +subsection {* Rounding Floats *} -lemma zero_le_pow2[simp]: "0 \ pow2 s" - apply (subgoal_tac "0 < pow2 s") - apply (auto simp only:) - apply auto - done +lift_definition float_up :: "int \ float \ float" is round_up by simp +declare float_up.rep_eq[simp] -lemma pow2_less_0_eq_False[simp]: "(pow2 s < 0) = False" - apply auto - apply (subgoal_tac "0 \ pow2 s") - apply simp - apply simp - done +lemma float_up_correct: + shows "real (float_up e f) - real f \ {0..2 powr -e}" +unfolding atLeastAtMost_iff +proof + have "round_up e f - f \ round_up e f - round_down e f" using round_down by simp + also have "\ \ 2 powr -e" using round_up_diff_round_down by simp + finally show "real (float_up e f) - real f \ 2 powr real (- e)" + by simp +qed (simp add: algebra_simps round_up) + +lift_definition float_down :: "int \ float \ float" is round_down by simp +declare float_down.rep_eq[simp] -lemma pow2_le_0_eq_False[simp]: "(pow2 s \ 0) = False" - apply auto - apply (subgoal_tac "0 < pow2 s") - apply simp - apply simp - done - -lemma float_pos_m_pos: "0 < Float m e \ 0 < m" - unfolding less_float_def real_of_float_simp real_of_float_0 zero_less_mult_iff - by auto +lemma float_down_correct: + shows "real f - real (float_down e f) \ {0..2 powr -e}" +unfolding atLeastAtMost_iff +proof + have "f - round_down e f \ round_up e f - round_down e f" using round_up by simp + also have "\ \ 2 powr -e" using round_up_diff_round_down by simp + finally show "real f - real (float_down e f) \ 2 powr real (- e)" + by simp +qed (simp add: algebra_simps round_down) -lemma float_pos_less1_e_neg: assumes "0 < Float m e" and "Float m e < 1" shows "e < 0" -proof - - have "0 < m" using float_pos_m_pos `0 < Float m e` by auto - hence "0 \ real m" and "1 \ real m" by auto - - show "e < 0" - proof (rule ccontr) - assume "\ e < 0" hence "0 \ e" by auto - hence "1 \ pow2 e" unfolding pow2_def by auto - from mult_mono[OF `1 \ real m` this `0 \ real m`] - have "1 \ Float m e" by (simp add: le_float_def real_of_float_simp) - thus False using `Float m e < 1` unfolding less_float_def le_float_def by auto - qed +lemma compute_float_down[code]: + "float_down p (Float m e) = + (if p + e < 0 then Float (m div 2^nat (-(p + e))) (-p) else Float m e)" +proof cases + assume "p + e < 0" + hence "real ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))" + using powr_realpow[of 2 "nat (-(p + e))"] by simp + also have "... = 1 / 2 powr p / 2 powr e" + unfolding powr_minus_divide real_of_int_minus by (simp add: powr_add) + finally show ?thesis + using `p + e < 0` + by transfer (simp add: ac_simps round_down_def floor_divide_eq_div[symmetric]) +next + assume "\ p + e < 0" + then have r: "real e + real p = real (nat (e + p))" by simp + have r: "\(m * 2 powr e) * 2 powr real p\ = (m * 2 powr e) * 2 powr real p" + by (auto intro: exI[where x="m*2^nat (e+p)"] + simp add: ac_simps powr_add[symmetric] r powr_realpow) + with `\ p + e < 0` show ?thesis + by transfer + (auto simp add: round_down_def field_simps powr_add powr_minus inverse_eq_divide) qed -lemma float_less1_mantissa_bound: assumes "0 < Float m e" "Float m e < 1" shows "m < 2^(nat (-e))" +lemma ceil_divide_floor_conv: +assumes "b \ 0" +shows "\real a / real b\ = (if b dvd a then a div b else \real a / real b\ + 1)" +proof cases + assume "\ b dvd a" + hence "a mod b \ 0" by auto + hence ne: "real (a mod b) / real b \ 0" using `b \ 0` by auto + have "\real a / real b\ = \real a / real b\ + 1" + apply (rule ceiling_eq) apply (auto simp: floor_divide_eq_div[symmetric]) + proof - + have "real \real a / real b\ \ real a / real b" by simp + moreover have "real \real a / real b\ \ real a / real b" + apply (subst (2) real_of_int_div_aux) unfolding floor_divide_eq_div using ne `b \ 0` by auto + ultimately show "real \real a / real b\ < real a / real b" by arith + qed + thus ?thesis using `\ b dvd a` by simp +qed (simp add: ceiling_def real_of_int_minus[symmetric] divide_minus_left[symmetric] + floor_divide_eq_div dvd_neg_div del: divide_minus_left real_of_int_minus) + +lemma compute_float_up[code]: + "float_up p (Float m e) = + (let P = 2^nat (-(p + e)); r = m mod P in + if p + e < 0 then Float (m div P + (if r = 0 then 0 else 1)) (-p) else Float m e)" +proof cases + assume "p + e < 0" + hence "real ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))" + using powr_realpow[of 2 "nat (-(p + e))"] by simp + also have "... = 1 / 2 powr p / 2 powr e" + unfolding powr_minus_divide real_of_int_minus by (simp add: powr_add) + finally have twopow_rewrite: + "real ((2::int) ^ nat (- (p + e))) = 1 / 2 powr real p / 2 powr real e" . + with `p + e < 0` have powr_rewrite: + "2 powr real e * 2 powr real p = 1 / real ((2::int) ^ nat (- (p + e)))" + unfolding powr_divide2 by simp + show ?thesis + proof cases + assume "2^nat (-(p + e)) dvd m" + with `p + e < 0` twopow_rewrite show ?thesis unfolding Let_def + by transfer (auto simp: ac_simps round_up_def floor_divide_eq_div dvd_eq_mod_eq_0) + next + assume ndvd: "\ 2 ^ nat (- (p + e)) dvd m" + have one_div: "real m * (1 / real ((2::int) ^ nat (- (p + e)))) = + real m / real ((2::int) ^ nat (- (p + e)))" + by (simp add: field_simps) + have "real \real m * (2 powr real e * 2 powr real p)\ = + real \real m * (2 powr real e * 2 powr real p)\ + 1" + using ndvd unfolding powr_rewrite one_div + by (subst ceil_divide_floor_conv) (auto simp: field_simps) + thus ?thesis using `p + e < 0` twopow_rewrite + unfolding Let_def + by transfer (auto simp: ac_simps round_up_def floor_divide_eq_div[symmetric]) + qed +next + assume "\ p + e < 0" + then have r1: "real e + real p = real (nat (e + p))" by simp + have r: "\(m * 2 powr e) * 2 powr real p\ = (m * 2 powr e) * 2 powr real p" + by (auto simp add: ac_simps powr_add[symmetric] r1 powr_realpow + intro: exI[where x="m*2^nat (e+p)"]) + then show ?thesis using `\ p + e < 0` + unfolding Let_def + by transfer + (simp add: round_up_def floor_divide_eq_div field_simps powr_add powr_minus inverse_eq_divide) +qed + +lemmas real_of_ints = + real_of_int_zero + real_of_one + real_of_int_add + real_of_int_minus + real_of_int_diff + real_of_int_mult + real_of_int_power + real_numeral +lemmas real_of_nats = + real_of_nat_zero + real_of_nat_one + real_of_nat_1 + real_of_nat_add + real_of_nat_mult + real_of_nat_power + +lemmas int_of_reals = real_of_ints[symmetric] +lemmas nat_of_reals = real_of_nats[symmetric] + +lemma two_real_int: "(2::real) = real (2::int)" by simp +lemma two_real_nat: "(2::real) = real (2::nat)" by simp + +lemma mult_cong: "a = c ==> b = d ==> a*b = c*d" by simp + +subsection {* Compute bitlen of integers *} + +definition bitlen :: "int \ int" where + "bitlen a = (if a > 0 then \log 2 a\ + 1 else 0)" + +lemma bitlen_nonneg: "0 \ bitlen x" proof - - have "e < 0" using float_pos_less1_e_neg assms by auto - have "\x. (0::real) < 2^x" by auto - have "real m < 2^(nat (-e))" using `Float m e < 1` - unfolding less_float_def real_of_float_neg_exp[OF `e < 0`] real_of_float_1 - real_mult_less_iff1[of _ _ 1, OF `0 < 2^(nat (-e))`, symmetric] - mult_assoc by auto - thus ?thesis unfolding real_of_int_less_iff[symmetric] by auto + { + assume "0 > x" + have "-1 = log 2 (inverse 2)" by (subst log_inverse) simp_all + also have "... < log 2 (-x)" using `0 > x` by auto + finally have "-1 < log 2 (-x)" . + } thus "0 \ bitlen x" unfolding bitlen_def by (auto intro!: add_nonneg_nonneg) qed -function bitlen :: "int \ int" where -"bitlen 0 = 0" | -"bitlen -1 = 1" | -"0 < x \ bitlen x = 1 + (bitlen (x div 2))" | -"x < -1 \ bitlen x = 1 + (bitlen (x div 2))" - apply (case_tac "x = 0 \ x = -1 \ x < -1 \ x > 0") - apply auto - done -termination by (relation "measure (nat o abs)", auto) +lemma bitlen_bounds: + assumes "x > 0" + shows "2 ^ nat (bitlen x - 1) \ x \ x < 2 ^ nat (bitlen x)" +proof + have "(2::real) ^ nat \log 2 (real x)\ = 2 powr real (floor (log 2 (real x)))" + using powr_realpow[symmetric, of 2 "nat \log 2 (real x)\"] `x > 0` + using real_nat_eq_real[of "floor (log 2 (real x))"] + by simp + also have "... \ 2 powr log 2 (real x)" + by simp + also have "... = real x" + using `0 < x` by simp + finally have "2 ^ nat \log 2 (real x)\ \ real x" by simp + thus "2 ^ nat (bitlen x - 1) \ x" using `x > 0` + by (simp add: bitlen_def) +next + have "x \ 2 powr (log 2 x)" using `x > 0` by simp + also have "... < 2 ^ nat (\log 2 (real x)\ + 1)" + apply (simp add: powr_realpow[symmetric]) + using `x > 0` by simp + finally show "x < 2 ^ nat (bitlen x)" using `x > 0` + by (simp add: bitlen_def ac_simps int_of_reals del: real_of_ints) +qed -lemma bitlen_ge0: "0 \ bitlen x" by (induct x rule: bitlen.induct, auto) -lemma bitlen_ge1: "x \ 0 \ 1 \ bitlen x" by (induct x rule: bitlen.induct, auto simp add: bitlen_ge0) +lemma bitlen_pow2[simp]: + assumes "b > 0" + shows "bitlen (b * 2 ^ c) = bitlen b + c" +proof - + from assms have "b * 2 ^ c > 0" by (auto intro: mult_pos_pos) + thus ?thesis + using floor_add[of "log 2 b" c] assms + by (auto simp add: log_mult log_nat_power bitlen_def) +qed + +lemma bitlen_Float: +fixes m e +defines "f \ Float m e" +shows "bitlen (\mantissa f\) + exponent f = (if m = 0 then 0 else bitlen \m\ + e)" +proof cases + assume "m \ 0" + hence "f \ float_of 0" + unfolding real_of_float_eq by (simp add: f_def) + hence "mantissa f \ 0" + by (simp add: mantissa_noteq_0) + moreover + from f_def[THEN denormalize_shift, OF `f \ float_of 0`] guess i . + ultimately show ?thesis by (simp add: abs_mult) +qed (simp add: f_def bitlen_def Float_def) -lemma bitlen_bounds': assumes "0 < x" shows "2^nat (bitlen x - 1) \ x \ x + 1 \ 2^nat (bitlen x)" (is "?P x") - using `0 < x` -proof (induct x rule: bitlen.induct) - fix x - assume "0 < x" and hyp: "0 < x div 2 \ ?P (x div 2)" hence "0 \ x" and "x \ 0" by auto - { fix x have "0 \ 1 + bitlen x" using bitlen_ge0[of x] by auto } note gt0_pls1 = this +lemma compute_bitlen[code]: + shows "bitlen x = (if x > 0 then bitlen (x div 2) + 1 else 0)" +proof - + { assume "2 \ x" + then have "\log 2 (x div 2)\ + 1 = \log 2 (x - x mod 2)\" + by (simp add: log_mult zmod_zdiv_equality') + also have "\ = \log 2 (real x)\" + proof cases + assume "x mod 2 = 0" then show ?thesis by simp + next + def n \ "\log 2 (real x)\" + then have "0 \ n" + using `2 \ x` by simp + assume "x mod 2 \ 0" + with `2 \ x` have "x mod 2 = 1" "\ 2 dvd x" by (auto simp add: dvd_eq_mod_eq_0) + with `2 \ x` have "x \ 2^nat n" by (cases "nat n") auto + moreover + { have "real (2^nat n :: int) = 2 powr (nat n)" + by (simp add: powr_realpow) + also have "\ \ 2 powr (log 2 x)" + using `2 \ x` by (simp add: n_def del: powr_log_cancel) + finally have "2^nat n \ x" using `2 \ x` by simp } + ultimately have "2^nat n \ x - 1" by simp + then have "2^nat n \ real (x - 1)" + unfolding real_of_int_le_iff[symmetric] by simp + { have "n = \log 2 (2^nat n)\" + using `0 \ n` by (simp add: log_nat_power) + also have "\ \ \log 2 (x - 1)\" + using `2^nat n \ real (x - 1)` `0 \ n` `2 \ x` by (auto intro: floor_mono) + finally have "n \ \log 2 (x - 1)\" . } + moreover have "\log 2 (x - 1)\ \ n" + using `2 \ x` by (auto simp add: n_def intro!: floor_mono) + ultimately show "\log 2 (x - x mod 2)\ = \log 2 x\" + unfolding n_def `x mod 2 = 1` by auto + qed + finally have "\log 2 (x div 2)\ + 1 = \log 2 x\" . } + moreover + { assume "x < 2" "0 < x" + then have "x = 1" by simp + then have "\log 2 (real x)\ = 0" by simp } + ultimately show ?thesis + unfolding bitlen_def + by (auto simp: pos_imp_zdiv_pos_iff not_le) +qed - have "0 < (2::int)" by auto - - show "?P x" - proof (cases "x = 1") - case True show "?P x" unfolding True by auto +lemma float_gt1_scale: assumes "1 \ Float m e" + shows "0 \ e + (bitlen m - 1)" +proof - + have "0 < Float m e" using assms by auto + hence "0 < m" using powr_gt_zero[of 2 e] + by (auto simp: zero_less_mult_iff) + hence "m \ 0" by auto + show ?thesis + proof (cases "0 \ e") + case True thus ?thesis using `0 < m` by (simp add: bitlen_def) next - case False hence "2 \ x" using `0 < x` `x \ 1` by auto - hence "2 div 2 \ x div 2" by (rule zdiv_mono1, auto) - hence "0 < x div 2" and "x div 2 \ 0" by auto - hence bitlen_s1_ge0: "0 \ bitlen (x div 2) - 1" using bitlen_ge1[OF `x div 2 \ 0`] by auto - - { from hyp[OF `0 < x div 2`] - have "2 ^ nat (bitlen (x div 2) - 1) \ x div 2" by auto - hence "2 ^ nat (bitlen (x div 2) - 1) * 2 \ x div 2 * 2" by (rule mult_right_mono, auto) - also have "\ \ x" using `0 < x` by auto - finally have "2^nat (1 + bitlen (x div 2) - 1) \ x" unfolding power_Suc2[symmetric] Suc_nat_eq_nat_zadd1[OF bitlen_s1_ge0] by auto - } moreover - { have "x + 1 \ x - x mod 2 + 2" - proof - - have "x mod 2 < 2" using `0 < x` by auto - hence "x < x - x mod 2 + 2" unfolding algebra_simps by auto - thus ?thesis by auto - qed - also have "x - x mod 2 + 2 = (x div 2 + 1) * 2" unfolding algebra_simps using `0 < x` div_mod_equality[of x 2 0] by auto - also have "\ \ 2^nat (bitlen (x div 2)) * 2" using hyp[OF `0 < x div 2`, THEN conjunct2] by (rule mult_right_mono, auto) - also have "\ = 2^(1 + nat (bitlen (x div 2)))" unfolding power_Suc2[symmetric] by auto - finally have "x + 1 \ 2^(1 + nat (bitlen (x div 2)))" . - } - ultimately show ?thesis - unfolding bitlen.simps(3)[OF `0 < x`] nat_add_distrib[OF zero_le_one bitlen_ge0] - unfolding add_commute nat_add_distrib[OF zero_le_one gt0_pls1] - by auto + have "(1::int) < 2" by simp + case False let ?S = "2^(nat (-e))" + have "inverse (2 ^ nat (- e)) = 2 powr e" using assms False powr_realpow[of 2 "nat (-e)"] + by (auto simp: powr_minus field_simps inverse_eq_divide) + hence "1 \ real m * inverse ?S" using assms False powr_realpow[of 2 "nat (-e)"] + by (auto simp: powr_minus) + hence "1 * ?S \ real m * inverse ?S * ?S" by (rule mult_right_mono, auto) + hence "?S \ real m" unfolding mult_assoc by auto + hence "?S \ m" unfolding real_of_int_le_iff[symmetric] by auto + from this bitlen_bounds[OF `0 < m`, THEN conjunct2] + have "nat (-e) < (nat (bitlen m))" unfolding power_strict_increasing_iff[OF `1 < 2`, symmetric] by (rule order_le_less_trans) + hence "-e < bitlen m" using False by auto + thus ?thesis by auto qed -next - fix x :: int assume "x < -1" and "0 < x" hence False by auto - thus "?P x" by auto -qed auto - -lemma bitlen_bounds: assumes "0 < x" shows "2^nat (bitlen x - 1) \ x \ x < 2^nat (bitlen x)" - using bitlen_bounds'[OF `0 real m / 2^nat (bitlen m - 1)" and "real m / 2^nat (bitlen m - 1) < 2" proof - @@ -514,840 +811,454 @@ thus "1 \ real m / ?B" by auto have "m \ 0" using assms by auto - have "0 \ bitlen m - 1" using bitlen_ge1[OF `m \ 0`] by auto + have "0 \ bitlen m - 1" using `0 < m` by (auto simp: bitlen_def) have "m < 2^nat(bitlen m)" using bitlen_bounds[OF `0 = 2^nat(bitlen m - 1 + 1)" using bitlen_ge1[OF `m \ 0`] by auto + also have "\ = 2^nat(bitlen m - 1 + 1)" using `0 < m` by (auto simp: bitlen_def) also have "\ = ?B * 2" unfolding nat_add_distrib[OF `0 \ bitlen m - 1` zero_le_one] by auto finally have "real m < 2 * ?B" unfolding real_of_int_less_iff[symmetric] by auto hence "real m / ?B < 2 * ?B / ?B" by (rule divide_strict_right_mono, auto) thus "real m / ?B < 2" by auto qed -lemma float_gt1_scale: assumes "1 \ Float m e" - shows "0 \ e + (bitlen m - 1)" -proof (cases "0 \ e") - have "0 < Float m e" using assms unfolding less_float_def le_float_def by auto - hence "0 < m" using float_pos_m_pos by auto - hence "m \ 0" by auto - case True with bitlen_ge1[OF `m \ 0`] show ?thesis by auto -next - have "0 < Float m e" using assms unfolding less_float_def le_float_def by auto - hence "0 < m" using float_pos_m_pos by auto - hence "m \ 0" and "1 < (2::int)" by auto - case False let ?S = "2^(nat (-e))" - have "1 \ real m * inverse ?S" using assms unfolding le_float_def real_of_float_nge0_exp[OF False] by auto - hence "1 * ?S \ real m * inverse ?S * ?S" by (rule mult_right_mono, auto) - hence "?S \ real m" unfolding mult_assoc by auto - hence "?S \ m" unfolding real_of_int_le_iff[symmetric] by auto - from this bitlen_bounds[OF `0 < m`, THEN conjunct2] - have "nat (-e) < (nat (bitlen m))" unfolding power_strict_increasing_iff[OF `1 < 2`, symmetric] by (rule order_le_less_trans) - hence "-e < bitlen m" using False bitlen_ge0 by auto - thus ?thesis by auto -qed +subsection {* Approximation of positive rationals *} + +lemma zdiv_zmult_twopow_eq: fixes a b::int shows "a div b div (2 ^ n) = a div (b * 2 ^ n)" +by (simp add: zdiv_zmult2_eq) -lemma normalized_float: assumes "m \ 0" shows "real (Float m (- (bitlen m - 1))) = real m / 2^nat (bitlen m - 1)" -proof (cases "- (bitlen m - 1) = 0") - case True show ?thesis unfolding real_of_float_simp pow2_def using True by auto -next - case False hence P: "\ 0 \ - (bitlen m - 1)" using bitlen_ge1[OF `m \ 0`] by auto - show ?thesis unfolding real_of_float_nge0_exp[OF P] divide_inverse by auto -qed +lemma div_mult_twopow_eq: fixes a b::nat shows "a div ((2::nat) ^ n) div b = a div (b * 2 ^ n)" + by (cases "b=0") (simp_all add: div_mult2_eq[symmetric] ac_simps) -(* BROKEN -lemma bitlen_Pls: "bitlen (Int.Pls) = Int.Pls" by (subst Pls_def, subst Pls_def, simp) +lemma real_div_nat_eq_floor_of_divide: + fixes a b::nat + shows "a div b = real (floor (a/b))" +by (metis floor_divide_eq_div real_of_int_of_nat_eq zdiv_int) -lemma bitlen_Min: "bitlen (Int.Min) = Int.Bit1 Int.Pls" by (subst Min_def, simp add: Bit1_def) +definition "rat_precision prec x y = int prec - (bitlen x - bitlen y)" -lemma bitlen_B0: "bitlen (Int.Bit0 b) = (if iszero b then Int.Pls else Int.succ (bitlen b))" - apply (auto simp add: iszero_def succ_def) - apply (simp add: Bit0_def Pls_def) - apply (subst Bit0_def) - apply simp - apply (subgoal_tac "0 < 2 * b \ 2 * b < -1") - apply auto - done +lift_definition lapprox_posrat :: "nat \ nat \ nat \ float" + is "\prec (x::nat) (y::nat). round_down (rat_precision prec x y) (x / y)" by simp -lemma bitlen_B1: "bitlen (Int.Bit1 b) = (if iszero (Int.succ b) then Int.Bit1 Int.Pls else Int.succ (bitlen b))" -proof - - have h: "! x. (2*x + 1) div 2 = (x::int)" - by arith - show ?thesis - apply (auto simp add: iszero_def succ_def) - apply (subst Bit1_def)+ - apply simp - apply (subgoal_tac "2 * b + 1 = -1") - apply (simp only:) - apply simp_all - apply (subst Bit1_def) - apply simp - apply (subgoal_tac "0 < 2 * b + 1 \ 2 * b + 1 < -1") - apply (auto simp add: h) - done -qed +lemma compute_lapprox_posrat[code]: + fixes prec x y + shows "lapprox_posrat prec x y = + (let + l = rat_precision prec x y; + d = if 0 \ l then x * 2^nat l div y else x div 2^nat (- l) div y + in normfloat (Float d (- l)))" + unfolding div_mult_twopow_eq Let_def normfloat_def + by transfer + (simp add: round_down_def powr_int real_div_nat_eq_floor_of_divide field_simps + del: two_powr_minus_int_float) -lemma bitlen_number_of: "bitlen (number_of w) = number_of (bitlen w)" - by (simp add: number_of_is_id) -BH *) - -lemma [code]: "bitlen x = - (if x = 0 then 0 - else if x = -1 then 1 - else (1 + (bitlen (x div 2))))" - by (cases "x = 0 \ x = -1 \ 0 < x") auto - -definition lapprox_posrat :: "nat \ int \ int \ float" -where - "lapprox_posrat prec x y = - (let - l = nat (int prec + bitlen y - bitlen x) ; - d = (x * 2^l) div y - in normfloat (Float d (- (int l))))" - -lemma pow2_minus: "pow2 (-x) = inverse (pow2 x)" - unfolding pow2_neg[of "-x"] by auto +lift_definition rapprox_posrat :: "nat \ nat \ nat \ float" + is "\prec (x::nat) (y::nat). round_up (rat_precision prec x y) (x / y)" by simp -lemma lapprox_posrat: - assumes x: "0 \ x" - and y: "0 < y" - shows "real (lapprox_posrat prec x y) \ real x / real y" -proof - - let ?l = "nat (int prec + bitlen y - bitlen x)" - - have "real (x * 2^?l div y) * inverse (2^?l) \ (real (x * 2^?l) / real y) * inverse (2^?l)" - by (rule mult_right_mono, fact real_of_int_div4, simp) - also have "\ \ (real x / real y) * 2^?l * inverse (2^?l)" by auto - finally have "real (x * 2^?l div y) * inverse (2^?l) \ real x / real y" unfolding mult_assoc by auto - thus ?thesis unfolding lapprox_posrat_def Let_def normfloat real_of_float_simp - unfolding pow2_minus pow2_int minus_minus . -qed - -lemma real_of_int_div_mult: - fixes x y c :: int assumes "0 < y" and "0 < c" - shows "real (x div y) \ real (x * c div y) * inverse (real c)" -proof - - have "c * (x div y) + 0 \ c * x div y" unfolding zdiv_zmult1_eq[of c x y] - by (rule add_left_mono, - auto intro!: mult_nonneg_nonneg - simp add: pos_imp_zdiv_nonneg_iff[OF `0 < y`] `0 < c`[THEN less_imp_le] pos_mod_sign[OF `0 < y`]) - hence "real (x div y) * real c \ real (x * c div y)" - unfolding real_of_int_mult[symmetric] real_of_int_le_iff mult_commute by auto - hence "real (x div y) * real c * inverse (real c) \ real (x * c div y) * inverse (real c)" - using `0 < c` by auto - thus ?thesis unfolding mult_assoc using `0 < c` by auto -qed - -lemma lapprox_posrat_bottom: assumes "0 < y" - shows "real (x div y) \ real (lapprox_posrat n x y)" -proof - - have pow: "\x. (0::int) < 2^x" by auto - show ?thesis - unfolding lapprox_posrat_def Let_def real_of_float_add normfloat real_of_float_simp pow2_minus pow2_int - using real_of_int_div_mult[OF `0 < y` pow] by auto -qed - -lemma lapprox_posrat_nonneg: assumes "0 \ x" and "0 < y" - shows "0 \ real (lapprox_posrat n x y)" -proof - +(* TODO: optimize using zmod_zmult2_eq, pdivmod ? *) +lemma compute_rapprox_posrat[code]: + fixes prec x y + defines "l \ rat_precision prec x y" + shows "rapprox_posrat prec x y = (let + l = l ; + X = if 0 \ l then (x * 2^nat l, y) else (x, y * 2^nat(-l)) ; + d = fst X div snd X ; + m = fst X mod snd X + in normfloat (Float (d + (if m = 0 \ y = 0 then 0 else 1)) (- l)))" +proof (cases "y = 0") + assume "y = 0" thus ?thesis unfolding Let_def normfloat_def by transfer simp +next + assume "y \ 0" show ?thesis - unfolding lapprox_posrat_def Let_def real_of_float_add normfloat real_of_float_simp pow2_minus pow2_int - using pos_imp_zdiv_nonneg_iff[OF `0 < y`] assms by (auto intro!: mult_nonneg_nonneg) -qed - -definition rapprox_posrat :: "nat \ int \ int \ float" -where - "rapprox_posrat prec x y = (let - l = nat (int prec + bitlen y - bitlen x) ; - X = x * 2^l ; - d = X div y ; - m = X mod y - in normfloat (Float (d + (if m = 0 then 0 else 1)) (- (int l))))" - -lemma rapprox_posrat: - assumes x: "0 \ x" - and y: "0 < y" - shows "real x / real y \ real (rapprox_posrat prec x y)" -proof - - let ?l = "nat (int prec + bitlen y - bitlen x)" let ?X = "x * 2^?l" - show ?thesis - proof (cases "?X mod y = 0") - case True hence "y dvd ?X" using `0 < y` by auto - from real_of_int_div[OF this] - have "real (?X div y) * inverse (2 ^ ?l) = real ?X / real y * inverse (2 ^ ?l)" by auto - also have "\ = real x / real y * (2^?l * inverse (2^?l))" by auto - finally have "real (?X div y) * inverse (2^?l) = real x / real y" by auto - thus ?thesis unfolding rapprox_posrat_def Let_def normfloat if_P[OF True] - unfolding real_of_float_simp pow2_minus pow2_int minus_minus by auto - next - case False - have "0 \ real y" and "real y \ 0" using `0 < y` by auto - have "0 \ real y * 2^?l" by (rule mult_nonneg_nonneg, rule `0 \ real y`, auto) - - have "?X = y * (?X div y) + ?X mod y" by auto - also have "\ \ y * (?X div y) + y" by (rule add_mono, auto simp add: pos_mod_bound[OF `0 < y`, THEN less_imp_le]) - also have "\ = y * (?X div y + 1)" unfolding right_distrib by auto - finally have "real ?X \ real y * real (?X div y + 1)" unfolding real_of_int_le_iff real_of_int_mult[symmetric] . - hence "real ?X / (real y * 2^?l) \ real y * real (?X div y + 1) / (real y * 2^?l)" - by (rule divide_right_mono, simp only: `0 \ real y * 2^?l`) - also have "\ = real y * real (?X div y + 1) / real y / 2^?l" by auto - also have "\ = real (?X div y + 1) * inverse (2^?l)" unfolding nonzero_mult_divide_cancel_left[OF `real y \ 0`] - unfolding divide_inverse .. - finally show ?thesis unfolding rapprox_posrat_def Let_def normfloat real_of_float_simp if_not_P[OF False] - unfolding pow2_minus pow2_int minus_minus by auto + proof (cases "0 \ l") + assume "0 \ l" + def x' == "x * 2 ^ nat l" + have "int x * 2 ^ nat l = x'" by (simp add: x'_def int_mult int_power) + moreover have "real x * 2 powr real l = real x'" + by (simp add: powr_realpow[symmetric] `0 \ l` x'_def) + ultimately show ?thesis + unfolding Let_def normfloat_def + using ceil_divide_floor_conv[of y x'] powr_realpow[of 2 "nat l"] `0 \ l` `y \ 0` + l_def[symmetric, THEN meta_eq_to_obj_eq] + by transfer + (simp add: floor_divide_eq_div[symmetric] dvd_eq_mod_eq_0 round_up_def) + next + assume "\ 0 \ l" + def y' == "y * 2 ^ nat (- l)" + from `y \ 0` have "y' \ 0" by (simp add: y'_def) + have "int y * 2 ^ nat (- l) = y'" by (simp add: y'_def int_mult int_power) + moreover have "real x * real (2::int) powr real l / real y = x / real y'" + using `\ 0 \ l` + by (simp add: powr_realpow[symmetric] powr_minus y'_def field_simps inverse_eq_divide) + ultimately show ?thesis + unfolding Let_def normfloat_def + using ceil_divide_floor_conv[of y' x] `\ 0 \ l` `y' \ 0` `y \ 0` + l_def[symmetric, THEN meta_eq_to_obj_eq] + by transfer + (simp add: round_up_def ceil_divide_floor_conv floor_divide_eq_div[symmetric] dvd_eq_mod_eq_0) qed qed -lemma rapprox_posrat_le1: assumes "0 \ x" and "0 < y" and "x \ y" - shows "real (rapprox_posrat n x y) \ 1" +lemma rat_precision_pos: + assumes "0 \ x" and "0 < y" and "2 * x < y" and "0 < n" + shows "rat_precision n (int x) (int y) > 0" proof - - let ?l = "nat (int n + bitlen y - bitlen x)" let ?X = "x * 2^?l" - show ?thesis - proof (cases "?X mod y = 0") - case True hence "y dvd ?X" using `0 < y` by auto - from real_of_int_div[OF this] - have "real (?X div y) * inverse (2 ^ ?l) = real ?X / real y * inverse (2 ^ ?l)" by auto - also have "\ = real x / real y * (2^?l * inverse (2^?l))" by auto - finally have "real (?X div y) * inverse (2^?l) = real x / real y" by auto - also have "real x / real y \ 1" using `0 \ x` and `0 < y` and `x \ y` by auto - finally show ?thesis unfolding rapprox_posrat_def Let_def normfloat if_P[OF True] - unfolding real_of_float_simp pow2_minus pow2_int minus_minus by auto - next - case False - have "x \ y" - proof (rule ccontr) - assume "\ x \ y" hence "x = y" by auto - have "?X mod y = 0" unfolding `x = y` using mod_mult_self1_is_0 by auto - thus False using False by auto - qed - hence "x < y" using `x \ y` by auto - hence "real x / real y < 1" using `0 < y` and `0 \ x` by auto - - from real_of_int_div4[of "?X" y] - have "real (?X div y) \ (real x / real y) * 2^?l" unfolding real_of_int_mult times_divide_eq_left real_of_int_power real_numeral . - also have "\ < 1 * 2^?l" using `real x / real y < 1` by (rule mult_strict_right_mono, auto) - finally have "?X div y < 2^?l" unfolding real_of_int_less_iff[of _ "2^?l", symmetric] by auto - hence "?X div y + 1 \ 2^?l" by auto - hence "real (?X div y + 1) * inverse (2^?l) \ 2^?l * inverse (2^?l)" - unfolding real_of_int_le_iff[of _ "2^?l", symmetric] real_of_int_power real_numeral - by (rule mult_right_mono, auto) - hence "real (?X div y + 1) * inverse (2^?l) \ 1" by auto - thus ?thesis unfolding rapprox_posrat_def Let_def normfloat real_of_float_simp if_not_P[OF False] - unfolding pow2_minus pow2_int minus_minus by auto - qed -qed - -lemma zdiv_greater_zero: fixes a b :: int assumes "0 < a" and "a \ b" - shows "0 < b div a" -proof (rule ccontr) - have "0 \ b" using assms by auto - assume "\ 0 < b div a" hence "b div a = 0" using `0 \ b`[unfolded pos_imp_zdiv_nonneg_iff[OF `0 b` by auto + { assume "0 < x" hence "log 2 x + 1 = log 2 (2 * x)" by (simp add: log_mult) } + hence "bitlen (int x) < bitlen (int y)" using assms + by (simp add: bitlen_def del: floor_add_one) + (auto intro!: floor_mono simp add: floor_add_one[symmetric] simp del: floor_add floor_add_one) + thus ?thesis + using assms by (auto intro!: pos_add_strict simp add: field_simps rat_precision_def) qed -lemma rapprox_posrat_less1: assumes "0 \ x" and "0 < y" and "2 * x < y" and "0 < n" - shows "real (rapprox_posrat n x y) < 1" -proof (cases "x = 0") - case True thus ?thesis unfolding rapprox_posrat_def True Let_def normfloat real_of_float_simp by auto -next - case False hence "0 < x" using `0 \ x` by auto - hence "x < y" using assms by auto - - let ?l = "nat (int n + bitlen y - bitlen x)" let ?X = "x * 2^?l" - show ?thesis - proof (cases "?X mod y = 0") - case True hence "y dvd ?X" using `0 < y` by auto - from real_of_int_div[OF this] - have "real (?X div y) * inverse (2 ^ ?l) = real ?X / real y * inverse (2 ^ ?l)" by auto - also have "\ = real x / real y * (2^?l * inverse (2^?l))" by auto - finally have "real (?X div y) * inverse (2^?l) = real x / real y" by auto - also have "real x / real y < 1" using `0 \ x` and `0 < y` and `x < y` by auto - finally show ?thesis unfolding rapprox_posrat_def Let_def normfloat real_of_float_simp if_P[OF True] - unfolding pow2_minus pow2_int minus_minus by auto - next - case False - hence "(real x / real y) < 1 / 2" using `0 < y` and `0 \ x` `2 * x < y` by auto - - have "0 < ?X div y" - proof - - have "2^nat (bitlen x - 1) \ y" and "y < 2^nat (bitlen y)" - using bitlen_bounds[OF `0 < x`, THEN conjunct1] bitlen_bounds[OF `0 < y`, THEN conjunct2] `x < y` by auto - hence "(2::int)^nat (bitlen x - 1) < 2^nat (bitlen y)" by (rule order_le_less_trans) - hence "bitlen x \ bitlen y" by auto - hence len_less: "nat (bitlen x - 1) \ nat (int (n - 1) + bitlen y)" by auto - - have "x \ 0" and "y \ 0" using `0 < x` `0 < y` by auto - - have exp_eq: "nat (int (n - 1) + bitlen y) - nat (bitlen x - 1) = ?l" - using `bitlen x \ bitlen y` bitlen_ge1[OF `x \ 0`] bitlen_ge1[OF `y \ 0`] `0 < n` by auto - - have "y * 2^nat (bitlen x - 1) \ y * x" - using bitlen_bounds[OF `0 < x`, THEN conjunct1] `0 < y`[THEN less_imp_le] by (rule mult_left_mono) - also have "\ \ 2^nat (bitlen y) * x" using bitlen_bounds[OF `0 < y`, THEN conjunct2, THEN less_imp_le] `0 \ x` by (rule mult_right_mono) - also have "\ \ x * 2^nat (int (n - 1) + bitlen y)" unfolding mult_commute[of x] by (rule mult_right_mono, auto simp add: `0 \ x`) - finally have "real y * 2^nat (bitlen x - 1) * inverse (2^nat (bitlen x - 1)) \ real x * 2^nat (int (n - 1) + bitlen y) * inverse (2^nat (bitlen x - 1))" - unfolding real_of_int_le_iff[symmetric] by auto - hence "real y \ real x * (2^nat (int (n - 1) + bitlen y) / (2^nat (bitlen x - 1)))" - unfolding mult_assoc divide_inverse by auto - also have "\ = real x * (2^(nat (int (n - 1) + bitlen y) - nat (bitlen x - 1)))" using power_diff[of "2::real", OF _ len_less] by auto - finally have "y \ x * 2^?l" unfolding exp_eq unfolding real_of_int_le_iff[symmetric] by auto - thus ?thesis using zdiv_greater_zero[OF `0 < y`] by auto - qed - - from real_of_int_div4[of "?X" y] - have "real (?X div y) \ (real x / real y) * 2^?l" unfolding real_of_int_mult times_divide_eq_left real_of_int_power real_numeral . - also have "\ < 1/2 * 2^?l" using `real x / real y < 1/2` by (rule mult_strict_right_mono, auto) - finally have "?X div y * 2 < 2^?l" unfolding real_of_int_less_iff[of _ "2^?l", symmetric] by auto - hence "?X div y + 1 < 2^?l" using `0 < ?X div y` by auto - hence "real (?X div y + 1) * inverse (2^?l) < 2^?l * inverse (2^?l)" - unfolding real_of_int_less_iff[of _ "2^?l", symmetric] real_of_int_power real_numeral - by (rule mult_strict_right_mono, auto) - hence "real (?X div y + 1) * inverse (2^?l) < 1" by auto - thus ?thesis unfolding rapprox_posrat_def Let_def normfloat real_of_float_simp if_not_P[OF False] - unfolding pow2_minus pow2_int minus_minus by auto - qed +lemma power_aux: assumes "x > 0" shows "(2::int) ^ nat (x - 1) \ 2 ^ nat x - 1" +proof - + def y \ "nat (x - 1)" moreover + have "(2::int) ^ y \ (2 ^ (y + 1)) - 1" by simp + ultimately show ?thesis using assms by simp qed -lemma approx_rat_pattern: fixes P and ps :: "nat * int * int" - assumes Y: "\y prec x. \y = 0; ps = (prec, x, 0)\ \ P" - and A: "\x y prec. \0 \ x; 0 < y; ps = (prec, x, y)\ \ P" - and B: "\x y prec. \x < 0; 0 < y; ps = (prec, x, y)\ \ P" - and C: "\x y prec. \x < 0; y < 0; ps = (prec, x, y)\ \ P" - and D: "\x y prec. \0 \ x; y < 0; ps = (prec, x, y)\ \ P" - shows P +lemma rapprox_posrat_less1: + assumes "0 \ x" and "0 < y" and "2 * x < y" and "0 < n" + shows "real (rapprox_posrat n x y) < 1" proof - - obtain prec x y where [simp]: "ps = (prec, x, y)" by (cases ps) auto - from Y have "y = 0 \ P" by auto - moreover { - assume "0 < y" - have P - proof (cases "0 \ x") - case True - with A and `0 < y` show P by auto - next - case False - with B and `0 < y` show P by auto - qed - } - moreover { - assume "y < 0" - have P - proof (cases "0 \ x") - case True - with D and `y < 0` show P by auto - next - case False - with C and `y < 0` show P by auto - qed - } - ultimately show P by (cases "y = 0 \ 0 < y \ y < 0") auto -qed - -function lapprox_rat :: "nat \ int \ int \ float" -where - "y = 0 \ lapprox_rat prec x y = 0" -| "0 \ x \ 0 < y \ lapprox_rat prec x y = lapprox_posrat prec x y" -| "x < 0 \ 0 < y \ lapprox_rat prec x y = - (rapprox_posrat prec (-x) y)" -| "x < 0 \ y < 0 \ lapprox_rat prec x y = lapprox_posrat prec (-x) (-y)" -| "0 \ x \ y < 0 \ lapprox_rat prec x y = - (rapprox_posrat prec x (-y))" -apply simp_all by (rule approx_rat_pattern) -termination by lexicographic_order - -lemma compute_lapprox_rat[code]: - "lapprox_rat prec x y = (if y = 0 then 0 else if 0 \ x then (if 0 < y then lapprox_posrat prec x y else - (rapprox_posrat prec x (-y))) - else (if 0 < y then - (rapprox_posrat prec (-x) y) else lapprox_posrat prec (-x) (-y)))" - by auto - -lemma lapprox_rat: "real (lapprox_rat prec x y) \ real x / real y" -proof - - have h[rule_format]: "! a b b'. b' \ b \ a \ b' \ a \ (b::real)" by auto - show ?thesis - apply (case_tac "y = 0") - apply simp - apply (case_tac "0 \ x \ 0 < y") - apply (simp add: lapprox_posrat) - apply (case_tac "x < 0 \ 0 < y") - apply simp - apply (subst minus_le_iff) - apply (rule h[OF rapprox_posrat]) - apply (simp_all) - apply (case_tac "x < 0 \ y < 0") - apply simp - apply (rule h[OF _ lapprox_posrat]) - apply (simp_all) - apply (case_tac "0 \ x \ y < 0") - apply (simp) - apply (subst minus_le_iff) - apply (rule h[OF rapprox_posrat]) - apply simp_all - apply arith - done -qed - -lemma lapprox_rat_bottom: assumes "0 \ x" and "0 < y" - shows "real (x div y) \ real (lapprox_rat n x y)" - unfolding lapprox_rat.simps(2)[OF assms] using lapprox_posrat_bottom[OF `0 int \ int \ float" -where - "y = 0 \ rapprox_rat prec x y = 0" -| "0 \ x \ 0 < y \ rapprox_rat prec x y = rapprox_posrat prec x y" -| "x < 0 \ 0 < y \ rapprox_rat prec x y = - (lapprox_posrat prec (-x) y)" -| "x < 0 \ y < 0 \ rapprox_rat prec x y = rapprox_posrat prec (-x) (-y)" -| "0 \ x \ y < 0 \ rapprox_rat prec x y = - (lapprox_posrat prec x (-y))" -apply simp_all by (rule approx_rat_pattern) -termination by lexicographic_order - -lemma compute_rapprox_rat[code]: - "rapprox_rat prec x y = (if y = 0 then 0 else if 0 \ x then (if 0 < y then rapprox_posrat prec x y else - (lapprox_posrat prec x (-y))) else - (if 0 < y then - (lapprox_posrat prec (-x) y) else rapprox_posrat prec (-x) (-y)))" - by auto - -lemma rapprox_rat: "real x / real y \ real (rapprox_rat prec x y)" -proof - - have h[rule_format]: "! a b b'. b' \ b \ a \ b' \ a \ (b::real)" by auto - show ?thesis - apply (case_tac "y = 0") - apply simp - apply (case_tac "0 \ x \ 0 < y") - apply (simp add: rapprox_posrat) - apply (case_tac "x < 0 \ 0 < y") - apply simp - apply (subst le_minus_iff) - apply (rule h[OF _ lapprox_posrat]) - apply (simp_all) - apply (case_tac "x < 0 \ y < 0") - apply simp - apply (rule h[OF rapprox_posrat]) - apply (simp_all) - apply (case_tac "0 \ x \ y < 0") - apply (simp) - apply (subst le_minus_iff) - apply (rule h[OF _ lapprox_posrat]) - apply simp_all - apply arith + have powr1: "2 powr real (rat_precision n (int x) (int y)) = + 2 ^ nat (rat_precision n (int x) (int y))" using rat_precision_pos[of x y n] assms + by (simp add: powr_realpow[symmetric]) + have "x * 2 powr real (rat_precision n (int x) (int y)) / y = (x / y) * + 2 powr real (rat_precision n (int x) (int y))" by simp + also have "... < (1 / 2) * 2 powr real (rat_precision n (int x) (int y))" + apply (rule mult_strict_right_mono) by (insert assms) auto + also have "\ = 2 powr real (rat_precision n (int x) (int y) - 1)" + by (simp add: powr_add diff_def powr_neg_numeral) + also have "\ = 2 ^ nat (rat_precision n (int x) (int y) - 1)" + using rat_precision_pos[of x y n] assms by (simp add: powr_realpow[symmetric]) + also have "\ \ 2 ^ nat (rat_precision n (int x) (int y)) - 1" + unfolding int_of_reals real_of_int_le_iff + using rat_precision_pos[OF assms] by (rule power_aux) + finally show ?thesis + apply (transfer fixing: n x y) + apply (simp add: round_up_def field_simps powr_minus inverse_eq_divide powr1) + unfolding int_of_reals real_of_int_less_iff + apply (simp add: ceiling_less_eq) done qed -lemma rapprox_rat_le1: assumes "0 \ x" and "0 < y" and "x \ y" - shows "real (rapprox_rat n x y) \ 1" - unfolding rapprox_rat.simps(2)[OF `0 \ x` `0 < y`] using rapprox_posrat_le1[OF assms] . +lift_definition lapprox_rat :: "nat \ int \ int \ float" is + "\prec (x::int) (y::int). round_down (rat_precision prec \x\ \y\) (x / y)" by simp + +lemma compute_lapprox_rat[code]: + "lapprox_rat prec x y = + (if y = 0 then 0 + else if 0 \ x then + (if 0 < y then lapprox_posrat prec (nat x) (nat y) + else - (rapprox_posrat prec (nat x) (nat (-y)))) + else (if 0 < y + then - (rapprox_posrat prec (nat (-x)) (nat y)) + else lapprox_posrat prec (nat (-x)) (nat (-y))))" + by transfer (auto simp: round_up_def round_down_def ceiling_def ac_simps) + +lift_definition rapprox_rat :: "nat \ int \ int \ float" is + "\prec (x::int) (y::int). round_up (rat_precision prec \x\ \y\) (x / y)" by simp -lemma rapprox_rat_neg: assumes "x < 0" and "0 < y" - shows "real (rapprox_rat n x y) \ 0" - unfolding rapprox_rat.simps(3)[OF assms] using lapprox_posrat_nonneg[of "-x" y n] assms by auto +lemma compute_rapprox_rat[code]: + "rapprox_rat prec x y = + (if y = 0 then 0 + else if 0 \ x then + (if 0 < y then rapprox_posrat prec (nat x) (nat y) + else - (lapprox_posrat prec (nat x) (nat (-y)))) + else (if 0 < y + then - (lapprox_posrat prec (nat (-x)) (nat y)) + else rapprox_posrat prec (nat (-x)) (nat (-y))))" + by transfer (auto simp: round_up_def round_down_def ceiling_def ac_simps) + +subsection {* Division *} + +lift_definition float_divl :: "nat \ float \ float \ float" is + "\(prec::nat) a b. round_down (prec + \ log 2 \b\ \ - \ log 2 \a\ \) (a / b)" by simp -lemma rapprox_rat_nonneg_neg: assumes "0 \ x" and "y < 0" - shows "real (rapprox_rat n x y) \ 0" - unfolding rapprox_rat.simps(5)[OF assms] using lapprox_posrat_nonneg[of x "-y" n] assms by auto +lemma compute_float_divl[code]: + "float_divl prec (Float m1 s1) (Float m2 s2) = lapprox_rat prec m1 m2 * Float 1 (s1 - s2)" +proof cases + let ?f1 = "real m1 * 2 powr real s1" and ?f2 = "real m2 * 2 powr real s2" + let ?m = "real m1 / real m2" and ?s = "2 powr real (s1 - s2)" + assume not_0: "m1 \ 0 \ m2 \ 0" + then have eq2: "(int prec + \log 2 \?f2\\ - \log 2 \?f1\\) = rat_precision prec \m1\ \m2\ + (s2 - s1)" + by (simp add: abs_mult log_mult rat_precision_def bitlen_def) + have eq1: "real m1 * 2 powr real s1 / (real m2 * 2 powr real s2) = ?m * ?s" + by (simp add: field_simps powr_divide2[symmetric]) + + show ?thesis + using not_0 + by (transfer fixing: m1 s1 m2 s2 prec) (unfold eq1 eq2 round_down_shift, simp add: field_simps) +qed (transfer, auto) -lemma rapprox_rat_nonpos_pos: assumes "x \ 0" and "0 < y" - shows "real (rapprox_rat n x y) \ 0" -proof (cases "x = 0") - case True - hence "0 \ x" by auto show ?thesis - unfolding rapprox_rat.simps(2)[OF `0 \ x` `0 < y`] - unfolding True rapprox_posrat_def Let_def - by auto -next - case False - hence "x < 0" using assms by auto - show ?thesis using rapprox_rat_neg[OF `x < 0` `0 < y`] . +lift_definition float_divr :: "nat \ float \ float \ float" is + "\(prec::nat) a b. round_up (prec + \ log 2 \b\ \ - \ log 2 \a\ \) (a / b)" by simp + +lemma compute_float_divr[code]: + "float_divr prec (Float m1 s1) (Float m2 s2) = rapprox_rat prec m1 m2 * Float 1 (s1 - s2)" +proof cases + let ?f1 = "real m1 * 2 powr real s1" and ?f2 = "real m2 * 2 powr real s2" + let ?m = "real m1 / real m2" and ?s = "2 powr real (s1 - s2)" + assume not_0: "m1 \ 0 \ m2 \ 0" + then have eq2: "(int prec + \log 2 \?f2\\ - \log 2 \?f1\\) = rat_precision prec \m1\ \m2\ + (s2 - s1)" + by (simp add: abs_mult log_mult rat_precision_def bitlen_def) + have eq1: "real m1 * 2 powr real s1 / (real m2 * 2 powr real s2) = ?m * ?s" + by (simp add: field_simps powr_divide2[symmetric]) + + show ?thesis + using not_0 + by (transfer fixing: m1 s1 m2 s2 prec) (unfold eq1 eq2 round_up_shift, simp add: field_simps) +qed (transfer, auto) + +subsection {* Lemmas needed by Approximate *} + +lemma Float_num[simp]: shows + "real (Float 1 0) = 1" and "real (Float 1 1) = 2" and "real (Float 1 2) = 4" and + "real (Float 1 -1) = 1/2" and "real (Float 1 -2) = 1/4" and "real (Float 1 -3) = 1/8" and + "real (Float -1 0) = -1" and "real (Float (number_of n) 0) = number_of n" +using two_powr_int_float[of 2] two_powr_int_float[of "-1"] two_powr_int_float[of "-2"] two_powr_int_float[of "-3"] +using powr_realpow[of 2 2] powr_realpow[of 2 3] +using powr_minus[of 2 1] powr_minus[of 2 2] powr_minus[of 2 3] +by auto + +lemma real_of_Float_int[simp]: "real (Float n 0) = real n" by simp + +lemma float_zero[simp]: "real (Float 0 e) = 0" by simp + +lemma abs_div_2_less: "a \ 0 \ a \ -1 \ abs((a::int) div 2) < abs a" +by arith + +lemma lapprox_rat: + shows "real (lapprox_rat prec x y) \ real x / real y" + using round_down by (simp add: lapprox_rat_def) + +lemma mult_div_le: fixes a b:: int assumes "b > 0" shows "a \ b * (a div b)" +proof - + from zmod_zdiv_equality'[of a b] + have "a = b * (a div b) + a mod b" by simp + also have "... \ b * (a div b) + 0" apply (rule add_left_mono) apply (rule pos_mod_sign) + using assms by simp + finally show ?thesis by simp qed -fun float_divl :: "nat \ float \ float \ float" -where - "float_divl prec (Float m1 s1) (Float m2 s2) = - (let - l = lapprox_rat prec m1 m2; - f = Float 1 (s1 - s2) - in - f * l)" +lemma lapprox_rat_nonneg: + fixes n x y + defines "p == int n - ((bitlen \x\) - (bitlen \y\))" + assumes "0 \ x" "0 < y" + shows "0 \ real (lapprox_rat n x y)" +using assms unfolding lapprox_rat_def p_def[symmetric] round_down_def real_of_int_minus[symmetric] + powr_int[of 2, simplified] + by (auto simp add: inverse_eq_divide intro!: mult_nonneg_nonneg divide_nonneg_pos mult_pos_pos) + +lemma rapprox_rat: "real x / real y \ real (rapprox_rat prec x y)" + using round_up by (simp add: rapprox_rat_def) + +lemma rapprox_rat_le1: + fixes n x y + assumes xy: "0 \ x" "0 < y" "x \ y" + shows "real (rapprox_rat n x y) \ 1" +proof - + have "bitlen \x\ \ bitlen \y\" + using xy unfolding bitlen_def by (auto intro!: floor_mono) + then have "0 \ rat_precision n \x\ \y\" by (simp add: rat_precision_def) + have "real \real x / real y * 2 powr real (rat_precision n \x\ \y\)\ + \ real \2 powr real (rat_precision n \x\ \y\)\" + using xy by (auto intro!: ceiling_mono simp: field_simps) + also have "\ = 2 powr real (rat_precision n \x\ \y\)" + using `0 \ rat_precision n \x\ \y\` + by (auto intro!: exI[of _ "2^nat (rat_precision n \x\ \y\)"] simp: powr_int) + finally show ?thesis + by (simp add: rapprox_rat_def round_up_def) + (simp add: powr_minus inverse_eq_divide) +qed + +lemma rapprox_rat_nonneg_neg: + "0 \ x \ y < 0 \ real (rapprox_rat n x y) \ 0" + unfolding rapprox_rat_def round_up_def + by (auto simp: field_simps mult_le_0_iff zero_le_mult_iff) + +lemma rapprox_rat_neg: + "x < 0 \ 0 < y \ real (rapprox_rat n x y) \ 0" + unfolding rapprox_rat_def round_up_def + by (auto simp: field_simps mult_le_0_iff) + +lemma rapprox_rat_nonpos_pos: + "x \ 0 \ 0 < y \ real (rapprox_rat n x y) \ 0" + unfolding rapprox_rat_def round_up_def + by (auto simp: field_simps mult_le_0_iff) lemma float_divl: "real (float_divl prec x y) \ real x / real y" - using lapprox_rat[of prec "mantissa x" "mantissa y"] - by (cases x y rule: float.exhaust[case_product float.exhaust]) - (simp split: split_if_asm - add: real_of_float_simp pow2_diff field_simps le_divide_eq mult_less_0_iff zero_less_mult_iff) + by transfer (simp add: round_down) + +lemma float_divl_lower_bound: + "0 \ x \ 0 < y \ 0 \ real (float_divl prec x y)" + by transfer (simp add: round_down_def zero_le_mult_iff zero_le_divide_iff) + +lemma exponent_1: "exponent 1 = 0" + using exponent_float[of 1 0] by (simp add: one_float_def) + +lemma mantissa_1: "mantissa 1 = 1" + using mantissa_float[of 1 0] by (simp add: one_float_def) + +lemma bitlen_1: "bitlen 1 = 1" + by (simp add: bitlen_def) -lemma float_divl_lower_bound: assumes "0 \ x" and "0 < y" shows "0 \ float_divl prec x y" -proof (cases x, cases y) - fix xm xe ym ye :: int - assume x_eq: "x = Float xm xe" and y_eq: "y = Float ym ye" - have "0 \ xm" - using `0 \ x`[unfolded x_eq le_float_def real_of_float_simp real_of_float_0 zero_le_mult_iff] - by auto - have "0 < ym" - using `0 < y`[unfolded y_eq less_float_def real_of_float_simp real_of_float_0 zero_less_mult_iff] - by auto +lemma mantissa_eq_zero_iff: "mantissa x = 0 \ x = 0" +proof + assume "mantissa x = 0" hence z: "0 = real x" using mantissa_exponent by simp + show "x = 0" by (simp add: zero_float_def z) +qed (simp add: zero_float_def) - have "\n. 0 \ real (Float 1 n)" - unfolding real_of_float_simp using zero_le_pow2 by auto - moreover have "0 \ real (lapprox_rat prec xm ym)" - apply (rule order_trans[OF _ lapprox_rat_bottom[OF `0 \ xm` `0 < ym`]]) - apply (auto simp add: `0 \ xm` pos_imp_zdiv_nonneg_iff[OF `0 < ym`]) - done - ultimately show "0 \ float_divl prec x y" - unfolding x_eq y_eq float_divl.simps Let_def le_float_def real_of_float_0 - by (auto intro!: mult_nonneg_nonneg) +lemma float_upper_bound: "x \ 2 powr (bitlen \mantissa x\ + exponent x)" +proof (cases "x = 0", simp) + assume "x \ 0" hence "mantissa x \ 0" using mantissa_eq_zero_iff by auto + have "x = mantissa x * 2 powr (exponent x)" by (rule mantissa_exponent) + also have "mantissa x \ \mantissa x\" by simp + also have "... \ 2 powr (bitlen \mantissa x\)" + using bitlen_bounds[of "\mantissa x\"] bitlen_nonneg `mantissa x \ 0` + by (simp add: powr_int) (simp only: two_real_int int_of_reals real_of_int_abs[symmetric] + real_of_int_le_iff less_imp_le) + finally show ?thesis by (simp add: powr_add) qed lemma float_divl_pos_less1_bound: - assumes "0 < x" and "x < 1" and "0 < prec" - shows "1 \ float_divl prec 1 x" -proof (cases x) - case (Float m e) - from `0 < x` `x < 1` have "0 < m" "e < 0" - using float_pos_m_pos float_pos_less1_e_neg unfolding Float by auto - let ?b = "nat (bitlen m)" and ?e = "nat (-e)" - have "1 \ m" and "m \ 0" using `0 < m` by auto - with bitlen_bounds[OF `0 < m`] have "m < 2^?b" and "(2::int) \ 2^?b" by auto - hence "1 \ bitlen m" using power_le_imp_le_exp[of "2::int" 1 ?b] by auto - hence pow_split: "nat (int prec + bitlen m - 1) = (prec - 1) + ?b" using `0 < prec` by auto - - have pow_not0: "\x. (2::real)^x \ 0" by auto + "0 < real x \ real x < 1 \ prec \ 1 \ 1 \ real (float_divl prec 1 x)" +proof transfer + fix prec :: nat and x :: real assume x: "0 < x" "x < 1" "x \ float" and prec: "1 \ prec" + def p \ "int prec + \log 2 \x\\" + show "1 \ round_down (int prec + \log 2 \x\\ - \log 2 \1\\) (1 / x) " + proof cases + assume nonneg: "0 \ p" + hence "2 powr real (p) = floor (real ((2::int) ^ nat p)) * floor (1::real)" + by (simp add: powr_int del: real_of_int_power) simp + also have "floor (1::real) \ floor (1 / x)" using x prec by simp + also have "floor (real ((2::int) ^ nat p)) * floor (1 / x) \ + floor (real ((2::int) ^ nat p) * (1 / x))" + by (rule le_mult_floor) (auto simp: x prec less_imp_le) + finally have "2 powr real p \ floor (2 powr nat p / x)" by (simp add: powr_realpow) + thus ?thesis unfolding p_def[symmetric] + using x prec nonneg by (simp add: powr_minus inverse_eq_divide round_down_def) + next + assume neg: "\ 0 \ p" - from float_less1_mantissa_bound `0 < x` `x < 1` Float - have "m < 2^?e" by auto - with bitlen_bounds[OF `0 < m`, THEN conjunct1] have "(2::int)^nat (bitlen m - 1) < 2^?e" - by (rule order_le_less_trans) - from power_less_imp_less_exp[OF _ this] - have "bitlen m \ - e" by auto - hence "(2::real)^?b \ 2^?e" by auto - hence "(2::real)^?b * inverse (2^?b) \ 2^?e * inverse (2^?b)" - by (rule mult_right_mono) auto - hence "(1::real) \ 2^?e * inverse (2^?b)" by auto - also - let ?d = "real (2 ^ nat (int prec + bitlen m - 1) div m) * inverse (2 ^ nat (int prec + bitlen m - 1))" - { - have "2^(prec - 1) * m \ 2^(prec - 1) * 2^?b" - using `m < 2^?b`[THEN less_imp_le] by (rule mult_left_mono) auto - also have "\ = 2 ^ nat (int prec + bitlen m - 1)" - unfolding pow_split power_add by auto - finally have "2^(prec - 1) * m div m \ 2 ^ nat (int prec + bitlen m - 1) div m" - using `0 < m` by (rule zdiv_mono1) - hence "2^(prec - 1) \ 2 ^ nat (int prec + bitlen m - 1) div m" - unfolding div_mult_self2_is_id[OF `m \ 0`] . - hence "2^(prec - 1) * inverse (2 ^ nat (int prec + bitlen m - 1)) \ ?d" - unfolding real_of_int_le_iff[of "2^(prec - 1)", symmetric] by auto - } - from mult_left_mono[OF this [unfolded pow_split power_add inverse_mult_distrib mult_assoc[symmetric] right_inverse[OF pow_not0] mult_1_left], of "2^?e"] - have "2^?e * inverse (2^?b) \ 2^?e * ?d" unfolding pow_split power_add by auto - finally have "1 \ 2^?e * ?d" . - - have e_nat: "0 - e = int (nat (-e))" using `e < 0` by auto - have "bitlen 1 = 1" using bitlen.simps by auto - - show ?thesis - unfolding one_float_def Float float_divl.simps Let_def - lapprox_rat.simps(2)[OF zero_le_one `0 < m`] - lapprox_posrat_def `bitlen 1 = 1` - unfolding le_float_def real_of_float_mult normfloat real_of_float_simp - pow2_minus pow2_int e_nat - using `1 \ 2^?e * ?d` by (auto simp add: pow2_def) + have "x = 2 powr (log 2 x)" + using x by simp + also have "2 powr (log 2 x) \ 2 powr p" + proof (rule powr_mono) + have "log 2 x \ \log 2 x\" + by simp + also have "\ \ \log 2 x\ + 1" + using ceiling_diff_floor_le_1[of "log 2 x"] by simp + also have "\ \ \log 2 x\ + prec" + using prec by simp + finally show "log 2 x \ real p" + using x by (simp add: p_def) + qed simp + finally have x_le: "x \ 2 powr p" . + + from neg have "2 powr real p \ 2 powr 0" + by (intro powr_mono) auto + also have "\ \ \2 powr 0\" by simp + also have "\ \ \2 powr real p / x\" unfolding real_of_int_le_iff + using x x_le by (intro floor_mono) (simp add: pos_le_divide_eq mult_pos_pos) + finally show ?thesis + using prec x unfolding p_def[symmetric] + by (simp add: round_down_def powr_minus_divide pos_le_divide_eq mult_pos_pos) + qed qed -fun float_divr :: "nat \ float \ float \ float" -where - "float_divr prec (Float m1 s1) (Float m2 s2) = - (let - r = rapprox_rat prec m1 m2; - f = Float 1 (s1 - s2) - in - f * r)" - lemma float_divr: "real x / real y \ real (float_divr prec x y)" - using rapprox_rat[of "mantissa x" "mantissa y" prec] - by (cases x y rule: float.exhaust[case_product float.exhaust]) - (simp split: split_if_asm - add: real_of_float_simp pow2_diff field_simps divide_le_eq mult_less_0_iff zero_less_mult_iff) + using round_up by transfer simp lemma float_divr_pos_less1_lower_bound: assumes "0 < x" and "x < 1" shows "1 \ float_divr prec 1 x" proof - - have "1 \ 1 / real x" using `0 < x` and `x < 1` unfolding less_float_def by auto + have "1 \ 1 / real x" using `0 < x` and `x < 1` by auto also have "\ \ real (float_divr prec 1 x)" using float_divr[where x=1 and y=x] by auto - finally show ?thesis unfolding le_float_def by auto -qed - -lemma float_divr_nonpos_pos_upper_bound: assumes "x \ 0" and "0 < y" shows "float_divr prec x y \ 0" -proof (cases x, cases y) - fix xm xe ym ye :: int - assume x_eq: "x = Float xm xe" and y_eq: "y = Float ym ye" - have "xm \ 0" using `x \ 0`[unfolded x_eq le_float_def real_of_float_simp real_of_float_0 mult_le_0_iff] by auto - have "0 < ym" using `0 < y`[unfolded y_eq less_float_def real_of_float_simp real_of_float_0 zero_less_mult_iff] by auto - - have "\n. 0 \ real (Float 1 n)" unfolding real_of_float_simp using zero_le_pow2 by auto - moreover have "real (rapprox_rat prec xm ym) \ 0" using rapprox_rat_nonpos_pos[OF `xm \ 0` `0 < ym`] . - ultimately show "float_divr prec x y \ 0" - unfolding x_eq y_eq float_divr.simps Let_def le_float_def real_of_float_0 real_of_float_mult by (auto intro!: mult_nonneg_nonpos) -qed - -lemma float_divr_nonneg_neg_upper_bound: assumes "0 \ x" and "y < 0" shows "float_divr prec x y \ 0" -proof (cases x, cases y) - fix xm xe ym ye :: int - assume x_eq: "x = Float xm xe" and y_eq: "y = Float ym ye" - have "0 \ xm" using `0 \ x`[unfolded x_eq le_float_def real_of_float_simp real_of_float_0 zero_le_mult_iff] by auto - have "ym < 0" using `y < 0`[unfolded y_eq less_float_def real_of_float_simp real_of_float_0 mult_less_0_iff] by auto - hence "0 < - ym" by auto - - have "\n. 0 \ real (Float 1 n)" unfolding real_of_float_simp using zero_le_pow2 by auto - moreover have "real (rapprox_rat prec xm ym) \ 0" using rapprox_rat_nonneg_neg[OF `0 \ xm` `ym < 0`] . - ultimately show "float_divr prec x y \ 0" - unfolding x_eq y_eq float_divr.simps Let_def le_float_def real_of_float_0 real_of_float_mult by (auto intro!: mult_nonneg_nonpos) -qed - -primrec round_down :: "nat \ float \ float" where -"round_down prec (Float m e) = (let d = bitlen m - int prec in - if 0 < d then let P = 2^nat d ; n = m div P in Float n (e + d) - else Float m e)" - -primrec round_up :: "nat \ float \ float" where -"round_up prec (Float m e) = (let d = bitlen m - int prec in - if 0 < d then let P = 2^nat d ; n = m div P ; r = m mod P in Float (n + (if r = 0 then 0 else 1)) (e + d) - else Float m e)" - -lemma round_up: "real x \ real (round_up prec x)" -proof (cases x) - case (Float m e) - let ?d = "bitlen m - int prec" - let ?p = "(2::int)^nat ?d" - have "0 < ?p" by auto - show "?thesis" - proof (cases "0 < ?d") - case True - hence pow_d: "pow2 ?d = real ?p" using pow2_int[symmetric] by simp - show ?thesis - proof (cases "m mod ?p = 0") - case True - have m: "m = m div ?p * ?p + 0" unfolding True[symmetric] using mod_div_equality [symmetric] . - have "real (Float m e) = real (Float (m div ?p) (e + ?d))" unfolding real_of_float_simp arg_cong[OF m, of real] - by (auto simp add: pow2_add `0 < ?d` pow_d) - thus ?thesis - unfolding Float round_up.simps Let_def if_P[OF `m mod ?p = 0`] if_P[OF `0 < ?d`] - by auto - next - case False - have "m = m div ?p * ?p + m mod ?p" unfolding mod_div_equality .. - also have "\ \ (m div ?p + 1) * ?p" unfolding left_distrib mult_1 by (rule add_left_mono, rule pos_mod_bound[OF `0 < ?p`, THEN less_imp_le]) - finally have "real (Float m e) \ real (Float (m div ?p + 1) (e + ?d))" unfolding real_of_float_simp add_commute[of e] - unfolding pow2_add mult_assoc[symmetric] real_of_int_le_iff[of m, symmetric] - by (auto intro!: mult_mono simp add: pow2_add `0 < ?d` pow_d) - thus ?thesis - unfolding Float round_up.simps Let_def if_not_P[OF `\ m mod ?p = 0`] if_P[OF `0 < ?d`] . - qed - next - case False - show ?thesis - unfolding Float round_up.simps Let_def if_not_P[OF False] .. - qed -qed - -lemma round_down: "real (round_down prec x) \ real x" -proof (cases x) - case (Float m e) - let ?d = "bitlen m - int prec" - let ?p = "(2::int)^nat ?d" - have "0 < ?p" by auto - show "?thesis" - proof (cases "0 < ?d") - case True - hence pow_d: "pow2 ?d = real ?p" using pow2_int[symmetric] by simp - have "m div ?p * ?p \ m div ?p * ?p + m mod ?p" by (auto simp add: pos_mod_bound[OF `0 < ?p`, THEN less_imp_le]) - also have "\ \ m" unfolding mod_div_equality .. - finally have "real (Float (m div ?p) (e + ?d)) \ real (Float m e)" unfolding real_of_float_simp add_commute[of e] - unfolding pow2_add mult_assoc[symmetric] real_of_int_le_iff[of _ m, symmetric] - by (auto intro!: mult_mono simp add: pow2_add `0 < ?d` pow_d) - thus ?thesis - unfolding Float round_down.simps Let_def if_P[OF `0 < ?d`] . - next - case False - show ?thesis - unfolding Float round_down.simps Let_def if_not_P[OF False] .. - qed -qed - -definition lb_mult :: "nat \ float \ float \ float" where - "lb_mult prec x y = - (case normfloat (x * y) of Float m e \ - let - l = bitlen m - int prec - in if l > 0 then Float (m div (2^nat l)) (e + l) - else Float m e)" - -definition ub_mult :: "nat \ float \ float \ float" where - "ub_mult prec x y = - (case normfloat (x * y) of Float m e \ - let - l = bitlen m - int prec - in if l > 0 then Float (m div (2^nat l) + 1) (e + l) - else Float m e)" - -lemma lb_mult: "real (lb_mult prec x y) \ real (x * y)" -proof (cases "normfloat (x * y)") - case (Float m e) - hence "odd m \ (m = 0 \ e = 0)" by (rule normfloat_imp_odd_or_zero) - let ?l = "bitlen m - int prec" - have "real (lb_mult prec x y) \ real (normfloat (x * y))" - proof (cases "?l > 0") - case False thus ?thesis unfolding lb_mult_def Float Let_def float.cases by auto - next - case True - have "real (m div 2^(nat ?l)) * pow2 ?l \ real m" - proof - - have "real (m div 2^(nat ?l)) * pow2 ?l = real (2^(nat ?l) * (m div 2^(nat ?l)))" unfolding real_of_int_mult real_of_int_power real_numeral unfolding pow2_int[symmetric] - using `?l > 0` by auto - also have "\ \ real (2^(nat ?l) * (m div 2^(nat ?l)) + m mod 2^(nat ?l))" unfolding real_of_int_add by auto - also have "\ = real m" unfolding zmod_zdiv_equality[symmetric] .. - finally show ?thesis by auto - qed - thus ?thesis unfolding lb_mult_def Float Let_def float.cases if_P[OF True] real_of_float_simp pow2_add mult_commute mult_assoc by auto - qed - also have "\ = real (x * y)" unfolding normfloat .. - finally show ?thesis . + finally show ?thesis by auto qed -lemma ub_mult: "real (x * y) \ real (ub_mult prec x y)" -proof (cases "normfloat (x * y)") - case (Float m e) - hence "odd m \ (m = 0 \ e = 0)" by (rule normfloat_imp_odd_or_zero) - let ?l = "bitlen m - int prec" - have "real (x * y) = real (normfloat (x * y))" unfolding normfloat .. - also have "\ \ real (ub_mult prec x y)" - proof (cases "?l > 0") - case False thus ?thesis unfolding ub_mult_def Float Let_def float.cases by auto - next - case True - have "real m \ real (m div 2^(nat ?l) + 1) * pow2 ?l" - proof - - have "m mod 2^(nat ?l) < 2^(nat ?l)" by (rule pos_mod_bound) auto - hence mod_uneq: "real (m mod 2^(nat ?l)) \ 1 * 2^(nat ?l)" unfolding mult_1 real_of_int_less_iff[symmetric] by auto - - have "real m = real (2^(nat ?l) * (m div 2^(nat ?l)) + m mod 2^(nat ?l))" unfolding zmod_zdiv_equality[symmetric] .. - also have "\ = real (m div 2^(nat ?l)) * 2^(nat ?l) + real (m mod 2^(nat ?l))" unfolding real_of_int_add by auto - also have "\ \ (real (m div 2^(nat ?l)) + 1) * 2^(nat ?l)" unfolding left_distrib using mod_uneq by auto - finally show ?thesis unfolding pow2_int[symmetric] using True by auto - qed - thus ?thesis unfolding ub_mult_def Float Let_def float.cases if_P[OF True] real_of_float_simp pow2_add mult_commute mult_assoc by auto - qed - finally show ?thesis . -qed +lemma float_divr_nonpos_pos_upper_bound: + "real x \ 0 \ 0 < real y \ real (float_divr prec x y) \ 0" + by transfer (auto simp: field_simps mult_le_0_iff divide_le_0_iff round_up_def) + +lemma float_divr_nonneg_neg_upper_bound: + "0 \ real x \ real y < 0 \ real (float_divr prec x y) \ 0" + by transfer (auto simp: field_simps mult_le_0_iff zero_le_mult_iff divide_le_0_iff round_up_def) + +lift_definition float_round_up :: "nat \ float \ float" is + "\(prec::nat) x. round_up (prec - \log 2 \x\\ - 1) x" by simp + +lemma float_round_up: "real x \ real (float_round_up prec x)" + using round_up by transfer simp + +lift_definition float_round_down :: "nat \ float \ float" is + "\(prec::nat) x. round_down (prec - \log 2 \x\\ - 1) x" by simp + +lemma float_round_down: "real (float_round_down prec x) \ real x" + using round_down by transfer simp + +lemma floor_add2[simp]: "\ real i + x \ = i + \ x \" + using floor_add[of x i] by (simp del: floor_add add: ac_simps) + +lemma compute_float_round_down[code]: + "float_round_down prec (Float m e) = (let d = bitlen (abs m) - int prec in + if 0 < d then let P = 2^nat d ; n = m div P in Float n (e + d) + else Float m e)" + using compute_float_down[of "prec - bitlen \m\ - e" m e, symmetric] + unfolding Let_def + by transfer (simp add: field_simps abs_mult log_mult bitlen_def cong del: if_weak_cong) + +lemma compute_float_round_up[code]: + "float_round_up prec (Float m e) = (let d = (bitlen (abs m) - int prec) in + if 0 < d then let P = 2^nat d ; n = m div P ; r = m mod P + in Float (n + (if r = 0 then 0 else 1)) (e + d) + else Float m e)" + using compute_float_up[of "prec - bitlen \m\ - e" m e, symmetric] + unfolding Let_def + by transfer (simp add: field_simps abs_mult log_mult bitlen_def cong del: if_weak_cong) -primrec float_abs :: "float \ float" where - "float_abs (Float m e) = Float \m\ e" +lemma Float_le_zero_iff: "Float a b \ 0 \ a \ 0" + apply (auto simp: zero_float_def mult_le_0_iff) + using powr_gt_zero[of 2 b] by simp + +(* TODO: how to use as code equation? -> pprt_float?! *) +lemma compute_pprt[code]: "pprt (Float a e) = (if a <= 0 then 0 else (Float a e))" +unfolding pprt_def sup_float_def max_def Float_le_zero_iff .. + +(* TODO: how to use as code equation? *) +lemma compute_nprt[code]: "nprt (Float a e) = (if a <= 0 then (Float a e) else 0)" +unfolding nprt_def inf_float_def min_def Float_le_zero_iff .. + +lemma of_float_pprt[simp]: fixes a::float shows "real (pprt a) = pprt (real a)" + unfolding pprt_def sup_float_def max_def sup_real_def by auto + +lemma of_float_nprt[simp]: fixes a::float shows "real (nprt a) = nprt (real a)" + unfolding nprt_def inf_float_def min_def inf_real_def by auto + +lift_definition int_floor_fl :: "float \ int" is floor by simp -instantiation float :: abs -begin -definition abs_float_def: "\x\ = float_abs x" -instance .. +lemma compute_int_floor_fl[code]: + "int_floor_fl (Float m e) = (if 0 \ e then m * 2 ^ nat e else m div (2 ^ (nat (-e))))" + by transfer (simp add: powr_int int_of_reals floor_divide_eq_div del: real_of_ints) + +lift_definition floor_fl :: "float \ float" is "\x. real (floor x)" by simp + +lemma compute_floor_fl[code]: + "floor_fl (Float m e) = (if 0 \ e then Float m e else Float (m div (2 ^ (nat (-e)))) 0)" + by transfer (simp add: powr_int int_of_reals floor_divide_eq_div del: real_of_ints) + +lemma floor_fl: "real (floor_fl x) \ real x" by transfer simp + +lemma int_floor_fl: "real (int_floor_fl x) \ real x" by transfer simp + +lemma floor_pos_exp: "exponent (floor_fl x) \ 0" +proof cases + assume nzero: "floor_fl x \ float_of 0" + have "floor_fl x = Float \real x\ 0" by transfer simp + from denormalize_shift[OF this[THEN eq_reflection] nzero] guess i . note i = this + thus ?thesis by simp +qed (simp add: floor_fl_def) + end -lemma real_of_float_abs: "real \x :: float\ = \real x\" -proof (cases x) - case (Float m e) - have "\real m\ * pow2 e = \real m * pow2 e\" unfolding abs_mult by auto - thus ?thesis unfolding Float abs_float_def float_abs.simps real_of_float_simp by auto -qed - -primrec floor_fl :: "float \ float" where - "floor_fl (Float m e) = (if 0 \ e then Float m e - else Float (m div (2 ^ (nat (-e)))) 0)" - -lemma floor_fl: "real (floor_fl x) \ real x" -proof (cases x) - case (Float m e) - show ?thesis - proof (cases "0 \ e") - case False - hence me_eq: "pow2 (-e) = pow2 (int (nat (-e)))" by auto - have "real (Float (m div (2 ^ (nat (-e)))) 0) = real (m div 2 ^ (nat (-e)))" unfolding real_of_float_simp by auto - also have "\ \ real m / real ((2::int) ^ (nat (-e)))" using real_of_int_div4 . - also have "\ = real m * inverse (2 ^ (nat (-e)))" unfolding real_of_int_power real_numeral divide_inverse .. - also have "\ = real (Float m e)" unfolding real_of_float_simp me_eq pow2_int pow2_neg[of e] .. - finally show ?thesis unfolding Float floor_fl.simps if_not_P[OF `\ 0 \ e`] . - next - case True thus ?thesis unfolding Float by auto - qed -qed - -lemma floor_pos_exp: assumes floor: "Float m e = floor_fl x" shows "0 \ e" -proof (cases x) - case (Float mx me) - from floor[unfolded Float floor_fl.simps] show ?thesis by (cases "0 \ me", auto) -qed - -declare floor_fl.simps[simp del] - -primrec ceiling_fl :: "float \ float" where - "ceiling_fl (Float m e) = (if 0 \ e then Float m e - else Float (m div (2 ^ (nat (-e))) + 1) 0)" - -lemma ceiling_fl: "real x \ real (ceiling_fl x)" -proof (cases x) - case (Float m e) - show ?thesis - proof (cases "0 \ e") - case False - hence me_eq: "pow2 (-e) = pow2 (int (nat (-e)))" by auto - have "real (Float m e) = real m * inverse (2 ^ (nat (-e)))" unfolding real_of_float_simp me_eq pow2_int pow2_neg[of e] .. - also have "\ = real m / real ((2::int) ^ (nat (-e)))" unfolding real_of_int_power real_numeral divide_inverse .. - also have "\ \ 1 + real (m div 2 ^ (nat (-e)))" using real_of_int_div3[unfolded diff_le_eq] . - also have "\ = real (Float (m div (2 ^ (nat (-e))) + 1) 0)" unfolding real_of_float_simp by auto - finally show ?thesis unfolding Float ceiling_fl.simps if_not_P[OF `\ 0 \ e`] . - next - case True thus ?thesis unfolding Float by auto - qed -qed - -declare ceiling_fl.simps[simp del] - -definition lb_mod :: "nat \ float \ float \ float \ float" where - "lb_mod prec x ub lb = x - ceiling_fl (float_divr prec x lb) * ub" - -definition ub_mod :: "nat \ float \ float \ float \ float" where - "ub_mod prec x ub lb = x - floor_fl (float_divl prec x ub) * lb" - -lemma lb_mod: fixes k :: int assumes "0 \ real x" and "real k * y \ real x" (is "?k * y \ ?x") - assumes "0 < real lb" "real lb \ y" (is "?lb \ y") "y \ real ub" (is "y \ ?ub") - shows "real (lb_mod prec x ub lb) \ ?x - ?k * y" -proof - - have "?lb \ ?ub" using assms by auto - have "0 \ ?lb" and "?lb \ 0" using assms by auto - have "?k * y \ ?x" using assms by auto - also have "\ \ ?x / ?lb * ?ub" by (metis mult_left_mono[OF `?lb \ ?ub` `0 \ ?x`] divide_right_mono[OF _ `0 \ ?lb` ] times_divide_eq_left nonzero_mult_divide_cancel_right[OF `?lb \ 0`]) - also have "\ \ real (ceiling_fl (float_divr prec x lb)) * ?ub" by (metis mult_right_mono order_trans `0 \ ?lb` `?lb \ ?ub` float_divr ceiling_fl) - finally show ?thesis unfolding lb_mod_def real_of_float_sub real_of_float_mult by auto -qed - -lemma ub_mod: - fixes k :: int and x :: float - assumes "0 \ real x" and "real x \ real k * y" (is "?x \ ?k * y") - assumes "0 < real lb" "real lb \ y" (is "?lb \ y") "y \ real ub" (is "y \ ?ub") - shows "?x - ?k * y \ real (ub_mod prec x ub lb)" -proof - - have "?lb \ ?ub" using assms by auto - hence "0 \ ?lb" and "0 \ ?ub" and "?ub \ 0" using assms by auto - have "real (floor_fl (float_divl prec x ub)) * ?lb \ ?x / ?ub * ?lb" by (metis mult_right_mono order_trans `0 \ ?lb` `?lb \ ?ub` float_divl floor_fl) - also have "\ \ ?x" by (metis mult_left_mono[OF `?lb \ ?ub` `0 \ ?x`] divide_right_mono[OF _ `0 \ ?ub` ] times_divide_eq_left nonzero_mult_divide_cancel_right[OF `?ub \ 0`]) - also have "\ \ ?k * y" using assms by auto - finally show ?thesis unfolding ub_mod_def real_of_float_sub real_of_float_mult by auto -qed - -lemma le_float_def'[code]: "f \ g = (case f - g of Float a b \ a \ 0)" -proof - - have le_transfer: "(f \ g) = (real (f - g) \ 0)" by (auto simp add: le_float_def) - from float_split[of "f - g"] obtain a b where f_diff_g: "f - g = Float a b" by auto - with le_transfer have le_transfer': "f \ g = (real (Float a b) \ 0)" by simp - show ?thesis by (simp add: le_transfer' f_diff_g float_le_zero) -qed - -lemma less_float_def'[code]: "f < g = (case f - g of Float a b \ a < 0)" -proof - - have less_transfer: "(f < g) = (real (f - g) < 0)" by (auto simp add: less_float_def) - from float_split[of "f - g"] obtain a b where f_diff_g: "f - g = Float a b" by auto - with less_transfer have less_transfer': "f < g = (real (Float a b) < 0)" by simp - show ?thesis by (simp add: less_transfer' f_diff_g float_less_zero) -qed - -end diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/Log.thy --- a/src/HOL/Log.thy Thu Apr 19 11:14:57 2012 +0200 +++ b/src/HOL/Log.thy Thu Apr 19 17:49:02 2012 +0200 @@ -168,6 +168,29 @@ "[| 1 < a; 0 < x; 0 < y |] ==> (log a x \ log a y) = (x \ y)" by (simp add: linorder_not_less [symmetric]) +lemma zero_less_log_cancel_iff[simp]: "1 < a \ 0 < x \ 0 < log a x \ 1 < x" + using log_less_cancel_iff[of a 1 x] by simp + +lemma zero_le_log_cancel_iff[simp]: "1 < a \ 0 < x \ 0 \ log a x \ 1 \ x" + using log_le_cancel_iff[of a 1 x] by simp + +lemma log_less_zero_cancel_iff[simp]: "1 < a \ 0 < x \ log a x < 0 \ x < 1" + using log_less_cancel_iff[of a x 1] by simp + +lemma log_le_zero_cancel_iff[simp]: "1 < a \ 0 < x \ log a x \ 0 \ x \ 1" + using log_le_cancel_iff[of a x 1] by simp + +lemma one_less_log_cancel_iff[simp]: "1 < a \ 0 < x \ 1 < log a x \ a < x" + using log_less_cancel_iff[of a a x] by simp + +lemma one_le_log_cancel_iff[simp]: "1 < a \ 0 < x \ 1 \ log a x \ a \ x" + using log_le_cancel_iff[of a a x] by simp + +lemma log_less_one_cancel_iff[simp]: "1 < a \ 0 < x \ log a x < 1 \ x < a" + using log_less_cancel_iff[of a x a] by simp + +lemma log_le_one_cancel_iff[simp]: "1 < a \ 0 < x \ log a x \ 1 \ x \ a" + using log_le_cancel_iff[of a x a] by simp lemma powr_realpow: "0 < x ==> x powr (real n) = x^n" apply (induct n, simp) @@ -176,12 +199,26 @@ apply (subst powr_add, simp, simp) done -lemma powr_realpow2: "0 <= x ==> 0 < n ==> x^n = (if (x = 0) then 0 - else x powr (real n))" +lemma powr_realpow2: "0 <= x ==> 0 < n ==> x^n = (if (x = 0) then 0 else x powr (real n))" apply (case_tac "x = 0", simp, simp) apply (rule powr_realpow [THEN sym], simp) done +lemma powr_int: + assumes "x > 0" + shows "x powr i = (if i \ 0 then x ^ nat i else 1 / x ^ nat (-i))" +proof cases + assume "i < 0" + have r: "x powr i = 1 / x powr (-i)" by (simp add: powr_minus field_simps) + show ?thesis using `i < 0` `x > 0` by (simp add: r field_simps powr_realpow[symmetric]) +qed (simp add: assms powr_realpow[symmetric]) + +lemma powr_numeral: "0 < x \ x powr numeral n = x^numeral n" + using powr_realpow[of x "numeral n"] by simp + +lemma powr_neg_numeral: "0 < x \ x powr neg_numeral n = 1 / x^numeral n" + using powr_int[of x "neg_numeral n"] by simp + lemma root_powr_inverse: "0 < n \ 0 < x \ root n x = x powr (1/n)" by (auto simp: root_def powr_realpow[symmetric] powr_powr) @@ -248,6 +285,10 @@ apply (rule powr_less_mono2, auto) done +lemma powr_inj: + "0 < a \ a \ 1 \ a powr x = a powr y \ x = y" + unfolding powr_def exp_inj_iff by simp + lemma ln_powr_bound: "1 <= x ==> 0 < a ==> ln x <= (x powr a) / a" apply (rule mult_imp_le_div_pos) apply (assumption) diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/Quotient.thy --- a/src/HOL/Quotient.thy Thu Apr 19 11:14:57 2012 +0200 +++ b/src/HOL/Quotient.thy Thu Apr 19 17:49:02 2012 +0200 @@ -38,10 +38,6 @@ definition "set_rel R xs ys \ \x y. R x y \ x \ xs \ y \ ys" -lemma vimage_id: - "vimage id = id" - unfolding vimage_def fun_eq_iff by auto - lemma set_rel_eq: "set_rel op = = op =" by (subst fun_eq_iff, subst fun_eq_iff) (simp add: set_eq_iff set_rel_def) diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/RComplete.thy --- a/src/HOL/RComplete.thy Thu Apr 19 11:14:57 2012 +0200 +++ b/src/HOL/RComplete.thy Thu Apr 19 17:49:02 2012 +0200 @@ -252,6 +252,16 @@ finally show ?thesis unfolding real_of_int_less_iff by simp qed +lemma floor_divide_eq_div: + "floor (real a / real b) = a div b" +proof cases + assume "b \ 0 \ b dvd a" + with real_of_int_div3[of a b] show ?thesis + by (auto simp: real_of_int_div[symmetric] intro!: floor_eq2 real_of_int_div4 neq_le_trans) + (metis add_left_cancel zero_neq_one real_of_int_div_aux real_of_int_inject + real_of_int_zero_cancel right_inverse_eq div_self mod_div_trivial) +qed (auto simp: real_of_int_div) + lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n" unfolding real_of_nat_def by simp diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/RealDef.thy --- a/src/HOL/RealDef.thy Thu Apr 19 11:14:57 2012 +0200 +++ b/src/HOL/RealDef.thy Thu Apr 19 17:49:02 2012 +0200 @@ -1159,6 +1159,18 @@ lemma real_of_int_le_zero_cancel_iff [simp, presburger]: "(real (n::int) <= 0) = (n <= 0)" by (simp add: real_of_int_def) +lemma one_less_real_of_int_cancel_iff: "1 < real (i :: int) \ 1 < i" + unfolding real_of_one[symmetric] real_of_int_less_iff .. + +lemma one_le_real_of_int_cancel_iff: "1 \ real (i :: int) \ 1 \ i" + unfolding real_of_one[symmetric] real_of_int_le_iff .. + +lemma real_of_int_less_one_cancel_iff: "real (i :: int) < 1 \ i < 1" + unfolding real_of_one[symmetric] real_of_int_less_iff .. + +lemma real_of_int_le_one_cancel_iff: "real (i :: int) \ 1 \ i \ 1" + unfolding real_of_one[symmetric] real_of_int_le_iff .. + lemma real_of_int_abs [simp]: "real (abs x) = abs(real (x::int))" by (auto simp add: abs_if) @@ -1555,6 +1567,30 @@ by (auto simp add: power2_eq_square) +lemma numeral_power_le_real_of_nat_cancel_iff[simp]: + "(numeral x::real) ^ n \ real a \ (numeral x::nat) ^ n \ a" + unfolding real_of_nat_le_iff[symmetric] by simp + +lemma real_of_nat_le_numeral_power_cancel_iff[simp]: + "real a \ (numeral x::real) ^ n \ a \ (numeral x::nat) ^ n" + unfolding real_of_nat_le_iff[symmetric] by simp + +lemma numeral_power_le_real_of_int_cancel_iff[simp]: + "(numeral x::real) ^ n \ real a \ (numeral x::int) ^ n \ a" + unfolding real_of_int_le_iff[symmetric] by simp + +lemma real_of_int_le_numeral_power_cancel_iff[simp]: + "real a \ (numeral x::real) ^ n \ a \ (numeral x::int) ^ n" + unfolding real_of_int_le_iff[symmetric] by simp + +lemma neg_numeral_power_le_real_of_int_cancel_iff[simp]: + "(neg_numeral x::real) ^ n \ real a \ (neg_numeral x::int) ^ n \ a" + unfolding real_of_int_le_iff[symmetric] by simp + +lemma real_of_int_le_neg_numeral_power_cancel_iff[simp]: + "real a \ (neg_numeral x::real) ^ n \ a \ (neg_numeral x::int) ^ n" + unfolding real_of_int_le_iff[symmetric] by simp + subsection{*Density of the Reals*} lemma real_lbound_gt_zero: diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/Tools/Lifting/lifting_setup.ML --- a/src/HOL/Tools/Lifting/lifting_setup.ML Thu Apr 19 11:14:57 2012 +0200 +++ b/src/HOL/Tools/Lifting/lifting_setup.ML Thu Apr 19 17:49:02 2012 +0200 @@ -97,7 +97,7 @@ val transfer_attr = Attrib.internal (K Transfer.transfer_add) val (_, qty) = Lifting_Term.quot_thm_rty_qty quot_thm val induct_attr = Attrib.internal (K (Induct.induct_type (fst (dest_Type qty)))) - val qty_name = (Binding.name o fst o dest_Type) qty + val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type) qty fun qualify suffix = Binding.qualified true suffix qty_name val lthy' = case maybe_reflp_thm of SOME reflp_thm => lthy @@ -142,7 +142,7 @@ [typedef_thm, T_def] MRSL @{thm typedef_to_Quotient} val (_, qty) = Lifting_Term.quot_thm_rty_qty quot_thm - val qty_name = (Binding.name o fst o dest_Type) qty + val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type) qty fun qualify suffix = Binding.qualified true suffix qty_name val lthy'' = case typedef_set of diff -r c0fe12591c93 -r 075b98ed1cab src/HOL/Tools/transfer.ML --- a/src/HOL/Tools/transfer.ML Thu Apr 19 11:14:57 2012 +0200 +++ b/src/HOL/Tools/transfer.ML Thu Apr 19 17:49:02 2012 +0200 @@ -14,6 +14,7 @@ val transfer_tac: Proof.context -> int -> tactic val correspondence_tac: Proof.context -> int -> tactic val setup: theory -> theory + val abs_tac: int -> tactic end structure Transfer : TRANSFER = @@ -93,21 +94,8 @@ (** Transfer proof method **) -fun bnames (t $ u) = bnames t @ bnames u - | bnames (Abs (x,_,t)) = x :: bnames t - | bnames _ = [] - -fun rename [] t = (t, []) - | rename (x::xs) ((c as Const (@{const_name Abs}, _)) $ Abs (_, T, t)) = - let val (t', xs') = rename xs t - in (c $ Abs (x, T, t'), xs') end - | rename xs0 (t $ u) = - let val (t', xs1) = rename xs0 t - val (u', xs2) = rename xs1 u - in (t' $ u', xs2) end - | rename xs t = (t, xs) - -fun cert ctxt t = cterm_of (Proof_Context.theory_of ctxt) t +fun dest_Rel (Const (@{const_name Rel}, _) $ r $ x $ y) = (r, x, y) + | dest_Rel t = raise TERM ("dest_Rel", [t]) fun RelT (Const (@{const_name Rel}, _) $ _ $ _ $ y) = fastype_of y @@ -135,23 +123,28 @@ Induct.arbitrary_tac ctxt 0 vs' i end) +(* Apply rule Rel_Abs with appropriate bound variable name *) +val abs_tac = SUBGOAL (fn (t, i) => + let + val pat = (#2 o dest_Rel o HOLogic.dest_Trueprop o Thm.concl_of) @{thm Rel_Abs} + val obj = (#3 o dest_Rel o HOLogic.dest_Trueprop o Logic.strip_assums_concl) t + val rule = Thm.rename_boundvars pat obj @{thm Rel_Abs} + in + rtac rule i + end handle TERM _ => no_tac) + fun transfer_tac ctxt = SUBGOAL (fn (t, i) => let - val bs = bnames t - val rules0 = @{thms Rel_App Rel_Abs} val rules = Data.get ctxt + val app_tac = rtac @{thm Rel_App} in EVERY [rewrite_goal_tac @{thms transfer_forall_eq transfer_implies_eq} i, CONVERSION (Trueprop_conv (fo_conv ctxt)) i, rtac @{thm transfer_start [rotated]} i, - REPEAT_ALL_NEW (resolve_tac rules0 ORELSE' atac + REPEAT_ALL_NEW (app_tac ORELSE' abs_tac ORELSE' atac ORELSE' SOLVED' (REPEAT_ALL_NEW (resolve_tac rules)) ORELSE' eq_tac ctxt) (i + 1), - (* Alpha-rename bound variables in goal *) - SUBGOAL (fn (u, i) => - rtac @{thm equal_elim_rule1} i THEN - rtac (Thm.reflexive (cert ctxt (fst (rename bs u)))) i) i, (* FIXME: rewrite_goal_tac does unwanted eta-contraction *) rewrite_goal_tac post_simps i, rtac @{thm _} i]