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,1076 @@ (* 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 - "HOL-Word.Misc_Arithmetic" + "Word_Lib.Misc_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]) 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/Berlekamp_Zassenhaus/ROOT b/thys/Berlekamp_Zassenhaus/ROOT --- a/thys/Berlekamp_Zassenhaus/ROOT +++ b/thys/Berlekamp_Zassenhaus/ROOT @@ -1,62 +1,61 @@ chapter AFP session Pre_BZ (AFP) in "Pre_BZ" = "Subresultants" + options [timeout = 600] sessions "HOL-Number_Theory" - "HOL-Word" "HOL-Types_To_Sets" Polynomial_Factorization Polynomial_Interpolation "Efficient-Mergesort" Show theories Jordan_Normal_Form.Matrix_Kernel Jordan_Normal_Form.Gauss_Jordan_IArray_Impl "Efficient-Mergesort.Efficient_Sort" "HOL-Number_Theory.Residues" "HOL-Types_To_Sets.Types_To_Sets" Native_Word.Uint32 Native_Word.Uint64 Native_Word.Code_Target_Bits_Int Polynomial_Factorization.Rational_Factorization session Berlekamp_Zassenhaus (AFP) = Pre_BZ + description "Berlekamp-Zassenhaus's Factorization Algorithm" options [timeout = 1200] theories (* Arithmetic for Finite Fields *) Finite_Field Finite_Field_Record_Based Arithmetic_Record_Based Matrix_Record_Based Poly_Mod Poly_Mod_Finite_Field Poly_Mod_Finite_Field_Record_Based Polynomial_Record_Based (* Berlekamp's Factorization Algorithm *) Chinese_Remainder_Poly Berlekamp_Type_Based Distinct_Degree_Factorization Finite_Field_Factorization_Record_Based (* Hensel Lifting *) Hensel_Lifting_Type_Based Hensel_Lifting Berlekamp_Hensel (* Zassenhaus Reconstruction *) Square_Free_Int_To_Square_Free_GFp Suitable_Prime Degree_Bound Mahler_Measure Factor_Bound Sublist_Iteration Reconstruction (* An Efficient Algorithm for Integer Polynomials *) Berlekamp_Zassenhaus Square_Free_Factorization_Int Factorize_Int_Poly Factorize_Rat_Poly document_files "algorithm2e.sty" "root.bib" "root.tex" diff --git a/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_State.thy b/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_State.thy --- a/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_State.thy +++ b/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_State.thy @@ -1,70 +1,70 @@ section \State for SM\ theory SM_State -imports SM_Syntax "HOL-Word.Word" "HOL-Word.Misc_Typedef" "HOL-Library.Multiset" +imports SM_Syntax "HOL-Library.Word" "Word_Lib.Misc_Typedef" "HOL-Library.Multiset" begin section \Values\ text \The primitive values are fixed-size signed integers\ type_synonym word_size = 32 \ \Word size\ type_synonym signed = "word_size Word.word" \ \Signed integer\ text \Currently, we only have signed integer values. This may change if we extend the language, and allow, i.e., channel pointers, pids or process references\ type_synonym val = signed \ \Value type\ section \Configurations\ type_synonym valuation = "ident \ val" record local_state = variables :: valuation record global_state = variables :: valuation text \The effect of actions is on focused states\ type_synonym focused_state = "local_state \ global_state" section \Utilities\ abbreviation "word_len \ LENGTH(word_size)" abbreviation "signeds \ sints (LENGTH(word_size))" definition min_signed :: int where "min_signed \ -(2^(word_len - 1))" definition max_signed :: int where "max_signed \ 2^(word_len - 1) - 1" definition signed_of_int :: "int \ signed" where "signed_of_int i \ word_of_int i" definition int_of_signed :: "signed \ int" where "int_of_signed i == sint i" lemma si_inv[simp]: "signed_of_int (int_of_signed i) = i" unfolding signed_of_int_def int_of_signed_def by simp lemma int_of_signed_in_range[simp]: "int_of_signed i \ min_signed" "int_of_signed i \ max_signed" unfolding int_of_signed_def min_signed_def max_signed_def apply - apply (rule sint_ge) using sint_lt[of i] by simp lemma is_inv[simp]: "\ i\min_signed; i\max_signed \ \ (int_of_signed (signed_of_int i)) = i" by (simp add: signed_take_bit_int_eq_self min_signed_def max_signed_def int_of_signed_def signed_of_int_def) primrec val_of_bool :: "bool \ val" where "val_of_bool False = 0" | "val_of_bool True = 1" definition bool_of_val :: "val \ bool" where "bool_of_val v \ v\0" lemma bool_of_val_invs[simp]: "bool_of_val (val_of_bool b) = b" "val_of_bool (bool_of_val v) = (if v=0 then 0 else 1)" unfolding bool_of_val_def by (cases b) auto end diff --git a/thys/CakeML/ROOT b/thys/CakeML/ROOT --- a/thys/CakeML/ROOT +++ b/thys/CakeML/ROOT @@ -1,66 +1,68 @@ chapter AFP -session LEM (AFP) in "generated" = "HOL-Word" + +session LEM (AFP) in "generated" = HOL + options [timeout = 300] + sessions + Word_Lib theories Lem_pervasives Lem_pervasives_extra Lem_list_extra Lem_set_extra Lem_string Lem_string_extra session CakeML (AFP) = LEM + options [timeout = 3600] sessions Coinductive IEEE_Floating_Point Show directories "doc" "generated/CakeML" "Tests" theories "doc/Doc_Generated" "generated/CakeML/Ast" "generated/CakeML/AstAuxiliary" "generated/CakeML/BigSmallInvariants" "generated/CakeML/BigStep" "generated/CakeML/Evaluate" "generated/CakeML/Lib" "generated/CakeML/LibAuxiliary" "generated/CakeML/Namespace" "generated/CakeML/NamespaceAuxiliary" "generated/CakeML/PrimTypes" "generated/CakeML/SemanticPrimitives" "generated/CakeML/SemanticPrimitivesAuxiliary" "generated/CakeML/SimpleIO" "generated/CakeML/Tokens" "generated/CakeML/TypeSystem" "generated/CakeML/TypeSystemAuxiliary" "doc/Doc_Proofs" Semantic_Extras Evaluate_Termination Evaluate_Clock Evaluate_Single Big_Step_Determ Big_Step_Fun_Equiv Big_Step_Total Big_Step_Unclocked Big_Step_Unclocked_Single Big_Step_Clocked Matching CakeML_Code CakeML_Quickcheck CakeML_Compiler theories [condition = "ISABELLE_CAKEML_HOME,ISABELLE_CC", document = false] "Tests/Compiler_Test" theories [condition = "ISABELLE_GHC", document = false] "Tests/Code_Test_Haskell" document_files - "root.tex" \ No newline at end of file + "root.tex" diff --git a/thys/CakeML/generated/Lem.thy b/thys/CakeML/generated/Lem.thy --- a/thys/CakeML/generated/Lem.thy +++ b/thys/CakeML/generated/Lem.thy @@ -1,109 +1,109 @@ (*========================================================================*) (* Lem *) (* *) (* Dominic Mulligan, University of Cambridge *) (* Francesco Zappa Nardelli, INRIA Paris-Rocquencourt *) (* Gabriel Kerneis, University of Cambridge *) (* Kathy Gray, University of Cambridge *) (* Peter Boehm, University of Cambridge (while working on Lem) *) (* Peter Sewell, University of Cambridge *) (* Scott Owens, University of Kent *) (* Thomas Tuerk, University of Cambridge *) (* Brian Campbell, University of Edinburgh *) (* Shaked Flur, University of Cambridge *) (* Thomas Bauereiss, University of Cambridge *) (* Stephen Kell, University of Cambridge *) (* Thomas Williams *) (* Lars Hupel *) (* Basile Clement *) (* *) (* The Lem sources are copyright 2010-2018 *) (* by the authors above and Institut National de Recherche en *) (* Informatique et en Automatique (INRIA). *) (* *) (* All files except ocaml-lib/pmap.{ml,mli} and ocaml-libpset.{ml,mli} *) (* are distributed under the license below. The former are distributed *) (* under the LGPLv2, as in the LICENSE file. *) (* *) (* *) (* Redistribution and use in source and binary forms, with or without *) (* modification, are permitted provided that the following conditions *) (* are met: *) (* 1. Redistributions of source code must retain the above copyright *) (* notice, this list of conditions and the following disclaimer. *) (* 2. Redistributions in binary form must reproduce the above copyright *) (* notice, this list of conditions and the following disclaimer in the *) (* documentation and/or other materials provided with the distribution. *) (* 3. The names of the authors may not be used to endorse or promote *) (* products derived from this software without specific prior written *) (* permission. *) (* *) (* THIS SOFTWARE IS PROVIDED BY THE AUTHORS ``AS IS'' AND ANY EXPRESS *) (* OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE IMPLIED *) (* WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE *) (* ARE DISCLAIMED. IN NO EVENT SHALL THE AUTHORS BE LIABLE FOR ANY *) (* DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL *) (* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE *) (* GOODS OR SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS *) (* INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER *) (* IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING NEGLIGENCE OR *) (* OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS SOFTWARE, EVEN *) (* IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. *) (*========================================================================*) chapter\Mappings of Syntax needed by Lem\ theory "Lem" imports LemExtraDefs - "HOL-Word.Word" - "HOL-Word.Reversed_Bit_Lists" + "HOL-Library.Word" + "Word_Lib.Reversed_Bit_Lists" begin type_synonym numeral = nat subsection \Finite Maps\ abbreviation (input) "map_find k m \ the (m k)" abbreviation (input) "map_update k v m \ m (k \ v)" abbreviation (input) "map_remove k m \ m |` (- {k})" abbreviation (input) "map_any P m \ \ (k, v) \ map_to_set m. P k v" abbreviation (input) "map_all P m \ \ (k, v) \ map_to_set m. P k v" subsection \Lists\ abbreviation (input) "list_mem e l \ (e \ set l)" abbreviation (input) "list_forall P l \ (\e\set l. P e)" abbreviation (input) "list_exists P l \ (\e\set l. P e)" abbreviation (input) "list_unzip l \ (map fst l, map snd l)" subsection \Sets\ abbreviation (input) "set_filter P (s::'a set) \ {x \ s. P x}" abbreviation (input) "set_bigunion S \ \ S" abbreviation (input) "set_biginter S \ \ S" subsection \Natural numbers\ subsection \Integers\ subsection \Dummy\ consts bitwise_xor :: "nat \ nat \ nat" num_asr :: "nat \ nat \ nat" num_lsl :: "nat \ nat \ nat" bitwise_or :: "nat \ nat \ nat" bitwise_not :: "nat \ nat" bitwise_and :: "nat \ nat \ nat" subsection \Machine words\ definition word_update :: "'a::len word \ nat \ nat \ 'b::len word \ 'a word" where "word_update v lo hi w = (let sz = size v in of_bl (take (sz-hi-1) (to_bl v) @ to_bl w @ drop (sz-lo) (to_bl v)))" end diff --git a/thys/CakeML/generated/Lem_machine_word.thy b/thys/CakeML/generated/Lem_machine_word.thy --- a/thys/CakeML/generated/Lem_machine_word.thy +++ b/thys/CakeML/generated/Lem_machine_word.thy @@ -1,452 +1,452 @@ chapter \Generated by Lem from \machine_word.lem\.\ theory "Lem_machine_word" imports Main "Lem_bool" "Lem_num" "Lem_basic_classes" "Lem_show" "Lem_function" - "HOL-Word.Word" - "HOL-Word.Misc_msb" + "HOL-Library.Word" + "Word_Lib.Misc_msb" begin \ \\open import Bool Num Basic_classes Show Function\\ -\ \\open import {isabelle} `HOL-Word.Word`\\ +\ \\open import {isabelle} `HOL-Library.Word`\\ \ \\open import {hol} `wordsTheory` `wordsLib` `bitstringTheory` `integer_wordTheory`\\ \ \\type mword 'a\\ \ \\class (Size 'a) val size : nat end\\ \ \\val native_size : forall 'a. nat\\ \ \\val ocaml_inject : forall 'a. nat * natural -> mword 'a\\ \ \\ A singleton type family that can be used to carry a size as the type parameter \\ \ \\type itself 'a\\ \ \\val the_value : forall 'a. itself 'a\\ \ \\val size_itself : forall 'a. Size 'a => itself 'a -> nat\\ definition size_itself :: "('a::len)itself \ nat " where " size_itself x = ( (len_of (TYPE(_) :: 'a itself)))" \ \\*****************************************************************\\ \ \\ Fixed bitwidths extracted from Anthony's models. \\ \ \\ \\ \ \\ If you need a size N that is not included here, put the lines \\ \ \\ \\ \ \\ type tyN \\ \ \\ instance (Size tyN) let size = N end \\ \ \\ declare isabelle target_rep type tyN = `N` \\ \ \\ declare hol target_rep type tyN = `N` \\ \ \\ \\ \ \\ in your project, replacing N in each line. \\ \ \\*****************************************************************\\ \ \\type ty1\\ \ \\type ty2\\ \ \\type ty3\\ \ \\type ty4\\ \ \\type ty5\\ \ \\type ty6\\ \ \\type ty7\\ \ \\type ty8\\ \ \\type ty9\\ \ \\type ty10\\ \ \\type ty11\\ \ \\type ty12\\ \ \\type ty13\\ \ \\type ty14\\ \ \\type ty15\\ \ \\type ty16\\ \ \\type ty17\\ \ \\type ty18\\ \ \\type ty19\\ \ \\type ty20\\ \ \\type ty21\\ \ \\type ty22\\ \ \\type ty23\\ \ \\type ty24\\ \ \\type ty25\\ \ \\type ty26\\ \ \\type ty27\\ \ \\type ty28\\ \ \\type ty29\\ \ \\type ty30\\ \ \\type ty31\\ \ \\type ty32\\ \ \\type ty33\\ \ \\type ty34\\ \ \\type ty35\\ \ \\type ty36\\ \ \\type ty37\\ \ \\type ty38\\ \ \\type ty39\\ \ \\type ty40\\ \ \\type ty41\\ \ \\type ty42\\ \ \\type ty43\\ \ \\type ty44\\ \ \\type ty45\\ \ \\type ty46\\ \ \\type ty47\\ \ \\type ty48\\ \ \\type ty49\\ \ \\type ty50\\ \ \\type ty51\\ \ \\type ty52\\ \ \\type ty53\\ \ \\type ty54\\ \ \\type ty55\\ \ \\type ty56\\ \ \\type ty57\\ \ \\type ty58\\ \ \\type ty59\\ \ \\type ty60\\ \ \\type ty61\\ \ \\type ty62\\ \ \\type ty63\\ \ \\type ty64\\ \ \\type ty65\\ \ \\type ty66\\ \ \\type ty67\\ \ \\type ty68\\ \ \\type ty69\\ \ \\type ty70\\ \ \\type ty71\\ \ \\type ty72\\ \ \\type ty73\\ \ \\type ty74\\ \ \\type ty75\\ \ \\type ty76\\ \ \\type ty77\\ \ \\type ty78\\ \ \\type ty79\\ \ \\type ty80\\ \ \\type ty81\\ \ \\type ty82\\ \ \\type ty83\\ \ \\type ty84\\ \ \\type ty85\\ \ \\type ty86\\ \ \\type ty87\\ \ \\type ty88\\ \ \\type ty89\\ \ \\type ty90\\ \ \\type ty91\\ \ \\type ty92\\ \ \\type ty93\\ \ \\type ty94\\ \ \\type ty95\\ \ \\type ty96\\ \ \\type ty97\\ \ \\type ty98\\ \ \\type ty99\\ \ \\type ty100\\ \ \\type ty101\\ \ \\type ty102\\ \ \\type ty103\\ \ \\type ty104\\ \ \\type ty105\\ \ \\type ty106\\ \ \\type ty107\\ \ \\type ty108\\ \ \\type ty109\\ \ \\type ty110\\ \ \\type ty111\\ \ \\type ty112\\ \ \\type ty113\\ \ \\type ty114\\ \ \\type ty115\\ \ \\type ty116\\ \ \\type ty117\\ \ \\type ty118\\ \ \\type ty119\\ \ \\type ty120\\ \ \\type ty121\\ \ \\type ty122\\ \ \\type ty123\\ \ \\type ty124\\ \ \\type ty125\\ \ \\type ty126\\ \ \\type ty127\\ \ \\type ty128\\ \ \\type ty129\\ \ \\type ty130\\ \ \\type ty131\\ \ \\type ty132\\ \ \\type ty133\\ \ \\type ty134\\ \ \\type ty135\\ \ \\type ty136\\ \ \\type ty137\\ \ \\type ty138\\ \ \\type ty139\\ \ \\type ty140\\ \ \\type ty141\\ \ \\type ty142\\ \ \\type ty143\\ \ \\type ty144\\ \ \\type ty145\\ \ \\type ty146\\ \ \\type ty147\\ \ \\type ty148\\ \ \\type ty149\\ \ \\type ty150\\ \ \\type ty151\\ \ \\type ty152\\ \ \\type ty153\\ \ \\type ty154\\ \ \\type ty155\\ \ \\type ty156\\ \ \\type ty157\\ \ \\type ty158\\ \ \\type ty159\\ \ \\type ty160\\ \ \\type ty161\\ \ \\type ty162\\ \ \\type ty163\\ \ \\type ty164\\ \ \\type ty165\\ \ \\type ty166\\ \ \\type ty167\\ \ \\type ty168\\ \ \\type ty169\\ \ \\type ty170\\ \ \\type ty171\\ \ \\type ty172\\ \ \\type ty173\\ \ \\type ty174\\ \ \\type ty175\\ \ \\type ty176\\ \ \\type ty177\\ \ \\type ty178\\ \ \\type ty179\\ \ \\type ty180\\ \ \\type ty181\\ \ \\type ty182\\ \ \\type ty183\\ \ \\type ty184\\ \ \\type ty185\\ \ \\type ty186\\ \ \\type ty187\\ \ \\type ty188\\ \ \\type ty189\\ \ \\type ty190\\ \ \\type ty191\\ \ \\type ty192\\ \ \\type ty193\\ \ \\type ty194\\ \ \\type ty195\\ \ \\type ty196\\ \ \\type ty197\\ \ \\type ty198\\ \ \\type ty199\\ \ \\type ty200\\ \ \\type ty201\\ \ \\type ty202\\ \ \\type ty203\\ \ \\type ty204\\ \ \\type ty205\\ \ \\type ty206\\ \ \\type ty207\\ \ \\type ty208\\ \ \\type ty209\\ \ \\type ty210\\ \ \\type ty211\\ \ \\type ty212\\ \ \\type ty213\\ \ \\type ty214\\ \ \\type ty215\\ \ \\type ty216\\ \ \\type ty217\\ \ \\type ty218\\ \ \\type ty219\\ \ \\type ty220\\ \ \\type ty221\\ \ \\type ty222\\ \ \\type ty223\\ \ \\type ty224\\ \ \\type ty225\\ \ \\type ty226\\ \ \\type ty227\\ \ \\type ty228\\ \ \\type ty229\\ \ \\type ty230\\ \ \\type ty231\\ \ \\type ty232\\ \ \\type ty233\\ \ \\type ty234\\ \ \\type ty235\\ \ \\type ty236\\ \ \\type ty237\\ \ \\type ty238\\ \ \\type ty239\\ \ \\type ty240\\ \ \\type ty241\\ \ \\type ty242\\ \ \\type ty243\\ \ \\type ty244\\ \ \\type ty245\\ \ \\type ty246\\ \ \\type ty247\\ \ \\type ty248\\ \ \\type ty249\\ \ \\type ty250\\ \ \\type ty251\\ \ \\type ty252\\ \ \\type ty253\\ \ \\type ty254\\ \ \\type ty255\\ \ \\type ty256\\ \ \\type ty257\\ \ \\val word_length : forall 'a. mword 'a -> nat\\ \ \\****************************************************************\\ \ \\ Conversions \\ \ \\****************************************************************\\ \ \\val signedIntegerFromWord : forall 'a. mword 'a -> integer\\ \ \\val unsignedIntegerFromWord : forall 'a. mword 'a -> integer\\ \ \\ Version without typeclass constraint so that we can derive operations in Lem for one of the theorem provers without requiring it. \\ \ \\val proverWordFromInteger : forall 'a. integer -> mword 'a\\ \ \\val wordFromInteger : forall 'a. Size 'a => integer -> mword 'a\\ \ \\ The OCaml version is defined after the arithmetic operations, below. \\ \ \\val naturalFromWord : forall 'a. mword 'a -> natural\\ \ \\val wordFromNatural : forall 'a. Size 'a => natural -> mword 'a\\ \ \\val wordToHex : forall 'a. mword 'a -> string\\ \ \\ Building libraries fails if we don't provide implementations for the type class. \\ definition wordToHex :: "('a::len)Word.word \ string " where " wordToHex w = ( (''wordToHex not yet implemented''))" definition instance_Show_Show_Machine_word_mword_dict :: "(('a::len)Word.word)Show_class " where " instance_Show_Show_Machine_word_mword_dict = ((| show_method = wordToHex |) )" \ \\val wordFromBitlist : forall 'a. Size 'a => list bool -> mword 'a\\ \ \\val bitlistFromWord : forall 'a. mword 'a -> list bool\\ \ \\val size_test_fn : forall 'a. Size 'a => mword 'a -> nat\\ definition size_test_fn :: "('a::len)Word.word \ nat " where " size_test_fn _ = ( (len_of (TYPE(_) :: 'a itself)))" \ \\****************************************************************\\ \ \\ Comparisons \\ \ \\****************************************************************\\ \ \\val mwordEq : forall 'a. mword 'a -> mword 'a -> bool\\ \ \\val signedLess : forall 'a. mword 'a -> mword 'a -> bool\\ \ \\val signedLessEq : forall 'a. mword 'a -> mword 'a -> bool\\ \ \\val unsignedLess : forall 'a. mword 'a -> mword 'a -> bool\\ \ \\val unsignedLessEq : forall 'a. mword 'a -> mword 'a -> bool\\ \ \\ Comparison tests are below, after the definition of wordFromInteger \\ \ \\****************************************************************\\ \ \\ Appending, splitting and probing words \\ \ \\****************************************************************\\ \ \\val word_concat : forall 'a 'b 'c. mword 'a -> mword 'b -> mword 'c\\ \ \\ Note that we assume the result type has the correct size, especially for Isabelle. \\ \ \\val word_extract : forall 'a 'b. nat -> nat -> mword 'a -> mword 'b\\ \ \\ Needs to be in the prover because we'd end up with unknown sizes in the types in Lem. \\ \ \\val word_update : forall 'a 'b. mword 'a -> nat -> nat -> mword 'b -> mword 'a\\ \ \\val setBit : forall 'a. mword 'a -> nat -> bool -> mword 'a\\ \ \\val getBit : forall 'a. mword 'a -> nat -> bool\\ \ \\val msb : forall 'a. mword 'a -> bool\\ \ \\val lsb : forall 'a. mword 'a -> bool\\ \ \\****************************************************************\\ \ \\ Bitwise operations, shifts, etc. \\ \ \\****************************************************************\\ \ \\val shiftLeft : forall 'a. mword 'a -> nat -> mword 'a\\ \ \\val shiftRight : forall 'a. mword 'a -> nat -> mword 'a\\ \ \\val arithShiftRight : forall 'a. mword 'a -> nat -> mword 'a\\ \ \\val lAnd : forall 'a. mword 'a -> mword 'a -> mword 'a\\ \ \\val lOr : forall 'a. mword 'a -> mword 'a -> mword 'a\\ \ \\val lXor : forall 'a. mword 'a -> mword 'a -> mword 'a\\ \ \\val lNot : forall 'a. mword 'a -> mword 'a\\ \ \\val rotateRight : forall 'a. nat -> mword 'a -> mword 'a\\ \ \\val rotateLeft : forall 'a. nat -> mword 'a -> mword 'a\\ \ \\val zeroExtend : forall 'a 'b. Size 'b => mword 'a -> mword 'b\\ \ \\val signExtend : forall 'a 'b. Size 'b => mword 'a -> mword 'b\\ \ \\ Sign extension tests are below, after the definition of wordFromInteger \\ \ \\***************************************************************\\ \ \\ Arithmetic \\ \ \\***************************************************************\\ \ \\val plus : forall 'a. mword 'a -> mword 'a -> mword 'a\\ \ \\val minus : forall 'a. mword 'a -> mword 'a -> mword 'a\\ \ \\val uminus : forall 'a. mword 'a -> mword 'a\\ \ \\val times : forall 'a. mword 'a -> mword 'a -> mword 'a\\ \ \\val unsignedDivide : forall 'a. mword 'a -> mword 'a -> mword 'a\\ \ \\val signedDivide : forall 'a. mword 'a -> mword 'a -> mword 'a\\ definition signedDivide :: "('a::len)Word.word \('a::len)Word.word \('a::len)Word.word " where " signedDivide x y = ( if msb x then if msb y then (- x) div (- y) else - ((- x) div y) else if msb y then - (x div (- y)) else x div y )" \ \\val modulo : forall 'a. mword 'a -> mword 'a -> mword 'a\\ end diff --git a/thys/CakeML/generated/Lem_num.thy b/thys/CakeML/generated/Lem_num.thy --- a/thys/CakeML/generated/Lem_num.thy +++ b/thys/CakeML/generated/Lem_num.thy @@ -1,1313 +1,1313 @@ chapter \Generated by Lem from \num.lem\.\ theory "Lem_num" imports Main "Lem_bool" "Lem_basic_classes" - "HOL-Word.Word" + "HOL-Library.Word" "Complex_Main" begin \ \\open import Bool Basic_classes\\ -\ \\open import {isabelle} `HOL-Word.Word` `Complex_Main`\\ +\ \\open import {isabelle} `HOL-Library.Word` `Complex_Main`\\ \ \\open import {hol} `integerTheory` `intReduce` `wordsTheory` `wordsLib` `ratTheory` `realTheory` `intrealTheory` `transcTheory`\\ \ \\open import {coq} `Coq.Numbers.BinNums` `Coq.ZArith.BinInt` `Coq.ZArith.Zpower` `Coq.ZArith.Zdiv` `Coq.ZArith.Zmax` `Coq.Numbers.Natural.Peano.NPeano` `Coq.QArith.Qabs` `Coq.QArith.Qminmax` `Coq.QArith.Qround` `Coq.Reals.ROrderedType` `Coq.Reals.Rbase` `Coq.Reals.Rfunctions`\\ \ \\class inline ( Numeral 'a ) val fromNumeral : numeral -> 'a end\\ \ \\ ========================================================================== \\ \ \\ Syntactic type-classes for common operations \\ \ \\ ========================================================================== \\ \ \\ Typeclasses can be used as a mean to overload constants like "+", "-", etc \\ record 'a NumNegate_class= numNegate_method ::" 'a \ 'a " record 'a NumAbs_class= abs_method ::" 'a \ 'a " record 'a NumAdd_class= numAdd_method ::" 'a \ 'a \ 'a " record 'a NumMinus_class= numMinus_method ::" 'a \ 'a \ 'a " record 'a NumMult_class= numMult_method ::" 'a \ 'a \ 'a " record 'a NumPow_class= numPow_method ::" 'a \ nat \ 'a " record 'a NumDivision_class= numDivision_method ::" 'a \ 'a \ 'a " record 'a NumIntegerDivision_class= div_method ::" 'a \ 'a \ 'a " record 'a NumRemainder_class= mod_method ::" 'a \ 'a \ 'a " record 'a NumSucc_class= succ_method ::" 'a \ 'a " record 'a NumPred_class= pred_method ::" 'a \ 'a " \ \\ ----------------------- \\ \ \\ natural \\ \ \\ ----------------------- \\ \ \\ unbounded size natural numbers \\ \ \\type natural\\ \ \\ ----------------------- \\ \ \\ int \\ \ \\ ----------------------- \\ \ \\ bounded size integers with uncertain length \\ \ \\type int\\ \ \\ ----------------------- \\ \ \\ integer \\ \ \\ ----------------------- \\ \ \\ unbounded size integers \\ \ \\type integer\\ \ \\ ----------------------- \\ \ \\ bint \\ \ \\ ----------------------- \\ \ \\ TODO the bounded ints are only partially implemented, use with care. \\ \ \\ 32 bit integers \\ \ \\type int32\\ \ \\ 64 bit integers \\ \ \\type int64\\ \ \\ ----------------------- \\ \ \\ rational \\ \ \\ ----------------------- \\ \ \\ unbounded size and precision rational numbers \\ \ \\type rational\\ \ \\ ???: better type for this in HOL? \\ \ \\ ----------------------- \\ \ \\ real \\ \ \\ ----------------------- \\ \ \\ real numbers \\ \ \\ Note that for OCaml, this is mapped to floats with 64 bits. \\ \ \\type real\\ \ \\ ???: better type for this in HOL? \\ \ \\ ----------------------- \\ \ \\ double \\ \ \\ ----------------------- \\ \ \\ double precision floating point (64 bits) \\ \ \\type float64\\ \ \\ ???: better type for this in HOL? \\ \ \\type float32\\ \ \\ ???: better type for this in HOL? \\ \ \\ ========================================================================== \\ \ \\ Binding the standard operations for the number types \\ \ \\ ========================================================================== \\ \ \\ ----------------------- \\ \ \\ nat \\ \ \\ ----------------------- \\ \ \\val natFromNumeral : numeral -> nat\\ \ \\val natEq : nat -> nat -> bool\\ \ \\val natLess : nat -> nat -> bool\\ \ \\val natLessEqual : nat -> nat -> bool\\ \ \\val natGreater : nat -> nat -> bool\\ \ \\val natGreaterEqual : nat -> nat -> bool\\ \ \\val natCompare : nat -> nat -> ordering\\ definition instance_Basic_classes_Ord_nat_dict :: "(nat)Ord_class " where " instance_Basic_classes_Ord_nat_dict = ((| compare_method = (genericCompare (<) (=)), isLess_method = (<), isLessEqual_method = (\), isGreater_method = (>), isGreaterEqual_method = (\)|) )" \ \\val natAdd : nat -> nat -> nat\\ definition instance_Num_NumAdd_nat_dict :: "(nat)NumAdd_class " where " instance_Num_NumAdd_nat_dict = ((| numAdd_method = (+)|) )" \ \\val natMinus : nat -> nat -> nat\\ definition instance_Num_NumMinus_nat_dict :: "(nat)NumMinus_class " where " instance_Num_NumMinus_nat_dict = ((| numMinus_method = (-)|) )" \ \\val natSucc : nat -> nat\\ \ \\let natSucc n= (Instance_Num_NumAdd_nat.+) n 1\\ definition instance_Num_NumSucc_nat_dict :: "(nat)NumSucc_class " where " instance_Num_NumSucc_nat_dict = ((| succ_method = Suc |) )" \ \\val natPred : nat -> nat\\ definition instance_Num_NumPred_nat_dict :: "(nat)NumPred_class " where " instance_Num_NumPred_nat_dict = ((| pred_method = (\ n. n -( 1 :: nat))|) )" \ \\val natMult : nat -> nat -> nat\\ definition instance_Num_NumMult_nat_dict :: "(nat)NumMult_class " where " instance_Num_NumMult_nat_dict = ((| numMult_method = (*)|) )" \ \\val natDiv : nat -> nat -> nat\\ definition instance_Num_NumIntegerDivision_nat_dict :: "(nat)NumIntegerDivision_class " where " instance_Num_NumIntegerDivision_nat_dict = ((| div_method = (div)|) )" definition instance_Num_NumDivision_nat_dict :: "(nat)NumDivision_class " where " instance_Num_NumDivision_nat_dict = ((| numDivision_method = (div)|) )" \ \\val natMod : nat -> nat -> nat\\ definition instance_Num_NumRemainder_nat_dict :: "(nat)NumRemainder_class " where " instance_Num_NumRemainder_nat_dict = ((| mod_method = (mod)|) )" \ \\val gen_pow_aux : forall 'a. ('a -> 'a -> 'a) -> 'a -> 'a -> nat -> 'a\\ fun gen_pow_aux :: "('a \ 'a \ 'a)\ 'a \ 'a \ nat \ 'a " where " gen_pow_aux (mul :: 'a \ 'a \ 'a) (a :: 'a) (b :: 'a) (e :: nat) = ( (case e of 0 => a \ \\ cannot happen, call discipline guarentees e >= 1 \\ | (Suc 0) => mul a b | ( (Suc(Suc e'))) => (let e'' = (e div( 2 :: nat)) in (let a' = (if (e mod( 2 :: nat)) =( 0 :: nat) then a else mul a b) in gen_pow_aux mul a' (mul b b) e'')) ))" definition gen_pow :: " 'a \('a \ 'a \ 'a)\ 'a \ nat \ 'a " where " gen_pow (one :: 'a) (mul :: 'a \ 'a \ 'a) (b :: 'a) (e :: nat) = ( if e <( 0 :: nat) then one else if (e =( 0 :: nat)) then one else gen_pow_aux mul one b e )" \ \\val natPow : nat -> nat -> nat\\ definition instance_Num_NumPow_nat_dict :: "(nat)NumPow_class " where " instance_Num_NumPow_nat_dict = ((| numPow_method = (^)|) )" \ \\val natMin : nat -> nat -> nat\\ \ \\val natMax : nat -> nat -> nat\\ definition instance_Basic_classes_OrdMaxMin_nat_dict :: "(nat)OrdMaxMin_class " where " instance_Basic_classes_OrdMaxMin_nat_dict = ((| max_method = max, min_method = min |) )" \ \\ ----------------------- \\ \ \\ natural \\ \ \\ ----------------------- \\ \ \\val naturalFromNumeral : numeral -> natural\\ \ \\val naturalEq : natural -> natural -> bool\\ \ \\val naturalLess : natural -> natural -> bool\\ \ \\val naturalLessEqual : natural -> natural -> bool\\ \ \\val naturalGreater : natural -> natural -> bool\\ \ \\val naturalGreaterEqual : natural -> natural -> bool\\ \ \\val naturalCompare : natural -> natural -> ordering\\ definition instance_Basic_classes_Ord_Num_natural_dict :: "(nat)Ord_class " where " instance_Basic_classes_Ord_Num_natural_dict = ((| compare_method = (genericCompare (<) (=)), isLess_method = (<), isLessEqual_method = (\), isGreater_method = (>), isGreaterEqual_method = (\)|) )" \ \\val naturalAdd : natural -> natural -> natural\\ definition instance_Num_NumAdd_Num_natural_dict :: "(nat)NumAdd_class " where " instance_Num_NumAdd_Num_natural_dict = ((| numAdd_method = (+)|) )" \ \\val naturalMinus : natural -> natural -> natural\\ definition instance_Num_NumMinus_Num_natural_dict :: "(nat)NumMinus_class " where " instance_Num_NumMinus_Num_natural_dict = ((| numMinus_method = (-)|) )" \ \\val naturalSucc : natural -> natural\\ \ \\let naturalSucc n= (Instance_Num_NumAdd_Num_natural.+) n 1\\ definition instance_Num_NumSucc_Num_natural_dict :: "(nat)NumSucc_class " where " instance_Num_NumSucc_Num_natural_dict = ((| succ_method = Suc |) )" \ \\val naturalPred : natural -> natural\\ definition instance_Num_NumPred_Num_natural_dict :: "(nat)NumPred_class " where " instance_Num_NumPred_Num_natural_dict = ((| pred_method = (\ n. n -( 1 :: nat))|) )" \ \\val naturalMult : natural -> natural -> natural\\ definition instance_Num_NumMult_Num_natural_dict :: "(nat)NumMult_class " where " instance_Num_NumMult_Num_natural_dict = ((| numMult_method = (*)|) )" \ \\val naturalPow : natural -> nat -> natural\\ definition instance_Num_NumPow_Num_natural_dict :: "(nat)NumPow_class " where " instance_Num_NumPow_Num_natural_dict = ((| numPow_method = (^)|) )" \ \\val naturalDiv : natural -> natural -> natural\\ definition instance_Num_NumIntegerDivision_Num_natural_dict :: "(nat)NumIntegerDivision_class " where " instance_Num_NumIntegerDivision_Num_natural_dict = ((| div_method = (div)|) )" definition instance_Num_NumDivision_Num_natural_dict :: "(nat)NumDivision_class " where " instance_Num_NumDivision_Num_natural_dict = ((| numDivision_method = (div)|) )" \ \\val naturalMod : natural -> natural -> natural\\ definition instance_Num_NumRemainder_Num_natural_dict :: "(nat)NumRemainder_class " where " instance_Num_NumRemainder_Num_natural_dict = ((| mod_method = (mod)|) )" \ \\val naturalMin : natural -> natural -> natural\\ \ \\val naturalMax : natural -> natural -> natural\\ definition instance_Basic_classes_OrdMaxMin_Num_natural_dict :: "(nat)OrdMaxMin_class " where " instance_Basic_classes_OrdMaxMin_Num_natural_dict = ((| max_method = max, min_method = min |) )" \ \\ ----------------------- \\ \ \\ int \\ \ \\ ----------------------- \\ \ \\val intFromNumeral : numeral -> int\\ \ \\val intEq : int -> int -> bool\\ \ \\val intLess : int -> int -> bool\\ \ \\val intLessEqual : int -> int -> bool\\ \ \\val intGreater : int -> int -> bool\\ \ \\val intGreaterEqual : int -> int -> bool\\ \ \\val intCompare : int -> int -> ordering\\ definition instance_Basic_classes_Ord_Num_int_dict :: "(int)Ord_class " where " instance_Basic_classes_Ord_Num_int_dict = ((| compare_method = (genericCompare (<) (=)), isLess_method = (<), isLessEqual_method = (\), isGreater_method = (>), isGreaterEqual_method = (\)|) )" \ \\val intNegate : int -> int\\ definition instance_Num_NumNegate_Num_int_dict :: "(int)NumNegate_class " where " instance_Num_NumNegate_Num_int_dict = ((| numNegate_method = (\ i. - i)|) )" \ \\val intAbs : int -> int\\ \ \\ TODO: check \\ definition instance_Num_NumAbs_Num_int_dict :: "(int)NumAbs_class " where " instance_Num_NumAbs_Num_int_dict = ((| abs_method = abs |) )" \ \\val intAdd : int -> int -> int\\ definition instance_Num_NumAdd_Num_int_dict :: "(int)NumAdd_class " where " instance_Num_NumAdd_Num_int_dict = ((| numAdd_method = (+)|) )" \ \\val intMinus : int -> int -> int\\ definition instance_Num_NumMinus_Num_int_dict :: "(int)NumMinus_class " where " instance_Num_NumMinus_Num_int_dict = ((| numMinus_method = (-)|) )" \ \\val intSucc : int -> int\\ definition instance_Num_NumSucc_Num_int_dict :: "(int)NumSucc_class " where " instance_Num_NumSucc_Num_int_dict = ((| succ_method = (\ n. n +( 1 :: int))|) )" \ \\val intPred : int -> int\\ definition instance_Num_NumPred_Num_int_dict :: "(int)NumPred_class " where " instance_Num_NumPred_Num_int_dict = ((| pred_method = (\ n. n -( 1 :: int))|) )" \ \\val intMult : int -> int -> int\\ definition instance_Num_NumMult_Num_int_dict :: "(int)NumMult_class " where " instance_Num_NumMult_Num_int_dict = ((| numMult_method = (*)|) )" \ \\val intPow : int -> nat -> int\\ definition instance_Num_NumPow_Num_int_dict :: "(int)NumPow_class " where " instance_Num_NumPow_Num_int_dict = ((| numPow_method = (^)|) )" \ \\val intDiv : int -> int -> int\\ definition instance_Num_NumIntegerDivision_Num_int_dict :: "(int)NumIntegerDivision_class " where " instance_Num_NumIntegerDivision_Num_int_dict = ((| div_method = (div)|) )" definition instance_Num_NumDivision_Num_int_dict :: "(int)NumDivision_class " where " instance_Num_NumDivision_Num_int_dict = ((| numDivision_method = (div)|) )" \ \\val intMod : int -> int -> int\\ definition instance_Num_NumRemainder_Num_int_dict :: "(int)NumRemainder_class " where " instance_Num_NumRemainder_Num_int_dict = ((| mod_method = (mod)|) )" \ \\val intMin : int -> int -> int\\ \ \\val intMax : int -> int -> int\\ definition instance_Basic_classes_OrdMaxMin_Num_int_dict :: "(int)OrdMaxMin_class " where " instance_Basic_classes_OrdMaxMin_Num_int_dict = ((| max_method = max, min_method = min |) )" \ \\ ----------------------- \\ \ \\ int32 \\ \ \\ ----------------------- \\ \ \\val int32FromNumeral : numeral -> int32\\ \ \\val int32Eq : int32 -> int32 -> bool\\ \ \\val int32Less : int32 -> int32 -> bool\\ \ \\val int32LessEqual : int32 -> int32 -> bool\\ \ \\val int32Greater : int32 -> int32 -> bool\\ \ \\val int32GreaterEqual : int32 -> int32 -> bool\\ \ \\val int32Compare : int32 -> int32 -> ordering\\ definition instance_Basic_classes_Ord_Num_int32_dict :: "( 32 word)Ord_class " where " instance_Basic_classes_Ord_Num_int32_dict = ((| compare_method = (genericCompare word_sless (=)), isLess_method = word_sless, isLessEqual_method = word_sle, isGreater_method = (\ x y. word_sless y x), isGreaterEqual_method = (\ x y. word_sle y x)|) )" \ \\val int32Negate : int32 -> int32\\ definition instance_Num_NumNegate_Num_int32_dict :: "( 32 word)NumNegate_class " where " instance_Num_NumNegate_Num_int32_dict = ((| numNegate_method = (\ i. - i)|) )" \ \\val int32Abs : int32 -> int32\\ definition int32Abs :: " 32 word \ 32 word " where " int32Abs i = ( (if word_sle(((word_of_int 0) :: 32 word)) i then i else - i))" definition instance_Num_NumAbs_Num_int32_dict :: "( 32 word)NumAbs_class " where " instance_Num_NumAbs_Num_int32_dict = ((| abs_method = int32Abs |) )" \ \\val int32Add : int32 -> int32 -> int32\\ definition instance_Num_NumAdd_Num_int32_dict :: "( 32 word)NumAdd_class " where " instance_Num_NumAdd_Num_int32_dict = ((| numAdd_method = (+)|) )" \ \\val int32Minus : int32 -> int32 -> int32\\ definition instance_Num_NumMinus_Num_int32_dict :: "( 32 word)NumMinus_class " where " instance_Num_NumMinus_Num_int32_dict = ((| numMinus_method = (-)|) )" \ \\val int32Succ : int32 -> int32\\ definition instance_Num_NumSucc_Num_int32_dict :: "( 32 word)NumSucc_class " where " instance_Num_NumSucc_Num_int32_dict = ((| succ_method = (\ n. n +((word_of_int 1) :: 32 word))|) )" \ \\val int32Pred : int32 -> int32\\ definition instance_Num_NumPred_Num_int32_dict :: "( 32 word)NumPred_class " where " instance_Num_NumPred_Num_int32_dict = ((| pred_method = (\ n. n -((word_of_int 1) :: 32 word))|) )" \ \\val int32Mult : int32 -> int32 -> int32\\ definition instance_Num_NumMult_Num_int32_dict :: "( 32 word)NumMult_class " where " instance_Num_NumMult_Num_int32_dict = ((| numMult_method = (*)|) )" \ \\val int32Pow : int32 -> nat -> int32\\ definition instance_Num_NumPow_Num_int32_dict :: "( 32 word)NumPow_class " where " instance_Num_NumPow_Num_int32_dict = ((| numPow_method = (^)|) )" \ \\val int32Div : int32 -> int32 -> int32\\ definition instance_Num_NumIntegerDivision_Num_int32_dict :: "( 32 word)NumIntegerDivision_class " where " instance_Num_NumIntegerDivision_Num_int32_dict = ((| div_method = (div)|) )" definition instance_Num_NumDivision_Num_int32_dict :: "( 32 word)NumDivision_class " where " instance_Num_NumDivision_Num_int32_dict = ((| numDivision_method = (div)|) )" \ \\val int32Mod : int32 -> int32 -> int32\\ definition instance_Num_NumRemainder_Num_int32_dict :: "( 32 word)NumRemainder_class " where " instance_Num_NumRemainder_Num_int32_dict = ((| mod_method = (mod)|) )" \ \\val int32Min : int32 -> int32 -> int32\\ \ \\val int32Max : int32 -> int32 -> int32\\ definition instance_Basic_classes_OrdMaxMin_Num_int32_dict :: "( 32 word)OrdMaxMin_class " where " instance_Basic_classes_OrdMaxMin_Num_int32_dict = ((| max_method = ((\ x y. if (word_sle y x) then x else y)), min_method = ((\ x y. if (word_sle x y) then x else y))|) )" \ \\ ----------------------- \\ \ \\ int64 \\ \ \\ ----------------------- \\ \ \\val int64FromNumeral : numeral -> int64\\ \ \\val int64Eq : int64 -> int64 -> bool\\ \ \\val int64Less : int64 -> int64 -> bool\\ \ \\val int64LessEqual : int64 -> int64 -> bool\\ \ \\val int64Greater : int64 -> int64 -> bool\\ \ \\val int64GreaterEqual : int64 -> int64 -> bool\\ \ \\val int64Compare : int64 -> int64 -> ordering\\ definition instance_Basic_classes_Ord_Num_int64_dict :: "( 64 word)Ord_class " where " instance_Basic_classes_Ord_Num_int64_dict = ((| compare_method = (genericCompare word_sless (=)), isLess_method = word_sless, isLessEqual_method = word_sle, isGreater_method = (\ x y. word_sless y x), isGreaterEqual_method = (\ x y. word_sle y x)|) )" \ \\val int64Negate : int64 -> int64\\ definition instance_Num_NumNegate_Num_int64_dict :: "( 64 word)NumNegate_class " where " instance_Num_NumNegate_Num_int64_dict = ((| numNegate_method = (\ i. - i)|) )" \ \\val int64Abs : int64 -> int64\\ definition int64Abs :: " 64 word \ 64 word " where " int64Abs i = ( (if word_sle(((word_of_int 0) :: 64 word)) i then i else - i))" definition instance_Num_NumAbs_Num_int64_dict :: "( 64 word)NumAbs_class " where " instance_Num_NumAbs_Num_int64_dict = ((| abs_method = int64Abs |) )" \ \\val int64Add : int64 -> int64 -> int64\\ definition instance_Num_NumAdd_Num_int64_dict :: "( 64 word)NumAdd_class " where " instance_Num_NumAdd_Num_int64_dict = ((| numAdd_method = (+)|) )" \ \\val int64Minus : int64 -> int64 -> int64\\ definition instance_Num_NumMinus_Num_int64_dict :: "( 64 word)NumMinus_class " where " instance_Num_NumMinus_Num_int64_dict = ((| numMinus_method = (-)|) )" \ \\val int64Succ : int64 -> int64\\ definition instance_Num_NumSucc_Num_int64_dict :: "( 64 word)NumSucc_class " where " instance_Num_NumSucc_Num_int64_dict = ((| succ_method = (\ n. n +((word_of_int 1) :: 64 word))|) )" \ \\val int64Pred : int64 -> int64\\ definition instance_Num_NumPred_Num_int64_dict :: "( 64 word)NumPred_class " where " instance_Num_NumPred_Num_int64_dict = ((| pred_method = (\ n. n -((word_of_int 1) :: 64 word))|) )" \ \\val int64Mult : int64 -> int64 -> int64\\ definition instance_Num_NumMult_Num_int64_dict :: "( 64 word)NumMult_class " where " instance_Num_NumMult_Num_int64_dict = ((| numMult_method = (*)|) )" \ \\val int64Pow : int64 -> nat -> int64\\ definition instance_Num_NumPow_Num_int64_dict :: "( 64 word)NumPow_class " where " instance_Num_NumPow_Num_int64_dict = ((| numPow_method = (^)|) )" \ \\val int64Div : int64 -> int64 -> int64\\ definition instance_Num_NumIntegerDivision_Num_int64_dict :: "( 64 word)NumIntegerDivision_class " where " instance_Num_NumIntegerDivision_Num_int64_dict = ((| div_method = (div)|) )" definition instance_Num_NumDivision_Num_int64_dict :: "( 64 word)NumDivision_class " where " instance_Num_NumDivision_Num_int64_dict = ((| numDivision_method = (div)|) )" \ \\val int64Mod : int64 -> int64 -> int64\\ definition instance_Num_NumRemainder_Num_int64_dict :: "( 64 word)NumRemainder_class " where " instance_Num_NumRemainder_Num_int64_dict = ((| mod_method = (mod)|) )" \ \\val int64Min : int64 -> int64 -> int64\\ \ \\val int64Max : int64 -> int64 -> int64\\ definition instance_Basic_classes_OrdMaxMin_Num_int64_dict :: "( 64 word)OrdMaxMin_class " where " instance_Basic_classes_OrdMaxMin_Num_int64_dict = ((| max_method = ((\ x y. if (word_sle y x) then x else y)), min_method = ((\ x y. if (word_sle x y) then x else y))|) )" \ \\ ----------------------- \\ \ \\ integer \\ \ \\ ----------------------- \\ \ \\val integerFromNumeral : numeral -> integer\\ \ \\val integerFromNat : nat -> integer\\ \ \\ TODO: check \\ \ \\val integerEq : integer -> integer -> bool\\ \ \\val integerLess : integer -> integer -> bool\\ \ \\val integerLessEqual : integer -> integer -> bool\\ \ \\val integerGreater : integer -> integer -> bool\\ \ \\val integerGreaterEqual : integer -> integer -> bool\\ \ \\val integerCompare : integer -> integer -> ordering\\ definition instance_Basic_classes_Ord_Num_integer_dict :: "(int)Ord_class " where " instance_Basic_classes_Ord_Num_integer_dict = ((| compare_method = (genericCompare (<) (=)), isLess_method = (<), isLessEqual_method = (\), isGreater_method = (>), isGreaterEqual_method = (\)|) )" \ \\val integerNegate : integer -> integer\\ definition instance_Num_NumNegate_Num_integer_dict :: "(int)NumNegate_class " where " instance_Num_NumNegate_Num_integer_dict = ((| numNegate_method = (\ i. - i)|) )" \ \\val integerAbs : integer -> integer\\ \ \\ TODO: check \\ definition instance_Num_NumAbs_Num_integer_dict :: "(int)NumAbs_class " where " instance_Num_NumAbs_Num_integer_dict = ((| abs_method = abs |) )" \ \\val integerAdd : integer -> integer -> integer\\ definition instance_Num_NumAdd_Num_integer_dict :: "(int)NumAdd_class " where " instance_Num_NumAdd_Num_integer_dict = ((| numAdd_method = (+)|) )" \ \\val integerMinus : integer -> integer -> integer\\ definition instance_Num_NumMinus_Num_integer_dict :: "(int)NumMinus_class " where " instance_Num_NumMinus_Num_integer_dict = ((| numMinus_method = (-)|) )" \ \\val integerSucc : integer -> integer\\ definition instance_Num_NumSucc_Num_integer_dict :: "(int)NumSucc_class " where " instance_Num_NumSucc_Num_integer_dict = ((| succ_method = (\ n. n +( 1 :: int))|) )" \ \\val integerPred : integer -> integer\\ definition instance_Num_NumPred_Num_integer_dict :: "(int)NumPred_class " where " instance_Num_NumPred_Num_integer_dict = ((| pred_method = (\ n. n -( 1 :: int))|) )" \ \\val integerMult : integer -> integer -> integer\\ definition instance_Num_NumMult_Num_integer_dict :: "(int)NumMult_class " where " instance_Num_NumMult_Num_integer_dict = ((| numMult_method = (*)|) )" \ \\val integerPow : integer -> nat -> integer\\ definition instance_Num_NumPow_Num_integer_dict :: "(int)NumPow_class " where " instance_Num_NumPow_Num_integer_dict = ((| numPow_method = (^)|) )" \ \\val integerDiv : integer -> integer -> integer\\ definition instance_Num_NumIntegerDivision_Num_integer_dict :: "(int)NumIntegerDivision_class " where " instance_Num_NumIntegerDivision_Num_integer_dict = ((| div_method = (div)|) )" definition instance_Num_NumDivision_Num_integer_dict :: "(int)NumDivision_class " where " instance_Num_NumDivision_Num_integer_dict = ((| numDivision_method = (div)|) )" \ \\val integerMod : integer -> integer -> integer\\ definition instance_Num_NumRemainder_Num_integer_dict :: "(int)NumRemainder_class " where " instance_Num_NumRemainder_Num_integer_dict = ((| mod_method = (mod)|) )" \ \\val integerMin : integer -> integer -> integer\\ \ \\val integerMax : integer -> integer -> integer\\ definition instance_Basic_classes_OrdMaxMin_Num_integer_dict :: "(int)OrdMaxMin_class " where " instance_Basic_classes_OrdMaxMin_Num_integer_dict = ((| max_method = max, min_method = min |) )" \ \\ ----------------------- \\ \ \\ rational \\ \ \\ ----------------------- \\ \ \\val rationalFromNumeral : numeral -> rational\\ \ \\val rationalFromInt : int -> rational\\ \ \\val rationalFromInteger : integer -> rational\\ \ \\val rationalEq : rational -> rational -> bool\\ \ \\val rationalLess : rational -> rational -> bool\\ \ \\val rationalLessEqual : rational -> rational -> bool\\ \ \\val rationalGreater : rational -> rational -> bool\\ \ \\val rationalGreaterEqual : rational -> rational -> bool\\ \ \\val rationalCompare : rational -> rational -> ordering\\ definition instance_Basic_classes_Ord_Num_rational_dict :: "(rat)Ord_class " where " instance_Basic_classes_Ord_Num_rational_dict = ((| compare_method = (genericCompare (<) (=)), isLess_method = (<), isLessEqual_method = (\), isGreater_method = (>), isGreaterEqual_method = (\)|) )" \ \\val rationalAdd : rational -> rational -> rational\\ definition instance_Num_NumAdd_Num_rational_dict :: "(rat)NumAdd_class " where " instance_Num_NumAdd_Num_rational_dict = ((| numAdd_method = (+)|) )" \ \\val rationalMinus : rational -> rational -> rational\\ definition instance_Num_NumMinus_Num_rational_dict :: "(rat)NumMinus_class " where " instance_Num_NumMinus_Num_rational_dict = ((| numMinus_method = (-)|) )" \ \\val rationalNegate : rational -> rational\\ definition instance_Num_NumNegate_Num_rational_dict :: "(rat)NumNegate_class " where " instance_Num_NumNegate_Num_rational_dict = ((| numNegate_method = (\ i. - i)|) )" \ \\val rationalAbs : rational -> rational\\ definition instance_Num_NumAbs_Num_rational_dict :: "(rat)NumAbs_class " where " instance_Num_NumAbs_Num_rational_dict = ((| abs_method = abs |) )" \ \\val rationalSucc : rational -> rational\\ definition instance_Num_NumSucc_Num_rational_dict :: "(rat)NumSucc_class " where " instance_Num_NumSucc_Num_rational_dict = ((| succ_method = (\ n. n +(Fract ( 1 :: int) (1 :: int)))|) )" \ \\val rationalPred : rational -> rational\\ definition instance_Num_NumPred_Num_rational_dict :: "(rat)NumPred_class " where " instance_Num_NumPred_Num_rational_dict = ((| pred_method = (\ n. n -(Fract ( 1 :: int) (1 :: int)))|) )" \ \\val rationalMult : rational -> rational -> rational\\ definition instance_Num_NumMult_Num_rational_dict :: "(rat)NumMult_class " where " instance_Num_NumMult_Num_rational_dict = ((| numMult_method = (*)|) )" \ \\val rationalDiv : rational -> rational -> rational\\ definition instance_Num_NumDivision_Num_rational_dict :: "(rat)NumDivision_class " where " instance_Num_NumDivision_Num_rational_dict = ((| numDivision_method = (div)|) )" \ \\val rationalFromFrac : int -> int -> rational\\ \ \\let rationalFromFrac n d= (Instance_Num_NumDivision_Num_rational./) (rationalFromInt n) (rationalFromInt d)\\ \ \\val rationalNumerator : rational -> integer\\ \ \\ TODO: test \\ \ \\val rationalDenominator : rational -> integer\\ \ \\ TODO: test \\ \ \\val rationalPowInteger : rational -> integer -> rational\\ fun rationalPowInteger :: " rat \ int \ rat " where " rationalPowInteger b e = ( if e =( 0 :: int) then(Fract ( 1 :: int) (1 :: int)) else if e >( 0 :: int) then rationalPowInteger b (e -( 1 :: int)) * b else rationalPowInteger b (e +( 1 :: int)) div b )" \ \\val rationalPowNat : rational -> nat -> rational\\ \ \\let rationalPowNat r e= rationalPowInteger r (integerFromNat e)\\ definition instance_Num_NumPow_Num_rational_dict :: "(rat)NumPow_class " where " instance_Num_NumPow_Num_rational_dict = ((| numPow_method = power |) )" \ \\val rationalMin : rational -> rational -> rational\\ \ \\val rationalMax : rational -> rational -> rational\\ definition instance_Basic_classes_OrdMaxMin_Num_rational_dict :: "(rat)OrdMaxMin_class " where " instance_Basic_classes_OrdMaxMin_Num_rational_dict = ((| max_method = max, min_method = min |) )" \ \\ ----------------------- \\ \ \\ real \\ \ \\ ----------------------- \\ \ \\val realFromNumeral : numeral -> real\\ \ \\val realFromInteger : integer -> real\\ \ \\val realEq : real -> real -> bool\\ \ \\val realLess : real -> real -> bool\\ \ \\val realLessEqual : real -> real -> bool\\ \ \\val realGreater : real -> real -> bool\\ \ \\val realGreaterEqual : real -> real -> bool\\ \ \\val realCompare : real -> real -> ordering\\ definition instance_Basic_classes_Ord_Num_real_dict :: "(real)Ord_class " where " instance_Basic_classes_Ord_Num_real_dict = ((| compare_method = (genericCompare (<) (=)), isLess_method = (<), isLessEqual_method = (\), isGreater_method = (>), isGreaterEqual_method = (\)|) )" \ \\val realAdd : real -> real -> real\\ definition instance_Num_NumAdd_Num_real_dict :: "(real)NumAdd_class " where " instance_Num_NumAdd_Num_real_dict = ((| numAdd_method = (+)|) )" \ \\val realMinus : real -> real -> real\\ definition instance_Num_NumMinus_Num_real_dict :: "(real)NumMinus_class " where " instance_Num_NumMinus_Num_real_dict = ((| numMinus_method = (-)|) )" \ \\val realNegate : real -> real\\ definition instance_Num_NumNegate_Num_real_dict :: "(real)NumNegate_class " where " instance_Num_NumNegate_Num_real_dict = ((| numNegate_method = (\ i. - i)|) )" \ \\val realAbs : real -> real\\ definition instance_Num_NumAbs_Num_real_dict :: "(real)NumAbs_class " where " instance_Num_NumAbs_Num_real_dict = ((| abs_method = abs |) )" \ \\val realSucc : real -> real\\ definition instance_Num_NumSucc_Num_real_dict :: "(real)NumSucc_class " where " instance_Num_NumSucc_Num_real_dict = ((| succ_method = (\ n. n +( 1 :: real))|) )" \ \\val realPred : real -> real\\ definition instance_Num_NumPred_Num_real_dict :: "(real)NumPred_class " where " instance_Num_NumPred_Num_real_dict = ((| pred_method = (\ n. n -( 1 :: real))|) )" \ \\val realMult : real -> real -> real\\ definition instance_Num_NumMult_Num_real_dict :: "(real)NumMult_class " where " instance_Num_NumMult_Num_real_dict = ((| numMult_method = (*)|) )" \ \\val realDiv : real -> real -> real\\ definition instance_Num_NumDivision_Num_real_dict :: "(real)NumDivision_class " where " instance_Num_NumDivision_Num_real_dict = ((| numDivision_method = (div)|) )" \ \\val realFromFrac : integer -> integer -> real\\ definition realFromFrac :: " int \ int \ real " where " realFromFrac n d = ( ((real_of_int n)) div ((real_of_int d)))" \ \\val realPowInteger : real -> integer -> real\\ fun realPowInteger :: " real \ int \ real " where " realPowInteger b e = ( if e =( 0 :: int) then( 1 :: real) else if e >( 0 :: int) then realPowInteger b (e -( 1 :: int)) * b else realPowInteger b (e +( 1 :: int)) div b )" \ \\val realPowNat : real -> nat -> real\\ \ \\let realPowNat r e= realPowInteger r (integerFromNat e)\\ definition instance_Num_NumPow_Num_real_dict :: "(real)NumPow_class " where " instance_Num_NumPow_Num_real_dict = ((| numPow_method = power |) )" \ \\val realSqrt : real -> real\\ \ \\val realMin : real -> real -> real\\ \ \\val realMax : real -> real -> real\\ definition instance_Basic_classes_OrdMaxMin_Num_real_dict :: "(real)OrdMaxMin_class " where " instance_Basic_classes_OrdMaxMin_Num_real_dict = ((| max_method = max, min_method = min |) )" \ \\val realCeiling : real -> integer\\ \ \\val realFloor : real -> integer\\ \ \\val integerSqrt : integer -> integer\\ definition integerSqrt :: " int \ int " where " integerSqrt i = ( floor (sqrt ((real_of_int i))))" \ \\ ========================================================================== \\ \ \\ Translation between number types \\ \ \\ ========================================================================== \\ \ \\****************\\ \ \\ integerFrom... \\ \ \\****************\\ \ \\val integerFromInt : int -> integer\\ \ \\val integerFromNatural : natural -> integer\\ \ \\val integerFromInt32 : int32 -> integer\\ \ \\val integerFromInt64 : int64 -> integer\\ \ \\****************\\ \ \\ naturalFrom... \\ \ \\****************\\ \ \\val naturalFromNat : nat -> natural\\ \ \\val naturalFromInteger : integer -> natural\\ \ \\****************\\ \ \\ intFrom ... \\ \ \\****************\\ \ \\val intFromInteger : integer -> int\\ \ \\val intFromNat : nat -> int\\ \ \\****************\\ \ \\ natFrom ... \\ \ \\****************\\ \ \\val natFromNatural : natural -> nat\\ \ \\val natFromInt : int -> nat\\ \ \\****************\\ \ \\ int32From ... \\ \ \\****************\\ \ \\val int32FromNat : nat -> int32\\ \ \\val int32FromNatural : natural -> int32\\ \ \\val int32FromInteger : integer -> int32\\ \ \\let int32FromInteger i= ( let abs_int32 = int32FromNatural (naturalFromInteger i) in if ((Instance_Basic_classes_Ord_Num_integer.<) i 0) then (Instance_Num_NumNegate_Num_int32.~ abs_int32) else abs_int32 )\\ \ \\val int32FromInt : int -> int32\\ \ \\let int32FromInt i= int32FromInteger (integerFromInt i)\\ \ \\val int32FromInt64 : int64 -> int32\\ \ \\let int32FromInt64 i= int32FromInteger (integerFromInt64 i)\\ \ \\****************\\ \ \\ int64From ... \\ \ \\****************\\ \ \\val int64FromNat : nat -> int64\\ \ \\val int64FromNatural : natural -> int64\\ \ \\val int64FromInteger : integer -> int64\\ \ \\let int64FromInteger i= ( let abs_int64 = int64FromNatural (naturalFromInteger i) in if ((Instance_Basic_classes_Ord_Num_integer.<) i 0) then (Instance_Num_NumNegate_Num_int64.~ abs_int64) else abs_int64 )\\ \ \\val int64FromInt : int -> int64\\ \ \\let int64FromInt i= int64FromInteger (integerFromInt i)\\ \ \\val int64FromInt32 : int32 -> int64\\ \ \\let int64FromInt32 i= int64FromInteger (integerFromInt32 i)\\ \ \\****************\\ \ \\ what's missing \\ \ \\****************\\ \ \\val naturalFromInt : int -> natural\\ \ \\val naturalFromInt32 : int32 -> natural\\ \ \\val naturalFromInt64 : int64 -> natural\\ \ \\val intFromNatural : natural -> int\\ \ \\val intFromInt32 : int32 -> int\\ \ \\val intFromInt64 : int64 -> int\\ \ \\val natFromInteger : integer -> nat\\ \ \\val natFromInt32 : int32 -> nat\\ \ \\val natFromInt64 : int64 -> nat\\ end diff --git a/thys/CakeML/generated/Lem_word.thy b/thys/CakeML/generated/Lem_word.thy --- a/thys/CakeML/generated/Lem_word.thy +++ b/thys/CakeML/generated/Lem_word.thy @@ -1,1024 +1,1024 @@ chapter \Generated by Lem from \word.lem\.\ theory "Lem_word" imports Main "Lem_bool" "Lem_maybe" "Lem_num" "Lem_basic_classes" "Lem_list" - "HOL-Word.Word" + "HOL-Library.Word" begin \ \\open import Bool Maybe Num Basic_classes List\\ -\ \\open import {isabelle} `HOL-Word.Word`\\ +\ \\open import {isabelle} `HOL-Library.Word`\\ \ \\open import {hol} `wordsTheory` `wordsLib`\\ \ \\ ========================================================================== \\ \ \\ Define general purpose word, i.e. sequences of bits of arbitrary length \\ \ \\ ========================================================================== \\ datatype bitSequence = BitSeq " nat option " " \ \\ length of the sequence, Nothing means infinite length \\ bool " " bool \ \\ sign of the word, used to fill up after concrete value is exhausted \\ list " \ \\ the initial part of the sequence, least significant bit first \\ \ \\val bitSeqEq : bitSequence -> bitSequence -> bool\\ \ \\val boolListFrombitSeq : nat -> bitSequence -> list bool\\ fun boolListFrombitSeqAux :: " nat \ 'a \ 'a list \ 'a list " where " boolListFrombitSeqAux n s bl = ( if n =( 0 :: nat) then [] else (case bl of [] => List.replicate n s | b # bl' => b # (boolListFrombitSeqAux (n-( 1 :: nat)) s bl') ))" fun boolListFrombitSeq :: " nat \ bitSequence \(bool)list " where " boolListFrombitSeq n (BitSeq _ s bl) = ( boolListFrombitSeqAux n s bl )" \ \\val bitSeqFromBoolList : list bool -> maybe bitSequence\\ definition bitSeqFromBoolList :: "(bool)list \(bitSequence)option " where " bitSeqFromBoolList bl = ( (case dest_init bl of None => None | Some (bl', s) => Some (BitSeq (Some (List.length bl)) s bl') ))" \ \\ cleans up the representation of a bitSequence without changing its semantics \\ \ \\val cleanBitSeq : bitSequence -> bitSequence\\ fun cleanBitSeq :: " bitSequence \ bitSequence " where " cleanBitSeq (BitSeq len s bl) = ( (case len of None => (BitSeq len s (List.rev (dropWhile ((\) s) (List.rev bl)))) | Some n => (BitSeq len s (List.rev (dropWhile ((\) s) (List.rev (List.take (n-( 1 :: nat)) bl))))) ))" \ \\val bitSeqTestBit : bitSequence -> nat -> maybe bool\\ fun bitSeqTestBit :: " bitSequence \ nat \(bool)option " where " bitSeqTestBit (BitSeq None s bl) pos = ( if pos < List.length bl then index bl pos else Some s )" |" bitSeqTestBit (BitSeq(Some l) s bl) pos = ( if (pos \ l) then None else if ((pos = (l -( 1 :: nat))) \ (pos \ List.length bl)) then Some s else index bl pos )" \ \\val bitSeqSetBit : bitSequence -> nat -> bool -> bitSequence\\ fun bitSeqSetBit :: " bitSequence \ nat \ bool \ bitSequence " where " bitSeqSetBit (BitSeq len s bl) pos v = ( (let bl' = (if (pos < List.length bl) then bl else bl @ List.replicate pos s) in (let bl'' = (List.list_update bl' pos v) in (let bs' = (BitSeq len s bl'') in cleanBitSeq bs'))))" \ \\val resizeBitSeq : maybe nat -> bitSequence -> bitSequence\\ definition resizeBitSeq :: "(nat)option \ bitSequence \ bitSequence " where " resizeBitSeq new_len bs = ( (case cleanBitSeq bs of (BitSeq len s bl) => (let shorten_opt = ((case (new_len, len) of (None, _) => None | (Some l1, None) => Some l1 | (Some l1, Some l2) => if (l1 < l2) then Some l1 else None )) in (case shorten_opt of None => BitSeq new_len s bl | Some l1 => ( (let bl' = (List.take l1 (bl @ [s])) in (case dest_init bl' of None => (BitSeq len s bl) \ \\ do nothing if size 0 is requested \\ | Some (bl'', s') => cleanBitSeq (BitSeq new_len s' bl'') ))) )) ) )" \ \\val bitSeqNot : bitSequence -> bitSequence\\ fun bitSeqNot :: " bitSequence \ bitSequence " where " bitSeqNot (BitSeq len s bl) = ( BitSeq len (\ s) (List.map (\ x. \ x) bl))" \ \\val bitSeqBinop : (bool -> bool -> bool) -> bitSequence -> bitSequence -> bitSequence\\ \ \\val bitSeqBinopAux : (bool -> bool -> bool) -> bool -> list bool -> bool -> list bool -> list bool\\ fun bitSeqBinopAux :: "(bool \ bool \ bool)\ bool \(bool)list \ bool \(bool)list \(bool)list " where " bitSeqBinopAux binop s1 ([]) s2 ([]) = ( [])" |" bitSeqBinopAux binop s1 (b1 # bl1') s2 ([]) = ( (binop b1 s2) # bitSeqBinopAux binop s1 bl1' s2 [])" |" bitSeqBinopAux binop s1 ([]) s2 (b2 # bl2') = ( (binop s1 b2) # bitSeqBinopAux binop s1 [] s2 bl2' )" |" bitSeqBinopAux binop s1 (b1 # bl1') s2 (b2 # bl2') = ( (binop b1 b2) # bitSeqBinopAux binop s1 bl1' s2 bl2' )" definition bitSeqBinop :: "(bool \ bool \ bool)\ bitSequence \ bitSequence \ bitSequence " where " bitSeqBinop binop bs1 bs2 = ( ( (case cleanBitSeq bs1 of (BitSeq len1 s1 bl1) => (case cleanBitSeq bs2 of (BitSeq len2 s2 bl2) => (let len = ((case (len1, len2) of (Some l1, Some l2) => Some (max l1 l2) | _ => None )) in (let s = (binop s1 s2) in (let bl = (bitSeqBinopAux binop s1 bl1 s2 bl2) in cleanBitSeq (BitSeq len s bl)))) ) ) ))" definition bitSeqAnd :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqAnd = ( bitSeqBinop (\))" definition bitSeqOr :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqOr = ( bitSeqBinop (\))" definition bitSeqXor :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqXor = ( bitSeqBinop (\ b1 b2. \ (b1 \ b2)))" \ \\val bitSeqShiftLeft : bitSequence -> nat -> bitSequence\\ fun bitSeqShiftLeft :: " bitSequence \ nat \ bitSequence " where " bitSeqShiftLeft (BitSeq len s bl) n = ( cleanBitSeq (BitSeq len s (List.replicate n False @ bl)))" \ \\val bitSeqArithmeticShiftRight : bitSequence -> nat -> bitSequence\\ definition bitSeqArithmeticShiftRight :: " bitSequence \ nat \ bitSequence " where " bitSeqArithmeticShiftRight bs n = ( (case cleanBitSeq bs of (BitSeq len s bl) => cleanBitSeq (BitSeq len s (List.drop n bl)) ) )" \ \\val bitSeqLogicalShiftRight : bitSequence -> nat -> bitSequence\\ definition bitSeqLogicalShiftRight :: " bitSequence \ nat \ bitSequence " where " bitSeqLogicalShiftRight bs n = ( if (n =( 0 :: nat)) then cleanBitSeq bs else (case cleanBitSeq bs of (BitSeq len s bl) => (case len of None => cleanBitSeq (BitSeq len s (List.drop n bl)) | Some l => cleanBitSeq (BitSeq len False ((List.drop n bl) @ List.replicate l s)) ) ) )" \ \\ integerFromBoolList sign bl creates an integer from a list of bits (least significant bit first) and an explicitly given sign bit. It uses two's complement encoding. \\ \ \\val integerFromBoolList : (bool * list bool) -> integer\\ fun integerFromBoolListAux :: " int \(bool)list \ int " where " integerFromBoolListAux (acc1 :: int) (([]) :: bool list) = ( acc1 )" |" integerFromBoolListAux (acc1 :: int) ((True # bl') :: bool list) = ( integerFromBoolListAux ((acc1 *( 2 :: int)) +( 1 :: int)) bl' )" |" integerFromBoolListAux (acc1 :: int) ((False # bl') :: bool list) = ( integerFromBoolListAux (acc1 *( 2 :: int)) bl' )" fun integerFromBoolList :: " bool*(bool)list \ int " where " integerFromBoolList (sign, bl) = ( if sign then - (integerFromBoolListAux(( 0 :: int)) (List.rev (List.map (\ x. \ x) bl)) +( 1 :: int)) else integerFromBoolListAux(( 0 :: int)) (List.rev bl))" \ \\ [boolListFromInteger i] creates a sign bit and a list of booleans from an integer. The len_opt tells it when to stop.\\ \ \\val boolListFromInteger : integer -> bool * list bool\\ fun boolListFromNatural :: "(bool)list \ nat \(bool)list " where " boolListFromNatural acc1 (remainder :: nat) = ( if (remainder >( 0 :: nat)) then (boolListFromNatural (((remainder mod( 2 :: nat)) =( 1 :: nat)) # acc1) (remainder div( 2 :: nat))) else List.rev acc1 )" definition boolListFromInteger :: " int \ bool*(bool)list " where " boolListFromInteger (i :: int) = ( if (i <( 0 :: int)) then (True, List.map (\ x. \ x) (boolListFromNatural [] (nat (abs (- (i +( 1 :: int))))))) else (False, boolListFromNatural [] (nat (abs i))))" \ \\ [bitSeqFromInteger len_opt i] encodes [i] as a bitsequence with [len_opt] bits. If there are not enough bits, truncation happens \\ \ \\val bitSeqFromInteger : maybe nat -> integer -> bitSequence\\ definition bitSeqFromInteger :: "(nat)option \ int \ bitSequence " where " bitSeqFromInteger len_opt i = ( (let (s, bl) = (boolListFromInteger i) in resizeBitSeq len_opt (BitSeq None s bl)))" \ \\val integerFromBitSeq : bitSequence -> integer\\ definition integerFromBitSeq :: " bitSequence \ int " where " integerFromBitSeq bs = ( (case cleanBitSeq bs of (BitSeq len s bl) => integerFromBoolList (s, bl) ) )" \ \\ Now we can via translation to integers map arithmetic operations to bitSequences \\ \ \\val bitSeqArithUnaryOp : (integer -> integer) -> bitSequence -> bitSequence\\ definition bitSeqArithUnaryOp :: "(int \ int)\ bitSequence \ bitSequence " where " bitSeqArithUnaryOp uop bs = ( (case bs of (BitSeq len _ _) => bitSeqFromInteger len (uop (integerFromBitSeq bs)) ) )" \ \\val bitSeqArithBinOp : (integer -> integer -> integer) -> bitSequence -> bitSequence -> bitSequence\\ definition bitSeqArithBinOp :: "(int \ int \ int)\ bitSequence \ bitSequence \ bitSequence " where " bitSeqArithBinOp binop bs1 bs2 = ( (case bs1 of (BitSeq len1 _ _) => (case bs2 of (BitSeq len2 _ _) => (let len = ((case (len1, len2) of (Some l1, Some l2) => Some (max l1 l2) | _ => None )) in bitSeqFromInteger len (binop (integerFromBitSeq bs1) (integerFromBitSeq bs2))) ) ) )" \ \\val bitSeqArithBinTest : forall 'a. (integer -> integer -> 'a) -> bitSequence -> bitSequence -> 'a\\ definition bitSeqArithBinTest :: "(int \ int \ 'a)\ bitSequence \ bitSequence \ 'a " where " bitSeqArithBinTest binop bs1 bs2 = ( binop (integerFromBitSeq bs1) (integerFromBitSeq bs2))" \ \\ now instantiate the number interface for bit-sequences \\ \ \\val bitSeqFromNumeral : numeral -> bitSequence\\ \ \\val bitSeqLess : bitSequence -> bitSequence -> bool\\ definition bitSeqLess :: " bitSequence \ bitSequence \ bool " where " bitSeqLess bs1 bs2 = ( bitSeqArithBinTest (<) bs1 bs2 )" \ \\val bitSeqLessEqual : bitSequence -> bitSequence -> bool\\ definition bitSeqLessEqual :: " bitSequence \ bitSequence \ bool " where " bitSeqLessEqual bs1 bs2 = ( bitSeqArithBinTest (\) bs1 bs2 )" \ \\val bitSeqGreater : bitSequence -> bitSequence -> bool\\ definition bitSeqGreater :: " bitSequence \ bitSequence \ bool " where " bitSeqGreater bs1 bs2 = ( bitSeqArithBinTest (>) bs1 bs2 )" \ \\val bitSeqGreaterEqual : bitSequence -> bitSequence -> bool\\ definition bitSeqGreaterEqual :: " bitSequence \ bitSequence \ bool " where " bitSeqGreaterEqual bs1 bs2 = ( bitSeqArithBinTest (\) bs1 bs2 )" \ \\val bitSeqCompare : bitSequence -> bitSequence -> ordering\\ definition bitSeqCompare :: " bitSequence \ bitSequence \ ordering " where " bitSeqCompare bs1 bs2 = ( bitSeqArithBinTest (genericCompare (<) (=)) bs1 bs2 )" definition instance_Basic_classes_Ord_Word_bitSequence_dict :: "(bitSequence)Ord_class " where " instance_Basic_classes_Ord_Word_bitSequence_dict = ((| compare_method = bitSeqCompare, isLess_method = bitSeqLess, isLessEqual_method = bitSeqLessEqual, isGreater_method = bitSeqGreater, isGreaterEqual_method = bitSeqGreaterEqual |) )" \ \\ arithmetic negation, don't mix up with bitwise negation \\ \ \\val bitSeqNegate : bitSequence -> bitSequence\\ definition bitSeqNegate :: " bitSequence \ bitSequence " where " bitSeqNegate bs = ( bitSeqArithUnaryOp (\ i. - i) bs )" definition instance_Num_NumNegate_Word_bitSequence_dict :: "(bitSequence)NumNegate_class " where " instance_Num_NumNegate_Word_bitSequence_dict = ((| numNegate_method = bitSeqNegate |) )" \ \\val bitSeqAdd : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqAdd :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqAdd bs1 bs2 = ( bitSeqArithBinOp (+) bs1 bs2 )" definition instance_Num_NumAdd_Word_bitSequence_dict :: "(bitSequence)NumAdd_class " where " instance_Num_NumAdd_Word_bitSequence_dict = ((| numAdd_method = bitSeqAdd |) )" \ \\val bitSeqMinus : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqMinus :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqMinus bs1 bs2 = ( bitSeqArithBinOp (-) bs1 bs2 )" definition instance_Num_NumMinus_Word_bitSequence_dict :: "(bitSequence)NumMinus_class " where " instance_Num_NumMinus_Word_bitSequence_dict = ((| numMinus_method = bitSeqMinus |) )" \ \\val bitSeqSucc : bitSequence -> bitSequence\\ definition bitSeqSucc :: " bitSequence \ bitSequence " where " bitSeqSucc bs = ( bitSeqArithUnaryOp (\ n. n +( 1 :: int)) bs )" definition instance_Num_NumSucc_Word_bitSequence_dict :: "(bitSequence)NumSucc_class " where " instance_Num_NumSucc_Word_bitSequence_dict = ((| succ_method = bitSeqSucc |) )" \ \\val bitSeqPred : bitSequence -> bitSequence\\ definition bitSeqPred :: " bitSequence \ bitSequence " where " bitSeqPred bs = ( bitSeqArithUnaryOp (\ n. n -( 1 :: int)) bs )" definition instance_Num_NumPred_Word_bitSequence_dict :: "(bitSequence)NumPred_class " where " instance_Num_NumPred_Word_bitSequence_dict = ((| pred_method = bitSeqPred |) )" \ \\val bitSeqMult : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqMult :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqMult bs1 bs2 = ( bitSeqArithBinOp (*) bs1 bs2 )" definition instance_Num_NumMult_Word_bitSequence_dict :: "(bitSequence)NumMult_class " where " instance_Num_NumMult_Word_bitSequence_dict = ((| numMult_method = bitSeqMult |) )" \ \\val bitSeqPow : bitSequence -> nat -> bitSequence\\ definition bitSeqPow :: " bitSequence \ nat \ bitSequence " where " bitSeqPow bs n = ( bitSeqArithUnaryOp (\ i . i ^ n) bs )" definition instance_Num_NumPow_Word_bitSequence_dict :: "(bitSequence)NumPow_class " where " instance_Num_NumPow_Word_bitSequence_dict = ((| numPow_method = bitSeqPow |) )" \ \\val bitSeqDiv : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqDiv :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqDiv bs1 bs2 = ( bitSeqArithBinOp (div) bs1 bs2 )" definition instance_Num_NumIntegerDivision_Word_bitSequence_dict :: "(bitSequence)NumIntegerDivision_class " where " instance_Num_NumIntegerDivision_Word_bitSequence_dict = ((| div_method = bitSeqDiv |) )" definition instance_Num_NumDivision_Word_bitSequence_dict :: "(bitSequence)NumDivision_class " where " instance_Num_NumDivision_Word_bitSequence_dict = ((| numDivision_method = bitSeqDiv |) )" \ \\val bitSeqMod : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqMod :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqMod bs1 bs2 = ( bitSeqArithBinOp (mod) bs1 bs2 )" definition instance_Num_NumRemainder_Word_bitSequence_dict :: "(bitSequence)NumRemainder_class " where " instance_Num_NumRemainder_Word_bitSequence_dict = ((| mod_method = bitSeqMod |) )" \ \\val bitSeqMin : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqMin :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqMin bs1 bs2 = ( bitSeqArithBinOp min bs1 bs2 )" \ \\val bitSeqMax : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqMax :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqMax bs1 bs2 = ( bitSeqArithBinOp max bs1 bs2 )" definition instance_Basic_classes_OrdMaxMin_Word_bitSequence_dict :: "(bitSequence)OrdMaxMin_class " where " instance_Basic_classes_OrdMaxMin_Word_bitSequence_dict = ((| max_method = bitSeqMax, min_method = bitSeqMin |) )" \ \\ ========================================================================== \\ \ \\ Interface for bitoperations \\ \ \\ ========================================================================== \\ record 'a WordNot_class= lnot_method ::" 'a \ 'a " record 'a WordAnd_class= land_method ::" 'a \ 'a \ 'a " record 'a WordOr_class= lor_method ::" 'a \ 'a \ 'a " record 'a WordXor_class= lxor_method ::" 'a \ 'a \ 'a " record 'a WordLsl_class= lsl_method ::" 'a \ nat \ 'a " record 'a WordLsr_class= lsr_method ::" 'a \ nat \ 'a " record 'a WordAsr_class= asr_method ::" 'a \ nat \ 'a " \ \\ ----------------------- \\ \ \\ bitSequence \\ \ \\ ----------------------- \\ definition instance_Word_WordNot_Word_bitSequence_dict :: "(bitSequence)WordNot_class " where " instance_Word_WordNot_Word_bitSequence_dict = ((| lnot_method = bitSeqNot |) )" definition instance_Word_WordAnd_Word_bitSequence_dict :: "(bitSequence)WordAnd_class " where " instance_Word_WordAnd_Word_bitSequence_dict = ((| land_method = bitSeqAnd |) )" definition instance_Word_WordOr_Word_bitSequence_dict :: "(bitSequence)WordOr_class " where " instance_Word_WordOr_Word_bitSequence_dict = ((| lor_method = bitSeqOr |) )" definition instance_Word_WordXor_Word_bitSequence_dict :: "(bitSequence)WordXor_class " where " instance_Word_WordXor_Word_bitSequence_dict = ((| lxor_method = bitSeqXor |) )" definition instance_Word_WordLsl_Word_bitSequence_dict :: "(bitSequence)WordLsl_class " where " instance_Word_WordLsl_Word_bitSequence_dict = ((| lsl_method = bitSeqShiftLeft |) )" definition instance_Word_WordLsr_Word_bitSequence_dict :: "(bitSequence)WordLsr_class " where " instance_Word_WordLsr_Word_bitSequence_dict = ((| lsr_method = bitSeqLogicalShiftRight |) )" definition instance_Word_WordAsr_Word_bitSequence_dict :: "(bitSequence)WordAsr_class " where " instance_Word_WordAsr_Word_bitSequence_dict = ((| asr_method = bitSeqArithmeticShiftRight |) )" \ \\ ----------------------- \\ \ \\ int32 \\ \ \\ ----------------------- \\ \ \\val int32Lnot : int32 -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordNot_Num_int32_dict :: "( 32 word)WordNot_class " where " instance_Word_WordNot_Num_int32_dict = ((| lnot_method = (\ w. (NOT w))|) )" \ \\val int32Lor : int32 -> int32 -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordOr_Num_int32_dict :: "( 32 word)WordOr_class " where " instance_Word_WordOr_Num_int32_dict = ((| lor_method = (OR)|) )" \ \\val int32Lxor : int32 -> int32 -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordXor_Num_int32_dict :: "( 32 word)WordXor_class " where " instance_Word_WordXor_Num_int32_dict = ((| lxor_method = (XOR)|) )" \ \\val int32Land : int32 -> int32 -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordAnd_Num_int32_dict :: "( 32 word)WordAnd_class " where " instance_Word_WordAnd_Num_int32_dict = ((| land_method = (AND)|) )" \ \\val int32Lsl : int32 -> nat -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordLsl_Num_int32_dict :: "( 32 word)WordLsl_class " where " instance_Word_WordLsl_Num_int32_dict = ((| lsl_method = (<<)|) )" \ \\val int32Lsr : int32 -> nat -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordLsr_Num_int32_dict :: "( 32 word)WordLsr_class " where " instance_Word_WordLsr_Num_int32_dict = ((| lsr_method = (>>)|) )" \ \\val int32Asr : int32 -> nat -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordAsr_Num_int32_dict :: "( 32 word)WordAsr_class " where " instance_Word_WordAsr_Num_int32_dict = ((| asr_method = (>>>)|) )" \ \\ ----------------------- \\ \ \\ int64 \\ \ \\ ----------------------- \\ \ \\val int64Lnot : int64 -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordNot_Num_int64_dict :: "( 64 word)WordNot_class " where " instance_Word_WordNot_Num_int64_dict = ((| lnot_method = (\ w. (NOT w))|) )" \ \\val int64Lor : int64 -> int64 -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordOr_Num_int64_dict :: "( 64 word)WordOr_class " where " instance_Word_WordOr_Num_int64_dict = ((| lor_method = (OR)|) )" \ \\val int64Lxor : int64 -> int64 -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordXor_Num_int64_dict :: "( 64 word)WordXor_class " where " instance_Word_WordXor_Num_int64_dict = ((| lxor_method = (XOR)|) )" \ \\val int64Land : int64 -> int64 -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordAnd_Num_int64_dict :: "( 64 word)WordAnd_class " where " instance_Word_WordAnd_Num_int64_dict = ((| land_method = (AND)|) )" \ \\val int64Lsl : int64 -> nat -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordLsl_Num_int64_dict :: "( 64 word)WordLsl_class " where " instance_Word_WordLsl_Num_int64_dict = ((| lsl_method = (<<)|) )" \ \\val int64Lsr : int64 -> nat -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordLsr_Num_int64_dict :: "( 64 word)WordLsr_class " where " instance_Word_WordLsr_Num_int64_dict = ((| lsr_method = (>>)|) )" \ \\val int64Asr : int64 -> nat -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordAsr_Num_int64_dict :: "( 64 word)WordAsr_class " where " instance_Word_WordAsr_Num_int64_dict = ((| asr_method = (>>>)|) )" \ \\ ----------------------- \\ \ \\ Words via bit sequences \\ \ \\ ----------------------- \\ \ \\val defaultLnot : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> 'a\\ definition defaultLnot :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ 'a " where " defaultLnot fromBitSeq toBitSeq x = ( fromBitSeq (bitSeqNegate (toBitSeq x)))" \ \\val defaultLand : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> 'a -> 'a\\ definition defaultLand :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ 'a \ 'a " where " defaultLand fromBitSeq toBitSeq x1 x2 = ( fromBitSeq (bitSeqAnd (toBitSeq x1) (toBitSeq x2)))" \ \\val defaultLor : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> 'a -> 'a\\ definition defaultLor :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ 'a \ 'a " where " defaultLor fromBitSeq toBitSeq x1 x2 = ( fromBitSeq (bitSeqOr (toBitSeq x1) (toBitSeq x2)))" \ \\val defaultLxor : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> 'a -> 'a\\ definition defaultLxor :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ 'a \ 'a " where " defaultLxor fromBitSeq toBitSeq x1 x2 = ( fromBitSeq (bitSeqXor (toBitSeq x1) (toBitSeq x2)))" \ \\val defaultLsl : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> nat -> 'a\\ definition defaultLsl :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ nat \ 'a " where " defaultLsl fromBitSeq toBitSeq x n = ( fromBitSeq (bitSeqShiftLeft (toBitSeq x) n))" \ \\val defaultLsr : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> nat -> 'a\\ definition defaultLsr :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ nat \ 'a " where " defaultLsr fromBitSeq toBitSeq x n = ( fromBitSeq (bitSeqLogicalShiftRight (toBitSeq x) n))" \ \\val defaultAsr : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> nat -> 'a\\ definition defaultAsr :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ nat \ 'a " where " defaultAsr fromBitSeq toBitSeq x n = ( fromBitSeq (bitSeqArithmeticShiftRight (toBitSeq x) n))" \ \\ ----------------------- \\ \ \\ integer \\ \ \\ ----------------------- \\ \ \\val integerLnot : integer -> integer\\ definition integerLnot :: " int \ int " where " integerLnot i = ( - (i +( 1 :: int)))" definition instance_Word_WordNot_Num_integer_dict :: "(int)WordNot_class " where " instance_Word_WordNot_Num_integer_dict = ((| lnot_method = integerLnot |) )" \ \\val integerLor : integer -> integer -> integer\\ definition integerLor :: " int \ int \ int " where " integerLor i1 i2 = ( defaultLor integerFromBitSeq (bitSeqFromInteger None) i1 i2 )" definition instance_Word_WordOr_Num_integer_dict :: "(int)WordOr_class " where " instance_Word_WordOr_Num_integer_dict = ((| lor_method = integerLor |) )" \ \\val integerLxor : integer -> integer -> integer\\ definition integerLxor :: " int \ int \ int " where " integerLxor i1 i2 = ( defaultLxor integerFromBitSeq (bitSeqFromInteger None) i1 i2 )" definition instance_Word_WordXor_Num_integer_dict :: "(int)WordXor_class " where " instance_Word_WordXor_Num_integer_dict = ((| lxor_method = integerLxor |) )" \ \\val integerLand : integer -> integer -> integer\\ definition integerLand :: " int \ int \ int " where " integerLand i1 i2 = ( defaultLand integerFromBitSeq (bitSeqFromInteger None) i1 i2 )" definition instance_Word_WordAnd_Num_integer_dict :: "(int)WordAnd_class " where " instance_Word_WordAnd_Num_integer_dict = ((| land_method = integerLand |) )" \ \\val integerLsl : integer -> nat -> integer\\ definition integerLsl :: " int \ nat \ int " where " integerLsl i n = ( defaultLsl integerFromBitSeq (bitSeqFromInteger None) i n )" definition instance_Word_WordLsl_Num_integer_dict :: "(int)WordLsl_class " where " instance_Word_WordLsl_Num_integer_dict = ((| lsl_method = integerLsl |) )" \ \\val integerAsr : integer -> nat -> integer\\ definition integerAsr :: " int \ nat \ int " where " integerAsr i n = ( defaultAsr integerFromBitSeq (bitSeqFromInteger None) i n )" definition instance_Word_WordLsr_Num_integer_dict :: "(int)WordLsr_class " where " instance_Word_WordLsr_Num_integer_dict = ((| lsr_method = integerAsr |) )" definition instance_Word_WordAsr_Num_integer_dict :: "(int)WordAsr_class " where " instance_Word_WordAsr_Num_integer_dict = ((| asr_method = integerAsr |) )" \ \\ ----------------------- \\ \ \\ int \\ \ \\ ----------------------- \\ \ \\ sometimes it is convenient to be able to perform bit-operations on ints. However, since int is not well-defined (it has different size on different systems), it should be used very carefully and only for operations that don't depend on the bitwidth of int \\ \ \\val intFromBitSeq : bitSequence -> int\\ definition intFromBitSeq :: " bitSequence \ int " where " intFromBitSeq bs = ( (integerFromBitSeq (resizeBitSeq (Some(( 31 :: nat))) bs)))" \ \\val bitSeqFromInt : int -> bitSequence\\ definition bitSeqFromInt :: " int \ bitSequence " where " bitSeqFromInt i = ( bitSeqFromInteger (Some(( 31 :: nat))) ( i))" \ \\val intLnot : int -> int\\ definition intLnot :: " int \ int " where " intLnot i = ( - (i +( 1 :: int)))" definition instance_Word_WordNot_Num_int_dict :: "(int)WordNot_class " where " instance_Word_WordNot_Num_int_dict = ((| lnot_method = intLnot |) )" \ \\val intLor : int -> int -> int\\ definition intLor :: " int \ int \ int " where " intLor i1 i2 = ( defaultLor intFromBitSeq bitSeqFromInt i1 i2 )" definition instance_Word_WordOr_Num_int_dict :: "(int)WordOr_class " where " instance_Word_WordOr_Num_int_dict = ((| lor_method = intLor |) )" \ \\val intLxor : int -> int -> int\\ definition intLxor :: " int \ int \ int " where " intLxor i1 i2 = ( defaultLxor intFromBitSeq bitSeqFromInt i1 i2 )" definition instance_Word_WordXor_Num_int_dict :: "(int)WordXor_class " where " instance_Word_WordXor_Num_int_dict = ((| lxor_method = intLxor |) )" \ \\val intLand : int -> int -> int\\ definition intLand :: " int \ int \ int " where " intLand i1 i2 = ( defaultLand intFromBitSeq bitSeqFromInt i1 i2 )" definition instance_Word_WordAnd_Num_int_dict :: "(int)WordAnd_class " where " instance_Word_WordAnd_Num_int_dict = ((| land_method = intLand |) )" \ \\val intLsl : int -> nat -> int\\ definition intLsl :: " int \ nat \ int " where " intLsl i n = ( defaultLsl intFromBitSeq bitSeqFromInt i n )" definition instance_Word_WordLsl_Num_int_dict :: "(int)WordLsl_class " where " instance_Word_WordLsl_Num_int_dict = ((| lsl_method = intLsl |) )" \ \\val intAsr : int -> nat -> int\\ definition intAsr :: " int \ nat \ int " where " intAsr i n = ( defaultAsr intFromBitSeq bitSeqFromInt i n )" definition instance_Word_WordAsr_Num_int_dict :: "(int)WordAsr_class " where " instance_Word_WordAsr_Num_int_dict = ((| asr_method = intAsr |) )" \ \\ ----------------------- \\ \ \\ natural \\ \ \\ ----------------------- \\ \ \\ some operations work also on positive numbers \\ \ \\val naturalFromBitSeq : bitSequence -> natural\\ definition naturalFromBitSeq :: " bitSequence \ nat " where " naturalFromBitSeq bs = ( nat (abs (integerFromBitSeq bs)))" \ \\val bitSeqFromNatural : maybe nat -> natural -> bitSequence\\ definition bitSeqFromNatural :: "(nat)option \ nat \ bitSequence " where " bitSeqFromNatural len n = ( bitSeqFromInteger len (int n))" \ \\val naturalLor : natural -> natural -> natural\\ definition naturalLor :: " nat \ nat \ nat " where " naturalLor i1 i2 = ( defaultLor naturalFromBitSeq (bitSeqFromNatural None) i1 i2 )" definition instance_Word_WordOr_Num_natural_dict :: "(nat)WordOr_class " where " instance_Word_WordOr_Num_natural_dict = ((| lor_method = naturalLor |) )" \ \\val naturalLxor : natural -> natural -> natural\\ definition naturalLxor :: " nat \ nat \ nat " where " naturalLxor i1 i2 = ( defaultLxor naturalFromBitSeq (bitSeqFromNatural None) i1 i2 )" definition instance_Word_WordXor_Num_natural_dict :: "(nat)WordXor_class " where " instance_Word_WordXor_Num_natural_dict = ((| lxor_method = naturalLxor |) )" \ \\val naturalLand : natural -> natural -> natural\\ definition naturalLand :: " nat \ nat \ nat " where " naturalLand i1 i2 = ( defaultLand naturalFromBitSeq (bitSeqFromNatural None) i1 i2 )" definition instance_Word_WordAnd_Num_natural_dict :: "(nat)WordAnd_class " where " instance_Word_WordAnd_Num_natural_dict = ((| land_method = naturalLand |) )" \ \\val naturalLsl : natural -> nat -> natural\\ definition naturalLsl :: " nat \ nat \ nat " where " naturalLsl i n = ( defaultLsl naturalFromBitSeq (bitSeqFromNatural None) i n )" definition instance_Word_WordLsl_Num_natural_dict :: "(nat)WordLsl_class " where " instance_Word_WordLsl_Num_natural_dict = ((| lsl_method = naturalLsl |) )" \ \\val naturalAsr : natural -> nat -> natural\\ definition naturalAsr :: " nat \ nat \ nat " where " naturalAsr i n = ( defaultAsr naturalFromBitSeq (bitSeqFromNatural None) i n )" definition instance_Word_WordLsr_Num_natural_dict :: "(nat)WordLsr_class " where " instance_Word_WordLsr_Num_natural_dict = ((| lsr_method = naturalAsr |) )" definition instance_Word_WordAsr_Num_natural_dict :: "(nat)WordAsr_class " where " instance_Word_WordAsr_Num_natural_dict = ((| asr_method = naturalAsr |) )" \ \\ ----------------------- \\ \ \\ nat \\ \ \\ ----------------------- \\ \ \\ sometimes it is convenient to be able to perform bit-operations on nats. However, since nat is not well-defined (it has different size on different systems), it should be used very carefully and only for operations that don't depend on the bitwidth of nat \\ \ \\val natFromBitSeq : bitSequence -> nat\\ definition natFromBitSeq :: " bitSequence \ nat " where " natFromBitSeq bs = ( (naturalFromBitSeq (resizeBitSeq (Some(( 31 :: nat))) bs)))" \ \\val bitSeqFromNat : nat -> bitSequence\\ definition bitSeqFromNat :: " nat \ bitSequence " where " bitSeqFromNat i = ( bitSeqFromNatural (Some(( 31 :: nat))) ( i))" \ \\val natLor : nat -> nat -> nat\\ definition natLor :: " nat \ nat \ nat " where " natLor i1 i2 = ( defaultLor natFromBitSeq bitSeqFromNat i1 i2 )" definition instance_Word_WordOr_nat_dict :: "(nat)WordOr_class " where " instance_Word_WordOr_nat_dict = ((| lor_method = natLor |) )" \ \\val natLxor : nat -> nat -> nat\\ definition natLxor :: " nat \ nat \ nat " where " natLxor i1 i2 = ( defaultLxor natFromBitSeq bitSeqFromNat i1 i2 )" definition instance_Word_WordXor_nat_dict :: "(nat)WordXor_class " where " instance_Word_WordXor_nat_dict = ((| lxor_method = natLxor |) )" \ \\val natLand : nat -> nat -> nat\\ definition natLand :: " nat \ nat \ nat " where " natLand i1 i2 = ( defaultLand natFromBitSeq bitSeqFromNat i1 i2 )" definition instance_Word_WordAnd_nat_dict :: "(nat)WordAnd_class " where " instance_Word_WordAnd_nat_dict = ((| land_method = natLand |) )" \ \\val natLsl : nat -> nat -> nat\\ definition natLsl :: " nat \ nat \ nat " where " natLsl i n = ( defaultLsl natFromBitSeq bitSeqFromNat i n )" definition instance_Word_WordLsl_nat_dict :: "(nat)WordLsl_class " where " instance_Word_WordLsl_nat_dict = ((| lsl_method = natLsl |) )" \ \\val natAsr : nat -> nat -> nat\\ definition natAsr :: " nat \ nat \ nat " where " natAsr i n = ( defaultAsr natFromBitSeq bitSeqFromNat i n )" definition instance_Word_WordAsr_nat_dict :: "(nat)WordAsr_class " where " instance_Word_WordAsr_nat_dict = ((| asr_method = natAsr |) )" end diff --git a/thys/Containers/ITP-2013/Benchmark_Set.thy b/thys/Containers/ITP-2013/Benchmark_Set.thy --- a/thys/Containers/ITP-2013/Benchmark_Set.thy +++ b/thys/Containers/ITP-2013/Benchmark_Set.thy @@ -1,82 +1,82 @@ theory Benchmark_Set imports - "HOL-Word.Word" + "HOL-Library.Word" "HOL-Library.Cardinality" begin instantiation word :: (len) card_UNIV begin definition "finite_UNIV = Phantom('a word) True" definition "card_UNIV = Phantom('a word) (2 ^ LENGTH('a))" instance by(intro_classes)(simp_all add: card_UNIV_word_def card_word finite_UNIV_word_def) end definition word_of_integer :: "integer \ 'a::len word" where "word_of_integer = word_of_int \ int_of_integer" lemma word_of_integer_code [code]: "word_of_integer k = (if k < 0 then - word_of_integer (- k) else if k = 0 then 0 else let (q, r) = divmod_integer k 2 in if r = 0 then 2 * word_of_integer q else 2 * word_of_integer q + 1)" apply (unfold word_of_integer_def) apply transfer apply (auto simp add: not_less comp_def split_beta) apply (subst int_of_integer_code) apply (clarsimp simp add: divmod_integer_def) done definition word_of :: "natural \ 'a::len word" where "word_of = word_of_integer o integer_of_natural" text \randomly generate a set of (up to) n elements drawn from 0 to bound\ fun gen_build1 :: "natural \ nat \ (32 word set \ Random.seed) \ (32 word set \ Random.seed)" where "gen_build1 bound n (A, seed) = (if n = 0 then (A, seed) else let (x, seed') = Random.range bound seed in gen_build1 bound (n - 1) (insert (word_of x) A, seed'))" declare gen_build1.simps[simp del] definition build1 :: "natural \ Random.seed \ (32 word set \ Random.seed)" where "build1 bound seed = (let (n', seed') = Random.range bound seed; (compl, seed'') = Random.range 2 seed; (x, seed''') = gen_build1 bound (Code_Numeral.nat_of_natural n') ({}, seed'') in (if compl = 0 then x else - x, seed'''))" text \randomly generate a set of (up to) n sets each with a random number between 0 and bound of elements between 0 and bound\ fun gen_build2 :: "natural \ nat \ (32 word set set \ Random.seed) \ (32 word set set \ Random.seed)" where "gen_build2 bound n (A, seed) = (if n = 0 then (A, seed) else let (x, seed') = build1 bound seed in gen_build2 bound (n - 1) (insert x A, seed'))" declare gen_build2.simps[simp del] definition build :: "nat \ nat \ Random.seed \ 32 word set set \ Random.seed" where "build n m seed = gen_build2 (of_nat m) n ({}, seed)" fun gen_lookup :: "32 word set set \ natural \ nat \ (nat \ Random.seed) \ (nat \ Random.seed)" where "gen_lookup A bound n (hits, seed) = (if n = 0 then (hits, seed) else let (x, seed') = build1 bound seed in gen_lookup A bound (n - 1) (if x : A then hits + 1 else hits, seed'))" declare gen_lookup.simps [simp del] primrec lookup :: "nat \ nat \ (32 word set set \ Random.seed) \ (nat \ Random.seed)" where "lookup n m (A, seed) = gen_lookup A (of_nat m) 100 (0, seed)" definition covered :: "32 word set set \ nat" where "covered A = card (\A)" definition complete :: "nat \ nat \ Random.seed \ (nat \ nat)" where "complete n m seed = (let (A, seed') = build n m seed in (fst (lookup n m (A, seed)), covered A))" end diff --git a/thys/Datatype_Order_Generator/ROOT b/thys/Datatype_Order_Generator/ROOT --- a/thys/Datatype_Order_Generator/ROOT +++ b/thys/Datatype_Order_Generator/ROOT @@ -1,14 +1,13 @@ chapter AFP session Datatype_Order_Generator (AFP) = "HOL-Library" + options [timeout = 600] sessions - "HOL-Word" Deriving Native_Word theories Derive Derive_Examples document_files "root.bib" "root.tex" diff --git a/thys/IP_Addresses/NumberWang_IPv4.thy b/thys/IP_Addresses/NumberWang_IPv4.thy --- a/thys/IP_Addresses/NumberWang_IPv4.thy +++ b/thys/IP_Addresses/NumberWang_IPv4.thy @@ -1,65 +1,65 @@ theory NumberWang_IPv4 imports Main - "HOL-Word.Word" + "HOL-Library.Word" begin lemma zdiv_mult_self: \m \ 0 \ (a + m * n) div m = a div m + n\ for a m n :: int by simp section\Helper Lemmas for Low-Level Operations on Machine Words\ text\Needed for IPv4 Syntax\ lemma mod256: "((d::nat) + 256 * c + 65536 * b + 16777216 * a) mod 256 = d mod 256" proof - from mod_mult_self2[where a="d + 256 * c + 65536 * b" and b="256" and c="65536 * a"] have "(d + 256 * c + 65536 * b + 256 * 65536 * a) mod 256 = (d + 256 * c + 65536 * b) mod 256" by simp also have "\ = (d + 256 * c) mod 256" using mod_mult_self2[where a="d + 256 * c" and b="256" and c="256 * b"] by simp also have "\ = d mod 256" using mod_mult_self2 by blast finally show ?thesis by simp qed lemma div65536: assumes "a < 256" "b < 256" "c < 256" "d < 256" shows "nat ((int d + int (256 * c) + int (65536 * b) + int (16777216 * a)) div 65536 mod 256) = b" proof - from zdiv_mult_self[where a="int d + int (256 * c) + int (65536 * b)" and m=65536 and n="256 * (int a)"] have "(int d + int (256 * c) + int (65536 * b) + int (16777216 * a)) div 65536 = ((int d + int (256 * c) + int (65536 * b)) div 65536) + 256 * int a" by linarith also from zdiv_mult_self[where a="int d + int (256 * c)" and m="65536" and n="int b"] have "\ = (int d + int (256 * c)) div 65536 + int b + 256 * int a" by linarith also from assms have "\ = int b + 256 * int a" by simp finally have helper: "(int d + int (256 * c) + int (65536 * b) + int (16777216 * a)) div 65536 = int b + 256 * int a" . with assms show ?thesis by simp qed lemma div256: assumes "a < 256" "b < 256" "c < 256" "d < 256" shows "nat ((int d + int (256 * c) + int (65536 * b) + int (16777216 * a)) div 256 mod 256) = c" proof - from zdiv_mult_self[where a="int d + int (256 * c) + int (65536 * b)" and m=256 and n="65536 * (int a)"] have "(int d + int (256 * c) + int (65536 * b) + int (16777216 * a)) div 256 = ((int d + int (256 * c) + int (65536 * b)) div 256) + 65536 * int a" by linarith also from zdiv_mult_self[where a="int d + int (256 * c)" and m="256" and n="256 * int b"] have "\ = (int d + int (256 * c)) div 256 + 256 * int b + 65536 * int a" by linarith also from zdiv_mult_self[where a="int d" and m="256" and n="int c"] have "\ = (int d) div 256 + int c + 256 * int b + 65536 * int a" by linarith also from assms have "\ = int c + 256 * int b + 65536 * int a" by simp finally have helper1: "(int d + int (256 * c) + int (65536 * b) + int (16777216 * a)) div 256 = int c + 256 * int b + 65536 * int a" . from mod_mult_self2[where a="int c + 256 * int b" and c="256 * int a" and b=256] have "(int c + 256 * int b + 65536 * int a) mod 256 = (int c + 256 * int b) mod 256" by simp also have "\ = int c mod 256" using mod_mult_self2[where a="int c" and b=256 and c="int b"] by simp also have "\ = int c" using assms apply(subst mod_pos_pos_trivial) by(simp_all) finally have helper2: "(int c + 256 * int b + 65536 * int a) mod 256 = int c" . from helper1 helper2 show ?thesis by simp qed end diff --git a/thys/Interval_Arithmetic_Word32/Interval_Word32.thy b/thys/Interval_Arithmetic_Word32/Interval_Word32.thy --- a/thys/Interval_Arithmetic_Word32/Interval_Word32.thy +++ b/thys/Interval_Arithmetic_Word32/Interval_Word32.thy @@ -1,4146 +1,4145 @@ (* Author: Brandon Bohrer *) theory Interval_Word32 imports Complex_Main - "HOL-Word.More_Word" Word_Lib.Word_Lemmas Word_Lib.Word_Lib Word_Lib.Word_Syntax Word_Lib.Bitwise begin abbreviation signed_real_of_word :: \'a::len word \ real\ where \signed_real_of_word \ signed\ lemma (in linordered_idom) signed_less_numeral_iff: \signed w < numeral n \ sint w < numeral n\ (is \?P \ ?Q\) proof - have \?Q \ of_int (sint w) < of_int (numeral n)\ by (simp only: of_int_less_iff) also have \\ \ ?P\ by (transfer fixing: less less_eq n) simp finally show ?thesis .. qed lemma (in linordered_idom) signed_less_neg_numeral_iff: \signed w < - numeral n \ sint w < - numeral n\ (is \?P \ ?Q\) proof - have \?Q \ of_int (sint w) < of_int (- numeral n)\ by (simp only: of_int_less_iff) also have \\ \ ?P\ by (transfer fixing: less less_eq uminus n) simp finally show ?thesis .. qed lemma (in linordered_idom) numeral_less_signed_iff: \numeral n < signed w \ numeral n < sint w\ (is \?P \ ?Q\) proof - have \?Q \ of_int (numeral n) < of_int (sint w)\ by (simp only: of_int_less_iff) also have \\ \ ?P\ by (transfer fixing: less less_eq n) simp finally show ?thesis .. qed lemma (in linordered_idom) neg_numeral_less_signed_iff: \- numeral n < signed w \ - numeral n < sint w\ (is \?P \ ?Q\) proof - have \?Q \ of_int (- numeral n) < of_int (sint w)\ by (simp only: of_int_less_iff) also have \\ \ ?P\ by (transfer fixing: less less_eq uminus n) simp finally show ?thesis .. qed lemma (in linordered_idom) signed_nonnegative_iff: \0 \ signed w \ 0 \ sint w\ (is \?P \ ?Q\) proof - have \?Q \ of_int 0 \ of_int (sint w)\ by (simp only: of_int_le_iff) also have \\ \ ?P\ by (transfer fixing: less_eq) simp finally show ?thesis .. qed lemma signed_real_of_word_plus_numeral_eq_signed_real_of_word_iff: \signed_real_of_word v + numeral n = signed_real_of_word w \ sint v + numeral n = sint w\ (is \?P \ ?Q\) proof - have \?Q \ real_of_int (sint v + numeral n) = real_of_int (sint w)\ by (simp only: of_int_eq_iff) also have \\ \ ?P\ by simp finally show ?thesis .. qed lemma signed_real_of_word_sum_less_eq_numeral_iff: \signed_real_of_word v + signed_real_of_word w \ numeral n \ sint v + sint w \ numeral n\ (is \?P \ ?Q\) proof - have \?Q \ real_of_int (sint v + sint w) \ real_of_int (numeral n)\ by (simp only: of_int_le_iff) also have \\ \ ?P\ by simp finally show ?thesis .. qed lemma signed_real_of_word_sum_less_eq_neg_numeral_iff: \signed_real_of_word v + signed_real_of_word w \ - numeral n \ sint v + sint w \ - numeral n\ (is \?P \ ?Q\) proof - have \?Q \ real_of_int (sint v + sint w) \ real_of_int (- numeral n)\ by (simp only: of_int_le_iff) also have \\ \ ?P\ by simp finally show ?thesis .. qed lemma signed_real_of_word_sum_less_numeral_iff: \signed_real_of_word v + signed_real_of_word w < numeral n \ sint v + sint w < numeral n\ (is \?P \ ?Q\) proof - have \?Q \ real_of_int (sint v + sint w) < real_of_int (numeral n)\ by (simp only: of_int_less_iff) also have \\ \ ?P\ by simp finally show ?thesis .. qed lemma signed_real_of_word_sum_less_neg_numeral_iff: \signed_real_of_word v + signed_real_of_word w < - numeral n \ sint v + sint w < - numeral n\ (is \?P \ ?Q\) proof - have \?Q \ real_of_int (sint v + sint w) < real_of_int (- numeral n)\ by (simp only: of_int_less_iff) also have \\ \ ?P\ by simp finally show ?thesis .. qed lemma numeral_less_eq_signed_real_of_word_sum: \numeral n \ signed_real_of_word v + signed_real_of_word w \ numeral n \ sint v + sint w\ (is \?P \ ?Q\) proof - have \?Q \ real_of_int (numeral n) \ real_of_int (sint v + sint w)\ by (simp only: of_int_le_iff) also have \\ \ ?P\ by simp finally show ?thesis .. qed lemma neg_numeral_less_eq_signed_real_of_word_sum: \- numeral n \ signed_real_of_word v + signed_real_of_word w \ - numeral n \ sint v + sint w\ (is \?P \ ?Q\) proof - have \?Q \ real_of_int (- numeral n) \ real_of_int (sint v + sint w)\ by (simp only: of_int_le_iff) also have \\ \ ?P\ by simp finally show ?thesis .. qed lemma numeral_less_signed_real_of_word_sum: \numeral n < signed_real_of_word v + signed_real_of_word w \ numeral n < sint v + sint w\ (is \?P \ ?Q\) proof - have \?Q \ real_of_int (numeral n) < real_of_int (sint v + sint w)\ by (simp only: of_int_less_iff) also have \\ \ ?P\ by simp finally show ?thesis .. qed lemma neg_numeral_less_signed_real_of_word_sum: \- numeral n < signed_real_of_word v + signed_real_of_word w \ - numeral n < sint v + sint w\ (is \?P \ ?Q\) proof - have \?Q \ real_of_int (- numeral n) < real_of_int (sint v + sint w)\ by (simp only: of_int_less_iff) also have \\ \ ?P\ by simp finally show ?thesis .. qed lemmas real_of_word_simps [simp] = signed_less_numeral_iff [where ?'a = real] numeral_less_signed_iff [where ?'a = real] signed_less_neg_numeral_iff [where ?'a = real] neg_numeral_less_signed_iff [where ?'a = real] signed_nonnegative_iff [where ?'a = real] lemmas more_real_of_word_simps = signed_real_of_word_plus_numeral_eq_signed_real_of_word_iff signed_real_of_word_sum_less_eq_numeral_iff signed_real_of_word_sum_less_eq_neg_numeral_iff signed_real_of_word_sum_less_numeral_iff signed_real_of_word_sum_less_neg_numeral_iff numeral_less_eq_signed_real_of_word_sum neg_numeral_less_eq_signed_real_of_word_sum numeral_less_signed_real_of_word_sum neg_numeral_less_signed_real_of_word_sum declare more_real_of_word_simps [simp] text\Interval-Word32.thy implements conservative interval arithmetic operators on 32-bit word values, with explicit infinities for values outside the representable bounds. It is suitable for use in interpreters for languages which must have a well-understood low-level behavior (see Interpreter.thy). This work was originally part of the paper by Bohrer \emph{et al.}~\cite{BohrerTMMP18}. It is worth noting that this is not the first formalization of interval arithmetic in Isabelle/HOL. This article is presented regardless because it has unique goals in mind which have led to unique design decisions. Our goal is generate code which can be used to perform conservative arithmetic in implementations extracted from a proof. The Isabelle standard library now features interval arithmetic, for example in Approximation.thy. Ours differs in two ways: 1) We use intervals with explicit positive and negative infinities, and with overflow checking. Such checking is often relevant in implementation-level code with unknown inputs. To promote memory-efficient implementations, we moreover use sentinel values for infinities, rather than datatype constructors. This is especially important in real-time settings where the garbarge collection required for datatypes can be a concern. 2) Our goal is not to use interval arithmetic to discharge Isabelle goals, but to generate useful proven-correct implementation code, see Interpreter.thy. On the other hand, we are not concerned with producing interval-based automation for arithmetic goals in HOL. In practice, much of the work in this theory comes down to sheer case-analysis. Bounds-checking requires many edge cases in arithmetic functions, which come with many cases in proofs. Where possible, we attempt to offload interesting facts about word representations of numbers into reusable lemmas, but even then main results require many subcases, each with a certain amount of arithmetic grunt work. \ section \Interval arithmetic definitions\ subsection \Syntax\ text\Words are 32-bit\ type_synonym word = "32 Word.word" text\Sentinel values for infinities. Note that we leave the maximum value ($2^31$) completed unused, so that negation of $(2^{31})-1$ is not an edge case\ definition NEG_INF::"word" where NEG_INF_def[simp]:"NEG_INF = -((2 ^ 31) -1)" definition NegInf::"real" where NegInf[simp]:"NegInf = real_of_int (sint NEG_INF)" definition POS_INF::"word" where POS_INF_def[simp]:"POS_INF = (2^31) - 1" definition PosInf::"real" where PosInf[simp]:"PosInf = real_of_int (sint POS_INF)" text\Subtype of words who represent a finite value. \ typedef bword = "{n::word. sint n \ sint NEG_INF \ sint n \ sint POS_INF}" apply(rule exI[where x=NEG_INF]) by (auto) text\Numeric literals\ type_synonym lit = bword setup_lifting type_definition_bword lift_definition bword_zero::"bword" is "0::32 Word.word" by auto lift_definition bword_one::"bword" is "1::32 Word.word" by(auto simp add: sint_uint) lift_definition bword_neg_one::"bword" is "-1::32 Word.word" by(auto) definition word::"word \ bool" where word_def[simp]:"word w \ w \ {NEG_INF..POS_INF}" named_theorems rep_simps "Simplifications for representation functions" text\Definitions of interval containment and word representation repe w r iff word w encodes real number r\ inductive repe ::"word \ real \ bool" (infix "\\<^sub>E" 10) where repPOS_INF:"r \ real_of_int (sint POS_INF) \ repe POS_INF r" | repNEG_INF:"r \ real_of_int (sint NEG_INF) \ repe NEG_INF r" | repINT:"(sint w) < real_of_int(sint POS_INF) \ (sint w) > real_of_int(sint NEG_INF) \ repe w (sint w)" inductive_simps repePos_simps[rep_simps]:"repe POS_INF r" and repeNeg_simps[rep_simps]:"repe NEG_INF r" and repeInt_simps[rep_simps]:"repe w (sint w)" text\repU w r if w represents an upper bound of r\ definition repU ::"word \ real \ bool" (infix "\\<^sub>U" 10) where "repU w r \ \ r'. r' \ r \ repe w r'" lemma repU_leq:"repU w r \ r' \ r \ repU w r'" unfolding repU_def using order_trans by auto text\repU w r if w represents a lower bound of r\ definition repL ::"word \ real \ bool" (infix "\\<^sub>L" 10) where "repL w r \ \ r'. r' \ r \ repe w r'" lemma repL_geq:"repL w r \ r' \ r \ repL w r'" unfolding repL_def using order_trans by auto text\repP (l,u) r iff l and u encode lower and upper bounds of r\ definition repP ::"word * word \ real \ bool" (infix "\\<^sub>P" 10) where "repP w r \ let (w1, w2) = w in repL w1 r \ repU w2 r" lemma int_not_posinf: assumes b1:"real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:"real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "ra \ POS_INF" using b1 b2 by auto lemma int_not_neginf: assumes b1:" real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:" real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "ra \ NEG_INF" using b1 b2 by auto lemma int_not_undef: assumes b1:"real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:"real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "ra \ NEG_INF-1" using b1 b2 by auto lemma sint_range: assumes b1:"real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:"real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "sint ra \ {i. i > -((2^31)-1) \ i < (2^31)-1}" using b1 b2 by auto lemma word_size_neg: fixes w :: "32 Word.word" shows "size (-w) = size w" using Word.word_size[of w] Word.word_size[of "-w"] by auto lemma uint_distinct: fixes w1 w2 shows "w1 \ w2 \ uint w1 \ uint w2" by auto section \Preliminary lemmas\ subsection \Case analysis lemmas\ text\Case analysis principle for pairs of intervals, used in proofs of arithmetic operations\ lemma ivl_zero_case: fixes l1 u1 l2 u2 :: real assumes ivl1:"l1 \ u1" assumes ivl2:"l2 \ u2" shows "(l1 \ 0 \ 0 \ u1 \ l2 \ 0 \ 0 \ u2) \(l1 \ 0 \ 0 \ u1 \ 0 \ l2) \(l1 \ 0 \ 0 \ u1 \ u2 \ 0) \(0 \ l1 \ l2 \ 0 \ 0 \ u2) \(u1 \ 0 \ l2 \ 0 \ 0 \ u2) \(u1 \ 0 \ u2 \ 0) \(u1 \ 0 \ 0 \ l2) \(0 \ l1 \ u2 \ 0) \(0 \ l1 \ 0 \ l2)" using ivl1 ivl2 by (metis le_cases) lemma case_ivl_zero [consumes 2, case_names ZeroZero ZeroPos ZeroNeg PosZero NegZero NegNeg NegPos PosNeg PosPos]: fixes l1 u1 l2 u2 :: real shows "l1 \ u1 \ l2 \ u2 \ ((l1 \ 0 \ 0 \ u1 \ l2 \ 0 \ 0 \ u2) \ P) \ ((l1 \ 0 \ 0 \ u1 \ 0 \ l2) \ P) \ ((l1 \ 0 \ 0 \ u1 \ u2 \ 0) \ P) \ ((0 \ l1 \ l2 \ 0 \ 0 \ u2) \ P) \ ((u1 \ 0 \ l2 \ 0 \ 0 \ u2) \ P) \ ((u1 \ 0 \ u2 \ 0) \ P) \ ((u1 \ 0 \ 0 \ l2) \ P) \ ((0 \ l1 \ u2 \ 0) \ P) \ ((0 \ l1 \ 0 \ l2) \ P) \ P" using ivl_zero_case[of l1 u1 l2 u2] by auto lemma case_inf2[case_names PosPos PosNeg PosNum NegPos NegNeg NegNum NumPos NumNeg NumNum]: shows "\w1 w2 P. (w1 = POS_INF \ w2 = POS_INF \ P w1 w2) \ (w1 = POS_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 = POS_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 = POS_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = POS_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ P w1 w2" by(auto) lemma case_pu_inf[case_names PosAny AnyPos NegNeg NegNum NumNeg NumNum]: shows "\w1 w2 P. (w1 = POS_INF \ P w1 w2) \ (w2 = POS_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ P w1 w2" by(auto) lemma case_pl_inf[case_names NegAny AnyNeg PosPos PosNum NumPos NumNum]: shows "\w1 w2 P. (w1 = NEG_INF \ P w1 w2) \ (w2 = NEG_INF \ P w1 w2) \ (w1 = POS_INF \ w2 = POS_INF \ P w1 w2) \ (w1 = POS_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = POS_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ P w1 w2" by(auto) lemma word_trichotomy[case_names Less Equal Greater]: fixes w1 w2 :: word shows "(w1 P w1 w2) \ (w1 = w2 \ P w1 w2) \ (w2 P w1 w2) \ P w1 w2" using signed.linorder_cases by auto lemma case_times_inf [case_names PosPos NegPos PosNeg NegNeg PosLo PosHi PosZero NegLo NegHi NegZero LoPos HiPos ZeroPos LoNeg HiNeg ZeroNeg AllFinite]: fixes w1 w2 P assumes pp:"(w1 = POS_INF \ w2 = POS_INF \ P w1 w2)" and np:"(w1 = NEG_INF \ w2 = POS_INF \ P w1 w2)" and pn:"(w1 = POS_INF \ w2 = NEG_INF \ P w1 w2)" and nn:"(w1 = NEG_INF \ w2 = NEG_INF \ P w1 w2)" and pl:"(w1 = POS_INF \ w2 \ NEG_INF \ w2 P w1 w2)" and ph:"(w1 = POS_INF \ w2 \ POS_INF \ 0 P w1 w2)" and pz:"(w1 = POS_INF \ w2 = 0 \ P w1 w2)" and nl:"(w1 = NEG_INF \ w2 \ NEG_INF \ w2 P w1 w2)" and nh:"(w1 = NEG_INF \ w2 \ POS_INF \ 0 P w1 w2)" and nz:"(w1 = NEG_INF \ 0 = w2 \ P w1 w2)" and lp:"(w1 \ NEG_INF \ w1 w2 = POS_INF \ P w1 w2)" and hp:"(w1 \ POS_INF \ 0 w2 = POS_INF \ P w1 w2)" and zp:"(0 = w1 \ w2 = POS_INF \ P w1 w2)" and ln:"(w1 \ NEG_INF \ w1 w2 = NEG_INF \ P w1 w2)" and hn:"(w1 \ POS_INF \ 0 w2 = NEG_INF \ P w1 w2)" and zn:"(0 = w1 \ w2 = NEG_INF \ P w1 w2)" and allFinite:"w1 \ NEG_INF \ w1 \ POS_INF \ w2 \ NEG_INF \ w2 \ POS_INF \ P w1 w2" shows " P w1 w2" proof (cases rule: word_trichotomy[of w1 0]) case Less then have w1l:"w1 Trivial arithmetic lemmas\ lemma max_diff_pos:"0 \ 9223372034707292161 + ((-(2 ^ 31))::real)" by auto lemma max_less:"2 ^ 31 < (9223372039002259455::int)" by auto lemma sints64:"sints 64 = {i. - (2 ^ 63) \ i \ i < 2 ^ 63}" using sints_def[of 64] range_sbintrunc[of 63] by auto lemma sints32:"sints 32 = {i. - (2 ^ 31) \ i \ i < 2 ^ 31}" using sints_def[of 32] range_sbintrunc[of 31] by auto lemma upcast_max:"sint((scast(0x80000001::word))::64 Word.word)=sint((0x80000001::32 Word.word))" by auto lemma upcast_min:"(0xFFFFFFFF80000001::64 Word.word) = ((scast (-0x7FFFFFFF::word))::64 Word.word)" by auto lemma min_extend_neg:"sint ((0xFFFFFFFF80000001)::64 Word.word) < 0" by auto lemma min_extend_val':"sint ((-0x7FFFFFFF)::64 Word.word) = (-0x7FFFFFFF)" by auto lemma min_extend_val:"(-0x7FFFFFFF::64 Word.word) = 0xFFFFFFFF80000001" by auto lemma range2s:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto section \Arithmetic operations\ text\This section defines operations which conservatively compute upper and lower bounds of arithmetic functions given upper and lower bounds on their arguments. Each function comes with a proof that it rounds in the advertised direction. \ subsection \Addition upper bound\ text\Upper bound of w1 + w2\ fun pu :: "word \ word \ word" where "pu w1 w2 = (if w1 = POS_INF then POS_INF else if w2 = POS_INF then POS_INF else if w1 = NEG_INF then (if w2 = NEG_INF then NEG_INF else (let sum::64 Word.word = ((scast w2)::64 Word.word) + ((scast NEG_INF)::64 Word.word) in if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum)) else if w2 = NEG_INF then (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast NEG_INF)::64 Word.word) in if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum) else (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s (sum::64 Word.word) then POS_INF else if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum))" lemma scast_down_range: fixes w::"'a::len Word.word" assumes "sint w \ sints (len_of (TYPE('b::len)))" shows "sint w = sint ((scast w)::'b Word.word)" using word_sint.Abs_inverse [OF assms] by simp lemma pu_lemma: fixes w1 w2 fixes r1 r2 :: real assumes up1:"w1 \\<^sub>U (r1::real)" assumes up2:"w2 \\<^sub>U (r2::real)" shows "pu w1 w2 \\<^sub>U (r1 + r2)" proof - have scast_eq1:"sint((scast w1)::64 Word.word) = sint w1" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have scast_eq2:"sint((scast (0x80000001::word))::64 Word.word) = sint ((0x80000001::32 Word.word))" by auto have scast_eq3:"sint((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have w2Geq:"sint ((scast w2)::64 Word.word) \ - (2 ^ 31) " using word_sint.Rep[of "(w2)::32 Word.word"] sints32 Word.word_size scast_eq1 upcast_max scast_eq3 len32 mem_Collect_eq by auto have "sint ((scast w2)::64 Word.word) \ 2 ^ 31" apply (auto simp add: word_sint.Rep[of "(w2)::32 Word.word"] sints32 scast_eq3 len32) using word_sint.Rep[of "(w2)::32 Word.word"] len32[of "TYPE(32)"] sints32 by auto then have w2Less:"sint ((scast w2)::64 Word.word) < 9223372039002259455" by auto have w2Range: "-(2 ^ (size ((scast w2)::64 Word.word) - 1)) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size scast_eq1 upcast_max sints64 max_less) using max_diff_pos max_less w2Less w2Geq by auto have w2case1a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w2Range]) have w1Lower:"sint ((scast w1)::64 Word.word) \ - (2 ^ 31) " using word_sint.Rep[of "(w1)::32 Word.word"] sints32 Word.word_size scast_eq1 scast_eq2 scast_eq3 len32 mem_Collect_eq by auto have w1Leq:"sint ((scast w1)::64 Word.word) \ 2 ^ 31" apply (auto simp add: word_sint.Rep[of "(w1)::32 Word.word"] sints32 scast_eq1 len32) using word_sint.Rep[of "(w1)::32 Word.word"] len32[of "TYPE(32)"] sints32 by auto then have w1Less:"sint ((scast w1)::64 Word.word) < 9223372039002259455" using max_less by auto have w1MinusBound:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size[of "(scast w1)::64 Word.word"] Word.word_size[of "(scast (-0x7FFFFFFF))::64 Word.word"] scast_eq3 scast_eq2 word_sint.Rep[of "(w1)::32 Word.word"] word_sint.Rep[of "0x80000001::32 Word.word"] word_sint.Rep[of "(scast w1)::64 Word.word"] word_sint.Rep[of "-0x7FFFFFFF::64 Word.word"] sints64 sints32) using w1Lower w1Less by auto have w1case1a:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w1)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by (rule signed_arith_sint(1)[of "(scast w1)::64 Word.word" "(- 0x7FFFFFFF)", OF w1MinusBound]) have w1case1a':"sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) = sint ((scast w1)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word)" using min_extend_val w1case1a by auto have w1Leq':"sint w1 \ 2^31 - 1" using word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have neg64:"(((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) = ((scast w2)::64 Word.word) + (-0x7FFFFFFF)" by auto have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto obtain r'\<^sub>1 and r'\<^sub>2 where geq1:"r'\<^sub>1\r1" and equiv1:"w1 \\<^sub>E r'\<^sub>1" and geq2:"r'\<^sub>2\r2" and equiv2:"w2 \\<^sub>E r'\<^sub>2" using up1 up2 unfolding repU_def by auto show ?thesis proof (cases rule: case_pu_inf[where ?w1.0=w1, where ?w2.0=w2]) case PosAny then show ?thesis apply (auto simp add: repU_def repe.simps) using linear by blast next case AnyPos then show ?thesis apply (auto simp add: repU_def repe.simps) using linear by blast next case NegNeg then show ?thesis using up1 up2 by (auto simp add: repU_def repe.simps) next case NegNum assume neq1:"w2 \ POS_INF" assume eq2:"w1 = NEG_INF" assume neq3:"w2 \ NEG_INF" let ?sum = "(scast w2 + scast NEG_INF)::64 Word.word" have leq1:"r'\<^sub>1 \ (real_of_int (sint NEG_INF))" using equiv1 neq1 eq2 neq3 by (auto simp add: repe.simps) have leq2:"r'\<^sub>2 = (real_of_int (sint w2))" using equiv2 neq1 eq2 neq3 by (auto simp add: repe.simps) have case1:"?sum <=s ((scast NEG_INF)::64 Word.word) \ NEG_INF \\<^sub>U r1 + r2" using up1 up2 apply (simp add: repU_def repe.simps word_sle_eq) apply (rule exI [where x= "r1 + r2"]) apply auto using w2case1a apply (auto simp add: eq2 scast_eq3) subgoal for r' proof - assume \r1 \ r'\ \r' \ - 2147483647\ \r2 \ signed w2\ \sint w2 \ 0\ from \sint w2 \ 0\ have \real_of_int (sint w2) \ real_of_int 0\ by (simp only: of_int_le_iff) then have \signed w2 \ (0::real)\ by simp from \r1 \ r'\ \r' \ - 2147483647\ have \r1 \ - 2147483647\ by (rule order_trans) moreover from \r2 \ signed w2\ \signed w2 \ (0::real)\ have \r2 \ 0\ by (rule order_trans) ultimately show \r1 + r2 \ - 2147483647\ by simp qed done have case2:"\(?sum <=s scast NEG_INF) \ scast ?sum \\<^sub>U r1 + r2" apply(simp add: repU_def repe.simps word_sle_def up1 up2) apply(rule exI[where x= "r'\<^sub>2 - 0x7FFFFFFF"]) apply(rule conjI) subgoal proof - assume " \ sint (scast w2 + 0xFFFFFFFF80000001) \ - 2147483647" have bound1:"r1 \ - 2147483647" using leq1 geq1 by (auto) have bound2:"r2 \ r'\<^sub>2" using leq2 geq2 by auto show "r1 + r2 \ r'\<^sub>2 - 2147483647" using bound1 bound2 by(linarith) qed apply(rule disjI2) apply(rule disjI2) apply(auto simp add: not_le) subgoal proof - assume a:"sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) > - 2147483647" then have sintw2_bound:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding min_extend_val by auto have case1a:" sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w2Range]) have "- 0x7FFFFFFF < sint w2 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq3 by linarith then have w2bound:"0 < sint w2" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" using case1a scast_eq3 min_extend_val' word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word)) " using scast_down_range[OF rightSize] by auto then have b:"sint ((scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001))::word) = sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001)" using min_extend_val by auto have c:"sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) = sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word)" using min_extend_val case1a by auto show \r'\<^sub>2 - 2147483647 = signed (SCAST(64 \ 32) (SCAST(32 \ 64) w2 + 0xFFFFFFFF80000001))\ using a b min_extend_val' scast_eq3 leq2 case1a [symmetric] apply (simp add: algebra_simps) apply transfer apply simp done qed subgoal proof - have range2a:" - (2 ^ (size ((scast w2)::64 Word.word) - 1)) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size scast_eq1 upcast_max sints64 sints32 max_less) using max_diff_pos max_less w2Geq w2Less by auto have case2a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF range2a]) have neg64:"(((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) = ((scast w2)::64 Word.word) + (-0x7FFFFFFF)" by auto assume "sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) > - 2147483647" then have sintw2_bound:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 by auto have a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((-0x7FFFFFFF)::64 Word.word)" using case2a by auto have b:"sint ((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have d:"sint w2 \ 2^31 - 1" using word_sint.Rep[of "(w2)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have "- 0x7FFFFFFF < sint w2 + (- 0x7FFFFFFF)" using sintw2_bound case2a min_extend_val' scast_eq3 by linarith then have w2bound:"0 < sint w2" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case2a b min_extend_val' using word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto have "sint (scast (((scast w2)::64 Word.word) + (-0x7FFFFFFF))::word) < 2147483647" unfolding downcast a b min_extend_val' using range2s[of "sint w2", OF d] by auto then show "sint (scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001)::word) < 2147483647" by auto qed subgoal proof - assume notLeq:"sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) > - 2147483647" then have gr:"sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) > - 2147483647" by auto have case2a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w2Range]) from neg64 have sintw2_bound:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 using notLeq by auto have a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((-0x7FFFFFFF)::64 Word.word)" using case2a by auto have c:"sint((-0x7FFFFFFF)::64 Word.word) = -0x7FFFFFFF" by auto have d:"sint w2 \ 2^31 - 1" using word_sint.Rep[of "(w2)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have "- 0x7FFFFFFF < sint w2 + (- 0x7FFFFFFF)" using sintw2_bound case2a c scast_eq3 by linarith then have w2bound:"0 < sint w2" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case2a scast_eq3 using word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto have sintEq:" sint ((scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001))::word) = sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) " using downcast by auto show "- 2147483647 < sint (SCAST(64 \ 32) (SCAST(32 \ 64) w2 + 0xFFFFFFFF80000001))" unfolding sintEq using gr of_int_less_iff of_int_minus of_int_numeral by linarith qed done have castEquiv:"\(?sum <=s scast NEG_INF) \ (scast ?sum) \\<^sub>U r1 + r2" using up1 up2 case1 case2 by fastforce have letRep:"(let sum = ?sum in if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>U r1 + r2" using case1 case2 by(cases "?sum <=s scast NEG_INF"; auto) show "pu w1 w2 \\<^sub>U r1 + r2" using letRep eq2 neq1 by(auto) next case NumNeg assume neq3:"w1 \ NEG_INF" assume neq1:"w1 \ POS_INF" assume eq2:"w2 = NEG_INF" let ?sum = "(scast w1 + scast NEG_INF)::64 Word.word" have case1:"?sum <=s ((scast NEG_INF)::64 Word.word) \ NEG_INF \\<^sub>U r1 + r2" using up1 up2 apply (simp add: repU_def repe.simps word_sle_def) apply(rule exI[where x= "r1 + r2"]) apply(auto) using w1case1a min_extend_neg apply (auto simp add: neq1 eq2 neq3 repINT repU_def repe.simps repeInt_simps up2 word_sless_alt) using repINT repU_def repe.simps repeInt_simps up2 word_sless_alt proof - fix r' assume a1:"sint ((scast w1)::64 Word.word) \ 0" then have "sint w1 \ 0" using scast_eq1 by auto then have h3: \signed w1 \ (0::real)\ by transfer simp assume a2:"r2 \ r'" assume a3:"r1 \ signed w1" assume a4:"r' \ (- 2147483647)" from a2 a4 have h1:"r2 \ - 2147483647" by auto from a1 a3 h3 have h2:"r1 \ 0" using dual_order.trans of_int_le_0_iff by blast show "r1 + r2 \ (- 2147483647)" using h1 h2 add.right_neutral add_mono by fastforce qed have leq1:"r'\<^sub>2 \ (real_of_int (sint NEG_INF))" and leq2:"r'\<^sub>1 = (real_of_int (sint w1))" using equiv1 equiv2 neq1 eq2 neq3 unfolding repe.simps by auto have case1a:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w1)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w1MinusBound]) have case2:"\(?sum <=s scast NEG_INF) \ scast ?sum \\<^sub>U r1 + r2" apply (simp add: repU_def repe.simps word_sle_def up1 up2) apply(rule exI[where x= "r'\<^sub>1 - 0x7FFFFFFF"]) (*r1 + r2*) apply(rule conjI) subgoal using leq1 leq2 geq1 geq2 by auto apply(rule disjI2) apply(rule disjI2) apply(auto) subgoal proof - have f:"r'\<^sub>1 = (real_of_int (sint w1))" by (simp add: leq1 leq2 ) assume a:"\ sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding min_extend_val by auto have "- 0x7FFFFFFF < sint w1 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq1 by linarith then have w2bound:"0 < sint w1" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding w1case1a using w2bound word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"] scast_eq1) have downcast:"sint ((scast (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto then have "sint ((scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001))::word) = sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001)" using min_extend_val by auto then have \signed (SCAST(64 \ 32) (SCAST(32 \ 64) w1 + 0xFFFFFFFF80000001)) = (signed (SCAST(32 \ 64) w1 + 0xFFFFFFFF80000001) :: real)\ by transfer simp moreover have "r'\<^sub>1 - (real_of_int 2147483647) = (real_of_int (sint ((scast w1)::64 Word.word ) - 2147483647))" by (simp add: scast_eq1 leq2) moreover from w1case1a' have \signed (SCAST(32 \ 64) w1 + 0xFFFFFFFF80000001) = signed (SCAST(32 \ 64) w1) + (signed (- 0x7FFFFFFF :: 64 Word.word) :: real)\ by transfer simp ultimately show "r'\<^sub>1 - 2147483647 = (signed ((scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001))::word))" by simp qed subgoal proof - assume "\ sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 by auto have "- 0x7FFFFFFF < sint w1 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq1 by linarith then have w2bound:"0 < sint w1" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a scast_eq1 w1case1a' using word_sint.Rep[of "(w1)::32 Word.word"] w2bound by(auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word)) " using scast_down_range[OF rightSize] by auto show "sint (scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001)::word) < 2147483647" using downcast min_extend_val' w1case1a' scast_eq1 arith[of "sint w1", OF w1Leq'] by auto qed subgoal proof - assume notLeq:"\ sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have gr:"sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) > - 2147483647" by auto then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 using notLeq by auto have "- 0x7FFFFFFF < sint w1 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq1 by linarith then have w2bound:"0 < sint w1" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a scast_eq1 w1case1a' using word_sint.Rep[of "(w1)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) show "- 2147483647 < sint ((scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001))::word)" using scast_down_range[OF rightSize] gr of_int_less_iff of_int_minus of_int_numeral by simp qed done have letUp:"(let sum=?sum in if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>U r1+r2" using case1 case2 by (auto simp add: Let_def) have puSimp:"pu w1 w2=(let sum = ?sum in if sum <=s scast NEG_INF then NEG_INF else scast sum)" using neq3 neq1 eq2 equiv1 leq2 repeInt_simps by force then show "pu w1 w2 \\<^sub>U r1 + r2" using letUp puSimp by auto next case NumNum assume notinf1:"w1 \ POS_INF" assume notinf2:"w2 \ POS_INF" assume notneginf1:"w1 \ NEG_INF" assume notneginf2:"w2 \ NEG_INF" let ?sum = "((scast w1)::64 Word.word) + ((scast w2):: 64 Word.word)" have inf_case:"scast POS_INF <=s ?sum \ POS_INF \\<^sub>U r1 + r2" using repU_def repePos_simps by (meson dual_order.strict_trans not_less order_refl) have truth:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" using Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast w1)::64 Word.word"] scast_eq1 scast_eq3 word_sint.Rep[of "(w1)::32 Word.word"] word_sint.Rep[of "(w2)::32 Word.word"] word_sint.Rep[of "(scast w1)::64 Word.word"] word_sint.Rep[of "(scast w2)::64 Word.word"] sints64 sints32 by auto have sint_eq:"sint((scast w1 + scast w2)::64 Word.word) = sint w1 + sint w2" using signed_arith_sint(1)[of "(scast w1)::64 Word.word" "(scast w2)::64 Word.word", OF truth] scast_eq1 scast_eq3 by auto have bigOne:"scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word) \ \r'\r1 + r2. r' \ (- 0x7FFFFFFF)" proof - assume "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" then have "sint w1 + sint w2 \ - 0x7FFFFFFF" using sint_eq unfolding word_sle_eq by auto then have sum_leq: \real_of_int (sint w1 + sint w2) \ real_of_int (- 0x7FFFFFFF)\ by (simp only: of_int_le_iff) obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using up1 up2 unfolding repU_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using \scast w1 + scast w2 <=s - 0x7FFFFFFF\ word_sle_def notinf1 notinf2 bound1 bound2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce show "\r'\r1 + r2. r' \ (- 0x7FFFFFFF)" apply(rule exI[where x = "r'\<^sub>1 + r'\<^sub>2"]) using bound1 bound2 add_mono something sum_leq apply (auto intro: order_trans [of _ \signed_real_of_word w1 + signed_real_of_word w2\]) done qed have anImp:"\r'. (r'\r1 + r2 \ r' \ (- 2147483647)) \ (\r. - (2 ^ 31 - 1) = - (2 ^ 31 - 1) \ r' = r \ r \ (real_of_int (sint ((- (2 ^ 31 - 1))::32 Word.word))))" by auto have anEq:"((scast ((- (2 ^ 31 - 1))::32 Word.word))::64 Word.word) = (- 0x7FFFFFFF)" by auto have bigTwo: "\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \(?sum <=s ((scast NEG_INF)::64 Word.word)) \ \r'\r1 + r2. r' = (real_of_int (sint (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word))) \ (r' < 0x7FFFFFFF \ (-0x7FFFFFFF) < r')" proof - assume "\(((scast POS_INF)::64 Word.word) <=s ?sum)" and "\(?sum <=s ((scast NEG_INF)::64 Word.word))" then have interval_int: "sint w1 + sint w2 < 0x7FFFFFFF" "(- 0x7FFFFFFF) < sint w1 + sint w2" unfolding word_sle_eq POS_INF_def NEG_INF_def using sint_eq by auto then have interval: \real_of_int (sint w1 + sint w2) < real_of_int (0x7FFFFFFF)\ \real_of_int (- 0x7FFFFFFF) < real_of_int (sint w1 + sint w2)\ by (simp_all only: of_int_less_iff) obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using up1 up2 unfolding repU_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using word_sle_def notinf1 notinf2 bound1 bound2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce have "(w1 \\<^sub>E r'\<^sub>1)" using bound1 by auto then have r1w1:"r'\<^sub>1 = (real_of_int (sint w1))" and w1U:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and w1L:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "(w2 \\<^sub>E r'\<^sub>2)" using bound2 by auto then have r2w1:"r'\<^sub>2 = (real_of_int (sint w2))" and w2U:" (real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and w2L:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "sint (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)) = sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)" apply(rule scast_down_range) unfolding sint_eq using sints32 interval_int by auto then have cast_eq:"(sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)) = sint w1 + sint w2" using scast_down_range sints32 interval_int sint_eq by auto from something and cast_eq have r12_sint_scast:"r'\<^sub>1 + r'\<^sub>2 = (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word))" using r1w1 w1U w1L r2w1 w2U w2L by (simp) show ?thesis using bound1 bound2 add_mono r12_sint_scast cast_eq interval \r'\<^sub>1 + r'\<^sub>2 = (real_of_int (sint (scast (scast w1 + scast w2))))\ by simp qed have neg_inf_case:"?sum <=s ((scast ((NEG_INF)::word))::64 Word.word) \ NEG_INF \\<^sub>U r1 + r2" proof (unfold repU_def NEG_INF_def repe.simps) assume "scast w1 + scast w2 <=s ((scast ((- (2 ^ 31 - 1))::word))::64 Word.word)" then have "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" by (metis anEq) then obtain r' where geq:"(r' \ r1 + r2)" and leq:"(r' \ (- 0x7FFFFFFF))" using bigOne by auto show "(\r'\plus r1 r2. (\r. uminus (minus(2 ^ 31) 1) = POS_INF \ r' = r \ (real_of_int (sint POS_INF)) \ r) \ (\r. uminus (minus(2 ^ 31) 1) = uminus (minus(2 ^ 31) 1) \ r' = r \ r \ real_of_int (sint ((uminus (minus(2 ^ 31) 1))::word))) \ (\w. uminus (minus(2 ^ 31) 1) = w \ r' = real_of_int (sint w) \ (real_of_int (sint w)) < (real_of_int (sint POS_INF)) \ less (real_of_int (sint (uminus (minus(2 ^ 31) 1)))) (real_of_int (sint w))))" using leq anImp geq by auto qed have int_case:"\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \ (?sum <=s ((scast NEG_INF)::64 Word.word)) \ ((scast ?sum)::word) \\<^sub>U r1 + r2" proof - assume bound1:"\ ((scast POS_INF)::64 Word.word) <=s scast w1 + scast w2" assume bound2:"\ scast w1 + scast w2 <=s ((scast NEG_INF)::64 Word.word)" obtain r'::real where rDef:"r' = (real_of_int (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)))" and r12:"r'\r1 + r2" and boundU:"r' < 0x7FFFFFFF" and boundL:"(-0x7FFFFFFF) < r'" using bigTwo[OF bound1 bound2] by auto obtain w::word where wdef:"w = (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word)" by auto then have wr:"r' = (real_of_int (sint w))" using r12 bound1 bound2 rDef by blast show ?thesis unfolding repU_def repe.simps using r12 wdef rDef boundU boundL wr by auto qed have almost:"(let sum::64 Word.word = scast w1 + scast w2 in if scast POS_INF <=s sum then POS_INF else if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>U r1 + r2" apply(cases "((scast POS_INF)::64 Word.word) <=s ((?sum)::64 Word.word)") subgoal using inf_case Let_def int_case neg_inf_case by auto apply(cases "?sum <=s scast NEG_INF") subgoal using inf_case Let_def int_case neg_inf_case proof - assume "\ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2" then have "\ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF \ \ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2 \ \ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF \ ((let w = scast w1 + scast w2 in if scast POS_INF <=s (w::64 Word.word) then POS_INF else if w <=s scast NEG_INF then NEG_INF else scast w) \\<^sub>U r1 + r2)" using neg_inf_case by presburger then show ?thesis using int_case by force qed subgoal using inf_case Let_def int_case neg_inf_case proof - assume a1: "\ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2" assume "\ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF" have "\ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF \ \ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2 \ ((let w = scast w1 + scast w2 in if scast POS_INF <=s (w::64 Word.word) then POS_INF else if w <=s scast NEG_INF then NEG_INF else scast w) \\<^sub>U r1 + r2)" using a1 neg_inf_case by presburger then show ?thesis using int_case by force qed done then show ?thesis using notinf1 notinf2 notneginf1 notneginf2 by auto qed qed text\Lower bound of w1 + w2\ fun pl :: "word \ word \ word" where "pl w1 w2 = (if w1 = NEG_INF then NEG_INF else if w2 = NEG_INF then NEG_INF else if w1 = POS_INF then (if w2 = POS_INF then POS_INF else (let sum::64 Word.word = ((scast w2)::64 Word.word) + ((scast POS_INF)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s(sum::64 Word.word) then POS_INF else scast sum)) else if w2 = POS_INF then (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast POS_INF)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s(sum::64 Word.word) then POS_INF else scast sum) else (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s (sum::64 Word.word) then POS_INF else if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum))" subsection \Addition lower bound\ text\Correctness of lower bound of w1 + w2\ lemma pl_lemma: assumes lo1:"w1 \\<^sub>L (r1::real)" assumes lo2:"w2 \\<^sub>L (r2::real)" shows "pl w1 w2 \\<^sub>L (r1 + r2)" proof - have scast_eq1:"sint((scast w1)::64 Word.word) = sint w1" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have scast_eq2:"sint((scast (0x80000001::word))::64 Word.word)=sint((0x80000001::32 Word.word))" by auto have scast_eq3:"sint((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have sints64:"sints 64 = {i. - (2 ^ 63) \ i \ i < 2 ^ 63}" using sints_def[of 64] range_sbintrunc[of 63] by auto have sints32:"sints 32 = {i. - (2 ^ 31) \ i \ i < 2 ^ 31}" using sints_def[of 32] range_sbintrunc[of 31] by auto have thing1:"0 \ 9223372034707292161 + ((-(2 ^ 31))::real)" by auto have "sint (( w2)) \ (-(2 ^ 31))" using word_sint.Rep[of "(w2)::32 Word.word"] sints32 mem_Collect_eq Word.word_size[of "(scast w2)::64 Word.word"] scast_eq1 scast_eq2 scast_eq3 len32 by auto then have thing4:"sint ((scast w2)::64 Word.word) \ (-(2 ^ 31))" using scast_down_range sint_ge sints_num using scast_eq3 by linarith have aLeq2:"(-(2 ^ 31)::int) \ -9223372039002259455" by auto then have thing2:" (0::int) \ 9223372039002259455 + sint ((scast w2)::64 Word.word)" using thing4 aLeq2 by (metis ab_group_add_class.ab_left_minus add.commute add_mono neg_le_iff_le) have aLeq:"2 ^ 31 \ (9223372039002259455::int)" by auto have bLeq:"sint ((scast w2)::64 Word.word) \ 2 ^ 31" apply ( auto simp add: word_sint.Rep[of "(w2)::32 Word.word"] sints32 scast_eq3 len32) using word_sint.Rep[of "(w2)::32 Word.word"] len32[of "TYPE(32)"] sints32 by auto have thing3:" sint ((scast w2)::64 Word.word) \ 9223372034707292160 " using aLeq bLeq by auto have truth:" - (2 ^ (size ((scast w2)::64 Word.word) - 1)) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" by(auto simp add: Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast (0x7FFFFFFF))::64 Word.word"] scast_eq1 scast_eq2 sints64 sints32 thing2 thing1 thing3) have case1a:" sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have case1b:"sint ((0xFFFFFFFF80000001)::64 Word.word) < 0" by auto have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto have neg64:"(((scast w2)::64 Word.word) + 0x7FFFFFFF) = ((scast w2)::64 Word.word) + (0x7FFFFFFF)" by auto obtain r'\<^sub>1 and r'\<^sub>2 where geq1:"r'\<^sub>1\r1" and equiv1:"w1 \\<^sub>E r'\<^sub>1" and geq2:"r'\<^sub>2\r2" and equiv2:"w2 \\<^sub>E r'\<^sub>2" using lo1 lo2 unfolding repL_def by auto show ?thesis proof (cases rule: case_pl_inf[where ?w1.0=w1, where ?w2.0=w2]) case NegAny then show ?thesis apply (auto simp add: repL_def repe.simps) using lo1 lo2 linear by auto next case AnyNeg then show ?thesis apply (auto simp add: repL_def repe.simps) using linear by auto next case PosPos then show ?thesis using lo1 lo2 by (auto simp add: repL_def repe.simps) next case PosNum assume neq1:"w2 \ POS_INF" assume eq2:"w1 = POS_INF" assume neq3:"w2 \ NEG_INF" let ?sum = "(scast w2 + scast POS_INF)::64 Word.word" have case1:"(((scast POS_INF)::64 Word.word) <=s ?sum) \ POS_INF \\<^sub>L r1 + r2" using lo1 lo2 apply (simp add: repL_def repe.simps word_sle_def) apply(rule exI[where x= "r1 + r2"]) using case1a case1b apply (auto simp add: neq1 eq2 neq3 repINT repL_def repe.simps repeInt_simps lo2 word_sless_alt) proof - fix r' assume a1:"0 \ sint (((scast w2)::64 Word.word))" from a1 have h3:"2147483647 \ sint w2 + 0x7FFFFFFF " using scast_eq3 by auto assume a2:"r' \ r1" assume a3:"signed w2 \ r2" assume a4:"(2147483647) \ r'" from a2 a4 have h1:"2147483647 \ r1" by auto from a1 a3 h3 have h2:"0 \ r2" using of_int_le_0_iff le_add_same_cancel2 apply simp apply transfer apply simp apply (metis (full_types) of_int_0 of_int_le_iff order_trans signed_take_bit_nonnegative_iff) done show "(2147483647) \ r1 + r2 " using h1 h2 h3 add.right_neutral add_mono by fastforce qed have leq1:"r'\<^sub>1 \ (real_of_int (sint POS_INF))" using equiv1 neq1 eq2 neq3 unfolding repe.simps by auto have leq2:"r'\<^sub>2 = (real_of_int (sint w2))" using equiv2 neq1 eq2 neq3 unfolding repe.simps by auto have case2:"\(scast POS_INF <=s ?sum) \ scast ?sum \\<^sub>L r1 + r2" apply (simp add: repL_def repe.simps word_sle_def lo1 lo2) apply(rule exI[where x= "r'\<^sub>2 + 0x7FFFFFFF"]) (*r1 + r2*) apply(rule conjI) subgoal proof - assume "\ 2147483647 \ sint (scast w2 + 0x7FFFFFFF)" have bound1:"2147483647 \ r1" using leq1 geq1 by (auto) have bound2:"r'\<^sub>2 \ r2 " using leq2 geq2 by auto show "r'\<^sub>2 + 2147483647 \ r1 + r2" using bound1 bound2 by linarith qed apply(rule disjI2) apply(rule disjI2) apply(auto) subgoal proof - assume a:"\ 2147483647 \ sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"2147483647 > sint (((scast w2)::64 Word.word) + (0x7FFFFFFF))" by auto have case1a:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have a1:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((0x7FFFFFFF)::64 Word.word)" using case1a by auto have c1:"sint((0x7FFFFFFF)::64 Word.word) = 0x7FFFFFFF" by auto have "sint w2 + ( 0x7FFFFFFF) < 0x7FFFFFFF" using sintw2_bound case1a c1 scast_eq3 by linarith then have w2bound:"sint w2 < 0" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a scast_eq3 c1 using word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto then have b:"sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word) = sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" by auto have c:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) = sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word)" using case1a by auto have d:"sint ((0x7FFFFFFF)::64 Word.word) = (0x7FFFFFFF)" by auto have f:"r'\<^sub>2 = (real_of_int (sint w2))" by (simp add: leq2) show "r'\<^sub>2 + 2147483647 = (signed ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word))" using a b c d scast_eq3 f leq2 of_int_numeral by auto qed subgoal proof - have truth2a:"-(2^(size ((scast w2)::64 Word.word)-1)) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast (0x7FFFFFFF))::64 Word.word"] scast_eq1 scast_eq2 sints64 sints32 thing2) using thing1 thing2 thing3 by auto have case2a:" sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth2a]) have min_cast:"(0x7FFFFFFF::64 Word.word) =((scast (0x7FFFFFFF::word))::64 Word.word)" by auto assume "\ 2147483647 \ sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"2147483647 > sint (((scast w2)::64 Word.word) + (0x7FFFFFFF))" using neg64 by auto have a:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((0x7FFFFFFF)::64 Word.word)" using case2a by auto have c:"sint((0x7FFFFFFF)::64 Word.word) = 0x7FFFFFFF" by auto have " 0x7FFFFFFF > sint w2 + ( 0x7FFFFFFF)" using sintw2_bound case2a c scast_eq3 by linarith then have w2bound:" sint w2 < 0" by simp have rightSize:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case2a scast_eq3 c apply (auto simp add: sints32 len32[of "TYPE(32)"]) using word_sint.Rep[of "(w2)::32 Word.word"] sints32 len32[of "TYPE(32)"] w2bound by auto have downcast:"sint ((scast (((scast w2)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto then show "sint (scast (((scast w2)::64 Word.word) + 0x7FFFFFFF)::word) < 2147483647" unfolding downcast a scast_eq3 c using w2bound by auto qed subgoal proof - assume notLeq:"\ 2147483647 \ sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" then have gr:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) < 2147483647" by auto have case2a:" sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have min_cast:"(0x7FFFFFFF::64 Word.word) =((scast (0x7FFFFFFF::word))::64 Word.word)" by auto have neg64:"(((scast w2)::64 Word.word) + 0x7FFFFFFF) = ((scast w2)::64 Word.word) + (0x7FFFFFFF)" by auto then have sintw2_bound:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) < 2147483647" using neg64 using notLeq by auto have a:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((0x7FFFFFFF)::64 Word.word)" using case2a by auto have c:"sint((0x7FFFFFFF)::64 Word.word) = 0x7FFFFFFF" by auto have "- 2147483647 \ w2" using neq3 unfolding NEG_INF_def by auto then have "sint((- 2147483647)::word) \ sint w2" using word_sint.Rep_inject by blast then have n1:"- 2147483647 \ sint w2" by auto have "- 2147483648 \ w2" apply(rule repe.cases[OF equiv2]) by auto then have "sint(- 2147483648::word) \ sint w2" using word_sint.Rep_inject by blast then have n2:"- 2147483648 \ sint w2" by auto then have d:"sint w2 > - 2147483647" using word_sint.Rep[of "(w2)::32 Word.word"] sints32 len32[of "TYPE(32)"] neq3 n1 n2 by auto have w2bound:"- 2147483647 < sint w2 + 0x7FFFFFFF" using sintw2_bound case2a c scast_eq3 d by linarith have rightSize:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" using sints32 len32[of "TYPE(32)"] w2bound word_sint.Rep[of "(w2)::32 Word.word"] c case2a scast_eq3 sintw2_bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto have sintEq:" sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word) = sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) " using downcast by auto show "- 2147483647 < sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word)" unfolding sintEq using gr of_int_less_iff of_int_minus of_int_numeral c case2a scast_eq3 w2bound by simp qed done have "(let sum = ?sum in if scast POS_INF <=s sum then POS_INF else scast sum) \\<^sub>L r1 + r2" using case1 case2 by (auto simp add: Let_def) then show ?thesis using lo1 lo2 neq1 eq2 neq3 by (auto) next case NumPos assume neq3:"w1 \ NEG_INF" assume neq1:"w1 \ POS_INF" assume eq2:"w2 = POS_INF" let ?sum = "(scast w1 + scast POS_INF)::64 Word.word" have thing1:"0 \ 9223372034707292161 + ((-(2 ^ 31))::real)" by auto have "sint (( w1)) \ (-(2 ^ 31))" using word_sint.Rep[of "(w1)::32 Word.word"] scast_eq1 scast_eq2 scast_eq3 Word.word_size[of "(scast w1)::64 Word.word"] sints32 len32 mem_Collect_eq by auto then have thing4:"sint ((scast w1)::64 Word.word) \ (-(2 ^ 31))" using scast_down_range sint_ge sints_num scast_eq3 scast_eq1 by linarith have aLeq2:"(-(2 ^ 31)::int) \ -9223372039002259455" by auto then have thing2:" (0::int) \ 9223372039002259455 + sint ((scast w1)::64 Word.word)" using thing4 aLeq2 by (metis ab_group_add_class.ab_left_minus add.commute add_mono neg_le_iff_le) have aLeq:"2 ^ 31 \ (9223372039002259455::int)" by auto have bLeq:"sint ((scast w1)::64 Word.word) \ 2 ^ 31" apply (auto simp add: word_sint.Rep[of "(w1)::32 Word.word"] sints32 scast_eq1 len32) using word_sint.Rep[of "(w1)::32 Word.word"] len32[of "TYPE(32)"] sints32 by clarsimp have thing3:" sint ((scast w1)::64 Word.word) \ 9223372034707292160 " using aLeq bLeq by auto have truth:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" by(auto simp add: Word.word_size[of "(scast w1)::64 Word.word"] Word.word_size[of "(scast (0x7FFFFFFF))::64 Word.word"] scast_eq3 scast_eq2 sints64 sints32 thing2 thing1 thing3) have case1a:"sint (((scast w1)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w1)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have case1b:"sint ((0xFFFFFFFF80000001)::64 Word.word) < 0" by auto have g:"(0x7FFFFFFF::64 Word.word) = 0x7FFFFFFF" by auto have c:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) = sint ((scast w1)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word)" using g case1a by blast have d:"sint ((0x7FFFFFFF)::64 Word.word) = (0x7FFFFFFF)" by auto have e:"sint ((scast w1)::64 Word.word) = sint w1" using scast_eq1 by blast have d2:"sint w1 \ 2^31 - 1" using word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have case1:"scast POS_INF <=s ?sum \ POS_INF \\<^sub>L r1 + r2" using lo1 lo2 apply (simp add: repL_def repe.simps word_sle_def) apply(rule exI[where x= "r1 + r2"]) apply(auto) using case1a case1b apply (auto simp add: neq1 eq2 neq3 repINT repL_def repe.simps repeInt_simps lo2 word_sless_alt) (* close 4 goals *) proof - fix r' have h3:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) = sint (((scast w1)::64 Word.word)) + sint(0x7FFFFFFF::64 Word.word)" using case1a by auto have h4:"sint(0x7FFFFFFF::64 Word.word) = 2147483647" by auto assume a1:"0 \ sint ((scast w1)::64 Word.word)" then have h3:"sint w1 \ 0" using scast_eq1 h3 h4 by auto assume a2:"r' \ r2" assume a3:"signed w1 \ r1" assume a4:"(2147483647) \ r'" from a2 a4 have h1:"r2 \ 2147483647" by auto from a3 h3 have h2:"r1 \ 0" by (auto intro: order_trans [of _ \signed_real_of_word w1\]) show " 2147483647 \ r1 + r2" using h1 h2 add.right_neutral add_mono by fastforce qed have leq1:"r'\<^sub>2 \ (real_of_int (sint POS_INF))" and leq2:"r'\<^sub>1 = (real_of_int (sint w1))" using equiv1 equiv2 neq1 eq2 neq3 unfolding repe.simps by auto have neg64:"(((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) = ((scast w1)::64 Word.word) + (-0x7FFFFFFF)" by auto have case2:"\(scast POS_INF <=s ?sum) \ scast ?sum \\<^sub>L r1 + r2" apply (simp add: repL_def repe.simps word_sle_def lo1 lo2) apply(rule exI[where x= "r'\<^sub>1 + 0x7FFFFFFF"]) apply(rule conjI) subgoal proof - assume "\ 2147483647 \ sint (scast w1 + 0x7FFFFFFF)" have bound1:"r2 \ 2147483647" using leq1 geq2 by (auto) have bound2:"r1 \ r'\<^sub>1" using leq2 geq1 by auto show "r'\<^sub>1 + 2147483647 \ r1 + r2" using bound1 bound2 by linarith qed apply(rule disjI2) apply(rule disjI2) apply(auto) subgoal proof - have f:"r'\<^sub>1 = (real_of_int (sint w1))" by (simp add: leq1 leq2 ) assume a:"\ 2147483647 \ sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"2147483647 > sint (((scast w1)::64 Word.word) + (0x7FFFFFFF))" by auto have "0x7FFFFFFF > sint w1 + (0x7FFFFFFF)" using sintw2_bound case1a d scast_eq1 by linarith then have w2bound:"0 > sint w1" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a e using w2bound word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"] ) have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto have downcast:"sint ((scast (((scast w1)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w1)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto then have b:"sint((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word) = sint(((scast w1)::64 Word.word) + 0x7FFFFFFF)" using g by auto show "r'\<^sub>1 + 2147483647 = ((signed_real_of_word ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word)))" using a b c d e f proof - have "(real_of_int (sint ((scast w1)::64 Word.word ) + 2147483647)) = r'\<^sub>1 + (real_of_int 2147483647)" using e leq2 by auto from this [symmetric] show ?thesis using b c d of_int_numeral by simp qed qed subgoal proof - assume "\ 2147483647 \ sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (0x7FFFFFFF)) < 2147483647" unfolding neg64 by auto have "0x7FFFFFFF > sint w1 + (0x7FFFFFFF)" using sintw2_bound case1a d scast_eq1 by linarith then have w2bound:"0 > sint w1" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a e c using word_sint.Rep[of "(w1)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto have downcast:"sint ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word) = sint (((scast w1)::64 Word.word) + ((0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto show "sint (scast (((scast w1)::64 Word.word) + 0x7FFFFFFF)::word) < 2147483647" using downcast d e c arith[of "sint w1", OF d2] sintw2_bound by linarith qed subgoal proof - assume notLeq:"\ 2147483647 \ sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" then have gr:"2147483647 > sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" by auto then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (0x7FFFFFFF)) < 2147483647" unfolding neg64 using notLeq by auto have "0x7FFFFFFF > sint w1 + ( 0x7FFFFFFF)" using sintw2_bound case1a d scast_eq1 by linarith then have useful:"0 > sint w1" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a e using word_sint.Rep[of "(w1)::32 Word.word"] sints32 len32[of "TYPE(32)"] useful by auto have "- 2147483647 \ w1" using neq3 unfolding NEG_INF_def by auto then have "sint((- 2147483647)::word) \ sint w1" using word_sint.Rep_inject by blast then have n1:"- 2147483647 \ sint w1" by auto have "- 2147483648 \ w1" apply(rule repe.cases[OF equiv1]) using int_not_undef[of w1] by auto then have "sint(- 2147483648::word) \ sint w1" using word_sint.Rep_inject by blast then have n2:"- 2147483648 \ sint w1" by auto then have d:"sint w1 > - 2147483647" using word_sint.Rep[of "(w1)::32 Word.word"] sints32 len32[of "TYPE(32)"] n1 n2 neq3 by (simp) have d2:"sint (0x7FFFFFFF::64 Word.word) > 0" by auto from d d2 have d3:"- 2147483647 < sint w1 + sint (0x7FFFFFFF::64 Word.word)" by auto have d4:"sint ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word) = sint w1 + sint (0x7FFFFFFF::64 Word.word)" using case1a rightSize scast_down_range scast_eq1 by fastforce then show "- 2147483647 < sint (SCAST(64 \ 32) (SCAST(32 \ 64) w1 + 0x7FFFFFFF))" using d3 d4 by auto qed done have "(let sum = ?sum in if scast POS_INF <=s sum then POS_INF else scast sum) \\<^sub>L r1 + r2" using case1 case2 by (auto simp add: Let_def) then show ?thesis using neq1 eq2 neq3 by (auto) next case NumNum assume notinf1:"w1 \ POS_INF" assume notinf2:"w2 \ POS_INF" assume notneginf1:"w1 \ NEG_INF" assume notneginf2:"w2 \ NEG_INF" let ?sum = "((scast w1)::64 Word.word) + ((scast w2):: 64 Word.word)" have truth:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" using Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast w1)::64 Word.word"] scast_eq1 scast_eq3 sints64 sints32 word_sint.Rep[of "(w1)::32 Word.word"] word_sint.Rep[of "(w2)::32 Word.word"] by auto have sint_eq:"sint((scast w1 + scast w2)::64 Word.word) = sint w1 + sint w2" using signed_arith_sint(1)[of "(scast w1)::64 Word.word" "(scast w2)::64 Word.word", OF truth] scast_eq1 scast_eq3 by auto have bigOne:"scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word) \ \r'\r1 + r2. r' \ -0x7FFFFFFF" proof - assume "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" then have sum_leq:"sint w1 + sint w2 \ - 0x7FFFFFFF" and sum_leq':" (sint w1 + sint w2) \ (- 2147483647)" using sint_eq unfolding word_sle_eq by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using lo1 lo2 unfolding repL_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using bound1 bound2 \scast w1 + scast w2 <=s -0x7FFFFFFF\ word_sle_def notinf1 notinf2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce show "\r'\r1 + r2. r' \ (-0x7FFFFFFF)" apply (rule exI[where x = "r'\<^sub>1 + r'\<^sub>2"]) using bound1 bound2 add_mono something sum_leq' apply (auto intro: order_trans [of _ \signed_real_of_word w1 + signed_real_of_word w2\]) done qed have anImp:"\r'. (r'\r1 + r2 \ r' \ (- 2147483647)) \ (\r. - (2 ^ 31 - 1) = - (2 ^ 31 - 1) \ r' = r \ r \ (real_of_int (sint ((- (2 ^ 31 - 1))::32 Word.word))))" by auto have anEq:"((scast ((- (2 ^ 31 - 1))::32 Word.word))::64 Word.word) = (- 0x7FFFFFFF)" by auto have bigTwo: "\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \(?sum <=s ((scast NEG_INF)::64 Word.word)) \ \r'\r1 + r2. r' = (real_of_int (sint (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word))) \ (r' < 0x7FFFFFFF \ (-0x7FFFFFFF) < r')" proof - assume "\(((scast POS_INF)::64 Word.word) <=s ?sum)" then have sum_leq:"sint w1 + sint w2 < 0x7FFFFFFF" unfolding word_sle_eq using sint_eq by auto then have sum_leq':" (sint w1 + sint w2) < (2147483647)" by auto assume "\(?sum <=s ((scast NEG_INF)::64 Word.word))" then have sum_geq:"(- 0x7FFFFFFF) < sint w1 + sint w2" unfolding word_sle_eq using sint_eq by auto then have sum_geq':" (- 2147483647) < (sint w1 + sint w2)" by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using lo1 lo2 unfolding repL_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using word_sle_def notinf1 notinf2 bound1 bound2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce have "(w1 \\<^sub>E r'\<^sub>1)" using bound1 by auto then have r1w1:"r'\<^sub>1 = (real_of_int (sint w1))" and w1U:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and w1L:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "(w2 \\<^sub>E r'\<^sub>2)" using bound2 by auto then have r2w1:"r'\<^sub>2 = (real_of_int (sint w2))" and w2U:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and w2L:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "sint (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)) = sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)" apply(rule scast_down_range) unfolding sint_eq using sints32 sum_geq sum_leq by auto then have cast_eq:"(sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)) = sint w1 + sint w2" using scast_down_range sints32 sum_geq sum_leq sint_eq by auto from something and cast_eq have r12_sint_scast:"r'\<^sub>1 + r'\<^sub>2 = (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word))" using r1w1 w1U w1L r2w1 w2U w2L by (simp) have leq_ref:"\x y ::real. x = y ==> x \ y" by auto show ?thesis apply(rule exI[where x="r'\<^sub>1 + r'\<^sub>2"]) apply(rule conjI) subgoal using r12_sint_scast cast_eq leq_ref r2w1 r1w1 add_mono[of r'\<^sub>1 r1 r'\<^sub>2 r2] bound1 bound2 by auto using bound1 bound2 add_mono r12_sint_scast cast_eq sum_leq sum_leq' sum_geq' sum_geq \r'\<^sub>1 + r'\<^sub>2 = (real_of_int (sint (scast (scast w1 + scast w2))))\ by auto qed have neg_inf_case:"?sum <=s ((scast ((NEG_INF)::word))::64 Word.word) \ NEG_INF \\<^sub>L r1 + r2" proof (unfold repL_def NEG_INF_def repe.simps) assume "scast w1 + scast w2 <=s ((scast ((- (2 ^ 31 - 1))::word))::64 Word.word)" then have "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" by (metis anEq) then obtain r' where geq:"(r' \ r1 + r2)" and leq:"(r' \ (-0x7FFFFFFF))" using bigOne by auto show "(\r'\plus r1 r2. (\r. uminus (minus(2 ^ 31) 1) = POS_INF \ r' = r \ (real_of_int (sint POS_INF)) \ r) \ (\r. uminus (minus(2 ^ 31) 1) = uminus (minus(2 ^ 31) 1) \ r' = r \ r \ (real_of_int (sint ((uminus (minus(2 ^ 31) 1))::word)))) \ (\w. uminus (minus(2 ^ 31) 1) = w \ r' = (real_of_int (sint w)) \ (real_of_int (sint w)) < (real_of_int (sint POS_INF)) \ less ( (real_of_int (sint (uminus (minus(2 ^ 31) 1))))) ((real_of_int (sint w)))))" using leq geq by auto qed have bigThree:"0x7FFFFFFF <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word) \ \r'\r1 + r2. 2147483647 \ r'" proof - assume "0x7FFFFFFF <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" then have sum_leq:"0x7FFFFFFF \ sint w1 + sint w2 " and sum_leq':" 2147483647 \ (sint w1 + sint w2)" using sint_eq unfolding word_sle_eq by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using lo1 lo2 unfolding repL_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using \ 0x7FFFFFFF <=s scast w1 + scast w2 \ word_sle_def notinf1 notinf2 bound1 bound2 repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono of_int_add by fastforce show "\r'\ r1 + r2. (2147483647) \ r'" apply(rule exI[where x = "r'\<^sub>1 + r'\<^sub>2"]) using bound1 bound2 add_mono something sum_leq' order.trans proof - have f1: " (real_of_int (sint w2)) = r'\<^sub>2" by (metis bound2 notinf2 notneginf2 repe.cases) have " (real_of_int (sint w1)) = r'\<^sub>1" by (metis bound1 notinf1 notneginf1 repe.cases) then have f2: " (real_of_int 2147483647) \ r'\<^sub>2 + r'\<^sub>1" using f1 sum_leq' by auto have "r'\<^sub>2 + r'\<^sub>1 \ r2 + r1" by (meson add_left_mono add_right_mono bound1 bound2 order.trans) then show "r'\<^sub>1 + r'\<^sub>2 \ r1 + r2 \ 2147483647 \ r'\<^sub>1 + r'\<^sub>2" using f2 by (simp add: add.commute) qed qed have inf_case:"((scast POS_INF)::64 Word.word) <=s ?sum \ POS_INF \\<^sub>L r1 + r2" proof - assume "((scast POS_INF)::64 Word.word) <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" then have "((scast ((2 ^ 31 - 1)::word))::64 Word.word) <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" unfolding repL_def repe.simps by auto then have "(0x7FFFFFFF::64 Word.word) <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" by auto then obtain r' where geq:"(r' \ r1 + r2)" and leq:"(0x7FFFFFFF \ r')" using bigThree by auto show "?thesis" unfolding repL_def repe.simps using leq geq by auto qed have int_case:"\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \ (?sum <=s ((scast NEG_INF)::64 Word.word)) \ ((scast ?sum)::word) \\<^sub>L r1 + r2" proof - assume bound1:"\ ((scast POS_INF)::64 Word.word) <=s scast w1 + scast w2" assume bound2:"\ scast w1 + scast w2 <=s ((scast NEG_INF)::64 Word.word)" obtain r'::real where rDef:"r' = (real_of_int (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)))" and r12:"r'\r1 + r2" and boundU:"r' < 0x7FFFFFFF" and boundL:" (-0x7FFFFFFF) < r'" using bigTwo[OF bound1 bound2] by auto obtain w::word where wdef:"w = (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word)" by auto then have wr:"r' = (real_of_int (sint w))" using r12 bound1 bound2 rDef by blast show ?thesis unfolding repL_def repe.simps using r12 wdef rDef boundU boundL wr by auto qed have "(let sum = ?sum in if scast POS_INF <=s sum then POS_INF else if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>L r1 + r2" apply(cases "((scast POS_INF)::64 Word.word) <=s ?sum") apply(cases "?sum <=s scast NEG_INF") using inf_case neg_inf_case int_case by (auto simp add: Let_def) then show ?thesis using notinf1 notinf2 notneginf1 notneginf2 by(auto) qed qed subsection \Max function\ text\Maximum of w1 + w2 in 2s-complement\ fun wmax :: "word \ word \ word" where "wmax w1 w2 = (if w1 Correctness of wmax\ lemma wmax_lemma: assumes eq1:"w1 \\<^sub>E (r1::real)" assumes eq2:"w2 \\<^sub>E (r2::real)" shows "wmax w1 w2 \\<^sub>E (max r1 r2)" proof(cases rule: case_inf2[where ?w1.0=w1, where ?w2.0=w2]) case PosPos from PosPos eq1 eq2 have bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2:"(real_of_int (sint POS_INF)) \ r2" by (auto simp add: repe.simps) have eqInf:"wmax w1 w2 = POS_INF" using PosPos unfolding wmax.simps by auto have pos_eq:"POS_INF \\<^sub>E max r1 r2" apply(rule repPOS_INF) using bound1 bound2 by linarith show ?thesis using pos_eq eqInf by auto next case PosNeg from PosNeg have bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2:"r2 \ (real_of_int (sint NEG_INF))" using eq1 eq2 by (auto simp add: repe.simps) have eqNeg:"wmax w1 w2 = POS_INF" unfolding eq1 eq2 wmax.simps PosNeg word_sless_def word_sle_def by(auto) have neg_eq:"POS_INF \\<^sub>E max r1 r2" apply(rule repPOS_INF) using bound1 bound2 eq1 eq2 by auto show "?thesis" using eqNeg neg_eq by auto next case PosNum from PosNum eq1 eq2 have bound1:" (real_of_int (sint POS_INF)) \ r1" and bound2a:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound2b:" (real_of_int (sint w2)) < (real_of_int (sint POS_INF))" by (auto simp add: repe.simps) have eqNeg:"wmax w1 w2 = POS_INF" using PosNum bound2b unfolding wmax.simps word_sless_def word_sle_def by auto have neg_eq:"POS_INF \\<^sub>E max r1 r2" apply (rule repPOS_INF) using bound1 bound2a bound2b word_sless_alt le_max_iff_disj unfolding eq1 eq2 by blast show "?thesis" using eqNeg neg_eq by auto next case NegPos from NegPos eq1 eq2 have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:" (real_of_int (sint POS_INF)) \ r2" by (auto simp add: repe.simps) have eqNeg:"wmax w1 w2 = POS_INF" unfolding NegPos word_sless_def word_sle_def by(auto) have neg_eq:"POS_INF \\<^sub>E max r1 r2" apply(rule repPOS_INF) using bound1 bound2 by auto show "wmax w1 w2 \\<^sub>E max r1 r2" using eqNeg neg_eq by auto next case NegNeg from NegNeg eq1 eq2 have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" by (auto simp add: repe.simps) have eqNeg:"NEG_INF \\<^sub>E max r1 r2" apply(rule repNEG_INF) using eq1 eq2 bound1 bound2 by(auto) have neg_eq:"wmax w1 w2 = NEG_INF" using NegNeg by auto show "wmax w1 w2 \\<^sub>E max r1 r2" using eqNeg neg_eq by auto next case NegNum from NegNum eq1 eq2 have eq3:"r2 = (real_of_int (sint w2))" and bound2a:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound1:"r1 \ (real_of_int (sint NEG_INF))" by (auto simp add: repe.simps) have eqNeg:"max r1 r2 = (real_of_int (sint w2))" using NegNum assms(2) bound2a eq3 repeInt_simps bound1 bound2a bound2b by (metis less_irrefl max.bounded_iff max_def not_less) then have extra_eq:"(wmax w1 w2) = w2" using assms(2) bound2a eq3 NegNum repeInt_simps by (simp add: word_sless_alt) have neg_eq:"wmax w1 w2 \\<^sub>E (real_of_int (sint (wmax w1 w2)))" apply(rule repINT) using extra_eq eq3 bound2a bound2b by(auto) show "wmax w1 w2 \\<^sub>E max r1 r2" using eqNeg neg_eq extra_eq by auto next case NumPos from NumPos eq1 eq2 have p2:"w2 = POS_INF" and eq1:"r1 = (real_of_int (sint w1))" and eq2:"r2 = r2" and bound1a:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:"(real_of_int (sint POS_INF)) \ r2" by (auto simp add: repe.simps) have res1:"wmax w1 w2 = POS_INF" using NumPos p2 eq1 eq2 assms(1) bound1b p2 repeInt_simps by (simp add: word_sless_alt) have res3:"POS_INF \\<^sub>E max r1 r2" using repPOS_INF NumPos bound2 le_max_iff_disj by blast show "wmax w1 w2 \\<^sub>E max r1 r2" using res1 res3 by auto next case NumNeg from NumNeg eq1 eq2 have n2:"w2 = NEG_INF" and rw1:"r1 = (real_of_int (sint w1))" and bound1a:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" by (auto simp add: repe.simps) have res1:"max r1 r2 = (real_of_int (sint (wmax w1 w2)))" using bound1b bound2 NumNeg less_trans wmax.simps of_int_less_iff word_sless_alt rw1 antisym_conv2 less_imp_le max_def by metis have res2:"wmax w1 w2 \\<^sub>E (real_of_int (sint (wmax w1 w2)))" apply(rule repINT) using bound1a bound1b bound2 NumNeg leD leI less_trans n2 wmax.simps by (auto simp add: word_sless_alt) show "wmax w1 w2 \\<^sub>E max r1 r2" using res1 res2 by auto next case NumNum from NumNum eq1 eq2 have eq1:"r1 = (real_of_int (sint w1))" and eq2:"r2 = (real_of_int (sint w2))" and bound1a:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2a:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" by (auto simp add: repe.simps) have res1:"max r1 r2 = (real_of_int (sint (wmax w1 w2)))" using eq1 eq2 bound1a bound1b bound2a bound2b apply (auto simp add: max_def word_sless_alt not_less; transfer) apply simp_all done have res2:"wmax w1 w2 \\<^sub>E (real_of_int (sint (wmax w1 w2)))" apply (rule repINT) using bound1a bound1b bound2a bound2b by (simp add: \max r1 r2 = (real_of_int (sint (wmax w1 w2)))\ eq2 min_less_iff_disj)+ show "wmax w1 w2 \\<^sub>E max r1 r2" using res1 res2 by auto qed lemma max_repU1: assumes "w1 \\<^sub>U x" assumes "w2 \\<^sub>U y" shows "wmax w1 w2 \\<^sub>U x " using wmax_lemma assms repU_def by (meson le_max_iff_disj) lemma max_repU2: assumes "w1 \\<^sub>U y" assumes "w2 \\<^sub>U x" shows "wmax w1 w2 \\<^sub>U x" using wmax_lemma assms repU_def by (meson le_max_iff_disj) text\Product of w1 * w2 with bounds checking\ fun wtimes :: "word \ word \ word" where "wtimes w1 w2 = (if w1 = POS_INF \ w2 = POS_INF then POS_INF else if w1 = NEG_INF \ w2 = POS_INF then NEG_INF else if w1 = POS_INF \ w2 = NEG_INF then NEG_INF else if w1 = NEG_INF \ w2 = NEG_INF then POS_INF else if w1 = POS_INF \ w2 0 0 = w2 then 0 else if w1 = NEG_INF \ w2 0 0 = w2 then 0 else if w1 w2 = POS_INF then NEG_INF else if 0 w2 = POS_INF then POS_INF else if 0 = w1 \ w2 = POS_INF then 0 else if w1 w2 = NEG_INF then POS_INF else if 0 w2 = NEG_INF then NEG_INF else if 0 = w1 \ w2 = NEG_INF then 0 else (let prod::64 Word.word = (scast w1) * (scast w2) in if prod <=s (scast NEG_INF) then NEG_INF else if (scast POS_INF) <=s prod then POS_INF else (scast prod)))" subsection \Multiplication upper bound\ text\Product of 32-bit numbers fits in 64 bits\ lemma times_upcast_lower: fixes x y::int assumes a1:"x \ -2147483648" assumes a2:"y \ -2147483648" assumes a3:"x \ 2147483648" assumes a4:"y \ 2147483648" shows "- 4611686018427387904 \ x * y" proof - let ?thesis = "- 4611686018427387904 \ x * y" have is_neg:"- 4611686018427387904 < (0::int)" by auto have case1:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"x * y \ 0" using a5 a6 by (simp add: zero_le_mult_iff) then show ?thesis using is_neg by auto qed have case2:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"-2147483648 * (2147483648) \ x * 2147483648" using a1 a2 a3 a4 a5 a6 by linarith have h2:"-2147483648 \ y" using a6 by auto have h3:"x * 2147483648 \ x * y" using a1 a2 a3 a4 a5 a6 h2 using mult_left_mono_neg by blast show ?thesis using h1 h2 h3 by auto qed have case3:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * (-2147483648) \ 2147483648 * y" using a1 a2 a3 a4 a5 a6 by linarith have h2:"-2147483648 \ x" using a5 by auto have h3:"2147483648 * y \ x * y" using a1 a2 a3 a4 a5 a6 h2 using mult_left_mono_neg mult_right_mono_neg by blast show ?thesis using h1 h2 h3 by auto qed have case4:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"x * y \ 0" using a5 a6 by (simp add: zero_le_mult_iff) then show ?thesis using is_neg by auto qed show ?thesis using case1 case2 case3 case4 by linarith qed text\Product of 32-bit numbers fits in 64 bits\ lemma times_upcast_upper: fixes x y ::int assumes a1:"x \ -2147483648" assumes a2:"y \ -2147483648" assumes a3:"x \ 2147483648" assumes a4:"y \ 2147483648" shows "x * y \ 4611686018427387904" proof - let ?thesis = "x * y \ 4611686018427387904" have case1:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * 2147483648 \ x * 2147483648" using a1 a2 a3 a4 a5 a6 by linarith have h2:"x * 2147483648 \ x * y" using a1 a2 a3 a4 a5 a6 by (simp add: mult_mono) show ?thesis using h1 h2 by auto qed have case2:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * 2147483648 \ (0::int)" by auto have h2:"0 \ x * y" using a5 a6 mult_nonneg_nonpos2 by blast show ?thesis using h1 h2 by auto qed have case3:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * 2147483648 \ (0::int)" by auto have h2:"0 \ x * y" using a5 a6 mult_nonneg_nonpos by blast show ?thesis using h1 h2 by auto qed have case4:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"-2147483648 * -2147483648 \ x * -2147483648" using a1 a2 a3 a4 a5 a6 by linarith have h2:"x * -2147483648 \ x * y" using a1 a2 a3 a4 a5 a6 mult_left_mono_neg by blast show ?thesis using h1 h2 by auto qed show "x * y \ 4611686018427387904" using case1 case2 case3 case4 by linarith qed text\Correctness of 32x32 bit multiplication\ subsection \Exact multiplication\ lemma wtimes_exact: assumes eq1:"w1 \\<^sub>E r1" assumes eq2:"w2 \\<^sub>E r2" shows "wtimes w1 w2 \\<^sub>E r1 * r2" proof - have POS_cast:"sint ((scast POS_INF)::64 Word.word) = sint POS_INF" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have POS_sint:"sint POS_INF = (2^31)-1" by auto have w1_cast:"sint ((scast w1)::64 Word.word) = sint w1" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have w2_cast:"sint ((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have NEG_cast:"sint ((scast NEG_INF)::64 Word.word) = sint NEG_INF" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have rangew1:"sint ((scast w1)::64 Word.word) \ {- (2 ^ 31).. (2^31)} " using word_sint.Rep[of "(w1)::32 Word.word"] sints32 len32 mem_Collect_eq POS_cast w1_cast by auto have rangew2:"sint ((scast w2)::64 Word.word) \ {- (2 ^ 31).. (2^31)} " using word_sint.Rep[of "(w2)::32 Word.word"] sints32 len32 mem_Collect_eq POS_cast w2_cast by auto show ?thesis proof (cases rule: case_times_inf[of w1 w2]) case PosPos then have a1: "PosInf \ r1" and a2: "PosInf \ r2" using "PosPos" eq1 eq2 repe.simps by (auto) have f3: "\n e::real. 1 \ max ( (numeral n)) e" by (simp add: le_max_iff_disj) have "\n e::real. 0 \ max ( (numeral n)) e" by (simp add: le_max_iff_disj) then have "r1 \ r1 * r2" using f3 "PosPos" eq1 eq2 repe.simps using eq1 eq2 by (auto simp add: repe.simps) then have "PosInf \ r1 * r2" using a1 by linarith then show ?thesis using "PosPos" by(auto simp add: repe.simps) next case NegPos from "NegPos" have notPos:"w1 \ POS_INF" unfolding POS_INF_def NEG_INF_def by auto have a1: "NegInf \ r1" using eq1 "NegPos" by (auto simp add: repe.simps) have a2: "PosInf \ r2" using eq2 "NegPos" by (auto simp add: repe.simps) have f1: "real_of_int Numeral1 = 1" by simp have f3: "(real_of_int 3) \ - r1" using a1 by (auto) have f4:"0 \ r2" using f1 a2 by(auto) have f5: "r1 \ - 1" using f3 by auto have fact:"r1 * r2 \ - r2" using f5 f4 mult_right_mono by fastforce show ?thesis using a1 a2 fact by (auto simp add: repe.simps "NegPos") next case PosNeg have a1: "PosInf \ r1" using eq1 "PosNeg" by (auto simp add: repe.simps) then have h1:"r1 \ 1" by (auto) have a2: " NegInf \ r2" using eq2 "PosNeg" by (auto simp add: repe.simps) have f1: "\ NegInf * (- 1) \ 1" by (auto) have f2: "\e ea::real. (e * (- 1) \ ea) = (ea * (- 1) \ e)" by force then have f3: "\ 1 * (- 1::real) \ NegInf" using f1 by blast have f4: "r1 * (- 1) \ NegInf" using f2 a1 by(auto) have f5: "\e ea eb. (if (ea::real) \ eb then e \ eb else e \ ea) = (e \ ea \ e \ eb)" by force have " 0 * (- 1::real) \ 1" by simp then have "r1 * (- 1) \ 0" using f5 f4 f3 f2 by meson then have f6: "0 \ r1" by fastforce have "1 * (- 1) \ (- 1::real)" using f2 by force then have fact:"r2 \ (- 1)" using f3 a2 by fastforce have rule:"\c. c > 0 \ r1 \ c \ r2 \ -1 \ r1 * r2 \ -c" apply auto by (metis (no_types, hide_lams) f5 mult_less_cancel_left_pos mult_minus1_right neg_le_iff_le not_less) have "r1 * r2 \ NegInf" using "PosNeg" f6 fact rule[of PosInf] a1 by(auto) then show ?thesis using "PosNeg" by (auto simp add: repe.simps) next case NegNeg have a1: "(-2147483647) \ r1" using eq1 "NegNeg" by (auto simp add: repe.simps) then have h1:"r1 \ -1" using max.bounded_iff max_def one_le_numeral by auto have a2: " (-2147483647) \ r2" using eq2 "NegNeg" by (auto simp add: repe.simps) have f1: "\e ea eb. \ (e::real) \ ea \ \ 0 \ eb \ eb * e \ eb * ea" using mult_left_mono by metis have f2: "- 1 = (- 1::real)" by force have f3: " 0 \ (1::real)" by simp have f4: "\e ea eb. (ea::real) \ e \ \ ea \ eb \ \ eb \ e" by (meson less_le_trans not_le) have f5: " 0 \ (2147483647::real)" by simp have f6: "- (- 2147483647) = (2147483647::real)" by force then have f7: "- ( (- 2147483647) * r1) = (2147483647 * r1)" by (metis mult_minus_left) have f8: "- ( (- 2147483647) * (- 1)) = 2147483647 * (- 1::real)" by simp have " 2147483647 = - 1 * (- 2147483647::real)" by simp then have f9: "r1 \ (- 1) \ 2147483647 \ r1 * (- 2147483647)" using f8 f7 f5 f2 f1 by linarith have f10: "- 2147483647 = (- 2147483647::real)" by fastforce have f11: "- (r2 * 1 * (r1 * (- 1))) = r1 * r2" by (simp add: mult.commute) have f12: "r1 * (- 1) = - (r1 * 1)" by simp have "r1 * 1 * ( (- 2147483647) * 1) = (- 2147483647) * r1" by (simp add: mult.commute) then have f13: "r1 * (- 1) * ( (- 2147483647) * 1) = 2147483647 * r1" using f12 f6 by (metis (no_types) mult_minus_left) have " 1 * r1 \ 1 * (- 2147483647)" using a1 by (auto simp add: a1) then have " 2147483647 \ r1 * (- 1)" by fastforce then have "0 \ r1 * (- 1)" using f5 f4 by (metis) then have "r1 \ (- 1) \ - (r1 * 2147483647) \ - (r2 * 1 * (r1 * (- 1)))" by (metis a2 f11 h1 mult_left_mono_neg minus_mult_right mult_minus1_right neg_0_le_iff_le) then have "r1 \ (- 1) \ r1 * (- 2147483647) \ r2 * r1" using f11 f10 by (metis mult_minus_left mult.commute) then have fact:" 2147483647 \ r2 * r1" using f9 f4 by blast show ?thesis using a1 a2 h1 fact by (auto simp add: repe.simps "NegNeg" mult.commute) next case PosLo from "PosLo" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "PosLo" have upper:" (real_of_int (sint POS_INF)) \ r1 " by (auto simp add: repe.simps) have lower1:"sint w2 < 0" using "PosLo" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower2:"sint w2 \ -1" by auto from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "PosLo" by (auto simp add: repe.simps) have f4: "r1 * (- 1) \ (- 2147483647)" using upper by (auto) have f5: "r2 \ (- 1)" using lower2 rw2 by transfer simp have "0 < r1" using upper by (auto) have "\r. r < - 2147483647 \ \ r < r1 * - 1" using f4 less_le_trans by blast then have "r1 * (real_of_int (sint w2)) \ (- 2147483647)" using f5 f4 upper lower2 rw2 mult_left_mono by (metis \0 < r1\ dual_order.order_iff_strict f5 mult_left_mono rw2) then have "r1 * r2 \ real_of_int (sint NEG_INF)" using upper lower2 rw2 by (auto) then show ?thesis using "PosLo" by (auto simp add: repe.simps) next case PosHi from "PosHi" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "PosHi" have upper:"(real_of_int (sint POS_INF)) \ r1 " by (auto simp add: repe.simps) have lower1:"sint w2 > 0" using "PosHi" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower2:"sint w2 \ 1" by auto from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "PosHi" by (auto) have "0 \ r1" using upper by (auto) then have "r1 \ r1 * r2" using rw2 lower2 by (metis (no_types) mult_left_mono mult.right_neutral of_int_1_le_iff) have "PosInf \ r1 * r2" using upper lower2 rw2 apply (auto) using \0 \ r1\ mult_numeral_1_right numeral_One of_int_1_le_iff zero_le_one apply simp using mult_mono [of 2147483647 r1 1 \signed_real_of_word (w2::32 Word.word)\] apply simp apply transfer apply simp done then show ?thesis using "PosHi" by (auto simp add: repe.simps) next case PosZero from "PosZero" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "PosZero" have upper:" (real_of_int (sint POS_INF)) \ r1 " by (auto simp add: repe.simps) have lower1:"sint w2 = 0" using "PosZero" by (auto simp add: word_sless_def word_sle_def) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "PosZero" by auto have "0 = r1 * r2" using "PosZero" rw2 by auto then show ?thesis using "PosZero" by (auto simp add: repe.simps) next case NegHi have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" using "NegHi" by (auto) from eq1 "NegHi" have upper:"(real_of_int (sint NEG_INF)) \ r1 " by (auto simp add: repe.simps) have low:"sint w2 > 0" using "NegHi" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower1:"(real_of_int (sint w2)) > 0" by transfer simp then have lower2:"(real_of_int (sint w2)) \ 1" using low by transfer simp from eq1 have rw1:"r1 \ (real_of_int (sint w1))" using repe.simps "NegHi" by (simp add: upper) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "NegHi" by (auto) have mylem:"\x y z::real. x \ -1 \ y \ 1 \ z \ -1 \ x \ z \ x * y \ z" proof - fix x y z::real assume a1:"x \ -1" assume a2:"y \ 1" then have h1:"-1 \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) assume a5:"x \ z" then have h6:"-x \ -z" by auto have h3:"-x * -z = x * z" by auto show "x * y \ z" using a1 a2 a3 a5 a4 h1 h2 h3 h6 h5 a5 dual_order.trans leD mult.right_neutral by (metis dual_order.order_iff_strict mult_less_cancel_left2) qed have prereqs:"r1 \ - 1" "1 \ (real_of_int (sint w2))" " (- 2147483647::real) \ - 1 " "r1 \ (-2147483647)" using rw1 rw2 "NegHi" lower2 by (auto simp add: word_sless_def word_sle_def) have "r1 * r2 \ real_of_int (sint NEG_INF)" using upper lower1 lower2 rw1 rw2 apply (auto simp add: word_sless_def word_sle_def) using mylem[of "r1" " (real_of_int (sint w2))" " (- 2147483647)"] prereqs by auto then show ?thesis using "NegHi" by (auto simp add: repe.simps) next case NegLo from "NegLo" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "NegLo" have upper:"(real_of_int (sint NEG_INF)) \ r1" by (auto simp add: repe.simps) have low:"sint w2 < 0" using "NegLo" by (auto simp add: word_sless_def word_sle_def dual_order.order_iff_strict) then have lower1:"(real_of_int (sint w2)) < 0" by transfer simp from eq1 have rw1:"r1 \ (real_of_int (sint w1))" using repe.simps "NegLo" by (simp add: upper) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "NegLo" by (auto) have hom:"(- 2147483647) = -(2147483647::real)" by auto have mylem:"\x y z::real. y < 0 \ x \ y \ z \ -1 \ -y \ x * z" proof - fix x y z::real assume a1:"y < 0" assume a2:"x \ y" then have h1:"-x \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) have h4:"-x * -z \ -y" using a1 a2 a3 a4 h1 h2 h5 dual_order.trans mult.right_neutral by (metis mult.commute neg_0_less_iff_less real_mult_le_cancel_iff1) have h3:"-x * -z = x * z" by auto show "- y \ x * z " using a1 a2 a3 a4 h1 h2 h3 h4 h5 by simp qed have prereqs:"- 2147483647 < (0::real)" " r1 \ - 2147483647" using rw1 rw2 "NegLo" by (auto simp add: word_sless_def word_sle_def) moreover have \sint w2 \ - 1\ using low by simp then have \real_of_int (sint w2) \ real_of_int (- 1)\ by (simp only: of_int_le_iff) then have \signed_real_of_word w2 \ - 1\ by simp ultimately have "2147483647 \ r1 * r2" using upper lower1 rw1 rw2 mylem[of "-2147483647" "r1" "(real_of_int (sint w2))"] by (auto simp add: word_sless_def word_sle_def) then show ?thesis using "NegLo" by (auto simp add: repe.simps) next case NegZero from "NegZero" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq2 "NegZero" have "r2 = 0" using repe.simps "NegZero" by (auto) then show ?thesis using "NegZero" by (auto simp add: repe.simps) next case LoPos from "LoPos" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "LoPos" have upper:"(real_of_int (sint POS_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 < 0" using "LoPos" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower2:"sint w1 \ -1" by auto from eq1 have rw1:"r1 = (real_of_int (sint w1))" using repe.simps "LoPos" by (auto simp add: repe.simps) have f4: "r2 * (- 1) \ (- 2147483647)" using upper by(auto) have f5: "r1 \ (- 1)" using lower2 rw1 by transfer simp have "0 < r2" using upper by(auto) then have "r2 * r1 \ r2 * (- 1)" by (metis dual_order.order_iff_strict mult_right_mono f5 mult.commute) then have "r2 * r1 \ (- 2147483647)" by (meson f4 less_le_trans not_le) then have "(real_of_int (sint w1)) * r2 \ (- 2147483647)" using f5 f4 rw1 less_le_trans not_le mult.commute rw1 by (auto simp add: mult.commute) then have "r1 * r2 \ NegInf" using rw1 by (auto) then show ?thesis using "LoPos" by (auto simp: repe.simps) next case HiPos from "HiPos" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "HiPos" have upper:"(real_of_int (sint POS_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 > 0" using "HiPos" by (auto simp add: word_sless_def word_sle_def dual_order.order_iff_strict) then have lower2:"sint w1 \ 1" by auto from eq1 have rw2:"r1 = (real_of_int (sint w1))" using "HiPos" by (auto simp add: repe.simps) have "0 \ r2" using upper by(auto) then have "r2 \ r2 * r1" using lower2 rw2 by (metis (no_types) mult_left_mono mult.right_neutral of_int_1_le_iff) have "2147483647 \ r1 * r2" using upper lower2 rw2 apply (simp add: word_sless_def word_sle_def) using mult_mono [of 1 \signed_real_of_word w1\ 2147483647 r2] apply simp apply transfer apply simp done then show ?thesis using "HiPos" by (auto simp add: repe.simps) next case ZeroPos from "ZeroPos" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "ZeroPos" have upper:" (real_of_int (sint POS_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 = 0" using "ZeroPos" by (auto simp add: word_sless_def word_sle_def) from eq1 have rw2:"r1 = (real_of_int (sint w1))" using repe.simps "ZeroPos" by (auto) have "r1 = 0" using lower1 rw2 by auto then show ?thesis using "ZeroPos" by (auto simp add: repe.simps) next case ZeroNeg from "ZeroNeg" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "ZeroNeg" have upper:"(real_of_int (sint NEG_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 = 0" using "ZeroNeg" by (auto simp add: word_sless_def word_sle_def) from eq1 have rw2:"r1 = (real_of_int (sint w1))" using repe.simps "ZeroNeg" by (auto) have "r1 = 0" using lower1 rw2 by auto then show ?thesis using "ZeroNeg" by (auto simp add: repe.simps) next case LoNeg from "LoNeg" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "LoNeg" have upper:" (real_of_int (sint NEG_INF)) \ r2 " by (auto simp add: repe.simps) have low:"sint w1 < 0" using "LoNeg" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower1:"(real_of_int (sint w1)) < 0" by transfer simp from low have \sint w1 \ - 1\ by simp then have lower2:"(real_of_int (sint w1)) \ -1" by transfer simp from eq1 have rw1:"r2 \ (real_of_int (sint w2))" using "LoNeg" upper by auto from eq1 have rw2:"r1 = (real_of_int (sint w1))" using "LoNeg" by (auto simp add: upper repe.simps) have hom:"(- 2147483647::real) = -(2147483647)" by auto have mylem:"\x y z::real. y < 0 \ x \ y \ z \ -1 \ -y \ x * z" proof - fix x y z::real assume a1:"y < 0" assume a2:"x \ y" then have h1:"-x \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) have h4:"-x * -z \ -y" using a1 a2 a3 a4 h1 h2 h5 dual_order.trans mult_left_mono mult.right_neutral mult.commute by (metis dual_order.order_iff_strict mult_minus_right mult_zero_right neg_le_iff_le) have h3:"-x * -z = x * z" by auto show "- y \ x * z " using a1 a2 a3 a4 h1 h2 h3 h4 h5 by simp qed have prereqs:"- 2147483647 < (0::real)" " r2 \ - 2147483647" " (real_of_int (sint w1)) \ - 1" using rw1 rw2 "LoNeg" lower2 by (auto simp add: word_sless_def word_sle_def lower2) have "2147483647 \ r1 * r2" using upper lower1 lower2 rw1 rw2 mylem[of "-2147483647" "r2" "(real_of_int (sint w1))"] prereqs by (auto simp add:word_sless_def word_sle_def mult.commute) then show ?thesis using "LoNeg" by (auto simp add: repe.simps) next case HiNeg from HiNeg have w1NotPinf:"w1 \ POS_INF" and w1NotNinf:"w1 \ NEG_INF" by (auto) have upper:" (real_of_int (sint NEG_INF)) \ r2 " using HiNeg eq2 by (auto simp add: repe.simps ) have low:"sint w1 > 0" using HiNeg apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower1:"(real_of_int (sint w1)) > 0" by transfer simp from low have \sint w1 \ 1\ by simp then have lower2:"(real_of_int (sint w1)) \ 1" by transfer simp from eq2 have rw1:"r2 \ (real_of_int (sint w2))" using repe.simps HiNeg by (simp add: upper) from eq1 have rw2:"r1 = (real_of_int (sint w1))" using repe.simps HiNeg by (auto) have mylem:"\x y z::real. x \ -1 \ y \ 1 \ z \ -1 \ x \ z \ x * y \ z" proof - fix x y z::real assume a1:"x \ -1" assume a2:"y \ 1" then have h1:"-1 \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) assume a5:"x \ z" then have h6:"-x \ -z" by auto have h3:"-x * -z = x * z" by auto show "x * y \ z" using a1 a2 a3 a4 h1 h2 h3 h6 h5 a5 dual_order.trans less_eq_real_def by (metis mult_less_cancel_left1 not_le) qed have prereqs:"r2 \ - 1" "1 \ (real_of_int (sint w1))" " (- 2147483647) \ - (1::real )" "r2 \ (- 2147483647)" using rw1 rw2 HiNeg lower2 by (auto simp add: word_sless_def word_sle_def) have "r1 * r2 \ - 2147483647" using upper lower1 lower2 rw1 rw2 apply (auto simp add: word_sless_def word_sle_def) using mylem[of "r2" "(real_of_int (sint w1))" " (- 2147483647)"] prereqs by (auto simp add: mult.commute) then show ?thesis using HiNeg by(auto simp add: repe.simps) next case AllFinite let ?prod = "(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))" consider (ProdNeg) "?prod <=s ((scast NEG_INF)::64 Word.word)" | (ProdPos) "(((scast POS_INF)::64 Word.word) <=s ?prod)" | (ProdFin) "\(?prod <=s ((scast NEG_INF)::64 Word.word)) \ \((scast POS_INF)::64 Word.word) <=s ?prod" by (auto) then show ?thesis proof (cases) case ProdNeg have bigLeq:"(4611686018427387904::real) \ 9223372036854775807" by auto have set_cast:"\x::int. (x \ {-(2^31)..2^31}) = ( (real_of_int x) \ {-(2^31)..2^31})" by auto have eq3:"sint(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) = sint ((scast w1)::64 Word.word) * sint ((scast w2)::64 Word.word)" apply(rule Word_Lemmas.signed_arith_sint(4)) using rangew1 rangew2 w1_cast w2_cast using Word.word_size[of "((scast w1)::64 Word.word)"] using Word.word_size[of "((scast w2)::64 Word.word)"] using times_upcast_upper[of "sint w1" "sint w2"] using times_upcast_lower[of "sint w1" "sint w2"] by auto assume "?prod <=s ((scast NEG_INF)::64 Word.word)" then have sint_leq:"sint ?prod \ sint ((scast NEG_INF)::64 Word.word)" using word_sle_def by blast have neqs:"w1 \ POS_INF" " w1 \ NEG_INF" "w2 \ POS_INF" "w2 \ NEG_INF" using AllFinite word_sless_def signed.not_less_iff_gr_or_eq by force+ from eq1 have rw1:"r1 = (real_of_int (sint w1))" using neqs by (auto simp add: repe.simps) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using neqs by (auto simp add: repe.simps) show ?thesis using AllFinite ProdNeg w1_cast w2_cast rw1 rw2 sint_leq apply (auto simp add: repe.simps eq3) apply (subst (asm) of_int_le_iff [symmetric, where ?'a = real]) apply simp done next case ProdPos have bigLeq:"(4611686018427387904::real) \ 9223372036854775807" by auto have set_cast:"\x::int. (x \ {-(2^31)..2^31}) = ( (real_of_int x) \ {-(2^31)..2^31})" by auto have eq3:"sint(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) = sint ((scast w1)::64 Word.word) * sint ((scast w2)::64 Word.word)" apply(rule Word_Lemmas.signed_arith_sint(4)) using rangew1 rangew2 POS_cast POS_sint w1_cast w2_cast using Word.word_size[of "((scast w1)::64 Word.word)"] using Word.word_size[of "((scast w2)::64 Word.word)"] using times_upcast_upper[of "sint w1" "sint w2"] using times_upcast_lower[of "sint w1" "sint w2"] by auto assume cast:"((scast POS_INF)::64 Word.word) <=s ?prod" then have sint_leq:"sint ((scast POS_INF)::64 Word.word) \ sint ?prod" using word_sle_def by blast have neqs:"w1 \ POS_INF" " w1 \ NEG_INF" "w2 \ POS_INF" "w2 \ NEG_INF" using AllFinite word_sless_def signed.not_less_iff_gr_or_eq by force+ from eq1 have rw1:"r1 = (real_of_int (sint w1))" using repe.simps AllFinite neqs by auto from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps AllFinite neqs by auto have prodHi:"r1 * r2 \ PosInf" using w1_cast w2_cast rw1 rw2 sint_leq apply (auto simp add: eq3) apply (subst (asm) of_int_le_iff [symmetric, where ?'a = real]) apply simp done have infs:"SCAST(32 \ 64) NEG_INF 64) POS_INF" by (auto) have casted:"SCAST(32 \ 64) POS_INF <=s SCAST(32 \ 64) w1 * SCAST(32 \ 64) w2" using cast by auto have almostContra:"SCAST(32 \ 64) NEG_INF 64) w1 * SCAST(32 \ 64) w2" using infs cast signed.order.strict_trans2 by blast have contra:"\(SCAST(32 \ 64) w1 * SCAST(32 \ 64) w2 <=s SCAST(32 \ 64) NEG_INF)" using eq3 almostContra by auto have wtimesCase:"wtimes w1 w2 = POS_INF" using neqs ProdPos almostContra wtimes.simps AllFinite ProdPos by (auto simp add: repe.simps Let_def) show ?thesis using prodHi apply(simp only: repe.simps) apply(rule disjI1) apply(rule exI[where x= "r1*r2"]) apply(rule conjI) apply(rule wtimesCase) using prodHi by auto next case ProdFin have bigLeq:"(4611686018427387904::real) \ 9223372036854775807" by auto have set_cast:"\x::int. (x \ {-(2^31)..2^31}) = ( (real_of_int x) \ {-(2^31)..2^31})" by auto have eq3:"sint(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) = sint ((scast w1)::64 Word.word) * sint ((scast w2)::64 Word.word)" apply(rule Word_Lemmas.signed_arith_sint(4)) using rangew1 rangew2 POS_cast POS_sint w1_cast w2_cast using Word.word_size[of "((scast w1)::64 Word.word)"] using Word.word_size[of "((scast w2)::64 Word.word)"] using times_upcast_upper[of "sint w1" "sint w2"] using times_upcast_lower[of "sint w1" "sint w2"] by auto from ProdFin have a1:"\(?prod <=s ((scast NEG_INF)::64 Word.word))" by auto then have sintGe:"sint (?prod) > sint (((scast NEG_INF)::64 Word.word))" using word_sle_def dual_order.order_iff_strict signed.linear by fastforce from ProdFin have a2:"\((scast POS_INF)::64 Word.word) <=s ?prod" by auto then have sintLe:"sint (((scast POS_INF)::64 Word.word)) > sint (?prod)" using word_sle_def dual_order.order_iff_strict signed.linear by fastforce have neqs:"w1 \ POS_INF" " w1 \ NEG_INF" "w2 \ POS_INF" "w2 \ NEG_INF" using AllFinite word_sless_def signed.not_less_iff_gr_or_eq by force+ from eq1 have rw1:"r1 = (real_of_int (sint w1))" using neqs by(auto simp add: repe.simps) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using neqs by(auto simp add: repe.simps) from rw1 rw2 have "r1 * r2 = (real_of_int ((sint w1) * (sint w2)))" by simp have rightSize:"sint (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) \ sints (len_of TYPE(32))" using sintLe sintGe sints32 by (simp) have downcast:"sint ((scast (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)))::word) = sint (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))" using scast_down_range[OF rightSize] by auto then have res_eq:"r1 * r2 = real_of_int(sint((scast (((scast w1)::64 Word.word)*((scast w2)::64 Word.word)))::word))" using rw1 rw2 eq3 POS_cast POS_sint w1_cast w2_cast downcast \r1 * r2 = (real_of_int (sint w1 * sint w2))\ by (auto) have res_up:"sint (scast (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))::word) < sint POS_INF" using rw1 rw2 eq3 POS_cast POS_sint w1_cast w2_cast downcast \r1 * r2 = (real_of_int (sint w1 * sint w2))\ \sint (scast w1 * scast w2) < sint (scast POS_INF)\ of_int_eq_iff res_eq by presburger have res_lo:"sint NEG_INF < sint (scast (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))::word)" using rw1 rw2 eq3 POS_cast POS_sint w1_cast w2_cast NEG_cast downcast \r1 * r2 = (real_of_int (sint w1 * sint w2))\ \sint (scast NEG_INF) < sint (scast w1 * scast w2)\ of_int_eq_iff res_eq by presburger have "scast ?prod \\<^sub>E (r1 * r2)" using res_eq res_up res_lo apply (auto simp add: rep_simps) using repeInt_simps by auto then show ?thesis using AllFinite ProdFin by(auto) qed qed qed subsection \Multiplication upper bound\ text\Upper bound of multiplication from upper and lower bounds\ fun tu :: "word \ word \ word \ word \ word" where "tu w1l w1u w2l w2u = wmax (wmax (wtimes w1l w2l) (wtimes w1u w2l)) (wmax (wtimes w1l w2u) (wtimes w1u w2u))" lemma tu_lemma: assumes u1:"u\<^sub>1 \\<^sub>U (r1::real)" assumes u2:"u\<^sub>2 \\<^sub>U (r2::real)" assumes l1:"l\<^sub>1 \\<^sub>L (r1::real)" assumes l2:"l\<^sub>2 \\<^sub>L (r2::real)" shows "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U (r1 * r2)" proof - obtain rl1 rl2 ru1 ru2 :: real where gru1:"ru1 \ r1" and gru2:"ru2 \ r2" and grl1:"rl1 \ r1" and grl2:"rl2 \ r2" and eru1:"u\<^sub>1 \\<^sub>E ru1" and eru2:"u\<^sub>2 \\<^sub>E ru2" and erl1:"l\<^sub>1 \\<^sub>E rl1" and erl2:"l\<^sub>2 \\<^sub>E rl2" using u1 u2 l1 l2 unfolding repU_def repL_def by auto have timesuu:"wtimes u\<^sub>1 u\<^sub>2 \\<^sub>E ru1 * ru2" using wtimes_exact[OF eru1 eru2] by auto have timesul:"wtimes u\<^sub>1 l\<^sub>2 \\<^sub>E ru1 * rl2" using wtimes_exact[OF eru1 erl2] by auto have timeslu:"wtimes l\<^sub>1 u\<^sub>2 \\<^sub>E rl1 * ru2" using wtimes_exact[OF erl1 eru2] by auto have timesll:"wtimes l\<^sub>1 l\<^sub>2 \\<^sub>E rl1 * rl2" using wtimes_exact[OF erl1 erl2] by auto have maxt12:"wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>E max (rl1 * rl2) (ru1 * rl2)" by (rule wmax_lemma[OF timesll timesul]) have maxt34:"wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>E max (rl1 * ru2) (ru1 * ru2)" by (rule wmax_lemma[OF timeslu timesuu]) have bigMax:"wmax (wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>E max (max (rl1 * rl2) (ru1 * rl2)) (max (rl1 * ru2) (ru1 * ru2))" by (rule wmax_lemma[OF maxt12 maxt34]) obtain maxt12val :: real where maxU12:"wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>U max (rl1 * rl2) (ru1 * rl2)" using maxt12 unfolding repU_def by blast obtain maxt34val :: real where maxU34:"wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>U max (rl1 * ru2) (ru1 * ru2)" using maxt34 unfolding repU_def by blast obtain bigMaxU:"wmax (wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>U max (max (rl1 * rl2) (ru1 * rl2)) (max (rl1 * ru2) (ru1 * ru2))" using bigMax unfolding repU_def by blast have ivl1:"rl1 \ ru1" using grl1 gru1 by auto have ivl2:"rl2 \ ru2" using grl2 gru2 by auto let ?thesis = "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" show ?thesis using ivl1 ivl2 proof(cases rule: case_ivl_zero) case ZeroZero assume "rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2" then have geq1:"ru1 \ 0" and geq2:"ru2 \ 0" by auto consider "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" using le_cases by auto then show "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" proof (cases) case 1 have g1:"ru1 * ru2 \ ru1 * r2" using "1" geq1 geq2 grl2 gru2 by (simp add: mult_left_mono) have g2:"ru1 * r2 \ r1 * r2" using "1" geq1 geq2 grl1 grl2 gru1 gru2 by (simp add: mult_right_mono) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up eru1 eru2 erl1 erl2 repU_def timesuu tu.simps max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] by (metis wmax.elims) next case 2 have g1:"ru1 * ru2 \ 0" using "2" geq1 geq2 grl2 gru2 by (simp) have g2:"0 \ r1 * r2" using "2" by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) next case 3 have g1:"ru1 * ru2 \ 0" using "3" geq1 geq2 by simp have g2:"0 \ r1 * r2" using "3" by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] repU_def tu.simps timesuu by (metis max.coboundedI1 max.commute maxt34) next case 4 have g1:"rl1 * rl2 \ rl1 * r2" using "4" geq1 geq2 grl1 grl2 gru1 gru2 using \rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2\ less_eq_real_def by (metis mult_left_mono_neg) have g2:"rl1 * r2 \ r1 * r2" using "4" geq1 geq2 grl1 grl2 gru1 gru2 \rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2\ by (metis mult_left_mono_neg mult.commute) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) qed next case ZeroPos assume bounds:"rl1 \ 0 \ 0 \ ru1 \ 0 \ rl2" have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast consider "r1 \ 0" | "r1 \ 0" using le_cases by (auto) then show ?thesis proof (cases) case 1 assume r1:"r1 \ 0" have g1:"ru1 * ru2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_left_mono by blast have g2:"ru1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by blast from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) next case 2 assume r1:"r1 \ 0" have g1:"ru1 * ru2 \ 0" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_left_mono by (simp add: mult_less_0_iff less_le_trans not_less) have g2:"0 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) qed next case ZeroNeg assume bounds:"rl1 \ 0 \ 0 \ ru1 \ ru2 \ 0" have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have case1:"r1 \ 0 \ ?thesis" proof - assume r1:"r1 \ 0" have g1:"rl1 * rl2 \ 0" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_less_0_iff less_le_trans not_less by metis have g2:"0 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by (simp add: mult_le_0_iff) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU2 max_repU1 repU_def timesll tu.simps) qed have case2:"r1 \ 0 \ ?thesis" proof - assume r1:"r1 \ 0" have g1:"rl1 * rl2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 by (metis mult_left_mono_neg) have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult.commute by (metis mult_left_mono_neg) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) qed show "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" using case1 case2 le_cases by blast next case PosZero assume bounds:"0 \ rl1 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"ru1 * ru2 \ ru1 * r2" using "1" bounds grl1 grl2 gru1 gru2 using mult_left_mono using leD leI less_le_trans by metis have g2:"ru1 * r2 \ r1 * r2" using "1" bounds grl1 grl2 gru1 gru2 using mult_right_mono by blast from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) next case 2 have g1:"ru1 * ru2 \ 0" using r1 bounds grl2 gru2 gru1 leD leI less_le_trans by auto have g2:"0 \ r1 * r2" using r1 "2" by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) qed next case NegZero assume bounds:"ru1 \ 0 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"ru1 * rl2 \ 0" using r1 "1" bounds grl1 grl2 gru1 gru2 mult_less_0_iff not_less by metis have g2:"0 \ r1 * r2" using r1 "1" bounds grl1 grl2 gru1 gru2 by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesul tu.simps) next case 2 have lower:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have g1:"rl1 * rl2 \ rl1 * r2" using r1 "2" bounds grl1 grl2 gru1 gru2 less_eq(1) less_le_trans not_less mult_le_cancel_left by metis have g2:"rl1 * r2 \ r1 * r2" using r1 "2" bounds grl1 grl2 gru1 gru2 mult.commute not_le lower mult_le_cancel_left by metis from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) qed next case NegNeg assume bounds:"ru1 \ 0 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds dual_order.trans grl2 r2 by blast have g1:"rl1 * rl2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 less_eq(1) mult_le_cancel_left less_le_trans not_less by metis have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult.commute not_le lower1 lower2 mult_le_cancel_left by metis from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) next case NegPos assume bounds:"ru1 \ 0 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using bounds by auto have upper2:"ru2 \ 0" using bounds dual_order.trans gru2 r2 by blast have g1:"ru1 * rl2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 not_less upper1 lower2 mult_le_cancel_left by metis have g2:"ru1 * r2 \ r1 * r2" using r1 upper1 r2 mult_right_mono gru1 by metis from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims maxt34 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] by (metis max_repU1 repU_def timesul tu.simps) next case PosNeg assume bounds:"0 \ rl1 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using dual_order.trans grl2 r2 by blast have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using bounds by auto have g1:"rl1 * ru2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 not_less upper2 lower1 mult_le_cancel_left by metis have g2:"rl1 * r2 \ r1 * r2" using r1 lower1 r2 not_less gru2 gru1 grl1 grl2 by (metis mult_le_cancel_left mult.commute) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" using up maxU12 maxU34 bigMaxU wmax.elims max.coboundedI1 max.commute maxt34 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] by (metis repU_def tu.simps) next case PosPos assume bounds:"0 \ rl1 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using dual_order.trans gru2 u2 r2 bounds by blast have g1:"ru1 * ru2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_left_mono leD leI less_le_trans by metis have g2:"ru1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_right_mono by metis from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max.coboundedI1 max.commute maxt34 max_repU2[OF bigMaxU] max_repU2[OF maxU12] max_repU2[OF maxU34] by (metis repU_def tu.simps) qed qed subsection \Minimum function\ text\Minimum of 2s-complement words\ fun wmin :: "word \ word \ word" where "wmin w1 w2 = (if w1 Correctness of wmin\ lemma wmin_lemma: assumes eq1:"w1 \\<^sub>E (r1::real)" assumes eq2:"w2 \\<^sub>E (r2::real)" shows "wmin w1 w2 \\<^sub>E (min r1 r2)" proof(cases rule: case_inf2[where ?w1.0=w1, where ?w2.0=w2]) case PosPos assume p1:"w1 = POS_INF" and p2:"w2 = POS_INF" then have bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2:"(real_of_int (sint POS_INF)) \ r2" using eq1 eq2 by (auto simp add: rep_simps repe.simps) have eqInf:"wmin w1 w2 = POS_INF" using p1 p2 unfolding wmin.simps by auto have pos_eq:"POS_INF \\<^sub>E min r1 r2" apply(rule repPOS_INF) using bound1 bound2 unfolding eq1 eq2 by auto show ?thesis using pos_eq eqInf by auto next case PosNeg assume p1:"w1 = POS_INF" assume n2:"w2 = NEG_INF" obtain r ra :: real where bound1:" (real_of_int (sint POS_INF)) \ r" and bound2:"ra \ (real_of_int (sint NEG_INF))" and eq1:"r1 = r" and eq2:"r2 = ra" using p1 n2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"wmin w1 w2 = NEG_INF" unfolding eq1 eq2 wmin.simps p1 n2 word_sless_def word_sle_def by(auto) have neg_eq:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using bound1 bound2 eq1 eq2 by auto show "?thesis" using eqNeg neg_eq by auto next case PosNum assume p1:"w1 = POS_INF" assume np2:"w2 \ POS_INF" assume nn2:"w2 \ NEG_INF" have eq2:"r2 = (real_of_int (sint w2))" and bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2a:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound2b:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" using p1 np2 nn2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"min r1 r2 = sint w2" using p1 by (metis bound1 bound2b dual_order.trans eq2 min_def not_less) have neg_eq:"wmin w1 w2 \\<^sub>E (real_of_int (sint (wmin w1 w2)))" apply (rule repINT) using bound1 bound2a bound2b bound2b p1 unfolding eq1 eq2 by (auto simp add: word_sless_alt) show "?thesis" using eqNeg neg_eq by (metis bound2b less_eq_real_def not_less of_int_less_iff p1 wmin.simps word_sless_alt) next case NegPos assume n1:"w1 = NEG_INF" assume p2:"w2 = POS_INF" have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:"(real_of_int (sint POS_INF)) \ r2" using n1 p2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"wmin w1 w2 = NEG_INF" unfolding eq1 eq2 wmin.simps n1 p2 word_sless_def word_sle_def by(auto) have neg_eq:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using bound1 bound2 unfolding eq1 eq2 by auto show "wmin w1 w2 \\<^sub>E min r1 r2" using eqNeg neg_eq by auto next case NegNeg assume n1:"w1 = NEG_INF" assume n2:"w2 = NEG_INF" have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" using n1 n2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using eq1 eq2 bound1 bound2 unfolding NEG_INF_def by (auto) have neg_eq:"wmin w1 w2 = NEG_INF" using n1 n2 unfolding NEG_INF_def wmin.simps by auto show "wmin w1 w2 \\<^sub>E min r1 r2" using eqNeg neg_eq by auto next case NegNum assume n1:"w1 = NEG_INF" and nn2:"w2 \ NEG_INF" and np2:"w2 \ POS_INF" have eq2:"r2 = (real_of_int (sint w2))" and bound2a:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound1:"r1 \ (real_of_int (sint NEG_INF))" using n1 nn2 np2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have eqNeg:"wmin w1 w2 = NEG_INF" using n1 assms(2) bound2a eq2 n1 repeInt_simps by (auto simp add: word_sless_alt) have neg_eq:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using bound1 bound2a bound2b eq1 min_le_iff_disj by blast show "wmin w1 w2 \\<^sub>E min r1 r2" using eqNeg neg_eq by auto next case NumPos assume p2:"w2 = POS_INF" and nn1:"w1 \ NEG_INF" and np1:"w1 \ POS_INF" have eq1:"r1 = (real_of_int (sint w1))" and bound1a:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:" (real_of_int (sint POS_INF)) \ r2" using nn1 np1 p2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have res1:"wmin w1 w2 = w1" using p2 eq1 eq2 assms(1) bound1b p2 repeInt_simps by (auto simp add: word_sless_alt) have res2:"min r1 r2 = (real_of_int (sint w1))" using eq1 eq2 bound1a bound1b bound2 by transfer (auto simp add: less_imp_le less_le_trans min_def) have res3:"wmin w1 w2 \\<^sub>E (real_of_int (sint (wmin w1 w2)))" apply(rule repINT) using p2 bound1a res1 bound1a bound1b bound2 by auto show "wmin w1 w2 \\<^sub>E min r1 r2" using res1 res2 res3 by auto next case NumNeg assume nn1:"w1 \ NEG_INF" assume np1:"w1 \ POS_INF" assume n2:"w2 = NEG_INF" have eq1:"r1 = (real_of_int (sint w1))" and bound1a:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" using nn1 np1 n2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have res1:"wmin w1 w2 = NEG_INF" using n2 bound1b by (metis min.absorb_iff2 min_def n2 not_less of_int_less_iff wmin.simps word_sless_alt) have res2:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using eq1 eq2 bound1a bound1b bound2 min_le_iff_disj by blast show "wmin w1 w2 \\<^sub>E min r1 r2" using res1 res2 by auto next case NumNum assume np1:"w1 \ POS_INF" assume nn1:"w1 \ NEG_INF" assume np2:"w2 \ POS_INF" assume nn2:"w2 \ NEG_INF" have eq1:"r1 = (real_of_int (sint w1))" and eq2:"r2 = (real_of_int (sint w2))" and bound1a:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2a:" (real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" using nn1 np1 nn2 np2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have res1:"min r1 r2 = (real_of_int (sint (wmin w1 w2)))" using eq1 eq2 bound1a bound1b bound2a bound2b apply (simp add: min_def word_sless_alt not_less) apply transfer apply simp done have res2:"wmin w1 w2 \\<^sub>E (real_of_int (sint (wmin w1 w2)))" apply (rule repINT) using bound1a bound1b bound2a bound2b by (simp add: \min r1 r2 = (real_of_int (sint (wmin w1 w2)))\ eq2 min_less_iff_disj)+ show "wmin w1 w2 \\<^sub>E min r1 r2" using res1 res2 by auto qed lemma min_repU1: assumes "w1 \\<^sub>L x" assumes "w2 \\<^sub>L y" shows "wmin w1 w2 \\<^sub>L x " using wmin_lemma assms repL_def by (meson min_le_iff_disj) lemma min_repU2: assumes "w1 \\<^sub>L y" assumes "w2 \\<^sub>L x" shows "wmin w1 w2 \\<^sub>L x" using wmin_lemma assms repL_def by (meson min_le_iff_disj) subsection \Multiplication lower bound\ text\Multiplication lower bound\ fun tl :: "word \ word \ word \ word \ word" where "tl w1l w1u w2l w2u = wmin (wmin (wtimes w1l w2l) (wtimes w1u w2l)) (wmin (wtimes w1l w2u) (wtimes w1u w2u))" text\Correctness of multiplication lower bound\ lemma tl_lemma: assumes u1:"u\<^sub>1 \\<^sub>U (r1::real)" assumes u2:"u\<^sub>2 \\<^sub>U (r2::real)" assumes l1:"l\<^sub>1 \\<^sub>L (r1::real)" assumes l2:"l\<^sub>2 \\<^sub>L (r2::real)" shows "tl l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>L (r1 * r2)" proof - obtain rl1 rl2 ru1 ru2 :: real where gru1:"ru1 \ r1" and gru2:"ru2 \ r2" and grl1:"rl1 \ r1" and grl2:"rl2 \ r2" and eru1:"u\<^sub>1 \\<^sub>E ru1" and eru2:"u\<^sub>2 \\<^sub>E ru2" and erl1:"l\<^sub>1 \\<^sub>E rl1" and erl2:"l\<^sub>2 \\<^sub>E rl2" using u1 u2 l1 l2 unfolding repU_def repL_def by auto have timesuu:"wtimes u\<^sub>1 u\<^sub>2 \\<^sub>E ru1 * ru2" using wtimes_exact[OF eru1 eru2] by auto have timesul:"wtimes u\<^sub>1 l\<^sub>2 \\<^sub>E ru1 * rl2" using wtimes_exact[OF eru1 erl2] by auto have timeslu:"wtimes l\<^sub>1 u\<^sub>2 \\<^sub>E rl1 * ru2" using wtimes_exact[OF erl1 eru2] by auto have timesll:"wtimes l\<^sub>1 l\<^sub>2 \\<^sub>E rl1 * rl2" using wtimes_exact[OF erl1 erl2] by auto have maxt12:"wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>E min (rl1 * rl2) (ru1 * rl2)" by (rule wmin_lemma[OF timesll timesul]) have maxt34:"wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>E min (rl1 * ru2) (ru1 * ru2)" by (rule wmin_lemma[OF timeslu timesuu]) have bigMax:"wmin (wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>E min (min(rl1 * rl2) (ru1 * rl2)) (min (rl1 * ru2) (ru1 * ru2))" by (rule wmin_lemma[OF maxt12 maxt34]) obtain maxt12val :: real where maxU12:"wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>L min (rl1 * rl2) (ru1 * rl2)" using maxt12 unfolding repL_def by blast obtain maxt34val :: real where maxU34:"wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>L min (rl1 * ru2) (ru1 * ru2)" using maxt34 unfolding repL_def by blast obtain bigMaxU:"wmin (wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>L min (min (rl1 * rl2) (ru1 * rl2)) (min (rl1 * ru2) (ru1 * ru2))" using bigMax unfolding repL_def by blast have ivl1:"rl1 \ ru1" using grl1 gru1 by auto have ivl2:"rl2 \ ru2" using grl2 gru2 by auto let ?thesis = "tl l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>L r1 * r2" show ?thesis using ivl1 ivl2 proof(cases rule: case_ivl_zero) case ZeroZero assume "rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2" then have geq1:"ru1 \ 0" and geq2:"ru2 \ 0" and geq3:"rl1 \ 0" and geq4:"rl2 \ 0" by auto consider "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"rl1 * ru2 \ 0" using "1" geq1 geq2 geq3 geq4 grl2 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using "1" geq1 geq2 grl1 grl2 gru1 gru2 by (simp) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis using up eru1 eru2 erl1 erl2 min_repU1 min_repU2 repL_def repU_def timeslu tl.simps wmin.elims by (metis bigMax min_le_iff_disj) next case 2 have g1:"rl1 * ru2 \ rl1 * r2" using "2" geq1 geq2 grl2 gru2 by (metis mult_le_cancel_left geq3 leD) have g2:"rl1 * r2 \ r1 * r2" using "2" geq1 geq2 grl2 gru2 by (simp add: mult_right_mono grl1) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 min_repU2 repL_def tl.simps min.coboundedI1 maxt34) next case 3 have g1:"ru1 * rl2 \ ru1 * r2" using "3" geq1 geq2 grl2 gru2 by (simp add: mult_left_mono) have g2:"ru1 * r2 \ r1 * r2" using "3" geq1 geq2 grl1 grl2 gru1 gru2 mult_minus_right mult_right_mono by (simp add: mult_right_mono_neg) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 maxt34 timesul by (metis repL_def tl.simps) next case 4 have g1:"ru1 * rl2 \ 0" using "4" geq1 geq2 grl1 grl2 gru1 gru2 \rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2\ mult_less_0_iff less_eq_real_def not_less by auto have g2:"0 \ r1 * r2" using "4" geq1 geq2 grl1 grl2 gru1 gru2 by (metis mult_less_0_iff not_less) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 maxU34 wmin.elims min_repU1 min_repU2 repL_def timesul tl.simps) qed next case ZeroPos assume bounds:"rl1 \ 0 \ 0 \ ru1 \ 0 \ rl2" have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast consider "r1 \ 0" | "r1 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"rl1 * rl2 \ 0" using "1" r2 bounds grl1 grl2 gru1 gru2 by (simp add: mult_le_0_iff) have g2:"0 \ r1 * r2" using "1" r2 bounds grl1 grl2 gru1 gru2 by (simp) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis by (metis repL_def timesll tl.simps up maxU12 maxU34 wmin.elims min_repU2 min_repU1) next case 2 have bound:"ru2 \ 0" using "2" r2 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto then have g1:"rl1 * ru2 \ rl1 * r2" using "2" r2 bounds grl1 grl2 gru1 gru2 mult_le_cancel_left by fastforce have g2:"rl1 * r2 \ r1 * r2" using "2" r2 bounds grl1 grl2 gru1 gru2 mult_le_0_iff mult_le_cancel_right by fastforce from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 wmin.elims min_repU2 min.coboundedI1 maxt34 repL_def tl.simps) qed next case ZeroNeg assume bounds:"rl1 \ 0 \ 0 \ ru1 \ ru2 \ 0" have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast consider (Pos) "r1 \ 0" | (Neg) "r1 \ 0" using le_cases by auto then show ?thesis proof (cases) case Pos have bound:"rl2 \ 0" using Pos r2 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto then have g1:"ru1 * rl2 \ ru1 * r2" using Pos bounds grl1 grl2 gru1 gru2 mult_le_cancel_left by fastforce have p1:"\a::real. (0 \ - a) = (a \ 0)" by(auto) have p2:"\a b::real. (- a \ - b) = (b \ a)" by auto have g2:"ru1 * r2 \ r1 * r2" using Pos r2 bounds grl1 grl2 gru1 gru2 p1 p2 by (simp add: mult_right_mono_neg) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 maxU34 wmin.elims min_repU2 min_repU1 repL_def timesul tl.simps) next case Neg have g1:"ru1 * ru2 \ 0" using Neg r2 bounds grl1 grl2 gru1 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using Neg r2 zero_le_mult_iff by blast from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt34 by (metis repL_def tl.simps) qed next case PosZero assume bounds:"0 \ rl1 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have bound:"0 \ ru1" using r1 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"rl1 * rl2 \ 0" using r1 "1" bounds grl1 grl2 gru1 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using r1 "1" bounds grl1 grl2 gru1 gru2 zero_le_mult_iff by blast from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt12 maxt34 repL_def timesll tl.simps by metis next case 2 have g1:"ru1 * rl2 \ ru1 * r2" using r1 "2" bounds bound grl1 grl2 gru1 gru2 using mult_left_mono by blast have g2:"ru1 * r2 \ r1 * r2" using r1 "2" bounds bound grl2 gru2 by (metis mult_left_mono_neg gru1 mult.commute) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 maxt34 by (metis repL_def timesul tl.simps) qed next case NegZero assume bounds:"ru1 \ 0 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have bound:"rl1 \ 0" using r1 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 assume r2:"r2 \ 0" have g1:"rl1 * ru2 \ rl1 * r2" using r1 r2 bounds bound grl1 grl2 gru1 gru2 by (metis mult_le_cancel_left leD) have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_right_mono by (simp add: mult_le_0_iff) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU min_repU2 min_repU1 min.coboundedI1 maxt34 by (metis min_repU2 repL_def tl.simps) next case 2 assume r2:"r2 \ 0" have lower:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have g1:"ru1 * ru2 \ 0" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using r1 r2 by (simp add: zero_le_mult_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt34 by (metis repL_def tl.simps) qed next case NegNeg assume bounds:"ru1 \ 0 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds dual_order.trans grl2 r2 by blast have g1:"ru1 * ru2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using not_less mult_le_cancel_left by metis have g2:"ru1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_le_cancel_left mult.commute not_le lower1 lower2 by metis from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt34 by (metis repL_def tl.simps) next case NegPos assume bounds:"ru1 \ 0 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using bounds by auto have upper2:"ru2 \ 0" using bounds dual_order.trans gru2 r2 by blast have g1:"rl1 * ru2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 less_le_trans upper1 lower2 by (metis mult_le_cancel_left not_less) have g2:"rl1 * r2 \ r1 * r2" using r1 upper1 r2 mult_right_mono mult_le_0_iff grl1 by blast from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 maxt12 maxt34 by (metis repL_def timeslu tl.simps) next case PosNeg assume bounds:"0 \ rl1 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using dual_order.trans grl2 r2 by blast have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using bounds by auto have g1:"ru1 * rl2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_left_mono less_le_trans not_less by metis have g2:"ru1 * r2 \ r1 * r2" using r1 lower1 r2 not_less gru2 gru1 grl1 grl2 by (metis mult_le_cancel_left mult.commute) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show "tl l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>L r1 * r2" using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 by (metis repL_def timesul tl.simps) next case PosPos assume bounds:"0 \ rl1 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using dual_order.trans gru2 u2 r2 bounds by blast have g1:"rl1 * rl2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_left_mono using leD leI less_le_trans by auto have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by blast from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU min_repU2 min_repU1 min.coboundedI1 maxt12 maxt34 by (metis repL_def tl.simps) qed qed text\Most significant bit only changes under successor when all other bits are 1\ lemma msb_succ: fixes w :: "32 Word.word" assumes neq1:"uint w \ 0xFFFFFFFF" assumes neq2:"uint w \ 0x7FFFFFFF" shows "msb (w + 1) = msb w" proof - have "w \ 0xFFFFFFFF" using neq1 by auto then have neqneg1:"w \ -1" by auto have "w \ 0x7FFFFFFF" using neq2 by auto then have neqneg2:"w \ (2^31)-1" by auto show ?thesis using neq1 neq2 unfolding msb_big using Word_Lemmas.word_le_make_less[of "w + 1" "0x80000000"] Word_Lemmas.word_le_make_less[of "w " "0x80000000"] neqneg1 neqneg2 by auto qed text\Negation commutes with msb except at edge cases\ lemma msb_non_min: fixes w :: "32 Word.word" assumes neq1:"uint w \ 0" assumes neq2:"uint w \ ((2^(len_of (TYPE(31)))))" shows "msb (uminus w) = HOL.Not(msb(w))" proof - have fact1:"uminus w = word_succ (~~ w)" by (rule twos_complement) have fact2:"msb (~~w) = HOL.Not(msb w)" using word_ops_msb[of w] by auto have neqneg1:"w \ 0" using neq1 by auto have not_undef:"w \ 0x80000000" using neq2 by auto then have neqneg2:"w \ (2^31)" by auto from \w \ 0\ have \~~ w \ ~~ 0\ by (simp only: bit.compl_eq_compl_iff) simp then have "(~~ w) \ 0xFFFFFFFF" by auto then have uintNeq1:"uint (~~ w) \ 0xFFFFFFFF" using uint_distinct[of "~~w" "0xFFFFFFFF"] by auto from \w \ 2 ^ 31\ have \~~ w \ ~~ 2 ^ 31\ by (simp only: bit.compl_eq_compl_iff) simp then have "(~~ w) \ 0x7FFFFFFF" by auto then have uintNeq2:" uint (~~ w) \ 0x7FFFFFFF" using uint_distinct[of "~~w" "0x7FFFFFFF"] by auto have fact3:"msb ((~~w) + 1) = msb (~~w)" apply(rule msb_succ[of "~~w"]) using neq1 neq2 uintNeq1 uintNeq2 by auto show "msb (uminus w) = HOL.Not(msb(w))" using fact1 fact2 fact3 by (simp add: word_succ_p1) qed text\Only 0x80000000 preserves msb=1 under negation\ lemma msb_min_neg: fixes w::"word" assumes msb1:"msb (- w)" assumes msb2:"msb w" shows "uint w = ((2^(len_of (TYPE(31)))))" proof (rule ccontr) from \msb w\ have \w \ 0\ using word_msb_0 by auto then have \uint w \ 0\ by transfer simp moreover assume \uint w \ 2 ^ LENGTH(31)\ ultimately have \msb (- w) \ \ msb w\ by (rule msb_non_min) with assms show False by simp qed text\Only 0x00000000 preserves msb=0 under negation\ lemma msb_zero: fixes w::"word" assumes msb1:"\ msb (- w)" assumes msb2:"\ msb w" shows "uint w = 0" proof - have neq:"w \ ((2 ^ len_of TYPE(31))::word)" using msb1 msb2 by auto have eq:"uint ((2 ^ len_of TYPE(31))::word) = 2 ^ len_of TYPE(31)" by auto then have neq:"uint w \ uint ((2 ^ len_of TYPE(31))::word)" using uint_distinct[of w "2^len_of TYPE(31)"] neq eq by auto show ?thesis using msb1 msb2 minus_zero msb_non_min[of w] neq by force qed text\Finite numbers alternate msb under negation\ lemma msb_pos: fixes w::"word" assumes msb1:"msb (- w)" assumes msb2:"\ msb w" shows "uint w \ {1 .. (2^((len_of TYPE(32)) - 1))-1}" proof - have main: "w \ {1 .. (2^((len_of TYPE(32)) - 1))-1}" using msb1 msb2 apply(clarsimp) unfolding word_msb_sint apply(rule conjI) apply (metis neg_equal_0_iff_equal not_le word_less_1) proof - have imp:"w \ 0x80000000 \ False" proof - assume geq:"w \ 0x80000000" then have "msb w" using Word_Lemmas.msb_big[of w] by auto then show False using msb2 by auto qed have mylem:"\w1 w2::word. uint w1 \ uint w2 \ w1 \ w2" subgoal for w1 w2 by (simp add: word_le_def) done have mylem2:"\w1 w2::word. w1 > w2 \ uint w1 > uint w2" subgoal for w1 w2 by (simp add: word_less_def) done have gr_to_geq:"w > 0x7FFFFFFF \ w \ 0x80000000" apply(rule mylem) using mylem2[of "0x7FFFFFFF" "w"] by auto have taut:"w \ 0x7FFFFFFF \ w > 0x7FFFFFFF" by auto then show "w \ 0x7FFFFFFF" using imp taut gr_to_geq by auto qed have set_eq:"(uint ` (({1..(minus(2 ^ (minus(len_of TYPE(32)) 1)) 1)})::word set)) = ({1..minus(2 ^ (minus (len_of TYPE(32)) 1)) 1}::int set)" apply(auto simp add: word_le_def) subgoal for xa proof - assume lower:"1 \ xa" and upper:"xa \ 2147483647" then have in_range:"xa \ {0 .. 2^32-1}" by auto then have "xa \ range (uint::word \ int)" unfolding word_uint.Rep_range uints_num by auto then obtain w::word where xaw:"xa = uint w" by auto then have "w \ {1..0x7FFFFFFF} " using lower upper apply(clarsimp, auto) by (auto simp add: word_le_def) then show ?thesis using uint_distinct uint_distinct main image_eqI word_le_def xaw by blast qed done then show "uint w \ {1..2 ^ (len_of TYPE(32) - 1) - 1}" using uint_distinct uint_distinct main image_eqI by blast qed lemma msb_neg: fixes w::"word" assumes msb1:"\ msb (- w)" assumes msb2:"msb w" shows "uint w \ {2^((len_of TYPE(32) - 1))+1 .. 2^((len_of TYPE(32)))-1}" proof - have mylem:"\w1 w2::word. uint w1 \ uint w2 \ w1 \ w2" by (simp add: word_le_def) have mylem2:"\w1 w2::word. w1 > w2 \ uint w1 > uint w2" by (simp add: word_less_def) have gr_to_geq:"w > 0x80000000 \ w \ 0x80000001" apply(rule mylem) using mylem2[of "0x80000000" "w"] by auto have taut:"w \ 0x80000000 \ 0x80000000 < w" by auto have imp:"w \ 0x80000000 \ False" proof - assume geq:"w \ 0x80000000" then have "(msb (-w))" using Word_Lemmas.msb_big[of "-w"] Word_Lemmas.msb_big[of "w"] by (simp add: msb2) then show False using msb1 by auto qed have main: "w \ {2^((len_of TYPE(32)) - 1)+1 .. 2^((len_of TYPE(32)))-1}" using msb1 msb2 apply(clarsimp) unfolding word_msb_sint proof - show "0x80000001 \ w" using imp taut gr_to_geq by auto qed have set_eq:"(uint ` (({2^((len_of TYPE(32) - 1))+1 .. 2^((len_of TYPE(32)))-1})::word set)) = {2^((len_of TYPE(32) - 1))+1 .. 2^((len_of TYPE(32)))-1}" apply(auto) subgoal for xa by (simp add: word_le_def) subgoal for w using uint_lt [of w] by simp subgoal for xa proof - assume lower:"2147483649 \ xa" and upper:"xa \ 4294967295" then have in_range:"xa \ {0x80000000 .. 0xFFFFFFFF}" by auto then have "xa \ range (uint::word \ int)" unfolding word_uint.Rep_range uints_num by auto then obtain w::word where xaw:"xa = uint w" by auto then have the_in:"w \ {0x80000001 .. 0xFFFFFFFF} " using lower upper by (auto simp add: word_le_def) have the_eq:"(0xFFFFFFFF::word) = -1" by auto from the_in the_eq have "w \ {0x80000001 .. -1}" by auto then show ?thesis using uint_distinct uint_distinct main image_eqI word_le_def xaw by blast qed done then show "uint w \ {2^((len_of TYPE(32)) - 1)+1 .. 2^((len_of TYPE(32)))-1}" using uint_distinct uint_distinct main image_eqI by blast qed text\2s-complement commutes with negation except edge cases\ lemma sint_neg_hom: fixes w :: "32 Word.word" shows "uint w \ ((2^(len_of (TYPE(31))))) \ (sint(-w) = -(sint w))" unfolding word_sint_msb_eq apply auto subgoal using msb_min_neg by auto prefer 3 subgoal using msb_zero[of w] by (simp add: msb_zero) proof - assume msb1:"msb (- w)" assume msb2:"\ msb w" have "uint w \ {1 .. (2^((len_of TYPE(32)) - 1))-1}" using msb_pos[OF msb1 msb2] by auto then have bound:"uint w \ {1 .. 0x7FFFFFFF}" by auto have size:"size (w::32 Word.word) = 32" using Word.word_size[of w] by auto have lem:"\x::int. \n::nat. x \ {1..(2^n)-1} \ ((- x) mod (2^n)) - (2^n) = - x" subgoal for x n apply(cases "x mod 2^n = 0") by(auto simp add: Divides.zmod_zminus1_eq_if[of x "2^n"]) done have lem_rule:"uint w \ {1..2 ^ 32 - 1} \ (- uint w mod 4294967296) - 4294967296 = - uint w" using lem[of "uint w" 32] by auto have almost:"- uint w mod 4294967296 - 4294967296 = - uint w" apply(rule lem_rule) using bound by auto show "uint (- w) - 2 ^ size (- w) = - uint w" using bound unfolding Word.uint_word_ariths word_size_neg by (auto simp add: size almost) next assume neq:"uint w \ 0x80000000" assume msb1:"\ msb (- w)" assume msb2:"msb w" have bound:"uint w \ {0x80000001.. 0xFFFFFFFF}" using msb1 msb2 msb_neg by auto have size:"size (w::32 Word.word) = 32" using Word.word_size[of w] by auto have lem:"\x::int. \n::nat. x \ {1..(2^n)-1} \ (-x mod (2^n)) = (2^n) - x" subgoal for x n apply(auto) apply(cases "x mod 2^n = 0") by (simp add: Divides.zmod_zminus1_eq_if[of x "2^n"])+ done from bound have wLeq: "uint w \ 4294967295" and wGeq: "2147483649 \ uint w" by auto from wLeq have wLeq':"uint w \ 4294967296" by fastforce have f3: "(0 \ 4294967296 + - 1 * uint w + - 1 * ((4294967296 + - 1 * uint w) mod 4294967296)) = (uint w + (4294967296 + - 1 * uint w) mod 4294967296 \ 4294967296)" by auto have f4: "(0 \ 4294967296 + - 1 * uint w) = (uint w \ 4294967296)" by auto have f5: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * (i mod ia)" by (simp add: zmod_le_nonneg_dividend) then have f6: "uint w + (4294967296 + - 1 * uint w) mod 4294967296 \ 4294967296" using f4 f3 wLeq' by blast have f7: "4294967296 + - 1 * uint w + - 4294967296 = - 1 * uint w" by auto have f8: "- (1::int) * 4294967296 = - 4294967296" by auto have f9: "(0 \ - 1 * uint w) = (uint w \ 0)" by auto have f10: "(4294967296 + -1 * uint w + -1 * ((4294967296 + -1 * uint w) mod 4294967296) \ 0) = (4294967296 \ uint w + (4294967296 + - 1 * uint w) mod 4294967296)" by auto have f11: "\ 4294967296 \ (0::int)" by auto have f12: "\x0. ((0::int) < x0) = (\ x0 \ 0)" by auto have f13: "\x0 x1. ((x1::int) < x0) = (\ 0 \ x1 + - 1 * x0)" by auto have f14: "\x0 x1. ((x1::int) \ x1 mod x0) = (x1 + - 1 * (x1 mod x0) \ 0)" by auto have "\ uint w \ 0" using wGeq by fastforce then have "4294967296 \ uint w + (4294967296 + - 1 * uint w) mod 4294967296" using f14 f13 f12 f11 f10 f9 f8 f7 by (metis (no_types) int_mod_ge) then show "uint (- w) = 2 ^ size w - uint w" using f6 unfolding Word.uint_word_ariths by (auto simp add: size f4) qed text\2s-complement encoding is injective\ lemma sint_dist: fixes x y ::word assumes "x \ y" shows "sint x \ sint y" by (simp add: assms) subsection\Negation\ fun wneg :: "word \ word" where "wneg w = (if w = NEG_INF then POS_INF else if w = POS_INF then NEG_INF else -w)" text\word negation is correct\ lemma wneg_lemma: assumes eq:"w \\<^sub>E (r::real)" shows "wneg w \\<^sub>E -r" apply(rule repe.cases[OF eq]) apply(auto intro!: repNEG_INF repPOS_INF simp add: repe.simps)[2] subgoal for ra proof - assume eq:"w = ra" assume i:"r = (real_of_int (sint ra))" assume bounda:" (real_of_int (sint ra)) < (real_of_int (sint POS_INF))" assume boundb:" (real_of_int (sint NEG_INF)) < (real_of_int (sint ra))" have raNeq:"ra \ 2147483647" using sint_range[OF bounda boundb] by (auto) have raNeqUndef:"ra \ 2147483648" using int_not_undef[OF bounda boundb] by (auto) have "uint ra \ uint ((2 ^ len_of TYPE(31))::word)" apply (rule uint_distinct) using raNeqUndef by auto then have raNeqUndefUint:"uint ra \ ((2 ^ len_of TYPE(31)))" by auto have res1:"wneg w \\<^sub>E (real_of_int (sint (wneg w)))" apply (rule repINT) using sint_range[OF bounda boundb] sint_neg_hom[of ra, OF raNeqUndefUint] raNeq raNeqUndefUint raNeqUndef eq by(auto) have res2:"- r = (real_of_int (sint (wneg w)))" using eq bounda boundb i sint_neg_hom[of ra, OF raNeqUndefUint] raNeq raNeqUndef eq apply auto apply transfer apply simp done show ?thesis using res1 res2 by auto qed done subsection\Comparison\ fun wgreater :: "word \ word \ bool" where "wgreater w1 w2 = (sint w1 > sint w2)" lemma neg_less_contra:"\x. Suc x < - (Suc x) \ False" by auto text\Comparison < is correct\ lemma wgreater_lemma:"w1 \\<^sub>L (r1::real) \ w2 \\<^sub>U r2 \ wgreater w1 w2 \ r1 > r2" proof (auto simp add: repU_def repL_def) fix r'\<^sub>1 r'\<^sub>2 assume sint_le:"sint w1 > sint w2" then have sless:"(w2 1 \ r1" assume r2_leq:"r2 \ r'\<^sub>2" assume wr1:"w1 \\<^sub>E r'\<^sub>1" assume wr2:"w2 \\<^sub>E r'\<^sub>2" have greater:"r'\<^sub>1 > r'\<^sub>2" using wr1 wr2 apply(auto simp add: repe.simps) prefer 4 using sless sint_le apply (auto simp add: less_le_trans not_le) apply transfer apply simp apply transfer apply simp apply transfer apply simp done show "r1 > r2" using r1_leq r2_leq greater by auto qed text\Comparison $\geq$ of words\ fun wgeq :: "word \ word \ bool" where "wgeq w1 w2 = ((\ ((w2 = NEG_INF \ w1 = NEG_INF) \(w2 = POS_INF \ w1 = POS_INF))) \ (sint w2 \ sint w1))" text\Comparison $\geq$ of words is correct\ lemma wgeq_lemma:"w1 \\<^sub>L r1 \ w2 \\<^sub>U (r2::real) \ wgeq w1 w2 \ r1 \ r2" proof (unfold wgeq.simps) assume assms:"\ (w2 = NEG_INF \ w1 = NEG_INF \ w2 = POS_INF \ w1 = POS_INF) \ sint w2 \ sint w1" assume a1:"w1 \\<^sub>L r1" and a2:"w2 \\<^sub>U (r2::real)" from assms have sint_le:"sint w2 \ sint w1" by auto then have sless:"(w2 <=s w1)" using word_sless_alt word_sle_def by auto obtain r'\<^sub>1 r'\<^sub>2 where r1_leq:"r'\<^sub>1 \ r1" and r2_leq:"r2 \ r'\<^sub>2" and wr1:"w1 \\<^sub>E r'\<^sub>1" and wr2:"w2 \\<^sub>E r'\<^sub>2" using a1 a2 unfolding repU_def repL_def by auto from assms have check1:"\ (w1 = NEG_INF \ w2 = NEG_INF)" by auto from assms have check2:"\ (w1 = POS_INF \ w2 = POS_INF)" by auto have less:"r'\<^sub>2 \ r'\<^sub>1" using sless sint_le check1 check2 repe.simps wr2 wr1 apply (auto simp add: repe.simps) apply transfer apply simp apply transfer apply simp apply transfer apply simp apply transfer apply simp apply transfer apply simp apply transfer apply simp apply transfer apply simp apply transfer apply simp done show "r1 \ r2" using r1_leq r2_leq less by auto qed subsection\Absolute value\ text\Absolute value of word\ fun wabs :: "word \ word" where "wabs l1 = (wmax l1 (wneg l1))" text\Correctness of wmax\ lemma wabs_lemma: assumes eq:"w \\<^sub>E (r::real)" shows "wabs w \\<^sub>E (abs r)" proof - have w:"wmax w (wneg w) \\<^sub>E max r (-r)" by (rule wmax_lemma[OF eq wneg_lemma[OF eq]]) have r:"max r (-r) = abs r" by auto from w r show ?thesis by auto qed declare more_real_of_word_simps [simp del] end diff --git a/thys/Iptables_Semantics/Primitive_Matchers/Ports.thy b/thys/Iptables_Semantics/Primitive_Matchers/Ports.thy --- a/thys/Iptables_Semantics/Primitive_Matchers/Ports.thy +++ b/thys/Iptables_Semantics/Primitive_Matchers/Ports.thy @@ -1,43 +1,43 @@ theory Ports imports - "HOL-Word.Word" + "HOL-Library.Word" "../Common/WordInterval_Lists" L4_Protocol_Flags begin section\Ports (layer 4)\ text\E.g. source and destination ports for TCP/UDP\ text\list of (start, end) port ranges\ type_synonym raw_ports = "(16 word \ 16 word) list" fun ports_to_set :: "raw_ports \ (16 word) set" where "ports_to_set [] = {}" | "ports_to_set ((s,e)#ps) = {s..e} \ ports_to_set ps" lemma ports_to_set: "ports_to_set pts = \ {{s..e} | s e . (s,e) \ set pts}" proof(induction pts) case Nil thus ?case by simp next case (Cons p pts) thus ?case by(cases p, simp, blast) qed text\We can reuse the wordinterval theory to reason about ports\ lemma ports_to_set_wordinterval: "ports_to_set ps = wordinterval_to_set (l2wi ps)" by(induction ps rule: l2wi.induct) (auto) text\inverting a raw listing of ports\ definition "raw_ports_invert" :: "raw_ports \ raw_ports" where "raw_ports_invert ps = wi2l (wordinterval_invert (l2wi ps))" lemma raw_ports_invert: "ports_to_set (raw_ports_invert ps) = - ports_to_set ps" by(auto simp add: raw_ports_invert_def l2wi_wi2l ports_to_set_wordinterval) text\A port always belongs to a protocol! We must not lose this information. You should never use @{typ raw_ports} directly\ datatype ipt_l4_ports = L4Ports primitive_protocol raw_ports end diff --git a/thys/JinjaThreads/Common/BinOp.thy b/thys/JinjaThreads/Common/BinOp.thy --- a/thys/JinjaThreads/Common/BinOp.thy +++ b/thys/JinjaThreads/Common/BinOp.thy @@ -1,589 +1,589 @@ (* Title: JinjaThreads/Common/BinOp.thy Author: Andreas Lochbihler *) section \Binary Operators\ theory BinOp imports - WellForm "HOL-Word.Traditional_Syntax" + WellForm "Word_Lib.Traditional_Syntax" begin datatype bop = \ \names of binary operations\ Eq | NotEq | LessThan | LessOrEqual | GreaterThan | GreaterOrEqual | Add | Subtract | Mult | Div | Mod | BinAnd | BinOr | BinXor | ShiftLeft | ShiftRightZeros | ShiftRightSigned subsection\The semantics of binary operators\ type_synonym 'addr binop_ret = "'addr val + 'addr" \ \a value or the address of an exception\ fun binop_LessThan :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_LessThan (Intg i1) (Intg i2) = Some (Inl (Bool (i1 'addr val \ 'addr binop_ret option" where "binop_LessOrEqual (Intg i1) (Intg i2) = Some (Inl (Bool (i1 <=s i2)))" | "binop_LessOrEqual v1 v2 = None" fun binop_GreaterThan :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_GreaterThan (Intg i1) (Intg i2) = Some (Inl (Bool (i2 'addr val \ 'addr binop_ret option" where "binop_GreaterOrEqual (Intg i1) (Intg i2) = Some (Inl (Bool (i2 <=s i1)))" | "binop_GreaterOrEqual v1 v2 = None" fun binop_Add :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_Add (Intg i1) (Intg i2) = Some (Inl (Intg (i1 + i2)))" | "binop_Add v1 v2 = None" fun binop_Subtract :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_Subtract (Intg i1) (Intg i2) = Some (Inl (Intg (i1 - i2)))" | "binop_Subtract v1 v2 = None" fun binop_Mult :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_Mult (Intg i1) (Intg i2) = Some (Inl (Intg (i1 * i2)))" | "binop_Mult v1 v2 = None" fun binop_BinAnd :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_BinAnd (Intg i1) (Intg i2) = Some (Inl (Intg (i1 AND i2)))" | "binop_BinAnd (Bool b1) (Bool b2) = Some (Inl (Bool (b1 \ b2)))" | "binop_BinAnd v1 v2 = None" fun binop_BinOr :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_BinOr (Intg i1) (Intg i2) = Some (Inl (Intg (i1 OR i2)))" | "binop_BinOr (Bool b1) (Bool b2) = Some (Inl (Bool (b1 \ b2)))" | "binop_BinOr v1 v2 = None" fun binop_BinXor :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_BinXor (Intg i1) (Intg i2) = Some (Inl (Intg (i1 XOR i2)))" | "binop_BinXor (Bool b1) (Bool b2) = Some (Inl (Bool (b1 \ b2)))" | "binop_BinXor v1 v2 = None" fun binop_ShiftLeft :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_ShiftLeft (Intg i1) (Intg i2) = Some (Inl (Intg (i1 << unat (i2 AND 0x1f))))" | "binop_ShiftLeft v1 v2 = None" fun binop_ShiftRightZeros :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_ShiftRightZeros (Intg i1) (Intg i2) = Some (Inl (Intg (i1 >> unat (i2 AND 0x1f))))" | "binop_ShiftRightZeros v1 v2 = None" fun binop_ShiftRightSigned :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_ShiftRightSigned (Intg i1) (Intg i2) = Some (Inl (Intg (i1 >>> unat (i2 AND 0x1f))))" | "binop_ShiftRightSigned v1 v2 = None" text \ Division on @{typ "'a word"} is unsigned, but JLS specifies signed division. \ definition word_sdiv :: "'a :: len word \ 'a word \ 'a word" (infixl "sdiv" 70) where [code]: "x sdiv y = (let x' = sint x; y' = sint y; negative = (x' < 0) \ (y' < 0); result = abs x' div abs y' in word_of_int (if negative then -result else result))" definition word_smod :: "'a :: len word \ 'a word \ 'a word" (infixl "smod" 70) where [code]: "x smod y = (let x' = sint x; y' = sint y; negative = (x' < 0); result = abs x' mod abs y' in word_of_int (if negative then -result else result))" declare word_sdiv_def [simp] word_smod_def [simp] lemma sdiv_smod_id: "(a sdiv b) * b + (a smod b) = a" proof - have F5: "\u::'a word. - (- u) = u" by simp have F7: "\v u::'a word. u + v = v + u" by (simp add: ac_simps) have F8: "\(w::'a word) (v::int) u::int. word_of_int u + word_of_int v * w = word_of_int (u + v * sint w)" by simp have "\u. u = - sint b \ word_of_int (sint a mod u + - (- u * (sint a div u))) = a" using F5 by simp hence "word_of_int (sint a mod - sint b + - (sint b * (sint a div - sint b))) = a" by (metis equation_minus_iff) hence "word_of_int (sint a mod - sint b) + word_of_int (- (sint a div - sint b)) * b = a" using F8 by (simp add: ac_simps) hence eq: "word_of_int (- (sint a div - sint b)) * b + word_of_int (sint a mod - sint b) = a" using F7 by simp show ?thesis proof(cases "sint a < 0") case True note a = this show ?thesis proof(cases "sint b < 0") case True with a show ?thesis by simp (metis F7 F8 eq minus_equation_iff minus_mult_minus mod_div_mult_eq) next case False from eq have "word_of_int (- (- sint a div sint b)) * b + word_of_int (- (- sint a mod sint b)) = a" by (metis div_minus_right mod_minus_right) with a False show ?thesis by simp qed next case False note a = this show ?thesis proof(cases "sint b < 0") case True with a eq show ?thesis by simp next case False with a show ?thesis by (simp add: F7 F8) qed qed qed notepad begin have " 5 sdiv ( 3 :: word32) = 1" and " 5 smod ( 3 :: word32) = 2" and " 5 sdiv (-3 :: word32) = -1" and " 5 smod (-3 :: word32) = 2" and "(-5) sdiv ( 3 :: word32) = -1" and "(-5) smod ( 3 :: word32) = -2" and "(-5) sdiv (-3 :: word32) = 1" and "(-5) smod (-3 :: word32) = -2" and "-2147483648 sdiv 1 = (-2147483648 :: word32)" by eval+ end context heap_base begin fun binop_Mod :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_Mod (Intg i1) (Intg i2) = Some (if i2 = 0 then Inr (addr_of_sys_xcpt ArithmeticException) else Inl (Intg (i1 smod i2)))" | "binop_Mod v1 v2 = None" fun binop_Div :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_Div (Intg i1) (Intg i2) = Some (if i2 = 0 then Inr (addr_of_sys_xcpt ArithmeticException) else Inl (Intg (i1 sdiv i2)))" | "binop_Div v1 v2 = None" primrec binop :: "bop \ 'addr val \ 'addr val \ 'addr binop_ret option" where "binop Eq v1 v2 = Some (Inl (Bool (v1 = v2)))" | "binop NotEq v1 v2 = Some (Inl (Bool (v1 \ v2)))" | "binop LessThan = binop_LessThan" | "binop LessOrEqual = binop_LessOrEqual" | "binop GreaterThan = binop_GreaterThan" | "binop GreaterOrEqual = binop_GreaterOrEqual" | "binop Add = binop_Add" | "binop Subtract = binop_Subtract" | "binop Mult = binop_Mult" | "binop Mod = binop_Mod" | "binop Div = binop_Div" | "binop BinAnd = binop_BinAnd" | "binop BinOr = binop_BinOr" | "binop BinXor = binop_BinXor" | "binop ShiftLeft = binop_ShiftLeft" | "binop ShiftRightZeros = binop_ShiftRightZeros" | "binop ShiftRightSigned = binop_ShiftRightSigned" end lemma [simp]: "(binop_LessThan v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Bool (i1 (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Bool (i1 <=s i2)))" by(cases "(v1, v2)" rule: binop_LessOrEqual.cases) auto lemma [simp]: "(binop_GreaterThan v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Bool (i2 (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Bool (i2 <=s i1)))" by(cases "(v1, v2)" rule: binop_GreaterOrEqual.cases) auto lemma [simp]: "(binop_Add v\<^sub>1 v\<^sub>2 = Some va) \ (\i\<^sub>1 i\<^sub>2. v\<^sub>1 = Intg i\<^sub>1 \ v\<^sub>2 = Intg i\<^sub>2 \ va = Inl (Intg (i\<^sub>1+i\<^sub>2)))" by(cases "(v\<^sub>1, v\<^sub>2)" rule: binop_Add.cases) auto lemma [simp]: "(binop_Subtract v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 - i2)))" by(cases "(v1, v2)" rule: binop_Subtract.cases) auto lemma [simp]: "(binop_Mult v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 * i2)))" by(cases "(v1, v2)" rule: binop_Mult.cases) auto lemma [simp]: "(binop_BinAnd v1 v2 = Some va) \ (\b1 b2. v1 = Bool b1 \ v2 = Bool b2 \ va = Inl (Bool (b1 \ b2))) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 AND i2)))" by(cases "(v1, v2)" rule: binop_BinAnd.cases) auto lemma [simp]: "(binop_BinOr v1 v2 = Some va) \ (\b1 b2. v1 = Bool b1 \ v2 = Bool b2 \ va = Inl (Bool (b1 \ b2))) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 OR i2)))" by(cases "(v1, v2)" rule: binop_BinOr.cases) auto lemma [simp]: "(binop_BinXor v1 v2 = Some va) \ (\b1 b2. v1 = Bool b1 \ v2 = Bool b2 \ va = Inl (Bool (b1 \ b2))) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 XOR i2)))" by(cases "(v1, v2)" rule: binop_BinXor.cases) auto lemma [simp]: "(binop_ShiftLeft v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 << unat (i2 AND 0x1f))))" by(cases "(v1, v2)" rule: binop_ShiftLeft.cases) auto lemma [simp]: "(binop_ShiftRightZeros v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 >> unat (i2 AND 0x1f))))" by(cases "(v1, v2)" rule: binop_ShiftRightZeros.cases) auto lemma [simp]: "(binop_ShiftRightSigned v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 >>> unat (i2 AND 0x1f))))" by(cases "(v1, v2)" rule: binop_ShiftRightSigned.cases) auto context heap_base begin lemma [simp]: "(binop_Mod v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = (if i2 = 0 then Inr (addr_of_sys_xcpt ArithmeticException) else Inl (Intg(i1 smod i2))))" by(cases "(v1, v2)" rule: binop_Mod.cases) auto lemma [simp]: "(binop_Div v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = (if i2 = 0 then Inr (addr_of_sys_xcpt ArithmeticException) else Inl (Intg(i1 sdiv i2))))" by(cases "(v1, v2)" rule: binop_Div.cases) auto end subsection \Typing for binary operators\ inductive WT_binop :: "'m prog \ ty \ bop \ ty \ ty \ bool" ("_ \ _\_\_ :: _" [51,0,0,0,51] 50) where WT_binop_Eq: "P \ T1 \ T2 \ P \ T2 \ T1 \ P \ T1\Eq\T2 :: Boolean" | WT_binop_NotEq: "P \ T1 \ T2 \ P \ T2 \ T1 \ P \ T1\NotEq\T2 :: Boolean" | WT_binop_LessThan: "P \ Integer\LessThan\Integer :: Boolean" | WT_binop_LessOrEqual: "P \ Integer\LessOrEqual\Integer :: Boolean" | WT_binop_GreaterThan: "P \ Integer\GreaterThan\Integer :: Boolean" | WT_binop_GreaterOrEqual: "P \ Integer\GreaterOrEqual\Integer :: Boolean" | WT_binop_Add: "P \ Integer\Add\Integer :: Integer" | WT_binop_Subtract: "P \ Integer\Subtract\Integer :: Integer" | WT_binop_Mult: "P \ Integer\Mult\Integer :: Integer" | WT_binop_Div: "P \ Integer\Div\Integer :: Integer" | WT_binop_Mod: "P \ Integer\Mod\Integer :: Integer" | WT_binop_BinAnd_Bool: "P \ Boolean\BinAnd\Boolean :: Boolean" | WT_binop_BinAnd_Int: "P \ Integer\BinAnd\Integer :: Integer" | WT_binop_BinOr_Bool: "P \ Boolean\BinOr\Boolean :: Boolean" | WT_binop_BinOr_Int: "P \ Integer\BinOr\Integer :: Integer" | WT_binop_BinXor_Bool: "P \ Boolean\BinXor\Boolean :: Boolean" | WT_binop_BinXor_Int: "P \ Integer\BinXor\Integer :: Integer" | WT_binop_ShiftLeft: "P \ Integer\ShiftLeft\Integer :: Integer" | WT_binop_ShiftRightZeros: "P \ Integer\ShiftRightZeros\Integer :: Integer" | WT_binop_ShiftRightSigned: "P \ Integer\ShiftRightSigned\Integer :: Integer" lemma WT_binopI [intro]: "P \ T1 \ T2 \ P \ T2 \ T1 \ P \ T1\Eq\T2 :: Boolean" "P \ T1 \ T2 \ P \ T2 \ T1 \ P \ T1\NotEq\T2 :: Boolean" "bop = Add \ bop = Subtract \ bop = Mult \ bop = Mod \ bop = Div \ bop = BinAnd \ bop = BinOr \ bop = BinXor \ bop = ShiftLeft \ bop = ShiftRightZeros \ bop = ShiftRightSigned \ P \ Integer\bop\Integer :: Integer" "bop = LessThan \ bop = LessOrEqual \ bop = GreaterThan \ bop = GreaterOrEqual \ P \ Integer\bop\Integer :: Boolean" "bop = BinAnd \ bop = BinOr \ bop = BinXor \ P \ Boolean\bop\Boolean :: Boolean" by(auto intro: WT_binop.intros) inductive_cases [elim]: "P \ T1\Eq\T2 :: T" "P \ T1\NotEq\T2 :: T" "P \ T1\LessThan\T2 :: T" "P \ T1\LessOrEqual\T2 :: T" "P \ T1\GreaterThan\T2 :: T" "P \ T1\GreaterOrEqual\T2 :: T" "P \ T1\Add\T2 :: T" "P \ T1\Subtract\T2 :: T" "P \ T1\Mult\T2 :: T" "P \ T1\Div\T2 :: T" "P \ T1\Mod\T2 :: T" "P \ T1\BinAnd\T2 :: T" "P \ T1\BinOr\T2 :: T" "P \ T1\BinXor\T2 :: T" "P \ T1\ShiftLeft\T2 :: T" "P \ T1\ShiftRightZeros\T2 :: T" "P \ T1\ShiftRightSigned\T2 :: T" lemma WT_binop_fun: "\ P \ T1\bop\T2 :: T; P \ T1\bop\T2 :: T' \ \ T = T'" by(cases bop)(auto) lemma WT_binop_is_type: "\ P \ T1\bop\T2 :: T; is_type P T1; is_type P T2 \ \ is_type P T" by(cases bop) auto inductive WTrt_binop :: "'m prog \ ty \ bop \ ty \ ty \ bool" ("_ \ _\_\_ : _" [51,0,0,0,51] 50) where WTrt_binop_Eq: "P \ T1\Eq\T2 : Boolean" | WTrt_binop_NotEq: "P \ T1\NotEq\T2 : Boolean" | WTrt_binop_LessThan: "P \ Integer\LessThan\Integer : Boolean" | WTrt_binop_LessOrEqual: "P \ Integer\LessOrEqual\Integer : Boolean" | WTrt_binop_GreaterThan: "P \ Integer\GreaterThan\Integer : Boolean" | WTrt_binop_GreaterOrEqual: "P \ Integer\GreaterOrEqual\Integer : Boolean" | WTrt_binop_Add: "P \ Integer\Add\Integer : Integer" | WTrt_binop_Subtract: "P \ Integer\Subtract\Integer : Integer" | WTrt_binop_Mult: "P \ Integer\Mult\Integer : Integer" | WTrt_binop_Div: "P \ Integer\Div\Integer : Integer" | WTrt_binop_Mod: "P \ Integer\Mod\Integer : Integer" | WTrt_binop_BinAnd_Bool: "P \ Boolean\BinAnd\Boolean : Boolean" | WTrt_binop_BinAnd_Int: "P \ Integer\BinAnd\Integer : Integer" | WTrt_binop_BinOr_Bool: "P \ Boolean\BinOr\Boolean : Boolean" | WTrt_binop_BinOr_Int: "P \ Integer\BinOr\Integer : Integer" | WTrt_binop_BinXor_Bool: "P \ Boolean\BinXor\Boolean : Boolean" | WTrt_binop_BinXor_Int: "P \ Integer\BinXor\Integer : Integer" | WTrt_binop_ShiftLeft: "P \ Integer\ShiftLeft\Integer : Integer" | WTrt_binop_ShiftRightZeros: "P \ Integer\ShiftRightZeros\Integer : Integer" | WTrt_binop_ShiftRightSigned: "P \ Integer\ShiftRightSigned\Integer : Integer" lemma WTrt_binopI [intro]: "P \ T1\Eq\T2 : Boolean" "P \ T1\NotEq\T2 : Boolean" "bop = Add \ bop = Subtract \ bop = Mult \ bop = Div \ bop = Mod \ bop = BinAnd \ bop = BinOr \ bop = BinXor \ bop = ShiftLeft \ bop = ShiftRightZeros \ bop = ShiftRightSigned \ P \ Integer\bop\Integer : Integer" "bop = LessThan \ bop = LessOrEqual \ bop = GreaterThan \ bop = GreaterOrEqual \ P \ Integer\bop\Integer : Boolean" "bop = BinAnd \ bop = BinOr \ bop = BinXor \ P \ Boolean\bop\Boolean : Boolean" by(auto intro: WTrt_binop.intros) inductive_cases WTrt_binop_cases [elim]: "P \ T1\Eq\T2 : T" "P \ T1\NotEq\T2 : T" "P \ T1\LessThan\T2 : T" "P \ T1\LessOrEqual\T2 : T" "P \ T1\GreaterThan\T2 : T" "P \ T1\GreaterOrEqual\T2 : T" "P \ T1\Add\T2 : T" "P \ T1\Subtract\T2 : T" "P \ T1\Mult\T2 : T" "P \ T1\Div\T2 : T" "P \ T1\Mod\T2 : T" "P \ T1\BinAnd\T2 : T" "P \ T1\BinOr\T2 : T" "P \ T1\BinXor\T2 : T" "P \ T1\ShiftLeft\T2 : T" "P \ T1\ShiftRightZeros\T2 : T" "P \ T1\ShiftRightSigned\T2 : T" inductive_simps WTrt_binop_simps [simp]: "P \ T1\Eq\T2 : T" "P \ T1\NotEq\T2 : T" "P \ T1\LessThan\T2 : T" "P \ T1\LessOrEqual\T2 : T" "P \ T1\GreaterThan\T2 : T" "P \ T1\GreaterOrEqual\T2 : T" "P \ T1\Add\T2 : T" "P \ T1\Subtract\T2 : T" "P \ T1\Mult\T2 : T" "P \ T1\Div\T2 : T" "P \ T1\Mod\T2 : T" "P \ T1\BinAnd\T2 : T" "P \ T1\BinOr\T2 : T" "P \ T1\BinXor\T2 : T" "P \ T1\ShiftLeft\T2 : T" "P \ T1\ShiftRightZeros\T2 : T" "P \ T1\ShiftRightSigned\T2 : T" fun binop_relevant_class :: "bop \ 'm prog \ cname \ bool" where "binop_relevant_class Div = (\P C. P \ ArithmeticException \\<^sup>* C )" | "binop_relevant_class Mod = (\P C. P \ ArithmeticException \\<^sup>* C )" | "binop_relevant_class _ = (\P C. False)" lemma WT_binop_WTrt_binop: "P \ T1\bop\T2 :: T \ P \ T1\bop\T2 : T" by(auto elim: WT_binop.cases) context heap begin lemma binop_progress: "\ typeof\<^bsub>h\<^esub> v1 = \T1\; typeof\<^bsub>h\<^esub> v2 = \T2\; P \ T1\bop\T2 : T \ \ \va. binop bop v1 v2 = \va\" by(cases bop)(auto del: disjCI split del: if_split) lemma binop_type: assumes wf: "wf_prog wf_md P" and pre: "preallocated h" and type: "typeof\<^bsub>h\<^esub> v1 = \T1\" "typeof\<^bsub>h\<^esub> v2 = \T2\" "P \ T1\bop\T2 : T" shows "binop bop v1 v2 = \Inl v\ \ P,h \ v :\ T" and "binop bop v1 v2 = \Inr a\ \ P,h \ Addr a :\ Class Throwable" using type apply(case_tac [!] bop) apply(auto split: if_split_asm simp add: conf_def wf_preallocatedD[OF wf pre]) done lemma binop_relevant_class: assumes wf: "wf_prog wf_md P" and pre: "preallocated h" and bop: "binop bop v1 v2 = \Inr a\" and sup: "P \ cname_of h a \\<^sup>* C" shows "binop_relevant_class bop P C" using assms by(cases bop)(auto split: if_split_asm) end lemma WTrt_binop_fun: "\ P \ T1\bop\T2 : T; P \ T1\bop\T2 : T' \ \ T = T'" by(cases bop)(auto) lemma WTrt_binop_THE [simp]: "P \ T1\bop\T2 : T \ The (WTrt_binop P T1 bop T2) = T" by(auto dest: WTrt_binop_fun) lemma WTrt_binop_widen_mono: "\ P \ T1\bop\T2 : T; P \ T1' \ T1; P \ T2' \ T2 \ \ \T'. P \ T1'\bop\T2' : T' \ P \ T' \ T" by(cases bop)(auto elim!: WTrt_binop_cases) lemma WTrt_binop_is_type: "\ P \ T1\bop\T2 : T; is_type P T1; is_type P T2 \ \ is_type P T" by(cases bop) auto subsection \Code generator setup\ lemmas [code] = heap_base.binop_Div.simps heap_base.binop_Mod.simps heap_base.binop.simps code_pred (modes: i \ i \ i \ i \ o \ bool, i \ i \ i \ i \ i \ bool) WT_binop . code_pred (modes: i \ i \ i \ i \ o \ bool, i \ i \ i \ i \ i \ bool) WTrt_binop . lemma eval_WTrt_binop_i_i_i_i_o: "Predicate.eval (WTrt_binop_i_i_i_i_o P T1 bop T2) T \ P \ T1\bop\T2 : T" by(auto elim: WTrt_binop_i_i_i_i_oE intro: WTrt_binop_i_i_i_i_oI) lemma the_WTrt_binop_code: "(THE T. P \ T1\bop\T2 : T) = Predicate.the (WTrt_binop_i_i_i_i_o P T1 bop T2)" by(simp add: Predicate.the_def eval_WTrt_binop_i_i_i_i_o) end diff --git a/thys/JinjaThreads/Common/Value.thy b/thys/JinjaThreads/Common/Value.thy --- a/thys/JinjaThreads/Common/Value.thy +++ b/thys/JinjaThreads/Common/Value.thy @@ -1,100 +1,100 @@ (* Title: JinjaThreads/Common/Value.thy Author: David von Oheimb, Tobias Nipkow, Andreas Lochbihler Based on the Jinja theory Common/Value.thy by David von Oheimb and Tobias Nipkow *) section \Jinja Values\ theory Value imports TypeRel - "HOL-Word.Word" + "HOL-Library.Word" begin no_notation floor ("\_\") type_synonym word32 = "32 word" datatype 'addr val = Unit \ \dummy result value of void expressions\ | Null \ \null reference\ | Bool bool \ \Boolean value\ | Intg word32 \ \integer value\ | Addr 'addr \ \addresses of objects, arrays and threads in the heap\ primrec default_val :: "ty \ 'addr val" \ \default value for all types\ where "default_val Void = Unit" | "default_val Boolean = Bool False" | "default_val Integer = Intg 0" | "default_val NT = Null" | "default_val (Class C) = Null" | "default_val (Array A) = Null" lemma default_val_not_Addr: "default_val T \ Addr a" by(cases T)(simp_all) lemma Addr_not_default_val: "Addr a \ default_val T" by(cases T)(simp_all) primrec the_Intg :: "'addr val \ word32" where "the_Intg (Intg i) = i" primrec the_Addr :: "'addr val \ 'addr" where "the_Addr (Addr a) = a" fun is_Addr :: "'addr val \ bool" where "is_Addr (Addr a) = True" | "is_Addr _ = False" lemma is_AddrE [elim!]: "\ is_Addr v; \a. v = Addr a \ thesis \ \ thesis" by(cases v, auto) fun is_Intg :: "'addr val \ bool" where "is_Intg (Intg i) = True" | "is_Intg _ = False" lemma is_IntgE [elim!]: "\ is_Intg v; \i. v = Intg i \ thesis \ \ thesis" by(cases v, auto) fun is_Bool :: "'addr val \ bool" where "is_Bool (Bool b) = True" | "is_Bool _ = False" lemma is_BoolE [elim!]: "\ is_Bool v; \a. v = Bool a \ thesis \ \ thesis" by(cases v, auto) definition is_Ref :: "'addr val \ bool" where "is_Ref v \ v = Null \ is_Addr v" lemma is_Ref_def2: "is_Ref v = (v = Null \ (\a. v = Addr a))" by (cases v) (auto simp add: is_Ref_def) lemma [iff]: "is_Ref Null" by (simp add: is_Ref_def2) definition undefined_value :: "'addr val" where "undefined_value = Unit" lemma undefined_value_not_Addr: "undefined_value \ Addr a" "Addr a \ undefined_value" by(simp_all add: undefined_value_def) class addr = fixes hash_addr :: "'a \ int" and monitor_finfun_to_list :: "('a \f nat) \ 'a list" assumes "set (monitor_finfun_to_list f) = Collect (($) (finfun_dom f))" locale addr_base = fixes addr2thread_id :: "'addr \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" end diff --git a/thys/JinjaThreads/ROOT b/thys/JinjaThreads/ROOT --- a/thys/JinjaThreads/ROOT +++ b/thys/JinjaThreads/ROOT @@ -1,34 +1,33 @@ chapter AFP session JinjaThreads (AFP slow) = HOL + options [timeout = 28800] sessions - "HOL-Word" "Binomial-Heaps" "Finger-Trees" Automatic_Refinement Coinductive Collections FinFun Native_Word Refine_Monadic Trie directories "BV" "Basic" "Common" "Compiler" "DFA" "Examples" "Execute" "Framework" "J" "JVM" "MM" theories [document = false] "Basic/Basic_Main" theories JinjaThreads document_files "root.bib" "root.tex" diff --git a/thys/Kleene_Algebra/Matrix.thy b/thys/Kleene_Algebra/Matrix.thy --- a/thys/Kleene_Algebra/Matrix.thy +++ b/thys/Kleene_Algebra/Matrix.thy @@ -1,356 +1,356 @@ (* Title: Matrix Model of Kleene Algebra Author: Alasdair Armstrong, Georg Struth, Tjark Weber Maintainer: Georg Struth Tjark Weber *) section \Matrices\ theory Matrix -imports "HOL-Word.Word" Dioid +imports "HOL-Library.Word" Dioid begin text \In this section we formalise a perhaps more natural version of matrices of fixed dimension ($m \times n$-matrices). It is well known that such matrices over a Kleene algebra form a Kleene algebra~\cite{conway71regular}.\ subsection \Type Definition\ typedef (overloaded) 'a atMost = "{..Our matrix type is similar to \mbox{\'a^'n^'m\} from {\em HOL/Multivariate\_Analysis/Finite\_Cartesian\_Product.thy}, but (i)~we explicitly define a type constructor for matrices and square matrices, and (ii)~in the definition of operations, e.g., matrix multiplication, we impose weaker sort requirements on the element type.\ context notes [[typedef_overloaded]] begin datatype ('a,'m,'n) matrix = Matrix "'m atMost \ 'n atMost \ 'a" datatype ('a,'m) sqmatrix = SqMatrix "'m atMost \ 'm atMost \ 'a" end fun sqmatrix_of_matrix where "sqmatrix_of_matrix (Matrix A) = SqMatrix A" fun matrix_of_sqmatrix where "matrix_of_sqmatrix (SqMatrix A) = Matrix A" subsection \0 and 1\ instantiation matrix :: (zero,type,type) zero begin definition zero_matrix_def: "0 \ Matrix (\i j. 0)" instance .. end instantiation sqmatrix :: (zero,type) zero begin definition zero_sqmatrix_def: "0 \ SqMatrix (\i j. 0)" instance .. end text \Tricky sort issues: compare @{term one_matrix} with @{term one_sqmatrix} \dots\ instantiation matrix :: ("{zero,one}",len,len) one begin definition one_matrix_def: "1 \ Matrix (\i j. if Rep_atMost i = Rep_atMost j then 1 else 0)" instance .. end instantiation sqmatrix :: ("{zero,one}",type) one begin definition one_sqmatrix_def: "1 \ SqMatrix (\i j. if i = j then 1 else 0)" instance .. end subsection \Matrix Addition\ fun matrix_plus where "matrix_plus (Matrix A) (Matrix B) = Matrix (\i j. A i j + B i j)" instantiation matrix :: (plus,type,type) plus begin definition plus_matrix_def: "A + B \ matrix_plus A B" instance .. end lemma plus_matrix_def' [simp]: "Matrix A + Matrix B = Matrix (\i j. A i j + B i j)" by (simp add: plus_matrix_def) instantiation sqmatrix :: (plus,type) plus begin definition plus_sqmatrix_def: "A + B \ sqmatrix_of_matrix (matrix_of_sqmatrix A + matrix_of_sqmatrix B)" instance .. end lemma plus_sqmatrix_def' [simp]: "SqMatrix A + SqMatrix B = SqMatrix (\i j. A i j + B i j)" by (simp add: plus_sqmatrix_def) lemma matrix_add_0_right [simp]: "A + 0 = (A::('a::monoid_add,'m,'n) matrix)" by (cases A, simp add: zero_matrix_def) lemma matrix_add_0_left [simp]: "0 + A = (A::('a::monoid_add,'m,'n) matrix)" by (cases A, simp add: zero_matrix_def) lemma matrix_add_commute [simp]: "(A::('a::ab_semigroup_add,'m,'n) matrix) + B = B + A" by (cases A, cases B, simp add: add.commute) lemma matrix_add_assoc: "(A::('a::semigroup_add,'m,'n) matrix) + B + C = A + (B + C)" by (cases A, cases B, cases C, simp add: add.assoc) lemma matrix_add_left_commute [simp]: "(A::('a::ab_semigroup_add,'m,'n) matrix) + (B + C) = B + (A + C)" by (metis matrix_add_assoc matrix_add_commute) lemma sqmatrix_add_0_right [simp]: "A + 0 = (A::('a::monoid_add,'m) sqmatrix)" by (cases A, simp add: zero_sqmatrix_def) lemma sqmatrix_add_0_left [simp]: "0 + A = (A::('a::monoid_add,'m) sqmatrix)" by (cases A, simp add: zero_sqmatrix_def) lemma sqmatrix_add_commute [simp]: "(A::('a::ab_semigroup_add,'m) sqmatrix) + B = B + A" by (cases A, cases B, simp add: add.commute) lemma sqmatrix_add_assoc: "(A::('a::semigroup_add,'m) sqmatrix) + B + C = A + (B + C)" by (cases A, cases B, cases C, simp add: add.assoc) lemma sqmatrix_add_left_commute [simp]: "(A::('a::ab_semigroup_add,'m) sqmatrix) + (B + C) = B + (A + C)" by (metis sqmatrix_add_commute sqmatrix_add_assoc) subsection \Order (via Addition)\ instantiation matrix :: (plus,type,type) plus_ord begin definition less_eq_matrix_def: "(A::('a, 'b, 'c) matrix) \ B \ A + B = B" definition less_matrix_def: "(A::('a, 'b, 'c) matrix) < B \ A \ B \ A \ B" instance proof fix A B :: "('a, 'b, 'c) matrix" show "A \ B \ A + B = B" by (metis less_eq_matrix_def) show "A < B \ A \ B \ A \ B" by (metis less_matrix_def) qed end instantiation sqmatrix :: (plus,type) plus_ord begin definition less_eq_sqmatrix_def: "(A::('a, 'b) sqmatrix) \ B \ A + B = B" definition less_sqmatrix_def: "(A::('a, 'b) sqmatrix) < B \ A \ B \ A \ B" instance proof fix A B :: "('a, 'b) sqmatrix" show "A \ B \ A + B = B" by (metis less_eq_sqmatrix_def) show "A < B \ A \ B \ A \ B" by (metis less_sqmatrix_def) qed end subsection \Matrix Multiplication\ fun matrix_times :: "('a::{comm_monoid_add,times},'m,'k) matrix \ ('a,'k,'n) matrix \ ('a,'m,'n) matrix" where "matrix_times (Matrix A) (Matrix B) = Matrix (\i j. sum (\k. A i k * B k j) (UNIV::'k atMost set))" notation matrix_times (infixl "*\<^sub>M" 70) instantiation sqmatrix :: ("{comm_monoid_add,times}",type) times begin definition times_sqmatrix_def: "A * B = sqmatrix_of_matrix (matrix_of_sqmatrix A *\<^sub>M matrix_of_sqmatrix B)" instance .. end lemma times_sqmatrix_def' [simp]: "SqMatrix A * SqMatrix B = SqMatrix (\i j. sum (\k. A i k * B k j) (UNIV::'k atMost set))" by (simp add: times_sqmatrix_def) lemma matrix_mult_0_right [simp]: "(A::('a::{comm_monoid_add,mult_zero},'m,'n) matrix) *\<^sub>M 0 = 0" by (cases A, simp add: zero_matrix_def) lemma matrix_mult_0_left [simp]: "0 *\<^sub>M (A::('a::{comm_monoid_add,mult_zero},'m,'n) matrix) = 0" by (cases A, simp add: zero_matrix_def) lemma sum_delta_r_0 [simp]: "\ finite S; j \ S \ \ (\k\S. f k * (if k = j then 1 else (0::'b::{semiring_0,monoid_mult}))) = 0" by (induct S rule: finite_induct, auto) lemma sum_delta_r_1 [simp]: "\ finite S; j \ S \ \ (\k\S. f k * (if k = j then 1 else (0::'b::{semiring_0,monoid_mult}))) = f j" by (induct S rule: finite_induct, auto) lemma matrix_mult_1_right [simp]: "(A::('a::{semiring_0,monoid_mult},'m::len,'n::len) matrix) *\<^sub>M 1 = A" by (cases A, simp add: one_matrix_def) lemma sum_delta_l_0 [simp]: "\ finite S; i \ S \ \ (\k\S. (if i = k then 1 else (0::'b::{semiring_0,monoid_mult})) * f k j) = 0" by (induct S rule: finite_induct, auto) lemma sum_delta_l_1 [simp]: "\ finite S; i \ S \ \ (\k\S. (if i = k then 1 else (0::'b::{semiring_0,monoid_mult})) * f k j) = f i j" by (induct S rule: finite_induct, auto) lemma matrix_mult_1_left [simp]: "1 *\<^sub>M (A::('a::{semiring_0,monoid_mult},'m::len,'n::len) matrix) = A" by (cases A, simp add: one_matrix_def) lemma matrix_mult_assoc: "(A::('a::semiring_0,'m,'n) matrix) *\<^sub>M B *\<^sub>M C = A *\<^sub>M (B *\<^sub>M C)" apply (cases A) apply (cases B) apply (cases C) apply (simp add: sum_distrib_right sum_distrib_left mult.assoc) apply (subst sum.swap) apply (rule refl) done lemma matrix_mult_distrib_left: "(A::('a::{comm_monoid_add,semiring},'m,'n::len) matrix) *\<^sub>M (B + C) = A *\<^sub>M B + A *\<^sub>M C" by (cases A, cases B, cases C, simp add: distrib_left sum.distrib) lemma matrix_mult_distrib_right: "((A::('a::{comm_monoid_add,semiring},'m,'n::len) matrix) + B) *\<^sub>M C = A *\<^sub>M C + B *\<^sub>M C" by (cases A, cases B, cases C, simp add: distrib_right sum.distrib) lemma sqmatrix_mult_0_right [simp]: "(A::('a::{comm_monoid_add,mult_zero},'m) sqmatrix) * 0 = 0" by (cases A, simp add: zero_sqmatrix_def) lemma sqmatrix_mult_0_left [simp]: "0 * (A::('a::{comm_monoid_add,mult_zero},'m) sqmatrix) = 0" by (cases A, simp add: zero_sqmatrix_def) lemma sqmatrix_mult_1_right [simp]: "(A::('a::{semiring_0,monoid_mult},'m::len) sqmatrix) * 1 = A" by (cases A, simp add: one_sqmatrix_def) lemma sqmatrix_mult_1_left [simp]: "1 * (A::('a::{semiring_0,monoid_mult},'m::len) sqmatrix) = A" by (cases A, simp add: one_sqmatrix_def) lemma sqmatrix_mult_assoc: "(A::('a::{semiring_0,monoid_mult},'m) sqmatrix) * B * C = A * (B * C)" apply (cases A) apply (cases B) apply (cases C) apply (simp add: sum_distrib_right sum_distrib_left mult.assoc) apply (subst sum.swap) apply (rule refl) done lemma sqmatrix_mult_distrib_left: "(A::('a::{comm_monoid_add,semiring},'m::len) sqmatrix) * (B + C) = A * B + A * C" by (cases A, cases B, cases C, simp add: distrib_left sum.distrib) lemma sqmatrix_mult_distrib_right: "((A::('a::{comm_monoid_add,semiring},'m::len) sqmatrix) + B) * C = A * C + B * C" by (cases A, cases B, cases C, simp add: distrib_right sum.distrib) subsection \Square-Matrix Model of Dioids\ text \The following subclass proofs are necessary to connect parts of our algebraic hierarchy to the hierarchy found in the Isabelle/HOL library.\ subclass (in ab_near_semiring_one_zerol) comm_monoid_add proof fix a :: 'a show "0 + a = a" by (fact add_zerol) qed subclass (in semiring_one_zero) semiring_0 proof fix a :: 'a show "0 * a = 0" by (fact annil) show "a * 0 = 0" by (fact annir) qed subclass (in ab_near_semiring_one) monoid_mult .. instantiation sqmatrix :: (dioid_one_zero,len) dioid_one_zero begin instance proof fix A B C :: "('a, 'b) sqmatrix" show "A + B + C = A + (B + C)" by (fact sqmatrix_add_assoc) show "A + B = B + A" by (fact sqmatrix_add_commute) show "A * B * C = A * (B * C)" by (fact sqmatrix_mult_assoc) show "(A + B) * C = A * C + B * C" by (fact sqmatrix_mult_distrib_right) show "1 * A = A" by (fact sqmatrix_mult_1_left) show "A * 1 = A" by (fact sqmatrix_mult_1_right) show "0 + A = A" by (fact sqmatrix_add_0_left) show "0 * A = 0" by (fact sqmatrix_mult_0_left) show "A * 0 = 0" by (fact sqmatrix_mult_0_right) show "A + A = A" by (cases A, simp) show "A * (B + C) = A * B + A * C" by (fact sqmatrix_mult_distrib_left) qed end subsection \Kleene Star for Matrices\ text \We currently do not implement the Kleene star of matrices, since this is complicated.\ end diff --git a/thys/Kleene_Algebra/ROOT b/thys/Kleene_Algebra/ROOT --- a/thys/Kleene_Algebra/ROOT +++ b/thys/Kleene_Algebra/ROOT @@ -1,23 +1,25 @@ chapter AFP -session "Kleene_Algebra" (AFP) = "HOL-Word" + +session "Kleene_Algebra" (AFP) = HOL + options [timeout = 600] + sessions + "HOL-Library" theories Signatures Dioid Dioid_Models Matrix Conway Kleene_Algebra Kleene_Algebra_Models Omega_Algebra Omega_Algebra_Models DRA PHL_KA PHL_DRA Finite_Suprema Formal_Power_Series Inf_Matrix document_files "root.bib" "root.tex" diff --git a/thys/Native_Word/Code_Symbolic_Bits_Int.thy b/thys/Native_Word/Code_Symbolic_Bits_Int.thy --- a/thys/Native_Word/Code_Symbolic_Bits_Int.thy +++ b/thys/Native_Word/Code_Symbolic_Bits_Int.thy @@ -1,120 +1,120 @@ (* Title: Code_Symbolic_Bits_Int.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Symbolic implementation of bit operations on int\ theory Code_Symbolic_Bits_Int imports - "HOL-Word.Misc_set_bit" "HOL-Word.Misc_lsb" More_Bits_Int + "Word_Lib.Misc_set_bit" "Word_Lib.Misc_lsb" More_Bits_Int begin section \Implementations of bit operations on \<^typ>\int\ operating on symbolic representation\ lemma test_bit_int_code [code]: "bit (0::int) n = False" "bit (Int.Neg num.One) n = True" "bit (Int.Pos num.One) 0 = True" "bit (Int.Pos (num.Bit0 m)) 0 = False" "bit (Int.Pos (num.Bit1 m)) 0 = True" "bit (Int.Neg (num.Bit0 m)) 0 = False" "bit (Int.Neg (num.Bit1 m)) 0 = True" "bit (Int.Pos num.One) (Suc n) = False" "bit (Int.Pos (num.Bit0 m)) (Suc n) = bit (Int.Pos m) n" "bit (Int.Pos (num.Bit1 m)) (Suc n) = bit (Int.Pos m) n" "bit (Int.Neg (num.Bit0 m)) (Suc n) = bit (Int.Neg m) n" "bit (Int.Neg (num.Bit1 m)) (Suc n) = bit (Int.Neg (Num.inc m)) n" by (simp_all add: Num.add_One bit_Suc) lemma int_not_code [code]: "NOT (0 :: int) = -1" "NOT (Int.Pos n) = Int.Neg (Num.inc n)" "NOT (Int.Neg n) = Num.sub n num.One" by(simp_all add: Num.add_One int_not_def) lemma int_and_code [code]: fixes i j :: int shows "0 AND j = 0" "i AND 0 = 0" "Int.Pos n AND Int.Pos m = (case bitAND_num n m of None \ 0 | Some n' \ Int.Pos n')" "Int.Neg n AND Int.Neg m = NOT (Num.sub n num.One OR Num.sub m num.One)" "Int.Pos n AND Int.Neg num.One = Int.Pos n" "Int.Pos n AND Int.Neg (num.Bit0 m) = Num.sub (bitORN_num (Num.BitM m) n) num.One" "Int.Pos n AND Int.Neg (num.Bit1 m) = Num.sub (bitORN_num (num.Bit0 m) n) num.One" "Int.Neg num.One AND Int.Pos m = Int.Pos m" "Int.Neg (num.Bit0 n) AND Int.Pos m = Num.sub (bitORN_num (Num.BitM n) m) num.One" "Int.Neg (num.Bit1 n) AND Int.Pos m = Num.sub (bitORN_num (num.Bit0 n) m) num.One" apply (simp_all add: int_numeral_bitAND_num Num.add_One sub_inc_One_eq inc_BitM_eq not_minus_numeral_inc_eq flip: int_not_neg_numeral int_or_not_bitORN_num split: option.split) apply (simp_all add: ac_simps) done lemma int_or_code [code]: fixes i j :: int shows "0 OR j = j" "i OR 0 = i" "Int.Pos n OR Int.Pos m = Int.Pos (bitOR_num n m)" "Int.Neg n OR Int.Neg m = NOT (Num.sub n num.One AND Num.sub m num.One)" "Int.Pos n OR Int.Neg num.One = Int.Neg num.One" "Int.Pos n OR Int.Neg (num.Bit0 m) = (case bitANDN_num (Num.BitM m) n of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" "Int.Pos n OR Int.Neg (num.Bit1 m) = (case bitANDN_num (num.Bit0 m) n of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" "Int.Neg num.One OR Int.Pos m = Int.Neg num.One" "Int.Neg (num.Bit0 n) OR Int.Pos m = (case bitANDN_num (Num.BitM n) m of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" "Int.Neg (num.Bit1 n) OR Int.Pos m = (case bitANDN_num (num.Bit0 n) m of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" apply (simp_all add: int_numeral_bitOR_num flip: int_not_neg_numeral) apply (simp_all add: or_int_def int_and_comm int_not_and_bitANDN_num del: int_not_simps(4) split: option.split) apply (simp_all add: Num.add_One) done lemma int_xor_code [code]: fixes i j :: int shows "0 XOR j = j" "i XOR 0 = i" "Int.Pos n XOR Int.Pos m = (case bitXOR_num n m of None \ 0 | Some n' \ Int.Pos n')" "Int.Neg n XOR Int.Neg m = Num.sub n num.One XOR Num.sub m num.One" "Int.Neg n XOR Int.Pos m = NOT (Num.sub n num.One XOR Int.Pos m)" "Int.Pos n XOR Int.Neg m = NOT (Int.Pos n XOR Num.sub m num.One)" by(fold int_not_neg_numeral)(simp_all add: int_numeral_bitXOR_num int_xor_not cong: option.case_cong) lemma bin_rest_code: "bin_rest i = i >> 1" by (simp add: shiftr_int_def) lemma set_bits_code [code]: "set_bits = Code.abort (STR ''set_bits is unsupported on type int'') (\_. set_bits :: _ \ int)" by simp lemma fixes i :: int shows int_set_bit_True_conv_OR [code]: "set_bit i n True = i OR (1 << n)" and int_set_bit_False_conv_NAND [code]: "set_bit i n False = i AND NOT (1 << n)" and int_set_bit_conv_ops: "set_bit i n b = (if b then i OR (1 << n) else i AND NOT (1 << n))" by(simp_all add: set_bit_int_def bin_set_conv_OR bin_clr_conv_NAND) declare [[code drop: \drop_bit :: nat \ int \ int\]] lemma drop_bit_int_code [code]: fixes i :: int shows "drop_bit 0 i = i" "drop_bit (Suc n) 0 = (0 :: int)" "drop_bit (Suc n) (Int.Pos num.One) = 0" "drop_bit (Suc n) (Int.Pos (num.Bit0 m)) = drop_bit n (Int.Pos m)" "drop_bit (Suc n) (Int.Pos (num.Bit1 m)) = drop_bit n (Int.Pos m)" "drop_bit (Suc n) (Int.Neg num.One) = - 1" "drop_bit (Suc n) (Int.Neg (num.Bit0 m)) = drop_bit n (Int.Neg m)" "drop_bit (Suc n) (Int.Neg (num.Bit1 m)) = drop_bit n (Int.Neg (Num.inc m))" by (simp_all add: shiftr_eq_drop_bit drop_bit_Suc add_One) declare [[code drop: \push_bit :: nat \ int \ int\]] lemma push_bit_int_code [code]: "push_bit 0 i = i" "push_bit (Suc n) i = push_bit n (Int.dup i)" by (simp_all add: ac_simps) lemma int_lsb_code [code]: "lsb (0 :: int) = False" "lsb (Int.Pos num.One) = True" "lsb (Int.Pos (num.Bit0 w)) = False" "lsb (Int.Pos (num.Bit1 w)) = True" "lsb (Int.Neg num.One) = True" "lsb (Int.Neg (num.Bit0 w)) = False" "lsb (Int.Neg (num.Bit1 w)) = True" by simp_all end diff --git a/thys/Native_Word/Code_Target_Word_Base.thy b/thys/Native_Word/Code_Target_Word_Base.thy --- a/thys/Native_Word/Code_Target_Word_Base.thy +++ b/thys/Native_Word/Code_Target_Word_Base.thy @@ -1,422 +1,422 @@ (* Title: Code_Target_Word_Base.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Common base for target language implementations of word types\ theory Code_Target_Word_Base imports - "HOL-Word.Word" + "HOL-Library.Word" "Word_Lib.Word_Lib" Bits_Integer begin text \More lemmas\ lemma nat_div_eq_Suc_0_iff: "n div m = Suc 0 \ m \ n \ n < 2 * m" apply auto using div_greater_zero_iff apply fastforce apply (metis One_nat_def div_greater_zero_iff dividend_less_div_times mult.right_neutral mult_Suc mult_numeral_1 numeral_2_eq_2 zero_less_numeral) apply (simp add: div_nat_eqI) done lemma Suc_0_lt_2p_len_of: "Suc 0 < 2 ^ LENGTH('a :: len)" by (metis One_nat_def len_gt_0 lessI numeral_2_eq_2 one_less_power) lemma div_half_nat: fixes x y :: nat assumes "y \ 0" shows "(x div y, x mod y) = (let q = 2 * (x div 2 div y); r = x - q * y in if y \ r then (q + 1, r - y) else (q, r))" proof - let ?q = "2 * (x div 2 div y)" have q: "?q = x div y - x div y mod 2" by(metis div_mult2_eq mult.commute minus_mod_eq_mult_div [symmetric]) let ?r = "x - ?q * y" have r: "?r = x mod y + x div y mod 2 * y" by(simp add: q diff_mult_distrib minus_mod_eq_div_mult [symmetric])(metis diff_diff_cancel mod_less_eq_dividend mod_mult2_eq add.commute mult.commute) show ?thesis proof(cases "y \ x - ?q * y") case True with assms q have "x div y mod 2 \ 0" unfolding r by (metis Nat.add_0_right diff_0_eq_0 diff_Suc_1 le_div_geq mod2_gr_0 mod_div_trivial mult_0 neq0_conv numeral_1_eq_Suc_0 numerals(1)) hence "x div y = ?q + 1" unfolding q by simp moreover hence "x mod y = ?r - y" by simp(metis minus_div_mult_eq_mod [symmetric] diff_commute diff_diff_left mult_Suc) ultimately show ?thesis using True by(simp add: Let_def) next case False hence "x div y mod 2 = 0" unfolding r by(simp add: not_le)(metis Nat.add_0_right assms div_less div_mult_self2 mod_div_trivial mult.commute) hence "x div y = ?q" unfolding q by simp moreover hence "x mod y = ?r" by (metis minus_div_mult_eq_mod [symmetric]) ultimately show ?thesis using False by(simp add: Let_def) qed qed lemma unat_p2: "n < LENGTH('a :: len) \ unat (2 ^ n :: 'a word) = 2 ^ n" proof(induct n) case 0 thus ?case by simp next case (Suc n) then obtain n' where "LENGTH('a) = Suc n'" by(cases "LENGTH('a)") simp_all with Suc show ?case by (simp add: unat_word_ariths bintrunc_mod2p) qed lemma word_div_lt_eq_0: "x < y \ x div y = 0" for x :: "'a :: len word" by (simp add: word_eq_iff word_less_def word_test_bit_def uint_div) lemma word_div_eq_1_iff: "n div m = 1 \ n \ m \ unat n < 2 * unat (m :: 'a :: len word)" apply(simp only: word_arith_nat_defs word_le_nat_alt nat_div_eq_Suc_0_iff[symmetric]) apply(rule word_unat.Abs_inject) apply(simp only: unat_div[symmetric] word_unat.Rep) apply(simp add: unats_def Suc_0_lt_2p_len_of) done lemma div_half_word: fixes x y :: "'a :: len word" assumes "y \ 0" shows "(x div y, x mod y) = (let q = (x >> 1) div y << 1; r = x - q * y in if y \ r then (q + 1, r - y) else (q, r))" proof - obtain n where n: "x = of_nat n" "n < 2 ^ LENGTH('a)" by (rule that [of \unat x\]) simp_all moreover obtain m where m: "y = of_nat m" "m < 2 ^ LENGTH('a)" by (rule that [of \unat y\]) simp_all ultimately have [simp]: \unat (of_nat n :: 'a word) = n\ \unat (of_nat m :: 'a word) = m\ by (transfer, simp add: take_bit_of_nat take_bit_nat_eq_self_iff)+ let ?q = "(x >> 1) div y << 1" let ?q' = "2 * (n div 2 div m)" have "n div 2 div m < 2 ^ LENGTH('a)" using n by (metis of_nat_inverse unat_lt2p uno_simps(2)) hence q: "?q = of_nat ?q'" using n m by (auto simp add: shiftr_word_eq drop_bit_eq_div shiftl_t2n word_arith_nat_div uno_simps take_bit_nat_eq_self) from assms have "m \ 0" using m by -(rule notI, simp) from n have "2 * (n div 2 div m) < 2 ^ LENGTH('a)" by(metis mult.commute div_mult2_eq minus_mod_eq_mult_div [symmetric] less_imp_diff_less of_nat_inverse unat_lt2p uno_simps(2)) moreover have "2 * (n div 2 div m) * m < 2 ^ LENGTH('a)" using n unfolding div_mult2_eq[symmetric] by(subst (2) mult.commute)(simp add: minus_mod_eq_div_mult [symmetric] diff_mult_distrib minus_mod_eq_mult_div [symmetric] div_mult2_eq) moreover have "2 * (n div 2 div m) * m \ n" by (metis div_mult2_eq dtle mult.assoc mult.left_commute) ultimately have r: "x - ?q * y = of_nat (n - ?q' * m)" and "y \ x - ?q * y \ of_nat (n - ?q' * m) - y = of_nat (n - ?q' * m - m)" using n m unfolding q apply (simp_all add: of_nat_diff) apply (subst of_nat_diff) apply (simp_all add: word_le_nat_alt take_bit_nat_eq_self unat_sub_if' unat_word_ariths) apply (cases \2 \ LENGTH('a)\) apply (simp_all add: unat_word_ariths take_bit_nat_eq_self) done then show ?thesis using n m div_half_nat [OF \m \ 0\, of n] unfolding q by (simp add: word_le_nat_alt word_div_def word_mod_def Let_def take_bit_nat_eq_self flip: zdiv_int zmod_int split del: if_split split: if_split_asm) qed lemma word_test_bit_set_bits: "(BITS n. f n :: 'a :: len word) !! n \ n < LENGTH('a) \ f n" by (simp add: test_bit_eq_bit bit_set_bits_word_iff) lemma word_of_int_conv_set_bits: "word_of_int i = (BITS n. i !! n)" by (rule word_eqI) (auto simp add: word_test_bit_set_bits test_bit.eq_norm) lemma word_and_mask_or_conv_and_mask: "n !! index \ (n AND mask index) OR (1 << index) = n AND mask (index + 1)" for n :: \'a::len word\ by(rule word_eqI)(auto simp add: word_ao_nth word_size nth_shiftl simp del: shiftl_1) lemma uint_and_mask_or_full: fixes n :: "'a :: len word" assumes "n !! (LENGTH('a) - 1)" and "mask1 = mask (LENGTH('a) - 1)" and "mask2 = 1 << LENGTH('a) - 1" shows "uint (n AND mask1) OR mask2 = uint n" proof - have "mask2 = uint (1 << LENGTH('a) - 1 :: 'a word)" using assms by (simp add: uint_shiftl word_size bintrunc_shiftl del: shiftl_1) hence "uint (n AND mask1) OR mask2 = uint (n AND mask1 OR (1 << LENGTH('a) - 1 :: 'a word))" by(simp add: uint_or) also have "\ = uint (n AND mask (LENGTH('a) - 1 + 1))" using assms by(simp only: word_and_mask_or_conv_and_mask) also have "\ = uint n" by simp finally show ?thesis . qed text \Division on @{typ "'a word"} is unsigned, but Scala and OCaml only have signed division and modulus.\ lemmas word_sdiv_def = sdiv_word_def lemmas word_smod_def = smod_word_def lemma [code]: "x sdiv y = (let x' = sint x; y' = sint y; negative = (x' < 0) \ (y' < 0); result = abs x' div abs y' in word_of_int (if negative then -result else result))" for x y :: \'a::len word\ by (simp add: sdiv_word_def signed_divide_int_def sgn_if Let_def not_less not_le) lemma [code]: "x smod y = (let x' = sint x; y' = sint y; negative = (x' < 0); result = abs x' mod abs y' in word_of_int (if negative then -result else result))" for x y :: \'a::len word\ proof - have *: \k mod l = k - k div l * l\ for k l :: int by (simp add: minus_div_mult_eq_mod) show ?thesis by (simp add: smod_word_def signed_modulo_int_def signed_divide_int_def * sgn_if) (simp add: signed_eq_0_iff) qed text \ This algorithm implements unsigned division in terms of signed division. Taken from Hacker's Delight. \ lemma divmod_via_sdivmod: fixes x y :: "'a :: len word" assumes "y \ 0" shows "(x div y, x mod y) = (if 1 << (LENGTH('a) - 1) \ y then if x < y then (0, x) else (1, x - y) else let q = ((x >> 1) sdiv y) << 1; r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" proof(cases "1 << (LENGTH('a) - 1) \ y") case True note y = this show ?thesis proof(cases "x < y") case True then have "x mod y = x" by (cases x, cases y) (simp add: word_less_def word_mod_def) thus ?thesis using True y by(simp add: word_div_lt_eq_0) next case False obtain n where n: "y = of_nat n" "n < 2 ^ LENGTH('a)" by (rule that [of \unat y\]) simp_all have "unat x < 2 ^ LENGTH('a)" by(rule unat_lt2p) also have "\ = 2 * 2 ^ (LENGTH('a) - 1)" by(metis Suc_pred len_gt_0 power_Suc One_nat_def) also have "\ \ 2 * n" using y n by transfer (simp add: push_bit_of_1 take_bit_eq_mod) finally have div: "x div of_nat n = 1" using False n by (simp add: word_div_eq_1_iff take_bit_nat_eq_self) moreover have "x mod y = x - x div y * y" by (simp add: minus_div_mult_eq_mod) with div n have "x mod y = x - y" by simp ultimately show ?thesis using False y n by simp qed next case False note y = this obtain n where n: "x = of_nat n" "n < 2 ^ LENGTH('a)" by (rule that [of \unat x\]) simp_all hence "int n div 2 + 2 ^ (LENGTH('a) - Suc 0) < 2 ^ LENGTH('a)" by (cases "LENGTH('a)") (simp_all, simp only: of_nat_numeral [where ?'a = int, symmetric] zdiv_int [symmetric] of_nat_power [symmetric]) with y n have "sint (x >> 1) = uint (x >> 1)" by (simp add: sint_uint sbintrunc_mod2p shiftr_div_2n take_bit_nat_eq_self) moreover have "uint y + 2 ^ (LENGTH('a) - Suc 0) < 2 ^ LENGTH('a)" using y by (cases "LENGTH('a)") (simp_all add: not_le word_2p_lem word_size) then have "sint y = uint y" by (simp add: sint_uint sbintrunc_mod2p) ultimately show ?thesis using y apply (subst div_half_word [OF assms]) apply (simp add: sdiv_word_def signed_divide_int_def flip: uint_div) done qed text \More implementations tailored towards target-language implementations\ context includes integer.lifting begin lift_definition word_of_integer :: "integer \ 'a :: len word" is word_of_int . lemma word_of_integer_code [code]: "word_of_integer n = word_of_int (int_of_integer n)" by(simp add: word_of_integer.rep_eq) end lemma word_of_int_code: "uint (word_of_int x :: 'a word) = x AND mask (LENGTH('a :: len))" by (simp add: take_bit_eq_mask) context fixes f :: "nat \ bool" begin definition set_bits_aux :: \'a word \ nat \ 'a :: len word\ where \set_bits_aux w n = push_bit n w OR take_bit n (set_bits f)\ lemma set_bits_aux_conv: \set_bits_aux w n = (w << n) OR (set_bits f AND mask n)\ for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: set_bits_aux_def shiftl_word_eq bit_and_iff bit_or_iff bit_push_bit_iff bit_take_bit_iff bit_mask_iff bit_set_bits_word_iff) corollary set_bits_conv_set_bits_aux: \set_bits f = (set_bits_aux 0 (LENGTH('a)) :: 'a :: len word)\ by (simp add: set_bits_aux_conv) lemma set_bits_aux_0 [simp]: \set_bits_aux w 0 = w\ by (simp add: set_bits_aux_conv) lemma set_bits_aux_Suc [simp]: \set_bits_aux w (Suc n) = set_bits_aux ((w << 1) OR (if f n then 1 else 0)) n\ by (simp add: set_bits_aux_def shiftl_word_eq bit_eq_iff bit_or_iff bit_push_bit_iff bit_take_bit_iff bit_set_bits_word_iff) (auto simp add: bit_exp_iff not_less bit_1_iff less_Suc_eq_le) lemma set_bits_aux_simps [code]: \set_bits_aux w 0 = w\ \set_bits_aux w (Suc n) = set_bits_aux ((w << 1) OR (if f n then 1 else 0)) n\ by simp_all end lemma word_of_int_via_signed: fixes mask assumes mask_def: "mask = Bit_Operations.mask (LENGTH('a))" and shift_def: "shift = 1 << LENGTH('a)" and index_def: "index = LENGTH('a) - 1" and overflow_def:"overflow = 1 << (LENGTH('a) - 1)" and least_def: "least = - overflow" shows "(word_of_int i :: 'a :: len word) = (let i' = i AND mask in if i' !! index then if i' - shift < least \ overflow \ i' - shift then arbitrary1 i' else word_of_int (i' - shift) else if i' < least \ overflow \ i' then arbitrary2 i' else word_of_int i')" proof - define i' where "i' = i AND mask" have "shift = mask + 1" unfolding assms by(simp add: bin_mask_p1_conv_shift) hence "i' < shift" by(simp add: mask_def i'_def int_and_le) show ?thesis proof(cases "i' !! index") case True then have unf: "i' = overflow OR i'" apply (simp add: assms i'_def shiftl_eq_push_bit push_bit_of_1 flip: take_bit_eq_mask) apply (rule bit_eqI) apply (auto simp add: bit_take_bit_iff bit_or_iff bit_exp_iff) done have "overflow \ i'" by(subst unf)(rule le_int_or, simp add: bin_sign_and assms i'_def) hence "i' - shift < least \ False" unfolding assms by(cases "LENGTH('a)")(simp_all add: not_less) moreover have "overflow \ i' - shift \ False" using \i' < shift\ unfolding assms by(cases "LENGTH('a)")(auto simp add: not_le elim: less_le_trans) moreover have "word_of_int (i' - shift) = (word_of_int i :: 'a word)" using \i' < shift\ by (simp add: i'_def shift_def mask_def shiftl_eq_push_bit push_bit_of_1 flip: take_bit_eq_mask) ultimately show ?thesis using True by(simp add: Let_def i'_def) next case False hence "i' = i AND Bit_Operations.mask (LENGTH('a) - 1)" unfolding assms i'_def by(clarsimp simp add: i'_def bin_nth_ops intro!: bin_eqI)(cases "LENGTH('a)", auto simp add: less_Suc_eq) also have "\ \ Bit_Operations.mask (LENGTH('a) - 1)" by(rule int_and_le) simp also have "\ < overflow" unfolding overflow_def by(simp add: bin_mask_p1_conv_shift[symmetric]) also have "least \ 0" unfolding least_def overflow_def by simp have "0 \ i'" by (simp add: i'_def mask_def) hence "least \ i'" using \least \ 0\ by simp moreover have "word_of_int i' = (word_of_int i :: 'a word)" by(rule word_eqI)(auto simp add: i'_def bin_nth_ops mask_def) ultimately show ?thesis using False by(simp add: Let_def i'_def) qed qed text \Quickcheck conversion functions\ notation scomp (infixl "\\" 60) definition qc_random_cnv :: "(natural \ 'a::term_of) \ natural \ Random.seed \ ('a \ (unit \ Code_Evaluation.term)) \ Random.seed" where "qc_random_cnv a_of_natural i = Random.range (i + 1) \\ (\k. Pair ( let n = a_of_natural k in (n, \_. Code_Evaluation.term_of n)))" no_notation scomp (infixl "\\" 60) definition qc_exhaustive_cnv :: "(natural \ 'a) \ ('a \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "qc_exhaustive_cnv a_of_natural f d = Quickcheck_Exhaustive.exhaustive (%x. f (a_of_natural x)) d" definition qc_full_exhaustive_cnv :: "(natural \ ('a::term_of)) \ ('a \ (unit \ term) \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "qc_full_exhaustive_cnv a_of_natural f d = Quickcheck_Exhaustive.full_exhaustive (%(x, xt). f (a_of_natural x, %_. Code_Evaluation.term_of (a_of_natural x))) d" declare [[quickcheck_narrowing_ghc_options = "-XTypeSynonymInstances"]] definition qc_narrowing_drawn_from :: "'a list \ integer \ _" where "qc_narrowing_drawn_from xs = foldr Quickcheck_Narrowing.sum (map Quickcheck_Narrowing.cons (butlast xs)) (Quickcheck_Narrowing.cons (last xs))" locale quickcheck_narrowing_samples = fixes a_of_integer :: "integer \ 'a \ 'a :: {partial_term_of, term_of}" and zero :: "'a" and tr :: "typerep" begin function narrowing_samples :: "integer \ 'a list" where "narrowing_samples i = (if i > 0 then let (a, a') = a_of_integer i in narrowing_samples (i - 1) @ [a, a'] else [zero])" by pat_completeness auto termination including integer.lifting proof(relation "measure nat_of_integer") fix i :: integer assume "0 < i" thus "(i - 1, i) \ measure nat_of_integer" by simp(transfer, simp) qed simp definition partial_term_of_sample :: "integer \ 'a" where "partial_term_of_sample i = (if i < 0 then undefined else if i = 0 then zero else if i mod 2 = 0 then snd (a_of_integer (i div 2)) else fst (a_of_integer (i div 2 + 1)))" lemma partial_term_of_code: "partial_term_of (ty :: 'a itself) (Quickcheck_Narrowing.Narrowing_variable p t) \ Code_Evaluation.Free (STR ''_'') tr" "partial_term_of (ty :: 'a itself) (Quickcheck_Narrowing.Narrowing_constructor i []) \ Code_Evaluation.term_of (partial_term_of_sample i)" by (rule partial_term_of_anything)+ end lemmas [code] = quickcheck_narrowing_samples.narrowing_samples.simps quickcheck_narrowing_samples.partial_term_of_sample_def text \ The separate code target \SML_word\ collects setups for the code generator that PolyML does not provide. \ setup \Code_Target.add_derived_target ("SML_word", [(Code_ML.target_SML, I)])\ code_identifier code_module Code_Target_Word_Base \ (SML) Word and (Haskell) Word and (OCaml) Word and (Scala) Word end diff --git a/thys/Native_Word/More_Bits_Int.thy b/thys/Native_Word/More_Bits_Int.thy --- a/thys/Native_Word/More_Bits_Int.thy +++ b/thys/Native_Word/More_Bits_Int.thy @@ -1,154 +1,154 @@ (* Title: Bits_Int.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \More bit operations on integers\ theory More_Bits_Int imports - "HOL-Word.Bits_Int" - "HOL-Word.Bit_Comprehension" + "Word_Lib.Bits_Int" + "Word_Lib.Bit_Comprehension" begin text \Preliminaries\ lemma last_rev' [simp]: "last (rev xs) = hd xs" \ \TODO define \last []\ as \hd []\?\ by (cases xs) (simp add: last_def hd_def, simp) lemma nat_LEAST_True: "(LEAST _ :: nat. True) = 0" by (rule Least_equality) simp_all text \ Use this function to convert numeral @{typ integer}s quickly into @{typ int}s. By default, it works only for symbolic evaluation; normally generated code raises an exception at run-time. If theory \Code_Target_Bits_Int\ is imported, it works again, because then @{typ int} is implemented in terms of @{typ integer} even for symbolic evaluation. \ definition int_of_integer_symbolic :: "integer \ int" where "int_of_integer_symbolic = int_of_integer" lemma int_of_integer_symbolic_aux_code [code nbe]: "int_of_integer_symbolic 0 = 0" "int_of_integer_symbolic (Code_Numeral.Pos n) = Int.Pos n" "int_of_integer_symbolic (Code_Numeral.Neg n) = Int.Neg n" by (simp_all add: int_of_integer_symbolic_def) section \Symbolic bit operations on numerals and @{typ int}s\ fun bitOR_num :: "num \ num \ num" where "bitOR_num num.One num.One = num.One" | "bitOR_num num.One (num.Bit0 n) = num.Bit1 n" | "bitOR_num num.One (num.Bit1 n) = num.Bit1 n" | "bitOR_num (num.Bit0 m) num.One = num.Bit1 m" | "bitOR_num (num.Bit0 m) (num.Bit0 n) = num.Bit0 (bitOR_num m n)" | "bitOR_num (num.Bit0 m) (num.Bit1 n) = num.Bit1 (bitOR_num m n)" | "bitOR_num (num.Bit1 m) num.One = num.Bit1 m" | "bitOR_num (num.Bit1 m) (num.Bit0 n) = num.Bit1 (bitOR_num m n)" | "bitOR_num (num.Bit1 m) (num.Bit1 n) = num.Bit1 (bitOR_num m n)" fun bitAND_num :: "num \ num \ num option" where "bitAND_num num.One num.One = Some num.One" | "bitAND_num num.One (num.Bit0 n) = None" | "bitAND_num num.One (num.Bit1 n) = Some num.One" | "bitAND_num (num.Bit0 m) num.One = None" | "bitAND_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (bitAND_num m n)" | "bitAND_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (bitAND_num m n)" | "bitAND_num (num.Bit1 m) num.One = Some num.One" | "bitAND_num (num.Bit1 m) (num.Bit0 n) = map_option num.Bit0 (bitAND_num m n)" | "bitAND_num (num.Bit1 m) (num.Bit1 n) = (case bitAND_num m n of None \ Some num.One | Some n' \ Some (num.Bit1 n'))" fun bitXOR_num :: "num \ num \ num option" where "bitXOR_num num.One num.One = None" | "bitXOR_num num.One (num.Bit0 n) = Some (num.Bit1 n)" | "bitXOR_num num.One (num.Bit1 n) = Some (num.Bit0 n)" | "bitXOR_num (num.Bit0 m) num.One = Some (num.Bit1 m)" | "bitXOR_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (bitXOR_num m n)" | "bitXOR_num (num.Bit0 m) (num.Bit1 n) = Some (case bitXOR_num m n of None \ num.One | Some n' \ num.Bit1 n')" | "bitXOR_num (num.Bit1 m) num.One = Some (num.Bit0 m)" | "bitXOR_num (num.Bit1 m) (num.Bit0 n) = Some (case bitXOR_num m n of None \ num.One | Some n' \ num.Bit1 n')" | "bitXOR_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (bitXOR_num m n)" fun bitORN_num :: "num \ num \ num" where "bitORN_num num.One num.One = num.One" | "bitORN_num num.One (num.Bit0 m) = num.Bit1 m" | "bitORN_num num.One (num.Bit1 m) = num.Bit1 m" | "bitORN_num (num.Bit0 n) num.One = num.Bit0 num.One" | "bitORN_num (num.Bit0 n) (num.Bit0 m) = Num.BitM (bitORN_num n m)" | "bitORN_num (num.Bit0 n) (num.Bit1 m) = num.Bit0 (bitORN_num n m)" | "bitORN_num (num.Bit1 n) num.One = num.One" | "bitORN_num (num.Bit1 n) (num.Bit0 m) = Num.BitM (bitORN_num n m)" | "bitORN_num (num.Bit1 n) (num.Bit1 m) = Num.BitM (bitORN_num n m)" fun bitANDN_num :: "num \ num \ num option" where "bitANDN_num num.One num.One = None" | "bitANDN_num num.One (num.Bit0 n) = Some num.One" | "bitANDN_num num.One (num.Bit1 n) = None" | "bitANDN_num (num.Bit0 m) num.One = Some (num.Bit0 m)" | "bitANDN_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (bitANDN_num m n)" | "bitANDN_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (bitANDN_num m n)" | "bitANDN_num (num.Bit1 m) num.One = Some (num.Bit0 m)" | "bitANDN_num (num.Bit1 m) (num.Bit0 n) = (case bitANDN_num m n of None \ Some num.One | Some n' \ Some (num.Bit1 n'))" | "bitANDN_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (bitANDN_num m n)" lemma int_numeral_bitOR_num: "numeral n OR numeral m = (numeral (bitOR_num n m) :: int)" by(induct n m rule: bitOR_num.induct) simp_all lemma int_numeral_bitAND_num: "numeral n AND numeral m = (case bitAND_num n m of None \ 0 :: int | Some n' \ numeral n')" by(induct n m rule: bitAND_num.induct)(simp_all split: option.split) lemma int_numeral_bitXOR_num: "numeral m XOR numeral n = (case bitXOR_num m n of None \ 0 :: int | Some n' \ numeral n')" by(induct m n rule: bitXOR_num.induct)(simp_all split: option.split) lemma int_or_not_bitORN_num: "numeral n OR NOT (numeral m) = (- numeral (bitORN_num n m) :: int)" by (induction n m rule: bitORN_num.induct) (simp_all add: add_One BitM_inc_eq) lemma int_and_not_bitANDN_num: "numeral n AND NOT (numeral m) = (case bitANDN_num n m of None \ 0 :: int | Some n' \ numeral n')" by (induction n m rule: bitANDN_num.induct) (simp_all add: add_One BitM_inc_eq split: option.split) lemma int_not_and_bitANDN_num: "NOT (numeral m) AND numeral n = (case bitANDN_num n m of None \ 0 :: int | Some n' \ numeral n')" by(simp add: int_and_not_bitANDN_num[symmetric] int_and_comm) section \Bit masks of type \<^typ>\int\\ lemma bin_mask_conv_pow2: "mask n = 2 ^ n - (1 :: int)" by (fact mask_eq_exp_minus_1) lemma bin_mask_ge0: "mask n \ (0 :: int)" by (fact mask_nonnegative_int) lemma and_bin_mask_conv_mod: "x AND mask n = x mod 2 ^ n" for x :: int by (simp flip: take_bit_eq_mod add: take_bit_eq_mask) lemma bin_mask_numeral: "mask (numeral n) = (1 :: int) + 2 * mask (pred_numeral n)" by (fact mask_numeral) lemma bin_nth_mask [simp]: "bit (mask n :: int) i \ i < n" by (simp add: bit_mask_iff) lemma bin_sign_mask [simp]: "bin_sign (mask n) = 0" by (simp add: bin_sign_def bin_mask_conv_pow2) lemma bin_mask_p1_conv_shift: "mask n + 1 = (1 :: int) << n" by (simp add: bin_mask_conv_pow2 shiftl_int_def) code_identifier code_module More_Bits_Int \ (SML) Bit_Operations and (OCaml) Bit_Operations and (Haskell) Bit_Operations and (Scala) Bit_Operations end diff --git a/thys/Native_Word/ROOT b/thys/Native_Word/ROOT --- a/thys/Native_Word/ROOT +++ b/thys/Native_Word/ROOT @@ -1,38 +1,37 @@ chapter AFP -session "Native_Word" (AFP) = "HOL-Word" + +session "Native_Word" (AFP) = Word_Lib + options [timeout = 2400] sessions - "Word_Lib" "HOL-Imperative_HOL" theories Code_Target_Bits_Int Uint64 Uint32 Uint16 Uint8 Uint Native_Cast Native_Cast_Uint Native_Word_Imperative_HOL Native_Word_Test_Emu Native_Word_Test_PolyML Native_Word_Test_PolyML2 Native_Word_Test_PolyML64 Native_Word_Test_Scala theories [condition = ISABELLE_GHC] Native_Word_Test_GHC theories [condition = ISABELLE_MLTON] Native_Word_Test_MLton Native_Word_Test_MLton2 theories [condition = ISABELLE_OCAMLFIND] Native_Word_Test_OCaml Native_Word_Test_OCaml2 theories [condition = ISABELLE_SMLNJ] Native_Word_Test_SMLNJ Native_Word_Test_SMLNJ2 theories Uint_Userguide document_files "root.tex" "root.bib" diff --git a/thys/Native_Word/Uint_Userguide.thy b/thys/Native_Word/Uint_Userguide.thy --- a/thys/Native_Word/Uint_Userguide.thy +++ b/thys/Native_Word/Uint_Userguide.thy @@ -1,343 +1,343 @@ (* Title: Uint_Userguide.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \User guide for native words\ (*<*) theory Uint_Userguide imports Uint32 Uint16 Code_Target_Bits_Int begin (*>*) text \ This tutorial explains how to best use the types for native words like @{typ "uint32"} in your formalisation. You can base your formalisation \begin{enumerate} \item either directly on these types, \item or on the generic @{typ "'a word"} and only introduce native words a posteriori via code generator refinement. \end{enumerate} The first option causes the least overhead if you have to prove only little about the words you use and start a fresh formalisation. Just use the native type @{typ uint32} instead of @{typ "32 word"} and similarly for \uint64\, \uint16\, and \uint8\. As native word types are meant only for code generation, the lemmas about @{typ "'a word"} have not been duplicated, but you can transfer theorems between native word types and @{typ "'a word"} using the transfer package. Note, however, that this option restricts your work a bit: your own functions cannot be ``polymorphic'' in the word length, but you have to define a separate function for every word length you need. The second option is recommended if you already have a formalisation based on @{typ "'a word"} or if your proofs involve words and their properties. It separates code generation from modelling and proving, i.e., you can work with words as usual. Consequently, you have to manually setup the code generator to use the native types wherever you want. The following describes how to achieve this with moderate effort. Note, however, that some target languages of the code generator (especially OCaml) do not support all the native word types provided. Therefore, you should only import those types that you need -- the theory file for each type mentions at the top the restrictions for code generation. For example, PolyML does not provide the Word16 structure, and OCaml provides neither Word8 nor Word16. You can still use these theories provided that you also import the theory @{theory Native_Word.Code_Target_Bits_Int} (which implements @{typ int} by target-language integers), but these words will - be implemented via Isabelle's \HOL-Word\ library, i.e., + be implemented via Isabelle's \Word\ library, i.e., you do not gain anything in terms of efficiency. \textbf{There is a separate code target \SML_word\ for SML.} If you use one of the native words that PolyML does not support (such as \uint16\ and \uint64\ in 32-bit mode), but would like to map its operations to the Standard Basis Library functions, make sure to use the target \SML_word\ instead of \SML\; if you only use native word sizes that PolyML supports, you can stick with \SML\. This ensures that code generation within Isabelle as used by \Quickcheck\, \value\ and @\{code\} in ML blocks continues to work. \ section \Lifting functions from @{typ "'a word"} to native words\ text \ This section shows how to convert functions from @{typ "'a word"} to native words. For example, the following function \sum_squares\ computes the sum of the first @{term n} square numbers in 16 bit arithmetic using a tail-recursive function \gen_sum_squares\ with accumulator; for convenience, \sum_squares_int\ takes an integer instead of a word. \ function gen_sum_squares :: "16 word \ 16 word \ 16 word" where (*<*)[simp del]:(*>*) "gen_sum_squares accum n = (if n = 0 then accum else gen_sum_squares (accum + n * n) (n - 1))" (*<*)by pat_completeness simp termination by (relation \measure (nat \ uint \ snd)\) (simp_all add: measure_unat)(*>*) definition sum_squares :: "16 word \ 16 word" where "sum_squares = gen_sum_squares 0" definition sum_squares_int :: "int \ 16 word" where "sum_squares_int n = sum_squares (word_of_int n)" text \ The generated code for @{term sum_squares} and @{term sum_squares_int} emulates words with unbounded integers and explicit modulus as specified - in the theory @{theory "HOL-Word.Word"}. But for efficiency, we want that the + in the theory @{theory "HOL-Library.Word"}. But for efficiency, we want that the generated code uses machine words and machine arithmetic. Unfortunately, as @{typ "'a word"} is polymorphic in the word length, the code generator can only do this if we use another type for machine words. The theory @{theory Native_Word.Uint16} defines the type @{typ uint16} for machine words of 16~bits. We just have to follow two steps to use it: First, we lift all our functions from @{typ "16 word"} to @{typ uint16}, i.e., @{term sum_squares}, @{term gen_sum_squares}, and @{term sum_squares_int} in our case. The theory @{theory Native_Word.Uint16} sets up the lifting package for this and has already taken care of the arithmetic and bit-wise operations. \ lift_definition gen_sum_squares_uint :: "uint16 \ uint16 \ uint16" is gen_sum_squares . lift_definition sum_squares_uint :: "uint16 \ uint16" is sum_squares . lift_definition sum_squares_int_uint :: "int \ uint16" is sum_squares_int . text \ Second, we also have to transfer the code equations for our functions. The attribute \Transfer.transferred\ takes care of that, but it is better to check that the transfer succeeded: inspect the theorem to check that the new constants are used throughout. \ lemmas [Transfer.transferred, code] = gen_sum_squares.simps sum_squares_def sum_squares_int_def text \ Finally, we export the code to standard ML. We use the target \SML_word\ instead of \SML\ to have the operations on @{typ uint16} mapped to the Standard Basis Library. As PolyML does not provide a Word16 type, the mapping for @{typ uint16} is only active in the refined target \SML_word\. \ export_code sum_squares_int_uint in SML_word text \ Nevertheless, we can still evaluate terms with @{term "uint16"} within Isabelle, i.e., PolyML, but this will be translated to @{typ "16 word"} and therefore less efficient. \ value "sum_squares_int_uint 40" section \Storing native words in datatypes\ text \ The above lifting is necessary for all functions whose type mentions the word type. Fortunately, we do not have to duplicate functions that merely operate on datatypes that contain words. Nevertheless, we have to tell the code generator that these functions should call the new ones, which operate on machine words. This section shows how to achieve this with data refinement. \ subsection \Example: expressions and two semantics\ text \ As the running example, we consider a language of expressions (literal values, less-than comparisions and conditional) where values are either booleans or 32-bit words. The original specification uses the type @{typ "32 word"}. \ datatype val = Bool bool | Word "32 word" datatype expr = Lit val | LT expr expr | IF expr expr expr abbreviation (input) word :: "32 word \ expr" where "word i \ Lit (Word i)" abbreviation (input) bool :: "bool \ expr" where "bool i \ Lit (Bool i)" \ \Denotational semantics of expressions, @{term None} denotes a type error\ fun eval :: "expr \ val option" where "eval (Lit v) = Some v" | "eval (LT e\<^sub>1 e\<^sub>2) = (case (eval e\<^sub>1, eval e\<^sub>2) of (Some (Word i\<^sub>1), Some (Word i\<^sub>2)) \ Some (Bool (i\<^sub>1 < i\<^sub>2)) | _ \ None)" | "eval (IF e\<^sub>1 e\<^sub>2 e\<^sub>3) = (case eval e\<^sub>1 of Some (Bool b) \ if b then eval e\<^sub>2 else eval e\<^sub>3 | _ \ None)" \ \Small-step semantics of expressions, it gets stuck upon type errors.\ inductive step :: "expr \ expr \ bool" ("_ \ _" [50, 50] 60) where "e \ e' \ LT e e\<^sub>2 \ LT e' e\<^sub>2" | "e \ e' \ LT (word i) e \ LT (word i) e'" | "LT (word i\<^sub>1) (word i\<^sub>2) \ bool (i\<^sub>1 < i\<^sub>2)" | "e \ e' \ IF e e\<^sub>1 e\<^sub>2 \ IF e' e\<^sub>1 e\<^sub>2" | "IF (bool True) e\<^sub>1 e\<^sub>2 \ e\<^sub>1" | "IF (bool False) e\<^sub>1 e\<^sub>2 \ e\<^sub>2" \ \Compile the inductive definition with the predicate compiler\ code_pred (modes: i \ o \ bool as reduce, i \ i \ bool as step') step . subsection \Change the datatype to use machine words\ text \ Now, we want to use @{typ uint32} instead of @{typ "32 word"}. The goal is to make the code generator use the new type without duplicating any of the types (@{typ val}, @{typ expr}) or the functions (@{term eval}, @{term reduce}) on such types. The constructor @{term Word} has @{typ "32 word"} in its type, so we have to lift it to \Word'\, and the same holds for the case combinator @{term case_val}, which @{term case_val'} replaces.% \footnote{% Note that we should not declare a case translation for the new case combinator because this will break parsing case expressions with old case combinator. } Next, we set up the code generator accordingly: @{term Bool} and @{term Word'} are the new constructors for @{typ val}, and @{term case_val'} is the new case combinator with an appropriate case certificate.% \footnote{% Case certificates tell the code generator to replace the HOL case combinator for a datatype with the case combinator of the target language. Without a case certificate, the code generator generates a function that re-implements the case combinator; in a strict languages like ML or Scala, this means that the code evaluates all possible cases before it decides which one is taken. Case certificates are described in Haftmann's PhD thesis \cite[Def.\ 27]{Haftmann2009PhD}. For a datatype \dt\ with constructors \C\<^sub>1\ to \C\<^sub>n\ where each constructor \C\<^sub>i\ takes \k\<^sub>i\ parameters, the certificate for the case combinator \case_dt\ looks as follows: { \isamarkuptrue\isacommand{lemma}\isamarkupfalse\isanewline% \ \ \isakeyword{assumes}\ {\isachardoublequoteopen}CASE\ {\isasymequiv}\ dt{\isacharunderscore}case\ c\isactrlsub {\isadigit{1}}\ c\isactrlsub {\isadigit{2}}\ \ldots\ c\isactrlsub{n}{\isachardoublequoteclose}\isanewline \ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isacharparenleft}CASE\ {\isacharparenleft}C\isactrlsub {\isadigit{1}}\ a\isactrlsub {\isadigit{1}}\isactrlsub {\isadigit{1}}\ a\isactrlsub {\isadigit{1}}\isactrlsub {\isadigit{2}}\ \ldots\ a\isactrlsub {\isadigit{1}}\isactrlsub {k\ensuremath{{}_1}}{\isacharparenright}\ {\isasymequiv}\ c\isactrlsub {\isadigit{1}}\ a\isactrlsub {\isadigit{1}}\isactrlsub {\isadigit{1}}\ a\isactrlsub {\isadigit{1}}\isactrlsub {\isadigit{2}}\ \ldots\ a\isactrlsub {\isadigit{1}}\isactrlsub {k\ensuremath{{}_1}}{\isacharparenright}\isanewline \ \ \ \ {\isacharampersand}{\isacharampersand}{\isacharampersand}\ {\isacharparenleft}CASE\ {\isacharparenleft}C\isactrlsub {\isadigit{2}}\ a\isactrlsub {\isadigit{2}}\isactrlsub {\isadigit{1}}\ a\isactrlsub {\isadigit{2}}\isactrlsub {\isadigit{2}}\ \ldots\ a\isactrlsub {\isadigit{2}}\isactrlsub {k\ensuremath{{}_2}}{\isacharparenright}\ {\isasymequiv}\ c\isactrlsub {\isadigit{2}}\ a\isactrlsub {\isadigit{2}}\isactrlsub {\isadigit{1}}\ a\isactrlsub {\isadigit{2}}\isactrlsub {\isadigit{2}}\ \ldots\ a\isactrlsub {\isadigit{2}}\isactrlsub {k\ensuremath{{}_2}}{\isacharparenright}\isanewline \ \ \ \ {\isacharampersand}{\isacharampersand}{\isacharampersand}\ \ldots\isanewline \ \ \ \ {\isacharampersand}{\isacharampersand}{\isacharampersand}\ {\isacharparenleft}CASE\ {\isacharparenleft}C\isactrlsub {n}\ a\isactrlsub {n}\isactrlsub {\isadigit{1}}\ a\isactrlsub {n}\isactrlsub {\isadigit{2}}\ \ldots\ a\isactrlsub {n}\isactrlsub {k\ensuremath{{}_n}}{\isacharparenright}\ {\isasymequiv}\ c\isactrlsub {n}\ a\isactrlsub {n}\isactrlsub {\isadigit{1}}\ a\isactrlsub {n}\isactrlsub {\isadigit{2}}\ \ldots\ a\isactrlsub {n}\isactrlsub {k\ensuremath{{}_n}}{\isacharparenright}{\isachardoublequoteclose}\isanewline } } We delete the code equations for the old constructor @{term Word} and case combinator @{term case_val} such that the code generator reports missing adaptations. \ lift_definition Word' :: "uint32 \ val" is Word . code_datatype Bool Word' lift_definition case_val' :: "(bool \ 'a) \ (uint32 \ 'a) \ val \ 'a" is case_val . lemmas [code, simp] = val.case [Transfer.transferred] lemma case_val'_cert: fixes bool word' b w assumes "CASE \ case_val' bool word'" shows "(CASE (Bool b) \ bool b) &&& (CASE (Word' w) \ word' w)" by (simp_all add: assms) setup \Code.declare_case_global @{thm case_val'_cert}\ declare [[code drop: case_val Word]] subsection \Make functions use functions on machine words\ text \ Finally, we merely have to change the code equations to use the new functions that operate on @{typ uint32}. As before, the attribute \Transfer.transferred\ does the job. In our example, we adapt the equality test on @{typ val} (code equations @{thm [source] val.eq.simps}) and the denotational and small-step semantics (code equations @{thm [source] eval.simps} and @{thm [source] step.equation}, respectively). We check that the adaptation has suceeded by exporting the functions. As we only use native word sizes that PolyML supports, we can use the usual target \SML\ instead of \SML_word\. \ lemmas [code] = val.eq.simps[THEN meta_eq_to_obj_eq, Transfer.transferred, THEN eq_reflection] eval.simps[Transfer.transferred] step.equation[Transfer.transferred] export_code reduce step' eval checking SML section \Troubleshooting\ text \ This section explains some possible problems when using native words. If you experience other difficulties, please contact the author. \ subsection \\export_code\ raises an exception \label{section:export_code:exception}\ text \ Probably, you have defined and are using a function on a native word type, but the code equation refers to emulated words. For example, the following defines a function \double\ that doubles a word. When we try to export code for \double\ without any further setup, \export_code\ will raise an exception or generate code that does not compile. \ lift_definition double :: "uint32 \ uint32" is "\x. x + x" . text \ We have to prove a code equation that only uses the existing operations on @{typ uint32}. Then, \export_code\ works again. \ lemma double_code [code]: "double n = n + n" by transfer simp subsection \The generated code does not compile\ text \ Probably, you have been exporting to a target language for which there is no setup, or your compiler does not provide the required API. Every theory for native words mentions at the start the limitations on code generation. Check that your concrete application meets all the requirements. Alternatively, this might be an instance of the problem described in \S\ref{section:export_code:exception}. For Haskell, you have to enable the extension TypeSynonymInstances with \texttt{-XTypeSynonymInstances} if you are using polymorphic bit operations on the native word types. \ subsection \The generated code is too slow\ text \ The generated code will most likely not be as fast as a direct implementation in the target language with manual tuning. This is because we want the configuration of the code generation to be sound (as it can be used to prove theorems in Isabelle). Therefore, the bit operations sometimes perform range checks before they call the target language API. Here are some examples: \begin{itemize} \item Shift distances and bit indices in target languages are often expected to fit into a bounded integer or word. However, the size of these types varies across target languages and platforms. Hence, no Isabelle/HOL type can model uniformly all of them. Instead, the bit operations use arbitrary-precision integers for such quantities and check at run-time that the values fit into a bounded integer or word, respectively -- if not, they raise an exception. \item Division and modulo operations explicitly test whether the divisor is $0$ and return the HOL value of division by $0$ in that case. This is necessary because some languages leave the behaviour of division by 0 unspecified. \end{itemize} If you have better ideas how to eliminate such checks and speed up the generated code without sacrificing soundness, please contact the author! \ (*<*)end(*>*) diff --git a/thys/ROBDD/ROOT b/thys/ROBDD/ROOT --- a/thys/ROBDD/ROOT +++ b/thys/ROBDD/ROOT @@ -1,29 +1,28 @@ chapter AFP session ROBDD (AFP) = Sepref_Prereq + options [timeout = 600] sessions "HOL-Imperative_HOL" "HOL-Library" - "HOL-Word" "HOL-ex" Automatic_Refinement Collections Native_Word theories Middle_Impl Bool_Func Conc_Impl Abstract_Impl BDT Pointer_Map_Impl Level_Collapse Array_List BDD_Examples Option_Helpers Pointer_Map theories [condition = ISABELLE_GHC] BDD_Code document_files "root.tex" "root.bib" diff --git a/thys/RSAPSS/ROOT b/thys/RSAPSS/ROOT --- a/thys/RSAPSS/ROOT +++ b/thys/RSAPSS/ROOT @@ -1,11 +1,11 @@ chapter AFP session RSAPSS (AFP) = "HOL-Number_Theory" + options [timeout = 600] theories [document = false] - Word (* local, not HOL-Word! *) + Word (* local, not HOL-Library! *) theories RSAPSS document_files "root.bib" "root.tex" diff --git a/thys/Refine_Monadic/ROOT b/thys/Refine_Monadic/ROOT --- a/thys/Refine_Monadic/ROOT +++ b/thys/Refine_Monadic/ROOT @@ -1,15 +1,13 @@ chapter AFP session Refine_Monadic (AFP) = Automatic_Refinement + options [timeout = 1200] - sessions - "HOL-Word" directories "Generic" "examples" theories Refine_Monadic "examples/Examples" document_files "root.bib" "root.tex" diff --git a/thys/Refine_Monadic/examples/WordRefine.thy b/thys/Refine_Monadic/examples/WordRefine.thy --- a/thys/Refine_Monadic/examples/WordRefine.thy +++ b/thys/Refine_Monadic/examples/WordRefine.thy @@ -1,79 +1,79 @@ section \Machine Words\ theory WordRefine -imports "../Refine_Monadic" "HOL-Word.Word" +imports "../Refine_Monadic" "HOL-Library.Word" begin text \This theory provides a simple example to show refinement of natural numbers to machine words. The setup is not yet very elaborated, but shows the direction to go. \ subsection \Setup\ definition [simp]: "word_nat_rel \ build_rel (unat) (\_. True)" lemma word_nat_RELEATES[refine_dref_RELATES]: "RELATES word_nat_rel" by (simp add: RELATES_def) lemma [simp, relator_props]: "single_valued word_nat_rel" unfolding word_nat_rel_def by blast lemma [simp]: "single_valuedp (\c a. a = unat c)" by (rule single_valuedpI) blast lemma [simp, relator_props]: "single_valued (converse word_nat_rel)" by (auto intro: injI) lemmas [refine_hsimp] = word_less_nat_alt word_le_nat_alt unat_sub iffD1[OF unat_add_lem] subsection \Example\ type_synonym word32 = "32 word" definition test :: "nat \ nat \ nat set nres" where "test x0 y0 \ do { let S={}; (S,_,_) \ WHILE (\(S,x,y). x>0) (\(S,x,y). do { let S=S\{y}; let x=x - 1; ASSERT (y0 \ test x0 y0 \ SPEC (\S. S={y0 .. y0 + x0 - 1})" \ \Choosen pre-condition to get least trouble when proving\ unfolding test_def apply (intro WHILE_rule[where I="\(S,x,y). x+y=x0+y0 \ x\x0 \ S={y0 .. y0 + (x0-x) - 1}"] refine_vcg) by auto definition test_impl :: "word32 \ word32 \ word32 set nres" where "test_impl x y \ do { let S={}; (S,_,_) \ WHILE (\(S,x,y). x>0) (\(S,x,y). do { let S=S\{y}; let x=x - 1; let y=y + 1; RETURN (S,x,y) }) (S,x,y); RETURN S }" lemma test_impl_refine: assumes "x'+y'<2^LENGTH(32)" assumes "(x,x')\word_nat_rel" assumes "(y,y')\word_nat_rel" shows "test_impl x y \ \(\word_nat_rel\set_rel) (test x' y')" proof - from assms show ?thesis unfolding test_impl_def test_def apply (refine_rcg) apply (refine_dref_type) apply (auto simp: refine_hsimp refine_rel_defs) done qed end diff --git a/thys/SPARCv8/ROOT b/thys/SPARCv8/ROOT --- a/thys/SPARCv8/ROOT +++ b/thys/SPARCv8/ROOT @@ -1,18 +1,19 @@ chapter AFP -session SPARCv8 (AFP) = "HOL-Word" + +session SPARCv8 (AFP) = HOL + options [timeout = 1200] sessions "HOL-Eisbach" + "Word_Lib" directories "lib" "lib/wp" "SparcModel_MMU" theories [document = false] "lib/WordDecl" theories "SparcModel_MMU/Sparc_Properties" theories [document = false] "SparcModel_MMU/Sparc_Code_Gen" document_files "root.tex" diff --git a/thys/SPARCv8/SparcModel_MMU/RegistersOps.thy b/thys/SPARCv8/SparcModel_MMU/RegistersOps.thy --- a/thys/SPARCv8/SparcModel_MMU/RegistersOps.thy +++ b/thys/SPARCv8/SparcModel_MMU/RegistersOps.thy @@ -1,79 +1,79 @@ section\Register Operations\ theory RegistersOps -imports Main "../lib/WordDecl" "HOL-Word.Traditional_Syntax" +imports Main "../lib/WordDecl" "Word_Lib.Traditional_Syntax" begin text\ This theory provides operations to get, set and clear bits in registers \ section "Getting Fields" text\ Get a field of type @{typ "'b::len word"} starting at @{term "index"} from @{term "addr"} of type @{typ "'a::len word"} \ definition get_field_from_word_a_b:: "'a::len word \ nat \ 'b::len word" where "get_field_from_word_a_b addr index \ let off = (size addr - LENGTH('b)) in ucast ((addr << (off-index)) >> off)" text\ Obtain, from addr of type @{typ "'a::len word"}, another @{typ "'a::len word"} containing the field of length \len\ starting at \index\ in \addr\. \ definition get_field_from_word_a_a:: "'a::len word \ nat \ nat \ 'a::len word" where "get_field_from_word_a_a addr index len \ (addr << (size addr - (index+len)) >> (size addr - len))" section "Setting Fields" text\ Set the field of type @{typ "'b::len word"} at \index\ from \record\ of type @{typ "'a::len word"}. \ definition set_field :: "'a::len word \ 'b::len word \ nat \ 'a::len word" where "set_field record field index \ let mask:: ('a::len word) = (mask (size field)) << index in (record AND (NOT mask)) OR ((ucast field) << index)" section "Clearing Fields" text\ Zero the \n\ initial bits of \addr\. \ definition clear_n_bits:: "'a::len word \ nat \ 'a::len word" where "clear_n_bits addr n \ addr AND (NOT (mask n))" text\ Gets the natural value of a 32 bit mask \ definition get_nat_from_mask::"word32 \ nat \ nat \ (word32 \ nat)" where " get_nat_from_mask w m v \ if (w AND (mask m) =0) then (w>>m, v+m) else (w,m) " definition get_nat_from_mask32::"word32\ nat" where "get_nat_from_mask32 w \ if (w=0) then len_of TYPE (word_length32) else let (w,res) = get_nat_from_mask w 16 0 in let (w,res)= get_nat_from_mask w 8 res in let (w,res) = get_nat_from_mask w 4 res in let (w,res) = get_nat_from_mask w 2 res in let (w,res) = get_nat_from_mask w 1 res in res " end diff --git a/thys/SPARCv8/SparcModel_MMU/Sparc_Types.thy b/thys/SPARCv8/SparcModel_MMU/Sparc_Types.thy --- a/thys/SPARCv8/SparcModel_MMU/Sparc_Types.thy +++ b/thys/SPARCv8/SparcModel_MMU/Sparc_Types.thy @@ -1,791 +1,791 @@ (* * Copyright 2016, NTU * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * Author: Zhe Hou, David Sanan. *) section \SPARC V8 architecture CPU model\ theory Sparc_Types -imports Main "../lib/WordDecl" "HOL-Word.Traditional_Syntax" +imports Main "../lib/WordDecl" "Word_Lib.Traditional_Syntax" begin text \The following type definitions are taken from David Sanan's definitions for SPARC machines.\ type_synonym machine_word = word32 type_synonym byte = word8 type_synonym phys_address = word36 type_synonym virtua_address = word32 type_synonym page_address = word24 type_synonym offset = word12 type_synonym table_entry = word8 definition page_size :: "word32" where "page_size \ 4096" type_synonym virtua_page_address = word20 type_synonym context_type = word8 type_synonym word_length_t1 = word_length8 type_synonym word_length_t2 = word_length6 type_synonym word_length_t3 = word_length6 type_synonym word_length_offset = word_length12 type_synonym word_length_page = word_length24 type_synonym word_length_phys_address = word_length36 type_synonym word_length_virtua_address = word_length32 type_synonym word_length_entry_type = word_length2 type_synonym word_length_machine_word = word_length32 definition length_machine_word :: "nat" where "length_machine_word \ LENGTH(word_length_machine_word)" text_raw \\newpage\ section \CPU Register Definitions\ text\ The definitions below come from the SPARC Architecture Manual, Version 8. The LEON3 processor has been certified SPARC V8 conformant (2005). \ definition leon3khz ::"word32" where "leon3khz \ 33000" text \The following type definitions for MMU is taken from David Sanan's definitions for MMU.\ text\ The definitions below come from the UT699 LEON 3FT/SPARC V8 Microprocessor Functional Manual, Aeroflex, June 20, 2012, p35. \ datatype MMU_register = CR \ \Control Register\ | CTP \ \ConText Pointer register\ | CNR \ \Context Register\ | FTSR \ \Fault Status Register\ | FAR \ \Fault Address Register\ lemma MMU_register_induct: "P CR \ P CTP \ P CNR \ P FTSR \ P FAR \ P x" by (cases x) auto lemma UNIV_MMU_register [no_atp]: "UNIV = {CR, CTP, CNR, FTSR, FAR}" apply (safe) apply (case_tac x) apply (auto intro:MMU_register_induct) done instantiation MMU_register :: enum begin definition "enum_MMU_register = [ CR, CTP, CNR, FTSR, FAR ]" definition "enum_all_MMU_register P \ P CR \ P CTP \ P CNR \ P FTSR \ P FAR " definition "enum_ex_MMU_register P \ P CR \ P CTP \ P CNR \ P FTSR \ P FAR" instance proof qed (simp_all only: enum_MMU_register_def enum_all_MMU_register_def enum_ex_MMU_register_def UNIV_MMU_register, simp_all) end type_synonym MMU_context = "MMU_register \ machine_word" text \\PTE_flags\ is the last 8 bits of a PTE. See page 242 of SPARCv8 manual. \<^item> C - bit 7 \<^item> M - bit 6, \<^item> R - bit 5 \<^item> ACC - bit 4~2 \<^item> ET - bit 1~0.\ type_synonym PTE_flags = word8 text \ @{term CPU_register} datatype is an enumeration with the CPU registers defined in the SPARC V8 architecture. \ datatype CPU_register = PSR \ \Processor State Register\ | WIM \ \Window Invalid Mask\ | TBR \ \Trap Base Register\ | Y \ \Multiply/Divide Register\ | PC \ \Program Counter\ | nPC \ \next Program Counter\ | DTQ \ \Deferred-Trap Queue\ | FSR \ \Floating-Point State Register\ | FQ \ \Floating-Point Deferred-Trap Queue\ | CSR \ \Coprocessor State Register\ | CQ \ \Coprocessor Deferred-Trap Queue\ (*| CCR -- "Cache Control Register"*) | ASR "word5" \ \Ancillary State Register\ text \The following two functions are dummies since we will not use ASRs. Future formalisation may add more details to this.\ definition privileged_ASR :: "word5 \ bool" where "privileged_ASR r \ False " definition illegal_instruction_ASR :: "word5 \ bool" where "illegal_instruction_ASR r \ False " definition get_tt :: "word32 \ word8" where "get_tt tbr \ ucast (((AND) tbr 0b00000000000000000000111111110000) >> 4) " text \Write the tt field of the TBR register. Return the new value of TBR.\ definition write_tt :: "word8 \ word32 \ word32" where "write_tt new_tt_val tbr_val \ let tmp = (AND) tbr_val 0b111111111111111111111000000001111 in (OR) tmp (((ucast new_tt_val)::word32) << 4) " text \Get the nth bit of WIM. This equals ((AND) WIM $2^n$). N.B. the first bit of WIM is the 0th bit.\ definition get_WIM_bit :: "nat \ word32 \ word1" where "get_WIM_bit n wim \ let mask = ((ucast (0b1::word1))::word32) << n in ucast (((AND) mask wim) >> n) " definition get_CWP :: "word32 \ word5" where "get_CWP psr \ ucast ((AND) psr 0b00000000000000000000000000011111) " definition get_ET :: "word32 \ word1" where "get_ET psr \ ucast (((AND) psr 0b00000000000000000000000000100000) >> 5) " definition get_PIL :: "word32 \ word4" where "get_PIL psr \ ucast (((AND) psr 0b00000000000000000000111100000000) >> 8) " definition get_PS :: "word32 \ word1" where "get_PS psr \ ucast (((AND) psr 0b00000000000000000000000001000000) >> 6) " definition get_S :: "word32 \ word1" where "get_S psr \ \<^cancel>\ucast (((AND) psr 0b00000000000000000000000010000000) >> 7)\ if ((AND) psr (0b00000000000000000000000010000000::word32)) = 0 then 0 else 1 " definition get_icc_N :: "word32 \ word1" where "get_icc_N psr \ ucast (((AND) psr 0b00000000100000000000000000000000) >> 23) " definition get_icc_Z :: "word32 \ word1" where "get_icc_Z psr \ ucast (((AND) psr 0b00000000010000000000000000000000) >> 22) " definition get_icc_V :: "word32 \ word1" where "get_icc_V psr \ ucast (((AND) psr 0b00000000001000000000000000000000) >> 21) " definition get_icc_C :: "word32 \ word1" where "get_icc_C psr \ ucast (((AND) psr 0b00000000000100000000000000000000) >> 20) " definition update_S :: "word1 \ word32 \ word32" where "update_S s_val psr_val \ let tmp0 = (AND) psr_val 0b11111111111111111111111101111111 in (OR) tmp0 (((ucast s_val)::word32) << 7) " text \Update the CWP field of PSR. Return the new value of PSR.\ definition update_CWP :: "word5 \ word32 \ word32" where "update_CWP cwp_val psr_val \ let tmp0 = (AND) psr_val (0b11111111111111111111111111100000::word32); s_val = ((ucast (get_S psr_val))::word1) in if s_val = 0 then (AND) ((OR) tmp0 ((ucast cwp_val)::word32)) (0b11111111111111111111111101111111::word32) else (OR) ((OR) tmp0 ((ucast cwp_val)::word32)) (0b00000000000000000000000010000000::word32) " text \Update the the ET, CWP, and S fields of PSR. Return the new value of PSR.\ definition update_PSR_rett :: "word5 \ word1 \ word1 \ word32 \ word32" where "update_PSR_rett cwp_val et_val s_val psr_val \ let tmp0 = (AND) psr_val 0b11111111111111111111111101000000; tmp1 = (OR) tmp0 ((ucast cwp_val)::word32); tmp2 = (OR) tmp1 (((ucast et_val)::word32) << 5); tmp3 = (OR) tmp2 (((ucast s_val)::word32) << 7) in tmp3 " definition update_PSR_exe_trap :: "word5 \ word1 \ word1 \ word32 \ word32" where "update_PSR_exe_trap cwp_val et_val ps_val psr_val \ let tmp0 = (AND) psr_val 0b11111111111111111111111110000000; tmp1 = (OR) tmp0 ((ucast cwp_val)::word32); tmp2 = (OR) tmp1 (((ucast et_val)::word32) << 5); tmp3 = (OR) tmp2 (((ucast ps_val)::word32) << 6) in tmp3 " text \Update the N, Z, V, C fields of PSR. Return the new value of PSR.\ definition update_PSR_icc :: "word1 \ word1 \ word1 \ word1 \ word32 \ word32" where "update_PSR_icc n_val z_val v_val c_val psr_val \ let n_val_32 = if n_val = 0 then 0 else (0b00000000100000000000000000000000::word32); z_val_32 = if z_val = 0 then 0 else (0b00000000010000000000000000000000::word32); v_val_32 = if v_val = 0 then 0 else (0b00000000001000000000000000000000::word32); c_val_32 = if c_val = 0 then 0 else (0b00000000000100000000000000000000::word32); tmp0 = (AND) psr_val (0b11111111000011111111111111111111::word32); tmp1 = (OR) tmp0 n_val_32; tmp2 = (OR) tmp1 z_val_32; tmp3 = (OR) tmp2 v_val_32; tmp4 = (OR) tmp3 c_val_32 in tmp4 " text \Update the ET, PIL fields of PSR. Return the new value of PSR.\ definition update_PSR_et_pil :: "word1 \ word4 \ word32 \ word32" where "update_PSR_et_pil et pil psr_val \ let tmp0 = (AND) psr_val 0b111111111111111111111000011011111; tmp1 = (OR) tmp0 (((ucast et)::word32) << 5); tmp2 = (OR) tmp1 (((ucast pil)::word32) << 8) in tmp2 " text \ SPARC V8 architecture is organized in windows of 32 user registers. The data stored in a register is defined as a 32 bits word @{term reg_type}: \ type_synonym reg_type = "word32" text \ The access to the value of a CPU register of type @{term CPU_register} is defined by a total function @{term cpu_context} \ type_synonym cpu_context = "CPU_register \ reg_type" text \ User registers are defined with the type @{term user_reg} represented by a 5 bits word. \ type_synonym user_reg_type = "word5" definition PSR_S ::"reg_type" where "PSR_S \ 6" text \ Each window context is defined by a total function @{term window_context} from @{term user_register} to @{term reg_type} (32 bits word storing the actual value of the register). \ type_synonym window_context = "user_reg_type \ reg_type" text \ The number of windows is implementation dependent. The LEON architecture is composed of 16 different windows (a 4 bits word). \ definition NWINDOWS :: "int" where "NWINDOWS \ 8" text \Maximum number of windows is 32 in SPARCv8.\ type_synonym ('a) window_size = "'a word" text \ Finally the user context is defined by another total function @{term user_context} from @{term window_size} to @{term window_context}. That is, the user context is a function taking as argument a register set window and a register within that window, and it returns the value stored in that user register. \ type_synonym ('a) user_context = "('a) window_size \ window_context" datatype sys_reg = CCR \ \Cache control register\ |ICCR \ \Instruction cache configuration register\ |DCCR \ \Data cache configuration register\ type_synonym sys_context = "sys_reg \ reg_type" text\ The memory model is defined by a total function from 32 bits words to 8 bits words \ type_synonym asi_type = "word8" text \ The memory is defined as a function from page address to page, which is also defined as a function from physical address to @{term "machine_word"} \ type_synonym mem_val_type = "word8" type_synonym mem_context = "asi_type \ phys_address \ mem_val_type option" type_synonym cache_tag = "word20" type_synonym cache_line_size = "word12" type_synonym cache_type = "(cache_tag \ cache_line_size)" type_synonym cache_context = "cache_type \ mem_val_type option" text \The delayed-write pool generated from write state register instructions.\ type_synonym delayed_write_pool = "(int \ reg_type \ CPU_register) list" definition DELAYNUM :: "int" where "DELAYNUM \ 0" text \Convert a set to a list.\ definition list_of_set :: "'a set \ 'a list" where "list_of_set s = (SOME l. set l = s)" lemma set_list_of_set: "finite s \ set (list_of_set s) = s" unfolding list_of_set_def by (metis (mono_tags) finite_list some_eq_ex) type_synonym ANNUL = "bool" type_synonym RESET_TRAP = "bool" type_synonym EXECUTE_MODE = "bool" type_synonym RESET_MODE = "bool" type_synonym ERROR_MODE = "bool" type_synonym TICC_TRAP_TYPE = "word7" type_synonym INTERRUPT_LEVEL = "word3" type_synonym STORE_BARRIER_PENDING = "bool" text \The processor asserts this signal to ensure that the memory system will not process another SWAP or LDSTUB operation to the same memory byte.\ type_synonym pb_block_ldst_byte = "virtua_address \ bool" text\The processor asserts this signal to ensure that the memory system will not process another SWAP or LDSTUB operation to the same memory word.\ type_synonym pb_block_ldst_word = "virtua_address \ bool" record sparc_state_var = annul:: ANNUL resett:: RESET_TRAP exe:: EXECUTE_MODE reset:: RESET_MODE err:: ERROR_MODE ticc:: TICC_TRAP_TYPE itrpt_lvl:: INTERRUPT_LEVEL st_bar:: STORE_BARRIER_PENDING atm_ldst_byte:: pb_block_ldst_byte atm_ldst_word:: pb_block_ldst_word definition get_annul :: "sparc_state_var \ bool" where "get_annul v \ annul v" definition get_reset_trap :: "sparc_state_var \ bool" where "get_reset_trap v \ resett v" definition get_exe_mode :: "sparc_state_var \ bool" where "get_exe_mode v \ exe v" definition get_reset_mode :: "sparc_state_var \ bool" where "get_reset_mode v \ reset v" definition get_err_mode :: "sparc_state_var \ bool" where "get_err_mode v \ err v" definition get_ticc_trap_type :: "sparc_state_var \ word7" where "get_ticc_trap_type v \ ticc v" definition get_interrupt_level :: "sparc_state_var \ word3" where "get_interrupt_level v \ itrpt_lvl v" definition get_store_barrier_pending :: "sparc_state_var \ bool" where "get_store_barrier_pending v \ st_bar v" definition write_annul :: "bool \ sparc_state_var \ sparc_state_var" where "write_annul b v \ v\annul := b\" definition write_reset_trap :: "bool \ sparc_state_var \ sparc_state_var" where "write_reset_trap b v \ v\resett := b\" definition write_exe_mode :: "bool \ sparc_state_var \ sparc_state_var" where "write_exe_mode b v \ v\exe := b\" definition write_reset_mode :: "bool \ sparc_state_var \ sparc_state_var" where "write_reset_mode b v \ v\reset := b\" definition write_err_mode :: "bool \ sparc_state_var \ sparc_state_var" where "write_err_mode b v \ v\err := b\" definition write_ticc_trap_type :: "word7 \ sparc_state_var \ sparc_state_var" where "write_ticc_trap_type w v \ v\ticc := w\" definition write_interrupt_level :: "word3 \ sparc_state_var \ sparc_state_var" where "write_interrupt_level w v \ v\itrpt_lvl := w\" definition write_store_barrier_pending :: "bool \ sparc_state_var \ sparc_state_var" where "write_store_barrier_pending b v \ v\st_bar := b\" text \Given a word7 value, find the highest bit, and fill the left bits to be the highest bit.\ definition sign_ext7::"word7 \ word32" where "sign_ext7 w \ let highest_bit = ((AND) w 0b1000000) >> 6 in if highest_bit = 0 then (ucast w)::word32 else (OR) ((ucast w)::word32) 0b11111111111111111111111110000000 " definition zero_ext8 :: "word8 \ word32" where "zero_ext8 w \ (ucast w)::word32 " text \Given a word8 value, find the highest bit, and fill the left bits to be the highest bit.\ definition sign_ext8::"word8 \ word32" where "sign_ext8 w \ let highest_bit = ((AND) w 0b10000000) >> 7 in if highest_bit = 0 then (ucast w)::word32 else (OR) ((ucast w)::word32) 0b11111111111111111111111100000000 " text \Given a word13 value, find the highest bit, and fill the left bits to be the highest bit.\ definition sign_ext13::"word13 \ word32" where "sign_ext13 w \ let highest_bit = ((AND) w 0b1000000000000) >> 12 in if highest_bit = 0 then (ucast w)::word32 else (OR) ((ucast w)::word32) 0b11111111111111111110000000000000 " definition zero_ext16 :: "word16 \ word32" where "zero_ext16 w \ (ucast w)::word32 " text \Given a word16 value, find the highest bit, and fill the left bits to be the highest bit.\ definition sign_ext16::"word16 \ word32" where "sign_ext16 w \ let highest_bit = ((AND) w 0b1000000000000000) >> 15 in if highest_bit = 0 then (ucast w)::word32 else (OR) ((ucast w)::word32) 0b11111111111111110000000000000000 " text \Given a word22 value, find the highest bit, and fill the left bits to tbe the highest bit.\ definition sign_ext22::"word22 \ word32" where "sign_ext22 w \ let highest_bit = ((AND) w 0b1000000000000000000000) >> 21 in if highest_bit = 0 then (ucast w)::word32 else (OR) ((ucast w)::word32) 0b11111111110000000000000000000000 " text \Given a word24 value, find the highest bit, and fill the left bits to tbe the highest bit.\ definition sign_ext24::"word24 \ word32" where "sign_ext24 w \ let highest_bit = ((AND) w 0b100000000000000000000000) >> 23 in if highest_bit = 0 then (ucast w)::word32 else (OR) ((ucast w)::word32) 0b11111111000000000000000000000000 " text\ Operations to be defined. The SPARC V8 architecture is composed of the following set of instructions: \<^item> Load Integer Instructions \<^item> Load Floating-point Instructions \<^item> Load Coprocessor Instructions \<^item> Store Integer Instructions \<^item> Store Floating-point Instructions \<^item> Store Coprocessor Instructions \<^item> Atomic Load-Store Unsigned Byte Instructions \<^item> SWAP Register With Memory Instruction \<^item> SETHI Instructions \<^item> NOP Instruction \<^item> Logical Instructions \<^item> Shift Instructions \<^item> Add Instructions \<^item> Tagged Add Instructions \<^item> Subtract Instructions \<^item> Tagged Subtract Instructions \<^item> Multiply Step Instruction \<^item> Multiply Instructions \<^item> Divide Instructions \<^item> SAVE and RESTORE Instructions \<^item> Branch on Integer Condition Codes Instructions \<^item> Branch on Floating-point Condition Codes Instructions \<^item> Branch on Coprocessor Condition Codes Instructions \<^item> Call and Link Instruction \<^item> Jump and Link Instruction \<^item> Return from Trap Instruction \<^item> Trap on Integer Condition Codes Instructions \<^item> Read State Register Instructions \<^item> Write State Register Instructions \<^item> STBAR Instruction \<^item> Unimplemented Instruction \<^item> Flush Instruction Memory \<^item> Floating-point Operate (FPop) Instructions \<^item> Convert Integer to Floating point Instructions \<^item> Convert Floating point to Integer Instructions \<^item> Convert Between Floating-point Formats Instructions \<^item> Floating-point Move Instructions \<^item> Floating-point Square Root Instructions \<^item> Floating-point Add and Subtract Instructions \<^item> Floating-point Multiply and Divide Instructions \<^item> Floating-point Compare Instructions \<^item> Coprocessor Operate Instructions \ text \The CALL instruction.\ datatype call_type = CALL \ \Call and Link\ text \The SETHI instruction.\ datatype sethi_type = SETHI \ \Set High 22 bits of r Register\ text \The NOP instruction.\ datatype nop_type = NOP \ \No Operation\ text \The Branch on integer condition codes instructions.\ datatype bicc_type = BE \ \Branch on Equal\ | BNE \ \Branch on Not Equal\ | BGU \ \Branch on Greater Unsigned\ | BLE \ \Branch on Less or Equal\ | BL \ \Branch on Less\ | BGE \ \Branch on Greater or Equal\ | BNEG \ \Branch on Negative\ | BG \ \Branch on Greater\ | BCS \ \Branch on Carry Set (Less than, Unsigned)\ | BLEU \ \Branch on Less or Equal Unsigned\ | BCC \ \Branch on Carry Clear (Greater than or Equal, Unsigned)\ | BA \ \Branch Always\ | BN \ \Branch Never\ \ \Added for unconditional branches\ | BPOS \ \Branch on Positive\ | BVC \ \Branch on Overflow Clear\ | BVS \ \Branch on Overflow Set\ text \Memory instructions. That is, load and store.\ datatype load_store_type = LDSB \ \Load Signed Byte\ | LDUB \ \Load Unsigned Byte\ | LDUBA \ \Load Unsigned Byte from Alternate space\ | LDUH \ \Load Unsigned Halfword\ | LD \ \Load Word\ | LDA \ \Load Word from Alternate space\ | LDD \ \Load Doubleword\ | STB \ \Store Byte\ | STH \ \Store Halfword\ | ST \ \Store Word\ | STA \ \Store Word into Alternate space\ | STD \ \Store Doubleword\ | LDSBA \ \Load Signed Byte from Alternate space\ | LDSH \ \Load Signed Halfword\ | LDSHA \ \Load Signed Halfword from Alternate space\ | LDUHA \ \Load Unsigned Halfword from Alternate space\ | LDDA \ \Load Doubleword from Alternate space\ | STBA \ \Store Byte into Alternate space\ | STHA \ \Store Halfword into Alternate space\ | STDA \ \Store Doubleword into Alternate space\ | LDSTUB \ \Atomic Load Store Unsigned Byte\ | LDSTUBA \ \Atomic Load Store Unsinged Byte in Alternate space\ | SWAP \ \Swap r Register with Mmemory\ | SWAPA \ \Swap r Register with Mmemory in Alternate space\ | FLUSH \ \Flush Instruction Memory\ | STBAR \ \Store Barrier\ text \Arithmetic instructions.\ datatype arith_type = ADD \ \Add\ | ADDcc \ \Add and modify icc\ | ADDX \ \Add with Carry\ | SUB \ \Subtract\ | SUBcc \ \Subtract and modify icc\ | SUBX \ \Subtract with Carry\ | UMUL \ \Unsigned Integer Multiply\ | SMUL \ \Signed Integer Multiply\ | SMULcc \ \Signed Integer Multiply and modify icc\ | UDIV \ \Unsigned Integer Divide\ | UDIVcc \ \Unsigned Integer Divide and modify icc\ | SDIV \ \Signed Integer Divide\ | ADDXcc \ \Add with Carry and modify icc\ | TADDcc \ \Tagged Add and modify icc\ | TADDccTV \ \Tagged Add and modify icc and Trap on overflow\ | SUBXcc \ \Subtract with Carry and modify icc\ | TSUBcc \ \Tagged Subtract and modify icc\ | TSUBccTV \ \Tagged Subtract and modify icc and Trap on overflow\ | MULScc \ \Multiply Step and modify icc\ | UMULcc \ \Unsigned Integer Multiply and modify icc\ | SDIVcc \ \Signed Integer Divide and modify icc\ text \Logical instructions.\ datatype logic_type = ANDs \ \And\ | ANDcc \ \And and modify icc\ | ANDN \ \And Not\ | ANDNcc \ \And Not and modify icc\ | ORs \ \Inclusive-Or\ | ORcc \ \Inclusive-Or and modify icc\ | ORN \ \Inclusive Or Not\ | XORs \ \Exclusive-Or\ | XNOR \ \Exclusive-Nor\ | ORNcc \ \Inclusive-Or Not and modify icc\ | XORcc \ \Exclusive-Or and modify icc\ | XNORcc \ \Exclusive-Nor and modify icc\ text \Shift instructions.\ datatype shift_type = SLL \ \Shift Left Logical\ | SRL \ \Shift Right Logical\ | SRA \ \Shift Right Arithmetic\ text \Other Control-transfer instructions.\ datatype ctrl_type = JMPL \ \Jump and Link\ | RETT \ \Return from Trap\ | SAVE \ \Save caller's window\ | RESTORE \ \Restore caller's window\ text \Access state registers instructions.\ datatype sreg_type = RDASR \ \Read Ancillary State Register\ | RDY \ \Read Y Register\ | RDPSR \ \Read Processor State Register\ | RDWIM \ \Read Window Invalid Mask Register\ | RDTBR \ \Read Trap Base Regiser\ | WRASR \ \Write Ancillary State Register\ | WRY \ \Write Y Register\ | WRPSR \ \Write Processor State Register\ | WRWIM \ \Write Window Invalid Mask Register\ | WRTBR \ \Write Trap Base Register\ text \Unimplemented instruction.\ datatype uimp_type = UNIMP \ \Unimplemented\ text \Trap on integer condition code instructions.\ datatype ticc_type = TA \ \Trap Always\ | TN \ \Trap Never\ | TNE \ \Trap on Not Equal\ | TE \ \Trap on Equal\ | TG \ \Trap on Greater\ | TLE \ \Trap on Less or Equal\ | TGE \ \Trap on Greater or Equal\ | TL \ \Trap on Less\ | TGU \ \Trap on Greater Unsigned\ | TLEU \ \Trap on Less or Equal Unsigned\ | TCC \ \Trap on Carry Clear (Greater than or Equal, Unsigned)\ | TCS \ \Trap on Carry Set (Less Than, Unsigned)\ | TPOS \ \Trap on Postive\ | TNEG \ \Trap on Negative\ | TVC \ \Trap on Overflow Clear\ | TVS \ \Trap on Overflow Set\ datatype sparc_operation = call_type call_type | sethi_type sethi_type | nop_type nop_type | bicc_type bicc_type | load_store_type load_store_type | arith_type arith_type | logic_type logic_type | shift_type shift_type | ctrl_type ctrl_type | sreg_type sreg_type | uimp_type uimp_type | ticc_type ticc_type datatype Trap = reset |data_store_error |instruction_access_MMU_miss |instruction_access_error |r_register_access_error |instruction_access_exception |privileged_instruction |illegal_instruction |unimplemented_FLUSH |watchpoint_detected |fp_disabled |cp_disabled |window_overflow |window_underflow |mem_address_not_aligned |fp_exception |cp_exception |data_access_error |data_access_MMU_miss |data_access_exception |tag_overflow |division_by_zero |trap_instruction |interrupt_level_n datatype Exception = \ \The following are processor states that are not in the instruction model,\ \ \but we MAY want to deal with these from hardware perspective.\ \<^cancel>\|execute_mode\ \<^cancel>\|reset_mode\ \<^cancel>\|error_mode\ \ \The following are self-defined exceptions.\ invalid_cond_f2 |invalid_op2_f2 |illegal_instruction2 \ \when \i = 0\ for load/store not from alternate space\ |invalid_op3_f3_op11 |case_impossible |invalid_op3_f3_op10 |invalid_op_f3 |unsupported_instruction |fetch_instruction_error |invalid_trap_cond end diff --git a/thys/SPARCv8/lib/WordDecl.thy b/thys/SPARCv8/lib/WordDecl.thy --- a/thys/SPARCv8/lib/WordDecl.thy +++ b/thys/SPARCv8/lib/WordDecl.thy @@ -1,412 +1,412 @@ (* Title: State.thy Author: David Sanán, Trinity College Dublin, 2012 *) section "Word Declarations" theory WordDecl -imports Main "HOL-Word.Word" +imports Main "HOL-Library.Word" begin text \ This theory provides \len0\ and \len\ instantiations for the most common used word sizes in the model (1,2,4,5,6,8,12,16,18,20,24,32,36). The \len0\ class defines lengths that range from 0 upwards, whilst the \len\ class caters for non-zero lengths. Bit-operators like \<<\, \>>\, \and\, or \or\ require \a'::len\ word instances, while other word operations such @{term "ucast"} (gets an integer from a word) can be defined for \len0\ words. For each length \N\, we: \<^enum> declare a type \word_lengthN\; \<^enum> make it an instance of both length classes; \<^enum> and introduce a short type synonym \wordN\, and a suitable translation. In essence, this theory is just a lot of boilerplate. \newpage \ section "Words of length 1" typedecl word_length1 instantiation word_length1 :: len0 begin definition len1 [simp]: "len_of (x::word_length1 itself) \ 1" instance .. end instantiation word_length1 :: len begin instance by intro_classes simp end type_synonym word1 = "word_length1 word" translations (type) "word1" <= (type) "word_length1 word" subsection "Words of length 1's constants" definition ONE::"word1" where "ONE \ 1" definition ZERO::"word1" where "ZERO \ 0" section "Words of length 2" typedecl word_length2 instantiation word_length2 :: len0 begin definition len2 [simp]: "len_of (x::word_length2 itself) \ 2" instance .. end instantiation word_length2 :: len begin instance by intro_classes simp end type_synonym word2 = "word_length2 word" translations (type) "word2" <= (type) "word_length2 word" section "Words of length 3" typedecl word_length3 instantiation word_length3 :: len0 begin definition len3 [simp]: "len_of (x::word_length3 itself) \ 3" instance .. end instantiation word_length3 :: len begin instance by intro_classes simp end type_synonym word3 = "word_length3 word" translations (type) "word3" <= (type) "word_length3 word" section "Words of length 4" typedecl word_length4 instantiation word_length4 :: len0 begin definition len4 [simp]: "len_of (x::word_length4 itself) \ 4" instance .. end instantiation word_length4 :: len begin instance by intro_classes simp end type_synonym word4 = "word_length4 word" translations (type) "word4" <= (type) "word_length4 word" section "Words of length 5" typedecl word_length5 instantiation word_length5 :: len0 begin definition len5 [simp]: "len_of (x::word_length5 itself) \ 5" instance .. end instantiation word_length5 :: len begin instance by intro_classes simp end type_synonym word5 = "word_length5 word" translations (type) "word5" <= (type) "word_length5 word" section "Words of length 6" typedecl word_length6 instantiation word_length6 :: len0 begin definition len6 [simp]: "len_of (x::word_length6 itself) \ 6" instance .. end instantiation word_length6 :: len begin instance by intro_classes simp end type_synonym word6 = "word_length6 word" translations (type) "word6" <= (type) "word_length6 word" section "Words of length 6" typedecl word_length7 instantiation word_length7 :: len0 begin definition len7 [simp]: "len_of (x::word_length7 itself) \ 7" instance .. end instantiation word_length7 :: len begin instance by intro_classes simp end type_synonym word7 = "word_length7 word" translations (type) "word7" <= (type) "word_length7 word" section "Words of length 8" typedecl word_length8 instantiation word_length8 :: len0 begin definition len8 [simp]: "len_of (x::word_length8 itself) \ 8" instance .. end instantiation word_length8 :: len begin instance by intro_classes simp end type_synonym word8 = "word_length8 word" translations (type) "word8" <= (type) "word_length8 word" section "Words of length 9" typedecl word_length9 instantiation word_length9 :: len0 begin definition len9 [simp]: "len_of (x::word_length9 itself) \ 9" instance .. end instantiation word_length9 :: len begin instance by intro_classes simp end type_synonym word9 = "word_length9 word" translations (type) "word9" <= (type) "word_length9 word" section "Words of length 10" typedecl word_length10 instantiation word_length10 :: len0 begin definition len10 [simp]: "len_of (x::word_length10 itself) \ 10" instance .. end instantiation word_length10 :: len begin instance by intro_classes simp end type_synonym word10 = "word_length10 word" translations (type) "word10" <= (type) "word_length10 word" section "Words of length 12" typedecl word_length12 instantiation word_length12 :: len0 begin definition len12 [simp]: "len_of (x::word_length12 itself) \ 12" instance .. end instantiation word_length12 :: len begin instance by intro_classes simp end type_synonym word12 = "word_length12 word" translations (type) "word12" <= (type) "word_length12 word" section "Words of length 13" typedecl word_length13 instantiation word_length13 :: len0 begin definition len13 [simp]: "len_of (x::word_length13 itself) \ 13" instance .. end instantiation word_length13 :: len begin instance by intro_classes simp end type_synonym word13 = "word_length13 word" translations (type) "word13" <= (type) "word_length13 word" section "Words of length 16" typedecl word_length16 instantiation word_length16 :: len0 begin definition len16 [simp]: "len_of (x::word_length16 itself) \ 16" instance .. end instantiation word_length16 :: len begin instance by intro_classes simp end type_synonym word16 = "word_length16 word" translations (type) "word16" <= (type) "word_length16 word" section "Words of length 18" typedecl word_length18 instantiation word_length18 :: len0 begin definition len18 [simp]: "len_of (x::word_length18 itself) \ 18" instance .. end instantiation word_length18 :: len begin instance by intro_classes simp end type_synonym word18 = "word_length18 word" translations (type) "word18" <= (type) "word_length18 word" section "Words of length 20" typedecl word_length20 instantiation word_length20 :: len0 begin definition len20 [simp]: "len_of (x::word_length20 itself) \ 20" instance .. end instantiation word_length20 :: len begin instance by intro_classes simp end type_synonym word20 = "word_length20 word" translations (type) "word20" <= (type) "word_length20 word" section "Words of length 22" typedecl word_length22 instantiation word_length22 :: len0 begin definition len22 [simp]: "len_of (x::word_length22 itself) \ 22" instance .. end instantiation word_length22 :: len begin instance by intro_classes simp end type_synonym word22 = "word_length22 word" translations (type) "word22" <= (type) "word_length22 word" section "Words of length 24" typedecl word_length24 instantiation word_length24 :: len0 begin definition len24 [simp]: "len_of (x::word_length24 itself) \ 24" instance .. end instantiation word_length24 :: len begin instance by intro_classes simp end type_synonym word24 = "word_length24 word" translations (type) "word24" <= (type) "word_length24 word" section "Words of length 30" typedecl word_length30 instantiation word_length30 :: len0 begin definition len30 [simp]: "len_of (x::word_length30 itself) \ 30" instance .. end instantiation word_length30 :: len begin instance by intro_classes simp end type_synonym word30 = "word_length30 word" translations (type) "word30" <= (type) "word_length30 word" section "Words of length 30" typedecl word_length31 instantiation word_length31 :: len0 begin definition len31 [simp]: "len_of (x::word_length31 itself) \ 31" instance .. end instantiation word_length31 :: len begin instance by intro_classes simp end type_synonym word31 = "word_length31 word" translations (type) "word31" <= (type) "word_length31 word" section "Words of length 32" typedecl word_length32 instantiation word_length32 :: len0 begin definition len32 [simp]: "len_of (x::word_length32 itself) \ 32" instance .. end instantiation word_length32 :: len begin instance by intro_classes simp end type_synonym word32 = "word_length32 word" translations (type) "word32" <= (type) "word_length32 word" section "Words of length 33" typedecl word_length33 instantiation word_length33 :: len0 begin definition len33 [simp]: "len_of (x::word_length33 itself) \ 33" instance .. end instantiation word_length33 :: len begin instance by intro_classes simp end type_synonym word33 = "word_length33 word" translations (type) "word33" <= (type) "word_length33 word" section "Words of length 36" typedecl word_length36 instantiation word_length36 :: len0 begin definition len36 [simp]: "len_of (x::word_length36 itself) \ 36" instance .. end instantiation word_length36 :: len begin instance by intro_classes simp end type_synonym word36 = "word_length36 word" translations (type) "word36" <= (type) "word_length36 word" section "Words of length 64" typedecl word_length64 instantiation word_length64 :: len0 begin definition len64 [simp]: "len_of (x::word_length64 itself) \ 64" instance .. end instantiation word_length64 :: len begin instance by intro_classes simp end type_synonym word64 = "word_length64 word" translations (type) "word64" <= (type) "word_length64 word" end diff --git a/thys/Separation_Algebra/ROOT b/thys/Separation_Algebra/ROOT --- a/thys/Separation_Algebra/ROOT +++ b/thys/Separation_Algebra/ROOT @@ -1,18 +1,19 @@ chapter AFP -session Separation_Algebra (AFP) = "HOL-Word" + +session Separation_Algebra (AFP) = HOL + options [timeout = 600] sessions "HOL-Hoare" + "HOL-Library" directories "ex" "ex/capDL" theories "ex/Simple_Separation_Example" "ex/Sep_Tactics_Test" "ex/VM_Example" Sep_Eq "ex/capDL/Separation_D" document_files "root.bib" "root.tex" diff --git a/thys/Separation_Algebra/ex/capDL/Types_D.thy b/thys/Separation_Algebra/ex/capDL/Types_D.thy --- a/thys/Separation_Algebra/ex/capDL/Types_D.thy +++ b/thys/Separation_Algebra/ex/capDL/Types_D.thy @@ -1,233 +1,233 @@ (* Author: Andrew Boyton, 2012 Maintainers: Gerwin Klein Rafal Kolanski *) section "A simplified version of the actual capDL specification." theory Types_D -imports "HOL-Word.Word" +imports "HOL-Library.Word" begin (* * Objects are named by 32 bit words. * This name may correspond to the memory address of the object. *) type_synonym cdl_object_id = "32 word" type_synonym cdl_object_set = "cdl_object_id set" (* The type we use to represent object sizes. *) type_synonym cdl_size_bits = nat (* An index into a CNode, TCB, or other kernel object that contains caps. *) type_synonym cdl_cnode_index = nat (* A reference to a capability slot. *) type_synonym cdl_cap_ref = "cdl_object_id \ cdl_cnode_index" (* The possible access-control rights that exist in the system. *) datatype cdl_right = AllowRead | AllowWrite | AllowGrant (* * Kernel capabilities. * * Such capabilities (or "caps") give the holder particular rights to * a kernel object or system hardware. * * Caps have attributes such as the object they point to, the rights * they give the holder, or how the holder is allowed to interact with * the target object. * * This is a simplified, cut-down version of this datatype for * demonstration purposes. *) datatype cdl_cap = NullCap | EndpointCap cdl_object_id "cdl_right set" | CNodeCap cdl_object_id | TcbCap cdl_object_id (* A mapping from capability identifiers to capabilities. *) type_synonym cdl_cap_map = "cdl_cnode_index \ cdl_cap option" translations (type) "cdl_cap_map" <= (type) "nat \ cdl_cap option" (type) "cdl_cap_ref" <= (type) "cdl_object_id \ nat" (* A user cap pointer. *) type_synonym cdl_cptr = "32 word" (* Kernel objects *) record cdl_tcb = cdl_tcb_caps :: cdl_cap_map cdl_tcb_fault_endpoint :: cdl_cptr record cdl_cnode = cdl_cnode_caps :: cdl_cap_map cdl_cnode_size_bits :: cdl_size_bits (* * Kernel objects. * * These are in-memory objects that may, over the course of the system * execution, be created or deleted by users. * * Again, a simplified version of the real datatype. *) datatype cdl_object = Endpoint | Tcb cdl_tcb | CNode cdl_cnode (* * The current state of the system. * * The state record contains the following primary pieces of information: * * objects: * The objects that currently exist in the system. * * current_thread: * The currently running thread. Operations will always be performed * on behalf of this thread. * * ghost_state: (Used for separation logic) * Which fields are owned by an object. * In capDL this is all of the fields (or none of them). * In any concrete state, this will be all of the fields. *) (* The ghost state tracks which components (fields and slots) are owned by an object. * Fields + slots are encoded as None + Some nat. *) type_synonym cdl_heap = "cdl_object_id \ cdl_object option" type_synonym cdl_component = "nat option" type_synonym cdl_components = "cdl_component set" type_synonym cdl_ghost_state = "cdl_object_id \ cdl_components" translations (type) "cdl_heap" <= (type) "cdl_object_id \ cdl_object option" (type) "cdl_ghost_state" <= (type) "cdl_object_id \ nat option set" record cdl_state = cdl_objects :: "cdl_heap" cdl_current_thread :: "cdl_object_id option" cdl_ghost_state :: "cdl_ghost_state" (* Kernel objects types. *) datatype cdl_object_type = EndpointType | TcbType | CNodeType (* Return the type of an object. *) definition object_type :: "cdl_object \ cdl_object_type" where "object_type x \ case x of Endpoint \ EndpointType | Tcb _ \ TcbType | CNode _ \ CNodeType" (* * Getters and setters for various data types. *) (* Capability getters / setters *) definition cap_objects :: "cdl_cap \ cdl_object_id set" where "cap_objects cap \ case cap of TcbCap x \ {x} | CNodeCap x \ {x} | EndpointCap x _ \ {x}" definition cap_has_object :: "cdl_cap \ bool" where "cap_has_object cap \ case cap of NullCap \ False | _ \ True" definition cap_object :: "cdl_cap \ cdl_object_id" where "cap_object cap \ if cap_has_object cap then THE obj_id. cap_objects cap = {obj_id} else undefined " lemma cap_object_simps: "cap_object (TcbCap x) = x" "cap_object (CNodeCap x) = x" "cap_object (EndpointCap x j) = x" by (simp_all add:cap_object_def cap_objects_def cap_has_object_def) definition cap_rights :: "cdl_cap \ cdl_right set" where "cap_rights c \ case c of EndpointCap _ x \ x | _ \ UNIV" definition update_cap_rights :: "cdl_right set \ cdl_cap \ cdl_cap" where "update_cap_rights r c \ case c of EndpointCap f1 _ \ EndpointCap f1 r | _ \ c" (* Kernel object getters / setters *) definition object_slots :: "cdl_object \ cdl_cap_map" where "object_slots obj \ case obj of CNode x \ cdl_cnode_caps x | Tcb x \ cdl_tcb_caps x | _ \ Map.empty" definition update_slots :: "cdl_cap_map \ cdl_object \ cdl_object" where "update_slots new_val obj \ case obj of CNode x \ CNode (x\cdl_cnode_caps := new_val\) | Tcb x \ Tcb (x\cdl_tcb_caps := new_val\) | _ \ obj" (* Adds new caps to an object. It won't overwrite on a collision. *) definition add_to_slots :: "cdl_cap_map \ cdl_object \ cdl_object" where "add_to_slots new_val obj \ update_slots (new_val ++ (object_slots obj)) obj" definition slots_of :: "cdl_heap \ cdl_object_id \ cdl_cap_map" where "slots_of h \ \obj_id. case h obj_id of None \ Map.empty | Some obj \ object_slots obj" definition has_slots :: "cdl_object \ bool" where "has_slots obj \ case obj of CNode _ \ True | Tcb _ \ True | _ \ False" definition object_at :: "(cdl_object \ bool) \ cdl_object_id \ cdl_heap \ bool" where "object_at P p s \ \object. s p = Some object \ P object" abbreviation "ko_at k \ object_at ((=) k)" end diff --git a/thys/Separation_Logic_Imperative_HOL/ROOT b/thys/Separation_Logic_Imperative_HOL/ROOT --- a/thys/Separation_Logic_Imperative_HOL/ROOT +++ b/thys/Separation_Logic_Imperative_HOL/ROOT @@ -1,39 +1,38 @@ chapter AFP session "Separation_Logic_Imperative_HOL" (AFP) = "HOL-Library" + options [timeout = 600] sessions "HOL-Imperative_HOL" - "HOL-Word" "HOL-ex" Automatic_Refinement Collections Native_Word directories "Tools" "Examples" theories [document = false] "Tools/Imperative_HOL_Add" "Tools/Syntax_Match" Collections.HashCode Collections.Code_Target_ICF Automatic_Refinement.Misc theories Sep_Main "Examples/Imp_List_Spec" "Examples/List_Seg" "Examples/Open_List" "Examples/Circ_List" "Examples/Imp_Map_Spec" "Examples/Hash_Map_Impl" "Examples/Imp_Set_Spec" "Examples/Hash_Set_Impl" "Examples/To_List_GA" "Examples/Union_Find" "Examples/Idioms" theories [document = false] "Examples/Default_Insts" Sep_Examples document_files "root.bib" "root.tex" diff --git a/thys/Simple_Firewall/Primitives/L4_Protocol.thy b/thys/Simple_Firewall/Primitives/L4_Protocol.thy --- a/thys/Simple_Firewall/Primitives/L4_Protocol.thy +++ b/thys/Simple_Firewall/Primitives/L4_Protocol.thy @@ -1,112 +1,112 @@ theory L4_Protocol -imports "../Common/Lib_Enum_toString" "HOL-Word.Word" +imports "../Common/Lib_Enum_toString" "HOL-Library.Word" begin section\Transport Layer Protocols\ type_synonym primitive_protocol = "8 word" definition "ICMP \ 1 :: 8 word" definition "TCP \ 6 :: 8 word" definition "UDP \ 17 :: 8 word" context begin (*let's not pollute the namespace too much*) qualified definition "SCTP \ 132 :: 8 word" qualified definition "IGMP \ 2 :: 8 word" qualified definition "GRE \ 47 :: 8 word" qualified definition "ESP \ 50 :: 8 word" qualified definition "AH \ 51 :: 8 word" qualified definition "IPv6ICMP \ 58 :: 8 word" end (* turn http://www.iana.org/assignments/protocol-numbers/protocol-numbers.xhtml into a separate file or so? *) datatype protocol = ProtoAny | Proto "primitive_protocol" fun match_proto :: "protocol \ primitive_protocol \ bool" where "match_proto ProtoAny _ \ True" | "match_proto (Proto (p)) p_p \ p_p = p" fun simple_proto_conjunct :: "protocol \ protocol \ protocol option" where "simple_proto_conjunct ProtoAny proto = Some proto" | "simple_proto_conjunct proto ProtoAny = Some proto" | "simple_proto_conjunct (Proto p1) (Proto p2) = (if p1 = p2 then Some (Proto p1) else None)" lemma simple_proto_conjunct_asimp[simp]: "simple_proto_conjunct proto ProtoAny = Some proto" by(cases proto) simp_all lemma simple_proto_conjunct_correct: "match_proto p1 pkt \ match_proto p2 pkt \ (case simple_proto_conjunct p1 p2 of None \ False | Some proto \ match_proto proto pkt)" apply(cases p1) apply(simp_all) apply(cases p2) apply(simp_all) done lemma simple_proto_conjunct_Some: "simple_proto_conjunct p1 p2 = Some proto \ match_proto proto pkt \ match_proto p1 pkt \ match_proto p2 pkt" using simple_proto_conjunct_correct by simp lemma simple_proto_conjunct_None: "simple_proto_conjunct p1 p2 = None \ \ (match_proto p1 pkt \ match_proto p2 pkt)" using simple_proto_conjunct_correct by simp lemma conjunctProtoD: "simple_proto_conjunct a (Proto b) = Some x \ x = Proto b \ (a = ProtoAny \ a = Proto b)" by(cases a) (simp_all split: if_splits) lemma conjunctProtoD2: "simple_proto_conjunct (Proto b) a = Some x \ x = Proto b \ (a = ProtoAny \ a = Proto b)" by(cases a) (simp_all split: if_splits) text\Originally, there was a @{typ nat} in the protocol definition, allowing infinitely many protocols This was intended behavior. We want to prevent things such as @{term "\TCP = UDP"}. So be careful with what you prove...\ lemma primitive_protocol_Ex_neq: "p = Proto pi \ \p'. p' \ pi" for pi proof show "pi + 1 \ pi" by simp qed lemma protocol_Ex_neq: "\p'. Proto p' \ p" by(cases p) (simp_all add: primitive_protocol_Ex_neq) section\TCP flags\ datatype tcp_flag = TCP_SYN | TCP_ACK | TCP_FIN | TCP_RST | TCP_URG | TCP_PSH (*| TCP_ALL | TCP_NONE*) lemma UNIV_tcp_flag: "UNIV = {TCP_SYN, TCP_ACK, TCP_FIN, TCP_RST, TCP_URG, TCP_PSH}" using tcp_flag.exhaust by auto instance tcp_flag :: finite proof from UNIV_tcp_flag show "finite (UNIV:: tcp_flag set)" using finite.simps by auto qed instantiation "tcp_flag" :: enum begin definition "enum_tcp_flag = [TCP_SYN, TCP_ACK, TCP_FIN, TCP_RST, TCP_URG, TCP_PSH]" definition "enum_all_tcp_flag P \ P TCP_SYN \ P TCP_ACK \ P TCP_FIN \ P TCP_RST \ P TCP_URG \ P TCP_PSH" definition "enum_ex_tcp_flag P \ P TCP_SYN \ P TCP_ACK \ P TCP_FIN \ P TCP_RST \ P TCP_URG \ P TCP_PSH" instance proof show "UNIV = set (enum_class.enum :: tcp_flag list)" by(simp add: UNIV_tcp_flag enum_tcp_flag_def) next show "distinct (enum_class.enum :: tcp_flag list)" by(simp add: enum_tcp_flag_def) next show "\P. (enum_class.enum_all :: (tcp_flag \ bool) \ bool) P = Ball UNIV P" by(simp add: UNIV_tcp_flag enum_all_tcp_flag_def) next show "\P. (enum_class.enum_ex :: (tcp_flag \ bool) \ bool) P = Bex UNIV P" by(simp add: UNIV_tcp_flag enum_ex_tcp_flag_def) qed end subsection\TCP Flags to String\ fun tcp_flag_toString :: "tcp_flag \ string" where "tcp_flag_toString TCP_SYN = ''TCP_SYN''" | "tcp_flag_toString TCP_ACK = ''TCP_ACK''" | "tcp_flag_toString TCP_FIN = ''TCP_FIN''" | "tcp_flag_toString TCP_RST = ''TCP_RST''" | "tcp_flag_toString TCP_URG = ''TCP_URG''" | "tcp_flag_toString TCP_PSH = ''TCP_PSH''" definition ipt_tcp_flags_toString :: "tcp_flag set \ char list" where "ipt_tcp_flags_toString flags \ list_toString tcp_flag_toString (enum_set_to_list flags)" lemma "ipt_tcp_flags_toString {TCP_SYN,TCP_SYN,TCP_ACK} = ''[TCP_SYN, TCP_ACK]''" by eval end diff --git a/thys/Word_Lib/Ancient_Numeral.thy b/thys/Word_Lib/Ancient_Numeral.thy new file mode 100644 --- /dev/null +++ b/thys/Word_Lib/Ancient_Numeral.thy @@ -0,0 +1,229 @@ +theory Ancient_Numeral + imports Main Bits_Int Misc_lsb Misc_msb +begin + +definition Bit :: "int \ bool \ int" (infixl "BIT" 90) + where "k BIT b = (if b then 1 else 0) + k + k" + +lemma Bit_B0: "k BIT False = k + k" + by (simp add: Bit_def) + +lemma Bit_B1: "k BIT True = k + k + 1" + by (simp add: Bit_def) + +lemma Bit_B0_2t: "k BIT False = 2 * k" + by (rule trans, rule Bit_B0) simp + +lemma Bit_B1_2t: "k BIT True = 2 * k + 1" + by (rule trans, rule Bit_B1) simp + +lemma uminus_Bit_eq: + "- k BIT b = (- k - of_bool b) BIT b" + by (cases b) (simp_all add: Bit_def) + +lemma power_BIT: "2 ^ Suc n - 1 = (2 ^ n - 1) BIT True" + by (simp add: Bit_B1) + +lemma bin_rl_simp [simp]: "bin_rest w BIT bin_last w = w" + by (simp add: Bit_def) + +lemma bin_rest_BIT [simp]: "bin_rest (x BIT b) = x" + by (simp add: Bit_def) + +lemma even_BIT [simp]: "even (x BIT b) \ \ b" + by (simp add: Bit_def) + +lemma bin_last_BIT [simp]: "bin_last (x BIT b) = b" + by simp + +lemma BIT_eq_iff [iff]: "u BIT b = v BIT c \ u = v \ b = c" + by (auto simp: Bit_def) arith+ + +lemma BIT_bin_simps [simp]: + "numeral k BIT False = numeral (Num.Bit0 k)" + "numeral k BIT True = numeral (Num.Bit1 k)" + "(- numeral k) BIT False = - numeral (Num.Bit0 k)" + "(- numeral k) BIT True = - numeral (Num.BitM k)" + by (simp_all only: Bit_B0 Bit_B1 numeral.simps numeral_BitM) + +lemma BIT_special_simps [simp]: + shows "0 BIT False = 0" + and "0 BIT True = 1" + and "1 BIT False = 2" + and "1 BIT True = 3" + and "(- 1) BIT False = - 2" + and "(- 1) BIT True = - 1" + by (simp_all add: Bit_def) + +lemma Bit_eq_0_iff: "w BIT b = 0 \ w = 0 \ \ b" + by (auto simp: Bit_def) arith + +lemma Bit_eq_m1_iff: "w BIT b = -1 \ w = -1 \ b" + by (auto simp: Bit_def) arith + +lemma expand_BIT: + "numeral (Num.Bit0 w) = numeral w BIT False" + "numeral (Num.Bit1 w) = numeral w BIT True" + "- numeral (Num.Bit0 w) = (- numeral w) BIT False" + "- numeral (Num.Bit1 w) = (- numeral (w + Num.One)) BIT True" + by (simp_all add: BitM_inc_eq add_One) + +lemma less_Bits: "v BIT b < w BIT c \ v < w \ v \ w \ \ b \ c" + by (auto simp: Bit_def) + +lemma le_Bits: "v BIT b \ w BIT c \ v < w \ v \ w \ (\ b \ c)" + by (auto simp: Bit_def) + +lemma pred_BIT_simps [simp]: + "x BIT False - 1 = (x - 1) BIT True" + "x BIT True - 1 = x BIT False" + by (simp_all add: Bit_B0_2t Bit_B1_2t) + +lemma succ_BIT_simps [simp]: + "x BIT False + 1 = x BIT True" + "x BIT True + 1 = (x + 1) BIT False" + by (simp_all add: Bit_B0_2t Bit_B1_2t) + +lemma add_BIT_simps [simp]: + "x BIT False + y BIT False = (x + y) BIT False" + "x BIT False + y BIT True = (x + y) BIT True" + "x BIT True + y BIT False = (x + y) BIT True" + "x BIT True + y BIT True = (x + y + 1) BIT False" + by (simp_all add: Bit_B0_2t Bit_B1_2t) + +lemma mult_BIT_simps [simp]: + "x BIT False * y = (x * y) BIT False" + "x * y BIT False = (x * y) BIT False" + "x BIT True * y = (x * y) BIT False + y" + by (simp_all add: Bit_B0_2t Bit_B1_2t algebra_simps) + +lemma B_mod_2': "X = 2 \ (w BIT True) mod X = 1 \ (w BIT False) mod X = 0" + by (simp add: Bit_B0 Bit_B1) + +lemma bin_ex_rl: "\w b. w BIT b = bin" + by (metis bin_rl_simp) + +lemma bin_exhaust: "(\x b. bin = x BIT b \ Q) \ Q" +by (metis bin_ex_rl) + +lemma bin_abs_lem: "bin = (w BIT b) \ bin \ -1 \ bin \ 0 \ nat \w\ < nat \bin\" + apply clarsimp + apply (unfold Bit_def) + apply (cases b) + apply (clarsimp, arith) + apply (clarsimp, arith) + done + +lemma bin_induct: + assumes PPls: "P 0" + and PMin: "P (- 1)" + and PBit: "\bin bit. P bin \ P (bin BIT bit)" + shows "P bin" + apply (rule_tac P=P and a=bin and f1="nat \ abs" in wf_measure [THEN wf_induct]) + apply (simp add: measure_def inv_image_def) + apply (case_tac x rule: bin_exhaust) + apply (frule bin_abs_lem) + apply (auto simp add : PPls PMin PBit) + done + +lemma Bit_div2: "(w BIT b) div 2 = w" + by (fact bin_rest_BIT) + +lemma twice_conv_BIT: "2 * x = x BIT False" + by (simp add: Bit_def) + +lemma BIT_lt0 [simp]: "x BIT b < 0 \ x < 0" +by(cases b)(auto simp add: Bit_def) + +lemma BIT_ge0 [simp]: "x BIT b \ 0 \ x \ 0" +by(cases b)(auto simp add: Bit_def) + +lemma bin_to_bl_aux_Bit_minus_simp [simp]: + "0 < n \ bin_to_bl_aux n (w BIT b) bl = bin_to_bl_aux (n - 1) w (b # bl)" + by (cases n) auto + +lemma bl_to_bin_BIT: + "bl_to_bin bs BIT b = bl_to_bin (bs @ [b])" + by (simp add: bl_to_bin_append Bit_def) + +lemma bin_nth_0_BIT: "bin_nth (w BIT b) 0 \ b" + by simp + +lemma bin_nth_Suc_BIT: "bin_nth (w BIT b) (Suc n) = bin_nth w n" + by (simp add: bit_Suc) + +lemma bin_nth_minus [simp]: "0 < n \ bin_nth (w BIT b) n = bin_nth w (n - 1)" + by (cases n) (simp_all add: bit_Suc) + +lemma bin_sign_simps [simp]: + "bin_sign (w BIT b) = bin_sign w" + by (simp add: bin_sign_def Bit_def) + +lemma bin_nth_Bit: "bin_nth (w BIT b) n \ n = 0 \ b \ (\m. n = Suc m \ bin_nth w m)" + by (cases n) auto + +lemmas sbintrunc_Suc_BIT [simp] = + signed_take_bit_Suc [where a="w BIT b", simplified bin_last_BIT bin_rest_BIT] for w b + +lemmas sbintrunc_0_BIT_B0 [simp] = + signed_take_bit_0 [where a="w BIT False", simplified bin_last_numeral_simps bin_rest_numeral_simps] + for w + +lemmas sbintrunc_0_BIT_B1 [simp] = + signed_take_bit_0 [where a="w BIT True", simplified bin_last_BIT bin_rest_numeral_simps] + for w + +lemma sbintrunc_Suc_minus_Is: + \0 < n \ + sbintrunc (n - 1) w = y \ + sbintrunc n (w BIT b) = y BIT b\ + by (cases n) (simp_all add: Bit_def signed_take_bit_Suc) + +lemma bin_cat_Suc_Bit: "bin_cat w (Suc n) (v BIT b) = bin_cat w n v BIT b" + by (auto simp add: Bit_def concat_bit_Suc) + +lemma int_not_BIT [simp]: "NOT (w BIT b) = (NOT w) BIT (\ b)" + by (simp add: not_int_def Bit_def) + +lemma int_and_Bits [simp]: "(x BIT b) AND (y BIT c) = (x AND y) BIT (b \ c)" + using and_int_rec [of \x BIT b\ \y BIT c\] by (auto simp add: Bit_B0_2t Bit_B1_2t) + +lemma int_or_Bits [simp]: "(x BIT b) OR (y BIT c) = (x OR y) BIT (b \ c)" + using or_int_rec [of \x BIT b\ \y BIT c\] by (auto simp add: Bit_B0_2t Bit_B1_2t) + +lemma int_xor_Bits [simp]: "(x BIT b) XOR (y BIT c) = (x XOR y) BIT ((b \ c) \ \ (b \ c))" + using xor_int_rec [of \x BIT b\ \y BIT c\] by (auto simp add: Bit_B0_2t Bit_B1_2t) + +lemma mod_BIT: + "bin BIT bit mod 2 ^ Suc n = (bin mod 2 ^ n) BIT bit" for bit +proof - + have "2 * (bin mod 2 ^ n) + 1 = (2 * bin mod 2 ^ Suc n) + 1" + by (simp add: mod_mult_mult1) + also have "\ = ((2 * bin mod 2 ^ Suc n) + 1) mod 2 ^ Suc n" + by (simp add: ac_simps pos_zmod_mult_2) + also have "\ = (2 * bin + 1) mod 2 ^ Suc n" + by (simp only: mod_simps) + finally show ?thesis + by (auto simp add: Bit_def) +qed + +lemma minus_BIT_0: fixes x y :: int shows "x BIT b - y BIT False = (x - y) BIT b" +by(simp add: Bit_def) + +lemma int_lsb_BIT [simp]: fixes x :: int shows + "lsb (x BIT b) \ b" +by(simp add: lsb_int_def) + +lemma int_shiftr_BIT [simp]: fixes x :: int + shows int_shiftr0: "x >> 0 = x" + and int_shiftr_Suc: "x BIT b >> Suc n = x >> n" +proof - + show "x >> 0 = x" by (simp add: shiftr_int_def) + show "x BIT b >> Suc n = x >> n" by (cases b) + (simp_all add: shiftr_int_def Bit_def add.commute pos_zdiv_mult_2) +qed + +lemma msb_BIT [simp]: "msb (x BIT b) = msb x" +by(simp add: msb_int_def) + +end \ No newline at end of file diff --git a/thys/Word_Lib/Bit_Comprehension.thy b/thys/Word_Lib/Bit_Comprehension.thy new file mode 100644 --- /dev/null +++ b/thys/Word_Lib/Bit_Comprehension.thy @@ -0,0 +1,248 @@ +(* Title: HOL/Word/Bit_Comprehension.thy + Author: Brian Huffman, PSU; Jeremy Dawson and Gerwin Klein, NICTA +*) + +section \Comprehension syntax for bit expressions\ + +theory Bit_Comprehension + imports "HOL-Library.Word" +begin + +class bit_comprehension = ring_bit_operations + + fixes set_bits :: \(nat \ bool) \ 'a\ (binder \BITS \ 10) + assumes set_bits_bit_eq: \set_bits (bit a) = a\ +begin + +lemma set_bits_False_eq [simp]: + \(BITS _. False) = 0\ + using set_bits_bit_eq [of 0] by (simp add: bot_fun_def) + +end + +instantiation int :: bit_comprehension +begin + +definition + \set_bits f = ( + if \n. \m\n. f m = f n then + let n = LEAST n. \m\n. f m = f n + in signed_take_bit n (horner_sum of_bool 2 (map f [0.. + +instance proof + fix k :: int + from int_bit_bound [of k] + obtain n where *: \\m. n \ m \ bit k m \ bit k n\ + and **: \n > 0 \ bit k (n - 1) \ bit k n\ + by blast + then have ***: \\n. \n'\n. bit k n' \ bit k n\ + by meson + have l: \(LEAST q. \m\q. bit k m \ bit k q) = n\ + apply (rule Least_equality) + using * apply blast + apply (metis "**" One_nat_def Suc_pred le_cases le0 neq0_conv not_less_eq_eq) + done + show \set_bits (bit k) = k\ + apply (simp only: *** set_bits_int_def horner_sum_bit_eq_take_bit l) + apply simp + apply (rule bit_eqI) + apply (simp add: bit_signed_take_bit_iff min_def) + apply (auto simp add: not_le bit_take_bit_iff dest: *) + done +qed + +end + +lemma int_set_bits_K_False [simp]: "(BITS _. False) = (0 :: int)" + by (simp add: set_bits_int_def) + +lemma int_set_bits_K_True [simp]: "(BITS _. True) = (-1 :: int)" + by (simp add: set_bits_int_def) + +instantiation word :: (len) bit_comprehension +begin + +definition word_set_bits_def: + \(BITS n. P n) = (horner_sum of_bool 2 (map P [0.. + +instance by standard + (simp add: word_set_bits_def horner_sum_bit_eq_take_bit) + +end + +lemma bit_set_bits_word_iff: + \bit (set_bits P :: 'a::len word) n \ n < LENGTH('a) \ P n\ + by (auto simp add: word_set_bits_def bit_horner_sum_bit_word_iff) + +lemma set_bits_K_False [simp]: + \set_bits (\_. False) = (0 :: 'a :: len word)\ + by (rule bit_word_eqI) (simp add: bit_set_bits_word_iff) + +lemma set_bits_int_unfold': + \set_bits f = + (if \n. \n'\n. \ f n' then + let n = LEAST n. \n'\n. \ f n' + in horner_sum of_bool 2 (map f [0..n. \n'\n. f n' then + let n = LEAST n. \n'\n. f n' + in signed_take_bit n (horner_sum of_bool 2 (map f [0.. +proof (cases \\n. \m\n. f m \ f n\) + case True + then obtain q where q: \\m\q. f m \ f q\ + by blast + define n where \n = (LEAST n. \m\n. f m \ f n)\ + have \\m\n. f m \ f n\ + unfolding n_def + using q by (rule LeastI [of _ q]) + then have n: \\m. n \ m \ f m \ f n\ + by blast + from n_def have n_eq: \(LEAST q. \m\q. f m \ f n) = n\ + by (smt Least_equality Least_le \\m\n. f m = f n\ dual_order.refl le_refl n order_refl) + show ?thesis + proof (cases \f n\) + case False + with n have *: \\n. \n'\n. \ f n'\ + by blast + have **: \(LEAST n. \n'\n. \ f n') = n\ + using False n_eq by simp + from * False show ?thesis + apply (simp add: set_bits_int_def n_def [symmetric] ** del: upt.upt_Suc) + apply (auto simp add: take_bit_horner_sum_bit_eq + bit_horner_sum_bit_iff take_map + signed_take_bit_def set_bits_int_def + horner_sum_bit_eq_take_bit simp del: upt.upt_Suc) + done + next + case True + with n have *: \\n. \n'\n. f n'\ + by blast + have ***: \\ (\n. \n'\n. \ f n')\ + apply (rule ccontr) + using * nat_le_linear by auto + have **: \(LEAST n. \n'\n. f n') = n\ + using True n_eq by simp + from * *** True show ?thesis + apply (simp add: set_bits_int_def n_def [symmetric] ** del: upt.upt_Suc) + apply (auto simp add: take_bit_horner_sum_bit_eq + bit_horner_sum_bit_iff take_map + signed_take_bit_def set_bits_int_def + horner_sum_bit_eq_take_bit nth_append simp del: upt.upt_Suc) + done + qed +next + case False + then show ?thesis + by (auto simp add: set_bits_int_def) +qed + +inductive wf_set_bits_int :: "(nat \ bool) \ bool" + for f :: "nat \ bool" +where + zeros: "\n' \ n. \ f n' \ wf_set_bits_int f" +| ones: "\n' \ n. f n' \ wf_set_bits_int f" + +lemma wf_set_bits_int_simps: "wf_set_bits_int f \ (\n. (\n'\n. \ f n') \ (\n'\n. f n'))" +by(auto simp add: wf_set_bits_int.simps) + +lemma wf_set_bits_int_const [simp]: "wf_set_bits_int (\_. b)" +by(cases b)(auto intro: wf_set_bits_int.intros) + +lemma wf_set_bits_int_fun_upd [simp]: + "wf_set_bits_int (f(n := b)) \ wf_set_bits_int f" (is "?lhs \ ?rhs") +proof + assume ?lhs + then obtain n' + where "(\n''\n'. \ (f(n := b)) n'') \ (\n''\n'. (f(n := b)) n'')" + by(auto simp add: wf_set_bits_int_simps) + hence "(\n''\max (Suc n) n'. \ f n'') \ (\n''\max (Suc n) n'. f n'')" by auto + thus ?rhs by(auto simp only: wf_set_bits_int_simps) +next + assume ?rhs + then obtain n' where "(\n''\n'. \ f n'') \ (\n''\n'. f n'')" (is "?wf f n'") + by(auto simp add: wf_set_bits_int_simps) + hence "?wf (f(n := b)) (max (Suc n) n')" by auto + thus ?lhs by(auto simp only: wf_set_bits_int_simps) +qed + +lemma wf_set_bits_int_Suc [simp]: + "wf_set_bits_int (\n. f (Suc n)) \ wf_set_bits_int f" (is "?lhs \ ?rhs") +by(auto simp add: wf_set_bits_int_simps intro: le_SucI dest: Suc_le_D) + +context + fixes f + assumes wff: "wf_set_bits_int f" +begin + +lemma int_set_bits_unfold_BIT: + "set_bits f = of_bool (f 0) + (2 :: int) * set_bits (f \ Suc)" +using wff proof cases + case (zeros n) + show ?thesis + proof(cases "\n. \ f n") + case True + hence "f = (\_. False)" by auto + thus ?thesis using True by(simp add: o_def) + next + case False + then obtain n' where "f n'" by blast + with zeros have "(LEAST n. \n'\n. \ f n') = Suc (LEAST n. \n'\Suc n. \ f n')" + by(auto intro: Least_Suc) + also have "(\n. \n'\Suc n. \ f n') = (\n. \n'\n. \ f (Suc n'))" by(auto dest: Suc_le_D) + also from zeros have "\n'\n. \ f (Suc n')" by auto + ultimately show ?thesis using zeros + apply (simp (no_asm_simp) add: set_bits_int_unfold' exI + del: upt.upt_Suc flip: map_map split del: if_split) + apply (simp only: map_Suc_upt upt_conv_Cons) + apply simp + done + qed +next + case (ones n) + show ?thesis + proof(cases "\n. f n") + case True + hence "f = (\_. True)" by auto + thus ?thesis using True by(simp add: o_def) + next + case False + then obtain n' where "\ f n'" by blast + with ones have "(LEAST n. \n'\n. f n') = Suc (LEAST n. \n'\Suc n. f n')" + by(auto intro: Least_Suc) + also have "(\n. \n'\Suc n. f n') = (\n. \n'\n. f (Suc n'))" by(auto dest: Suc_le_D) + also from ones have "\n'\n. f (Suc n')" by auto + moreover from ones have "(\n. \n'\n. \ f n') = False" + by(auto intro!: exI[where x="max n m" for n m] simp add: max_def split: if_split_asm) + moreover hence "(\n. \n'\n. \ f (Suc n')) = False" + by(auto elim: allE[where x="Suc n" for n] dest: Suc_le_D) + ultimately show ?thesis using ones + apply (simp (no_asm_simp) add: set_bits_int_unfold' exI split del: if_split) + apply (auto simp add: Let_def hd_map map_tl[symmetric] map_map[symmetric] map_Suc_upt upt_conv_Cons signed_take_bit_Suc + not_le simp del: map_map) + done + qed +qed + +lemma bin_last_set_bits [simp]: + "odd (set_bits f :: int) = f 0" + by (subst int_set_bits_unfold_BIT) simp_all + +lemma bin_rest_set_bits [simp]: + "set_bits f div (2 :: int) = set_bits (f \ Suc)" + by (subst int_set_bits_unfold_BIT) simp_all + +lemma bin_nth_set_bits [simp]: + "bit (set_bits f :: int) m \ f m" +using wff proof (induction m arbitrary: f) + case 0 + then show ?case + by (simp add: Bit_Comprehension.bin_last_set_bits) +next + case Suc + from Suc.IH [of "f \ Suc"] Suc.prems show ?case + by (simp add: Bit_Comprehension.bin_rest_set_bits comp_def bit_Suc) +qed + +end + +end diff --git a/thys/Word_Lib/Bits_Int.thy b/thys/Word_Lib/Bits_Int.thy new file mode 100644 --- /dev/null +++ b/thys/Word_Lib/Bits_Int.thy @@ -0,0 +1,1405 @@ +(* Title: HOL/Word/Bits_Int.thy + Author: Jeremy Dawson and Gerwin Klein, NICTA +*) + +section \Bitwise Operations on integers\ + +theory Bits_Int + imports + "HOL-Library.Word" + Traditional_Syntax +begin + +subsection \Implicit bit representation of \<^typ>\int\\ + +abbreviation (input) bin_last :: "int \ bool" + where "bin_last \ odd" + +lemma bin_last_def: + "bin_last w \ w mod 2 = 1" + by (fact odd_iff_mod_2_eq_one) + +abbreviation (input) bin_rest :: "int \ int" + where "bin_rest w \ w div 2" + +lemma bin_last_numeral_simps [simp]: + "\ bin_last 0" + "bin_last 1" + "bin_last (- 1)" + "bin_last Numeral1" + "\ bin_last (numeral (Num.Bit0 w))" + "bin_last (numeral (Num.Bit1 w))" + "\ bin_last (- numeral (Num.Bit0 w))" + "bin_last (- numeral (Num.Bit1 w))" + by simp_all + +lemma bin_rest_numeral_simps [simp]: + "bin_rest 0 = 0" + "bin_rest 1 = 0" + "bin_rest (- 1) = - 1" + "bin_rest Numeral1 = 0" + "bin_rest (numeral (Num.Bit0 w)) = numeral w" + "bin_rest (numeral (Num.Bit1 w)) = numeral w" + "bin_rest (- numeral (Num.Bit0 w)) = - numeral w" + "bin_rest (- numeral (Num.Bit1 w)) = - numeral (w + Num.One)" + by simp_all + +lemma bin_rl_eqI: "\bin_rest x = bin_rest y; bin_last x = bin_last y\ \ x = y" + by (auto elim: oddE) + +lemma [simp]: + shows bin_rest_lt0: "bin_rest i < 0 \ i < 0" + and bin_rest_ge_0: "bin_rest i \ 0 \ i \ 0" + by auto + +lemma bin_rest_gt_0 [simp]: "bin_rest x > 0 \ x > 1" + by auto + + +subsection \Bit projection\ + +abbreviation (input) bin_nth :: \int \ nat \ bool\ + where \bin_nth \ bit\ + +lemma bin_nth_eq_iff: "bin_nth x = bin_nth y \ x = y" + by (simp add: bit_eq_iff fun_eq_iff) + +lemma bin_eqI: + "x = y" if "\n. bin_nth x n \ bin_nth y n" + using that bin_nth_eq_iff [of x y] by (simp add: fun_eq_iff) + +lemma bin_eq_iff: "x = y \ (\n. bin_nth x n = bin_nth y n)" + by (fact bit_eq_iff) + +lemma bin_nth_zero [simp]: "\ bin_nth 0 n" + by simp + +lemma bin_nth_1 [simp]: "bin_nth 1 n \ n = 0" + by (cases n) (simp_all add: bit_Suc) + +lemma bin_nth_minus1 [simp]: "bin_nth (- 1) n" + by (induction n) (simp_all add: bit_Suc) + +lemma bin_nth_numeral: "bin_rest x = y \ bin_nth x (numeral n) = bin_nth y (pred_numeral n)" + by (simp add: numeral_eq_Suc bit_Suc) + +lemmas bin_nth_numeral_simps [simp] = + bin_nth_numeral [OF bin_rest_numeral_simps(2)] + bin_nth_numeral [OF bin_rest_numeral_simps(5)] + bin_nth_numeral [OF bin_rest_numeral_simps(6)] + bin_nth_numeral [OF bin_rest_numeral_simps(7)] + bin_nth_numeral [OF bin_rest_numeral_simps(8)] + +lemmas bin_nth_simps = + bit_0 bit_Suc bin_nth_zero bin_nth_minus1 + bin_nth_numeral_simps + +lemma nth_2p_bin: "bin_nth (2 ^ n) m = (m = n)" \ \for use when simplifying with \bin_nth_Bit\\ + by (auto simp add: bit_exp_iff) + +lemma nth_rest_power_bin: "bin_nth ((bin_rest ^^ k) w) n = bin_nth w (n + k)" + apply (induct k arbitrary: n) + apply clarsimp + apply clarsimp + apply (simp only: bit_Suc [symmetric] add_Suc) + done + +lemma bin_nth_numeral_unfold: + "bin_nth (numeral (num.Bit0 x)) n \ n > 0 \ bin_nth (numeral x) (n - 1)" + "bin_nth (numeral (num.Bit1 x)) n \ (n > 0 \ bin_nth (numeral x) (n - 1))" + by (cases n; simp)+ + + +subsection \Truncating\ + +definition bin_sign :: "int \ int" + where "bin_sign k = (if k \ 0 then 0 else - 1)" + +lemma bin_sign_simps [simp]: + "bin_sign 0 = 0" + "bin_sign 1 = 0" + "bin_sign (- 1) = - 1" + "bin_sign (numeral k) = 0" + "bin_sign (- numeral k) = -1" + by (simp_all add: bin_sign_def) + +lemma bin_sign_rest [simp]: "bin_sign (bin_rest w) = bin_sign w" + by (simp add: bin_sign_def) + +abbreviation (input) bintrunc :: \nat \ int \ int\ + where \bintrunc \ take_bit\ + +lemma bintrunc_mod2p: "bintrunc n w = w mod 2 ^ n" + by (fact take_bit_eq_mod) + +abbreviation (input) sbintrunc :: \nat \ int \ int\ + where \sbintrunc \ signed_take_bit\ + +abbreviation (input) norm_sint :: \nat \ int \ int\ + where \norm_sint n \ signed_take_bit (n - 1)\ + +lemma sbintrunc_mod2p: "sbintrunc n w = (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n" + by (simp add: bintrunc_mod2p signed_take_bit_eq_take_bit_shift) + +lemma sbintrunc_eq_take_bit: + \sbintrunc n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\ + by (fact signed_take_bit_eq_take_bit_shift) + +lemma sign_bintr: "bin_sign (bintrunc n w) = 0" + by (simp add: bin_sign_def) + +lemma bintrunc_n_0: "bintrunc n 0 = 0" + by (fact take_bit_of_0) + +lemma sbintrunc_n_0: "sbintrunc n 0 = 0" + by (fact signed_take_bit_of_0) + +lemma sbintrunc_n_minus1: "sbintrunc n (- 1) = -1" + by (fact signed_take_bit_of_minus_1) + +lemma bintrunc_Suc_numeral: + "bintrunc (Suc n) 1 = 1" + "bintrunc (Suc n) (- 1) = 1 + 2 * bintrunc n (- 1)" + "bintrunc (Suc n) (numeral (Num.Bit0 w)) = 2 * bintrunc n (numeral w)" + "bintrunc (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * bintrunc n (numeral w)" + "bintrunc (Suc n) (- numeral (Num.Bit0 w)) = 2 * bintrunc n (- numeral w)" + "bintrunc (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * bintrunc n (- numeral (w + Num.One))" + by (simp_all add: take_bit_Suc) + +lemma sbintrunc_0_numeral [simp]: + "sbintrunc 0 1 = -1" + "sbintrunc 0 (numeral (Num.Bit0 w)) = 0" + "sbintrunc 0 (numeral (Num.Bit1 w)) = -1" + "sbintrunc 0 (- numeral (Num.Bit0 w)) = 0" + "sbintrunc 0 (- numeral (Num.Bit1 w)) = -1" + by simp_all + +lemma sbintrunc_Suc_numeral: + "sbintrunc (Suc n) 1 = 1" + "sbintrunc (Suc n) (numeral (Num.Bit0 w)) = 2 * sbintrunc n (numeral w)" + "sbintrunc (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * sbintrunc n (numeral w)" + "sbintrunc (Suc n) (- numeral (Num.Bit0 w)) = 2 * sbintrunc n (- numeral w)" + "sbintrunc (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * sbintrunc n (- numeral (w + Num.One))" + by (simp_all add: signed_take_bit_Suc) + +lemma bin_sign_lem: "(bin_sign (sbintrunc n bin) = -1) = bit bin n" + by (simp add: bin_sign_def) + +lemma nth_bintr: "bin_nth (bintrunc m w) n \ n < m \ bin_nth w n" + by (fact bit_take_bit_iff) + +lemma nth_sbintr: "bin_nth (sbintrunc m w) n = (if n < m then bin_nth w n else bin_nth w m)" + by (simp add: bit_signed_take_bit_iff min_def) + +lemma bin_nth_Bit0: + "bin_nth (numeral (Num.Bit0 w)) n \ + (\m. n = Suc m \ bin_nth (numeral w) m)" + using bit_double_iff [of \numeral w :: int\ n] + by (auto intro: exI [of _ \n - 1\]) + +lemma bin_nth_Bit1: + "bin_nth (numeral (Num.Bit1 w)) n \ + n = 0 \ (\m. n = Suc m \ bin_nth (numeral w) m)" + using even_bit_succ_iff [of \2 * numeral w :: int\ n] + bit_double_iff [of \numeral w :: int\ n] + by auto + +lemma bintrunc_bintrunc_l: "n \ m \ bintrunc m (bintrunc n w) = bintrunc n w" + by (simp add: min.absorb2) + +lemma sbintrunc_sbintrunc_l: "n \ m \ sbintrunc m (sbintrunc n w) = sbintrunc n w" + by (simp add: min_def) + +lemma bintrunc_bintrunc_ge: "n \ m \ bintrunc n (bintrunc m w) = bintrunc n w" + by (rule bin_eqI) (auto simp: nth_bintr) + +lemma bintrunc_bintrunc_min [simp]: "bintrunc m (bintrunc n w) = bintrunc (min m n) w" + by (rule bin_eqI) (auto simp: nth_bintr) + +lemma sbintrunc_sbintrunc_min [simp]: "sbintrunc m (sbintrunc n w) = sbintrunc (min m n) w" + by (rule bin_eqI) (auto simp: nth_sbintr min.absorb1 min.absorb2) + +lemmas sbintrunc_Suc_Pls = + signed_take_bit_Suc [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] + +lemmas sbintrunc_Suc_Min = + signed_take_bit_Suc [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] + +lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min + sbintrunc_Suc_numeral + +lemmas sbintrunc_Pls = + signed_take_bit_0 [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] + +lemmas sbintrunc_Min = + signed_take_bit_0 [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] + +lemmas sbintrunc_0_simps = + sbintrunc_Pls sbintrunc_Min + +lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs + +lemma bintrunc_minus: "0 < n \ bintrunc (Suc (n - 1)) w = bintrunc n w" + by auto + +lemma sbintrunc_minus: "0 < n \ sbintrunc (Suc (n - 1)) w = sbintrunc n w" + by auto + +lemmas sbintrunc_minus_simps = + sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans]] + +lemma sbintrunc_BIT_I: + \0 < n \ + sbintrunc (n - 1) 0 = y \ + sbintrunc n 0 = 2 * y\ + by simp + +lemma sbintrunc_Suc_Is: + \sbintrunc n (- 1) = y \ + sbintrunc (Suc n) (- 1) = 1 + 2 * y\ + by auto + +lemma sbintrunc_Suc_lem: "sbintrunc (Suc n) x = y \ m = Suc n \ sbintrunc m x = y" + by auto + +lemmas sbintrunc_Suc_Ialts = + sbintrunc_Suc_Is [THEN sbintrunc_Suc_lem] + +lemma sbintrunc_bintrunc_lt: "m > n \ sbintrunc n (bintrunc m w) = sbintrunc n w" + by (rule bin_eqI) (auto simp: nth_sbintr nth_bintr) + +lemma bintrunc_sbintrunc_le: "m \ Suc n \ bintrunc m (sbintrunc n w) = bintrunc m w" + apply (rule bin_eqI) + using le_Suc_eq less_Suc_eq_le apply (auto simp: nth_sbintr nth_bintr) + done + +lemmas bintrunc_sbintrunc [simp] = order_refl [THEN bintrunc_sbintrunc_le] +lemmas sbintrunc_bintrunc [simp] = lessI [THEN sbintrunc_bintrunc_lt] +lemmas bintrunc_bintrunc [simp] = order_refl [THEN bintrunc_bintrunc_l] +lemmas sbintrunc_sbintrunc [simp] = order_refl [THEN sbintrunc_sbintrunc_l] + +lemma bintrunc_sbintrunc' [simp]: "0 < n \ bintrunc n (sbintrunc (n - 1) w) = bintrunc n w" + by (cases n) simp_all + +lemma sbintrunc_bintrunc' [simp]: "0 < n \ sbintrunc (n - 1) (bintrunc n w) = sbintrunc (n - 1) w" + by (cases n) simp_all + +lemma bin_sbin_eq_iff: "bintrunc (Suc n) x = bintrunc (Suc n) y \ sbintrunc n x = sbintrunc n y" + apply (rule iffI) + apply (rule box_equals [OF _ sbintrunc_bintrunc sbintrunc_bintrunc]) + apply simp + apply (rule box_equals [OF _ bintrunc_sbintrunc bintrunc_sbintrunc]) + apply simp + done + +lemma bin_sbin_eq_iff': + "0 < n \ bintrunc n x = bintrunc n y \ sbintrunc (n - 1) x = sbintrunc (n - 1) y" + by (cases n) (simp_all add: bin_sbin_eq_iff) + +lemmas bintrunc_sbintruncS0 [simp] = bintrunc_sbintrunc' [unfolded One_nat_def] +lemmas sbintrunc_bintruncS0 [simp] = sbintrunc_bintrunc' [unfolded One_nat_def] + +lemmas bintrunc_bintrunc_l' = le_add1 [THEN bintrunc_bintrunc_l] +lemmas sbintrunc_sbintrunc_l' = le_add1 [THEN sbintrunc_sbintrunc_l] + +(* although bintrunc_minus_simps, if added to default simpset, + tends to get applied where it's not wanted in developing the theories, + we get a version for when the word length is given literally *) + +lemmas nat_non0_gr = + trans [OF iszero_def [THEN Not_eq_iff [THEN iffD2]] refl] + +lemma bintrunc_numeral: + "bintrunc (numeral k) x = of_bool (odd x) + 2 * bintrunc (pred_numeral k) (x div 2)" + by (simp add: numeral_eq_Suc take_bit_Suc mod_2_eq_odd) + +lemma sbintrunc_numeral: + "sbintrunc (numeral k) x = of_bool (odd x) + 2 * sbintrunc (pred_numeral k) (x div 2)" + by (simp add: numeral_eq_Suc signed_take_bit_Suc mod2_eq_if) + +lemma bintrunc_numeral_simps [simp]: + "bintrunc (numeral k) (numeral (Num.Bit0 w)) = + 2 * bintrunc (pred_numeral k) (numeral w)" + "bintrunc (numeral k) (numeral (Num.Bit1 w)) = + 1 + 2 * bintrunc (pred_numeral k) (numeral w)" + "bintrunc (numeral k) (- numeral (Num.Bit0 w)) = + 2 * bintrunc (pred_numeral k) (- numeral w)" + "bintrunc (numeral k) (- numeral (Num.Bit1 w)) = + 1 + 2 * bintrunc (pred_numeral k) (- numeral (w + Num.One))" + "bintrunc (numeral k) 1 = 1" + by (simp_all add: bintrunc_numeral) + +lemma sbintrunc_numeral_simps [simp]: + "sbintrunc (numeral k) (numeral (Num.Bit0 w)) = + 2 * sbintrunc (pred_numeral k) (numeral w)" + "sbintrunc (numeral k) (numeral (Num.Bit1 w)) = + 1 + 2 * sbintrunc (pred_numeral k) (numeral w)" + "sbintrunc (numeral k) (- numeral (Num.Bit0 w)) = + 2 * sbintrunc (pred_numeral k) (- numeral w)" + "sbintrunc (numeral k) (- numeral (Num.Bit1 w)) = + 1 + 2 * sbintrunc (pred_numeral k) (- numeral (w + Num.One))" + "sbintrunc (numeral k) 1 = 1" + by (simp_all add: sbintrunc_numeral) + +lemma no_bintr_alt1: "bintrunc n = (\w. w mod 2 ^ n :: int)" + by (rule ext) (rule bintrunc_mod2p) + +lemma range_bintrunc: "range (bintrunc n) = {i. 0 \ i \ i < 2 ^ n}" + by (auto simp add: take_bit_eq_mod image_iff) (metis mod_pos_pos_trivial) + +lemma no_sbintr_alt2: "sbintrunc n = (\w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" + by (rule ext) (simp add : sbintrunc_mod2p) + +lemma range_sbintrunc: "range (sbintrunc n) = {i. - (2 ^ n) \ i \ i < 2 ^ n}" +proof - + have \surj (\k::int. k + 2 ^ n)\ + by (rule surjI [of _ \(\k. k - 2 ^ n)\]) simp + moreover have \sbintrunc n = ((\k. k - 2 ^ n) \ take_bit (Suc n) \ (\k. k + 2 ^ n))\ + by (simp add: sbintrunc_eq_take_bit fun_eq_iff) + ultimately show ?thesis + apply (simp only: fun.set_map range_bintrunc) + apply (auto simp add: image_iff) + apply presburger + done +qed + +lemma sbintrunc_inc: + \k + 2 ^ Suc n \ sbintrunc n k\ if \k < - (2 ^ n)\ + using that by (fact signed_take_bit_int_greater_eq) + +lemma sbintrunc_dec: + \sbintrunc n k \ k - 2 ^ (Suc n)\ if \k \ 2 ^ n\ + using that by (fact signed_take_bit_int_less_eq) + +lemma bintr_ge0: "0 \ bintrunc n w" + by (simp add: bintrunc_mod2p) + +lemma bintr_lt2p: "bintrunc n w < 2 ^ n" + by (simp add: bintrunc_mod2p) + +lemma bintr_Min: "bintrunc n (- 1) = 2 ^ n - 1" + by (simp add: stable_imp_take_bit_eq) + +lemma sbintr_ge: "- (2 ^ n) \ sbintrunc n w" + by (simp add: sbintrunc_mod2p) + +lemma sbintr_lt: "sbintrunc n w < 2 ^ n" + by (simp add: sbintrunc_mod2p) + +lemma sign_Pls_ge_0: "bin_sign bin = 0 \ bin \ 0" + for bin :: int + by (simp add: bin_sign_def) + +lemma sign_Min_lt_0: "bin_sign bin = -1 \ bin < 0" + for bin :: int + by (simp add: bin_sign_def) + +lemma bin_rest_trunc: "bin_rest (bintrunc n bin) = bintrunc (n - 1) (bin_rest bin)" + by (simp add: take_bit_rec [of n bin]) + +lemma bin_rest_power_trunc: + "(bin_rest ^^ k) (bintrunc n bin) = bintrunc (n - k) ((bin_rest ^^ k) bin)" + by (induct k) (auto simp: bin_rest_trunc) + +lemma bin_rest_trunc_i: "bintrunc n (bin_rest bin) = bin_rest (bintrunc (Suc n) bin)" + by (auto simp add: take_bit_Suc) + +lemma bin_rest_strunc: "bin_rest (sbintrunc (Suc n) bin) = sbintrunc n (bin_rest bin)" + by (simp add: signed_take_bit_Suc) + +lemma bintrunc_rest [simp]: "bintrunc n (bin_rest (bintrunc n bin)) = bin_rest (bintrunc n bin)" + by (induct n arbitrary: bin) (simp_all add: take_bit_Suc) + +lemma sbintrunc_rest [simp]: "sbintrunc n (bin_rest (sbintrunc n bin)) = bin_rest (sbintrunc n bin)" + by (induct n arbitrary: bin) (simp_all add: signed_take_bit_Suc mod2_eq_if) + +lemma bintrunc_rest': "bintrunc n \ bin_rest \ bintrunc n = bin_rest \ bintrunc n" + by (rule ext) auto + +lemma sbintrunc_rest': "sbintrunc n \ bin_rest \ sbintrunc n = bin_rest \ sbintrunc n" + by (rule ext) auto + +lemma rco_lem: "f \ g \ f = g \ f \ f \ (g \ f) ^^ n = g ^^ n \ f" + apply (rule ext) + apply (induct_tac n) + apply (simp_all (no_asm)) + apply (drule fun_cong) + apply (unfold o_def) + apply (erule trans) + apply simp + done + +lemmas rco_bintr = bintrunc_rest' + [THEN rco_lem [THEN fun_cong], unfolded o_def] +lemmas rco_sbintr = sbintrunc_rest' + [THEN rco_lem [THEN fun_cong], unfolded o_def] + + +subsection \Splitting and concatenation\ + +definition bin_split :: \nat \ int \ int \ int\ + where [simp]: \bin_split n k = (drop_bit n k, take_bit n k)\ + +lemma [code]: + "bin_split (Suc n) w = (let (w1, w2) = bin_split n (w div 2) in (w1, of_bool (odd w) + 2 * w2))" + "bin_split 0 w = (w, 0)" + by (simp_all add: drop_bit_Suc take_bit_Suc mod_2_eq_odd) + +abbreviation (input) bin_cat :: \int \ nat \ int \ int\ + where \bin_cat k n l \ concat_bit n l k\ + +lemma bin_cat_eq_push_bit_add_take_bit: + \bin_cat k n l = push_bit n k + take_bit n l\ + by (simp add: concat_bit_eq) + +lemma bin_sign_cat: "bin_sign (bin_cat x n y) = bin_sign x" +proof - + have \0 \ x\ if \0 \ x * 2 ^ n + y mod 2 ^ n\ + proof - + have \y mod 2 ^ n < 2 ^ n\ + using pos_mod_bound [of \2 ^ n\ y] by simp + then have \\ y mod 2 ^ n \ 2 ^ n\ + by (simp add: less_le) + with that have \x \ - 1\ + by auto + have *: \- 1 \ (- (y mod 2 ^ n)) div 2 ^ n\ + by (simp add: zdiv_zminus1_eq_if) + from that have \- (y mod 2 ^ n) \ x * 2 ^ n\ + by simp + then have \(- (y mod 2 ^ n)) div 2 ^ n \ (x * 2 ^ n) div 2 ^ n\ + using zdiv_mono1 zero_less_numeral zero_less_power by blast + with * have \- 1 \ x * 2 ^ n div 2 ^ n\ by simp + with \x \ - 1\ show ?thesis + by simp + qed + then show ?thesis + by (simp add: bin_sign_def not_le not_less bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult take_bit_eq_mod) +qed + +lemma bin_cat_assoc: "bin_cat (bin_cat x m y) n z = bin_cat x (m + n) (bin_cat y n z)" + by (fact concat_bit_assoc) + +lemma bin_cat_assoc_sym: "bin_cat x m (bin_cat y n z) = bin_cat (bin_cat x (m - n) y) (min m n) z" + by (fact concat_bit_assoc_sym) + +definition bin_rcat :: \nat \ int list \ int\ + where \bin_rcat n = horner_sum (take_bit n) (2 ^ n) \ rev\ + +lemma bin_rcat_eq_foldl: + \bin_rcat n = foldl (\u v. bin_cat u n v) 0\ +proof + fix ks :: \int list\ + show \bin_rcat n ks = foldl (\u v. bin_cat u n v) 0 ks\ + by (induction ks rule: rev_induct) + (simp_all add: bin_rcat_def concat_bit_eq push_bit_eq_mult) +qed + +fun bin_rsplit_aux :: "nat \ nat \ int \ int list \ int list" + where "bin_rsplit_aux n m c bs = + (if m = 0 \ n = 0 then bs + else + let (a, b) = bin_split n c + in bin_rsplit_aux n (m - n) a (b # bs))" + +definition bin_rsplit :: "nat \ nat \ int \ int list" + where "bin_rsplit n w = bin_rsplit_aux n (fst w) (snd w) []" + +value \bin_rsplit 1705 (3, 88)\ + +fun bin_rsplitl_aux :: "nat \ nat \ int \ int list \ int list" + where "bin_rsplitl_aux n m c bs = + (if m = 0 \ n = 0 then bs + else + let (a, b) = bin_split (min m n) c + in bin_rsplitl_aux n (m - n) a (b # bs))" + +definition bin_rsplitl :: "nat \ nat \ int \ int list" + where "bin_rsplitl n w = bin_rsplitl_aux n (fst w) (snd w) []" + +declare bin_rsplit_aux.simps [simp del] +declare bin_rsplitl_aux.simps [simp del] + +lemma bin_nth_cat: + "bin_nth (bin_cat x k y) n = + (if n < k then bin_nth y n else bin_nth x (n - k))" + by (simp add: bit_concat_bit_iff) + +lemma bin_nth_drop_bit_iff: + \bin_nth (drop_bit n c) k \ bin_nth c (n + k)\ + by (simp add: bit_drop_bit_eq) + +lemma bin_nth_take_bit_iff: + \bin_nth (take_bit n c) k \ k < n \ bin_nth c k\ + by (fact bit_take_bit_iff) + +lemma bin_nth_split: + "bin_split n c = (a, b) \ + (\k. bin_nth a k = bin_nth c (n + k)) \ + (\k. bin_nth b k = (k < n \ bin_nth c k))" + by (auto simp add: bin_nth_drop_bit_iff bin_nth_take_bit_iff) + +lemma bin_cat_zero [simp]: "bin_cat 0 n w = bintrunc n w" + by (simp add: bin_cat_eq_push_bit_add_take_bit) + +lemma bintr_cat1: "bintrunc (k + n) (bin_cat a n b) = bin_cat (bintrunc k a) n b" + by (metis bin_cat_assoc bin_cat_zero) + +lemma bintr_cat: "bintrunc m (bin_cat a n b) = + bin_cat (bintrunc (m - n) a) n (bintrunc (min m n) b)" + + by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr) + +lemma bintr_cat_same [simp]: "bintrunc n (bin_cat a n b) = bintrunc n b" + by (auto simp add : bintr_cat) + +lemma cat_bintr [simp]: "bin_cat a n (bintrunc n b) = bin_cat a n b" + by (simp add: bin_cat_eq_push_bit_add_take_bit) + +lemma split_bintrunc: "bin_split n c = (a, b) \ b = bintrunc n c" + by simp + +lemma bin_cat_split: "bin_split n w = (u, v) \ w = bin_cat u n v" + by (auto simp add: bin_cat_eq_push_bit_add_take_bit bits_ident) + +lemma drop_bit_bin_cat_eq: + \drop_bit n (bin_cat v n w) = v\ + by (rule bit_eqI) (simp add: bit_drop_bit_eq bit_concat_bit_iff) + +lemma take_bit_bin_cat_eq: + \take_bit n (bin_cat v n w) = take_bit n w\ + by (rule bit_eqI) (simp add: bit_concat_bit_iff) + +lemma bin_split_cat: "bin_split n (bin_cat v n w) = (v, bintrunc n w)" + by (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) + +lemma bin_split_zero [simp]: "bin_split n 0 = (0, 0)" + by simp + +lemma bin_split_minus1 [simp]: + "bin_split n (- 1) = (- 1, bintrunc n (- 1))" + by simp + +lemma bin_split_trunc: + "bin_split (min m n) c = (a, b) \ + bin_split n (bintrunc m c) = (bintrunc (m - n) a, b)" + apply (induct n arbitrary: m b c, clarsimp) + apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) + apply (case_tac m) + apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) + done + +lemma bin_split_trunc1: + "bin_split n c = (a, b) \ + bin_split n (bintrunc m c) = (bintrunc (m - n) a, bintrunc m b)" + apply (induct n arbitrary: m b c, clarsimp) + apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) + apply (case_tac m) + apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) + done + +lemma bin_cat_num: "bin_cat a n b = a * 2 ^ n + bintrunc n b" + by (simp add: bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult) + +lemma bin_split_num: "bin_split n b = (b div 2 ^ n, b mod 2 ^ n)" + by (simp add: drop_bit_eq_div take_bit_eq_mod) + +lemmas bin_rsplit_aux_simps = bin_rsplit_aux.simps bin_rsplitl_aux.simps +lemmas rsplit_aux_simps = bin_rsplit_aux_simps + +lemmas th_if_simp1 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct1, THEN mp] for l +lemmas th_if_simp2 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct2, THEN mp] for l + +lemmas rsplit_aux_simp1s = rsplit_aux_simps [THEN th_if_simp1] + +lemmas rsplit_aux_simp2ls = rsplit_aux_simps [THEN th_if_simp2] +\ \these safe to \[simp add]\ as require calculating \m - n\\ +lemmas bin_rsplit_aux_simp2s [simp] = rsplit_aux_simp2ls [unfolded Let_def] +lemmas rbscl = bin_rsplit_aux_simp2s (2) + +lemmas rsplit_aux_0_simps [simp] = + rsplit_aux_simp1s [OF disjI1] rsplit_aux_simp1s [OF disjI2] + +lemma bin_rsplit_aux_append: "bin_rsplit_aux n m c (bs @ cs) = bin_rsplit_aux n m c bs @ cs" + apply (induct n m c bs rule: bin_rsplit_aux.induct) + apply (subst bin_rsplit_aux.simps) + apply (subst bin_rsplit_aux.simps) + apply (clarsimp split: prod.split) + done + +lemma bin_rsplitl_aux_append: "bin_rsplitl_aux n m c (bs @ cs) = bin_rsplitl_aux n m c bs @ cs" + apply (induct n m c bs rule: bin_rsplitl_aux.induct) + apply (subst bin_rsplitl_aux.simps) + apply (subst bin_rsplitl_aux.simps) + apply (clarsimp split: prod.split) + done + +lemmas rsplit_aux_apps [where bs = "[]"] = + bin_rsplit_aux_append bin_rsplitl_aux_append + +lemmas rsplit_def_auxs = bin_rsplit_def bin_rsplitl_def + +lemmas rsplit_aux_alts = rsplit_aux_apps + [unfolded append_Nil rsplit_def_auxs [symmetric]] + +lemma bin_split_minus: "0 < n \ bin_split (Suc (n - 1)) w = bin_split n w" + by auto + +lemma bin_split_pred_simp [simp]: + "(0::nat) < numeral bin \ + bin_split (numeral bin) w = + (let (w1, w2) = bin_split (numeral bin - 1) (bin_rest w) + in (w1, of_bool (odd w) + 2 * w2))" + by (simp add: take_bit_rec drop_bit_rec mod_2_eq_odd) + +lemma bin_rsplit_aux_simp_alt: + "bin_rsplit_aux n m c bs = + (if m = 0 \ n = 0 then bs + else let (a, b) = bin_split n c in bin_rsplit n (m - n, a) @ b # bs)" + apply (simp add: bin_rsplit_aux.simps [of n m c bs]) + apply (subst rsplit_aux_alts) + apply (simp add: bin_rsplit_def) + done + +lemmas bin_rsplit_simp_alt = + trans [OF bin_rsplit_def bin_rsplit_aux_simp_alt] + +lemmas bthrs = bin_rsplit_simp_alt [THEN [2] trans] + +lemma bin_rsplit_size_sign' [rule_format]: + "n > 0 \ rev sw = bin_rsplit n (nw, w) \ \v\set sw. bintrunc n v = v" + apply (induct sw arbitrary: nw w) + apply clarsimp + apply clarsimp + apply (drule bthrs) + apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) + apply clarify + apply simp + done + +lemmas bin_rsplit_size_sign = bin_rsplit_size_sign' [OF asm_rl + rev_rev_ident [THEN trans] set_rev [THEN equalityD2 [THEN subsetD]]] + +lemma bin_nth_rsplit [rule_format] : + "n > 0 \ m < n \ + \w k nw. + rev sw = bin_rsplit n (nw, w) \ + k < size sw \ bin_nth (sw ! k) m = bin_nth w (k * n + m)" + apply (induct sw) + apply clarsimp + apply clarsimp + apply (drule bthrs) + apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) + apply (erule allE, erule impE, erule exI) + apply (case_tac k) + apply clarsimp + prefer 2 + apply clarsimp + apply (erule allE) + apply (erule (1) impE) + apply (simp add: bit_drop_bit_eq ac_simps) + apply (simp add: bit_take_bit_iff ac_simps) + done + +lemma bin_rsplit_all: "0 < nw \ nw \ n \ bin_rsplit n (nw, w) = [bintrunc n w]" + by (auto simp: bin_rsplit_def rsplit_aux_simp2ls split: prod.split dest!: split_bintrunc) + +lemma bin_rsplit_l [rule_format]: + "\bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, bintrunc m bin)" + apply (rule_tac a = "m" in wf_less_than [THEN wf_induct]) + apply (simp (no_asm) add: bin_rsplitl_def bin_rsplit_def) + apply (rule allI) + apply (subst bin_rsplitl_aux.simps) + apply (subst bin_rsplit_aux.simps) + apply (clarsimp simp: Let_def split: prod.split) + apply (simp add: ac_simps) + apply (subst rsplit_aux_alts(1)) + apply (subst rsplit_aux_alts(2)) + apply clarsimp + unfolding bin_rsplit_def bin_rsplitl_def + apply (simp add: drop_bit_take_bit) + apply (case_tac \x < n\) + apply (simp_all add: not_less min_def) + done + +lemma bin_rsplit_rcat [rule_format]: + "n > 0 \ bin_rsplit n (n * size ws, bin_rcat n ws) = map (bintrunc n) ws" + apply (unfold bin_rsplit_def bin_rcat_eq_foldl) + apply (rule_tac xs = ws in rev_induct) + apply clarsimp + apply clarsimp + apply (subst rsplit_aux_alts) + apply (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) + done + +lemma bin_rsplit_aux_len_le [rule_format] : + "\ws m. n \ 0 \ ws = bin_rsplit_aux n nw w bs \ + length ws \ m \ nw + length bs * n \ m * n" +proof - + have *: R + if d: "i \ j \ m < j'" + and R1: "i * k \ j * k \ R" + and R2: "Suc m * k' \ j' * k' \ R" + for i j j' k k' m :: nat and R + using d + apply safe + apply (rule R1, erule mult_le_mono1) + apply (rule R2, erule Suc_le_eq [THEN iffD2 [THEN mult_le_mono1]]) + done + have **: "0 < sc \ sc - n + (n + lb * n) \ m * n \ sc + lb * n \ m * n" + for sc m n lb :: nat + apply safe + apply arith + apply (case_tac "sc \ n") + apply arith + apply (insert linorder_le_less_linear [of m lb]) + apply (erule_tac k=n and k'=n in *) + apply arith + apply simp + done + show ?thesis + apply (induct n nw w bs rule: bin_rsplit_aux.induct) + apply (subst bin_rsplit_aux.simps) + apply (simp add: ** Let_def split: prod.split) + done +qed + +lemma bin_rsplit_len_le: "n \ 0 \ ws = bin_rsplit n (nw, w) \ length ws \ m \ nw \ m * n" + by (auto simp: bin_rsplit_def bin_rsplit_aux_len_le) + +lemma bin_rsplit_aux_len: + "n \ 0 \ length (bin_rsplit_aux n nw w cs) = (nw + n - 1) div n + length cs" + apply (induct n nw w cs rule: bin_rsplit_aux.induct) + apply (subst bin_rsplit_aux.simps) + apply (clarsimp simp: Let_def split: prod.split) + apply (erule thin_rl) + apply (case_tac m) + apply simp + apply (case_tac "m \ n") + apply (auto simp add: div_add_self2) + done + +lemma bin_rsplit_len: "n \ 0 \ length (bin_rsplit n (nw, w)) = (nw + n - 1) div n" + by (auto simp: bin_rsplit_def bin_rsplit_aux_len) + +lemma bin_rsplit_aux_len_indep: + "n \ 0 \ length bs = length cs \ + length (bin_rsplit_aux n nw v bs) = + length (bin_rsplit_aux n nw w cs)" +proof (induct n nw w cs arbitrary: v bs rule: bin_rsplit_aux.induct) + case (1 n m w cs v bs) + show ?case + proof (cases "m = 0") + case True + with \length bs = length cs\ show ?thesis by simp + next + case False + from "1.hyps" [of \bin_split n w\ \drop_bit n w\ \take_bit n w\] \m \ 0\ \n \ 0\ + have hyp: "\v bs. length bs = Suc (length cs) \ + length (bin_rsplit_aux n (m - n) v bs) = + length (bin_rsplit_aux n (m - n) (drop_bit n w) (take_bit n w # cs))" + using bin_rsplit_aux_len by fastforce + from \length bs = length cs\ \n \ 0\ show ?thesis + by (auto simp add: bin_rsplit_aux_simp_alt Let_def bin_rsplit_len split: prod.split) + qed +qed + +lemma bin_rsplit_len_indep: + "n \ 0 \ length (bin_rsplit n (nw, v)) = length (bin_rsplit n (nw, w))" + apply (unfold bin_rsplit_def) + apply (simp (no_asm)) + apply (erule bin_rsplit_aux_len_indep) + apply (rule refl) + done + + +subsection \Logical operations\ + +primrec bin_sc :: "nat \ bool \ int \ int" + where + Z: "bin_sc 0 b w = of_bool b + 2 * bin_rest w" + | Suc: "bin_sc (Suc n) b w = of_bool (odd w) + 2 * bin_sc n b (w div 2)" + +lemma bin_nth_sc [simp]: "bit (bin_sc n b w) n \ b" + by (induction n arbitrary: w) (simp_all add: bit_Suc) + +lemma bin_sc_sc_same [simp]: "bin_sc n c (bin_sc n b w) = bin_sc n c w" + by (induction n arbitrary: w) (simp_all add: bit_Suc) + +lemma bin_sc_sc_diff: "m \ n \ bin_sc m c (bin_sc n b w) = bin_sc n b (bin_sc m c w)" + apply (induct n arbitrary: w m) + apply (case_tac [!] m) + apply auto + done + +lemma bin_nth_sc_gen: "bin_nth (bin_sc n b w) m = (if m = n then b else bin_nth w m)" + apply (induct n arbitrary: w m) + apply (case_tac m; simp add: bit_Suc) + apply (case_tac m; simp add: bit_Suc) + done + +lemma bin_sc_eq: + \bin_sc n False = unset_bit n\ + \bin_sc n True = Bit_Operations.set_bit n\ + by (simp_all add: fun_eq_iff bit_eq_iff) + (simp_all add: bin_nth_sc_gen bit_set_bit_iff bit_unset_bit_iff) + +lemma bin_sc_nth [simp]: "bin_sc n (bin_nth w n) w = w" + by (rule bit_eqI) (simp add: bin_nth_sc_gen) + +lemma bin_sign_sc [simp]: "bin_sign (bin_sc n b w) = bin_sign w" +proof (induction n arbitrary: w) + case 0 + then show ?case + by (auto simp add: bin_sign_def) (use bin_rest_ge_0 in fastforce) +next + case (Suc n) + from Suc [of \w div 2\] + show ?case by (auto simp add: bin_sign_def split: if_splits) +qed + +lemma bin_sc_bintr [simp]: + "bintrunc m (bin_sc n x (bintrunc m w)) = bintrunc m (bin_sc n x w)" + apply (cases x) + apply (simp_all add: bin_sc_eq bit_eq_iff) + apply (auto simp add: bit_take_bit_iff bit_set_bit_iff bit_unset_bit_iff) + done + +lemma bin_clr_le: "bin_sc n False w \ w" + by (simp add: bin_sc_eq unset_bit_less_eq) + +lemma bin_set_ge: "bin_sc n True w \ w" + by (simp add: bin_sc_eq set_bit_greater_eq) + +lemma bintr_bin_clr_le: "bintrunc n (bin_sc m False w) \ bintrunc n w" + by (simp add: bin_sc_eq take_bit_unset_bit_eq unset_bit_less_eq) + +lemma bintr_bin_set_ge: "bintrunc n (bin_sc m True w) \ bintrunc n w" + by (simp add: bin_sc_eq take_bit_set_bit_eq set_bit_greater_eq) + +lemma bin_sc_FP [simp]: "bin_sc n False 0 = 0" + by (induct n) auto + +lemma bin_sc_TM [simp]: "bin_sc n True (- 1) = - 1" + by (induct n) auto + +lemmas bin_sc_simps = bin_sc.Z bin_sc.Suc bin_sc_TM bin_sc_FP + +lemma bin_sc_minus: "0 < n \ bin_sc (Suc (n - 1)) b w = bin_sc n b w" + by auto + +lemmas bin_sc_Suc_minus = + trans [OF bin_sc_minus [symmetric] bin_sc.Suc] + +lemma bin_sc_numeral [simp]: + "bin_sc (numeral k) b w = + of_bool (odd w) + 2 * bin_sc (pred_numeral k) b (w div 2)" + by (simp add: numeral_eq_Suc) + +instance int :: semiring_bit_syntax .. + +lemma test_bit_int_def [iff]: + "i !! n \ bin_nth i n" + by (simp add: test_bit_eq_bit) + +lemma shiftl_int_def: + "shiftl x n = x * 2 ^ n" for x :: int + by (simp add: push_bit_int_def shiftl_eq_push_bit) + +lemma shiftr_int_def: + "shiftr x n = x div 2 ^ n" for x :: int + by (simp add: drop_bit_int_def shiftr_eq_drop_bit) + + +subsubsection \Basic simplification rules\ + +lemmas int_not_def = not_int_def + +lemma int_not_simps [simp]: + "NOT (0::int) = -1" + "NOT (1::int) = -2" + "NOT (- 1::int) = 0" + "NOT (numeral w::int) = - numeral (w + Num.One)" + "NOT (- numeral (Num.Bit0 w)::int) = numeral (Num.BitM w)" + "NOT (- numeral (Num.Bit1 w)::int) = numeral (Num.Bit0 w)" + by (simp_all add: not_int_def) + +lemma int_not_not: "NOT (NOT x) = x" + for x :: int + by (fact bit.double_compl) + +lemma int_and_0 [simp]: "0 AND x = 0" + for x :: int + by (fact bit.conj_zero_left) + +lemma int_and_m1 [simp]: "-1 AND x = x" + for x :: int + by (fact bit.conj_one_left) + +lemma int_or_zero [simp]: "0 OR x = x" + for x :: int + by (fact bit.disj_zero_left) + +lemma int_or_minus1 [simp]: "-1 OR x = -1" + for x :: int + by (fact bit.disj_one_left) + +lemma int_xor_zero [simp]: "0 XOR x = x" + for x :: int + by (fact bit.xor_zero_left) + + +subsubsection \Binary destructors\ + +lemma bin_rest_NOT [simp]: "bin_rest (NOT x) = NOT (bin_rest x)" + by (fact not_int_div_2) + +lemma bin_last_NOT [simp]: "bin_last (NOT x) \ \ bin_last x" + by simp + +lemma bin_rest_AND [simp]: "bin_rest (x AND y) = bin_rest x AND bin_rest y" + by (subst and_int_rec) auto + +lemma bin_last_AND [simp]: "bin_last (x AND y) \ bin_last x \ bin_last y" + by (subst and_int_rec) auto + +lemma bin_rest_OR [simp]: "bin_rest (x OR y) = bin_rest x OR bin_rest y" + by (subst or_int_rec) auto + +lemma bin_last_OR [simp]: "bin_last (x OR y) \ bin_last x \ bin_last y" + by (subst or_int_rec) auto + +lemma bin_rest_XOR [simp]: "bin_rest (x XOR y) = bin_rest x XOR bin_rest y" + by (subst xor_int_rec) auto + +lemma bin_last_XOR [simp]: "bin_last (x XOR y) \ (bin_last x \ bin_last y) \ \ (bin_last x \ bin_last y)" + by (subst xor_int_rec) auto + +lemma bin_nth_ops: + "\x y. bin_nth (x AND y) n \ bin_nth x n \ bin_nth y n" + "\x y. bin_nth (x OR y) n \ bin_nth x n \ bin_nth y n" + "\x y. bin_nth (x XOR y) n \ bin_nth x n \ bin_nth y n" + "\x. bin_nth (NOT x) n \ \ bin_nth x n" + by (simp_all add: bit_and_iff bit_or_iff bit_xor_iff bit_not_iff) + + +subsubsection \Derived properties\ + +lemma int_xor_minus1 [simp]: "-1 XOR x = NOT x" + for x :: int + by (fact bit.xor_one_left) + +lemma int_xor_extra_simps [simp]: + "w XOR 0 = w" + "w XOR -1 = NOT w" + for w :: int + by simp_all + +lemma int_or_extra_simps [simp]: + "w OR 0 = w" + "w OR -1 = -1" + for w :: int + by simp_all + +lemma int_and_extra_simps [simp]: + "w AND 0 = 0" + "w AND -1 = w" + for w :: int + by simp_all + +text \Commutativity of the above.\ +lemma bin_ops_comm: + fixes x y :: int + shows int_and_comm: "x AND y = y AND x" + and int_or_comm: "x OR y = y OR x" + and int_xor_comm: "x XOR y = y XOR x" + by (simp_all add: ac_simps) + +lemma bin_ops_same [simp]: + "x AND x = x" + "x OR x = x" + "x XOR x = 0" + for x :: int + by simp_all + +lemmas bin_log_esimps = + int_and_extra_simps int_or_extra_simps int_xor_extra_simps + int_and_0 int_and_m1 int_or_zero int_or_minus1 int_xor_zero int_xor_minus1 + + +subsubsection \Basic properties of logical (bit-wise) operations\ + +lemma bbw_ao_absorb: "x AND (y OR x) = x \ x OR (y AND x) = x" + for x y :: int + by (auto simp add: bin_eq_iff bin_nth_ops) + +lemma bbw_ao_absorbs_other: + "x AND (x OR y) = x \ (y AND x) OR x = x" + "(y OR x) AND x = x \ x OR (x AND y) = x" + "(x OR y) AND x = x \ (x AND y) OR x = x" + for x y :: int + by (auto simp add: bin_eq_iff bin_nth_ops) + +lemmas bbw_ao_absorbs [simp] = bbw_ao_absorb bbw_ao_absorbs_other + +lemma int_xor_not: "(NOT x) XOR y = NOT (x XOR y) \ x XOR (NOT y) = NOT (x XOR y)" + for x y :: int + by (auto simp add: bin_eq_iff bin_nth_ops) + +lemma int_and_assoc: "(x AND y) AND z = x AND (y AND z)" + for x y z :: int + by (auto simp add: bin_eq_iff bin_nth_ops) + +lemma int_or_assoc: "(x OR y) OR z = x OR (y OR z)" + for x y z :: int + by (auto simp add: bin_eq_iff bin_nth_ops) + +lemma int_xor_assoc: "(x XOR y) XOR z = x XOR (y XOR z)" + for x y z :: int + by (auto simp add: bin_eq_iff bin_nth_ops) + +lemmas bbw_assocs = int_and_assoc int_or_assoc int_xor_assoc + +(* BH: Why are these declared as simp rules??? *) +lemma bbw_lcs [simp]: + "y AND (x AND z) = x AND (y AND z)" + "y OR (x OR z) = x OR (y OR z)" + "y XOR (x XOR z) = x XOR (y XOR z)" + for x y :: int + by (auto simp add: bin_eq_iff bin_nth_ops) + +lemma bbw_not_dist: + "NOT (x OR y) = (NOT x) AND (NOT y)" + "NOT (x AND y) = (NOT x) OR (NOT y)" + for x y :: int + by (auto simp add: bin_eq_iff bin_nth_ops) + +lemma bbw_oa_dist: "(x AND y) OR z = (x OR z) AND (y OR z)" + for x y z :: int + by (auto simp add: bin_eq_iff bin_nth_ops) + +lemma bbw_ao_dist: "(x OR y) AND z = (x AND z) OR (y AND z)" + for x y z :: int + by (auto simp add: bin_eq_iff bin_nth_ops) + +(* +Why were these declared simp??? +declare bin_ops_comm [simp] bbw_assocs [simp] +*) + + +subsubsection \Simplification with numerals\ + +text \Cases for \0\ and \-1\ are already covered by other simp rules.\ + +lemma bin_rest_neg_numeral_BitM [simp]: + "bin_rest (- numeral (Num.BitM w)) = - numeral w" + by simp + +lemma bin_last_neg_numeral_BitM [simp]: + "bin_last (- numeral (Num.BitM w))" + by simp + + +subsubsection \Interactions with arithmetic\ + +lemma le_int_or: "bin_sign y = 0 \ x \ x OR y" + for x y :: int + by (simp add: bin_sign_def or_greater_eq split: if_splits) + +lemmas int_and_le = + xtrans(3) [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] + +text \Interaction between bit-wise and arithmetic: good example of \bin_induction\.\ +lemma bin_add_not: "x + NOT x = (-1::int)" + by (simp add: not_int_def) + +lemma AND_mod: "x AND (2 ^ n - 1) = x mod 2 ^ n" + for x :: int + by (simp flip: take_bit_eq_mod add: take_bit_eq_mask mask_eq_exp_minus_1) + + +subsubsection \Truncating results of bit-wise operations\ + +lemma bin_trunc_ao: + "bintrunc n x AND bintrunc n y = bintrunc n (x AND y)" + "bintrunc n x OR bintrunc n y = bintrunc n (x OR y)" + by simp_all + +lemma bin_trunc_xor: "bintrunc n (bintrunc n x XOR bintrunc n y) = bintrunc n (x XOR y)" + by simp + +lemma bin_trunc_not: "bintrunc n (NOT (bintrunc n x)) = bintrunc n (NOT x)" + by (fact take_bit_not_take_bit) + +text \Want theorems of the form of \bin_trunc_xor\.\ +lemma bintr_bintr_i: "x = bintrunc n y \ bintrunc n x = bintrunc n y" + by auto + +lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i] +lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i] + + +subsubsection \More lemmas\ + +lemma not_int_cmp_0 [simp]: + fixes i :: int shows + "0 < NOT i \ i < -1" + "0 \ NOT i \ i < 0" + "NOT i < 0 \ i \ 0" + "NOT i \ 0 \ i \ -1" +by(simp_all add: int_not_def) arith+ + +lemma bbw_ao_dist2: "(x :: int) AND (y OR z) = x AND y OR x AND z" + by (fact bit.conj_disj_distrib) + +lemmas int_and_ac = bbw_lcs(1) int_and_comm int_and_assoc + +lemma int_nand_same [simp]: fixes x :: int shows "x AND NOT x = 0" + by simp + +lemma int_nand_same_middle: fixes x :: int shows "x AND y AND NOT x = 0" + by (simp add: bit_eq_iff bit_and_iff bit_not_iff) + +lemma and_xor_dist: fixes x :: int shows + "x AND (y XOR z) = (x AND y) XOR (x AND z)" + by (fact bit.conj_xor_distrib) + +lemma int_and_lt0 [simp]: + \x AND y < 0 \ x < 0 \ y < 0\ for x y :: int + by (fact and_negative_int_iff) + +lemma int_and_ge0 [simp]: + \x AND y \ 0 \ x \ 0 \ y \ 0\ for x y :: int + by (fact and_nonnegative_int_iff) + +lemma int_and_1: fixes x :: int shows "x AND 1 = x mod 2" + by (fact and_one_eq) + +lemma int_1_and: fixes x :: int shows "1 AND x = x mod 2" + by (fact one_and_eq) + +lemma int_or_lt0 [simp]: + \x OR y < 0 \ x < 0 \ y < 0\ for x y :: int + by (fact or_negative_int_iff) + +lemma int_or_ge0 [simp]: + \x OR y \ 0 \ x \ 0 \ y \ 0\ for x y :: int + by (fact or_nonnegative_int_iff) + +lemma int_xor_lt0 [simp]: + \x XOR y < 0 \ (x < 0) \ (y < 0)\ for x y :: int + by (fact xor_negative_int_iff) + +lemma int_xor_ge0 [simp]: + \x XOR y \ 0 \ (x \ 0 \ y \ 0)\ for x y :: int + by (fact xor_nonnegative_int_iff) + +lemma even_conv_AND: + \even i \ i AND 1 = 0\ for i :: int + by (simp add: and_one_eq mod2_eq_if) + +lemma bin_last_conv_AND: + "bin_last i \ i AND 1 \ 0" + by (simp add: and_one_eq mod2_eq_if) + +lemma bitval_bin_last: + "of_bool (bin_last i) = i AND 1" + by (simp add: and_one_eq mod2_eq_if) + +lemma bin_sign_and: + "bin_sign (i AND j) = - (bin_sign i * bin_sign j)" +by(simp add: bin_sign_def) + +lemma int_not_neg_numeral: "NOT (- numeral n) = (Num.sub n num.One :: int)" +by(simp add: int_not_def) + +lemma int_neg_numeral_pOne_conv_not: "- numeral (n + num.One) = (NOT (numeral n) :: int)" +by(simp add: int_not_def) + + +subsection \Setting and clearing bits\ + +lemma int_shiftl_BIT: fixes x :: int + shows int_shiftl0 [simp]: "x << 0 = x" + and int_shiftl_Suc [simp]: "x << Suc n = 2 * (x << n)" + by (auto simp add: shiftl_int_def) + +lemma int_0_shiftl [simp]: "0 << n = (0 :: int)" +by(induct n) simp_all + +lemma bin_last_shiftl: "bin_last (x << n) \ n = 0 \ bin_last x" +by(cases n)(simp_all) + +lemma bin_rest_shiftl: "bin_rest (x << n) = (if n > 0 then x << (n - 1) else bin_rest x)" +by(cases n)(simp_all) + +lemma bin_nth_shiftl [simp]: "bin_nth (x << n) m \ n \ m \ bin_nth x (m - n)" + by (simp add: bit_push_bit_iff_int shiftl_eq_push_bit) + +lemma bin_last_shiftr: "odd (x >> n) \ x !! n" for x :: int + by (simp add: shiftr_eq_drop_bit bit_iff_odd_drop_bit) + +lemma bin_rest_shiftr [simp]: "bin_rest (x >> n) = x >> Suc n" + by (simp add: bit_eq_iff shiftr_eq_drop_bit drop_bit_Suc bit_drop_bit_eq drop_bit_half) + +lemma bin_nth_shiftr [simp]: "bin_nth (x >> n) m = bin_nth x (n + m)" + by (simp add: shiftr_eq_drop_bit bit_drop_bit_eq) + +lemma bin_nth_conv_AND: + fixes x :: int shows + "bin_nth x n \ x AND (1 << n) \ 0" + by (simp add: bit_eq_iff) + (auto simp add: shiftl_eq_push_bit bit_and_iff bit_push_bit_iff bit_exp_iff) + +lemma int_shiftl_numeral [simp]: + "(numeral w :: int) << numeral w' = numeral (num.Bit0 w) << pred_numeral w'" + "(- numeral w :: int) << numeral w' = - numeral (num.Bit0 w) << pred_numeral w'" +by(simp_all add: numeral_eq_Suc shiftl_int_def) + (metis add_One mult_inc semiring_norm(11) semiring_norm(13) semiring_norm(2) semiring_norm(6) semiring_norm(87))+ + +lemma int_shiftl_One_numeral [simp]: + "(1 :: int) << numeral w = 2 << pred_numeral w" + using int_shiftl_numeral [of Num.One w] by simp + +lemma shiftl_ge_0 [simp]: fixes i :: int shows "i << n \ 0 \ i \ 0" +by(induct n) simp_all + +lemma shiftl_lt_0 [simp]: fixes i :: int shows "i << n < 0 \ i < 0" +by (metis not_le shiftl_ge_0) + +lemma int_shiftl_test_bit: "(n << i :: int) !! m \ m \ i \ n !! (m - i)" + by simp + +lemma int_0shiftr [simp]: "(0 :: int) >> x = 0" +by(simp add: shiftr_int_def) + +lemma int_minus1_shiftr [simp]: "(-1 :: int) >> x = -1" +by(simp add: shiftr_int_def div_eq_minus1) + +lemma int_shiftr_ge_0 [simp]: fixes i :: int shows "i >> n \ 0 \ i \ 0" + by (simp add: shiftr_eq_drop_bit) + +lemma int_shiftr_lt_0 [simp]: fixes i :: int shows "i >> n < 0 \ i < 0" +by (metis int_shiftr_ge_0 not_less) + +lemma int_shiftr_numeral [simp]: + "(1 :: int) >> numeral w' = 0" + "(numeral num.One :: int) >> numeral w' = 0" + "(numeral (num.Bit0 w) :: int) >> numeral w' = numeral w >> pred_numeral w'" + "(numeral (num.Bit1 w) :: int) >> numeral w' = numeral w >> pred_numeral w'" + "(- numeral (num.Bit0 w) :: int) >> numeral w' = - numeral w >> pred_numeral w'" + "(- numeral (num.Bit1 w) :: int) >> numeral w' = - numeral (Num.inc w) >> pred_numeral w'" + by (simp_all add: shiftr_eq_drop_bit numeral_eq_Suc add_One drop_bit_Suc) + +lemma int_shiftr_numeral_Suc0 [simp]: + "(1 :: int) >> Suc 0 = 0" + "(numeral num.One :: int) >> Suc 0 = 0" + "(numeral (num.Bit0 w) :: int) >> Suc 0 = numeral w" + "(numeral (num.Bit1 w) :: int) >> Suc 0 = numeral w" + "(- numeral (num.Bit0 w) :: int) >> Suc 0 = - numeral w" + "(- numeral (num.Bit1 w) :: int) >> Suc 0 = - numeral (Num.inc w)" + by (simp_all add: shiftr_eq_drop_bit drop_bit_Suc add_One) + +lemma bin_nth_minus_p2: + assumes sign: "bin_sign x = 0" + and y: "y = 1 << n" + and m: "m < n" + and x: "x < y" + shows "bin_nth (x - y) m = bin_nth x m" +proof - + from sign y x have \x \ 0\ and \y = 2 ^ n\ and \x < 2 ^ n\ + by (simp_all add: bin_sign_def shiftl_eq_push_bit push_bit_eq_mult split: if_splits) + from \0 \ x\ \x < 2 ^ n\ \m < n\ have \bit x m \ bit (x - 2 ^ n) m\ + proof (induction m arbitrary: x n) + case 0 + then show ?case + by simp + next + case (Suc m) + moreover define q where \q = n - 1\ + ultimately have n: \n = Suc q\ + by simp + have \(x - 2 ^ Suc q) div 2 = x div 2 - 2 ^ q\ + by simp + moreover from Suc.IH [of \x div 2\ q] Suc.prems + have \bit (x div 2) m \ bit (x div 2 - 2 ^ q) m\ + by (simp add: n) + ultimately show ?case + by (simp add: bit_Suc n) + qed + with \y = 2 ^ n\ show ?thesis + by simp +qed + +lemma bin_clr_conv_NAND: + "bin_sc n False i = i AND NOT (1 << n)" + by (induct n arbitrary: i) (rule bin_rl_eqI; simp)+ + +lemma bin_set_conv_OR: + "bin_sc n True i = i OR (1 << n)" + by (induct n arbitrary: i) (rule bin_rl_eqI; simp)+ + + +subsection \More lemmas on words\ + +lemma word_rcat_eq: + \word_rcat ws = word_of_int (bin_rcat (LENGTH('a::len)) (map uint ws))\ + for ws :: \'a::len word list\ + apply (simp add: word_rcat_def bin_rcat_def rev_map) + apply transfer + apply (simp add: horner_sum_foldr foldr_map comp_def) + done + +lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" + by (simp add: sign_Pls_ge_0) + +lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or + +\ \following definitions require both arithmetic and bit-wise word operations\ + +\ \to get \word_no_log_defs\ from \word_log_defs\, using \bin_log_bintrs\\ +lemmas wils1 = bin_log_bintrs [THEN word_of_int_eq_iff [THEN iffD2], + folded uint_word_of_int_eq, THEN eq_reflection] + +\ \the binary operations only\ (* BH: why is this needed? *) +lemmas word_log_binary_defs = + word_and_def word_or_def word_xor_def + +lemma setBit_no [simp]: "setBit (numeral bin) n = word_of_int (bin_sc n True (numeral bin))" + by transfer (simp add: bin_sc_eq) + +lemma clearBit_no [simp]: + "clearBit (numeral bin) n = word_of_int (bin_sc n False (numeral bin))" + by transfer (simp add: bin_sc_eq) + +lemma eq_mod_iff: "0 < n \ b = b mod n \ 0 \ b \ b < n" + for b n :: int + by auto (metis pos_mod_conj)+ + +lemma split_uint_lem: "bin_split n (uint w) = (a, b) \ + a = take_bit (LENGTH('a) - n) a \ b = take_bit (LENGTH('a)) b" + for w :: "'a::len word" + by transfer (simp add: drop_bit_take_bit ac_simps) + +\ \limited hom result\ +lemma word_cat_hom: + "LENGTH('a::len) \ LENGTH('b::len) + LENGTH('c::len) \ + (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = + word_of_int (bin_cat w (size b) (uint b))" + by transfer (simp add: take_bit_concat_bit_eq) + +lemma bintrunc_shiftl: + "take_bit n (m << i) = take_bit (n - i) m << i" + for m :: int + by (rule bit_eqI) (auto simp add: bit_take_bit_iff) + +lemma uint_shiftl: + "uint (n << i) = take_bit (size n) (uint n << i)" + by transfer (simp add: push_bit_take_bit shiftl_eq_push_bit) + + +code_identifier + code_module Bits_Int \ + (SML) Bit_Operations and (OCaml) Bit_Operations and (Haskell) Bit_Operations and (Scala) Bit_Operations + +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,527 @@ (* Authors: Thomas Sewell, NICTA and Sascha Boehme, TU Muenchen *) theory Bitwise - imports "HOL-Word.Word" "HOL-Word.Misc_Arithmetic" "HOL-Word.Misc_msb" + imports "HOL-Library.Word" Misc_Arithmetic Misc_msb 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\ 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/Bitwise_Signed.thy b/thys/Word_Lib/Bitwise_Signed.thy --- a/thys/Word_Lib/Bitwise_Signed.thy +++ b/thys/Word_Lib/Bitwise_Signed.thy @@ -1,30 +1,30 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Bitwise tactic for Signed Words" theory Bitwise_Signed imports - "HOL-Word.Word" + "HOL-Library.Word" Bitwise Signed_Words begin ML \fun bw_tac_signed ctxt = let val (ss, sss) = Word_Bitwise_Tac.expand_word_eq_sss val sss = nth_map 2 (fn ss => put_simpset ss ctxt addsimps @{thms len_signed} |> simpset_of) 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; \ method_setup word_bitwise_signed = \Scan.succeed (fn ctxt => Method.SIMPLE_METHOD (bw_tac_signed ctxt 1))\ "decomposer for word equalities and inequalities into bit propositions" end diff --git a/thys/Word_Lib/Guide.thy b/thys/Word_Lib/Guide.thy --- a/thys/Word_Lib/Guide.thy +++ b/thys/Word_Lib/Guide.thy @@ -1,312 +1,312 @@ (*<*) theory Guide imports Main Word_Lemmas Word_Lemmas_32 Word_Lemmas_64 begin hide_const (open) Misc_set_bit.set_bit (*>*) section \A short overview over bit operations and word types\ subsection \Basic theories and key ideas\ text \ When formalizing bit operations, it is tempting to represent bit values as explicit lists over a binary type. This however is a bad idea, mainly due to the inherent ambiguities in representation concerning repeating leading bits. Hence this approach avoids such explicit lists altogether following an algebraic path: \<^item> Bit values are represented by numeric types: idealized unbounded bit values can be represented by type \<^typ>\int\, bounded bit values by quotient types over \<^typ>\int\, aka \<^typ>\'a word\. \<^item> (A special case are idealized unbounded bit values ending in @{term [source] 0} which can be represented by type \<^typ>\nat\ but only support a restricted set of operations). The most fundamental ideas are developed in theory \<^theory>\HOL.Parity\ (which is part of \<^theory>\Main\): \<^item> Multiplication by \<^term>\2 :: int\ is a bit shift to the left and \<^item> Division by \<^term>\2 :: int\ is a bit shift to the right. \<^item> Concerning bounded bit values, iterated shifts to the left may result in eliminating all bits by shifting them all beyond the boundary. The property \<^prop>\(2 :: int) ^ n \ 0\ represents that \<^term>\n\ is \<^emph>\not\ beyond that boundary. \<^item> The projection on a single bit is then @{thm [mode=iff] bit_iff_odd [where ?'a = int, no_vars]}. \<^item> This leads to the most fundamental properties of bit values: \<^item> Equality rule: @{thm [display, mode=iff] bit_eq_iff [where ?'a = int, no_vars]} \<^item> Induction rule: @{thm [display, mode=iff] bits_induct [where ?'a = int, no_vars]} On top of this, the following generic operations are provided after import of theory \<^theory>\HOL-Library.Bit_Operations\: \<^item> Singleton \<^term>\n\th bit: \<^term>\(2 :: int) ^ n\ \<^item> Bit mask upto bit \<^term>\n\: @{thm mask_eq_exp_minus_1 [where ?'a = int, no_vars]} \<^item> Left shift: @{thm push_bit_eq_mult [where ?'a = int, no_vars]} \<^item> Right shift: @{thm drop_bit_eq_div [where ?'a = int, no_vars]} \<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]} \<^item> Negation: @{thm [mode=iff] bit_not_iff [where ?'a = int, no_vars]} \<^item> And: @{thm [mode=iff] bit_and_iff [where ?'a = int, no_vars]} \<^item> Or: @{thm [mode=iff] bit_or_iff [where ?'a = int, no_vars]} \<^item> Xor: @{thm [mode=iff] bit_xor_iff [where ?'a = int, no_vars]} \<^item> Set a single bit: @{thm set_bit_def [where ?'a = int, no_vars]} \<^item> Unset a single bit: @{thm unset_bit_def [where ?'a = int, no_vars]} \<^item> Flip a single bit: @{thm flip_bit_def [where ?'a = int, no_vars]} \<^item> Signed truncation, or modulus centered around \<^term>\0::int\: @{thm [display] signed_take_bit_def [where ?'a = int, no_vars]} \<^item> (Bounded) conversion from and to a list of bits: @{thm [display] horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]} - Proper word types are introduced in theory \<^theory>\HOL-Word.Word\, with + Proper word types are introduced in theory \<^theory>\HOL-Library.Word\, with the following specific operations: \<^item> Standard arithmetic: @{term \(+) :: 'a::len word \ 'a word \ 'a word\}, @{term \uminus :: 'a::len word \ 'a word\}, @{term \(-) :: 'a::len word \ 'a word \ 'a word\}, @{term \(*) :: 'a::len word \ 'a word \ 'a word\}, @{term \0 :: 'a::len word\}, @{term \1 :: 'a::len word\}, numerals etc. \<^item> Standard bit operations: see above. \<^item> Conversion with unsigned interpretation of words: \<^item> @{term [source] \unsigned :: 'a::len word \ 'b::semiring_1\} \<^item> Important special cases as abbreviations: \<^item> @{term [source] \unat :: 'a::len word \ nat\} \<^item> @{term [source] \uint :: 'a::len word \ int\} \<^item> @{term [source] \ucast :: 'a::len word \ 'b::len word\} \<^item> Conversion with signed interpretation of words: \<^item> @{term [source] \signed :: 'a::len word \ 'b::ring_1\} \<^item> Important special cases as abbreviations: \<^item> @{term [source] \sint :: 'a::len word \ int\} \<^item> @{term [source] \scast :: 'a::len word \ 'b::len word\} \<^item> Operations with unsigned interpretation of words: \<^item> @{thm [mode=iff] word_le_nat_alt [no_vars]} \<^item> @{thm [mode=iff] word_less_nat_alt [no_vars]} \<^item> @{thm unat_div_distrib [no_vars]} \<^item> @{thm unat_drop_bit_eq [no_vars]} \<^item> @{thm unat_mod_distrib [no_vars]} \<^item> @{thm [mode=iff] udvd_iff_dvd [no_vars]} \<^item> Operations with signed interpretation of words: \<^item> @{thm [mode=iff] word_sle_eq [no_vars]} \<^item> @{thm [mode=iff] word_sless_alt [no_vars]} \<^item> @{thm sint_signed_drop_bit_eq [no_vars]} \<^item> Rotation and reversal: \<^item> @{term [source] \word_rotl :: nat \ 'a::len word \ 'a word\} \<^item> @{term [source] \word_rotr :: nat \ 'a::len word \ 'a word\} \<^item> @{term [source] \word_roti :: int \ 'a::len word \ 'a word\} \<^item> @{term [source] \word_reverse :: 'a::len word \ 'a word\} \<^item> Concatenation: @{term [source, display] \word_cat :: 'a::len word \ 'b::len word \ 'c::len word\} For proofs about words the following default strategies are applicable: \<^item> Using bit extensionality (facts \<^text>\bit_eq_iff\, \<^text>\bit_eqI\). \<^item> Using the @{method transfer} method. \ subsection \More library theories\ text \ Note: currently, the theories listed here are hardly separate entities since they import each other in various ways. Always inspect them to understand what you pull in if you want to import one. \<^descr>[Syntax] \<^descr>[\<^theory>\Word_Lib.Hex_Words\] Printing word numerals as hexadecimal numerals. \<^descr>[\<^theory>\Word_Lib.Word_Type_Syntax\] Pretty type-sensitive syntax for cast operations. \<^descr>[\<^theory>\Word_Lib.Word_Syntax\] Specific ASCII syntax for prominent bit operations on word. \<^descr>[Proof tools] \<^descr>[\<^theory>\Word_Lib.Norm_Words\] Rewriting word numerals to normal forms. \<^descr>[\<^theory>\Word_Lib.Bitwise\] Method @{method word_bitwise} decomposes word equalities and inequalities into bit propositions. \<^descr>[\<^theory>\Word_Lib.Word_EqI\] Method @{method word_eqI_solve} decomposes word equalities and inequalities into bit propositions. \<^descr>[Operations] - \<^descr>[\<^theory>\HOL-Word.Misc_lsb\] + \<^descr>[\<^theory>\Word_Lib.Misc_lsb\] The least significant bit as an alias: @{thm [mode=iff] lsb_odd [where ?'a = int, no_vars]} - \<^descr>[\<^theory>\HOL-Word.Misc_msb\] + \<^descr>[\<^theory>\Word_Lib.Misc_msb\] The most significant bit: \<^item> @{thm [mode=iff] msb_int_def [of k]} \<^item> @{thm [mode=iff] word_msb_sint [no_vars]} \<^item> @{thm [mode=iff] msb_word_iff_sless_0 [no_vars]} \<^item> @{thm [mode=iff] msb_word_iff_bit [no_vars]} - \<^descr>[\<^theory>\HOL-Word.Traditional_Syntax\] + \<^descr>[\<^theory>\Word_Lib.Traditional_Syntax\] Clones of existing operations decorated with traditional syntax: \<^item> @{thm test_bit_eq_bit [no_vars]} \<^item> @{thm shiftl_eq_push_bit [no_vars]} \<^item> @{thm shiftr_eq_drop_bit [no_vars]} \<^item> @{thm sshiftr_eq [no_vars]} \<^descr>[\<^theory>\Word_Lib.Word_Lib\] Various operations on word, particularly: \<^item> @{term [source] \(sdiv) :: 'a::len word \ 'a word \ 'a word\} \<^item> @{term [source] \(smod) :: 'a::len word \ 'a word \ 'a word\} \<^descr>[\<^theory>\Word_Lib.Aligned\] \ \<^item> @{thm [mode=iff] is_aligned_iff_udvd [no_vars]} \<^descr>[\<^theory>\Word_Lib.Word_Next\] \ \<^item> @{thm word_next_unfold [no_vars]} \<^item> @{thm word_prev_unfold [no_vars]} \<^descr>[Types] \<^descr>[\<^theory>\Word_Lib.Signed_Words\] Formal tagging of word types with a \<^text>\signed\ marker. \<^descr>[Mechanisms] \<^descr>[\<^theory>\Word_Lib.Word_Enum\] More on explicit enumeration of word types. \<^descr>[Lemmas] Collections of lemmas: \<^descr>[\<^theory>\Word_Lib.Word_Lemmas\] generic. \<^descr>[\<^theory>\Word_Lib.Word_Lemmas_32\] for 32-bit words. \<^descr>[\<^theory>\Word_Lib.Word_Lemmas_64\] for 64-bit words. \ subsection \More library sessions\ text \ \<^descr>[\<^text>\Native_Word\] Makes machine words and machine arithmetic available for code generation. It provides a common abstraction that hides the differences between the different target languages. The code generator maps these operations to the APIs of the target languages. \ subsection \Legacy theories\ text \ The following theories contain material which has been factored out since it is not recommended to use it in new applications, mostly because matters can be expressed succinctly using already existing operations. This section gives some indication how to migrate away from those theories. However theorem coverage may still be terse in some cases. - \<^descr>[\<^theory>\HOL-Word.Misc_set_bit\] + \<^descr>[\<^theory>\Word_Lib.Misc_set_bit\] Kind of an alias: @{thm set_bit_eq [no_vars]} - \<^descr>[\<^theory>\HOL-Word.Misc_Typedef\] + \<^descr>[\<^theory>\Word_Lib.Misc_Typedef\] A low-level extension to HOL typedef providing conversions along type morphisms. The @{method transfer} method seems to be sufficient for most applications though. - \<^descr>[\<^theory>\HOL-Word.Bit_Comprehension\] + \<^descr>[\<^theory>\Word_Lib.Bit_Comprehension\] Comprehension syntax for bit values over predicates \<^typ>\nat \ bool\. For \<^typ>\'a::len word\, straightforward alternatives exist; difficult to handle for \<^typ>\int\. - \<^descr>[\<^theory>\HOL-Word.Reversed_Bit_Lists\] + \<^descr>[\<^theory>\Word_Lib.Reversed_Bit_Lists\] Representation of bit values as explicit list in \<^emph>\reversed\ order. This should rarely be necessary: the \<^const>\bit\ projection should be sufficient in most cases. In case explicit lists are needed, existing operations can be used: @{thm [display] horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]} \ (*<*) end (*>*) diff --git a/thys/Word_Lib/Hex_Words.thy b/thys/Word_Lib/Hex_Words.thy --- a/thys/Word_Lib/Hex_Words.thy +++ b/thys/Word_Lib/Hex_Words.thy @@ -1,48 +1,48 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Print Words in Hex" theory Hex_Words imports - "HOL-Word.Word" + "HOL-Library.Word" begin text \Print words in hex.\ (* mostly clagged from Num.thy *) typed_print_translation \ let fun dest_num (Const (@{const_syntax Num.Bit0}, _) $ n) = 2 * dest_num n | dest_num (Const (@{const_syntax Num.Bit1}, _) $ n) = 2 * dest_num n + 1 | dest_num (Const (@{const_syntax Num.One}, _)) = 1; fun dest_bin_hex_str tm = let val num = dest_num tm; val pre = if num < 10 then "" else "0x" in pre ^ (Int.fmt StringCvt.HEX num) end; fun num_tr' sign ctxt T [n] = let val k = dest_bin_hex_str n; val t' = Syntax.const @{syntax_const "_Numeral"} $ Syntax.free (sign ^ k); in case T of Type (@{type_name fun}, [_, T' as Type("Word.word",_)]) => if not (Config.get ctxt show_types) andalso can Term.dest_Type T' then t' else Syntax.const @{syntax_const "_constrain"} $ t' $ Syntax_Phases.term_of_typ ctxt T' | T' => if T' = dummyT then t' else raise Match end; in [(@{const_syntax numeral}, num_tr' "")] end \ end diff --git a/thys/Word_Lib/Misc_Arithmetic.thy b/thys/Word_Lib/Misc_Arithmetic.thy new file mode 100644 --- /dev/null +++ b/thys/Word_Lib/Misc_Arithmetic.thy @@ -0,0 +1,462 @@ +(* Title: HOL/Word/Misc_Arithmetic.thy *) + +section \Miscellaneous lemmas, mostly for arithmetic\ + +theory Misc_Arithmetic + imports Main Bits_Int +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 + +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) + +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 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] + +end diff --git a/thys/Word_Lib/Misc_Auxiliary.thy b/thys/Word_Lib/Misc_Auxiliary.thy new file mode 100644 --- /dev/null +++ b/thys/Word_Lib/Misc_Auxiliary.thy @@ -0,0 +1,25 @@ +(* Title: HOL/Word/Misc_Auxiliary.thy + Author: Jeremy Dawson, NICTA +*) + +section \Generic auxiliary\ + +theory Misc_Auxiliary + imports Main +begin + +subsection \Lemmas on list operations\ + +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 + +end diff --git a/thys/Word_Lib/Misc_Typedef.thy b/thys/Word_Lib/Misc_Typedef.thy new file mode 100644 --- /dev/null +++ b/thys/Word_Lib/Misc_Typedef.thy @@ -0,0 +1,356 @@ +(* + Author: Jeremy Dawson and Gerwin Klein, NICTA + + Consequences of type definition theorems, and of extended type definition. +*) + +section \Type Definition Theorems\ + +theory Misc_Typedef + imports Main "HOL-Library.Word" Bit_Comprehension Bits_Int +begin + +subsection "More lemmas about normal type definitions" + +lemma tdD1: "type_definition Rep Abs A \ \x. Rep x \ A" + and tdD2: "type_definition Rep Abs A \ \x. Abs (Rep x) = x" + and tdD3: "type_definition Rep Abs A \ \y. y \ A \ Rep (Abs y) = y" + by (auto simp: type_definition_def) + +lemma td_nat_int: "type_definition int nat (Collect ((\) 0))" + unfolding type_definition_def by auto + +context type_definition +begin + +declare Rep [iff] Rep_inverse [simp] Rep_inject [simp] + +lemma Abs_eqD: "Abs x = Abs y \ x \ A \ y \ A \ x = y" + by (simp add: Abs_inject) + +lemma Abs_inverse': "r \ A \ Abs r = a \ Rep a = r" + by (safe elim!: Abs_inverse) + +lemma Rep_comp_inverse: "Rep \ f = g \ Abs \ g = f" + using Rep_inverse by auto + +lemma Rep_eqD [elim!]: "Rep x = Rep y \ x = y" + by simp + +lemma Rep_inverse': "Rep a = r \ Abs r = a" + by (safe intro!: Rep_inverse) + +lemma comp_Abs_inverse: "f \ Abs = g \ g \ Rep = f" + using Rep_inverse by auto + +lemma set_Rep: "A = range Rep" +proof (rule set_eqI) + show "x \ A \ x \ range Rep" for x + by (auto dest: Abs_inverse [of x, symmetric]) +qed + +lemma set_Rep_Abs: "A = range (Rep \ Abs)" +proof (rule set_eqI) + show "x \ A \ x \ range (Rep \ Abs)" for x + by (auto dest: Abs_inverse [of x, symmetric]) +qed + +lemma Abs_inj_on: "inj_on Abs A" + unfolding inj_on_def + by (auto dest: Abs_inject [THEN iffD1]) + +lemma image: "Abs ` A = UNIV" + by (fact Abs_image) + +lemmas td_thm = type_definition_axioms + +lemma fns1: "Rep \ fa = fr \ Rep \ fa \ Abs = Abs \ fr \ Abs \ fr \ Rep = fa" + by (auto dest: Rep_comp_inverse elim: comp_Abs_inverse simp: o_assoc) + +lemmas fns1a = disjI1 [THEN fns1] +lemmas fns1b = disjI2 [THEN fns1] + +lemma fns4: "Rep \ fa \ Abs = fr \ Rep \ fa = fr \ Rep \ fa \ Abs = Abs \ fr" + by auto + +end + +interpretation nat_int: type_definition int nat "Collect ((\) 0)" + by (rule td_nat_int) + +declare + nat_int.Rep_cases [cases del] + nat_int.Abs_cases [cases del] + nat_int.Rep_induct [induct del] + nat_int.Abs_induct [induct del] + + +subsection "Extended form of type definition predicate" + +lemma td_conds: + "norm \ norm = norm \ + fr \ norm = norm \ fr \ norm \ fr \ norm = fr \ norm \ norm \ fr \ norm = norm \ fr" + apply safe + apply (simp_all add: comp_assoc) + apply (simp_all add: o_assoc) + done + +lemma fn_comm_power: "fa \ tr = tr \ fr \ fa ^^ n \ tr = tr \ fr ^^ n" + apply (rule ext) + apply (induct n) + apply (auto dest: fun_cong) + done + +lemmas fn_comm_power' = + ext [THEN fn_comm_power, THEN fun_cong, unfolded o_def] + + +locale td_ext = type_definition + + fixes norm + assumes eq_norm: "\x. Rep (Abs x) = norm x" +begin + +lemma Abs_norm [simp]: "Abs (norm x) = Abs x" + using eq_norm [of x] by (auto elim: Rep_inverse') + +lemma td_th: "g \ Abs = f \ f (Rep x) = g x" + by (drule comp_Abs_inverse [symmetric]) simp + +lemma eq_norm': "Rep \ Abs = norm" + by (auto simp: eq_norm) + +lemma norm_Rep [simp]: "norm (Rep x) = Rep x" + by (auto simp: eq_norm' intro: td_th) + +lemmas td = td_thm + +lemma set_iff_norm: "w \ A \ w = norm w" + by (auto simp: set_Rep_Abs eq_norm' eq_norm [symmetric]) + +lemma inverse_norm: "Abs n = w \ Rep w = norm n" + apply (rule iffI) + apply (clarsimp simp add: eq_norm) + apply (simp add: eq_norm' [symmetric]) + done + +lemma norm_eq_iff: "norm x = norm y \ Abs x = Abs y" + by (simp add: eq_norm' [symmetric]) + +lemma norm_comps: + "Abs \ norm = Abs" + "norm \ Rep = Rep" + "norm \ norm = norm" + by (auto simp: eq_norm' [symmetric] o_def) + +lemmas norm_norm [simp] = norm_comps + +lemma fns5: "Rep \ fa \ Abs = fr \ fr \ norm = fr \ norm \ fr = fr" + by (fold eq_norm') auto + +text \ + following give conditions for converses to \td_fns1\ + \<^item> the condition \norm \ fr \ norm = fr \ norm\ says that + \fr\ takes normalised arguments to normalised results + \<^item> \norm \ fr \ norm = norm \ fr\ says that \fr\ + takes norm-equivalent arguments to norm-equivalent results + \<^item> \fr \ norm = fr\ says that \fr\ + takes norm-equivalent arguments to the same result + \<^item> \norm \ fr = fr\ says that \fr\ takes any argument to a normalised result +\ +lemma fns2: "Abs \ fr \ Rep = fa \ norm \ fr \ norm = fr \ norm \ Rep \ fa = fr \ Rep" + apply (fold eq_norm') + apply safe + prefer 2 + apply (simp add: o_assoc) + apply (rule ext) + apply (drule_tac x="Rep x" in fun_cong) + apply auto + done + +lemma fns3: "Abs \ fr \ Rep = fa \ norm \ fr \ norm = norm \ fr \ fa \ Abs = Abs \ fr" + apply (fold eq_norm') + apply safe + prefer 2 + apply (simp add: comp_assoc) + apply (rule ext) + apply (drule_tac f="a \ b" for a b in fun_cong) + apply simp + done + +lemma fns: "fr \ norm = norm \ fr \ fa \ Abs = Abs \ fr \ Rep \ fa = fr \ Rep" + apply safe + apply (frule fns1b) + prefer 2 + apply (frule fns1a) + apply (rule fns3 [THEN iffD1]) + prefer 3 + apply (rule fns2 [THEN iffD1]) + apply (simp_all add: comp_assoc) + apply (simp_all add: o_assoc) + done + +lemma range_norm: "range (Rep \ Abs) = A" + by (simp add: set_Rep_Abs) + +end + +lemmas td_ext_def' = + td_ext_def [unfolded type_definition_def td_ext_axioms_def] + + +subsection \Type-definition locale instantiations\ + +definition uints :: "nat \ int set" + \ \the sets of integers representing the words\ + where "uints n = range (take_bit n)" + +definition sints :: "nat \ int set" + where "sints n = range (signed_take_bit (n - 1))" + +lemma uints_num: "uints n = {i. 0 \ i \ i < 2 ^ n}" + by (simp add: uints_def range_bintrunc) + +lemma sints_num: "sints n = {i. - (2 ^ (n - 1)) \ i \ i < 2 ^ (n - 1)}" + by (simp add: sints_def range_sbintrunc) + +definition unats :: "nat \ nat set" + where "unats n = {i. i < 2 ^ n}" + +\ \naturals\ +lemma uints_unats: "uints n = int ` unats n" + apply (unfold unats_def uints_num) + apply safe + apply (rule_tac image_eqI) + apply (erule_tac nat_0_le [symmetric]) + by auto + +lemma unats_uints: "unats n = nat ` uints n" + by (auto simp: uints_unats image_iff) + +lemma td_ext_uint: + "td_ext (uint :: 'a word \ int) word_of_int (uints (LENGTH('a::len))) + (\w::int. w mod 2 ^ LENGTH('a))" + apply (unfold td_ext_def') + apply transfer + apply (simp add: uints_num take_bit_eq_mod) + done + +interpretation word_uint: + td_ext + "uint::'a::len word \ int" + word_of_int + "uints (LENGTH('a::len))" + "\w. w mod 2 ^ LENGTH('a::len)" + by (fact td_ext_uint) + +lemmas td_uint = word_uint.td_thm +lemmas int_word_uint = word_uint.eq_norm + +lemma td_ext_ubin: + "td_ext (uint :: 'a word \ int) word_of_int (uints (LENGTH('a::len))) + (take_bit (LENGTH('a)))" + apply standard + apply transfer + apply simp + done + +interpretation word_ubin: + td_ext + "uint::'a::len word \ int" + word_of_int + "uints (LENGTH('a::len))" + "take_bit (LENGTH('a::len))" + by (fact td_ext_ubin) + +lemma td_ext_unat [OF refl]: + "n = LENGTH('a::len) \ + td_ext (unat :: 'a word \ nat) of_nat (unats n) (\i. i mod 2 ^ n)" + apply (standard; transfer) + apply (simp_all add: unats_def take_bit_of_nat take_bit_nat_eq_self_iff + flip: take_bit_eq_mod) + done + +lemmas unat_of_nat = td_ext_unat [THEN td_ext.eq_norm] + +interpretation word_unat: + td_ext + "unat::'a::len word \ nat" + of_nat + "unats (LENGTH('a::len))" + "\i. i mod 2 ^ LENGTH('a::len)" + by (rule td_ext_unat) + +lemmas td_unat = word_unat.td_thm + +lemmas unat_lt2p [iff] = word_unat.Rep [unfolded unats_def mem_Collect_eq] + +lemma unat_le: "y \ unat z \ y \ unats (LENGTH('a))" + for z :: "'a::len word" + apply (unfold unats_def) + apply clarsimp + apply (rule xtrans, rule unat_lt2p, assumption) + done + +lemma td_ext_sbin: + "td_ext (sint :: 'a word \ int) word_of_int (sints (LENGTH('a::len))) + (signed_take_bit (LENGTH('a) - 1))" + by (standard; transfer) (auto simp add: sints_def) + +lemma td_ext_sint: + "td_ext (sint :: 'a word \ int) word_of_int (sints (LENGTH('a::len))) + (\w. (w + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - + 2 ^ (LENGTH('a) - 1))" + using td_ext_sbin [where ?'a = 'a] by (simp add: no_sbintr_alt2) + +text \ + We do \sint\ before \sbin\, before \sint\ is the user version + and interpretations do not produce thm duplicates. I.e. + we get the name \word_sint.Rep_eqD\, but not \word_sbin.Req_eqD\, + because the latter is the same thm as the former. +\ +interpretation word_sint: + td_ext + "sint ::'a::len word \ int" + word_of_int + "sints (LENGTH('a::len))" + "\w. (w + 2^(LENGTH('a::len) - 1)) mod 2^LENGTH('a::len) - + 2 ^ (LENGTH('a::len) - 1)" + by (rule td_ext_sint) + +interpretation word_sbin: + td_ext + "sint ::'a::len word \ int" + word_of_int + "sints (LENGTH('a::len))" + "signed_take_bit (LENGTH('a::len) - 1)" + by (rule td_ext_sbin) + +lemmas int_word_sint = td_ext_sint [THEN td_ext.eq_norm] + +lemmas td_sint = word_sint.td + +lemma uints_mod: "uints n = range (\w. w mod 2 ^ n)" + by (fact uints_def [unfolded no_bintr_alt1]) + +lemmas uint_range' = word_uint.Rep [unfolded uints_num mem_Collect_eq] +lemmas sint_range' = word_sint.Rep [unfolded One_nat_def sints_num mem_Collect_eq] + +lemmas bintr_num = + word_ubin.norm_eq_iff [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b +lemmas sbintr_num = + word_sbin.norm_eq_iff [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b + +lemmas uint_div_alt = word_div_def [THEN trans [OF uint_cong int_word_uint]] +lemmas uint_mod_alt = word_mod_def [THEN trans [OF uint_cong int_word_uint]] + +interpretation test_bit: + td_ext + "(!!) :: 'a::len word \ nat \ bool" + set_bits + "{f. \i. f i \ i < LENGTH('a::len)}" + "(\h i. h i \ i < LENGTH('a::len))" + by standard (auto simp add: test_bit_word_eq bit_imp_le_length bit_set_bits_word_iff set_bits_bit_eq) + +lemmas td_nth = test_bit.td_thm + +end diff --git a/thys/Word_Lib/Misc_lsb.thy b/thys/Word_Lib/Misc_lsb.thy new file mode 100644 --- /dev/null +++ b/thys/Word_Lib/Misc_lsb.thy @@ -0,0 +1,89 @@ +(* Author: Jeremy Dawson, NICTA +*) + +section \Operation variant for the least significant bit\ + +theory Misc_lsb + imports + "HOL-Library.Word" + Reversed_Bit_Lists +begin + +class lsb = semiring_bits + + fixes lsb :: \'a \ bool\ + assumes lsb_odd: \lsb = odd\ + +instantiation int :: lsb +begin + +definition lsb_int :: \int \ bool\ + where \lsb i = i !! 0\ for i :: int + +instance + by standard (simp add: fun_eq_iff lsb_int_def) + +end + +lemma bin_last_conv_lsb: "bin_last = lsb" + by (simp add: lsb_odd) + +lemma int_lsb_numeral [simp]: + "lsb (0 :: int) = False" + "lsb (1 :: int) = True" + "lsb (Numeral1 :: int) = True" + "lsb (- 1 :: int) = True" + "lsb (- Numeral1 :: int) = True" + "lsb (numeral (num.Bit0 w) :: int) = False" + "lsb (numeral (num.Bit1 w) :: int) = True" + "lsb (- numeral (num.Bit0 w) :: int) = False" + "lsb (- numeral (num.Bit1 w) :: int) = True" + by (simp_all add: lsb_int_def) + +instantiation word :: (len) lsb +begin + +definition lsb_word :: \'a word \ bool\ + where word_lsb_def: \lsb a \ odd (uint a)\ for a :: \'a word\ + +instance + apply standard + apply (simp add: fun_eq_iff word_lsb_def) + apply transfer apply simp + done + +end + +lemma lsb_word_eq: + \lsb = (odd :: 'a word \ bool)\ for w :: \'a::len word\ + by (fact lsb_odd) + +lemma word_lsb_alt: "lsb w = test_bit w 0" + for w :: "'a::len word" + by (auto simp: word_test_bit_def word_lsb_def) + +lemma word_lsb_1_0 [simp]: "lsb (1::'a::len word) \ \ lsb (0::'b::len word)" + unfolding word_lsb_def by simp + +lemma word_lsb_last: + \lsb w \ last (to_bl w)\ + for w :: \'a::len word\ + using nth_to_bl [of \LENGTH('a) - Suc 0\ w] + by (simp add: lsb_odd last_conv_nth) + +lemma word_lsb_int: "lsb w \ uint w mod 2 = 1" + apply (simp add: lsb_odd flip: odd_iff_mod_2_eq_one) + apply transfer + apply simp + done + +lemmas word_ops_lsb = lsb0 [unfolded word_lsb_alt] + +lemma word_lsb_numeral [simp]: + "lsb (numeral bin :: 'a::len word) \ bin_last (numeral bin)" + unfolding word_lsb_alt test_bit_numeral by simp + +lemma word_lsb_neg_numeral [simp]: + "lsb (- numeral bin :: 'a::len word) \ bin_last (- numeral bin)" + by (simp add: word_lsb_alt) + +end diff --git a/thys/Word_Lib/Misc_msb.thy b/thys/Word_Lib/Misc_msb.thy new file mode 100644 --- /dev/null +++ b/thys/Word_Lib/Misc_msb.thy @@ -0,0 +1,137 @@ +(* Author: Jeremy Dawson, NICTA +*) + +section \Operation variant for the most significant bit\ + +theory Misc_msb + imports + "HOL-Library.Word" + Reversed_Bit_Lists +begin + +class msb = + fixes msb :: \'a \ bool\ + +instantiation int :: msb +begin + +definition \msb x \ x < 0\ for x :: int + +instance .. + +end + +lemma msb_conv_bin_sign: "msb x \ bin_sign x = -1" +by(simp add: bin_sign_def not_le msb_int_def) + +lemma msb_bin_rest [simp]: "msb (bin_rest x) = msb x" +by(simp add: msb_int_def) + +lemma int_msb_and [simp]: "msb ((x :: int) AND y) \ msb x \ msb y" +by(simp add: msb_int_def) + +lemma int_msb_or [simp]: "msb ((x :: int) OR y) \ msb x \ msb y" +by(simp add: msb_int_def) + +lemma int_msb_xor [simp]: "msb ((x :: int) XOR y) \ msb x \ msb y" +by(simp add: msb_int_def) + +lemma int_msb_not [simp]: "msb (NOT (x :: int)) \ \ msb x" +by(simp add: msb_int_def not_less) + +lemma msb_shiftl [simp]: "msb ((x :: int) << n) \ msb x" +by(simp add: msb_int_def) + +lemma msb_shiftr [simp]: "msb ((x :: int) >> r) \ msb x" +by(simp add: msb_int_def) + +lemma msb_bin_sc [simp]: "msb (bin_sc n b x) \ msb x" +by(simp add: msb_conv_bin_sign) + +lemma msb_0 [simp]: "msb (0 :: int) = False" +by(simp add: msb_int_def) + +lemma msb_1 [simp]: "msb (1 :: int) = False" +by(simp add: msb_int_def) + +lemma msb_numeral [simp]: + "msb (numeral n :: int) = False" + "msb (- numeral n :: int) = True" +by(simp_all add: msb_int_def) + +instantiation word :: (len) msb +begin + +definition msb_word :: \'a word \ bool\ + where \msb a \ bin_sign (sbintrunc (LENGTH('a) - 1) (uint a)) = - 1\ + +lemma msb_word_eq: + \msb w \ bit w (LENGTH('a) - 1)\ for w :: \'a::len word\ + by (simp add: msb_word_def bin_sign_lem bit_uint_iff) + +instance .. + +end + +lemma msb_word_iff_bit: + \msb w \ bit w (LENGTH('a) - Suc 0)\ + for w :: \'a::len word\ + by (simp add: msb_word_def bin_sign_def bit_uint_iff) + +lemma word_msb_def: + "msb a \ bin_sign (sint a) = - 1" + by (simp add: msb_word_def sint_uint) + +lemma word_msb_sint: "msb w \ sint w < 0" + by (simp add: msb_word_eq bit_last_iff) + +lemma msb_word_iff_sless_0: + \msb w \ w + by (simp add: word_msb_sint word_sless_alt) + +lemma msb_word_of_int: "msb (word_of_int x::'a::len word) = bin_nth x (LENGTH('a) - 1)" + by (simp add: word_msb_def word_sbin.eq_norm bin_sign_lem) + +lemma word_msb_numeral [simp]: + "msb (numeral w::'a::len word) = bin_nth (numeral w) (LENGTH('a) - 1)" + unfolding word_numeral_alt by (rule msb_word_of_int) + +lemma word_msb_neg_numeral [simp]: + "msb (- numeral w::'a::len word) = bin_nth (- numeral w) (LENGTH('a) - 1)" + unfolding word_neg_numeral_alt by (rule msb_word_of_int) + +lemma word_msb_0 [simp]: "\ msb (0::'a::len word)" + by (simp add: word_msb_def bin_sign_def sint_uint sbintrunc_eq_take_bit) + +lemma word_msb_1 [simp]: "msb (1::'a::len word) \ LENGTH('a) = 1" + unfolding word_1_wi msb_word_of_int eq_iff [where 'a=nat] + by (simp add: Suc_le_eq) + +lemma word_msb_nth: "msb w = bin_nth (uint w) (LENGTH('a) - 1)" + for w :: "'a::len word" + by (simp add: word_msb_def sint_uint bin_sign_lem) + +lemma word_msb_alt: "msb w \ hd (to_bl w)" + for w :: "'a::len word" + apply (simp add: msb_word_eq) + apply (subst hd_conv_nth) + apply simp + apply (subst nth_to_bl) + apply simp + apply simp + done + +lemma msb_nth: "msb w = w !! (LENGTH('a) - 1)" + for w :: "'a::len word" + by (simp add: word_msb_nth word_test_bit_def) + +lemma word_msb_n1 [simp]: "msb (-1::'a::len word)" + by (simp add: msb_word_eq exp_eq_zero_iff not_le) + +lemma msb_shift: "msb w \ w >> (LENGTH('a) - 1) \ 0" + for w :: "'a::len word" + by (simp add: msb_word_eq shiftr_word_eq bit_iff_odd_drop_bit drop_bit_eq_zero_iff_not_bit_last) + +lemmas word_ops_msb = msb1 [unfolded msb_nth [symmetric, unfolded One_nat_def]] + +end diff --git a/thys/Word_Lib/Misc_set_bit.thy b/thys/Word_Lib/Misc_set_bit.thy new file mode 100644 --- /dev/null +++ b/thys/Word_Lib/Misc_set_bit.thy @@ -0,0 +1,167 @@ +(* Author: Jeremy Dawson, NICTA +*) + +section \Operation variant for setting and unsetting bits\ + +theory Misc_set_bit + imports "HOL-Library.Word" Misc_msb +begin + +class set_bit = ring_bit_operations + + fixes set_bit :: \'a \ nat \ bool \ 'a\ + assumes set_bit_eq: \set_bit a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ + +instantiation int :: set_bit +begin + +definition set_bit_int :: \int \ nat \ bool \ int\ + where \set_bit i n b = bin_sc n b i\ + +instance + by standard (simp add: set_bit_int_def bin_sc_eq) + +end + +lemma int_set_bit_0 [simp]: fixes x :: int shows + "set_bit x 0 b = of_bool b + 2 * (x div 2)" + by (auto simp add: set_bit_int_def intro: bin_rl_eqI) + +lemma int_set_bit_Suc: fixes x :: int shows + "set_bit x (Suc n) b = of_bool (odd x) + 2 * set_bit (x div 2) n b" + by (auto simp add: set_bit_int_def intro: bin_rl_eqI) + +lemma bin_last_set_bit: + "bin_last (set_bit x n b) = (if n > 0 then bin_last x else b)" + by (cases n) (simp_all add: int_set_bit_Suc) + +lemma bin_rest_set_bit: + "bin_rest (set_bit x n b) = (if n > 0 then set_bit (x div 2) (n - 1) b else x div 2)" + by (cases n) (simp_all add: int_set_bit_Suc) + +lemma int_set_bit_numeral: fixes x :: int shows + "set_bit x (numeral w) b = of_bool (odd x) + 2 * set_bit (x div 2) (pred_numeral w) b" + by (simp add: set_bit_int_def) + +lemmas int_set_bit_numerals [simp] = + int_set_bit_numeral[where x="numeral w'"] + int_set_bit_numeral[where x="- numeral w'"] + int_set_bit_numeral[where x="Numeral1"] + int_set_bit_numeral[where x="1"] + int_set_bit_numeral[where x="0"] + int_set_bit_Suc[where x="numeral w'"] + int_set_bit_Suc[where x="- numeral w'"] + int_set_bit_Suc[where x="Numeral1"] + int_set_bit_Suc[where x="1"] + int_set_bit_Suc[where x="0"] + for w' + +lemma msb_set_bit [simp]: "msb (set_bit (x :: int) n b) \ msb x" +by(simp add: msb_conv_bin_sign set_bit_int_def) + +instantiation word :: (len) set_bit +begin + +definition set_bit_word :: \'a word \ nat \ bool \ 'a word\ + where word_set_bit_def: \set_bit a n x = word_of_int (bin_sc n x (uint a))\ + +instance + apply standard + apply (simp add: word_set_bit_def bin_sc_eq Bit_Operations.set_bit_def unset_bit_def) + apply transfer + apply simp + done + +end + +lemma set_bit_unfold: + \set_bit w n b = (if b then Bit_Operations.set_bit n w else unset_bit n w)\ + for w :: \'a::len word\ + by (simp add: set_bit_eq) + +lemma bit_set_bit_word_iff: + \bit (set_bit w m b) n \ (if m = n then n < LENGTH('a) \ b else bit w n)\ + for w :: \'a::len word\ + by (auto simp add: set_bit_unfold bit_unset_bit_iff bit_set_bit_iff exp_eq_zero_iff not_le bit_imp_le_length) + +lemma word_set_nth [simp]: "set_bit w n (test_bit w n) = w" + for w :: "'a::len word" + by (auto simp: word_test_bit_def word_set_bit_def) + +lemma test_bit_set: "(set_bit w n x) !! n \ n < size w \ x" + for w :: "'a::len word" + by (auto simp: word_size word_test_bit_def word_set_bit_def word_ubin.eq_norm nth_bintr) + +lemma test_bit_set_gen: + "test_bit (set_bit w n x) m = (if m = n then n < size w \ x else test_bit w m)" + for w :: "'a::len word" + apply (unfold word_size word_test_bit_def word_set_bit_def) + apply (clarsimp simp add: word_ubin.eq_norm nth_bintr bin_nth_sc_gen) + apply (auto elim!: test_bit_size [unfolded word_size] + simp add: word_test_bit_def [symmetric]) + done + +lemma word_set_set_same [simp]: "set_bit (set_bit w n x) n y = set_bit w n y" + for w :: "'a::len word" + by (rule word_eqI) (simp add : test_bit_set_gen word_size) + +lemma word_set_set_diff: + fixes w :: "'a::len word" + assumes "m \ n" + shows "set_bit (set_bit w m x) n y = set_bit (set_bit w n y) m x" + by (rule word_eqI) (auto simp: test_bit_set_gen word_size assms) + +lemma set_bit_word_of_int: "set_bit (word_of_int x) n b = word_of_int (bin_sc n b x)" + unfolding word_set_bit_def + by (rule word_eqI)(simp add: word_size bin_nth_sc_gen word_ubin.eq_norm nth_bintr) + +lemma word_set_numeral [simp]: + "set_bit (numeral bin::'a::len word) n b = + word_of_int (bin_sc n b (numeral bin))" + unfolding word_numeral_alt by (rule set_bit_word_of_int) + +lemma word_set_neg_numeral [simp]: + "set_bit (- numeral bin::'a::len word) n b = + word_of_int (bin_sc n b (- numeral bin))" + unfolding word_neg_numeral_alt by (rule set_bit_word_of_int) + +lemma word_set_bit_0 [simp]: "set_bit 0 n b = word_of_int (bin_sc n b 0)" + unfolding word_0_wi by (rule set_bit_word_of_int) + +lemma word_set_bit_1 [simp]: "set_bit 1 n b = word_of_int (bin_sc n b 1)" + unfolding word_1_wi by (rule set_bit_word_of_int) + +lemma word_set_nth_iff: "set_bit w n b = w \ w !! n = b \ n \ size w" + for w :: "'a::len word" + apply (rule iffI) + apply (rule disjCI) + apply (drule word_eqD) + apply (erule sym [THEN trans]) + apply (simp add: test_bit_set) + apply (erule disjE) + apply clarsimp + apply (rule word_eqI) + apply (clarsimp simp add : test_bit_set_gen) + apply (drule test_bit_size) + apply force + done + +lemma word_clr_le: "w \ set_bit w n False" + for w :: "'a::len word" + apply (unfold word_set_bit_def word_le_def word_ubin.eq_norm) + apply (rule order_trans) + apply (rule bintr_bin_clr_le) + apply simp + done + +lemma word_set_ge: "w \ set_bit w n True" + for w :: "'a::len word" + apply (unfold word_set_bit_def word_le_def word_ubin.eq_norm) + apply (rule order_trans [OF _ bintr_bin_set_ge]) + apply simp + done + +lemma set_bit_beyond: + "size x \ n \ set_bit x n b = x" for x :: "'a :: len word" + by (auto intro: word_eqI simp add: test_bit_set_gen word_size) + +end diff --git a/thys/Word_Lib/More_Word.thy b/thys/Word_Lib/More_Word.thy new file mode 100644 --- /dev/null +++ b/thys/Word_Lib/More_Word.thy @@ -0,0 +1,58 @@ +(* Title: HOL/Word/More_thy +*) + +section \Ancient comprehensive Word Library\ + +theory More_Word +imports + "HOL-Library.Word" + Ancient_Numeral + Reversed_Bit_Lists + Bits_Int + Bit_Comprehension + Misc_Auxiliary + Misc_Arithmetic + Misc_set_bit + Misc_lsb + Misc_Typedef + Word_rsplit +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 + +lemma shiftl_transfer [transfer_rule]: + includes lifting_syntax + shows "(pcr_word ===> (=) ===> pcr_word) (<<) (<<)" + by (unfold shiftl_eq_push_bit) transfer_prover + +end diff --git a/thys/Word_Lib/Norm_Words.thy b/thys/Word_Lib/Norm_Words.thy --- a/thys/Word_Lib/Norm_Words.thy +++ b/thys/Word_Lib/Norm_Words.thy @@ -1,110 +1,110 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Normalising Word Numerals" theory Norm_Words -imports "Signed_Words" "HOL-Word.More_Word" + imports Signed_Words More_Word begin text \ Normalise word numerals, including negative ones apart from @{term "-1"}, to the interval \[0..2^len_of 'a)\. Only for concrete word lengths. \ lemma neg_num_bintr: "(- numeral x :: 'a::len word) = word_of_int (bintrunc (LENGTH('a)) (-numeral x))" by transfer simp ML \ fun is_refl (Const (@{const_name Pure.eq}, _) $ x $ y) = (x = y) | is_refl _ = false; fun signed_dest_wordT (Type (@{type_name word}, [Type (@{type_name signed}, [T])])) = Word_Lib.dest_binT T | signed_dest_wordT T = Word_Lib.dest_wordT T fun typ_size_of t = signed_dest_wordT (type_of (Thm.term_of t)); fun num_len (Const (@{const_name Num.Bit0}, _) $ n) = num_len n + 1 | num_len (Const (@{const_name Num.Bit1}, _) $ n) = num_len n + 1 | num_len (Const (@{const_name Num.One}, _)) = 1 | num_len (Const (@{const_name numeral}, _) $ t) = num_len t | num_len (Const (@{const_name uminus}, _) $ t) = num_len t | num_len t = raise TERM ("num_len", [t]) fun unsigned_norm is_neg _ ctxt ct = (if is_neg orelse num_len (Thm.term_of ct) > typ_size_of ct then let val btr = if is_neg then @{thm neg_num_bintr} else @{thm num_abs_bintr} val th = [Thm.reflexive ct, mk_eq btr] MRS transitive_thm (* will work in context of theory Word as well *) val ss = simpset_of (@{context} addsimps @{thms bintrunc_numeral}) val cnv = simplify (put_simpset ss ctxt) th in if is_refl (Thm.prop_of cnv) then NONE else SOME cnv end else NONE) handle TERM ("num_len", _) => NONE | TYPE ("dest_binT", _, _) => NONE \ simproc_setup unsigned_norm ("numeral n::'a::len word") = \unsigned_norm false\ simproc_setup unsigned_norm_neg0 ("-numeral (num.Bit0 num)::'a::len word") = \unsigned_norm true\ simproc_setup unsigned_norm_neg1 ("-numeral (num.Bit1 num)::'a::len word") = \unsigned_norm true\ lemma minus_one_norm: "(-1 :: 'a :: len word) = of_nat (2 ^ LENGTH('a) - 1)" by (simp add:of_nat_diff) lemmas minus_one_norm_num = minus_one_norm [where 'a="'b::len bit0"] minus_one_norm [where 'a="'b::len0 bit1"] lemma "f (7 :: 2 word) = f 3" by simp lemma "f 7 = f (3 :: 2 word)" by simp lemma "f (-2) = f (21 + 1 :: 3 word)" by simp lemma "f (-2) = f (13 + 1 :: 'a::len word)" apply simp (* does not touch generic word length *) oops lemma "f (-2) = f (0xFFFFFFFE :: 32 word)" by simp lemma "(-1 :: 2 word) = 3" by simp lemma "f (-2) = f (0xFFFFFFFE :: 32 signed word)" by simp text \ We leave @{term "-1"} untouched by default, because it is often useful and its normal form can be large. To include it in the normalisation, add @{thm [source] minus_one_norm_num}. The additional normalisation is restricted to concrete numeral word lengths, like the rest. \ context notes minus_one_norm_num [simp] begin lemma "f (-1) = f (15 :: 4 word)" by simp lemma "f (-1) = f (7 :: 3 word)" by simp lemma "f (-1) = f (0xFFFF :: 16 word)" by simp lemma "f (-1) = f (0xFFFF + 1 :: 'a::len word)" apply simp (* does not touch generic -1 *) oops end end diff --git a/thys/Word_Lib/ROOT b/thys/Word_Lib/ROOT --- a/thys/Word_Lib/ROOT +++ b/thys/Word_Lib/ROOT @@ -1,13 +1,13 @@ chapter AFP -session Word_Lib (AFP) = "HOL-Word" + +session Word_Lib (AFP) = HOL + options [timeout = 300] sessions "HOL-Eisbach" theories Word_Lemmas_32 Word_Lemmas_64 Examples Guide document_files "root.tex" diff --git a/thys/Word_Lib/Reversed_Bit_Lists.thy b/thys/Word_Lib/Reversed_Bit_Lists.thy new file mode 100644 --- /dev/null +++ b/thys/Word_Lib/Reversed_Bit_Lists.thy @@ -0,0 +1,1870 @@ +(* Title: HOL/Word/Reversed_Bit_Lists.thy + Author: Jeremy Dawson, NICTA +*) + +section \Bit values as reversed lists of bools\ + +theory Reversed_Bit_Lists + imports "HOL-Library.Word"Misc_Typedef +begin + +context comm_semiring_1 +begin + +lemma horner_sum_append: + \horner_sum f a (xs @ ys) = horner_sum f a xs + a ^ length xs * horner_sum f a ys\ + using sum.atLeastLessThan_shift_bounds [of _ 0 \length xs\ \length ys\] + atLeastLessThan_add_Un [of 0 \length xs\ \length ys\] + by (simp add: horner_sum_eq_sum sum_distrib_left sum.union_disjoint ac_simps nth_append power_add) + +end + +lemma horner_sum_of_bool_2_concat: + \horner_sum of_bool 2 (concat (map (\x. map (bit x) [0.. + for ws :: \'a::len word list\ +proof (induction ws) + case Nil + then show ?case + by simp +next + case (Cons w ws) + moreover have \horner_sum of_bool 2 (map (bit w) [0.. + proof transfer + fix k :: int + have \map (\n. n < LENGTH('a) \ bit k n) [0.. + by simp + then show \horner_sum of_bool 2 (map (\n. n < LENGTH('a) \ bit k n) [0.. + by (simp only: horner_sum_bit_eq_take_bit) + qed + ultimately show ?case + by (simp add: horner_sum_append) +qed + + +subsection \Implicit augmentation of list prefixes\ + +primrec takefill :: "'a \ nat \ 'a list \ 'a list" +where + Z: "takefill fill 0 xs = []" + | Suc: "takefill fill (Suc n) xs = + (case xs of + [] \ fill # takefill fill n xs + | y # ys \ y # takefill fill n ys)" + +lemma nth_takefill: "m < n \ takefill fill n l ! m = (if m < length l then l ! m else fill)" + apply (induct n arbitrary: m l) + apply clarsimp + apply clarsimp + apply (case_tac m) + apply (simp split: list.split) + apply (simp split: list.split) + done + +lemma takefill_alt: "takefill fill n l = take n l @ replicate (n - length l) fill" + by (induct n arbitrary: l) (auto split: list.split) + +lemma takefill_replicate [simp]: "takefill fill n (replicate m fill) = replicate n fill" + by (simp add: takefill_alt replicate_add [symmetric]) + +lemma takefill_le': "n = m + k \ takefill x m (takefill x n l) = takefill x m l" + by (induct m arbitrary: l n) (auto split: list.split) + +lemma length_takefill [simp]: "length (takefill fill n l) = n" + by (simp add: takefill_alt) + +lemma take_takefill': "n = k + m \ take k (takefill fill n w) = takefill fill k w" + by (induct k arbitrary: w n) (auto split: list.split) + +lemma drop_takefill: "drop k (takefill fill (m + k) w) = takefill fill m (drop k w)" + by (induct k arbitrary: w) (auto split: list.split) + +lemma takefill_le [simp]: "m \ n \ takefill x m (takefill x n l) = takefill x m l" + by (auto simp: le_iff_add takefill_le') + +lemma take_takefill [simp]: "m \ n \ take m (takefill fill n w) = takefill fill m w" + by (auto simp: le_iff_add take_takefill') + +lemma takefill_append: "takefill fill (m + length xs) (xs @ w) = xs @ (takefill fill m w)" + by (induct xs) auto + +lemma takefill_same': "l = length xs \ takefill fill l xs = xs" + by (induct xs arbitrary: l) auto + +lemmas takefill_same [simp] = takefill_same' [OF refl] + +lemma tf_rev: + "n + k = m + length bl \ takefill x m (rev (takefill y n bl)) = + rev (takefill y m (rev (takefill x k (rev bl))))" + apply (rule nth_equalityI) + apply (auto simp add: nth_takefill rev_nth) + apply (rule_tac f = "\n. bl ! n" in arg_cong) + apply arith + done + +lemma takefill_minus: "0 < n \ takefill fill (Suc (n - 1)) w = takefill fill n w" + by auto + +lemmas takefill_Suc_cases = + list.cases [THEN takefill.Suc [THEN trans]] + +lemmas takefill_Suc_Nil = takefill_Suc_cases (1) +lemmas takefill_Suc_Cons = takefill_Suc_cases (2) + +lemmas takefill_minus_simps = takefill_Suc_cases [THEN [2] + takefill_minus [symmetric, THEN trans]] + +lemma takefill_numeral_Nil [simp]: + "takefill fill (numeral k) [] = fill # takefill fill (pred_numeral k) []" + by (simp add: numeral_eq_Suc) + +lemma takefill_numeral_Cons [simp]: + "takefill fill (numeral k) (x # xs) = x # takefill fill (pred_numeral k) xs" + by (simp add: numeral_eq_Suc) + + +subsection \Range projection\ + +definition bl_of_nth :: "nat \ (nat \ 'a) \ 'a list" + where "bl_of_nth n f = map f (rev [0.. rev (bl_of_nth n f) ! m = f m" + by (simp add: bl_of_nth_def rev_map) + +lemma bl_of_nth_inj: "(\k. k < n \ f k = g k) \ bl_of_nth n f = bl_of_nth n g" + by (simp add: bl_of_nth_def) + +lemma bl_of_nth_nth_le: "n \ length xs \ bl_of_nth n (nth (rev xs)) = drop (length xs - n) xs" + apply (induct n arbitrary: xs) + apply clarsimp + apply clarsimp + apply (rule trans [OF _ hd_Cons_tl]) + apply (frule Suc_le_lessD) + apply (simp add: rev_nth trans [OF drop_Suc drop_tl, symmetric]) + apply (subst hd_drop_conv_nth) + apply force + apply simp_all + apply (rule_tac f = "\n. drop n xs" in arg_cong) + apply simp + done + +lemma bl_of_nth_nth [simp]: "bl_of_nth (length xs) ((!) (rev xs)) = xs" + by (simp add: bl_of_nth_nth_le) + + +subsection \More\ + +definition rotater1 :: "'a list \ 'a list" + where "rotater1 ys = + (case ys of [] \ [] | x # xs \ last ys # butlast ys)" + +definition rotater :: "nat \ 'a list \ 'a list" + where "rotater n = rotater1 ^^ n" + +lemmas rotater_0' [simp] = rotater_def [where n = "0", simplified] + +lemma rotate1_rl': "rotater1 (l @ [a]) = a # l" + by (cases l) (auto simp: rotater1_def) + +lemma rotate1_rl [simp] : "rotater1 (rotate1 l) = l" + apply (unfold rotater1_def) + apply (cases "l") + apply (case_tac [2] "list") + apply auto + done + +lemma rotate1_lr [simp] : "rotate1 (rotater1 l) = l" + by (cases l) (auto simp: rotater1_def) + +lemma rotater1_rev': "rotater1 (rev xs) = rev (rotate1 xs)" + by (cases "xs") (simp add: rotater1_def, simp add: rotate1_rl') + +lemma rotater_rev': "rotater n (rev xs) = rev (rotate n xs)" + by (induct n) (auto simp: rotater_def intro: rotater1_rev') + +lemma rotater_rev: "rotater n ys = rev (rotate n (rev ys))" + using rotater_rev' [where xs = "rev ys"] by simp + +lemma rotater_drop_take: + "rotater n xs = + drop (length xs - n mod length xs) xs @ + take (length xs - n mod length xs) xs" + by (auto simp: rotater_rev rotate_drop_take rev_take rev_drop) + +lemma rotater_Suc [simp]: "rotater (Suc n) xs = rotater1 (rotater n xs)" + unfolding rotater_def by auto + +lemma nth_rotater: + \rotater m xs ! n = xs ! ((n + (length xs - m mod length xs)) mod length xs)\ if \n < length xs\ + using that by (simp add: rotater_drop_take nth_append not_less less_diff_conv ac_simps le_mod_geq) + +lemma nth_rotater1: + \rotater1 xs ! n = xs ! ((n + (length xs - 1)) mod length xs)\ if \n < length xs\ + using that nth_rotater [of n xs 1] by simp + +lemma rotate_inv_plus [rule_format]: + "\k. k = m + n \ rotater k (rotate n xs) = rotater m xs \ + rotate k (rotater n xs) = rotate m xs \ + rotater n (rotate k xs) = rotate m xs \ + rotate n (rotater k xs) = rotater m xs" + by (induct n) (auto simp: rotater_def rotate_def intro: funpow_swap1 [THEN trans]) + +lemmas rotate_inv_rel = le_add_diff_inverse2 [symmetric, THEN rotate_inv_plus] + +lemmas rotate_inv_eq = order_refl [THEN rotate_inv_rel, simplified] + +lemmas rotate_lr [simp] = rotate_inv_eq [THEN conjunct1] +lemmas rotate_rl [simp] = rotate_inv_eq [THEN conjunct2, THEN conjunct1] + +lemma rotate_gal: "rotater n xs = ys \ rotate n ys = xs" + by auto + +lemma rotate_gal': "ys = rotater n xs \ xs = rotate n ys" + by auto + +lemma length_rotater [simp]: "length (rotater n xs) = length xs" + by (simp add : rotater_rev) + +lemma rotate_eq_mod: "m mod length xs = n mod length xs \ rotate m xs = rotate n xs" + apply (rule box_equals) + defer + apply (rule rotate_conv_mod [symmetric])+ + apply simp + done + +lemma restrict_to_left: "x = y \ x = z \ y = z" + by simp + +lemmas rotate_eqs = + trans [OF rotate0 [THEN fun_cong] id_apply] + rotate_rotate [symmetric] + rotate_id + rotate_conv_mod + rotate_eq_mod + +lemmas rrs0 = rotate_eqs [THEN restrict_to_left, + simplified rotate_gal [symmetric] rotate_gal' [symmetric]] +lemmas rrs1 = rrs0 [THEN refl [THEN rev_iffD1]] +lemmas rotater_eqs = rrs1 [simplified length_rotater] +lemmas rotater_0 = rotater_eqs (1) +lemmas rotater_add = rotater_eqs (2) + +lemma butlast_map: "xs \ [] \ butlast (map f xs) = map f (butlast xs)" + by (induct xs) auto + +lemma rotater1_map: "rotater1 (map f xs) = map f (rotater1 xs)" + by (cases xs) (auto simp: rotater1_def last_map butlast_map) + +lemma rotater_map: "rotater n (map f xs) = map f (rotater n xs)" + by (induct n) (auto simp: rotater_def rotater1_map) + +lemma but_last_zip [rule_format] : + "\ys. length xs = length ys \ xs \ [] \ + last (zip xs ys) = (last xs, last ys) \ + butlast (zip xs ys) = zip (butlast xs) (butlast ys)" + apply (induct xs) + apply auto + apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ + done + +lemma but_last_map2 [rule_format] : + "\ys. length xs = length ys \ xs \ [] \ + last (map2 f xs ys) = f (last xs) (last ys) \ + butlast (map2 f xs ys) = map2 f (butlast xs) (butlast ys)" + apply (induct xs) + apply auto + apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ + done + +lemma rotater1_zip: + "length xs = length ys \ + rotater1 (zip xs ys) = zip (rotater1 xs) (rotater1 ys)" + apply (unfold rotater1_def) + apply (cases xs) + apply auto + apply ((case_tac ys, auto simp: neq_Nil_conv but_last_zip)[1])+ + done + +lemma rotater1_map2: + "length xs = length ys \ + rotater1 (map2 f xs ys) = map2 f (rotater1 xs) (rotater1 ys)" + by (simp add: rotater1_map rotater1_zip) + +lemmas lrth = + box_equals [OF asm_rl length_rotater [symmetric] + length_rotater [symmetric], + THEN rotater1_map2] + +lemma rotater_map2: + "length xs = length ys \ + rotater n (map2 f xs ys) = map2 f (rotater n xs) (rotater n ys)" + by (induct n) (auto intro!: lrth) + +lemma rotate1_map2: + "length xs = length ys \ + rotate1 (map2 f xs ys) = map2 f (rotate1 xs) (rotate1 ys)" + by (cases xs; cases ys) auto + +lemmas lth = box_equals [OF asm_rl length_rotate [symmetric] + length_rotate [symmetric], THEN rotate1_map2] + +lemma rotate_map2: + "length xs = length ys \ + rotate n (map2 f xs ys) = map2 f (rotate n xs) (rotate n ys)" + by (induct n) (auto intro!: lth) + + +subsection \Explicit bit representation of \<^typ>\int\\ + +primrec bl_to_bin_aux :: "bool list \ int \ int" + where + Nil: "bl_to_bin_aux [] w = w" + | Cons: "bl_to_bin_aux (b # bs) w = bl_to_bin_aux bs (of_bool b + 2 * w)" + +definition bl_to_bin :: "bool list \ int" + where "bl_to_bin bs = bl_to_bin_aux bs 0" + +primrec bin_to_bl_aux :: "nat \ int \ bool list \ bool list" + where + Z: "bin_to_bl_aux 0 w bl = bl" + | Suc: "bin_to_bl_aux (Suc n) w bl = bin_to_bl_aux n (bin_rest w) ((bin_last w) # bl)" + +definition bin_to_bl :: "nat \ int \ bool list" + where "bin_to_bl n w = bin_to_bl_aux n w []" + +lemma bin_to_bl_aux_zero_minus_simp [simp]: + "0 < n \ bin_to_bl_aux n 0 bl = bin_to_bl_aux (n - 1) 0 (False # bl)" + by (cases n) auto + +lemma bin_to_bl_aux_minus1_minus_simp [simp]: + "0 < n \ bin_to_bl_aux n (- 1) bl = bin_to_bl_aux (n - 1) (- 1) (True # bl)" + by (cases n) auto + +lemma bin_to_bl_aux_one_minus_simp [simp]: + "0 < n \ bin_to_bl_aux n 1 bl = bin_to_bl_aux (n - 1) 0 (True # bl)" + by (cases n) auto + +lemma bin_to_bl_aux_Bit0_minus_simp [simp]: + "0 < n \ + bin_to_bl_aux n (numeral (Num.Bit0 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (False # bl)" + by (cases n) simp_all + +lemma bin_to_bl_aux_Bit1_minus_simp [simp]: + "0 < n \ + bin_to_bl_aux n (numeral (Num.Bit1 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (True # bl)" + by (cases n) simp_all + +lemma bl_to_bin_aux_append: "bl_to_bin_aux (bs @ cs) w = bl_to_bin_aux cs (bl_to_bin_aux bs w)" + by (induct bs arbitrary: w) auto + +lemma bin_to_bl_aux_append: "bin_to_bl_aux n w bs @ cs = bin_to_bl_aux n w (bs @ cs)" + by (induct n arbitrary: w bs) auto + +lemma bl_to_bin_append: "bl_to_bin (bs @ cs) = bl_to_bin_aux cs (bl_to_bin bs)" + unfolding bl_to_bin_def by (rule bl_to_bin_aux_append) + +lemma bin_to_bl_aux_alt: "bin_to_bl_aux n w bs = bin_to_bl n w @ bs" + by (simp add: bin_to_bl_def bin_to_bl_aux_append) + +lemma bin_to_bl_0 [simp]: "bin_to_bl 0 bs = []" + by (auto simp: bin_to_bl_def) + +lemma size_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" + by (induct n arbitrary: w bs) auto + +lemma size_bin_to_bl [simp]: "length (bin_to_bl n w) = n" + by (simp add: bin_to_bl_def size_bin_to_bl_aux) + +lemma bl_bin_bl': "bin_to_bl (n + length bs) (bl_to_bin_aux bs w) = bin_to_bl_aux n w bs" + apply (induct bs arbitrary: w n) + apply auto + apply (simp_all only: add_Suc [symmetric]) + apply (auto simp add: bin_to_bl_def) + done + +lemma bl_bin_bl [simp]: "bin_to_bl (length bs) (bl_to_bin bs) = bs" + unfolding bl_to_bin_def + apply (rule box_equals) + apply (rule bl_bin_bl') + prefer 2 + apply (rule bin_to_bl_aux.Z) + apply simp + done + +lemma bl_to_bin_inj: "bl_to_bin bs = bl_to_bin cs \ length bs = length cs \ bs = cs" + apply (rule_tac box_equals) + defer + apply (rule bl_bin_bl) + apply (rule bl_bin_bl) + apply simp + done + +lemma bl_to_bin_False [simp]: "bl_to_bin (False # bl) = bl_to_bin bl" + by (auto simp: bl_to_bin_def) + +lemma bl_to_bin_Nil [simp]: "bl_to_bin [] = 0" + by (auto simp: bl_to_bin_def) + +lemma bin_to_bl_zero_aux: "bin_to_bl_aux n 0 bl = replicate n False @ bl" + by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) + +lemma bin_to_bl_zero: "bin_to_bl n 0 = replicate n False" + by (simp add: bin_to_bl_def bin_to_bl_zero_aux) + +lemma bin_to_bl_minus1_aux: "bin_to_bl_aux n (- 1) bl = replicate n True @ bl" + by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) + +lemma bin_to_bl_minus1: "bin_to_bl n (- 1) = replicate n True" + by (simp add: bin_to_bl_def bin_to_bl_minus1_aux) + + +subsection \Semantic interpretation of \<^typ>\bool list\ as \<^typ>\int\\ + +lemma bin_bl_bin': "bl_to_bin (bin_to_bl_aux n w bs) = bl_to_bin_aux bs (bintrunc n w)" + by (induct n arbitrary: w bs) (auto simp: bl_to_bin_def take_bit_Suc ac_simps mod_2_eq_odd) + +lemma bin_bl_bin [simp]: "bl_to_bin (bin_to_bl n w) = bintrunc n w" + by (auto simp: bin_to_bl_def bin_bl_bin') + +lemma bl_to_bin_rep_F: "bl_to_bin (replicate n False @ bl) = bl_to_bin bl" + by (simp add: bin_to_bl_zero_aux [symmetric] bin_bl_bin') (simp add: bl_to_bin_def) + +lemma bin_to_bl_trunc [simp]: "n \ m \ bin_to_bl n (bintrunc m w) = bin_to_bl n w" + by (auto intro: bl_to_bin_inj) + +lemma bin_to_bl_aux_bintr: + "bin_to_bl_aux n (bintrunc m bin) bl = + replicate (n - m) False @ bin_to_bl_aux (min n m) bin bl" + apply (induct n arbitrary: m bin bl) + apply clarsimp + apply clarsimp + apply (case_tac "m") + apply (clarsimp simp: bin_to_bl_zero_aux) + apply (erule thin_rl) + apply (induct_tac n) + apply (auto simp add: take_bit_Suc) + done + +lemma bin_to_bl_bintr: + "bin_to_bl n (bintrunc m bin) = replicate (n - m) False @ bin_to_bl (min n m) bin" + unfolding bin_to_bl_def by (rule bin_to_bl_aux_bintr) + +lemma bl_to_bin_rep_False: "bl_to_bin (replicate n False) = 0" + by (induct n) auto + +lemma len_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" + by (fact size_bin_to_bl_aux) + +lemma len_bin_to_bl: "length (bin_to_bl n w) = n" + by (fact size_bin_to_bl) (* FIXME: duplicate *) + +lemma sign_bl_bin': "bin_sign (bl_to_bin_aux bs w) = bin_sign w" + by (induction bs arbitrary: w) (simp_all add: bin_sign_def) + +lemma sign_bl_bin: "bin_sign (bl_to_bin bs) = 0" + by (simp add: bl_to_bin_def sign_bl_bin') + +lemma bl_sbin_sign_aux: "hd (bin_to_bl_aux (Suc n) w bs) = (bin_sign (sbintrunc n w) = -1)" + by (induction n arbitrary: w bs) (auto simp add: bin_sign_def even_iff_mod_2_eq_zero bit_Suc) + +lemma bl_sbin_sign: "hd (bin_to_bl (Suc n) w) = (bin_sign (sbintrunc n w) = -1)" + unfolding bin_to_bl_def by (rule bl_sbin_sign_aux) + +lemma bin_nth_of_bl_aux: + "bin_nth (bl_to_bin_aux bl w) n = + (n < size bl \ rev bl ! n \ n \ length bl \ bin_nth w (n - size bl))" + apply (induction bl arbitrary: w) + apply simp_all + apply safe + apply (simp_all add: not_le nth_append bit_double_iff even_bit_succ_iff split: if_splits) + done + +lemma bin_nth_of_bl: "bin_nth (bl_to_bin bl) n = (n < length bl \ rev bl ! n)" + by (simp add: bl_to_bin_def bin_nth_of_bl_aux) + +lemma bin_nth_bl: "n < m \ bin_nth w n = nth (rev (bin_to_bl m w)) n" + apply (induct n arbitrary: m w) + apply clarsimp + apply (case_tac m, clarsimp) + apply (clarsimp simp: bin_to_bl_def) + apply (simp add: bin_to_bl_aux_alt) + apply (case_tac m, clarsimp) + apply (clarsimp simp: bin_to_bl_def) + apply (simp add: bin_to_bl_aux_alt bit_Suc) + done + +lemma nth_bin_to_bl_aux: + "n < m + length bl \ (bin_to_bl_aux m w bl) ! n = + (if n < m then bit w (m - 1 - n) else bl ! (n - m))" + apply (induction bl arbitrary: w) + apply simp_all + apply (simp add: bin_nth_bl [of \m - Suc n\ m] rev_nth flip: bin_to_bl_def) + apply (metis One_nat_def Suc_pred add_diff_cancel_left' + add_diff_cancel_right' bin_to_bl_aux_alt bin_to_bl_def + diff_Suc_Suc diff_is_0_eq diff_zero less_Suc_eq_0_disj + less_antisym less_imp_Suc_add list.size(3) nat_less_le nth_append size_bin_to_bl_aux) + done + +lemma nth_bin_to_bl: "n < m \ (bin_to_bl m w) ! n = bin_nth w (m - Suc n)" + by (simp add: bin_to_bl_def nth_bin_to_bl_aux) + +lemma takefill_bintrunc: "takefill False n bl = rev (bin_to_bl n (bl_to_bin (rev bl)))" + apply (rule nth_equalityI) + apply simp + apply (clarsimp simp: nth_takefill rev_nth nth_bin_to_bl bin_nth_of_bl) + done + +lemma bl_bin_bl_rtf: "bin_to_bl n (bl_to_bin bl) = rev (takefill False n (rev bl))" + by (simp add: takefill_bintrunc) + +lemma bl_to_bin_lt2p_aux: "bl_to_bin_aux bs w < (w + 1) * (2 ^ length bs)" +proof (induction bs arbitrary: w) + case Nil + then show ?case + by simp +next + case (Cons b bs) + from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] + show ?case + apply (auto simp add: algebra_simps) + apply (subst mult_2 [of \2 ^ length bs\]) + apply (simp only: add.assoc) + apply (rule pos_add_strict) + apply simp_all + done +qed + +lemma bl_to_bin_lt2p_drop: "bl_to_bin bs < 2 ^ length (dropWhile Not bs)" +proof (induct bs) + case Nil + then show ?case by simp +next + case (Cons b bs) + with bl_to_bin_lt2p_aux[where w=1] show ?case + by (simp add: bl_to_bin_def) +qed + +lemma bl_to_bin_lt2p: "bl_to_bin bs < 2 ^ length bs" + by (metis bin_bl_bin bintr_lt2p bl_bin_bl) + +lemma bl_to_bin_ge2p_aux: "bl_to_bin_aux bs w \ w * (2 ^ length bs)" +proof (induction bs arbitrary: w) + case Nil + then show ?case + by simp +next + case (Cons b bs) + from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] + show ?case + apply (auto simp add: algebra_simps) + apply (rule add_le_imp_le_left [of \2 ^ length bs\]) + apply (rule add_increasing) + apply simp_all + done +qed + +lemma bl_to_bin_ge0: "bl_to_bin bs \ 0" + apply (unfold bl_to_bin_def) + apply (rule xtrans(4)) + apply (rule bl_to_bin_ge2p_aux) + apply simp + done + +lemma butlast_rest_bin: "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (bin_rest w)" + apply (unfold bin_to_bl_def) + apply (cases n, clarsimp) + apply clarsimp + apply (auto simp add: bin_to_bl_aux_alt) + done + +lemma butlast_bin_rest: "butlast bl = bin_to_bl (length bl - Suc 0) (bin_rest (bl_to_bin bl))" + using butlast_rest_bin [where w="bl_to_bin bl" and n="length bl"] by simp + +lemma butlast_rest_bl2bin_aux: + "bl \ [] \ bl_to_bin_aux (butlast bl) w = bin_rest (bl_to_bin_aux bl w)" + by (induct bl arbitrary: w) auto + +lemma butlast_rest_bl2bin: "bl_to_bin (butlast bl) = bin_rest (bl_to_bin bl)" + by (cases bl) (auto simp: bl_to_bin_def butlast_rest_bl2bin_aux) + +lemma trunc_bl2bin_aux: + "bintrunc m (bl_to_bin_aux bl w) = + bl_to_bin_aux (drop (length bl - m) bl) (bintrunc (m - length bl) w)" +proof (induct bl arbitrary: w) + case Nil + show ?case by simp +next + case (Cons b bl) + show ?case + proof (cases "m - length bl") + case 0 + then have "Suc (length bl) - m = Suc (length bl - m)" by simp + with Cons show ?thesis by simp + next + case (Suc n) + then have "m - Suc (length bl) = n" by simp + with Cons Suc show ?thesis by (simp add: take_bit_Suc ac_simps) + qed +qed + +lemma trunc_bl2bin: "bintrunc m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)" + by (simp add: bl_to_bin_def trunc_bl2bin_aux) + +lemma trunc_bl2bin_len [simp]: "bintrunc (length bl) (bl_to_bin bl) = bl_to_bin bl" + by (simp add: trunc_bl2bin) + +lemma bl2bin_drop: "bl_to_bin (drop k bl) = bintrunc (length bl - k) (bl_to_bin bl)" + apply (rule trans) + prefer 2 + apply (rule trunc_bl2bin [symmetric]) + apply (cases "k \ length bl") + apply auto + done + +lemma take_rest_power_bin: "m \ n \ take m (bin_to_bl n w) = bin_to_bl m ((bin_rest ^^ (n - m)) w)" + apply (rule nth_equalityI) + apply simp + apply (clarsimp simp add: nth_bin_to_bl nth_rest_power_bin) + done + +lemma last_bin_last': "size xs > 0 \ last xs \ bin_last (bl_to_bin_aux xs w)" + by (induct xs arbitrary: w) auto + +lemma last_bin_last: "size xs > 0 \ last xs \ bin_last (bl_to_bin xs)" + unfolding bl_to_bin_def by (erule last_bin_last') + +lemma bin_last_last: "bin_last w \ last (bin_to_bl (Suc n) w)" + by (simp add: bin_to_bl_def) (auto simp: bin_to_bl_aux_alt) + +lemma drop_bin2bl_aux: + "drop m (bin_to_bl_aux n bin bs) = + bin_to_bl_aux (n - m) bin (drop (m - n) bs)" + apply (induction n arbitrary: m bin bs) + apply auto + apply (case_tac "m \ n") + apply (auto simp add: not_le Suc_diff_le) + apply (case_tac "m - n") + apply auto + apply (use Suc_diff_Suc in fastforce) + done + +lemma drop_bin2bl: "drop m (bin_to_bl n bin) = bin_to_bl (n - m) bin" + by (simp add: bin_to_bl_def drop_bin2bl_aux) + +lemma take_bin2bl_lem1: "take m (bin_to_bl_aux m w bs) = bin_to_bl m w" + apply (induct m arbitrary: w bs) + apply clarsimp + apply clarsimp + apply (simp add: bin_to_bl_aux_alt) + apply (simp add: bin_to_bl_def) + apply (simp add: bin_to_bl_aux_alt) + done + +lemma take_bin2bl_lem: "take m (bin_to_bl_aux (m + n) w bs) = take m (bin_to_bl (m + n) w)" + by (induct n arbitrary: w bs) (simp_all (no_asm) add: bin_to_bl_def take_bin2bl_lem1, simp) + +lemma bin_split_take: "bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl (m + n) c)" + apply (induct n arbitrary: b c) + apply clarsimp + apply (clarsimp simp: Let_def split: prod.split_asm) + apply (simp add: bin_to_bl_def) + apply (simp add: take_bin2bl_lem drop_bit_Suc) + done + +lemma bin_to_bl_drop_bit: + "k = m + n \ bin_to_bl m (drop_bit n c) = take m (bin_to_bl k c)" + using bin_split_take by simp + +lemma bin_split_take1: + "k = m + n \ bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl k c)" + using bin_split_take by simp + +lemma bl_bin_bl_rep_drop: + "bin_to_bl n (bl_to_bin bl) = + replicate (n - length bl) False @ drop (length bl - n) bl" + by (simp add: bl_to_bin_inj bl_to_bin_rep_F trunc_bl2bin) + +lemma bl_to_bin_aux_cat: + "bl_to_bin_aux bs (bin_cat w nv v) = + bin_cat w (nv + length bs) (bl_to_bin_aux bs v)" + by (rule bit_eqI) + (auto simp add: bin_nth_of_bl_aux bin_nth_cat algebra_simps) + +lemma bin_to_bl_aux_cat: + "bin_to_bl_aux (nv + nw) (bin_cat v nw w) bs = + bin_to_bl_aux nv v (bin_to_bl_aux nw w bs)" + by (induction nw arbitrary: w bs) (simp_all add: concat_bit_Suc) + +lemma bl_to_bin_aux_alt: "bl_to_bin_aux bs w = bin_cat w (length bs) (bl_to_bin bs)" + using bl_to_bin_aux_cat [where nv = "0" and v = "0"] + by (simp add: bl_to_bin_def [symmetric]) + +lemma bin_to_bl_cat: + "bin_to_bl (nv + nw) (bin_cat v nw w) = + bin_to_bl_aux nv v (bin_to_bl nw w)" + by (simp add: bin_to_bl_def bin_to_bl_aux_cat) + +lemmas bl_to_bin_aux_app_cat = + trans [OF bl_to_bin_aux_append bl_to_bin_aux_alt] + +lemmas bin_to_bl_aux_cat_app = + trans [OF bin_to_bl_aux_cat bin_to_bl_aux_alt] + +lemma bl_to_bin_app_cat: + "bl_to_bin (bsa @ bs) = bin_cat (bl_to_bin bsa) (length bs) (bl_to_bin bs)" + by (simp only: bl_to_bin_aux_app_cat bl_to_bin_def) + +lemma bin_to_bl_cat_app: + "bin_to_bl (n + nw) (bin_cat w nw wa) = bin_to_bl n w @ bin_to_bl nw wa" + by (simp only: bin_to_bl_def bin_to_bl_aux_cat_app) + +text \\bl_to_bin_app_cat_alt\ and \bl_to_bin_app_cat\ are easily interderivable.\ +lemma bl_to_bin_app_cat_alt: "bin_cat (bl_to_bin cs) n w = bl_to_bin (cs @ bin_to_bl n w)" + by (simp add: bl_to_bin_app_cat) + +lemma mask_lem: "(bl_to_bin (True # replicate n False)) = bl_to_bin (replicate n True) + 1" + apply (unfold bl_to_bin_def) + apply (induct n) + apply simp + apply (simp only: Suc_eq_plus1 replicate_add append_Cons [symmetric] bl_to_bin_aux_append) + apply simp + done + +lemma bin_exhaust: + "(\x b. bin = of_bool b + 2 * x \ Q) \ Q" for bin :: int + apply (cases \even bin\) + apply (auto elim!: evenE oddE) + apply fastforce + apply fastforce + done + +primrec rbl_succ :: "bool list \ bool list" + where + Nil: "rbl_succ Nil = Nil" + | Cons: "rbl_succ (x # xs) = (if x then False # rbl_succ xs else True # xs)" + +primrec rbl_pred :: "bool list \ bool list" + where + Nil: "rbl_pred Nil = Nil" + | Cons: "rbl_pred (x # xs) = (if x then False # xs else True # rbl_pred xs)" + +primrec rbl_add :: "bool list \ bool list \ bool list" + where \ \result is length of first arg, second arg may be longer\ + Nil: "rbl_add Nil x = Nil" + | Cons: "rbl_add (y # ys) x = + (let ws = rbl_add ys (tl x) + in (y \ hd x) # (if hd x \ y then rbl_succ ws else ws))" + +primrec rbl_mult :: "bool list \ bool list \ bool list" + where \ \result is length of first arg, second arg may be longer\ + Nil: "rbl_mult Nil x = Nil" + | Cons: "rbl_mult (y # ys) x = + (let ws = False # rbl_mult ys x + in if y then rbl_add ws x else ws)" + +lemma size_rbl_pred: "length (rbl_pred bl) = length bl" + by (induct bl) auto + +lemma size_rbl_succ: "length (rbl_succ bl) = length bl" + by (induct bl) auto + +lemma size_rbl_add: "length (rbl_add bl cl) = length bl" + by (induct bl arbitrary: cl) (auto simp: Let_def size_rbl_succ) + +lemma size_rbl_mult: "length (rbl_mult bl cl) = length bl" + by (induct bl arbitrary: cl) (auto simp add: Let_def size_rbl_add) + +lemmas rbl_sizes [simp] = + size_rbl_pred size_rbl_succ size_rbl_add size_rbl_mult + +lemmas rbl_Nils = + rbl_pred.Nil rbl_succ.Nil rbl_add.Nil rbl_mult.Nil + +lemma rbl_add_app2: "length blb \ length bla \ rbl_add bla (blb @ blc) = rbl_add bla blb" + apply (induct bla arbitrary: blb) + apply simp + apply clarsimp + apply (case_tac blb, clarsimp) + apply (clarsimp simp: Let_def) + done + +lemma rbl_add_take2: + "length blb \ length bla \ rbl_add bla (take (length bla) blb) = rbl_add bla blb" + apply (induct bla arbitrary: blb) + apply simp + apply clarsimp + apply (case_tac blb, clarsimp) + apply (clarsimp simp: Let_def) + done + +lemma rbl_mult_app2: "length blb \ length bla \ rbl_mult bla (blb @ blc) = rbl_mult bla blb" + apply (induct bla arbitrary: blb) + apply simp + apply clarsimp + apply (case_tac blb, clarsimp) + apply (clarsimp simp: Let_def rbl_add_app2) + done + +lemma rbl_mult_take2: + "length blb \ length bla \ rbl_mult bla (take (length bla) blb) = rbl_mult bla blb" + apply (rule trans) + apply (rule rbl_mult_app2 [symmetric]) + apply simp + apply (rule_tac f = "rbl_mult bla" in arg_cong) + apply (rule append_take_drop_id) + done + +lemma rbl_add_split: + "P (rbl_add (y # ys) (x # xs)) = + (\ws. length ws = length ys \ ws = rbl_add ys xs \ + (y \ ((x \ P (False # rbl_succ ws)) \ (\ x \ P (True # ws)))) \ + (\ y \ P (x # ws)))" + by (cases y) (auto simp: Let_def) + +lemma rbl_mult_split: + "P (rbl_mult (y # ys) xs) = + (\ws. length ws = Suc (length ys) \ ws = False # rbl_mult ys xs \ + (y \ P (rbl_add ws xs)) \ (\ y \ P ws))" + by (auto simp: Let_def) + +lemma rbl_pred: "rbl_pred (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin - 1))" +proof (unfold bin_to_bl_def, induction n arbitrary: bin) + case 0 + then show ?case + by simp +next + case (Suc n) + obtain b k where \bin = of_bool b + 2 * k\ + using bin_exhaust by blast + moreover have \(2 * k - 1) div 2 = k - 1\ + using even_succ_div_2 [of \2 * (k - 1)\] + by simp + ultimately show ?case + using Suc [of \bin div 2\] + by simp (simp add: bin_to_bl_aux_alt) +qed + +lemma rbl_succ: "rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin + 1))" + apply (unfold bin_to_bl_def) + apply (induction n arbitrary: bin) + apply simp_all + apply (case_tac bin rule: bin_exhaust) + apply simp + apply (simp add: bin_to_bl_aux_alt ac_simps) + done + +lemma rbl_add: + "\bina binb. rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = + rev (bin_to_bl n (bina + binb))" + apply (unfold bin_to_bl_def) + apply (induct n) + apply simp + apply clarsimp + apply (case_tac bina rule: bin_exhaust) + apply (case_tac binb rule: bin_exhaust) + apply (case_tac b) + apply (case_tac [!] "ba") + apply (auto simp: rbl_succ bin_to_bl_aux_alt Let_def ac_simps) + done + +lemma rbl_add_long: + "m \ n \ rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = + rev (bin_to_bl n (bina + binb))" + apply (rule box_equals [OF _ rbl_add_take2 rbl_add]) + apply (rule_tac f = "rbl_add (rev (bin_to_bl n bina))" in arg_cong) + apply (rule rev_swap [THEN iffD1]) + apply (simp add: rev_take drop_bin2bl) + apply simp + done + +lemma rbl_mult_gt1: + "m \ length bl \ + rbl_mult bl (rev (bin_to_bl m binb)) = + rbl_mult bl (rev (bin_to_bl (length bl) binb))" + apply (rule trans) + apply (rule rbl_mult_take2 [symmetric]) + apply simp_all + apply (rule_tac f = "rbl_mult bl" in arg_cong) + apply (rule rev_swap [THEN iffD1]) + apply (simp add: rev_take drop_bin2bl) + done + +lemma rbl_mult_gt: + "m > n \ + rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = + rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb))" + by (auto intro: trans [OF rbl_mult_gt1]) + +lemmas rbl_mult_Suc = lessI [THEN rbl_mult_gt] + +lemma rbbl_Cons: "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (of_bool b + 2 * x))" + by (simp add: bin_to_bl_def) (simp add: bin_to_bl_aux_alt) + +lemma rbl_mult: + "rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = + rev (bin_to_bl n (bina * binb))" + apply (induct n arbitrary: bina binb) + apply simp_all + apply (unfold bin_to_bl_def) + apply clarsimp + apply (case_tac bina rule: bin_exhaust) + apply (case_tac binb rule: bin_exhaust) + apply simp + apply (simp add: bin_to_bl_aux_alt) + apply (simp add: rbbl_Cons rbl_mult_Suc rbl_add algebra_simps) + done + +lemma sclem: "size (concat (map (bin_to_bl n) xs)) = length xs * n" + by (induct xs) auto + +lemma bin_cat_foldl_lem: + "foldl (\u. bin_cat u n) x xs = + bin_cat x (size xs * n) (foldl (\u. bin_cat u n) y xs)" + apply (induct xs arbitrary: x) + apply simp + apply (simp (no_asm)) + apply (frule asm_rl) + apply (drule meta_spec) + apply (erule trans) + apply (drule_tac x = "bin_cat y n a" in meta_spec) + apply (simp add: bin_cat_assoc_sym min.absorb2) + done + +lemma bin_rcat_bl: "bin_rcat n wl = bl_to_bin (concat (map (bin_to_bl n) wl))" + apply (unfold bin_rcat_eq_foldl) + apply (rule sym) + apply (induct wl) + apply (auto simp add: bl_to_bin_append) + apply (simp add: bl_to_bin_aux_alt sclem) + apply (simp add: bin_cat_foldl_lem [symmetric]) + done + +lemma bin_last_bl_to_bin: "bin_last (bl_to_bin bs) \ bs \ [] \ last bs" +by(cases "bs = []")(auto simp add: bl_to_bin_def last_bin_last'[where w=0]) + +lemma bin_rest_bl_to_bin: "bin_rest (bl_to_bin bs) = bl_to_bin (butlast bs)" +by(cases "bs = []")(simp_all add: bl_to_bin_def butlast_rest_bl2bin_aux) + +lemma bl_xor_aux_bin: + "map2 (\x y. x \ y) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = + bin_to_bl_aux n (v XOR w) (map2 (\x y. x \ y) bs cs)" + apply (induction n arbitrary: v w bs cs) + apply auto + apply (case_tac v rule: bin_exhaust) + apply (case_tac w rule: bin_exhaust) + apply clarsimp + done + +lemma bl_or_aux_bin: + "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = + bin_to_bl_aux n (v OR w) (map2 (\) bs cs)" + by (induct n arbitrary: v w bs cs) simp_all + +lemma bl_and_aux_bin: + "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = + bin_to_bl_aux n (v AND w) (map2 (\) bs cs)" + by (induction n arbitrary: v w bs cs) simp_all + +lemma bl_not_aux_bin: "map Not (bin_to_bl_aux n w cs) = bin_to_bl_aux n (NOT w) (map Not cs)" + by (induct n arbitrary: w cs) auto + +lemma bl_not_bin: "map Not (bin_to_bl n w) = bin_to_bl n (NOT w)" + by (simp add: bin_to_bl_def bl_not_aux_bin) + +lemma bl_and_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v AND w)" + by (simp add: bin_to_bl_def bl_and_aux_bin) + +lemma bl_or_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v OR w)" + by (simp add: bin_to_bl_def bl_or_aux_bin) + +lemma bl_xor_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v XOR w)" + using bl_xor_aux_bin by (simp add: bin_to_bl_def) + + +subsection \Type \<^typ>\'a word\\ + +lift_definition of_bl :: \bool list \ 'a::len word\ + is bl_to_bin . + +lift_definition to_bl :: \'a::len word \ bool list\ + is \bin_to_bl LENGTH('a)\ + by (simp add: bl_to_bin_inj) + +lemma to_bl_eq: + \to_bl w = bin_to_bl (LENGTH('a)) (uint w)\ + for w :: \'a::len word\ + by transfer simp + +lemma bit_of_bl_iff: + \bit (of_bl bs :: 'a word) n \ rev bs ! n \ n < LENGTH('a::len) \ n < length bs\ + by transfer (simp add: bin_nth_of_bl ac_simps) + +lemma rev_to_bl_eq: + \rev (to_bl w) = map (bit w) [0.. + for w :: \'a::len word\ + apply (rule nth_equalityI) + apply (simp add: to_bl.rep_eq) + apply (simp add: bin_nth_bl bit_word.rep_eq to_bl.rep_eq) + done + +lemma to_bl_eq_rev: + \to_bl w = map (bit w) (rev [0.. + for w :: \'a::len word\ + using rev_to_bl_eq [of w] + apply (subst rev_is_rev_conv [symmetric]) + apply (simp add: rev_map) + done + +lemma of_bl_rev_eq: + \of_bl (rev bs) = horner_sum of_bool 2 bs\ + apply (rule bit_word_eqI) + apply (simp add: bit_of_bl_iff) + apply transfer + apply (simp add: bit_horner_sum_bit_iff ac_simps) + done + +lemma of_bl_eq: + \of_bl bs = horner_sum of_bool 2 (rev bs)\ + using of_bl_rev_eq [of \rev bs\] by simp + +lemma bshiftr1_eq: + \bshiftr1 b w = of_bl (b # butlast (to_bl w))\ + apply transfer + apply auto + apply (subst bl_to_bin_app_cat [of \[True]\, simplified]) + apply simp + apply (metis One_nat_def add.commute bin_bl_bin bin_last_bl_to_bin bin_rest_bl_to_bin butlast_bin_rest concat_bit_eq last.simps list.distinct(1) list.size(3) list.size(4) odd_iff_mod_2_eq_one plus_1_eq_Suc power_Suc0_right push_bit_of_1 size_bin_to_bl take_bit_eq_mod trunc_bl2bin_len) + apply (simp add: butlast_rest_bl2bin) + done + +lemma length_to_bl_eq: + \length (to_bl w) = LENGTH('a)\ + for w :: \'a::len word\ + by transfer simp + +lemma word_rotr_eq: + \word_rotr n w = of_bl (rotater n (to_bl w))\ + apply (rule bit_word_eqI) + subgoal for n + apply (cases \n < LENGTH('a)\) + apply (simp_all add: bit_word_rotr_iff bit_of_bl_iff rotater_rev length_to_bl_eq nth_rotate rev_to_bl_eq ac_simps) + done + done + +lemma word_rotl_eq: + \word_rotl n w = of_bl (rotate n (to_bl w))\ +proof - + have \rotate n (to_bl w) = rev (rotater n (rev (to_bl w)))\ + by (simp add: rotater_rev') + then show ?thesis + apply (simp add: word_rotl_eq_word_rotr bit_of_bl_iff length_to_bl_eq rev_to_bl_eq) + apply (rule bit_word_eqI) + subgoal for n + apply (cases \n < LENGTH('a)\) + apply (simp_all add: bit_word_rotr_iff bit_of_bl_iff nth_rotater) + done + done +qed + +lemma to_bl_def': "(to_bl :: 'a::len word \ bool list) = bin_to_bl (LENGTH('a)) \ uint" + by transfer (simp add: fun_eq_iff) + +\ \type definitions theorem for in terms of equivalent bool list\ +lemma td_bl: + "type_definition + (to_bl :: 'a::len word \ bool list) + of_bl + {bl. length bl = LENGTH('a)}" + apply (standard; transfer) + apply (auto dest: sym) + done + +interpretation word_bl: + type_definition + "to_bl :: 'a::len word \ bool list" + of_bl + "{bl. length bl = LENGTH('a::len)}" + by (fact td_bl) + +lemmas word_bl_Rep' = word_bl.Rep [unfolded mem_Collect_eq, iff] + +lemma word_size_bl: "size w = size (to_bl w)" + by (auto simp: word_size) + +lemma to_bl_use_of_bl: "to_bl w = bl \ w = of_bl bl \ length bl = length (to_bl w)" + by (fastforce elim!: word_bl.Abs_inverse [unfolded mem_Collect_eq]) + +lemma length_bl_gt_0 [iff]: "0 < length (to_bl x)" + for x :: "'a::len word" + unfolding word_bl_Rep' by (rule len_gt_0) + +lemma bl_not_Nil [iff]: "to_bl x \ []" + for x :: "'a::len word" + by (fact length_bl_gt_0 [unfolded length_greater_0_conv]) + +lemma length_bl_neq_0 [iff]: "length (to_bl x) \ 0" + for x :: "'a::len word" + by (fact length_bl_gt_0 [THEN gr_implies_not0]) + +lemma hd_bl_sign_sint: "hd (to_bl w) = (bin_sign (sint w) = -1)" + apply transfer + apply (auto simp add: bin_sign_def) + using bin_sign_lem bl_sbin_sign apply fastforce + using bin_sign_lem bl_sbin_sign apply force + done + +lemma of_bl_drop': + "lend = length bl - LENGTH('a::len) \ + of_bl (drop lend bl) = (of_bl bl :: 'a word)" + by transfer (simp flip: trunc_bl2bin) + +lemma test_bit_of_bl: + "(of_bl bl::'a::len word) !! n = (rev bl ! n \ n < LENGTH('a) \ n < length bl)" + by transfer (simp add: bin_nth_of_bl ac_simps) + +lemma no_of_bl: "(numeral bin ::'a::len word) = of_bl (bin_to_bl (LENGTH('a)) (numeral bin))" + by transfer simp + +lemma uint_bl: "to_bl w = bin_to_bl (size w) (uint w)" + by transfer simp + +lemma to_bl_bin: "bl_to_bin (to_bl w) = uint w" + by (simp add: uint_bl word_size) + +lemma to_bl_of_bin: "to_bl (word_of_int bin::'a::len word) = bin_to_bl (LENGTH('a)) bin" + by (auto simp: uint_bl word_ubin.eq_norm word_size) + +lemma to_bl_numeral [simp]: + "to_bl (numeral bin::'a::len word) = + bin_to_bl (LENGTH('a)) (numeral bin)" + unfolding word_numeral_alt by (rule to_bl_of_bin) + +lemma to_bl_neg_numeral [simp]: + "to_bl (- numeral bin::'a::len word) = + bin_to_bl (LENGTH('a)) (- numeral bin)" + unfolding word_neg_numeral_alt by (rule to_bl_of_bin) + +lemma to_bl_to_bin [simp] : "bl_to_bin (to_bl w) = uint w" + by (simp add: uint_bl word_size) + +lemma uint_bl_bin: "bl_to_bin (bin_to_bl (LENGTH('a)) (uint x)) = uint x" + for x :: "'a::len word" + by (rule trans [OF bin_bl_bin word_ubin.norm_Rep]) + +lemma ucast_bl: "ucast w = of_bl (to_bl w)" + by transfer simp + +lemma ucast_down_bl: + \(ucast :: 'a::len word \ 'b::len word) (of_bl bl) = of_bl bl\ + if \is_down (ucast :: 'a::len word \ 'b::len word)\ + using that by transfer simp + +lemma of_bl_append_same: "of_bl (X @ to_bl w) = w" + by transfer (simp add: bl_to_bin_app_cat) + +lemma ucast_of_bl_up: + \ucast (of_bl bl :: 'a::len word) = of_bl bl\ + if \size bl \ size (of_bl bl :: 'a::len word)\ + using that + apply transfer + apply (rule bit_eqI) + apply (auto simp add: bit_take_bit_iff) + apply (subst (asm) trunc_bl2bin_len [symmetric]) + apply (auto simp only: bit_take_bit_iff) + done + +lemma word_rev_tf: + "to_bl (of_bl bl::'a::len word) = + rev (takefill False (LENGTH('a)) (rev bl))" + by transfer (simp add: bl_bin_bl_rtf) + +lemma word_rep_drop: + "to_bl (of_bl bl::'a::len word) = + replicate (LENGTH('a) - length bl) False @ + drop (length bl - LENGTH('a)) bl" + by (simp add: word_rev_tf takefill_alt rev_take) + +lemma to_bl_ucast: + "to_bl (ucast (w::'b::len word) ::'a::len word) = + replicate (LENGTH('a) - LENGTH('b)) False @ + drop (LENGTH('b) - LENGTH('a)) (to_bl w)" + apply (unfold ucast_bl) + apply (rule trans) + apply (rule word_rep_drop) + apply simp + done + +lemma ucast_up_app: + \to_bl (ucast w :: 'b::len word) = replicate n False @ (to_bl w)\ + if \source_size (ucast :: 'a word \ 'b word) + n = target_size (ucast :: 'a word \ 'b word)\ + for w :: \'a::len word\ + using that + by (auto simp add : source_size target_size to_bl_ucast) + +lemma ucast_down_drop [OF refl]: + "uc = ucast \ source_size uc = target_size uc + n \ + to_bl (uc w) = drop n (to_bl w)" + by (auto simp add : source_size target_size to_bl_ucast) + +lemma scast_down_drop [OF refl]: + "sc = scast \ source_size sc = target_size sc + n \ + to_bl (sc w) = drop n (to_bl w)" + apply (subgoal_tac "sc = ucast") + apply safe + apply simp + apply (erule ucast_down_drop) + apply (rule down_cast_same [symmetric]) + apply (simp add : source_size target_size is_down) + done + +lemma word_0_bl [simp]: "of_bl [] = 0" + by transfer simp + +lemma word_1_bl: "of_bl [True] = 1" + by transfer (simp add: bl_to_bin_def) + +lemma of_bl_0 [simp]: "of_bl (replicate n False) = 0" + by transfer (simp add: bl_to_bin_rep_False) + +lemma to_bl_0 [simp]: "to_bl (0::'a::len word) = replicate (LENGTH('a)) False" + by (simp add: uint_bl word_size bin_to_bl_zero) + +\ \links with \rbl\ operations\ +lemma word_succ_rbl: "to_bl w = bl \ to_bl (word_succ w) = rev (rbl_succ (rev bl))" + by transfer (simp add: rbl_succ) + +lemma word_pred_rbl: "to_bl w = bl \ to_bl (word_pred w) = rev (rbl_pred (rev bl))" + by transfer (simp add: rbl_pred) + +lemma word_add_rbl: + "to_bl v = vbl \ to_bl w = wbl \ + to_bl (v + w) = rev (rbl_add (rev vbl) (rev wbl))" + apply transfer + apply (drule sym) + apply (drule sym) + apply (simp add: rbl_add) + done + +lemma word_mult_rbl: + "to_bl v = vbl \ to_bl w = wbl \ + to_bl (v * w) = rev (rbl_mult (rev vbl) (rev wbl))" + apply transfer + apply (drule sym) + apply (drule sym) + apply (simp add: rbl_mult) + done + +lemma rtb_rbl_ariths: + "rev (to_bl w) = ys \ rev (to_bl (word_succ w)) = rbl_succ ys" + "rev (to_bl w) = ys \ rev (to_bl (word_pred w)) = rbl_pred ys" + "rev (to_bl v) = ys \ rev (to_bl w) = xs \ rev (to_bl (v * w)) = rbl_mult ys xs" + "rev (to_bl v) = ys \ rev (to_bl w) = xs \ rev (to_bl (v + w)) = rbl_add ys xs" + by (auto simp: rev_swap [symmetric] word_succ_rbl word_pred_rbl word_mult_rbl word_add_rbl) + +lemma of_bl_length_less: + \(of_bl x :: 'a::len word) < 2 ^ k\ + if \length x = k\ \k < LENGTH('a)\ +proof - + from that have \length x < LENGTH('a)\ + by simp + then have \(of_bl x :: 'a::len word) < 2 ^ length x\ + apply (simp add: of_bl_eq) + apply transfer + apply (simp add: take_bit_horner_sum_bit_eq) + apply (subst length_rev [symmetric]) + apply (simp only: horner_sum_of_bool_2_less) + done + with that show ?thesis + by simp +qed + +lemma word_eq_rbl_eq: "x = y \ rev (to_bl x) = rev (to_bl y)" + by simp + +lemma bl_word_not: "to_bl (NOT w) = map Not (to_bl w)" + by transfer (simp add: bl_not_bin) + +lemma bl_word_xor: "to_bl (v XOR w) = map2 (\) (to_bl v) (to_bl w)" + by transfer (simp flip: bl_xor_bin) + +lemma bl_word_or: "to_bl (v OR w) = map2 (\) (to_bl v) (to_bl w)" + by transfer (simp flip: bl_or_bin) + +lemma bl_word_and: "to_bl (v AND w) = map2 (\) (to_bl v) (to_bl w)" + by transfer (simp flip: bl_and_bin) + +lemma bin_nth_uint': "bin_nth (uint w) n \ rev (bin_to_bl (size w) (uint w)) ! n \ n < size w" + apply (unfold word_size) + apply (safe elim!: bin_nth_uint_imp) + apply (frule bin_nth_uint_imp) + apply (fast dest!: bin_nth_bl)+ + done + +lemmas bin_nth_uint = bin_nth_uint' [unfolded word_size] + +lemma test_bit_bl: "w !! n \ rev (to_bl w) ! n \ n < size w" + by transfer (auto simp add: bin_nth_bl) + +lemma to_bl_nth: "n < size w \ to_bl w ! n = w !! (size w - Suc n)" + by (simp add: word_size rev_nth test_bit_bl) + +lemma map_bit_interval_eq: + \map (bit w) [0.. for w :: \'a::len word\ +proof (rule nth_equalityI) + show \length (map (bit w) [0.. + by simp + fix m + assume \m < length (map (bit w) [0.. + then have \m < n\ + by simp + then have \bit w m \ takefill False n (rev (to_bl w)) ! m\ + by (auto simp add: nth_takefill not_less rev_nth to_bl_nth word_size test_bit_word_eq dest: bit_imp_le_length) + with \m < n \show \map (bit w) [0.. takefill False n (rev (to_bl w)) ! m\ + by simp +qed + +lemma to_bl_unfold: + \to_bl w = rev (map (bit w) [0.. for w :: \'a::len word\ + by (simp add: map_bit_interval_eq takefill_bintrunc to_bl_def flip: bin_to_bl_def) + +lemma nth_rev_to_bl: + \rev (to_bl w) ! n \ bit w n\ + if \n < LENGTH('a)\ for w :: \'a::len word\ + using that by (simp add: to_bl_unfold) + +lemma nth_to_bl: + \to_bl w ! n \ bit w (LENGTH('a) - Suc n)\ + if \n < LENGTH('a)\ for w :: \'a::len word\ + using that by (simp add: to_bl_unfold rev_nth) + +lemma of_bl_rep_False: "of_bl (replicate n False @ bs) = of_bl bs" + by (auto simp: of_bl_def bl_to_bin_rep_F) + +lemma [code abstract]: + \uint (of_bl bs :: 'a word) = horner_sum of_bool 2 (take LENGTH('a::len) (rev bs))\ + apply (simp add: of_bl_eq flip: take_bit_horner_sum_bit_eq) + apply transfer + apply simp + done + +lemma [code]: + \to_bl w = map (bit w) (rev [0.. + for w :: \'a::len word\ + by (simp add: to_bl_unfold rev_map) + +lemma word_reverse_eq_of_bl_rev_to_bl: + \word_reverse w = of_bl (rev (to_bl w))\ + by (rule bit_word_eqI) + (auto simp add: bit_word_reverse_iff bit_of_bl_iff nth_to_bl) + +lemmas word_reverse_no_def [simp] = + word_reverse_eq_of_bl_rev_to_bl [of "numeral w"] for w + +lemma to_bl_word_rev: "to_bl (word_reverse w) = rev (to_bl w)" + by (rule nth_equalityI) (simp_all add: nth_rev_to_bl word_reverse_def word_rep_drop flip: of_bl_eq) + +lemma to_bl_n1 [simp]: "to_bl (-1::'a::len word) = replicate (LENGTH('a)) True" + apply (rule word_bl.Abs_inverse') + apply simp + apply (rule word_eqI) + apply (clarsimp simp add: word_size) + apply (auto simp add: word_bl.Abs_inverse test_bit_bl word_size) + done + +lemma rbl_word_or: "rev (to_bl (x OR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" + by (simp add: zip_rev bl_word_or rev_map) + +lemma rbl_word_and: "rev (to_bl (x AND y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" + by (simp add: zip_rev bl_word_and rev_map) + +lemma rbl_word_xor: "rev (to_bl (x XOR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" + by (simp add: zip_rev bl_word_xor rev_map) + +lemma rbl_word_not: "rev (to_bl (NOT x)) = map Not (rev (to_bl x))" + by (simp add: bl_word_not rev_map) + +lemma bshiftr1_numeral [simp]: + \bshiftr1 b (numeral w :: 'a word) = of_bl (b # butlast (bin_to_bl LENGTH('a::len) (numeral w)))\ + by (simp add: bshiftr1_eq) + +lemma bshiftr1_bl: "to_bl (bshiftr1 b w) = b # butlast (to_bl w)" + unfolding bshiftr1_eq by (rule word_bl.Abs_inverse) simp + +lemma shiftl1_of_bl: "shiftl1 (of_bl bl) = of_bl (bl @ [False])" + by transfer (simp add: bl_to_bin_append) + +lemma shiftl1_bl: "shiftl1 w = of_bl (to_bl w @ [False])" + for w :: "'a::len word" +proof - + have "shiftl1 w = shiftl1 (of_bl (to_bl w))" + by simp + also have "\ = of_bl (to_bl w @ [False])" + by (rule shiftl1_of_bl) + finally show ?thesis . +qed + +lemma bl_shiftl1: "to_bl (shiftl1 w) = tl (to_bl w) @ [False]" + for w :: "'a::len word" + by (simp add: shiftl1_bl word_rep_drop drop_Suc drop_Cons') (fast intro!: Suc_leI) + +\ \Generalized version of \bl_shiftl1\. Maybe this one should replace it?\ +lemma bl_shiftl1': "to_bl (shiftl1 w) = tl (to_bl w @ [False])" + by (simp add: shiftl1_bl word_rep_drop drop_Suc del: drop_append) + +lemma shiftr1_bl: + \shiftr1 w = of_bl (butlast (to_bl w))\ +proof (rule bit_word_eqI) + fix n + assume \n < LENGTH('a)\ + show \bit (shiftr1 w) n \ bit (of_bl (butlast (to_bl w)) :: 'a word) n\ + proof (cases \n = LENGTH('a) - 1\) + case True + then show ?thesis + by (simp add: bit_shiftr1_iff bit_of_bl_iff) + next + case False + with \n < LENGTH('a)\ + have \n < LENGTH('a) - 1\ + by simp + with \n < LENGTH('a)\ show ?thesis + by (simp add: bit_shiftr1_iff bit_of_bl_iff rev_nth nth_butlast + word_size test_bit_word_eq to_bl_nth) + qed +qed + +lemma bl_shiftr1: "to_bl (shiftr1 w) = False # butlast (to_bl w)" + for w :: "'a::len word" + by (simp add: shiftr1_bl word_rep_drop len_gt_0 [THEN Suc_leI]) + +\ \Generalized version of \bl_shiftr1\. Maybe this one should replace it?\ +lemma bl_shiftr1': "to_bl (shiftr1 w) = butlast (False # to_bl w)" + apply (rule word_bl.Abs_inverse') + apply (simp del: butlast.simps) + apply (simp add: shiftr1_bl of_bl_def) + done + +lemma bl_sshiftr1: "to_bl (sshiftr1 w) = hd (to_bl w) # butlast (to_bl w)" + for w :: "'a::len word" +proof (rule nth_equalityI) + fix n + assume \n < length (to_bl (sshiftr1 w))\ + then have \n < LENGTH('a)\ + by simp + then show \to_bl (sshiftr1 w) ! n \ (hd (to_bl w) # butlast (to_bl w)) ! n\ + apply (cases n) + apply (simp_all add: to_bl_nth word_size hd_conv_nth test_bit_eq_bit bit_sshiftr1_iff nth_butlast Suc_diff_Suc nth_to_bl) + done +qed simp + +lemma drop_shiftr: "drop n (to_bl (w >> n)) = take (size w - n) (to_bl w)" + for w :: "'a::len word" + apply (unfold shiftr_def) + apply (induct n) + prefer 2 + apply (simp add: drop_Suc bl_shiftr1 butlast_drop [symmetric]) + apply (rule butlast_take [THEN trans]) + apply (auto simp: word_size) + done + +lemma drop_sshiftr: "drop n (to_bl (w >>> n)) = take (size w - n) (to_bl w)" + for w :: "'a::len word" + apply (simp_all add: word_size sshiftr_eq) + apply (rule nth_equalityI) + apply (simp_all add: word_size nth_to_bl bit_signed_drop_bit_iff) + done + +lemma take_shiftr: "n \ size w \ take n (to_bl (w >> n)) = replicate n False" + apply (unfold shiftr_def) + apply (induct n) + prefer 2 + apply (simp add: bl_shiftr1' length_0_conv [symmetric] word_size) + apply (rule take_butlast [THEN trans]) + apply (auto simp: word_size) + done + +lemma take_sshiftr': + "n \ size w \ hd (to_bl (w >>> n)) = hd (to_bl w) \ + take n (to_bl (w >>> n)) = replicate n (hd (to_bl w))" + for w :: "'a::len word" + apply (auto simp add: sshiftr_eq hd_bl_sign_sint bin_sign_def not_le word_size sint_signed_drop_bit_eq) + apply (rule nth_equalityI) + apply (auto simp add: nth_to_bl bit_signed_drop_bit_iff bit_last_iff) + apply (rule nth_equalityI) + apply (auto simp add: nth_to_bl bit_signed_drop_bit_iff bit_last_iff) + done + +lemmas hd_sshiftr = take_sshiftr' [THEN conjunct1] +lemmas take_sshiftr = take_sshiftr' [THEN conjunct2] + +lemma atd_lem: "take n xs = t \ drop n xs = d \ xs = t @ d" + by (auto intro: append_take_drop_id [symmetric]) + +lemmas bl_shiftr = atd_lem [OF take_shiftr drop_shiftr] +lemmas bl_sshiftr = atd_lem [OF take_sshiftr drop_sshiftr] + +lemma shiftl_of_bl: "of_bl bl << n = of_bl (bl @ replicate n False)" + by (induct n) (auto simp: shiftl_def shiftl1_of_bl replicate_app_Cons_same) + +lemma shiftl_bl: "w << n = of_bl (to_bl w @ replicate n False)" + for w :: "'a::len word" +proof - + have "w << n = of_bl (to_bl w) << n" + by simp + also have "\ = of_bl (to_bl w @ replicate n False)" + by (rule shiftl_of_bl) + finally show ?thesis . +qed + +lemma bl_shiftl: "to_bl (w << n) = drop n (to_bl w) @ replicate (min (size w) n) False" + by (simp add: shiftl_bl word_rep_drop word_size) + +lemma shiftr1_bl_of: + "length bl \ LENGTH('a) \ + shiftr1 (of_bl bl::'a::len word) = of_bl (butlast bl)" + by transfer (simp add: butlast_rest_bl2bin trunc_bl2bin) + +lemma shiftr_bl_of: + "length bl \ LENGTH('a) \ + (of_bl bl::'a::len word) >> n = of_bl (take (length bl - n) bl)" + apply (unfold shiftr_def) + apply (induct n) + apply clarsimp + apply clarsimp + apply (subst shiftr1_bl_of) + apply simp + apply (simp add: butlast_take) + done + +lemma shiftr_bl: "x >> n \ of_bl (take (LENGTH('a) - n) (to_bl x))" + for x :: "'a::len word" + using shiftr_bl_of [where 'a='a, of "to_bl x"] by simp + +lemma aligned_bl_add_size [OF refl]: + "size x - n = m \ n \ size x \ drop m (to_bl x) = replicate n False \ + take m (to_bl y) = replicate m False \ + to_bl (x + y) = take m (to_bl x) @ drop m (to_bl y)" for x :: \'a::len word\ + apply (subgoal_tac "x AND y = 0") + prefer 2 + apply (rule word_bl.Rep_eqD) + apply (simp add: bl_word_and) + apply (rule align_lem_and [THEN trans]) + apply (simp_all add: word_size)[5] + apply simp + apply (subst word_plus_and_or [symmetric]) + apply (simp add : bl_word_or) + apply (rule align_lem_or) + apply (simp_all add: word_size) + done + +lemma mask_bl: "mask n = of_bl (replicate n True)" + by (auto simp add : test_bit_of_bl word_size intro: word_eqI) + +lemma bl_and_mask': + "to_bl (w AND mask n :: 'a::len word) = + replicate (LENGTH('a) - n) False @ + drop (LENGTH('a) - n) (to_bl w)" + apply (rule nth_equalityI) + apply simp + apply (clarsimp simp add: to_bl_nth word_size) + apply (simp add: word_size word_ops_nth_size) + apply (auto simp add: word_size test_bit_bl nth_append rev_nth) + done + +lemma slice1_eq_of_bl: + \(slice1 n w :: 'b::len word) = of_bl (takefill False n (to_bl w))\ + for w :: \'a::len word\ +proof (rule bit_word_eqI) + fix m + assume \m < LENGTH('b)\ + show \bit (slice1 n w :: 'b::len word) m \ bit (of_bl (takefill False n (to_bl w)) :: 'b word) m\ + by (cases \m \ n\; cases \LENGTH('a) \ n\) + (auto simp add: bit_slice1_iff bit_of_bl_iff not_less rev_nth not_le nth_takefill nth_to_bl algebra_simps) +qed + +lemma slice1_no_bin [simp]: + "slice1 n (numeral w :: 'b word) = of_bl (takefill False n (bin_to_bl (LENGTH('b::len)) (numeral w)))" + by (simp add: slice1_eq_of_bl) (* TODO: neg_numeral *) + +lemma slice_no_bin [simp]: + "slice n (numeral w :: 'b word) = of_bl (takefill False (LENGTH('b::len) - n) + (bin_to_bl (LENGTH('b::len)) (numeral w)))" + by (simp add: slice_def) (* TODO: neg_numeral *) + +lemma slice_take': "slice n w = of_bl (take (size w - n) (to_bl w))" + by (simp add: slice_def word_size slice1_eq_of_bl takefill_alt) + +lemmas slice_take = slice_take' [unfolded word_size] + +\ \shiftr to a word of the same size is just slice, + slice is just shiftr then ucast\ +lemmas shiftr_slice = trans [OF shiftr_bl [THEN meta_eq_to_obj_eq] slice_take [symmetric]] + +lemma slice1_down_alt': + "sl = slice1 n w \ fs = size sl \ fs + k = n \ + to_bl sl = takefill False fs (drop k (to_bl w))" + apply (simp add: slice1_eq_of_bl) + apply transfer + apply (simp add: bl_bin_bl_rep_drop) + using drop_takefill + apply force + done + +lemma slice1_up_alt': + "sl = slice1 n w \ fs = size sl \ fs = n + k \ + to_bl sl = takefill False fs (replicate k False @ (to_bl w))" + apply (simp add: slice1_eq_of_bl) + apply transfer + apply (simp add: bl_bin_bl_rep_drop flip: takefill_append) + apply (metis diff_add_inverse) + done + +lemmas sd1 = slice1_down_alt' [OF refl refl, unfolded word_size] +lemmas su1 = slice1_up_alt' [OF refl refl, unfolded word_size] +lemmas slice1_down_alt = le_add_diff_inverse [THEN sd1] +lemmas slice1_up_alts = + le_add_diff_inverse [symmetric, THEN su1] + le_add_diff_inverse2 [symmetric, THEN su1] + +lemma slice1_tf_tf': + "to_bl (slice1 n w :: 'a::len word) = + rev (takefill False (LENGTH('a)) (rev (takefill False n (to_bl w))))" + unfolding slice1_eq_of_bl by (rule word_rev_tf) + +lemmas slice1_tf_tf = slice1_tf_tf' [THEN word_bl.Rep_inverse', symmetric] + +lemma revcast_eq_of_bl: + \(revcast w :: 'b::len word) = of_bl (takefill False (LENGTH('b)) (to_bl w))\ + for w :: \'a::len word\ + by (simp add: revcast_def slice1_eq_of_bl) + +lemmas revcast_no_def [simp] = revcast_eq_of_bl [where w="numeral w", unfolded word_size] for w + +lemma to_bl_revcast: + "to_bl (revcast w :: 'a::len word) = takefill False (LENGTH('a)) (to_bl w)" + apply (rule nth_equalityI) + apply simp + apply (cases \LENGTH('a) \ LENGTH('b)\) + apply (auto simp add: nth_to_bl nth_takefill bit_revcast_iff) + done + +lemma word_cat_bl: "word_cat a b = of_bl (to_bl a @ to_bl b)" + apply (rule bit_word_eqI) + apply (simp add: bit_word_cat_iff bit_of_bl_iff nth_append not_less nth_rev_to_bl) + apply (meson bit_word.rep_eq less_diff_conv2 nth_rev_to_bl) + done + +lemma of_bl_append: + "(of_bl (xs @ ys) :: 'a::len word) = of_bl xs * 2^(length ys) + of_bl ys" + apply transfer + apply (simp add: bl_to_bin_app_cat bin_cat_num) + done + +lemma of_bl_False [simp]: "of_bl (False#xs) = of_bl xs" + by (rule word_eqI) (auto simp: test_bit_of_bl nth_append) + +lemma of_bl_True [simp]: "(of_bl (True # xs) :: 'a::len word) = 2^length xs + of_bl xs" + by (subst of_bl_append [where xs="[True]", simplified]) (simp add: word_1_bl) + +lemma of_bl_Cons: "of_bl (x#xs) = of_bool x * 2^length xs + of_bl xs" + by (cases x) simp_all + +lemma word_split_bl': + "std = size c - size b \ (word_split c = (a, b)) \ + (a = of_bl (take std (to_bl c)) \ b = of_bl (drop std (to_bl c)))" + apply (simp add: word_split_def) + apply transfer + apply (cases \LENGTH('b) \ LENGTH('a)\) + apply (auto simp add: drop_bit_take_bit drop_bin2bl bin_to_bl_drop_bit [symmetric, of \LENGTH('a)\ \LENGTH('a) - LENGTH('b)\ \LENGTH('b)\] min_absorb2) + done + +lemma word_split_bl: "std = size c - size b \ + (a = of_bl (take std (to_bl c)) \ b = of_bl (drop std (to_bl c))) \ + word_split c = (a, b)" + apply (rule iffI) + defer + apply (erule (1) word_split_bl') + apply (case_tac "word_split c") + apply (auto simp add: word_size) + apply (frule word_split_bl' [rotated]) + apply (auto simp add: word_size) + done + +lemma word_split_bl_eq: + "(word_split c :: ('c::len word \ 'd::len word)) = + (of_bl (take (LENGTH('a::len) - LENGTH('d::len)) (to_bl c)), + of_bl (drop (LENGTH('a) - LENGTH('d)) (to_bl c)))" + for c :: "'a::len word" + apply (rule word_split_bl [THEN iffD1]) + apply (unfold word_size) + apply (rule refl conjI)+ + done + +lemma word_rcat_bl: + \word_rcat wl = of_bl (concat (map to_bl wl))\ +proof - + define ws where \ws = rev wl\ + moreover have \word_rcat (rev ws) = of_bl (concat (map to_bl (rev ws)))\ + apply (simp add: word_rcat_def of_bl_eq rev_concat rev_map comp_def rev_to_bl_eq flip: horner_sum_of_bool_2_concat) + apply transfer + apply simp + done + ultimately show ?thesis + by simp +qed + +lemma size_rcat_lem': "size (concat (map to_bl wl)) = length wl * size (hd wl)" + by (induct wl) (auto simp: word_size) + +lemmas size_rcat_lem = size_rcat_lem' [unfolded word_size] + +lemma nth_rcat_lem: + "n < length (wl::'a word list) * LENGTH('a::len) \ + rev (concat (map to_bl wl)) ! n = + rev (to_bl (rev wl ! (n div LENGTH('a)))) ! (n mod LENGTH('a))" + apply (induct wl) + apply clarsimp + apply (clarsimp simp add : nth_append size_rcat_lem) + apply (simp flip: mult_Suc minus_div_mult_eq_mod add: less_Suc_eq_le not_less) + apply (metis (no_types, lifting) diff_is_0_eq div_le_mono len_not_eq_0 less_Suc_eq less_mult_imp_div_less nonzero_mult_div_cancel_right not_le nth_Cons_0) + done + +lemma foldl_eq_foldr: "foldl (+) x xs = foldr (+) (x # xs) 0" + for x :: "'a::comm_monoid_add" + by (induct xs arbitrary: x) (auto simp: add.assoc) + +lemmas word_cat_bl_no_bin [simp] = + word_cat_bl [where a="numeral a" and b="numeral b", unfolded to_bl_numeral] + for a b (* FIXME: negative numerals, 0 and 1 *) + +lemmas word_split_bl_no_bin [simp] = + word_split_bl_eq [where c="numeral c", unfolded to_bl_numeral] for c + +lemmas word_rot_defs = word_roti_eq_word_rotr_word_rotl word_rotr_eq word_rotl_eq + +lemma to_bl_rotl: "to_bl (word_rotl n w) = rotate n (to_bl w)" + by (simp add: word_rotl_eq to_bl_use_of_bl) + +lemmas blrs0 = rotate_eqs [THEN to_bl_rotl [THEN trans]] + +lemmas word_rotl_eqs = + blrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotl [symmetric]] + +lemma to_bl_rotr: "to_bl (word_rotr n w) = rotater n (to_bl w)" + by (simp add: word_rotr_eq to_bl_use_of_bl) + +lemmas brrs0 = rotater_eqs [THEN to_bl_rotr [THEN trans]] + +lemmas word_rotr_eqs = + brrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotr [symmetric]] + +declare word_rotr_eqs (1) [simp] +declare word_rotl_eqs (1) [simp] + +lemmas abl_cong = arg_cong [where f = "of_bl"] + +locale word_rotate +begin + +lemmas word_rot_defs' = to_bl_rotl to_bl_rotr + +lemmas blwl_syms [symmetric] = bl_word_not bl_word_and bl_word_or bl_word_xor + +lemmas lbl_lbl = trans [OF word_bl_Rep' word_bl_Rep' [symmetric]] + +lemmas ths_map2 [OF lbl_lbl] = rotate_map2 rotater_map2 + +lemmas ths_map [where xs = "to_bl v"] = rotate_map rotater_map for v + +lemmas th1s [simplified word_rot_defs' [symmetric]] = ths_map2 ths_map + +end + +lemmas bl_word_rotl_dt = trans [OF to_bl_rotl rotate_drop_take, + simplified word_bl_Rep'] + +lemmas bl_word_rotr_dt = trans [OF to_bl_rotr rotater_drop_take, + simplified word_bl_Rep'] + +lemma bl_word_roti_dt': + "n = nat ((- i) mod int (size (w :: 'a::len word))) \ + to_bl (word_roti i w) = drop n (to_bl w) @ take n (to_bl w)" + apply (unfold word_roti_eq_word_rotr_word_rotl) + apply (simp add: bl_word_rotl_dt bl_word_rotr_dt word_size) + apply safe + apply (simp add: zmod_zminus1_eq_if) + apply safe + apply (simp add: nat_mult_distrib) + apply (simp add: nat_diff_distrib [OF pos_mod_sign pos_mod_conj + [THEN conjunct2, THEN order_less_imp_le]] + nat_mod_distrib) + apply (simp add: nat_mod_distrib) + done + +lemmas bl_word_roti_dt = bl_word_roti_dt' [unfolded word_size] + +lemmas word_rotl_dt = bl_word_rotl_dt [THEN word_bl.Rep_inverse' [symmetric]] +lemmas word_rotr_dt = bl_word_rotr_dt [THEN word_bl.Rep_inverse' [symmetric]] +lemmas word_roti_dt = bl_word_roti_dt [THEN word_bl.Rep_inverse' [symmetric]] + +lemmas word_rotr_dt_no_bin' [simp] = + word_rotr_dt [where w="numeral w", unfolded to_bl_numeral] for w + (* FIXME: negative numerals, 0 and 1 *) + +lemmas word_rotl_dt_no_bin' [simp] = + word_rotl_dt [where w="numeral w", unfolded to_bl_numeral] for w + (* FIXME: negative numerals, 0 and 1 *) + +lemma max_word_bl: "to_bl (max_word::'a::len word) = replicate LENGTH('a) True" + by (fact to_bl_n1) + +lemma to_bl_mask: + "to_bl (mask n :: 'a::len word) = + replicate (LENGTH('a) - n) False @ + replicate (min (LENGTH('a)) n) True" + by (simp add: mask_bl word_rep_drop min_def) + +lemma map_replicate_True: + "n = length xs \ + map (\(x,y). x \ y) (zip xs (replicate n True)) = xs" + by (induct xs arbitrary: n) auto + +lemma map_replicate_False: + "n = length xs \ map (\(x,y). x \ y) + (zip xs (replicate n False)) = replicate n False" + by (induct xs arbitrary: n) auto + +lemma bl_and_mask: + fixes w :: "'a::len word" + and n :: nat + defines "n' \ LENGTH('a) - n" + shows "to_bl (w AND mask n) = replicate n' False @ drop n' (to_bl w)" +proof - + note [simp] = map_replicate_True map_replicate_False + have "to_bl (w AND mask n) = map2 (\) (to_bl w) (to_bl (mask n::'a::len word))" + by (simp add: bl_word_and) + also have "to_bl w = take n' (to_bl w) @ drop n' (to_bl w)" + by simp + also have "map2 (\) \ (to_bl (mask n::'a::len word)) = + replicate n' False @ drop n' (to_bl w)" + unfolding to_bl_mask n'_def by (subst zip_append) auto + finally show ?thesis . +qed + +lemma drop_rev_takefill: + "length xs \ n \ + drop (n - length xs) (rev (takefill False n (rev xs))) = xs" + by (simp add: takefill_alt rev_take) + +declare bin_to_bl_def [simp] + +end diff --git a/thys/Word_Lib/Signed_Words.thy b/thys/Word_Lib/Signed_Words.thy --- a/thys/Word_Lib/Signed_Words.thy +++ b/thys/Word_Lib/Signed_Words.thy @@ -1,40 +1,40 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Signed Words" theory Signed_Words -imports "HOL-Word.Word" + imports "HOL-Library.Word" begin text \Signed words as separate (isomorphic) word length class. Useful for tagging words in C.\ typedef ('a::len0) signed = "UNIV :: 'a set" .. lemma card_signed [simp]: "CARD (('a::len0) signed) = CARD('a)" unfolding type_definition.card [OF type_definition_signed] by simp instantiation signed :: (len0) len0 begin definition len_signed [simp]: "len_of (x::'a::len0 signed itself) = LENGTH('a)" instance .. end instance signed :: (len) len by (intro_classes, simp) type_synonym 'a sword = "'a signed word" type_synonym sword8 = "8 sword" type_synonym sword16 = "16 sword" type_synonym sword32 = "32 sword" type_synonym sword64 = "64 sword" end diff --git a/thys/Word_Lib/Traditional_Syntax.thy b/thys/Word_Lib/Traditional_Syntax.thy new file mode 100644 --- /dev/null +++ b/thys/Word_Lib/Traditional_Syntax.thy @@ -0,0 +1,534 @@ +(* Author: Jeremy Dawson, NICTA +*) + +section \Operation variants with traditional syntax\ + +theory Traditional_Syntax + imports "HOL-Library.Word" +begin + +class semiring_bit_syntax = semiring_bit_shifts +begin + +definition test_bit :: \'a \ nat \ bool\ (infixl "!!" 100) + where test_bit_eq_bit: \test_bit = bit\ + +definition shiftl :: \'a \ nat \ 'a\ (infixl "<<" 55) + where shiftl_eq_push_bit: \a << n = push_bit n a\ + +definition shiftr :: \'a \ nat \ 'a\ (infixl ">>" 55) + where shiftr_eq_drop_bit: \a >> n = drop_bit n a\ + +end + +instance word :: (len) semiring_bit_syntax .. + +context + includes lifting_syntax +begin + +lemma test_bit_word_transfer [transfer_rule]: + \(pcr_word ===> (=)) (\k n. n < LENGTH('a) \ bit k n) (test_bit :: 'a::len word \ _)\ + by (unfold test_bit_eq_bit) transfer_prover + +lemma shiftl_word_transfer [transfer_rule]: + \(pcr_word ===> (=) ===> pcr_word) (\k n. push_bit n k) shiftl\ + by (unfold shiftl_eq_push_bit) transfer_prover + +lemma shiftr_word_transfer [transfer_rule]: + \(pcr_word ===> (=) ===> pcr_word) (\k n. (drop_bit n \ take_bit LENGTH('a)) k) (shiftr :: 'a::len word \ _)\ + by (unfold shiftr_eq_drop_bit) transfer_prover + +end + +lemma test_bit_word_eq: + \test_bit = (bit :: 'a::len word \ _)\ + by (fact test_bit_eq_bit) + +lemma shiftl_word_eq: + \w << n = push_bit n w\ for w :: \'a::len word\ + by (fact shiftl_eq_push_bit) + +lemma shiftr_word_eq: + \w >> n = drop_bit n w\ for w :: \'a::len word\ + by (fact shiftr_eq_drop_bit) + +lemma test_bit_eq_iff: "test_bit u = test_bit v \ u = v" + for u v :: "'a::len word" + by (simp add: bit_eq_iff test_bit_eq_bit fun_eq_iff) + +lemma test_bit_size: "w !! n \ n < size w" + for w :: "'a::len word" + by transfer simp + +lemma word_eq_iff: "x = y \ (\n?P \ ?Q\) + for x y :: "'a::len word" + by transfer (auto simp add: bit_eq_iff bit_take_bit_iff) + +lemma word_eqI: "(\n. n < size u \ u !! n = v !! n) \ u = v" + for u :: "'a::len word" + by (simp add: word_size word_eq_iff) + +lemma word_eqD: "u = v \ u !! x = v !! x" + for u v :: "'a::len word" + by simp + +lemma test_bit_bin': "w !! n \ n < size w \ bit (uint w) n" + by transfer (simp add: bit_take_bit_iff) + +lemmas test_bit_bin = test_bit_bin' [unfolded word_size] + +lemma word_test_bit_def: + \test_bit a = bit (uint a)\ + by transfer (simp add: fun_eq_iff bit_take_bit_iff) + +lemmas test_bit_def' = word_test_bit_def [THEN fun_cong] + +lemma word_test_bit_transfer [transfer_rule]: + "(rel_fun pcr_word (rel_fun (=) (=))) + (\x n. n < LENGTH('a) \ bit x n) (test_bit :: 'a::len word \ _)" + by (simp only: test_bit_eq_bit) transfer_prover + +lemma test_bit_wi [simp]: + "(word_of_int x :: 'a::len word) !! n \ n < LENGTH('a) \ bit x n" + by transfer simp + +lemma word_ops_nth_size: + "n < size x \ + (x OR y) !! n = (x !! n | y !! n) \ + (x AND y) !! n = (x !! n \ y !! n) \ + (x XOR y) !! n = (x !! n \ y !! n) \ + (NOT x) !! n = (\ x !! n)" + for x :: "'a::len word" + by transfer (simp add: bit_or_iff bit_and_iff bit_xor_iff bit_not_iff) + +lemma word_ao_nth: + "(x OR y) !! n = (x !! n | y !! n) \ + (x AND y) !! n = (x !! n \ y !! n)" + for x :: "'a::len word" + by transfer (auto simp add: bit_or_iff bit_and_iff) + +lemmas msb0 = len_gt_0 [THEN diff_Suc_less, THEN word_ops_nth_size [unfolded word_size]] +lemmas msb1 = msb0 [where i = 0] + +lemma test_bit_numeral [simp]: + "(numeral w :: 'a::len word) !! n \ + n < LENGTH('a) \ bit (numeral w :: int) n" + by transfer (rule refl) + +lemma test_bit_neg_numeral [simp]: + "(- numeral w :: 'a::len word) !! n \ + n < LENGTH('a) \ bit (- numeral w :: int) n" + by transfer (rule refl) + +lemma test_bit_1 [simp]: "(1 :: 'a::len word) !! n \ n = 0" + by transfer (auto simp add: bit_1_iff) + +lemma nth_0 [simp]: "\ (0 :: 'a::len word) !! n" + by transfer simp + +lemma nth_minus1 [simp]: "(-1 :: 'a::len word) !! n \ n < LENGTH('a)" + by transfer simp + +lemma shiftl1_code [code]: + \shiftl1 w = push_bit 1 w\ + by transfer (simp add: ac_simps) + +lemma uint_shiftr_eq: + \uint (w >> n) = uint w div 2 ^ n\ + by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit min_def le_less less_diff_conv) + +lemma shiftr1_code [code]: + \shiftr1 w = drop_bit 1 w\ + by transfer (simp add: drop_bit_Suc) + +lemma shiftl_def: + \w << n = (shiftl1 ^^ n) w\ +proof - + have \push_bit n = (((*) 2 ^^ n) :: int \ int)\ for n + by (induction n) (simp_all add: fun_eq_iff funpow_swap1, simp add: ac_simps) + then show ?thesis + by transfer simp +qed + +lemma shiftr_def: + \w >> n = (shiftr1 ^^ n) w\ +proof - + have \shiftr1 ^^ n = (drop_bit n :: 'a word \ 'a word)\ + apply (induction n) + apply simp + apply (simp only: shiftr1_eq_div_2 [abs_def] drop_bit_eq_div [abs_def] funpow_Suc_right) + apply (use div_exp_eq [of _ 1, where ?'a = \'a word\] in simp) + done + then show ?thesis + by (simp add: shiftr_eq_drop_bit) +qed + +lemma bit_shiftl_word_iff: + \bit (w << m) n \ m \ n \ n < LENGTH('a) \ bit w (n - m)\ + for w :: \'a::len word\ + by (simp add: shiftl_word_eq bit_push_bit_iff exp_eq_zero_iff not_le) + +lemma bit_shiftr_word_iff: + \bit (w >> m) n \ bit w (m + n)\ + for w :: \'a::len word\ + by (simp add: shiftr_word_eq bit_drop_bit_eq) + +lift_definition sshiftr :: \'a::len word \ nat \ 'a word\ (infixl \>>>\ 55) + is \\k n. take_bit LENGTH('a) (drop_bit n (signed_take_bit (LENGTH('a) - Suc 0) k))\ + by (simp flip: signed_take_bit_decr_length_iff) + +lemma sshiftr_eq [code]: + \w >>> n = signed_drop_bit n w\ + by transfer simp + +lemma sshiftr_eq_funpow_sshiftr1: + \w >>> n = (sshiftr1 ^^ n) w\ + apply (rule sym) + apply (simp add: sshiftr1_eq_signed_drop_bit_Suc_0 sshiftr_eq) + apply (induction n) + apply simp_all + done + +lemma uint_sshiftr_eq: + \uint (w >>> n) = take_bit LENGTH('a) (sint w div 2 ^ n)\ + for w :: \'a::len word\ + by transfer (simp flip: drop_bit_eq_div) + +lemma sshift1_code [code]: + \sshiftr1 w = signed_drop_bit 1 w\ + by transfer (simp add: drop_bit_Suc) + +lemma sshiftr_0 [simp]: "0 >>> n = 0" + by transfer simp + +lemma sshiftr_n1 [simp]: "-1 >>> n = -1" + by transfer simp + +lemma bit_sshiftr_word_iff: + \bit (w >>> m) n \ bit w (if LENGTH('a) - m \ n \ n < LENGTH('a) then LENGTH('a) - 1 else (m + n))\ + for w :: \'a::len word\ + apply transfer + apply (auto simp add: bit_take_bit_iff bit_drop_bit_eq bit_signed_take_bit_iff min_def not_le simp flip: bit_Suc) + using le_less_Suc_eq apply fastforce + using le_less_Suc_eq apply fastforce + done + +lemma nth_sshiftr : + "(w >>> m) !! n = + (n < size w \ (if n + m \ size w then w !! (size w - 1) else w !! (n + m)))" + apply transfer + apply (auto simp add: bit_take_bit_iff bit_drop_bit_eq bit_signed_take_bit_iff min_def not_le ac_simps) + using le_less_Suc_eq apply fastforce + using le_less_Suc_eq apply fastforce + done + +lemma sshiftr_numeral [simp]: + \(numeral k >>> numeral n :: 'a::len word) = + word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (numeral k)))\ + apply (rule word_eqI) + apply (cases \LENGTH('a)\) + apply (simp_all add: word_size bit_drop_bit_eq nth_sshiftr bit_signed_take_bit_iff min_def not_le not_less less_Suc_eq_le ac_simps) + done + +lemma revcast_down_us [OF refl]: + "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (w >>> n)" + for w :: "'a::len word" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_revcast_iff bit_ucast_iff bit_sshiftr_word_iff ac_simps) + done + +lemma revcast_down_ss [OF refl]: + "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (w >>> n)" + for w :: "'a::len word" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_revcast_iff bit_word_scast_iff bit_sshiftr_word_iff ac_simps) + done + +lemma sshiftr_div_2n: "sint (w >>> n) = sint w div 2 ^ n" + using sint_signed_drop_bit_eq [of n w] + by (simp add: drop_bit_eq_div sshiftr_eq) + +lemmas lsb0 = len_gt_0 [THEN word_ops_nth_size [unfolded word_size]] + +lemma nth_sint: + fixes w :: "'a::len word" + defines "l \ LENGTH('a)" + shows "bit (sint w) n = (if n < l - 1 then w !! n else w !! (l - 1))" + unfolding sint_uint l_def + by (auto simp: bit_signed_take_bit_iff word_test_bit_def not_less min_def) + +lemma test_bit_2p: "(word_of_int (2 ^ n)::'a::len word) !! m \ m = n \ m < LENGTH('a)" + by transfer (auto simp add: bit_exp_iff) + +lemma nth_w2p: "((2::'a::len word) ^ n) !! m \ m = n \ m < LENGTH('a::len)" + by transfer (auto simp add: bit_exp_iff) + +lemma bang_is_le: "x !! m \ 2 ^ m \ x" + for x :: "'a::len word" + apply (rule xtrans(3)) + apply (rule_tac [2] y = "x" in le_word_or2) + apply (rule word_eqI) + apply (auto simp add: word_ao_nth nth_w2p word_size) + done + +lemma mask_eq: + \mask n = (1 << n) - (1 :: 'a::len word)\ + by transfer (simp add: mask_eq_exp_minus_1 push_bit_of_1) + +lemma nth_ucast: "(ucast w::'a::len word) !! n = (w !! n \ n < LENGTH('a))" + by transfer (simp add: bit_take_bit_iff ac_simps) + +lemma shiftl_0 [simp]: "(0::'a::len word) << n = 0" + by transfer simp + +lemma shiftr_0 [simp]: "(0::'a::len word) >> n = 0" + by transfer simp + +lemma nth_shiftl1: "shiftl1 w !! n \ n < size w \ n > 0 \ w !! (n - 1)" + by transfer (auto simp add: bit_double_iff) + +lemma nth_shiftl': "(w << m) !! n \ n < size w \ n >= m \ w !! (n - m)" + for w :: "'a::len word" + by transfer (auto simp add: bit_push_bit_iff) + +lemmas nth_shiftl = nth_shiftl' [unfolded word_size] + +lemma nth_shiftr1: "shiftr1 w !! n = w !! Suc n" + by transfer (auto simp add: bit_take_bit_iff simp flip: bit_Suc) + +lemma nth_shiftr: "(w >> m) !! n = w !! (n + m)" + for w :: "'a::len word" + apply (unfold shiftr_def) + apply (induct "m" arbitrary: n) + apply (auto simp add: nth_shiftr1) + done + +lemma nth_sshiftr1: "sshiftr1 w !! n = (if n = size w - 1 then w !! n else w !! Suc n)" + apply transfer + apply (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def simp flip: bit_Suc) + using le_less_Suc_eq apply fastforce + using le_less_Suc_eq apply fastforce + done + +lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" + by (fact uint_shiftr_eq) + +lemma shiftl_rev: "shiftl w n = word_reverse (shiftr (word_reverse w) n)" + by (induct n) (auto simp add: shiftl_def shiftr_def shiftl1_rev) + +lemma rev_shiftl: "word_reverse w << n = word_reverse (w >> n)" + by (simp add: shiftl_rev) + +lemma shiftr_rev: "w >> n = word_reverse (word_reverse w << n)" + by (simp add: rev_shiftl) + +lemma rev_shiftr: "word_reverse w >> n = word_reverse (w << n)" + by (simp add: shiftr_rev) + +lemma shiftl_numeral [simp]: + \numeral k << numeral l = (push_bit (numeral l) (numeral k) :: 'a::len word)\ + by (fact shiftl_word_eq) + +lemma shiftl_zero_size: "size x \ n \ x << n = 0" + for x :: "'a::len word" + apply transfer + apply (simp add: take_bit_push_bit) + done + +lemma shiftl_t2n: "shiftl w n = 2 ^ n * w" + for w :: "'a::len word" + by (induct n) (auto simp: shiftl_def shiftl1_2t) + +lemma shiftr_numeral [simp]: + \(numeral k >> numeral n :: 'a::len word) = drop_bit (numeral n) (numeral k)\ + by (fact shiftr_word_eq) + +lemma nth_mask [simp]: + \(mask n :: 'a::len word) !! i \ i < n \ i < size (mask n :: 'a word)\ + by (auto simp add: test_bit_word_eq word_size Word.bit_mask_iff) + +lemma slice_shiftr: "slice n w = ucast (w >> n)" + apply (rule bit_word_eqI) + apply (cases \n \ LENGTH('b)\) + apply (auto simp add: bit_slice_iff bit_ucast_iff bit_shiftr_word_iff ac_simps + dest: bit_imp_le_length) + done + +lemma nth_slice: "(slice n w :: 'a::len word) !! m = (w !! (m + n) \ m < LENGTH('a))" + by (simp add: slice_shiftr nth_ucast nth_shiftr) + +lemma revcast_down_uu [OF refl]: + "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (w >> n)" + for w :: "'a::len word" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_revcast_iff bit_ucast_iff bit_shiftr_word_iff ac_simps) + done + +lemma revcast_down_su [OF refl]: + "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (w >> n)" + for w :: "'a::len word" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_revcast_iff bit_word_scast_iff bit_shiftr_word_iff ac_simps) + done + +lemma cast_down_rev [OF refl]: + "uc = ucast \ source_size uc = target_size uc + n \ uc w = revcast (w << n)" + for w :: "'a::len word" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_revcast_iff bit_word_ucast_iff bit_shiftl_word_iff) + done + +lemma revcast_up [OF refl]: + "rc = revcast \ source_size rc + n = target_size rc \ + rc w = (ucast w :: 'a::len word) << n" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_revcast_iff bit_word_ucast_iff bit_shiftl_word_iff) + apply auto + apply (metis add.commute add_diff_cancel_right) + apply (metis diff_add_inverse2 diff_diff_add) + done + +lemmas rc1 = revcast_up [THEN + revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] +lemmas rc2 = revcast_down_uu [THEN + revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] + +lemmas ucast_up = + rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]] +lemmas ucast_down = + rc2 [simplified rev_shiftr revcast_ucast [symmetric]] + +\ \problem posed by TPHOLs referee: + criterion for overflow of addition of signed integers\ + +lemma sofl_test: + \sint x + sint y = sint (x + y) \ + (x + y XOR x) AND (x + y XOR y) >> (size x - 1) = 0\ + for x y :: \'a::len word\ +proof - + obtain n where n: \LENGTH('a) = Suc n\ + by (cases \LENGTH('a)\) simp_all + have *: \sint x + sint y + 2 ^ Suc n > signed_take_bit n (sint x + sint y) \ sint x + sint y \ - (2 ^ n)\ + \signed_take_bit n (sint x + sint y) > sint x + sint y - 2 ^ Suc n \ 2 ^ n > sint x + sint y\ + using signed_take_bit_int_greater_eq [of \sint x + sint y\ n] signed_take_bit_int_less_eq [of n \sint x + sint y\] + by (auto intro: ccontr) + have \sint x + sint y = sint (x + y) \ + (sint (x + y) < 0 \ sint x < 0) \ + (sint (x + y) < 0 \ sint y < 0)\ + using sint_less [of x] sint_greater_eq [of x] sint_less [of y] sint_greater_eq [of y] + signed_take_bit_int_eq_self [of \LENGTH('a) - 1\ \sint x + sint y\] + apply (auto simp add: not_less) + apply (unfold sint_word_ariths) + apply (subst signed_take_bit_int_eq_self) + prefer 4 + apply (subst signed_take_bit_int_eq_self) + prefer 7 + apply (subst signed_take_bit_int_eq_self) + prefer 10 + apply (subst signed_take_bit_int_eq_self) + apply (auto simp add: signed_take_bit_int_eq_self signed_take_bit_eq_take_bit_minus take_bit_Suc_from_most n not_less intro!: *) + done + then show ?thesis + apply (simp only: One_nat_def word_size shiftr_word_eq drop_bit_eq_zero_iff_not_bit_last bit_and_iff bit_xor_iff) + apply (simp add: bit_last_iff) + done +qed + +lemma shiftr_zero_size: "size x \ n \ x >> n = 0" + for x :: "'a :: len word" + by (rule word_eqI) (auto simp add: nth_shiftr dest: test_bit_size) + +lemma test_bit_cat [OF refl]: + "wc = word_cat a b \ wc !! n = (n < size wc \ + (if n < size b then b !! n else a !! (n - size b)))" + apply (simp add: word_size not_less; transfer) + apply (auto simp add: bit_concat_bit_iff bit_take_bit_iff) + done + +\ \keep quantifiers for use in simplification\ +lemma test_bit_split': + "word_split c = (a, b) \ + (\n m. + b !! n = (n < size b \ c !! n) \ + a !! m = (m < size a \ c !! (m + size b)))" + by (auto simp add: word_split_bin' test_bit_bin bit_unsigned_iff word_size + bit_drop_bit_eq ac_simps exp_eq_zero_iff + dest: bit_imp_le_length) + +lemma test_bit_split: + "word_split c = (a, b) \ + (\n::nat. b !! n \ n < size b \ c !! n) \ + (\m::nat. a !! m \ m < size a \ c !! (m + size b))" + by (simp add: test_bit_split') + +lemma test_bit_split_eq: + "word_split c = (a, b) \ + ((\n::nat. b !! n = (n < size b \ c !! n)) \ + (\m::nat. a !! m = (m < size a \ c !! (m + size b))))" + apply (rule_tac iffI) + apply (rule_tac conjI) + apply (erule test_bit_split [THEN conjunct1]) + apply (erule test_bit_split [THEN conjunct2]) + apply (case_tac "word_split c") + apply (frule test_bit_split) + apply (erule trans) + apply (fastforce intro!: word_eqI simp add: word_size) + done + +lemma test_bit_rcat: + "sw = size (hd wl) \ rc = word_rcat wl \ rc !! n = + (n < size rc \ n div sw < size wl \ (rev wl) ! (n div sw) !! (n mod sw))" + for wl :: "'a::len word list" + by (simp add: word_size word_rcat_def foldl_map rev_map bit_horner_sum_uint_exp_iff) + (simp add: test_bit_eq_bit) + +lemmas test_bit_cong = arg_cong [where f = "test_bit", THEN fun_cong] + +lemma max_test_bit: "(max_word::'a::len word) !! n \ n < LENGTH('a)" + by (fact nth_minus1) + +lemma shiftr_x_0 [iff]: "x >> 0 = x" + for x :: "'a::len word" + by transfer simp + +lemma shiftl_x_0 [simp]: "x << 0 = x" + for x :: "'a::len word" + by (simp add: shiftl_t2n) + +lemma shiftl_1 [simp]: "(1::'a::len word) << n = 2^n" + by (simp add: shiftl_t2n) + +lemma shiftr_1[simp]: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" + by (induct n) (auto simp: shiftr_def) + +lemma map_nth_0 [simp]: "map ((!!) (0::'a::len word)) xs = replicate (length xs) False" + by (induct xs) auto + +lemma word_and_1: + "n AND 1 = (if n !! 0 then 1 else 0)" for n :: "_ word" + by (rule bit_word_eqI) (auto simp add: bit_and_iff test_bit_eq_bit bit_1_iff intro: gr0I) + +lemma test_bit_1' [simp]: + "(1 :: 'a :: len word) !! n \ 0 < LENGTH('a) \ n = 0" + by simp + +lemma shiftl0: + "x << 0 = (x :: 'a :: len word)" + by (fact shiftl_x_0) + +setup \ + Context.theory_map (fold SMT_Word.add_word_shift' [ + (\<^term>\shiftl :: 'a::len word \ _\, "bvshl"), + (\<^term>\shiftr :: 'a::len word \ _\, "bvlshr"), + (\<^term>\sshiftr :: 'a::len word \ _\, "bvashr") + ]) +\ + +end diff --git a/thys/Word_Lib/Word_Examples.thy b/thys/Word_Lib/Word_Examples.thy new file mode 100644 --- /dev/null +++ b/thys/Word_Lib/Word_Examples.thy @@ -0,0 +1,164 @@ +(* Title: HOL/Word/Word_Examples.thy + Authors: Gerwin Klein and Thomas Sewell, NICTA + +Examples demonstrating and testing various word operations. +*) + +section "Examples of word operations" + +theory Word_Examples + imports Word Misc_lsb Misc_set_bit +begin + +type_synonym word32 = "32 word" +type_synonym word8 = "8 word" +type_synonym byte = word8 + +text "modulus" + +lemma "(27 :: 4 word) = -5" by simp + +lemma "(27 :: 4 word) = 11" by simp + +lemma "27 \ (11 :: 6 word)" by simp + +text "signed" + +lemma "(127 :: 6 word) = -1" by simp + +text "number ring simps" + +lemma + "27 + 11 = (38::'a::len word)" + "27 + 11 = (6::5 word)" + "7 * 3 = (21::'a::len word)" + "11 - 27 = (-16::'a::len word)" + "- (- 11) = (11::'a::len word)" + "-40 + 1 = (-39::'a::len word)" + by simp_all + +lemma "word_pred 2 = 1" by simp + +lemma "word_succ (- 3) = -2" by simp + +lemma "23 < (27::8 word)" by simp +lemma "23 \ (27::8 word)" by simp +lemma "\ 23 < (27::2 word)" by simp +lemma "0 < (4::3 word)" by simp +lemma "1 < (4::3 word)" by simp +lemma "0 < (1::3 word)" by simp + +text "ring operations" + +lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" by simp + +text "casting" + +lemma "uint (234567 :: 10 word) = 71" by simp +lemma "uint (-234567 :: 10 word) = 953" by simp +lemma "sint (234567 :: 10 word) = 71" by simp +lemma "sint (-234567 :: 10 word) = -71" by simp +lemma "uint (1 :: 10 word) = 1" by simp + +lemma "unat (-234567 :: 10 word) = 953" by simp +lemma "unat (1 :: 10 word) = 1" by simp + +lemma "ucast (0b1010 :: 4 word) = (0b10 :: 2 word)" by simp +lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" by simp +lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" by simp +lemma "ucast (1 :: 4 word) = (1 :: 2 word)" by simp + +text "reducing goals to nat or int and arith:" +lemma "i < x \ i < i + 1" for i x :: "'a::len word" + by unat_arith +lemma "i < x \ i < i + 1" for i x :: "'a::len word" + by unat_arith + +text "bool lists" + +lemma "of_bl [True, False, True, True] = (0b1011::'a::len word)" by simp + +lemma "to_bl (0b110::4 word) = [False, True, True, False]" by simp + +lemma "of_bl (replicate 32 True) = (0xFFFFFFFF::32 word)" + by (simp add: numeral_eq_Suc) + + +text "bit operations" + +lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by simp +lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by simp +lemma "0xF0 XOR 0xFF = (0x0F :: byte)" by simp +lemma "NOT (0xF0 :: 16 word) = 0xFF0F" by simp +lemma "0 AND 5 = (0 :: byte)" by simp +lemma "1 AND 1 = (1 :: byte)" by simp +lemma "1 AND 0 = (0 :: byte)" by simp +lemma "1 AND 5 = (1 :: byte)" by simp +lemma "1 OR 6 = (7 :: byte)" by simp +lemma "1 OR 1 = (1 :: byte)" by simp +lemma "1 XOR 7 = (6 :: byte)" by simp +lemma "1 XOR 1 = (0 :: byte)" by simp +lemma "NOT 1 = (254 :: byte)" by simp +lemma "NOT 0 = (255 :: byte)" by simp + +lemma "(-1 :: 32 word) = 0xFFFFFFFF" by simp + +lemma "(0b0010 :: 4 word) !! 1" by simp +lemma "\ (0b0010 :: 4 word) !! 0" by simp +lemma "\ (0b1000 :: 3 word) !! 4" by simp +lemma "\ (1 :: 3 word) !! 2" by simp + +lemma "(0b11000 :: 10 word) !! n = (n = 4 \ n = 3)" + by (auto simp add: bin_nth_Bit0 bin_nth_Bit1) + +lemma "set_bit 55 7 True = (183::'a::len word)" by simp +lemma "set_bit 0b0010 7 True = (0b10000010::'a::len word)" by simp +lemma "set_bit 0b0010 1 False = (0::'a::len word)" by simp +lemma "set_bit 1 3 True = (0b1001::'a::len word)" by simp +lemma "set_bit 1 0 False = (0::'a::len word)" by simp +lemma "set_bit 0 3 True = (0b1000::'a::len word)" by simp +lemma "set_bit 0 3 False = (0::'a::len word)" by simp + +lemma "lsb (0b0101::'a::len word)" by simp +lemma "\ lsb (0b1000::'a::len word)" by simp +lemma "lsb (1::'a::len word)" by simp +lemma "\ lsb (0::'a::len word)" by simp + +lemma "\ msb (0b0101::4 word)" by simp +lemma "msb (0b1000::4 word)" by simp +lemma "\ msb (1::4 word)" by simp +lemma "\ msb (0::4 word)" by simp + +lemma "word_cat (27::4 word) (27::8 word) = (2843::'a::len word)" by simp +lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)" + by simp + +lemma "0b1011 << 2 = (0b101100::'a::len word)" by simp +lemma "0b1011 >> 2 = (0b10::8 word)" by simp +lemma "0b1011 >>> 2 = (0b10::8 word)" by simp +lemma "1 << 2 = (0b100::'a::len word)" apply simp? oops + +lemma "slice 3 (0b101111::6 word) = (0b101::3 word)" by simp +lemma "slice 3 (1::6 word) = (0::3 word)" apply simp? oops + +lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by simp +lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by simp +lemma "word_roti 2 0b1110 = (0b1011::4 word)" by simp +lemma "word_roti (- 2) 0b0110 = (0b1001::4 word)" by simp +lemma "word_rotr 2 0 = (0::4 word)" by simp +lemma "word_rotr 2 1 = (0b0100::4 word)" apply simp? oops +lemma "word_rotl 2 1 = (0b0100::4 word)" apply simp? oops +lemma "word_roti (- 2) 1 = (0b0100::4 word)" apply simp? oops + +lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" +proof - + have "(x AND 0xff00) OR (x AND 0x00ff) = x AND (0xff00 OR 0x00ff)" + by (simp only: word_ao_dist2) + also have "0xff00 OR 0x00ff = (-1::16 word)" + by simp + also have "x AND -1 = x" + by simp + finally show ?thesis . +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,6168 +1,6168 @@ (* * 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 "HOL-Library.Sublist" - "HOL-Word.Misc_lsb" + Misc_lsb Word_EqI Word_Enum Norm_Words Word_Type_Syntax Bitwise_Signed Hex_Words begin lemmas word_next_def = word_next_unfold lemmas word_prev_def = word_prev_unfold lemmas is_aligned_def = is_aligned_iff_dvd_nat 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 \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 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 by simp 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" by word_eqI 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'[simp]: "is_aligned 0 n" by (simp add: is_aligned_def) 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 sublist_equal_part: "prefix xs ys \ take (length xs) ys = xs" by (clarsimp simp: prefix_def) 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 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 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 map_prefixI: "prefix xs ys \ prefix (map f xs) (map f ys)" by (clarsimp simp: prefix_def) 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 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 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 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 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) 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 word_lsb_nat:"lsb w = (unat w mod 2 = 1)" apply (simp add: word_lsb_def Groebner_Basis.algebra(31)) apply transfer apply (simp add: even_nat_iff) done lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = x !! 0" apply (simp add:even_iff_mod_2_eq_zero) apply (subst word_lsb_nat[unfolded One_nat_def, symmetric]) apply (rule word_lsb_alt) done 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" using word_overflow_unat [of x] apply (simp only: shiftr1_is_div_2 flip: odd_iff_lsb) apply (cases \2 \ LENGTH('a)\) apply (auto simp add: test_bit_def' word_arith_nat_div dest: overflow_imp_lsb) using odd_iff_lsb overflow_imp_lsb by blast 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) unfolding is_aligned_def by (metis le_unat_uoi nat_dvd_not_less order_less_imp_le unat_power_lower) 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) 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) apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (clarsimp simp: is_aligned_def dvd_def word_le_nat_alt) apply (drule le_imp_less_Suc) apply (simp add: Suc_2p_unat_mask) by (metis Groups.mult_ac(2) Suc_leI linorder_not_less mult_le_mono order_refl times_nat.simps(2)) 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: is_aligned_def dvd_def) 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_def) 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) 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 (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 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,735 +1,735 @@ (* * 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" - "HOL-Word.Misc_set_bit" + Misc_set_bit Word_Syntax Signed_Words begin 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 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 (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) lemma one_bit_shiftl: "set_bit 0 n True = (1 :: 'a :: len word) << n" apply (rule word_eqI) apply (auto simp add: test_bit_set_gen nth_shiftl word_size simp del: word_set_bit_0 shiftl_1) done lemmas one_bit_pow = trans [OF one_bit_shiftl shiftl_1] 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_Syntax.thy b/thys/Word_Lib/Word_Syntax.thy --- a/thys/Word_Lib/Word_Syntax.thy +++ b/thys/Word_Lib/Word_Syntax.thy @@ -1,60 +1,60 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Additional Syntax for Word Bit Operations" theory Word_Syntax imports - "HOL-Word.Word" + "HOL-Library.Word" Bitwise begin text \Additional bit and type syntax that forces word types.\ type_synonym word8 = "8 word" type_synonym word16 = "16 word" type_synonym word32 = "32 word" type_synonym word64 = "64 word" lemma len8: "len_of (x :: 8 itself) = 8" by simp lemma len16: "len_of (x :: 16 itself) = 16" by simp lemma len32: "len_of (x :: 32 itself) = 32" by simp lemma len64: "len_of (x :: 64 itself) = 64" by simp abbreviation wordNOT :: "'a::len word \ 'a word" ("~~ _" [70] 71) where "~~ x == NOT x" abbreviation wordAND :: "'a::len word \ 'a word \ 'a word" (infixr "&&" 64) where "a && b == a AND b" abbreviation wordOR :: "'a::len word \ 'a word \ 'a word" (infixr "||" 59) where "a || b == a OR b" abbreviation wordXOR :: "'a::len word \ 'a word \ 'a word" (infixr "xor" 59) where "a xor b == a XOR b" (* testing for presence of word_bitwise *) lemma "((x :: word32) >> 3) AND 7 = (x AND 56) >> 3" by word_bitwise (* FIXME: move to Word distribution *) lemma bin_nth_minus_Bit0[simp]: "0 < n \ bin_nth (numeral (num.Bit0 w)) n = bin_nth (numeral w) (n - 1)" by (cases n; simp) lemma bin_nth_minus_Bit1[simp]: "0 < n \ bin_nth (numeral (num.Bit1 w)) n = bin_nth (numeral w) (n - 1)" by (cases n; simp) end diff --git a/thys/Word_Lib/Word_Type_Syntax.thy b/thys/Word_Lib/Word_Type_Syntax.thy --- a/thys/Word_Lib/Word_Type_Syntax.thy +++ b/thys/Word_Lib/Word_Type_Syntax.thy @@ -1,58 +1,58 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Displaying Phantom Types for Word Operations" theory Word_Type_Syntax -imports "HOL-Word.Word" + imports "HOL-Library.Word" begin ML \ structure Word_Syntax = struct val show_word_types = Attrib.setup_config_bool @{binding show_word_types} (K true) fun tr' cnst ctxt typ ts = if Config.get ctxt show_word_types then case (Term.binder_types typ, Term.body_type typ) of ([Type (@{type_name "word"}, [S])], Type (@{type_name "word"}, [T])) => list_comb (Syntax.const cnst $ Syntax_Phases.term_of_typ ctxt S $ Syntax_Phases.term_of_typ ctxt T , ts) | _ => raise Match else raise Match end \ syntax "_Ucast" :: "type \ type \ logic" ("(1UCAST/(1'(_ \ _')))") translations "UCAST('s \ 't)" => "CONST ucast :: ('s word \ 't word)" typed_print_translation \ [(@{const_syntax ucast}, Word_Syntax.tr' @{syntax_const "_Ucast"})] \ syntax "_Scast" :: "type \ type \ logic" ("(1SCAST/(1'(_ \ _')))") translations "SCAST('s \ 't)" => "CONST scast :: ('s word \ 't word)" typed_print_translation \ [(@{const_syntax scast}, Word_Syntax.tr' @{syntax_const "_Scast"})] \ syntax "_Revcast" :: "type \ type \ logic" ("(1REVCAST/(1'(_ \ _')))") translations "REVCAST('s \ 't)" => "CONST revcast :: ('s word \ 't word)" typed_print_translation \ [(@{const_syntax revcast}, Word_Syntax.tr' @{syntax_const "_Revcast"})] \ (* Further candidates for showing word sizes, but with different arities: slice, word_cat, word_split, word_rcat, word_rsplit *) end diff --git a/thys/Word_Lib/Word_rsplit.thy b/thys/Word_Lib/Word_rsplit.thy new file mode 100644 --- /dev/null +++ b/thys/Word_Lib/Word_rsplit.thy @@ -0,0 +1,149 @@ +(* Title: HOL/Word/Word_rsplit.thy + Author: Jeremy Dawson and Gerwin Klein, NICTA +*) + +theory Word_rsplit + imports Bits_Int "HOL-Library.Word" +begin + +definition word_rsplit :: "'a::len word \ 'b::len word list" + where "word_rsplit w = map word_of_int (bin_rsplit (LENGTH('b)) (LENGTH('a), uint w))" + +lemma word_rsplit_no: + "(word_rsplit (numeral bin :: 'b::len word) :: 'a word list) = + map word_of_int (bin_rsplit (LENGTH('a::len)) + (LENGTH('b), take_bit (LENGTH('b)) (numeral bin)))" + by (simp add: word_rsplit_def of_nat_take_bit) + +lemmas word_rsplit_no_cl [simp] = word_rsplit_no + [unfolded bin_rsplitl_def bin_rsplit_l [symmetric]] + +text \ + This odd result arises from the fact that the statement of the + result implies that the decoded words are of the same type, + and therefore of the same length, as the original word.\ + +lemma word_rsplit_same: "word_rsplit w = [w]" + apply (simp add: word_rsplit_def bin_rsplit_all) + apply transfer + apply simp + done + +lemma word_rsplit_empty_iff_size: "word_rsplit w = [] \ size w = 0" + by (simp add: word_rsplit_def bin_rsplit_def word_size bin_rsplit_aux_simp_alt Let_def + split: prod.split) + +lemma test_bit_rsplit: + "sw = word_rsplit w \ m < size (hd sw) \ + k < length sw \ (rev sw ! k) !! m = w !! (k * size (hd sw) + m)" + for sw :: "'a::len word list" + apply (unfold word_rsplit_def word_test_bit_def) + apply (rule trans) + apply (rule_tac f = "\x. bin_nth x m" in arg_cong) + apply (rule nth_map [symmetric]) + apply simp + apply (rule bin_nth_rsplit) + apply simp_all + apply (simp add : word_size rev_map) + apply (rule trans) + defer + apply (rule map_ident [THEN fun_cong]) + apply (rule refl [THEN map_cong]) + apply simp + using bin_rsplit_size_sign take_bit_int_eq_self_iff by blast + +lemma test_bit_rsplit_alt: + \(word_rsplit w :: 'b::len word list) ! i !! m \ + w !! ((length (word_rsplit w :: 'b::len word list) - Suc i) * size (hd (word_rsplit w :: 'b::len word list)) + m)\ + if \i < length (word_rsplit w :: 'b::len word list)\ \m < size (hd (word_rsplit w :: 'b::len word list))\ \0 < length (word_rsplit w :: 'b::len word list)\ + for w :: \'a::len word\ + apply (rule trans) + apply (rule test_bit_cong) + apply (rule rev_nth [of _ \rev (word_rsplit w)\, simplified rev_rev_ident]) + apply simp + apply (rule that(1)) + apply simp + apply (rule test_bit_rsplit) + apply (rule refl) + apply (rule asm_rl) + apply (rule that(2)) + apply (rule diff_Suc_less) + apply (rule that(3)) + done + +lemma word_rsplit_len_indep [OF refl refl refl refl]: + "[u,v] = p \ [su,sv] = q \ word_rsplit u = su \ + word_rsplit v = sv \ length su = length sv" + by (auto simp: word_rsplit_def bin_rsplit_len_indep) + +lemma length_word_rsplit_size: + "n = LENGTH('a::len) \ + length (word_rsplit w :: 'a word list) \ m \ size w \ m * n" + by (auto simp: word_rsplit_def word_size bin_rsplit_len_le) + +lemmas length_word_rsplit_lt_size = + length_word_rsplit_size [unfolded Not_eq_iff linorder_not_less [symmetric]] + +lemma length_word_rsplit_exp_size: + "n = LENGTH('a::len) \ + length (word_rsplit w :: 'a word list) = (size w + n - 1) div n" + by (auto simp: word_rsplit_def word_size bin_rsplit_len) + +lemma length_word_rsplit_even_size: + "n = LENGTH('a::len) \ size w = m * n \ + length (word_rsplit w :: 'a word list) = m" + by (cases \LENGTH('a)\) (simp_all add: length_word_rsplit_exp_size div_nat_eqI) + +lemmas length_word_rsplit_exp_size' = refl [THEN length_word_rsplit_exp_size] + +\ \alternative proof of \word_rcat_rsplit\\ +lemmas tdle = times_div_less_eq_dividend +lemmas dtle = xtrans(4) [OF tdle mult.commute] + +lemma word_rcat_rsplit: "word_rcat (word_rsplit w) = w" + apply (rule word_eqI) + apply (clarsimp simp: test_bit_rcat word_size) + apply (subst refl [THEN test_bit_rsplit]) + apply (simp_all add: word_size + refl [THEN length_word_rsplit_size [simplified not_less [symmetric], simplified]]) + apply safe + apply (erule xtrans(7), rule dtle)+ + done + +lemma size_word_rsplit_rcat_size: + "word_rcat ws = frcw \ size frcw = length ws * LENGTH('a) + \ length (word_rsplit frcw::'a word list) = length ws" + for ws :: "'a::len word list" and frcw :: "'b::len word" + by (cases \LENGTH('a)\) (simp_all add: word_size length_word_rsplit_exp_size' div_nat_eqI) + +lemma msrevs: + "0 < n \ (k * n + m) div n = m div n + k" + "(k * n + m) mod n = m mod n" + for n :: nat + by (auto simp: add.commute) + +lemma word_rsplit_rcat_size [OF refl]: + "word_rcat ws = frcw \ + size frcw = length ws * LENGTH('a) \ word_rsplit frcw = ws" + for ws :: "'a::len word list" + apply (frule size_word_rsplit_rcat_size, assumption) + apply (clarsimp simp add : word_size) + apply (rule nth_equalityI, assumption) + apply clarsimp + apply (rule word_eqI [rule_format]) + apply (rule trans) + apply (rule test_bit_rsplit_alt) + apply (clarsimp simp: word_size)+ + apply (rule trans) + apply (rule test_bit_rcat [OF refl refl]) + apply (simp add: word_size) + apply (subst rev_nth) + apply arith + apply (simp add: le0 [THEN [2] xtrans(7), THEN diff_Suc_less]) + apply safe + apply (simp add: diff_mult_distrib) + apply (cases "size ws") + apply simp_all + done + +end \ No newline at end of file