diff --git a/thys/Berlekamp_Zassenhaus/Mahler_Measure.thy b/thys/Berlekamp_Zassenhaus/Mahler_Measure.thy --- a/thys/Berlekamp_Zassenhaus/Mahler_Measure.thy +++ b/thys/Berlekamp_Zassenhaus/Mahler_Measure.thy @@ -1,867 +1,868 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada *) subsection \Mahler Measure\ text \This part contains a definition of the Mahler measure, it contains Landau's inequality and the Graeffe-transformation. We also assemble a heuristic to approximate the Mahler's measure.\ theory Mahler_Measure imports Sqrt_Babylonian.Sqrt_Babylonian Poly_Mod_Finite_Field_Record_Based (* stuff about polynomials *) Polynomial_Factorization.Fundamental_Theorem_Algebra_Factorized Polynomial_Factorization.Missing_Multiset begin context comm_monoid_list begin lemma induct_gen_abs: assumes "\ a r. a\set lst \ P (f (h a) r) (f (g a) r)" "\ x y z. P x y \ P y z \ P x z" "P (F (map g lst)) (F (map g lst))" shows "P (F (map h lst)) (F (map g lst)) " using assms proof(induct lst arbitrary:P) case (Cons a as P) have inl:"a\set (a#as)" by auto let ?uf = "\ v w. P (f (g a) v) (f (g a) w)" have p_suc:"?uf (F (map g as)) (F (map g as))" using Cons.prems(3) by auto { fix r aa assume "aa \ set as" hence ins:"aa \ set (a#as)" by auto have "P (f (g a) (f (h aa) r)) (f (g a) (f (g aa) r))" using Cons.prems(1)[of aa "f r (g a)",OF ins] by (auto simp: assoc commute left_commute) } note h = this from Cons.hyps(1)[of ?uf, OF h Cons.prems(2)[simplified] p_suc] have e1:"P (f (g a) (F (map h as))) (f (g a) (F (map g as)))" by simp have e2:"P (f (h a) (F (map h as))) (f (g a) (F (map h as)))" using Cons.prems(1)[OF inl] by blast from Cons(3)[OF e2 e1] show ?case by auto next qed auto end lemma prod_induct_gen: assumes "\ a r. f (h a * r :: 'a :: {comm_monoid_mult}) = f (g a * r)" shows "f (\v\lst. h v) = f (\v\lst. g v)" proof - let "?P x y" = "f x = f y" show ?thesis using comm_monoid_mult_class.prod_list.induct_gen_abs[of _ ?P,OF assms] by auto qed abbreviation complex_of_int::"int \ complex" where "complex_of_int \ of_int" definition l2norm_list :: "int list \ int" where "l2norm_list lst = \sqrt (sum_list (map (\ a. a * a) lst))\" abbreviation l2norm :: "int poly \ int" where "l2norm p \ l2norm_list (coeffs p)" abbreviation "norm2 p \ \a\coeffs p. (cmod a)\<^sup>2" (* the square of the Euclidean/l2-norm *) abbreviation l2norm_complex where "l2norm_complex p \ sqrt (norm2 p)" abbreviation height :: "int poly \ int" where "height p \ max_list (map (nat \ abs) (coeffs p))" definition complex_roots_complex where "complex_roots_complex (p::complex poly) = (SOME as. smult (coeff p (degree p)) (\a\as. [:- a, 1:]) = p \ length as = degree p)" lemma complex_roots: "smult (lead_coeff p) (\a\complex_roots_complex p. [:- a, 1:]) = p" "length (complex_roots_complex p) = degree p" using someI_ex[OF fundamental_theorem_algebra_factorized] unfolding complex_roots_complex_def by simp_all lemma complex_roots_c [simp]: "complex_roots_complex [:c:] = []" using complex_roots(2) [of "[:c:]"] by simp declare complex_roots(2)[simp] lemma complex_roots_1 [simp]: "complex_roots_complex 1 = []" using complex_roots_c [of 1] by (simp add: pCons_one) lemma linear_term_irreducible\<^sub>d[simp]: "irreducible\<^sub>d [: a, 1:]" by (rule linear_irreducible\<^sub>d, simp) definition complex_roots_int where "complex_roots_int (p::int poly) = complex_roots_complex (map_poly of_int p)" lemma complex_roots_int: "smult (lead_coeff p) (\a\complex_roots_int p. [:- a, 1:]) = map_poly of_int p" "length (complex_roots_int p) = degree p" proof - show "smult (lead_coeff p) (\a\complex_roots_int p. [:- a, 1:]) = map_poly of_int p" "length (complex_roots_int p) = degree p" using complex_roots[of "map_poly of_int p"] unfolding complex_roots_int_def by auto qed text \The measure for polynomials, after K. Mahler\ definition mahler_measure_poly where "mahler_measure_poly p = cmod (lead_coeff p) * (\a\complex_roots_complex p. (max 1 (cmod a)))" definition mahler_measure where "mahler_measure p = mahler_measure_poly (map_poly complex_of_int p)" definition mahler_measure_monic where "mahler_measure_monic p = (\a\complex_roots_complex p. (max 1 (cmod a)))" lemma mahler_measure_poly_via_monic : "mahler_measure_poly p = cmod (lead_coeff p) * mahler_measure_monic p" unfolding mahler_measure_poly_def mahler_measure_monic_def by simp lemma smult_inj[simp]: assumes "(a::'a::idom) \ 0" shows "inj (smult a)" proof- interpret map_poly_inj_zero_hom "(*) a" using assms by (unfold_locales, auto) show ?thesis unfolding smult_as_map_poly by (rule inj_f) qed definition reconstruct_poly::"'a::idom \ 'a list \ 'a poly" where "reconstruct_poly c roots = smult c (\a\roots. [:- a, 1:])" lemma reconstruct_is_original_poly: "reconstruct_poly (lead_coeff p) (complex_roots_complex p) = p" using complex_roots(1) by (simp add: reconstruct_poly_def) lemma reconstruct_with_type_conversion: "smult (lead_coeff (map_poly of_int f)) (prod_list (map (\ a. [:- a, 1:]) (complex_roots_int f))) = map_poly of_int f" unfolding complex_roots_int_def complex_roots(1) by simp lemma reconstruct_prod: shows "reconstruct_poly (a::complex) as * reconstruct_poly b bs = reconstruct_poly (a * b) (as @ bs)" unfolding reconstruct_poly_def by auto lemma linear_term_inj[simplified,simp]: "inj (\ a. [:- a, 1::'a::idom:])" unfolding inj_on_def by simp lemma reconstruct_poly_monic_defines_mset: assumes "(\a\as. [:- a, 1:]) = (\a\bs. [:- a, 1::'a::field:])" shows "mset as = mset bs" proof - let ?as = "mset (map (\ a. [:- a, 1:]) as)" let ?bs = "mset (map (\ a. [:- a, 1:]) bs)" have eq_smult:"prod_mset ?as = prod_mset ?bs" using assms by (metis prod_mset_prod_list) have irr:"\ as::'a list. set_mset (mset (map (\ a. [:- a, 1:]) as)) \ {q. irreducible q \ monic q}" by (auto intro!: linear_term_irreducible\<^sub>d[of "-_::'a", simplified]) from monic_factorization_unique_mset[OF eq_smult irr irr] show ?thesis apply (subst inj_eq[OF multiset.inj_map,symmetric]) by auto qed lemma reconstruct_poly_defines_mset_of_argument: assumes "(a::'a::field) \ 0" "reconstruct_poly a as = reconstruct_poly a bs" shows "mset as = mset bs" proof - have eq_smult:"smult a (\a\as. [:- a, 1:]) = smult a (\a\bs. [:- a, 1:])" using assms(2) by (auto simp:reconstruct_poly_def) from reconstruct_poly_monic_defines_mset[OF Fun.injD[OF smult_inj[OF assms(1)] eq_smult]] show ?thesis by simp qed lemma complex_roots_complex_prod [simp]: assumes "f \ 0" "g \ 0" shows "mset (complex_roots_complex (f * g)) = mset (complex_roots_complex f) + mset (complex_roots_complex g)" proof - let ?p = "f * g" let "?lc v" = "(lead_coeff (v:: complex poly))" have nonzero_prod:"?lc ?p \ 0" using assms by auto from reconstruct_prod[of "?lc f" "complex_roots_complex f" "?lc g" "complex_roots_complex g"] have "reconstruct_poly (?lc ?p) (complex_roots_complex ?p) = reconstruct_poly (?lc ?p) (complex_roots_complex f @ complex_roots_complex g)" unfolding lead_coeff_mult[symmetric] reconstruct_is_original_poly by auto from reconstruct_poly_defines_mset_of_argument[OF nonzero_prod this] show ?thesis by simp qed lemma mset_mult_add: assumes "mset (a::'a::field list) = mset b + mset c" shows "prod_list a = prod_list b * prod_list c" unfolding prod_mset_prod_list[symmetric] using prod_mset_Un[of "mset b" "mset c",unfolded assms[symmetric]]. lemma mset_mult_add_2: assumes "mset a = mset b + mset c" shows "prod_list (map i a::'b::field list) = prod_list (map i b) * prod_list (map i c)" proof - have r:"mset (map i a) = mset (map i b) + mset (map i c) " using assms by (metis map_append mset_append mset_map) show ?thesis using mset_mult_add[OF r] by auto qed lemma measure_mono_eq_prod: assumes "f \ 0" "g \ 0" shows "mahler_measure_monic (f * g) = mahler_measure_monic f * mahler_measure_monic g" unfolding mahler_measure_monic_def using mset_mult_add_2[OF complex_roots_complex_prod[OF assms],of "\ a. max 1 (cmod a)"] by simp lemma mahler_measure_poly_0[simp]: "mahler_measure_poly 0 = 0" unfolding mahler_measure_poly_via_monic by auto lemma measure_eq_prod: (* Remark 10.2 *) "mahler_measure_poly (f * g) = mahler_measure_poly f * mahler_measure_poly g" proof - consider "f = 0" | "g = 0" | (both) "f \ 0" "g \ 0" by auto thus ?thesis proof(cases) case both show ?thesis unfolding mahler_measure_poly_via_monic norm_mult lead_coeff_mult by (auto simp: measure_mono_eq_prod[OF both]) qed (simp_all) qed lemma prod_cmod[simp]: "cmod (\a\lst. f a) = (\a\lst. cmod (f a))" by(induct lst,auto simp:real_normed_div_algebra_class.norm_mult) lemma lead_coeff_of_prod[simp]: "lead_coeff (\a\lst. f a::'a::idom poly) = (\a\lst. lead_coeff (f a))" by(induct lst,auto simp:lead_coeff_mult) lemma ineq_about_squares:assumes "x \ (y::real)" shows "x \ c^2 + y" using assms by (simp add: add.commute add_increasing2) lemma first_coeff_le_tail:"(cmod (lead_coeff g))^2 \ (\a\coeffs g. (cmod a)^2)" proof(induct g) case (pCons a p) thus ?case proof(cases "p = 0") case False show ?thesis using pCons unfolding lead_coeff_pCons(1)[OF False] by(cases "a = 0",simp_all add:ineq_about_squares) qed simp qed simp lemma square_prod_cmod[simp]: "(cmod (a * b))^2 = cmod a ^ 2 * cmod b ^ 2" by (simp add: norm_mult power_mult_distrib) lemma sum_coeffs_smult_cmod: "(\a\coeffs (smult v p). (cmod a)^2) = (cmod v)^2 * (\a\coeffs p. (cmod a)^2)" (is "?l = ?r") proof - have "?l = (\a\coeffs p. (cmod v)^2 * (cmod a)^2)" by(cases "v=0";induct p,auto) thus ?thesis by (auto simp:sum_list_const_mult) qed abbreviation "linH a \ if (cmod a > 1) then [:- 1,cnj a:] else [:- a,1:]" lemma coeffs_cong_1[simp]: "cCons a v = cCons b v \ a = b" unfolding cCons_def by auto lemma strip_while_singleton[simp]: "strip_while ((=) 0) [v * a] = cCons (v * a) []" unfolding cCons_def strip_while_def by auto lemma coeffs_times_linterm: shows "coeffs (pCons 0 (smult a p) + smult b p) = strip_while (HOL.eq (0::'a::{comm_ring_1})) (map (\(c,d).b*d+c*a) (zip (0 # coeffs p) (coeffs p @ [0])))" proof - {fix v have "coeffs (smult b p + pCons (a* v) (smult a p)) = strip_while (HOL.eq 0) (map (\(c,d).b*d+c*a) (zip ([v] @ coeffs p) (coeffs p @ [0])))" proof(induct p arbitrary:v) case (pCons pa ps) thus ?case by auto qed auto (* just putting ;auto does not work *) } from this[of 0] show ?thesis by (simp add: add.commute) qed lemma filter_distr_rev[simp]: shows "filter f (rev lst) = rev (filter f lst)" by(induct lst;auto) lemma strip_while_filter: shows "filter ((\) 0) (strip_while ((=) 0) (lst::'a::zero list)) = filter ((\) 0) lst" proof - {fix lst::"'a list" have "filter ((\) 0) (dropWhile ((=) 0) lst) = filter ((\) 0) lst" by (induct lst;auto) hence "(filter ((\) 0) (strip_while ((=) 0) (rev lst))) = filter ((\) 0) (rev lst)" unfolding strip_while_def by(simp)} from this[of "rev lst"] show ?thesis by simp qed lemma sum_stripwhile[simp]: assumes "f 0 = 0" shows "(\a\strip_while ((=) 0) lst. f a) = (\a\lst. f a)" proof - {fix lst have "(\a\filter ((\) 0) lst. f a) = (\a\lst. f a)" by(induct lst,auto simp:assms)} note f=this have "sum_list (map f (filter ((\) 0) (strip_while ((=) 0) lst))) = sum_list (map f (filter ((\) 0) lst))" using strip_while_filter[of lst] by(simp) thus ?thesis unfolding f. qed lemma complex_split : "Complex a b = c \ (a = Re c \ b = Im c)" using complex_surj by auto lemma norm_times_const:"(\y\lst. (cmod (a * y))\<^sup>2) = (cmod a)\<^sup>2 * (\y\lst. (cmod y)\<^sup>2)" by(induct lst,auto simp:ring_distribs) fun bisumTail where (* Used for Landau's lemma *) "bisumTail f (Cons a (Cons b bs)) = f a b + bisumTail f (Cons b bs)" | "bisumTail f (Cons a Nil) = f a 0" | "bisumTail f Nil = f 1 0" (* never called, not used in proofs *) fun bisum where "bisum f (Cons a as) = f 0 a + bisumTail f (Cons a as)" | "bisum f Nil = f 0 0" lemma bisumTail_is_map_zip: "(\x\zip (v # l1) (l1 @ [0]). f x) = bisumTail (\x y .f (x,y)) (v#l1)" by(induct l1 arbitrary:v,auto) (* converting to and from bisum *) lemma bisum_is_map_zip: "(\x\zip (0 # l1) (l1 @ [0]). f x) = bisum (\x y. f (x,y)) l1" using bisumTail_is_map_zip[of f "hd l1" "tl l1"] by(cases l1,auto) lemma map_zip_is_bisum: "bisum f l1 = (\(x,y)\zip (0 # l1) (l1 @ [0]). f x y)" using bisum_is_map_zip[of "\(x,y). f x y"] by auto lemma bisum_outside : "(bisum (\ x y. f1 x - f2 x y + f3 y) lst :: 'a :: field) = sum_list (map f1 lst) + f1 0 - bisum f2 lst + sum_list (map f3 lst) + f3 0" proof(cases lst) case (Cons a lst) show ?thesis unfolding map_zip_is_bisum Cons by(induct lst arbitrary:a,auto) qed auto lemma Landau_lemma: "(\a\coeffs (\a\lst. [:- a, 1:]). (cmod a)\<^sup>2) = (\a\coeffs (\a\lst. linH a). (cmod a)\<^sup>2)" (is "norm2 ?l = norm2 ?r") proof - have a:"\ a. (cmod a)\<^sup>2 = Re (a * cnj a) " using complex_norm_square unfolding complex_split complex_of_real_def by simp have b:"\ x a y. (cmod (x - a * y))^2 = (cmod x)\<^sup>2 - Re (a * y * cnj x + x * cnj (a * y)) + (cmod (a * y))^2" unfolding left_diff_distrib right_diff_distrib a complex_cnj_diff by simp have c:"\ y a x. (cmod (cnj a * x - y))\<^sup>2 = (cmod (a * x))\<^sup>2 - Re (a * y * cnj x + x * cnj (a * y)) + (cmod y)^2" unfolding left_diff_distrib right_diff_distrib a complex_cnj_diff by (simp add: mult.assoc mult.left_commute) { fix f1 a have "norm2 ([:- a, 1 :] * f1) = bisum (\x y. cmod (x - a * y)^2) (coeffs f1)" by(simp add: bisum_is_map_zip[of _ "coeffs f1"] coeffs_times_linterm[of 1 _ "-a",simplified]) also have "\ = norm2 f1 + cmod a^2*norm2 f1 - bisum (\x y. Re (a * y * cnj x + x * cnj (a * y))) (coeffs f1)" unfolding b bisum_outside norm_times_const by simp also have "\ = bisum (\x y. cmod (cnj a * x - y)^2) (coeffs f1)" unfolding c bisum_outside norm_times_const by auto also have "\ = norm2 ([:- 1, cnj a :] * f1)" using coeffs_times_linterm[of "cnj a" _ "-1"] by(simp add: bisum_is_map_zip[of _ "coeffs f1"] mult.commute) finally have "norm2 ([:- a, 1 :] * f1) = \".} hence h:"\ a f1. norm2 ([:- a, 1 :] * f1) = norm2 (linH a * f1)" by auto show ?thesis by(rule prod_induct_gen[OF h]) qed lemma Landau_inequality: "mahler_measure_poly f \ l2norm_complex f" proof - let ?f = "reconstruct_poly (lead_coeff f) (complex_roots_complex f)" let ?roots = "(complex_roots_complex f)" let ?g = "\a\?roots. linH a" (* g is chosen such that lead_coeff_g holds, and its l2 norm is equal to f's l2 norm *) have max:"\a. cmod (if 1 < cmod a then cnj a else 1) = max 1 (cmod a)" by(simp add:if_split,auto) have "\a. 1 < cmod a \ a \ 0" by auto hence "\a. lead_coeff (linH a) = (if (cmod a > 1) then cnj a else 1)" by(auto simp:if_split) hence lead_coeff_g:"cmod (lead_coeff ?g) = (\a\?roots. max 1 (cmod a))" by(auto simp:max) have "norm2 f = (\a\coeffs ?f. (cmod a)^2)" unfolding reconstruct_is_original_poly.. also have "\ = cmod (lead_coeff f)^2 * (\a\coeffs (\a\?roots. [:- a, 1:]). (cmod a)\<^sup>2)" unfolding reconstruct_poly_def using sum_coeffs_smult_cmod. finally have fg_norm:"norm2 f = cmod (lead_coeff f)^2 * (\a\coeffs ?g. (cmod a)^2)" unfolding Landau_lemma by auto have "(cmod (lead_coeff ?g))^2 \ (\a\coeffs ?g. (cmod a)^2)" using first_coeff_le_tail by blast from ordered_comm_semiring_class.comm_mult_left_mono[OF this] have "(cmod (lead_coeff f) * cmod (lead_coeff ?g))^2 \ (\a\coeffs f. (cmod a)^2)" unfolding fg_norm by (simp add:power_mult_distrib) hence "cmod (lead_coeff f) * (\a\?roots. max 1 (cmod a)) \ sqrt (norm2 f)" using NthRoot.real_le_rsqrt lead_coeff_g by auto thus "mahler_measure_poly f \ sqrt (norm2 f)" using reconstruct_with_type_conversion[unfolded complex_roots_int_def] by (simp add: mahler_measure_poly_via_monic mahler_measure_monic_def complex_roots_int_def) qed lemma prod_list_ge1: assumes "Ball (set x) (\ (a::real). a \ 1)" shows "prod_list x \ 1" using assms proof(induct x) case (Cons a as) have "\a\set as. 1 \ a" "1 \ a" using Cons(2) by auto thus ?case using Cons.hyps mult_mono' by fastforce qed auto lemma mahler_measure_monic_ge_1: "mahler_measure_monic p \ 1" unfolding mahler_measure_monic_def by(rule prod_list_ge1,simp) lemma mahler_measure_monic_ge_0: "mahler_measure_monic p \ 0" using mahler_measure_monic_ge_1 le_numeral_extra(1) order_trans by blast lemma mahler_measure_ge_0: "0 \ mahler_measure h" unfolding mahler_measure_def mahler_measure_poly_via_monic by (simp add: mahler_measure_monic_ge_0) lemma mahler_measure_constant[simp]: "mahler_measure_poly [:c:] = cmod c" proof - have main: "complex_roots_complex [:c:] = []" unfolding complex_roots_complex_def by (rule some_equality, auto) show ?thesis unfolding mahler_measure_poly_def main by auto qed lemma mahler_measure_factor[simplified,simp]: "mahler_measure_poly [:- a, 1:] = max 1 (cmod a)" proof - have main: "complex_roots_complex [:- a, 1:] = [a]" unfolding complex_roots_complex_def proof (rule some_equality, auto, goal_cases) case (1 as) thus ?case by (cases as, auto) qed show ?thesis unfolding mahler_measure_poly_def main by auto qed lemma mahler_measure_poly_explicit: "mahler_measure_poly (smult c (\a\as. [:- a, 1:])) = cmod c * (\a\as. (max 1 (cmod a)))" proof (cases "c = 0") case True thus ?thesis by auto next case False note c = this show ?thesis proof (induct as) case (Cons a as) have "mahler_measure_poly (smult c (\a\a # as. [:- a, 1:])) = mahler_measure_poly (smult c (\a\as. [:- a, 1:]) * [: -a, 1 :])" by (rule arg_cong[of _ _ mahler_measure_poly], unfold list.simps prod_list.Cons mult_smult_left, simp) also have "\ = mahler_measure_poly (smult c (\a\as. [:- a, 1:])) * mahler_measure_poly ([:- a, 1:])" (is "_ = ?l * ?r") by (rule measure_eq_prod) also have "?l = cmod c * (\a\as. max 1 (cmod a))" unfolding Cons by simp also have "?r = max 1 (cmod a)" by simp finally show ?case by simp next case Nil show ?case by simp qed qed lemma mahler_measure_poly_ge_1: assumes "h \ 0" shows "(1::real) \ mahler_measure h" proof - have rc: "\real_of_int i\ = of_int \i\" for i by simp from assms have "cmod (lead_coeff (map_poly complex_of_int h)) > 0" by simp hence "cmod (lead_coeff (map_poly complex_of_int h)) \ 1" by(cases "lead_coeff h = 0", auto simp del: leading_coeff_0_iff) from mult_mono[OF this mahler_measure_monic_ge_1 norm_ge_zero] show ?thesis unfolding mahler_measure_def mahler_measure_poly_via_monic by auto qed lemma mahler_measure_dvd: assumes "f \ 0" and "h dvd f" shows "mahler_measure h \ mahler_measure f" proof - from assms obtain g where f: "f = g * h" unfolding dvd_def by auto from f assms have g0: "g \ 0" by auto hence mg: "mahler_measure g \ 1" by (rule mahler_measure_poly_ge_1) have "1 * mahler_measure h \ mahler_measure f" unfolding mahler_measure_def f measure_eq_prod of_int_poly_hom.hom_mult unfolding mahler_measure_def[symmetric] by (rule mult_right_mono[OF mg mahler_measure_ge_0]) thus ?thesis by simp qed definition graeffe_poly :: "'a \ 'a :: comm_ring_1 list \ nat \ 'a poly" where "graeffe_poly c as m = smult (c ^ (2^m)) (\a\as. [:- (a ^ (2^m)), 1:])" context fixes f :: "complex poly" and c as assumes f: "f = smult c (\a\as. [:- a, 1:])" begin lemma mahler_graeffe: "mahler_measure_poly (graeffe_poly c as m) = (mahler_measure_poly f)^(2^m)" proof - have graeffe: "graeffe_poly c as m = smult (c ^ 2 ^ m) (\a\(map (\ a. a ^ 2 ^ m) as). [:- a, 1:])" unfolding graeffe_poly_def by (rule arg_cong[of _ _ "smult (c ^ 2 ^ m)"], induct as, auto) { fix n :: nat assume n: "n > 0" have id: "max 1 (cmod a ^ n) = max 1 (cmod a) ^ n" for a proof (cases "cmod a \ 1") case True hence "cmod a ^ n \ 1" by (simp add: power_le_one) with True show ?thesis by (simp add: max_def) qed (auto simp: max_def) have "(\x\as. max 1 (cmod x ^ n)) = (\a\as. max 1 (cmod a)) ^ n" by (induct as, auto simp: field_simps n id) } thus ?thesis unfolding f mahler_measure_poly_explicit graeffe by (auto simp: o_def field_simps norm_power) qed end fun drop_half :: "'a list \ 'a list" where "drop_half (x # y # ys) = x # drop_half ys" | "drop_half xs = xs" fun alternate :: "'a list \ 'a list \ 'a list" where "alternate (x # y # ys) = (case alternate ys of (evn, od) \ (x # evn, y # od))" | "alternate xs = (xs,[])" definition poly_square_subst :: "'a :: comm_ring_1 poly \ 'a poly" where "poly_square_subst f = poly_of_list (drop_half (coeffs f))" definition poly_even_odd :: "'a :: comm_ring_1 poly \ 'a poly \ 'a poly" where "poly_even_odd f = (case alternate (coeffs f) of (evn,od) \ (poly_of_list evn, poly_of_list od))" lemma poly_square_subst_coeff: "coeff (poly_square_subst f) i = coeff f (2 * i)" proof - have id: "coeff f (2 * i) = coeff (Poly (coeffs f)) (2 * i)" by simp obtain xs where xs: "coeffs f = xs" by auto show ?thesis unfolding poly_square_subst_def poly_of_list_def coeff_Poly_eq id xs proof (induct xs arbitrary: i rule: drop_half.induct) case (1 x y ys i) thus ?case by (cases i, auto) next case ("2_2" x i) thus ?case by (cases i, auto) qed auto qed lemma poly_even_odd_coeff: assumes "poly_even_odd f = (ev,od)" shows "coeff ev i = coeff f (2 * i)" "coeff od i = coeff f (2 * i + 1)" proof - have id: "\ i. coeff f i = coeff (Poly (coeffs f)) i" by simp obtain xs where xs: "coeffs f = xs" by auto from assms[unfolded poly_even_odd_def] have ev_od: "ev = Poly (fst (alternate xs))" "od = Poly (snd (alternate xs))" by (auto simp: xs split: prod.splits) have "coeff ev i = coeff f (2 * i) \ coeff od i = coeff f (2 * i + 1)" unfolding poly_of_list_def coeff_Poly_eq id xs ev_od proof (induct xs arbitrary: i rule: alternate.induct) case (1 x y ys i) thus ?case by (cases "alternate ys"; cases i, auto) next case ("2_2" x i) thus ?case by (cases i, auto) qed auto thus "coeff ev i = coeff f (2 * i)" "coeff od i = coeff f (2 * i + 1)" by auto qed lemma poly_square_subst: "poly_square_subst (f \\<^sub>p (monom 1 2)) = f" by (rule poly_eqI, unfold poly_square_subst_coeff, subst coeff_pcompose_x_pow_n, auto) lemma poly_even_odd: assumes "poly_even_odd f = (g,h)" shows "f = g \\<^sub>p monom 1 2 + monom 1 1 * (h \\<^sub>p monom 1 2)" proof - note id = poly_even_odd_coeff[OF assms] show ?thesis proof (rule poly_eqI, unfold coeff_add coeff_monom_mult) fix n :: nat obtain m i where mi: "m = n div 2" "i = n mod 2" by auto have nmi: "n = 2 * m + i" "i < 2" "0 < (2 :: nat)" "1 < (2 :: nat)" unfolding mi by auto have "(2 :: nat) \ 0" by auto show "coeff f n = coeff (g \\<^sub>p monom 1 2) n + (if 1 \ n then 1 * coeff (h \\<^sub>p monom 1 2) (n - 1) else 0)" proof (cases "i = 1") case True hence id1: "2 * m + i - 1 = 2 * m + 0" by auto show ?thesis unfolding nmi id id1 coeff_pcompose_monom[OF nmi(2)] coeff_pcompose_monom[OF nmi(3)] unfolding True by auto next case False with nmi have i0: "i = 0" by auto show ?thesis proof (cases m) case (Suc k) hence id1: "2 * m + i - 1 = 2 * k + 1" using i0 by auto show ?thesis unfolding nmi id coeff_pcompose_monom[OF nmi(2)] coeff_pcompose_monom[OF nmi(4)] id1 unfolding Suc i0 by auto next case 0 show ?thesis unfolding nmi id coeff_pcompose_monom[OF nmi(2)] unfolding i0 0 by auto qed qed qed qed context fixes f :: "'a :: idom poly" begin lemma graeffe_0: "f = smult c (\a\as. [:- a, 1:]) \ graeffe_poly c as 0 = f" unfolding graeffe_poly_def by auto lemma graeffe_recursion: assumes "graeffe_poly c as m = f" shows "graeffe_poly c as (Suc m) = smult ((-1)^(degree f)) (poly_square_subst (f * f \\<^sub>p [:0,-1:]))" proof - let ?g = "graeffe_poly c as m" have "f * f \\<^sub>p [:0,-1:] = ?g * ?g \\<^sub>p [:0,-1:]" unfolding assms by simp also have "?g \\<^sub>p [:0,-1:] = smult ((- 1) ^ length as) (smult (c ^ 2 ^ m) (\a\as. [:a ^ 2 ^ m, 1:]))" unfolding graeffe_poly_def proof (induct as) case (Cons a as) have "?case = ((smult (c ^ 2 ^ m) ([:- (a ^ 2 ^ m), 1:] \\<^sub>p [:0, - 1:] * (\a\as. [:- (a ^ 2 ^ m), 1:]) \\<^sub>p [:0, - 1:]) = smult (-1 * (- 1) ^ length as) (smult (c ^ 2 ^ m) ([: a ^ 2 ^ m, 1:] * (\a\as. [:a ^ 2 ^ m, 1:])))))" unfolding list.simps prod_list.Cons pcompose_smult pcompose_mult by simp also have "smult (c ^ 2 ^ m) ([:- (a ^ 2 ^ m), 1:] \\<^sub>p [:0, - 1:] * (\a\as. [:- (a ^ 2 ^ m), 1:]) \\<^sub>p [:0, - 1:]) = smult (c ^ 2 ^ m) ((\a\as. [:- (a ^ 2 ^ m), 1:]) \\<^sub>p [:0, - 1:]) * [:- (a ^ 2 ^ m), 1:] \\<^sub>p [:0, - 1:]" unfolding mult_smult_left by simp also have "smult (c ^ 2 ^ m) ((\a\as. [:- (a ^ 2 ^ m), 1:]) \\<^sub>p [:0, - 1:]) = smult ((- 1) ^ length as) (smult (c ^ 2 ^ m) (\a\as. [:a ^ 2 ^ m, 1:]))" unfolding pcompose_smult[symmetric] Cons .. also have "[:- (a ^ 2 ^ m), 1:] \\<^sub>p [:0, - 1:] = smult (-1) [: a^2^m, 1:]" by simp finally have id: "?case = (smult ((- 1) ^ length as) (smult (c ^ 2 ^ m) (\a\as. [:a ^ 2 ^ m, 1:])) * smult (- 1) [:a ^ 2 ^ m, 1:] = smult (- 1 * (- 1) ^ length as) (smult (c ^ 2 ^ m) ([:a ^ 2 ^ m, 1:] * (\a\as. [:a ^ 2 ^ m, 1:]))))" by simp obtain c d where id': "(\a\as. [:a ^ 2 ^ m, 1:]) = c" "[:a ^ 2 ^ m, 1:] = d" by auto show ?case unfolding id unfolding id' by (simp add: ac_simps) qed simp finally have "f * f \\<^sub>p [:0, - 1:] = smult ((- 1) ^ length as * (c ^ 2 ^ m * c ^ 2 ^ m)) ((\a\as. [:- (a ^ 2 ^ m), 1:]) * (\a\as. [:a ^ 2 ^ m, 1:]))" unfolding graeffe_poly_def by (simp add: ac_simps) also have "c ^ 2 ^ m * c ^ 2 ^ m = c ^ 2 ^ (Suc m)" by (simp add: semiring_normalization_rules(36)) also have "(\a\as. [:- (a ^ 2 ^ m), 1:]) * (\a\as. [:a ^ 2 ^ m, 1:]) = (\a\as. [:- (a ^ 2 ^ (Suc m)), 1:]) \\<^sub>p monom 1 2" proof (induct as) case (Cons a as) have id: "(monom 1 2 :: 'a poly) = [:0,0,1:]" by (metis monom_altdef pCons_0_as_mult power2_eq_square smult_1_left) have "(\a\a # as. [:- (a ^ 2 ^ m), 1:]) * (\a\a # as. [:a ^ 2 ^ m, 1:]) = ([:- (a ^ 2 ^ m), 1:] * [: a ^ 2 ^ m, 1:]) * ((\a\ as. [:- (a ^ 2 ^ m), 1:]) * (\a\ as. [:a ^ 2 ^ m, 1:]))" (is "_ = ?a * ?b") unfolding list.simps prod_list.Cons by (simp only: ac_simps) also have "?b = (\a\as. [:- (a ^ 2 ^ Suc m), 1:]) \\<^sub>p monom 1 2" unfolding Cons by simp also have "?a = [: - (a ^ 2 ^ (Suc m)), 0 , 1:]" by (simp add: semiring_normalization_rules(36)) also have "\ = [: - (a ^ 2 ^ (Suc m)), 1:] \\<^sub>p monom 1 2" by (simp add: id) also have "[: - (a ^ 2 ^ (Suc m)), 1:] \\<^sub>p monom 1 2 * (\a\as. [:- (a ^ 2 ^ Suc m), 1:]) \\<^sub>p monom 1 2 = (\a\a # as. [:- (a ^ 2 ^ Suc m), 1:]) \\<^sub>p monom 1 2" unfolding pcompose_mult[symmetric] by simp finally show ?case . qed simp finally have "f * f \\<^sub>p [:0, - 1:] = (smult ((- 1) ^ length as) (graeffe_poly c as (Suc m)) \\<^sub>p monom 1 2)" unfolding graeffe_poly_def pcompose_smult by simp from arg_cong[OF this, of "\ f. smult ((- 1) ^ length as) (poly_square_subst f)", unfolded poly_square_subst] have "graeffe_poly c as (Suc m) = smult ((- 1) ^ length as) (poly_square_subst (f * f \\<^sub>p [:0, - 1:]))" by simp also have "\ = smult ((- 1) ^ degree f) (poly_square_subst (f * f \\<^sub>p [:0, - 1:]))" proof (cases "f = 0") case True thus ?thesis by (auto simp: poly_square_subst_def) next case False with assms have c0: "c \ 0" unfolding graeffe_poly_def by auto from arg_cong[OF assms, of degree] have "degree f = degree (smult (c ^ 2 ^ m) (\a\as. [:- (a ^ 2 ^ m), 1:]))" unfolding graeffe_poly_def by auto also have "\ = degree (\a\as. [:- (a ^ 2 ^ m), 1:])" unfolding degree_smult_eq using c0 by auto also have "\ = length as" unfolding degree_linear_factors by simp finally show ?thesis by simp qed finally show ?thesis . qed end definition graeffe_one_step :: "'a \ 'a :: idom poly \ 'a poly" where "graeffe_one_step c f = smult c (poly_square_subst (f * f \\<^sub>p [:0,-1:]))" lemma graeffe_one_step_code[code]: fixes c :: "'a :: idom" shows "graeffe_one_step c f = (case poly_even_odd f of (g,h) \ smult c (g * g - monom 1 1 * h * h))" proof - obtain g h where eo: "poly_even_odd f = (g,h)" by force from poly_even_odd[OF eo] have fgh: "f = g \\<^sub>p monom 1 2 + monom 1 1 * h \\<^sub>p monom 1 2 " by auto have m2: "monom (1 :: 'a) 2 = [:0,0,1:]" "monom (1 :: 'a) 1 = [:0,1:]" - unfolding coeffs_eq_iff coeffs_monom by auto + unfolding coeffs_eq_iff coeffs_monom + by (auto simp add: numeral_2_eq_2) show ?thesis unfolding eo split graeffe_one_step_def proof (rule arg_cong[of _ _ "smult c"]) let ?g = "g \\<^sub>p monom 1 2" let ?h = "h \\<^sub>p monom 1 2" let ?x = "monom (1 :: 'a) 1" have 2: "2 = Suc (Suc 0)" by simp have "f * f \\<^sub>p [:0, - 1:] = (g \\<^sub>p monom 1 2 + monom 1 1 * h \\<^sub>p monom 1 2) * (g \\<^sub>p monom 1 2 + monom 1 1 * h \\<^sub>p monom 1 2) \\<^sub>p [:0, - 1:]" unfolding fgh by simp also have "(g \\<^sub>p monom 1 2 + monom 1 1 * h \\<^sub>p monom 1 2) \\<^sub>p [:0, - 1:] = g \\<^sub>p (monom 1 2 \\<^sub>p [:0, - 1:]) + monom 1 1 \\<^sub>p [:0, - 1:] * h \\<^sub>p (monom 1 2 \\<^sub>p [:0, - 1:])" unfolding pcompose_add pcompose_mult pcompose_assoc by simp also have "monom (1 :: 'a) 2 \\<^sub>p [:0, - 1:] = monom 1 2" unfolding m2 by auto also have "?x \\<^sub>p [:0, - 1:] = [:0, -1:]" unfolding m2 by auto also have "[:0, - 1:] * h \\<^sub>p monom 1 2 = (-?x) * ?h" unfolding m2 by simp also have "(?g + ?x * ?h) * (?g + (- ?x) * ?h) = (?g * ?g - (?x * ?x) * ?h * ?h)" by (auto simp: field_simps) also have "?x * ?x = ?x \\<^sub>p monom 1 2" unfolding mult_monom by (insert m2, simp add: 2) also have "(?g * ?g - \ * ?h * ?h) = (g * g - ?x * h * h) \\<^sub>p monom 1 2" unfolding pcompose_diff pcompose_mult by auto finally have "poly_square_subst (f * f \\<^sub>p [:0, - 1:]) = poly_square_subst ((g * g - ?x * h * h) \\<^sub>p monom 1 2)" by simp also have "\ = g * g - ?x * h * h" unfolding poly_square_subst by simp finally show "poly_square_subst (f * f \\<^sub>p [:0, - 1:]) = g * g - ?x * h * h" . qed qed fun graeffe_poly_impl_main :: "'a \ 'a :: idom poly \ nat \ 'a poly" where "graeffe_poly_impl_main c f 0 = f" | "graeffe_poly_impl_main c f (Suc m) = graeffe_one_step c (graeffe_poly_impl_main c f m)" lemma graeffe_poly_impl_main: assumes "f = smult c (\a\as. [:- a, 1:])" shows "graeffe_poly_impl_main ((-1)^degree f) f m = graeffe_poly c as m" proof (induct m) case 0 show ?case using graeffe_0[OF assms] by simp next case (Suc m) have [simp]: "degree (graeffe_poly c as m) = degree f" unfolding graeffe_poly_def degree_smult_eq assms degree_linear_factors by auto from arg_cong[OF Suc, of degree] show ?case unfolding graeffe_recursion[OF Suc[symmetric]] by (simp add: graeffe_one_step_def) qed definition graeffe_poly_impl :: "'a :: idom poly \ nat \ 'a poly" where "graeffe_poly_impl f = graeffe_poly_impl_main ((-1)^(degree f)) f" lemma graeffe_poly_impl: assumes "f = smult c (\a\as. [:- a, 1:])" shows "graeffe_poly_impl f m = graeffe_poly c as m" using graeffe_poly_impl_main[OF assms] unfolding graeffe_poly_impl_def . lemma drop_half_map: "drop_half (map f xs) = map f (drop_half xs)" by (induct xs rule: drop_half.induct, auto) lemma (in inj_comm_ring_hom) map_poly_poly_square_subst: "map_poly hom (poly_square_subst f) = poly_square_subst (map_poly hom f)" unfolding poly_square_subst_def coeffs_map_poly_hom drop_half_map poly_of_list_def by (rule poly_eqI, auto simp: nth_default_map_eq) context inj_idom_hom begin lemma graeffe_poly_impl_hom: "map_poly hom (graeffe_poly_impl f m) = graeffe_poly_impl (map_poly hom f) m" proof - interpret mh: map_poly_inj_idom_hom.. obtain c where c: "(((- 1) ^ degree f) :: 'a) = c" by auto have c': "(((- 1) ^ degree f) :: 'b) = hom c" unfolding c[symmetric] by (simp add:hom_distribs) show ?thesis unfolding graeffe_poly_impl_def degree_map_poly_hom c c' apply (induct m arbitrary: f; simp) by (unfold graeffe_one_step_def hom_distribs map_poly_poly_square_subst map_poly_pcompose,simp) qed end lemma graeffe_poly_impl_mahler: "mahler_measure (graeffe_poly_impl f m) = mahler_measure f ^ 2 ^ m" proof - let ?c = "complex_of_int" let ?cc = "map_poly ?c" let ?f = "?cc f" note eq = complex_roots(1)[of ?f] interpret inj_idom_hom complex_of_int by (standard, auto) show ?thesis unfolding mahler_measure_def mahler_graeffe[OF eq[symmetric], symmetric] graeffe_poly_impl[OF eq[symmetric], symmetric] by (simp add: of_int_hom.graeffe_poly_impl_hom) qed definition mahler_landau_graeffe_approximation :: "nat \ nat \ int poly \ int" where "mahler_landau_graeffe_approximation kk dd f = (let no = sum_list (map (\ a. a * a) (coeffs f)) in root_int_floor kk (dd * no))" lemma mahler_landau_graeffe_approximation_core: assumes g: "g = graeffe_poly_impl f k" shows "mahler_measure f \ root (2 ^ Suc k) (real_of_int (\a\coeffs g. a * a))" proof - have "mahler_measure f = root (2^k) (mahler_measure f ^ (2^k))" by (simp add: real_root_power_cancel mahler_measure_ge_0) also have "\ = root (2^k) (mahler_measure g)" unfolding graeffe_poly_impl_mahler g by simp also have "\ = root (2^k) (root 2 (((mahler_measure g)^2)))" by (simp add: real_root_power_cancel mahler_measure_ge_0) also have "\ = root (2^Suc k) (((mahler_measure g)^2))" by (metis power_Suc2 real_root_mult_exp) also have "\ \ root (2 ^ Suc k) (real_of_int (\a\coeffs g. a * a))" proof (rule real_root_le_mono, force) have square_mono: "0 \ (x :: real) \ x \ y \ x * x \ y * y" for x y by (simp add: mult_mono') obtain gs where gs: "coeffs g = gs" by auto have "(mahler_measure g)\<^sup>2 \ real_of_int \\a\coeffs g. a * a\" using square_mono[OF mahler_measure_ge_0 Landau_inequality[of "of_int_poly g", folded mahler_measure_def]] by (auto simp: power2_eq_square coeffs_map_poly o_def of_int_hom.hom_sum_list) also have "\\a\coeffs g. a * a\ = (\a\coeffs g. a * a)" unfolding gs by (induct gs, auto) finally show "(mahler_measure g)\<^sup>2 \ real_of_int (\a\coeffs g. a * a)" . qed finally show "mahler_measure f \ root (2 ^ Suc k) (real_of_int (\a\coeffs g. a * a))" . qed lemma Landau_inequality_mahler_measure: "mahler_measure f \ sqrt (real_of_int (\a\coeffs f. a * a))" by (rule order.trans[OF mahler_landau_graeffe_approximation_core[OF refl, of _ 0]], auto simp: graeffe_poly_impl_def sqrt_def) lemma mahler_landau_graeffe_approximation: assumes g: "g = graeffe_poly_impl f k" "dd = d^(2^(Suc k))" "kk = 2^(Suc k)" shows "\real d * mahler_measure f\ \ mahler_landau_graeffe_approximation kk dd g" proof - have id1: "real_of_int (int (d ^ 2 ^ Suc k)) = (real d) ^ 2 ^ Suc k" by simp have id2: "root (2 ^ Suc k) (real d ^ 2 ^ Suc k) = real d" by (simp add: real_root_power_cancel) show ?thesis unfolding mahler_landau_graeffe_approximation_def Let_def root_int_floor of_int_mult g(2-3) by (rule floor_mono, unfold real_root_mult id1 id2, rule mult_left_mono, rule mahler_landau_graeffe_approximation_core[OF g(1)], auto) qed context fixes bnd :: nat begin (* "dd = d^(2^(Suc k))" "kk = 2^(Suc k)" *) function mahler_approximation_main :: "nat \ int \ int poly \ int \ nat \ nat \ int" where "mahler_approximation_main dd c g mm k kk = (let mmm = mahler_landau_graeffe_approximation kk dd g; new_mm = (if k = 0 then mmm else min mm mmm) in (if k \ bnd then new_mm else \ \abort after \bnd\ iterations of Graeffe transformation\ mahler_approximation_main (dd * dd) c (graeffe_one_step c g) new_mm (Suc k) (2 * kk)))" by pat_completeness auto termination by (relation "measure (\ (dd,c,f,mm,k,kk). Suc bnd - k)", auto) declare mahler_approximation_main.simps[simp del] lemma mahler_approximation_main: assumes "k \ 0 \ \real d * mahler_measure f\ \ mm" and "c = (-1)^(degree f)" and "g = graeffe_poly_impl_main c f k" "dd = d^(2^(Suc k))" "kk = 2^(Suc k)" shows "\real d * mahler_measure f\ \ mahler_approximation_main dd c g mm k kk" using assms proof (induct c g mm k kk rule: mahler_approximation_main.induct) case (1 dd c g mm k kk) let ?df = "\real d * mahler_measure f\" note dd = 1(5) note kk = 1(6) note g = 1(4) note c = 1(3) note mm = 1(2) note IH = 1(1) note mahl = mahler_approximation_main.simps[of dd c g mm k kk] define mmm where "mmm = mahler_landau_graeffe_approximation kk dd g" define new_mm where "new_mm = (if k = 0 then mmm else min mm mmm)" let ?cond = "bnd \ k" have id: "mahler_approximation_main dd c g mm k kk = (if ?cond then new_mm else mahler_approximation_main (dd * dd) c (graeffe_one_step c g) new_mm (Suc k) (2 * kk))" unfolding mahl mmm_def[symmetric] Let_def new_mm_def[symmetric] by simp have gg: "g = (graeffe_poly_impl f k)" unfolding g graeffe_poly_impl_def c .. from mahler_landau_graeffe_approximation[OF gg dd kk, folded mmm_def] have mmm: "?df \ mmm" . with mm have new_mm: "?df \ new_mm" unfolding new_mm_def by auto show ?case proof (cases ?cond) case True show ?thesis unfolding id using True new_mm by auto next case False hence id: "mahler_approximation_main dd c g mm k kk = mahler_approximation_main (dd * dd) c (graeffe_one_step c g) new_mm (Suc k) (2 * kk)" unfolding id by auto have id': "graeffe_one_step c g = graeffe_poly_impl_main c f (Suc k)" unfolding g by simp have "dd * dd = d ^ 2 ^ Suc (Suc k)" "2 * kk = 2 ^ Suc (Suc k)" unfolding dd kk semiring_normalization_rules(26) by auto from IH[OF mmm_def new_mm_def False new_mm c id' this] show ?thesis unfolding id . qed qed definition mahler_approximation :: "nat \ int poly \ int" where "mahler_approximation d f = mahler_approximation_main (d * d) ((-1)^(degree f)) f (-1) 0 2" lemma mahler_approximation: "\real d * mahler_measure f\ \ mahler_approximation d f" unfolding mahler_approximation_def by (rule mahler_approximation_main, auto simp: semiring_normalization_rules(29)) end end diff --git a/thys/Berlekamp_Zassenhaus/Poly_Mod.thy b/thys/Berlekamp_Zassenhaus/Poly_Mod.thy --- a/thys/Berlekamp_Zassenhaus/Poly_Mod.thy +++ b/thys/Berlekamp_Zassenhaus/Poly_Mod.thy @@ -1,1076 +1,1075 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada *) section \Polynomials in Rings and Fields\ subsection \Polynomials in Rings\ text \We use a locale to work with polynomials in some integer-modulo ring.\ theory Poly_Mod imports "HOL-Computational_Algebra.Primes" Polynomial_Factorization.Square_Free_Factorization Unique_Factorization_Poly - "Word_Lib.More_Arithmetic" begin locale poly_mod = fixes m :: "int" begin definition M :: "int \ int" where "M x = x mod m" lemma M_0[simp]: "M 0 = 0" by (auto simp add: M_def) lemma M_M[simp]: "M (M x) = M x" by (auto simp add: M_def) lemma M_plus[simp]: "M (M x + y) = M (x + y)" "M (x + M y) = M (x + y)" by (auto simp add: M_def mod_simps) lemma M_minus[simp]: "M (M x - y) = M (x - y)" "M (x - M y) = M (x - y)" by (auto simp add: M_def mod_simps) lemma M_times[simp]: "M (M x * y) = M (x * y)" "M (x * M y) = M (x * y)" by (auto simp add: M_def mod_simps) lemma M_sum: "M (sum (\ x. M (f x)) A) = M (sum f A)" proof (induct A rule: infinite_finite_induct) case (insert x A) from insert(1-2) have "M (\x\insert x A. M (f x)) = M (f x + M ((\x\A. M (f x))))" by simp also have "M ((\x\A. M (f x))) = M ((\x\A. f x))" using insert by simp finally show ?case using insert by simp qed auto definition inv_M :: "int \ int" where "inv_M = (\ x. if x + x \ m then x else x - m)" lemma M_inv_M_id[simp]: "M (inv_M x) = M x" unfolding inv_M_def M_def by simp definition Mp :: "int poly \ int poly" where "Mp = map_poly M" lemma Mp_0[simp]: "Mp 0 = 0" unfolding Mp_def by auto lemma Mp_coeff: "coeff (Mp f) i = M (coeff f i)" unfolding Mp_def by (simp add: M_def coeff_map_poly) abbreviation eq_m :: "int poly \ int poly \ bool" (infixl "=m" 50) where "f =m g \ (Mp f = Mp g)" notation eq_m (infixl "=m" 50) abbreviation degree_m :: "int poly \ nat" where "degree_m f \ degree (Mp f)" lemma mult_Mp[simp]: "Mp (Mp f * g) = Mp (f * g)" "Mp (f * Mp g) = Mp (f * g)" proof - { fix f g have "Mp (Mp f * g) = Mp (f * g)" unfolding poly_eq_iff Mp_coeff unfolding coeff_mult Mp_coeff proof fix n show "M (\i\n. M (coeff f i) * coeff g (n - i)) = M (\i\n. coeff f i * coeff g (n - i))" by (subst M_sum[symmetric], rule sym, subst M_sum[symmetric], unfold M_times, simp) qed } from this[of f g] this[of g f] show "Mp (Mp f * g) = Mp (f * g)" "Mp (f * Mp g) = Mp (f * g)" by (auto simp: ac_simps) qed lemma plus_Mp[simp]: "Mp (Mp f + g) = Mp (f + g)" "Mp (f + Mp g) = Mp (f + g)" unfolding poly_eq_iff Mp_coeff unfolding coeff_mult Mp_coeff by (auto simp add: Mp_coeff) lemma minus_Mp[simp]: "Mp (Mp f - g) = Mp (f - g)" "Mp (f - Mp g) = Mp (f - g)" unfolding poly_eq_iff Mp_coeff unfolding coeff_mult Mp_coeff by (auto simp add: Mp_coeff) lemma Mp_smult[simp]: "Mp (smult (M a) f) = Mp (smult a f)" "Mp (smult a (Mp f)) = Mp (smult a f)" unfolding Mp_def smult_as_map_poly by (rule poly_eqI, auto simp: coeff_map_poly)+ lemma Mp_Mp[simp]: "Mp (Mp f) = Mp f" unfolding Mp_def by (intro poly_eqI, auto simp: coeff_map_poly) lemma Mp_smult_m_0[simp]: "Mp (smult m f) = 0" by (intro poly_eqI, auto simp: Mp_coeff, auto simp: M_def) definition dvdm :: "int poly \ int poly \ bool" (infix "dvdm" 50) where "f dvdm g = (\ h. g =m f * h)" notation dvdm (infix "dvdm" 50) lemma dvdmE: assumes fg: "f dvdm g" and main: "\h. g =m f * h \ Mp h = h \ thesis" shows "thesis" proof- from fg obtain h where "g =m f * h" by (auto simp: dvdm_def) then have "g =m f * Mp h" by auto from main[OF this] show thesis by auto qed lemma Mp_dvdm[simp]: "Mp f dvdm g \ f dvdm g" and dvdm_Mp[simp]: "f dvdm Mp g \ f dvdm g" by (auto simp: dvdm_def) definition irreducible_m where "irreducible_m f = (\f =m 0 \ \ f dvdm 1 \ (\a b. f =m a * b \ a dvdm 1 \ b dvdm 1))" definition irreducible\<^sub>d_m :: "int poly \ bool" where "irreducible\<^sub>d_m f \ degree_m f > 0 \ (\ g h. degree_m g < degree_m f \ degree_m h < degree_m f \ \ f =m g * h)" definition prime_elem_m where "prime_elem_m f \ \ f =m 0 \ \ f dvdm 1 \ (\g h. f dvdm g * h \ f dvdm g \ f dvdm h)" lemma degree_m_le_degree [intro!]: "degree_m f \ degree f" by (simp add: Mp_def degree_map_poly_le) lemma irreducible\<^sub>d_mI: assumes f0: "degree_m f > 0" and main: "\g h. Mp g = g \ Mp h = h \ degree g > 0 \ degree g < degree_m f \ degree h > 0 \ degree h < degree_m f \ f =m g * h \ False" shows "irreducible\<^sub>d_m f" proof (unfold irreducible\<^sub>d_m_def, intro conjI allI impI f0 notI) fix g h assume deg: "degree_m g < degree_m f" "degree_m h < degree_m f" and "f =m g * h" then have f: "f =m Mp g * Mp h" by simp have "degree_m f \ degree_m g + degree_m h" unfolding f using degree_mult_le order.trans by blast with main[of "Mp g" "Mp h"] deg f show False by auto qed lemma irreducible\<^sub>d_mE: assumes "irreducible\<^sub>d_m f" and "degree_m f > 0 \ (\g h. degree_m g < degree_m f \ degree_m h < degree_m f \ \ f =m g * h) \ thesis" shows thesis using assms by (unfold irreducible\<^sub>d_m_def, auto) lemma irreducible\<^sub>d_mD: assumes "irreducible\<^sub>d_m f" shows "degree_m f > 0" and "\g h. degree_m g < degree_m f \ degree_m h < degree_m f \ \ f =m g * h" using assms by (auto elim: irreducible\<^sub>d_mE) definition square_free_m :: "int poly \ bool" where "square_free_m f = (\ f =m 0 \ (\ g. degree_m g \ 0 \ \ (g * g dvdm f)))" definition coprime_m :: "int poly \ int poly \ bool" where "coprime_m f g = (\ h. h dvdm f \ h dvdm g \ h dvdm 1)" lemma Mp_square_free_m[simp]: "square_free_m (Mp f) = square_free_m f" unfolding square_free_m_def dvdm_def by simp lemma square_free_m_cong: "square_free_m f \ Mp f = Mp g \ square_free_m g" unfolding square_free_m_def dvdm_def by simp lemma Mp_prod_mset[simp]: "Mp (prod_mset (image_mset Mp b)) = Mp (prod_mset b)" proof (induct b) case (add x b) have "Mp (prod_mset (image_mset Mp ({#x#}+b))) = Mp (Mp x * prod_mset (image_mset Mp b))" by simp also have "\ = Mp (Mp x * Mp (prod_mset (image_mset Mp b)))" by simp also have "\ = Mp ( Mp x * Mp (prod_mset b))" unfolding add by simp finally show ?case by simp qed simp lemma Mp_prod_list: "Mp (prod_list (map Mp b)) = Mp (prod_list b)" proof (induct b) case (Cons b xs) have "Mp (prod_list (map Mp (b # xs))) = Mp (Mp b * prod_list (map Mp xs))" by simp also have "\ = Mp (Mp b * Mp (prod_list (map Mp xs)))" by simp also have "\ = Mp (Mp b * Mp (prod_list xs))" unfolding Cons by simp finally show ?case by simp qed simp text \Polynomial evaluation modulo\ definition "M_poly p x \ M (poly p x)" lemma M_poly_Mp[simp]: "M_poly (Mp p) = M_poly p" proof(intro ext, induct p) case 0 show ?case by auto next case IH: (pCons a p) from IH(1) have "M_poly (Mp (pCons a p)) x = M (a + M(x * M_poly (Mp p) x))" by (simp add: M_poly_def Mp_def) also note IH(2)[of x] finally show ?case by (simp add: M_poly_def) qed lemma Mp_lift_modulus: assumes "f =m g" shows "poly_mod.eq_m (m * k) (smult k f) (smult k g)" using assms unfolding poly_eq_iff poly_mod.Mp_coeff coeff_smult unfolding poly_mod.M_def by simp lemma Mp_ident_product: "n > 0 \ Mp f = f \ poly_mod.Mp (m * n) f = f" unfolding poly_eq_iff poly_mod.Mp_coeff poly_mod.M_def by (auto simp add: zmod_zmult2_eq) (metis mod_div_trivial mod_0) lemma Mp_shrink_modulus: assumes "poly_mod.eq_m (m * k) f g" "k \ 0" shows "f =m g" proof - from assms have a: "\ n. coeff f n mod (m * k) = coeff g n mod (m * k)" unfolding poly_eq_iff poly_mod.Mp_coeff unfolding poly_mod.M_def by auto show ?thesis unfolding poly_eq_iff poly_mod.Mp_coeff unfolding poly_mod.M_def proof fix n show "coeff f n mod m = coeff g n mod m" using a[of n] \k \ 0\ by (metis mod_mult_right_eq mult.commute mult_cancel_left mult_mod_right) qed qed lemma degree_m_le: "degree_m f \ degree f" unfolding Mp_def by (rule degree_map_poly_le) lemma degree_m_eq: "coeff f (degree f) mod m \ 0 \ m > 1 \ degree_m f = degree f" using degree_m_le[of f] unfolding Mp_def by (auto intro: degree_map_poly simp: Mp_def poly_mod.M_def) lemma degree_m_mult_le: assumes eq: "f =m g * h" shows "degree_m f \ degree_m g + degree_m h" proof - have "degree_m f = degree_m (Mp g * Mp h)" using eq by simp also have "\ \ degree (Mp g * Mp h)" by (rule degree_m_le) also have "\ \ degree_m g + degree_m h" by (rule degree_mult_le) finally show ?thesis by auto qed lemma degree_m_smult_le: "degree_m (smult c f) \ degree_m f" by (metis Mp_0 coeff_0 degree_le degree_m_le degree_smult_eq poly_mod.Mp_smult(2) smult_eq_0_iff) lemma irreducible_m_Mp[simp]: "irreducible_m (Mp f) \ irreducible_m f" by (simp add: irreducible_m_def) lemma eq_m_irreducible_m: "f =m g \ irreducible_m f \ irreducible_m g" using irreducible_m_Mp by metis definition mset_factors_m where "mset_factors_m F p \ F \ {#} \ (\f. f \# F \ irreducible_m f) \ p =m prod_mset F" end declare poly_mod.M_def[code] declare poly_mod.Mp_def[code] declare poly_mod.inv_M_def[code] definition Irr_Mon :: "'a :: comm_semiring_1 poly set" where "Irr_Mon = {x. irreducible x \ monic x}" definition factorization :: "'a :: comm_semiring_1 poly set \ 'a poly \ ('a \ 'a poly multiset) \ bool" where "factorization Factors f cfs \ (case cfs of (c,fs) \ f = (smult c (prod_mset fs)) \ (set_mset fs \ Factors))" definition unique_factorization :: "'a :: comm_semiring_1 poly set \ 'a poly \ ('a \ 'a poly multiset) \ bool" where "unique_factorization Factors f cfs = (Collect (factorization Factors f) = {cfs})" lemma irreducible_multD: assumes l: "irreducible (a*b)" shows "a dvd 1 \ irreducible b \ b dvd 1 \ irreducible a" proof- from l have "a dvd 1 \ b dvd 1" by auto then show ?thesis proof(elim disjE) assume a: "a dvd 1" with l have "irreducible b" unfolding irreducible_def by (meson is_unit_mult_iff mult.left_commute mult_not_zero) with a show ?thesis by auto next assume a: "b dvd 1" with l have "irreducible a" unfolding irreducible_def by (meson is_unit_mult_iff mult_not_zero semiring_normalization_rules(16)) with a show ?thesis by auto qed qed lemma irreducible_dvd_prod_mset: fixes p :: "'a :: field poly" assumes irr: "irreducible p" and dvd: "p dvd prod_mset as" shows "\ a \# as. p dvd a" proof - from irr[unfolded irreducible_def] have deg: "degree p \ 0" by auto hence p1: "\ p dvd 1" unfolding dvd_def by (metis degree_1 nonzero_mult_div_cancel_left div_poly_less linorder_neqE_nat mult_not_zero not_less0 zero_neq_one) from dvd show ?thesis proof (induct as) case (add a as) hence "prod_mset (add_mset a as) = a * prod_mset as" by auto from add(2)[unfolded this] add(1) irr show ?case by auto qed (insert p1, auto) qed lemma monic_factorization_unique_mset: fixes P::"'a::field poly multiset" assumes eq: "prod_mset P = prod_mset Q" and P: "set_mset P \ {q. irreducible q \ monic q}" and Q: "set_mset Q \ {q. irreducible q \ monic q}" shows "P = Q" proof - { fix P Q :: "'a poly multiset" assume id: "prod_mset P = prod_mset Q" and P: "set_mset P \ {q. irreducible q \ monic q}" and Q: "set_mset Q \ {q. irreducible q \ monic q}" hence "P \# Q" proof (induct P arbitrary: Q) case (add x P Q') from add(3) have irr: "irreducible x" and mon: "monic x" by auto have "\ a \# Q'. x dvd a" proof (rule irreducible_dvd_prod_mset[OF irr]) show "x dvd prod_mset Q'" unfolding add(2)[symmetric] by simp qed then obtain y Q where Q': "Q' = add_mset y Q" and xy: "x dvd y" by (meson mset_add) from add(4) Q' have irr': "irreducible y" and mon': "monic y" by auto have "x = y" using irr irr' xy mon mon' by (metis irreducibleD' irreducible_not_unit poly_dvd_antisym) hence Q': "Q' = Q + {#x#}" using Q' by auto from mon have x0: "x \ 0" by auto from arg_cong[OF add(2)[unfolded Q'], of "\ z. z div x"] have eq: "prod_mset P = prod_mset Q" using x0 by auto from add(3-4)[unfolded Q'] have "set_mset P \ {q. irreducible q \ monic q}" "set_mset Q \ {q. irreducible q \ monic q}" by auto from add(1)[OF eq this] show ?case unfolding Q' by auto qed auto } from this[OF eq P Q] this[OF eq[symmetric] Q P] show ?thesis by auto qed lemma exactly_one_monic_factorization: assumes mon: "monic (f :: 'a :: field poly)" shows "\! fs. f = prod_mset fs \ set_mset fs \ {q. irreducible q \ monic q}" proof - from monic_irreducible_factorization[OF mon] obtain gs g where fin: "finite gs" and f: "f = (\a\gs. a ^ Suc (g a))" and gs: "gs \ {q. irreducible q \ monic q}" by blast from fin have "\ fs. set_mset fs \ gs \ prod_mset fs = (\a\gs. a ^ Suc (g a))" proof (induct gs) case (insert a gs) from insert(3) obtain fs where *: "set_mset fs \ gs" "prod_mset fs = (\a\gs. a ^ Suc (g a))" by auto let ?fs = "fs + replicate_mset (Suc (g a)) a" show ?case proof (rule exI[of _ "fs + replicate_mset (Suc (g a)) a"], intro conjI) show "set_mset ?fs \ insert a gs" using *(1) by auto show "prod_mset ?fs = (\a\insert a gs. a ^ Suc (g a))" by (subst prod.insert[OF insert(1-2)], auto simp: *(2)) qed qed simp then obtain fs where "set_mset fs \ gs" "prod_mset fs = (\a\gs. a ^ Suc (g a))" by auto with gs f have ex: "\fs. f = prod_mset fs \ set_mset fs \ {q. irreducible q \ monic q}" by (intro exI[of _ fs], auto) thus ?thesis using monic_factorization_unique_mset by blast qed lemma monic_prod_mset: fixes as :: "'a :: idom poly multiset" assumes "\ a. a \ set_mset as \ monic a" shows "monic (prod_mset as)" using assms by (induct as, auto intro: monic_mult) lemma exactly_one_factorization: assumes f: "f \ (0 :: 'a :: field poly)" shows "\! cfs. factorization Irr_Mon f cfs" proof - let ?a = "coeff f (degree f)" let ?b = "inverse ?a" let ?g = "smult ?b f" define g where "g = ?g" from f have a: "?a \ 0" "?b \ 0" by (auto simp: field_simps) hence "monic g" unfolding g_def by simp note ex1 = exactly_one_monic_factorization[OF this, folded Irr_Mon_def] then obtain fs where g: "g = prod_mset fs" "set_mset fs \ Irr_Mon" by auto let ?cfs = "(?a,fs)" have cfs: "factorization Irr_Mon f ?cfs" unfolding factorization_def split g(1)[symmetric] using g(2) unfolding g_def by (simp add: a field_simps) show ?thesis proof (rule, rule cfs) fix dgs assume fact: "factorization Irr_Mon f dgs" obtain d gs where dgs: "dgs = (d,gs)" by force from fact[unfolded factorization_def dgs split] have fd: "f = smult d (prod_mset gs)" and gs: "set_mset gs \ Irr_Mon" by auto have "monic (prod_mset gs)" by (rule monic_prod_mset, insert gs[unfolded Irr_Mon_def], auto) hence d: "d = ?a" unfolding fd by auto from arg_cong[OF fd, of "\ x. smult ?b x", unfolded d g_def[symmetric]] have "g = prod_mset gs" using a by (simp add: field_simps) with ex1 g gs have "gs = fs" by auto thus "dgs = ?cfs" unfolding dgs d by auto qed qed lemma mod_ident_iff: "m > 0 \ (x :: int) mod m = x \ x \ {0 ..< m}" by (metis Divides.pos_mod_bound Divides.pos_mod_sign atLeastLessThan_iff mod_pos_pos_trivial) declare prod_mset_prod_list[simp] lemma mult_1_is_id[simp]: "(*) (1 :: 'a :: ring_1) = id" by auto context poly_mod begin lemma degree_m_eq_monic: "monic f \ m > 1 \ degree_m f = degree f" by (rule degree_m_eq) auto lemma monic_degree_m_lift: assumes "monic f" "k > 1" "m > 1" shows "monic (poly_mod.Mp (m * k) f)" proof - have deg: "degree (poly_mod.Mp (m * k) f) = degree f" by (rule poly_mod.degree_m_eq_monic[of f "m * k"], insert assms, auto simp: less_1_mult) show ?thesis unfolding poly_mod.Mp_coeff deg assms poly_mod.M_def using assms(2-) by (simp add: less_1_mult) qed end locale poly_mod_2 = poly_mod m for m + assumes m1: "m > 1" begin lemma M_1[simp]: "M 1 = 1" unfolding M_def using m1 by auto lemma Mp_1[simp]: "Mp 1 = 1" unfolding Mp_def by simp lemma monic_degree_m[simp]: "monic f \ degree_m f = degree f" using degree_m_eq_monic[of f] using m1 by auto lemma monic_Mp: "monic f \ monic (Mp f)" by (auto simp: Mp_coeff) lemma Mp_0_smult_sdiv_poly: assumes "Mp f = 0" shows "smult m (sdiv_poly f m) = f" proof (intro poly_eqI, unfold Mp_coeff coeff_smult sdiv_poly_def, subst coeff_map_poly, force) fix n from assms have "coeff (Mp f) n = 0" by simp hence 0: "coeff f n mod m = 0" unfolding Mp_coeff M_def . thus "m * (coeff f n div m) = coeff f n" by auto qed lemma Mp_product_modulus: "m' = m * k \ k > 0 \ Mp (poly_mod.Mp m' f) = Mp f" by (intro poly_eqI, unfold poly_mod.Mp_coeff poly_mod.M_def, auto simp: mod_mod_cancel) lemma inv_M_rev: assumes bnd: "2 * abs c < m" shows "inv_M (M c) = c" proof (cases "c \ 0") case True with bnd show ?thesis unfolding M_def inv_M_def by auto next case False have 2: "\ v :: int. 2 * v = v + v" by auto from False have c: "c < 0" by auto from bnd c have "c + m > 0" "c + m < m" by auto with c have cm: "c mod m = c + m" by (metis le_less mod_add_self2 mod_pos_pos_trivial) from c bnd have "2 * (c mod m) > m" unfolding cm by auto with bnd c show ?thesis unfolding M_def inv_M_def cm by auto qed end lemma (in poly_mod) degree_m_eq_prime: assumes f0: "Mp f \ 0" and deg: "degree_m f = degree f" and eq: "f =m g * h" and p: "prime m" shows "degree_m f = degree_m g + degree_m h" proof - interpret poly_mod_2 m using prime_ge_2_int[OF p] unfolding poly_mod_2_def by simp from f0 eq have "Mp (Mp g * Mp h) \ 0" by auto hence "Mp g * Mp h \ 0" using Mp_0 by (cases "Mp g * Mp h", auto) hence g0: "Mp g \ 0" and h0: "Mp h \ 0" by auto have "degree (Mp (g * h)) = degree_m (Mp g * Mp h)" by simp also have "\ = degree (Mp g * Mp h)" proof (rule degree_m_eq[OF _ m1], rule) have id: "\ g. coeff (Mp g) (degree (Mp g)) mod m = coeff (Mp g) (degree (Mp g))" unfolding M_def[symmetric] Mp_coeff by simp from p have p': "prime m" unfolding prime_int_nat_transfer unfolding prime_nat_iff by auto assume "coeff (Mp g * Mp h) (degree (Mp g * Mp h)) mod m = 0" from this[unfolded coeff_degree_mult] have "coeff (Mp g) (degree (Mp g)) mod m = 0 \ coeff (Mp h) (degree (Mp h)) mod m = 0" unfolding dvd_eq_mod_eq_0[symmetric] using m1 prime_dvd_mult_int[OF p'] by auto with g0 h0 show False unfolding id by auto qed also have "\ = degree (Mp g) + degree (Mp h)" by (rule degree_mult_eq[OF g0 h0]) finally show ?thesis using eq by simp qed lemma monic_smult_add_small: assumes "f = 0 \ degree f < degree g" and mon: "monic g" shows "monic (g + smult q f)" proof (cases "f = 0") case True thus ?thesis using mon by auto next case False with assms have "degree f < degree g" by auto hence "degree (smult q f) < degree g" by (meson degree_smult_le not_less order_trans) thus ?thesis using mon using coeff_eq_0 degree_add_eq_left by fastforce qed context poly_mod begin definition factorization_m :: "int poly \ (int \ int poly multiset) \ bool" where "factorization_m f cfs \ (case cfs of (c,fs) \ f =m (smult c (prod_mset fs)) \ (\ f \ set_mset fs. irreducible\<^sub>d_m f \ monic (Mp f)))" definition Mf :: "int \ int poly multiset \ int \ int poly multiset" where "Mf cfs \ case cfs of (c,fs) \ (M c, image_mset Mp fs)" lemma Mf_Mf[simp]: "Mf (Mf x) = Mf x" proof (cases x, auto simp: Mf_def, goal_cases) case (1 c fs) show ?case by (induct fs, auto) qed definition equivalent_fact_m :: "int \ int poly multiset \ int \ int poly multiset \ bool" where "equivalent_fact_m cfs dgs = (Mf cfs = Mf dgs)" definition unique_factorization_m :: "int poly \ (int \ int poly multiset) \ bool" where "unique_factorization_m f cfs = (Mf ` Collect (factorization_m f) = {Mf cfs})" lemma Mp_irreducible\<^sub>d_m[simp]: "irreducible\<^sub>d_m (Mp f) = irreducible\<^sub>d_m f" unfolding irreducible\<^sub>d_m_def dvdm_def by simp lemma Mf_factorization_m[simp]: "factorization_m f (Mf cfs) = factorization_m f cfs" unfolding factorization_m_def Mf_def proof (cases cfs, simp, goal_cases) case (1 c fs) have "Mp (smult c (prod_mset fs)) = Mp (smult (M c) (Mp (prod_mset fs)))" by simp also have "\ = Mp (smult (M c) (Mp (prod_mset (image_mset Mp fs))))" unfolding Mp_prod_mset by simp also have "\ = Mp (smult (M c) (prod_mset (image_mset Mp fs)))" unfolding Mp_smult .. finally show ?case by auto qed lemma unique_factorization_m_imp_factorization: assumes "unique_factorization_m f cfs" shows "factorization_m f cfs" proof - from assms[unfolded unique_factorization_m_def] obtain dfs where fact: "factorization_m f dfs" and id: "Mf cfs = Mf dfs" by blast from fact have "factorization_m f (Mf dfs)" by simp from this[folded id] show ?thesis by simp qed lemma unique_factorization_m_alt_def: "unique_factorization_m f cfs = (factorization_m f cfs \ (\ dgs. factorization_m f dgs \ Mf dgs = Mf cfs))" using unique_factorization_m_imp_factorization[of f cfs] unfolding unique_factorization_m_def by auto end context poly_mod_2 begin lemma factorization_m_lead_coeff: assumes "factorization_m f (c,fs)" shows "lead_coeff (Mp f) = M c" proof - note * = assms[unfolded factorization_m_def split] have "monic (prod_mset (image_mset Mp fs))" by (rule monic_prod_mset, insert *, auto) hence "monic (Mp (prod_mset (image_mset Mp fs)))" by (rule monic_Mp) from this[unfolded Mp_prod_mset] have monic: "monic (Mp (prod_mset fs))" by simp from * have "lead_coeff (Mp f) = lead_coeff (Mp (smult c (prod_mset fs)))" by simp also have "Mp (smult c (prod_mset fs)) = Mp (smult (M c) (Mp (prod_mset fs)))" by simp finally show ?thesis using monic \smult c (prod_mset fs) =m smult (M c) (Mp (prod_mset fs))\ by (metis M_M M_def Mp_0 Mp_coeff lead_coeff_smult m1 mult_cancel_left2 poly_mod.degree_m_eq smult_eq_0_iff) qed lemma factorization_m_smult: assumes "factorization_m f (c,fs)" shows "factorization_m (smult d f) (c * d,fs)" proof - note * = assms[unfolded factorization_m_def split] from * have f: "Mp f = Mp (smult c (prod_mset fs))" by simp have "Mp (smult d f) = Mp (smult d (Mp f))" by simp also have "\ = Mp (smult (c * d) (prod_mset fs))" unfolding f by (simp add: ac_simps) finally show ?thesis using assms unfolding factorization_m_def split by auto qed lemma factorization_m_prod: assumes "factorization_m f (c,fs)" "factorization_m g (d,gs)" shows "factorization_m (f * g) (c * d, fs + gs)" proof - note * = assms[unfolded factorization_m_def split] have "Mp (f * g) = Mp (Mp f * Mp g)" by simp also have "Mp f = Mp (smult c (prod_mset fs))" using * by simp also have "Mp g = Mp (smult d (prod_mset gs))" using * by simp finally have "Mp (f * g) = Mp (smult (c * d) (prod_mset (fs + gs)))" unfolding mult_Mp by (simp add: ac_simps) with * show ?thesis unfolding factorization_m_def split by auto qed lemma Mp_factorization_m[simp]: "factorization_m (Mp f) cfs = factorization_m f cfs" unfolding factorization_m_def by simp lemma Mp_unique_factorization_m[simp]: "unique_factorization_m (Mp f) cfs = unique_factorization_m f cfs" unfolding unique_factorization_m_alt_def by simp lemma unique_factorization_m_cong: "unique_factorization_m f cfs \ Mp f = Mp g \ unique_factorization_m g cfs" unfolding Mp_unique_factorization_m[of f, symmetric] by simp lemma unique_factorization_mI: assumes "factorization_m f (c,fs)" and "\ d gs. factorization_m f (d,gs) \ Mf (d,gs) = Mf (c,fs)" shows "unique_factorization_m f (c,fs)" unfolding unique_factorization_m_alt_def by (intro conjI[OF assms(1)] allI impI, insert assms(2), auto) lemma unique_factorization_m_smult: assumes uf: "unique_factorization_m f (c,fs)" and d: "M (di * d) = 1" shows "unique_factorization_m (smult d f) (c * d,fs)" proof (rule unique_factorization_mI[OF factorization_m_smult]) show "factorization_m f (c, fs)" using uf[unfolded unique_factorization_m_alt_def] by auto fix e gs assume fact: "factorization_m (smult d f) (e,gs)" from factorization_m_smult[OF this, of di] have "factorization_m (Mp (smult di (smult d f))) (e * di, gs)" by simp also have "Mp (smult di (smult d f)) = Mp (smult (M (di * d)) f)" by simp also have "\ = Mp f" unfolding d by simp finally have fact: "factorization_m f (e * di, gs)" by simp with uf[unfolded unique_factorization_m_alt_def] have eq: "Mf (e * di, gs) = Mf (c, fs)" by blast from eq[unfolded Mf_def] have "M (e * di) = M c" by simp from arg_cong[OF this, of "\ x. M (x * d)"] have "M (e * M (di * d)) = M (c * d)" by (simp add: ac_simps) from this[unfolded d] have e: "M e = M (c * d)" by simp with eq show "Mf (e,gs) = Mf (c * d, fs)" unfolding Mf_def split by simp qed lemma unique_factorization_m_smultD: assumes uf: "unique_factorization_m (smult d f) (c,fs)" and d: "M (di * d) = 1" shows "unique_factorization_m f (c * di,fs)" proof - from d have d': "M (d * di) = 1" by (simp add: ac_simps) show ?thesis proof (rule unique_factorization_m_cong[OF unique_factorization_m_smult[OF uf d']], rule poly_eqI, unfold Mp_coeff coeff_smult) fix n have "M (di * (d * coeff f n)) = M (M (di * d) * coeff f n)" by (auto simp: ac_simps) from this[unfolded d] show "M (di * (d * coeff f n)) = M (coeff f n)" by simp qed qed lemma degree_m_eq_lead_coeff: "degree_m f = degree f \ lead_coeff (Mp f) = M (lead_coeff f)" by (simp add: Mp_coeff) lemma unique_factorization_m_zero: assumes "unique_factorization_m f (c,fs)" shows "M c \ 0" proof assume c: "M c = 0" from unique_factorization_m_imp_factorization[OF assms] have "Mp f = Mp (smult (M c) (prod_mset fs))" unfolding factorization_m_def split by simp from this[unfolded c] have f: "Mp f = 0" by simp have "factorization_m f (0,{#})" unfolding factorization_m_def split f by auto moreover have "Mf (0,{#}) = (0,{#})" unfolding Mf_def by auto ultimately have fact1: "(0, {#}) \ Mf ` Collect (factorization_m f)" by force define g :: "int poly" where "g = [:0,1:]" have mpg: "Mp g = [:0,1:]" unfolding Mp_def by (auto simp: g_def) { fix g h assume *: "degree (Mp g) = 0" "degree (Mp h) = 0" "[:0, 1:] = Mp (g * h)" from arg_cong[OF *(3), of degree] have "1 = degree_m (Mp g * Mp h)" by simp also have "\ \ degree (Mp g * Mp h)" by (rule degree_m_le) also have "\ \ degree (Mp g) + degree (Mp h)" by (rule degree_mult_le) also have "\ \ 0" using * by simp finally have False by simp } note irr = this have "factorization_m f (0,{# g #})" unfolding factorization_m_def split using irr by (auto simp: irreducible\<^sub>d_m_def f mpg) moreover have "Mf (0,{# g #}) = (0,{# g #})" unfolding Mf_def by (auto simp: mpg, simp add: g_def) ultimately have fact2: "(0, {#g#}) \ Mf ` Collect (factorization_m f)" by force note [simp] = assms[unfolded unique_factorization_m_def] from fact1[simplified, folded fact2[simplified]] show False by auto qed end context poly_mod begin lemma dvdm_smult: assumes "f dvdm g" shows "f dvdm smult c g" proof - from assms[unfolded dvdm_def] obtain h where g: "g =m f * h" by auto show ?thesis unfolding dvdm_def proof (intro exI[of _ "smult c h"]) have "Mp (smult c g) = Mp (smult c (Mp g))" by simp also have "Mp g = Mp (f * h)" using g by simp finally show "Mp (smult c g) = Mp (f * smult c h)" by simp qed qed lemma dvdm_factor: assumes "f dvdm g" shows "f dvdm g * h" proof - from assms[unfolded dvdm_def] obtain k where g: "g =m f * k" by auto show ?thesis unfolding dvdm_def proof (intro exI[of _ "h * k"]) have "Mp (g * h) = Mp (Mp g * h)" by simp also have "Mp g = Mp (f * k)" using g by simp finally show "Mp (g * h) = Mp (f * (h * k))" by (simp add: ac_simps) qed qed lemma square_free_m_smultD: assumes "square_free_m (smult c f)" shows "square_free_m f" unfolding square_free_m_def proof (intro conjI allI impI) fix g assume "degree_m g \ 0" with assms[unfolded square_free_m_def] have "\ g * g dvdm smult c f" by auto thus "\ g * g dvdm f" using dvdm_smult[of "g * g" f c] by blast next from assms[unfolded square_free_m_def] have "\ smult c f =m 0" by simp thus "\ f =m 0" by (metis Mp_smult(2) smult_0_right) qed lemma square_free_m_smultI: assumes sf: "square_free_m f" and inv: "M (ci * c) = 1" shows "square_free_m (smult c f)" proof - have "square_free_m (smult ci (smult c f))" proof (rule square_free_m_cong[OF sf], rule poly_eqI, unfold Mp_coeff coeff_smult) fix n have "M (ci * (c * coeff f n)) = M ( M (ci * c) * coeff f n)" by (simp add: ac_simps) from this[unfolded inv] show "M (coeff f n) = M (ci * (c * coeff f n))" by simp qed from square_free_m_smultD[OF this] show ?thesis . qed lemma square_free_m_factor: assumes "square_free_m (f * g)" shows "square_free_m f" "square_free_m g" proof - { fix f g assume sf: "square_free_m (f * g)" have "square_free_m f" unfolding square_free_m_def proof (intro conjI allI impI) fix h assume "degree_m h \ 0" with sf[unfolded square_free_m_def] have "\ h * h dvdm f * g" by auto thus "\ h * h dvdm f" using dvdm_factor[of "h * h" f g] by blast next from sf[unfolded square_free_m_def] have "\ f * g =m 0" by simp thus "\ f =m 0" by (metis mult.commute mult_zero_right poly_mod.mult_Mp(2)) qed } from this[of f g] this[of g f] assms show "square_free_m f" "square_free_m g" by (auto simp: ac_simps) qed end context poly_mod_2 begin lemma Mp_ident_iff: "Mp f = f \ (\ n. coeff f n \ {0 ..< m})" proof - have m0: "m > 0" using m1 by simp show ?thesis unfolding poly_eq_iff Mp_coeff M_def mod_ident_iff[OF m0] by simp qed lemma Mp_ident_iff': "Mp f = f \ (set (coeffs f) \ {0 ..< m})" proof - have 0: "0 \ {0 ..< m}" using m1 by auto have ran: "(\n. coeff f n \ {0.. range (coeff f) \ {0 ..< m}" by blast show ?thesis unfolding Mp_ident_iff ran using range_coeff[of f] 0 by auto qed end lemma Mp_Mp_pow_is_Mp: "n \ 0 \ p > 1 \ poly_mod.Mp p (poly_mod.Mp (p^n) f) = poly_mod.Mp p f" using poly_mod_2.Mp_product_modulus poly_mod_2_def by(subst power_eq_if, auto) lemma M_M_pow_is_M: "n \ 0 \ p > 1 \ poly_mod.M p (poly_mod.M (p^n) f) = poly_mod.M p f" using Mp_Mp_pow_is_Mp[of n p "[:f:]"] by (metis coeff_pCons_0 poly_mod.Mp_coeff) definition inverse_mod :: "int \ int \ int" where "inverse_mod x m = fst (bezout_coefficients x m)" lemma inverse_mod: "(inverse_mod x m * x) mod m = 1" if "coprime x m" "m > 1" proof - from bezout_coefficients [of x m "inverse_mod x m" "snd (bezout_coefficients x m)"] have "inverse_mod x m * x + snd (bezout_coefficients x m) * m = gcd x m" by (simp add: inverse_mod_def) with that have "inverse_mod x m * x + snd (bezout_coefficients x m) * m = 1" by simp then have "(inverse_mod x m * x + snd (bezout_coefficients x m) * m) mod m = 1 mod m" by simp with \m > 1\ show ?thesis by simp qed lemma inverse_mod_pow: "(inverse_mod x (p ^ n) * x) mod (p ^ n) = 1" if "coprime x p" "p > 1" "n \ 0" using that by (auto intro: inverse_mod) lemma (in poly_mod) inverse_mod_coprime: assumes p: "prime m" and cop: "coprime x m" shows "M (inverse_mod x m * x) = 1" unfolding M_def using inverse_mod_pow[OF cop, of 1] p by (auto simp: prime_int_iff) lemma (in poly_mod) inverse_mod_coprime_exp: assumes m: "m = p^n" and p: "prime p" and n: "n \ 0" and cop: "coprime x p" shows "M (inverse_mod x m * x) = 1" unfolding M_def unfolding m using inverse_mod_pow[OF cop _ n] p by (auto simp: prime_int_iff) locale poly_mod_prime = poly_mod p for p :: int + assumes prime: "prime p" begin sublocale poly_mod_2 p using prime unfolding poly_mod_2_def using prime_gt_1_int by force lemma square_free_m_prod_imp_coprime_m: assumes sf: "square_free_m (A * B)" shows "coprime_m A B" unfolding coprime_m_def proof (intro allI impI) fix h assume dvd: "h dvdm A" "h dvdm B" then obtain ha hb where *: "Mp A = Mp (h * ha)" "Mp B = Mp (h * hb)" unfolding dvdm_def by auto have AB: "Mp (A * B) = Mp (Mp A * Mp B)" by simp from this[unfolded *, simplified] have eq: "Mp (A * B) = Mp (h * h * (ha * hb))" by (simp add: ac_simps) hence dvd_hh: "(h * h) dvdm (A * B)" unfolding dvdm_def by auto { assume "degree_m h \ 0" from sf[unfolded square_free_m_def, THEN conjunct2, rule_format, OF this] have "\ h * h dvdm A * B" . with dvd_hh have False by simp } hence "degree (Mp h) = 0" by auto then obtain c where hc: "Mp h = [: c :]" by (rule degree_eq_zeroE) { assume "c = 0" hence "Mp h = 0" unfolding hc by auto with *(1) have "Mp A = 0" by (metis Mp_0 mult_zero_left poly_mod.mult_Mp(1)) with sf[unfolded square_free_m_def, THEN conjunct1] have False by (simp add: AB) } hence c0: "c \ 0" by auto with arg_cong[OF hc[symmetric], of "\ f. coeff f 0", unfolded Mp_coeff M_def] m1 have "c \ 0" "c < p" by auto with c0 have c_props:"c > 0" "c < p" by auto with prime have "prime p" by simp with c_props have "coprime p c" by (auto intro: prime_imp_coprime dest: zdvd_not_zless) then have "coprime c p" by (simp add: ac_simps) from inverse_mod_coprime[OF prime this] obtain d where d: "M (c * d) = 1" by (auto simp: ac_simps) show "h dvdm 1" unfolding dvdm_def proof (intro exI[of _ "[:d:]"]) have "Mp (h * [: d :]) = Mp (Mp h * [: d :])" by simp also have "\ = Mp ([: c * d :])" unfolding hc by (auto simp: ac_simps) also have "\ = [: M (c * d) :]" unfolding Mp_def by (metis (no_types) M_0 map_poly_pCons Mp_0 Mp_def d zero_neq_one) also have "\ = 1" unfolding d by simp finally show "Mp 1 = Mp (h * [:d:])" by simp qed qed lemma coprime_exp_mod: "coprime lu p \ n \ 0 \ lu mod p ^ n \ 0" using prime by fastforce end context poly_mod begin definition Dp :: "int poly \ int poly" where "Dp f = map_poly (\ a. a div m) f" lemma Dp_Mp_eq: "f = Mp f + smult m (Dp f)" by (rule poly_eqI, auto simp: Mp_coeff M_def Dp_def coeff_map_poly) lemma dvd_imp_dvdm: assumes "a dvd b" shows "a dvdm b" by (metis assms dvd_def dvdm_def) lemma dvdm_add: assumes a: "u dvdm a" and b: "u dvdm b" shows "u dvdm (a+b)" proof - obtain a' where a: "a =m u*a'" using a unfolding dvdm_def by auto obtain b' where b: "b =m u*b'" using b unfolding dvdm_def by auto have "Mp (a + b) = Mp (u*a'+u*b')" using a b by (metis poly_mod.plus_Mp(1) poly_mod.plus_Mp(2)) also have "... = Mp (u * (a'+ b'))" by (simp add: distrib_left) finally show ?thesis unfolding dvdm_def by auto qed lemma monic_dvdm_constant: assumes uk: "u dvdm [:k:]" and u1: "monic u" and u2: "degree u > 0" shows "k mod m = 0" proof - have d1: "degree_m [:k:] = degree [:k:]" by (metis degree_pCons_0 le_zero_eq poly_mod.degree_m_le) obtain h where h: "Mp [:k:] = Mp (u * h)" using uk unfolding dvdm_def by auto have d2: "degree_m [:k:] = degree_m (u*h)" using h by metis have d3: "degree (map_poly M (u * map_poly M h)) = degree (u * map_poly M h)" by (rule degree_map_poly) (metis coeff_degree_mult leading_coeff_0_iff mult.right_neutral M_M Mp_coeff Mp_def u1) thus ?thesis using assms d1 d2 d3 by (auto, metis M_def map_poly_pCons degree_mult_right_le h leD map_poly_0 mult_poly_0_right pCons_eq_0_iff M_0 Mp_def mult_Mp(2)) qed lemma div_mod_imp_dvdm: assumes "\q r. b = q * a + Polynomial.smult m r" shows "a dvdm b" proof - from assms obtain q r where b:"b = a * q + smult m r" by (metis mult.commute) have a: "Mp (Polynomial.smult m r) = 0" by auto show ?thesis proof (unfold dvdm_def, rule exI[of _ q]) have "Mp (a * q + smult m r) = Mp (a * q + Mp (smult m r))" using plus_Mp(2)[of "a*q" "smult m r"] by auto also have "... = Mp (a*q)" by auto finally show "eq_m b (a * q)" using b by auto qed qed lemma lead_coeff_monic_mult: fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" assumes "monic p" shows "lead_coeff (p * q) = lead_coeff q" using assms by (simp add: lead_coeff_mult) lemma degree_m_mult_eq: assumes p: "monic p" and q: "lead_coeff q mod m \ 0" and m1: "m > 1" shows "degree (Mp (p * q)) = degree p + degree q" proof- have "lead_coeff (p * q) mod m \ 0" using q p by (auto simp: lead_coeff_monic_mult) with m1 show ?thesis by (auto simp: degree_m_eq intro!: degree_mult_eq) qed lemma dvdm_imp_degree_le: assumes pq: "p dvdm q" and p: "monic p" and q0: "Mp q \ 0" and m1: "m > 1" shows "degree p \ degree q" proof- from q0 have q: "lead_coeff (Mp q) mod m \ 0" by (metis Mp_Mp Mp_coeff leading_coeff_neq_0 M_def) from pq obtain r where Mpq: "Mp q = Mp (p * Mp r)" by (auto elim: dvdmE) with p q have "lead_coeff (Mp r) mod m \ 0" by (metis Mp_Mp Mp_coeff leading_coeff_0_iff mult_poly_0_right M_def) from degree_m_mult_eq[OF p this m1] Mpq have "degree p \ degree_m q" by simp thus ?thesis using degree_m_le le_trans by blast qed lemma dvdm_uminus [simp]: "p dvdm -q \ p dvdm q" by (metis add.inverse_inverse dvdm_smult smult_1_left smult_minus_left) (*TODO: simp?*) lemma Mp_const_poly: "Mp [:a:] = [:a mod m:]" by (simp add: Mp_def M_def Polynomial.map_poly_pCons) lemma dvdm_imp_div_mod: assumes "u dvdm g" shows "\q r. g = q*u + smult m r" proof - obtain q where q: "Mp g = Mp (u*q)" using assms unfolding dvdm_def by fast have "(u*q) = Mp (u*q) + smult m (Dp (u*q))" by (simp add: poly_mod.Dp_Mp_eq[of "u*q"]) hence uq: "Mp (u*q) = (u*q) - smult m (Dp (u*q))" by auto have g: "g = Mp g + smult m (Dp g)" by (simp add: poly_mod.Dp_Mp_eq[of "g"]) also have "... = poly_mod.Mp m (u*q) + smult m (Dp g)" using q by simp also have "... = u * q - smult m (Dp (u * q)) + smult m (Dp g)" unfolding uq by auto also have "... = u * q + smult m (-Dp (u*q)) + smult m (Dp g)" by auto also have "... = u * q + smult m (-Dp (u*q) + Dp g)" unfolding smult_add_right by auto also have "... = q * u + smult m (-Dp (u*q) + Dp g)" by auto finally show ?thesis by auto qed corollary div_mod_iff_dvdm: shows "a dvdm b = (\q r. b = q * a + Polynomial.smult m r)" using div_mod_imp_dvdm dvdm_imp_div_mod by blast lemma dvdmE': assumes "p dvdm q" and "\r. q =m p * Mp r \ thesis" shows thesis using assms by (auto simp: dvdm_def) end context poly_mod_2 begin lemma factorization_m_mem_dvdm: assumes fact: "factorization_m f (c,fs)" and mem: "Mp g \# image_mset Mp fs" shows "g dvdm f" proof - from fact have "factorization_m f (Mf (c, fs))" by auto then obtain l where f: "factorization_m f (l, image_mset Mp fs)" by (auto simp: Mf_def) from multi_member_split[OF mem] obtain ls where fs: "image_mset Mp fs = {# Mp g #} + ls" by auto from f[unfolded fs split factorization_m_def] show "g dvdm f" unfolding dvdm_def by (intro exI[of _ "smult l (prod_mset ls)"], auto simp del: Mp_smult simp add: Mp_smult(2)[of _ "Mp g * prod_mset ls", symmetric], simp) qed lemma dvdm_degree: "monic u \ u dvdm f \ Mp f \ 0 \ degree u \ degree f" using dvdm_imp_degree_le m1 by blast end lemma (in poly_mod_prime) pl_dvdm_imp_p_dvdm: assumes l0: "l \ 0" and pl_dvdm: "poly_mod.dvdm (p^l) a b" shows "a dvdm b" proof - from l0 have l_gt_0: "l > 0" by auto with m1 interpret pl: poly_mod_2 "p^l" by (unfold_locales, auto) - have p_rw: "p * p ^ (l - 1) = p ^ l" by (rule power_minus_simp[symmetric, OF l_gt_0]) + from l_gt_0 have p_rw: "p * p ^ (l - 1) = p ^ l" + by (cases l) simp_all obtain q r where b: "b = q * a + smult (p^l) r" using pl.dvdm_imp_div_mod[OF pl_dvdm] by auto have "smult (p^l) r = smult p (smult (p ^ (l - 1)) r)" unfolding smult_smult p_rw .. hence b2: "b = q * a + smult p (smult (p ^ (l - 1)) r)" using b by auto show ?thesis by (rule div_mod_imp_dvdm, rule exI[of _ q], rule exI[of _ "(smult (p ^ (l - 1)) r)"], auto simp add: b2) qed - end \ No newline at end of file diff --git a/thys/IP_Addresses/NumberWang_IPv6.thy b/thys/IP_Addresses/NumberWang_IPv6.thy --- a/thys/IP_Addresses/NumberWang_IPv6.thy +++ b/thys/IP_Addresses/NumberWang_IPv6.thy @@ -1,231 +1,231 @@ theory NumberWang_IPv6 imports Word_Lib.Word_Lemmas begin section\Helper Lemmas for Low-Level Operations on Machine Words\ text\Needed for IPv6 Syntax\ lemma length_drop_bl: "length (dropWhile Not (to_bl (of_bl bs))) \ length bs" proof - have length_takeWhile_Not_replicate_False: "length (takeWhile Not (replicate n False @ ls)) = n + length (takeWhile Not ls)" for n ls by(subst takeWhile_append2) simp+ show ?thesis by(simp add: word_rep_drop dropWhile_eq_drop length_takeWhile_Not_replicate_False) qed lemma bl_drop_leading_zeros: "(of_bl:: bool list \ 'a::len word) (dropWhile Not bs) = (of_bl:: bool list \ 'a::len word) bs" by(induction bs) simp_all lemma bl_length_drop_bound: assumes "length (dropWhile Not bs) \ n" shows "length (dropWhile Not (to_bl ((of_bl:: bool list \ 'a::len word) bs))) \ n" proof - have bl_length_drop_twice: "length (dropWhile Not (to_bl ((of_bl:: bool list \ 'a::len word) (dropWhile Not bs)))) = length (dropWhile Not (to_bl ((of_bl:: bool list \ 'a::len word) bs)))" by(simp add: bl_drop_leading_zeros) from length_drop_bl have *: "length (dropWhile Not (to_bl ((of_bl:: bool list \ 'a::len word) bs))) \ length (dropWhile Not bs)" apply(rule dual_order.trans) apply(subst bl_length_drop_twice) .. show ?thesis apply(rule order.trans, rule *) using assms by(simp) qed lemma length_drop_mask_outer: fixes ip::"'a::len word" shows "LENGTH('a) - n' = len \ length (dropWhile Not (to_bl (ip AND (mask n << n') >> n'))) \ len" apply(subst Word_Lemmas.word_and_mask_shiftl) apply(subst Word_Lib.shiftl_shiftr1) apply(simp; fail) apply(simp) apply(subst Word_Lib.and_mask) apply(simp add: word_size) apply(simp add: length_drop_mask) done lemma length_drop_mask_inner: fixes ip::"'a::len word" shows "n \ LENGTH('a) - n' \ length (dropWhile Not (to_bl (ip AND (mask n << n') >> n'))) \ n" apply(subst Word_Lemmas.word_and_mask_shiftl) apply(subst Word_Lemmas.shiftl_shiftr3) apply(simp; fail) apply(simp) apply(simp add: word_size) apply(simp add: Word_Lemmas.mask_twice) apply(simp add: length_drop_mask) done lemma mask128: "0xFFFFFFFFFFFFFFFFFFFFFFFFFFFFFFFF = (mask 128 :: 'a::len word)" by (simp add: mask_eq) (*-------------- things for ipv6 syntax round trip property two ------------------*) (*n small, m large*) lemma helper_masked_ucast_generic: fixes b::"16 word" assumes "n + 16 \ m" and "m < 128" shows "((ucast:: 16 word \ 128 word) b << n) && (mask 16 << m) = 0" proof - have "x < 2 ^ (m - n)" if mnh2: "x < 0x10000" for x::"128 word" proof - from assms(1) have mnh3: "16 \ m - n" by fastforce have power_2_16_nat: "(16::nat) \ n \ (65535::nat) < 2 ^ n" if a:"16 \ n"for n proof - have power2_rule: "a \ b \ (2::nat)^a \ 2 ^ b" for a b by fastforce show ?thesis apply(subgoal_tac "65536 \ 2 ^ n") apply(subst Nat.less_eq_Suc_le) apply(simp; fail) apply(subgoal_tac "(65536::nat) = 2^16") subgoal using power2_rule \16 \ n\ by presburger by(simp) qed have "65536 = unat (65536::128 word)" by auto moreover from mnh2 have "unat x < unat (65536::128 word)" by(rule Word.unat_mono) ultimately have x: "unat x < 65536" by simp with mnh3 have "unat x < 2 ^ (m - n)" apply(rule_tac b=65535 in Orderings.order_class.order.strict_trans1) apply(simp_all) using power_2_16_nat apply blast done with assms(2) show ?thesis by(subst word_less_nat_alt) simp qed hence mnhelper2: "(of_bl::bool list \ 128 word) (to_bl b) < 2 ^ (m - n)" apply(subgoal_tac "(of_bl::bool list \ 128 word) (to_bl b) < 2^(LENGTH(16))") apply(simp; fail) by(rule of_bl_length_less) simp+ have mnhelper3: "(of_bl::bool list \ 128 word) (to_bl b) * 2 ^ n < 2 ^ m" apply(rule Word.div_lt_mult) apply(rule Word_Lemmas.word_less_two_pow_divI) using assms by(simp_all add: mnhelper2 Word_Lib.p2_gt_0) from assms show ?thesis apply(subst ucast_bl)+ apply(subst shiftl_of_bl) apply(subst of_bl_append) apply simp apply(subst word_and_mask_shiftl) apply(subst shiftr_div_2n_w) subgoal by(simp add: word_size; fail) apply(subst word_div_less) subgoal by(rule mnhelper3) apply simp done qed lemma unat_of_bl_128_16_less_helper: fixes b::"16 word" shows "unat ((of_bl::bool list \ 128 word) (to_bl b)) < 2^16" proof - from word_bl_Rep' have 1: "length (to_bl b) = 16" by simp have "unat ((of_bl::bool list \ 128 word) (to_bl b)) < 2^(length (to_bl b))" by(fact unat_of_bl_length) with 1 show ?thesis by auto qed lemma unat_of_bl_128_16_le_helper: "unat ((of_bl:: bool list \ 128 word) (to_bl (b::16 word))) \ 65535" proof - from unat_of_bl_128_16_less_helper[of b] have "unat ((of_bl:: bool list \ 128 word) (to_bl b)) < 65536" by simp from Suc_leI[OF this] show ?thesis by simp qed (*reverse*) lemma helper_masked_ucast_reverse_generic: fixes b::"16 word" assumes "m + 16 \ n" and "n \ 128 - 16" shows "((ucast:: 16 word \ 128 word) b << n) && (mask 16 << m) = 0" proof - have power_less_128_helper: "2 ^ n * unat ((of_bl::bool list \ 128 word) (to_bl b)) < 2 ^ LENGTH(128)" if n: "n \ 128 - 16" for n proof - have help_mult: "n \ l \ 2 ^ n * x < 2 ^ l \ x < 2 ^ (l - n)" for x::nat and l by (simp add: nat_mult_power_less_eq semiring_normalization_rules(7)) from n show ?thesis apply(subst help_mult) subgoal by (simp) apply(rule order_less_le_trans) apply(rule unat_of_bl_128_16_less_helper) apply(rule Power.power_increasing) apply(simp_all) done qed have *: "2 ^ m * (2 ^ (n - m) * unat ((of_bl::bool list \ 128 word) (to_bl b))) = 2 ^ n * unat ((of_bl::bool list \ 128 word) (to_bl b))" proof(cases "unat ((of_bl::bool list \ 128 word) (to_bl b)) = 0") case True thus ?thesis by simp next case False have help_mult: "x \ 0 \ b * (c * x) = a * (x::nat) \ b * c = a" for x a b c by simp from assms show ?thesis apply(subst help_mult[OF False]) apply(subst Power.monoid_mult_class.power_add[symmetric]) apply(simp) done qed from assms have "unat ((2 ^ n)::128 word) * unat ((of_bl::bool list \ 128 word) (to_bl b)) mod 2 ^ LENGTH(128) = 2 ^ m * (2 ^ (n - m) * unat ((of_bl::bool list \ 128 word) (to_bl b)) mod 2 ^ LENGTH(128))" apply(subst nat_mod_eq') - subgoal apply(subst Aligned.unat_power_lower) + subgoal apply(subst unat_power_lower) subgoal by(simp; fail) subgoal by (rule power_less_128_helper) simp done apply(subst nat_mod_eq') subgoal by(rule power_less_128_helper) simp - apply(subst Aligned.unat_power_lower) + apply(subst unat_power_lower) apply(simp; fail) apply(simp only: *) done hence ex_k: "\k. unat ((2 ^ n)::128 word) * unat ((of_bl::bool list \ 128 word) (to_bl b)) mod 2 ^ LENGTH(128) = 2 ^ m * k" by blast hence aligned: "is_aligned ((of_bl::bool list \ 128 word) (to_bl b) << n) m" unfolding is_aligned_iff_dvd_nat unfolding dvd_def unfolding shiftl_t2n unfolding Word.unat_word_ariths(2) by assumption from assms have of_bl_to_bl_shift_mask: "((of_bl::bool list \ 128 word) (to_bl b) << n) && mask (16 + m) = 0" using is_aligned_mask is_aligned_shiftl by force (*sledgehammer*) show ?thesis apply(subst ucast_bl)+ apply(subst word_and_mask_shiftl) apply(subst aligned_shiftr_mask_shiftl) subgoal by (fact aligned) subgoal by (fact of_bl_to_bl_shift_mask) done qed lemma helper_masked_ucast_equal_generic: fixes b::"16 word" assumes "n \ 128 - 16" shows "ucast (((ucast:: 16 word \ 128 word) b << n) && (mask 16 << n) >> n) = b" proof - have ucast_mask: "(ucast:: 16 word \ 128 word) b && mask 16 = ucast b" by transfer (simp flip: take_bit_eq_mask) from assms have "ucast (((ucast:: 16 word \ 128 word) b && mask (128 - n) && mask 16) && mask (128 - n)) = b" by (auto simp add: nth_ucast word_size intro: word_eqI) thus ?thesis apply(subst word_and_mask_shiftl) apply(subst shiftl_shiftr3) apply(simp; fail) apply(simp) apply(subst shiftl_shiftr3) apply(simp_all add: word_size and.assoc) done qed end diff --git a/thys/Native_Word/Uint.thy b/thys/Native_Word/Uint.thy --- a/thys/Native_Word/Uint.thy +++ b/thys/Native_Word/Uint.thy @@ -1,899 +1,898 @@ (* Title: Uint.thy Author: Peter Lammich, TU Munich Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of default size\ theory Uint imports Code_Target_Word_Base begin text \ This theory provides access to words in the target languages of the code generator whose bit width is the default of the target language. To that end, the type \uint\ models words of width \dflt_size\, but \dflt_size\ is known only to be positive. Usage restrictions: Default-size words (type \uint\) cannot be used for evaluation, because the results depend on the particular choice of word size in the target language and implementation. Symbolic evaluation has not yet been set up for \uint\. \ text \The default size type\ typedecl dflt_size instantiation dflt_size :: typerep begin definition "typerep_class.typerep \ \_ :: dflt_size itself. Typerep.Typerep (STR ''Uint.dflt_size'') []" instance .. end consts dflt_size_aux :: "nat" specification (dflt_size_aux) dflt_size_aux_g0: "dflt_size_aux > 0" by auto hide_fact dflt_size_aux_def instantiation dflt_size :: len begin definition "len_of_dflt_size (_ :: dflt_size itself) \ dflt_size_aux" instance by(intro_classes)(simp add: len_of_dflt_size_def dflt_size_aux_g0) end abbreviation "dflt_size \ len_of (TYPE (dflt_size))" context includes integer.lifting begin lift_definition dflt_size_integer :: integer is "int dflt_size" . declare dflt_size_integer_def[code del] \ \The code generator will substitute a machine-dependent value for this constant\ lemma dflt_size_by_int[code]: "dflt_size = nat_of_integer dflt_size_integer" by transfer simp lemma dflt_size[simp]: "dflt_size > 0" "dflt_size \ Suc 0" "\ dflt_size < Suc 0" using len_gt_0[where 'a=dflt_size] by (simp_all del: len_gt_0) end declare prod.Quotient[transfer_rule] section \Type definition and primitive operations\ typedef uint = "UNIV :: dflt_size word set" .. setup_lifting type_definition_uint text \Use an abstract type for code generation to disable pattern matching on @{term Abs_uint}.\ declare Rep_uint_inverse[code abstype] declare Quotient_uint[transfer_rule] instantiation uint :: comm_ring_1 begin lift_definition zero_uint :: uint is "0 :: dflt_size word" . lift_definition one_uint :: uint is "1" . lift_definition plus_uint :: "uint \ uint \ uint" is "(+) :: dflt_size word \ _" . lift_definition minus_uint :: "uint \ uint \ uint" is "(-)" . lift_definition uminus_uint :: "uint \ uint" is uminus . lift_definition times_uint :: "uint \ uint \ uint" is "(*)" . instance by (standard; transfer) (simp_all add: algebra_simps) end instantiation uint :: semiring_modulo begin lift_definition divide_uint :: "uint \ uint \ uint" is "(div)" . lift_definition modulo_uint :: "uint \ uint \ uint" is "(mod)" . instance by (standard; transfer) (fact word_mod_div_equality) end instantiation uint :: linorder begin lift_definition less_uint :: "uint \ uint \ bool" is "(<)" . lift_definition less_eq_uint :: "uint \ uint \ bool" is "(\)" . instance by (standard; transfer) (simp_all add: less_le_not_le linear) end lemmas [code] = less_uint.rep_eq less_eq_uint.rep_eq context includes lifting_syntax notes transfer_rule_of_bool [transfer_rule] transfer_rule_numeral [transfer_rule] begin lemma [transfer_rule]: "((=) ===> cr_uint) of_bool of_bool" by transfer_prover lemma transfer_rule_numeral_uint [transfer_rule]: "((=) ===> cr_uint) numeral numeral" by transfer_prover lemma [transfer_rule]: \(cr_uint ===> (\)) even ((dvd) 2 :: uint \ bool)\ by (unfold dvd_def) transfer_prover end instantiation uint :: semiring_bits begin lift_definition bit_uint :: \uint \ nat \ bool\ is bit . instance by (standard; transfer) (fact bit_iff_odd even_iff_mod_2_eq_zero odd_iff_mod_2_eq_one odd_one bits_induct bits_div_0 bits_div_by_1 bits_mod_div_trivial even_succ_div_2 even_mask_div_iff exp_div_exp_eq div_exp_eq mod_exp_eq mult_exp_mod_exp_eq div_exp_mod_exp_eq even_mult_exp_div_exp_iff)+ end instantiation uint :: semiring_bit_shifts begin lift_definition push_bit_uint :: \nat \ uint \ uint\ is push_bit . lift_definition drop_bit_uint :: \nat \ uint \ uint\ is drop_bit . lift_definition take_bit_uint :: \nat \ uint \ uint\ is take_bit . instance by (standard; transfer) (fact push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod)+ end instantiation uint :: ring_bit_operations begin lift_definition not_uint :: \uint \ uint\ is NOT . lift_definition and_uint :: \uint \ uint \ uint\ is \(AND)\ . lift_definition or_uint :: \uint \ uint \ uint\ is \(OR)\ . lift_definition xor_uint :: \uint \ uint \ uint\ is \(XOR)\ . lift_definition mask_uint :: \nat \ uint\ is mask . instance by (standard; transfer) (simp_all add: bit_and_iff bit_or_iff bit_xor_iff bit_not_iff minus_eq_not_minus_1 mask_eq_decr_exp) end lemma [code]: \take_bit n a = a AND mask n\ for a :: uint by (fact take_bit_eq_mask) lemma [code]: \mask (Suc n) = push_bit n (1 :: uint) OR mask n\ \mask 0 = (0 :: uint)\ by (simp_all add: mask_Suc_exp push_bit_of_1) instance uint :: semiring_bit_syntax .. context includes lifting_syntax begin lemma test_bit_uint_transfer [transfer_rule]: \(cr_uint ===> (=)) bit (!!)\ unfolding test_bit_eq_bit by transfer_prover lemma shiftl_uint_transfer [transfer_rule]: \(cr_uint ===> (=) ===> cr_uint) (\k n. push_bit n k) (<<)\ unfolding shiftl_eq_push_bit by transfer_prover lemma shiftr_uint_transfer [transfer_rule]: \(cr_uint ===> (=) ===> cr_uint) (\k n. drop_bit n k) (>>)\ unfolding shiftr_eq_drop_bit by transfer_prover end instantiation uint :: lsb begin lift_definition lsb_uint :: \uint \ bool\ is lsb . instance by (standard; transfer) (fact lsb_odd) end instantiation uint :: msb begin lift_definition msb_uint :: \uint \ bool\ is msb . instance .. end instantiation uint :: set_bit begin lift_definition set_bit_uint :: \uint \ nat \ bool \ uint\ is set_bit . instance apply standard apply (unfold Bit_Operations.set_bit_def unset_bit_def) apply transfer apply (simp add: set_bit_eq Bit_Operations.set_bit_def unset_bit_def) done end instantiation uint :: bit_comprehension begin lift_definition set_bits_uint :: "(nat \ bool) \ uint" is "set_bits" . instance by (standard; transfer) (fact set_bits_bit_eq) end lemmas [code] = bit_uint.rep_eq lsb_uint.rep_eq msb_uint.rep_eq instantiation uint :: equal begin lift_definition equal_uint :: "uint \ uint \ bool" is "equal_class.equal" . instance by standard (transfer, simp add: equal_eq) end lemmas [code] = equal_uint.rep_eq instantiation uint :: size begin lift_definition size_uint :: "uint \ nat" is "size" . instance .. end lemmas [code] = size_uint.rep_eq lift_definition sshiftr_uint :: "uint \ nat \ uint" (infixl ">>>" 55) is \\w n. signed_drop_bit n w\ . lift_definition uint_of_int :: "int \ uint" is "word_of_int" . text \Use pretty numerals from integer for pretty printing\ context includes integer.lifting begin lift_definition Uint :: "integer \ uint" is "word_of_int" . lemma Rep_uint_numeral [simp]: "Rep_uint (numeral n) = numeral n" by(induction n)(simp_all add: one_uint_def Abs_uint_inverse numeral.simps plus_uint_def) lemma numeral_uint_transfer [transfer_rule]: "(rel_fun (=) cr_uint) numeral numeral" by(auto simp add: cr_uint_def) lemma numeral_uint [code_unfold]: "numeral n = Uint (numeral n)" by transfer simp lemma Rep_uint_neg_numeral [simp]: "Rep_uint (- numeral n) = - numeral n" by(simp only: uminus_uint_def)(simp add: Abs_uint_inverse) lemma neg_numeral_uint [code_unfold]: "- numeral n = Uint (- numeral n)" by transfer(simp add: cr_uint_def) end lemma Abs_uint_numeral [code_post]: "Abs_uint (numeral n) = numeral n" by(induction n)(simp_all add: one_uint_def numeral.simps plus_uint_def Abs_uint_inverse) lemma Abs_uint_0 [code_post]: "Abs_uint 0 = 0" by(simp add: zero_uint_def) lemma Abs_uint_1 [code_post]: "Abs_uint 1 = 1" by(simp add: one_uint_def) section \Code setup\ code_printing code_module Uint \ (SML) \ structure Uint : sig val set_bit : Word.word -> IntInf.int -> bool -> Word.word val shiftl : Word.word -> IntInf.int -> Word.word val shiftr : Word.word -> IntInf.int -> Word.word val shiftr_signed : Word.word -> IntInf.int -> Word.word val test_bit : Word.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word.orb (x, mask) else Word.andb (x, Word.notb mask) end fun shiftl x n = Word.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word.andb (x, Word.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word.fromInt 0 end; (* struct Uint *)\ code_reserved SML Uint code_printing code_module Uint \ (Haskell) \module Uint(Int, Word, dflt_size) where import qualified Prelude import Data.Int(Int) import Data.Word(Word) import qualified Data.Bits dflt_size :: Prelude.Integer dflt_size = Prelude.toInteger (bitSize_aux (0::Word)) where bitSize_aux :: (Data.Bits.Bits a, Prelude.Bounded a) => a -> Int bitSize_aux = Data.Bits.bitSize\ and (Haskell_Quickcheck) \module Uint(Int, Word, dflt_size) where import qualified Prelude import Data.Int(Int) import Data.Word(Word) import qualified Data.Bits dflt_size :: Prelude.Int dflt_size = bitSize_aux (0::Word) where bitSize_aux :: (Data.Bits.Bits a, Prelude.Bounded a) => a -> Int bitSize_aux = Data.Bits.bitSize \ code_reserved Haskell Uint dflt_size text \ OCaml and Scala provide only signed bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module "Uint" \ (OCaml) \module Uint : sig type t = int val dflt_size : Z.t val less : t -> t -> bool val less_eq : t -> t -> bool val set_bit : t -> Z.t -> bool -> t val shiftl : t -> Z.t -> t val shiftr : t -> Z.t -> t val shiftr_signed : t -> Z.t -> t val test_bit : t -> Z.t -> bool val int_mask : int val int32_mask : int32 val int64_mask : int64 end = struct type t = int let dflt_size = Z.of_int Sys.int_size;; (* negative numbers have their highest bit set, so they are greater than positive ones *) let less x y = if x<0 then y<0 && x 0;; let int_mask = if Sys.int_size < 32 then lnot 0 else 0xFFFFFFFF;; let int32_mask = if Sys.int_size < 32 then Int32.pred (Int32.shift_left Int32.one Sys.int_size) else Int32.of_string "0xFFFFFFFF";; let int64_mask = if Sys.int_size < 64 then Int64.pred (Int64.shift_left Int64.one Sys.int_size) else Int64.of_string "0xFFFFFFFFFFFFFFFF";; end;; (*struct Uint*)\ code_reserved OCaml Uint code_printing code_module Uint \ (Scala) \object Uint { def dflt_size : BigInt = BigInt(32) def less(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Int, n: BigInt, b: Boolean) : Int = if (b) x | (1 << n.intValue) else x & (1 << n.intValue).unary_~ def shiftl(x: Int, n: BigInt) : Int = x << n.intValue def shiftr(x: Int, n: BigInt) : Int = x >>> n.intValue def shiftr_signed(x: Int, n: BigInt) : Int = x >> n.intValue def test_bit(x: Int, n: BigInt) : Boolean = (x & (1 << n.intValue)) != 0 } /* object Uint */\ code_reserved Scala Uint text \ OCaml's conversion from Big\_int to int demands that the value fits into a signed integer. The following justifies the implementation. \ context includes integer.lifting begin definition wivs_mask :: int where "wivs_mask = 2^ dflt_size - 1" lift_definition wivs_mask_integer :: integer is wivs_mask . lemma [code]: "wivs_mask_integer = 2 ^ dflt_size - 1" by transfer (simp add: wivs_mask_def) definition wivs_shift :: int where "wivs_shift = 2 ^ dflt_size" lift_definition wivs_shift_integer :: integer is wivs_shift . lemma [code]: "wivs_shift_integer = 2 ^ dflt_size" by transfer (simp add: wivs_shift_def) definition wivs_index :: nat where "wivs_index == dflt_size - 1" lift_definition wivs_index_integer :: integer is "int wivs_index". lemma wivs_index_integer_code[code]: "wivs_index_integer = dflt_size_integer - 1" by transfer (simp add: wivs_index_def of_nat_diff) definition wivs_overflow :: int where "wivs_overflow == 2^ (dflt_size - 1)" lift_definition wivs_overflow_integer :: integer is wivs_overflow . lemma [code]: "wivs_overflow_integer = 2 ^ (dflt_size - 1)" by transfer (simp add: wivs_overflow_def) definition wivs_least :: int where "wivs_least == - wivs_overflow" lift_definition wivs_least_integer :: integer is wivs_least . lemma [code]: "wivs_least_integer = - (2 ^ (dflt_size - 1))" by transfer (simp add: wivs_overflow_def wivs_least_def) definition Uint_signed :: "integer \ uint" where "Uint_signed i = (if i < wivs_least_integer \ wivs_overflow_integer \ i then undefined Uint i else Uint i)" lemma Uint_code [code]: "Uint i = (let i' = i AND wivs_mask_integer in if bit i' wivs_index then Uint_signed (i' - wivs_shift_integer) else Uint_signed i')" including undefined_transfer unfolding Uint_signed_def apply transfer apply (subst word_of_int_via_signed) apply (auto simp add: shiftl_eq_push_bit push_bit_of_1 mask_eq_exp_minus_1 word_of_int_via_signed wivs_mask_def wivs_index_def wivs_overflow_def wivs_least_def wivs_shift_def) done lemma Uint_signed_code [code abstract]: "Rep_uint (Uint_signed i) = (if i < wivs_least_integer \ i \ wivs_overflow_integer then Rep_uint (undefined Uint i) else word_of_int (int_of_integer_symbolic i))" unfolding Uint_signed_def Uint_def int_of_integer_symbolic_def word_of_integer_def by(simp add: Abs_uint_inverse) end text \ Avoid @{term Abs_uint} in generated code, use @{term Rep_uint'} instead. The symbolic implementations for code\_simp use @{term Rep_uint}. The new destructor @{term Rep_uint'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint} ([code abstract] equations for @{typ uint} may use @{term Rep_uint} because these instances will be folded away.) \ definition Rep_uint' where [simp]: "Rep_uint' = Rep_uint" lemma Rep_uint'_code [code]: "Rep_uint' x = (BITS n. bit x n)" unfolding Rep_uint'_def by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint' :: "dflt_size word \ uint" is "\x :: dflt_size word. x" . lemma Abs_uint'_code [code]: "Abs_uint' x = Uint (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint \ _"]] lemma term_of_uint_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint.uint.Abs_uint'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR (STR ''Uint.dflt_size'') []], TR (STR ''Uint.uint'') []])) (term_of_class.term_of (Rep_uint' x))" by(simp add: term_of_anything) text \Important: We must prevent the reflection oracle (eval-tac) to use our machine-dependent type. \ code_printing type_constructor uint \ (SML) "Word.word" and (Haskell) "Uint.Word" and (OCaml) "Uint.t" and (Scala) "Int" and (Eval) "*** \"Error: Machine dependent type\" ***" and (Quickcheck) "Word.word" | constant dflt_size_integer \ (SML) "(IntInf.fromLarge (Int.toLarge Word.wordSize))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.wordSize" and (Haskell) "Uint.dflt'_size" and (OCaml) "Uint.dflt'_size" and (Scala) "Uint.dflt'_size" | constant Uint \ (SML) "Word.fromLargeInt (IntInf.toLarge _)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.fromInt" and (Haskell) "(Prelude.fromInteger _ :: Uint.Word)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint.Word)" and (Scala) "_.intValue" | constant Uint_signed \ (OCaml) "Z.to'_int" | constant "0 :: uint" \ (SML) "(Word.fromInt 0)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "(Word.fromInt 0)" and (Haskell) "(0 :: Uint.Word)" and (OCaml) "0" and (Scala) "0" | constant "1 :: uint" \ (SML) "(Word.fromInt 1)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "(Word.fromInt 1)" and (Haskell) "(1 :: Uint.Word)" and (OCaml) "1" and (Scala) "1" | constant "plus :: uint \ _ " \ (SML) "Word.+ ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.+ ((_), (_))" and (Haskell) infixl 6 "+" and (OCaml) "Pervasives.(+)" and (Scala) infixl 7 "+" | constant "uminus :: uint \ _" \ (SML) "Word.~" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.~" and (Haskell) "negate" and (OCaml) "Pervasives.(~-)" and (Scala) "!(- _)" | constant "minus :: uint \ _" \ (SML) "Word.- ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.- ((_), (_))" and (Haskell) infixl 6 "-" and (OCaml) "Pervasives.(-)" and (Scala) infixl 7 "-" | constant "times :: uint \ _ \ _" \ (SML) "Word.* ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.* ((_), (_))" and (Haskell) infixl 7 "*" and (OCaml) "Pervasives.( * )" and (Scala) infixl 8 "*" | constant "HOL.equal :: uint \ _ \ bool" \ (SML) "!((_ : Word.word) = _)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "!((_ : Word.word) = _)" and (Haskell) infix 4 "==" and (OCaml) "(Pervasives.(=):Uint.t -> Uint.t -> bool)" and (Scala) infixl 5 "==" | class_instance uint :: equal \ (Haskell) - | constant "less_eq :: uint \ _ \ bool" \ (SML) "Word.<= ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.<= ((_), (_))" and (Haskell) infix 4 "<=" and (OCaml) "Uint.less'_eq" and (Scala) "Uint.less'_eq" | constant "less :: uint \ _ \ bool" \ (SML) "Word.< ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.< ((_), (_))" and (Haskell) infix 4 "<" and (OCaml) "Uint.less" and (Scala) "Uint.less" | constant "NOT :: uint \ _" \ (SML) "Word.notb" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.notb" and (Haskell) "Data'_Bits.complement" and (OCaml) "Pervasives.lnot" and (Scala) "_.unary'_~" | constant "(AND) :: uint \ _" \ (SML) "Word.andb ((_),/ (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (OCaml) "Pervasives.(land)" and (Scala) infixl 3 "&" | constant "(OR) :: uint \ _" \ (SML) "Word.orb ((_),/ (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (OCaml) "Pervasives.(lor)" and (Scala) infixl 1 "|" | constant "(XOR) :: uint \ _" \ (SML) "Word.xorb ((_),/ (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.xorb ((_),/ (_))" and (Haskell) "Data'_Bits.xor" and (OCaml) "Pervasives.(lxor)" and (Scala) infixl 2 "^" definition uint_divmod :: "uint \ uint \ uint \ uint" where "uint_divmod x y = (if y = 0 then (undefined ((div) :: uint \ _) x (0 :: uint), undefined ((mod) :: uint \ _) x (0 :: uint)) else (x div y, x mod y))" definition uint_div :: "uint \ uint \ uint" where "uint_div x y = fst (uint_divmod x y)" definition uint_mod :: "uint \ uint \ uint" where "uint_mod x y = snd (uint_divmod x y)" lemma div_uint_code [code]: "x div y = (if y = 0 then 0 else uint_div x y)" including undefined_transfer unfolding uint_divmod_def uint_div_def by transfer(simp add: word_div_def) lemma mod_uint_code [code]: "x mod y = (if y = 0 then x else uint_mod x y)" including undefined_transfer unfolding uint_mod_def uint_divmod_def by transfer(simp add: word_mod_def) definition uint_sdiv :: "uint \ uint \ uint" where [code del]: "uint_sdiv x y = (if y = 0 then undefined ((div) :: uint \ _) x (0 :: uint) else Abs_uint (Rep_uint x sdiv Rep_uint y))" definition div0_uint :: "uint \ uint" where [code del]: "div0_uint x = undefined ((div) :: uint \ _) x (0 :: uint)" declare [[code abort: div0_uint]] definition mod0_uint :: "uint \ uint" where [code del]: "mod0_uint x = undefined ((mod) :: uint \ _) x (0 :: uint)" declare [[code abort: mod0_uint]] definition wivs_overflow_uint :: uint where "wivs_overflow_uint \ push_bit (dflt_size - 1) 1" lemma uint_divmod_code [code]: "uint_divmod x y = (if wivs_overflow_uint \ y then if x < y then (0, x) else (1, x - y) else if y = 0 then (div0_uint x, mod0_uint x) else let q = push_bit 1 (uint_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" proof (cases \y = 0\) case True moreover have \x \ 0\ by transfer simp moreover have \wivs_overflow_uint > 0\ apply (simp add: wivs_overflow_uint_def push_bit_of_1) apply transfer apply transfer apply simp done ultimately show ?thesis by (auto simp add: uint_divmod_def div0_uint_def mod0_uint_def not_less) next case False then show ?thesis including undefined_transfer unfolding uint_divmod_def uint_sdiv_def div0_uint_def mod0_uint_def wivs_overflow_uint_def - apply (simp only: if_simps) apply transfer apply (simp add: divmod_via_sdivmod push_bit_of_1 shiftl_eq_push_bit shiftr_eq_drop_bit) done qed lemma uint_sdiv_code [code abstract]: "Rep_uint (uint_sdiv x y) = (if y = 0 then Rep_uint (undefined ((div) :: uint \ _) x (0 :: uint)) else Rep_uint x sdiv Rep_uint y)" unfolding uint_sdiv_def by(simp add: Abs_uint_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint_divmod_code} computes both with division only. \ code_printing constant uint_div \ (SML) "Word.div ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.div ((_), (_))" and (Haskell) "Prelude.div" | constant uint_mod \ (SML) "Word.mod ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.mod ((_), (_))" and (Haskell) "Prelude.mod" | constant uint_divmod \ (Haskell) "divmod" | constant uint_sdiv \ (OCaml) "Pervasives.('/)" and (Scala) "_ '/ _" definition uint_test_bit :: "uint \ integer \ bool" where [code del]: "uint_test_bit x n = (if n < 0 \ dflt_size_integer \ n then undefined (bit :: uint \ _) x n else bit x (nat_of_integer n))" lemma test_bit_uint_code [code]: "bit x n \ n < dflt_size \ uint_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint_test_bit_def by (transfer, simp, transfer, simp) lemma uint_test_bit_code [code]: "uint_test_bit w n = (if n < 0 \ dflt_size_integer \ n then undefined (bit :: uint \ _) w n else bit (Rep_uint w) (nat_of_integer n))" unfolding uint_test_bit_def by(simp add: bit_uint.rep_eq) code_printing constant uint_test_bit \ (SML) "Uint.test'_bit" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (OCaml) "Uint.test'_bit" and (Scala) "Uint.test'_bit" definition uint_set_bit :: "uint \ integer \ bool \ uint" where [code del]: "uint_set_bit x n b = (if n < 0 \ dflt_size_integer \ n then undefined (set_bit :: uint \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint_code [code]: "set_bit x n b = (if n < dflt_size then uint_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint_set_bit_def by (transfer) (auto cong: conj_cong simp add: not_less set_bit_beyond word_size) lemma uint_set_bit_code [code abstract]: "Rep_uint (uint_set_bit w n b) = (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined (set_bit :: uint \ _) w n b) else set_bit (Rep_uint w) (nat_of_integer n) b)" including undefined_transfer integer.lifting unfolding uint_set_bit_def by transfer simp code_printing constant uint_set_bit \ (SML) "Uint.set'_bit" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (OCaml) "Uint.set'_bit" and (Scala) "Uint.set'_bit" lift_definition uint_set_bits :: "(nat \ bool) \ uint \ nat \ uint" is set_bits_aux . lemma uint_set_bits_code [code]: "uint_set_bits f w n = (if n = 0 then w else let n' = n - 1 in uint_set_bits f (push_bit 1 w OR (if f n' then 1 else 0)) n')" apply (transfer fixing: n) apply (cases n) apply (simp_all add: shiftl_eq_push_bit) done lemma set_bits_uint [code]: "(BITS n. f n) = uint_set_bits f 0 dflt_size" by transfer (simp add: set_bits_conv_set_bits_aux) lemma lsb_code [code]: fixes x :: uint shows "lsb x = bit x 0" by transfer (simp add: lsb_word_eq) definition uint_shiftl :: "uint \ integer \ uint" where [code del]: "uint_shiftl x n = (if n < 0 \ dflt_size_integer \ n then undefined (push_bit :: nat \ uint \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint_code [code]: "push_bit n x = (if n < dflt_size then uint_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint_shiftl_def by (transfer fixing: n) simp lemma uint_shiftl_code [code abstract]: "Rep_uint (uint_shiftl w n) = (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined (push_bit :: nat \ uint \ _) w n) else push_bit (nat_of_integer n) (Rep_uint w))" including undefined_transfer integer.lifting unfolding uint_shiftl_def by transfer simp code_printing constant uint_shiftl \ (SML) "Uint.shiftl" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (OCaml) "Uint.shiftl" and (Scala) "Uint.shiftl" definition uint_shiftr :: "uint \ integer \ uint" where [code del]: "uint_shiftr x n = (if n < 0 \ dflt_size_integer \ n then undefined (drop_bit :: nat \ uint \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint_code [code]: "drop_bit n x = (if n < dflt_size then uint_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint_shiftr_def by (transfer fixing: n) simp lemma uint_shiftr_code [code abstract]: "Rep_uint (uint_shiftr w n) = (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined (drop_bit :: nat \ uint \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint w))" including undefined_transfer unfolding uint_shiftr_def by transfer simp code_printing constant uint_shiftr \ (SML) "Uint.shiftr" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (OCaml) "Uint.shiftr" and (Scala) "Uint.shiftr" definition uint_sshiftr :: "uint \ integer \ uint" where [code del]: "uint_sshiftr x n = (if n < 0 \ dflt_size_integer \ n then undefined sshiftr_uint x n else sshiftr_uint x (nat_of_integer n))" lemma sshiftr_uint_code [code]: "x >>> n = (if n < dflt_size then uint_sshiftr x (integer_of_nat n) else if bit x wivs_index then -1 else 0)" including undefined_transfer integer.lifting unfolding uint_sshiftr_def by transfer(simp add: not_less signed_drop_bit_beyond word_size wivs_index_def) lemma uint_sshiftr_code [code abstract]: "Rep_uint (uint_sshiftr w n) = (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined sshiftr_uint w n) else signed_drop_bit (nat_of_integer n) (Rep_uint w))" including undefined_transfer unfolding uint_sshiftr_def by transfer simp code_printing constant uint_sshiftr \ (SML) "Uint.shiftr'_signed" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint.Int) _)) :: Uint.Word)" and (OCaml) "Uint.shiftr'_signed" and (Scala) "Uint.shiftr'_signed" lemma uint_msb_test_bit: "msb x \ bit (x :: uint) wivs_index" by transfer (simp add: msb_word_iff_bit wivs_index_def) lemma msb_uint_code [code]: "msb x \ uint_test_bit x wivs_index_integer" apply(simp add: uint_test_bit_def uint_msb_test_bit wivs_index_integer_code dflt_size_integer_def wivs_index_def) by (metis (full_types) One_nat_def dflt_size(2) less_iff_diff_less_0 nat_of_integer_of_nat of_nat_1 of_nat_diff of_nat_less_0_iff wivs_index_def) lemma uint_of_int_code [code]: "uint_of_int i = (BITS n. bit i n)" by transfer (simp add: word_of_int_conv_set_bits) section \Quickcheck setup\ definition uint_of_natural :: "natural \ uint" where "uint_of_natural x \ Uint (integer_of_natural x)" instantiation uint :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint \ qc_random_cnv uint_of_natural" definition "exhaustive_uint \ qc_exhaustive_cnv uint_of_natural" definition "full_exhaustive_uint \ qc_full_exhaustive_cnv uint_of_natural" instance .. end instantiation uint :: narrowing begin interpretation quickcheck_narrowing_samples "\i. (Uint i, Uint (- i))" "0" "Typerep.Typerep (STR ''Uint.uint'') []" . definition "narrowing_uint d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint itself \ _"]] lemmas partial_term_of_uint [code] = partial_term_of_code instance .. end no_notation sshiftr_uint (infixl ">>>" 55) end diff --git a/thys/VerifyThis2018/Challenge1.thy b/thys/VerifyThis2018/Challenge1.thy --- a/thys/VerifyThis2018/Challenge1.thy +++ b/thys/VerifyThis2018/Challenge1.thy @@ -1,409 +1,408 @@ chapter \Gap Buffer\ section \Challenge\ text_raw \{\upshape A gap buffer is a data structure for the implementation of text editors, which can efficiently move the cursor, as well add and delete characters. The idea is simple: the editor's content is represented as a character array~$a$ of length~$n$, which has a gap of unused entries $a[l], \ldots, a[r-1]$, with respect to two indices $l \le r$. The data it represents is composed as $a[0], \ldots, a[l-1], a[r], ..., a[n-1]$. The current cursor position is at the left index $l$, and if we type a character, it is written to $a[l]$ and $l$ is increased. When the gap becomes empty, the array is enlarged and the data from $r$ is shifted to the right. \paragraph{Implementation task.} Implement the following four operations in the language of your tool: Procedures \verb|left()| and \verb|right()| move the cursor by one character; \verb|insert()| places a character at the beginning of the gap $a[l]$; \verb|delete()| removes the character at $a[l]$ from the range of text. \begin{multicols}{2} \begin{lstlisting}[language=C,morekeywords={procedure,function,end,to,in,var,then,not,mod}] procedure left() if l != 0 then l := l - 1 r := r - 1 a[r] := a[l] end-if end-procedure procedure right() // your task: similar to left() // but pay attention to the // order of statements end-procedure \end{lstlisting} \begin{lstlisting}[language=C,morekeywords={procedure,function,end,to,in,var,then,not,mod}] procedure insert(x: char) if l == r then // see extended task grow() end-if a[l] := x l := l + 1 end-procedure procedure delete() if l != 0 then l := l - 1 end-if end-procedure \end{lstlisting} \end{multicols} \paragraph{Verification task.} Specify the intended behavior of the buffer in terms of a contiguous representation of the editor content. This can for example be based on strings, functional arrays, sequences, or lists. Verify that the gap buffer implementation satisfies this specification, and that every access to the array is within bounds. \emph{Hint:} For this task you may assume that \verb|insert()| has the precondition $l < r$ and remove the call to \verb|grow()|. Alternatively, assume a contract for \verb|grow()| that ensures that this call does not change the abstract representation. \paragraph{Extended verification task.} Implement the operation \verb|grow()|, specify its behavior in a way that lets you verify \verb|insert()| in a modular way (i.e. not by referring to the implementation of \verb|grow()|), and verify that \verb|grow()| satisfies this specification. \emph{Hint}: You may assume that the allocation of the new buffer always succeeds. If~your tool/language supports copying array ranges (such as \verb|System.arraycopy()| in Java), consider using these primitives instead of the loops in the pseudo-code below. \begin{lstlisting}[language=C,morekeywords={procedure,function,end,to,in,new,var,then,not,mod}] procedure grow() var b := new char[a.length + K] // b[0..l] := a[0..l] for i = 0 to l - 1 do b[i] := a[i] end-for // b[r + K..] := a[r..] for i = r to a.length - 1 do b[i + K] := a[i] end-for r := r + K a := b end-procedure \end{lstlisting} \paragraph{Resources} \begin{itemize} \item \url{https://en.wikipedia.org/wiki/Gap_buffer} \item \url{http://scienceblogs.com/goodmath/2009/02/18/gap-buffers-or-why-bother-with-1} \end{itemize} } \clearpage \ section \Solution\ theory Challenge1 imports "lib/VTcomp" begin text \Fully fledged specification of textbuffer ADT, and its implementation by a gap buffer. \ subsection \Abstract Specification\ text \ Initially, we modelled the abstract text as a cursor position and a list. However, this gives you an invariant on the abstract level. An isomorphic but invariant free formulation is a pair of lists, representing the text before and after the cursor. \ datatype 'a textbuffer = BUF "'a list" "'a list" text \The primitive operations are the empty textbuffer, and to extract the text and the cursor position\ definition empty :: "'a textbuffer" where "empty = BUF [] []" primrec get_text :: "'a textbuffer \ 'a list" where "get_text (BUF a b) = a@b" primrec get_pos :: "'a textbuffer \ nat" where "get_pos (BUF a b) = length a" text \These are the operations that were specified in the challenge\ primrec move_left :: "'a textbuffer \ 'a textbuffer" where "move_left (BUF a b) = (if a\[] then BUF (butlast a) (last a#b) else BUF a b)" primrec move_right :: "'a textbuffer \ 'a textbuffer" where "move_right (BUF a b) = (if b\[] then BUF (a@[hd b]) (tl b) else BUF a b)" primrec insert :: "'a \ 'a textbuffer \ 'a textbuffer" where "insert x (BUF a b) = BUF (a@[x]) b" primrec delete :: "'a textbuffer \ 'a textbuffer" where "delete (BUF a b) = BUF (butlast a) b" \ \Note that @{lemma \butlast [] = []\ by simp} in Isabelle\ text \We can also assign them a meaning wrt position and text\ lemma empty_pos[simp]: "get_pos empty = 0" unfolding empty_def by auto lemma empty_text[simp]: "get_text empty = []" unfolding empty_def by auto lemma move_left_pos[simp]: "get_pos (move_left b) = get_pos b - 1" \ \Note that @{lemma \0-(1::nat)=0\ by simp} in Isabelle\ by (cases b) auto lemma move_left_text[simp]: "get_text (move_left b) = get_text b" by (cases b) auto lemma move_right_pos[simp]: "get_pos (move_right b) = min (get_pos b+1) (length (get_text b))" by (cases b) auto lemma move_right_text[simp]: "get_text (move_right b) = get_text b" by (cases b) auto lemma insert_pos[simp]: "get_pos (insert x b) = get_pos b + 1" by (cases b) auto lemma insert_text: "get_text (insert x b) = take (get_pos b) (get_text b)@x#drop (get_pos b) (get_text b)" by (cases b) auto lemma delete_pos[simp]: "get_pos (delete b) = get_pos b - 1" by (cases b) auto lemma delete_text: "get_text (delete b) = take (get_pos b-1) (get_text b)@drop (get_pos b) (get_text b)" by (cases b) auto text \For the zero case, we can prove a simpler (equivalent) lemma\ lemma delete_text0[simp]: "get_pos b=0 \ get_text (delete b) = get_text b" by (cases b) auto text \To fully exploit the capabilities of our tool, we can (optionally) show that the operations of a text buffer are parametric in its content. Then, we can automatically refine the representation of the content. \ definition [to_relAPP]: "textbuffer_rel A \ {(BUF a b, BUF a' b') | a b a' b'. (a,a')\\A\list_rel \ (b,b')\\A\list_rel}" lemma [param]: "(BUF,BUF) \ \A\list_rel \ \A\list_rel \ \A\textbuffer_rel" by (auto simp: textbuffer_rel_def) lemma [param]: "(rec_textbuffer,rec_textbuffer) \ (\A\list_rel \ \A\list_rel\B) \ \A\textbuffer_rel \ B" by (auto simp: textbuffer_rel_def) parametricity context notes[simp] = empty_def get_text_def get_pos_def move_left_def move_right_def insert_def delete_def conv_to_is_Nil begin sepref_decl_op (no_def) empty :: "\A\textbuffer_rel" . sepref_decl_op (no_def) get_text :: "\A\textbuffer_rel \ \A\list_rel" . sepref_decl_op (no_def) get_pos :: "\A\textbuffer_rel \ nat_rel" . sepref_decl_op (no_def) move_left :: "\A\textbuffer_rel \ \A\textbuffer_rel" . sepref_decl_op (no_def) move_right :: "\A\textbuffer_rel \ \A\textbuffer_rel" . sepref_decl_op (no_def) insert :: "A\\A\textbuffer_rel \ \A\textbuffer_rel" . sepref_decl_op (no_def) delete :: "\A\textbuffer_rel \ \A\textbuffer_rel" . end subsection \Refinement 1: List with Gap\ subsection \Implementation on List-Level\ type_synonym 'a gap_buffer = "nat \ nat \ 'a list" subsubsection \Abstraction Relation\ text \Also called coupling relation sometimes. Can be any relation, here we define it by an invariant and an abstraction function. \ definition "gap_\ \ \(l,r,buf). BUF (take l buf) (drop r buf)" definition "gap_invar \ \(l,r,buf). l\r \ r\length buf" abbreviation "gap_rel \ br gap_\ gap_invar" subsubsection \Empty\ definition "empty1 \ RETURN (0,0,[])" lemma empty1_correct: "(empty1, RETURN empty) \ \gap_rel\nres_rel" unfolding empty1_def empty_def apply refine_vcg by (auto simp: in_br_conv gap_\_def gap_invar_def) subsubsection \Left\ definition "move_left1 \ \(l,r,buf). doN { if l\0 then doN { ASSERT(r-1 l-1 gap_rel \ \gap_rel\nres_rel" apply clarsimp unfolding move_left1_def apply refine_vcg apply (auto simp: in_br_conv gap_\_def gap_invar_def move_left1_def split: prod.splits) subgoal by (simp add: butlast_take) subgoal by (smt Cons_nth_drop_Suc One_nat_def Suc_pred diff_Suc_less drop_update_cancel last_take_nth_conv le_trans length_list_update less_le_trans neq0_conv nth_list_update_eq) done subsubsection \Right\ definition "move_right1 \ \(l,r,buf). doN { if r gap_rel \ \gap_rel\nres_rel" apply clarsimp unfolding move_right1_def apply refine_vcg unfolding gap_\_def gap_invar_def apply (auto simp: in_br_conv hd_drop_conv_nth take_update_last split: prod.split) by (simp add: drop_Suc tl_drop) subsubsection \Insert and Grow\ definition "can_insert \ \(l,r,buf). l \(l,r,buf). doN { let b = op_array_replicate (length buf + K) default; b \ mop_list_blit buf 0 b 0 l; b \ mop_list_blit buf r b (r+K) (length buf - r); RETURN (l,r+K,b) }" lemma grow1_correct[THEN SPEC_trans, refine_vcg]: assumes "gap_invar gb" shows "grow1 K gb \ (SPEC (\gb'. gap_invar gb' \ gap_\ gb' = gap_\ gb \ (K>0 \ can_insert gb')))" unfolding grow1_def apply refine_vcg using assms unfolding gap_\_def gap_invar_def can_insert_def - apply clarsimp_all apply (auto simp: op_list_blit_def) - by (simp add: min_def) + done definition "insert1 x \ \(l,r,buf). doN { (l,r,buf) \ if (l=r) then grow1 (length buf+1) (l,r,buf) else RETURN (l,r,buf); ASSERT (l Id \ gap_rel \ \gap_rel\nres_rel" apply clarsimp unfolding insert1_def apply refine_vcg unfolding insert_def gap_\_def gap_invar_def can_insert_def apply (auto simp: in_br_conv take_update_last split: prod.split) done subsubsection \Delete\ definition "delete1 \ \(l,r,buf). if l>0 then RETURN (l-1,r,buf) else RETURN (l,r,buf)" lemma delete1_correct: "(delete1,RETURN o delete) \ gap_rel \ \gap_rel\nres_rel" apply clarsimp unfolding delete1_def apply refine_vcg unfolding gap_\_def gap_invar_def by (auto simp: in_br_conv butlast_take split: prod.split) subsection \Imperative Arrays and Executable Code\ abbreviation "gap_impl_assn \ nat_assn \\<^sub>a nat_assn \\<^sub>a array_assn id_assn" definition "gap_assn A \ hr_comp (hr_comp gap_impl_assn gap_rel) (\the_pure A\textbuffer_rel)" context notes gap_assn_def[symmetric,fcomp_norm_unfold] begin sepref_definition empty_impl is "uncurry0 empty1" :: "unit_assn\<^sup>k\\<^sub>agap_impl_assn" unfolding empty1_def array.fold_custom_empty by sepref sepref_decl_impl empty_impl: empty_impl.refine[FCOMP empty1_correct] . sepref_definition move_left_impl is move_left1 :: "gap_impl_assn\<^sup>d\\<^sub>agap_impl_assn" unfolding move_left1_def by sepref sepref_decl_impl move_left_impl: move_left_impl.refine[FCOMP move_left1_correct] . sepref_definition move_right_impl is move_right1 :: "gap_impl_assn\<^sup>d\\<^sub>agap_impl_assn" unfolding move_right1_def by sepref sepref_decl_impl move_right_impl: move_right_impl.refine[FCOMP move_right1_correct] . sepref_definition insert_impl is "uncurry insert1" :: "id_assn\<^sup>k*\<^sub>agap_impl_assn\<^sup>d\\<^sub>agap_impl_assn" unfolding insert1_def grow1_def by sepref \ \We inline @{const grow1} here\ sepref_decl_impl insert_impl: insert_impl.refine[FCOMP insert1_correct] . sepref_definition delete_impl is delete1 :: "gap_impl_assn\<^sup>d\\<^sub>agap_impl_assn" unfolding delete1_def by sepref sepref_decl_impl delete_impl: delete_impl.refine[FCOMP delete1_correct] . end text \ The above setup generated the following refinement theorems, connecting the implementations with our abstract specification: @{thm [display] empty_impl_hnr move_left_impl_hnr move_right_impl_hnr insert_impl_hnr delete_impl_hnr } \ export_code move_left_impl move_right_impl insert_impl delete_impl in SML_imp module_name Gap_Buffer in OCaml_imp module_name Gap_Buffer in Haskell module_name Gap_Buffer in Scala module_name Gap_Buffer subsection \Simple Client\ definition "client \ RETURN (fold (\f. f) [ insert (1::int), insert (2::int), insert (3::int), insert (5::int), move_left, insert (4::int), move_right, insert (6::int), delete ] empty)" lemma "client \ SPEC (\r. get_text r=[1,2,3,4,5])" unfolding client_def by (simp add: delete_text insert_text) sepref_definition client_impl is "uncurry0 client" :: "unit_assn\<^sup>k \\<^sub>a gap_assn id_assn" unfolding client_def fold.simps id_def comp_def by sepref ML_val \ @{code client_impl} () \ end diff --git a/thys/VerifyThis2018/Challenge1_short.thy b/thys/VerifyThis2018/Challenge1_short.thy --- a/thys/VerifyThis2018/Challenge1_short.thy +++ b/thys/VerifyThis2018/Challenge1_short.thy @@ -1,199 +1,198 @@ section \Shorter Solution\ theory Challenge1_short imports "lib/VTcomp" begin text \Small specification of textbuffer ADT, and its implementation by a gap buffer. Annotated and elaborated version of just the challenge requirements. \ subsection \Abstract Specification\ datatype 'a textbuffer = BUF ("pos": "nat") ("text": "'a list") \ \Note that we do not model the abstract invariant --- pos in range --- here, as it is not strictly required for the challenge spec.\ text \These are the operations that were specified in the challenge. Note: Isabelle has type inference, so we do not need to specify types. Note: We exploit that, in Isabelle, we have @{lemma "(0::nat)-1=0" by simp}. \ primrec move_left where "move_left (BUF p t) = BUF (p-1) t" primrec move_right where "move_right (BUF p t) = BUF (min (length t) (p+1)) t" primrec insert where "insert x (BUF p t) = BUF (p+1) (take p t@x#drop p t)" primrec delete where "delete (BUF p t) = BUF (p-1) (take (p-1) t@drop p t)" subsection \Refinement 1: List with Gap\ subsection \Implementation on List-Level\ type_synonym 'a gap_buffer = "nat \ nat \ 'a list" subsubsection \Abstraction Relation\ text \We define an invariant on the concrete gap-buffer, and its mapping to the abstract model. From these two, we define a relation \gap_rel\ between concrete and abstract buffers. \ definition "gap_\ \ \(l,r,buf). BUF l (take l buf @ drop r buf)" definition "gap_invar \ \(l,r,buf). l\r \ r\length buf" abbreviation "gap_rel \ br gap_\ gap_invar" subsubsection \Left\ text \For the operations, we insert assertions. These are not required to prove the list-level specification correct (during the proof, they are inferred easily). However, they are required in the subsequent automatic refinement step to arrays, to give our tool the information that all indexes are, indeed, in bounds.\ definition "move_left1 \ \(l,r,buf). doN { if l\0 then doN { ASSERT(r-1 l-1 gap_rel \ \gap_rel\nres_rel" apply clarsimp unfolding move_left1_def apply refine_vcg apply (auto simp: in_br_conv gap_\_def gap_invar_def move_left1_def split: prod.splits) (* sledgehammer! *) by (smt Cons_nth_drop_Suc Suc_pred append.assoc append_Cons append_Nil diff_Suc_less drop_update_cancel hd_drop_conv_nth length_list_update less_le_trans nth_list_update_eq take_hd_drop) subsubsection \Right\ definition "move_right1 \ \(l,r,buf). doN { if r gap_rel \ \gap_rel\nres_rel" apply clarsimp unfolding move_right1_def apply refine_vcg unfolding gap_\_def gap_invar_def apply (auto simp: in_br_conv split: prod.split) (* sledgehammer *) by (metis Cons_eq_appendI Cons_nth_drop_Suc append_eq_append_conv2 atd_lem drop_0 dual_order.strict_trans2 take_eq_Nil take_update_last) (* Manual: by (simp add: hd_drop_conv_nth take_update_last Cons_nth_drop_Suc) *) subsubsection \Insert and Grow\ definition "can_insert \ \(l,r,buf). l \(l,r,buf). doN { let b = op_array_replicate (length buf + K) default; b \ mop_list_blit buf 0 b 0 l; b \ mop_list_blit buf r b (r+K) (length buf - r); RETURN (l,r+K,b) }" \ \Note: Most operations have also a variant prefixed with \mop\. These are defined in the refinement monad and already contain the assertion of their precondition. The backside is that they cannot be easily used in as part of expressions, e.g., in @{term "buf[l:=buf!r]"}, we would have to explicitly bind each intermediate value: @{term "doN { v \ mop_list_get buf r; mop_list_set buf l v }"}. \ lemma grow1_correct[THEN SPEC_trans, refine_vcg]: \ \Declares this as a rule to be used by the VCG\ assumes "gap_invar gb" shows "grow1 K gb \ (SPEC (\gb'. gap_invar gb' \ gap_\ gb' = gap_\ gb \ (K>0 \ can_insert gb')))" unfolding grow1_def apply refine_vcg using assms unfolding gap_\_def gap_invar_def can_insert_def - apply clarsimp_all apply (auto simp: op_list_blit_def) - by (simp add: min_def) + done definition "insert1 x \ \(l,r,buf). doN { (l,r,buf) \ if (l=r) then grow1 (length buf+1) (l,r,buf) else RETURN (l,r,buf); ASSERT (l Id \ gap_rel \ \gap_rel\nres_rel" apply clarsimp unfolding insert1_def apply refine_vcg \ \VCG knows the rule for grow1 already\ unfolding insert_def gap_\_def gap_invar_def can_insert_def apply (auto simp: in_br_conv take_update_last split: prod.split) done subsubsection \Delete\ definition "delete1 \ \(l,r,buf). if l>0 then RETURN (l-1,r,buf) else RETURN (l,r,buf)" lemma delete1_correct: "(delete1,RETURN o delete) \ gap_rel \ \gap_rel\nres_rel" apply clarsimp unfolding delete1_def apply refine_vcg unfolding gap_\_def gap_invar_def by (auto simp: in_br_conv butlast_take split: prod.split) subsection \Imperative Arrays\ text \The following indicates how we will further refine the gap-buffer: The list will become an array, the indices and the content will not be refined (expressed by @{const nat_assn} and @{const id_assn}). \ abbreviation "gap_impl_assn \ nat_assn \\<^sub>a nat_assn \\<^sub>a array_assn id_assn" sepref_definition move_left_impl is move_left1 :: "gap_impl_assn\<^sup>d\\<^sub>agap_impl_assn" unfolding move_left1_def by sepref sepref_definition move_right_impl is move_right1 :: "gap_impl_assn\<^sup>d\\<^sub>agap_impl_assn" unfolding move_right1_def by sepref sepref_definition insert_impl is "uncurry insert1" :: "id_assn\<^sup>k*\<^sub>agap_impl_assn\<^sup>d\\<^sub>agap_impl_assn" unfolding insert1_def grow1_def by sepref \ \We inline @{const grow1} here\ sepref_definition delete_impl is delete1 :: "gap_impl_assn\<^sup>d\\<^sub>agap_impl_assn" unfolding delete1_def by sepref text \Finally, we combine the two refinement steps, to get overall correctness theorems\ definition "gap_assn \ hr_comp gap_impl_assn gap_rel" \ \@{const hr_comp} is composition of refinement relations\ context notes gap_assn_def[symmetric,fcomp_norm_unfold] begin lemmas move_left_impl_correct = move_left_impl.refine[FCOMP move_left1_correct] and move_right_impl_correct = move_right_impl.refine[FCOMP move_right1_correct] and insert_impl_correct = insert_impl.refine[FCOMP insert1_correct] and delete_impl_correct = delete_impl.refine[FCOMP delete1_correct] text \Proves: @{thm [display] move_left_impl_correct} @{thm [display] move_right_impl_correct} @{thm [display] insert_impl_correct} @{thm [display] delete_impl_correct} \ end subsection \Executable Code\ text \Isabelle/HOL can generate code in various target languages.\ export_code move_left_impl move_right_impl insert_impl delete_impl in SML_imp module_name Gap_Buffer in OCaml_imp module_name Gap_Buffer in Haskell module_name Gap_Buffer in Scala module_name Gap_Buffer end diff --git a/thys/Word_Lib/Aligned.thy b/thys/Word_Lib/Aligned.thy --- a/thys/Word_Lib/Aligned.thy +++ b/thys/Word_Lib/Aligned.thy @@ -1,912 +1,907 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Word Alignment" theory Aligned imports "HOL-Library.Word" Word_Lib - More_Misc + More_Divides begin lift_definition is_aligned :: \'a::len word \ nat \ bool\ is \\k n. 2 ^ n dvd take_bit LENGTH('a) k\ by simp lemma is_aligned_iff_udvd: \is_aligned w n \ 2 ^ n udvd w\ by transfer (simp flip: take_bit_eq_0_iff add: min_def) lemma is_aligned_iff_take_bit_eq_0: \is_aligned w n \ take_bit n w = 0\ by (simp add: is_aligned_iff_udvd take_bit_eq_0_iff exp_dvd_iff_exp_udvd) lemma is_aligned_iff_dvd_int: \is_aligned ptr n \ 2 ^ n dvd uint ptr\ by transfer simp lemma is_aligned_iff_dvd_nat: \is_aligned ptr n \ 2 ^ n dvd unat ptr\ proof - have \unat ptr = nat \uint ptr\\ by transfer simp then have \2 ^ n dvd unat ptr \ 2 ^ n dvd uint ptr\ by (simp only: dvd_nat_abs_iff) simp then show ?thesis by (simp add: is_aligned_iff_dvd_int) qed lemma is_aligned_0 [simp]: \is_aligned 0 n\ by transfer simp lemma is_aligned_at_0 [simp]: \is_aligned w 0\ by transfer simp lemma is_aligned_beyond_length: \is_aligned w n \ w = 0\ if \LENGTH('a) \ n\ for w :: \'a::len word\ using that apply (simp add: is_aligned_iff_udvd) apply transfer apply auto done lemma is_alignedI [intro?]: \is_aligned x n\ if \x = 2 ^ n * k\ for x :: \'a::len word\ proof (unfold is_aligned_iff_udvd) from that show \2 ^ n udvd x\ using dvd_triv_left exp_dvd_iff_exp_udvd by blast qed lemma is_alignedE [elim?]: fixes w :: \'a::len word\ assumes \is_aligned w n\ obtains q where \w = 2 ^ n * word_of_nat q\ \q < 2 ^ (LENGTH('a) - n)\ proof (cases \n < LENGTH('a)\) case False with assms have \w = 0\ by (simp add: is_aligned_beyond_length) with that [of 0] show thesis by simp next case True moreover define m where \m = LENGTH('a) - n\ ultimately have l: \LENGTH('a) = n + m\ and \m \ 0\ by simp_all from \n < LENGTH('a)\ have *: \unat (2 ^ n :: 'a word) = 2 ^ n\ by transfer simp from assms have \2 ^ n udvd w\ by (simp add: is_aligned_iff_udvd) then obtain v :: \'a word\ where \unat w = unat (2 ^ n :: 'a word) * unat v\ .. moreover define q where \q = unat v\ ultimately have unat_w: \unat w = 2 ^ n * q\ by (simp add: *) then have \word_of_nat (unat w) = (word_of_nat (2 ^ n * q) :: 'a word)\ by simp then have w: \w = 2 ^ n * word_of_nat q\ by simp moreover have \q < 2 ^ (LENGTH('a) - n)\ proof (rule ccontr) assume \\ q < 2 ^ (LENGTH('a) - n)\ then have \2 ^ (LENGTH('a) - n) \ q\ by simp then have \2 ^ LENGTH('a) \ 2 ^ n * q\ by (simp add: l power_add) with unat_w [symmetric] show False by (metis le_antisym nat_less_le unsigned_less) qed ultimately show thesis using that by blast qed lemma is_aligned_mask: \is_aligned w n \ w && mask n = 0\ by (simp add: is_aligned_iff_take_bit_eq_0 take_bit_eq_mask) lemma is_aligned_weaken: "\ is_aligned w x; x \ y \ \ is_aligned w y" unfolding is_aligned_iff_dvd_nat by (erule dvd_trans [rotated]) (simp add: le_imp_power_dvd) lemma nat_power_less_diff: assumes lt: "(2::nat) ^ n * q < 2 ^ m" shows "q < 2 ^ (m - n)" using lt proof (induct n arbitrary: m) case 0 then show ?case by simp next case (Suc n) have ih: "\m. 2 ^ n * q < 2 ^ m \ q < 2 ^ (m - n)" and prem: "2 ^ Suc n * q < 2 ^ m" by fact+ show ?case proof (cases m) case 0 then show ?thesis using Suc by simp next case (Suc m') then show ?thesis using prem by (simp add: ac_simps ih) qed qed lemma is_alignedE_pre: fixes w::"'a::len word" assumes aligned: "is_aligned w n" shows rl: "\q. w = 2 ^ n * (of_nat q) \ q < 2 ^ (LENGTH('a) - n)" using aligned is_alignedE by blast lemma is_aligned_to_bl: "is_aligned (w :: 'a :: len word) n = (True \ set (drop (size w - n) (to_bl w)))" apply (simp add: is_aligned_mask eq_zero_set_bl) apply (clarsimp simp: in_set_conv_nth word_size) apply (simp add: to_bl_nth word_size cong: conj_cong) apply (simp add: diff_diff_less) apply safe apply (case_tac "n \ LENGTH('a)") prefer 2 apply (rule_tac x=i in exI) apply clarsimp apply (subgoal_tac "\j < LENGTH('a). j < n \ LENGTH('a) - n + j = i") apply (erule exE) apply (rule_tac x=j in exI) apply clarsimp apply (thin_tac "w !! t" for t) apply (rule_tac x="i + n - LENGTH('a)" in exI) apply clarsimp apply arith apply (rule_tac x="LENGTH('a) - n + i" in exI) apply clarsimp apply arith done -lemma unat_power_lower [simp]: - assumes nv: "n < LENGTH('a::len)" - shows "unat ((2::'a::len word) ^ n) = 2 ^ n" - using assms by transfer simp - lemma power_overflow: "n \ LENGTH('a) \ 2 ^ n = (0 :: 'a::len word)" by simp lemma is_aligned_replicate: fixes w::"'a::len word" assumes aligned: "is_aligned w n" and nv: "n \ LENGTH('a)" shows "to_bl w = (take (LENGTH('a) - n) (to_bl w)) @ replicate n False" proof - from nv have rl: "\q. q < 2 ^ (LENGTH('a) - n) \ to_bl (2 ^ n * (of_nat q :: 'a word)) = drop n (to_bl (of_nat q :: 'a word)) @ replicate n False" by (metis bl_shiftl le_antisym min_def shiftl_t2n wsst_TYs(3)) show ?thesis using aligned by (auto simp: rl elim: is_alignedE) qed lemma is_aligned_drop: fixes w::"'a::len word" assumes "is_aligned w n" "n \ LENGTH('a)" shows "drop (LENGTH('a) - n) (to_bl w) = replicate n False" proof - have "to_bl w = take (LENGTH('a) - n) (to_bl w) @ replicate n False" by (rule is_aligned_replicate) fact+ then have "drop (LENGTH('a) - n) (to_bl w) = drop (LENGTH('a) - n) \" by simp also have "\ = replicate n False" by simp finally show ?thesis . qed lemma less_is_drop_replicate: fixes x::"'a::len word" assumes lt: "x < 2 ^ n" shows "to_bl x = replicate (LENGTH('a) - n) False @ drop (LENGTH('a) - n) (to_bl x)" by (metis assms bl_and_mask' less_mask_eq) lemma is_aligned_add_conv: fixes off::"'a::len word" assumes aligned: "is_aligned w n" and offv: "off < 2 ^ n" shows "to_bl (w + off) = (take (LENGTH('a) - n) (to_bl w)) @ (drop (LENGTH('a) - n) (to_bl off))" proof cases assume nv: "n \ LENGTH('a)" show ?thesis proof (subst aligned_bl_add_size, simp_all only: word_size) show "drop (LENGTH('a) - n) (to_bl w) = replicate n False" by (subst is_aligned_replicate [OF aligned nv]) (simp add: word_size) from offv show "take (LENGTH('a) - n) (to_bl off) = replicate (LENGTH('a) - n) False" by (subst less_is_drop_replicate, assumption) simp qed fact next assume "\ n \ LENGTH('a)" with offv show ?thesis by (simp add: power_overflow) qed lemma aligned_add_aligned: fixes x::"'a::len word" assumes aligned1: "is_aligned x n" and aligned2: "is_aligned y m" and lt: "m \ n" shows "is_aligned (x + y) m" proof cases assume nlt: "n < LENGTH('a)" show ?thesis unfolding is_aligned_iff_dvd_nat dvd_def proof - from aligned2 obtain q2 where yv: "y = 2 ^ m * of_nat q2" and q2v: "q2 < 2 ^ (LENGTH('a) - m)" by (auto elim: is_alignedE) from lt obtain k where kv: "m + k = n" by (auto simp: le_iff_add) with aligned1 obtain q1 where xv: "x = 2 ^ (m + k) * of_nat q1" and q1v: "q1 < 2 ^ (LENGTH('a) - (m + k))" by (auto elim: is_alignedE) have l1: "2 ^ (m + k) * q1 < 2 ^ LENGTH('a)" by (rule nat_less_power_trans [OF q1v]) (subst kv, rule order_less_imp_le [OF nlt]) have l2: "2 ^ m * q2 < 2 ^ LENGTH('a)" by (rule nat_less_power_trans [OF q2v], rule order_less_imp_le [OF order_le_less_trans]) fact+ have "x = of_nat (2 ^ (m + k) * q1)" using xv by simp moreover have "y = of_nat (2 ^ m * q2)" using yv by simp ultimately have upls: "unat x + unat y = 2 ^ m * (2 ^ k * q1 + q2)" proof - have f1: "unat x = 2 ^ (m + k) * q1" by (metis (no_types) \x = of_nat (2 ^ (m + k) * q1)\ l1 nat_mod_lem word_unat.inverse_norm zero_less_numeral zero_less_power) have "unat y = 2 ^ m * q2" by (metis (no_types) \y = of_nat (2 ^ m * q2)\ l2 nat_mod_lem word_unat.inverse_norm zero_less_numeral zero_less_power) then show ?thesis using f1 by (simp add: power_add semiring_normalization_rules(34)) qed (* (2 ^ k * q1 + q2) *) show "\d. unat (x + y) = 2 ^ m * d" proof (cases "unat x + unat y < 2 ^ LENGTH('a)") case True have "unat (x + y) = unat x + unat y" by (subst unat_plus_if', rule if_P) fact also have "\ = 2 ^ m * (2 ^ k * q1 + q2)" by (rule upls) finally show ?thesis .. next case False then have "unat (x + y) = (unat x + unat y) mod 2 ^ LENGTH('a)" by (subst unat_word_ariths(1)) simp also have "\ = (2 ^ m * (2 ^ k * q1 + q2)) mod 2 ^ LENGTH('a)" by (subst upls, rule refl) also have "\ = 2 ^ m * ((2 ^ k * q1 + q2) mod 2 ^ (LENGTH('a) - m))" proof - have "m \ len_of (TYPE('a))" by (meson le_trans less_imp_le_nat lt nlt) then show ?thesis by (metis mult_mod_right ordered_cancel_comm_monoid_diff_class.add_diff_inverse power_add) qed finally show ?thesis .. qed qed next assume "\ n < LENGTH('a)" with assms show ?thesis by (simp add: is_aligned_mask not_less take_bit_eq_mod power_overflow word_arith_nat_defs(7) flip: take_bit_eq_mask) qed corollary aligned_sub_aligned: "\is_aligned (x::'a::len word) n; is_aligned y m; m \ n\ \ is_aligned (x - y) m" apply (simp del: add_uminus_conv_diff add:diff_conv_add_uminus) apply (erule aligned_add_aligned, simp_all) apply (erule is_alignedE) apply (rule_tac k="- of_nat q" in is_alignedI) apply simp done lemma is_aligned_shift: fixes k::"'a::len word" shows "is_aligned (k << m) m" proof cases assume mv: "m < LENGTH('a)" from mv obtain q where mq: "m + q = LENGTH('a)" and "0 < q" by (auto dest: less_imp_add_positive) have "(2::nat) ^ m dvd unat (k << m)" proof have kv: "(unat k div 2 ^ q) * 2 ^ q + unat k mod 2 ^ q = unat k" by (rule div_mult_mod_eq) have "unat (k << m) = unat (2 ^ m * k)" by (simp add: shiftl_t2n) also have "\ = (2 ^ m * unat k) mod (2 ^ LENGTH('a))" using mv by (subst unat_word_ariths(2))+ simp also have "\ = 2 ^ m * (unat k mod 2 ^ q)" by (subst mq [symmetric], subst power_add, subst mod_mult2_eq) simp finally show "unat (k << m) = 2 ^ m * (unat k mod 2 ^ q)" . qed then show ?thesis by (unfold is_aligned_iff_dvd_nat) next assume "\ m < LENGTH('a)" then show ?thesis by (simp add: not_less power_overflow is_aligned_mask shiftl_zero_size word_size) qed lemma word_mod_by_0: "k mod (0::'a::len word) = k" by (simp add: word_arith_nat_mod) lemma aligned_mod_eq_0: fixes p::"'a::len word" assumes al: "is_aligned p sz" shows "p mod 2 ^ sz = 0" proof cases assume szv: "sz < LENGTH('a)" with al show ?thesis unfolding is_aligned_iff_dvd_nat by (simp add: and_mask_dvd_nat p2_gt_0 word_mod_2p_is_mask) next assume "\ sz < LENGTH('a)" with al show ?thesis by (simp add: is_aligned_mask flip: take_bit_eq_mask take_bit_eq_mod) qed lemma is_aligned_triv: "is_aligned (2 ^ n ::'a::len word) n" by (rule is_alignedI [where k = 1], simp) lemma is_aligned_mult_triv1: "is_aligned (2 ^ n * x ::'a::len word) n" by (rule is_alignedI [OF refl]) lemma is_aligned_mult_triv2: "is_aligned (x * 2 ^ n ::'a::len word) n" by (subst mult.commute, simp add: is_aligned_mult_triv1) lemma word_power_less_0_is_0: fixes x :: "'a::len word" shows "x < a ^ 0 \ x = 0" by simp lemma nat_add_offset_less: fixes x :: nat assumes yv: "y < 2 ^ n" and xv: "x < 2 ^ m" and mn: "sz = m + n" shows "x * 2 ^ n + y < 2 ^ sz" proof (subst mn) from yv obtain qy where "y + qy = 2 ^ n" and "0 < qy" by (auto dest: less_imp_add_positive) have "x * 2 ^ n + y < x * 2 ^ n + 2 ^ n" by simp fact+ also have "\ = (x + 1) * 2 ^ n" by simp also have "\ \ 2 ^ (m + n)" using xv by (subst power_add) (rule mult_le_mono1, simp) finally show "x * 2 ^ n + y < 2 ^ (m + n)" . qed lemma is_aligned_no_wrap: fixes off :: "'a::len word" fixes ptr :: "'a::len word" assumes al: "is_aligned ptr sz" and off: "off < 2 ^ sz" shows "unat ptr + unat off < 2 ^ LENGTH('a)" proof - have szv: "sz < LENGTH('a)" using off p2_gt_0 word_neq_0_conv by fastforce from al obtain q where ptrq: "ptr = 2 ^ sz * of_nat q" and qv: "q < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis proof (cases "sz = 0") case True then show ?thesis using off ptrq qv by simp next case False then have sne: "0 < sz" .. show ?thesis proof - have uq: "unat (of_nat q ::'a::len word) = q" apply (subst unat_of_nat) apply (rule mod_less) apply (rule order_less_trans [OF qv]) apply (rule power_strict_increasing [OF diff_less [OF sne]]) apply (simp_all) done have uptr: "unat ptr = 2 ^ sz * q" apply (subst ptrq) apply (subst iffD1 [OF unat_mult_lem]) apply (subst unat_power_lower [OF szv]) apply (subst uq) apply (rule nat_less_power_trans [OF qv order_less_imp_le [OF szv]]) apply (subst uq) apply (subst unat_power_lower [OF szv]) apply simp done show "unat ptr + unat off < 2 ^ LENGTH('a)" using szv apply (subst uptr) apply (subst mult.commute, rule nat_add_offset_less [OF _ qv]) apply (rule order_less_le_trans [OF unat_mono [OF off] order_eq_refl]) apply simp_all done qed qed qed lemma is_aligned_no_wrap': fixes ptr :: "'a::len word" assumes al: "is_aligned ptr sz" and off: "off < 2 ^ sz" shows "ptr \ ptr + off" by (subst no_plus_overflow_unat_size, subst word_size, rule is_aligned_no_wrap) fact+ lemma is_aligned_no_overflow': fixes p :: "'a::len word" assumes al: "is_aligned p n" shows "p \ p + (2 ^ n - 1)" proof cases assume "n n ptr \ ptr + 2^sz - 1" by (drule is_aligned_no_overflow') (simp add: field_simps) lemma replicate_not_True: "\n. xs = replicate n False \ True \ set xs" by (induct xs) auto lemma is_aligned_replicateI: "to_bl p = addr @ replicate n False \ is_aligned (p::'a::len word) n" apply (simp add: is_aligned_to_bl word_size) apply (subgoal_tac "length addr = LENGTH('a) - n") apply (simp add: replicate_not_True) apply (drule arg_cong [where f=length]) apply simp done lemma to_bl_2p: "n < LENGTH('a) \ to_bl ((2::'a::len word) ^ n) = replicate (LENGTH('a) - Suc n) False @ True # replicate n False" apply (subst shiftl_1 [symmetric]) apply (subst bl_shiftl) apply (simp add: to_bl_1 min_def word_size) done lemma map_zip_replicate_False_xor: "n = length xs \ map (\(x, y). x = (\ y)) (zip xs (replicate n False)) = xs" by (induct xs arbitrary: n, auto) lemma drop_minus_lem: "\ n \ length xs; 0 < n; n' = length xs \ \ drop (n' - n) xs = rev xs ! (n - 1) # drop (Suc (n' - n)) xs" proof (induct xs arbitrary: n n') case Nil then show ?case by simp next case (Cons y ys) from Cons.prems show ?case apply simp apply (cases "n = Suc (length ys)") apply (simp add: nth_append) apply (simp add: Suc_diff_le Cons.hyps nth_append) apply clarsimp apply arith done qed lemma drop_minus: "\ n < length xs; n' = length xs \ \ drop (n' - Suc n) xs = rev xs ! n # drop (n' - n) xs" apply (subst drop_minus_lem) apply simp apply simp apply simp apply simp apply (cases "length xs", simp) apply (simp add: Suc_diff_le) done lemma xor_2p_to_bl: fixes x::"'a::len word" shows "to_bl (x xor 2^n) = (if n < LENGTH('a) then take (LENGTH('a)-Suc n) (to_bl x) @ (\rev (to_bl x)!n) # drop (LENGTH('a)-n) (to_bl x) else to_bl x)" proof - have x: "to_bl x = take (LENGTH('a)-Suc n) (to_bl x) @ drop (LENGTH('a)-Suc n) (to_bl x)" by simp show ?thesis apply simp apply (rule conjI) apply (clarsimp simp: word_size) apply (simp add: bl_word_xor to_bl_2p) apply (subst x) apply (subst zip_append) apply simp apply (simp add: map_zip_replicate_False_xor drop_minus) apply (auto simp add: word_size nth_w2p intro!: word_eqI) done qed lemma aligned_add_xor: assumes al: "is_aligned (x::'a::len word) n'" and le: "n < n'" shows "(x + 2^n) xor 2^n = x" proof cases assume "n' < LENGTH('a)" with assms show ?thesis apply - apply (rule word_bl.Rep_eqD) apply (subst xor_2p_to_bl) apply simp apply (subst is_aligned_add_conv, simp, simp add: word_less_nat_alt)+ apply (simp add: to_bl_2p nth_append) apply (cases "n' = Suc n") apply simp apply (subst is_aligned_replicate [where n="Suc n", simplified, symmetric]; simp) apply (subgoal_tac "\ LENGTH('a) - Suc n \ LENGTH('a) - n'") prefer 2 apply arith apply (subst replicate_Suc [symmetric]) apply (subst replicate_add [symmetric]) apply (simp add: is_aligned_replicate [simplified, symmetric]) done next assume "\ n' < LENGTH('a)" show ?thesis using al apply (rule is_alignedE) using \\ n' < LENGTH('a)\ by auto qed lemma is_aligned_replicateD: "\ is_aligned (w::'a::len word) n; n \ LENGTH('a) \ \ \xs. to_bl w = xs @ replicate n False \ length xs = size w - n" apply (subst is_aligned_replicate, assumption+) apply (rule exI, rule conjI, rule refl) apply (simp add: word_size) done lemma is_aligned_add_mult_multI: fixes p :: "'a::len word" shows "\is_aligned p m; n \ m; n' = n\ \ is_aligned (p + x * 2 ^ n * z) n'" apply (erule aligned_add_aligned) apply (auto intro: is_alignedI [where k="x*z"]) done lemma is_aligned_add_multI: fixes p :: "'a::len word" shows "\is_aligned p m; n \ m; n' = n\ \ is_aligned (p + x * 2 ^ n) n'" apply (erule aligned_add_aligned) apply (auto intro: is_alignedI [where k="x"]) done lemma unat_of_nat_len: "x < 2 ^ LENGTH('a) \ unat (of_nat x :: 'a::len word) = x" by (simp add: take_bit_nat_eq_self_iff) lemma is_aligned_no_wrap''': fixes ptr :: "'a::len word" shows"\ is_aligned ptr sz; sz < LENGTH('a); off < 2 ^ sz \ \ unat ptr + off < 2 ^ LENGTH('a)" apply (drule is_aligned_no_wrap[where off="of_nat off"]) apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) apply (simp add: take_bit_eq_mod) apply (subst(asm) unat_of_nat_len) apply (erule order_less_trans) apply (erule power_strict_increasing) apply simp apply assumption done lemma is_aligned_get_word_bits: fixes p :: "'a::len word" shows "\ is_aligned p n; \ is_aligned p n; n < LENGTH('a) \ \ P; \ p = 0; n \ LENGTH('a) \ \ P \ \ P" apply (cases "n < LENGTH('a)") apply simp apply simp apply (erule meta_mp) apply (simp add: is_aligned_mask power_add power_overflow not_less flip: take_bit_eq_mask) apply (metis take_bit_length_eq take_bit_of_0 take_bit_tightened) done lemma aligned_small_is_0: "\ is_aligned x n; x < 2 ^ n \ \ x = 0" apply (erule is_aligned_get_word_bits) apply (frule is_aligned_add_conv [rotated, where w=0]) apply (simp add: is_aligned_iff_dvd_nat) apply simp apply (drule is_aligned_replicateD) apply simp apply (clarsimp simp: word_size) apply (subst (asm) replicate_add [symmetric]) apply (drule arg_cong[where f="of_bl :: bool list \ 'a::len word"]) apply simp apply (simp only: replicate.simps[symmetric, where x=False] drop_replicate) done corollary is_aligned_less_sz: "\is_aligned a sz; a \ 0\ \ \ a < 2 ^ sz" by (rule notI, drule(1) aligned_small_is_0, erule(1) notE) lemma aligned_at_least_t2n_diff: "\is_aligned x n; is_aligned y n; x < y\ \ x \ y - 2 ^ n" apply (erule is_aligned_get_word_bits[where p=y]) apply (rule ccontr) apply (clarsimp simp: linorder_not_le) apply (subgoal_tac "y - x = 0") apply clarsimp apply (rule aligned_small_is_0) apply (erule(1) aligned_sub_aligned) apply simp apply unat_arith apply simp done lemma word_sub_1_le: "x \ 0 \ x - 1 \ (x :: ('a :: len) word)" apply (subst no_ulen_sub) apply simp apply (cases "uint x = 0") apply (simp add: uint_0_iff) apply (insert uint_ge_0[where x=x]) apply arith done lemma is_aligned_no_overflow'': "\is_aligned x n; x + 2 ^ n \ 0\ \ x \ x + 2 ^ n" apply (frule is_aligned_no_overflow') apply (erule order_trans) apply (simp add: field_simps) apply (erule word_sub_1_le) done lemma is_aligned_nth: "is_aligned p m = (\n < m. \p !! n)" apply (clarsimp simp: is_aligned_mask bang_eq word_size) apply (rule iffI) apply clarsimp apply (case_tac "n < size p") apply (simp add: word_size) apply (drule test_bit_size) apply simp apply clarsimp done lemma range_inter: "({a..b} \ {c..d} = {}) = (\x. \(a \ x \ x \ b \ c \ x \ x \ d))" by auto lemma aligned_inter_non_empty: "\ {p..p + (2 ^ n - 1)} \ {p..p + 2 ^ m - 1} = {}; is_aligned p n; is_aligned p m\ \ False" apply (clarsimp simp only: range_inter) apply (erule_tac x=p in allE) apply simp apply (erule impE) apply (erule is_aligned_no_overflow') apply (erule notE) apply (erule is_aligned_no_overflow) done lemma not_aligned_mod_nz: assumes al: "\ is_aligned a n" shows "a mod 2 ^ n \ 0" apply (rule ccontr) using al apply (rule notE) apply simp apply (rule is_alignedI [of _ _ \a div 2 ^ n\]) apply (metis add.right_neutral mult.commute word_mod_div_equality) done lemma nat_add_offset_le: fixes x :: nat assumes yv: "y \ 2 ^ n" and xv: "x < 2 ^ m" and mn: "sz = m + n" shows "x * 2 ^ n + y \ 2 ^ sz" proof (subst mn) from yv obtain qy where "y + qy = 2 ^ n" by (auto simp: le_iff_add) have "x * 2 ^ n + y \ x * 2 ^ n + 2 ^ n" using yv xv by simp also have "\ = (x + 1) * 2 ^ n" by simp also have "\ \ 2 ^ (m + n)" using xv by (subst power_add) (rule mult_le_mono1, simp) finally show "x * 2 ^ n + y \ 2 ^ (m + n)" . qed lemma is_aligned_no_wrap_le: fixes ptr::"'a::len word" assumes al: "is_aligned ptr sz" and szv: "sz < LENGTH('a)" and off: "off \ 2 ^ sz" shows "unat ptr + off \ 2 ^ LENGTH('a)" proof - from al obtain q where ptrq: "ptr = 2 ^ sz * of_nat q" and qv: "q < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis proof (cases "sz = 0") case True then show ?thesis using off ptrq qv by (auto simp add: le_Suc_eq Suc_le_eq) (simp add: le_less) next case False then have sne: "0 < sz" .. show ?thesis proof - have uq: "unat (of_nat q :: 'a word) = q" apply (subst unat_of_nat) apply (rule mod_less) apply (rule order_less_trans [OF qv]) apply (rule power_strict_increasing [OF diff_less [OF sne]]) apply simp_all done have uptr: "unat ptr = 2 ^ sz * q" apply (subst ptrq) apply (subst iffD1 [OF unat_mult_lem]) apply (subst unat_power_lower [OF szv]) apply (subst uq) apply (rule nat_less_power_trans [OF qv order_less_imp_le [OF szv]]) apply (subst uq) apply (subst unat_power_lower [OF szv]) apply simp done show "unat ptr + off \ 2 ^ LENGTH('a)" using szv apply (subst uptr) apply (subst mult.commute, rule nat_add_offset_le [OF off qv]) apply simp done qed qed qed lemma is_aligned_neg_mask: "m \ n \ is_aligned (x && ~~ (mask n)) m" by (metis and_not_mask is_aligned_shift is_aligned_weaken) lemma unat_minus: "unat (- (x :: 'a :: len word)) = (if x = 0 then 0 else 2 ^ size x - unat x)" using unat_sub_if_size[where x="2 ^ size x" and y=x] by (simp add: unat_eq_0 word_size) lemma is_aligned_minus: \is_aligned (- p) n\ if \is_aligned p n\ for p :: \'a::len word\ using that apply (cases \n < LENGTH('a)\) apply (simp_all add: not_less is_aligned_beyond_length) apply transfer apply (simp flip: take_bit_eq_0_iff) apply (subst take_bit_minus [symmetric]) apply simp done lemma add_mask_lower_bits: "\is_aligned (x :: 'a :: len word) n; \n' \ n. n' < LENGTH('a) \ \ p !! n'\ \ x + p && ~~ (mask n) = x" apply (subst word_plus_and_or_coroll) apply (rule word_eqI) apply (clarsimp simp: word_size is_aligned_nth) apply (erule_tac x=na in allE)+ apply simp apply (rule word_eqI) apply (clarsimp simp: word_size is_aligned_nth word_ops_nth_size le_def) apply blast done lemma is_aligned_andI1: "is_aligned x n \ is_aligned (x && y) n" by (simp add: is_aligned_nth) lemma is_aligned_andI2: "is_aligned y n \ is_aligned (x && y) n" by (simp add: is_aligned_nth) lemma is_aligned_shiftl: "is_aligned w (n - m) \ is_aligned (w << m) n" by (simp add: is_aligned_nth nth_shiftl) lemma is_aligned_shiftr: "is_aligned w (n + m) \ is_aligned (w >> m) n" by (simp add: is_aligned_nth nth_shiftr) lemma is_aligned_shiftl_self: "is_aligned (p << n) n" by (rule is_aligned_shift) lemma is_aligned_neg_mask_eq: "is_aligned p n \ p && ~~ (mask n) = p" by (metis add.left_neutral is_aligned_mask word_plus_and_or_coroll2) lemma is_aligned_shiftr_shiftl: "is_aligned w n \ w >> n << n = w" by (metis and_not_mask is_aligned_neg_mask_eq) lemma aligned_shiftr_mask_shiftl: "is_aligned x n \ ((x >> n) && mask v) << n = x && mask (v + n)" apply (rule word_eqI) apply (simp add: word_size nth_shiftl nth_shiftr) apply (subgoal_tac "\m. x !! m \ m \ n") apply auto[1] apply (clarsimp simp: is_aligned_mask) apply (drule_tac x=m in word_eqD) apply (frule test_bit_size) apply (simp add: word_size) done lemma mask_zero: "is_aligned x a \ x && mask a = 0" by (metis is_aligned_mask) lemma is_aligned_neg_mask_eq_concrete: "\ is_aligned p n; msk && ~~(mask n) = ~~(mask n) \ \ p && msk = p" by (metis word_bw_assocs(1) word_bw_comms(1) is_aligned_neg_mask_eq) lemma is_aligned_and_not_zero: "\ is_aligned n k; n \ 0 \ \ 2 ^ k \ n" using is_aligned_less_sz leI by blast lemma is_aligned_and_2_to_k: "(n && 2 ^ k - 1) = 0 \ is_aligned (n :: 'a :: len word) k" by (simp add: is_aligned_mask mask_eq_decr_exp) lemma is_aligned_power2: "b \ a \ is_aligned (2 ^ a) b" by (metis is_aligned_triv is_aligned_weaken) lemma aligned_sub_aligned': "\ is_aligned (a :: 'a :: len word) n; is_aligned b n; n < LENGTH('a) \ \ is_aligned (a - b) n" by (simp add: aligned_sub_aligned) lemma is_aligned_neg_mask_weaken: "\ is_aligned p n; m \ n \ \ p && ~~(mask m) = p" using is_aligned_neg_mask_eq is_aligned_weaken by blast -lemma is_aligned_neg_mask2[simp]: +lemma is_aligned_neg_mask2 [simp]: "is_aligned (a && ~~(mask n)) n" by (simp add: and_not_mask is_aligned_shift) end diff --git a/thys/Word_Lib/Bitwise.thy b/thys/Word_Lib/Bitwise.thy --- a/thys/Word_Lib/Bitwise.thy +++ b/thys/Word_Lib/Bitwise.thy @@ -1,527 +1,535 @@ (* Authors: Thomas Sewell, NICTA and Sascha Boehme, TU Muenchen *) theory Bitwise - imports "HOL-Library.Word" More_Arithmetic Most_significant_bit + imports + "HOL-Library.Word" + Most_significant_bit + More_Arithmetic begin text \Helper constants used in defining addition\ definition xor3 :: "bool \ bool \ bool \ bool" where "xor3 a b c = (a = (b = c))" definition carry :: "bool \ bool \ bool \ bool" where "carry a b c = ((a \ (b \ c)) \ (b \ c))" lemma carry_simps: "carry True a b = (a \ b)" "carry a True b = (a \ b)" "carry a b True = (a \ b)" "carry False a b = (a \ b)" "carry a False b = (a \ b)" "carry a b False = (a \ b)" by (auto simp add: carry_def) lemma xor3_simps: "xor3 True a b = (a = b)" "xor3 a True b = (a = b)" "xor3 a b True = (a = b)" "xor3 False a b = (a \ b)" "xor3 a False b = (a \ b)" "xor3 a b False = (a \ b)" by (simp_all add: xor3_def) text \Breaking up word equalities into equalities on their bit lists. Equalities are generated and manipulated in the reverse order to \<^const>\to_bl\.\ lemma bl_word_sub: "to_bl (x - y) = to_bl (x + (- y))" by simp lemma rbl_word_1: "rev (to_bl (1 :: 'a::len word)) = takefill False (LENGTH('a)) [True]" apply (rule_tac s="rev (to_bl (word_succ (0 :: 'a word)))" in trans) apply simp apply (simp only: rtb_rbl_ariths(1)[OF refl]) apply simp apply (case_tac "LENGTH('a)") apply simp apply (simp add: takefill_alt) done lemma rbl_word_if: "rev (to_bl (if P then x else y)) = map2 (If P) (rev (to_bl x)) (rev (to_bl y))" by (simp add: split_def) lemma rbl_add_carry_Cons: "(if car then rbl_succ else id) (rbl_add (x # xs) (y # ys)) = xor3 x y car # (if carry x y car then rbl_succ else id) (rbl_add xs ys)" by (simp add: carry_def xor3_def) lemma rbl_add_suc_carry_fold: "length xs = length ys \ \car. (if car then rbl_succ else id) (rbl_add xs ys) = (foldr (\(x, y) res car. xor3 x y car # res (carry x y car)) (zip xs ys) (\_. [])) car" apply (erule list_induct2) apply simp apply (simp only: rbl_add_carry_Cons) apply simp done lemma to_bl_plus_carry: "to_bl (x + y) = rev (foldr (\(x, y) res car. xor3 x y car # res (carry x y car)) (rev (zip (to_bl x) (to_bl y))) (\_. []) False)" using rbl_add_suc_carry_fold[where xs="rev (to_bl x)" and ys="rev (to_bl y)"] apply (simp add: word_add_rbl[OF refl refl]) apply (drule_tac x=False in spec) apply (simp add: zip_rev) done definition "rbl_plus cin xs ys = foldr (\(x, y) res car. xor3 x y car # res (carry x y car)) (zip xs ys) (\_. []) cin" lemma rbl_plus_simps: "rbl_plus cin (x # xs) (y # ys) = xor3 x y cin # rbl_plus (carry x y cin) xs ys" "rbl_plus cin [] ys = []" "rbl_plus cin xs [] = []" by (simp_all add: rbl_plus_def) lemma rbl_word_plus: "rev (to_bl (x + y)) = rbl_plus False (rev (to_bl x)) (rev (to_bl y))" by (simp add: rbl_plus_def to_bl_plus_carry zip_rev) definition "rbl_succ2 b xs = (if b then rbl_succ xs else xs)" lemma rbl_succ2_simps: "rbl_succ2 b [] = []" "rbl_succ2 b (x # xs) = (b \ x) # rbl_succ2 (x \ b) xs" by (simp_all add: rbl_succ2_def) lemma twos_complement: "- x = word_succ (NOT x)" using arg_cong[OF word_add_not[where x=x], where f="\a. a - x + 1"] by (simp add: word_succ_p1 word_sp_01[unfolded word_succ_p1] del: word_add_not) lemma rbl_word_neg: "rev (to_bl (- x)) = rbl_succ2 True (map Not (rev (to_bl x)))" for x :: \'a::len word\ by (simp add: twos_complement word_succ_rbl[OF refl] bl_word_not rev_map rbl_succ2_def) lemma rbl_word_cat: "rev (to_bl (word_cat x y :: 'a::len word)) = takefill False (LENGTH('a)) (rev (to_bl y) @ rev (to_bl x))" by (simp add: word_cat_bl word_rev_tf) lemma rbl_word_slice: "rev (to_bl (slice n w :: 'a::len word)) = takefill False (LENGTH('a)) (drop n (rev (to_bl w)))" apply (simp add: slice_take word_rev_tf rev_take) apply (cases "n < LENGTH('b)", simp_all) done lemma rbl_word_ucast: "rev (to_bl (ucast x :: 'a::len word)) = takefill False (LENGTH('a)) (rev (to_bl x))" apply (simp add: to_bl_ucast takefill_alt) apply (simp add: rev_drop) apply (cases "LENGTH('a) < LENGTH('b)") apply simp_all done lemma rbl_shiftl: "rev (to_bl (w << n)) = takefill False (size w) (replicate n False @ rev (to_bl w))" by (simp add: bl_shiftl takefill_alt word_size rev_drop) lemma rbl_shiftr: "rev (to_bl (w >> n)) = takefill False (size w) (drop n (rev (to_bl w)))" by (simp add: shiftr_slice rbl_word_slice word_size) definition "drop_nonempty v n xs = (if n < length xs then drop n xs else [last (v # xs)])" lemma drop_nonempty_simps: "drop_nonempty v (Suc n) (x # xs) = drop_nonempty x n xs" "drop_nonempty v 0 (x # xs) = (x # xs)" "drop_nonempty v n [] = [v]" by (simp_all add: drop_nonempty_def) definition "takefill_last x n xs = takefill (last (x # xs)) n xs" lemma takefill_last_simps: "takefill_last z (Suc n) (x # xs) = x # takefill_last x n xs" "takefill_last z 0 xs = []" "takefill_last z n [] = replicate n z" by (simp_all add: takefill_last_def) (simp_all add: takefill_alt) lemma rbl_sshiftr: "rev (to_bl (w >>> n)) = takefill_last False (size w) (drop_nonempty False n (rev (to_bl w)))" apply (cases "n < size w") apply (simp add: bl_sshiftr takefill_last_def word_size takefill_alt rev_take last_rev drop_nonempty_def) apply (subgoal_tac "(w >>> n) = of_bl (replicate (size w) (msb w))") apply (simp add: word_size takefill_last_def takefill_alt last_rev word_msb_alt word_rev_tf drop_nonempty_def take_Cons') apply (case_tac "LENGTH('a)", simp_all) apply (rule word_eqI) apply (simp add: nth_sshiftr word_size test_bit_of_bl msb_nth) done lemma nth_word_of_int: "(word_of_int x :: 'a::len word) !! n = (n < LENGTH('a) \ bin_nth x n)" apply (simp add: test_bit_bl word_size to_bl_of_bin) apply (subst conj_cong[OF refl], erule bin_nth_bl) apply auto done lemma nth_scast: "(scast (x :: 'a::len word) :: 'b::len word) !! n = (n < LENGTH('b) \ (if n < LENGTH('a) - 1 then x !! n else x !! (LENGTH('a) - 1)))" apply transfer apply (auto simp add: bit_signed_take_bit_iff min_def) done lemma rbl_word_scast: "rev (to_bl (scast x :: 'a::len word)) = takefill_last False (LENGTH('a)) (rev (to_bl x))" apply (rule nth_equalityI) apply (simp add: word_size takefill_last_def) apply (clarsimp simp: nth_scast takefill_last_def nth_takefill word_size rev_nth to_bl_nth) apply (cases "LENGTH('b)") apply simp apply (clarsimp simp: less_Suc_eq_le linorder_not_less last_rev word_msb_alt[symmetric] msb_nth) done definition rbl_mul :: "bool list \ bool list \ bool list" where "rbl_mul xs ys = foldr (\x sm. rbl_plus False (map ((\) x) ys) (False # sm)) xs []" lemma rbl_mul_simps: "rbl_mul (x # xs) ys = rbl_plus False (map ((\) x) ys) (False # rbl_mul xs ys)" "rbl_mul [] ys = []" by (simp_all add: rbl_mul_def) lemma takefill_le2: "length xs \ n \ takefill x m (takefill x n xs) = takefill x m xs" by (simp add: takefill_alt replicate_add[symmetric]) lemma take_rbl_plus: "\n b. take n (rbl_plus b xs ys) = rbl_plus b (take n xs) (take n ys)" apply (simp add: rbl_plus_def take_zip[symmetric]) apply (rule_tac list="zip xs ys" in list.induct) apply simp apply (clarsimp simp: split_def) apply (case_tac n, simp_all) done lemma word_rbl_mul_induct: "length xs \ size y \ rbl_mul xs (rev (to_bl y)) = take (length xs) (rev (to_bl (of_bl (rev xs) * y)))" for y :: "'a::len word" proof (induct xs) case Nil show ?case by (simp add: rbl_mul_simps) next case (Cons z zs) have rbl_word_plus': "to_bl (x + y) = rev (rbl_plus False (rev (to_bl x)) (rev (to_bl y)))" for x y :: "'a word" by (simp add: rbl_word_plus[symmetric]) have mult_bit: "to_bl (of_bl [z] * y) = map ((\) z) (to_bl y)" by (cases z) (simp cong: map_cong, simp add: map_replicate_const cong: map_cong) have shiftl: "of_bl xs * 2 * y = (of_bl xs * y) << 1" for xs by (simp add: shiftl_t2n) have zip_take_triv: "\xs ys n. n = length ys \ zip (take n xs) ys = zip xs ys" by (rule nth_equalityI) simp_all from Cons show ?case apply (simp add: trans [OF of_bl_append add.commute] rbl_mul_simps rbl_word_plus' distrib_right mult_bit shiftl rbl_shiftl) apply (simp add: takefill_alt word_size rev_map take_rbl_plus min_def) apply (simp add: rbl_plus_def zip_take_triv) done qed lemma rbl_word_mul: "rev (to_bl (x * y)) = rbl_mul (rev (to_bl x)) (rev (to_bl y))" for x :: "'a::len word" using word_rbl_mul_induct[where xs="rev (to_bl x)" and y=y] by (simp add: word_size) text \Breaking up inequalities into bitlist properties.\ definition "rev_bl_order F xs ys = (length xs = length ys \ ((xs = ys \ F) \ (\n < length xs. drop (Suc n) xs = drop (Suc n) ys \ \ xs ! n \ ys ! n)))" lemma rev_bl_order_simps: "rev_bl_order F [] [] = F" "rev_bl_order F (x # xs) (y # ys) = rev_bl_order ((y \ \ x) \ ((y \ \ x) \ F)) xs ys" apply (simp_all add: rev_bl_order_def) apply (rule conj_cong[OF refl]) apply (cases "xs = ys") apply (simp add: nth_Cons') apply blast apply (simp add: nth_Cons') apply safe apply (rule_tac x="n - 1" in exI) apply simp apply (rule_tac x="Suc n" in exI) apply simp done lemma rev_bl_order_rev_simp: "length xs = length ys \ rev_bl_order F (xs @ [x]) (ys @ [y]) = ((y \ \ x) \ ((y \ \ x) \ rev_bl_order F xs ys))" by (induct arbitrary: F rule: list_induct2) (auto simp: rev_bl_order_simps) lemma rev_bl_order_bl_to_bin: "length xs = length ys \ rev_bl_order True xs ys = (bl_to_bin (rev xs) \ bl_to_bin (rev ys)) \ rev_bl_order False xs ys = (bl_to_bin (rev xs) < bl_to_bin (rev ys))" apply (induct xs ys rule: list_induct2) apply (simp_all add: rev_bl_order_simps bl_to_bin_app_cat concat_bit_Suc) apply (auto simp add: bl_to_bin_def add1_zle_eq) done lemma word_le_rbl: "x \ y \ rev_bl_order True (rev (to_bl x)) (rev (to_bl y))" for x y :: "'a::len word" by (simp add: rev_bl_order_bl_to_bin word_le_def) lemma word_less_rbl: "x < y \ rev_bl_order False (rev (to_bl x)) (rev (to_bl y))" for x y :: "'a::len word" by (simp add: word_less_alt rev_bl_order_bl_to_bin) lemma word_sint_msb_eq: "sint x = uint x - (if msb x then 2 ^ size x else 0)" apply (cases "msb x") apply (rule word_sint.Abs_eqD[where 'a='a], simp_all) apply (simp add: word_size wi_hom_syms word_of_int_2p_len) apply (simp add: sints_num word_size) apply (rule conjI) apply (simp add: le_diff_eq') apply (rule order_trans[where y="2 ^ (LENGTH('a) - 1)"]) apply (simp add: power_Suc[symmetric]) apply (simp add: linorder_not_less[symmetric] mask_eq_iff[symmetric]) apply (rule notI, drule word_eqD[where x="size x - 1"]) apply (simp add: msb_nth word_ops_nth_size word_size) apply (simp add: order_less_le_trans[where y=0]) apply (rule word_uint.Abs_eqD[where 'a='a], simp_all) apply (simp add: linorder_not_less uints_num word_msb_sint) apply (rule order_less_le_trans[OF sint_lt]) apply simp done lemma word_sle_msb_le: "x <=s y \ (msb y \ msb x) \ ((msb x \ \ msb y) \ x \ y)" apply (simp add: word_sle_eq word_sint_msb_eq word_size word_le_def) apply safe apply (rule order_trans[OF _ uint_ge_0]) apply (simp add: order_less_imp_le) apply (erule notE[OF leD]) apply (rule order_less_le_trans[OF _ uint_ge_0]) apply simp done lemma word_sless_msb_less: "x (msb y \ msb x) \ ((msb x \ \ msb y) \ x < y)" by (auto simp add: word_sless_eq word_sle_msb_le) definition "map_last f xs = (if xs = [] then [] else butlast xs @ [f (last xs)])" lemma map_last_simps: "map_last f [] = []" "map_last f [x] = [f x]" "map_last f (x # y # zs) = x # map_last f (y # zs)" by (simp_all add: map_last_def) lemma word_sle_rbl: "x <=s y \ rev_bl_order True (map_last Not (rev (to_bl x))) (map_last Not (rev (to_bl y)))" using word_msb_alt[where w=x] word_msb_alt[where w=y] apply (simp add: word_sle_msb_le word_le_rbl) apply (subgoal_tac "length (to_bl x) = length (to_bl y)") apply (cases "to_bl x", simp) apply (cases "to_bl y", simp) apply (clarsimp simp: map_last_def rev_bl_order_rev_simp) apply auto done lemma word_sless_rbl: "x rev_bl_order False (map_last Not (rev (to_bl x))) (map_last Not (rev (to_bl y)))" using word_msb_alt[where w=x] word_msb_alt[where w=y] apply (simp add: word_sless_msb_less word_less_rbl) apply (subgoal_tac "length (to_bl x) = length (to_bl y)") apply (cases "to_bl x", simp) apply (cases "to_bl y", simp) apply (clarsimp simp: map_last_def rev_bl_order_rev_simp) apply auto done text \Lemmas for unpacking \<^term>\rev (to_bl n)\ for numerals n and also for irreducible values and expressions.\ lemma rev_bin_to_bl_simps: "rev (bin_to_bl 0 x) = []" "rev (bin_to_bl (Suc n) (numeral (num.Bit0 nm))) = False # rev (bin_to_bl n (numeral nm))" "rev (bin_to_bl (Suc n) (numeral (num.Bit1 nm))) = True # rev (bin_to_bl n (numeral nm))" "rev (bin_to_bl (Suc n) (numeral (num.One))) = True # replicate n False" "rev (bin_to_bl (Suc n) (- numeral (num.Bit0 nm))) = False # rev (bin_to_bl n (- numeral nm))" "rev (bin_to_bl (Suc n) (- numeral (num.Bit1 nm))) = True # rev (bin_to_bl n (- numeral (nm + num.One)))" "rev (bin_to_bl (Suc n) (- numeral (num.One))) = True # replicate n True" "rev (bin_to_bl (Suc n) (- numeral (num.Bit0 nm + num.One))) = True # rev (bin_to_bl n (- numeral (nm + num.One)))" "rev (bin_to_bl (Suc n) (- numeral (num.Bit1 nm + num.One))) = False # rev (bin_to_bl n (- numeral (nm + num.One)))" "rev (bin_to_bl (Suc n) (- numeral (num.One + num.One))) = False # rev (bin_to_bl n (- numeral num.One))" by (simp_all add: bin_to_bl_aux_append bin_to_bl_zero_aux bin_to_bl_minus1_aux replicate_append_same) lemma to_bl_upt: "to_bl x = rev (map ((!!) x) [0 ..< size x])" apply (rule nth_equalityI) apply (simp add: word_size) apply (auto simp: to_bl_nth word_size rev_nth) done lemma rev_to_bl_upt: "rev (to_bl x) = map ((!!) x) [0 ..< size x]" by (simp add: to_bl_upt) lemma upt_eq_list_intros: "j \ i \ [i ..< j] = []" "i = x \ x < j \ [x + 1 ..< j] = xs \ [i ..< j] = (x # xs)" by (simp_all add: upt_eq_Cons_conv) subsection \Tactic definition\ +lemma if_bool_simps: + "If p True y = (p \ y) \ If p False y = (\ p \ y) \ + If p y True = (p \ y) \ If p y False = (p \ y)" + by auto + ML \ structure Word_Bitwise_Tac = struct val word_ss = simpset_of \<^theory_context>\Word\; fun mk_nat_clist ns = fold_rev (Thm.mk_binop \<^cterm>\Cons :: nat \ _\) ns \<^cterm>\[] :: nat list\; fun upt_conv ctxt ct = case Thm.term_of ct of (\<^const>\upt\ $ n $ m) => let val (i, j) = apply2 (snd o HOLogic.dest_number) (n, m); val ns = map (Numeral.mk_cnumber \<^ctyp>\nat\) (i upto (j - 1)) |> mk_nat_clist; val prop = Thm.mk_binop \<^cterm>\(=) :: nat list \ _\ ct ns |> Thm.apply \<^cterm>\Trueprop\; in try (fn () => Goal.prove_internal ctxt [] prop (K (REPEAT_DETERM (resolve_tac ctxt @{thms upt_eq_list_intros} 1 ORELSE simp_tac (put_simpset word_ss ctxt) 1))) |> mk_meta_eq) () end | _ => NONE; val expand_upt_simproc = Simplifier.make_simproc \<^context> "expand_upt" {lhss = [\<^term>\upt x y\], proc = K upt_conv}; fun word_len_simproc_fn ctxt ct = (case Thm.term_of ct of Const (\<^const_name>\len_of\, _) $ t => (let val T = fastype_of t |> dest_Type |> snd |> the_single val n = Numeral.mk_cnumber \<^ctyp>\nat\ (Word_Lib.dest_binT T); val prop = Thm.mk_binop \<^cterm>\(=) :: nat \ _\ ct n |> Thm.apply \<^cterm>\Trueprop\; in Goal.prove_internal ctxt [] prop (K (simp_tac (put_simpset word_ss ctxt) 1)) |> mk_meta_eq |> SOME end handle TERM _ => NONE | TYPE _ => NONE) | _ => NONE); val word_len_simproc = Simplifier.make_simproc \<^context> "word_len" {lhss = [\<^term>\len_of x\], proc = K word_len_simproc_fn}; (* convert 5 or nat 5 to Suc 4 when n_sucs = 1, Suc (Suc 4) when n_sucs = 2, or just 5 (discarding nat) when n_sucs = 0 *) fun nat_get_Suc_simproc_fn n_sucs ctxt ct = let val (f $ arg) = Thm.term_of ct; val n = (case arg of \<^term>\nat\ $ n => n | n => n) |> HOLogic.dest_number |> snd; val (i, j) = if n > n_sucs then (n_sucs, n - n_sucs) else (n, 0); val arg' = funpow i HOLogic.mk_Suc (HOLogic.mk_number \<^typ>\nat\ j); val _ = if arg = arg' then raise TERM ("", []) else (); fun propfn g = HOLogic.mk_eq (g arg, g arg') |> HOLogic.mk_Trueprop |> Thm.cterm_of ctxt; val eq1 = Goal.prove_internal ctxt [] (propfn I) (K (simp_tac (put_simpset word_ss ctxt) 1)); in Goal.prove_internal ctxt [] (propfn (curry (op $) f)) (K (simp_tac (put_simpset HOL_ss ctxt addsimps [eq1]) 1)) |> mk_meta_eq |> SOME end handle TERM _ => NONE; fun nat_get_Suc_simproc n_sucs ts = Simplifier.make_simproc \<^context> "nat_get_Suc" {lhss = map (fn t => t $ \<^term>\n :: nat\) ts, proc = K (nat_get_Suc_simproc_fn n_sucs)}; val no_split_ss = simpset_of (put_simpset HOL_ss \<^context> |> Splitter.del_split @{thm if_split}); val expand_word_eq_sss = (simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms word_eq_rbl_eq word_le_rbl word_less_rbl word_sle_rbl word_sless_rbl}), map simpset_of [ put_simpset no_split_ss \<^context> addsimps @{thms rbl_word_plus rbl_word_and rbl_word_or rbl_word_not rbl_word_neg bl_word_sub rbl_word_xor rbl_word_cat rbl_word_slice rbl_word_scast rbl_word_ucast rbl_shiftl rbl_shiftr rbl_sshiftr rbl_word_if}, put_simpset no_split_ss \<^context> addsimps @{thms to_bl_numeral to_bl_neg_numeral to_bl_0 rbl_word_1}, put_simpset no_split_ss \<^context> addsimps @{thms rev_rev_ident rev_replicate rev_map to_bl_upt word_size} addsimprocs [word_len_simproc], put_simpset no_split_ss \<^context> addsimps @{thms list.simps split_conv replicate.simps list.map zip_Cons_Cons zip_Nil drop_Suc_Cons drop_0 drop_Nil foldr.simps list.map zip.simps(1) zip_Nil zip_Cons_Cons takefill_Suc_Cons takefill_Suc_Nil takefill.Z rbl_succ2_simps rbl_plus_simps rev_bin_to_bl_simps append.simps takefill_last_simps drop_nonempty_simps rev_bl_order_simps} addsimprocs [expand_upt_simproc, nat_get_Suc_simproc 4 [\<^term>\replicate\, \<^term>\takefill x\, \<^term>\drop\, \<^term>\bin_to_bl\, \<^term>\takefill_last x\, \<^term>\drop_nonempty x\]], put_simpset no_split_ss \<^context> addsimps @{thms xor3_simps carry_simps if_bool_simps} ]) fun tac ctxt = let val (ss, sss) = expand_word_eq_sss; in foldr1 (op THEN_ALL_NEW) ((CHANGED o safe_full_simp_tac (put_simpset ss ctxt)) :: map (fn ss => safe_full_simp_tac (put_simpset ss ctxt)) sss) end; end \ method_setup word_bitwise = \Scan.succeed (fn ctxt => Method.SIMPLE_METHOD (Word_Bitwise_Tac.tac ctxt 1))\ "decomposer for word equalities and inequalities into bit propositions" end diff --git a/thys/Word_Lib/More_Arithmetic.thy b/thys/Word_Lib/More_Arithmetic.thy --- a/thys/Word_Lib/More_Arithmetic.thy +++ b/thys/Word_Lib/More_Arithmetic.thy @@ -1,461 +1,252 @@ -section \Miscellaneous lemmas (mostly) for arithmetic\ +section \Arithmetic lemmas\ theory More_Arithmetic - imports Main Bits_Int + imports Main begin -lemma int_mod_lem: "0 < n \ 0 \ b \ b < n \ b mod n = b" - for b n :: int - apply safe - apply (erule (1) mod_pos_pos_trivial) - apply (erule_tac [!] subst) - apply auto - done - -lemma int_mod_ge': "b < 0 \ 0 < n \ b + n \ b mod n" - for b n :: int - by (metis add_less_same_cancel2 int_mod_ge mod_add_self2) - -lemma int_mod_le': "0 \ b - n \ b mod n \ b - n" - for b n :: int - by (metis minus_mod_self2 zmod_le_nonneg_dividend) - -lemma emep1: "even n \ even d \ 0 \ d \ (n + 1) mod d = (n mod d) + 1" - for n d :: int - by (auto simp add: pos_zmod_mult_2 add.commute dvd_def) - -lemma m1mod2k: "- 1 mod 2 ^ n = (2 ^ n - 1 :: int)" - by (rule zmod_minus1) simp - -lemma sb_inc_lem: "a + 2^k < 0 \ a + 2^k + 2^(Suc k) \ (a + 2^k) mod 2^(Suc k)" - for a :: int - using int_mod_ge' [where n = "2 ^ (Suc k)" and b = "a + 2 ^ k"] - by simp - -lemma sb_inc_lem': "a < - (2^k) \ a + 2^k + 2^(Suc k) \ (a + 2^k) mod 2^(Suc k)" - for a :: int - by (rule sb_inc_lem) simp - -lemma sb_dec_lem: "0 \ - (2 ^ k) + a \ (a + 2 ^ k) mod (2 * 2 ^ k) \ - (2 ^ k) + a" - for a :: int - using int_mod_le'[where n = "2 ^ (Suc k)" and b = "a + 2 ^ k"] by simp - -lemma sb_dec_lem': "2 ^ k \ a \ (a + 2 ^ k) mod (2 * 2 ^ k) \ - (2 ^ k) + a" - for a :: int - by (rule sb_dec_lem) simp - -lemma one_mod_exp_eq_one [simp]: - "1 mod (2 * 2 ^ n) = (1::int)" - using power_gt1 [of 2 n] by (auto intro: mod_pos_pos_trivial) - -lemma mod_2_neq_1_eq_eq_0: "k mod 2 \ 1 \ k mod 2 = 0" - for k :: int - by (fact not_mod_2_eq_1_eq_0) - -lemma z1pmod2: "(2 * b + 1) mod 2 = (1::int)" - for b :: int - by arith - -lemma diff_le_eq': "a - b \ c \ a \ b + c" - for a b c :: int - by arith - -lemma zless2: "0 < (2 :: int)" - by (fact zero_less_numeral) - -lemma zless2p: "0 < (2 ^ n :: int)" - by arith - -lemma zle2p: "0 \ (2 ^ n :: int)" - by arith - -lemma p1mod22k': "(1 + 2 * b) mod (2 * 2 ^ n) = 1 + 2 * (b mod 2 ^ n)" - for b :: int - using zle2p by (rule pos_zmod_mult_2) - -lemma p1mod22k: "(2 * b + 1) mod (2 * 2 ^ n) = 2 * (b mod 2 ^ n) + 1" - for b :: int - by (simp add: p1mod22k' add.commute) - -lemma ex_eq_or: "(\m. n = Suc m \ (m = k \ P m)) \ n = Suc k \ (\m. n = Suc m \ P m)" - by auto - -lemma power_minus_simp: "0 < n \ a ^ n = a * a ^ (n - 1)" - by (auto dest: gr0_implies_Suc) - -lemma funpow_minus_simp: "0 < n \ f ^^ n = f \ f ^^ (n - 1)" - by (auto dest: gr0_implies_Suc) - -lemma power_numeral: "a ^ numeral k = a * a ^ (pred_numeral k)" - by (simp add: numeral_eq_Suc) - -lemma funpow_numeral [simp]: "f ^^ numeral k = f \ f ^^ (pred_numeral k)" - by (simp add: numeral_eq_Suc) - -lemma replicate_numeral [simp]: "replicate (numeral k) x = x # replicate (pred_numeral k) x" - by (simp add: numeral_eq_Suc) - -lemma rco_alt: "(f \ g) ^^ n \ f = f \ (g \ f) ^^ n" - apply (rule ext) - apply (induct n) - apply (simp_all add: o_def) - done - -lemma list_exhaust_size_gt0: - assumes "\a list. y = a # list \ P" - shows "0 < length y \ P" - apply (cases y) - apply simp - apply (rule assms) - apply fastforce - done - -lemma list_exhaust_size_eq0: - assumes "y = [] \ P" - shows "length y = 0 \ P" - apply (cases y) - apply (rule assms) - apply simp - apply simp - done - -lemma size_Cons_lem_eq: "y = xa # list \ size y = Suc k \ size list = k" - by auto - -lemmas ls_splits = prod.split prod.split_asm if_split_asm - -\ \simplifications for specific word lengths\ -lemmas n2s_ths [THEN eq_reflection] = add_2_eq_Suc add_2_eq_Suc' - -lemmas s2n_ths = n2s_ths [symmetric] - -lemma and_len: "xs = ys \ xs = ys \ length xs = length ys" - by auto - -lemma size_if: "size (if p then xs else ys) = (if p then size xs else size ys)" - by auto - -lemma tl_if: "tl (if p then xs else ys) = (if p then tl xs else tl ys)" - by auto - -lemma hd_if: "hd (if p then xs else ys) = (if p then hd xs else hd ys)" - by auto - -lemma if_Not_x: "(if p then \ x else x) = (p = (\ x))" - by auto - -lemma if_x_Not: "(if p then x else \ x) = (p = x)" - by auto - -lemma if_same_and: "(If p x y \ If p u v) = (if p then x \ u else y \ v)" - by auto - -lemma if_same_eq: "(If p x y = (If p u v)) = (if p then x = u else y = v)" - by auto - -lemma if_same_eq_not: "(If p x y = (\ If p u v)) = (if p then x = (\ u) else y = (\ v))" - by auto - -\ \note -- \if_Cons\ can cause blowup in the size, if \p\ is complex, so make a simproc\ -lemma if_Cons: "(if p then x # xs else y # ys) = If p x y # If p xs ys" - by auto - -lemma if_single: "(if xc then [xab] else [an]) = [if xc then xab else an]" - by auto - -lemma if_bool_simps: - "If p True y = (p \ y) \ If p False y = (\ p \ y) \ - If p y True = (p \ y) \ If p y False = (p \ y)" - by auto - -lemmas if_simps = - if_x_Not if_Not_x if_cancel if_True if_False if_bool_simps - -lemmas seqr = eq_reflection [where x = "size w"] for w (* FIXME: delete *) - -lemma the_elemI: "y = {x} \ the_elem y = x" - by simp - -lemma nonemptyE: "S \ {} \ (\x. x \ S \ R) \ R" - by auto - -lemma gt_or_eq_0: "0 < y \ 0 = y" - for y :: nat - by arith - -lemmas xtr1 = xtrans(1) -lemmas xtr2 = xtrans(2) -lemmas xtr3 = xtrans(3) -lemmas xtr4 = xtrans(4) -lemmas xtr5 = xtrans(5) -lemmas xtr6 = xtrans(6) -lemmas xtr7 = xtrans(7) -lemmas xtr8 = xtrans(8) - -lemmas nat_simps = diff_add_inverse2 diff_add_inverse -lemmas nat_iffs = le_add1 le_add2 - -lemma sum_imp_diff: "j = k + i \ j - i = k" - for k :: nat - by arith - -lemmas pos_mod_sign2 = zless2 [THEN pos_mod_sign [where b = "2::int"]] -lemmas pos_mod_bound2 = zless2 [THEN pos_mod_bound [where b = "2::int"]] - -lemma nmod2: "n mod 2 = 0 \ n mod 2 = 1" - for n :: int - by arith - -lemma eme1p: - "even n \ even d \ 0 \ d \ (1 + n) mod d = 1 + n mod d" for n d :: int - using emep1 [of n d] by (simp add: ac_simps) - -lemma le_diff_eq': "a \ c - b \ b + a \ c" - for a b c :: int - by arith +declare iszero_0 [intro] -lemma less_diff_eq': "a < c - b \ b + a < c" - for a b c :: int - by arith - -lemma diff_less_eq': "a - b < c \ a < b + c" - for a b c :: int - by arith - -lemmas m1mod22k = mult_pos_pos [OF zless2 zless2p, THEN zmod_minus1] - -lemma z1pdiv2: "(2 * b + 1) div 2 = b" - for b :: int - by arith - -lemmas zdiv_le_dividend = xtr3 [OF div_by_1 [symmetric] zdiv_mono2, - simplified int_one_le_iff_zero_less, simplified] - -lemma axxbyy: "a + m + m = b + n + n \ a = 0 \ a = 1 \ b = 0 \ b = 1 \ a = b \ m = n" - for a b m n :: int - by arith - -lemma axxmod2: "(1 + x + x) mod 2 = 1 \ (0 + x + x) mod 2 = 0" - for x :: int - by arith - -lemma axxdiv2: "(1 + x + x) div 2 = x \ (0 + x + x) div 2 = x" - for x :: int - by arith - -lemmas iszero_minus = - trans [THEN trans, OF iszero_def neg_equal_0_iff_equal iszero_def [symmetric]] - -lemmas zadd_diff_inverse = - trans [OF diff_add_cancel [symmetric] add.commute] - -lemmas add_diff_cancel2 = - add.commute [THEN diff_eq_eq [THEN iffD2]] - -lemmas rdmods [symmetric] = mod_minus_eq - mod_diff_left_eq mod_diff_right_eq mod_add_left_eq - mod_add_right_eq mod_mult_right_eq mod_mult_left_eq - -lemma mod_plus_right: "(a + x) mod m = (b + x) mod m \ a mod m = b mod m" - for a b m x :: nat - by (induct x) (simp_all add: mod_Suc, arith) - -lemma nat_minus_mod: "(n - n mod m) mod m = 0" - for m n :: nat - by (induct n) (simp_all add: mod_Suc) - -lemmas nat_minus_mod_plus_right = - trans [OF nat_minus_mod mod_0 [symmetric], - THEN mod_plus_right [THEN iffD2], simplified] - -lemmas push_mods' = mod_add_eq - mod_mult_eq mod_diff_eq - mod_minus_eq - -lemmas push_mods = push_mods' [THEN eq_reflection] -lemmas pull_mods = push_mods [symmetric] rdmods [THEN eq_reflection] - -lemma nat_mod_eq: "b < n \ a mod n = b mod n \ a mod n = b" - for a b n :: nat - by (induct a) auto - -lemmas nat_mod_eq' = refl [THEN [2] nat_mod_eq] - -lemma nat_mod_lem: "0 < n \ b < n \ b mod n = b" - for b n :: nat - apply safe - apply (erule nat_mod_eq') - apply (erule subst) - apply (erule mod_less_divisor) - done - -lemma mod_nat_add: "x < z \ y < z \ (x + y) mod z = (if x + y < z then x + y else x + y - z)" - for x y z :: nat - apply (rule nat_mod_eq) - apply auto - apply (rule trans) - apply (rule le_mod_geq) - apply simp - apply (rule nat_mod_eq') - apply arith - done - -lemma mod_nat_sub: "x < z \ (x - y) mod z = x - y" - for x y :: nat - by (rule nat_mod_eq') arith - -lemma int_mod_eq: "0 \ b \ b < n \ a mod n = b mod n \ a mod n = b" - for a b n :: int - by (metis mod_pos_pos_trivial) - -lemmas int_mod_eq' = mod_pos_pos_trivial (* FIXME delete *) - -lemmas int_mod_le = zmod_le_nonneg_dividend (* FIXME: delete *) - -lemma mod_add_if_z: - "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ - (x + y) mod z = (if x + y < z then x + y else x + y - z)" - for x y z :: int - by (auto intro: int_mod_eq) +declare min.absorb1 [simp] min.absorb2 [simp] -lemma mod_sub_if_z: - "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ - (x - y) mod z = (if y \ x then x - y else x - y + z)" - for x y z :: int - by (auto intro: int_mod_eq) - -lemmas zmde = mult_div_mod_eq [symmetric, THEN diff_eq_eq [THEN iffD2], symmetric] -lemmas mcl = mult_cancel_left [THEN iffD1, THEN make_pos_rule] - -(* already have this for naturals, div_mult_self1/2, but not for ints *) -lemma zdiv_mult_self: "m \ 0 \ (a + m * n) div m = a div m + n" - for a m n :: int - apply (rule mcl) - prefer 2 - apply (erule asm_rl) - apply (simp add: zmde ring_distribs) - done - -lemma mod_power_lem: "a > 1 \ a ^ n mod a ^ m = (if m \ n then 0 else a ^ n)" - for a :: int - by (simp add: mod_eq_0_iff_dvd le_imp_power_dvd) - -lemma pl_pl_rels: "a + b = c + d \ a \ c \ b \ d \ a \ c \ b \ d" - for a b c d :: nat - by arith - -lemmas pl_pl_rels' = add.commute [THEN [2] trans, THEN pl_pl_rels] - -lemma minus_eq: "m - k = m \ k = 0 \ m = 0" - for k m :: nat - by arith - -lemma pl_pl_mm: "a + b = c + d \ a - c = d - b" - for a b c d :: nat - by arith - -lemmas pl_pl_mm' = add.commute [THEN [2] trans, THEN pl_pl_mm] - -lemmas dme = div_mult_mod_eq -lemmas dtle = div_times_less_eq_dividend -lemmas th2 = order_trans [OF order_refl [THEN [2] mult_le_mono] div_times_less_eq_dividend] - -lemma td_gal: "0 < c \ a \ b * c \ a div c \ b" - for a b c :: nat - apply safe - apply (erule (1) xtr4 [OF div_le_mono div_mult_self_is_m]) - apply (erule th2) - done - -lemmas td_gal_lt = td_gal [simplified not_less [symmetric], simplified] - -lemmas div_mult_le = div_times_less_eq_dividend - -lemmas sdl = div_nat_eqI - -lemma given_quot: "f > 0 \ (f * l + (f - 1)) div f = l" - for f l :: nat - by (rule div_nat_eqI) (simp_all) - -lemma given_quot_alt: "f > 0 \ (l * f + f - Suc 0) div f = l" - for f l :: nat - apply (frule given_quot) - apply (rule trans) - prefer 2 - apply (erule asm_rl) - apply (rule_tac f="\n. n div f" in arg_cong) - apply (simp add : ac_simps) - done - -lemma diff_mod_le: "a < d \ b dvd d \ a - a mod b \ d - b" - for a b d :: nat - apply (unfold dvd_def) - apply clarify - apply (case_tac k) - apply clarsimp - apply clarify - apply (cases "b > 0") - apply (drule mult.commute [THEN xtr1]) - apply (frule (1) td_gal_lt [THEN iffD1]) - apply (clarsimp simp: le_simps) - apply (rule minus_mod_eq_mult_div [symmetric, THEN [2] xtr4]) - apply (rule mult_mono) - apply auto - done - -lemma less_le_mult': "w * c < b * c \ 0 \ c \ (w + 1) * c \ b * c" - for b c w :: int - apply (rule mult_right_mono) - apply (rule zless_imp_add1_zle) - apply (erule (1) mult_right_less_imp_less) - apply assumption - done - -lemma less_le_mult: "w * c < b * c \ 0 \ c \ w * c + c \ b * c" - for b c w :: int - using less_le_mult' [of w c b] by (simp add: algebra_simps) - -lemmas less_le_mult_minus = iffD2 [OF le_diff_eq less_le_mult, - simplified left_diff_distrib] - -lemma gen_minus: "0 < n \ f n = f (Suc (n - 1))" - by auto - -lemma mpl_lem: "j \ i \ k < j \ i - j + k < i" - for i j k :: nat - by arith - -lemma nonneg_mod_div: "0 \ a \ 0 \ b \ 0 \ (a mod b) \ 0 \ a div b" - for a b :: int - by (cases "b = 0") (auto intro: pos_imp_zdiv_nonneg_iff [THEN iffD2]) - -declare iszero_0 [intro] +lemma n_less_equal_power_2 [simp]: + "n < 2 ^ n" + by (induction n) simp_all lemma min_pm [simp]: "min a b + (a - b) = a" for a b :: nat by arith lemma min_pm1 [simp]: "a - b + min a b = a" for a b :: nat by arith lemma rev_min_pm [simp]: "min b a + (a - b) = a" for a b :: nat by arith lemma rev_min_pm1 [simp]: "a - b + min b a = a" for a b :: nat by arith lemma min_minus [simp]: "min m (m - k) = m - k" for m k :: nat by arith lemma min_minus' [simp]: "min (m - k) m = m - k" for m k :: nat by arith -lemmas m2pths = pos_mod_sign pos_mod_bound [OF zless2p] +lemma nat_min_simps: + "(a::nat) \ b \ min b a = a" + "a \ b \ min a b = a" + by simp_all + +lemma iszero_minus: + \iszero (- z) \ iszero z\ + by (simp add: iszero_def) + +lemma diff_le_eq': "a - b \ c \ a \ b + c" + for a b c :: int + by arith + +lemma zless2: "0 < (2 :: int)" + by (fact zero_less_numeral) + +lemma zless2p: "0 < (2 ^ n :: int)" + by arith + +lemma zle2p: "0 \ (2 ^ n :: int)" + by arith + +lemma ex_eq_or: "(\m. n = Suc m \ (m = k \ P m)) \ n = Suc k \ (\m. n = Suc m \ P m)" + by auto + +lemma power_minus_simp: "0 < n \ a ^ n = a * a ^ (n - 1)" + by (auto dest: gr0_implies_Suc) + +lemma n2s_ths: + \2 + n = Suc (Suc n)\ + \n + 2 = Suc (Suc n)\ + by (fact add_2_eq_Suc add_2_eq_Suc')+ + +lemma s2n_ths: + \Suc (Suc n) = 2 + n\ + \Suc (Suc n) = n + 2\ + by simp_all + +lemma gt_or_eq_0: "0 < y \ 0 = y" + for y :: nat + by arith + +lemmas nat_simps = diff_add_inverse2 diff_add_inverse + +lemmas nat_iffs = le_add1 le_add2 + +lemma sum_imp_diff: "j = k + i \ j - i = k" + for k :: nat + by simp + +lemma le_diff_eq': "a \ c - b \ b + a \ c" + for a b c :: int + by arith + +lemma less_diff_eq': "a < c - b \ b + a < c" + for a b c :: int + by arith + +lemma diff_less_eq': "a - b < c \ a < b + c" + for a b c :: int + by arith + +lemma axxbyy: "a + m + m = b + n + n \ a = 0 \ a = 1 \ b = 0 \ b = 1 \ a = b \ m = n" + for a b m n :: int + by arith + +lemmas zadd_diff_inverse = + trans [OF diff_add_cancel [symmetric] add.commute] + +lemmas add_diff_cancel2 = + add.commute [THEN diff_eq_eq [THEN iffD2]] + +lemmas mcl = mult_cancel_left [THEN iffD1, THEN make_pos_rule] + +lemma pl_pl_rels: "a + b = c + d \ a \ c \ b \ d \ a \ c \ b \ d" + for a b c d :: nat + by arith + +lemmas pl_pl_rels' = add.commute [THEN [2] trans, THEN pl_pl_rels] + +lemma minus_eq: "m - k = m \ k = 0 \ m = 0" + for k m :: nat + by arith + +lemma pl_pl_mm: "a + b = c + d \ a - c = d - b" + for a b c d :: nat + by arith + +lemmas pl_pl_mm' = add.commute [THEN [2] trans, THEN pl_pl_mm] + +lemma less_le_mult': "w * c < b * c \ 0 \ c \ (w + 1) * c \ b * c" + for b c w :: int + apply (rule mult_right_mono) + apply (rule zless_imp_add1_zle) + apply (erule (1) mult_right_less_imp_less) + apply assumption + done + +lemma less_le_mult: "w * c < b * c \ 0 \ c \ w * c + c \ b * c" + for b c w :: int + using less_le_mult' [of w c b] by (simp add: algebra_simps) + +lemmas less_le_mult_minus = iffD2 [OF le_diff_eq less_le_mult, + simplified left_diff_distrib] + +lemma gen_minus: "0 < n \ f n = f (Suc (n - 1))" + by auto + +lemma mpl_lem: "j \ i \ k < j \ i - j + k < i" + for i j k :: nat + by arith + +lemmas dme = div_mult_mod_eq +lemmas dtle = div_times_less_eq_dividend +lemmas th2 = order_trans [OF order_refl [THEN [2] mult_le_mono] div_times_less_eq_dividend] + +lemma td_gal: "0 < c \ a \ b * c \ a div c \ b" + for a b c :: nat + apply safe + apply (erule (1) xtrans(4) [OF div_le_mono div_mult_self_is_m]) + apply (erule th2) + done + +lemmas td_gal_lt = td_gal [simplified not_less [symmetric], simplified] + +lemmas sdl = div_nat_eqI + +lemma given_quot: "f > 0 \ (f * l + (f - 1)) div f = l" + for f l :: nat + by (rule div_nat_eqI) (simp_all) + +lemma given_quot_alt: "f > 0 \ (l * f + f - Suc 0) div f = l" + for f l :: nat + apply (frule given_quot) + apply (rule trans) + prefer 2 + apply (erule asm_rl) + apply (rule_tac f="\n. n div f" in arg_cong) + apply (simp add : ac_simps) + done + +lemma nat_less_power_trans: + fixes n :: nat + assumes nv: "n < 2 ^ (m - k)" + and kv: "k \ m" + shows "2 ^ k * n < 2 ^ m" +proof (rule order_less_le_trans) + show "2 ^ k * n < 2 ^ k * 2 ^ (m - k)" + by (rule mult_less_mono2 [OF nv zero_less_power]) simp + + show "(2::nat) ^ k * 2 ^ (m - k) \ 2 ^ m" using nv kv + by (subst power_add [symmetric]) simp +qed + +lemma nat_le_power_trans: + fixes n :: nat + shows "\n \ 2 ^ (m - k); k \ m\ \ 2 ^ k * n \ 2 ^ m" + by (metis le_add_diff_inverse mult_le_mono2 semiring_normalization_rules(26)) + +lemma x_power_minus_1: + fixes x :: "'a :: {ab_group_add, power, numeral, one}" + shows "x + (2::'a) ^ n - (1::'a) = x + (2 ^ n - 1)" by simp + +lemma nat_diff_add: + fixes i :: nat + shows "\ i + j = k \ \ i = k - j" + by arith + +lemma pow_2_gt: "n \ 2 \ (2::int) < 2 ^ n" + by (induct n) auto + +lemma sum_to_zero: + "(a :: 'a :: ring) + b = 0 \ a = (- b)" + by (drule arg_cong[where f="\ x. x - a"], simp) + +lemma arith_is_1: + "\ x \ Suc 0; x > 0 \ \ x = 1" + by arith + +lemma suc_le_pow_2: + "1 < (n::nat) \ Suc n < 2 ^ n" + by (induct n; clarsimp) + (case_tac "n = 1"; clarsimp) + +lemma nat_le_Suc_less_imp: + "x < y \ x \ y - Suc 0" + by arith + +lemma power_sub_int: + "\ m \ n; 0 < b \ \ b ^ n div b ^ m = (b ^ (n - m) :: int)" + apply (subgoal_tac "\n'. n = m + n'") + apply (clarsimp simp: power_add) + apply (rule exI[where x="n - m"]) + apply simp + done + +lemma nat_Suc_less_le_imp: + "(k::nat) < Suc n \ k \ n" + by auto + +lemma nat_add_less_by_max: + "\ (x::nat) \ xmax ; y < k - xmax \ \ x + y < k" + by simp + +lemma nat_le_Suc_less: + "0 < y \ (x \ y - Suc 0) = (x < y)" + by arith + +lemma nat_power_minus_less: + "a < 2 ^ (x - n) \ (a :: nat) < 2 ^ x" + by (erule order_less_le_trans) simp end diff --git a/thys/Word_Lib/More_Divides.thy b/thys/Word_Lib/More_Divides.thy --- a/thys/Word_Lib/More_Divides.thy +++ b/thys/Word_Lib/More_Divides.thy @@ -1,55 +1,269 @@ -(* - * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) - * - * SPDX-License-Identifier: BSD-2-Clause - *) -section "More Lemmas on Division" +section \Lemmas on division\ theory More_Divides -imports Main + imports + "HOL-Library.Word" begin -lemma div_mult_le: - \a div b * b \ a\ for a b :: nat - by (fact div_times_less_eq_dividend) +declare div_eq_dividend_iff [simp] + +lemma int_div_same_is_1 [simp]: + \a div b = a \ b = 1\ if \0 < a\ for a b :: int + using that by (metis div_by_1 abs_ge_zero abs_of_pos int_div_less_self neq_iff + nonneg1_imp_zdiv_pos_iff zabs_less_one_iff) + +lemma int_div_minus_is_minus1 [simp]: + \a div b = - a \ b = - 1\ if \0 > a\ for a b :: int + using that by (metis div_minus_right equation_minus_iff int_div_same_is_1 neg_0_less_iff_less) lemma diff_mod_le: \a - a mod b \ d - b\ if \a < d\ \b dvd d\ for a b d :: nat using that apply(subst minus_mod_eq_mult_div) apply(clarsimp simp: dvd_def) apply(cases \b = 0\) apply simp apply(subgoal_tac "a div b \ k - 1") prefer 2 apply(subgoal_tac "a div b < k") apply(simp add: less_Suc_eq_le [symmetric]) apply(subgoal_tac "b * (a div b) < b * ((b * k) div b)") apply clarsimp apply(subst div_mult_self1_is_m) apply arith apply(rule le_less_trans) apply simp apply(subst mult.commute) - apply(rule div_mult_le) + apply(rule div_times_less_eq_dividend) apply assumption apply clarsimp apply(subgoal_tac "b * (a div b) \ b * (k - 1)") apply(erule le_trans) apply(simp add: diff_mult_distrib2) apply simp done -lemma int_div_same_is_1 [simp]: - \a div b = a \ b = 1\ if \0 < a\ for a b :: int - using that by (metis div_by_1 abs_ge_zero abs_of_pos int_div_less_self neq_iff - nonneg1_imp_zdiv_pos_iff zabs_less_one_iff) +lemma one_mod_exp_eq_one [simp]: + "1 mod (2 * 2 ^ n) = (1::int)" + using power_gt1 [of 2 n] by (auto intro: mod_pos_pos_trivial) -lemma int_div_minus_is_minus1 [simp]: - \a div b = - a \ b = - 1\ if \0 > a\ for a b :: int - using that by (metis div_minus_right equation_minus_iff int_div_same_is_1 neg_0_less_iff_less) +lemma int_mod_lem: "0 < n \ 0 \ b \ b < n \ b mod n = b" + for b n :: int + apply safe + apply (erule (1) mod_pos_pos_trivial) + apply (erule_tac [!] subst) + apply auto + done -declare div_eq_dividend_iff [simp] +lemma int_mod_ge': "b < 0 \ 0 < n \ b + n \ b mod n" + for b n :: int + by (metis add_less_same_cancel2 int_mod_ge mod_add_self2) + +lemma int_mod_le': "0 \ b - n \ b mod n \ b - n" + for b n :: int + by (metis minus_mod_self2 zmod_le_nonneg_dividend) + +lemma emep1: "even n \ even d \ 0 \ d \ (n + 1) mod d = (n mod d) + 1" + for n d :: int + by (auto simp add: pos_zmod_mult_2 add.commute dvd_def) + +lemma m1mod2k: "- 1 mod 2 ^ n = (2 ^ n - 1 :: int)" + by (rule zmod_minus1) simp + +lemma sb_inc_lem: "a + 2^k < 0 \ a + 2^k + 2^(Suc k) \ (a + 2^k) mod 2^(Suc k)" + for a :: int + using int_mod_ge' [where n = "2 ^ (Suc k)" and b = "a + 2 ^ k"] + by simp + +lemma sb_inc_lem': "a < - (2^k) \ a + 2^k + 2^(Suc k) \ (a + 2^k) mod 2^(Suc k)" + for a :: int + by (rule sb_inc_lem) simp + +lemma sb_dec_lem: "0 \ - (2 ^ k) + a \ (a + 2 ^ k) mod (2 * 2 ^ k) \ - (2 ^ k) + a" + for a :: int + using int_mod_le'[where n = "2 ^ (Suc k)" and b = "a + 2 ^ k"] by simp + +lemma sb_dec_lem': "2 ^ k \ a \ (a + 2 ^ k) mod (2 * 2 ^ k) \ - (2 ^ k) + a" + for a :: int + by (rule sb_dec_lem) simp + +lemma mod_2_neq_1_eq_eq_0: "k mod 2 \ 1 \ k mod 2 = 0" + for k :: int + by (fact not_mod_2_eq_1_eq_0) + +lemma z1pmod2: "(2 * b + 1) mod 2 = (1::int)" + for b :: int + by arith + +lemma p1mod22k': "(1 + 2 * b) mod (2 * 2 ^ n) = 1 + 2 * (b mod 2 ^ n)" + for b :: int + by (rule pos_zmod_mult_2) simp + +lemma p1mod22k: "(2 * b + 1) mod (2 * 2 ^ n) = 2 * (b mod 2 ^ n) + 1" + for b :: int + by (simp add: p1mod22k' add.commute) + +lemma pos_mod_sign2: + \0 \ a mod 2\ for a :: int + by simp + +lemma pos_mod_bound2: + \a mod 2 < 2\ for a :: int + by simp + +lemma nmod2: "n mod 2 = 0 \ n mod 2 = 1" + for n :: int + by arith + +lemma eme1p: + "even n \ even d \ 0 \ d \ (1 + n) mod d = 1 + n mod d" for n d :: int + using emep1 [of n d] by (simp add: ac_simps) + +lemma m1mod22k: + \- 1 mod (2 * 2 ^ n) = 2 * 2 ^ n - (1::int)\ + by (simp add: zmod_minus1) + +lemma z1pdiv2: "(2 * b + 1) div 2 = b" + for b :: int + by arith + +lemma zdiv_le_dividend: + \0 \ a \ 0 < b \ a div b \ a\ for a b :: int + by (metis div_by_1 int_one_le_iff_zero_less zdiv_mono2 zero_less_one) + +lemma axxmod2: "(1 + x + x) mod 2 = 1 \ (0 + x + x) mod 2 = 0" + for x :: int + by arith + +lemma axxdiv2: "(1 + x + x) div 2 = x \ (0 + x + x) div 2 = x" + for x :: int + by arith + +lemmas rdmods = + mod_minus_eq [symmetric] + mod_diff_left_eq [symmetric] + mod_diff_right_eq [symmetric] + mod_add_left_eq [symmetric] + mod_add_right_eq [symmetric] + mod_mult_right_eq [symmetric] + mod_mult_left_eq [symmetric] + +lemma mod_plus_right: "(a + x) mod m = (b + x) mod m \ a mod m = b mod m" + for a b m x :: nat + by (induct x) (simp_all add: mod_Suc, arith) + +lemma nat_minus_mod: "(n - n mod m) mod m = 0" + for m n :: nat + by (induct n) (simp_all add: mod_Suc) + +lemmas nat_minus_mod_plus_right = + trans [OF nat_minus_mod mod_0 [symmetric], + THEN mod_plus_right [THEN iffD2], simplified] + +lemmas push_mods' = mod_add_eq + mod_mult_eq mod_diff_eq + mod_minus_eq + +lemmas push_mods = push_mods' [THEN eq_reflection] +lemmas pull_mods = push_mods [symmetric] rdmods [THEN eq_reflection] + +lemma nat_mod_eq: "b < n \ a mod n = b mod n \ a mod n = b" + for a b n :: nat + by (induct a) auto + +lemmas nat_mod_eq' = refl [THEN [2] nat_mod_eq] + +lemma nat_mod_lem: "0 < n \ b < n \ b mod n = b" + for b n :: nat + apply safe + apply (erule nat_mod_eq') + apply (erule subst) + apply (erule mod_less_divisor) + done + +lemma mod_nat_add: "x < z \ y < z \ (x + y) mod z = (if x + y < z then x + y else x + y - z)" + for x y z :: nat + apply (rule nat_mod_eq) + apply auto + apply (rule trans) + apply (rule le_mod_geq) + apply simp + apply (rule nat_mod_eq') + apply arith + done + +lemma mod_nat_sub: "x < z \ (x - y) mod z = x - y" + for x y :: nat + by (rule nat_mod_eq') arith + +lemma int_mod_eq: "0 \ b \ b < n \ a mod n = b mod n \ a mod n = b" + for a b n :: int + by (metis mod_pos_pos_trivial) + +lemma zmde: + \b * (a div b) = a - a mod b\ for a b :: \'a::{group_add,semiring_modulo}\ + using mult_div_mod_eq [of b a] by (simp add: eq_diff_eq) + +(* already have this for naturals, div_mult_self1/2, but not for ints *) +lemma zdiv_mult_self: "m \ 0 \ (a + m * n) div m = a div m + n" + for a m n :: int + by simp + +lemma mod_power_lem: "a > 1 \ a ^ n mod a ^ m = (if m \ n then 0 else a ^ n)" + for a :: int + by (simp add: mod_eq_0_iff_dvd le_imp_power_dvd) + +lemma nonneg_mod_div: "0 \ a \ 0 \ b \ 0 \ (a mod b) \ 0 \ a div b" + for a b :: int + by (cases "b = 0") (auto intro: pos_imp_zdiv_nonneg_iff [THEN iffD2]) + +lemma mod_exp_less_eq_exp: + \a mod 2 ^ n < 2 ^ n\ for a :: int + by (rule pos_mod_bound) simp + +lemma div_mult_le: + \a div b * b \ a\ for a b :: nat + by (fact div_times_less_eq_dividend) + +lemma power_sub: + fixes a :: nat + assumes lt: "n \ m" + and av: "0 < a" + shows "a ^ (m - n) = a ^ m div a ^ n" +proof (subst nat_mult_eq_cancel1 [symmetric]) + show "(0::nat) < a ^ n" using av by simp +next + from lt obtain q where mv: "n + q = m" + by (auto simp: le_iff_add) + + have "a ^ n * (a ^ m div a ^ n) = a ^ m" + proof (subst mult.commute) + have "a ^ m = (a ^ m div a ^ n) * a ^ n + a ^ m mod a ^ n" + by (rule div_mult_mod_eq [symmetric]) + + moreover have "a ^ m mod a ^ n = 0" + by (subst mod_eq_0_iff_dvd, subst dvd_def, rule exI [where x = "a ^ q"], + (subst power_add [symmetric] mv)+, rule refl) + + ultimately show "(a ^ m div a ^ n) * a ^ n = a ^ m" by simp + qed + + then show "a ^ n * a ^ (m - n) = a ^ n * (a ^ m div a ^ n)" using lt + by (simp add: power_add [symmetric]) +qed + +lemma mod_lemma: "[| (0::nat) < c; r < b |] ==> b * (q mod c) + r < b * c" + apply (cut_tac m = q and n = c in mod_less_divisor) + apply (drule_tac [2] m = "q mod c" in less_imp_Suc_add, auto) + apply (erule_tac P = "%x. lhs < rhs x" for lhs rhs in ssubst) + apply (simp add: add_mult_distrib2) + done + +lemmas m2pths = pos_mod_sign mod_exp_less_eq_exp + +lemmas int_mod_eq' = mod_pos_pos_trivial (* FIXME delete *) + +lemmas int_mod_le = zmod_le_nonneg_dividend (* FIXME: delete *) end diff --git a/thys/Word_Lib/More_List.thy b/thys/Word_Lib/More_List.thy --- a/thys/Word_Lib/More_List.thy +++ b/thys/Word_Lib/More_List.thy @@ -1,22 +1,183 @@ -(* Author: Jeremy Dawson, NICTA -*) section \Lemmas on list operations\ theory More_List - imports Main + imports + Main + "HOL-Library.Sublist" begin +lemma replicate_numeral [simp]: "replicate (numeral k) x = x # replicate (pred_numeral k) x" + by (simp add: numeral_eq_Suc) + +lemma list_exhaust_size_gt0: + assumes "\a list. y = a # list \ P" + shows "0 < length y \ P" + apply (cases y) + apply simp + apply (rule assms) + apply fastforce + done + +lemma list_exhaust_size_eq0: + assumes "y = [] \ P" + shows "length y = 0 \ P" + apply (cases y) + apply (rule assms) + apply simp + apply simp + done + +lemma size_Cons_lem_eq: "y = xa # list \ size y = Suc k \ size list = k" + by auto + +lemma takeWhile_take_has_property: + "n \ length (takeWhile P xs) \ \x \ set (take n xs). P x" + by (induct xs arbitrary: n; simp split: if_split_asm) (case_tac n, simp_all) + +lemma takeWhile_take_has_property_nth: + "\ n < length (takeWhile P xs) \ \ P (xs ! n)" + by (induct xs arbitrary: n; simp split: if_split_asm) (case_tac n, simp_all) + +lemma takeWhile_replicate: + "takeWhile f (replicate len x) = (if f x then replicate len x else [])" + by (induct_tac len) auto + +lemma takeWhile_replicate_empty: + "\ f x \ takeWhile f (replicate len x) = []" + by (simp add: takeWhile_replicate) + +lemma takeWhile_replicate_id: + "f x \ takeWhile f (replicate len x) = replicate len x" + by (simp add: takeWhile_replicate) + lemma butlast_power: "(butlast ^^ n) bl = take (length bl - n) bl" by (induct n) (auto simp: butlast_take) lemma nth_rev: "n < length xs \ rev xs ! n = xs ! (length xs - 1 - n)" using rev_nth by simp lemma nth_rev_alt: "n < length ys \ ys ! n = rev ys ! (length ys - Suc n)" by (simp add: nth_rev) lemma hd_butlast: "length xs > 1 \ hd (butlast xs) = hd xs" by (cases xs) auto +lemma upt_add_eq_append': + assumes "i \ j" and "j \ k" + shows "[i.. [0 ..< m] = [0 ..< n] @ [n] @ [Suc n ..< m]" + by (metis append_Cons append_Nil less_Suc_eq_le less_imp_le_nat upt_add_eq_append' + upt_rec zero_less_Suc) + +lemma length_takeWhile_less: + "\x\set xs. \ P x \ length (takeWhile P xs) < length xs" + by (induct xs) (auto split: if_splits) + +lemma drop_eq_mono: + assumes le: "m \ n" + assumes drop: "drop m xs = drop m ys" + shows "drop n xs = drop n ys" +proof - + have ex: "\p. n = p + m" by (rule exI[of _ "n - m"]) (simp add: le) + then obtain p where p: "n = p + m" by blast + show ?thesis unfolding p drop_drop[symmetric] drop by simp +qed + +lemma drop_Suc_nth: + "n < length xs \ drop n xs = xs!n # drop (Suc n) xs" + by (simp add: Cons_nth_drop_Suc) + +lemma and_len: "xs = ys \ xs = ys \ length xs = length ys" + by auto + +lemma tl_if: "tl (if p then xs else ys) = (if p then tl xs else tl ys)" + by auto + +lemma hd_if: "hd (if p then xs else ys) = (if p then hd xs else hd ys)" + by auto + +lemma if_single: "(if xc then [xab] else [an]) = [if xc then xab else an]" + by auto + +\ \note -- \if_Cons\ can cause blowup in the size, if \p\ is complex, so make a simproc\ +lemma if_Cons: "(if p then x # xs else y # ys) = If p x y # If p xs ys" + by auto + +lemma same_length_is_parallel: + assumes len: "\y \ set as. length y = x" + shows "\x \ set as. \y \ set as - {x}. x \ y" +proof (rule, rule) + fix x y + assume xi: "x \ set as" and yi: "y \ set as - {x}" + from len obtain q where len': "\y \ set as. length y = q" .. + + show "x \ y" + proof (rule not_equal_is_parallel) + from xi yi show "x \ y" by auto + from xi yi len' show "length x = length y" by (auto dest: bspec) + qed +qed + +lemma sublist_equal_part: + "prefix xs ys \ take (length xs) ys = xs" + by (clarsimp simp: prefix_def) + +lemma prefix_length_less: + "strict_prefix xs ys \ length xs < length ys" + apply (clarsimp simp: strict_prefix_def) + apply (frule prefix_length_le) + apply (rule ccontr, simp) + apply (clarsimp simp: prefix_def) + done + +lemmas take_less = take_strict_prefix + +lemma not_prefix_longer: + "\ length xs > length ys \ \ \ prefix xs ys" + by (clarsimp dest!: prefix_length_le) + +lemma map_prefixI: + "prefix xs ys \ prefix (map f xs) (map f ys)" + by (clarsimp simp: prefix_def) + +lemma list_all2_induct_suffixeq [consumes 1, case_names Nil Cons]: + assumes lall: "list_all2 Q as bs" + and nilr: "P [] []" + and consr: "\x xs y ys. + \list_all2 Q xs ys; Q x y; P xs ys; suffix (x # xs) as; suffix (y # ys) bs\ + \ P (x # xs) (y # ys)" + shows "P as bs" +proof - + define as' where "as' == as" + define bs' where "bs' == bs" + + have "suffix as as' \ suffix bs bs'" unfolding as'_def bs'_def by simp + then show ?thesis using lall + proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) + case 1 show ?case by fact + next + case (2 x xs y ys) + + show ?case + proof (rule consr) + from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all + then show "P xs ys" using "2.hyps" "2.prems" by (auto dest: suffix_ConsD) + from "2.prems" show "suffix (x # xs) as" and "suffix (y # ys) bs" + by (auto simp: as'_def bs'_def) + qed + qed +qed + +lemma take_prefix: + "(take (length xs) ys = xs) = prefix xs ys" +proof (induct xs arbitrary: ys) + case Nil then show ?case by simp +next + case Cons then show ?case by (cases ys) auto +qed + end diff --git a/thys/Word_Lib/More_Misc.thy b/thys/Word_Lib/More_Misc.thy --- a/thys/Word_Lib/More_Misc.thy +++ b/thys/Word_Lib/More_Misc.thy @@ -1,229 +1,87 @@ -(* - * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) - * - * SPDX-License-Identifier: BSD-2-Clause - *) -section "Generic Lemmas used in the Word Library" +section \Miscellaneous lemmas\ theory More_Misc imports Main begin -definition - strict_part_mono :: "'a set \ ('a :: order \ 'b :: order) \ bool" where - "strict_part_mono S f \ \A\S. \B\S. A < B \ f A < f B" - -lemma strict_part_mono_by_steps: - "strict_part_mono {..n :: nat} f = (n \ 0 \ f (n - 1) < f n \ strict_part_mono {.. n - 1} f)" - apply (cases n; simp add: strict_part_mono_def) - apply (safe; clarsimp) - apply (case_tac "B = Suc nat"; simp) - apply (case_tac "A = nat"; clarsimp) - apply (erule order_less_trans [rotated]) - apply simp - done - -lemma strict_part_mono_singleton[simp]: - "strict_part_mono {x} f" - by (simp add: strict_part_mono_def) - -lemma strict_part_mono_lt: - "\ x < f 0; strict_part_mono {.. n :: nat} f \ \ \m \ n. x < f m" - by (metis atMost_iff le_0_eq le_cases neq0_conv order.strict_trans strict_part_mono_def) - -lemma strict_part_mono_reverseE: - "\ f n \ f m; strict_part_mono {.. N :: nat} f; n \ N \ \ n \ m" - by (rule ccontr) (fastforce simp: linorder_not_le strict_part_mono_def) - -lemma takeWhile_take_has_property: - "n \ length (takeWhile P xs) \ \x \ set (take n xs). P x" - by (induct xs arbitrary: n; simp split: if_split_asm) (case_tac n, simp_all) - -lemma takeWhile_take_has_property_nth: - "\ n < length (takeWhile P xs) \ \ P (xs ! n)" - by (induct xs arbitrary: n; simp split: if_split_asm) (case_tac n, simp_all) +lemma power_numeral: "a ^ numeral k = a * a ^ (pred_numeral k)" + by (simp add: numeral_eq_Suc) -lemma takeWhile_replicate: - "takeWhile f (replicate len x) = (if f x then replicate len x else [])" - by (induct_tac len) auto - -lemma takeWhile_replicate_empty: - "\ f x \ takeWhile f (replicate len x) = []" - by (simp add: takeWhile_replicate) - -lemma takeWhile_replicate_id: - "f x \ takeWhile f (replicate len x) = replicate len x" - by (simp add: takeWhile_replicate) +lemma funpow_numeral [simp]: "f ^^ numeral k = f \ f ^^ (pred_numeral k)" + by (simp add: numeral_eq_Suc) -lemma power_sub: - fixes a :: nat - assumes lt: "n \ m" - and av: "0 < a" - shows "a ^ (m - n) = a ^ m div a ^ n" -proof (subst nat_mult_eq_cancel1 [symmetric]) - show "(0::nat) < a ^ n" using av by simp -next - from lt obtain q where mv: "n + q = m" - by (auto simp: le_iff_add) +lemma funpow_minus_simp: "0 < n \ f ^^ n = f \ f ^^ (n - 1)" + by (auto dest: gr0_implies_Suc) - have "a ^ n * (a ^ m div a ^ n) = a ^ m" - proof (subst mult.commute) - have "a ^ m = (a ^ m div a ^ n) * a ^ n + a ^ m mod a ^ n" - by (rule div_mult_mod_eq [symmetric]) - - moreover have "a ^ m mod a ^ n = 0" - by (subst mod_eq_0_iff_dvd, subst dvd_def, rule exI [where x = "a ^ q"], - (subst power_add [symmetric] mv)+, rule refl) - - ultimately show "(a ^ m div a ^ n) * a ^ n = a ^ m" by simp - qed - - then show "a ^ n * a ^ (m - n) = a ^ n * (a ^ m div a ^ n)" using lt - by (simp add: power_add [symmetric]) -qed - +lemma rco_alt: "(f \ g) ^^ n \ f = f \ (g \ f) ^^ n" + apply (rule ext) + apply (induct n) + apply (simp_all add: o_def) + done lemma union_sub: "\B \ A; C \ B\ \ (A - B) \ (B - C) = (A - C)" by fastforce lemma insert_sub: "x \ xs \ (insert x (xs - ys)) = (xs - (ys - {x}))" by blast lemma ran_upd: "\ inj_on f (dom f); f y = Some z \ \ ran (\x. if x = y then None else f x) = ran f - {z}" unfolding ran_def apply (rule set_eqI) apply simp by (metis domI inj_on_eq_iff option.sel) -lemma nat_less_power_trans: - fixes n :: nat - assumes nv: "n < 2 ^ (m - k)" - and kv: "k \ m" - shows "2 ^ k * n < 2 ^ m" -proof (rule order_less_le_trans) - show "2 ^ k * n < 2 ^ k * 2 ^ (m - k)" - by (rule mult_less_mono2 [OF nv zero_less_power]) simp - - show "(2::nat) ^ k * 2 ^ (m - k) \ 2 ^ m" using nv kv - by (subst power_add [symmetric]) simp -qed - -lemma nat_le_power_trans: - fixes n :: nat - shows "\n \ 2 ^ (m - k); k \ m\ \ 2 ^ k * n \ 2 ^ m" - by (metis le_add_diff_inverse mult_le_mono2 semiring_normalization_rules(26)) - -lemma x_power_minus_1: - fixes x :: "'a :: {ab_group_add, power, numeral, one}" - shows "x + (2::'a) ^ n - (1::'a) = x + (2 ^ n - 1)" by simp - -lemma nat_diff_add: - fixes i :: nat - shows "\ i + j = k \ \ i = k - j" - by arith - -lemma pow_2_gt: "n \ 2 \ (2::int) < 2 ^ n" - by (induct n) auto - lemma if_apply_def2: "(if P then F else G) = (\x. (P \ F x) \ (\ P \ G x))" by simp lemma case_bool_If: "case_bool P Q b = (if b then P else Q)" by simp -lemma sum_to_zero: - "(a :: 'a :: ring) + b = 0 \ a = (- b)" - by (drule arg_cong[where f="\ x. x - a"], simp) - -lemma arith_is_1: - "\ x \ Suc 0; x > 0 \ \ x = 1" - by arith - lemma if_f: "(if a then f b else f c) = f (if a then b else c)" by simp -lemma upt_add_eq_append': - assumes "i \ j" and "j \ k" - shows "[i.. [0 ..< m] = [0 ..< n] @ [n] @ [Suc n ..< m]" - by (metis append_Cons append_Nil less_Suc_eq_le less_imp_le_nat upt_add_eq_append' - upt_rec zero_less_Suc) +lemmas ls_splits = prod.split prod.split_asm if_split_asm -lemma drop_Suc_nth: - "n < length xs \ drop n xs = xs!n # drop (Suc n) xs" - by (simp add: Cons_nth_drop_Suc) +lemma size_if: "size (if p then xs else ys) = (if p then size xs else size ys)" + by (fact if_distrib) -lemma n_less_equal_power_2 [simp]: - "n < 2 ^ n" - by (induct n; simp) - -lemma nat_min_simps [simp]: - "(a::nat) \ b \ min b a = a" - "a \ b \ min a b = a" +lemma if_Not_x: "(if p then \ x else x) = (p = (\ x))" by auto -lemma power_sub_int: - "\ m \ n; 0 < b \ \ b ^ n div b ^ m = (b ^ (n - m) :: int)" - apply (subgoal_tac "\n'. n = m + n'") - apply (clarsimp simp: power_add) - apply (rule exI[where x="n - m"]) - apply simp - done - -lemma suc_le_pow_2: - "1 < (n::nat) \ Suc n < 2 ^ n" - by (induct n; clarsimp) - (case_tac "n = 1"; clarsimp) - -lemma nat_le_Suc_less_imp: - "x < y \ x \ y - Suc 0" - by arith - -lemma length_takeWhile_less: - "\x\set xs. \ P x \ length (takeWhile P xs) < length xs" - by (induct xs) (auto split: if_splits) - -lemma drop_eq_mono: - assumes le: "m \ n" - assumes drop: "drop m xs = drop m ys" - shows "drop n xs = drop n ys" -proof - - have ex: "\p. n = p + m" by (rule exI[of _ "n - m"]) (simp add: le) - then obtain p where p: "n = p + m" by blast - show ?thesis unfolding p drop_drop[symmetric] drop by simp -qed - -lemma nat_Suc_less_le_imp: - "(k::nat) < Suc n \ k \ n" +lemma if_x_Not: "(if p then x else \ x) = (p = x)" by auto -lemma nat_add_less_by_max: - "\ (x::nat) \ xmax ; y < k - xmax \ \ x + y < k" +lemma if_same_and: "(If p x y \ If p u v) = (if p then x \ u else y \ v)" + by auto + +lemma if_same_eq: "(If p x y = (If p u v)) = (if p then x = u else y = v)" + by auto + +lemma if_same_eq_not: "(If p x y = (\ If p u v)) = (if p then x = (\ u) else y = (\ v))" + by auto + +lemma the_elemI: "y = {x} \ the_elem y = x" by simp -lemma mod_lemma: "[| (0::nat) < c; r < b |] ==> b * (q mod c) + r < b * c" - apply (cut_tac m = q and n = c in mod_less_divisor) - apply (drule_tac [2] m = "q mod c" in less_imp_Suc_add, auto) - apply (erule_tac P = "%x. lhs < rhs x" for lhs rhs in ssubst) - apply (simp add: add_mult_distrib2) - done +lemma nonemptyE: "S \ {} \ (\x. x \ S \ R) \ R" + by auto -lemma nat_le_Suc_less: - "0 < y \ (x \ y - Suc 0) = (x < y)" - by arith +lemmas xtr1 = xtrans(1) +lemmas xtr2 = xtrans(2) +lemmas xtr3 = xtrans(3) +lemmas xtr4 = xtrans(4) +lemmas xtr5 = xtrans(5) +lemmas xtr6 = xtrans(6) +lemmas xtr7 = xtrans(7) +lemmas xtr8 = xtrans(8) -lemma nat_power_minus_less: - "a < 2 ^ (x - n) \ (a :: nat) < 2 ^ x" - by (erule order_less_le_trans) simp +lemmas if_fun_split = if_apply_def2 end diff --git a/thys/Word_Lib/Strict_part_mono.thy b/thys/Word_Lib/Strict_part_mono.thy new file mode 100644 --- /dev/null +++ b/thys/Word_Lib/Strict_part_mono.thy @@ -0,0 +1,51 @@ + +theory Strict_part_mono + imports "HOL-Library.Word" Word_Lib +begin + +definition + strict_part_mono :: "'a set \ ('a :: order \ 'b :: order) \ bool" where + "strict_part_mono S f \ \A\S. \B\S. A < B \ f A < f B" + +lemma strict_part_mono_by_steps: + "strict_part_mono {..n :: nat} f = (n \ 0 \ f (n - 1) < f n \ strict_part_mono {.. n - 1} f)" + apply (cases n; simp add: strict_part_mono_def) + apply (safe; clarsimp) + apply (case_tac "B = Suc nat"; simp) + apply (case_tac "A = nat"; clarsimp) + apply (erule order_less_trans [rotated]) + apply simp + done + +lemma strict_part_mono_singleton[simp]: + "strict_part_mono {x} f" + by (simp add: strict_part_mono_def) + +lemma strict_part_mono_lt: + "\ x < f 0; strict_part_mono {.. n :: nat} f \ \ \m \ n. x < f m" + by (metis atMost_iff le_0_eq le_cases neq0_conv order.strict_trans strict_part_mono_def) + +lemma strict_part_mono_reverseE: + "\ f n \ f m; strict_part_mono {.. N :: nat} f; n \ N \ \ n \ m" + by (rule ccontr) (fastforce simp: linorder_not_le strict_part_mono_def) + +lemma two_power_strict_part_mono: + "strict_part_mono {..LENGTH('a) - 1} (\x. (2 :: 'a :: len word) ^ x)" +proof - + { fix n + have "n < LENGTH('a) \ strict_part_mono {..n} (\x. (2 :: 'a :: len word) ^ x)" + proof (induct n) + case 0 then show ?case by simp + next + case (Suc n) + from Suc.prems + have "2 ^ n < (2 :: 'a :: len word) ^ Suc n" + using power_strict_increasing unat_power_lower word_less_nat_alt by fastforce + with Suc + show ?case by (subst strict_part_mono_by_steps) simp + qed + } + then show ?thesis by simp +qed + +end diff --git a/thys/Word_Lib/Word_Lemmas.thy b/thys/Word_Lib/Word_Lemmas.thy --- a/thys/Word_Lib/Word_Lemmas.thy +++ b/thys/Word_Lib/Word_Lemmas.thy @@ -1,6164 +1,6069 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Lemmas with Generic Word Length" theory Word_Lemmas imports Type_Syntax Next_and_Prev Enumeration_Word - More_Divides Word_EqI Rsplit - "HOL-Library.Sublist" + More_Divides + More_List begin lemma nat_mask_eq: \nat (mask n) = mask n\ by (simp add: nat_eq_iff of_nat_mask_eq) lemma unat_mask_eq: \unat (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ by transfer (simp add: nat_mask_eq) lemma word_plus_mono_left: fixes x :: "'a :: len word" shows "\y \ z; x \ x + z\ \ y + x \ z + x" by unat_arith lemma word_shiftl_add_distrib: fixes x :: "'a :: len word" shows "(x + y) << n = (x << n) + (y << n)" by (simp add: shiftl_t2n ring_distribs) lemma less_Suc_unat_less_bound: "n < Suc (unat (x :: 'a :: len word)) \ n < 2 ^ LENGTH('a)" by (auto elim!: order_less_le_trans intro: Suc_leI) lemma up_ucast_inj: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ len_of TYPE ('b) \ \ x = (y::'a::len word)" by (subst (asm) bang_eq) (fastforce simp: nth_ucast word_size intro: word_eqI) lemmas ucast_up_inj = up_ucast_inj lemma up_ucast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (ucast x = (ucast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_ucast_inj) lemma no_plus_overflow_neg: "(x :: 'a :: len word) < -y \ x \ x + y" by (metis diff_minus_eq_add less_imp_le sub_wrap_lt) lemma ucast_ucast_eq: "\ ucast x = (ucast (ucast y::'a word)::'c::len word); LENGTH('a) \ LENGTH('b); LENGTH('b) \ LENGTH('c) \ \ x = ucast y" for x :: "'a::len word" and y :: "'b::len word" by (fastforce intro: word_eqI simp: bang_eq nth_ucast word_size) lemma ucast_0_I: "x = 0 \ ucast x = 0" by simp text \Lemmas about words\ lemmas and_bang = word_and_nth lemma of_drop_to_bl: "of_bl (drop n (to_bl x)) = (x && mask (size x - n))" by (simp add: of_bl_drop word_size_bl) lemma word_add_offset_less: fixes x :: "'a :: len word" assumes yv: "y < 2 ^ n" and xv: "x < 2 ^ m" and mnv: "sz < LENGTH('a :: len)" and xv': "x < 2 ^ (LENGTH('a :: len) - n)" and mn: "sz = m + n" shows "x * 2 ^ n + y < 2 ^ sz" proof (subst mn) from mnv mn have nv: "n < LENGTH('a)" and mv: "m < LENGTH('a)" by auto have uy: "unat y < 2 ^ n" by (rule order_less_le_trans [OF unat_mono [OF yv] order_eq_refl], rule unat_power_lower[OF nv]) have ux: "unat x < 2 ^ m" by (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl], rule unat_power_lower[OF mv]) then show "x * 2 ^ n + y < 2 ^ (m + n)" using ux uy nv mnv xv' apply (subst word_less_nat_alt) apply (subst unat_word_ariths)+ apply (subst mod_less) apply simp apply (subst mult.commute) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv']]) apply (cases "n = 0"; simp) apply (subst unat_power_lower[OF nv]) apply (subst mod_less) apply (erule order_less_le_trans [OF nat_add_offset_less], assumption) apply (rule mn) apply simp apply (simp add: mn mnv) apply (erule nat_add_offset_less; simp) done qed lemma word_less_power_trans: fixes n :: "'a :: len word" assumes nv: "n < 2 ^ (m - k)" and kv: "k \ m" and mv: "m < len_of TYPE ('a)" shows "2 ^ k * n < 2 ^ m" using nv kv mv apply - apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply simp apply (rule nat_less_power_trans) apply (erule order_less_trans [OF unat_mono]) apply simp apply simp apply simp apply (rule nat_less_power_trans) apply (subst unat_power_lower[where 'a = 'a, symmetric]) apply simp apply (erule unat_mono) apply simp done lemma Suc_unat_diff_1: fixes x :: "'a :: len word" assumes lt: "1 \ x" shows "Suc (unat (x - 1)) = unat x" proof - have "0 < unat x" by (rule order_less_le_trans [where y = 1], simp, subst unat_1 [symmetric], rule iffD1 [OF word_le_nat_alt lt]) then show ?thesis by ((subst unat_sub [OF lt])+, simp only: unat_1) qed lemma word_div_sub: fixes x :: "'a :: len word" assumes yx: "y \ x" and y0: "0 < y" shows "(x - y) div y = x div y - 1" apply (rule word_unat.Rep_eqD) apply (subst unat_div) apply (subst unat_sub [OF yx]) apply (subst unat_sub) apply (subst word_le_nat_alt) apply (subst unat_div) apply (subst le_div_geq) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply (subst word_le_nat_alt [symmetric], rule yx) apply simp apply (subst unat_div) apply (subst le_div_geq [OF _ iffD1 [OF word_le_nat_alt yx]]) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply simp done lemma word_mult_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k < j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto intro: order_less_subst2 simp: word_less_nat_alt elim: mult_less_mono1) then show ?thesis using ujk knz ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_less_dest: fixes i :: "'a :: len word" assumes ij: "i * k < j * k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem] elim: mult_less_mono1) lemma word_mult_less_cancel: fixes k :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" by (rule iffI [OF word_mult_less_dest [OF _ uik ujk] word_mult_less_mono1 [OF _ knz ujk]]) lemma Suc_div_unat_helper: assumes szv: "sz < LENGTH('a :: len)" and usszv: "us \ sz" shows "2 ^ (sz - us) = Suc (unat (((2::'a :: len word) ^ sz - 1) div 2 ^ us))" proof - note usv = order_le_less_trans [OF usszv szv] from usszv obtain q where qv: "sz = us + q" by (auto simp: le_iff_add) have "Suc (unat (((2:: 'a word) ^ sz - 1) div 2 ^ us)) = (2 ^ us + unat ((2:: 'a word) ^ sz - 1)) div 2 ^ us" apply (subst unat_div unat_power_lower[OF usv])+ apply (subst div_add_self1, simp+) done also have "\ = ((2 ^ us - 1) + 2 ^ sz) div 2 ^ us" using szv by (simp add: unat_minus_one) also have "\ = 2 ^ q + ((2 ^ us - 1) div 2 ^ us)" apply (subst qv) apply (subst power_add) apply (subst div_mult_self2; simp) done also have "\ = 2 ^ (sz - us)" using qv by simp finally show ?thesis .. qed lemma set_enum_word8_def: "(set enum::word8 set) = {0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27, 28, 29, 30, 31, 32, 33, 34, 35, 36, 37, 38, 39, 40, 41, 42, 43, 44, 45, 46, 47, 48, 49, 50, 51, 52, 53, 54, 55, 56, 57, 58, 59, 60, 61, 62, 63, 64, 65, 66, 67, 68, 69, 70, 71, 72, 73, 74, 75, 76, 77, 78, 79, 80, 81, 82, 83, 84, 85, 86, 87, 88, 89, 90, 91, 92, 93, 94, 95, 96, 97, 98, 99, 100, 101, 102, 103, 104, 105, 106, 107, 108, 109, 110, 111, 112, 113, 114, 115, 116, 117, 118, 119, 120, 121, 122, 123, 124, 125, 126, 127, 128, 129, 130, 131, 132, 133, 134, 135, 136, 137, 138, 139, 140, 141, 142, 143, 144, 145, 146, 147, 148, 149, 150, 151, 152, 153, 154, 155, 156, 157, 158, 159, 160, 161, 162, 163, 164, 165, 166, 167, 168, 169, 170, 171, 172, 173, 174, 175, 176, 177, 178, 179, 180, 181, 182, 183, 184, 185, 186, 187, 188, 189, 190, 191, 192, 193, 194, 195, 196, 197, 198, 199, 200, 201, 202, 203, 204, 205, 206, 207, 208, 209, 210, 211, 212, 213, 214, 215, 216, 217, 218, 219, 220, 221, 222, 223, 224, 225, 226, 227, 228, 229, 230, 231, 232, 233, 234, 235, 236, 237, 238, 239, 240, 241, 242, 243, 244, 245, 246, 247, 248, 249, 250, 251, 252, 253, 254, 255}" by eval lemma set_strip_insert: "\ x \ insert a S; x \ a \ \ x \ S" by simp lemma word8_exhaust: fixes x :: word8 shows "\x \ 0; x \ 1; x \ 2; x \ 3; x \ 4; x \ 5; x \ 6; x \ 7; x \ 8; x \ 9; x \ 10; x \ 11; x \ 12; x \ 13; x \ 14; x \ 15; x \ 16; x \ 17; x \ 18; x \ 19; x \ 20; x \ 21; x \ 22; x \ 23; x \ 24; x \ 25; x \ 26; x \ 27; x \ 28; x \ 29; x \ 30; x \ 31; x \ 32; x \ 33; x \ 34; x \ 35; x \ 36; x \ 37; x \ 38; x \ 39; x \ 40; x \ 41; x \ 42; x \ 43; x \ 44; x \ 45; x \ 46; x \ 47; x \ 48; x \ 49; x \ 50; x \ 51; x \ 52; x \ 53; x \ 54; x \ 55; x \ 56; x \ 57; x \ 58; x \ 59; x \ 60; x \ 61; x \ 62; x \ 63; x \ 64; x \ 65; x \ 66; x \ 67; x \ 68; x \ 69; x \ 70; x \ 71; x \ 72; x \ 73; x \ 74; x \ 75; x \ 76; x \ 77; x \ 78; x \ 79; x \ 80; x \ 81; x \ 82; x \ 83; x \ 84; x \ 85; x \ 86; x \ 87; x \ 88; x \ 89; x \ 90; x \ 91; x \ 92; x \ 93; x \ 94; x \ 95; x \ 96; x \ 97; x \ 98; x \ 99; x \ 100; x \ 101; x \ 102; x \ 103; x \ 104; x \ 105; x \ 106; x \ 107; x \ 108; x \ 109; x \ 110; x \ 111; x \ 112; x \ 113; x \ 114; x \ 115; x \ 116; x \ 117; x \ 118; x \ 119; x \ 120; x \ 121; x \ 122; x \ 123; x \ 124; x \ 125; x \ 126; x \ 127; x \ 128; x \ 129; x \ 130; x \ 131; x \ 132; x \ 133; x \ 134; x \ 135; x \ 136; x \ 137; x \ 138; x \ 139; x \ 140; x \ 141; x \ 142; x \ 143; x \ 144; x \ 145; x \ 146; x \ 147; x \ 148; x \ 149; x \ 150; x \ 151; x \ 152; x \ 153; x \ 154; x \ 155; x \ 156; x \ 157; x \ 158; x \ 159; x \ 160; x \ 161; x \ 162; x \ 163; x \ 164; x \ 165; x \ 166; x \ 167; x \ 168; x \ 169; x \ 170; x \ 171; x \ 172; x \ 173; x \ 174; x \ 175; x \ 176; x \ 177; x \ 178; x \ 179; x \ 180; x \ 181; x \ 182; x \ 183; x \ 184; x \ 185; x \ 186; x \ 187; x \ 188; x \ 189; x \ 190; x \ 191; x \ 192; x \ 193; x \ 194; x \ 195; x \ 196; x \ 197; x \ 198; x \ 199; x \ 200; x \ 201; x \ 202; x \ 203; x \ 204; x \ 205; x \ 206; x \ 207; x \ 208; x \ 209; x \ 210; x \ 211; x \ 212; x \ 213; x \ 214; x \ 215; x \ 216; x \ 217; x \ 218; x \ 219; x \ 220; x \ 221; x \ 222; x \ 223; x \ 224; x \ 225; x \ 226; x \ 227; x \ 228; x \ 229; x \ 230; x \ 231; x \ 232; x \ 233; x \ 234; x \ 235; x \ 236; x \ 237; x \ 238; x \ 239; x \ 240; x \ 241; x \ 242; x \ 243; x \ 244; x \ 245; x \ 246; x \ 247; x \ 248; x \ 249; x \ 250; x \ 251; x \ 252; x \ 253; x \ 254; x \ 255\ \ P" apply (subgoal_tac "x \ set enum", subst (asm) set_enum_word8_def) apply (drule set_strip_insert, assumption)+ apply (erule emptyE) apply (subst enum_UNIV, rule UNIV_I) done lemma upto_enum_red': assumes lt: "1 \ X" shows "[(0::'a :: len word) .e. X - 1] = map of_nat [0 ..< unat X]" proof - have lt': "unat X < 2 ^ LENGTH('a)" by (rule unat_lt2p) show ?thesis apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_unat_diff_1 [OF lt]) apply (rule map_cong [OF refl]) apply (rule toEnum_of_nat) apply simp apply (erule order_less_trans [OF _ lt']) done qed lemma upto_enum_red2: assumes szv: "sz < LENGTH('a :: len)" shows "[(0:: 'a :: len word) .e. 2 ^ sz - 1] = map of_nat [0 ..< 2 ^ sz]" using szv apply (subst unat_power_lower[OF szv, symmetric]) apply (rule upto_enum_red') apply (subst word_le_nat_alt, simp) done lemma upto_enum_step_red: assumes szv: "sz < LENGTH('a)" and usszv: "us \ sz" shows "[0 :: 'a :: len word , 2 ^ us .e. 2 ^ sz - 1] = map (\x. of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" using szv unfolding upto_enum_step_def apply (subst if_not_P) apply (rule leD) apply (subst word_le_nat_alt) apply (subst unat_minus_one) apply simp apply simp apply simp apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_div_unat_helper [where 'a = 'a, OF szv usszv, symmetric]) apply clarsimp apply (subst toEnum_of_nat) apply (erule order_less_trans) using szv apply simp apply simp done lemma upto_enum_word: "[x .e. y] = map of_nat [unat x ..< Suc (unat y)]" apply (subst upto_enum_red) apply clarsimp apply (subst toEnum_of_nat) prefer 2 apply (rule refl) apply (erule disjE, simp) apply clarsimp apply (erule order_less_trans) apply simp done lemma word_upto_Cons_eq: "x < y \ [x::'a::len word .e. y] = x # [x + 1 .e. y]" apply (subst upto_enum_red) apply (subst upt_conv_Cons) apply simp_all apply unat_arith apply (simp only: list.map list.inject upto_enum_red to_from_enum simp_thms) apply simp_all apply unat_arith done lemma distinct_enum_upto: "distinct [(0 :: 'a::len word) .e. b]" proof - have "\(b::'a word). [0 .e. b] = nths enum {..< Suc (fromEnum b)}" apply (subst upto_enum_red) apply (subst nths_upt_eq_take) apply (subst enum_word_def) apply (subst take_map) apply (subst take_upt) apply (simp only: add_0 fromEnum_unat) apply (rule order_trans [OF _ order_eq_refl]) apply (rule Suc_leI [OF unat_lt2p]) apply simp apply clarsimp apply (rule toEnum_of_nat) apply (erule order_less_trans [OF _ unat_lt2p]) done then show ?thesis by (rule ssubst) (rule distinct_nthsI, simp) qed lemma upto_enum_set_conv [simp]: fixes a :: "'a :: len word" shows "set [a .e. b] = {x. a \ x \ x \ b}" apply (subst upto_enum_red) apply (subst set_map) apply safe apply simp apply clarsimp apply (erule disjE) apply simp apply (erule iffD2 [OF word_le_nat_alt]) apply clarsimp apply (erule word_unat.Rep_cases [OF unat_le [OF order_less_imp_le]]) apply simp apply (erule iffD2 [OF word_le_nat_alt]) apply simp apply clarsimp apply (erule disjE) apply simp apply clarsimp apply (rule word_unat.Rep_cases [OF unat_le [OF order_less_imp_le]]) apply assumption apply simp apply (erule order_less_imp_le [OF iffD2 [OF word_less_nat_alt]]) apply clarsimp apply (rule_tac x="fromEnum x" in image_eqI) apply clarsimp apply clarsimp apply (rule conjI) apply (subst word_le_nat_alt [symmetric]) apply simp apply safe apply (simp add: word_le_nat_alt [symmetric]) apply (simp add: word_less_nat_alt [symmetric]) done lemma upto_enum_less: assumes xin: "x \ set [(a::'a::len word).e.2 ^ n - 1]" and nv: "n < LENGTH('a::len)" shows "x < 2 ^ n" proof (cases n) case 0 then show ?thesis using xin by simp next case (Suc m) show ?thesis using xin nv le_m1_iff_lt p2_gt_0 by auto qed lemma upto_enum_len_less: "\ n \ length [a, b .e. c]; n \ 0 \ \ a \ c" unfolding upto_enum_step_def by (simp split: if_split_asm) lemma length_upto_enum_step: fixes x :: "'a :: len word" shows "x \ z \ length [x , y .e. z] = (unat ((z - x) div (y - x))) + 1" unfolding upto_enum_step_def by (simp add: upto_enum_red) lemma map_length_unfold_one: fixes x :: "'a::len word" assumes xv: "Suc (unat x) < 2 ^ LENGTH('a)" and ax: "a < x" shows "map f [a .e. x] = f a # map f [a + 1 .e. x]" by (subst word_upto_Cons_eq, auto, fact+) lemma upto_enum_set_conv2: fixes a :: "'a::len word" shows "set [a .e. b] = {a .. b}" by auto lemma of_nat_unat [simp]: "of_nat \ unat = id" by (rule ext, simp) lemma Suc_unat_minus_one [simp]: "x \ 0 \ Suc (unat (x - 1)) = unat x" by (metis Suc_diff_1 unat_gt_0 unat_minus_one) lemma word_add_le_dest: fixes i :: "'a :: len word" assumes le: "i + k \ j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i \ j" using uik ujk le by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem] elim: add_le_mono1) lemma mask_shift: "(x && ~~ (mask y)) >> y = x >> y" apply (rule bit_eqI) apply (simp add: bit_and_iff bit_not_iff bit_shiftr_word_iff bit_mask_iff not_le) using bit_imp_le_length apply auto done lemma word_add_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k \ j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: add_le_mono1) then show ?thesis using ujk ij by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_le_mono2: fixes i :: "'a :: len word" shows "\i \ j; unat j + unat k < 2 ^ LENGTH('a)\ \ k + i \ k + j" by (subst field_simps, subst field_simps, erule (1) word_add_le_mono1) lemma word_add_le_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k \ j + k) = (i \ j)" proof assume "i \ j" show "i + k \ j + k" by (rule word_add_le_mono1) fact+ next assume "i + k \ j + k" show "i \ j" by (rule word_add_le_dest) fact+ qed lemma word_add_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k < j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_less_nat_alt elim: add_less_mono1) then show ?thesis using ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_less_dest: fixes i :: "'a :: len word" assumes le: "i + k < j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk le by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem] elim: add_less_mono1) lemma word_add_less_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k < j + k) = (i < j)" proof assume "i < j" show "i + k < j + k" by (rule word_add_less_mono1) fact+ next assume "i + k < j + k" show "i < j" by (rule word_add_less_dest) fact+ qed lemma shiftr_div_2n': "unat (w >> n) = unat w div 2 ^ n" apply (unfold unat_eq_nat_uint) apply (subst shiftr_div_2n) apply (subst nat_div_distrib) apply simp apply (simp add: nat_power_eq) done lemma shiftl_shiftr_id: assumes nv: "n < LENGTH('a)" and xv: "x < 2 ^ (LENGTH('a) - n)" shows "x << n >> n = (x::'a::len word)" apply (simp add: shiftl_t2n) apply (rule word_unat.Rep_eqD) apply (subst shiftr_div_2n') apply (cases n) apply simp apply (subst iffD1 [OF unat_mult_lem])+ apply (subst unat_power_lower[OF nv]) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl]) apply (rule unat_power_lower) apply simp apply (subst unat_power_lower[OF nv]) apply simp done lemma ucast_shiftl_eq_0: fixes w :: "'a :: len word" shows "\ n \ LENGTH('b) \ \ ucast (w << n) = (0 :: 'b :: len word)" by (case_tac "size w \ n", clarsimp simp: shiftl_zero_size) (clarsimp simp: not_le ucast_bl bl_shiftl bang_eq test_bit_of_bl rev_nth nth_append) lemma word_mult_less_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" using assms by (rule word_mult_less_cancel) lemma word_le_imp_diff_le: fixes n :: "'a::len word" shows "\k \ n; n \ m\ \ n - k \ m" by (auto simp: unat_sub word_le_nat_alt) lemma word_less_imp_diff_less: fixes n :: "'a::len word" shows "\k \ n; n < m\ \ n - k < m" by (clarsimp simp: unat_sub word_less_nat_alt intro!: less_imp_diff_less) lemma word_mult_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k \ j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: mult_le_mono1) then show ?thesis using ujk knz ij by (auto simp: word_le_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_le_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k \ j * k) = (i \ j)" proof assume "i \ j" show "i * k \ j * k" by (rule word_mult_le_mono1) fact+ next assume p: "i * k \ j * k" have "0 < unat k" using knz by (simp add: word_less_nat_alt) then show "i \ j" using p by (clarsimp simp: word_le_nat_alt iffD1 [OF unat_mult_lem uik] iffD1 [OF unat_mult_lem ujk]) qed lemma word_diff_less: fixes n :: "'a :: len word" shows "\0 < n; 0 < m; n \ m\ \ m - n < m" apply (subst word_less_nat_alt) apply (subst unat_sub) apply assumption apply (rule diff_less) apply (simp_all add: word_less_nat_alt) done lemma MinI: assumes fa: "finite A" and ne: "A \ {}" and xv: "m \ A" and min: "\y \ A. m \ y" shows "Min A = m" using fa ne xv min proof (induct A arbitrary: m rule: finite_ne_induct) case singleton then show ?case by simp next case (insert y F) from insert.prems have yx: "m \ y" and fx: "\y \ F. m \ y" by auto have "m \ insert y F" by fact then show ?case proof assume mv: "m = y" have mlt: "m \ Min F" by (rule iffD2 [OF Min_ge_iff [OF insert.hyps(1) insert.hyps(2)] fx]) show ?case apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)]) apply (subst mv [symmetric]) apply (auto simp: min_def mlt) done next assume "m \ F" then have mf: "Min F = m" by (rule insert.hyps(4) [OF _ fx]) show ?case apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)]) apply (subst mf) apply (rule iffD2 [OF _ yx]) apply (auto simp: min_def) done qed qed lemma length_upto_enum [simp]: fixes a :: "'a :: len word" shows "length [a .e. b] = Suc (unat b) - unat a" apply (simp add: word_le_nat_alt upto_enum_red) apply (clarsimp simp: Suc_diff_le) done lemma length_upto_enum_cases: fixes a :: "'a::len word" shows "length [a .e. b] = (if a \ b then Suc (unat b) - unat a else 0)" apply (case_tac "a \ b") apply (clarsimp) apply (clarsimp simp: upto_enum_def) apply unat_arith done lemma length_upto_enum_less_one: "\a \ b; b \ 0\ \ length [a .e. b - 1] = unat (b - a)" apply clarsimp apply (subst unat_sub[symmetric], assumption) apply clarsimp done lemma drop_upto_enum: "drop (unat n) [0 .e. m] = [n .e. m]" apply (clarsimp simp: upto_enum_def) apply (induct m, simp) by (metis drop_map drop_upt plus_nat.add_0) lemma distinct_enum_upto' [simp]: "distinct [a::'a::len word .e. b]" apply (subst drop_upto_enum [symmetric]) apply (rule distinct_drop) apply (rule distinct_enum_upto) done lemma length_interval: "\set xs = {x. (a::'a::len word) \ x \ x \ b}; distinct xs\ \ length xs = Suc (unat b) - unat a" apply (frule distinct_card) apply (subgoal_tac "set xs = set [a .e. b]") apply (cut_tac distinct_card [where xs="[a .e. b]"]) apply (subst (asm) length_upto_enum) apply clarsimp apply (rule distinct_enum_upto') apply simp done lemma not_empty_eq: "(S \ {}) = (\x. x \ S)" by auto lemma range_subset_lower: fixes c :: "'a ::linorder" shows "\ {a..b} \ {c..d}; x \ {a..b} \ \ c \ a" apply (frule (1) subsetD) apply (rule classical) apply clarsimp done lemma range_subset_upper: fixes c :: "'a ::linorder" shows "\ {a..b} \ {c..d}; x \ {a..b} \ \ b \ d" apply (frule (1) subsetD) apply (rule classical) apply clarsimp done lemma range_subset_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} \ {c..d}) = (c \ a \ b \ d)" apply (insert non_empty) apply (rule iffI) apply (frule range_subset_lower [where x=a], simp) apply (drule range_subset_upper [where x=a], simp) apply simp apply auto done lemma range_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} = {c..d}) = (a = c \ b = d)" by (metis atLeastatMost_subset_iff eq_iff non_empty) lemma range_strict_subset_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} \ {c..d}) = (c \ a \ b \ d \ (a = c \ b \ d))" apply (insert non_empty) apply (subst psubset_eq) apply (subst range_subset_eq, assumption+) apply (subst range_eq, assumption+) apply simp done lemma range_subsetI: fixes x :: "'a :: order" assumes xX: "X \ x" and yY: "y \ Y" shows "{x .. y} \ {X .. Y}" using xX yY by auto lemma set_False [simp]: "(set bs \ {False}) = (True \ set bs)" by auto declare of_nat_power [simp del] (* TODO: move to word *) lemma unat_of_bl_length: "unat (of_bl xs :: 'a::len word) < 2 ^ (length xs)" proof (cases "length xs < LENGTH('a)") case True then have "(of_bl xs::'a::len word) < 2 ^ length xs" by (simp add: of_bl_length_less) with True show ?thesis by (simp add: word_less_nat_alt word_unat_power unat_of_nat) next case False have "unat (of_bl xs::'a::len word) < 2 ^ LENGTH('a)" by (simp split: unat_split) also from False have "LENGTH('a) \ length xs" by simp then have "2 ^ LENGTH('a) \ (2::nat) ^ length xs" by (rule power_increasing) simp finally show ?thesis . qed lemma is_aligned_0': "is_aligned 0 n" by (fact is_aligned_0) lemma p_assoc_help: fixes p :: "'a::{ring,power,numeral,one}" shows "p + 2^sz - 1 = p + (2^sz - 1)" by simp lemma word_add_increasing: fixes x :: "'a :: len word" shows "\ p + w \ x; p \ p + w \ \ p \ x" by unat_arith lemma word_random: fixes x :: "'a :: len word" shows "\ p \ p + x'; x \ x' \ \ p \ p + x" by unat_arith lemma word_sub_mono: "\ a \ c; d \ b; a - b \ a; c - d \ c \ \ (a - b) \ (c - d :: 'a :: len word)" by unat_arith lemma power_not_zero: "n < LENGTH('a::len) \ (2 :: 'a word) ^ n \ 0" by (metis p2_gt_0 word_neq_0_conv) lemma word_gt_a_gt_0: "a < n \ (0 :: 'a::len word) < n" apply (case_tac "n = 0") apply clarsimp apply (clarsimp simp: word_neq_0_conv) done lemma word_shift_nonzero: "\ (x::'a::len word) \ 2 ^ m; m + n < LENGTH('a::len); x \ 0\ \ x << n \ 0" apply (simp only: word_neq_0_conv word_less_nat_alt shiftl_t2n mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (subst mod_less) apply (rule order_le_less_trans) apply (erule mult_le_mono2) apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply simp done lemma word_power_less_1 [simp]: "sz < LENGTH('a::len) \ (2::'a word) ^ sz - 1 < 2 ^ sz" apply (simp add: word_less_nat_alt) apply (subst unat_minus_one) apply (simp add: word_unat.Rep_inject [symmetric]) apply simp done lemma nasty_split_lt: "\ (x :: 'a:: len word) < 2 ^ (m - n); n \ m; m < LENGTH('a::len) \ \ x * 2 ^ n + (2 ^ n - 1) \ 2 ^ m - 1" apply (simp only: add_diff_eq) apply (subst mult_1[symmetric], subst distrib_right[symmetric]) apply (rule word_sub_mono) apply (rule order_trans) apply (rule word_mult_le_mono1) apply (rule inc_le) apply assumption apply (subst word_neq_0_conv[symmetric]) apply (rule power_not_zero) apply simp apply (subst unat_power_lower, simp)+ apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply (subst power_add[symmetric]) apply simp apply simp apply (rule word_sub_1_le) apply (subst mult.commute) apply (subst shiftl_t2n[symmetric]) apply (rule word_shift_nonzero) apply (erule inc_le) apply simp apply (unat_arith) apply (drule word_power_less_1) apply simp done lemma nasty_split_less: "\m \ n; n \ nm; nm < LENGTH('a::len); x < 2 ^ (nm - n)\ \ (x :: 'a word) * 2 ^ n + (2 ^ m - 1) < 2 ^ nm" apply (simp only: word_less_sub_le[symmetric]) apply (rule order_trans [OF _ nasty_split_lt]) apply (rule word_plus_mono_right) apply (rule word_sub_mono) apply (simp add: word_le_nat_alt) apply simp apply (simp add: word_sub_1_le[OF power_not_zero]) apply (simp add: word_sub_1_le[OF power_not_zero]) apply (rule is_aligned_no_wrap') apply (rule is_aligned_mult_triv2) apply simp apply (erule order_le_less_trans, simp) apply simp+ done lemma int_not_emptyD: "A \ B \ {} \ \x. x \ A \ x \ B" by (erule contrapos_np, clarsimp simp: disjoint_iff_not_equal) lemma unat_less_power: fixes k :: "'a::len word" assumes szv: "sz < LENGTH('a)" and kv: "k < 2 ^ sz" shows "unat k < 2 ^ sz" using szv unat_mono [OF kv] by simp lemma unat_mult_power_lem: assumes kv: "k < 2 ^ (LENGTH('a::len) - sz)" shows "unat (2 ^ sz * of_nat k :: (('a::len) word)) = 2 ^ sz * k" proof (cases \sz < LENGTH('a)\) case True with assms show ?thesis by (simp add: unat_word_ariths take_bit_eq_mod mod_simps) (simp add: take_bit_nat_eq_self_iff nat_less_power_trans flip: take_bit_eq_mod) next case False with assms show ?thesis by simp qed lemma aligned_add_offset_no_wrap: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off < 2 ^ sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis using szv apply (subst xv) apply (subst unat_mult_power_lem[OF kl]) apply (subst mult.commute, rule nat_add_offset_less) apply (rule less_le_trans[OF unat_mono[OF offv, simplified]]) apply (erule eq_imp_le[OF unat_power_lower]) apply (rule kl) apply simp done next assume "\ sz < LENGTH('a)" with offv show ?thesis by (simp add: not_less power_overflow ) qed lemma aligned_add_offset_mod: fixes x :: "('a::len) word" assumes al: "is_aligned x sz" and kv: "k < 2 ^ sz" shows "(x + k) mod 2 ^ sz = k" proof cases assume szv: "sz < LENGTH('a)" have ux: "unat x + unat k < 2 ^ LENGTH('a)" by (rule aligned_add_offset_no_wrap) fact+ show ?thesis using al szv apply - apply (erule is_alignedE) apply (subst word_unat.Rep_inject [symmetric]) apply (subst unat_mod) apply (subst iffD1 [OF unat_add_lem], rule ux) apply simp apply (subst unat_mult_power_lem, assumption+) apply (simp) apply (rule mod_less[OF less_le_trans[OF unat_mono], OF kv]) apply (erule eq_imp_le[OF unat_power_lower]) done next assume "\ sz < LENGTH('a)" with al show ?thesis by (simp add: not_less power_overflow is_aligned_mask mask_eq_decr_exp word_mod_by_0) qed lemma word_plus_mcs_4: "\v + x \ w + x; x \ v + x\ \ v \ (w::'a::len word)" by uint_arith lemma word_plus_mcs_3: "\v \ w; x \ w + x\ \ v + x \ w + (x::'a::len word)" by unat_arith lemma aligned_neq_into_no_overlap: fixes x :: "'a::len word" assumes neq: "x \ y" and alx: "is_aligned x sz" and aly: "is_aligned y sz" shows "{x .. x + (2 ^ sz - 1)} \ {y .. y + (2 ^ sz - 1)} = {}" proof cases assume szv: "sz < LENGTH('a)" show ?thesis proof (rule equals0I, clarsimp) fix z assume xb: "x \ z" and xt: "z \ x + (2 ^ sz - 1)" and yb: "y \ z" and yt: "z \ y + (2 ^ sz - 1)" have rl: "\(p::'a word) k w. \uint p + uint k < 2 ^ LENGTH('a); w = p + k; w \ p + (2 ^ sz - 1) \ \ k < 2 ^ sz" apply - apply simp apply (subst (asm) add.commute, subst (asm) add.commute, drule word_plus_mcs_4) apply (subst add.commute, subst no_plus_overflow_uint_size) apply (simp add: word_size_bl) apply (erule iffD1 [OF word_less_sub_le[OF szv]]) done from xb obtain kx where kx: "z = x + kx" and kxl: "uint x + uint kx < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') from yb obtain ky where ky: "z = y + ky" and kyl: "uint y + uint ky < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') have "x = y" proof - have "kx = z mod 2 ^ sz" proof (subst kx, rule sym, rule aligned_add_offset_mod) show "kx < 2 ^ sz" by (rule rl) fact+ qed fact+ also have "\ = ky" proof (subst ky, rule aligned_add_offset_mod) show "ky < 2 ^ sz" using kyl ky yt by (rule rl) qed fact+ finally have kxky: "kx = ky" . moreover have "x + kx = y + ky" by (simp add: kx [symmetric] ky [symmetric]) ultimately show ?thesis by simp qed then show False using neq by simp qed next assume "\ sz < LENGTH('a)" with neq alx aly have False by (simp add: is_aligned_mask mask_eq_decr_exp power_overflow) then show ?thesis .. qed lemma less_two_pow_divD: "\ (x :: nat) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (rule context_conjI) apply (rule ccontr) apply (simp add: power_strict_increasing) apply (simp add: power_sub) done lemma less_two_pow_divI: "\ (x :: nat) < 2 ^ (n - m); m \ n \ \ x < 2 ^ n div 2 ^ m" by (simp add: power_sub) lemma word_less_two_pow_divI: "\ (x :: 'a::len word) < 2 ^ (n - m); m \ n; n < LENGTH('a) \ \ x < 2 ^ n div 2 ^ m" apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (simp add: power_sub) done lemma word_less_two_pow_divD: "\ (x :: 'a::len word) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (cases "n < LENGTH('a)") apply (cases "m < LENGTH('a)") apply (simp add: word_less_nat_alt) apply (subst(asm) unat_word_ariths) apply (subst(asm) mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (clarsimp dest!: less_two_pow_divD) apply (simp add: power_overflow) apply (simp add: word_div_def) apply (simp add: power_overflow word_div_def) done lemma of_nat_less_two_pow_div_set: "\ n < LENGTH('a) \ \ {x. x < (2 ^ n div 2 ^ m :: 'a::len word)} = of_nat ` {k. k < 2 ^ n div 2 ^ m}" apply (simp add: image_def) apply (safe dest!: word_less_two_pow_divD less_two_pow_divD intro!: word_less_two_pow_divI) apply (rule_tac x="unat x" in exI) apply (simp add: power_sub[symmetric]) apply (subst unat_power_lower[symmetric, where 'a='a]) apply simp apply (erule unat_mono) apply (subst word_unat_power) apply (rule of_nat_mono_maybe) apply (rule power_strict_increasing) apply simp apply simp apply assumption done lemma word_less_power_trans2: fixes n :: "'a::len word" shows "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ n * 2 ^ k < 2 ^ m" by (subst field_simps, rule word_less_power_trans) (* shadows the slightly weaker Word.nth_ucast *) lemma nth_ucast: "(ucast (w::'a::len word)::'b::len word) !! n = (w !! n \ n < min LENGTH('a) LENGTH('b))" by transfer (simp add: bit_take_bit_iff ac_simps) lemma ucast_less: "LENGTH('b) < LENGTH('a) \ (ucast (x :: 'b :: len word) :: ('a :: len word)) < 2 ^ LENGTH('b)" by transfer simp lemma ucast_range_less: "LENGTH('a :: len) < LENGTH('b :: len) \ range (ucast :: 'a word \ 'b word) = {x. x < 2 ^ len_of TYPE ('a)}" apply safe apply (erule ucast_less) apply (simp add: image_def) apply (rule_tac x="ucast x" in exI) by word_eqI_solve lemma word_power_less_diff: "\2 ^ n * q < (2::'a::len word) ^ m; q < 2 ^ (LENGTH('a) - n)\ \ q < 2 ^ (m - n)" apply (case_tac "m \ LENGTH('a)") apply (simp add: power_overflow) apply (case_tac "n \ LENGTH('a)") apply (simp add: power_overflow) apply (cases "n = 0") apply simp apply (subst word_less_nat_alt) apply (subst unat_power_lower) apply simp apply (rule nat_power_less_diff) apply (simp add: word_less_nat_alt) apply (subst (asm) iffD1 [OF unat_mult_lem]) apply (simp add:nat_less_power_trans) apply simp done lemmas word_diff_ls'' = word_diff_ls [where xa=x and x=x for x] lemmas word_diff_ls' = word_diff_ls'' [simplified] lemmas word_l_diffs' = word_l_diffs [where xa=x and x=x for x] lemmas word_l_diffs = word_l_diffs' [simplified] lemma is_aligned_diff: fixes m :: "'a::len word" assumes alm: "is_aligned m s1" and aln: "is_aligned n s2" and s2wb: "s2 < LENGTH('a)" and nm: "m \ {n .. n + (2 ^ s2 - 1)}" and s1s2: "s1 \ s2" and s10: "0 < s1" (* Probably can be folded into the proof \ *) shows "\q. m - n = of_nat q * 2 ^ s1 \ q < 2 ^ (s2 - s1)" proof - have rl: "\m s. \ m < 2 ^ (LENGTH('a) - s); s < LENGTH('a) \ \ unat ((2::'a word) ^ s * of_nat m) = 2 ^ s * m" proof - fix m :: nat and s assume m: "m < 2 ^ (LENGTH('a) - s)" and s: "s < LENGTH('a)" then have "unat ((of_nat m) :: 'a word) = m" apply (subst unat_of_nat) apply (subst mod_less) apply (erule order_less_le_trans) apply (rule power_increasing) apply simp_all done then show "?thesis m s" using s m apply (subst iffD1 [OF unat_mult_lem]) apply (simp add: nat_less_power_trans)+ done qed have s1wb: "s1 < LENGTH('a)" using s2wb s1s2 by simp from alm obtain mq where mmq: "m = 2 ^ s1 * of_nat mq" and mq: "mq < 2 ^ (LENGTH('a) - s1)" by (auto elim: is_alignedE simp: field_simps) from aln obtain nq where nnq: "n = 2 ^ s2 * of_nat nq" and nq: "nq < 2 ^ (LENGTH('a) - s2)" by (auto elim: is_alignedE simp: field_simps) from s1s2 obtain sq where sq: "s2 = s1 + sq" by (auto simp: le_iff_add) note us1 = rl [OF mq s1wb] note us2 = rl [OF nq s2wb] from nm have "n \ m" by clarsimp then have "(2::'a word) ^ s2 * of_nat nq \ 2 ^ s1 * of_nat mq" using nnq mmq by simp then have "2 ^ s2 * nq \ 2 ^ s1 * mq" using s1wb s2wb by (simp add: word_le_nat_alt us1 us2) then have nqmq: "2 ^ sq * nq \ mq" using sq by (simp add: power_add) have "m - n = 2 ^ s1 * of_nat mq - 2 ^ s2 * of_nat nq" using mmq nnq by simp also have "\ = 2 ^ s1 * of_nat mq - 2 ^ s1 * 2 ^ sq * of_nat nq" using sq by (simp add: power_add) also have "\ = 2 ^ s1 * (of_nat mq - 2 ^ sq * of_nat nq)" by (simp add: field_simps) also have "\ = 2 ^ s1 * of_nat (mq - 2 ^ sq * nq)" using s1wb s2wb us1 us2 nqmq by (simp add: word_unat_power) finally have mn: "m - n = of_nat (mq - 2 ^ sq * nq) * 2 ^ s1" by simp moreover from nm have "m - n \ 2 ^ s2 - 1" by - (rule word_diff_ls', (simp add: field_simps)+) then have "(2::'a word) ^ s1 * of_nat (mq - 2 ^ sq * nq) < 2 ^ s2" using mn s2wb by (simp add: field_simps) then have "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (s2 - s1)" proof (rule word_power_less_diff) have mm: "mq - 2 ^ sq * nq < 2 ^ (LENGTH('a) - s1)" using mq by simp moreover from s10 have "LENGTH('a) - s1 < LENGTH('a)" by (rule diff_less, simp) ultimately show "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (LENGTH('a) - s1)" using take_bit_nat_less_self_iff [of \LENGTH('a)\ \mq - 2 ^ sq * nq\] apply (auto simp add: word_less_nat_alt not_le not_less) apply (metis take_bit_nat_eq_self_iff) done qed then have "mq - 2 ^ sq * nq < 2 ^ (s2 - s1)" using mq s2wb apply (simp add: word_less_nat_alt take_bit_eq_mod) apply (subst (asm) mod_less) apply auto apply (rule order_le_less_trans) apply (rule diff_le_self) apply (erule order_less_le_trans) apply simp done ultimately show ?thesis by auto qed lemma word_less_sub_1: "x < (y :: 'a :: len word) \ x \ y - 1" by (fact word_le_minus_one_leq) lemma word_sub_mono2: "\ a + b \ c + d; c \ a; b \ a + b; d \ c + d \ \ b \ (d :: 'a :: len word)" apply (drule(1) word_sub_mono) apply simp apply simp apply simp done lemma word_not_le: "(\ x \ (y :: 'a :: len word)) = (y < x)" by fastforce lemma word_subset_less: "\ {x .. x + r - 1} \ {y .. y + s - 1}; x \ x + r - 1; y \ y + (s :: 'a :: len word) - 1; s \ 0 \ \ r \ s" apply (frule subsetD[where c=x]) apply simp apply (drule subsetD[where c="x + r - 1"]) apply simp apply (clarsimp simp: add_diff_eq[symmetric]) apply (drule(1) word_sub_mono2) apply (simp_all add: olen_add_eqv[symmetric]) apply (erule word_le_minus_cancel) apply (rule ccontr) apply (simp add: word_not_le) done lemma uint_power_lower: "n < LENGTH('a) \ uint (2 ^ n :: 'a :: len word) = (2 ^ n :: int)" by (rule uint_2p_alt) lemma power_le_mono: "\2 ^ n \ (2::'a::len word) ^ m; n < LENGTH('a); m < LENGTH('a)\ \ n \ m" apply (clarsimp simp add: le_less) apply safe apply (simp add: word_less_nat_alt) apply (simp only: uint_arith_simps(3)) apply (drule uint_power_lower)+ apply simp done lemma two_power_eq: "\n < LENGTH('a); m < LENGTH('a)\ \ ((2::'a::len word) ^ n = 2 ^ m) = (n = m)" apply safe apply (rule order_antisym) apply (simp add: power_le_mono[where 'a='a])+ done lemma of_bl_length: "length xs < LENGTH('a) \ of_bl xs < (2 :: 'a::len word) ^ length xs" by (simp add: of_bl_length_less) lemma unat_of_nat_eq: "x < 2 ^ LENGTH('a) \ unat (of_nat x ::'a::len word) = x" by (rule unat_of_nat_len) lemma unat_eq_of_nat: "n < 2 ^ LENGTH('a) \ (unat (x :: 'a::len word) = n) = (x = of_nat n)" by transfer (auto simp add: take_bit_of_nat nat_eq_iff take_bit_nat_eq_self_iff intro: sym) lemma unat_less_helper: "x < of_nat n \ unat x < n" apply (simp add: word_less_nat_alt) apply (erule order_less_le_trans) apply (simp add: take_bit_eq_mod) done lemma nat_uint_less_helper: "nat (uint y) = z \ x < y \ nat (uint x) < z" apply (erule subst) apply (subst unat_eq_nat_uint [symmetric]) apply (subst unat_eq_nat_uint [symmetric]) by (simp add: unat_mono) lemma of_nat_0: "\of_nat n = (0::'a::len word); n < 2 ^ LENGTH('a)\ \ n = 0" by transfer (simp add: take_bit_eq_mod) lemma word_leq_le_minus_one: "\ x \ y; x \ 0 \ \ x - 1 < (y :: 'a :: len word)" apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst unat_minus_one) apply assumption apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma of_nat_inj: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (of_nat x = (of_nat y :: 'a :: len word)) = (x = y)" by (simp add: word_unat.norm_eq_iff [symmetric]) lemma if_Some_None_eq_None: "((if P then Some v else None) = None) = (\ P)" by simp lemma CollectPairFalse [iff]: "{(a,b). False} = {}" by (simp add: split_def) lemma if_conj_dist: "((if b then w else x) \ (if b then y else z) \ X) = ((if b then w \ y else x \ z) \ X)" by simp lemma if_P_True1: "Q \ (if P then True else Q)" by simp lemma if_P_True2: "Q \ (if P then Q else True)" by simp lemma list_all2_induct [consumes 1, case_names Nil Cons]: assumes lall: "list_all2 Q xs ys" and nilr: "P [] []" and consr: "\x xs y ys. \list_all2 Q xs ys; Q x y; P xs ys\ \ P (x # xs) (y # ys)" shows "P xs ys" using lall proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) case 1 then show ?case by auto fact+ next case (2 x xs y ys) show ?case proof (rule consr) from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all then show "P xs ys" by (intro "2.hyps") qed qed lemma upto_enum_step_shift: "\ is_aligned p n \ \ ([p , p + 2 ^ m .e. p + 2 ^ n - 1]) = map ((+) p) [0, 2 ^ m .e. 2 ^ n - 1]" apply (erule is_aligned_get_word_bits) prefer 2 apply (simp add: map_idI) apply (clarsimp simp: upto_enum_step_def) apply (frule is_aligned_no_overflow) apply (simp add: linorder_not_le [symmetric]) done lemma upto_enum_step_shift_red: "\ is_aligned p sz; sz < LENGTH('a); us \ sz \ \ [p :: 'a :: len word, p + 2 ^ us .e. p + 2 ^ sz - 1] = map (\x. p + of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" apply (subst upto_enum_step_shift, assumption) apply (simp add: upto_enum_step_red) done lemma div_to_mult_word_lt: "\ (x :: 'a :: len word) \ y div z \ \ x * z \ y" apply (cases "z = 0") apply simp apply (simp add: word_neq_0_conv) apply (rule order_trans) apply (erule(1) word_mult_le_mono1) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply simp apply (rule word_div_mult_le) done lemma upto_enum_step_subset: "set [x, y .e. z] \ {x .. z}" apply (clarsimp simp: upto_enum_step_def linorder_not_less) apply (drule div_to_mult_word_lt) apply (rule conjI) apply (erule word_random[rotated]) apply simp apply (rule order_trans) apply (erule word_plus_mono_right) apply simp apply simp done lemma shiftr_less_t2n': "\ x && mask (n + m) = x; m < LENGTH('a) \ \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (simp add: word_size mask_eq_iff_w2p[symmetric]) apply word_eqI apply (erule_tac x="na + n" in allE) apply fastforce done lemma shiftr_less_t2n: "x < 2 ^ (n + m) \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (rule shiftr_less_t2n') apply (erule less_mask_eq) apply (rule ccontr) apply (simp add: not_less) apply (subst (asm) p2_eq_0[symmetric]) apply (simp add: power_add) done lemma shiftr_eq_0: "n \ LENGTH('a) \ ((w::'a::len word) >> n) = 0" apply (cut_tac shiftr_less_t2n'[of w n 0], simp) apply (simp add: mask_eq_iff) apply (simp add: lt2p_lem) apply simp done lemma shiftr_not_mask_0: "n+m \ LENGTH('a :: len) \ ((w::'a::len word) >> n) && ~~ (mask m) = 0" apply (simp add: and_not_mask shiftr_less_t2n shiftr_shiftr) apply (subgoal_tac "w >> n + m = 0", simp) apply (simp add: le_mask_iff[symmetric] mask_eq_decr_exp le_def) apply (subst (asm) p2_gt_0[symmetric]) apply (simp add: power_add not_less) done lemma shiftl_less_t2n: fixes x :: "'a :: len word" shows "\ x < (2 ^ (m - n)); m < LENGTH('a) \ \ (x << n) < 2 ^ m" apply (simp add: word_size mask_eq_iff_w2p[symmetric]) apply word_eqI apply (erule_tac x="na - n" in allE) apply auto done lemma shiftl_less_t2n': "(x::'a::len word) < 2 ^ m \ m+n < LENGTH('a) \ x << n < 2 ^ (m + n)" by (rule shiftl_less_t2n) simp_all lemma ucast_ucast_mask: "(ucast :: 'a :: len word \ 'b :: len word) (ucast x) = x && mask (len_of TYPE ('a))" by word_eqI lemma ucast_ucast_len: "\ x < 2 ^ LENGTH('b) \ \ ucast (ucast x::'b::len word) = (x::'a::len word)" apply (subst ucast_ucast_mask) apply (erule less_mask_eq) done lemma ucast_ucast_id: "LENGTH('a) < LENGTH('b) \ ucast (ucast (x::'a::len word)::'b::len word) = x" by (auto intro: ucast_up_ucast_id simp: is_up_def source_size_def target_size_def word_size) lemma unat_ucast: "unat (ucast x :: ('a :: len) word) = unat x mod 2 ^ (LENGTH('a))" proof - have \2 ^ LENGTH('a) = nat (2 ^ LENGTH('a))\ by simp moreover have \unat (UCAST('b \ 'a) x) = unat x mod nat (2 ^ LENGTH('a))\ by transfer (simp flip: nat_mod_distrib take_bit_eq_mod) ultimately show ?thesis by (simp only:) qed lemma ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ (ucast x < ((ucast (y :: 'a::len word)) :: 'b::len word)) = (x < y)" apply (simp add: word_less_nat_alt unat_ucast) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply simp done \ \This weaker version was previously called @{text ucast_less_ucast}. We retain it to support existing proofs.\ lemmas ucast_less_ucast_weak = ucast_less_ucast[OF order.strict_implies_order] lemma sints_subset: "m \ n \ sints m \ sints n" apply (simp add: sints_num) apply clarsimp apply (rule conjI) apply (erule order_trans[rotated]) apply simp apply (erule order_less_le_trans) apply simp done lemma up_scast_inj: "\ scast x = (scast y :: 'b :: len word); size x \ LENGTH('b) \ \ x = y" apply (unfold scast_eq) apply (subst(asm) word_sint.Abs_inject) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply simp done lemma up_scast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (scast x = (scast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_scast_inj simp: word_size) lemma nth_bounded: "\(x :: 'a :: len word) !! n; x < 2 ^ m; m \ len_of TYPE ('a)\ \ n < m" apply (frule test_bit_size) apply (clarsimp simp: test_bit_bl word_size) apply (simp add: rev_nth) apply (subst(asm) is_aligned_add_conv[OF is_aligned_0', simplified add_0_left, rotated]) apply assumption+ apply (simp only: to_bl_0) apply (simp add: nth_append split: if_split_asm) done lemma is_aligned_add_or: "\is_aligned p n; d < 2 ^ n\ \ p + d = p || d" by (rule word_plus_and_or_coroll, word_eqI) blast lemma two_power_increasing: "\ n \ m; m < LENGTH('a) \ \ (2 :: 'a :: len word) ^ n \ 2 ^ m" by (simp add: word_le_nat_alt) lemma is_aligned_add_less_t2n: "\is_aligned (p::'a::len word) n; d < 2^n; n \ m; p < 2^m\ \ p + d < 2^m" apply (case_tac "m < LENGTH('a)") apply (subst mask_eq_iff_w2p[symmetric]) apply (simp add: word_size) apply (simp add: is_aligned_add_or word_ao_dist less_mask_eq) apply (subst less_mask_eq) apply (erule order_less_le_trans) apply (erule(1) two_power_increasing) apply simp apply (simp add: power_overflow) done lemma aligned_offset_non_zero: "\ is_aligned x n; y < 2 ^ n; x \ 0 \ \ x + y \ 0" apply (cases "y = 0") apply simp apply (subst word_neq_0_conv) apply (subst gt0_iff_gem1) apply (erule is_aligned_get_word_bits) apply (subst field_simps[symmetric], subst plus_le_left_cancel_nowrap) apply (rule is_aligned_no_wrap') apply simp apply (rule word_leq_le_minus_one) apply simp apply assumption apply (erule (1) is_aligned_no_wrap') apply (simp add: gt0_iff_gem1 [symmetric] word_neq_0_conv) apply simp done lemmas mask_inner_mask = mask_eqs(1) lemma mask_add_aligned: "is_aligned p n \ (p + q) && mask n = q && mask n" apply (simp add: is_aligned_mask) apply (subst mask_inner_mask [symmetric]) apply simp done lemma cart_singleton_empty: "(S \ {e} = {}) = (S = {})" by blast lemma word_div_1: "(n :: 'a :: len word) div 1 = n" by (simp add: word_div_def) lemma word_minus_one_le: "-1 \ (x :: 'a :: len word) = (x = -1)" apply (insert word_n1_ge[where y=x]) apply safe apply (erule(1) order_antisym) done lemma mask_out_sub_mask: "(x && ~~ (mask n)) = x - (x && (mask n))" by (simp add: field_simps word_plus_and_or_coroll2) lemma is_aligned_addD1: assumes al1: "is_aligned (x + y) n" and al2: "is_aligned (x::'a::len word) n" shows "is_aligned y n" using al2 proof (rule is_aligned_get_word_bits) assume "x = 0" then show ?thesis using al1 by simp next assume nv: "n < LENGTH('a)" from al1 obtain q1 where xy: "x + y = 2 ^ n * of_nat q1" and "q1 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) moreover from al2 obtain q2 where x: "x = 2 ^ n * of_nat q2" and "q2 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) ultimately have "y = 2 ^ n * (of_nat q1 - of_nat q2)" by (simp add: field_simps) then show ?thesis using nv by (simp add: is_aligned_mult_triv1) qed lemmas is_aligned_addD2 = is_aligned_addD1[OF subst[OF add.commute, of "%x. is_aligned x n" for n]] lemma is_aligned_add: "\is_aligned p n; is_aligned q n\ \ is_aligned (p + q) n" by (simp add: is_aligned_mask mask_add_aligned) lemma word_le_add: fixes x :: "'a :: len word" shows "x \ y \ \n. y = x + of_nat n" by (rule exI [where x = "unat (y - x)"]) simp lemma word_plus_mcs_4': fixes x :: "'a :: len word" shows "\x + v \ x + w; x \ x + v\ \ v \ w" apply (rule word_plus_mcs_4) apply (simp add: add.commute) apply (simp add: add.commute) done lemma shiftl_mask_is_0[simp]: "(x << n) && mask n = 0" apply (rule iffD1 [OF is_aligned_mask]) apply (rule is_aligned_shiftl_self) done definition sum_map :: "('a \ 'b) \ ('c \ 'd) \ 'a + 'c \ 'b + 'd" where "sum_map f g x \ case x of Inl v \ Inl (f v) | Inr v' \ Inr (g v')" lemma sum_map_simps[simp]: "sum_map f g (Inl v) = Inl (f v)" "sum_map f g (Inr w) = Inr (g w)" by (simp add: sum_map_def)+ lemma if_and_helper: "(If x v v') && v'' = If x (v && v'') (v' && v'')" by (rule if_distrib) lemma unat_Suc2: fixes n :: "'a :: len word" shows "n \ -1 \ unat (n + 1) = Suc (unat n)" apply (subst add.commute, rule unatSuc) apply (subst eq_diff_eq[symmetric], simp add: minus_equation_iff) done lemmas word_unat_Rep_inject1 = word_unat.Rep_inject[where y=1] lemmas unat_eq_1 = unat_eq_0 word_unat_Rep_inject1[simplified] lemma rshift_sub_mask_eq: "(a >> (size a - b)) && mask b = a >> (size a - b)" using shiftl_shiftr2[where a=a and b=0 and c="size a - b"] apply (cases "b < size a") apply simp apply (simp add: linorder_not_less mask_eq_decr_exp word_size p2_eq_0[THEN iffD2]) done lemma shiftl_shiftr3: "b \ c \ a << b >> c = (a >> c - b) && mask (size a - c)" apply (cases "b = c") apply (simp add: shiftl_shiftr1) apply (simp add: shiftl_shiftr2) done lemma and_mask_shiftr_comm: "m \ size w \ (w && mask m) >> n = (w >> n) && mask (m-n)" by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3) lemma and_mask_shiftl_comm: "m+n \ size w \ (w && mask m) << n = (w << n) && mask (m+n)" by (simp add: and_mask word_size shiftl_shiftl) (simp add: shiftl_shiftr1) lemma le_mask_shiftl_le_mask: "s = m + n \ x \ mask n \ x << m \ mask s" for x :: \'a::len word\ by (simp add: le_mask_iff shiftl_shiftr3) lemma and_not_mask_twice: "(w && ~~ (mask n)) && ~~ (mask m) = w && ~~ (mask (max m n))" apply (simp add: and_not_mask) apply (case_tac "n x = y - 1 \ x < y - (1 ::'a::len word)" apply (drule word_less_sub_1) apply (drule order_le_imp_less_or_eq) apply auto done lemma eq_eqI: "a = b \ (a = x) = (b = x)" by simp lemma mask_and_mask: "mask a && mask b = mask (min a b)" by word_eqI lemma mask_eq_0_eq_x: "(x && w = 0) = (x && ~~ w = x)" using word_plus_and_or_coroll2[where x=x and w=w] by auto lemma mask_eq_x_eq_0: "(x && w = x) = (x && ~~ w = 0)" using word_plus_and_or_coroll2[where x=x and w=w] by auto definition "limited_and (x :: 'a :: len word) y = (x && y = x)" lemma limited_and_eq_0: "\ limited_and x z; y && ~~ z = y \ \ x && y = 0" unfolding limited_and_def apply (subst arg_cong2[where f="(&&)"]) apply (erule sym)+ apply (simp(no_asm) add: word_bw_assocs word_bw_comms word_bw_lcs) done lemma limited_and_eq_id: "\ limited_and x z; y && z = z \ \ x && y = x" unfolding limited_and_def by (erule subst, fastforce simp: word_bw_lcs word_bw_assocs word_bw_comms) lemma lshift_limited_and: "limited_and x z \ limited_and (x << n) (z << n)" unfolding limited_and_def by (simp add: shiftl_over_and_dist[symmetric]) lemma rshift_limited_and: "limited_and x z \ limited_and (x >> n) (z >> n)" unfolding limited_and_def by (simp add: shiftr_over_and_dist[symmetric]) lemmas limited_and_simps1 = limited_and_eq_0 limited_and_eq_id lemmas is_aligned_limited_and = is_aligned_neg_mask_eq[unfolded mask_eq_decr_exp, folded limited_and_def] lemma compl_of_1: "~~ 1 = (-2 :: 'a :: len word)" by (fact not_one) lemmas limited_and_simps = limited_and_simps1 limited_and_simps1[OF is_aligned_limited_and] limited_and_simps1[OF lshift_limited_and] limited_and_simps1[OF rshift_limited_and] limited_and_simps1[OF rshift_limited_and, OF is_aligned_limited_and] compl_of_1 shiftl_shiftr1[unfolded word_size mask_eq_decr_exp] shiftl_shiftr2[unfolded word_size mask_eq_decr_exp] lemma split_word_eq_on_mask: "(x = y) = (x && m = y && m \ x && ~~ m = y && ~~ m)" by safe word_eqI_solve lemma map2_Cons_2_3: "(map2 f xs (y # ys) = (z # zs)) = (\x xs'. xs = x # xs' \ f x y = z \ map2 f xs' ys = zs)" by (case_tac xs, simp_all) lemma map2_xor_replicate_False: "map2 (\x y. x \ \ y) xs (replicate n False) = take n xs" apply (induct xs arbitrary: n, simp) apply (case_tac n; simp) done lemma word_and_1_shiftl: "x && (1 << n) = (if x !! n then (1 << n) else 0)" for x :: "'a :: len word" by word_eqI_solve lemmas word_and_1_shiftls' = word_and_1_shiftl[where n=0] word_and_1_shiftl[where n=1] word_and_1_shiftl[where n=2] lemmas word_and_1_shiftls = word_and_1_shiftls' [simplified] lemma word_and_mask_shiftl: "x && (mask n << m) = ((x >> m) && mask n) << m" by word_eqI_solve lemma plus_Collect_helper: "(+) x ` {xa. P (xa :: 'a :: len word)} = {xa. P (xa - x)}" by (fastforce simp add: image_def) lemma plus_Collect_helper2: "(+) (- x) ` {xa. P (xa :: 'a :: len word)} = {xa. P (x + xa)}" using plus_Collect_helper [of "- x" P] by (simp add: ac_simps) lemma word_FF_is_mask: "0xFF = (mask 8 :: 'a::len word)" by (simp add: mask_eq_decr_exp) lemma word_1FF_is_mask: "0x1FF = (mask 9 :: 'a::len word)" by (simp add: mask_eq_decr_exp) lemma ucast_of_nat_small: "x < 2 ^ LENGTH('a) \ ucast (of_nat x :: 'a :: len word) = (of_nat x :: 'b :: len word)" apply (rule sym, subst word_unat.inverse_norm) apply (simp add: ucast_eq of_nat_nat[symmetric] take_bit_eq_mod) done lemma word_le_make_less: fixes x :: "'a :: len word" shows "y \ -1 \ (x \ y) = (x < (y + 1))" apply safe apply (erule plus_one_helper2) apply (simp add: eq_diff_eq[symmetric]) done lemmas finite_word = finite [where 'a="'a::len word"] lemma word_to_1_set: "{0 ..< (1 :: 'a :: len word)} = {0}" by fastforce lemma range_subset_eq2: "{a :: 'a :: len word .. b} \ {} \ ({a .. b} \ {c .. d}) = (c \ a \ b \ d)" by simp lemma word_leq_minus_one_le: fixes x :: "'a::len word" shows "\y \ 0; x \ y - 1 \ \ x < y" using le_m1_iff_lt word_neq_0_conv by blast lemma word_count_from_top: "n \ 0 \ {0 ..< n :: 'a :: len word} = {0 ..< n - 1} \ {n - 1}" apply (rule set_eqI, rule iffI) apply simp apply (drule word_le_minus_one_leq) apply (rule disjCI) apply simp apply simp apply (erule word_leq_minus_one_le) apply fastforce done lemma word_minus_one_le_leq: "\ x - 1 < y \ \ x \ (y :: 'a :: len word)" apply (cases "x = 0") apply simp apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst(asm) unat_minus_one) apply (simp add: word_less_nat_alt) apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma mod_mod_power: fixes k :: nat shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" proof (cases "m \ n") case True then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ m" apply - apply (subst mod_less [where n = "2 ^ n"]) apply (rule order_less_le_trans [OF mod_less_divisor]) apply simp+ done also have "\ = k mod 2 ^ (min m n)" using True by simp finally show ?thesis . next case False then have "n < m" by simp then obtain d where md: "m = n + d" by (auto dest: less_imp_add_positive) then have "k mod 2 ^ m = 2 ^ n * (k div 2 ^ n mod 2 ^ d) + k mod 2 ^ n" by (simp add: mod_mult2_eq power_add) then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ n" by (simp add: mod_add_left_eq) then show ?thesis using False by simp qed lemma word_div_less: "m < n \ m div n = 0" for m :: "'a :: len word" by (simp add: unat_mono word_arith_nat_defs(6)) lemma word_must_wrap: "\ x \ n - 1; n \ x \ \ n = (0 :: 'a :: len word)" using dual_order.trans sub_wrap word_less_1 by blast lemma range_subset_card: "\ {a :: 'a :: len word .. b} \ {c .. d}; b \ a \ \ d \ c \ d - c \ b - a" using word_sub_le word_sub_mono by fastforce lemma less_1_simp: "n - 1 < m = (n \ (m :: 'a :: len word) \ n \ 0)" by unat_arith lemma alignUp_div_helper: fixes a :: "'a::len word" assumes kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" and le: "a \ x" and sz: "n < LENGTH('a)" and anz: "a mod 2 ^ n \ 0" shows "a div 2 ^ n < of_nat k" proof - have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using xk kv sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst unat_power_lower, simp) apply (subst mult.commute) apply (rule nat_less_power_trans) apply simp apply simp done have "unat a div 2 ^ n * 2 ^ n \ unat a" proof - have "unat a = unat a div 2 ^ n * 2 ^ n + unat a mod 2 ^ n" by (simp add: div_mult_mod_eq) also have "\ \ unat a div 2 ^ n * 2 ^ n" using sz anz by (simp add: unat_arith_simps) finally show ?thesis .. qed then have "a div 2 ^ n * 2 ^ n < a" using sz anz apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst unat_div) apply simp apply (rule order_le_less_trans [OF mod_less_eq_dividend]) apply (erule order_le_neq_trans [OF div_mult_le]) done also from xk le have "\ \ of_nat k * 2 ^ n" by (simp add: field_simps) finally show ?thesis using sz kv apply - apply (erule word_mult_less_dest [OF _ _ kn]) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply (rule unat_lt2p) done qed lemma nat_mod_power_lem: fixes a :: nat shows "1 < a \ a ^ n mod a ^ m = (if m \ n then 0 else a ^ n)" apply (clarsimp) apply (clarsimp simp add: le_iff_add power_add) done lemma power_mod_div: fixes x :: "nat" shows "x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" (is "?LHS = ?RHS") proof (cases "n \ m") case True then have "?LHS = 0" apply - apply (rule div_less) apply (rule order_less_le_trans [OF mod_less_divisor]; simp) done also have "\ = ?RHS" using True by simp finally show ?thesis . next case False then have lt: "m < n" by simp then obtain q where nv: "n = m + q" and "0 < q" by (auto dest: less_imp_Suc_add) then have "x mod 2 ^ n = 2 ^ m * (x div 2 ^ m mod 2 ^ q) + x mod 2 ^ m" by (simp add: power_add mod_mult2_eq) then have "?LHS = x div 2 ^ m mod 2 ^ q" by (simp add: div_add1_eq) also have "\ = ?RHS" using nv by simp finally show ?thesis . qed lemma word_power_mod_div: fixes x :: "'a::len word" shows "\ n < LENGTH('a); m < LENGTH('a)\ \ x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" apply (simp add: word_arith_nat_div unat_mod power_mod_div) apply (subst unat_arith_simps(3)) apply (subst unat_mod) apply (subst unat_of_nat)+ apply (simp add: mod_mod_power min.commute) done lemma word_range_minus_1': fixes a :: "'a :: len word" shows "a \ 0 \ {a - 1<..b} = {a..b}" by (simp add: greaterThanAtMost_def atLeastAtMost_def greaterThan_def atLeast_def less_1_simp) lemma word_range_minus_1: fixes a :: "'a :: len word" shows "b \ 0 \ {a..b - 1} = {a.. 'b :: len word) x" by transfer simp -lemmas if_fun_split = if_apply_def2 - lemma i_hate_words_helper: "i \ (j - k :: nat) \ i \ j" by simp lemma i_hate_words: "unat (a :: 'a word) \ unat (b :: 'a :: len word) - Suc 0 \ a \ -1" apply (frule i_hate_words_helper) apply (subst(asm) word_le_nat_alt[symmetric]) apply (clarsimp simp only: word_minus_one_le) apply (simp only: linorder_not_less[symmetric]) apply (erule notE) apply (rule diff_Suc_less) apply (subst neq0_conv[symmetric]) apply (subst unat_eq_0) apply (rule notI, drule arg_cong[where f="(+) 1"]) apply simp done lemma overflow_plus_one_self: "(1 + p \ p) = (p = (-1 :: 'a :: len word))" apply rule apply (rule ccontr) apply (drule plus_one_helper2) apply (rule notI) apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply (simp add: field_simps) apply simp done lemma plus_1_less: "(x + 1 \ (x :: 'a :: len word)) = (x = -1)" apply (rule iffI) apply (rule ccontr) apply (cut_tac plus_one_helper2[where x=x, OF order_refl]) apply simp apply clarsimp apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply simp done lemma pos_mult_pos_ge: "[|x > (0::int); n>=0 |] ==> n * x >= n*1" apply (simp only: mult_left_mono) done lemma If_eq_obvious: "x \ z \ ((if P then x else y) = z) = (\ P \ y = z)" by simp lemma Some_to_the: "v = Some x \ x = the v" by simp lemma dom_if_Some: "dom (\x. if P x then Some (f x) else g x) = {x. P x} \ dom g" by fastforce lemma dom_insert_absorb: "x \ dom f \ insert x (dom f) = dom f" by auto lemma emptyE2: "\ S = {}; x \ S \ \ P" by simp lemma mod_div_equality_div_eq: "a div b * b = (a - (a mod b) :: int)" by (simp add: field_simps) lemma zmod_helper: "n mod m = k \ ((n :: int) + a) mod m = (k + a) mod m" by (metis add.commute mod_add_right_eq) lemma int_div_sub_1: "\ m \ 1 \ \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)" apply (subgoal_tac "m = 0 \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)") apply fastforce apply (subst mult_cancel_right[symmetric]) apply (simp only: left_diff_distrib split: if_split) apply (simp only: mod_div_equality_div_eq) apply (clarsimp simp: field_simps) apply (clarsimp simp: dvd_eq_mod_eq_0) apply (cases "m = 1") apply simp apply (subst mod_diff_eq[symmetric], simp add: zmod_minus1) apply clarsimp apply (subst diff_add_cancel[where b=1, symmetric]) apply (subst mod_add_eq[symmetric]) apply (simp add: field_simps) apply (rule mod_pos_pos_trivial) apply (subst add_0_right[where a=0, symmetric]) apply (rule add_mono) apply simp apply simp apply (cases "(n - 1) mod m = m - 1") apply (drule zmod_helper[where a=1]) apply simp apply (subgoal_tac "1 + (n - 1) mod m \ m") apply simp apply (subst field_simps, rule zless_imp_add1_zle) apply simp done lemma ptr_add_image_multI: "\ \x y. (x * val = y * val') = (x * val'' = y); x * val'' \ S \ \ ptr_add ptr (x * val) \ (\p. ptr_add ptr (p * val')) ` S" apply (simp add: image_def) apply (erule rev_bexI) apply (rule arg_cong[where f="ptr_add ptr"]) apply simp done lemma shift_times_fold: "(x :: 'a :: len word) * (2 ^ n) << m = x << (m + n)" by (simp add: shiftl_t2n ac_simps power_add) lemma word_plus_strict_mono_right: fixes x :: "'a :: len word" shows "\y < z; x \ x + z\ \ x + y < x + z" by unat_arith lemma replicate_minus: "k < n \ replicate n False = replicate (n - k) False @ replicate k False" by (subst replicate_add [symmetric]) simp lemmas map_prod_split_imageI' = map_prod_imageI[where f="case_prod f" and g="case_prod g" and a="(a, b)" and b="(c, d)" for a b c d f g] lemmas map_prod_split_imageI = map_prod_split_imageI'[simplified] lemma word_div_mult: "0 < c \ a < b * c \ a div c < b" for a b c :: "'a::len word" by (rule classical) (use div_to_mult_word_lt [of b a c] in \auto simp add: word_less_nat_alt word_le_nat_alt unat_div\) lemma word_less_power_trans_ofnat: "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ of_nat n * 2 ^ k < (2::'a::len word) ^ m" apply (subst mult.commute) apply (rule word_less_power_trans) apply (simp_all add: word_less_nat_alt less_le_trans take_bit_eq_mod) done lemma word_1_le_power: "n < LENGTH('a) \ (1 :: 'a :: len word) \ 2 ^ n" by (rule inc_le[where i=0, simplified], erule iffD2[OF p2_gt_0]) lemma enum_word_div: fixes v :: "'a :: len word" shows "\xs ys. enum = xs @ [v] @ ys \ (\x \ set xs. x < v) \ (\y \ set ys. v < y)" apply (simp only: enum_word_def) apply (subst upt_add_eq_append'[where j="unat v"]) apply simp apply (rule order_less_imp_le, simp) apply (simp add: upt_conv_Cons) apply (intro exI conjI) apply fastforce apply clarsimp apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp apply (clarsimp simp: Suc_le_eq) apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp done lemma of_bool_nth: "of_bool (x !! v) = (x >> v) && 1" by (simp add: test_bit_word_eq shiftr_word_eq bit_eq_iff) (auto simp add: bit_1_iff bit_and_iff bit_drop_bit_eq intro: ccontr) lemma unat_1_0: "1 \ (x::'a::len word) = (0 < unat x)" by (auto simp add: word_le_nat_alt) lemma x_less_2_0_1': fixes x :: "'a::len word" shows "\LENGTH('a) \ 1; x < 2\ \ x = 0 \ x = 1" apply (cases \2 \ LENGTH('a)\) apply simp_all apply transfer apply auto apply (metis add.commute add.right_neutral even_two_times_div_two mod_div_trivial mod_pos_pos_trivial mult.commute mult_zero_left not_less not_take_bit_negative odd_two_times_div_two_succ) done lemmas word_add_le_iff2 = word_add_le_iff [folded no_olen_add_nat] lemma of_nat_power: shows "\ p < 2 ^ x; x < len_of TYPE ('a) \ \ of_nat p < (2 :: 'a :: len word) ^ x" apply (rule order_less_le_trans) apply (rule of_nat_mono_maybe) apply (erule power_strict_increasing) apply simp apply assumption apply (simp add: word_unat_power) done lemma of_nat_n_less_equal_power_2: "n < LENGTH('a::len) \ ((of_nat n)::'a word) < 2 ^ n" apply (induct n) apply clarsimp apply clarsimp apply (metis of_nat_power n_less_equal_power_2 of_nat_Suc power_Suc) done lemma eq_mask_less: fixes w :: "'a::len word" assumes eqm: "w = w && mask n" and sz: "n < len_of TYPE ('a)" shows "w < (2::'a word) ^ n" by (subst eqm, rule and_mask_less' [OF sz]) lemma of_nat_mono_maybe': fixes Y :: "nat" assumes xlt: "x < 2 ^ len_of TYPE ('a)" assumes ylt: "y < 2 ^ len_of TYPE ('a)" shows "(y < x) = (of_nat y < (of_nat x :: 'a :: len word))" apply (subst word_less_nat_alt) apply (subst unat_of_nat)+ apply (subst mod_less) apply (rule ylt) apply (subst mod_less) apply (rule xlt) apply simp done lemma shiftr_mask_eq: "(x >> n) && mask (size x - n) = x >> n" for x :: "'a :: len word" by word_eqI_solve lemma shiftr_mask_eq': "m = (size x - n) \ (x >> n) && mask m = x >> n" for x :: "'a :: len word" by (simp add: shiftr_mask_eq) lemma dom_if: "dom (\a. if a \ addrs then Some (f a) else g a) = addrs \ dom g" by (auto simp: dom_def split: if_split) lemma less_is_non_zero_p1: fixes a :: "'a :: len word" shows "a < k \ a + 1 \ 0" apply (erule contrapos_pn) apply (drule max_word_wrap) apply (simp add: not_less) done lemma of_nat_mono_maybe_le: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (y \ x) = ((of_nat y :: 'a :: len word) \ of_nat x)" apply (clarsimp simp: le_less) apply (rule disj_cong) apply (rule of_nat_mono_maybe', assumption+) apply (simp add: word_unat.norm_eq_iff [symmetric]) done lemma mask_AND_NOT_mask: "(w && ~~ (mask n)) && mask n = 0" by word_eqI lemma AND_NOT_mask_plus_AND_mask_eq: "(w && ~~ (mask n)) + (w && mask n) = w" by (subst word_plus_and_or_coroll; word_eqI_solve) lemma mask_eqI: fixes x :: "'a :: len word" assumes m1: "x && mask n = y && mask n" and m2: "x && ~~ (mask n) = y && ~~ (mask n)" shows "x = y" proof (subst bang_eq, rule allI) fix m show "x !! m = y !! m" proof (cases "m < n") case True then have "x !! m = ((x && mask n) !! m)" by (simp add: word_size test_bit_conj_lt) also have "\ = ((y && mask n) !! m)" using m1 by simp also have "\ = y !! m" using True by (simp add: word_size test_bit_conj_lt) finally show ?thesis . next case False then have "x !! m = ((x && ~~ (mask n)) !! m)" by (simp add: neg_mask_test_bit test_bit_conj_lt) also have "\ = ((y && ~~ (mask n)) !! m)" using m2 by simp also have "\ = y !! m" using False by (simp add: neg_mask_test_bit test_bit_conj_lt) finally show ?thesis . qed qed lemma nat_less_power_trans2: fixes n :: nat shows "\n < 2 ^ (m - k); k \ m\ \ n * 2 ^ k < 2 ^ m" by (subst mult.commute, erule (1) nat_less_power_trans) lemma nat_move_sub_le: "(a::nat) + b \ c \ a \ c - b" by arith lemma neq_0_no_wrap: fixes x :: "'a :: len word" shows "\ x \ x + y; x \ 0 \ \ x + y \ 0" by clarsimp lemma plus_minus_one_rewrite: "v + (- 1 :: ('a :: {ring, one, uminus})) \ v - 1" by (simp add: field_simps) lemma power_minus_is_div: "b \ a \ (2 :: nat) ^ (a - b) = 2 ^ a div 2 ^ b" apply (induct a arbitrary: b) apply simp apply (erule le_SucE) apply (clarsimp simp:Suc_diff_le le_iff_add power_add) apply simp done lemma two_pow_div_gt_le: "v < 2 ^ n div (2 ^ m :: nat) \ m \ n" by (clarsimp dest!: less_two_pow_divD) lemma unatSuc2: fixes n :: "'a :: len word" shows "n + 1 \ 0 \ unat (n + 1) = Suc (unat n)" by (simp add: add.commute unatSuc) lemma word_of_nat_less: "\ n < unat x \ \ of_nat n < x" apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) apply (simp add: take_bit_eq_mod) done lemma word_of_nat_le: "n \ unat x \ of_nat n \ x" apply (simp add: word_le_nat_alt unat_of_nat) apply (erule order_trans[rotated]) apply (simp add: take_bit_eq_mod) done lemma word_unat_less_le: "a \ of_nat b \ unat a \ b" by (metis eq_iff le_cases le_unat_uoi word_of_nat_le) lemma and_eq_0_is_nth: fixes x :: "'a :: len word" shows "y = 1 << n \ ((x && y) = 0) = (\ (x !! n))" apply safe apply (drule_tac u="(x && (1 << n))" and x=n in word_eqD) apply (simp add: nth_w2p) apply (simp add: test_bit_bin) apply word_eqI done lemmas arg_cong_Not = arg_cong [where f=Not] lemmas and_neq_0_is_nth = arg_cong_Not [OF and_eq_0_is_nth, simplified] lemma nth_is_and_neq_0: "(x::'a::len word) !! n = (x && 2 ^ n \ 0)" by (subst and_neq_0_is_nth; rule refl) lemma mask_Suc_0 : "mask (Suc 0) = (1 :: 'a::len word)" by (simp add: mask_eq_decr_exp) lemma ucast_ucast_add: fixes x :: "'a :: len word" fixes y :: "'b :: len word" shows "LENGTH('b) \ LENGTH('a) \ ucast (ucast x + y) = x + ucast y" apply (rule word_unat.Rep_eqD) apply (simp add: unat_ucast unat_word_ariths mod_mod_power min.absorb2 unat_of_nat) apply (subst mod_add_left_eq[symmetric]) apply (simp add: mod_mod_power min.absorb2) apply (subst mod_add_right_eq) apply simp done lemma word_shift_zero: "\ x << n = 0; x \ 2^m; m + n < LENGTH('a)\ \ (x::'a::len word) = 0" apply (rule ccontr) apply (drule (2) word_shift_nonzero) apply simp done lemma bool_mask': fixes x :: "'a :: len word" shows "2 < LENGTH('a) \ (0 < x && 1) = (x && 1 = 1)" by (simp add: and_one_eq mod_2_eq_odd) lemma sint_eq_uint: "\ msb x \ sint x = uint x" apply (rule word_uint.Abs_eqD, subst word_sint.Rep_inverse) apply simp_all apply (cut_tac x=x in word_sint.Rep) apply (clarsimp simp add: uints_num sints_num) apply (rule conjI) apply (rule ccontr) apply (simp add: linorder_not_le word_msb_sint[symmetric]) apply (erule order_less_le_trans) apply simp done lemma scast_eq_ucast: "\ msb x \ scast x = ucast x" apply (cases \LENGTH('a)\) apply simp apply (rule bit_word_eqI) apply (auto simp add: bit_signed_iff bit_unsigned_iff min_def msb_word_eq) apply (erule notE) apply (metis le_less_Suc_eq test_bit_bin test_bit_word_eq) done lemma lt1_neq0: fixes x :: "'a :: len word" shows "(1 \ x) = (x \ 0)" by unat_arith lemma word_plus_one_nonzero: fixes x :: "'a :: len word" shows "\x \ x + y; y \ 0\ \ x + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (erule word_random) apply (simp add: lt1_neq0) done lemma word_sub_plus_one_nonzero: fixes n :: "'a :: len word" shows "\n' \ n; n' \ 0\ \ (n - n') + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (rule word_random [where x' = n']) apply simp apply (erule word_sub_le) apply (simp add: lt1_neq0) done lemma word_le_minus_mono_right: fixes x :: "'a :: len word" shows "\ z \ y; y \ x; z \ x \ \ x - y \ x - z" apply (rule word_sub_mono) apply simp apply assumption apply (erule word_sub_le) apply (erule word_sub_le) done lemma drop_append_miracle: "n = length xs \ drop n (xs @ ys) = ys" by simp lemma foldr_does_nothing_to_xf: "\ \x s. x \ set xs \ xf (f x s) = xf s \ \ xf (foldr f xs s) = xf s" by (induct xs, simp_all) lemma nat_less_mult_monoish: "\ a < b; c < (d :: nat) \ \ (a + 1) * (c + 1) <= b * d" apply (drule Suc_leI)+ apply (drule(1) mult_le_mono) apply simp done lemma word_0_sle_from_less[unfolded word_size]: "\ x < 2 ^ (size x - 1) \ \ 0 <=s x" apply (clarsimp simp: word_sle_msb_le) apply (simp add: word_msb_nth) apply (subst (asm) word_test_bit_def [symmetric]) apply (drule less_mask_eq) apply (drule_tac x="size x - 1" in word_eqD) apply (simp add: word_size) done lemma not_msb_from_less: "(v :: 'a word) < 2 ^ (LENGTH('a :: len) - 1) \ \ msb v" apply (clarsimp simp add: msb_nth) apply (drule less_mask_eq) apply (drule word_eqD, drule(1) iffD2) apply simp done lemma distinct_lemma: "f x \ f y \ x \ y" by auto lemma ucast_sub_ucast: fixes x :: "'a::len word" assumes "y \ x" assumes T: "LENGTH('a) \ LENGTH('b)" shows "ucast (x - y) = (ucast x - ucast y :: 'b::len word)" proof - from T have P: "unat x < 2 ^ LENGTH('b)" "unat y < 2 ^ LENGTH('b)" by (fastforce intro!: less_le_trans[OF unat_lt2p])+ then show ?thesis by (simp add: unat_arith_simps unat_ucast assms[simplified unat_arith_simps]) qed lemma word_1_0: "\a + (1::('a::len) word) \ b; a < of_nat x\ \ a < b" apply transfer apply (subst (asm) take_bit_incr_eq) apply (auto simp add: diff_less_eq) using take_bit_int_less_exp le_less_trans by blast lemma unat_of_nat_less:"\ a < b; unat b = c \ \ a < of_nat c" by fastforce lemma word_le_plus_1: "\ (y::('a::len) word) < y + n; a < n \ \ y + a \ y + a + 1" by unat_arith lemma word_le_plus:"\(a::('a::len) word) < a + b; c < b\ \ a \ a + c" by (metis order_less_imp_le word_random) (* * Basic signed arithemetic properties. *) lemma sint_minus1 [simp]: "(sint x = -1) = (x = -1)" by (metis sint_n1 word_sint.Rep_inverse') lemma sint_0 [simp]: "(sint x = 0) = (x = 0)" by (metis sint_0 word_sint.Rep_inverse') (* It is not always that case that "sint 1 = 1", because of 1-bit word sizes. * This lemma produces the different cases. *) lemma sint_1_cases: P if \\ len_of TYPE ('a::len) = 1; (a::'a word) = 0; sint a = 0 \ \ P\ \\ len_of TYPE ('a) = 1; a = 1; sint (1 :: 'a word) = -1 \ \ P\ \\ len_of TYPE ('a) > 1; sint (1 :: 'a word) = 1 \ \ P\ proof (cases \LENGTH('a) = 1\) case True then have \a = 0 \ a = 1\ by transfer auto with True that show ?thesis by auto next case False with that show ?thesis by (simp add: less_le Suc_le_eq) qed lemma sint_int_min: "sint (- (2 ^ (LENGTH('a) - Suc 0)) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (subst word_sint.Abs_inverse' [where r="- (2 ^ (LENGTH('a) - Suc 0))"]) apply (clarsimp simp: sints_num) apply (clarsimp simp: wi_hom_syms word_of_int_2p) apply clarsimp done lemma sint_int_max_plus_1: "sint (2 ^ (LENGTH('a) - Suc 0) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (cases \LENGTH('a)\) apply simp_all apply (subst word_of_int_2p [symmetric]) apply (subst int_word_sint) apply simp done lemma sbintrunc_eq_in_range: "(sbintrunc n x = x) = (x \ range (sbintrunc n))" "(x = sbintrunc n x) = (x \ range (sbintrunc n))" apply (simp_all add: image_def) apply (metis sbintrunc_sbintrunc)+ done lemma sbintrunc_If: "- 3 * (2 ^ n) \ x \ x < 3 * (2 ^ n) \ sbintrunc n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n) else if x \ 2 ^ n then x - 2 * (2 ^ n) else x)" apply (simp add: no_sbintr_alt2, safe) apply (simp add: mod_pos_geq) apply (subst mod_add_self1[symmetric], simp) done lemma signed_arith_eq_checks_to_ord: "(sint a + sint b = sint (a + b )) = ((a <=s a + b) = (0 <=s b))" "(sint a - sint b = sint (a - b )) = ((0 <=s a - b) = (b <=s a))" "(- sint a = sint (- a)) = (0 <=s (- a) = (a <=s 0))" using sint_range'[where x=a] sint_range'[where x=b] by (simp_all add: sint_word_ariths word_sle_eq word_sless_alt sbintrunc_If) (* Basic proofs that signed word div/mod operations are * truncations of their integer counterparts. *) lemma signed_div_arith: "sint ((a::('a::len) word) sdiv b) = sbintrunc (LENGTH('a) - 1) (sint a sdiv sint b)" apply (subst word_sbin.norm_Rep [symmetric]) apply (subst bin_sbin_eq_iff' [symmetric]) apply simp apply (subst uint_sint [symmetric]) apply (clarsimp simp: sdiv_int_def sdiv_word_def) apply transfer apply simp done lemma signed_mod_arith: "sint ((a::('a::len) word) smod b) = sbintrunc (LENGTH('a) - 1) (sint a smod sint b)" apply (subst word_sbin.norm_Rep [symmetric]) apply (subst bin_sbin_eq_iff' [symmetric]) apply simp apply (subst uint_sint [symmetric]) apply (clarsimp simp: smod_int_def smod_word_def) using word_ubin.inverse_norm by force (* Signed word arithmetic overflow constraints. *) lemma signed_arith_ineq_checks_to_eq: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) = (sint a + sint b = sint (a + b ))" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) = (sint a - sint b = sint (a - b))" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) = ((- sint a) = sint (- a))" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) = (sint a * sint b = sint (a * b))" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) = (sint a sdiv sint b = sint (a sdiv b))" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) = (sint a smod sint b = sint (a smod b))" by (auto simp: sint_word_ariths word_size signed_div_arith signed_mod_arith sbintrunc_eq_in_range range_sbintrunc) lemma signed_arith_sint: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) \ sint (a + b) = (sint a + sint b)" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) \ sint (a - b) = (sint a - sint b)" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) \ sint (- a) = (- sint a)" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) \ sint (a * b) = (sint a * sint b)" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) \ sint (a sdiv b) = (sint a sdiv sint b)" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) \ sint (a smod b) = (sint a smod sint b)" by (subst (asm) signed_arith_ineq_checks_to_eq; simp)+ lemma signed_mult_eq_checks_double_size: assumes mult_le: "(2 ^ (len_of TYPE ('a) - 1) + 1) ^ 2 \ (2 :: int) ^ (len_of TYPE ('b) - 1)" and le: "2 ^ (LENGTH('a) - 1) \ (2 :: int) ^ (len_of TYPE ('b) - 1)" shows "(sint (a :: 'a :: len word) * sint b = sint (a * b)) = (scast a * scast b = (scast (a * b) :: 'b :: len word))" proof - have P: "sbintrunc (size a - 1) (sint a * sint b) \ range (sbintrunc (size a - 1))" by simp have abs: "!! x :: 'a word. abs (sint x) < 2 ^ (size a - 1) + 1" apply (cut_tac x=x in sint_range') apply (simp add: abs_le_iff word_size) done have abs_ab: "abs (sint a * sint b) < 2 ^ (LENGTH('b) - 1)" using abs_mult_less[OF abs[where x=a] abs[where x=b]] mult_le by (simp add: abs_mult power2_eq_square word_size) define r s where \r = LENGTH('a) - 1\ \s = LENGTH('b) - 1\ then have \LENGTH('a) = Suc r\ \LENGTH('b) = Suc s\ \size a = Suc r\ \size b = Suc r\ by (simp_all add: word_size) then show ?thesis using P[unfolded range_sbintrunc] abs_ab le apply clarsimp apply (transfer fixing: r s) apply (auto simp add: signed_take_bit_int_eq_self simp flip: signed_take_bit_eq_iff_take_bit_eq) done qed (* Properties about signed division. *) lemma int_sdiv_simps [simp]: "(a :: int) sdiv 1 = a" "(a :: int) sdiv 0 = 0" "(a :: int) sdiv -1 = -a" apply (auto simp: sdiv_int_def sgn_if) done lemma sgn_div_eq_sgn_mult: "a div b \ 0 \ sgn ((a :: int) div b) = sgn (a * b)" apply (clarsimp simp: sgn_if zero_le_mult_iff neg_imp_zdiv_nonneg_iff not_less) apply (metis less_le mult_le_0_iff neg_imp_zdiv_neg_iff not_less pos_imp_zdiv_neg_iff zdiv_eq_0_iff) done lemma sgn_sdiv_eq_sgn_mult: "a sdiv b \ 0 \ sgn ((a :: int) sdiv b) = sgn (a * b)" by (auto simp: sdiv_int_def sgn_div_eq_sgn_mult sgn_mult) lemma int_sdiv_same_is_1 [simp]: "a \ 0 \ ((a :: int) sdiv b = a) = (b = 1)" apply (rule iffI) apply (clarsimp simp: sdiv_int_def) apply (subgoal_tac "b > 0") apply (case_tac "a > 0") apply (clarsimp simp: sgn_if) apply (clarsimp simp: algebra_split_simps not_less) apply (metis int_div_same_is_1 le_neq_trans minus_minus neg_0_le_iff_le neg_equal_0_iff_equal) apply (case_tac "a > 0") apply (case_tac "b = 0") apply clarsimp apply (rule classical) apply (clarsimp simp: sgn_mult not_less) apply (metis le_less neg_0_less_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) apply (rule classical) apply (clarsimp simp: algebra_split_simps sgn_mult not_less sgn_if split: if_splits) apply (metis antisym less_le neg_imp_zdiv_nonneg_iff) apply (clarsimp simp: sdiv_int_def sgn_if) done lemma int_sdiv_negated_is_minus1 [simp]: "a \ 0 \ ((a :: int) sdiv b = - a) = (b = -1)" apply (clarsimp simp: sdiv_int_def) apply (rule iffI) apply (subgoal_tac "b < 0") apply (case_tac "a > 0") apply (clarsimp simp: sgn_if algebra_split_simps not_less) apply (case_tac "sgn (a * b) = -1") apply (clarsimp simp: not_less algebra_split_simps) apply (clarsimp simp: algebra_split_simps not_less) apply (rule classical) apply (case_tac "b = 0") apply (clarsimp simp: not_less sgn_mult) apply (case_tac "a > 0") apply (clarsimp simp: not_less sgn_mult) apply (metis less_le neg_less_0_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) apply (clarsimp simp: not_less sgn_mult) apply (metis antisym_conv div_minus_right neg_imp_zdiv_nonneg_iff neg_le_0_iff_le not_less) apply (clarsimp simp: sgn_if) done lemma sdiv_int_range: "(a :: int) sdiv b \ { - (abs a) .. (abs a) }" apply (unfold sdiv_int_def) apply (subgoal_tac "(abs a) div (abs b) \ (abs a)") apply (auto simp add: sgn_if not_less) apply (metis le_less le_less_trans neg_equal_0_iff_equal neg_less_iff_less not_le pos_imp_zdiv_neg_iff) apply (metis add.inverse_neutral div_int_pos_iff le_less neg_le_iff_le order_trans) apply (metis div_minus_right le_less_trans neg_imp_zdiv_neg_iff neg_less_0_iff_less not_le) using div_int_pos_iff apply fastforce apply (metis abs_0_eq abs_ge_zero div_by_0 zdiv_le_dividend zero_less_abs_iff) done lemma word_sdiv_div1 [simp]: "(a :: ('a::len) word) sdiv 1 = a" apply (rule sint_1_cases [where a=a]) apply (clarsimp simp: sdiv_word_def sdiv_int_def) apply (clarsimp simp: sdiv_word_def sdiv_int_def simp del: sint_minus1) apply (clarsimp simp: sdiv_word_def) done lemma sdiv_int_div_0 [simp]: "(x :: int) sdiv 0 = 0" by (clarsimp simp: sdiv_int_def) lemma sdiv_int_0_div [simp]: "0 sdiv (x :: int) = 0" by (clarsimp simp: sdiv_int_def) lemma word_sdiv_div0 [simp]: "(a :: ('a::len) word) sdiv 0 = 0" apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) done lemma word_sdiv_div_minus1 [simp]: "(a :: ('a::len) word) sdiv -1 = -a" apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) done lemmas word_sdiv_0 = word_sdiv_div0 lemma sdiv_word_min: "- (2 ^ (size a - 1)) \ sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word)" apply (clarsimp simp: word_size) apply (cut_tac sint_range' [where x=a]) apply (cut_tac sint_range' [where x=b]) apply clarsimp apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: max_def abs_if split: if_split_asm) done lemma sdiv_word_max: "(sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word) < (2 ^ (size a - 1))) = ((a \ - (2 ^ (size a - 1)) \ (b \ -1)))" (is "?lhs = (\ ?a_int_min \ \ ?b_minus1)") proof (rule classical) assume not_thesis: "\ ?thesis" have not_zero: "b \ 0" using not_thesis by (clarsimp) have result_range: "sint a sdiv sint b \ (sints (size a)) \ {2 ^ (size a - 1)}" apply (cut_tac sdiv_int_range [where a="sint a" and b="sint b"]) apply (erule rev_subsetD) using sint_range' [where x=a] sint_range' [where x=b] apply (auto simp: max_def abs_if word_size sints_num) done have result_range_overflow: "(sint a sdiv sint b = 2 ^ (size a - 1)) = (?a_int_min \ ?b_minus1)" apply (rule iffI [rotated]) apply (clarsimp simp: sdiv_int_def sgn_if word_size sint_int_min) apply (rule classical) apply (case_tac "?a_int_min") apply (clarsimp simp: word_size sint_int_min) apply (metis diff_0_right int_sdiv_negated_is_minus1 minus_diff_eq minus_int_code(2) power_eq_0_iff sint_minus1 zero_neq_numeral) apply (subgoal_tac "abs (sint a) < 2 ^ (size a - 1)") apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (clarsimp simp: word_size) apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (insert word_sint.Rep [where x="a"])[1] apply (clarsimp simp: minus_le_iff word_size abs_if sints_num split: if_split_asm) apply (metis minus_minus sint_int_min word_sint.Rep_inject) done have result_range_simple: "(sint a sdiv sint b \ (sints (size a))) \ ?thesis" apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: word_size sints_num sint_int_min) done show ?thesis apply (rule UnE [OF result_range result_range_simple]) apply simp apply (clarsimp simp: word_size) using result_range_overflow apply (clarsimp simp: word_size) done qed lemmas sdiv_word_min' = sdiv_word_min [simplified word_size, simplified] lemmas sdiv_word_max' = sdiv_word_max [simplified word_size, simplified] lemmas word_sdiv_numerals_lhs = sdiv_word_def[where v="numeral x" for x] sdiv_word_def[where v=0] sdiv_word_def[where v=1] lemmas word_sdiv_numerals = word_sdiv_numerals_lhs[where w="numeral y" for y] word_sdiv_numerals_lhs[where w=0] word_sdiv_numerals_lhs[where w=1] (* * Signed modulo properties. *) lemma smod_int_alt_def: "(a::int) smod b = sgn (a) * (abs a mod abs b)" apply (clarsimp simp: smod_int_def sdiv_int_def) apply (clarsimp simp: minus_div_mult_eq_mod [symmetric] abs_sgn sgn_mult sgn_if algebra_split_simps) done lemma smod_int_range: "b \ 0 \ (a::int) smod b \ { - abs b + 1 .. abs b - 1 }" apply (case_tac "b > 0") apply (insert pos_mod_conj [where a=a and b=b])[1] apply (insert pos_mod_conj [where a="-a" and b=b])[1] apply (auto simp: smod_int_alt_def algebra_simps sgn_if abs_if not_less add1_zle_eq [simplified add.commute])[1] apply (metis add_nonneg_nonneg int_one_le_iff_zero_less le_less less_add_same_cancel2 not_le pos_mod_conj) apply (metis (full_types) add.inverse_inverse eucl_rel_int eucl_rel_int_iff le_less_trans neg_0_le_iff_le) apply (insert neg_mod_conj [where a=a and b="b"])[1] apply (insert neg_mod_conj [where a="-a" and b="b"])[1] apply (clarsimp simp: smod_int_alt_def algebra_simps sgn_if abs_if not_less add1_zle_eq [simplified add.commute]) apply (metis neg_0_less_iff_less neg_mod_conj not_le not_less_iff_gr_or_eq order_trans pos_mod_conj) done lemma smod_int_compares: "\ 0 \ a; 0 < b \ \ (a :: int) smod b < b" "\ 0 \ a; 0 < b \ \ 0 \ (a :: int) smod b" "\ a \ 0; 0 < b \ \ -b < (a :: int) smod b" "\ a \ 0; 0 < b \ \ (a :: int) smod b \ 0" "\ 0 \ a; b < 0 \ \ (a :: int) smod b < - b" "\ 0 \ a; b < 0 \ \ 0 \ (a :: int) smod b" "\ a \ 0; b < 0 \ \ (a :: int) smod b \ 0" "\ a \ 0; b < 0 \ \ b \ (a :: int) smod b" apply (insert smod_int_range [where a=a and b=b]) apply (auto simp: add1_zle_eq smod_int_alt_def sgn_if) done lemma smod_int_mod_0 [simp]: "x smod (0 :: int) = x" by (clarsimp simp: smod_int_def) lemma smod_int_0_mod [simp]: "0 smod (x :: int) = 0" by (clarsimp simp: smod_int_alt_def) lemma smod_word_mod_0 [simp]: "x smod (0 :: ('a::len) word) = x" by (clarsimp simp: smod_word_def) lemma smod_word_0_mod [simp]: "0 smod (x :: ('a::len) word) = 0" by (clarsimp simp: smod_word_def) lemma smod_word_max: "sint (a::'a word) smod sint (b::'a word) < 2 ^ (LENGTH('a::len) - Suc 0)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply (clarsimp) apply (insert word_sint.Rep [where x="b", simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if split: if_split_asm) done lemma smod_word_min: "- (2 ^ (LENGTH('a::len) - Suc 0)) \ sint (a::'a word) smod sint (b::'a word)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply clarsimp apply (insert word_sint.Rep [where x=b, simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if add1_zle_eq split: if_split_asm) done lemma smod_word_alt_def: "(a :: ('a::len) word) smod b = a - (a sdiv b) * b" apply (cases \a \ - (2 ^ (LENGTH('a) - 1)) \ b \ - 1\) apply (clarsimp simp: smod_word_def sdiv_word_def smod_int_def simp flip: wi_hom_sub wi_hom_mult) apply (clarsimp simp: smod_word_def smod_int_def) done lemmas word_smod_numerals_lhs = smod_word_def[where v="numeral x" for x] smod_word_def[where v=0] smod_word_def[where v=1] lemmas word_smod_numerals = word_smod_numerals_lhs[where w="numeral y" for y] word_smod_numerals_lhs[where w=0] word_smod_numerals_lhs[where w=1] lemma sint_of_int_eq: "\ - (2 ^ (LENGTH('a) - 1)) \ x; x < 2 ^ (LENGTH('a) - 1) \ \ sint (of_int x :: ('a::len) word) = x" by (simp add: signed_take_bit_int_eq_self) lemma of_int_sint: "of_int (sint a) = a" by simp lemma nth_w2p_scast [simp]: "((scast ((2::'a::len signed word) ^ n) :: 'a word) !! m) \ ((((2::'a::len word) ^ n) :: 'a word) !! m)" apply (subst nth_w2p) apply (case_tac "n \ LENGTH('a)") apply (subst power_overflow, simp) apply clarsimp apply (metis nth_w2p scast_eq test_bit_conj_lt len_signed nth_word_of_int word_sint.Rep_inverse) done lemma scast_2_power [simp]: "scast ((2 :: 'a::len signed word) ^ x) = ((2 :: 'a word) ^ x)" by (clarsimp simp: word_eq_iff) lemma scast_bit_test [simp]: "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" by (clarsimp simp: word_eq_iff) lemma ucast_nat_def': "of_nat (unat x) = (ucast :: 'a :: len word \ ('b :: len) signed word) x" by (fact ucast_nat_def) lemma mod_mod_power_int: fixes k :: int shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" by (metis bintrunc_bintrunc_min bintrunc_mod2p min.commute) (* Normalise combinations of scast and ucast. *) lemma ucast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (ucast :: 'a word \ 'b word)" shows "ucast (M a b) = M' (ucast a) (ucast b)" apply (simp only: ucast_eq) apply (subst lift_M) apply (subst of_int_uint [symmetric], subst lift_M') apply (subst (1 2) int_word_uint) apply (subst word_ubin.norm_eq_iff [symmetric]) apply (subst (1 2) bintrunc_mod2p) apply (insert is_down) apply (unfold is_down_def) apply (clarsimp simp: target_size source_size) apply (clarsimp simp: mod_mod_power_int min_def) apply (rule distrib [symmetric]) done lemma ucast_down_add: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) + b) = (ucast a + ucast b :: 'b::len word)" by (rule ucast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma ucast_down_minus: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) - b) = (ucast a - ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma ucast_down_mult: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) * b) = (ucast a * ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (scast :: 'a word \ 'b word)" shows "scast (M a b) = M' (scast a) (scast b)" apply (subst (1 2 3) down_cast_same [symmetric]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) apply (rule ucast_distrib [where L=L, OF lift_M lift_M' distrib]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) done lemma scast_down_add: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) + b) = (scast a + scast b :: 'b::len word)" by (rule scast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma scast_down_minus: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) - b) = (scast a - scast b :: 'b::len word)" apply (rule scast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma scast_down_mult: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) * b) = (scast a * scast b :: 'b::len word)" apply (rule scast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_ucast_1: "\ is_down (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_ucast_3: "\ is_down (ucast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_ucast_4: "\ is_up (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_scast_b: "\ is_up (scast :: 'a word \ 'b word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq sint_up_scast) lemma ucast_scast_1: "\ is_down (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq ucast_down_wi) lemma ucast_scast_3: "\ is_down (scast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq ucast_down_wi) lemma ucast_scast_4: "\ is_up (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis down_cast_same scast_eq sint_up_scast) lemma ucast_ucast_a: "\ is_down (ucast :: 'b word \ 'c word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma ucast_ucast_b: "\ is_up (ucast :: 'a word \ 'b word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis ucast_up_ucast) lemma scast_scast_a: "\ is_down (scast :: 'b word \ 'c word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" apply (simp only: scast_eq) apply (metis down_cast_same is_up_down scast_eq ucast_down_wi) done lemma scast_down_wi [OF refl]: "uc = scast \ is_down uc \ uc (word_of_int x) = word_of_int x" by (metis down_cast_same is_up_down ucast_down_wi) lemmas cast_simps = is_down is_up scast_down_add scast_down_minus scast_down_mult ucast_down_add ucast_down_minus ucast_down_mult scast_ucast_1 scast_ucast_3 scast_ucast_4 ucast_scast_1 ucast_scast_3 ucast_scast_4 ucast_ucast_a ucast_ucast_b scast_scast_a scast_scast_b ucast_down_bl ucast_down_wi scast_down_wi ucast_of_nat scast_of_nat uint_up_ucast sint_up_scast up_scast_surj up_ucast_surj lemma smod_mod_positive: "\ 0 \ (a :: int); 0 \ b \ \ a smod b = a mod b" by (clarsimp simp: smod_int_alt_def zsgn_def) lemma nat_mult_power_less_eq: "b > 0 \ (a * b ^ n < (b :: nat) ^ m) = (a < b ^ (m - n))" using mult_less_cancel2[where m = a and k = "b ^ n" and n="b ^ (m - n)"] mult_less_cancel2[where m="a * b ^ (n - m)" and k="b ^ m" and n=1] apply (simp only: power_add[symmetric] nat_minus_add_max) apply (simp only: power_add[symmetric] nat_minus_add_max ac_simps) apply (simp add: max_def split: if_split_asm) done lemma signed_shift_guard_to_word: "\ n < len_of TYPE ('a); n > 0 \ \ (unat (x :: 'a :: len word) * 2 ^ y < 2 ^ n) = (x = 0 \ x < (1 << n >> y))" apply (simp only: nat_mult_power_less_eq) apply (cases "y \ n") apply (simp only: shiftl_shiftr1) apply (subst less_mask_eq) apply (simp add: word_less_nat_alt word_size) apply (rule order_less_le_trans[rotated], rule power_increasing[where n=1]) apply simp apply simp apply simp apply (simp add: nat_mult_power_less_eq word_less_nat_alt word_size) apply auto[1] apply (simp only: shiftl_shiftr2, simp add: unat_eq_0) done lemma sint_ucast_eq_uint: "\ \ is_down (ucast :: ('a::len word \ 'b::len word)) \ \ sint ((ucast :: ('a::len word \ 'b::len word)) x) = uint x" apply (subst sint_eq_uint) apply (simp add: msb_word_eq) apply transfer apply (simp add: bit_take_bit_iff) apply transfer apply simp done lemma word_less_nowrapI': "(x :: 'a :: len word) \ z - k \ k \ z \ 0 < k \ x < x + k" by uint_arith lemma mask_plus_1: "mask n + 1 = (2 ^ n :: 'a::len word)" by (clarsimp simp: mask_eq_decr_exp) lemma unat_inj: "inj unat" by (metis eq_iff injI word_le_nat_alt) lemma unat_ucast_upcast: "is_up (ucast :: 'b word \ 'a word) \ unat (ucast x :: ('a::len) word) = unat (x :: ('b::len) word)" unfolding ucast_eq unat_eq_nat_uint apply (subst int_word_uint) apply (subst mod_pos_pos_trivial) apply simp apply (rule lt2p_lem) apply (clarsimp simp: is_up) apply simp done lemma ucast_mono: "\ (x :: 'b :: len word) < y; y < 2 ^ LENGTH('a) \ \ ucast x < ((ucast y) :: 'a :: len word)" apply (simp only: flip: ucast_nat_def) apply (rule of_nat_mono_maybe) apply (rule unat_less_helper) apply (simp add: Power.of_nat_power) apply (simp add: word_less_nat_alt) done lemma ucast_mono_le: "\x \ y; y < 2 ^ LENGTH('b)\ \ (ucast (x :: 'a :: len word) :: 'b :: len word) \ ucast y" apply (simp only: flip: ucast_nat_def) apply (subst of_nat_mono_maybe_le[symmetric]) apply (rule unat_less_helper) apply (simp add: Power.of_nat_power) apply (rule unat_less_helper) apply (erule le_less_trans) apply (simp add: Power.of_nat_power) apply (simp add: word_le_nat_alt) done lemma ucast_mono_le': "\ unat y < 2 ^ LENGTH('b); LENGTH('b::len) < LENGTH('a::len); x \ y \ \ UCAST('a \ 'b) x \ UCAST('a \ 'b) y" by (auto simp: word_less_nat_alt intro: ucast_mono_le) lemma zero_sle_ucast_up: "\ is_down (ucast :: 'a word \ 'b signed word) \ (0 <=s ((ucast (b::('a::len) word)) :: ('b::len) signed word))" apply (subgoal_tac "\ msb (ucast b :: 'b signed word)") apply (clarsimp simp: word_sle_msb_le) apply (clarsimp simp: is_down not_le msb_nth nth_ucast) done lemma word_le_ucast_sless: "\ x \ y; y \ -1; LENGTH('a) < LENGTH('b) \ \ UCAST (('a :: len) \ ('b :: len) signed) x msb (ucast x :: ('a::len) word) = msb (x :: ('b::len) word)" apply (clarsimp simp: word_msb_alt) apply (subst ucast_down_drop [where n=0]) apply (clarsimp simp: source_size_def target_size_def word_size) apply clarsimp done lemma msb_big: "msb (a :: ('a::len) word) = (a \ 2 ^ (LENGTH('a) - Suc 0))" apply (rule iffI) apply (clarsimp simp: msb_nth) apply (drule bang_is_le) apply simp apply (rule ccontr) apply (subgoal_tac "a = a && mask (LENGTH('a) - Suc 0)") apply (cut_tac and_mask_less' [where w=a and n="LENGTH('a) - Suc 0"]) apply (clarsimp simp: word_not_le [symmetric]) apply clarsimp apply (rule sym, subst and_mask_eq_iff_shiftr_0) apply (clarsimp simp: msb_shift) done lemma zero_sle_ucast: "(0 <=s ((ucast (b::('a::len) word)) :: ('a::len) signed word)) = (uint b < 2 ^ (LENGTH('a) - 1))" apply (case_tac "msb b") apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI) apply (clarsimp simp: msb_big word_le_def uint_2p_alt) apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI) apply (clarsimp simp: msb_big word_le_def uint_2p_alt) done (* to_bool / from_bool. *) definition from_bool :: "bool \ 'a::len word" where "from_bool b \ case b of True \ of_nat 1 | False \ of_nat 0" lemma from_bool_eq: \from_bool = of_bool\ by (simp add: fun_eq_iff from_bool_def) lemma from_bool_0: "(from_bool x = 0) = (\ x)" by (simp add: from_bool_def split: bool.split) definition to_bool :: "'a::len word \ bool" where "to_bool \ (\) 0" lemma to_bool_and_1: "to_bool (x && 1) = (x !! 0)" by (simp add: test_bit_word_eq to_bool_def and_one_eq mod_2_eq_odd) lemma to_bool_from_bool [simp]: "to_bool (from_bool r) = r" unfolding from_bool_def to_bool_def by (simp split: bool.splits) lemma from_bool_neq_0 [simp]: "(from_bool b \ 0) = b" by (simp add: from_bool_def split: bool.splits) lemma from_bool_mask_simp [simp]: "(from_bool r :: 'a::len word) && 1 = from_bool r" unfolding from_bool_def by (clarsimp split: bool.splits) lemma from_bool_1 [simp]: "(from_bool P = 1) = P" by (simp add: from_bool_def split: bool.splits) lemma ge_0_from_bool [simp]: "(0 < from_bool P) = P" by (simp add: from_bool_def split: bool.splits) lemma limited_and_from_bool: "limited_and (from_bool b) 1" by (simp add: from_bool_def limited_and_def split: bool.split) lemma to_bool_1 [simp]: "to_bool 1" by (simp add: to_bool_def) lemma to_bool_0 [simp]: "\to_bool 0" by (simp add: to_bool_def) lemma from_bool_eq_if: "(from_bool Q = (if P then 1 else 0)) = (P = Q)" - by (simp add: case_bool_If from_bool_def split: if_split) + by (cases Q) (simp_all add: from_bool_def) lemma to_bool_eq_0: "(\ to_bool x) = (x = 0)" by (simp add: to_bool_def) lemma to_bool_neq_0: "(to_bool x) = (x \ 0)" by (simp add: to_bool_def) lemma from_bool_all_helper: "(\bool. from_bool bool = val \ P bool) = ((\bool. from_bool bool = val) \ P (val \ 0))" by (auto simp: from_bool_0) lemma from_bool_to_bool_iff: "w = from_bool b \ to_bool w = b \ (w = 0 \ w = 1)" by (cases b) (auto simp: from_bool_def to_bool_def) lemma from_bool_eqI: "from_bool x = from_bool y \ x = y" unfolding from_bool_def by (auto split: bool.splits) lemma word_rsplit_upt: "\ size x = LENGTH('a :: len) * n; n \ 0 \ \ word_rsplit x = map (\i. ucast (x >> i * len_of TYPE ('a)) :: 'a word) (rev [0 ..< n])" apply (subgoal_tac "length (word_rsplit x :: 'a word list) = n") apply (rule nth_equalityI, simp) apply (intro allI word_eqI impI) apply (simp add: test_bit_rsplit_alt word_size) apply (simp add: nth_ucast nth_shiftr rev_nth field_simps) apply (simp add: length_word_rsplit_exp_size) apply (metis mult.commute given_quot_alt word_size word_size_gt_0) done lemma aligned_shift: "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ x + y >> n = y >> n" by (subst word_plus_and_or_coroll; word_eqI, blast) lemma aligned_shift': "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ y + x >> n = y >> n" by (subst word_plus_and_or_coroll; word_eqI, blast) lemma neg_mask_add_mask: "((x:: 'a :: len word) && ~~ (mask n)) + (2 ^ n - 1) = x || mask n" unfolding mask_2pm1[symmetric] by (subst word_plus_and_or_coroll; word_eqI_solve) lemma subtract_mask: "p - (p && mask n) = (p && ~~ (mask n))" "p - (p && ~~ (mask n)) = (p && mask n)" by (simp add: field_simps word_plus_and_or_coroll2)+ lemma and_neg_mask_plus_mask_mono: "(p && ~~ (mask n)) + mask n \ p" apply (rule word_le_minus_cancel[where x = "p && ~~ (mask n)"]) apply (clarsimp simp: subtract_mask) using word_and_le1[where a = "mask n" and y = p] apply (clarsimp simp: mask_eq_decr_exp word_le_less_eq) apply (rule is_aligned_no_overflow'[folded mask_2pm1]) apply (clarsimp simp: is_aligned_neg_mask) done lemma word_neg_and_le: "ptr \ (ptr && ~~ (mask n)) + (2 ^ n - 1)" by (simp add: and_neg_mask_plus_mask_mono mask_2pm1[symmetric]) lemma aligned_less_plus_1: "\ is_aligned x n; n > 0 \ \ x < x + 1" apply (rule plus_one_helper2) apply (rule order_refl) apply (clarsimp simp: field_simps) apply (drule arg_cong[where f="\x. x - 1"]) apply (clarsimp simp: is_aligned_mask) apply (drule word_eqD[where x=0]) apply simp done lemma aligned_add_offset_less: "\is_aligned x n; is_aligned y n; x < y; z < 2 ^ n\ \ x + z < y" apply (cases "y = 0") apply simp apply (erule is_aligned_get_word_bits[where p=y], simp_all) apply (cases "z = 0", simp_all) apply (drule(2) aligned_at_least_t2n_diff[rotated -1]) apply (drule plus_one_helper2) apply (rule less_is_non_zero_p1) apply (rule aligned_less_plus_1) apply (erule aligned_sub_aligned[OF _ _ order_refl], simp_all add: is_aligned_triv)[1] apply (cases n, simp_all)[1] apply (simp only: trans[OF diff_add_eq diff_diff_eq2[symmetric]]) apply (drule word_less_add_right) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0, erule order_less_trans) apply (clarsimp simp: power_overflow) apply simp apply (erule order_le_less_trans[rotated], rule word_plus_mono_right) apply (erule word_le_minus_one_leq) apply (simp add: is_aligned_no_wrap' is_aligned_no_overflow field_simps) done lemma is_aligned_add_helper: "\ is_aligned p n; d < 2 ^ n \ \ (p + d && mask n = d) \ (p + d && (~~ (mask n)) = p)" apply (subst(asm) is_aligned_mask) apply (drule less_mask_eq) apply (rule context_conjI) apply (subst word_plus_and_or_coroll; word_eqI; blast) using word_plus_and_or_coroll2[where x="p + d" and w="mask n"] by simp lemma is_aligned_sub_helper: "\ is_aligned (p - d) n; d < 2 ^ n \ \ (p && mask n = d) \ (p && (~~ (mask n)) = p - d)" by (drule(1) is_aligned_add_helper, simp) lemma mask_twice: "(x && mask n) && mask m = x && mask (min m n)" by word_eqI_solve lemma is_aligned_after_mask: "\is_aligned k m;m\ n\ \ is_aligned (k && mask n) m" by (rule is_aligned_andI1) lemma and_mask_plus: "\is_aligned ptr m; m \ n; a < 2 ^ m\ \ ptr + a && mask n = (ptr && mask n) + a" apply (rule mask_eqI[where n = m]) apply (simp add:mask_twice min_def) apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct1]) apply (erule is_aligned_after_mask) apply simp apply simp apply simp apply (subgoal_tac "(ptr + a && mask n) && ~~ (mask m) = (ptr + a && ~~ (mask m) ) && mask n") apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct2]) apply (simp add:is_aligned_after_mask) apply simp apply simp apply (simp add:word_bw_comms word_bw_lcs) done lemma le_step_down_word:"\(i::('a::len) word) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by unat_arith lemma le_step_down_word_2: fixes x :: "'a::len word" shows "\x \ y; x \ y\ \ x \ y - 1" by (subst (asm) word_le_less_eq, clarsimp, simp add: word_le_minus_one_leq) lemma NOT_mask_AND_mask[simp]: "(w && mask n) && ~~ (mask n) = 0" by (clarsimp simp add: mask_eq_decr_exp Parity.bit_eq_iff bit_and_iff bit_not_iff bit_mask_iff) lemma and_and_not[simp]:"(a && b) && ~~ b = 0" apply (subst word_bw_assocs(1)) apply clarsimp done lemma mask_shift_and_negate[simp]:"(w && mask n << m) && ~~ (mask n << m) = 0" by (clarsimp simp add: mask_eq_decr_exp Parity.bit_eq_iff bit_and_iff bit_not_iff shiftl_word_eq bit_push_bit_iff) lemma le_step_down_nat:"\(i::nat) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith lemma le_step_down_int:"\(i::int) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith lemma ex_mask_1[simp]: "(\x. mask x = (1 :: 'a::len word))" apply (rule_tac x=1 in exI) apply (simp add:mask_eq_decr_exp) done lemma not_switch:"~~ a = x \ a = ~~ x" by auto (* The seL4 bitfield generator produces functions containing mask and shift operations, such that * invoking two of them consecutively can produce something like the following. *) lemma bitfield_op_twice: "(x && ~~ (mask n << m) || ((y && mask n) << m)) && ~~ (mask n << m) = x && ~~ (mask n << m)" by (induct n arbitrary: m) (auto simp: word_ao_dist) lemma bitfield_op_twice'': "\~~ a = b << c; \x. b = mask x\ \ (x && a || (y && b << c)) && a = x && a" apply clarsimp apply (cut_tac n=xa and m=c and x=x and y=y in bitfield_op_twice) apply (clarsimp simp:mask_eq_decr_exp) apply (drule not_switch) apply clarsimp done lemma bit_twiddle_min: "(y::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = min x y" by (auto simp add: Parity.bit_eq_iff bit_xor_iff min_def) lemma bit_twiddle_max: "(x::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = max x y" by (auto simp add: Parity.bit_eq_iff bit_xor_iff max_def) lemma swap_with_xor: "\(x::'a::len word) = a xor b; y = b xor x; z = x xor y\ \ z = b \ y = a" by (auto simp add: Parity.bit_eq_iff bit_xor_iff max_def) lemma scast_nop1 [simp]: "((scast ((of_int x)::('a::len) word))::'a sword) = of_int x" apply (simp only: scast_eq) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemma scast_nop2 [simp]: "((scast ((of_int x)::('a::len) sword))::'a word) = of_int x" apply (simp only: scast_eq) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemmas scast_nop = scast_nop1 scast_nop2 scast_id lemma le_mask_imp_and_mask: "(x::'a::len word) \ mask n \ x && mask n = x" by (metis and_mask_eq_iff_le_mask) lemma or_not_mask_nop: "((x::'a::len word) || ~~ (mask n)) && mask n = x && mask n" by (metis word_and_not word_ao_dist2 word_bw_comms(1) word_log_esimps(3)) lemma mask_subsume: "\n \ m\ \ ((x::'a::len word) || y && mask n) && ~~ (mask m) = x && ~~ (mask m)" by (auto simp add: Parity.bit_eq_iff bit_not_iff bit_or_iff bit_and_iff bit_mask_iff) lemma and_mask_0_iff_le_mask: fixes w :: "'a::len word" shows "(w && ~~(mask n) = 0) = (w \ mask n)" by (simp add: mask_eq_0_eq_x le_mask_imp_and_mask and_mask_eq_iff_le_mask) lemma mask_twice2: "n \ m \ ((x::'a::len word) && mask m) && mask n = x && mask n" by (metis mask_twice min_def) lemma uint_2_id: "LENGTH('a) \ 2 \ uint (2::('a::len) word) = 2" by simp lemma bintrunc_id: "\m \ of_nat n; 0 < m\ \ bintrunc n m = m" by (simp add: bintrunc_mod2p le_less_trans) lemma shiftr1_unfold: "shiftr1 x = x >> 1" by (metis One_nat_def comp_apply funpow.simps(1) funpow.simps(2) id_apply shiftr_def) lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2" by transfer (simp add: drop_bit_Suc) lemma shiftl1_is_mult: "(x << 1) = (x :: 'a::len word) * 2" by (metis One_nat_def mult_2 mult_2_right one_add_one power_0 power_Suc shiftl_t2n) lemma div_of_0_id[simp]:"(0::('a::len) word) div n = 0" by (simp add: word_div_def) lemma degenerate_word:"LENGTH('a) = 1 \ (x::('a::len) word) = 0 \ x = 1" by (metis One_nat_def less_irrefl_nat sint_1_cases) lemma div_by_0_word:"(x::('a::len) word) div 0 = 0" by (metis div_0 div_by_0 unat_0 word_arith_nat_defs(6) word_div_1) lemma div_less_dividend_word:"\x \ 0; n \ 1\ \ (x::('a::len) word) div n < x" apply (cases \n = 0\) apply clarsimp apply (simp add:word_neq_0_conv) apply (subst word_arith_nat_div) apply (rule word_of_nat_less) apply (rule div_less_dividend) using unat_eq_zero word_unat_Rep_inject1 apply force apply (simp add:unat_gt_0) done lemma shiftr1_lt:"x \ 0 \ (x::('a::len) word) >> 1 < x" apply (subst shiftr1_is_div_2) apply (rule div_less_dividend_word) apply simp+ done lemma word_less_div: fixes x :: "('a::len) word" and y :: "('a::len) word" shows "x div y = 0 \ y = 0 \ x < y" apply (case_tac "y = 0", clarsimp+) by (metis One_nat_def Suc_le_mono le0 le_div_geq not_less unat_0 unat_div unat_gt_0 word_less_nat_alt zero_less_one) lemma not_degenerate_imp_2_neq_0:"LENGTH('a) > 1 \ (2::('a::len) word) \ 0" by (metis numerals(1) power_not_zero power_zero_numeral) lemma shiftr1_0_or_1:"(x::('a::len) word) >> 1 = 0 \ x = 0 \ x = 1" apply (subst (asm) shiftr1_is_div_2) apply (drule word_less_div) apply (case_tac "LENGTH('a) = 1") apply (simp add:degenerate_word) apply (erule disjE) apply (subgoal_tac "(2::'a word) \ 0") apply simp apply (rule not_degenerate_imp_2_neq_0) apply (subgoal_tac "LENGTH('a) \ 0") apply arith apply simp apply (rule x_less_2_0_1', simp+) done lemma word_overflow:"(x::('a::len) word) + 1 > x \ x + 1 = 0" apply clarsimp by (metis diff_0 eq_diff_eq less_x_plus_1) lemma word_overflow_unat:"unat ((x::('a::len) word) + 1) = unat x + 1 \ x + 1 = 0" by (metis Suc_eq_plus1 add.commute unatSuc) lemma even_word_imp_odd_next:"even (unat (x::('a::len) word)) \ x + 1 = 0 \ odd (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma odd_word_imp_even_next:"odd (unat (x::('a::len) word)) \ x + 1 = 0 \ even (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma overflow_imp_lsb:"(x::('a::len) word) + 1 = 0 \ x !! 0" using even_plus_one_iff [of x] by (simp add: test_bit_word_eq) lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = x !! 0" by transfer (simp add: even_nat_iff) lemma of_nat_neq_iff_word: "x mod 2 ^ LENGTH('a) \ y mod 2 ^ LENGTH('a) \ (((of_nat x)::('a::len) word) \ of_nat y) = (x \ y)" apply (rule iffI) apply (case_tac "x = y") apply (subst (asm) of_nat_eq_iff[symmetric]) apply simp+ apply (case_tac "((of_nat x)::('a::len) word) = of_nat y") apply (subst (asm) word_unat.norm_eq_iff[symmetric]) apply simp+ done lemma shiftr1_irrelevant_lsb:"(x::('a::len) word) !! 0 \ x >> 1 = (x + 1) >> 1" apply (cases \LENGTH('a)\; transfer) apply (simp_all add: take_bit_drop_bit) apply (simp add: drop_bit_take_bit drop_bit_Suc) done lemma shiftr1_0_imp_only_lsb:"((x::('a::len) word) + 1) >> 1 = 0 \ x = 0 \ x + 1 = 0" by (metis One_nat_def shiftr1_0_or_1 word_less_1 word_overflow) lemma shiftr1_irrelevant_lsb':"\((x::('a::len) word) !! 0) \ x >> 1 = (x + 1) >> 1" by (metis shiftr1_irrelevant_lsb) lemma lsb_this_or_next:"\(((x::('a::len) word) + 1) !! 0) \ x !! 0" by (metis (poly_guards_query) even_word_imp_odd_next odd_iff_lsb overflow_imp_lsb) (* Perhaps this one should be a simp lemma, but it seems a little dangerous. *) lemma cast_chunk_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((ucast (x::'b word))::'a word))::'b word) || (((ucast ((ucast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((ucast ((ucast (x >> n))::'a word))::'b word) = x >> n") apply clarsimp apply (subst and_not_mask[symmetric]) apply (subst ucast_ucast_mask) apply (subst word_ao_dist2[symmetric]) apply clarsimp apply (rule ucast_ucast_len) apply (rule shiftr_less_t2n') apply (subst and_mask_eq_iff_le_mask) apply (simp_all add: mask_eq_decr_exp flip: mult_2_right) apply (metis add_diff_cancel_left' len_gt_0 mult_2_right zero_less_diff) done lemma cast_chunk_scast_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((scast (x::'b word))::'a word))::'b word) || (((ucast ((scast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((scast x)::'a word) = ((ucast x)::'a word)") apply (subgoal_tac "((scast (x >> n))::'a word) = ((ucast (x >> n))::'a word)") apply (simp add:cast_chunk_assemble_id) apply (subst down_cast_same[symmetric], subst is_down, arith, simp)+ done lemma mask_or_not_mask: "x && mask n || x && ~~ (mask n) = x" apply (subst word_oa_dist, simp) apply (subst word_oa_dist2, simp) done lemma is_aligned_add_not_aligned: "\is_aligned (p::'a::len word) n; \ is_aligned (q::'a::len word) n\ \ \ is_aligned (p + q) n" by (metis is_aligned_addD1) lemma word_gr0_conv_Suc: "(m::'a::len word) > 0 \ \n. m = n + 1" by (metis add.commute add_minus_cancel) lemma neg_mask_add_aligned: "\ is_aligned p n; q < 2 ^ n \ \ (p + q) && ~~ (mask n) = p && ~~ (mask n)" by (metis is_aligned_add_helper is_aligned_neg_mask_eq) lemma word_sless_sint_le:"x sint x \ sint y - 1" by (metis word_sless_alt zle_diff1_eq) lemma upper_trivial: fixes x :: "'a::len word" shows "x \ 2 ^ LENGTH('a) - 1 \ x < 2 ^ LENGTH('a) - 1" by (simp add: less_le) lemma constraint_expand: fixes x :: "'a::len word" shows "x \ {y. lower \ y \ y \ upper} = (lower \ x \ x \ upper)" by (rule mem_Collect_eq) lemma card_map_elide: "card ((of_nat :: nat \ 'a::len word) ` {0.. CARD('a::len word)" proof - let ?of_nat = "of_nat :: nat \ 'a word" from word_unat.Abs_inj_on have "inj_on ?of_nat {i. i < CARD('a word)}" by (simp add: unats_def card_word) moreover have "{0.. {i. i < CARD('a word)}" using that by auto ultimately have "inj_on ?of_nat {0.. CARD('a::len word) \ card ((of_nat::nat \ 'a::len word) ` {0..UCAST('b \ 'a) (UCAST('a \ 'b) x) = x\ if \x \ UCAST('b::len \ 'a) (- 1)\ for x :: \'a::len word\ proof - from that have a1: \x \ word_of_int (uint (word_of_int (2 ^ LENGTH('b) - 1) :: 'b word))\ by simp have f2: "((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (0::int) \ - 1 + 2 ^ LENGTH('b) \ (0::int) \ - 1 + 2 ^ LENGTH('b) + - 1 * 2 ^ LENGTH('b) \ (- (1::int) + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b) = - 1 + 2 ^ LENGTH('b)) = ((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (1::int) \ 2 ^ LENGTH('b) \ 2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)) = 1)" by force have f3: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * ia \ i mod ia = i" using mod_pos_pos_trivial by force have "(1::int) \ 2 ^ LENGTH('b)" by simp then have "2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ len_of TYPE ('b)) = 1" using f3 f2 by blast then have f4: "- (1::int) + 2 ^ LENGTH('b) = (- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)" by linarith have f5: "x \ word_of_int (uint (word_of_int (- 1 + 2 ^ LENGTH('b))::'b word))" using a1 by force have f6: "2 ^ LENGTH('b) + - (1::int) = - 1 + 2 ^ LENGTH('b)" by force have f7: "- (1::int) * 1 = - 1" by auto have "\x0 x1. (x1::int) - x0 = x1 + - 1 * x0" by force then have "x \ 2 ^ LENGTH('b) - 1" using f7 f6 f5 f4 by (metis uint_word_of_int wi_homs(2) word_arith_wis(8) word_of_int_2p) then have \uint x \ uint (2 ^ LENGTH('b) - (1 :: 'a word))\ by (simp add: word_le_def) then have \uint x \ 2 ^ LENGTH('b) - 1\ by (simp add: uint_word_ariths) (metis \1 \ 2 ^ LENGTH('b)\ \uint x \ uint (2 ^ LENGTH('b) - 1)\ linorder_not_less lt2p_lem uint_1 uint_minus_simple_alt uint_power_lower word_le_def zle_diff1_eq) then show ?thesis apply (simp add: word_ubin.eq_norm bintrunc_mod2p unsigned_ucast_eq) by (metis \x \ 2 ^ LENGTH('b) - 1\ le_def take_bit_word_eq_self ucast_ucast_len unsigned_take_bit_eq word_less_sub_le word_ubin.norm_Rep word_uint_eqI) qed lemma remdups_enum_upto: fixes s::"'a::len word" shows "remdups [s .e. e] = [s .e. e]" by simp lemma card_enum_upto: fixes s::"'a::len word" shows "card (set [s .e. e]) = Suc (unat e) - unat s" by (subst List.card_set) (simp add: remdups_enum_upto) lemma unat_mask: "unat (mask n :: 'a :: len word) = 2 ^ (min n (LENGTH('a))) - 1" apply (subst min.commute) apply (simp add: mask_eq_decr_exp not_less min_def split: if_split_asm) apply (intro conjI impI) apply (simp add: unat_sub_if_size) apply (simp add: power_overflow word_size) apply (simp add: unat_sub_if_size) done lemma word_shiftr_lt: fixes w :: "'a::len word" shows "unat (w >> n) < (2 ^ (LENGTH('a) - n))" apply (subst shiftr_div_2n') by (metis nat_mod_lem nat_zero_less_power_iff power_mod_div word_unat.Rep_inverse word_unat.eq_norm zero_less_numeral) lemma complement_nth_w2p: shows "n' < LENGTH('a) \ (~~ (2 ^ n :: 'a::len word)) !! n' = (n' \ n)" by (fastforce simp: word_ops_nth_size word_size nth_w2p) lemma word_unat_and_lt: "unat x < n \ unat y < n \ unat (x && y) < n" by (meson le_less_trans word_and_le1 word_and_le2 word_le_nat_alt) lemma word_unat_mask_lt: "m \ size w \ unat ((w::'a::len word) && mask m) < 2 ^ m" by (rule word_unat_and_lt) (simp add: unat_mask word_size) lemma unat_shiftr_less_t2n: fixes x :: "'a :: len word" shows "unat x < 2 ^ (n + m) \ unat (x >> n) < 2 ^ m" by (simp add: shiftr_div_2n' power_add mult.commute td_gal_lt) lemma le_or_mask: "w \ w' \ w || mask x \ w' || mask x" by (metis neg_mask_add_mask add.commute le_word_or1 mask_2pm1 neg_mask_mono_le word_plus_mono_left) lemma le_shiftr1': "\ shiftr1 u \ shiftr1 v ; shiftr1 u \ shiftr1 v \ \ u \ v" apply transfer apply simp done lemma le_shiftr': "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len word) \ v" apply (induct n; simp add: shiftr_def) apply (case_tac "(shiftr1 ^^ n) u = (shiftr1 ^^ n) v", simp) apply (fastforce dest: le_shiftr1') done lemma word_log2_nth_same: "w \ 0 \ w !! word_log2 w" unfolding word_log2_def using nth_length_takeWhile[where P=Not and xs="to_bl w"] apply (simp add: word_clz_def word_size to_bl_nth) apply (fastforce simp: linorder_not_less eq_zero_set_bl dest: takeWhile_take_has_property) done lemma word_log2_nth_not_set: "\ word_log2 w < i ; i < size w \ \ \ w !! i" unfolding word_log2_def word_clz_def using takeWhile_take_has_property_nth[where P=Not and xs="to_bl w" and n="size w - Suc i"] by (fastforce simp add: to_bl_nth word_size) lemma word_log2_highest: assumes a: "w !! i" shows "i \ word_log2 w" proof - from a have "i < size w" by - (rule test_bit_size) with a show ?thesis by - (rule ccontr, simp add: word_log2_nth_not_set) qed lemma word_log2_max: "word_log2 w < size w" unfolding word_log2_def word_clz_def by simp lemma word_clz_0[simp]: "word_clz (0::'a::len word) = LENGTH('a)" unfolding word_clz_def by (simp add: takeWhile_replicate) lemma word_clz_minus_one[simp]: "word_clz (-1::'a::len word) = 0" unfolding word_clz_def by (simp add: takeWhile_replicate) lemma word_add_no_overflow:"(x::'a::len word) < max_word \ x < x + 1" using less_x_plus_1 order_less_le by blast lemma lt_plus_1_le_word: fixes x :: "'a::len word" assumes bound:"n < unat (maxBound::'a word)" shows "x < 1 + of_nat n = (x \ of_nat n)" by (metis add.commute bound max_word_max word_Suc_leq word_not_le word_of_nat_less) lemma unat_ucast_up_simp: fixes x :: "'a::len word" assumes "LENGTH('a) \ LENGTH('b)" shows "unat (ucast x :: 'b::len word) = unat x" unfolding ucast_eq unat_eq_nat_uint apply (subst int_word_uint) apply (subst mod_pos_pos_trivial; simp?) apply (rule lt2p_lem) apply (simp add: assms) done lemma unat_ucast_less_no_overflow: "\n < 2 ^ LENGTH('a); unat f < n\ \ (f::('a::len) word) < of_nat n" by (erule (1) order_le_less_trans[OF _ of_nat_mono_maybe,rotated]) simp lemma unat_ucast_less_no_overflow_simp: "n < 2 ^ LENGTH('a) \ (unat f < n) = ((f::('a::len) word) < of_nat n)" using unat_less_helper unat_ucast_less_no_overflow by blast lemma unat_ucast_no_overflow_le: assumes no_overflow: "unat b < (2 :: nat) ^ LENGTH('a)" and upward_cast: "LENGTH('a) < LENGTH('b)" shows "(ucast (f::'a::len word) < (b :: 'b :: len word)) = (unat f < unat b)" proof - have LR: "ucast f < b \ unat f < unat b" apply (rule unat_less_helper) apply (simp add:ucast_nat_def) apply (rule_tac 'b1 = 'b in ucast_less_ucast[OF order.strict_implies_order, THEN iffD1]) apply (rule upward_cast) apply (simp add: ucast_ucast_mask less_mask_eq word_less_nat_alt unat_power_lower[OF upward_cast] no_overflow) done have RL: "unat f < unat b \ ucast f < b" proof- assume ineq: "unat f < unat b" have "ucast (f::'a::len word) < ((ucast (ucast b ::'a::len word)) :: 'b :: len word)" apply (simp add: ucast_less_ucast[OF order.strict_implies_order] upward_cast) apply (simp only: flip: ucast_nat_def) apply (rule unat_ucast_less_no_overflow[OF no_overflow ineq]) done then show ?thesis apply (rule order_less_le_trans) apply (simp add:ucast_ucast_mask word_and_le2) done qed then show ?thesis by (simp add:RL LR iffI) qed lemmas ucast_up_mono = ucast_less_ucast[THEN iffD2] (* casting a long word to a shorter word and casting back to the long word is equal to the original long word -- if the word is small enough. 'l is the longer word. 's is the shorter word. *) lemma bl_cast_long_short_long_ingoreLeadingZero_generic: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (of_bl :: _ \ 'l::len word) (to_bl ((of_bl::_ \ 's::len word) (to_bl w))) = w" by (rule word_uint_eqI) (simp add: uint_of_bl_is_bl_to_bin uint_of_bl_is_bl_to_bin_drop) (* Casting between longer and shorter word. 'l is the longer word. 's is the shorter word. For example: 'l::len word is 128 word (full ipv6 address) 's::len word is 16 word (address piece of ipv6 address in colon-text-representation) *) corollary ucast_short_ucast_long_ingoreLeadingZero: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (ucast:: 's::len word \ 'l::len word) ((ucast:: 'l::len word \ 's::len word) w) = w" apply (subst ucast_bl)+ apply (rule bl_cast_long_short_long_ingoreLeadingZero_generic; simp) done lemma length_drop_mask: fixes w::"'a::len word" shows "length (dropWhile Not (to_bl (w AND mask n))) \ n" proof - have "length (takeWhile Not (replicate n False @ ls)) = n + length (takeWhile Not ls)" for ls n by(subst takeWhile_append2) simp+ then show ?thesis unfolding bl_and_mask by (simp add: dropWhile_eq_drop) qed lemma minus_one_word: "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by simp lemma mask_exceed: "n \ LENGTH('a) \ (x::'a::len word) && ~~ (mask n) = 0" by (simp add: and_not_mask shiftr_eq_0) -lemma two_power_strict_part_mono: - "strict_part_mono {..LENGTH('a) - 1} (\x. (2 :: 'a :: len word) ^ x)" -proof - - { fix n - have "n < LENGTH('a) \ strict_part_mono {..n} (\x. (2 :: 'a :: len word) ^ x)" - proof (induct n) - case 0 then show ?case by simp - next - case (Suc n) - from Suc.prems - have "2 ^ n < (2 :: 'a :: len word) ^ Suc n" - using power_strict_increasing unat_power_lower word_less_nat_alt by fastforce - with Suc - show ?case by (subst strict_part_mono_by_steps) simp - qed - } - then show ?thesis by simp -qed - lemma word_shift_by_2: "x * 4 = (x::'a::len word) << 2" by (simp add: shiftl_t2n) lemma le_2p_upper_bits: "\ (p::'a::len word) \ 2^n - 1; n < LENGTH('a) \ \ \n'\n. n' < LENGTH('a) \ \ p !! n'" by (subst upper_bits_unset_is_l2p; simp) lemma le2p_bits_unset: "p \ 2 ^ n - 1 \ \n'\n. n' < LENGTH('a) \ \ (p::'a::len word) !! n'" using upper_bits_unset_is_l2p [where p=p] by (cases "n < LENGTH('a)") auto lemma ucast_less_shiftl_helper: "\ LENGTH('b) + 2 < LENGTH('a); 2 ^ (LENGTH('b) + 2) \ n\ \ (ucast (x :: 'b::len word) << 2) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done lemma word_power_nonzero: "\ (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x \ 0 \ \ x * 2 ^ n \ 0" by (metis and_mask_eq_iff_shiftr_0 less_mask_eq p2_gt_0 semiring_normalization_rules(7) shiftl_shiftr_id shiftl_t2n) lemma less_1_helper: "n \ m \ (n - 1 :: int) < m" by arith lemma div_power_helper: "\ x \ y; y < LENGTH('a) \ \ (2 ^ y - 1) div (2 ^ x :: 'a::len word) = 2 ^ (y - x) - 1" apply (rule word_uint.Rep_eqD) apply (simp only: uint_word_ariths uint_div uint_power_lower) apply (subst mod_pos_pos_trivial, fastforce, fastforce)+ apply (subst mod_pos_pos_trivial) apply (simp add: le_diff_eq uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst mod_pos_pos_trivial) apply (simp add: uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst int_div_sub_1; simp add: uint_2p_alt) apply (subst power_0[symmetric]) apply (simp add: uint_2p_alt le_imp_power_dvd power_sub_int) done lemma word_add_power_off: fixes a :: "'a :: len word" assumes ak: "a < k" and kw: "k < 2 ^ (LENGTH('a) - m)" and mw: "m < LENGTH('a)" and off: "off < 2 ^ m" shows "(a * 2 ^ m) + off < k * 2 ^ m" proof (cases "m = 0") case True then show ?thesis using off ak by simp next case False from ak have ak1: "a + 1 \ k" by (rule inc_le) then have "(a + 1) * 2 ^ m \ 0" apply - apply (rule word_power_nonzero) apply (erule order_le_less_trans [OF _ kw]) apply (rule mw) apply (rule less_is_non_zero_p1 [OF ak]) done then have "(a * 2 ^ m) + off < ((a + 1) * 2 ^ m)" using kw mw apply - apply (simp add: distrib_right) apply (rule word_plus_strict_mono_right [OF off]) apply (rule is_aligned_no_overflow'') apply (rule is_aligned_mult_triv2) apply assumption done also have "\ \ k * 2 ^ m" using ak1 mw kw False apply - apply (erule word_mult_le_mono1) apply (simp add: p2_gt_0) apply (simp add: word_less_nat_alt) apply (rule nat_less_power_trans2[simplified]) apply (simp add: word_less_nat_alt) apply simp done finally show ?thesis . qed lemma offset_not_aligned: "\ is_aligned (p::'a::len word) n; i > 0; i < 2 ^ n; n < LENGTH('a)\ \ \ is_aligned (p + of_nat i) n" apply (erule is_aligned_add_not_aligned) apply transfer apply (auto simp add: is_aligned_iff_udvd) apply (metis bintrunc_bintrunc_ge int_ops(1) nat_int_comparison(1) nat_less_le take_bit_eq_0_iff take_bit_nat_eq_self_iff take_bit_of_nat) done lemma length_upto_enum_one: fixes x :: "'a :: len word" assumes lt1: "x < y" and lt2: "z < y" and lt3: "x \ z" shows "[x , y .e. z] = [x]" unfolding upto_enum_step_def proof (subst upto_enum_red, subst if_not_P [OF leD [OF lt3]], clarsimp, rule conjI) show "unat ((z - x) div (y - x)) = 0" proof (subst unat_div, rule div_less) have syx: "unat (y - x) = unat y - unat x" by (rule unat_sub [OF order_less_imp_le]) fact moreover have "unat (z - x) = unat z - unat x" by (rule unat_sub) fact ultimately show "unat (z - x) < unat (y - x)" using lt2 lt3 unat_mono word_less_minus_mono_left by blast qed then show "(z - x) div (y - x) * (y - x) = 0" by (metis mult_zero_left unat_0 word_unat.Rep_eqD) qed lemma max_word_mask: "(max_word :: 'a::len word) = mask LENGTH('a)" unfolding mask_eq_decr_exp by simp lemmas mask_len_max = max_word_mask[symmetric] lemma is_aligned_alignUp[simp]: "is_aligned (alignUp p n) n" by (simp add: alignUp_def complement_def is_aligned_mask mask_eq_decr_exp word_bw_assocs) lemma alignUp_le[simp]: "alignUp p n \ p + 2 ^ n - 1" unfolding alignUp_def by (rule word_and_le2) lemma complement_mask: "complement (2 ^ n - 1) = ~~ (mask n)" unfolding complement_def mask_eq_decr_exp by simp lemma alignUp_idem: fixes a :: "'a::len word" assumes "is_aligned a n" "n < LENGTH('a)" shows "alignUp a n = a" using assms unfolding alignUp_def by (metis complement_mask is_aligned_add_helper p_assoc_help power_2_ge_iff) lemma alignUp_not_aligned_eq: fixes a :: "'a :: len word" assumes al: "\ is_aligned a n" and sz: "n < LENGTH('a)" shows "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" proof - have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) fact+ then have um: "unat (a mod 2 ^ n - 1) div 2 ^ n = 0" using sz by (meson Euclidean_Division.div_eq_0_iff le_m1_iff_lt measure_unat order_less_trans unat_less_power word_less_sub_le word_mod_less_divisor) have "a + 2 ^ n - 1 = (a div 2 ^ n) * 2 ^ n + (a mod 2 ^ n) + 2 ^ n - 1" by (simp add: word_mod_div_equality) also have "\ = (a mod 2 ^ n - 1) + (a div 2 ^ n + 1) * 2 ^ n" by (simp add: field_simps) finally show "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" using sz unfolding alignUp_def apply (subst complement_mask) apply (erule ssubst) apply (subst neg_mask_is_div) apply (simp add: word_arith_nat_div) apply (subst unat_word_ariths(1) unat_word_ariths(2))+ apply (subst uno_simps) apply (subst unat_1) apply (subst mod_add_right_eq) apply simp apply (subst power_mod_div) apply (subst div_mult_self1) apply simp apply (subst um) apply simp apply (subst mod_mod_power) apply simp apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst mult_mod_left) apply (subst power_add [symmetric]) apply simp apply (subst Abs_fnat_hom_1) apply (subst Abs_fnat_hom_add) apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst word_unat.Rep_inverse[symmetric], subst Abs_fnat_hom_mult) apply simp done qed lemma alignUp_ge: fixes a :: "'a :: len word" assumes sz: "n < LENGTH('a)" and nowrap: "alignUp a n \ 0" shows "a \ alignUp a n" proof (cases "is_aligned a n") case True then show ?thesis using sz by (subst alignUp_idem, simp_all) next case False have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis shiftr_div_2n' word_shiftr_lt) have"2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans nat_less_le) moreover have "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using nowrap sz apply - apply (erule contrapos_nn) apply (subst alignUp_not_aligned_eq [OF False sz]) apply (subst unat_arith_simps) apply (subst unat_word_ariths) apply (subst unat_word_ariths) apply simp apply (subst mult_mod_left) apply (simp add: unat_div field_simps power_add[symmetric] mod_mod_power min.absorb2) done ultimately have lt: "2 ^ n * (unat a div 2 ^ n + 1) < 2 ^ LENGTH('a)" by simp have "a = a div 2 ^ n * 2 ^ n + a mod 2 ^ n" by (rule word_mod_div_equality [symmetric]) also have "\ < (a div 2 ^ n + 1) * 2 ^ n" using sz lt apply (simp add: field_simps) apply (rule word_add_less_mono1) apply (rule word_mod_less_divisor) apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (simp add: unat_div) done also have "\ = alignUp a n" by (rule alignUp_not_aligned_eq [symmetric]) fact+ finally show ?thesis by (rule order_less_imp_le) qed lemma alignUp_le_greater_al: fixes x :: "'a :: len word" assumes le: "a \ x" and sz: "n < LENGTH('a)" and al: "is_aligned x n" shows "alignUp a n \ x" proof (cases "is_aligned a n") case True then show ?thesis using sz le by (simp add: alignUp_idem) next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) from al obtain k where xk: "x = 2 ^ n * of_nat k" and kv: "k < 2 ^ (LENGTH('a) - n)" by (auto elim!: is_alignedE) then have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst mult.commute) apply simp apply (rule nat_less_power_trans) apply simp apply simp done have au: "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ also have "\ \ of_nat k * 2 ^ n" proof (rule word_mult_le_mono1 [OF inc_le _ kn]) show "a div 2 ^ n < of_nat k" using kv xk le sz anz by (simp add: alignUp_div_helper) show "(0:: 'a word) < 2 ^ n" using sz by (simp add: p2_gt_0 sz) qed finally show ?thesis using xk by (simp add: field_simps) qed lemma alignUp_is_aligned_nz: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and ax: "a \ x" and az: "a \ 0" shows "alignUp (a::'a :: len word) n \ 0" proof (cases "is_aligned a n") case True then have "alignUp a n = a" using sz by (simp add: alignUp_idem) then show ?thesis using az by simp next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) { assume asm: "alignUp a n = 0" have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis shiftr_div_2n' word_shiftr_lt) have leq: "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans order_less_imp_le) from al obtain k where kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" by (auto elim!: is_alignedE) then have "a div 2 ^ n < of_nat k" using ax sz anz by (rule alignUp_div_helper) then have r: "unat a div 2 ^ n < k" using sz by (metis unat_div unat_less_helper unat_power_lower) have "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ then have "\ = 0" using asm by simp then have "2 ^ LENGTH('a) dvd 2 ^ n * (unat a div 2 ^ n + 1)" using sz by (simp add: unat_arith_simps ac_simps) (simp add: unat_word_ariths mod_simps mod_eq_0_iff_dvd) with leq have "2 ^ n * (unat a div 2 ^ n + 1) = 2 ^ LENGTH('a)" by (force elim!: le_SucE) then have "unat a div 2 ^ n = 2 ^ LENGTH('a) div 2 ^ n - 1" by (metis (no_types, hide_lams) Groups.add_ac(2) add.right_neutral add_diff_cancel_left' div_le_dividend div_mult_self4 gr_implies_not0 le_neq_implies_less power_eq_0_iff zero_neq_numeral) then have "unat a div 2 ^ n = 2 ^ (LENGTH('a) - n) - 1" using sz by (simp add: power_sub) then have "2 ^ (LENGTH('a) - n) - 1 < k" using r by simp then have False using kv by simp } then show ?thesis by clarsimp qed lemma alignUp_ar_helper: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and sub: "{x..x + 2 ^ n - 1} \ {a..b}" and anz: "a \ 0" shows "a \ alignUp a n \ alignUp a n + 2 ^ n - 1 \ b" proof from al have xl: "x \ x + 2 ^ n - 1" by (simp add: is_aligned_no_overflow) from xl sub have ax: "a \ x" by (clarsimp elim!: range_subset_lower [where x = x]) show "a \ alignUp a n" proof (rule alignUp_ge) show "alignUp a n \ 0" using al sz ax anz by (rule alignUp_is_aligned_nz) qed fact+ show "alignUp a n + 2 ^ n - 1 \ b" proof (rule order_trans) from xl show tp: "x + 2 ^ n - 1 \ b" using sub by (clarsimp elim!: range_subset_upper [where x = x]) from ax have "alignUp a n \ x" by (rule alignUp_le_greater_al) fact+ then have "alignUp a n + (2 ^ n - 1) \ x + (2 ^ n - 1)" using xl al is_aligned_no_overflow' olen_add_eqv word_plus_mcs_3 by blast then show "alignUp a n + 2 ^ n - 1 \ x + 2 ^ n - 1" by (simp add: field_simps) qed qed lemma alignUp_def2: "alignUp a sz = a + 2 ^ sz - 1 && ~~ (mask sz)" unfolding alignUp_def[unfolded complement_def] by (simp add:mask_eq_decr_exp[symmetric,unfolded shiftl_t2n,simplified]) lemma mask_out_first_mask_some: "\ x && ~~ (mask n) = y; n \ m \ \ x && ~~ (mask m) = y && ~~ (mask m)" by word_eqI_solve lemma gap_between_aligned: "\a < (b :: 'a ::len word); is_aligned a n; is_aligned b n; n < LENGTH('a) \ \ a + (2^n - 1) < b" by (simp add: aligned_add_offset_less) lemma mask_out_add_aligned: assumes al: "is_aligned p n" shows "p + (q && ~~ (mask n)) = (p + q) && ~~ (mask n)" using mask_add_aligned [OF al] by (simp add: mask_out_sub_mask) lemma alignUp_def3: "alignUp a sz = 2^ sz + (a - 1 && ~~ (mask sz))" by (simp add: alignUp_def2 is_aligned_triv field_simps mask_out_add_aligned) lemma alignUp_plus: "is_aligned w us \ alignUp (w + a) us = w + alignUp a us" by (clarsimp simp: alignUp_def2 mask_out_add_aligned field_simps) lemma mask_lower_twice: "n \ m \ (x && ~~ (mask n)) && ~~ (mask m) = x && ~~ (mask m)" by word_eqI_solve lemma mask_lower_twice2: "(a && ~~ (mask n)) && ~~ (mask m) = a && ~~ (mask (max n m))" by word_eqI_solve lemma ucast_and_neg_mask: "ucast (x && ~~ (mask n)) = ucast x && ~~ (mask n)" by word_eqI_solve lemma ucast_and_mask: "ucast (x && mask n) = ucast x && mask n" by word_eqI_solve lemma ucast_mask_drop: "LENGTH('a :: len) \ n \ (ucast (x && mask n) :: 'a word) = ucast x" by word_eqI lemma alignUp_distance: "alignUp (q :: 'a :: len word) sz - q \ mask sz" by (metis (no_types) add.commute add_diff_cancel_left alignUp_def2 diff_add_cancel mask_2pm1 subtract_mask(2) word_and_le1 word_sub_le_iff) lemma is_aligned_diff_neg_mask: "is_aligned p sz \ (p - q && ~~ (mask sz)) = (p - ((alignUp q sz) && ~~ (mask sz)))" apply (clarsimp simp only:word_and_le2 diff_conv_add_uminus) apply (subst mask_out_add_aligned[symmetric]; simp) apply (rule sum_to_zero) apply (subst add.commute) by (simp add: alignUp_distance and_mask_0_iff_le_mask is_aligned_neg_mask_eq mask_out_add_aligned) lemma map_bits_rev_to_bl: "map ((!!) x) [0.. LENGTH('a) \ x = ucast y \ ucast x = y" for x :: "'a::len word" and y :: "'b::len word" by (simp add: is_down ucast_ucast_a) lemma le_ucast_ucast_le: "x \ ucast y \ ucast x \ y" for x :: "'a::len word" and y :: "'b::len word" by (smt le_unat_uoi linorder_not_less order_less_imp_le ucast_nat_def unat_arith_simps(1)) lemma less_ucast_ucast_less: "LENGTH('b) \ LENGTH('a) \ x < ucast y \ ucast x < y" for x :: "'a::len word" and y :: "'b::len word" by (metis ucast_nat_def unat_mono unat_ucast_up_simp word_of_nat_less) lemma ucast_le_ucast: "LENGTH('a) \ LENGTH('b) \ (ucast x \ (ucast y::'b::len word)) = (x \ y)" for x :: "'a::len word" by (simp add: unat_arith_simps(1) unat_ucast_up_simp) lemmas ucast_up_mono_le = ucast_le_ucast[THEN iffD2] lemma ucast_le_ucast_eq: fixes x y :: "'a::len word" assumes x: "x < 2 ^ n" assumes y: "y < 2 ^ n" assumes n: "n = LENGTH('b::len)" shows "(UCAST('a \ 'b) x \ UCAST('a \ 'b) y) = (x \ y)" apply (rule iffI) apply (cases "LENGTH('b) < LENGTH('a)") apply (subst less_mask_eq[OF x, symmetric]) apply (subst less_mask_eq[OF y, symmetric]) apply (unfold n) apply (subst ucast_ucast_mask[symmetric])+ apply (simp add: ucast_le_ucast)+ apply (erule ucast_mono_le[OF _ y[unfolded n]]) done lemma word_le_not_less: "((b::'a::len word) \ a) = (\(a < b))" by fastforce lemma ucast_or_distrib: fixes x :: "'a::len word" fixes y :: "'a::len word" shows "(ucast (x || y) :: ('b::len) word) = ucast x || ucast y" by transfer simp lemma shiftr_less: "(w::'a::len word) < k \ w >> n < k" by (metis div_le_dividend le_less_trans shiftr_div_2n' unat_arith_simps(2)) lemma word_and_notzeroD: "w && w' \ 0 \ w \ 0 \ w' \ 0" by auto lemma word_clz_max: "word_clz w \ size (w::'a::len word)" unfolding word_clz_def by (metis length_takeWhile_le word_size_bl) lemma word_clz_nonzero_max: fixes w :: "'a::len word" assumes nz: "w \ 0" shows "word_clz w < size (w::'a::len word)" proof - { assume a: "word_clz w = size (w::'a::len word)" hence "length (takeWhile Not (to_bl w)) = length (to_bl w)" by (simp add: word_clz_def word_size) hence allj: "\j\set(to_bl w). \ j" by (metis a length_takeWhile_less less_irrefl_nat word_clz_def) hence "to_bl w = replicate (length (to_bl w)) False" by (fastforce intro!: list_of_false) hence "w = 0" by (metis to_bl_0 word_bl.Rep_eqD word_bl_Rep') with nz have False by simp } thus ?thesis using word_clz_max by (fastforce intro: le_neq_trans) qed lemma unat_add_lem': "(unat x + unat y < 2 ^ LENGTH('a)) \ (unat (x + y :: 'a :: len word) = unat x + unat y)" by (subst unat_add_lem[symmetric], assumption) lemma from_bool_eq_if': "((if P then 1 else 0) = from_bool Q) = (P = Q)" - by (simp add: case_bool_If from_bool_def split: if_split) + by (cases Q) (simp_all add: from_bool_def) lemma word_exists_nth: "(w::'a::len word) \ 0 \ \i. w !! i" using word_log2_nth_same by blast lemma shiftr_le_0: "unat (w::'a::len word) < 2 ^ n \ w >> n = (0::'a::len word)" by (rule word_unat.Rep_eqD) (simp add: shiftr_div_2n') lemma of_nat_shiftl: "(of_nat x << n) = (of_nat (x * 2 ^ n) :: ('a::len) word)" proof - have "(of_nat x::'a word) << n = of_nat (2 ^ n) * of_nat x" using shiftl_t2n by (metis word_unat_power) thus ?thesis by simp qed lemma shiftl_1_not_0: "n < LENGTH('a) \ (1::'a::len word) << n \ 0" by (simp add: shiftl_t2n) lemma max_word_not_0 [simp]: "- 1 \ (0 :: 'a::len word)" by simp lemma ucast_zero_is_aligned: "UCAST('a::len \ 'b::len) w = 0 \ n \ LENGTH('b) \ is_aligned w n" by (clarsimp simp: is_aligned_mask word_eq_iff word_size nth_ucast) lemma unat_ucast_eq_unat_and_mask: "unat (UCAST('b::len \ 'a::len) w) = unat (w && mask LENGTH('a))" proof - have "unat (UCAST('b \ 'a) w) = unat (UCAST('a \ 'b) (UCAST('b \ 'a) w))" by (cases "LENGTH('a) < LENGTH('b)"; simp add: is_down ucast_ucast_a unat_ucast_up_simp) thus ?thesis using ucast_ucast_mask by simp qed lemma unat_max_word_pos[simp]: "0 < unat (- 1 :: 'a::len word)" using unat_gt_0 [of \- 1 :: 'a::len word\] by simp (* Miscellaneous conditional injectivity rules. *) lemma mult_pow2_inj: assumes ws: "m + n \ LENGTH('a)" assumes le: "x \ mask m" "y \ mask m" assumes eq: "x * 2^n = y * (2^n::'a::len word)" shows "x = y" proof (cases n) case 0 thus ?thesis using eq by simp next case (Suc n') have m_lt: "m < LENGTH('a)" using Suc ws by simp have xylt: "x < 2^m" "y < 2^m" using le m_lt unfolding mask_2pm1 by auto have lenm: "n \ LENGTH('a) - m" using ws by simp show ?thesis using eq xylt apply (fold shiftl_t2n[where n=n, simplified mult.commute]) apply (simp only: word_bl.Rep_inject[symmetric] bl_shiftl) apply (erule ssubst[OF less_is_drop_replicate])+ apply (clarsimp elim!: drop_eq_mono[OF lenm]) done qed lemma word_of_nat_inj: assumes bounded: "x < 2 ^ LENGTH('a)" "y < 2 ^ LENGTH('a)" assumes of_nats: "of_nat x = (of_nat y :: 'a::len word)" shows "x = y" by (rule contrapos_pp[OF of_nats]; cases "x < y"; cases "y < x") (auto dest: bounded[THEN of_nat_mono_maybe]) (* Sign extension from bit n. *) lemma sign_extend_bitwise_if: "i < size w \ sign_extend e w !! i \ (if i < e then w !! i else w !! e)" by (simp add: sign_extend_def neg_mask_test_bit word_size) lemma sign_extend_bitwise_disj: "i < size w \ sign_extend e w !! i \ i \ e \ w !! i \ e \ i \ w !! e" by (auto simp: sign_extend_bitwise_if) lemma sign_extend_bitwise_cases: "i < size w \ sign_extend e w !! i \ (i \ e \ w !! i) \ (e \ i \ w !! e)" by (auto simp: sign_extend_bitwise_if) lemmas sign_extend_bitwise_if'[word_eqI_simps] = sign_extend_bitwise_if[simplified word_size] lemmas sign_extend_bitwise_disj' = sign_extend_bitwise_disj[simplified word_size] lemmas sign_extend_bitwise_cases' = sign_extend_bitwise_cases[simplified word_size] (* Often, it is easier to reason about an operation which does not overwrite the bit which determines which mask operation to apply. *) lemma sign_extend_def': "sign_extend n w = (if w !! n then w || ~~ (mask (Suc n)) else w && mask (Suc n))" by word_eqI (auto dest: less_antisym) lemma sign_extended_sign_extend: "sign_extended n (sign_extend n w)" by (clarsimp simp: sign_extended_def word_size sign_extend_bitwise_if') lemma sign_extended_iff_sign_extend: "sign_extended n w \ sign_extend n w = w" apply (rule iffI) apply (word_eqI, rename_tac i) apply (case_tac "n < i"; simp add: sign_extended_def word_size) apply (erule subst, rule sign_extended_sign_extend) done lemma sign_extended_weaken: "sign_extended n w \ n \ m \ sign_extended m w" unfolding sign_extended_def by (cases "n < m") auto lemma sign_extend_sign_extend_eq: "sign_extend m (sign_extend n w) = sign_extend (min m n) w" by word_eqI lemma sign_extended_high_bits: "\ sign_extended e p; j < size p; e \ i; i < j \ \ p !! i = p !! j" by (drule (1) sign_extended_weaken; simp add: sign_extended_def) lemma sign_extend_eq: "w && mask (Suc n) = v && mask (Suc n) \ sign_extend n w = sign_extend n v" by word_eqI_solve lemma sign_extended_add: assumes p: "is_aligned p n" assumes f: "f < 2 ^ n" assumes e: "n \ e" assumes "sign_extended e p" shows "sign_extended e (p + f)" proof (cases "e < size p") case True note and_or = is_aligned_add_or[OF p f] have "\ f !! e" using True e less_2p_is_upper_bits_unset[THEN iffD1, OF f] by (fastforce simp: word_size) hence i: "(p + f) !! e = p !! e" by (simp add: and_or) have fm: "f && mask e = f" by (fastforce intro: subst[where P="\f. f && mask e = f", OF less_mask_eq[OF f]] simp: mask_twice e) show ?thesis using assms apply (simp add: sign_extended_iff_sign_extend sign_extend_def i) apply (simp add: and_or word_bw_comms[of p f]) apply (clarsimp simp: word_ao_dist fm word_bw_assocs split: if_splits) done next case False thus ?thesis by (simp add: sign_extended_def word_size) qed lemma sign_extended_neq_mask: "\sign_extended n ptr; m \ n\ \ sign_extended n (ptr && ~~ (mask m))" by (fastforce simp: sign_extended_def word_size neg_mask_test_bit) (* Uints *) lemma uints_mono_iff: "uints l \ uints m \ l \ m" using power_increasing_iff[of "2::int" l m] apply (auto simp: uints_num subset_iff simp del: power_increasing_iff) by (meson less_irrefl not_less zle2p) lemmas uints_monoI = uints_mono_iff[THEN iffD2] lemma Bit_in_uints_Suc: "of_bool c + 2 * w \ uints (Suc m)" if "w \ uints m" using that by (auto simp: uints_num) lemma Bit_in_uintsI: "of_bool c + 2 * w \ uints m" if "w \ uints (m - 1)" "m > 0" using Bit_in_uints_Suc[OF that(1)] that(2) by auto lemma bin_cat_in_uintsI: \bin_cat a n b \ uints m\ if \a \ uints l\ \m \ l + n\ proof - from \m \ l + n\ obtain q where \m = l + n + q\ using le_Suc_ex by blast then have \(2::int) ^ m = 2 ^ n * 2 ^ (l + q)\ by (simp add: ac_simps power_add) moreover have \a mod 2 ^ (l + q) = a\ using \a \ uints l\ by (auto simp add: uints_def take_bit_eq_mod power_add Divides.mod_mult2_eq) ultimately have \concat_bit n b a = take_bit m (concat_bit n b a)\ by (simp add: concat_bit_eq take_bit_eq_mod push_bit_eq_mult Divides.mod_mult2_eq) then show ?thesis by (simp add: uints_def) qed lemma bin_cat_cong: "bin_cat a n b = bin_cat c m d" if "n = m" "a = c" "bintrunc m b = bintrunc m d" using that(3) unfolding that(1,2) by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bin_cat_eqD1: "bin_cat a n b = bin_cat c n d \ a = c" by (metis drop_bit_bin_cat_eq) lemma bin_cat_eqD2: "bin_cat a n b = bin_cat c n d \ bintrunc n b = bintrunc n d" by (metis take_bit_bin_cat_eq) lemma bin_cat_inj: "(bin_cat a n b) = bin_cat c n d \ a = c \ bintrunc n b = bintrunc n d" by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) lemma word_of_int_bin_cat_eq_iff: "(word_of_int (bin_cat (uint a) LENGTH('b) (uint b))::'c::len word) = word_of_int (bin_cat (uint c) LENGTH('b) (uint d)) \ b = d \ a = c" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len word" and b::"'b::len word" by (subst word_uint.Abs_inject) (auto simp: bin_cat_inj intro!: that bin_cat_in_uintsI) lemma word_cat_inj: "(word_cat a b::'c::len word) = word_cat c d \ a = c \ b = d" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len word" and b::"'b::len word" using word_of_int_bin_cat_eq_iff [OF that, of b a d c] by transfer auto lemma p2_eq_1: "2 ^ n = (1::'a::len word) \ n = 0" proof - have "2 ^ n = (1::'a word) \ n = 0" by (metis One_nat_def not_less one_less_numeral_iff p2_eq_0 p2_gt_0 power_0 power_0 power_inject_exp semiring_norm(76) unat_power_lower zero_neq_one) then show ?thesis by auto qed (* usually: x,y = (len_of TYPE ('a)) *) lemma bitmagic_zeroLast_leq_or1Last: "(a::('a::len) word) AND (mask len << x - len) \ a OR mask (y - len)" by (meson le_word_or2 order_trans word_and_le2) lemma zero_base_lsb_imp_set_eq_as_bit_operation: fixes base ::"'a::len word" assumes valid_prefix: "mask (LENGTH('a) - len) AND base = 0" shows "(base = NOT (mask (LENGTH('a) - len)) AND a) \ (a \ {base .. base OR mask (LENGTH('a) - len)})" proof have helper3: "x OR y = x OR y AND NOT x" for x y ::"'a::len word" by (simp add: word_oa_dist2) from assms show "base = NOT (mask (LENGTH('a) - len)) AND a \ a \ {base..base OR mask (LENGTH('a) - len)}" apply(simp add: word_and_le1) apply(metis helper3 le_word_or2 word_bw_comms(1) word_bw_comms(2)) done next assume "a \ {base..base OR mask (LENGTH('a) - len)}" hence a: "base \ a \ a \ base OR mask (LENGTH('a) - len)" by simp show "base = NOT (mask (LENGTH('a) - len)) AND a" proof - have f2: "\x\<^sub>0. base AND NOT (mask x\<^sub>0) \ a AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f3: "\x\<^sub>0. a AND NOT (mask x\<^sub>0) \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f4: "base = base AND NOT (mask (LENGTH('a) - len))" using valid_prefix by (metis mask_eq_0_eq_x word_bw_comms(1)) hence f5: "\x\<^sub>6. (base OR x\<^sub>6) AND NOT (mask (LENGTH('a) - len)) = base OR x\<^sub>6 AND NOT (mask (LENGTH('a) - len))" using word_ao_dist by (metis) have f6: "\x\<^sub>2 x\<^sub>3. a AND NOT (mask x\<^sub>2) \ x\<^sub>3 \ \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>2) \ x\<^sub>3" using f3 dual_order.trans by auto have "base = (base OR mask (LENGTH('a) - len)) AND NOT (mask (LENGTH('a) - len))" using f5 by auto hence "base = a AND NOT (mask (LENGTH('a) - len))" using f2 f4 f6 by (metis eq_iff) thus "base = NOT (mask (LENGTH('a) - len)) AND a" by (metis word_bw_comms(1)) qed qed lemma unat_minus_one_word: "unat (-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by (subst minus_one_word) (subst unat_sub_if', clarsimp) lemma of_nat_eq_signed_scast: "(of_nat x = (y :: ('a::len) signed word)) = (of_nat x = (scast y :: 'a word))" by (metis scast_of_nat scast_scast_id(2)) lemma word_ctz_le: "word_ctz (w :: ('a::len word)) \ LENGTH('a)" apply (clarsimp simp: word_ctz_def) apply (rule nat_le_Suc_less_imp[where y="LENGTH('a) + 1" , simplified]) apply (rule order_le_less_trans[OF List.length_takeWhile_le]) apply simp done lemma word_ctz_less: "w \ 0 \ word_ctz (w :: ('a::len word)) < LENGTH('a)" apply (clarsimp simp: word_ctz_def eq_zero_set_bl) apply (rule order_less_le_trans[OF length_takeWhile_less]) apply fastforce+ done lemma word_ctz_not_minus_1: "1 < LENGTH('a) \ of_nat (word_ctz (w :: 'a :: len word)) \ (- 1 :: 'a::len word)" by (metis (mono_tags) One_nat_def add.right_neutral add_Suc_right le_diff_conv le_less_trans n_less_equal_power_2 not_le suc_le_pow_2 unat_minus_one_word unat_of_nat_len word_ctz_le) lemma word_aligned_add_no_wrap_bounded: "\ w + 2^n \ x; w + 2^n \ 0; is_aligned w n \ \ (w::'a::len word) < x" by (blast dest: is_aligned_no_overflow le_less_trans word_leq_le_minus_one) lemma mask_Suc: "mask (Suc n) = (2 :: 'a::len word) ^ n + mask n" by (simp add: mask_eq_decr_exp) lemma is_aligned_no_overflow_mask: "is_aligned x n \ x \ x + mask n" by (simp add: mask_eq_decr_exp) (erule is_aligned_no_overflow') lemma is_aligned_mask_offset_unat: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off \ mask sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) from offv szv have offv': "unat off < 2 ^ sz" by (simp add: mask_2pm1 unat_less_power) show ?thesis using szv using al is_aligned_no_wrap''' offv' by blast next assume "\ sz < LENGTH('a)" with al have "x = 0" by - word_eqI thus ?thesis by simp qed lemma of_bl_max: "(of_bl xs :: 'a::len word) \ mask (length xs)" apply (induct xs) apply simp apply (simp add: of_bl_Cons mask_Suc) apply (rule conjI; clarsimp) apply (erule word_plus_mono_right) apply (rule is_aligned_no_overflow_mask) apply (rule is_aligned_triv) apply (simp add: word_le_nat_alt) apply (subst unat_add_lem') apply (rule is_aligned_mask_offset_unat) apply (rule is_aligned_triv) apply (simp add: mask_eq_decr_exp) apply simp done lemma mask_over_length: "LENGTH('a) \ n \ mask n = (-1::'a::len word)" by (simp add: mask_eq_decr_exp) lemma is_aligned_over_length: "\ is_aligned p n; LENGTH('a) \ n \ \ (p::'a::len word) = 0" by (simp add: is_aligned_mask mask_over_length) lemma Suc_2p_unat_mask: "n < LENGTH('a) \ Suc (2 ^ n * k + unat (mask n :: 'a::len word)) = 2 ^ n * (k+1)" by (simp add: unat_mask) lemma is_aligned_add_step_le: "\ is_aligned (a::'a::len word) n; is_aligned b n; a < b; b \ a + mask n \ \ False" apply (simp flip: not_le) apply (erule notE) apply (cases "LENGTH('a) \ n") apply (drule (1) is_aligned_over_length)+ apply (drule mask_over_length) apply clarsimp apply (clarsimp simp: word_le_nat_alt not_less not_le) apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (subst (asm) unat_add_lem' [symmetric]) apply (simp add: is_aligned_mask_offset_unat) apply (metis gap_between_aligned linorder_not_less mask_eq_decr_exp unat_arith_simps(2)) done lemma power_2_mult_step_le: "\n' \ n; 2 ^ n' * k' < 2 ^ n * k\ \ 2 ^ n' * (k' + 1) \ 2 ^ n * (k::nat)" apply (cases "n'=n", simp) apply (metis Suc_leI le_refl mult_Suc_right mult_le_mono semiring_normalization_rules(7)) apply (drule (1) le_neq_trans) apply clarsimp apply (subgoal_tac "\m. n = n' + m") prefer 2 apply (simp add: le_Suc_ex) apply (clarsimp simp: power_add) by (metis Suc_leI mult.assoc mult_Suc_right nat_mult_le_cancel_disj) lemma aligned_mask_step: "\ n' \ n; p' \ p + mask n; is_aligned p n; is_aligned p' n' \ \ (p'::'a::len word) + mask n' \ p + mask n" apply (cases "LENGTH('a) \ n") apply (frule (1) is_aligned_over_length) apply (drule mask_over_length) apply clarsimp apply (simp add: not_le) apply (simp add: word_le_nat_alt unat_plus_simple) apply (subst unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask)+ apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (clarsimp simp: dvd_def is_aligned_iff_dvd_nat) apply (rename_tac k k') apply (thin_tac "unat p = x" for p x)+ apply (subst Suc_le_mono[symmetric]) apply (simp only: Suc_2p_unat_mask) apply (drule le_imp_less_Suc, subst (asm) Suc_2p_unat_mask, assumption) apply (erule (1) power_2_mult_step_le) done lemma mask_mono: "sz' \ sz \ mask sz' \ (mask sz :: 'a::len word)" by (simp add: le_mask_iff shiftr_mask_le) lemma aligned_mask_disjoint: "\ is_aligned (a :: 'a :: len word) n; b \ mask n \ \ a && b = 0" by word_eqI_solve lemma word_and_or_mask_aligned: "\ is_aligned a n; b \ mask n \ \ a + b = a || b" by (simp add: aligned_mask_disjoint word_plus_and_or_coroll) lemma word_and_or_mask_aligned2: \is_aligned b n \ a \ mask n \ a + b = a || b\ using word_and_or_mask_aligned [of b n a] by (simp add: ac_simps) lemma is_aligned_ucastI: "is_aligned w n \ is_aligned (ucast w) n" by (clarsimp simp: word_eqI_simps) lemma ucast_le_maskI: "a \ mask n \ UCAST('a::len \ 'b::len) a \ mask n" by (metis and_mask_eq_iff_le_mask ucast_and_mask) lemma ucast_add_mask_aligned: "\ a \ mask n; is_aligned b n \ \ UCAST ('a::len \ 'b::len) (a + b) = ucast a + ucast b" by (metis add.commute is_aligned_ucastI ucast_le_maskI ucast_or_distrib word_and_or_mask_aligned) lemma ucast_shiftl: "LENGTH('b) \ LENGTH ('a) \ UCAST ('a::len \ 'b::len) x << n = ucast (x << n)" by word_eqI_solve lemma ucast_leq_mask: "LENGTH('a) \ n \ ucast (x::'a::len word) \ mask n" by (clarsimp simp: le_mask_high_bits word_size nth_ucast) lemma shiftl_inj: "\ x << n = y << n; x \ mask (LENGTH('a)-n); y \ mask (LENGTH('a)-n) \ \ x = (y :: 'a :: len word)" apply word_eqI apply (rename_tac n') apply (case_tac "LENGTH('a) - n \ n'", simp) by (metis add.commute add.right_neutral diff_add_inverse le_diff_conv linorder_not_less zero_order(1)) lemma distinct_word_add_ucast_shift_inj: "\ p + (UCAST('a::len \ 'b::len) off << n) = p' + (ucast off' << n); is_aligned p n'; is_aligned p' n'; n' = n + LENGTH('a); n' < LENGTH('b) \ \ p' = p \ off' = off" apply (simp add: word_and_or_mask_aligned le_mask_shiftl_le_mask[where n="LENGTH('a)"] ucast_leq_mask) apply (simp add: is_aligned_nth) apply (rule conjI; word_eqI) apply (metis add.commute test_bit_conj_lt diff_add_inverse le_diff_conv nat_less_le) apply (rename_tac i) apply (erule_tac x="i+n" in allE) apply simp done lemma aligned_add_mask_lessD: "\ x + mask n < y; is_aligned x n \ \ x < y" for y::"'a::len word" by (metis is_aligned_no_overflow' mask_2pm1 order_le_less_trans) lemma aligned_add_mask_less_eq: "\ is_aligned x n; is_aligned y n; n < LENGTH('a) \ \ (x + mask n < y) = (x < y)" for y::"'a::len word" using aligned_add_mask_lessD is_aligned_add_step_le word_le_not_less by blast lemma word_upto_Nil: "y < x \ [x .e. y ::'a::len word] = []" by (simp add: upto_enum_red not_le word_less_nat_alt) lemma word_enum_decomp_elem: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y" proof - have "set as \ set [x .e. y] \ a \ set [x .e. y]" using assms by (auto dest: arg_cong[where f=set]) then show ?thesis by auto qed lemma max_word_not_less[simp]: "\ max_word < x" by (simp add: not_less) lemma word_enum_prefix: "[x .e. (y ::'a::len word)] = as @ a # bs \ as = (if x < a then [x .e. a - 1] else [])" apply (induct as arbitrary: x; clarsimp) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (simp add: word_upto_Cons_eq) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (clarsimp simp: word_upto_Cons_eq) apply (frule word_enum_decomp_elem) apply clarsimp apply (rule conjI) prefer 2 apply (subst word_Suc_le[symmetric]; clarsimp) apply (drule meta_spec) apply (drule (1) meta_mp) apply clarsimp apply (rule conjI; clarsimp) apply (subst (2) word_upto_Cons_eq) apply unat_arith apply simp done lemma word_enum_decomp_set: "[x .e. (y ::'a::len word)] = as @ a # bs \ a \ set as" by (metis distinct_append distinct_enum_upto' not_distinct_conv_prefix) lemma word_enum_decomp: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y \ a \ set as \ (\z \ set as. x \ z \ z \ y)" proof - from assms have "set as \ set [x .e. y] \ a \ set [x .e. y]" by (auto dest: arg_cong[where f=set]) with word_enum_decomp_set[OF assms] show ?thesis by auto qed lemma of_nat_unat_le_mask_ucast: "\of_nat (unat t) = w; t \ mask LENGTH('a)\ \ t = UCAST('a::len \ 'b::len) w" by (clarsimp simp: ucast_nat_def ucast_ucast_mask simp flip: and_mask_eq_iff_le_mask) lemma fold_eq_0_to_bool: "(v = 0) = (\ to_bool v)" by (simp add: to_bool_def) lemma less_diff_gt0: "a < b \ (0 :: 'a :: len word) < b - a" by unat_arith lemma unat_plus_gt: "unat ((a :: 'a :: len word) + b) \ unat a + unat b" by (clarsimp simp: unat_plus_if_size) lemma const_less: "\ (a :: 'a :: len word) - 1 < b; a \ b \ \ a < b" by (metis less_1_simp word_le_less_eq) lemma add_mult_aligned_neg_mask: \(x + y * m) && ~~(mask n) = (x && ~~(mask n)) + y * m\ if \m && (2 ^ n - 1) = 0\ by (metis (no_types, hide_lams) add.assoc add.commute add.right_neutral add_uminus_conv_diff mask_eq_decr_exp mask_eqs(2) mask_eqs(6) mult.commute mult_zero_left subtract_mask(1) that) lemma unat_of_nat_minus_1: "\ n < 2 ^ LENGTH('a); n \ 0 \ \ unat ((of_nat n:: 'a :: len word) - 1) = n - 1" by (simp add: unat_eq_of_nat) lemma word_eq_zeroI: "a \ a - 1 \ a = 0" for a :: "'a :: len word" by (simp add: word_must_wrap) lemma word_add_format: "(-1 :: 'a :: len word) + b + c = b + (c - 1)" by simp lemma upto_enum_word_nth: "\ i \ j; k \ unat (j - i) \ \ [i .e. j] ! k = i + of_nat k" apply (clarsimp simp: upto_enum_def nth_append) apply (clarsimp simp: word_le_nat_alt[symmetric]) apply (rule conjI, clarsimp) apply (subst toEnum_of_nat, unat_arith) apply unat_arith apply (clarsimp simp: not_less unat_sub[symmetric]) apply unat_arith done lemma upto_enum_step_nth: "\ a \ c; n \ unat ((c - a) div (b - a)) \ \ [a, b .e. c] ! n = a + of_nat n * (b - a)" by (clarsimp simp: upto_enum_step_def not_less[symmetric] upto_enum_word_nth) lemma upto_enum_inc_1_len: "a < - 1 \ [(0 :: 'a :: len word) .e. 1 + a] = [0 .e. a] @ [1 + a]" apply (simp add: upto_enum_word) apply (subgoal_tac "unat (1+a) = 1 + unat a") apply simp apply (subst unat_plus_simple[THEN iffD1]) apply (metis add.commute no_plus_overflow_neg olen_add_eqv) apply unat_arith done lemma neg_mask_add: "y && mask n = 0 \ x + y && ~~(mask n) = (x && ~~(mask n)) + y" by (clarsimp simp: mask_out_sub_mask mask_eqs(7)[symmetric] mask_twice) lemma shiftr_shiftl_shiftr[simp]: "(x :: 'a :: len word) >> a << a >> a = x >> a" by word_eqI_solve lemma add_right_shift: "\ x && mask n = 0; y && mask n = 0; x \ x + y \ \ (x + y :: ('a :: len) word) >> n = (x >> n) + (y >> n)" apply (simp add: no_olen_add_nat is_aligned_mask[symmetric]) apply (simp add: unat_arith_simps shiftr_div_2n' split del: if_split) apply (subst if_P) apply (erule order_le_less_trans[rotated]) apply (simp add: add_mono) apply (simp add: shiftr_div_2n' is_aligned_iff_dvd_nat) done lemma sub_right_shift: "\ x && mask n = 0; y && mask n = 0; y \ x \ \ (x - y) >> n = (x >> n :: 'a :: len word) - (y >> n)" using add_right_shift[where x="x - y" and y=y and n=n] by (simp add: aligned_sub_aligned is_aligned_mask[symmetric] word_sub_le) lemma and_and_mask_simple: "y && mask n = mask n \ (x && y) && mask n = x && mask n" by (simp add: ac_simps) lemma and_and_mask_simple_not: "y && mask n = 0 \ (x && y) && mask n = 0" by (simp add: ac_simps) lemma word_and_le': "b \ c \ (a :: 'a :: len word) && b \ c" by (metis word_and_le1 order_trans) lemma word_and_less': "b < c \ (a :: 'a :: len word) && b < c" - by (metis word_and_le1 xtr7) + by transfer simp lemma shiftr_w2p: "x < LENGTH('a) \ 2 ^ x = (2 ^ (LENGTH('a) - 1) >> (LENGTH('a) - 1 - x) :: 'a :: len word)" by word_eqI_solve lemma t2p_shiftr: "\ b \ a; a < LENGTH('a) \ \ (2 :: 'a :: len word) ^ a >> b = 2 ^ (a - b)" by word_eqI_solve lemma scast_1[simp]: "scast (1 :: 'a :: len signed word) = (1 :: 'a word)" by simp lemma ucast_ucast_mask_eq: "\ UCAST('a::len \ 'b::len) x = y; x && mask LENGTH('b) = x \ \ x = ucast y" by word_eqI_solve lemma ucast_up_eq: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ LENGTH ('b) \ \ ucast x = (ucast y::'a::len word)" by word_eqI_solve lemma ucast_up_neq: "\ ucast x \ (ucast y::'b::len word); LENGTH('b) \ LENGTH ('a) \ \ ucast x \ (ucast y::'a::len word)" by (fastforce dest: ucast_up_eq) lemma mask_AND_less_0: "\ x && mask n = 0; m \ n \ \ x && mask m = 0" by (metis mask_twice2 word_and_notzeroD) lemma mask_len_id [simp]: "(x :: 'a :: len word) && mask LENGTH('a) = x" using uint_lt2p [of x] by (simp add: mask_eq_iff) lemma scast_ucast_down_same: "LENGTH('b) \ LENGTH('a) \ SCAST('a \ 'b) = UCAST('a::len \ 'b::len)" by (simp add: down_cast_same is_down) lemma word_aligned_0_sum: "\ a + b = 0; is_aligned (a :: 'a :: len word) n; b \ mask n; n < LENGTH('a) \ \ a = 0 \ b = 0" by (simp add: word_plus_and_or_coroll aligned_mask_disjoint word_or_zero) lemma mask_eq1_nochoice: "\ LENGTH('a) > 1; (x :: 'a :: len word) && 1 = x \ \ x = 0 \ x = 1" by (metis word_and_1) lemma pow_mono_leq_imp_lt: "x \ y \ x < 2 ^ y" by (simp add: le_less_trans) lemma unat_of_nat_ctz_mw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len word) = word_ctz w" using word_ctz_le[where w=w, simplified] unat_of_nat_eq[where x="word_ctz w" and 'a="'a"] pow_mono_leq_imp_lt by simp lemma unat_of_nat_ctz_smw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len sword) = word_ctz w" using word_ctz_le[where w=w, simplified] unat_of_nat_eq[where x="word_ctz w" and 'a="'a"] pow_mono_leq_imp_lt by (metis le_unat_uoi le_unat_uoi linorder_neqE_nat nat_less_le scast_of_nat word_unat.Rep_inverse) lemma shiftr_and_eq_shiftl: "(w >> n) && x = y \ w && (x << n) = (y << n)" for y :: "'a:: len word" by (metis (no_types, lifting) and_not_mask bit.conj_ac(1) bit.conj_ac(2) mask_eq_0_eq_x shiftl_mask_is_0 shiftl_over_and_dist) lemma neg_mask_combine: "~~(mask a) && ~~(mask b) = ~~(mask (max a b))" by (auto simp: word_ops_nth_size word_size intro!: word_eqI) lemma neg_mask_twice: "x && ~~(mask n) && ~~(mask m) = x && ~~(mask (max n m))" by (metis neg_mask_combine) lemma multiple_mask_trivia: "n \ m \ (x && ~~(mask n)) + (x && mask n && ~~(mask m)) = x && ~~(mask m)" apply (rule trans[rotated], rule_tac w="mask n" in word_plus_and_or_coroll2) apply (simp add: word_bw_assocs word_bw_comms word_bw_lcs neg_mask_twice max_absorb2) done lemma add_mask_lower_bits': "\ len = LENGTH('a); is_aligned (x :: 'a :: len word) n; \n' \ n. n' < len \ \ p !! n' \ \ x + p && ~~(mask n) = x" using add_mask_lower_bits by auto lemma neg_mask_in_mask_range: "is_aligned ptr bits \ (ptr' && ~~(mask bits) = ptr) = (ptr' \ mask_range ptr bits)" apply (erule is_aligned_get_word_bits) apply (rule iffI) apply (drule sym) apply (simp add: word_and_le2) apply (subst word_plus_and_or_coroll, word_eqI_solve) apply (metis le_word_or2 neg_mask_add_mask and.right_idem) apply clarsimp apply (smt add.right_neutral eq_iff is_aligned_neg_mask_eq mask_out_add_aligned neg_mask_mono_le word_and_not) apply (simp add: power_overflow mask_eq_decr_exp) done lemma aligned_offset_in_range: "\ is_aligned (x :: 'a :: len word) m; y < 2 ^ m; is_aligned p n; n \ m; n < LENGTH('a) \ \ (x + y \ {p .. p + mask n}) = (x \ mask_range p n)" apply (simp only: is_aligned_add_or flip: neg_mask_in_mask_range) by (metis less_mask_eq mask_subsume) lemma mask_range_to_bl': "\ is_aligned (ptr :: 'a :: len word) bits; bits < LENGTH('a) \ \ mask_range ptr bits = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}" apply (rule set_eqI, rule iffI) apply clarsimp apply (subgoal_tac "\y. x = ptr + y \ y < 2 ^ bits") apply clarsimp apply (subst is_aligned_add_conv) apply assumption apply simp apply simp apply (rule_tac x="x - ptr" in exI) apply (simp add: add_diff_eq[symmetric]) apply (simp only: word_less_sub_le[symmetric]) apply (rule word_diff_ls') apply (simp add: field_simps mask_eq_decr_exp) apply assumption apply simp apply (subgoal_tac "\y. y < 2 ^ bits \ to_bl (ptr + y) = to_bl x") apply clarsimp apply (rule conjI) apply (erule(1) is_aligned_no_wrap') apply (simp only: add_diff_eq[symmetric] mask_eq_decr_exp) apply (rule word_plus_mono_right) apply simp apply (erule is_aligned_no_wrap') apply simp apply (rule_tac x="of_bl (drop (LENGTH('a) - bits) (to_bl x))" in exI) apply (rule context_conjI) apply (rule order_less_le_trans [OF of_bl_length]) apply simp apply simp apply (subst is_aligned_add_conv) apply assumption apply simp apply (drule sym) apply (simp add: word_rep_drop) done lemma mask_range_to_bl: "is_aligned (ptr :: 'a :: len word) bits \ mask_range ptr bits = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}" apply (erule is_aligned_get_word_bits) apply (erule(1) mask_range_to_bl') apply (rule set_eqI) apply (simp add: power_overflow mask_eq_decr_exp) done lemma aligned_mask_range_cases: "\ is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n' \ \ mask_range p n \ mask_range p' n' = {} \ mask_range p n \ mask_range p' n' \ mask_range p n \ mask_range p' n'" apply (simp add: mask_range_to_bl) apply (rule Meson.disj_comm, rule disjCI) - apply (erule nonemptyE) - apply simp + apply auto apply (subgoal_tac "(\n''. LENGTH('a) - n = (LENGTH('a) - n') + n'') \ (\n''. LENGTH('a) - n' = (LENGTH('a) - n) + n'')") apply (fastforce simp: take_add) apply arith done lemma aligned_mask_range_offset_subset: assumes al: "is_aligned (ptr :: 'a :: len word) sz" and al': "is_aligned x sz'" and szv: "sz' \ sz" and xsz: "x < 2 ^ sz" shows "mask_range (ptr+x) sz' \ mask_range ptr sz" using al proof (rule is_aligned_get_word_bits) assume p0: "ptr = 0" and szv': "LENGTH ('a) \ sz" then have "(2 ::'a word) ^ sz = 0" by simp show ?thesis using p0 by (simp add: \2 ^ sz = 0\ mask_eq_decr_exp) next assume szv': "sz < LENGTH('a)" hence blah: "2 ^ (sz - sz') < (2 :: nat) ^ LENGTH('a)" using szv by auto show ?thesis using szv szv' apply (intro range_subsetI) apply (rule is_aligned_no_wrap' [OF al xsz]) apply (simp only: flip: add_diff_eq add_mask_fold) apply (subst add.assoc, rule word_plus_mono_right) using al' is_aligned_add_less_t2n xsz apply fastforce apply (simp add: field_simps szv al is_aligned_no_overflow) done qed lemma aligned_mask_diff: "\ is_aligned (dest :: 'a :: len word) bits; is_aligned (ptr :: 'a :: len word) sz; bits \ sz; sz < LENGTH('a); dest < ptr \ \ mask bits + dest < ptr" apply (frule_tac p' = ptr in aligned_mask_range_cases, assumption) apply (elim disjE) apply (drule_tac is_aligned_no_overflow_mask, simp)+ apply (simp add: algebra_split_simps word_le_not_less) apply (drule is_aligned_no_overflow_mask; fastforce) by (simp add: aligned_add_mask_less_eq is_aligned_weaken algebra_split_simps) lemma aligned_mask_ranges_disjoint: "\ is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n'; p && ~~(mask n') \ p'; p' && ~~(mask n) \ p \ \ mask_range p n \ mask_range p' n' = {}" using aligned_mask_range_cases by (auto simp: neg_mask_in_mask_range) lemma aligned_mask_ranges_disjoint2: "\ is_aligned p n; is_aligned ptr bits; n \ m; n < size p; m \ bits; (\y < 2 ^ (n - m). p + (y << m) \ mask_range ptr bits) \ \ mask_range p n \ mask_range ptr bits = {}" apply safe apply (simp only: flip: neg_mask_in_mask_range) apply (drule_tac x="x && mask n >> m" in spec) apply (clarsimp simp: shiftr_less_t2n and_mask_less_size wsst_TYs multiple_mask_trivia word_bw_assocs neg_mask_twice max_absorb2 shiftr_shiftl1) done lemma leq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + high_bits) \ (x >> low_bits) \ mask high_bits" by (simp add: le_mask_iff shiftr_shiftr) lemma ucast_ucast_eq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + LENGTH('b)) \ ucast((ucast (x >> low_bits)) :: 'b :: len word) = x >> low_bits" by (meson and_mask_eq_iff_le_mask eq_ucast_ucast_eq not_le_imp_less shiftr_less_t2n' ucast_ucast_len) lemma const_le_unat: "\ b < 2 ^ LENGTH('a); of_nat b \ a \ \ b \ unat (a :: 'a :: len word)" apply (simp add: word_le_def) apply (simp only: uint_nat zle_int) apply transfer apply (simp add: take_bit_nat_eq_self) done lemma upt_enum_offset_trivial: "\ x < 2 ^ LENGTH('a) - 1 ; n \ unat x \ \ ([(0 :: 'a :: len word) .e. x] ! n) = of_nat n" apply (induct x arbitrary: n) apply simp by (simp add: upto_enum_word_nth) lemma word_le_mask_out_plus_2sz: "x \ (x && ~~(mask sz)) + 2 ^ sz - 1" by (metis add_diff_eq word_neg_and_le) lemma ucast_add: "ucast (a + (b :: 'a :: len word)) = ucast a + (ucast b :: ('a signed word))" by transfer (simp add: take_bit_add) lemma ucast_minus: "ucast (a - (b :: 'a :: len word)) = ucast a - (ucast b :: ('a signed word))" apply (insert ucast_add[where a=a and b="-b"]) apply (metis (no_types, hide_lams) add_diff_eq diff_add_cancel ucast_add) done lemma scast_ucast_add_one [simp]: "scast (ucast (x :: 'a::len word) + (1 :: 'a signed word)) = x + 1" apply (subst ucast_1[symmetric]) apply (subst ucast_add[symmetric]) apply clarsimp done lemma word_and_le_plus_one: "a > 0 \ (x :: 'a :: len word) && (a - 1) < a" by (simp add: gt0_iff_gem1 word_and_less') lemma unat_of_ucast_then_shift_eq_unat_of_shift[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) >> n) = unat (x >> n)" by (simp add: shiftr_div_2n' unat_ucast_up_simp) lemma unat_of_ucast_then_mask_eq_unat_of_mask[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) && mask m) = unat (x && mask m)" by (metis ucast_and_mask unat_ucast_up_simp) lemma small_powers_of_2: "x \ 3 \ x < 2 ^ (x - 1)" by (induct x; simp add: suc_le_pow_2) lemma word_clz_sint_upper[simp]: "LENGTH('a) \ 3 \ sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a sword) \ int (LENGTH('a))" using small_powers_of_2 by (smt One_nat_def diff_less le_less_trans len_gt_0 len_signed lessI n_less_equal_power_2 not_msb_from_less of_nat_mono sint_eq_uint uint_nat unat_of_nat_eq unat_power_lower word_clz_max word_of_nat_less wsst_TYs(3)) lemma word_clz_sint_lower[simp]: "LENGTH('a) \ 3 \ - sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a signed word) \ int (LENGTH('a))" apply (subst sint_eq_uint) using small_powers_of_2 uint_nat apply (simp add: order_le_less_trans[OF word_clz_max] not_msb_from_less word_of_nat_less word_size) by (simp add: uint_nat) lemma shiftr_less_t2n3: "\ (2 :: 'a word) ^ (n + m) = 0; m < LENGTH('a) \ \ (x :: 'a :: len word) >> n < 2 ^ m" by (fastforce intro: shiftr_less_t2n' simp: mask_eq_decr_exp power_overflow) lemma unat_shiftr_le_bound: "\ 2 ^ (LENGTH('a :: len) - n) - 1 \ bnd; 0 < n \ \ unat ((x :: 'a word) >> n) \ bnd" using less_not_refl3 le_step_down_nat le_trans less_or_eq_imp_le word_shiftr_lt by (metis (no_types, lifting)) lemma shiftr_eqD: "\ x >> n = y >> n; is_aligned x n; is_aligned y n \ \ x = y" by (metis is_aligned_shiftr_shiftl) lemma word_shiftr_shiftl_shiftr_eq_shiftr: "a \ b \ (x :: 'a :: len word) >> a << b >> b = x >> a" by (simp add: mask_shift multi_shift_simps(5) shiftr_shiftr) lemma of_int_uint_ucast: "of_int (uint (x :: 'a::len word)) = (ucast x :: 'b::len word)" by (fact Word.of_int_uint) lemma mod_mask_drop: "\ m = 2 ^ n; 0 < m; mask n && msk = mask n \ \ (x mod m) && msk = x mod m" by (simp add: word_mod_2p_is_mask word_bw_assocs) lemma mask_eq_ucast_eq: "\ x && mask LENGTH('a) = (x :: ('c :: len word)); LENGTH('a) \ LENGTH('b)\ \ ucast (ucast x :: ('a :: len word)) = (ucast x :: ('b :: len word))" by (metis ucast_and_mask ucast_id ucast_ucast_mask ucast_up_eq) lemma of_nat_less_t2n: "of_nat i < (2 :: ('a :: len) word) ^ n \ n < LENGTH('a) \ unat (of_nat i :: 'a word) < 2 ^ n" by (metis order_less_trans p2_gt_0 unat_less_power word_neq_0_conv) lemma two_power_increasing_less_1: "\ n \ m; m \ LENGTH('a) \ \ (2 :: 'a :: len word) ^ n - 1 \ 2 ^ m - 1" by (metis diff_diff_cancel le_m1_iff_lt less_imp_diff_less p2_gt_0 two_power_increasing word_1_le_power word_le_minus_mono_left word_less_sub_1) lemma word_sub_mono4: "\ y + x \ z + x; y \ y + x; z \ z + x \ \ y \ z" for y :: "'a :: len word" by (simp add: word_add_le_iff2) lemma eq_or_less_helperD: "\ n = unat (2 ^ m - 1 :: 'a :: len word) \ n < unat (2 ^ m - 1 :: 'a word); m < LENGTH('a) \ \ n < 2 ^ m" by (meson le_less_trans nat_less_le unat_less_power word_power_less_1) lemma mask_sub: "n \ m \ mask m - mask n = mask m && ~~(mask n)" by (metis (full_types) and_mask_eq_iff_shiftr_0 mask_out_sub_mask shiftr_mask_le word_bw_comms(1)) lemma neg_mask_diff_bound: "sz'\ sz \ (ptr && ~~(mask sz')) - (ptr && ~~(mask sz)) \ 2 ^ sz - 2 ^ sz'" (is "_ \ ?lhs \ ?rhs") proof - assume lt: "sz' \ sz" hence "?lhs = ptr && (mask sz && ~~(mask sz'))" by (metis add_diff_cancel_left' multiple_mask_trivia) also have "\ \ ?rhs" using lt by (metis (mono_tags) add_diff_eq diff_eq_eq eq_iff mask_2pm1 mask_sub word_and_le') finally show ?thesis by simp qed lemma mask_range_subsetD: "\ p' \ mask_range p n; x' \ mask_range p' n'; n' \ n; is_aligned p n; is_aligned p' n' \ \ x' \ mask_range p n" using aligned_mask_step by fastforce lemma add_mult_in_mask_range: "\ is_aligned (base :: 'a :: len word) n; n < LENGTH('a); bits \ n; x < 2 ^ (n - bits) \ \ base + x * 2^bits \ mask_range base n" by (simp add: is_aligned_no_wrap' mask_2pm1 nasty_split_lt word_less_power_trans2 word_plus_mono_right) lemma of_bl_length2: "length xs + c < LENGTH('a) \ of_bl xs * 2^c < (2::'a::len word) ^ (length xs + c)" by (simp add: of_bl_length word_less_power_trans2) lemma mask_out_eq_0: "\ idx < 2 ^ sz; sz < LENGTH('a) \ \ (of_nat idx :: 'a :: len word) && ~~(mask sz) = 0" by (simp add: Word_Lemmas.of_nat_power less_mask_eq mask_eq_0_eq_x) lemma is_aligned_neg_mask_eq': "is_aligned ptr sz = (ptr && ~~(mask sz) = ptr)" using is_aligned_mask mask_eq_0_eq_x by blast lemma neg_mask_mask_unat: "sz < LENGTH('a) \ unat ((ptr :: 'a :: len word) && ~~(mask sz)) + unat (ptr && mask sz) = unat ptr" by (metis AND_NOT_mask_plus_AND_mask_eq unat_plus_simple word_and_le2) lemma unat_pow_le_intro: "LENGTH('a) \ n \ unat (x :: 'a :: len word) < 2 ^ n" by (metis lt2p_lem not_le of_nat_le_iff of_nat_numeral semiring_1_class.of_nat_power uint_nat) lemma unat_shiftl_less_t2n: "\ unat (x :: 'a :: len word) < 2 ^ (m - n); m < LENGTH('a) \ \ unat (x << n) < 2 ^ m" by (metis (no_types) Word_Lemmas.of_nat_power diff_le_self le_less_trans shiftl_less_t2n unat_less_power word_unat.Rep_inverse) lemma unat_is_aligned_add: "\ is_aligned p n; unat d < 2 ^ n \ \ unat (p + d && mask n) = unat d \ unat (p + d && ~~(mask n)) = unat p" by (metis add.right_neutral and_mask_eq_iff_le_mask and_not_mask le_mask_iff mask_add_aligned mask_out_add_aligned mult_zero_right shiftl_t2n shiftr_le_0) lemma unat_shiftr_shiftl_mask_zero: "\ c + a \ LENGTH('a) + b ; c < LENGTH('a) \ \ unat (((q :: 'a :: len word) >> a << b) && ~~(mask c)) = 0" by (fastforce intro: unat_is_aligned_add[where p=0 and n=c, simplified, THEN conjunct2] unat_shiftl_less_t2n unat_shiftr_less_t2n unat_pow_le_intro) lemmas of_nat_ucast = ucast_of_nat[symmetric] lemma shift_then_mask_eq_shift_low_bits: "x \ mask (low_bits + high_bits) \ (x >> low_bits) && mask high_bits = x >> low_bits" by (simp add: leq_mask_shift le_mask_imp_and_mask) lemma leq_low_bits_iff_zero: "\ x \ mask (low bits + high bits); x >> low_bits = 0 \ \ (x && mask low_bits = 0) = (x = 0)" using and_mask_eq_iff_shiftr_0 by force lemma unat_less_iff: "\ unat (a :: 'a :: len word) = b; c < 2 ^ LENGTH('a) \ \ (a < of_nat c) = (b < c)" using unat_ucast_less_no_overflow_simp by blast lemma is_aligned_no_overflow3: "\ is_aligned (a :: 'a :: len word) n; n < LENGTH('a); b < 2 ^ n; c \ 2 ^ n; b < c \ \ a + b \ a + (c - 1)" by (meson is_aligned_no_wrap' le_m1_iff_lt not_le word_less_sub_1 word_plus_mono_right) lemma mask_add_aligned_right: "is_aligned p n \ (q + p) && mask n = q && mask n" by (simp add: mask_add_aligned add.commute) lemma leq_high_bits_shiftr_low_bits_leq_bits_mask: "x \ mask high_bits \ (x :: 'a :: len word) << low_bits \ mask (low_bits + high_bits)" by (metis le_mask_shiftl_le_mask) lemma from_to_bool_last_bit: "from_bool (to_bool (x && 1)) = x && 1" by (metis from_bool_to_bool_iff word_and_1) lemma word_two_power_neg_ineq: "2 ^ m \ (0 :: 'a word) \ 2 ^ n \ - (2 ^ m :: 'a :: len word)" apply (cases "n < LENGTH('a)"; simp add: power_overflow) apply (cases "m < LENGTH('a)"; simp add: power_overflow) apply (simp add: word_le_nat_alt unat_minus word_size) apply (cases "LENGTH('a)"; simp) apply (simp add: less_Suc_eq_le) apply (drule power_increasing[where a=2 and n=n] power_increasing[where a=2 and n=m], simp)+ apply (drule(1) add_le_mono) apply simp done lemma unat_shiftl_absorb: "\ x \ 2 ^ p; p + k < LENGTH('a) \ \ unat (x :: 'a :: len word) * 2 ^ k = unat (x * 2 ^ k)" by (smt add_diff_cancel_right' add_lessD1 le_add2 le_less_trans mult.commute nat_le_power_trans unat_lt2p unat_mult_lem unat_power_lower word_le_nat_alt) lemma word_plus_mono_right_split: "\ unat ((x :: 'a :: len word) && mask sz) + unat z < 2 ^ sz; sz < LENGTH('a) \ \ x \ x + z" apply (subgoal_tac "(x && ~~(mask sz)) + (x && mask sz) \ (x && ~~(mask sz)) + ((x && mask sz) + z)") apply (simp add:word_plus_and_or_coroll2 field_simps) apply (rule word_plus_mono_right) apply (simp add: less_le_trans no_olen_add_nat) using Word_Lemmas.of_nat_power is_aligned_no_wrap' by force lemma mul_not_mask_eq_neg_shiftl: "~~(mask n) = -1 << n" by (simp add: NOT_mask shiftl_t2n) lemma shiftr_mul_not_mask_eq_and_not_mask: "(x >> n) * ~~(mask n) = - (x && ~~(mask n))" by (metis NOT_mask and_not_mask mult_minus_left semiring_normalization_rules(7) shiftl_t2n) lemma mask_eq_n1_shiftr: "n \ LENGTH('a) \ (mask n :: 'a :: len word) = -1 >> (LENGTH('a) - n)" by (metis diff_diff_cancel eq_refl mask_full shiftr_mask2) lemma is_aligned_mask_out_add_eq: "is_aligned p n \ (p + x) && ~~(mask n) = p + (x && ~~(mask n))" by (simp add: mask_out_sub_mask mask_add_aligned) lemmas is_aligned_mask_out_add_eq_sub = is_aligned_mask_out_add_eq[where x="a - b" for a b, simplified field_simps] lemma aligned_bump_down: "is_aligned x n \ (x - 1) && ~~(mask n) = x - 2 ^ n" by (drule is_aligned_mask_out_add_eq[where x="-1"]) (simp add: NOT_mask) lemma unat_2tp_if: "unat (2 ^ n :: ('a :: len) word) = (if n < LENGTH ('a) then 2 ^ n else 0)" by (split if_split, simp_all add: power_overflow) lemma mask_of_mask: "mask (n::nat) && mask (m::nat) = mask (min m n)" by word_eqI_solve lemma unat_signed_ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ unat (ucast (x :: 'a :: len word) :: 'b :: len signed word) = unat x" by (simp add: unat_ucast_up_simp) lemma toEnum_of_ucast: "LENGTH('b) \ LENGTH('a) \ (toEnum (unat (b::'b :: len word))::'a :: len word) = of_nat (unat b)" by (simp add: unat_pow_le_intro) lemmas unat_ucast_mask = unat_ucast_eq_unat_and_mask[where w=a for a] lemma t2n_mask_eq_if: "2 ^ n && mask m = (if n < m then 2 ^ n else 0)" by (rule word_eqI, auto simp add: word_size nth_w2p split: if_split) lemma unat_ucast_le: "unat (ucast (x :: 'a :: len word) :: 'b :: len word) \ unat x" by (simp add: ucast_nat_def word_unat_less_le) lemma ucast_le_up_down_iff: "\ LENGTH('a) \ LENGTH('b); (x :: 'b :: len word) \ ucast (max_word :: 'a :: len word) \ \ (ucast x \ (y :: 'a word)) = (x \ ucast y)" using le_max_word_ucast_id ucast_le_ucast by metis lemma ucast_ucast_mask_shift: "a \ LENGTH('a) + b \ ucast (ucast (p && mask a >> b) :: 'a :: len word) = p && mask a >> b" by (metis add.commute le_mask_iff shiftr_mask_le ucast_ucast_eq_mask_shift word_and_le') lemma unat_ucast_mask_shift: "a \ LENGTH('a) + b \ unat (ucast (p && mask a >> b) :: 'a :: len word) = unat (p && mask a >> b)" by (metis linear ucast_ucast_mask_shift unat_ucast_up_simp) lemma mask_overlap_zero: "a \ b \ (p && mask a) && ~~(mask b) = 0" by (metis NOT_mask_AND_mask mask_lower_twice2 max_def) lemma mask_shifl_overlap_zero: "a + c \ b \ (p && mask a << c) && ~~(mask b) = 0" by (metis and_mask_0_iff_le_mask mask_mono mask_shiftl_decompose order_trans shiftl_over_and_dist word_and_le' word_and_le2) lemma mask_overlap_zero': "a \ b \ (p && ~~(mask a)) && mask b = 0" using mask_AND_NOT_mask mask_AND_less_0 by blast lemma mask_rshift_mult_eq_rshift_lshift: "((a :: 'a :: len word) >> b) * (1 << c) = (a >> b << c)" by (simp add: shiftl_t2n) lemma shift_alignment: "a \ b \ is_aligned (p >> a << a) b" using is_aligned_shift is_aligned_weaken by blast lemma mask_split_sum_twice: "a \ b \ (p && ~~(mask a)) + ((p && mask a) && ~~(mask b)) + (p && mask b) = p" by (simp add: add.commute multiple_mask_trivia word_bw_comms(1) word_bw_lcs(1) word_plus_and_or_coroll2) lemma mask_shift_eq_mask_mask: "(p && mask a >> b << b) = (p && mask a) && ~~(mask b)" by (simp add: and_not_mask) lemma mask_shift_sum: "\ a \ b; unat n = unat (p && mask b) \ \ (p && ~~(mask a)) + (p && mask a >> b) * (1 << b) + n = (p :: 'a :: len word)" by (metis and_not_mask mask_rshift_mult_eq_rshift_lshift mask_split_sum_twice word_unat.Rep_eqD) lemma is_up_compose: "\ is_up uc; is_up uc' \ \ is_up (uc' \ uc)" unfolding is_up_def by (simp add: Word.target_size Word.source_size) lemma of_int_sint_scast: "of_int (sint (x :: 'a :: len word)) = (scast x :: 'b :: len word)" by (fact Word.of_int_sint) lemma scast_of_nat_to_signed [simp]: "scast (of_nat x :: 'a :: len word) = (of_nat x :: 'a signed word)" by (metis cast_simps(23) scast_scast_id(2)) lemma scast_of_nat_signed_to_unsigned_add: "scast (of_nat x + of_nat y :: 'a :: len signed word) = (of_nat x + of_nat y :: 'a :: len word)" by (metis of_nat_add scast_of_nat) lemma scast_of_nat_unsigned_to_signed_add: "(scast (of_nat x + of_nat y :: 'a :: len word)) = (of_nat x + of_nat y :: 'a :: len signed word)" by (metis Abs_fnat_hom_add scast_of_nat_to_signed) lemma and_mask_cases: fixes x :: "'a :: len word" assumes len: "n < LENGTH('a)" shows "x && mask n \ of_nat ` set [0 ..< 2 ^ n]" apply (simp flip: take_bit_eq_mask) apply (rule image_eqI [of _ _ \unat (take_bit n x)\]) using len apply simp_all apply transfer apply simp done lemma sint_of_nat_ge_zero: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) \ 0" by (simp add: bit_iff_odd) lemma sint_eq_uint_2pl: "\ (a :: 'a :: len word) < 2 ^ (LENGTH('a) - 1) \ \ sint a = uint a" by (simp add: not_msb_from_less sint_eq_uint word_2p_lem word_size) lemma sint_of_nat_le: "\ b < 2 ^ (LENGTH('a) - 1); a \ b \ \ sint (of_nat a :: 'a :: len word) \ sint (of_nat b :: 'a :: len word)" by (smt Word_Lemmas.of_nat_power diff_less le_less_trans len_gt_0 len_of_numeral_defs(2) nat_power_minus_less of_nat_le_iff sint_eq_uint_2pl uint_nat unat_of_nat_len) lemma int_eq_sint: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) = int x" by (smt Word_Lemmas.of_nat_power diff_less le_less_trans len_gt_0 len_of_numeral_defs(2) nat_less_le sint_eq_uint_2pl uint_nat unat_lt2p unat_of_nat_len unat_power_lower) lemma sint_ctz: "LENGTH('a) > 2 \ 0 \ sint (of_nat (word_ctz (x :: 'a :: len word)) :: 'a signed word) \ sint (of_nat (word_ctz x) :: 'a signed word) \ int (LENGTH('a))" apply (subgoal_tac "LENGTH('a) < 2 ^ (LENGTH('a) - 1)") apply (rule conjI) apply (metis len_signed order_le_less_trans sint_of_nat_ge_zero word_ctz_le) apply (metis int_eq_sint len_signed sint_of_nat_le word_ctz_le) by (rule small_powers_of_2, simp) lemma pow_sub_less: "\ a + b \ LENGTH('a); unat (x :: 'a :: len word) = 2 ^ a \ \ unat (x * 2 ^ b - 1) < 2 ^ (a + b)" by (metis (mono_tags) eq_or_less_helperD not_less of_nat_numeral power_add semiring_1_class.of_nat_power unat_pow_le_intro word_unat.Rep_inverse) lemma sle_le_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a \ b \ \ a <=s b" by (simp add: not_msb_from_less word_sle_msb_le) lemma sless_less_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a < b \ \ a > n = w && mask (size w - n)" by (cases "n \ size w"; clarsimp simp: word_and_le2 and_mask shiftl_zero_size) lemma unat_of_nat_word_log2: "LENGTH('a) < 2 ^ LENGTH('b) \ unat (of_nat (word_log2 (n :: 'a :: len word)) :: 'b :: len word) = word_log2 n" by (metis less_trans unat_of_nat_eq word_log2_max word_size) lemma aligned_sub_aligned_simple: "\ is_aligned a n; is_aligned b n \ \ is_aligned (a - b) n" by (simp add: aligned_sub_aligned) lemma minus_one_shift: "- (1 << n) = (-1 << n :: 'a::len word)" by (simp add: mask_eq_decr_exp NOT_eq flip: mul_not_mask_eq_neg_shiftl) lemma ucast_eq_mask: "(UCAST('a::len \ 'b::len) x = UCAST('a \ 'b) y) = (x && mask LENGTH('b) = y && mask LENGTH('b))" by (rule iffI; word_eqI_solve) context fixes w :: "'a::len word" begin private lemma sbintrunc_uint_ucast: assumes "Suc n = LENGTH('b::len)" shows "sbintrunc n (uint (ucast w :: 'b word)) = sbintrunc n (uint w)" by (metis assms sbintrunc_bintrunc ucast_eq word_ubin.eq_norm) private lemma test_bit_sbintrunc: assumes "i < LENGTH('a)" shows "(word_of_int (sbintrunc n (uint w)) :: 'a word) !! i = (if n < i then w !! n else w !! i)" using assms by (simp add: nth_sbintr) (simp add: test_bit_bin) private lemma test_bit_sbintrunc_ucast: assumes len_a: "i < LENGTH('a)" shows "(word_of_int (sbintrunc (LENGTH('b) - 1) (uint (ucast w :: 'b word))) :: 'a word) !! i = (if LENGTH('b::len) \ i then w !! (LENGTH('b) - 1) else w !! i)" apply (subst sbintrunc_uint_ucast) apply simp apply (subst test_bit_sbintrunc) apply (rule len_a) apply (rule if_cong[OF _ refl refl]) using leD less_linear by fastforce lemma scast_ucast_high_bits: \scast (ucast w :: 'b::len word) = w \ (\ i \ {LENGTH('b) ..< size w}. w !! i = w !! (LENGTH('b) - 1))\ proof (cases \LENGTH('a) \ LENGTH('b)\) case True moreover define m where \m = LENGTH('b) - LENGTH('a)\ ultimately have \LENGTH('b) = m + LENGTH('a)\ by simp then show ?thesis apply (simp_all add: signed_ucast_eq word_size) apply (rule bit_word_eqI) apply (simp add: bit_signed_take_bit_iff) done next case False define q where \q = LENGTH('b) - 1\ then have \LENGTH('b) = Suc q\ by simp moreover define m where \m = Suc LENGTH('a) - LENGTH('b)\ with False \LENGTH('b) = Suc q\ have \LENGTH('a) = m + q\ by (simp add: not_le) ultimately show ?thesis apply (simp_all add: signed_ucast_eq word_size) apply (transfer fixing: m q) apply (simp add: signed_take_bit_take_bit) apply rule apply (subst bit_eq_iff) apply (simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def) apply (auto simp add: Suc_le_eq) using less_imp_le_nat apply blast using less_imp_le_nat apply blast done qed lemma scast_ucast_mask_compare: "scast (ucast w :: 'b::len word) = w \ (w \ mask (LENGTH('b) - 1) \ ~~(mask (LENGTH('b) - 1)) \ w)" apply (clarsimp simp: le_mask_high_bits neg_mask_le_high_bits scast_ucast_high_bits word_size) apply (rule iffI; clarsimp) apply (rename_tac i j; case_tac "i = LENGTH('b) - 1"; case_tac "j = LENGTH('b) - 1") by auto lemma ucast_less_shiftl_helper': "\ LENGTH('b) + (a::nat) < LENGTH('a); 2 ^ (LENGTH('b) + a) \ n\ \ (ucast (x :: 'b::len word) << a) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done end lemma ucast_ucast_mask2: "is_down (UCAST ('a \ 'b)) \ UCAST ('b::len \ 'c::len) (UCAST ('a::len \ 'b::len) x) = UCAST ('a \ 'c) (x && mask LENGTH('b))" by word_eqI_solve lemma ucast_NOT: "ucast (~~x) = ~~(ucast x) && mask (LENGTH('a))" for x::"'a::len word" by word_eqI lemma ucast_NOT_down: "is_down UCAST('a::len \ 'b::len) \ UCAST('a \ 'b) (~~x) = ~~(UCAST('a \ 'b) x)" by word_eqI lemma of_bl_mult_and_not_mask_eq: "\is_aligned (a :: 'a::len word) n; length b + m \ n\ \ a + of_bl b * (2^m) && ~~(mask n) = a" by (smt add.left_neutral add_diff_cancel_right' add_mask_lower_bits and_mask_plus is_aligned_mask is_aligned_weaken le_less_trans of_bl_length2 subtract_mask(1)) lemma bin_to_bl_of_bl_eq: "\is_aligned (a::'a::len word) n; length b + c \ n; length b + c < LENGTH('a)\ \ bin_to_bl (length b) (uint ((a + of_bl b * 2^c) >> c)) = b" apply (subst word_plus_and_or_coroll) apply (erule is_aligned_get_word_bits) apply (rule is_aligned_AND_less_0) apply (simp add: is_aligned_mask) apply (rule order_less_le_trans) apply (rule of_bl_length2) apply simp apply (simp add: two_power_increasing) apply simp apply (rule nth_equalityI) apply (simp only: len_bin_to_bl) apply (clarsimp simp only: len_bin_to_bl nth_bin_to_bl word_test_bit_def[symmetric]) apply (simp add: nth_shiftr nth_shiftl shiftl_t2n[where n=c, simplified mult.commute, simplified, symmetric]) apply (simp add: is_aligned_nth[THEN iffD1, rule_format] test_bit_of_bl rev_nth) apply arith done text \right-padding a word to a certain length\ definition "bl_pad_to bl sz \ bl @ (replicate (sz - length bl) False)" lemma bl_pad_to_length: assumes lbl: "length bl \ sz" shows "length (bl_pad_to bl sz) = sz" using lbl by (simp add: bl_pad_to_def) lemma bl_pad_to_prefix: "prefix bl (bl_pad_to bl sz)" by (simp add: bl_pad_to_def) -lemma same_length_is_parallel: - assumes len: "\y \ set as. length y = x" - shows "\x \ set as. \y \ set as - {x}. x \ y" -proof (rule, rule) - fix x y - assume xi: "x \ set as" and yi: "y \ set as - {x}" - from len obtain q where len': "\y \ set as. length y = q" .. - - show "x \ y" - proof (rule not_equal_is_parallel) - from xi yi show "x \ y" by auto - from xi yi len' show "length x = length y" by (auto dest: bspec) - qed -qed - -lemma sublist_equal_part: - "prefix xs ys \ take (length xs) ys = xs" - by (clarsimp simp: prefix_def) - -lemma prefix_length_less: - "strict_prefix xs ys \ length xs < length ys" - apply (clarsimp simp: strict_prefix_def) - apply (frule prefix_length_le) - apply (rule ccontr, simp) - apply (clarsimp simp: prefix_def) - done - -lemmas take_less = take_strict_prefix - -lemma not_prefix_longer: - "\ length xs > length ys \ \ \ prefix xs ys" - by (clarsimp dest!: prefix_length_le) - -lemma map_prefixI: - "prefix xs ys \ prefix (map f xs) (map f ys)" - by (clarsimp simp: prefix_def) - -lemma list_all2_induct_suffixeq [consumes 1, case_names Nil Cons]: - assumes lall: "list_all2 Q as bs" - and nilr: "P [] []" - and consr: "\x xs y ys. - \list_all2 Q xs ys; Q x y; P xs ys; suffix (x # xs) as; suffix (y # ys) bs\ - \ P (x # xs) (y # ys)" - shows "P as bs" -proof - - define as' where "as' == as" - define bs' where "bs' == bs" - - have "suffix as as' \ suffix bs bs'" unfolding as'_def bs'_def by simp - then show ?thesis using lall - proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) - case 1 show ?case by fact - next - case (2 x xs y ys) - - show ?case - proof (rule consr) - from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all - then show "P xs ys" using "2.hyps" "2.prems" by (auto dest: suffix_ConsD) - from "2.prems" show "suffix (x # xs) as" and "suffix (y # ys) bs" - by (auto simp: as'_def bs'_def) - qed - qed -qed - -lemma take_prefix: - "(take (length xs) ys = xs) = prefix xs ys" -proof (induct xs arbitrary: ys) - case Nil then show ?case by simp -next - case Cons then show ?case by (cases ys) auto -qed - end diff --git a/thys/Word_Lib/Word_Lib.thy b/thys/Word_Lib/Word_Lib.thy --- a/thys/Word_Lib/Word_Lib.thy +++ b/thys/Word_Lib/Word_Lib.thy @@ -1,726 +1,731 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Additional Word Operations" theory Word_Lib imports "HOL-Library.Signed_Division" Word_Syntax Signed_Words begin +lemma unat_power_lower [simp]: + assumes nv: "n < LENGTH('a::len)" + shows "unat ((2::'a::len word) ^ n) = 2 ^ n" + using assms by transfer simp + definition ptr_add :: "'a :: len word \ nat \ 'a word" where "ptr_add ptr n \ ptr + of_nat n" definition complement :: "'a :: len word \ 'a word" where "complement x \ ~~ x" definition alignUp :: "'a::len word \ nat \ 'a word" where "alignUp x n \ x + 2 ^ n - 1 && complement (2 ^ n - 1)" (* standard notation for blocks of 2^n-1 words, usually aligned; abbreviation so it simplifies directly *) abbreviation mask_range :: "'a::len word \ nat \ 'a word set" where "mask_range p n \ {p .. p + mask n}" (* Haskellish names/syntax *) notation (input) test_bit ("testBit") definition w2byte :: "'a :: len word \ 8 word" where "w2byte \ ucast" lemmas sdiv_int_def = signed_divide_int_def lemmas smod_int_def = signed_modulo_int_def instantiation word :: (len) signed_division begin lift_definition signed_divide_word :: \'a::len word \ 'a word \ 'a word\ is \\k l. signed_take_bit (LENGTH('a) - Suc 0) k sdiv signed_take_bit (LENGTH('a) - Suc 0) l\ by (simp flip: signed_take_bit_decr_length_iff) lift_definition signed_modulo_word :: \'a::len word \ 'a word \ 'a word\ is \\k l. signed_take_bit (LENGTH('a) - Suc 0) k smod signed_take_bit (LENGTH('a) - Suc 0) l\ by (simp flip: signed_take_bit_decr_length_iff) instance .. end lemma sdiv_word_def [code]: \v sdiv w = word_of_int (sint v sdiv sint w)\ for v w :: \'a::len word\ by transfer simp lemma smod_word_def [code]: \v smod w = word_of_int (sint v smod sint w)\ for v w :: \'a::len word\ by transfer simp lemma sdiv_smod_id: \(a sdiv b) * b + (a smod b) = a\ for a b :: \'a::len word\ by (cases \sint a < 0\; cases \sint b < 0\) (simp_all add: signed_modulo_int_def sdiv_word_def smod_word_def) (* Tests *) lemma "( 4 :: word32) sdiv 4 = 1" "(-4 :: word32) sdiv 4 = -1" "(-3 :: word32) sdiv 4 = 0" "( 3 :: word32) sdiv -4 = 0" "(-3 :: word32) sdiv -4 = 0" "(-5 :: word32) sdiv -4 = 1" "( 5 :: word32) sdiv -4 = -1" by (simp_all add: sdiv_word_def signed_divide_int_def) lemma "( 4 :: word32) smod 4 = 0" "( 3 :: word32) smod 4 = 3" "(-3 :: word32) smod 4 = -3" "( 3 :: word32) smod -4 = 3" "(-3 :: word32) smod -4 = -3" "(-5 :: word32) smod -4 = -1" "( 5 :: word32) smod -4 = 1" by (simp_all add: smod_word_def signed_modulo_int_def signed_divide_int_def) (* Count leading zeros *) definition word_clz :: "'a::len word \ nat" where "word_clz w \ length (takeWhile Not (to_bl w))" (* Count trailing zeros *) definition word_ctz :: "'a::len word \ nat" where "word_ctz w \ length (takeWhile Not (rev (to_bl w)))" definition word_log2 :: "'a::len word \ nat" where "word_log2 (w::'a::len word) \ size w - 1 - word_clz w" (* Bit population count. Equivalent of __builtin_popcount. *) definition pop_count :: "('a::len) word \ nat" where "pop_count w \ length (filter id (to_bl w))" (* Sign extension from bit n *) definition sign_extend :: "nat \ 'a::len word \ 'a word" where "sign_extend n w \ if w !! n then w || ~~ (mask n) else w && mask n" lemma sign_extend_eq_signed_take_bit: \sign_extend = signed_take_bit\ proof (rule ext)+ fix n and w :: \'a::len word\ show \sign_extend n w = signed_take_bit n w\ proof (rule bit_word_eqI) fix q assume \q < LENGTH('a)\ then show \bit (sign_extend n w) q \ bit (signed_take_bit n w) q\ by (auto simp add: test_bit_eq_bit bit_signed_take_bit_iff sign_extend_def bit_and_iff bit_or_iff bit_not_iff bit_mask_iff not_less exp_eq_0_imp_not_bit not_le min_def) qed qed definition sign_extended :: "nat \ 'a::len word \ bool" where "sign_extended n w \ \i. n < i \ i < size w \ w !! i = w !! n" lemma ptr_add_0 [simp]: "ptr_add ref 0 = ref " unfolding ptr_add_def by simp lemma shiftl_power: "(shiftl1 ^^ x) (y::'a::len word) = 2 ^ x * y" apply (induct x) apply simp apply (simp add: shiftl1_2t) done lemmas of_bl_reasoning = to_bl_use_of_bl of_bl_append lemma uint_of_bl_is_bl_to_bin_drop: "length (dropWhile Not l) \ LENGTH('a) \ uint (of_bl l :: 'a::len word) = bl_to_bin l" apply transfer apply (simp add: take_bit_eq_mod) apply (rule Divides.mod_less) apply (rule bl_to_bin_ge0) using bl_to_bin_lt2p_drop apply (rule order.strict_trans2) apply simp done corollary uint_of_bl_is_bl_to_bin: "length l\LENGTH('a) \ uint ((of_bl::bool list\ ('a :: len) word) l) = bl_to_bin l" apply(rule uint_of_bl_is_bl_to_bin_drop) using le_trans length_dropWhile_le by blast lemma bin_to_bl_or: "bin_to_bl n (a OR b) = map2 (\) (bin_to_bl n a) (bin_to_bl n b)" using bl_or_aux_bin[where n=n and v=a and w=b and bs="[]" and cs="[]"] by simp lemma word_ops_nth [simp]: shows word_or_nth: "(x || y) !! n = (x !! n \ y !! n)" and word_and_nth: "(x && y) !! n = (x !! n \ y !! n)" and word_xor_nth: "(x xor y) !! n = (x !! n \ y !! n)" by ((cases "n < size x", auto dest: test_bit_size simp: word_ops_nth_size word_size)[1])+ (* simp del to avoid warning on the simp add in iff *) declare test_bit_1 [simp del, iff] (* test: *) lemma "1 < (1024::32 word) \ 1 \ (1024::32 word)" by simp lemma and_not_mask: "w AND NOT (mask n) = (w >> n) << n" for w :: \'a::len word\ apply (rule word_eqI) apply (simp add : word_ops_nth_size word_size) apply (simp add : nth_shiftr nth_shiftl) by auto lemma and_mask: "w AND mask n = (w << (size w - n)) >> (size w - n)" for w :: \'a::len word\ apply (rule word_eqI) apply (simp add : word_ops_nth_size word_size) apply (simp add : nth_shiftr nth_shiftl) by auto lemma AND_twice [simp]: "(w && m) && m = w && m" by (fact and.right_idem) lemma word_combine_masks: "w && m = z \ w && m' = z' \ w && (m || m') = (z || z')" by (auto simp: word_eq_iff) lemma nth_w2p_same: "(2^n :: 'a :: len word) !! n = (n < LENGTH('a))" by (simp add : nth_w2p) lemma p2_gt_0: "(0 < (2 ^ n :: 'a :: len word)) = (n < LENGTH('a))" by (simp add : word_gt_0 not_le) lemmas uint_2p_alt = uint_2p [unfolded p2_gt_0] lemma shiftr_div_2n_w: "n < size w \ w >> n = w div (2^n :: 'a :: len word)" apply (unfold word_div_def) apply (simp add : uint_2p_alt word_size) apply (rule word_uint.Rep_inverse' [THEN sym]) apply (rule shiftr_div_2n) done lemmas less_def = less_eq [symmetric] lemmas le_def = not_less [symmetric, where ?'a = nat] lemmas p2_eq_0 [simp] = trans [OF eq_commute iffD2 [OF Not_eq_iff p2_gt_0, folded le_def, unfolded word_gt_0 not_not]] lemma neg_mask_is_div': "n < size w \ w AND NOT (mask n) = ((w div (2 ^ n)) * (2 ^ n))" for w :: \'a::len word\ by (simp add : and_not_mask shiftr_div_2n_w shiftl_t2n word_size) lemma neg_mask_is_div: "w AND NOT (mask n) = (w div 2^n) * 2^n" for w :: \'a::len word\ apply (cases "n < size w") apply (erule neg_mask_is_div') apply (simp add: word_size) apply (frule p2_gt_0 [THEN Not_eq_iff [THEN iffD2], THEN iffD2]) apply (simp add: word_gt_0 del: p2_eq_0) apply (rule word_eqI) apply (simp add: word_ops_nth_size word_size) done lemma and_mask_arith': "0 < n \ w AND mask n = (w * 2^(size w - n)) div 2^(size w - n)" for w :: \'a::len word\ by (simp add: and_mask shiftr_div_2n_w shiftl_t2n word_size mult.commute) lemmas p2len = iffD2 [OF p2_eq_0 order_refl] lemma and_mask_arith: "w AND mask n = (w * 2^(size w - n)) div 2^(size w - n)" for w :: \'a::len word\ apply (cases "0 < n") apply (auto elim!: and_mask_arith') apply (simp add: word_size) done lemma mask_2pm1: "mask n = 2 ^ n - (1 :: 'a::len word)" by (fact mask_eq_decr_exp) lemma add_mask_fold: "x + 2 ^ n - 1 = x + mask n" for x :: \'a::len word\ by (simp add: mask_eq_decr_exp) lemma word_and_mask_le_2pm1: "w && mask n \ 2 ^ n - 1" by (simp add: mask_2pm1[symmetric] word_and_le1) lemma is_aligned_AND_less_0: "u && mask n = 0 \ v < 2^n \ u && v = 0" apply (drule less_mask_eq) apply (simp add: mask_2pm1) apply (rule word_eqI) apply (clarsimp simp add: word_size) apply (drule_tac x=na in word_eqD) apply (drule_tac x=na in word_eqD) apply simp done lemma le_shiftr1: "u <= v \ shiftr1 u <= shiftr1 v" apply (unfold word_le_def shiftr1_eq word_ubin.eq_norm) apply (unfold bin_rest_trunc_i trans [OF bintrunc_bintrunc_l word_ubin.norm_Rep, unfolded word_ubin.norm_Rep, OF order_refl [THEN le_SucI]]) apply (case_tac "uint u" rule: bin_exhaust) apply (rename_tac bs bit) apply (case_tac "uint v" rule: bin_exhaust) apply (rename_tac bs' bit') apply (case_tac "bit") apply (case_tac "bit'", auto simp: less_eq_int_code)[1] apply (case_tac bit') apply (simp add: less_eq_int_code) apply (simp add: less_eq_int_code) done lemma le_shiftr: "u \ v \ u >> (n :: nat) \ (v :: 'a :: len word) >> n" apply (unfold shiftr_def) apply (induct_tac "n") apply auto apply (erule le_shiftr1) done lemma shiftr_mask_le: "n <= m \ mask n >> m = (0 :: 'a::len word)" apply (rule word_eqI) apply (simp add: word_size nth_shiftr) done lemmas shiftr_mask = order_refl [THEN shiftr_mask_le, simp] lemma word_leI: "(\n. \n < size (u::'a::len word); u !! n \ \ (v::'a::len word) !! n) \ u <= v" - apply (rule xtr4) + apply (rule xtrans(4)) apply (rule word_and_le2) apply (rule word_eqI) apply (simp add: word_ao_nth) apply safe apply assumption apply (erule_tac [2] asm_rl) apply (unfold word_size) by auto lemma le_mask_iff: "(w \ mask n) = (w >> n = 0)" for w :: \'a::len word\ apply safe apply (rule word_le_0_iff [THEN iffD1]) - apply (rule xtr3) + apply (rule xtrans(3)) apply (erule_tac [2] le_shiftr) apply simp apply (rule word_leI) apply (rename_tac n') apply (drule_tac x = "n' - n" in word_eqD) apply (simp add : nth_shiftr word_size) apply (case_tac "n <= n'") by auto lemma and_mask_eq_iff_shiftr_0: "(w AND mask n = w) = (w >> n = 0)" for w :: \'a::len word\ apply (unfold test_bit_eq_iff [THEN sym]) apply (rule iffI) apply (rule ext) apply (rule_tac [2] ext) apply (auto simp add : word_ao_nth nth_shiftr) apply (drule arg_cong) apply (drule iffD2) apply assumption apply (simp add : word_ao_nth) prefer 2 apply (simp add : word_size test_bit_bin) apply (drule_tac f = "%u. u !! (x - n)" in arg_cong) apply (simp add : nth_shiftr) apply (case_tac "n <= x") apply auto done lemmas and_mask_eq_iff_le_mask = trans [OF and_mask_eq_iff_shiftr_0 le_mask_iff [THEN sym]] lemma mask_shiftl_decompose: "mask m << n = mask (m + n) && ~~ (mask n)" by (auto intro!: word_eqI simp: and_not_mask nth_shiftl nth_shiftr word_size) lemmas bin_sc_minus_simps = bin_sc_simps (2,3,4) [THEN [2] trans, OF bin_sc_minus [THEN sym]] lemma NOT_eq: "NOT (x :: 'a :: len word) = - x - 1" apply (cut_tac x = "x" in word_add_not) apply (drule add.commute [THEN trans]) apply (drule eq_diff_eq [THEN iffD2]) by simp lemma NOT_mask: "NOT (mask n :: 'a::len word) = - (2 ^ n)" by (simp add : NOT_eq mask_2pm1) lemma le_m1_iff_lt: "(x > (0 :: 'a :: len word)) = ((y \ x - 1) = (y < x))" by uint_arith lemmas gt0_iff_gem1 = iffD1 [OF iffD1 [OF iff_left_commute le_m1_iff_lt] order_refl] lemmas power_2_ge_iff = trans [OF gt0_iff_gem1 [THEN sym] p2_gt_0] lemma le_mask_iff_lt_2n: "n < len_of TYPE ('a) = (((w :: 'a :: len word) \ mask n) = (w < 2 ^ n))" unfolding mask_2pm1 by (rule trans [OF p2_gt_0 [THEN sym] le_m1_iff_lt]) lemmas mask_lt_2pn = le_mask_iff_lt_2n [THEN iffD1, THEN iffD1, OF _ order_refl] lemma bang_eq: fixes x :: "'a::len word" shows "(x = y) = (\n. x !! n = y !! n)" by (subst test_bit_eq_iff[symmetric]) fastforce lemma word_unat_power: "(2 :: 'a :: len word) ^ n = of_nat (2 ^ n)" by simp lemma of_nat_mono_maybe: assumes xlt: "x < 2 ^ len_of TYPE ('a)" shows "y < x \ of_nat y < (of_nat x :: 'a :: len word)" apply (subst word_less_nat_alt) apply (subst unat_of_nat)+ apply (subst mod_less) apply (erule order_less_trans [OF _ xlt]) apply (subst mod_less [OF xlt]) apply assumption done lemma shiftl_over_and_dist: fixes a::"'a::len word" shows "(a AND b) << c = (a << c) AND (b << c)" apply(rule word_eqI) apply(simp add: word_ao_nth nth_shiftl, safe) done lemma shiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >> c = (a >> c) AND (b >> c)" apply(rule word_eqI) apply(simp add:nth_shiftr word_ao_nth) done lemma sshiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >>> c = (a >>> c) AND (b >>> c)" apply(rule word_eqI) apply(simp add:nth_sshiftr word_ao_nth word_size) done lemma shiftl_over_or_dist: fixes a::"'a::len word" shows "a OR b << c = (a << c) OR (b << c)" apply(rule word_eqI) apply(simp add:nth_shiftl word_ao_nth, safe) done lemma shiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >> c = (a >> c) OR (b >> c)" apply(rule word_eqI) apply(simp add:nth_shiftr word_ao_nth) done lemma sshiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >>> c = (a >>> c) OR (b >>> c)" apply(rule word_eqI) apply(simp add:nth_sshiftr word_ao_nth word_size) done lemmas shift_over_ao_dists = shiftl_over_or_dist shiftr_over_or_dist sshiftr_over_or_dist shiftl_over_and_dist shiftr_over_and_dist sshiftr_over_and_dist lemma shiftl_shiftl: fixes a::"'a::len word" shows "a << b << c = a << (b + c)" apply(rule word_eqI) apply(auto simp:word_size nth_shiftl add.commute add.left_commute) done lemma shiftr_shiftr: fixes a::"'a::len word" shows "a >> b >> c = a >> (b + c)" apply(rule word_eqI) apply(simp add:word_size nth_shiftr add.left_commute add.commute) done lemma shiftl_shiftr1: fixes a::"'a::len word" shows "c \ b \ a << b >> c = a AND (mask (size a - b)) << (b - c)" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ao_nth) done lemma shiftl_shiftr2: fixes a::"'a::len word" shows "b < c \ a << b >> c = (a >> (c - b)) AND (mask (size a - c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ao_nth) done lemma shiftr_shiftl1: fixes a::"'a::len word" shows "c \ b \ a >> b << c = (a >> (b - c)) AND (NOT (mask c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ops_nth_size) done lemma shiftr_shiftl2: fixes a::"'a::len word" shows "b < c \ a >> b << c = (a << (c - b)) AND (NOT (mask c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ops_nth_size) done lemmas multi_shift_simps = shiftl_shiftl shiftr_shiftr shiftl_shiftr1 shiftl_shiftr2 shiftr_shiftl1 shiftr_shiftl2 lemma word_and_max_word: fixes a::"'a::len word" shows "x = max_word \ a AND x = a" by simp lemma word_and_full_mask_simp: \x && Bit_Operations.mask LENGTH('a) = x\ for x :: \'a::len word\ proof (rule bit_eqI) fix n assume \2 ^ n \ (0 :: 'a word)\ then have \n < LENGTH('a)\ by simp then show \bit (x && Bit_Operations.mask LENGTH('a)) n \ bit x n\ by (simp add: bit_and_iff bit_mask_iff) qed lemma word8_and_max_simp: \x && 0xFF = x\ for x :: \8 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) lemma word16_and_max_simp: \x && 0xFFFF = x\ for x :: \16 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) lemma word32_and_max_simp: \x && 0xFFFFFFFF = x\ for x :: \32 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) lemma word64_and_max_simp: \x && 0xFFFFFFFFFFFFFFFF = x\ for x :: \64 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) lemmas word_and_max_simps = word8_and_max_simp word16_and_max_simp word32_and_max_simp word64_and_max_simp lemma word_and_1_bl: fixes x::"'a::len word" shows "(x AND 1) = of_bl [x !! 0]" by (simp add: mod_2_eq_odd test_bit_word_eq and_one_eq) lemma word_1_and_bl: fixes x::"'a::len word" shows "(1 AND x) = of_bl [x !! 0]" by (simp add: mod_2_eq_odd test_bit_word_eq one_and_eq) lemma scast_scast_id [simp]: "scast (scast x :: ('a::len) signed word) = (x :: 'a word)" "scast (scast y :: ('a::len) word) = (y :: 'a signed word)" by (auto simp: is_up scast_up_scast_id) lemma scast_ucast_id [simp]: "scast (ucast (x :: 'a::len word) :: 'a signed word) = x" by (metis down_cast_same is_down len_signed order_refl scast_scast_id(1)) lemma ucast_scast_id [simp]: "ucast (scast (x :: 'a::len signed word) :: 'a word) = x" by (metis scast_scast_id(2) scast_ucast_id) lemma scast_of_nat [simp]: "scast (of_nat x :: 'a::len signed word) = (of_nat x :: 'a word)" by transfer simp lemma ucast_of_nat: "is_down (ucast :: 'a :: len word \ 'b :: len word) \ ucast (of_nat n :: 'a word) = (of_nat n :: 'b word)" by transfer simp (* shortcut for some specific lengths *) lemma word_fixed_sint_1[simp]: "sint (1::8 word) = 1" "sint (1::16 word) = 1" "sint (1::32 word) = 1" "sint (1::64 word) = 1" by (auto simp: sint_word_ariths) lemma word_sint_1 [simp]: "sint (1::'a::len word) = (if LENGTH('a) = 1 then -1 else 1)" by (cases \LENGTH('a)\) (simp_all add: not_le sint_uint le_Suc_eq sbintrunc_minus_simps) lemma scast_1': "(scast (1::'a::len word) :: 'b::len word) = (word_of_int (sbintrunc (LENGTH('a::len) - Suc 0) (1::int)))" by transfer simp lemma scast_1: "(scast (1::'a::len word) :: 'b::len word) = (if LENGTH('a) = 1 then -1 else 1)" by (fact signed_1) lemma scast_eq_scast_id [simp]: "((scast (a :: 'a::len signed word) :: 'a word) = scast b) = (a = b)" by (metis ucast_scast_id) lemma ucast_eq_ucast_id [simp]: "((ucast (a :: 'a::len word) :: 'a signed word) = ucast b) = (a = b)" by (metis scast_ucast_id) lemma scast_ucast_norm [simp]: "(ucast (a :: 'a::len word) = (b :: 'a signed word)) = (a = scast b)" "((b :: 'a signed word) = ucast (a :: 'a::len word)) = (a = scast b)" by (metis scast_ucast_id ucast_scast_id)+ lemma of_bl_drop: "of_bl (drop n xs) = (of_bl xs && mask (length xs - n))" apply (clarsimp simp: bang_eq test_bit_of_bl rev_nth cong: rev_conj_cong) apply (safe; simp add: word_size to_bl_nth) done lemma of_int_uint: "of_int (uint x) = x" by (fact word_of_int_uint) lemma shiftr_mask2: "n \ LENGTH('a) \ (mask n >> m :: ('a :: len) word) = mask (n - m)" apply (rule word_eqI) apply (simp add: nth_shiftr word_size) apply arith done corollary word_plus_and_or_coroll: "x && y = 0 \ x + y = x || y" using word_plus_and_or[where x=x and y=y] by simp corollary word_plus_and_or_coroll2: "(x && w) + (x && ~~ w) = x" apply (subst word_plus_and_or_coroll) apply (rule word_eqI, simp add: word_size word_ops_nth_size) apply (rule word_eqI, simp add: word_size word_ops_nth_size) apply blast done lemma less_le_mult_nat': "w * c < b * c ==> 0 \ c ==> Suc w * c \ b * (c::nat)" apply (rule mult_right_mono) apply (rule Suc_leI) apply (erule (1) mult_right_less_imp_less) apply assumption done lemmas less_le_mult_nat = less_le_mult_nat'[simplified distrib_right, simplified] (* FIXME: these should eventually be moved to HOL/Word. *) lemmas extra_sle_sless_unfolds [simp] = word_sle_eq[where a=0 and b=1] word_sle_eq[where a=0 and b="numeral n"] word_sle_eq[where a=1 and b=0] word_sle_eq[where a=1 and b="numeral n"] word_sle_eq[where a="numeral n" and b=0] word_sle_eq[where a="numeral n" and b=1] word_sless_alt[where a=0 and b=1] word_sless_alt[where a=0 and b="numeral n"] word_sless_alt[where a=1 and b=0] word_sless_alt[where a=1 and b="numeral n"] word_sless_alt[where a="numeral n" and b=0] word_sless_alt[where a="numeral n" and b=1] for n lemma to_bl_1: "to_bl (1::'a::len word) = replicate (LENGTH('a) - 1) False @ [True]" by (rule nth_equalityI) (auto simp add: to_bl_unfold nth_append rev_nth bit_1_iff not_less not_le) lemma list_of_false: "True \ set xs \ xs = replicate (length xs) False" by (induct xs, simp_all) lemma eq_zero_set_bl: "(w = 0) = (True \ set (to_bl w))" using list_of_false word_bl.Rep_inject by fastforce lemma diff_diff_less: "(i < m - (m - (n :: nat))) = (i < m \ i < n)" by auto lemma pop_count_0[simp]: "pop_count 0 = 0" by (clarsimp simp:pop_count_def) lemma pop_count_1[simp]: "pop_count 1 = 1" by (clarsimp simp:pop_count_def to_bl_1) lemma pop_count_0_imp_0: "(pop_count w = 0) = (w = 0)" apply (rule iffI) apply (clarsimp simp:pop_count_def) apply (subst (asm) filter_empty_conv) apply (clarsimp simp:eq_zero_set_bl) apply fast apply simp done end diff --git a/thys/Word_Lib/Word_Lib_Sumo.thy b/thys/Word_Lib/Word_Lib_Sumo.thy --- a/thys/Word_Lib/Word_Lib_Sumo.thy +++ b/thys/Word_Lib/Word_Lib_Sumo.thy @@ -1,83 +1,84 @@ section \Ancient comprehensive Word Library\ theory Word_Lib_Sumo imports "HOL-Library.Word" Aligned Ancient_Numeral Bit_Comprehension Bits_Int Bitwise_Signed Bitwise Enumeration_Word Generic_set_bit Hex_Words Least_significant_bit More_Arithmetic More_Divides More_List More_Misc + Strict_part_mono Most_significant_bit Next_and_Prev Norm_Words Reversed_Bit_Lists Rsplit Signed_Words Traditional_Infix_Syntax Typedef_Morphisms Type_Syntax Word_EqI Word_Lemmas_32 Word_Lemmas_64 Word_Lemmas Word_Lib Word_Setup_32 Word_Setup_64 Word_Syntax begin declare signed_take_bit_Suc [simp] lemmas bshiftr1_def = bshiftr1_eq lemmas is_down_def = is_down_eq lemmas is_up_def = is_up_eq lemmas mask_def = mask_eq_decr_exp lemmas scast_def = scast_eq lemmas shiftl1_def = shiftl1_eq lemmas shiftr1_def = shiftr1_eq lemmas sshiftr1_def = sshiftr1_eq lemmas sshiftr_def = sshiftr_eq_funpow_sshiftr1 lemmas to_bl_def = to_bl_eq lemmas ucast_def = ucast_eq lemmas unat_def = unat_eq_nat_uint lemmas word_cat_def = word_cat_eq lemmas word_reverse_def = word_reverse_eq_of_bl_rev_to_bl lemmas word_roti_def = word_roti_eq_word_rotr_word_rotl lemmas word_rotl_def = word_rotl_eq lemmas word_rotr_def = word_rotr_eq lemmas word_sle_def = word_sle_eq lemmas word_sless_def = word_sless_eq lemmas uint_0 = uint_nonnegative lemmas uint_lt = uint_bounded lemmas uint_mod_same = uint_idem lemmas of_nth_def = word_set_bits_def lemmas of_nat_word_eq_iff = word_of_nat_eq_iff lemmas of_nat_word_eq_0_iff = word_of_nat_eq_0_iff lemmas of_int_word_eq_iff = word_of_int_eq_iff lemmas of_int_word_eq_0_iff = word_of_int_eq_0_iff lemmas word_next_def = word_next_unfold lemmas word_prev_def = word_prev_unfold lemmas is_aligned_def = is_aligned_iff_dvd_nat lemma shiftl_transfer [transfer_rule]: includes lifting_syntax shows "(pcr_word ===> (=) ===> pcr_word) (<<) (<<)" by (unfold shiftl_eq_push_bit) transfer_prover end