diff --git a/src/HOL/Word/Bits_Int.thy b/src/HOL/Word/Bits_Int.thy --- a/src/HOL/Word/Bits_Int.thy +++ b/src/HOL/Word/Bits_Int.thy @@ -1,1548 +1,1550 @@ (* Title: HOL/Word/Bits_Int.thy Author: Jeremy Dawson and Gerwin Klein, NICTA *) section \Bitwise Operations on integers\ theory Bits_Int imports "HOL-Library.Bit_Operations" Traditional_Syntax begin subsection \Generic auxiliary\ lemma int_mod_ge: "a < n \ 0 < n \ a \ a mod n" for a n :: int by (metis dual_order.trans le_cases mod_pos_pos_trivial pos_mod_conj) 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) instantiation int :: semiring_bit_syntax begin definition [iff]: "i !! n \ bin_nth i n" definition "shiftl x n = x * 2 ^ n" for x :: int definition "shiftr x n = x div 2 ^ n" for x :: int instance by standard (simp_all add: fun_eq_iff shiftl_int_def shiftr_int_def push_bit_eq_mult drop_bit_eq_div) end lemma shiftl_eq_push_bit: \k << n = push_bit n k\ for k :: int by (fact shiftl_eq_push_bit) lemma shiftr_eq_drop_bit: \k >> n = drop_bit n k\ for k :: int by (fact 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 (* FIXME: The rule sets below are very large (24 rules for each operator). Is there a simpler way to do this? *) lemma int_and_numerals [simp]: "numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)" "numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND numeral y)" "numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)" "numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND numeral y)" "numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)" "numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND - numeral (y + Num.One))" "numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)" "numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND - numeral (y + Num.One))" "- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND numeral y)" "- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND numeral y)" "- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND numeral y)" "- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND numeral y)" "- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND - numeral y)" "- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND - numeral (y + Num.One))" "- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND - numeral y)" "- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND - numeral (y + Num.One))" "(1::int) AND numeral (Num.Bit0 y) = 0" "(1::int) AND numeral (Num.Bit1 y) = 1" "(1::int) AND - numeral (Num.Bit0 y) = 0" "(1::int) AND - numeral (Num.Bit1 y) = 1" "numeral (Num.Bit0 x) AND (1::int) = 0" "numeral (Num.Bit1 x) AND (1::int) = 1" "- numeral (Num.Bit0 x) AND (1::int) = 0" "- numeral (Num.Bit1 x) AND (1::int) = 1" by (rule bin_rl_eqI; simp)+ lemma int_or_numerals [simp]: "numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR numeral y)" "numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)" "numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR numeral y)" "numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)" "numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR - numeral y)" "numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))" "numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR - numeral y)" "numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))" "- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR numeral y)" "- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR numeral y)" "- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)" "- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)" "- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR - numeral y)" "- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR - numeral (y + Num.One))" "- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral y)" "- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral (y + Num.One))" "(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" "(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)" "(1::int) OR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)" "(1::int) OR - numeral (Num.Bit1 y) = - numeral (Num.Bit1 y)" "numeral (Num.Bit0 x) OR (1::int) = numeral (Num.Bit1 x)" "numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)" "- numeral (Num.Bit0 x) OR (1::int) = - numeral (Num.BitM x)" "- numeral (Num.Bit1 x) OR (1::int) = - numeral (Num.Bit1 x)" by (rule bin_rl_eqI; simp)+ lemma int_xor_numerals [simp]: "numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR numeral y)" "numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR numeral y)" "numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR numeral y)" "numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR numeral y)" "numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR - numeral y)" "numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR - numeral (y + Num.One))" "numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR - numeral y)" "numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR - numeral (y + Num.One))" "- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR numeral y)" "- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR numeral y)" "- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR numeral y)" "- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR numeral y)" "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR - numeral y)" "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR - numeral (y + Num.One))" "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR - numeral y)" "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR - numeral (y + Num.One))" "(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" "(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)" "(1::int) XOR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)" "(1::int) XOR - numeral (Num.Bit1 y) = - numeral (Num.Bit0 (y + Num.One))" "numeral (Num.Bit0 x) XOR (1::int) = numeral (Num.Bit1 x)" "numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)" "- numeral (Num.Bit0 x) XOR (1::int) = - numeral (Num.BitM x)" "- numeral (Num.Bit1 x) XOR (1::int) = - numeral (Num.Bit0 (x + Num.One))" by (rule bin_rl_eqI; simp)+ subsubsection \Interactions with arithmetic\ lemma plus_and_or: "(x AND y) + (x OR y) = x + y" for x y :: int proof (induction x arbitrary: y rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even x) from even.IH [of \y div 2\] show ?case by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE) next case (odd x) from odd.IH [of \y div 2\] show ?case by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE) qed 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 \Comparison\ lemma AND_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" shows "0 \ x AND y" using assms by simp lemma OR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" "0 \ y" shows "0 \ x OR y" using assms by simp lemma XOR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" "0 \ y" shows "0 \ x XOR y" using assms by simp lemma AND_upper1 [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" shows "x AND y \ x" using assms by (induction x arbitrary: y rule: int_bit_induct) (simp_all add: and_int_rec [of \_ * 2\] and_int_rec [of \1 + _ * 2\] add_increasing) lemmas AND_upper1' [simp] = order_trans [OF AND_upper1] \<^marker>\contributor \Stefan Berghofer\\ lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] \<^marker>\contributor \Stefan Berghofer\\ lemma AND_upper2 [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ y" shows "x AND y \ y" using assms AND_upper1 [of y x] by (simp add: ac_simps) lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] \<^marker>\contributor \Stefan Berghofer\\ lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] \<^marker>\contributor \Stefan Berghofer\\ lemma OR_upper: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" "x < 2 ^ n" "y < 2 ^ n" shows "x OR y < 2 ^ n" using assms proof (induction x arbitrary: y n rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even x) from even.IH [of \n - 1\ \y div 2\] even.prems even.hyps show ?case by (cases n) (auto simp add: or_int_rec [of \_ * 2\] elim: oddE) next case (odd x) from odd.IH [of \n - 1\ \y div 2\] odd.prems odd.hyps show ?case by (cases n) (auto simp add: or_int_rec [of \1 + _ * 2\], linarith) qed lemma XOR_upper: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" "x < 2 ^ n" "y < 2 ^ n" shows "x XOR y < 2 ^ n" using assms proof (induction x arbitrary: y n rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even x) from even.IH [of \n - 1\ \y div 2\] even.prems even.hyps show ?case by (cases n) (auto simp add: xor_int_rec [of \_ * 2\] elim: oddE) next case (odd x) from odd.IH [of \n - 1\ \y div 2\] odd.prems odd.hyps show ?case by (cases n) (auto simp add: xor_int_rec [of \1 + _ * 2\]) qed 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 (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) lemma bin_trunc_xor: "bintrunc n (bintrunc n x XOR bintrunc n y) = bintrunc n (x XOR y)" by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) lemma bin_trunc_not: "bintrunc n (NOT (bintrunc n x)) = bintrunc n (NOT x)" by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) 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)+ end diff --git a/src/HOL/Word/More_Word.thy b/src/HOL/Word/More_Word.thy --- a/src/HOL/Word/More_Word.thy +++ b/src/HOL/Word/More_Word.thy @@ -1,57 +1,58 @@ (* Title: HOL/Word/More_thy *) section \Ancient comprehensive Word Library\ theory More_Word imports 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/src/HOL/Word/Word.thy b/src/HOL/Word/Word.thy --- a/src/HOL/Word/Word.thy +++ b/src/HOL/Word/Word.thy @@ -1,5310 +1,5170 @@ (* Title: HOL/Word/Word.thy Author: Jeremy Dawson and Gerwin Klein, NICTA *) section \A type of finite bit strings\ theory Word imports "HOL-Library.Type_Length" "HOL-Library.Boolean_Algebra" "HOL-Library.Bit_Operations" Bits_Int Traditional_Syntax begin subsection \Preliminaries\ lemma signed_take_bit_decr_length_iff: \signed_take_bit (LENGTH('a::len) - Suc 0) k = signed_take_bit (LENGTH('a) - Suc 0) l \ take_bit LENGTH('a) k = take_bit LENGTH('a) l\ by (cases \LENGTH('a)\) (simp_all add: signed_take_bit_eq_iff_take_bit_eq) subsection \Fundamentals\ subsubsection \Type definition\ quotient_type (overloaded) 'a word = int / \\k l. take_bit LENGTH('a) k = take_bit LENGTH('a::len) l\ morphisms rep Word by (auto intro!: equivpI reflpI sympI transpI) hide_const (open) rep \ \only for foundational purpose\ hide_const (open) Word \ \only for code generation\ subsubsection \Basic arithmetic\ instantiation word :: (len) comm_ring_1 begin lift_definition zero_word :: \'a word\ is 0 . lift_definition one_word :: \'a word\ is 1 . lift_definition plus_word :: \'a word \ 'a word \ 'a word\ is \(+)\ by (auto simp add: take_bit_eq_mod intro: mod_add_cong) lift_definition minus_word :: \'a word \ 'a word \ 'a word\ is \(-)\ by (auto simp add: take_bit_eq_mod intro: mod_diff_cong) lift_definition uminus_word :: \'a word \ 'a word\ is uminus by (auto simp add: take_bit_eq_mod intro: mod_minus_cong) lift_definition times_word :: \'a word \ 'a word \ 'a word\ is \(*)\ by (auto simp add: take_bit_eq_mod intro: mod_mult_cong) instance by (standard; transfer) (simp_all add: algebra_simps) end context includes lifting_syntax notes power_transfer [transfer_rule] transfer_rule_of_bool [transfer_rule] transfer_rule_numeral [transfer_rule] transfer_rule_of_nat [transfer_rule] transfer_rule_of_int [transfer_rule] begin lemma power_transfer_word [transfer_rule]: \(pcr_word ===> (=) ===> pcr_word) (^) (^)\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) of_bool of_bool\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) numeral numeral\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) int of_nat\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) (\k. k) of_int\ proof - have \((=) ===> pcr_word) of_int of_int\ by transfer_prover then show ?thesis by (simp add: id_def) qed lemma [transfer_rule]: \(pcr_word ===> (\)) even ((dvd) 2 :: 'a::len word \ bool)\ proof - have even_word_unfold: "even k \ (\l. take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l))" (is "?P \ ?Q") for k :: int proof assume ?P then show ?Q by auto next assume ?Q then obtain l where "take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l)" .. then have "even (take_bit LENGTH('a) k)" by simp then show ?P by simp qed show ?thesis by (simp only: even_word_unfold [abs_def] dvd_def [where ?'a = "'a word", abs_def]) transfer_prover qed end lemma word_exp_length_eq_0 [simp]: \(2 :: 'a::len word) ^ LENGTH('a) = 0\ by transfer simp lemma exp_eq_zero_iff: \2 ^ n = (0 :: 'a::len word) \ n \ LENGTH('a)\ by transfer simp subsubsection \Basic code generation setup\ context begin qualified lift_definition the_int :: \'a::len word \ int\ is \take_bit LENGTH('a)\ . end lemma [code abstype]: \Word.Word (Word.the_int w) = w\ by transfer simp lemma Word_eq_word_of_int [code_post, simp]: \Word.Word = of_int\ by (rule; transfer) simp quickcheck_generator word constructors: \0 :: 'a::len word\, \numeral :: num \ 'a::len word\ instantiation word :: (len) equal begin lift_definition equal_word :: \'a word \ 'a word \ bool\ is \\k l. take_bit LENGTH('a) k = take_bit LENGTH('a) l\ by simp instance by (standard; transfer) rule end lemma [code]: \HOL.equal v w \ HOL.equal (Word.the_int v) (Word.the_int w)\ by transfer (simp add: equal) lemma [code]: \Word.the_int 0 = 0\ by transfer simp lemma [code]: \Word.the_int 1 = 1\ by transfer simp lemma [code]: \Word.the_int (v + w) = take_bit LENGTH('a) (Word.the_int v + Word.the_int w)\ for v w :: \'a::len word\ by transfer (simp add: take_bit_add) lemma [code]: \Word.the_int (- w) = (let k = Word.the_int w in if w = 0 then 0 else 2 ^ LENGTH('a) - k)\ for w :: \'a::len word\ by transfer (auto simp add: take_bit_eq_mod zmod_zminus1_eq_if) lemma [code]: \Word.the_int (v - w) = take_bit LENGTH('a) (Word.the_int v - Word.the_int w)\ for v w :: \'a::len word\ by transfer (simp add: take_bit_diff) lemma [code]: \Word.the_int (v * w) = take_bit LENGTH('a) (Word.the_int v * Word.the_int w)\ for v w :: \'a::len word\ by transfer (simp add: take_bit_mult) subsubsection \Basic conversions\ abbreviation word_of_nat :: \nat \ 'a::len word\ where \word_of_nat \ of_nat\ abbreviation word_of_int :: \int \ 'a::len word\ where \word_of_int \ of_int\ lemma word_of_nat_eq_iff: \word_of_nat m = (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m = take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma word_of_int_eq_iff: \word_of_int k = (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k = take_bit LENGTH('a) l\ by transfer rule lemma word_of_nat_eq_0_iff [simp]: \word_of_nat n = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd n\ using word_of_nat_eq_iff [where ?'a = 'a, of n 0] by (simp add: take_bit_eq_0_iff) lemma word_of_int_eq_0_iff [simp]: \word_of_int k = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd k\ using word_of_int_eq_iff [where ?'a = 'a, of k 0] by (simp add: take_bit_eq_0_iff) context semiring_1 begin lift_definition unsigned :: \'b::len word \ 'a\ is \of_nat \ nat \ take_bit LENGTH('b)\ by simp lemma unsigned_0 [simp]: \unsigned 0 = 0\ by transfer simp lemma unsigned_1 [simp]: \unsigned 1 = 1\ by transfer simp lemma unsigned_numeral [simp]: \unsigned (numeral n :: 'b::len word) = of_nat (take_bit LENGTH('b) (numeral n))\ by transfer (simp add: nat_take_bit_eq) lemma unsigned_neg_numeral [simp]: \unsigned (- numeral n :: 'b::len word) = of_nat (nat (take_bit LENGTH('b) (- numeral n)))\ by transfer simp end context semiring_1 begin lemma unsigned_of_nat [simp]: \unsigned (word_of_nat n :: 'b::len word) = of_nat (take_bit LENGTH('b) n)\ by transfer (simp add: nat_eq_iff take_bit_of_nat) lemma unsigned_of_int [simp]: \unsigned (word_of_int k :: 'b::len word) = of_nat (nat (take_bit LENGTH('b) k))\ by transfer simp end context semiring_char_0 begin lemma unsigned_word_eqI: \v = w\ if \unsigned v = unsigned w\ using that by transfer (simp add: eq_nat_nat_iff) lemma word_eq_iff_unsigned: \v = w \ unsigned v = unsigned w\ by (auto intro: unsigned_word_eqI) lemma inj_unsigned [simp]: \inj unsigned\ by (rule injI) (simp add: unsigned_word_eqI) lemma unsigned_eq_0_iff: \unsigned w = 0 \ w = 0\ using word_eq_iff_unsigned [of w 0] by simp end context ring_1 begin lift_definition signed :: \'b::len word \ 'a\ is \of_int \ signed_take_bit (LENGTH('b) - Suc 0)\ by (simp flip: signed_take_bit_decr_length_iff) lemma signed_0 [simp]: \signed 0 = 0\ by transfer simp lemma signed_1 [simp]: \signed (1 :: 'b::len word) = (if LENGTH('b) = 1 then - 1 else 1)\ by (transfer fixing: uminus; cases \LENGTH('b)\) (simp_all add: sbintrunc_minus_simps) lemma signed_minus_1 [simp]: \signed (- 1 :: 'b::len word) = - 1\ by (transfer fixing: uminus) simp lemma signed_numeral [simp]: \signed (numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (numeral n))\ by transfer simp lemma signed_neg_numeral [simp]: \signed (- numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (- numeral n))\ by transfer simp lemma signed_of_nat [simp]: \signed (word_of_nat n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - Suc 0) (int n))\ by transfer simp lemma signed_of_int [simp]: \signed (word_of_int n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - Suc 0) n)\ by transfer simp end context ring_char_0 begin lemma signed_word_eqI: \v = w\ if \signed v = signed w\ using that by transfer (simp flip: signed_take_bit_decr_length_iff) lemma word_eq_iff_signed: \v = w \ signed v = signed w\ by (auto intro: signed_word_eqI) lemma inj_signed [simp]: \inj signed\ by (rule injI) (simp add: signed_word_eqI) lemma signed_eq_0_iff: \signed w = 0 \ w = 0\ using word_eq_iff_signed [of w 0] by simp end abbreviation unat :: \'a::len word \ nat\ where \unat \ unsigned\ abbreviation uint :: \'a::len word \ int\ where \uint \ unsigned\ abbreviation sint :: \'a::len word \ int\ where \sint \ signed\ abbreviation ucast :: \'a::len word \ 'b::len word\ where \ucast \ unsigned\ abbreviation scast :: \'a::len word \ 'b::len word\ where \scast \ signed\ context includes lifting_syntax begin lemma [transfer_rule]: \(pcr_word ===> (=)) (nat \ take_bit LENGTH('a)) (unat :: 'a::len word \ nat)\ using unsigned.transfer [where ?'a = nat] by simp lemma [transfer_rule]: \(pcr_word ===> (=)) (take_bit LENGTH('a)) (uint :: 'a::len word \ int)\ using unsigned.transfer [where ?'a = int] by (simp add: comp_def) lemma [transfer_rule]: \(pcr_word ===> (=)) (signed_take_bit (LENGTH('a) - Suc 0)) (sint :: 'a::len word \ int)\ using signed.transfer [where ?'a = int] by simp lemma [transfer_rule]: \(pcr_word ===> pcr_word) (take_bit LENGTH('a)) (ucast :: 'a::len word \ 'b::len word)\ proof (rule rel_funI) fix k :: int and w :: \'a word\ assume \pcr_word k w\ then have \w = word_of_int k\ by (simp add: pcr_word_def cr_word_def relcompp_apply) moreover have \pcr_word (take_bit LENGTH('a) k) (ucast (word_of_int k :: 'a word))\ by transfer (simp add: pcr_word_def cr_word_def relcompp_apply) ultimately show \pcr_word (take_bit LENGTH('a) k) (ucast w)\ by simp qed lemma [transfer_rule]: \(pcr_word ===> pcr_word) (signed_take_bit (LENGTH('a) - Suc 0)) (scast :: 'a::len word \ 'b::len word)\ proof (rule rel_funI) fix k :: int and w :: \'a word\ assume \pcr_word k w\ then have \w = word_of_int k\ by (simp add: pcr_word_def cr_word_def relcompp_apply) moreover have \pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast (word_of_int k :: 'a word))\ by transfer (simp add: pcr_word_def cr_word_def relcompp_apply) ultimately show \pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast w)\ by simp qed end lemma of_nat_unat [simp]: \of_nat (unat w) = unsigned w\ by transfer simp lemma of_int_uint [simp]: \of_int (uint w) = unsigned w\ by transfer simp lemma of_int_sint [simp]: \of_int (sint a) = signed a\ by transfer (simp_all add: take_bit_signed_take_bit) lemma nat_uint_eq [simp]: \nat (uint w) = unat w\ by transfer simp lemma sgn_uint_eq [simp]: \sgn (uint w) = of_bool (w \ 0)\ by transfer (simp add: less_le) text \Aliasses only for code generation\ context begin qualified lift_definition of_int :: \int \ 'a::len word\ is \take_bit LENGTH('a)\ . qualified lift_definition of_nat :: \nat \ 'a::len word\ is \int \ take_bit LENGTH('a)\ . qualified lift_definition the_nat :: \'a::len word \ nat\ is \nat \ take_bit LENGTH('a)\ by simp qualified lift_definition the_signed_int :: \'a::len word \ int\ is \signed_take_bit (LENGTH('a) - Suc 0)\ by (simp add: signed_take_bit_decr_length_iff) qualified lift_definition cast :: \'a::len word \ 'b::len word\ is \take_bit LENGTH('a)\ by simp qualified lift_definition signed_cast :: \'a::len word \ 'b::len word\ is \signed_take_bit (LENGTH('a) - Suc 0)\ by (metis signed_take_bit_decr_length_iff) end lemma [code_abbrev, simp]: \Word.the_int = uint\ by transfer rule lemma [code]: \Word.the_int (Word.of_int k :: 'a::len word) = take_bit LENGTH('a) k\ by transfer simp lemma [code_abbrev, simp]: \Word.of_int = word_of_int\ by (rule; transfer) simp lemma [code]: \Word.the_int (Word.of_nat n :: 'a::len word) = take_bit LENGTH('a) (int n)\ by transfer (simp add: take_bit_of_nat) lemma [code_abbrev, simp]: \Word.of_nat = word_of_nat\ by (rule; transfer) (simp add: take_bit_of_nat) lemma [code]: \Word.the_nat w = nat (Word.the_int w)\ by transfer simp lemma [code_abbrev, simp]: \Word.the_nat = unat\ by (rule; transfer) simp lemma [code]: \Word.the_signed_int w = signed_take_bit (LENGTH('a) - Suc 0) (Word.the_int w)\ for w :: \'a::len word\ by transfer simp lemma [code_abbrev, simp]: \Word.the_signed_int = sint\ by (rule; transfer) simp lemma [code]: \Word.the_int (Word.cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_int w)\ for w :: \'a::len word\ by transfer simp lemma [code_abbrev, simp]: \Word.cast = ucast\ by (rule; transfer) simp lemma [code]: \Word.the_int (Word.signed_cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_signed_int w)\ for w :: \'a::len word\ by transfer simp lemma [code_abbrev, simp]: \Word.signed_cast = scast\ by (rule; transfer) simp lemma [code]: \unsigned w = of_nat (nat (Word.the_int w))\ by transfer simp lemma [code]: \signed w = of_int (Word.the_signed_int w)\ by transfer simp subsubsection \Basic ordering\ instantiation word :: (len) linorder begin lift_definition less_eq_word :: "'a word \ 'a word \ bool" is "\a b. take_bit LENGTH('a) a \ take_bit LENGTH('a) b" by simp lift_definition less_word :: "'a word \ 'a word \ bool" is "\a b. take_bit LENGTH('a) a < take_bit LENGTH('a) b" by simp instance by (standard; transfer) auto end interpretation word_order: ordering_top \(\)\ \(<)\ \- 1 :: 'a::len word\ by (standard; transfer) (simp add: take_bit_eq_mod zmod_minus1) interpretation word_coorder: ordering_top \(\)\ \(>)\ \0 :: 'a::len word\ by (standard; transfer) simp lemma word_of_nat_less_eq_iff: \word_of_nat m \ (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m \ take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma word_of_int_less_eq_iff: \word_of_int k \ (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ by transfer rule lemma word_of_nat_less_iff: \word_of_nat m < (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m < take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma word_of_int_less_iff: \word_of_int k < (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k < take_bit LENGTH('a) l\ by transfer rule lemma word_le_def [code]: "a \ b \ uint a \ uint b" by transfer rule lemma word_less_def [code]: "a < b \ uint a < uint b" by transfer rule lemma word_greater_zero_iff: \a > 0 \ a \ 0\ for a :: \'a::len word\ by transfer (simp add: less_le) lemma of_nat_word_less_eq_iff: \of_nat m \ (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m \ take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma of_nat_word_less_iff: \of_nat m < (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m < take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma of_int_word_less_eq_iff: \of_int k \ (of_int l :: 'a::len word) \ take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ by transfer rule lemma of_int_word_less_iff: \of_int k < (of_int l :: 'a::len word) \ take_bit LENGTH('a) k < take_bit LENGTH('a) l\ by transfer rule subsection \Enumeration\ lemma inj_on_word_of_nat: \inj_on (word_of_nat :: nat \ 'a::len word) {0..<2 ^ LENGTH('a)}\ by (rule inj_onI; transfer) (simp_all add: take_bit_int_eq_self) lemma UNIV_word_eq_word_of_nat: \(UNIV :: 'a::len word set) = word_of_nat ` {0..<2 ^ LENGTH('a)}\ (is \_ = ?A\) proof show \word_of_nat ` {0..<2 ^ LENGTH('a)} \ UNIV\ by simp show \UNIV \ ?A\ proof fix w :: \'a word\ show \w \ (word_of_nat ` {0..<2 ^ LENGTH('a)} :: 'a word set)\ by (rule image_eqI [of _ _ \unat w\]; transfer) simp_all qed qed instantiation word :: (len) enum begin definition enum_word :: \'a word list\ where \enum_word = map word_of_nat [0..<2 ^ LENGTH('a)]\ definition enum_all_word :: \('a word \ bool) \ bool\ where \enum_all_word = Ball UNIV\ definition enum_ex_word :: \('a word \ bool) \ bool\ where \enum_ex_word = Bex UNIV\ lemma [code]: \Enum.enum_all P \ Ball UNIV P\ \Enum.enum_ex P \ Bex UNIV P\ for P :: \'a word \ bool\ by (simp_all add: enum_all_word_def enum_ex_word_def) instance by standard (simp_all add: UNIV_word_eq_word_of_nat inj_on_word_of_nat enum_word_def enum_all_word_def enum_ex_word_def distinct_map) end subsection \Bit-wise operations\ instantiation word :: (len) semiring_modulo begin lift_definition divide_word :: \'a word \ 'a word \ 'a word\ is \\a b. take_bit LENGTH('a) a div take_bit LENGTH('a) b\ by simp lift_definition modulo_word :: \'a word \ 'a word \ 'a word\ is \\a b. take_bit LENGTH('a) a mod take_bit LENGTH('a) b\ by simp instance proof show "a div b * b + a mod b = a" for a b :: "'a word" proof transfer fix k l :: int define r :: int where "r = 2 ^ LENGTH('a)" then have r: "take_bit LENGTH('a) k = k mod r" for k by (simp add: take_bit_eq_mod) have "k mod r = ((k mod r) div (l mod r) * (l mod r) + (k mod r) mod (l mod r)) mod r" by (simp add: div_mult_mod_eq) also have "... = (((k mod r) div (l mod r) * (l mod r)) mod r + (k mod r) mod (l mod r)) mod r" by (simp add: mod_add_left_eq) also have "... = (((k mod r) div (l mod r) * l) mod r + (k mod r) mod (l mod r)) mod r" by (simp add: mod_mult_right_eq) finally have "k mod r = ((k mod r) div (l mod r) * l + (k mod r) mod (l mod r)) mod r" by (simp add: mod_simps) with r show "take_bit LENGTH('a) (take_bit LENGTH('a) k div take_bit LENGTH('a) l * l + take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = take_bit LENGTH('a) k" by simp qed qed end instance word :: (len) semiring_parity proof show "\ 2 dvd (1::'a word)" by transfer simp show even_iff_mod_2_eq_0: "2 dvd a \ a mod 2 = 0" for a :: "'a word" by transfer (simp_all add: mod_2_eq_odd take_bit_Suc) show "\ 2 dvd a \ a mod 2 = 1" for a :: "'a word" by transfer (simp_all add: mod_2_eq_odd take_bit_Suc) qed lemma word_bit_induct [case_names zero even odd]: \P a\ if word_zero: \P 0\ and word_even: \\a. P a \ 0 < a \ a < 2 ^ (LENGTH('a) - Suc 0) \ P (2 * a)\ and word_odd: \\a. P a \ a < 2 ^ (LENGTH('a) - Suc 0) \ P (1 + 2 * a)\ for P and a :: \'a::len word\ proof - define m :: nat where \m = LENGTH('a) - Suc 0\ then have l: \LENGTH('a) = Suc m\ by simp define n :: nat where \n = unat a\ then have \n < 2 ^ LENGTH('a)\ by transfer (simp add: take_bit_eq_mod) then have \n < 2 * 2 ^ m\ by (simp add: l) then have \P (of_nat n)\ proof (induction n rule: nat_bit_induct) case zero show ?case by simp (rule word_zero) next case (even n) then have \n < 2 ^ m\ by simp with even.IH have \P (of_nat n)\ by simp moreover from \n < 2 ^ m\ even.hyps have \0 < (of_nat n :: 'a word)\ by (auto simp add: word_greater_zero_iff l) moreover from \n < 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\ using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] by (simp add: l take_bit_eq_mod) ultimately have \P (2 * of_nat n)\ by (rule word_even) then show ?case by simp next case (odd n) then have \Suc n \ 2 ^ m\ by simp with odd.IH have \P (of_nat n)\ by simp moreover from \Suc n \ 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\ using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] by (simp add: l take_bit_eq_mod) ultimately have \P (1 + 2 * of_nat n)\ by (rule word_odd) then show ?case by simp qed moreover have \of_nat (nat (uint a)) = a\ by transfer simp ultimately show ?thesis by (simp add: n_def) qed lemma bit_word_half_eq: \(of_bool b + a * 2) div 2 = a\ if \a < 2 ^ (LENGTH('a) - Suc 0)\ for a :: \'a::len word\ proof (cases \2 \ LENGTH('a::len)\) case False have \of_bool (odd k) < (1 :: int) \ even k\ for k :: int by auto with False that show ?thesis by transfer (simp add: eq_iff) next case True obtain n where length: \LENGTH('a) = Suc n\ by (cases \LENGTH('a)\) simp_all show ?thesis proof (cases b) case False moreover have \a * 2 div 2 = a\ using that proof transfer fix k :: int from length have \k * 2 mod 2 ^ LENGTH('a) = (k mod 2 ^ n) * 2\ by simp moreover assume \take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\ with \LENGTH('a) = Suc n\ have \k mod 2 ^ LENGTH('a) = k mod 2 ^ n\ by (simp add: take_bit_eq_mod divmod_digit_0) ultimately have \take_bit LENGTH('a) (k * 2) = take_bit LENGTH('a) k * 2\ by (simp add: take_bit_eq_mod) with True show \take_bit LENGTH('a) (take_bit LENGTH('a) (k * 2) div take_bit LENGTH('a) 2) = take_bit LENGTH('a) k\ by simp qed ultimately show ?thesis by simp next case True moreover have \(1 + a * 2) div 2 = a\ using that proof transfer fix k :: int from length have \(1 + k * 2) mod 2 ^ LENGTH('a) = 1 + (k mod 2 ^ n) * 2\ using pos_zmod_mult_2 [of \2 ^ n\ k] by (simp add: ac_simps) moreover assume \take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\ with \LENGTH('a) = Suc n\ have \k mod 2 ^ LENGTH('a) = k mod 2 ^ n\ by (simp add: take_bit_eq_mod divmod_digit_0) ultimately have \take_bit LENGTH('a) (1 + k * 2) = 1 + take_bit LENGTH('a) k * 2\ by (simp add: take_bit_eq_mod) with True show \take_bit LENGTH('a) (take_bit LENGTH('a) (1 + k * 2) div take_bit LENGTH('a) 2) = take_bit LENGTH('a) k\ by (auto simp add: take_bit_Suc) qed ultimately show ?thesis by simp qed qed lemma even_mult_exp_div_word_iff: \even (a * 2 ^ m div 2 ^ n) \ \ ( m \ n \ n < LENGTH('a) \ odd (a div 2 ^ (n - m)))\ for a :: \'a::len word\ by transfer (auto simp flip: drop_bit_eq_div simp add: even_drop_bit_iff_not_bit bit_take_bit_iff, simp_all flip: push_bit_eq_mult add: bit_push_bit_iff_int) instantiation word :: (len) semiring_bits begin lift_definition bit_word :: \'a word \ nat \ bool\ is \\k n. n < LENGTH('a) \ bit k n\ proof fix k l :: int and n :: nat assume *: \take_bit LENGTH('a) k = take_bit LENGTH('a) l\ show \n < LENGTH('a) \ bit k n \ n < LENGTH('a) \ bit l n\ proof (cases \n < LENGTH('a)\) case True from * have \bit (take_bit LENGTH('a) k) n \ bit (take_bit LENGTH('a) l) n\ by simp then show ?thesis by (simp add: bit_take_bit_iff) next case False then show ?thesis by simp qed qed instance proof show \P a\ if stable: \\a. a div 2 = a \ P a\ and rec: \\a b. P a \ (of_bool b + 2 * a) div 2 = a \ P (of_bool b + 2 * a)\ for P and a :: \'a word\ proof (induction a rule: word_bit_induct) case zero have \0 div 2 = (0::'a word)\ by transfer simp with stable [of 0] show ?case by simp next case (even a) with rec [of a False] show ?case using bit_word_half_eq [of a False] by (simp add: ac_simps) next case (odd a) with rec [of a True] show ?case using bit_word_half_eq [of a True] by (simp add: ac_simps) qed show \bit a n \ odd (a div 2 ^ n)\ for a :: \'a word\ and n by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit bit_iff_odd_drop_bit) show \0 div a = 0\ for a :: \'a word\ by transfer simp show \a div 1 = a\ for a :: \'a word\ by transfer simp show \a mod b div b = 0\ for a b :: \'a word\ apply transfer apply (simp add: take_bit_eq_mod) apply (subst (3) mod_pos_pos_trivial [of _ \2 ^ LENGTH('a)\]) apply simp_all apply (metis le_less mod_by_0 pos_mod_conj zero_less_numeral zero_less_power) using pos_mod_bound [of \2 ^ LENGTH('a)\] apply simp proof - fix aa :: int and ba :: int have f1: "\i n. (i::int) mod 2 ^ n = 0 \ 0 < i mod 2 ^ n" by (metis le_less take_bit_eq_mod take_bit_nonnegative) have "(0::int) < 2 ^ len_of (TYPE('a)::'a itself) \ ba mod 2 ^ len_of (TYPE('a)::'a itself) \ 0 \ aa mod 2 ^ len_of (TYPE('a)::'a itself) mod (ba mod 2 ^ len_of (TYPE('a)::'a itself)) < 2 ^ len_of (TYPE('a)::'a itself)" by (metis (no_types) mod_by_0 unique_euclidean_semiring_numeral_class.pos_mod_bound zero_less_numeral zero_less_power) then show "aa mod 2 ^ len_of (TYPE('a)::'a itself) mod (ba mod 2 ^ len_of (TYPE('a)::'a itself)) < 2 ^ len_of (TYPE('a)::'a itself)" using f1 by (meson le_less less_le_trans unique_euclidean_semiring_numeral_class.pos_mod_bound) qed show \(1 + a) div 2 = a div 2\ if \even a\ for a :: \'a word\ using that by transfer (auto dest: le_Suc_ex simp add: mod_2_eq_odd take_bit_Suc elim!: evenE) show \(2 :: 'a word) ^ m div 2 ^ n = of_bool ((2 :: 'a word) ^ m \ 0 \ n \ m) * 2 ^ (m - n)\ for m n :: nat by transfer (simp, simp add: exp_div_exp_eq) show "a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)" for a :: "'a word" and m n :: nat apply transfer apply (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: drop_bit_eq_div) apply (simp add: drop_bit_take_bit) done show "a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n" for a :: "'a word" and m n :: nat by transfer (auto simp flip: take_bit_eq_mod simp add: ac_simps) show \a * 2 ^ m mod 2 ^ n = a mod 2 ^ (n - m) * 2 ^ m\ if \m \ n\ for a :: "'a word" and m n :: nat using that apply transfer apply (auto simp flip: take_bit_eq_mod) apply (auto simp flip: push_bit_eq_mult simp add: push_bit_take_bit split: split_min_lin) done show \a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\ for a :: "'a word" and m n :: nat by transfer (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: take_bit_eq_mod drop_bit_eq_div split: split_min_lin) show \even ((2 ^ m - 1) div (2::'a word) ^ n) \ 2 ^ n = (0::'a word) \ m \ n\ for m n :: nat by transfer (auto simp add: take_bit_of_mask even_mask_div_iff) show \even (a * 2 ^ m div 2 ^ n) \ n < m \ (2::'a word) ^ n = 0 \ m \ n \ even (a div 2 ^ (n - m))\ for a :: \'a word\ and m n :: nat proof transfer show \even (take_bit LENGTH('a) (k * 2 ^ m) div take_bit LENGTH('a) (2 ^ n)) \ n < m \ take_bit LENGTH('a) ((2::int) ^ n) = take_bit LENGTH('a) 0 \ (m \ n \ even (take_bit LENGTH('a) k div take_bit LENGTH('a) (2 ^ (n - m))))\ for m n :: nat and k l :: int by (auto simp flip: take_bit_eq_mod drop_bit_eq_div push_bit_eq_mult simp add: div_push_bit_of_1_eq_drop_bit drop_bit_take_bit drop_bit_push_bit_int [of n m]) qed qed end lemma bit_word_eqI: \a = b\ if \\n. n < LENGTH('a) \ bit a n \ bit b n\ for a b :: \'a::len word\ using that by transfer (auto simp add: nat_less_le bit_eq_iff bit_take_bit_iff) lemma bit_imp_le_length: \n < LENGTH('a)\ if \bit w n\ for w :: \'a::len word\ using that by transfer simp lemma not_bit_length [simp]: \\ bit w LENGTH('a)\ for w :: \'a::len word\ by transfer simp instantiation word :: (len) semiring_bit_shifts begin lift_definition push_bit_word :: \nat \ 'a word \ 'a word\ is push_bit proof - show \take_bit LENGTH('a) (push_bit n k) = take_bit LENGTH('a) (push_bit n l)\ if \take_bit LENGTH('a) k = take_bit LENGTH('a) l\ for k l :: int and n :: nat proof - from that have \take_bit (LENGTH('a) - n) (take_bit LENGTH('a) k) = take_bit (LENGTH('a) - n) (take_bit LENGTH('a) l)\ by simp moreover have \min (LENGTH('a) - n) LENGTH('a) = LENGTH('a) - n\ by simp ultimately show ?thesis by (simp add: take_bit_push_bit) qed qed lift_definition drop_bit_word :: \nat \ 'a word \ 'a word\ is \\n. drop_bit n \ take_bit LENGTH('a)\ by (simp add: take_bit_eq_mod) lift_definition take_bit_word :: \nat \ 'a word \ 'a word\ is \\n. take_bit (min LENGTH('a) n)\ by (simp add: ac_simps) (simp only: flip: take_bit_take_bit) instance proof show \push_bit n a = a * 2 ^ n\ for n :: nat and a :: \'a word\ by transfer (simp add: push_bit_eq_mult) show \drop_bit n a = a div 2 ^ n\ for n :: nat and a :: \'a word\ by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit) show \take_bit n a = a mod 2 ^ n\ for n :: nat and a :: \'a word\ by transfer (auto simp flip: take_bit_eq_mod) qed end lemma [code]: \Word.the_int (take_bit n w) = (if n < LENGTH('a::len) then take_bit n (Word.the_int w) else Word.the_int w)\ for w :: \'a::len word\ by transfer (simp add: not_le not_less ac_simps min_absorb2) instantiation word :: (len) ring_bit_operations begin lift_definition not_word :: \'a word \ 'a word\ is not by (simp add: take_bit_not_iff) lift_definition and_word :: \'a word \ 'a word \ 'a word\ is \and\ by simp lift_definition or_word :: \'a word \ 'a word \ 'a word\ is or by simp lift_definition xor_word :: \'a word \ 'a word \ 'a word\ is xor by simp lift_definition mask_word :: \nat \ 'a word\ is mask . instance by (standard; transfer) (auto simp add: minus_eq_not_minus_1 mask_eq_exp_minus_1 bit_not_iff bit_and_iff bit_or_iff bit_xor_iff) end lemma [code_abbrev]: \push_bit n 1 = (2 :: 'a::len word) ^ n\ by (fact push_bit_of_1) lemma [code]: \NOT w = Word.of_int (NOT (Word.the_int w))\ for w :: \'a::len word\ by transfer (simp add: take_bit_not_take_bit) lemma [code]: \Word.the_int (v AND w) = Word.the_int v AND Word.the_int w\ by transfer simp lemma [code]: \Word.the_int (v OR w) = Word.the_int v OR Word.the_int w\ by transfer simp lemma [code]: \Word.the_int (v XOR w) = Word.the_int v XOR Word.the_int w\ by transfer simp lemma [code]: \Word.the_int (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ by transfer simp context includes lifting_syntax begin lemma set_bit_word_transfer [transfer_rule]: \((=) ===> pcr_word ===> pcr_word) set_bit set_bit\ by (unfold set_bit_def) transfer_prover lemma unset_bit_word_transfer [transfer_rule]: \((=) ===> pcr_word ===> pcr_word) unset_bit unset_bit\ by (unfold unset_bit_def) transfer_prover lemma flip_bit_word_transfer [transfer_rule]: \((=) ===> pcr_word ===> pcr_word) flip_bit flip_bit\ by (unfold flip_bit_def) transfer_prover lemma signed_take_bit_word_transfer [transfer_rule]: \((=) ===> pcr_word ===> pcr_word) (\n k. signed_take_bit n (take_bit LENGTH('a::len) k)) (signed_take_bit :: nat \ 'a word \ 'a word)\ proof - let ?K = \\n (k :: int). take_bit (min LENGTH('a) n) k OR of_bool (n < LENGTH('a) \ bit k n) * NOT (mask n)\ let ?W = \\n (w :: 'a word). take_bit n w OR of_bool (bit w n) * NOT (mask n)\ have \((=) ===> pcr_word ===> pcr_word) ?K ?W\ by transfer_prover also have \?K = (\n k. signed_take_bit n (take_bit LENGTH('a::len) k))\ by (simp add: fun_eq_iff signed_take_bit_def bit_take_bit_iff ac_simps) also have \?W = signed_take_bit\ by (simp add: fun_eq_iff signed_take_bit_def) finally show ?thesis . qed end subsection \Conversions including casts\ subsubsection \Generic unsigned conversion\ context semiring_bits begin lemma bit_unsigned_iff: \bit (unsigned w) n \ 2 ^ n \ 0 \ bit w n\ for w :: \'b::len word\ by (transfer fixing: bit) (simp add: bit_of_nat_iff bit_nat_iff bit_take_bit_iff) end context semiring_bit_shifts begin lemma unsigned_push_bit_eq: \unsigned (push_bit n w) = take_bit LENGTH('b) (push_bit n (unsigned w))\ for w :: \'b::len word\ proof (rule bit_eqI) fix m assume \2 ^ m \ 0\ show \bit (unsigned (push_bit n w)) m = bit (take_bit LENGTH('b) (push_bit n (unsigned w))) m\ proof (cases \n \ m\) case True with \2 ^ m \ 0\ have \2 ^ (m - n) \ 0\ by (metis (full_types) diff_add exp_add_not_zero_imp) with True show ?thesis by (simp add: bit_unsigned_iff bit_push_bit_iff Parity.bit_push_bit_iff bit_take_bit_iff not_le exp_eq_zero_iff ac_simps) next case False then show ?thesis by (simp add: not_le bit_unsigned_iff bit_push_bit_iff Parity.bit_push_bit_iff bit_take_bit_iff) qed qed lemma unsigned_take_bit_eq: \unsigned (take_bit n w) = take_bit n (unsigned w)\ for w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_unsigned_iff bit_take_bit_iff Parity.bit_take_bit_iff) end context semiring_bit_operations begin lemma unsigned_and_eq: \unsigned (v AND w) = unsigned v AND unsigned w\ for v w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_unsigned_iff bit_and_iff Bit_Operations.bit_and_iff) lemma unsigned_or_eq: \unsigned (v OR w) = unsigned v OR unsigned w\ for v w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_unsigned_iff bit_or_iff Bit_Operations.bit_or_iff) lemma unsigned_xor_eq: \unsigned (v XOR w) = unsigned v XOR unsigned w\ for v w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_unsigned_iff bit_xor_iff Bit_Operations.bit_xor_iff) end context ring_bit_operations begin lemma unsigned_not_eq: \unsigned (NOT w) = take_bit LENGTH('b) (NOT (unsigned w))\ for w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_unsigned_iff bit_take_bit_iff bit_not_iff Bit_Operations.bit_not_iff exp_eq_zero_iff not_le) end context unique_euclidean_semiring_numeral begin lemma unsigned_greater_eq [simp]: \0 \ unsigned w\ for w :: \'b::len word\ by (transfer fixing: less_eq) simp lemma unsigned_less [simp]: \unsigned w < 2 ^ LENGTH('b)\ for w :: \'b::len word\ by (transfer fixing: less) simp end context linordered_semidom begin lemma word_less_eq_iff_unsigned: "a \ b \ unsigned a \ unsigned b" by (transfer fixing: less_eq) (simp add: nat_le_eq_zle) lemma word_less_iff_unsigned: "a < b \ unsigned a < unsigned b" by (transfer fixing: less) (auto dest: preorder_class.le_less_trans [OF take_bit_nonnegative]) end subsubsection \Generic signed conversion\ context ring_bit_operations begin lemma bit_signed_iff: \bit (signed w) n \ 2 ^ n \ 0 \ bit w (min (LENGTH('b) - Suc 0) n)\ for w :: \'b::len word\ by (transfer fixing: bit) (auto simp add: bit_of_int_iff Bit_Operations.bit_signed_take_bit_iff min_def) lemma signed_push_bit_eq: \signed (push_bit n w) = signed_take_bit (LENGTH('b) - Suc 0) (push_bit n (signed w :: 'a))\ for w :: \'b::len word\ proof (rule bit_eqI) fix m assume \2 ^ m \ 0\ define q where \q = LENGTH('b) - Suc 0\ then have *: \LENGTH('b) = Suc q\ by simp show \bit (signed (push_bit n w)) m \ bit (signed_take_bit (LENGTH('b) - Suc 0) (push_bit n (signed w :: 'a))) m\ proof (cases \q \ m\) case True moreover define r where \r = m - q\ ultimately have \m = q + r\ by simp moreover from \m = q + r\ \2 ^ m \ 0\ have \2 ^ q \ 0\ \2 ^ r \ 0\ using exp_add_not_zero_imp_left [of q r] exp_add_not_zero_imp_right [of q r] by simp_all moreover from \2 ^ q \ 0\ have \2 ^ (q - n) \ 0\ by (rule exp_not_zero_imp_exp_diff_not_zero) ultimately show ?thesis by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Parity.bit_push_bit_iff min_def * exp_eq_zero_iff le_diff_conv2) next case False then show ?thesis using exp_not_zero_imp_exp_diff_not_zero [of m n] by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Parity.bit_push_bit_iff min_def not_le not_less * le_diff_conv2 less_diff_conv2 Parity.exp_eq_0_imp_not_bit exp_eq_0_imp_not_bit exp_eq_zero_iff) qed qed lemma signed_take_bit_eq: \signed (take_bit n w) = (if n < LENGTH('b) then take_bit n (signed w) else signed w)\ for w :: \'b::len word\ by (transfer fixing: take_bit; cases \LENGTH('b)\) (auto simp add: Bit_Operations.signed_take_bit_take_bit Bit_Operations.take_bit_signed_take_bit take_bit_of_int min_def less_Suc_eq) lemma signed_not_eq: \signed (NOT w) = signed_take_bit LENGTH('b) (NOT (signed w))\ for w :: \'b::len word\ proof (rule bit_eqI) fix n assume \2 ^ n \ 0\ define q where \q = LENGTH('b) - Suc 0\ then have *: \LENGTH('b) = Suc q\ by simp show \bit (signed (NOT w)) n \ bit (signed_take_bit LENGTH('b) (NOT (signed w))) n\ proof (cases \q < n\) case True moreover define r where \r = n - Suc q\ ultimately have \n = r + Suc q\ by simp moreover from \2 ^ n \ 0\ \n = r + Suc q\ have \2 ^ Suc q \ 0\ using exp_add_not_zero_imp_right by blast ultimately show ?thesis by (simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def exp_eq_zero_iff) next case False then show ?thesis by (auto simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def exp_eq_zero_iff) qed qed lemma signed_and_eq: \signed (v AND w) = signed v AND signed w\ for v w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_signed_iff bit_and_iff Bit_Operations.bit_and_iff) lemma signed_or_eq: \signed (v OR w) = signed v OR signed w\ for v w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_signed_iff bit_or_iff Bit_Operations.bit_or_iff) lemma signed_xor_eq: \signed (v XOR w) = signed v XOR signed w\ for v w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_signed_iff bit_xor_iff Bit_Operations.bit_xor_iff) end subsubsection \More\ lemma sint_greater_eq: \- (2 ^ (LENGTH('a) - Suc 0)) \ sint w\ for w :: \'a::len word\ proof (cases \bit w (LENGTH('a) - Suc 0)\) case True then show ?thesis by transfer (simp add: signed_take_bit_eq_if_negative minus_exp_eq_not_mask or_greater_eq ac_simps) next have *: \- (2 ^ (LENGTH('a) - Suc 0)) \ (0::int)\ by simp case False then show ?thesis by transfer (auto simp add: signed_take_bit_eq intro: order_trans *) qed lemma sint_less: \sint w < 2 ^ (LENGTH('a) - Suc 0)\ for w :: \'a::len word\ by (cases \bit w (LENGTH('a) - Suc 0)\; transfer) (simp_all add: signed_take_bit_eq signed_take_bit_def not_eq_complement mask_eq_exp_minus_1 OR_upper) lemma unat_div_distrib: \unat (v div w) = unat v div unat w\ proof transfer fix k l have \nat (take_bit LENGTH('a) k) div nat (take_bit LENGTH('a) l) \ nat (take_bit LENGTH('a) k)\ by (rule div_le_dividend) also have \nat (take_bit LENGTH('a) k) < 2 ^ LENGTH('a)\ by (simp add: nat_less_iff) finally show \(nat \ take_bit LENGTH('a)) (take_bit LENGTH('a) k div take_bit LENGTH('a) l) = (nat \ take_bit LENGTH('a)) k div (nat \ take_bit LENGTH('a)) l\ by (simp add: nat_take_bit_eq div_int_pos_iff nat_div_distrib take_bit_nat_eq_self_iff) qed lemma unat_mod_distrib: \unat (v mod w) = unat v mod unat w\ proof transfer fix k l have \nat (take_bit LENGTH('a) k) mod nat (take_bit LENGTH('a) l) \ nat (take_bit LENGTH('a) k)\ by (rule mod_less_eq_dividend) also have \nat (take_bit LENGTH('a) k) < 2 ^ LENGTH('a)\ by (simp add: nat_less_iff) finally show \(nat \ take_bit LENGTH('a)) (take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = (nat \ take_bit LENGTH('a)) k mod (nat \ take_bit LENGTH('a)) l\ by (simp add: nat_take_bit_eq mod_int_pos_iff less_le nat_mod_distrib take_bit_nat_eq_self_iff) qed lemma uint_div_distrib: \uint (v div w) = uint v div uint w\ proof - have \int (unat (v div w)) = int (unat v div unat w)\ by (simp add: unat_div_distrib) then show ?thesis by (simp add: of_nat_div) qed lemma unat_drop_bit_eq: \unat (drop_bit n w) = drop_bit n (unat w)\ by (rule bit_eqI) (simp add: bit_unsigned_iff bit_drop_bit_eq) lemma uint_mod_distrib: \uint (v mod w) = uint v mod uint w\ proof - have \int (unat (v mod w)) = int (unat v mod unat w)\ by (simp add: unat_mod_distrib) then show ?thesis by (simp add: of_nat_mod) qed context semiring_bit_shifts begin lemma unsigned_ucast_eq: \unsigned (ucast w :: 'c::len word) = take_bit LENGTH('c) (unsigned w)\ for w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_unsigned_iff Word.bit_unsigned_iff bit_take_bit_iff exp_eq_zero_iff not_le) end context ring_bit_operations begin lemma signed_ucast_eq: \signed (ucast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (unsigned w)\ for w :: \'b::len word\ proof (rule bit_eqI) fix n assume \2 ^ n \ 0\ then have \2 ^ (min (LENGTH('c) - Suc 0) n) \ 0\ by (simp add: min_def) (metis (mono_tags) diff_diff_cancel exp_not_zero_imp_exp_diff_not_zero) then show \bit (signed (ucast w :: 'c::len word)) n \ bit (signed_take_bit (LENGTH('c) - Suc 0) (unsigned w)) n\ by (simp add: bit_signed_iff bit_unsigned_iff Word.bit_unsigned_iff bit_signed_take_bit_iff exp_eq_zero_iff not_le) qed lemma signed_scast_eq: \signed (scast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (signed w)\ for w :: \'b::len word\ proof (rule bit_eqI) fix n assume \2 ^ n \ 0\ then have \2 ^ (min (LENGTH('c) - Suc 0) n) \ 0\ by (simp add: min_def) (metis (mono_tags) diff_diff_cancel exp_not_zero_imp_exp_diff_not_zero) then show \bit (signed (scast w :: 'c::len word)) n \ bit (signed_take_bit (LENGTH('c) - Suc 0) (signed w)) n\ by (simp add: bit_signed_iff bit_unsigned_iff Word.bit_signed_iff bit_signed_take_bit_iff exp_eq_zero_iff not_le) qed end lemma uint_nonnegative: "0 \ uint w" by (fact unsigned_greater_eq) lemma uint_bounded: "uint w < 2 ^ LENGTH('a)" for w :: "'a::len word" by (fact unsigned_less) lemma uint_idem: "uint w mod 2 ^ LENGTH('a) = uint w" for w :: "'a::len word" by transfer (simp add: take_bit_eq_mod) lemma word_uint_eqI: "uint a = uint b \ a = b" by (fact unsigned_word_eqI) lemma word_uint_eq_iff: "a = b \ uint a = uint b" by (fact word_eq_iff_unsigned) lemma uint_word_of_int_eq: \uint (word_of_int k :: 'a::len word) = take_bit LENGTH('a) k\ by transfer rule lemma uint_word_of_int: "uint (word_of_int k :: 'a::len word) = k mod 2 ^ LENGTH('a)" by (simp add: uint_word_of_int_eq take_bit_eq_mod) lemma word_of_int_uint: "word_of_int (uint w) = w" by transfer simp lemma word_div_def [code]: "a div b = word_of_int (uint a div uint b)" by transfer rule lemma word_mod_def [code]: "a mod b = word_of_int (uint a mod uint b)" by transfer rule lemma split_word_all: "(\x::'a::len word. PROP P x) \ (\x. PROP P (word_of_int x))" proof fix x :: "'a word" assume "\x. PROP P (word_of_int x)" then have "PROP P (word_of_int (uint x))" . then show "PROP P x" by (simp only: word_of_int_uint) qed lemma sint_uint: \sint w = signed_take_bit (LENGTH('a) - Suc 0) (uint w)\ for w :: \'a::len word\ by (cases \LENGTH('a)\; transfer) (simp_all add: signed_take_bit_take_bit) lemma unat_eq_nat_uint: \unat w = nat (uint w)\ by simp lemma ucast_eq: \ucast w = word_of_int (uint w)\ by transfer simp lemma scast_eq: \scast w = word_of_int (sint w)\ by transfer simp lemma uint_0_eq: \uint 0 = 0\ by (fact unsigned_0) lemma uint_1_eq: \uint 1 = 1\ by (fact unsigned_1) lemma word_m1_wi: "- 1 = word_of_int (- 1)" by simp lemma uint_0_iff: "uint x = 0 \ x = 0" by (auto simp add: unsigned_word_eqI) lemma unat_0_iff: "unat x = 0 \ x = 0" by (auto simp add: unsigned_word_eqI) lemma unat_0: "unat 0 = 0" by (fact unsigned_0) lemma unat_gt_0: "0 < unat x \ x \ 0" by (auto simp: unat_0_iff [symmetric]) lemma ucast_0: "ucast 0 = 0" by (fact unsigned_0) lemma sint_0: "sint 0 = 0" by (fact signed_0) lemma scast_0: "scast 0 = 0" by (fact signed_0) lemma sint_n1: "sint (- 1) = - 1" by (fact signed_minus_1) lemma scast_n1: "scast (- 1) = - 1" by (fact signed_minus_1) lemma uint_1: "uint (1::'a::len word) = 1" by (fact uint_1_eq) lemma unat_1: "unat (1::'a::len word) = 1" by (fact unsigned_1) lemma ucast_1: "ucast (1::'a::len word) = 1" by (fact unsigned_1) instantiation word :: (len) size begin lift_definition size_word :: \'a word \ nat\ is \\_. LENGTH('a)\ .. instance .. end lemma word_size [code]: \size w = LENGTH('a)\ for w :: \'a::len word\ by (fact size_word.rep_eq) lemma word_size_gt_0 [iff]: "0 < size w" for w :: "'a::len word" by (simp add: word_size) lemmas lens_gt_0 = word_size_gt_0 len_gt_0 lemma lens_not_0 [iff]: \size w \ 0\ for w :: \'a::len word\ by auto lift_definition source_size :: \('a::len word \ 'b) \ nat\ is \\_. LENGTH('a)\ . lift_definition target_size :: \('a \ 'b::len word) \ nat\ is \\_. LENGTH('b)\ .. lift_definition is_up :: \('a::len word \ 'b::len word) \ bool\ is \\_. LENGTH('a) \ LENGTH('b)\ .. lift_definition is_down :: \('a::len word \ 'b::len word) \ bool\ is \\_. LENGTH('a) \ LENGTH('b)\ .. lemma is_up_eq: \is_up f \ source_size f \ target_size f\ for f :: \'a::len word \ 'b::len word\ by (simp add: source_size.rep_eq target_size.rep_eq is_up.rep_eq) lemma is_down_eq: \is_down f \ target_size f \ source_size f\ for f :: \'a::len word \ 'b::len word\ by (simp add: source_size.rep_eq target_size.rep_eq is_down.rep_eq) lift_definition word_int_case :: \(int \ 'b) \ 'a::len word \ 'b\ is \\f. f \ take_bit LENGTH('a)\ by simp lemma word_int_case_eq_uint [code]: \word_int_case f w = f (uint w)\ by transfer simp translations "case x of XCONST of_int y \ b" \ "CONST word_int_case (\y. b) x" "case x of (XCONST of_int :: 'a) y \ b" \ "CONST word_int_case (\y. b) x" subsection \Arithmetic operations\ text \Legacy theorems:\ lemma word_add_def [code]: "a + b = word_of_int (uint a + uint b)" by transfer (simp add: take_bit_add) lemma word_sub_wi [code]: "a - b = word_of_int (uint a - uint b)" by transfer (simp add: take_bit_diff) lemma word_mult_def [code]: "a * b = word_of_int (uint a * uint b)" by transfer (simp add: take_bit_eq_mod mod_simps) lemma word_minus_def [code]: "- a = word_of_int (- uint a)" by transfer (simp add: take_bit_minus) lemma word_0_wi: "0 = word_of_int 0" by transfer simp lemma word_1_wi: "1 = word_of_int 1" by transfer simp lift_definition word_succ :: "'a::len word \ 'a word" is "\x. x + 1" by (auto simp add: take_bit_eq_mod intro: mod_add_cong) lift_definition word_pred :: "'a::len word \ 'a word" is "\x. x - 1" by (auto simp add: take_bit_eq_mod intro: mod_diff_cong) lemma word_succ_alt [code]: "word_succ a = word_of_int (uint a + 1)" by transfer (simp add: take_bit_eq_mod mod_simps) lemma word_pred_alt [code]: "word_pred a = word_of_int (uint a - 1)" by transfer (simp add: take_bit_eq_mod mod_simps) lemmas word_arith_wis = word_add_def word_sub_wi word_mult_def word_minus_def word_succ_alt word_pred_alt word_0_wi word_1_wi lemma wi_homs: shows wi_hom_add: "word_of_int a + word_of_int b = word_of_int (a + b)" and wi_hom_sub: "word_of_int a - word_of_int b = word_of_int (a - b)" and wi_hom_mult: "word_of_int a * word_of_int b = word_of_int (a * b)" and wi_hom_neg: "- word_of_int a = word_of_int (- a)" and wi_hom_succ: "word_succ (word_of_int a) = word_of_int (a + 1)" and wi_hom_pred: "word_pred (word_of_int a) = word_of_int (a - 1)" by (transfer, simp)+ lemmas wi_hom_syms = wi_homs [symmetric] lemmas word_of_int_homs = wi_homs word_0_wi word_1_wi lemmas word_of_int_hom_syms = word_of_int_homs [symmetric] lemma double_eq_zero_iff: \2 * a = 0 \ a = 0 \ a = 2 ^ (LENGTH('a) - Suc 0)\ for a :: \'a::len word\ proof - define n where \n = LENGTH('a) - Suc 0\ then have *: \LENGTH('a) = Suc n\ by simp have \a = 0\ if \2 * a = 0\ and \a \ 2 ^ (LENGTH('a) - Suc 0)\ using that by transfer (auto simp add: take_bit_eq_0_iff take_bit_eq_mod *) moreover have \2 ^ LENGTH('a) = (0 :: 'a word)\ by transfer simp then have \2 * 2 ^ (LENGTH('a) - Suc 0) = (0 :: 'a word)\ by (simp add: *) ultimately show ?thesis by auto qed subsection \Ordering\ lift_definition word_sle :: \'a::len word \ 'a word \ bool\ is \\k l. signed_take_bit (LENGTH('a) - Suc 0) k \ signed_take_bit (LENGTH('a) - Suc 0) l\ by (simp flip: signed_take_bit_decr_length_iff) lift_definition word_sless :: \'a::len word \ 'a word \ bool\ is \\k l. signed_take_bit (LENGTH('a) - Suc 0) k < signed_take_bit (LENGTH('a) - Suc 0) l\ by (simp flip: signed_take_bit_decr_length_iff) notation word_sle ("'(\s')") and word_sle ("(_/ \s _)" [51, 51] 50) and word_sless ("'(a <=s b \ sint a \ sint b\ by transfer simp lemma [code]: \a sint a < sint b\ by transfer simp lemma signed_ordering: \ordering word_sle word_sless\ apply (standard; transfer) apply simp_all using signed_take_bit_decr_length_iff apply force using signed_take_bit_decr_length_iff apply force done lemma signed_linorder: \class.linorder word_sle word_sless\ by (standard; transfer) (auto simp add: signed_take_bit_decr_length_iff) interpretation signed: linorder word_sle word_sless by (fact signed_linorder) lemma word_sless_eq: \x x <=s y \ x \ y\ by (fact signed.less_le) lemma word_less_alt: "a < b \ uint a < uint b" by (fact word_less_def) lemma word_zero_le [simp]: "0 \ y" for y :: "'a::len word" by (fact word_coorder.extremum) lemma word_m1_ge [simp] : "word_pred 0 \ y" (* FIXME: delete *) by transfer (simp add: take_bit_minus_one_eq_mask mask_eq_exp_minus_1 bintr_lt2p) lemma word_n1_ge [simp]: "y \ -1" for y :: "'a::len word" by (fact word_order.extremum) lemmas word_not_simps [simp] = word_zero_le [THEN leD] word_m1_ge [THEN leD] word_n1_ge [THEN leD] lemma word_gt_0: "0 < y \ 0 \ y" for y :: "'a::len word" by (simp add: less_le) lemmas word_gt_0_no [simp] = word_gt_0 [of "numeral y"] for y lemma word_sless_alt: "a sint a < sint b" by transfer simp lemma word_le_nat_alt: "a \ b \ unat a \ unat b" by transfer (simp add: nat_le_eq_zle) lemma word_less_nat_alt: "a < b \ unat a < unat b" by transfer (auto simp add: less_le [of 0]) lemmas unat_mono = word_less_nat_alt [THEN iffD1] instance word :: (len) wellorder proof fix P :: "'a word \ bool" and a assume *: "(\b. (\a. a < b \ P a) \ P b)" have "wf (measure unat)" .. moreover have "{(a, b :: ('a::len) word). a < b} \ measure unat" by (auto simp add: word_less_nat_alt) ultimately have "wf {(a, b :: ('a::len) word). a < b}" by (rule wf_subset) then show "P a" using * by induction blast qed lemma wi_less: "(word_of_int n < (word_of_int m :: 'a::len word)) = (n mod 2 ^ LENGTH('a) < m mod 2 ^ LENGTH('a))" by transfer (simp add: take_bit_eq_mod) lemma wi_le: "(word_of_int n \ (word_of_int m :: 'a::len word)) = (n mod 2 ^ LENGTH('a) \ m mod 2 ^ LENGTH('a))" by transfer (simp add: take_bit_eq_mod) subsection \Bit-wise operations\ lemma uint_take_bit_eq: \uint (take_bit n w) = take_bit n (uint w)\ by transfer (simp add: ac_simps) lemma take_bit_word_eq_self: \take_bit n w = w\ if \LENGTH('a) \ n\ for w :: \'a::len word\ using that by transfer simp lemma take_bit_length_eq [simp]: \take_bit LENGTH('a) w = w\ for w :: \'a::len word\ by (rule take_bit_word_eq_self) simp lemma bit_word_of_int_iff: \bit (word_of_int k :: 'a::len word) n \ n < LENGTH('a) \ bit k n\ by transfer rule lemma bit_uint_iff: \bit (uint w) n \ n < LENGTH('a) \ bit w n\ for w :: \'a::len word\ by transfer (simp add: bit_take_bit_iff) lemma bit_sint_iff: \bit (sint w) n \ n \ LENGTH('a) \ bit w (LENGTH('a) - 1) \ bit w n\ for w :: \'a::len word\ by transfer (auto simp add: bit_signed_take_bit_iff min_def le_less not_less) lemma bit_word_ucast_iff: \bit (ucast w :: 'b::len word) n \ n < LENGTH('a) \ n < LENGTH('b) \ bit w n\ for w :: \'a::len word\ by transfer (simp add: bit_take_bit_iff ac_simps) lemma bit_word_scast_iff: \bit (scast w :: 'b::len word) n \ n < LENGTH('b) \ (bit w n \ LENGTH('a) \ n \ bit w (LENGTH('a) - Suc 0))\ for w :: \'a::len word\ by transfer (auto simp add: bit_signed_take_bit_iff le_less min_def) lift_definition shiftl1 :: \'a::len word \ 'a word\ is \(*) 2\ by (auto simp add: take_bit_eq_mod intro: mod_mult_cong) lemma shiftl1_eq: \shiftl1 w = word_of_int (2 * uint w)\ by transfer (simp add: take_bit_eq_mod mod_simps) lemma shiftl1_eq_mult_2: \shiftl1 = (*) 2\ by (rule ext, transfer) simp lemma bit_shiftl1_iff: \bit (shiftl1 w) n \ 0 < n \ n < LENGTH('a) \ bit w (n - 1)\ for w :: \'a::len word\ by (simp add: shiftl1_eq_mult_2 bit_double_iff exp_eq_zero_iff not_le) (simp add: ac_simps) lift_definition shiftr1 :: \'a::len word \ 'a word\ \ \shift right as unsigned or as signed, ie logical or arithmetic\ is \\k. take_bit LENGTH('a) k div 2\ by simp lemma shiftr1_eq_div_2: \shiftr1 w = w div 2\ by transfer simp lemma bit_shiftr1_iff: \bit (shiftr1 w) n \ bit w (Suc n)\ by transfer (auto simp flip: bit_Suc simp add: bit_take_bit_iff) lemma shiftr1_eq: \shiftr1 w = word_of_int (uint w div 2)\ by transfer simp instantiation word :: (len) semiring_bit_syntax begin lift_definition test_bit_word :: \'a::len word \ nat \ bool\ is \\k n. n < LENGTH('a) \ bit k n\ proof fix k l :: int and n :: nat assume *: \take_bit LENGTH('a) k = take_bit LENGTH('a) l\ show \n < LENGTH('a) \ bit k n \ n < LENGTH('a) \ bit l n\ proof (cases \n < LENGTH('a)\) case True from * have \bit (take_bit LENGTH('a) k) n \ bit (take_bit LENGTH('a) l) n\ by simp then show ?thesis by (simp add: bit_take_bit_iff) next case False then show ?thesis by simp qed qed lemma test_bit_word_eq: \test_bit = (bit :: 'a word \ _)\ by transfer simp lemma bit_word_iff_drop_bit_and [code]: \bit a n \ drop_bit n a AND 1 = 1\ for a :: \'a::len word\ by (simp add: bit_iff_odd_drop_bit odd_iff_mod_2_eq_one and_one_eq) lemma [code]: \test_bit a n \ drop_bit n a AND 1 = 1\ for a :: \'a::len word\ by (simp add: test_bit_word_eq bit_word_iff_drop_bit_and) lift_definition shiftl_word :: \'a::len word \ nat \ 'a word\ is \\k n. push_bit n k\ proof - show \take_bit LENGTH('a) (push_bit n k) = take_bit LENGTH('a) (push_bit n l)\ if \take_bit LENGTH('a) k = take_bit LENGTH('a) l\ for k l :: int and n :: nat proof - from that have \take_bit (LENGTH('a) - n) (take_bit LENGTH('a) k) = take_bit (LENGTH('a) - n) (take_bit LENGTH('a) l)\ by simp moreover have \min (LENGTH('a) - n) LENGTH('a) = LENGTH('a) - n\ by simp ultimately show ?thesis by (simp add: take_bit_push_bit) qed qed lemma shiftl_word_eq: \w << n = push_bit n w\ for w :: \'a::len word\ by transfer rule lift_definition shiftr_word :: \'a::len word \ nat \ 'a word\ is \\k n. drop_bit n (take_bit LENGTH('a) k)\ by simp lemma shiftr_word_eq: \w >> n = drop_bit n w\ for w :: \'a::len word\ by transfer simp instance by (standard; transfer) simp_all end lemma shiftl_code [code]: \w << n = w * 2 ^ n\ for w :: \'a::len word\ by transfer (simp add: push_bit_eq_mult) lemma shiftl1_code [code]: \shiftl1 w = w << 1\ by transfer (simp add: push_bit_eq_mult 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 [code]: \Word.the_int (w >> n) = drop_bit n (Word.the_int w)\ by transfer (simp add: drop_bit_take_bit min_def le_less less_diff_conv) lemma shiftr1_code [code]: \shiftr1 w = w >> 1\ by transfer (simp add: drop_bit_Suc) lemma word_test_bit_def: \test_bit a = bit (uint a)\ by transfer (simp add: fun_eq_iff bit_take_bit_iff) 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 \drop_bit n = (((\k::int. k div 2) ^^ n))\ for n by (rule sym, induction n) (simp_all add: fun_eq_iff drop_bit_Suc flip: drop_bit_half) then show ?thesis apply transfer apply simp apply (metis bintrunc_bintrunc rco_bintr) done 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 [code]: \push_bit n w = w << n\ for w :: \'a::len word\ by (simp add: shiftl_word_eq) lemma [code]: \drop_bit n w = w >> n\ for w :: \'a::len word\ by (simp add: shiftr_word_eq) 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) lemma word_not_def: "NOT (a::'a::len word) = word_of_int (NOT (uint a))" and word_and_def: "(a::'a word) AND b = word_of_int (uint a AND uint b)" and word_or_def: "(a::'a word) OR b = word_of_int (uint a OR uint b)" and word_xor_def: "(a::'a word) XOR b = word_of_int (uint a XOR uint b)" by (transfer, simp add: take_bit_not_take_bit)+ lift_definition setBit :: \'a::len word \ nat \ 'a word\ is \\k n. set_bit n k\ by (simp add: take_bit_set_bit_eq) lemma set_Bit_eq: \setBit w n = set_bit n w\ by transfer simp lemma bit_setBit_iff: \bit (setBit w m) n \ (m = n \ n < LENGTH('a) \ bit w n)\ for w :: \'a::len word\ by transfer (auto simp add: bit_set_bit_iff) lift_definition clearBit :: \'a::len word \ nat \ 'a word\ is \\k n. unset_bit n k\ by (simp add: take_bit_unset_bit_eq) lemma clear_Bit_eq: \clearBit w n = unset_bit n w\ by transfer simp lemma bit_clearBit_iff: \bit (clearBit w m) n \ m \ n \ bit w n\ for w :: \'a::len word\ by transfer (auto simp add: bit_unset_bit_iff) definition even_word :: \'a::len word \ bool\ where [code_abbrev]: \even_word = even\ lemma even_word_iff [code]: \even_word a \ a AND 1 = 0\ by (simp add: and_one_eq even_iff_mod_2_eq_zero even_word_def) lemma map_bit_range_eq_if_take_bit_eq: \map (bit k) [0.. if \take_bit n k = take_bit n l\ for k l :: int using that proof (induction n arbitrary: k l) case 0 then show ?case by simp next case (Suc n) from Suc.prems have \take_bit n (k div 2) = take_bit n (l div 2)\ by (simp add: take_bit_Suc) then have \map (bit (k div 2)) [0.. by (rule Suc.IH) moreover have \bit (r div 2) = bit r \ Suc\ for r :: int by (simp add: fun_eq_iff bit_Suc) moreover from Suc.prems have \even k \ even l\ by (auto simp add: take_bit_Suc elim!: evenE oddE) arith+ ultimately show ?case by (simp only: map_Suc_upt upt_conv_Cons flip: list.map_comp) simp qed lemma take_bit_word_Bit0_eq [simp]: \take_bit (numeral n) (numeral (num.Bit0 m) :: 'a::len word) = 2 * take_bit (pred_numeral n) (numeral m)\ (is ?P) and take_bit_word_Bit1_eq [simp]: \take_bit (numeral n) (numeral (num.Bit1 m) :: 'a::len word) = 1 + 2 * take_bit (pred_numeral n) (numeral m)\ (is ?Q) and take_bit_word_minus_Bit0_eq [simp]: \take_bit (numeral n) (- numeral (num.Bit0 m) :: 'a::len word) = 2 * take_bit (pred_numeral n) (- numeral m)\ (is ?R) and take_bit_word_minus_Bit1_eq [simp]: \take_bit (numeral n) (- numeral (num.Bit1 m) :: 'a::len word) = 1 + 2 * take_bit (pred_numeral n) (- numeral (Num.inc m))\ (is ?S) proof - define w :: \'a::len word\ where \w = numeral m\ moreover define q :: nat where \q = pred_numeral n\ ultimately have num: \numeral m = w\ \numeral (num.Bit0 m) = 2 * w\ \numeral (num.Bit1 m) = 1 + 2 * w\ \numeral (Num.inc m) = 1 + w\ \pred_numeral n = q\ \numeral n = Suc q\ by (simp_all only: w_def q_def numeral_Bit0 [of m] numeral_Bit1 [of m] ac_simps numeral_inc numeral_eq_Suc flip: mult_2) have even: \take_bit (Suc q) (2 * w) = 2 * take_bit q w\ for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_take_bit_iff bit_double_iff) have odd: \take_bit (Suc q) (1 + 2 * w) = 1 + 2 * take_bit q w\ for w :: \'a::len word\ by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_double_iff even_bit_succ_iff) show ?P using even [of w] by (simp add: num) show ?Q using odd [of w] by (simp add: num) show ?R using even [of \- w\] by (simp add: num) show ?S using odd [of \- (1 + w)\] by (simp add: num) qed subsection \More shift operations\ lift_definition signed_drop_bit :: \nat \ 'a word \ 'a::len word\ is \\n. drop_bit n \ signed_take_bit (LENGTH('a) - Suc 0)\ using signed_take_bit_decr_length_iff by (simp add: take_bit_drop_bit) force lemma bit_signed_drop_bit_iff: \bit (signed_drop_bit m w) 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_drop_bit_eq bit_signed_take_bit_iff not_le min_def) apply (metis add.commute le_antisym less_diff_conv less_eq_decr_length_iff) apply (metis le_antisym less_eq_decr_length_iff) done lemma signed_drop_bit_0 [simp]: \signed_drop_bit 0 w = w\ by transfer simp lemma sint_signed_drop_bit_eq: \sint (signed_drop_bit n w) = drop_bit n (sint w)\ apply (cases \LENGTH('a)\; cases n) apply simp_all apply (rule bit_eqI) apply (auto simp add: bit_sint_iff bit_drop_bit_eq bit_signed_drop_bit_iff dest: bit_imp_le_length) done lift_definition sshiftr1 :: \'a::len word \ 'a word\ is \\k. take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - 1) k div 2)\ by (simp flip: signed_take_bit_decr_length_iff) 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) - 1) k))\ by (simp flip: signed_take_bit_decr_length_iff) lift_definition bshiftr1 :: \bool \ 'a::len word \ 'a word\ is \\b k. take_bit LENGTH('a) k div 2 + of_bool b * 2 ^ (LENGTH('a) - Suc 0)\ by (fact arg_cong) lemma sshiftr1_eq: \sshiftr1 w = word_of_int (sint w div 2)\ by transfer simp lemma sshiftr_eq_funpow_sshiftr1: \w >>> n = (sshiftr1 ^^ n) w\ proof - have *: \(\k::int. take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - Suc 0) k div 2)) ^^ Suc n = take_bit LENGTH('a) \ drop_bit (Suc n) \ signed_take_bit (LENGTH('a) - Suc 0)\ for n apply (induction n) apply (auto simp add: fun_eq_iff drop_bit_Suc) apply (metis (no_types, lifting) Suc_pred funpow_swap1 len_gt_0 sbintrunc_bintrunc sbintrunc_rest) done show ?thesis apply transfer apply simp subgoal for k n apply (cases n) apply (simp_all only: *) apply simp_all done done qed 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 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 [code]: \Word.the_int (w >>> n) = take_bit LENGTH('a) (drop_bit n (Word.the_signed_int w))\ for w :: \'a::len word\ by transfer simp lemma sshift1_code [code]: \sshiftr1 w = w >>> 1\ by transfer (simp add: drop_bit_Suc) subsection \Rotation\ lift_definition word_rotr :: \nat \ 'a::len word \ 'a::len word\ is \\n k. concat_bit (LENGTH('a) - n mod LENGTH('a)) (drop_bit (n mod LENGTH('a)) (take_bit LENGTH('a) k)) (take_bit (n mod LENGTH('a)) k)\ subgoal for n k l apply (simp add: concat_bit_def nat_le_iff less_imp_le take_bit_tightened [of \LENGTH('a)\ k l \n mod LENGTH('a::len)\]) done done lift_definition word_rotl :: \nat \ 'a::len word \ 'a::len word\ is \\n k. concat_bit (n mod LENGTH('a)) (drop_bit (LENGTH('a) - n mod LENGTH('a)) (take_bit LENGTH('a) k)) (take_bit (LENGTH('a) - n mod LENGTH('a)) k)\ subgoal for n k l apply (simp add: concat_bit_def nat_le_iff less_imp_le take_bit_tightened [of \LENGTH('a)\ k l \LENGTH('a) - n mod LENGTH('a::len)\]) done done lift_definition word_roti :: \int \ 'a::len word \ 'a::len word\ is \\r k. concat_bit (LENGTH('a) - nat (r mod int LENGTH('a))) (drop_bit (nat (r mod int LENGTH('a))) (take_bit LENGTH('a) k)) (take_bit (nat (r mod int LENGTH('a))) k)\ subgoal for r k l apply (simp add: concat_bit_def nat_le_iff less_imp_le take_bit_tightened [of \LENGTH('a)\ k l \nat (r mod int LENGTH('a::len))\]) done done lemma word_rotl_eq_word_rotr [code]: \word_rotl n = (word_rotr (LENGTH('a) - n mod LENGTH('a)) :: 'a::len word \ 'a word)\ by (rule ext, cases \n mod LENGTH('a) = 0\; transfer) simp_all lemma word_roti_eq_word_rotr_word_rotl [code]: \word_roti i w = (if i \ 0 then word_rotr (nat i) w else word_rotl (nat (- i)) w)\ proof (cases \i \ 0\) case True moreover define n where \n = nat i\ ultimately have \i = int n\ by simp moreover have \word_roti (int n) = (word_rotr n :: _ \ 'a word)\ by (rule ext, transfer) (simp add: nat_mod_distrib) ultimately show ?thesis by simp next case False moreover define n where \n = nat (- i)\ ultimately have \i = - int n\ \n > 0\ by simp_all moreover have \word_roti (- int n) = (word_rotl n :: _ \ 'a word)\ by (rule ext, transfer) (simp add: zmod_zminus1_eq_if flip: of_nat_mod of_nat_diff) ultimately show ?thesis by simp qed lemma bit_word_rotr_iff: \bit (word_rotr m w) n \ n < LENGTH('a) \ bit w ((n + m) mod LENGTH('a))\ for w :: \'a::len word\ proof transfer fix k :: int and m n :: nat define q where \q = m mod LENGTH('a)\ have \q < LENGTH('a)\ by (simp add: q_def) then have \q \ LENGTH('a)\ by simp have \m mod LENGTH('a) = q\ by (simp add: q_def) moreover have \(n + m) mod LENGTH('a) = (n + q) mod LENGTH('a)\ by (subst mod_add_right_eq [symmetric]) (simp add: \m mod LENGTH('a) = q\) moreover have \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - q) (drop_bit q (take_bit LENGTH('a) k)) (take_bit q k)) n \ n < LENGTH('a) \ bit k ((n + q) mod LENGTH('a))\ using \q < LENGTH('a)\ by (cases \q + n \ LENGTH('a)\) (auto simp add: bit_concat_bit_iff bit_drop_bit_eq bit_take_bit_iff le_mod_geq ac_simps) ultimately show \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - m mod LENGTH('a)) (drop_bit (m mod LENGTH('a)) (take_bit LENGTH('a) k)) (take_bit (m mod LENGTH('a)) k)) n \ n < LENGTH('a) \ (n + m) mod LENGTH('a) < LENGTH('a) \ bit k ((n + m) mod LENGTH('a))\ by simp qed lemma bit_word_rotl_iff: \bit (word_rotl m w) n \ n < LENGTH('a) \ bit w ((n + (LENGTH('a) - m mod LENGTH('a))) mod LENGTH('a))\ for w :: \'a::len word\ by (simp add: word_rotl_eq_word_rotr bit_word_rotr_iff) lemma bit_word_roti_iff: \bit (word_roti k w) n \ n < LENGTH('a) \ bit w (nat ((int n + k) mod int LENGTH('a)))\ for w :: \'a::len word\ proof transfer fix k l :: int and n :: nat define m where \m = nat (k mod int LENGTH('a))\ have \m < LENGTH('a)\ by (simp add: nat_less_iff m_def) then have \m \ LENGTH('a)\ by simp have \k mod int LENGTH('a) = int m\ by (simp add: nat_less_iff m_def) moreover have \(int n + k) mod int LENGTH('a) = int ((n + m) mod LENGTH('a))\ by (subst mod_add_right_eq [symmetric]) (simp add: of_nat_mod \k mod int LENGTH('a) = int m\) moreover have \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - m) (drop_bit m (take_bit LENGTH('a) l)) (take_bit m l)) n \ n < LENGTH('a) \ bit l ((n + m) mod LENGTH('a))\ using \m < LENGTH('a)\ by (cases \m + n \ LENGTH('a)\) (auto simp add: bit_concat_bit_iff bit_drop_bit_eq bit_take_bit_iff nat_less_iff not_le not_less ac_simps le_diff_conv le_mod_geq) ultimately show \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - nat (k mod int LENGTH('a))) (drop_bit (nat (k mod int LENGTH('a))) (take_bit LENGTH('a) l)) (take_bit (nat (k mod int LENGTH('a))) l)) n \ n < LENGTH('a) \ nat ((int n + k) mod int LENGTH('a)) < LENGTH('a) \ bit l (nat ((int n + k) mod int LENGTH('a)))\ by simp qed lemma uint_word_rotr_eq: \uint (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a)) (drop_bit (n mod LENGTH('a)) (uint w)) (uint (take_bit (n mod LENGTH('a)) w))\ for w :: \'a::len word\ apply transfer apply (simp add: concat_bit_def take_bit_drop_bit push_bit_take_bit min_def) using mod_less_divisor not_less apply blast done lemma [code]: \Word.the_int (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a)) (drop_bit (n mod LENGTH('a)) (Word.the_int w)) (Word.the_int (take_bit (n mod LENGTH('a)) w))\ for w :: \'a::len word\ using uint_word_rotr_eq [of n w] by simp subsection \Split and cat operations\ lift_definition word_cat :: \'a::len word \ 'b::len word \ 'c::len word\ is \\k l. concat_bit LENGTH('b) l (take_bit LENGTH('a) k)\ by (simp add: bit_eq_iff bit_concat_bit_iff bit_take_bit_iff) lemma word_cat_eq: \(word_cat v w :: 'c::len word) = push_bit LENGTH('b) (ucast v) + ucast w\ for v :: \'a::len word\ and w :: \'b::len word\ by transfer (simp add: concat_bit_eq ac_simps) lemma word_cat_eq' [code]: \word_cat a b = word_of_int (concat_bit LENGTH('b) (uint b) (uint a))\ for a :: \'a::len word\ and b :: \'b::len word\ by transfer simp lemma bit_word_cat_iff: \bit (word_cat v w :: 'c::len word) n \ n < LENGTH('c) \ (if n < LENGTH('b) then bit w n else bit v (n - LENGTH('b)))\ for v :: \'a::len word\ and w :: \'b::len word\ by transfer (simp add: bit_concat_bit_iff bit_take_bit_iff) definition word_split :: "'a::len word \ 'b::len word \ 'c::len word" where "word_split a = (case bin_split (LENGTH('c)) (uint a) of (u, v) \ (word_of_int u, word_of_int v))" definition word_rcat :: \'a::len word list \ 'b::len word\ where \word_rcat = word_of_int \ horner_sum uint (2 ^ LENGTH('a)) \ rev\ 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 -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))" - abbreviation (input) max_word :: \'a::len word\ \ \Largest representable machine integer.\ where "max_word \ - 1" subsection \More on conversions\ lemma int_word_sint: \sint (word_of_int x :: 'a::len word) = (x + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)\ by transfer (simp add: no_sbintr_alt2) lemma sint_sbintrunc': "sint (word_of_int bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) bin" by simp lemma uint_sint: "uint w = take_bit (LENGTH('a)) (sint w)" for w :: "'a::len word" by transfer simp lemma bintr_uint: "LENGTH('a) \ n \ take_bit n (uint w) = uint w" for w :: "'a::len word" by transfer (simp add: min_def) lemma wi_bintr: "LENGTH('a::len) \ n \ word_of_int (take_bit n w) = (word_of_int w :: 'a word)" by transfer simp lemma word_numeral_alt: "numeral b = word_of_int (numeral b)" by (induct b, simp_all only: numeral.simps word_of_int_homs) declare word_numeral_alt [symmetric, code_abbrev] lemma word_neg_numeral_alt: "- numeral b = word_of_int (- numeral b)" by (simp only: word_numeral_alt wi_hom_neg) declare word_neg_numeral_alt [symmetric, code_abbrev] lemma uint_bintrunc [simp]: "uint (numeral bin :: 'a word) = take_bit (LENGTH('a::len)) (numeral bin)" by transfer rule lemma uint_bintrunc_neg [simp]: "uint (- numeral bin :: 'a word) = take_bit (LENGTH('a::len)) (- numeral bin)" by transfer rule lemma sint_sbintrunc [simp]: "sint (numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (numeral bin)" by transfer simp lemma sint_sbintrunc_neg [simp]: "sint (- numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (- numeral bin)" by transfer simp lemma unat_bintrunc [simp]: "unat (numeral bin :: 'a::len word) = nat (take_bit (LENGTH('a)) (numeral bin))" by transfer simp lemma unat_bintrunc_neg [simp]: "unat (- numeral bin :: 'a::len word) = nat (take_bit (LENGTH('a)) (- numeral bin))" by transfer simp lemma size_0_eq: "size w = 0 \ v = w" for v w :: "'a::len word" by transfer simp lemma uint_ge_0 [iff]: "0 \ uint x" by (fact unsigned_greater_eq) lemma uint_lt2p [iff]: "uint x < 2 ^ LENGTH('a)" for x :: "'a::len word" by (fact unsigned_less) lemma sint_ge: "- (2 ^ (LENGTH('a) - 1)) \ sint x" for x :: "'a::len word" using sint_greater_eq [of x] by simp lemma sint_lt: "sint x < 2 ^ (LENGTH('a) - 1)" for x :: "'a::len word" using sint_less [of x] by simp lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" by (simp add: sign_Pls_ge_0) lemma uint_m2p_neg: "uint x - 2 ^ LENGTH('a) < 0" for x :: "'a::len word" by (simp only: diff_less_0_iff_less uint_lt2p) lemma uint_m2p_not_non_neg: "\ 0 \ uint x - 2 ^ LENGTH('a)" for x :: "'a::len word" by (simp only: not_le uint_m2p_neg) lemma lt2p_lem: "LENGTH('a) \ n \ uint w < 2 ^ n" for w :: "'a::len word" by (metis bintr_lt2p bintr_uint) lemma uint_le_0_iff [simp]: "uint x \ 0 \ uint x = 0" by (fact uint_ge_0 [THEN leD, THEN antisym_conv1]) lemma uint_nat: "uint w = int (unat w)" by transfer simp lemma uint_numeral: "uint (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)" by (simp flip: take_bit_eq_mod add: of_nat_take_bit) lemma uint_neg_numeral: "uint (- numeral b :: 'a::len word) = - numeral b mod 2 ^ LENGTH('a)" by (simp flip: take_bit_eq_mod add: of_nat_take_bit) lemma unat_numeral: "unat (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)" by transfer (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq) lemma sint_numeral: "sint (numeral b :: 'a::len word) = (numeral b + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)" apply (transfer fixing: b) using int_word_sint [of \numeral b\] apply simp done lemma word_of_int_0 [simp, code_post]: "word_of_int 0 = 0" by (fact of_int_0) lemma word_of_int_1 [simp, code_post]: "word_of_int 1 = 1" by (fact of_int_1) lemma word_of_int_neg_1 [simp]: "word_of_int (- 1) = - 1" by (simp add: wi_hom_syms) lemma word_of_int_numeral [simp] : "(word_of_int (numeral bin) :: 'a::len word) = numeral bin" by (fact of_int_numeral) lemma word_of_int_neg_numeral [simp]: "(word_of_int (- numeral bin) :: 'a::len word) = - numeral bin" by (fact of_int_neg_numeral) lemma word_int_case_wi: "word_int_case f (word_of_int i :: 'b word) = f (i mod 2 ^ LENGTH('b::len))" by transfer (simp add: take_bit_eq_mod) lemma word_int_split: "P (word_int_case f x) = (\i. x = (word_of_int i :: 'b::len word) \ 0 \ i \ i < 2 ^ LENGTH('b) \ P (f i))" by transfer (auto simp add: take_bit_eq_mod) lemma word_int_split_asm: "P (word_int_case f x) = (\n. x = (word_of_int n :: 'b::len word) \ 0 \ n \ n < 2 ^ LENGTH('b::len) \ \ P (f n))" by transfer (auto simp add: take_bit_eq_mod) lemma uint_range_size: "0 \ uint w \ uint w < 2 ^ size w" by transfer simp lemma sint_range_size: "- (2 ^ (size w - Suc 0)) \ sint w \ sint w < 2 ^ (size w - Suc 0)" by transfer (use sbintr_ge sbintr_lt in blast) lemma sint_above_size: "2 ^ (size w - 1) \ x \ sint w < x" for w :: "'a::len word" unfolding word_size by (rule less_le_trans [OF sint_lt]) lemma sint_below_size: "x \ - (2 ^ (size w - 1)) \ x \ sint w" for w :: "'a::len word" unfolding word_size by (rule order_trans [OF _ sint_ge]) subsection \Testing bits\ 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 \ bin_nth (uint w) n" by transfer (simp add: bit_take_bit_iff) lemmas test_bit_bin = test_bit_bin' [unfolded word_size] lemma bin_nth_uint_imp: "bin_nth (uint w) n \ n < LENGTH('a)" for w :: "'a::len word" by transfer (simp add: bit_take_bit_iff) lemma bin_nth_sint: "LENGTH('a) \ n \ bin_nth (sint w) n = bin_nth (sint w) (LENGTH('a) - 1)" for w :: "'a::len word" by (transfer fixing: n) (simp add: bit_signed_take_bit_iff le_diff_conv min_def) lemma num_of_bintr': "take_bit (LENGTH('a::len)) (numeral a :: int) = (numeral b) \ numeral a = (numeral b :: 'a word)" proof (transfer fixing: a b) assume \take_bit LENGTH('a) (numeral a :: int) = numeral b\ then have \take_bit LENGTH('a) (take_bit LENGTH('a) (numeral a :: int)) = take_bit LENGTH('a) (numeral b)\ by simp then show \take_bit LENGTH('a) (numeral a :: int) = take_bit LENGTH('a) (numeral b)\ by simp qed lemma num_of_sbintr': "signed_take_bit (LENGTH('a::len) - 1) (numeral a :: int) = (numeral b) \ numeral a = (numeral b :: 'a word)" proof (transfer fixing: a b) assume \signed_take_bit (LENGTH('a) - 1) (numeral a :: int) = numeral b\ then have \take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - 1) (numeral a :: int)) = take_bit LENGTH('a) (numeral b)\ by simp then show \take_bit LENGTH('a) (numeral a :: int) = take_bit LENGTH('a) (numeral b)\ by simp qed lemma num_abs_bintr: "(numeral x :: 'a word) = word_of_int (take_bit (LENGTH('a::len)) (numeral x))" by transfer simp lemma num_abs_sbintr: "(numeral x :: 'a word) = word_of_int (signed_take_bit (LENGTH('a::len) - 1) (numeral x))" by transfer simp text \ \cast\ -- note, no arg for new length, as it's determined by type of result, thus in \cast w = w\, the type means cast to length of \w\! \ lemma bit_ucast_iff: \bit (ucast a :: 'a::len word) n \ n < LENGTH('a::len) \ Parity.bit a n\ by transfer (simp add: bit_take_bit_iff) lemma ucast_id [simp]: "ucast w = w" by transfer simp lemma scast_id [simp]: "scast w = w" by transfer simp 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 ucast_mask_eq: \ucast (mask n :: 'b word) = mask (min LENGTH('b::len) n)\ by (simp add: bit_eq_iff) (auto simp add: bit_mask_iff bit_ucast_iff exp_eq_zero_iff) \ \literal u(s)cast\ lemma ucast_bintr [simp]: "ucast (numeral w :: 'a::len word) = word_of_int (take_bit (LENGTH('a)) (numeral w))" by transfer simp (* TODO: neg_numeral *) lemma scast_sbintr [simp]: "scast (numeral w ::'a::len word) = word_of_int (signed_take_bit (LENGTH('a) - Suc 0) (numeral w))" by transfer simp lemma source_size: "source_size (c::'a::len word \ _) = LENGTH('a)" by transfer simp lemma target_size: "target_size (c::_ \ 'b::len word) = LENGTH('b)" by transfer simp lemma is_down: "is_down c \ LENGTH('b) \ LENGTH('a)" for c :: "'a::len word \ 'b::len word" by transfer simp lemma is_up: "is_up c \ LENGTH('a) \ LENGTH('b)" for c :: "'a::len word \ 'b::len word" by transfer simp lemma is_up_down: \is_up c \ is_down d\ for c :: \'a::len word \ 'b::len word\ and d :: \'b::len word \ 'a::len word\ by transfer simp context fixes dummy_types :: \'a::len \ 'b::len\ begin private abbreviation (input) UCAST :: \'a::len word \ 'b::len word\ where \UCAST == ucast\ private abbreviation (input) SCAST :: \'a::len word \ 'b::len word\ where \SCAST == scast\ lemma down_cast_same: \UCAST = scast\ if \is_down UCAST\ by (rule ext, use that in transfer) (simp add: take_bit_signed_take_bit) lemma sint_up_scast: \sint (SCAST w) = sint w\ if \is_up SCAST\ using that by transfer (simp add: min_def Suc_leI le_diff_iff) lemma uint_up_ucast: \uint (UCAST w) = uint w\ if \is_up UCAST\ using that by transfer (simp add: min_def) lemma ucast_up_ucast: \ucast (UCAST w) = ucast w\ if \is_up UCAST\ using that by transfer (simp add: ac_simps) lemma ucast_up_ucast_id: \ucast (UCAST w) = w\ if \is_up UCAST\ using that by (simp add: ucast_up_ucast) lemma scast_up_scast: \scast (SCAST w) = scast w\ if \is_up SCAST\ using that by transfer (simp add: ac_simps) lemma scast_up_scast_id: \scast (SCAST w) = w\ if \is_up SCAST\ using that by (simp add: scast_up_scast) lemma isduu: \is_up UCAST\ if \is_down d\ for d :: \'b word \ 'a word\ using that is_up_down [of UCAST d] by simp lemma isdus: \is_up SCAST\ if \is_down d\ for d :: \'b word \ 'a word\ using that is_up_down [of SCAST d] by simp lemmas ucast_down_ucast_id = isduu [THEN ucast_up_ucast_id] lemmas scast_down_scast_id = isdus [THEN scast_up_scast_id] lemma up_ucast_surj: \surj (ucast :: 'b word \ 'a word)\ if \is_up UCAST\ by (rule surjI) (use that in \rule ucast_up_ucast_id\) lemma up_scast_surj: \surj (scast :: 'b word \ 'a word)\ if \is_up SCAST\ by (rule surjI) (use that in \rule scast_up_scast_id\) lemma down_ucast_inj: \inj_on UCAST A\ if \is_down (ucast :: 'b word \ 'a word)\ by (rule inj_on_inverseI) (use that in \rule ucast_down_ucast_id\) lemma down_scast_inj: \inj_on SCAST A\ if \is_down (scast :: 'b word \ 'a word)\ by (rule inj_on_inverseI) (use that in \rule scast_down_scast_id\) lemma ucast_down_wi: \UCAST (word_of_int x) = word_of_int x\ if \is_down UCAST\ using that by transfer simp lemma ucast_down_no: \UCAST (numeral bin) = numeral bin\ if \is_down UCAST\ using that by transfer simp end lemmas test_bit_def' = word_test_bit_def [THEN fun_cong] lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def lemma bit_last_iff: \bit w (LENGTH('a) - Suc 0) \ sint w < 0\ (is \?P \ ?Q\) for w :: \'a::len word\ proof - have \?P \ bit (uint w) (LENGTH('a) - Suc 0)\ by (simp add: bit_uint_iff) also have \\ \ ?Q\ by (simp add: sint_uint) finally show ?thesis . qed lemma drop_bit_eq_zero_iff_not_bit_last: \drop_bit (LENGTH('a) - Suc 0) w = 0 \ \ bit w (LENGTH('a) - Suc 0)\ for w :: "'a::len word" apply (cases \LENGTH('a)\) apply simp_all apply (simp add: bit_iff_odd_drop_bit) apply transfer apply (simp add: take_bit_drop_bit) apply (auto simp add: drop_bit_eq_div take_bit_eq_mod min_def) apply (auto elim!: evenE) apply (metis div_exp_eq mod_div_trivial mult.commute nonzero_mult_div_cancel_left power_Suc0_right power_add zero_neq_numeral) done subsection \Word Arithmetic\ lemmas word_div_no [simp] = word_div_def [of "numeral a" "numeral b"] for a b lemmas word_mod_no [simp] = word_mod_def [of "numeral a" "numeral b"] for a b lemmas word_less_no [simp] = word_less_def [of "numeral a" "numeral b"] for a b lemmas word_le_no [simp] = word_le_def [of "numeral a" "numeral b"] for a b lemmas word_sless_no [simp] = word_sless_eq [of "numeral a" "numeral b"] for a b lemmas word_sle_no [simp] = word_sle_eq [of "numeral a" "numeral b"] for a b lemma size_0_same': "size w = 0 \ w = v" for v w :: "'a::len word" by (unfold word_size) simp lemmas size_0_same = size_0_same' [unfolded word_size] lemmas unat_eq_0 = unat_0_iff lemmas unat_eq_zero = unat_0_iff subsection \Transferring goals from words to ints\ lemma word_ths: shows word_succ_p1: "word_succ a = a + 1" and word_pred_m1: "word_pred a = a - 1" and word_pred_succ: "word_pred (word_succ a) = a" and word_succ_pred: "word_succ (word_pred a) = a" and word_mult_succ: "word_succ a * b = b + a * b" by (transfer, simp add: algebra_simps)+ lemma uint_cong: "x = y \ uint x = uint y" by simp lemma uint_word_ariths: fixes a b :: "'a::len word" shows "uint (a + b) = (uint a + uint b) mod 2 ^ LENGTH('a::len)" and "uint (a - b) = (uint a - uint b) mod 2 ^ LENGTH('a)" and "uint (a * b) = uint a * uint b mod 2 ^ LENGTH('a)" and "uint (- a) = - uint a mod 2 ^ LENGTH('a)" and "uint (word_succ a) = (uint a + 1) mod 2 ^ LENGTH('a)" and "uint (word_pred a) = (uint a - 1) mod 2 ^ LENGTH('a)" and "uint (0 :: 'a word) = 0 mod 2 ^ LENGTH('a)" and "uint (1 :: 'a word) = 1 mod 2 ^ LENGTH('a)" by (simp_all only: word_arith_wis uint_word_of_int_eq flip: take_bit_eq_mod) lemma uint_word_arith_bintrs: fixes a b :: "'a::len word" shows "uint (a + b) = take_bit (LENGTH('a)) (uint a + uint b)" and "uint (a - b) = take_bit (LENGTH('a)) (uint a - uint b)" and "uint (a * b) = take_bit (LENGTH('a)) (uint a * uint b)" and "uint (- a) = take_bit (LENGTH('a)) (- uint a)" and "uint (word_succ a) = take_bit (LENGTH('a)) (uint a + 1)" and "uint (word_pred a) = take_bit (LENGTH('a)) (uint a - 1)" and "uint (0 :: 'a word) = take_bit (LENGTH('a)) 0" and "uint (1 :: 'a word) = take_bit (LENGTH('a)) 1" by (simp_all add: uint_word_ariths take_bit_eq_mod) lemma sint_word_ariths: fixes a b :: "'a::len word" shows "sint (a + b) = signed_take_bit (LENGTH('a) - 1) (sint a + sint b)" and "sint (a - b) = signed_take_bit (LENGTH('a) - 1) (sint a - sint b)" and "sint (a * b) = signed_take_bit (LENGTH('a) - 1) (sint a * sint b)" and "sint (- a) = signed_take_bit (LENGTH('a) - 1) (- sint a)" and "sint (word_succ a) = signed_take_bit (LENGTH('a) - 1) (sint a + 1)" and "sint (word_pred a) = signed_take_bit (LENGTH('a) - 1) (sint a - 1)" and "sint (0 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 0" and "sint (1 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 1" prefer 8 apply (simp add: Suc_lessI sbintrunc_minus_simps(3)) prefer 7 apply simp apply transfer apply (simp add: signed_take_bit_add) apply transfer apply (simp add: signed_take_bit_diff) apply transfer apply (simp add: signed_take_bit_mult) apply transfer apply (simp add: signed_take_bit_minus) apply (metis One_nat_def id_apply of_int_eq_id sbintrunc_sbintrunc signed.rep_eq signed_word_eqI sint_sbintrunc' wi_hom_succ) apply (metis (no_types, lifting) One_nat_def signed_take_bit_decr_length_iff sint_uint uint_sint uint_word_of_int_eq wi_hom_pred word_of_int_uint) done lemma word_pred_0_n1: "word_pred 0 = word_of_int (- 1)" unfolding word_pred_m1 by simp lemma succ_pred_no [simp]: "word_succ (numeral w) = numeral w + 1" "word_pred (numeral w) = numeral w - 1" "word_succ (- numeral w) = - numeral w + 1" "word_pred (- numeral w) = - numeral w - 1" by (simp_all add: word_succ_p1 word_pred_m1) lemma word_sp_01 [simp]: "word_succ (- 1) = 0 \ word_succ 0 = 1 \ word_pred 0 = - 1 \ word_pred 1 = 0" by (simp_all add: word_succ_p1 word_pred_m1) \ \alternative approach to lifting arithmetic equalities\ lemma word_of_int_Ex: "\y. x = word_of_int y" by (rule_tac x="uint x" in exI) simp subsection \Order on fixed-length words\ lift_definition udvd :: \'a::len word \ 'a::len word \ bool\ (infixl \udvd\ 50) is \\k l. take_bit LENGTH('a) k dvd take_bit LENGTH('a) l\ by simp lemma udvd_iff_dvd: \x udvd y \ unat x dvd unat y\ by transfer (simp add: nat_dvd_iff) lemma udvd_iff_dvd_int: \v udvd w \ uint v dvd uint w\ by transfer rule lemma udvdI [intro]: \v udvd w\ if \unat w = unat v * unat u\ proof - from that have \unat v dvd unat w\ .. then show ?thesis by (simp add: udvd_iff_dvd) qed lemma udvdE [elim]: fixes v w :: \'a::len word\ assumes \v udvd w\ obtains u :: \'a word\ where \unat w = unat v * unat u\ proof (cases \v = 0\) case True moreover from True \v udvd w\ have \w = 0\ by transfer simp ultimately show thesis using that by simp next case False then have \unat v > 0\ by (simp add: unat_gt_0) from \v udvd w\ have \unat v dvd unat w\ by (simp add: udvd_iff_dvd) then obtain n where \unat w = unat v * n\ .. moreover have \n < 2 ^ LENGTH('a)\ proof (rule ccontr) assume \\ n < 2 ^ LENGTH('a)\ then have \n \ 2 ^ LENGTH('a)\ by (simp add: not_le) then have \unat v * n \ 2 ^ LENGTH('a)\ using \unat v > 0\ mult_le_mono [of 1 \unat v\ \2 ^ LENGTH('a)\ n] by simp with \unat w = unat v * n\ have \unat w \ 2 ^ LENGTH('a)\ by simp with unsigned_less [of w, where ?'a = nat] show False by linarith qed ultimately have \unat w = unat v * unat (word_of_nat n :: 'a word)\ by (auto simp add: take_bit_nat_eq_self_iff intro: sym) with that show thesis . qed lemma udvd_imp_mod_eq_0: \w mod v = 0\ if \v udvd w\ using that by transfer simp lemma mod_eq_0_imp_udvd [intro?]: \v udvd w\ if \w mod v = 0\ proof - from that have \unat (w mod v) = unat 0\ by simp then have \unat w mod unat v = 0\ by (simp add: unat_mod_distrib) then have \unat v dvd unat w\ .. then show ?thesis by (simp add: udvd_iff_dvd) qed lemma udvd_imp_dvd: \v dvd w\ if \v udvd w\ for v w :: \'a::len word\ proof - from that obtain u :: \'a word\ where \unat w = unat v * unat u\ .. then have \(word_of_nat (unat w) :: 'a word) = word_of_nat (unat v * unat u)\ by simp then have \w = v * u\ by simp then show \v dvd w\ .. qed lemma exp_dvd_iff_exp_udvd: \2 ^ n dvd w \ 2 ^ n udvd w\ for v w :: \'a::len word\ proof assume \2 ^ n udvd w\ then show \2 ^ n dvd w\ by (rule udvd_imp_dvd) next assume \2 ^ n dvd w\ then obtain u :: \'a word\ where \w = 2 ^ n * u\ .. then have \w = push_bit n u\ by (simp add: push_bit_eq_mult) then show \2 ^ n udvd w\ by transfer (simp add: take_bit_push_bit dvd_eq_mod_eq_0 flip: take_bit_eq_mod) qed lemma udvd_nat_alt: \a udvd b \ (\n. unat b = n * unat a)\ by (auto simp add: udvd_iff_dvd) lemma udvd_unfold_int: \a udvd b \ (\n\0. uint b = n * uint a)\ apply (auto elim!: dvdE simp add: udvd_iff_dvd) apply (simp only: uint_nat) apply auto using of_nat_0_le_iff apply blast apply (simp only: unat_eq_nat_uint) apply (simp add: nat_mult_distrib) done lemma unat_minus_one: \unat (w - 1) = unat w - 1\ if \w \ 0\ proof - have "0 \ uint w" by (fact uint_nonnegative) moreover from that have "0 \ uint w" by (simp add: uint_0_iff) ultimately have "1 \ uint w" by arith from uint_lt2p [of w] have "uint w - 1 < 2 ^ LENGTH('a)" by arith with \1 \ uint w\ have "(uint w - 1) mod 2 ^ LENGTH('a) = uint w - 1" by (auto intro: mod_pos_pos_trivial) with \1 \ uint w\ have "nat ((uint w - 1) mod 2 ^ LENGTH('a)) = nat (uint w) - 1" by (auto simp del: nat_uint_eq) then show ?thesis by (simp only: unat_eq_nat_uint word_arith_wis mod_diff_right_eq) (metis of_int_1 uint_word_of_int unsigned_1) qed lemma measure_unat: "p \ 0 \ unat (p - 1) < unat p" by (simp add: unat_minus_one) (simp add: unat_0_iff [symmetric]) lemmas uint_add_ge0 [simp] = add_nonneg_nonneg [OF uint_ge_0 uint_ge_0] lemmas uint_mult_ge0 [simp] = mult_nonneg_nonneg [OF uint_ge_0 uint_ge_0] lemma uint_sub_lt2p [simp]: "uint x - uint y < 2 ^ LENGTH('a)" for x :: "'a::len word" and y :: "'b::len word" using uint_ge_0 [of y] uint_lt2p [of x] by arith subsection \Conditions for the addition (etc) of two words to overflow\ lemma uint_add_lem: "(uint x + uint y < 2 ^ LENGTH('a)) = (uint (x + y) = uint x + uint y)" for x y :: "'a::len word" by (metis add.right_neutral add_mono_thms_linordered_semiring(1) mod_pos_pos_trivial of_nat_0_le_iff uint_lt2p uint_nat uint_word_ariths(1)) lemma uint_mult_lem: "(uint x * uint y < 2 ^ LENGTH('a)) = (uint (x * y) = uint x * uint y)" for x y :: "'a::len word" by (metis mod_pos_pos_trivial uint_lt2p uint_mult_ge0 uint_word_ariths(3)) lemma uint_sub_lem: "uint x \ uint y \ uint (x - y) = uint x - uint y" by (metis diff_ge_0_iff_ge of_nat_0_le_iff uint_nat uint_sub_lt2p uint_word_of_int unique_euclidean_semiring_numeral_class.mod_less word_sub_wi) lemma uint_add_le: "uint (x + y) \ uint x + uint y" unfolding uint_word_ariths by (simp add: zmod_le_nonneg_dividend) lemma uint_sub_ge: "uint (x - y) \ uint x - uint y" unfolding uint_word_ariths by (simp add: int_mod_ge) 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 apply (auto simp add: not_less) apply (rule antisym) apply (metis diff_ge_0_iff_ge minus_mod_self2 zmod_le_nonneg_dividend) apply (simp only: flip: minus_mod_self2 [of \x + y\ z]) apply (rule int_mod_ge) apply simp_all done lemma uint_plus_if': "uint (a + b) = (if uint a + uint b < 2 ^ LENGTH('a) then uint a + uint b else uint a + uint b - 2 ^ LENGTH('a))" for a b :: "'a::len word" using mod_add_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths) 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 apply (auto simp add: not_le) apply (rule antisym) apply (simp only: flip: mod_add_self2 [of \x - y\ z]) apply (rule zmod_le_nonneg_dividend) apply simp apply (metis add.commute add.right_neutral add_le_cancel_left diff_ge_0_iff_ge int_mod_ge le_less le_less_trans mod_add_self1 not_less) done lemma uint_sub_if': "uint (a - b) = (if uint b \ uint a then uint a - uint b else uint a - uint b + 2 ^ LENGTH('a))" for a b :: "'a::len word" using mod_sub_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths) subsection \Definition of \uint_arith\\ lemma word_of_int_inverse: "word_of_int r = a \ 0 \ r \ r < 2 ^ LENGTH('a) \ uint a = r" for a :: "'a::len word" apply transfer apply (drule sym) apply (simp add: take_bit_int_eq_self) done lemma uint_split: "P (uint x) = (\i. word_of_int i = x \ 0 \ i \ i < 2^LENGTH('a) \ P i)" for x :: "'a::len word" by transfer (auto simp add: take_bit_eq_mod) lemma uint_split_asm: "P (uint x) = (\i. word_of_int i = x \ 0 \ i \ i < 2^LENGTH('a) \ \ P i)" for x :: "'a::len word" by auto (metis take_bit_int_eq_self_iff) lemmas uint_splits = uint_split uint_split_asm lemmas uint_arith_simps = word_le_def word_less_alt word_uint_eq_iff uint_sub_if' uint_plus_if' \ \use this to stop, eg. \2 ^ LENGTH(32)\ being simplified\ lemma power_False_cong: "False \ a ^ b = c ^ d" by auto \ \\uint_arith_tac\: reduce to arithmetic on int, try to solve by arith\ ML \ val uint_arith_simpset = @{context} |> fold Simplifier.add_simp @{thms uint_arith_simps} |> fold Splitter.add_split @{thms if_split_asm} |> fold Simplifier.add_cong @{thms power_False_cong} |> simpset_of; fun uint_arith_tacs ctxt = let fun arith_tac' n t = Arith_Data.arith_tac ctxt n t handle Cooper.COOPER _ => Seq.empty; in [ clarify_tac ctxt 1, full_simp_tac (put_simpset uint_arith_simpset ctxt) 1, ALLGOALS (full_simp_tac (put_simpset HOL_ss ctxt |> fold Splitter.add_split @{thms uint_splits} |> fold Simplifier.add_cong @{thms power_False_cong})), rewrite_goals_tac ctxt @{thms word_size}, ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN REPEAT (eresolve_tac ctxt [conjE] n) THEN REPEAT (dresolve_tac ctxt @{thms word_of_int_inverse} n THEN assume_tac ctxt n THEN assume_tac ctxt n)), TRYALL arith_tac' ] end fun uint_arith_tac ctxt = SELECT_GOAL (EVERY (uint_arith_tacs ctxt)) \ method_setup uint_arith = \Scan.succeed (SIMPLE_METHOD' o uint_arith_tac)\ "solving word arithmetic via integers and arith" subsection \More on overflows and monotonicity\ lemma no_plus_overflow_uint_size: "x \ x + y \ uint x + uint y < 2 ^ size x" for x y :: "'a::len word" unfolding word_size by uint_arith lemmas no_olen_add = no_plus_overflow_uint_size [unfolded word_size] lemma no_ulen_sub: "x \ x - y \ uint y \ uint x" for x y :: "'a::len word" by uint_arith lemma no_olen_add': "x \ y + x \ uint y + uint x < 2 ^ LENGTH('a)" for x y :: "'a::len word" by (simp add: ac_simps no_olen_add) lemmas olen_add_eqv = trans [OF no_olen_add no_olen_add' [symmetric]] lemmas uint_plus_simple_iff = trans [OF no_olen_add uint_add_lem] lemmas uint_plus_simple = uint_plus_simple_iff [THEN iffD1] lemmas uint_minus_simple_iff = trans [OF no_ulen_sub uint_sub_lem] lemmas uint_minus_simple_alt = uint_sub_lem [folded word_le_def] lemmas word_sub_le_iff = no_ulen_sub [folded word_le_def] lemmas word_sub_le = word_sub_le_iff [THEN iffD2] lemma word_less_sub1: "x \ 0 \ 1 < x \ 0 < x - 1" for x :: "'a::len word" by uint_arith lemma word_le_sub1: "x \ 0 \ 1 \ x \ 0 \ x - 1" for x :: "'a::len word" by uint_arith lemma sub_wrap_lt: "x < x - z \ x < z" for x z :: "'a::len word" by uint_arith lemma sub_wrap: "x \ x - z \ z = 0 \ x < z" for x z :: "'a::len word" by uint_arith lemma plus_minus_not_NULL_ab: "x \ ab - c \ c \ ab \ c \ 0 \ x + c \ 0" for x ab c :: "'a::len word" by uint_arith lemma plus_minus_no_overflow_ab: "x \ ab - c \ c \ ab \ x \ x + c" for x ab c :: "'a::len word" by uint_arith lemma le_minus': "a + c \ b \ a \ a + c \ c \ b - a" for a b c :: "'a::len word" by uint_arith lemma le_plus': "a \ b \ c \ b - a \ a + c \ b" for a b c :: "'a::len word" by uint_arith lemmas le_plus = le_plus' [rotated] lemmas le_minus = leD [THEN thin_rl, THEN le_minus'] (* FIXME *) lemma word_plus_mono_right: "y \ z \ x \ x + z \ x + y \ x + z" for x y z :: "'a::len word" by uint_arith lemma word_less_minus_cancel: "y - x < z - x \ x \ z \ y < z" for x y z :: "'a::len word" by uint_arith lemma word_less_minus_mono_left: "y < z \ x \ y \ y - x < z - x" for x y z :: "'a::len word" by uint_arith lemma word_less_minus_mono: "a < c \ d < b \ a - b < a \ c - d < c \ a - b < c - d" for a b c d :: "'a::len word" by uint_arith lemma word_le_minus_cancel: "y - x \ z - x \ x \ z \ y \ z" for x y z :: "'a::len word" by uint_arith lemma word_le_minus_mono_left: "y \ z \ x \ y \ y - x \ z - x" for x y z :: "'a::len word" by uint_arith lemma word_le_minus_mono: "a \ c \ d \ b \ a - b \ a \ c - d \ c \ a - b \ c - d" for a b c d :: "'a::len word" by uint_arith lemma plus_le_left_cancel_wrap: "x + y' < x \ x + y < x \ x + y' < x + y \ y' < y" for x y y' :: "'a::len word" by uint_arith lemma plus_le_left_cancel_nowrap: "x \ x + y' \ x \ x + y \ x + y' < x + y \ y' < y" for x y y' :: "'a::len word" by uint_arith lemma word_plus_mono_right2: "a \ a + b \ c \ b \ a \ a + c" for a b c :: "'a::len word" by uint_arith lemma word_less_add_right: "x < y - z \ z \ y \ x + z < y" for x y z :: "'a::len word" by uint_arith lemma word_less_sub_right: "x < y + z \ y \ x \ x - y < z" for x y z :: "'a::len word" by uint_arith lemma word_le_plus_either: "x \ y \ x \ z \ y \ y + z \ x \ y + z" for x y z :: "'a::len word" by uint_arith lemma word_less_nowrapI: "x < z - k \ k \ z \ 0 < k \ x < x + k" for x z k :: "'a::len word" by uint_arith lemma inc_le: "i < m \ i + 1 \ m" for i m :: "'a::len word" by uint_arith lemma inc_i: "1 \ i \ i < m \ 1 \ i + 1 \ i + 1 \ m" for i m :: "'a::len word" by uint_arith lemma udvd_incr_lem: "up < uq \ up = ua + n * uint K \ uq = ua + n' * uint K \ up + uint K \ uq" by auto (metis int_distrib(1) linorder_not_less mult.left_neutral mult_right_mono uint_nonnegative zless_imp_add1_zle) lemma udvd_incr': "p < q \ uint p = ua + n * uint K \ uint q = ua + n' * uint K \ p + K \ q" apply (unfold word_less_alt word_le_def) apply (drule (2) udvd_incr_lem) apply (erule uint_add_le [THEN order_trans]) done lemma udvd_decr': "p < q \ uint p = ua + n * uint K \ uint q = ua + n' * uint K \ p \ q - K" apply (unfold word_less_alt word_le_def) apply (drule (2) udvd_incr_lem) apply (drule le_diff_eq [THEN iffD2]) apply (erule order_trans) apply (rule uint_sub_ge) done lemmas udvd_incr_lem0 = udvd_incr_lem [where ua=0, unfolded add_0_left] lemmas udvd_incr0 = udvd_incr' [where ua=0, unfolded add_0_left] lemmas udvd_decr0 = udvd_decr' [where ua=0, unfolded add_0_left] lemma udvd_minus_le': "xy < k \ z udvd xy \ z udvd k \ xy \ k - z" apply (unfold udvd_unfold_int) apply clarify apply (erule (2) udvd_decr0) done lemma udvd_incr2_K: "p < a + s \ a \ a + s \ K udvd s \ K udvd p - a \ a \ p \ 0 < K \ p \ p + K \ p + K \ a + s" supply [[simproc del: linordered_ring_less_cancel_factor]] apply (unfold udvd_unfold_int) apply clarify apply (simp add: uint_arith_simps split: if_split_asm) prefer 2 using uint_lt2p [of s] apply simp apply (drule add.commute [THEN xtrans(1)]) apply (simp flip: diff_less_eq) apply (subst (asm) mult_less_cancel_right) apply simp apply (simp add: diff_eq_eq not_less) apply (subst (asm) (3) zless_iff_Suc_zadd) apply auto apply (auto simp add: algebra_simps) apply (drule less_le_trans [of _ \2 ^ LENGTH('a)\]) apply assumption apply (simp add: mult_less_0_iff) done subsection \Arithmetic type class instantiations\ lemmas word_le_0_iff [simp] = word_zero_le [THEN leD, THEN antisym_conv1] lemma word_of_int_nat: "0 \ x \ word_of_int x = of_nat (nat x)" by simp text \ note that \iszero_def\ is only for class \comm_semiring_1_cancel\, which requires word length \\ 1\, ie \'a::len word\ \ lemma iszero_word_no [simp]: "iszero (numeral bin :: 'a::len word) = iszero (take_bit LENGTH('a) (numeral bin :: int))" apply (simp add: iszero_def) apply transfer apply simp done text \Use \iszero\ to simplify equalities between word numerals.\ lemmas word_eq_numeral_iff_iszero [simp] = eq_numeral_iff_iszero [where 'a="'a::len word"] subsection \Word and nat\ lemma word_nchotomy: "\w :: 'a::len word. \n. w = of_nat n \ n < 2 ^ LENGTH('a)" apply (rule allI) apply (rule exI [of _ \unat w\ for w :: \'a word\]) apply simp done lemma of_nat_eq: "of_nat n = w \ (\q. n = unat w + q * 2 ^ LENGTH('a))" for w :: "'a::len word" using mod_div_mult_eq [of n "2 ^ LENGTH('a)", symmetric] by (auto simp flip: take_bit_eq_mod) lemma of_nat_eq_size: "of_nat n = w \ (\q. n = unat w + q * 2 ^ size w)" unfolding word_size by (rule of_nat_eq) lemma of_nat_0: "of_nat m = (0::'a::len word) \ (\q. m = q * 2 ^ LENGTH('a))" by (simp add: of_nat_eq) lemma of_nat_2p [simp]: "of_nat (2 ^ LENGTH('a)) = (0::'a::len word)" by (fact mult_1 [symmetric, THEN iffD2 [OF of_nat_0 exI]]) lemma of_nat_gt_0: "of_nat k \ 0 \ 0 < k" by (cases k) auto lemma of_nat_neq_0: "0 < k \ k < 2 ^ LENGTH('a::len) \ of_nat k \ (0 :: 'a word)" by (auto simp add : of_nat_0) lemma Abs_fnat_hom_add: "of_nat a + of_nat b = of_nat (a + b)" by simp lemma Abs_fnat_hom_mult: "of_nat a * of_nat b = (of_nat (a * b) :: 'a::len word)" by (simp add: wi_hom_mult) lemma Abs_fnat_hom_Suc: "word_succ (of_nat a) = of_nat (Suc a)" by transfer (simp add: ac_simps) lemma Abs_fnat_hom_0: "(0::'a::len word) = of_nat 0" by simp lemma Abs_fnat_hom_1: "(1::'a::len word) = of_nat (Suc 0)" by simp lemmas Abs_fnat_homs = Abs_fnat_hom_add Abs_fnat_hom_mult Abs_fnat_hom_Suc Abs_fnat_hom_0 Abs_fnat_hom_1 lemma word_arith_nat_add: "a + b = of_nat (unat a + unat b)" by simp lemma word_arith_nat_mult: "a * b = of_nat (unat a * unat b)" by simp lemma word_arith_nat_Suc: "word_succ a = of_nat (Suc (unat a))" by (subst Abs_fnat_hom_Suc [symmetric]) simp lemma word_arith_nat_div: "a div b = of_nat (unat a div unat b)" by (metis of_int_of_nat_eq of_nat_unat of_nat_div word_div_def) lemma word_arith_nat_mod: "a mod b = of_nat (unat a mod unat b)" by (metis of_int_of_nat_eq of_nat_mod of_nat_unat word_mod_def) lemmas word_arith_nat_defs = word_arith_nat_add word_arith_nat_mult word_arith_nat_Suc Abs_fnat_hom_0 Abs_fnat_hom_1 word_arith_nat_div word_arith_nat_mod lemma unat_cong: "x = y \ unat x = unat y" by (fact arg_cong) lemma unat_of_nat: \unat (word_of_nat x :: 'a::len word) = x mod 2 ^ LENGTH('a)\ by transfer (simp flip: take_bit_eq_mod add: nat_take_bit_eq) lemmas unat_word_ariths = word_arith_nat_defs [THEN trans [OF unat_cong unat_of_nat]] lemmas word_sub_less_iff = word_sub_le_iff [unfolded linorder_not_less [symmetric] Not_eq_iff] lemma unat_add_lem: "unat x + unat y < 2 ^ LENGTH('a) \ unat (x + y) = unat x + unat y" for x y :: "'a::len word" apply (auto simp: unat_word_ariths) apply (drule sym) apply (metis unat_of_nat unsigned_less) done lemma unat_mult_lem: "unat x * unat y < 2 ^ LENGTH('a) \ unat (x * y) = unat x * unat y" for x y :: "'a::len word" apply (auto simp: unat_word_ariths) apply (drule sym) apply (metis unat_of_nat unsigned_less) done lemma unat_plus_if': \unat (a + b) = (if unat a + unat b < 2 ^ LENGTH('a) then unat a + unat b else unat a + unat b - 2 ^ LENGTH('a))\ for a b :: \'a::len word\ apply (auto simp: unat_word_ariths not_less) apply (subst (asm) le_iff_add) apply auto apply (simp flip: take_bit_eq_mod add: take_bit_nat_eq_self_iff) apply (metis add.commute add_less_cancel_right le_less_trans less_imp_le unsigned_less) done lemma le_no_overflow: "x \ b \ a \ a + b \ x \ a + b" for a b x :: "'a::len word" apply (erule order_trans) apply (erule olen_add_eqv [THEN iffD1]) done lemmas un_ui_le = trans [OF word_le_nat_alt [symmetric] word_le_def] lemma unat_sub_if_size: "unat (x - y) = (if unat y \ unat x then unat x - unat y else unat x + 2 ^ size x - unat y)" supply nat_uint_eq [simp del] apply (unfold word_size) apply (simp add: un_ui_le) apply (auto simp add: unat_eq_nat_uint uint_sub_if') apply (rule nat_diff_distrib) prefer 3 apply (simp add: algebra_simps) apply (rule nat_diff_distrib [THEN trans]) prefer 3 apply (subst nat_add_distrib) prefer 3 apply (simp add: nat_power_eq) apply auto apply uint_arith done lemmas unat_sub_if' = unat_sub_if_size [unfolded word_size] lemma uint_div: \uint (x div y) = uint x div uint y\ by (fact uint_div_distrib) lemma unat_div: \unat (x div y) = unat x div unat y\ by (fact unat_div_distrib) lemma uint_mod: \uint (x mod y) = uint x mod uint y\ by (fact uint_mod_distrib) lemma unat_mod: \unat (x mod y) = unat x mod unat y\ by (fact unat_mod_distrib) text \Definition of \unat_arith\ tactic\ lemma unat_split: "P (unat x) \ (\n. of_nat n = x \ n < 2^LENGTH('a) \ P n)" for x :: "'a::len word" by auto (metis take_bit_nat_eq_self_iff) lemma unat_split_asm: "P (unat x) \ (\n. of_nat n = x \ n < 2^LENGTH('a) \ \ P n)" for x :: "'a::len word" by auto (metis take_bit_nat_eq_self_iff) lemma of_nat_inverse: \word_of_nat r = a \ r < 2 ^ LENGTH('a) \ unat a = r\ for a :: \'a::len word\ apply (drule sym) apply transfer apply (simp add: take_bit_int_eq_self) done lemma word_unat_eq_iff: \v = w \ unat v = unat w\ for v w :: \'a::len word\ by (fact word_eq_iff_unsigned) lemmas unat_splits = unat_split unat_split_asm lemmas unat_arith_simps = word_le_nat_alt word_less_nat_alt word_unat_eq_iff unat_sub_if' unat_plus_if' unat_div unat_mod \ \\unat_arith_tac\: tactic to reduce word arithmetic to \nat\, try to solve via \arith\\ ML \ val unat_arith_simpset = @{context} |> fold Simplifier.add_simp @{thms unat_arith_simps} |> fold Splitter.add_split @{thms if_split_asm} |> fold Simplifier.add_cong @{thms power_False_cong} |> simpset_of fun unat_arith_tacs ctxt = let fun arith_tac' n t = Arith_Data.arith_tac ctxt n t handle Cooper.COOPER _ => Seq.empty; in [ clarify_tac ctxt 1, full_simp_tac (put_simpset unat_arith_simpset ctxt) 1, ALLGOALS (full_simp_tac (put_simpset HOL_ss ctxt |> fold Splitter.add_split @{thms unat_splits} |> fold Simplifier.add_cong @{thms power_False_cong})), rewrite_goals_tac ctxt @{thms word_size}, ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN REPEAT (eresolve_tac ctxt [conjE] n) THEN REPEAT (dresolve_tac ctxt @{thms of_nat_inverse} n THEN assume_tac ctxt n)), TRYALL arith_tac' ] end fun unat_arith_tac ctxt = SELECT_GOAL (EVERY (unat_arith_tacs ctxt)) \ method_setup unat_arith = \Scan.succeed (SIMPLE_METHOD' o unat_arith_tac)\ "solving word arithmetic via natural numbers and arith" lemma no_plus_overflow_unat_size: "x \ x + y \ unat x + unat y < 2 ^ size x" for x y :: "'a::len word" unfolding word_size by unat_arith lemmas no_olen_add_nat = no_plus_overflow_unat_size [unfolded word_size] lemmas unat_plus_simple = trans [OF no_olen_add_nat unat_add_lem] lemma word_div_mult: "0 < y \ unat x * unat y < 2 ^ LENGTH('a) \ x * y div y = x" for x y :: "'a::len word" apply unat_arith apply clarsimp apply (subst unat_mult_lem [THEN iffD1]) apply auto done lemma div_lt': "i \ k div x \ unat i * unat x < 2 ^ LENGTH('a)" for i k x :: "'a::len word" apply unat_arith apply clarsimp apply (drule mult_le_mono1) apply (erule order_le_less_trans) apply (metis add_lessD1 div_mult_mod_eq unsigned_less) done lemmas div_lt'' = order_less_imp_le [THEN div_lt'] lemma div_lt_mult: "i < k div x \ 0 < x \ i * x < k" for i k x :: "'a::len word" apply (frule div_lt'' [THEN unat_mult_lem [THEN iffD1]]) apply (simp add: unat_arith_simps) apply (drule (1) mult_less_mono1) apply (erule order_less_le_trans) apply auto done lemma div_le_mult: "i \ k div x \ 0 < x \ i * x \ k" for i k x :: "'a::len word" apply (frule div_lt' [THEN unat_mult_lem [THEN iffD1]]) apply (simp add: unat_arith_simps) apply (drule mult_le_mono1) apply (erule order_trans) apply auto done lemma div_lt_uint': "i \ k div x \ uint i * uint x < 2 ^ LENGTH('a)" for i k x :: "'a::len word" apply (unfold uint_nat) apply (drule div_lt') apply (metis of_nat_less_iff of_nat_mult of_nat_numeral of_nat_power) done lemmas div_lt_uint'' = order_less_imp_le [THEN div_lt_uint'] lemma word_le_exists': "x \ y \ \z. y = x + z \ uint x + uint z < 2 ^ LENGTH('a)" for x y z :: "'a::len word" by (metis add_diff_cancel_left' add_diff_eq uint_add_lem uint_plus_simple) lemmas plus_minus_not_NULL = order_less_imp_le [THEN plus_minus_not_NULL_ab] lemmas plus_minus_no_overflow = order_less_imp_le [THEN plus_minus_no_overflow_ab] lemmas mcs = word_less_minus_cancel word_less_minus_mono_left word_le_minus_cancel word_le_minus_mono_left lemmas word_l_diffs = mcs [where y = "w + x", unfolded add_diff_cancel] for w x lemmas word_diff_ls = mcs [where z = "w + x", unfolded add_diff_cancel] for w x lemmas word_plus_mcs = word_diff_ls [where y = "v + x", unfolded add_diff_cancel] for v x lemma le_unat_uoi: \y \ unat z \ unat (word_of_nat y :: 'a word) = y\ for z :: \'a::len word\ by transfer (simp add: nat_take_bit_eq take_bit_nat_eq_self_iff le_less_trans) lemmas thd = times_div_less_eq_dividend lemmas uno_simps [THEN le_unat_uoi] = mod_le_divisor div_le_dividend lemma word_mod_div_equality: "(n div b) * b + (n mod b) = n" for n b :: "'a::len word" by (fact div_mult_mod_eq) lemma word_div_mult_le: "a div b * b \ a" for a b :: "'a::len word" by (metis div_le_mult mult_not_zero order.not_eq_order_implies_strict order_refl word_zero_le) lemma word_mod_less_divisor: "0 < n \ m mod n < n" for m n :: "'a::len word" by (simp add: unat_arith_simps) lemma word_of_int_power_hom: "word_of_int a ^ n = (word_of_int (a ^ n) :: 'a::len word)" by (induct n) (simp_all add: wi_hom_mult [symmetric]) lemma word_arith_power_alt: "a ^ n = (word_of_int (uint a ^ n) :: 'a::len word)" by (simp add : word_of_int_power_hom [symmetric]) lemma unatSuc: "1 + n \ 0 \ unat (1 + n) = Suc (unat n)" for n :: "'a::len word" by unat_arith subsection \Cardinality, finiteness of set of words\ lemma inj_on_word_of_int: \inj_on (word_of_int :: int \ 'a word) {0..<2 ^ LENGTH('a::len)}\ apply (rule inj_onI) apply transfer apply (simp add: take_bit_eq_mod) done lemma inj_uint: \inj uint\ by (fact inj_unsigned) lemma range_uint: \range (uint :: 'a word \ int) = {0..<2 ^ LENGTH('a::len)}\ by transfer (auto simp add: bintr_lt2p range_bintrunc) lemma UNIV_eq: \(UNIV :: 'a word set) = word_of_int ` {0..<2 ^ LENGTH('a::len)}\ proof - have \uint ` (UNIV :: 'a word set) = uint ` (word_of_int :: int \ 'a word) ` {0..<2 ^ LENGTH('a::len)}\ by transfer (simp add: range_bintrunc, auto simp add: take_bit_eq_mod) then show ?thesis using inj_image_eq_iff [of \uint :: 'a word \ int\ \UNIV :: 'a word set\ \word_of_int ` {0..<2 ^ LENGTH('a)} :: 'a word set\, OF inj_uint] by simp qed lemma card_word: "CARD('a word) = 2 ^ LENGTH('a::len)" by (simp add: UNIV_eq card_image inj_on_word_of_int) lemma card_word_size: "CARD('a word) = 2 ^ size x" for x :: "'a::len word" unfolding word_size by (rule card_word) instance word :: (len) finite by standard (simp add: UNIV_eq) subsection \Bitwise Operations on Words\ 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 word_wi_log_defs: "NOT (word_of_int a) = word_of_int (NOT a)" "word_of_int a AND word_of_int b = word_of_int (a AND b)" "word_of_int a OR word_of_int b = word_of_int (a OR b)" "word_of_int a XOR word_of_int b = word_of_int (a XOR b)" by (transfer, rule refl)+ lemma word_no_log_defs [simp]: "NOT (numeral a) = word_of_int (NOT (numeral a))" "NOT (- numeral a) = word_of_int (NOT (- numeral a))" "numeral a AND numeral b = word_of_int (numeral a AND numeral b)" "numeral a AND - numeral b = word_of_int (numeral a AND - numeral b)" "- numeral a AND numeral b = word_of_int (- numeral a AND numeral b)" "- numeral a AND - numeral b = word_of_int (- numeral a AND - numeral b)" "numeral a OR numeral b = word_of_int (numeral a OR numeral b)" "numeral a OR - numeral b = word_of_int (numeral a OR - numeral b)" "- numeral a OR numeral b = word_of_int (- numeral a OR numeral b)" "- numeral a OR - numeral b = word_of_int (- numeral a OR - numeral b)" "numeral a XOR numeral b = word_of_int (numeral a XOR numeral b)" "numeral a XOR - numeral b = word_of_int (numeral a XOR - numeral b)" "- numeral a XOR numeral b = word_of_int (- numeral a XOR numeral b)" "- numeral a XOR - numeral b = word_of_int (- numeral a XOR - numeral b)" by (transfer, rule refl)+ text \Special cases for when one of the arguments equals 1.\ lemma word_bitwise_1_simps [simp]: "NOT (1::'a::len word) = -2" "1 AND numeral b = word_of_int (1 AND numeral b)" "1 AND - numeral b = word_of_int (1 AND - numeral b)" "numeral a AND 1 = word_of_int (numeral a AND 1)" "- numeral a AND 1 = word_of_int (- numeral a AND 1)" "1 OR numeral b = word_of_int (1 OR numeral b)" "1 OR - numeral b = word_of_int (1 OR - numeral b)" "numeral a OR 1 = word_of_int (numeral a OR 1)" "- numeral a OR 1 = word_of_int (- numeral a OR 1)" "1 XOR numeral b = word_of_int (1 XOR numeral b)" "1 XOR - numeral b = word_of_int (1 XOR - numeral b)" "numeral a XOR 1 = word_of_int (numeral a XOR 1)" "- numeral a XOR 1 = word_of_int (- numeral a XOR 1)" by (transfer, simp)+ text \Special cases for when one of the arguments equals -1.\ lemma word_bitwise_m1_simps [simp]: "NOT (-1::'a::len word) = 0" "(-1::'a::len word) AND x = x" "x AND (-1::'a::len word) = x" "(-1::'a::len word) OR x = -1" "x OR (-1::'a::len word) = -1" " (-1::'a::len word) XOR x = NOT x" "x XOR (-1::'a::len word) = NOT x" by (transfer, simp)+ lemma uint_and: \uint (x AND y) = uint x AND uint y\ by transfer simp lemma uint_or: \uint (x OR y) = uint x OR uint y\ by transfer simp lemma uint_xor: \uint (x XOR y) = uint x XOR uint y\ by transfer simp 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) \ bin_nth 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" unfolding word_size by transfer (simp add: bin_nth_ops) 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: bin_nth_ops) 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) \ bin_nth (numeral w) n" by transfer (rule refl) lemma test_bit_neg_numeral [simp]: "(- numeral w :: 'a::len word) !! n \ n < LENGTH('a) \ bin_nth (- numeral w) n" by transfer (rule refl) lemma test_bit_1 [simp]: "(1 :: 'a::len word) !! n \ n = 0" by transfer auto 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 \ \get from commutativity, associativity etc of \int_and\ etc to same for \word_and etc\\ lemmas bwsimps = wi_hom_add word_wi_log_defs lemma word_bw_assocs: "(x AND y) AND z = x AND y AND z" "(x OR y) OR z = x OR y OR z" "(x XOR y) XOR z = x XOR y XOR z" for x :: "'a::len word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_bw_comms: "x AND y = y AND x" "x OR y = y OR x" "x XOR y = y XOR x" for x :: "'a::len word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_bw_lcs: "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 :: "'a::len word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_log_esimps: "x AND 0 = 0" "x AND -1 = x" "x OR 0 = x" "x OR -1 = -1" "x XOR 0 = x" "x XOR -1 = NOT x" "0 AND x = 0" "-1 AND x = x" "0 OR x = x" "-1 OR x = -1" "0 XOR x = x" "-1 XOR x = NOT x" for x :: "'a::len word" by simp_all lemma word_not_dist: "NOT (x OR y) = NOT x AND NOT y" "NOT (x AND y) = NOT x OR NOT y" for x :: "'a::len word" by simp_all lemma word_bw_same: "x AND x = x" "x OR x = x" "x XOR x = 0" for x :: "'a::len word" by simp_all lemma word_ao_absorbs [simp]: "x AND (y OR x) = x" "x OR y AND x = x" "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 :: "'a::len word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_not_not [simp]: "NOT (NOT x) = x" for x :: "'a::len word" by simp lemma word_ao_dist: "(x OR y) AND z = x AND z OR y AND z" for x :: "'a::len word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_oa_dist: "x AND y OR z = (x OR z) AND (y OR z)" for x :: "'a::len word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_add_not [simp]: "x + NOT x = -1" for x :: "'a::len word" by transfer (simp add: bin_add_not) lemma word_plus_and_or [simp]: "(x AND y) + (x OR y) = x + y" for x :: "'a::len word" by transfer (simp add: plus_and_or) lemma leoa: "w = x OR y \ y = w AND y" for x :: "'a::len word" by auto lemma leao: "w' = x' AND y' \ x' = x' OR w'" for x' :: "'a::len word" by auto lemma word_ao_equiv: "w = w OR w' \ w' = w AND w'" for w w' :: "'a::len word" by (auto intro: leoa leao) lemma le_word_or2: "x \ x OR y" for x y :: "'a::len word" by (auto simp: word_le_def uint_or intro: le_int_or) lemmas le_word_or1 = xtrans(3) [OF word_bw_comms (2) le_word_or2] lemmas word_and_le1 = xtrans(3) [OF word_ao_absorbs (4) [symmetric] le_word_or2] lemmas word_and_le2 = xtrans(3) [OF word_ao_absorbs (8) [symmetric] le_word_or2] lemma bit_horner_sum_bit_word_iff: \bit (horner_sum of_bool (2 :: 'a::len word) bs) n \ n < min LENGTH('a) (length bs) \ bs ! n\ by transfer (simp add: bit_horner_sum_bit_iff) definition word_reverse :: \'a::len word \ 'a word\ where \word_reverse w = horner_sum of_bool 2 (rev (map (bit w) [0.. lemma bit_word_reverse_iff: \bit (word_reverse w) n \ n < LENGTH('a) \ bit w (LENGTH('a) - Suc n)\ for w :: \'a::len word\ by (cases \n < LENGTH('a)\) (simp_all add: word_reverse_def bit_horner_sum_bit_word_iff rev_nth) lemma word_rev_rev [simp] : "word_reverse (word_reverse w) = w" by (rule bit_word_eqI) (auto simp add: bit_word_reverse_iff bit_imp_le_length Suc_diff_Suc) lemma word_rev_gal: "word_reverse w = u \ word_reverse u = w" by (metis word_rev_rev) lemma word_rev_gal': "u = word_reverse w \ w = word_reverse u" by simp 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 "bin_nth (sint w) n = (if n < l - 1 then w !! n else w !! (l - 1))" unfolding sint_uint l_def by (auto simp: nth_sbintr word_test_bit_def [symmetric]) 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 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 uint_2p: "(0::'a::len word) < 2 ^ n \ uint (2 ^ n::'a::len word) = 2 ^ n" apply (cases \n < LENGTH('a)\; transfer) apply auto done lemma word_of_int_2p: "(word_of_int (2 ^ n) :: 'a::len word) = 2 ^ n" by (induct n) (simp_all add: wi_hom_syms) 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 subsection \Shifting, Rotating, and Splitting Words\ lemma shiftl1_wi [simp]: "shiftl1 (word_of_int w) = word_of_int (2 * w)" by transfer simp lemma shiftl1_numeral [simp]: "shiftl1 (numeral w) = numeral (Num.Bit0 w)" unfolding word_numeral_alt shiftl1_wi by simp lemma shiftl1_neg_numeral [simp]: "shiftl1 (- numeral w) = - numeral (Num.Bit0 w)" unfolding word_neg_numeral_alt shiftl1_wi by simp lemma shiftl1_0 [simp] : "shiftl1 0 = 0" by transfer simp lemma shiftl1_def_u: "shiftl1 w = word_of_int (2 * uint w)" by (fact shiftl1_eq) lemma shiftl1_def_s: "shiftl1 w = word_of_int (2 * sint w)" by (simp add: shiftl1_def_u wi_hom_syms) lemma shiftr1_0 [simp]: "shiftr1 0 = 0" by transfer simp lemma sshiftr1_0 [simp]: "sshiftr1 0 = 0" by transfer simp lemma sshiftr1_n1 [simp]: "sshiftr1 (- 1) = - 1" by transfer simp 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 sshiftr_0 [simp]: "0 >>> n = 0" by transfer simp lemma sshiftr_n1 [simp]: "-1 >>> n = -1" 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 text \ see paper page 10, (1), (2), \shiftr1_def\ is of the form of (1), where \f\ (ie \bin_rest\) takes normal arguments to normal results, thus we get (2) from (1) \ lemma uint_shiftr1: "uint (shiftr1 w) = bin_rest (uint w)" by transfer simp lemma bit_sshiftr1_iff: \bit (sshiftr1 w) n \ bit w (if n = LENGTH('a) - 1 then LENGTH('a) - 1 else Suc n)\ for w :: \'a::len word\ 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 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 sshiftr_eq: \w >>> m = signed_drop_bit m w\ by (rule bit_eqI) (simp add: bit_signed_drop_bit_iff bit_sshiftr_word_iff) 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 nth_sshiftr : "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 shiftr1_div_2: "uint (shiftr1 w) = uint w div 2" by (fact uint_shiftr1) lemma sshiftr1_div_2: "sint (sshiftr1 w) = sint w div 2" by transfer simp lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" apply (unfold shiftr_def) apply (induct n) apply simp apply (simp add: shiftr1_div_2 mult.commute zdiv_zmult2_eq [symmetric]) done lemma sshiftr_div_2n: "sint (sshiftr w n) = sint w div 2 ^ n" apply transfer apply (rule bit_eqI) apply (simp add: bit_signed_take_bit_iff bit_drop_bit_eq min_def flip: drop_bit_eq_div) done lemma bit_bshiftr1_iff: \bit (bshiftr1 b w) n \ b \ n = LENGTH('a) - 1 \ bit w (Suc n)\ for w :: \'a::len word\ apply transfer apply (simp add: bit_take_bit_iff flip: bit_Suc) apply (subst disjunctive_add) apply (auto simp add: bit_take_bit_iff bit_or_iff bit_exp_iff simp flip: bit_Suc) done subsubsection \shift functions in terms of lists of bools\ lemma shiftl1_rev: "shiftl1 w = word_reverse (shiftr1 (word_reverse w))" apply (rule bit_word_eqI) apply (auto simp add: bit_shiftl1_iff bit_word_reverse_iff bit_shiftr1_iff Suc_diff_Suc) done 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 \ \note -- the following results use \'a::len word < number_ring\\ lemma shiftl1_2t: "shiftl1 w = 2 * w" for w :: "'a::len word" by (simp add: shiftl1_eq wi_hom_mult [symmetric]) lemma shiftl1_p: "shiftl1 w = w + w" for w :: "'a::len word" by (simp add: shiftl1_2t) lemma shiftl_t2n: "shiftl w n = 2 ^ n * w" for w :: "'a::len word" by (induct n) (auto simp: shiftl_def shiftl1_2t) lemma shiftr1_bintr [simp]: "(shiftr1 (numeral w) :: 'a::len word) = word_of_int (bin_rest (take_bit (LENGTH('a)) (numeral w)))" by transfer simp lemma sshiftr1_sbintr [simp]: "(sshiftr1 (numeral w) :: 'a::len word) = word_of_int (bin_rest (signed_take_bit (LENGTH('a) - 1) (numeral w)))" by transfer simp text \TODO: rules for \<^term>\- (numeral n)\\ lemma drop_bit_word_numeral [simp]: \drop_bit (numeral n) (numeral k) = (word_of_int (drop_bit (numeral n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\ by transfer simp lemma shiftr_numeral [simp]: \(numeral k >> numeral n :: 'a::len word) = drop_bit (numeral n) (numeral k)\ by (fact shiftr_word_eq) 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 nth_sbintr not_le not_less less_Suc_eq_le ac_simps) done lemma zip_replicate: "n \ length ys \ zip (replicate n x) ys = map (\y. (x, y)) ys" apply (induct ys arbitrary: n) apply simp_all apply (case_tac n) apply simp_all done lemma align_lem_or [rule_format] : "\x m. length x = n + m \ length y = n + m \ drop m x = replicate n False \ take m y = replicate m False \ map2 (|) x y = take m x @ drop m y" apply (induct y) apply force apply clarsimp apply (case_tac x) apply force apply (case_tac m) apply auto apply (drule_tac t="length xs" for xs in sym) apply (auto simp: zip_replicate o_def) done lemma align_lem_and [rule_format] : "\x m. length x = n + m \ length y = n + m \ drop m x = replicate n False \ take m y = replicate m False \ map2 (\) x y = replicate (n + m) False" apply (induct y) apply force apply clarsimp apply (case_tac x) apply force apply (case_tac m) apply auto apply (drule_tac t="length xs" for xs in sym) apply (auto simp: zip_replicate o_def map_replicate_const) done subsubsection \Mask\ lemma minus_1_eq_mask: \- 1 = (mask LENGTH('a) :: 'a::len word)\ by (rule bit_eqI) (simp add: bit_exp_iff bit_mask_iff exp_eq_zero_iff) lemma mask_eq_decr_exp: \mask n = 2 ^ n - (1 :: 'a::len word)\ by (fact mask_eq_exp_minus_1) lemma mask_Suc_rec: \mask (Suc n) = 2 * mask n + (1 :: 'a::len word)\ by (simp add: mask_eq_exp_minus_1) context begin qualified lemma bit_mask_iff: \bit (mask m :: 'a::len word) n \ n < min LENGTH('a) m\ by (simp add: bit_mask_iff exp_eq_zero_iff not_le) end 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 mask_bin: "mask n = word_of_int (take_bit n (- 1))" by (auto simp add: nth_bintr word_size intro: word_eqI) lemma and_mask_bintr: "w AND mask n = word_of_int (take_bit n (uint w))" apply (rule word_eqI) apply (simp add: nth_bintr word_size word_ops_nth_size) apply (auto simp add: test_bit_bin) done lemma and_mask_wi: "word_of_int i AND mask n = word_of_int (take_bit n i)" by (auto simp add: nth_bintr word_size word_ops_nth_size word_eq_iff) lemma and_mask_wi': "word_of_int i AND mask n = (word_of_int (take_bit (min LENGTH('a) n) i) :: 'a::len word)" by (auto simp add: nth_bintr word_size word_ops_nth_size word_eq_iff) lemma and_mask_no: "numeral i AND mask n = word_of_int (take_bit n (numeral i))" unfolding word_numeral_alt by (rule and_mask_wi) lemma and_mask_mod_2p: "w AND mask n = word_of_int (uint w mod 2 ^ n)" by (simp only: and_mask_bintr take_bit_eq_mod) lemma uint_mask_eq: \uint (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ by transfer simp lemma and_mask_lt_2p: "uint (w AND mask n) < 2 ^ n" apply (simp add: uint_and uint_mask_eq) apply (rule AND_upper2'') apply simp apply (auto simp add: mask_eq_exp_minus_1 min_def power_add algebra_simps dest!: le_Suc_ex) apply (metis add_minus_cancel le_add2 one_le_numeral power_add power_increasing uminus_add_conv_diff zle_diff1_eq) done 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 mask_eq_iff: "w AND mask n = w \ uint w < 2 ^ n" apply (auto simp flip: take_bit_eq_mask) apply (metis take_bit_int_eq_self_iff uint_take_bit_eq) apply (simp add: take_bit_int_eq_self unsigned_take_bit_eq word_uint_eqI) done lemma and_mask_dvd: "2 ^ n dvd uint w \ w AND mask n = 0" by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 uint_0_iff) lemma and_mask_dvd_nat: "2 ^ n dvd unat w \ w AND mask n = 0" by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 unat_0_iff uint_0_iff) lemma word_2p_lem: "n < size w \ w < 2 ^ n = (uint w < 2 ^ n)" for w :: "'a::len word" by transfer simp lemma less_mask_eq: "x < 2 ^ n \ x AND mask n = x" for x :: "'a::len word" apply (simp add: and_mask_bintr) apply transfer apply (simp add: ac_simps) apply (auto simp add: min_def) apply (metis bintrunc_bintrunc_ge mod_pos_pos_trivial mult.commute mult.left_neutral mult_zero_left not_le of_bool_def take_bit_eq_mod take_bit_nonnegative) done lemmas mask_eq_iff_w2p = trans [OF mask_eq_iff word_2p_lem [symmetric]] lemmas and_mask_less' = iffD2 [OF word_2p_lem and_mask_lt_2p, simplified word_size] lemma and_mask_less_size: "n < size x \ x AND mask n < 2 ^ n" for x :: \'a::len word\ unfolding word_size by (erule and_mask_less') lemma word_mod_2p_is_mask [OF refl]: "c = 2 ^ n \ c > 0 \ x mod c = x AND mask n" for c x :: "'a::len word" by (auto simp: word_mod_def uint_2p and_mask_mod_2p) lemma mask_eqs: "(a AND mask n) + b AND mask n = a + b AND mask n" "a + (b AND mask n) AND mask n = a + b AND mask n" "(a AND mask n) - b AND mask n = a - b AND mask n" "a - (b AND mask n) AND mask n = a - b AND mask n" "a * (b AND mask n) AND mask n = a * b AND mask n" "(b AND mask n) * a AND mask n = b * a AND mask n" "(a AND mask n) + (b AND mask n) AND mask n = a + b AND mask n" "(a AND mask n) - (b AND mask n) AND mask n = a - b AND mask n" "(a AND mask n) * (b AND mask n) AND mask n = a * b AND mask n" "- (a AND mask n) AND mask n = - a AND mask n" "word_succ (a AND mask n) AND mask n = word_succ a AND mask n" "word_pred (a AND mask n) AND mask n = word_pred a AND mask n" using word_of_int_Ex [where x=a] word_of_int_Ex [where x=b] apply (auto simp flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) done lemma mask_power_eq: "(x AND mask n) ^ k AND mask n = x ^ k AND mask n" for x :: \'a::len word\ using word_of_int_Ex [where x=x] apply (auto simp flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_eq_mod mod_simps) done lemma mask_full [simp]: "mask LENGTH('a) = (- 1 :: 'a::len word)" by transfer (simp add: take_bit_minus_one_eq_mask) subsubsection \Slices\ definition slice1 :: \nat \ 'a::len word \ 'b::len word\ where \slice1 n w = (if n < LENGTH('a) then ucast (drop_bit (LENGTH('a) - n) w) else push_bit (n - LENGTH('a)) (ucast w))\ lemma bit_slice1_iff: \bit (slice1 m w :: 'b::len word) n \ m - LENGTH('a) \ n \ n < min LENGTH('b) m \ bit w (n + (LENGTH('a) - m) - (m - LENGTH('a)))\ for w :: \'a::len word\ by (auto simp add: slice1_def bit_ucast_iff bit_drop_bit_eq bit_push_bit_iff exp_eq_zero_iff not_less not_le ac_simps dest: bit_imp_le_length) definition slice :: \nat \ 'a::len word \ 'b::len word\ where \slice n = slice1 (LENGTH('a) - n)\ lemma bit_slice_iff: \bit (slice m w :: 'b::len word) n \ n < min LENGTH('b) (LENGTH('a) - m) \ bit w (n + LENGTH('a) - (LENGTH('a) - m))\ for w :: \'a::len word\ by (simp add: slice_def word_size bit_slice1_iff) lemma slice1_0 [simp] : "slice1 n 0 = 0" unfolding slice1_def by simp lemma slice_0 [simp] : "slice n 0 = 0" unfolding slice_def by auto 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 ucast_slice1: "ucast w = slice1 (size w) w" apply (simp add: slice1_def) apply transfer apply simp done lemma ucast_slice: "ucast w = slice 0 w" by (simp add: slice_def slice1_def) lemma slice_id: "slice 0 t = t" by (simp only: ucast_slice [symmetric] ucast_id) lemma rev_slice1: \slice1 n (word_reverse w :: 'b::len word) = word_reverse (slice1 k w :: 'a::len word)\ if \n + k = LENGTH('a) + LENGTH('b)\ proof (rule bit_word_eqI) fix m assume *: \m < LENGTH('a)\ from that have **: \LENGTH('b) = n + k - LENGTH('a)\ by simp show \bit (slice1 n (word_reverse w :: 'b word) :: 'a word) m \ bit (word_reverse (slice1 k w :: 'a word)) m\ apply (simp add: bit_slice1_iff bit_word_reverse_iff) using * ** apply (cases \n \ LENGTH('a)\; cases \k \ LENGTH('a)\) apply auto done qed lemma rev_slice: "n + k + LENGTH('a::len) = LENGTH('b::len) \ slice n (word_reverse (w::'b word)) = word_reverse (slice k w :: 'a word)" apply (unfold slice_def word_size) apply (rule rev_slice1) apply arith done subsubsection \Revcast\ definition revcast :: \'a::len word \ 'b::len word\ where \revcast = slice1 LENGTH('b)\ lemma bit_revcast_iff: \bit (revcast w :: 'b::len word) n \ LENGTH('b) - LENGTH('a) \ n \ n < LENGTH('b) \ bit w (n + (LENGTH('a) - LENGTH('b)) - (LENGTH('b) - LENGTH('a)))\ for w :: \'a::len word\ by (simp add: revcast_def bit_slice1_iff) lemma revcast_slice1 [OF refl]: "rc = revcast w \ slice1 (size rc) w = rc" by (simp add: revcast_def word_size) lemma revcast_rev_ucast [OF refl refl refl]: "cs = [rc, uc] \ rc = revcast (word_reverse w) \ uc = ucast w \ rc = word_reverse uc" apply auto apply (rule bit_word_eqI) apply (cases \LENGTH('a) \ LENGTH('b)\) apply (simp_all add: bit_revcast_iff bit_word_reverse_iff bit_ucast_iff not_le bit_imp_le_length) using bit_imp_le_length apply fastforce using bit_imp_le_length apply fastforce done lemma revcast_ucast: "revcast w = word_reverse (ucast (word_reverse w))" using revcast_rev_ucast [of "word_reverse w"] by simp lemma ucast_revcast: "ucast w = word_reverse (revcast (word_reverse w))" by (fact revcast_rev_ucast [THEN word_rev_gal']) lemma ucast_rev_revcast: "ucast (word_reverse w) = word_reverse (revcast w)" by (fact revcast_ucast [THEN word_rev_gal']) text "linking revcast and cast via shift" lemmas wsst_TYs = source_size target_size word_size 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_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_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 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 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]] lemmas sym_notr = not_iff [THEN iffD2, THEN not_sym, THEN not_iff [THEN iffD1]] \ \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) subsection \Split and cat\ lemmas word_split_bin' = word_split_def lemmas word_cat_bin' = word_cat_eq -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]] - 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 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) \ \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)))" apply (unfold word_split_bin' test_bit_bin) apply (clarify) apply simp apply (erule conjE) apply (drule sym) apply (drule sym) apply (simp add: bit_take_bit_iff bit_drop_bit_eq) apply transfer apply (simp add: bit_take_bit_iff ac_simps) done 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 \ \this odd result is analogous to \ucast_id\, result to the length given by the result type\ lemma word_cat_id: "word_cat a b = b" by transfer simp \ \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))" apply transfer using bintr_cat by auto lemma word_cat_split_alt: "size w \ size u + size v \ word_split w = (u, v) \ word_cat u v = w" apply (rule word_eqI) apply (drule test_bit_split) apply (clarsimp simp add : test_bit_cat word_size) apply safe apply arith done lemmas word_cat_split_size = sym [THEN [2] word_cat_split_alt [symmetric]] subsubsection \Split and slice\ lemma split_slices: "word_split w = (u, v) \ u = slice (size v) w \ v = slice 0 w" apply (drule test_bit_split) apply (rule conjI) apply (rule word_eqI, clarsimp simp: nth_slice word_size)+ done lemma slice_cat1 [OF refl]: "wc = word_cat a b \ size wc >= size a + size b \ slice (size b) wc = a" apply safe apply (rule word_eqI) apply (simp add: nth_slice test_bit_cat word_size) done lemmas slice_cat2 = trans [OF slice_id word_cat_id] lemma cat_slices: "a = slice n c \ b = slice 0 c \ n = size b \ size a + size b >= size c \ word_cat a b = c" apply safe apply (rule word_eqI) apply (simp add: nth_slice test_bit_cat word_size) apply safe apply arith done lemma word_split_cat_alt: "w = word_cat u v \ size u + size v \ size w \ word_split w = (u, v)" apply (case_tac "word_split w") apply (rule trans, assumption) apply (drule test_bit_split) apply safe apply (rule word_eqI, clarsimp simp: test_bit_cat word_size)+ done -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 horner_sum_uint_exp_Cons_eq: \horner_sum uint (2 ^ LENGTH('a)) (w # ws) = concat_bit LENGTH('a) (uint w) (horner_sum uint (2 ^ LENGTH('a)) ws)\ for ws :: \'a::len word list\ apply (simp add: concat_bit_eq push_bit_eq_mult) apply transfer apply simp done lemma bit_horner_sum_uint_exp_iff: \bit (horner_sum uint (2 ^ LENGTH('a)) ws) n \ n div LENGTH('a) < length ws \ bit (ws ! (n div LENGTH('a))) (n mod LENGTH('a))\ for ws :: \'a::len word list\ proof (induction ws arbitrary: n) case Nil then show ?case by simp next case (Cons w ws) then show ?case by (cases \n \ LENGTH('a)\) (simp_all only: horner_sum_uint_exp_Cons_eq, simp_all add: bit_concat_bit_iff le_div_geq le_mod_geq bit_uint_iff Cons) qed 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 bin_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 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 - subsection \Rotation\ lemma word_rotr_word_rotr_eq: \word_rotr m (word_rotr n w) = word_rotr (m + n) w\ by (rule bit_word_eqI) (simp add: bit_word_rotr_iff ac_simps mod_add_right_eq) lemma word_rot_rl [simp]: \word_rotl k (word_rotr k v) = v\ apply (rule bit_word_eqI) apply (simp add: word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps) apply (auto dest: bit_imp_le_length) apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_if mod_mult_self2_is_0) apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_less mod_mult_self2_is_0) done lemma word_rot_lr [simp]: \word_rotr k (word_rotl k v) = v\ apply (rule bit_word_eqI) apply (simp add: word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps) apply (auto dest: bit_imp_le_length) apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_if mod_mult_self2_is_0) apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_less mod_mult_self2_is_0) done lemma word_rot_gal: \word_rotr n v = w \ word_rotl n w = v\ by auto lemma word_rot_gal': \w = word_rotr n v \ v = word_rotl n w\ by auto lemma word_rotr_rev: \word_rotr n w = word_reverse (word_rotl n (word_reverse w))\ proof (rule bit_word_eqI) fix m assume \m < LENGTH('a)\ moreover have \1 + ((int m + int n mod int LENGTH('a)) mod int LENGTH('a) + ((int LENGTH('a) * 2) mod int LENGTH('a) - (1 + (int m + int n mod int LENGTH('a)))) mod int LENGTH('a)) = int LENGTH('a)\ apply (cases \(1 + (int m + int n mod int LENGTH('a))) mod int LENGTH('a) = 0\) using zmod_zminus1_eq_if [of \1 + (int m + int n mod int LENGTH('a))\ \int LENGTH('a)\] apply simp_all apply (auto simp add: algebra_simps) apply (simp add: minus_equation_iff [of \int m\]) apply (drule sym [of _ \int m\]) apply simp apply (metis add.commute add_minus_cancel diff_minus_eq_add len_gt_0 less_imp_of_nat_less less_nat_zero_code mod_mult_self3 of_nat_gt_0 zmod_minus1) apply (metis (no_types, hide_lams) Abs_fnat_hom_add less_not_refl mod_Suc of_nat_Suc of_nat_gt_0 of_nat_mod) done then have \int ((m + n) mod LENGTH('a)) = int (LENGTH('a) - Suc ((LENGTH('a) - Suc m + LENGTH('a) - n mod LENGTH('a)) mod LENGTH('a)))\ using \m < LENGTH('a)\ by (simp only: of_nat_mod mod_simps) (simp add: of_nat_diff of_nat_mod Suc_le_eq add_less_mono algebra_simps mod_simps) then have \(m + n) mod LENGTH('a) = LENGTH('a) - Suc ((LENGTH('a) - Suc m + LENGTH('a) - n mod LENGTH('a)) mod LENGTH('a))\ by simp ultimately show \bit (word_rotr n w) m \ bit (word_reverse (word_rotl n (word_reverse w))) m\ by (simp add: word_rotl_eq_word_rotr bit_word_rotr_iff bit_word_reverse_iff) qed lemma word_roti_0 [simp]: "word_roti 0 w = w" by transfer simp lemma word_roti_add: "word_roti (m + n) w = word_roti m (word_roti n w)" by (rule bit_word_eqI) (simp add: bit_word_roti_iff nat_less_iff mod_simps ac_simps) lemma word_roti_conv_mod': "word_roti n w = word_roti (n mod int (size w)) w" by transfer simp lemmas word_roti_conv_mod = word_roti_conv_mod' [unfolded word_size] subsubsection \"Word rotation commutes with bit-wise operations\ \ \using locale to not pollute lemma namespace\ locale word_rotate begin lemma word_rot_logs: "word_rotl n (NOT v) = NOT (word_rotl n v)" "word_rotr n (NOT v) = NOT (word_rotr n v)" "word_rotl n (x AND y) = word_rotl n x AND word_rotl n y" "word_rotr n (x AND y) = word_rotr n x AND word_rotr n y" "word_rotl n (x OR y) = word_rotl n x OR word_rotl n y" "word_rotr n (x OR y) = word_rotr n x OR word_rotr n y" "word_rotl n (x XOR y) = word_rotl n x XOR word_rotl n y" "word_rotr n (x XOR y) = word_rotr n x XOR word_rotr n y" apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotl_iff bit_not_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotr_iff bit_not_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotl_iff bit_and_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotr_iff bit_and_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotl_iff bit_or_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotr_iff bit_or_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotl_iff bit_xor_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotr_iff bit_xor_iff algebra_simps exp_eq_zero_iff not_le) done end lemmas word_rot_logs = word_rotate.word_rot_logs lemma word_rotx_0 [simp] : "word_rotr i 0 = 0 \ word_rotl i 0 = 0" by transfer simp_all lemma word_roti_0' [simp] : "word_roti n 0 = 0" by transfer simp declare word_roti_eq_word_rotr_word_rotl [simp] subsection \Maximum machine word\ lemma word_int_cases: fixes x :: "'a::len word" obtains n where "x = word_of_int n" and "0 \ n" and "n < 2^LENGTH('a)" by (rule that [of \uint x\]) simp_all lemma word_nat_cases [cases type: word]: fixes x :: "'a::len word" obtains n where "x = of_nat n" and "n < 2^LENGTH('a)" by (rule that [of \unat x\]) simp_all lemma max_word_max [intro!]: "n \ max_word" by (fact word_order.extremum) lemma word_of_int_2p_len: "word_of_int (2 ^ LENGTH('a)) = (0::'a::len word)" by simp lemma word_pow_0: "(2::'a::len word) ^ LENGTH('a) = 0" by (fact word_exp_length_eq_0) lemma max_word_wrap: "x + 1 = 0 \ x = max_word" by (simp add: eq_neg_iff_add_eq_0) lemma max_test_bit: "(max_word::'a::len word) !! n \ n < LENGTH('a)" by (fact nth_minus1) lemma word_and_max: "x AND max_word = x" by (fact word_log_esimps) lemma word_or_max: "x OR max_word = max_word" by (fact word_log_esimps) lemma word_ao_dist2: "x AND (y OR z) = x AND y OR x AND z" for x y z :: "'a::len word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_oa_dist2: "x OR y AND z = (x OR y) AND (x OR z)" for x y z :: "'a::len word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_and_not [simp]: "x AND NOT x = 0" for x :: "'a::len word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_or_not [simp]: "x OR NOT x = max_word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_xor_and_or: "x XOR y = x AND NOT y OR NOT x AND y" for x y :: "'a::len word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) 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 uint_lt_0 [simp]: "uint x < 0 = False" by (simp add: linorder_not_less) lemma shiftr1_1 [simp]: "shiftr1 (1::'a::len word) = 0" by transfer simp lemma shiftr_1[simp]: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" by (induct n) (auto simp: shiftr_def) lemma word_less_1 [simp]: "x < 1 \ x = 0" for x :: "'a::len word" by (simp add: word_less_nat_alt unat_0_iff) lemma map_nth_0 [simp]: "map ((!!) (0::'a::len word)) xs = replicate (length xs) False" by (induct xs) auto lemma uint_plus_if_size: "uint (x + y) = (if uint x + uint y < 2^size x then uint x + uint y else uint x + uint y - 2^size x)" apply (simp only: word_arith_wis word_size uint_word_of_int_eq) apply (auto simp add: not_less take_bit_int_eq_self_iff) apply (metis not_less take_bit_eq_mod uint_plus_if' uint_word_ariths(1)) done lemma unat_plus_if_size: "unat (x + y) = (if unat x + unat y < 2^size x then unat x + unat y else unat x + unat y - 2^size x)" for x y :: "'a::len word" apply (subst word_arith_nat_defs) apply (subst unat_of_nat) apply (auto simp add: not_less word_size) apply (metis not_le unat_plus_if' unat_word_ariths(1)) done lemma word_neq_0_conv: "w \ 0 \ 0 < w" for w :: "'a::len word" by (fact word_coorder.not_eq_extremum) lemma max_lt: "unat (max a b div c) = unat (max a b) div unat c" for c :: "'a::len word" by (fact unat_div) lemma uint_sub_if_size: "uint (x - y) = (if uint y \ uint x then uint x - uint y else uint x - uint y + 2^size x)" apply (simp only: word_arith_wis word_size uint_word_of_int_eq) apply (auto simp add: take_bit_int_eq_self_iff not_le) apply (metis not_less uint_sub_if' uint_word_arith_bintrs(2)) done lemma unat_sub: \unat (a - b) = unat a - unat b\ if \b \ a\ proof - from that have \unat b \ unat a\ by transfer simp with that show ?thesis apply transfer apply simp apply (subst take_bit_diff [symmetric]) apply (subst nat_take_bit_eq) apply (simp add: nat_le_eq_zle) apply (simp add: nat_diff_distrib take_bit_nat_eq_self_iff less_imp_diff_less bintr_lt2p) done qed lemmas word_less_sub1_numberof [simp] = word_less_sub1 [of "numeral w"] for w lemmas word_le_sub1_numberof [simp] = word_le_sub1 [of "numeral w"] for w lemma word_of_int_minus: "word_of_int (2^LENGTH('a) - i) = (word_of_int (-i)::'a::len word)" apply transfer apply (subst take_bit_diff [symmetric]) apply (simp add: take_bit_minus) done lemma word_of_int_inj: \(word_of_int x :: 'a::len word) = word_of_int y \ x = y\ if \0 \ x \ x < 2 ^ LENGTH('a)\ \0 \ y \ y < 2 ^ LENGTH('a)\ using that by (transfer fixing: x y) (simp add: take_bit_int_eq_self) lemma word_le_less_eq: "x \ y \ x = y \ x < y" for x y :: "'z::len word" by (auto simp add: order_class.le_less) lemma mod_plus_cong: fixes b b' :: int assumes 1: "b = b'" and 2: "x mod b' = x' mod b'" and 3: "y mod b' = y' mod b'" and 4: "x' + y' = z'" shows "(x + y) mod b = z' mod b'" proof - from 1 2[symmetric] 3[symmetric] have "(x + y) mod b = (x' mod b' + y' mod b') mod b'" by (simp add: mod_add_eq) also have "\ = (x' + y') mod b'" by (simp add: mod_add_eq) finally show ?thesis by (simp add: 4) qed lemma mod_minus_cong: fixes b b' :: int assumes "b = b'" and "x mod b' = x' mod b'" and "y mod b' = y' mod b'" and "x' - y' = z'" shows "(x - y) mod b = z' mod b'" using assms [symmetric] by (auto intro: mod_diff_cong) lemma word_induct_less: \P m\ if zero: \P 0\ and less: \\n. n < m \ P n \ P (1 + n)\ for m :: \'a::len word\ proof - define q where \q = unat m\ with less have \\n. n < word_of_nat q \ P n \ P (1 + n)\ by simp then have \P (word_of_nat q :: 'a word)\ proof (induction q) case 0 show ?case by (simp add: zero) next case (Suc q) show ?case proof (cases \1 + word_of_nat q = (0 :: 'a word)\) case True then show ?thesis by (simp add: zero) next case False then have *: \word_of_nat q < (word_of_nat (Suc q) :: 'a word)\ by (simp add: unatSuc word_less_nat_alt) then have **: \n < (1 + word_of_nat q :: 'a word) \ n \ (word_of_nat q :: 'a word)\ for n by (metis (no_types, lifting) add.commute inc_le le_less_trans not_less of_nat_Suc) have \P (word_of_nat q)\ apply (rule Suc.IH) apply (rule Suc.prems) apply (erule less_trans) apply (rule *) apply assumption done with * have \P (1 + word_of_nat q)\ by (rule Suc.prems) then show ?thesis by simp qed qed with \q = unat m\ show ?thesis by simp qed lemma word_induct: "P 0 \ (\n. P n \ P (1 + n)) \ P m" for P :: "'a::len word \ bool" by (rule word_induct_less) lemma word_induct2 [induct type]: "P 0 \ (\n. 1 + n \ 0 \ P n \ P (1 + n)) \ P n" for P :: "'b::len word \ bool" apply (rule word_induct_less) apply simp_all apply (case_tac "1 + na = 0") apply auto done subsection \Recursion combinator for words\ definition word_rec :: "'a \ ('b::len word \ 'a \ 'a) \ 'b word \ 'a" where "word_rec forZero forSuc n = rec_nat forZero (forSuc \ of_nat) (unat n)" lemma word_rec_0: "word_rec z s 0 = z" by (simp add: word_rec_def) lemma word_rec_Suc: "1 + n \ 0 \ word_rec z s (1 + n) = s n (word_rec z s n)" for n :: "'a::len word" apply (auto simp add: word_rec_def unat_word_ariths) apply (metis (mono_tags, lifting) Abs_fnat_hom_add add_diff_cancel_left' o_def of_nat_1 old.nat.simps(7) plus_1_eq_Suc unatSuc unat_word_ariths(1) unsigned_1 word_arith_nat_add) done lemma word_rec_Pred: "n \ 0 \ word_rec z s n = s (n - 1) (word_rec z s (n - 1))" apply (rule subst[where t="n" and s="1 + (n - 1)"]) apply simp apply (subst word_rec_Suc) apply simp apply simp done lemma word_rec_in: "f (word_rec z (\_. f) n) = word_rec (f z) (\_. f) n" by (induct n) (simp_all add: word_rec_0 word_rec_Suc) lemma word_rec_in2: "f n (word_rec z f n) = word_rec (f 0 z) (f \ (+) 1) n" by (induct n) (simp_all add: word_rec_0 word_rec_Suc) lemma word_rec_twice: "m \ n \ word_rec z f n = word_rec (word_rec z f (n - m)) (f \ (+) (n - m)) m" apply (erule rev_mp) apply (rule_tac x=z in spec) apply (rule_tac x=f in spec) apply (induct n) apply (simp add: word_rec_0) apply clarsimp apply (rule_tac t="1 + n - m" and s="1 + (n - m)" in subst) apply simp apply (case_tac "1 + (n - m) = 0") apply (simp add: word_rec_0) apply (rule_tac f = "word_rec a b" for a b in arg_cong) apply (rule_tac t="m" and s="m + (1 + (n - m))" in subst) apply simp apply (simp (no_asm_use)) apply (simp add: word_rec_Suc word_rec_in2) apply (erule impE) apply uint_arith apply (drule_tac x="x \ (+) 1" in spec) apply (drule_tac x="x 0 xa" in spec) apply simp apply (rule_tac t="\a. x (1 + (n - m + a))" and s="\a. x (1 + (n - m) + a)" in subst) apply (clarsimp simp add: fun_eq_iff) apply (rule_tac t="(1 + (n - m + xb))" and s="1 + (n - m) + xb" in subst) apply simp apply (rule refl) apply (rule refl) done lemma word_rec_id: "word_rec z (\_. id) n = z" by (induct n) (auto simp add: word_rec_0 word_rec_Suc) lemma word_rec_id_eq: "\m < n. f m = id \ word_rec z f n = z" apply (erule rev_mp) apply (induct n) apply (auto simp add: word_rec_0 word_rec_Suc) apply (drule spec, erule mp) apply uint_arith apply (drule_tac x=n in spec, erule impE) apply uint_arith apply simp done lemma word_rec_max: "\m\n. m \ - 1 \ f m = id \ word_rec z f (- 1) = word_rec z f n" apply (subst word_rec_twice[where n="-1" and m="-1 - n"]) apply simp apply simp apply (rule word_rec_id_eq) apply clarsimp apply (drule spec, rule mp, erule mp) apply (rule word_plus_mono_right2[OF _ order_less_imp_le]) prefer 2 apply assumption apply simp apply (erule contrapos_pn) apply simp apply (drule arg_cong[where f="\x. x - n"]) apply simp done subsection \More\ 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) lemma mask_1: "mask 1 = 1" by transfer (simp add: min_def mask_Suc_exp) lemma mask_Suc_0: "mask (Suc 0) = 1" using mask_1 by simp lemma bin_last_bintrunc: "bin_last (take_bit l n) = (l > 0 \ bin_last n)" by simp 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 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) subsection \Misc\ ML_file \Tools/word_lib.ML\ ML_file \Tools/smt_word.ML\ end diff --git a/src/HOL/Word/Word_rsplit.thy b/src/HOL/Word/Word_rsplit.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Word/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 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