# HG changeset patch # User haftmann # Date 1610010864 0 # Node ID dc62ecc7e59ae2fb017fca91223e875071dcc074 # Parent 131ab1a941dd7bf76c6684a13b897a414070c875 dropped junk diff -r 131ab1a941dd -r dc62ecc7e59a src/HOL/Word/Bits_Int.thy --- a/src/HOL/Word/Bits_Int.thy Thu Jan 07 00:04:13 2021 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1406 +0,0 @@ -(* 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 - Word -begin - -subsection \Implicit bit representation of \<^typ>\int\\ - -abbreviation (input) bin_last :: "int \ bool" - where "bin_last \ odd" - -lemma bin_last_def: - "bin_last w \ w mod 2 = 1" - by (fact odd_iff_mod_2_eq_one) - -abbreviation (input) bin_rest :: "int \ int" - where "bin_rest w \ w div 2" - -lemma bin_last_numeral_simps [simp]: - "\ bin_last 0" - "bin_last 1" - "bin_last (- 1)" - "bin_last Numeral1" - "\ bin_last (numeral (Num.Bit0 w))" - "bin_last (numeral (Num.Bit1 w))" - "\ bin_last (- numeral (Num.Bit0 w))" - "bin_last (- numeral (Num.Bit1 w))" - by simp_all - -lemma bin_rest_numeral_simps [simp]: - "bin_rest 0 = 0" - "bin_rest 1 = 0" - "bin_rest (- 1) = - 1" - "bin_rest Numeral1 = 0" - "bin_rest (numeral (Num.Bit0 w)) = numeral w" - "bin_rest (numeral (Num.Bit1 w)) = numeral w" - "bin_rest (- numeral (Num.Bit0 w)) = - numeral w" - "bin_rest (- numeral (Num.Bit1 w)) = - numeral (w + Num.One)" - by simp_all - -lemma bin_rl_eqI: "\bin_rest x = bin_rest y; bin_last x = bin_last y\ \ x = y" - by (auto elim: oddE) - -lemma [simp]: - shows bin_rest_lt0: "bin_rest i < 0 \ i < 0" - and bin_rest_ge_0: "bin_rest i \ 0 \ i \ 0" - by auto - -lemma bin_rest_gt_0 [simp]: "bin_rest x > 0 \ x > 1" - by auto - - -subsection \Bit projection\ - -abbreviation (input) bin_nth :: \int \ nat \ bool\ - where \bin_nth \ bit\ - -lemma bin_nth_eq_iff: "bin_nth x = bin_nth y \ x = y" - by (simp add: bit_eq_iff fun_eq_iff) - -lemma bin_eqI: - "x = y" if "\n. bin_nth x n \ bin_nth y n" - using that bin_nth_eq_iff [of x y] by (simp add: fun_eq_iff) - -lemma bin_eq_iff: "x = y \ (\n. bin_nth x n = bin_nth y n)" - by (fact bit_eq_iff) - -lemma bin_nth_zero [simp]: "\ bin_nth 0 n" - by simp - -lemma bin_nth_1 [simp]: "bin_nth 1 n \ n = 0" - by (cases n) (simp_all add: bit_Suc) - -lemma bin_nth_minus1 [simp]: "bin_nth (- 1) n" - by (induction n) (simp_all add: bit_Suc) - -lemma bin_nth_numeral: "bin_rest x = y \ bin_nth x (numeral n) = bin_nth y (pred_numeral n)" - by (simp add: numeral_eq_Suc bit_Suc) - -lemmas bin_nth_numeral_simps [simp] = - bin_nth_numeral [OF bin_rest_numeral_simps(2)] - bin_nth_numeral [OF bin_rest_numeral_simps(5)] - bin_nth_numeral [OF bin_rest_numeral_simps(6)] - bin_nth_numeral [OF bin_rest_numeral_simps(7)] - bin_nth_numeral [OF bin_rest_numeral_simps(8)] - -lemmas bin_nth_simps = - bit_0 bit_Suc bin_nth_zero bin_nth_minus1 - bin_nth_numeral_simps - -lemma nth_2p_bin: "bin_nth (2 ^ n) m = (m = n)" \ \for use when simplifying with \bin_nth_Bit\\ - by (auto simp add: bit_exp_iff) - -lemma nth_rest_power_bin: "bin_nth ((bin_rest ^^ k) w) n = bin_nth w (n + k)" - apply (induct k arbitrary: n) - apply clarsimp - apply clarsimp - apply (simp only: bit_Suc [symmetric] add_Suc) - done - -lemma bin_nth_numeral_unfold: - "bin_nth (numeral (num.Bit0 x)) n \ n > 0 \ bin_nth (numeral x) (n - 1)" - "bin_nth (numeral (num.Bit1 x)) n \ (n > 0 \ bin_nth (numeral x) (n - 1))" - by (cases n; simp)+ - - -subsection \Truncating\ - -definition bin_sign :: "int \ int" - where "bin_sign k = (if k \ 0 then 0 else - 1)" - -lemma bin_sign_simps [simp]: - "bin_sign 0 = 0" - "bin_sign 1 = 0" - "bin_sign (- 1) = - 1" - "bin_sign (numeral k) = 0" - "bin_sign (- numeral k) = -1" - by (simp_all add: bin_sign_def) - -lemma bin_sign_rest [simp]: "bin_sign (bin_rest w) = bin_sign w" - by (simp add: bin_sign_def) - -abbreviation (input) bintrunc :: \nat \ int \ int\ - where \bintrunc \ take_bit\ - -lemma bintrunc_mod2p: "bintrunc n w = w mod 2 ^ n" - by (fact take_bit_eq_mod) - -abbreviation (input) sbintrunc :: \nat \ int \ int\ - where \sbintrunc \ signed_take_bit\ - -abbreviation (input) norm_sint :: \nat \ int \ int\ - where \norm_sint n \ signed_take_bit (n - 1)\ - -lemma sbintrunc_mod2p: "sbintrunc n w = (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n" - by (simp add: bintrunc_mod2p signed_take_bit_eq_take_bit_shift) - -lemma sbintrunc_eq_take_bit: - \sbintrunc n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\ - by (fact signed_take_bit_eq_take_bit_shift) - -lemma sign_bintr: "bin_sign (bintrunc n w) = 0" - by (simp add: bin_sign_def) - -lemma bintrunc_n_0: "bintrunc n 0 = 0" - by (fact take_bit_of_0) - -lemma sbintrunc_n_0: "sbintrunc n 0 = 0" - by (fact signed_take_bit_of_0) - -lemma sbintrunc_n_minus1: "sbintrunc n (- 1) = -1" - by (fact signed_take_bit_of_minus_1) - -lemma bintrunc_Suc_numeral: - "bintrunc (Suc n) 1 = 1" - "bintrunc (Suc n) (- 1) = 1 + 2 * bintrunc n (- 1)" - "bintrunc (Suc n) (numeral (Num.Bit0 w)) = 2 * bintrunc n (numeral w)" - "bintrunc (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * bintrunc n (numeral w)" - "bintrunc (Suc n) (- numeral (Num.Bit0 w)) = 2 * bintrunc n (- numeral w)" - "bintrunc (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * bintrunc n (- numeral (w + Num.One))" - by (simp_all add: take_bit_Suc) - -lemma sbintrunc_0_numeral [simp]: - "sbintrunc 0 1 = -1" - "sbintrunc 0 (numeral (Num.Bit0 w)) = 0" - "sbintrunc 0 (numeral (Num.Bit1 w)) = -1" - "sbintrunc 0 (- numeral (Num.Bit0 w)) = 0" - "sbintrunc 0 (- numeral (Num.Bit1 w)) = -1" - by simp_all - -lemma sbintrunc_Suc_numeral: - "sbintrunc (Suc n) 1 = 1" - "sbintrunc (Suc n) (numeral (Num.Bit0 w)) = 2 * sbintrunc n (numeral w)" - "sbintrunc (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * sbintrunc n (numeral w)" - "sbintrunc (Suc n) (- numeral (Num.Bit0 w)) = 2 * sbintrunc n (- numeral w)" - "sbintrunc (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * sbintrunc n (- numeral (w + Num.One))" - by (simp_all add: signed_take_bit_Suc) - -lemma bin_sign_lem: "(bin_sign (sbintrunc n bin) = -1) = bit bin n" - by (simp add: bin_sign_def) - -lemma nth_bintr: "bin_nth (bintrunc m w) n \ n < m \ bin_nth w n" - by (fact bit_take_bit_iff) - -lemma nth_sbintr: "bin_nth (sbintrunc m w) n = (if n < m then bin_nth w n else bin_nth w m)" - by (simp add: bit_signed_take_bit_iff min_def) - -lemma bin_nth_Bit0: - "bin_nth (numeral (Num.Bit0 w)) n \ - (\m. n = Suc m \ bin_nth (numeral w) m)" - using bit_double_iff [of \numeral w :: int\ n] - by (auto intro: exI [of _ \n - 1\]) - -lemma bin_nth_Bit1: - "bin_nth (numeral (Num.Bit1 w)) n \ - n = 0 \ (\m. n = Suc m \ bin_nth (numeral w) m)" - using even_bit_succ_iff [of \2 * numeral w :: int\ n] - bit_double_iff [of \numeral w :: int\ n] - by auto - -lemma bintrunc_bintrunc_l: "n \ m \ bintrunc m (bintrunc n w) = bintrunc n w" - by (simp add: min.absorb2) - -lemma sbintrunc_sbintrunc_l: "n \ m \ sbintrunc m (sbintrunc n w) = sbintrunc n w" - by (simp add: min_def) - -lemma bintrunc_bintrunc_ge: "n \ m \ bintrunc n (bintrunc m w) = bintrunc n w" - by (rule bin_eqI) (auto simp: nth_bintr) - -lemma bintrunc_bintrunc_min [simp]: "bintrunc m (bintrunc n w) = bintrunc (min m n) w" - by (rule bin_eqI) (auto simp: nth_bintr) - -lemma sbintrunc_sbintrunc_min [simp]: "sbintrunc m (sbintrunc n w) = sbintrunc (min m n) w" - by (rule bin_eqI) (auto simp: nth_sbintr min.absorb1 min.absorb2) - -lemmas sbintrunc_Suc_Pls = - signed_take_bit_Suc [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] - -lemmas sbintrunc_Suc_Min = - signed_take_bit_Suc [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] - -lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min - sbintrunc_Suc_numeral - -lemmas sbintrunc_Pls = - signed_take_bit_0 [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] - -lemmas sbintrunc_Min = - signed_take_bit_0 [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] - -lemmas sbintrunc_0_simps = - sbintrunc_Pls sbintrunc_Min - -lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs - -lemma bintrunc_minus: "0 < n \ bintrunc (Suc (n - 1)) w = bintrunc n w" - by auto - -lemma sbintrunc_minus: "0 < n \ sbintrunc (Suc (n - 1)) w = sbintrunc n w" - by auto - -lemmas sbintrunc_minus_simps = - sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans]] - -lemma sbintrunc_BIT_I: - \0 < n \ - sbintrunc (n - 1) 0 = y \ - sbintrunc n 0 = 2 * y\ - by simp - -lemma sbintrunc_Suc_Is: - \sbintrunc n (- 1) = y \ - sbintrunc (Suc n) (- 1) = 1 + 2 * y\ - by auto - -lemma sbintrunc_Suc_lem: "sbintrunc (Suc n) x = y \ m = Suc n \ sbintrunc m x = y" - by auto - -lemmas sbintrunc_Suc_Ialts = - sbintrunc_Suc_Is [THEN sbintrunc_Suc_lem] - -lemma sbintrunc_bintrunc_lt: "m > n \ sbintrunc n (bintrunc m w) = sbintrunc n w" - by (rule bin_eqI) (auto simp: nth_sbintr nth_bintr) - -lemma bintrunc_sbintrunc_le: "m \ Suc n \ bintrunc m (sbintrunc n w) = bintrunc m w" - apply (rule bin_eqI) - using le_Suc_eq less_Suc_eq_le apply (auto simp: nth_sbintr nth_bintr) - done - -lemmas bintrunc_sbintrunc [simp] = order_refl [THEN bintrunc_sbintrunc_le] -lemmas sbintrunc_bintrunc [simp] = lessI [THEN sbintrunc_bintrunc_lt] -lemmas bintrunc_bintrunc [simp] = order_refl [THEN bintrunc_bintrunc_l] -lemmas sbintrunc_sbintrunc [simp] = order_refl [THEN sbintrunc_sbintrunc_l] - -lemma bintrunc_sbintrunc' [simp]: "0 < n \ bintrunc n (sbintrunc (n - 1) w) = bintrunc n w" - by (cases n) simp_all - -lemma sbintrunc_bintrunc' [simp]: "0 < n \ sbintrunc (n - 1) (bintrunc n w) = sbintrunc (n - 1) w" - by (cases n) simp_all - -lemma bin_sbin_eq_iff: "bintrunc (Suc n) x = bintrunc (Suc n) y \ sbintrunc n x = sbintrunc n y" - apply (rule iffI) - apply (rule box_equals [OF _ sbintrunc_bintrunc sbintrunc_bintrunc]) - apply simp - apply (rule box_equals [OF _ bintrunc_sbintrunc bintrunc_sbintrunc]) - apply simp - done - -lemma bin_sbin_eq_iff': - "0 < n \ bintrunc n x = bintrunc n y \ sbintrunc (n - 1) x = sbintrunc (n - 1) y" - by (cases n) (simp_all add: bin_sbin_eq_iff) - -lemmas bintrunc_sbintruncS0 [simp] = bintrunc_sbintrunc' [unfolded One_nat_def] -lemmas sbintrunc_bintruncS0 [simp] = sbintrunc_bintrunc' [unfolded One_nat_def] - -lemmas bintrunc_bintrunc_l' = le_add1 [THEN bintrunc_bintrunc_l] -lemmas sbintrunc_sbintrunc_l' = le_add1 [THEN sbintrunc_sbintrunc_l] - -(* although bintrunc_minus_simps, if added to default simpset, - tends to get applied where it's not wanted in developing the theories, - we get a version for when the word length is given literally *) - -lemmas nat_non0_gr = - trans [OF iszero_def [THEN Not_eq_iff [THEN iffD2]] refl] - -lemma bintrunc_numeral: - "bintrunc (numeral k) x = of_bool (odd x) + 2 * bintrunc (pred_numeral k) (x div 2)" - by (simp add: numeral_eq_Suc take_bit_Suc mod_2_eq_odd) - -lemma sbintrunc_numeral: - "sbintrunc (numeral k) x = of_bool (odd x) + 2 * sbintrunc (pred_numeral k) (x div 2)" - by (simp add: numeral_eq_Suc signed_take_bit_Suc mod2_eq_if) - -lemma bintrunc_numeral_simps [simp]: - "bintrunc (numeral k) (numeral (Num.Bit0 w)) = - 2 * bintrunc (pred_numeral k) (numeral w)" - "bintrunc (numeral k) (numeral (Num.Bit1 w)) = - 1 + 2 * bintrunc (pred_numeral k) (numeral w)" - "bintrunc (numeral k) (- numeral (Num.Bit0 w)) = - 2 * bintrunc (pred_numeral k) (- numeral w)" - "bintrunc (numeral k) (- numeral (Num.Bit1 w)) = - 1 + 2 * bintrunc (pred_numeral k) (- numeral (w + Num.One))" - "bintrunc (numeral k) 1 = 1" - by (simp_all add: bintrunc_numeral) - -lemma sbintrunc_numeral_simps [simp]: - "sbintrunc (numeral k) (numeral (Num.Bit0 w)) = - 2 * sbintrunc (pred_numeral k) (numeral w)" - "sbintrunc (numeral k) (numeral (Num.Bit1 w)) = - 1 + 2 * sbintrunc (pred_numeral k) (numeral w)" - "sbintrunc (numeral k) (- numeral (Num.Bit0 w)) = - 2 * sbintrunc (pred_numeral k) (- numeral w)" - "sbintrunc (numeral k) (- numeral (Num.Bit1 w)) = - 1 + 2 * sbintrunc (pred_numeral k) (- numeral (w + Num.One))" - "sbintrunc (numeral k) 1 = 1" - by (simp_all add: sbintrunc_numeral) - -lemma no_bintr_alt1: "bintrunc n = (\w. w mod 2 ^ n :: int)" - by (rule ext) (rule bintrunc_mod2p) - -lemma range_bintrunc: "range (bintrunc n) = {i. 0 \ i \ i < 2 ^ n}" - by (auto simp add: take_bit_eq_mod image_iff) (metis mod_pos_pos_trivial) - -lemma no_sbintr_alt2: "sbintrunc n = (\w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" - by (rule ext) (simp add : sbintrunc_mod2p) - -lemma range_sbintrunc: "range (sbintrunc n) = {i. - (2 ^ n) \ i \ i < 2 ^ n}" -proof - - have \surj (\k::int. k + 2 ^ n)\ - by (rule surjI [of _ \(\k. k - 2 ^ n)\]) simp - moreover have \sbintrunc n = ((\k. k - 2 ^ n) \ take_bit (Suc n) \ (\k. k + 2 ^ n))\ - by (simp add: sbintrunc_eq_take_bit fun_eq_iff) - ultimately show ?thesis - apply (simp only: fun.set_map range_bintrunc) - apply (auto simp add: image_iff) - apply presburger - done -qed - -lemma sbintrunc_inc: - \k + 2 ^ Suc n \ sbintrunc n k\ if \k < - (2 ^ n)\ - using that by (fact signed_take_bit_int_greater_eq) - -lemma sbintrunc_dec: - \sbintrunc n k \ k - 2 ^ (Suc n)\ if \k \ 2 ^ n\ - using that by (fact signed_take_bit_int_less_eq) - -lemma bintr_ge0: "0 \ bintrunc n w" - by (simp add: bintrunc_mod2p) - -lemma bintr_lt2p: "bintrunc n w < 2 ^ n" - by (simp add: bintrunc_mod2p) - -lemma bintr_Min: "bintrunc n (- 1) = 2 ^ n - 1" - by (simp add: stable_imp_take_bit_eq) - -lemma sbintr_ge: "- (2 ^ n) \ sbintrunc n w" - by (simp add: sbintrunc_mod2p) - -lemma sbintr_lt: "sbintrunc n w < 2 ^ n" - by (simp add: sbintrunc_mod2p) - -lemma sign_Pls_ge_0: "bin_sign bin = 0 \ bin \ 0" - for bin :: int - by (simp add: bin_sign_def) - -lemma sign_Min_lt_0: "bin_sign bin = -1 \ bin < 0" - for bin :: int - by (simp add: bin_sign_def) - -lemma bin_rest_trunc: "bin_rest (bintrunc n bin) = bintrunc (n - 1) (bin_rest bin)" - by (simp add: take_bit_rec [of n bin]) - -lemma bin_rest_power_trunc: - "(bin_rest ^^ k) (bintrunc n bin) = bintrunc (n - k) ((bin_rest ^^ k) bin)" - by (induct k) (auto simp: bin_rest_trunc) - -lemma bin_rest_trunc_i: "bintrunc n (bin_rest bin) = bin_rest (bintrunc (Suc n) bin)" - by (auto simp add: take_bit_Suc) - -lemma bin_rest_strunc: "bin_rest (sbintrunc (Suc n) bin) = sbintrunc n (bin_rest bin)" - by (simp add: signed_take_bit_Suc) - -lemma bintrunc_rest [simp]: "bintrunc n (bin_rest (bintrunc n bin)) = bin_rest (bintrunc n bin)" - by (induct n arbitrary: bin) (simp_all add: take_bit_Suc) - -lemma sbintrunc_rest [simp]: "sbintrunc n (bin_rest (sbintrunc n bin)) = bin_rest (sbintrunc n bin)" - by (induct n arbitrary: bin) (simp_all add: signed_take_bit_Suc mod2_eq_if) - -lemma bintrunc_rest': "bintrunc n \ bin_rest \ bintrunc n = bin_rest \ bintrunc n" - by (rule ext) auto - -lemma sbintrunc_rest': "sbintrunc n \ bin_rest \ sbintrunc n = bin_rest \ sbintrunc n" - by (rule ext) auto - -lemma rco_lem: "f \ g \ f = g \ f \ f \ (g \ f) ^^ n = g ^^ n \ f" - apply (rule ext) - apply (induct_tac n) - apply (simp_all (no_asm)) - apply (drule fun_cong) - apply (unfold o_def) - apply (erule trans) - apply simp - done - -lemmas rco_bintr = bintrunc_rest' - [THEN rco_lem [THEN fun_cong], unfolded o_def] -lemmas rco_sbintr = sbintrunc_rest' - [THEN rco_lem [THEN fun_cong], unfolded o_def] - - -subsection \Splitting and concatenation\ - -definition bin_split :: \nat \ int \ int \ int\ - where [simp]: \bin_split n k = (drop_bit n k, take_bit n k)\ - -lemma [code]: - "bin_split (Suc n) w = (let (w1, w2) = bin_split n (w div 2) in (w1, of_bool (odd w) + 2 * w2))" - "bin_split 0 w = (w, 0)" - by (simp_all add: drop_bit_Suc take_bit_Suc mod_2_eq_odd) - -abbreviation (input) bin_cat :: \int \ nat \ int \ int\ - where \bin_cat k n l \ concat_bit n l k\ - -lemma bin_cat_eq_push_bit_add_take_bit: - \bin_cat k n l = push_bit n k + take_bit n l\ - by (simp add: concat_bit_eq) - -lemma bin_sign_cat: "bin_sign (bin_cat x n y) = bin_sign x" -proof - - have \0 \ x\ if \0 \ x * 2 ^ n + y mod 2 ^ n\ - proof - - have \y mod 2 ^ n < 2 ^ n\ - using pos_mod_bound [of \2 ^ n\ y] by simp - then have \\ y mod 2 ^ n \ 2 ^ n\ - by (simp add: less_le) - with that have \x \ - 1\ - by auto - have *: \- 1 \ (- (y mod 2 ^ n)) div 2 ^ n\ - by (simp add: zdiv_zminus1_eq_if) - from that have \- (y mod 2 ^ n) \ x * 2 ^ n\ - by simp - then have \(- (y mod 2 ^ n)) div 2 ^ n \ (x * 2 ^ n) div 2 ^ n\ - using zdiv_mono1 zero_less_numeral zero_less_power by blast - with * have \- 1 \ x * 2 ^ n div 2 ^ n\ by simp - with \x \ - 1\ show ?thesis - by simp - qed - then show ?thesis - by (simp add: bin_sign_def not_le not_less bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult take_bit_eq_mod) -qed - -lemma bin_cat_assoc: "bin_cat (bin_cat x m y) n z = bin_cat x (m + n) (bin_cat y n z)" - by (fact concat_bit_assoc) - -lemma bin_cat_assoc_sym: "bin_cat x m (bin_cat y n z) = bin_cat (bin_cat x (m - n) y) (min m n) z" - by (fact concat_bit_assoc_sym) - -definition bin_rcat :: \nat \ int list \ int\ - where \bin_rcat n = horner_sum (take_bit n) (2 ^ n) \ rev\ - -lemma bin_rcat_eq_foldl: - \bin_rcat n = foldl (\u v. bin_cat u n v) 0\ -proof - fix ks :: \int list\ - show \bin_rcat n ks = foldl (\u v. bin_cat u n v) 0 ks\ - by (induction ks rule: rev_induct) - (simp_all add: bin_rcat_def concat_bit_eq push_bit_eq_mult) -qed - -fun bin_rsplit_aux :: "nat \ nat \ int \ int list \ int list" - where "bin_rsplit_aux n m c bs = - (if m = 0 \ n = 0 then bs - else - let (a, b) = bin_split n c - in bin_rsplit_aux n (m - n) a (b # bs))" - -definition bin_rsplit :: "nat \ nat \ int \ int list" - where "bin_rsplit n w = bin_rsplit_aux n (fst w) (snd w) []" - -value \bin_rsplit 1705 (3, 88)\ - -fun bin_rsplitl_aux :: "nat \ nat \ int \ int list \ int list" - where "bin_rsplitl_aux n m c bs = - (if m = 0 \ n = 0 then bs - else - let (a, b) = bin_split (min m n) c - in bin_rsplitl_aux n (m - n) a (b # bs))" - -definition bin_rsplitl :: "nat \ nat \ int \ int list" - where "bin_rsplitl n w = bin_rsplitl_aux n (fst w) (snd w) []" - -declare bin_rsplit_aux.simps [simp del] -declare bin_rsplitl_aux.simps [simp del] - -lemma bin_nth_cat: - "bin_nth (bin_cat x k y) n = - (if n < k then bin_nth y n else bin_nth x (n - k))" - by (simp add: bit_concat_bit_iff) - -lemma bin_nth_drop_bit_iff: - \bin_nth (drop_bit n c) k \ bin_nth c (n + k)\ - by (simp add: bit_drop_bit_eq) - -lemma bin_nth_take_bit_iff: - \bin_nth (take_bit n c) k \ k < n \ bin_nth c k\ - by (fact bit_take_bit_iff) - -lemma bin_nth_split: - "bin_split n c = (a, b) \ - (\k. bin_nth a k = bin_nth c (n + k)) \ - (\k. bin_nth b k = (k < n \ bin_nth c k))" - by (auto simp add: bin_nth_drop_bit_iff bin_nth_take_bit_iff) - -lemma bin_cat_zero [simp]: "bin_cat 0 n w = bintrunc n w" - by (simp add: bin_cat_eq_push_bit_add_take_bit) - -lemma bintr_cat1: "bintrunc (k + n) (bin_cat a n b) = bin_cat (bintrunc k a) n b" - by (metis bin_cat_assoc bin_cat_zero) - -lemma bintr_cat: "bintrunc m (bin_cat a n b) = - bin_cat (bintrunc (m - n) a) n (bintrunc (min m n) b)" - - by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr) - -lemma bintr_cat_same [simp]: "bintrunc n (bin_cat a n b) = bintrunc n b" - by (auto simp add : bintr_cat) - -lemma cat_bintr [simp]: "bin_cat a n (bintrunc n b) = bin_cat a n b" - by (simp add: bin_cat_eq_push_bit_add_take_bit) - -lemma split_bintrunc: "bin_split n c = (a, b) \ b = bintrunc n c" - by simp - -lemma bin_cat_split: "bin_split n w = (u, v) \ w = bin_cat u n v" - by (auto simp add: bin_cat_eq_push_bit_add_take_bit bits_ident) - -lemma drop_bit_bin_cat_eq: - \drop_bit n (bin_cat v n w) = v\ - by (rule bit_eqI) (simp add: bit_drop_bit_eq bit_concat_bit_iff) - -lemma take_bit_bin_cat_eq: - \take_bit n (bin_cat v n w) = take_bit n w\ - by (rule bit_eqI) (simp add: bit_concat_bit_iff) - -lemma bin_split_cat: "bin_split n (bin_cat v n w) = (v, bintrunc n w)" - by (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) - -lemma bin_split_zero [simp]: "bin_split n 0 = (0, 0)" - by simp - -lemma bin_split_minus1 [simp]: - "bin_split n (- 1) = (- 1, bintrunc n (- 1))" - by simp - -lemma bin_split_trunc: - "bin_split (min m n) c = (a, b) \ - bin_split n (bintrunc m c) = (bintrunc (m - n) a, b)" - apply (induct n arbitrary: m b c, clarsimp) - apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) - apply (case_tac m) - apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) - done - -lemma bin_split_trunc1: - "bin_split n c = (a, b) \ - bin_split n (bintrunc m c) = (bintrunc (m - n) a, bintrunc m b)" - apply (induct n arbitrary: m b c, clarsimp) - apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) - apply (case_tac m) - apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) - done - -lemma bin_cat_num: "bin_cat a n b = a * 2 ^ n + bintrunc n b" - by (simp add: bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult) - -lemma bin_split_num: "bin_split n b = (b div 2 ^ n, b mod 2 ^ n)" - by (simp add: drop_bit_eq_div take_bit_eq_mod) - -lemmas bin_rsplit_aux_simps = bin_rsplit_aux.simps bin_rsplitl_aux.simps -lemmas rsplit_aux_simps = bin_rsplit_aux_simps - -lemmas th_if_simp1 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct1, THEN mp] for l -lemmas th_if_simp2 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct2, THEN mp] for l - -lemmas rsplit_aux_simp1s = rsplit_aux_simps [THEN th_if_simp1] - -lemmas rsplit_aux_simp2ls = rsplit_aux_simps [THEN th_if_simp2] -\ \these safe to \[simp add]\ as require calculating \m - n\\ -lemmas bin_rsplit_aux_simp2s [simp] = rsplit_aux_simp2ls [unfolded Let_def] -lemmas rbscl = bin_rsplit_aux_simp2s (2) - -lemmas rsplit_aux_0_simps [simp] = - rsplit_aux_simp1s [OF disjI1] rsplit_aux_simp1s [OF disjI2] - -lemma bin_rsplit_aux_append: "bin_rsplit_aux n m c (bs @ cs) = bin_rsplit_aux n m c bs @ cs" - apply (induct n m c bs rule: bin_rsplit_aux.induct) - apply (subst bin_rsplit_aux.simps) - apply (subst bin_rsplit_aux.simps) - apply (clarsimp split: prod.split) - done - -lemma bin_rsplitl_aux_append: "bin_rsplitl_aux n m c (bs @ cs) = bin_rsplitl_aux n m c bs @ cs" - apply (induct n m c bs rule: bin_rsplitl_aux.induct) - apply (subst bin_rsplitl_aux.simps) - apply (subst bin_rsplitl_aux.simps) - apply (clarsimp split: prod.split) - done - -lemmas rsplit_aux_apps [where bs = "[]"] = - bin_rsplit_aux_append bin_rsplitl_aux_append - -lemmas rsplit_def_auxs = bin_rsplit_def bin_rsplitl_def - -lemmas rsplit_aux_alts = rsplit_aux_apps - [unfolded append_Nil rsplit_def_auxs [symmetric]] - -lemma bin_split_minus: "0 < n \ bin_split (Suc (n - 1)) w = bin_split n w" - by auto - -lemma bin_split_pred_simp [simp]: - "(0::nat) < numeral bin \ - bin_split (numeral bin) w = - (let (w1, w2) = bin_split (numeral bin - 1) (bin_rest w) - in (w1, of_bool (odd w) + 2 * w2))" - by (simp add: take_bit_rec drop_bit_rec mod_2_eq_odd) - -lemma bin_rsplit_aux_simp_alt: - "bin_rsplit_aux n m c bs = - (if m = 0 \ n = 0 then bs - else let (a, b) = bin_split n c in bin_rsplit n (m - n, a) @ b # bs)" - apply (simp add: bin_rsplit_aux.simps [of n m c bs]) - apply (subst rsplit_aux_alts) - apply (simp add: bin_rsplit_def) - done - -lemmas bin_rsplit_simp_alt = - trans [OF bin_rsplit_def bin_rsplit_aux_simp_alt] - -lemmas bthrs = bin_rsplit_simp_alt [THEN [2] trans] - -lemma bin_rsplit_size_sign' [rule_format]: - "n > 0 \ rev sw = bin_rsplit n (nw, w) \ \v\set sw. bintrunc n v = v" - apply (induct sw arbitrary: nw w) - apply clarsimp - apply clarsimp - apply (drule bthrs) - apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) - apply clarify - apply simp - done - -lemmas bin_rsplit_size_sign = bin_rsplit_size_sign' [OF asm_rl - rev_rev_ident [THEN trans] set_rev [THEN equalityD2 [THEN subsetD]]] - -lemma bin_nth_rsplit [rule_format] : - "n > 0 \ m < n \ - \w k nw. - rev sw = bin_rsplit n (nw, w) \ - k < size sw \ bin_nth (sw ! k) m = bin_nth w (k * n + m)" - apply (induct sw) - apply clarsimp - apply clarsimp - apply (drule bthrs) - apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) - apply (erule allE, erule impE, erule exI) - apply (case_tac k) - apply clarsimp - prefer 2 - apply clarsimp - apply (erule allE) - apply (erule (1) impE) - apply (simp add: bit_drop_bit_eq ac_simps) - apply (simp add: bit_take_bit_iff ac_simps) - done - -lemma bin_rsplit_all: "0 < nw \ nw \ n \ bin_rsplit n (nw, w) = [bintrunc n w]" - by (auto simp: bin_rsplit_def rsplit_aux_simp2ls split: prod.split dest!: split_bintrunc) - -lemma bin_rsplit_l [rule_format]: - "\bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, bintrunc m bin)" - apply (rule_tac a = "m" in wf_less_than [THEN wf_induct]) - apply (simp (no_asm) add: bin_rsplitl_def bin_rsplit_def) - apply (rule allI) - apply (subst bin_rsplitl_aux.simps) - apply (subst bin_rsplit_aux.simps) - apply (clarsimp simp: Let_def split: prod.split) - apply (simp add: ac_simps) - apply (subst rsplit_aux_alts(1)) - apply (subst rsplit_aux_alts(2)) - apply clarsimp - unfolding bin_rsplit_def bin_rsplitl_def - apply (simp add: drop_bit_take_bit) - apply (case_tac \x < n\) - apply (simp_all add: not_less min_def) - done - -lemma bin_rsplit_rcat [rule_format]: - "n > 0 \ bin_rsplit n (n * size ws, bin_rcat n ws) = map (bintrunc n) ws" - apply (unfold bin_rsplit_def bin_rcat_eq_foldl) - apply (rule_tac xs = ws in rev_induct) - apply clarsimp - apply clarsimp - apply (subst rsplit_aux_alts) - apply (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) - done - -lemma bin_rsplit_aux_len_le [rule_format] : - "\ws m. n \ 0 \ ws = bin_rsplit_aux n nw w bs \ - length ws \ m \ nw + length bs * n \ m * n" -proof - - have *: R - if d: "i \ j \ m < j'" - and R1: "i * k \ j * k \ R" - and R2: "Suc m * k' \ j' * k' \ R" - for i j j' k k' m :: nat and R - using d - apply safe - apply (rule R1, erule mult_le_mono1) - apply (rule R2, erule Suc_le_eq [THEN iffD2 [THEN mult_le_mono1]]) - done - have **: "0 < sc \ sc - n + (n + lb * n) \ m * n \ sc + lb * n \ m * n" - for sc m n lb :: nat - apply safe - apply arith - apply (case_tac "sc \ n") - apply arith - apply (insert linorder_le_less_linear [of m lb]) - apply (erule_tac k=n and k'=n in *) - apply arith - apply simp - done - show ?thesis - apply (induct n nw w bs rule: bin_rsplit_aux.induct) - apply (subst bin_rsplit_aux.simps) - apply (simp add: ** Let_def split: prod.split) - done -qed - -lemma bin_rsplit_len_le: "n \ 0 \ ws = bin_rsplit n (nw, w) \ length ws \ m \ nw \ m * n" - by (auto simp: bin_rsplit_def bin_rsplit_aux_len_le) - -lemma bin_rsplit_aux_len: - "n \ 0 \ length (bin_rsplit_aux n nw w cs) = (nw + n - 1) div n + length cs" - apply (induct n nw w cs rule: bin_rsplit_aux.induct) - apply (subst bin_rsplit_aux.simps) - apply (clarsimp simp: Let_def split: prod.split) - apply (erule thin_rl) - apply (case_tac m) - apply simp - apply (case_tac "m \ n") - apply (auto simp add: div_add_self2) - done - -lemma bin_rsplit_len: "n \ 0 \ length (bin_rsplit n (nw, w)) = (nw + n - 1) div n" - by (auto simp: bin_rsplit_def bin_rsplit_aux_len) - -lemma bin_rsplit_aux_len_indep: - "n \ 0 \ length bs = length cs \ - length (bin_rsplit_aux n nw v bs) = - length (bin_rsplit_aux n nw w cs)" -proof (induct n nw w cs arbitrary: v bs rule: bin_rsplit_aux.induct) - case (1 n m w cs v bs) - show ?case - proof (cases "m = 0") - case True - with \length bs = length cs\ show ?thesis by simp - next - case False - from "1.hyps" [of \bin_split n w\ \drop_bit n w\ \take_bit n w\] \m \ 0\ \n \ 0\ - have hyp: "\v bs. length bs = Suc (length cs) \ - length (bin_rsplit_aux n (m - n) v bs) = - length (bin_rsplit_aux n (m - n) (drop_bit n w) (take_bit n w # cs))" - using bin_rsplit_aux_len by fastforce - from \length bs = length cs\ \n \ 0\ show ?thesis - by (auto simp add: bin_rsplit_aux_simp_alt Let_def bin_rsplit_len split: prod.split) - qed -qed - -lemma bin_rsplit_len_indep: - "n \ 0 \ length (bin_rsplit n (nw, v)) = length (bin_rsplit n (nw, w))" - apply (unfold bin_rsplit_def) - apply (simp (no_asm)) - apply (erule bin_rsplit_aux_len_indep) - apply (rule refl) - done - - -subsection \Logical operations\ - -primrec bin_sc :: "nat \ bool \ int \ int" - where - Z: "bin_sc 0 b w = of_bool b + 2 * bin_rest w" - | Suc: "bin_sc (Suc n) b w = of_bool (odd w) + 2 * bin_sc n b (w div 2)" - -lemma bin_nth_sc [simp]: "bit (bin_sc n b w) n \ b" - by (induction n arbitrary: w) (simp_all add: bit_Suc) - -lemma bin_sc_sc_same [simp]: "bin_sc n c (bin_sc n b w) = bin_sc n c w" - by (induction n arbitrary: w) (simp_all add: bit_Suc) - -lemma bin_sc_sc_diff: "m \ n \ bin_sc m c (bin_sc n b w) = bin_sc n b (bin_sc m c w)" - apply (induct n arbitrary: w m) - apply (case_tac [!] m) - apply auto - done - -lemma bin_nth_sc_gen: "bin_nth (bin_sc n b w) m = (if m = n then b else bin_nth w m)" - apply (induct n arbitrary: w m) - apply (case_tac m; simp add: bit_Suc) - apply (case_tac m; simp add: bit_Suc) - done - -lemma bin_sc_eq: - \bin_sc n False = unset_bit n\ - \bin_sc n True = Bit_Operations.set_bit n\ - by (simp_all add: fun_eq_iff bit_eq_iff) - (simp_all add: bin_nth_sc_gen bit_set_bit_iff bit_unset_bit_iff) - -lemma bin_sc_nth [simp]: "bin_sc n (bin_nth w n) w = w" - by (rule bit_eqI) (simp add: bin_nth_sc_gen) - -lemma bin_sign_sc [simp]: "bin_sign (bin_sc n b w) = bin_sign w" -proof (induction n arbitrary: w) - case 0 - then show ?case - by (auto simp add: bin_sign_def) (use bin_rest_ge_0 in fastforce) -next - case (Suc n) - from Suc [of \w div 2\] - show ?case by (auto simp add: bin_sign_def split: if_splits) -qed - -lemma bin_sc_bintr [simp]: - "bintrunc m (bin_sc n x (bintrunc m w)) = bintrunc m (bin_sc n x w)" - apply (cases x) - apply (simp_all add: bin_sc_eq bit_eq_iff) - apply (auto simp add: bit_take_bit_iff bit_set_bit_iff bit_unset_bit_iff) - done - -lemma bin_clr_le: "bin_sc n False w \ w" - by (simp add: bin_sc_eq unset_bit_less_eq) - -lemma bin_set_ge: "bin_sc n True w \ w" - by (simp add: bin_sc_eq set_bit_greater_eq) - -lemma bintr_bin_clr_le: "bintrunc n (bin_sc m False w) \ bintrunc n w" - by (simp add: bin_sc_eq take_bit_unset_bit_eq unset_bit_less_eq) - -lemma bintr_bin_set_ge: "bintrunc n (bin_sc m True w) \ bintrunc n w" - by (simp add: bin_sc_eq take_bit_set_bit_eq set_bit_greater_eq) - -lemma bin_sc_FP [simp]: "bin_sc n False 0 = 0" - by (induct n) auto - -lemma bin_sc_TM [simp]: "bin_sc n True (- 1) = - 1" - by (induct n) auto - -lemmas bin_sc_simps = bin_sc.Z bin_sc.Suc bin_sc_TM bin_sc_FP - -lemma bin_sc_minus: "0 < n \ bin_sc (Suc (n - 1)) b w = bin_sc n b w" - by auto - -lemmas bin_sc_Suc_minus = - trans [OF bin_sc_minus [symmetric] bin_sc.Suc] - -lemma bin_sc_numeral [simp]: - "bin_sc (numeral k) b w = - of_bool (odd w) + 2 * bin_sc (pred_numeral k) b (w div 2)" - by (simp add: numeral_eq_Suc) - -instance int :: semiring_bit_syntax .. - -lemma test_bit_int_def [iff]: - "i !! n \ bin_nth i n" - by (simp add: test_bit_eq_bit) - -lemma shiftl_int_def: - "shiftl x n = x * 2 ^ n" for x :: int - by (simp add: push_bit_int_def shiftl_eq_push_bit) - -lemma shiftr_int_def: - "shiftr x n = x div 2 ^ n" for x :: int - by (simp add: drop_bit_int_def shiftr_eq_drop_bit) - - -subsubsection \Basic simplification rules\ - -lemmas int_not_def = not_int_def - -lemma int_not_simps [simp]: - "NOT (0::int) = -1" - "NOT (1::int) = -2" - "NOT (- 1::int) = 0" - "NOT (numeral w::int) = - numeral (w + Num.One)" - "NOT (- numeral (Num.Bit0 w)::int) = numeral (Num.BitM w)" - "NOT (- numeral (Num.Bit1 w)::int) = numeral (Num.Bit0 w)" - by (simp_all add: not_int_def) - -lemma int_not_not: "NOT (NOT x) = x" - for x :: int - by (fact bit.double_compl) - -lemma int_and_0 [simp]: "0 AND x = 0" - for x :: int - by (fact bit.conj_zero_left) - -lemma int_and_m1 [simp]: "-1 AND x = x" - for x :: int - by (fact bit.conj_one_left) - -lemma int_or_zero [simp]: "0 OR x = x" - for x :: int - by (fact bit.disj_zero_left) - -lemma int_or_minus1 [simp]: "-1 OR x = -1" - for x :: int - by (fact bit.disj_one_left) - -lemma int_xor_zero [simp]: "0 XOR x = x" - for x :: int - by (fact bit.xor_zero_left) - - -subsubsection \Binary destructors\ - -lemma bin_rest_NOT [simp]: "bin_rest (NOT x) = NOT (bin_rest x)" - by (fact not_int_div_2) - -lemma bin_last_NOT [simp]: "bin_last (NOT x) \ \ bin_last x" - by simp - -lemma bin_rest_AND [simp]: "bin_rest (x AND y) = bin_rest x AND bin_rest y" - by (subst and_int_rec) auto - -lemma bin_last_AND [simp]: "bin_last (x AND y) \ bin_last x \ bin_last y" - by (subst and_int_rec) auto - -lemma bin_rest_OR [simp]: "bin_rest (x OR y) = bin_rest x OR bin_rest y" - by (subst or_int_rec) auto - -lemma bin_last_OR [simp]: "bin_last (x OR y) \ bin_last x \ bin_last y" - by (subst or_int_rec) auto - -lemma bin_rest_XOR [simp]: "bin_rest (x XOR y) = bin_rest x XOR bin_rest y" - by (subst xor_int_rec) auto - -lemma bin_last_XOR [simp]: "bin_last (x XOR y) \ (bin_last x \ bin_last y) \ \ (bin_last x \ bin_last y)" - by (subst xor_int_rec) auto - -lemma bin_nth_ops: - "\x y. bin_nth (x AND y) n \ bin_nth x n \ bin_nth y n" - "\x y. bin_nth (x OR y) n \ bin_nth x n \ bin_nth y n" - "\x y. bin_nth (x XOR y) n \ bin_nth x n \ bin_nth y n" - "\x. bin_nth (NOT x) n \ \ bin_nth x n" - by (simp_all add: bit_and_iff bit_or_iff bit_xor_iff bit_not_iff) - - -subsubsection \Derived properties\ - -lemma int_xor_minus1 [simp]: "-1 XOR x = NOT x" - for x :: int - by (fact bit.xor_one_left) - -lemma int_xor_extra_simps [simp]: - "w XOR 0 = w" - "w XOR -1 = NOT w" - for w :: int - by simp_all - -lemma int_or_extra_simps [simp]: - "w OR 0 = w" - "w OR -1 = -1" - for w :: int - by simp_all - -lemma int_and_extra_simps [simp]: - "w AND 0 = 0" - "w AND -1 = w" - for w :: int - by simp_all - -text \Commutativity of the above.\ -lemma bin_ops_comm: - fixes x y :: int - shows int_and_comm: "x AND y = y AND x" - and int_or_comm: "x OR y = y OR x" - and int_xor_comm: "x XOR y = y XOR x" - by (simp_all add: ac_simps) - -lemma bin_ops_same [simp]: - "x AND x = x" - "x OR x = x" - "x XOR x = 0" - for x :: int - by simp_all - -lemmas bin_log_esimps = - int_and_extra_simps int_or_extra_simps int_xor_extra_simps - int_and_0 int_and_m1 int_or_zero int_or_minus1 int_xor_zero int_xor_minus1 - - -subsubsection \Basic properties of logical (bit-wise) operations\ - -lemma bbw_ao_absorb: "x AND (y OR x) = x \ x OR (y AND x) = x" - for x y :: int - by (auto simp add: bin_eq_iff bin_nth_ops) - -lemma bbw_ao_absorbs_other: - "x AND (x OR y) = x \ (y AND x) OR x = x" - "(y OR x) AND x = x \ x OR (x AND y) = x" - "(x OR y) AND x = x \ (x AND y) OR x = x" - for x y :: int - by (auto simp add: bin_eq_iff bin_nth_ops) - -lemmas bbw_ao_absorbs [simp] = bbw_ao_absorb bbw_ao_absorbs_other - -lemma int_xor_not: "(NOT x) XOR y = NOT (x XOR y) \ x XOR (NOT y) = NOT (x XOR y)" - for x y :: int - by (auto simp add: bin_eq_iff bin_nth_ops) - -lemma int_and_assoc: "(x AND y) AND z = x AND (y AND z)" - for x y z :: int - by (auto simp add: bin_eq_iff bin_nth_ops) - -lemma int_or_assoc: "(x OR y) OR z = x OR (y OR z)" - for x y z :: int - by (auto simp add: bin_eq_iff bin_nth_ops) - -lemma int_xor_assoc: "(x XOR y) XOR z = x XOR (y XOR z)" - for x y z :: int - by (auto simp add: bin_eq_iff bin_nth_ops) - -lemmas bbw_assocs = int_and_assoc int_or_assoc int_xor_assoc - -(* BH: Why are these declared as simp rules??? *) -lemma bbw_lcs [simp]: - "y AND (x AND z) = x AND (y AND z)" - "y OR (x OR z) = x OR (y OR z)" - "y XOR (x XOR z) = x XOR (y XOR z)" - for x y :: int - by (auto simp add: bin_eq_iff bin_nth_ops) - -lemma bbw_not_dist: - "NOT (x OR y) = (NOT x) AND (NOT y)" - "NOT (x AND y) = (NOT x) OR (NOT y)" - for x y :: int - by (auto simp add: bin_eq_iff bin_nth_ops) - -lemma bbw_oa_dist: "(x AND y) OR z = (x OR z) AND (y OR z)" - for x y z :: int - by (auto simp add: bin_eq_iff bin_nth_ops) - -lemma bbw_ao_dist: "(x OR y) AND z = (x AND z) OR (y AND z)" - for x y z :: int - by (auto simp add: bin_eq_iff bin_nth_ops) - -(* -Why were these declared simp??? -declare bin_ops_comm [simp] bbw_assocs [simp] -*) - - -subsubsection \Simplification with numerals\ - -text \Cases for \0\ and \-1\ are already covered by other simp rules.\ - -lemma bin_rest_neg_numeral_BitM [simp]: - "bin_rest (- numeral (Num.BitM w)) = - numeral w" - by simp - -lemma bin_last_neg_numeral_BitM [simp]: - "bin_last (- numeral (Num.BitM w))" - by simp - - -subsubsection \Interactions with arithmetic\ - -lemma le_int_or: "bin_sign y = 0 \ x \ x OR y" - for x y :: int - by (simp add: bin_sign_def or_greater_eq split: if_splits) - -lemmas int_and_le = - xtrans(3) [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] - -text \Interaction between bit-wise and arithmetic: good example of \bin_induction\.\ -lemma bin_add_not: "x + NOT x = (-1::int)" - by (simp add: not_int_def) - -lemma AND_mod: "x AND (2 ^ n - 1) = x mod 2 ^ n" - for x :: int - by (simp flip: take_bit_eq_mod add: take_bit_eq_mask mask_eq_exp_minus_1) - - -subsubsection \Truncating results of bit-wise operations\ - -lemma bin_trunc_ao: - "bintrunc n x AND bintrunc n y = bintrunc n (x AND y)" - "bintrunc n x OR bintrunc n y = bintrunc n (x OR y)" - by simp_all - -lemma bin_trunc_xor: "bintrunc n (bintrunc n x XOR bintrunc n y) = bintrunc n (x XOR y)" - by simp - -lemma bin_trunc_not: "bintrunc n (NOT (bintrunc n x)) = bintrunc n (NOT x)" - by (fact take_bit_not_take_bit) - -text \Want theorems of the form of \bin_trunc_xor\.\ -lemma bintr_bintr_i: "x = bintrunc n y \ bintrunc n x = bintrunc n y" - by auto - -lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i] -lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i] - - -subsubsection \More lemmas\ - -lemma not_int_cmp_0 [simp]: - fixes i :: int shows - "0 < NOT i \ i < -1" - "0 \ NOT i \ i < 0" - "NOT i < 0 \ i \ 0" - "NOT i \ 0 \ i \ -1" -by(simp_all add: int_not_def) arith+ - -lemma bbw_ao_dist2: "(x :: int) AND (y OR z) = x AND y OR x AND z" - by (fact bit.conj_disj_distrib) - -lemmas int_and_ac = bbw_lcs(1) int_and_comm int_and_assoc - -lemma int_nand_same [simp]: fixes x :: int shows "x AND NOT x = 0" - by simp - -lemma int_nand_same_middle: fixes x :: int shows "x AND y AND NOT x = 0" - by (simp add: bit_eq_iff bit_and_iff bit_not_iff) - -lemma and_xor_dist: fixes x :: int shows - "x AND (y XOR z) = (x AND y) XOR (x AND z)" - by (fact bit.conj_xor_distrib) - -lemma int_and_lt0 [simp]: - \x AND y < 0 \ x < 0 \ y < 0\ for x y :: int - by (fact and_negative_int_iff) - -lemma int_and_ge0 [simp]: - \x AND y \ 0 \ x \ 0 \ y \ 0\ for x y :: int - by (fact and_nonnegative_int_iff) - -lemma int_and_1: fixes x :: int shows "x AND 1 = x mod 2" - by (fact and_one_eq) - -lemma int_1_and: fixes x :: int shows "1 AND x = x mod 2" - by (fact one_and_eq) - -lemma int_or_lt0 [simp]: - \x OR y < 0 \ x < 0 \ y < 0\ for x y :: int - by (fact or_negative_int_iff) - -lemma int_or_ge0 [simp]: - \x OR y \ 0 \ x \ 0 \ y \ 0\ for x y :: int - by (fact or_nonnegative_int_iff) - -lemma int_xor_lt0 [simp]: - \x XOR y < 0 \ (x < 0) \ (y < 0)\ for x y :: int - by (fact xor_negative_int_iff) - -lemma int_xor_ge0 [simp]: - \x XOR y \ 0 \ (x \ 0 \ y \ 0)\ for x y :: int - by (fact xor_nonnegative_int_iff) - -lemma even_conv_AND: - \even i \ i AND 1 = 0\ for i :: int - by (simp add: and_one_eq mod2_eq_if) - -lemma bin_last_conv_AND: - "bin_last i \ i AND 1 \ 0" - by (simp add: and_one_eq mod2_eq_if) - -lemma bitval_bin_last: - "of_bool (bin_last i) = i AND 1" - by (simp add: and_one_eq mod2_eq_if) - -lemma bin_sign_and: - "bin_sign (i AND j) = - (bin_sign i * bin_sign j)" -by(simp add: bin_sign_def) - -lemma int_not_neg_numeral: "NOT (- numeral n) = (Num.sub n num.One :: int)" -by(simp add: int_not_def) - -lemma int_neg_numeral_pOne_conv_not: "- numeral (n + num.One) = (NOT (numeral n) :: int)" -by(simp add: int_not_def) - - -subsection \Setting and clearing bits\ - -lemma int_shiftl_BIT: fixes x :: int - shows int_shiftl0 [simp]: "x << 0 = x" - and int_shiftl_Suc [simp]: "x << Suc n = 2 * (x << n)" - by (auto simp add: shiftl_int_def) - -lemma int_0_shiftl [simp]: "0 << n = (0 :: int)" -by(induct n) simp_all - -lemma bin_last_shiftl: "bin_last (x << n) \ n = 0 \ bin_last x" -by(cases n)(simp_all) - -lemma bin_rest_shiftl: "bin_rest (x << n) = (if n > 0 then x << (n - 1) else bin_rest x)" -by(cases n)(simp_all) - -lemma bin_nth_shiftl [simp]: "bin_nth (x << n) m \ n \ m \ bin_nth x (m - n)" - by (simp add: bit_push_bit_iff_int shiftl_eq_push_bit) - -lemma bin_last_shiftr: "odd (x >> n) \ x !! n" for x :: int - by (simp add: shiftr_eq_drop_bit bit_iff_odd_drop_bit) - -lemma bin_rest_shiftr [simp]: "bin_rest (x >> n) = x >> Suc n" - by (simp add: bit_eq_iff shiftr_eq_drop_bit drop_bit_Suc bit_drop_bit_eq drop_bit_half) - -lemma bin_nth_shiftr [simp]: "bin_nth (x >> n) m = bin_nth x (n + m)" - by (simp add: shiftr_eq_drop_bit bit_drop_bit_eq) - -lemma bin_nth_conv_AND: - fixes x :: int shows - "bin_nth x n \ x AND (1 << n) \ 0" - by (simp add: bit_eq_iff) - (auto simp add: shiftl_eq_push_bit bit_and_iff bit_push_bit_iff bit_exp_iff) - -lemma int_shiftl_numeral [simp]: - "(numeral w :: int) << numeral w' = numeral (num.Bit0 w) << pred_numeral w'" - "(- numeral w :: int) << numeral w' = - numeral (num.Bit0 w) << pred_numeral w'" -by(simp_all add: numeral_eq_Suc shiftl_int_def) - (metis add_One mult_inc semiring_norm(11) semiring_norm(13) semiring_norm(2) semiring_norm(6) semiring_norm(87))+ - -lemma int_shiftl_One_numeral [simp]: - "(1 :: int) << numeral w = 2 << pred_numeral w" - using int_shiftl_numeral [of Num.One w] by simp - -lemma shiftl_ge_0 [simp]: fixes i :: int shows "i << n \ 0 \ i \ 0" -by(induct n) simp_all - -lemma shiftl_lt_0 [simp]: fixes i :: int shows "i << n < 0 \ i < 0" -by (metis not_le shiftl_ge_0) - -lemma int_shiftl_test_bit: "(n << i :: int) !! m \ m \ i \ n !! (m - i)" - by simp - -lemma int_0shiftr [simp]: "(0 :: int) >> x = 0" -by(simp add: shiftr_int_def) - -lemma int_minus1_shiftr [simp]: "(-1 :: int) >> x = -1" -by(simp add: shiftr_int_def div_eq_minus1) - -lemma int_shiftr_ge_0 [simp]: fixes i :: int shows "i >> n \ 0 \ i \ 0" - by (simp add: shiftr_eq_drop_bit) - -lemma int_shiftr_lt_0 [simp]: fixes i :: int shows "i >> n < 0 \ i < 0" -by (metis int_shiftr_ge_0 not_less) - -lemma int_shiftr_numeral [simp]: - "(1 :: int) >> numeral w' = 0" - "(numeral num.One :: int) >> numeral w' = 0" - "(numeral (num.Bit0 w) :: int) >> numeral w' = numeral w >> pred_numeral w'" - "(numeral (num.Bit1 w) :: int) >> numeral w' = numeral w >> pred_numeral w'" - "(- numeral (num.Bit0 w) :: int) >> numeral w' = - numeral w >> pred_numeral w'" - "(- numeral (num.Bit1 w) :: int) >> numeral w' = - numeral (Num.inc w) >> pred_numeral w'" - by (simp_all add: shiftr_eq_drop_bit numeral_eq_Suc add_One drop_bit_Suc) - -lemma int_shiftr_numeral_Suc0 [simp]: - "(1 :: int) >> Suc 0 = 0" - "(numeral num.One :: int) >> Suc 0 = 0" - "(numeral (num.Bit0 w) :: int) >> Suc 0 = numeral w" - "(numeral (num.Bit1 w) :: int) >> Suc 0 = numeral w" - "(- numeral (num.Bit0 w) :: int) >> Suc 0 = - numeral w" - "(- numeral (num.Bit1 w) :: int) >> Suc 0 = - numeral (Num.inc w)" - by (simp_all add: shiftr_eq_drop_bit drop_bit_Suc add_One) - -lemma bin_nth_minus_p2: - assumes sign: "bin_sign x = 0" - and y: "y = 1 << n" - and m: "m < n" - and x: "x < y" - shows "bin_nth (x - y) m = bin_nth x m" -proof - - from sign y x have \x \ 0\ and \y = 2 ^ n\ and \x < 2 ^ n\ - by (simp_all add: bin_sign_def shiftl_eq_push_bit push_bit_eq_mult split: if_splits) - from \0 \ x\ \x < 2 ^ n\ \m < n\ have \bit x m \ bit (x - 2 ^ n) m\ - proof (induction m arbitrary: x n) - case 0 - then show ?case - by simp - next - case (Suc m) - moreover define q where \q = n - 1\ - ultimately have n: \n = Suc q\ - by simp - have \(x - 2 ^ Suc q) div 2 = x div 2 - 2 ^ q\ - by simp - moreover from Suc.IH [of \x div 2\ q] Suc.prems - have \bit (x div 2) m \ bit (x div 2 - 2 ^ q) m\ - by (simp add: n) - ultimately show ?case - by (simp add: bit_Suc n) - qed - with \y = 2 ^ n\ show ?thesis - by simp -qed - -lemma bin_clr_conv_NAND: - "bin_sc n False i = i AND NOT (1 << n)" - by (induct n arbitrary: i) (rule bin_rl_eqI; simp)+ - -lemma bin_set_conv_OR: - "bin_sc n True i = i OR (1 << n)" - by (induct n arbitrary: i) (rule bin_rl_eqI; simp)+ - - -subsection \More lemmas on words\ - -lemma word_rcat_eq: - \word_rcat ws = word_of_int (bin_rcat (LENGTH('a::len)) (map uint ws))\ - for ws :: \'a::len word list\ - apply (simp add: word_rcat_def bin_rcat_def rev_map) - apply transfer - apply (simp add: horner_sum_foldr foldr_map comp_def) - done - -lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" - by (simp add: sign_Pls_ge_0) - -lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or - -\ \following definitions require both arithmetic and bit-wise word operations\ - -\ \to get \word_no_log_defs\ from \word_log_defs\, using \bin_log_bintrs\\ -lemmas wils1 = bin_log_bintrs [THEN word_of_int_eq_iff [THEN iffD2], - folded uint_word_of_int_eq, THEN eq_reflection] - -\ \the binary operations only\ (* BH: why is this needed? *) -lemmas word_log_binary_defs = - word_and_def word_or_def word_xor_def - -lemma setBit_no [simp]: "setBit (numeral bin) n = word_of_int (bin_sc n True (numeral bin))" - by transfer (simp add: bin_sc_eq) - -lemma clearBit_no [simp]: - "clearBit (numeral bin) n = word_of_int (bin_sc n False (numeral bin))" - by transfer (simp add: bin_sc_eq) - -lemma eq_mod_iff: "0 < n \ b = b mod n \ 0 \ b \ b < n" - for b n :: int - by auto (metis pos_mod_conj)+ - -lemma split_uint_lem: "bin_split n (uint w) = (a, b) \ - a = take_bit (LENGTH('a) - n) a \ b = take_bit (LENGTH('a)) b" - for w :: "'a::len word" - by transfer (simp add: drop_bit_take_bit ac_simps) - -\ \limited hom result\ -lemma word_cat_hom: - "LENGTH('a::len) \ LENGTH('b::len) + LENGTH('c::len) \ - (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = - word_of_int (bin_cat w (size b) (uint b))" - by transfer (simp add: take_bit_concat_bit_eq) - -lemma bintrunc_shiftl: - "take_bit n (m << i) = take_bit (n - i) m << i" - for m :: int - by (rule bit_eqI) (auto simp add: bit_take_bit_iff) - -lemma uint_shiftl: - "uint (n << i) = take_bit (size n) (uint n << i)" - by transfer (simp add: push_bit_take_bit shiftl_eq_push_bit) - - -code_identifier - code_module Bits_Int \ - (SML) Bit_Operations and (OCaml) Bit_Operations and (Haskell) Bit_Operations and (Scala) Bit_Operations - -end diff -r 131ab1a941dd -r dc62ecc7e59a src/HOL/Word/Traditional_Syntax.thy --- a/src/HOL/Word/Traditional_Syntax.thy Thu Jan 07 00:04:13 2021 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,526 +0,0 @@ -(* Author: Jeremy Dawson, NICTA -*) - -section \Operation variants with traditional syntax\ - -theory Traditional_Syntax - imports Word -begin - -class semiring_bit_syntax = semiring_bit_shifts -begin - -definition test_bit :: \'a \ nat \ bool\ (infixl "!!" 100) - where test_bit_eq_bit: \test_bit = bit\ - -definition shiftl :: \'a \ nat \ 'a\ (infixl "<<" 55) - where shiftl_eq_push_bit: \a << n = push_bit n a\ - -definition shiftr :: \'a \ nat \ 'a\ (infixl ">>" 55) - where shiftr_eq_drop_bit: \a >> n = drop_bit n a\ - -end - -instance word :: (len) semiring_bit_syntax .. - -context - includes lifting_syntax -begin - -lemma test_bit_word_transfer [transfer_rule]: - \(pcr_word ===> (=)) (\k n. n < LENGTH('a) \ bit k n) (test_bit :: 'a::len word \ _)\ - by (unfold test_bit_eq_bit) transfer_prover - -lemma shiftl_word_transfer [transfer_rule]: - \(pcr_word ===> (=) ===> pcr_word) (\k n. push_bit n k) shiftl\ - by (unfold shiftl_eq_push_bit) transfer_prover - -lemma shiftr_word_transfer [transfer_rule]: - \(pcr_word ===> (=) ===> pcr_word) (\k n. (drop_bit n \ take_bit LENGTH('a)) k) (shiftr :: 'a::len word \ _)\ - by (unfold shiftr_eq_drop_bit) transfer_prover - -end - -lemma test_bit_word_eq: - \test_bit = (bit :: 'a::len word \ _)\ - by (fact test_bit_eq_bit) - -lemma shiftl_word_eq: - \w << n = push_bit n w\ for w :: \'a::len word\ - by (fact shiftl_eq_push_bit) - -lemma shiftr_word_eq: - \w >> n = drop_bit n w\ for w :: \'a::len word\ - by (fact shiftr_eq_drop_bit) - -lemma test_bit_eq_iff: "test_bit u = test_bit v \ u = v" - for u v :: "'a::len word" - by (simp add: bit_eq_iff test_bit_eq_bit fun_eq_iff) - -lemma test_bit_size: "w !! n \ n < size w" - for w :: "'a::len word" - by transfer simp - -lemma word_eq_iff: "x = y \ (\n?P \ ?Q\) - for x y :: "'a::len word" - by transfer (auto simp add: bit_eq_iff bit_take_bit_iff) - -lemma word_eqI: "(\n. n < size u \ u !! n = v !! n) \ u = v" - for u :: "'a::len word" - by (simp add: word_size word_eq_iff) - -lemma word_eqD: "u = v \ u !! x = v !! x" - for u v :: "'a::len word" - by simp - -lemma test_bit_bin': "w !! n \ n < size w \ bit (uint w) n" - by transfer (simp add: bit_take_bit_iff) - -lemmas test_bit_bin = test_bit_bin' [unfolded word_size] - -lemma word_test_bit_def: - \test_bit a = bit (uint a)\ - by transfer (simp add: fun_eq_iff bit_take_bit_iff) - -lemmas test_bit_def' = word_test_bit_def [THEN fun_cong] - -lemma word_test_bit_transfer [transfer_rule]: - "(rel_fun pcr_word (rel_fun (=) (=))) - (\x n. n < LENGTH('a) \ bit x n) (test_bit :: 'a::len word \ _)" - by (simp only: test_bit_eq_bit) transfer_prover - -lemma test_bit_wi [simp]: - "(word_of_int x :: 'a::len word) !! n \ n < LENGTH('a) \ bit x n" - by transfer simp - -lemma word_ops_nth_size: - "n < size x \ - (x OR y) !! n = (x !! n | y !! n) \ - (x AND y) !! n = (x !! n \ y !! n) \ - (x XOR y) !! n = (x !! n \ y !! n) \ - (NOT x) !! n = (\ x !! n)" - for x :: "'a::len word" - by transfer (simp add: bit_or_iff bit_and_iff bit_xor_iff bit_not_iff) - -lemma word_ao_nth: - "(x OR y) !! n = (x !! n | y !! n) \ - (x AND y) !! n = (x !! n \ y !! n)" - for x :: "'a::len word" - by transfer (auto simp add: bit_or_iff bit_and_iff) - -lemmas msb0 = len_gt_0 [THEN diff_Suc_less, THEN word_ops_nth_size [unfolded word_size]] -lemmas msb1 = msb0 [where i = 0] - -lemma test_bit_numeral [simp]: - "(numeral w :: 'a::len word) !! n \ - n < LENGTH('a) \ bit (numeral w :: int) n" - by transfer (rule refl) - -lemma test_bit_neg_numeral [simp]: - "(- numeral w :: 'a::len word) !! n \ - n < LENGTH('a) \ bit (- numeral w :: int) n" - by transfer (rule refl) - -lemma test_bit_1 [simp]: "(1 :: 'a::len word) !! n \ n = 0" - by transfer (auto simp add: bit_1_iff) - -lemma nth_0 [simp]: "\ (0 :: 'a::len word) !! n" - by transfer simp - -lemma nth_minus1 [simp]: "(-1 :: 'a::len word) !! n \ n < LENGTH('a)" - by transfer simp - -lemma shiftl1_code [code]: - \shiftl1 w = push_bit 1 w\ - by transfer (simp add: ac_simps) - -lemma uint_shiftr_eq: - \uint (w >> n) = uint w div 2 ^ n\ - by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit min_def le_less less_diff_conv) - -lemma shiftr1_code [code]: - \shiftr1 w = drop_bit 1 w\ - by transfer (simp add: drop_bit_Suc) - -lemma shiftl_def: - \w << n = (shiftl1 ^^ n) w\ -proof - - have \push_bit n = (((*) 2 ^^ n) :: int \ int)\ for n - by (induction n) (simp_all add: fun_eq_iff funpow_swap1, simp add: ac_simps) - then show ?thesis - by transfer simp -qed - -lemma shiftr_def: - \w >> n = (shiftr1 ^^ n) w\ -proof - - have \shiftr1 ^^ n = (drop_bit n :: 'a word \ 'a word)\ - apply (induction n) - apply simp - apply (simp only: shiftr1_eq_div_2 [abs_def] drop_bit_eq_div [abs_def] funpow_Suc_right) - apply (use div_exp_eq [of _ 1, where ?'a = \'a word\] in simp) - done - then show ?thesis - by (simp add: shiftr_eq_drop_bit) -qed - -lemma bit_shiftl_word_iff: - \bit (w << m) n \ m \ n \ n < LENGTH('a) \ bit w (n - m)\ - for w :: \'a::len word\ - by (simp add: shiftl_word_eq bit_push_bit_iff exp_eq_zero_iff not_le) - -lemma bit_shiftr_word_iff: - \bit (w >> m) n \ bit w (m + n)\ - for w :: \'a::len word\ - by (simp add: shiftr_word_eq bit_drop_bit_eq) - -lift_definition sshiftr :: \'a::len word \ nat \ 'a word\ (infixl \>>>\ 55) - is \\k n. take_bit LENGTH('a) (drop_bit n (signed_take_bit (LENGTH('a) - Suc 0) k))\ - by (simp flip: signed_take_bit_decr_length_iff) - -lemma sshiftr_eq [code]: - \w >>> n = signed_drop_bit n w\ - by transfer simp - -lemma sshiftr_eq_funpow_sshiftr1: - \w >>> n = (sshiftr1 ^^ n) w\ - apply (rule sym) - apply (simp add: sshiftr1_eq_signed_drop_bit_Suc_0 sshiftr_eq) - apply (induction n) - apply simp_all - done - -lemma uint_sshiftr_eq: - \uint (w >>> n) = take_bit LENGTH('a) (sint w div 2 ^ n)\ - for w :: \'a::len word\ - by transfer (simp flip: drop_bit_eq_div) - -lemma sshift1_code [code]: - \sshiftr1 w = signed_drop_bit 1 w\ - by transfer (simp add: drop_bit_Suc) - -lemma sshiftr_0 [simp]: "0 >>> n = 0" - by transfer simp - -lemma sshiftr_n1 [simp]: "-1 >>> n = -1" - by transfer simp - -lemma bit_sshiftr_word_iff: - \bit (w >>> m) n \ bit w (if LENGTH('a) - m \ n \ n < LENGTH('a) then LENGTH('a) - 1 else (m + n))\ - for w :: \'a::len word\ - apply transfer - apply (auto simp add: bit_take_bit_iff bit_drop_bit_eq bit_signed_take_bit_iff min_def not_le simp flip: bit_Suc) - using le_less_Suc_eq apply fastforce - using le_less_Suc_eq apply fastforce - done - -lemma nth_sshiftr : - "(w >>> m) !! n = - (n < size w \ (if n + m \ size w then w !! (size w - 1) else w !! (n + m)))" - apply transfer - apply (auto simp add: bit_take_bit_iff bit_drop_bit_eq bit_signed_take_bit_iff min_def not_le ac_simps) - using le_less_Suc_eq apply fastforce - using le_less_Suc_eq apply fastforce - done - -lemma sshiftr_numeral [simp]: - \(numeral k >>> numeral n :: 'a::len word) = - word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (numeral k)))\ - apply (rule word_eqI) - apply (cases \LENGTH('a)\) - apply (simp_all add: word_size bit_drop_bit_eq nth_sshiftr bit_signed_take_bit_iff min_def not_le not_less less_Suc_eq_le ac_simps) - done - -lemma revcast_down_us [OF refl]: - "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (w >>> n)" - for w :: "'a::len word" - apply (simp add: source_size_def target_size_def) - apply (rule bit_word_eqI) - apply (simp add: bit_revcast_iff bit_ucast_iff bit_sshiftr_word_iff ac_simps) - done - -lemma revcast_down_ss [OF refl]: - "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (w >>> n)" - for w :: "'a::len word" - apply (simp add: source_size_def target_size_def) - apply (rule bit_word_eqI) - apply (simp add: bit_revcast_iff bit_word_scast_iff bit_sshiftr_word_iff ac_simps) - done - -lemma sshiftr_div_2n: "sint (w >>> n) = sint w div 2 ^ n" - using sint_signed_drop_bit_eq [of n w] - by (simp add: drop_bit_eq_div sshiftr_eq) - -lemmas lsb0 = len_gt_0 [THEN word_ops_nth_size [unfolded word_size]] - -lemma nth_sint: - fixes w :: "'a::len word" - defines "l \ LENGTH('a)" - shows "bit (sint w) n = (if n < l - 1 then w !! n else w !! (l - 1))" - unfolding sint_uint l_def - by (auto simp: bit_signed_take_bit_iff word_test_bit_def not_less min_def) - -lemma test_bit_2p: "(word_of_int (2 ^ n)::'a::len word) !! m \ m = n \ m < LENGTH('a)" - by transfer (auto simp add: bit_exp_iff) - -lemma nth_w2p: "((2::'a::len word) ^ n) !! m \ m = n \ m < LENGTH('a::len)" - by transfer (auto simp add: bit_exp_iff) - -lemma bang_is_le: "x !! m \ 2 ^ m \ x" - for x :: "'a::len word" - apply (rule xtrans(3)) - apply (rule_tac [2] y = "x" in le_word_or2) - apply (rule word_eqI) - apply (auto simp add: word_ao_nth nth_w2p word_size) - done - -lemma mask_eq: - \mask n = (1 << n) - (1 :: 'a::len word)\ - by transfer (simp add: mask_eq_exp_minus_1 push_bit_of_1) - -lemma nth_ucast: "(ucast w::'a::len word) !! n = (w !! n \ n < LENGTH('a))" - by transfer (simp add: bit_take_bit_iff ac_simps) - -lemma shiftl_0 [simp]: "(0::'a::len word) << n = 0" - by transfer simp - -lemma shiftr_0 [simp]: "(0::'a::len word) >> n = 0" - by transfer simp - -lemma nth_shiftl1: "shiftl1 w !! n \ n < size w \ n > 0 \ w !! (n - 1)" - by transfer (auto simp add: bit_double_iff) - -lemma nth_shiftl': "(w << m) !! n \ n < size w \ n >= m \ w !! (n - m)" - for w :: "'a::len word" - by transfer (auto simp add: bit_push_bit_iff) - -lemmas nth_shiftl = nth_shiftl' [unfolded word_size] - -lemma nth_shiftr1: "shiftr1 w !! n = w !! Suc n" - by transfer (auto simp add: bit_take_bit_iff simp flip: bit_Suc) - -lemma nth_shiftr: "(w >> m) !! n = w !! (n + m)" - for w :: "'a::len word" - apply (unfold shiftr_def) - apply (induct "m" arbitrary: n) - apply (auto simp add: nth_shiftr1) - done - -lemma nth_sshiftr1: "sshiftr1 w !! n = (if n = size w - 1 then w !! n else w !! Suc n)" - apply transfer - apply (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def simp flip: bit_Suc) - using le_less_Suc_eq apply fastforce - using le_less_Suc_eq apply fastforce - done - -lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" - by (fact uint_shiftr_eq) - -lemma shiftl_rev: "shiftl w n = word_reverse (shiftr (word_reverse w) n)" - by (induct n) (auto simp add: shiftl_def shiftr_def shiftl1_rev) - -lemma rev_shiftl: "word_reverse w << n = word_reverse (w >> n)" - by (simp add: shiftl_rev) - -lemma shiftr_rev: "w >> n = word_reverse (word_reverse w << n)" - by (simp add: rev_shiftl) - -lemma rev_shiftr: "word_reverse w >> n = word_reverse (w << n)" - by (simp add: shiftr_rev) - -lemma shiftl_numeral [simp]: - \numeral k << numeral l = (push_bit (numeral l) (numeral k) :: 'a::len word)\ - by (fact shiftl_word_eq) - -lemma shiftl_zero_size: "size x \ n \ x << n = 0" - for x :: "'a::len word" - apply transfer - apply (simp add: take_bit_push_bit) - done - -lemma shiftl_t2n: "shiftl w n = 2 ^ n * w" - for w :: "'a::len word" - by (induct n) (auto simp: shiftl_def shiftl1_2t) - -lemma shiftr_numeral [simp]: - \(numeral k >> numeral n :: 'a::len word) = drop_bit (numeral n) (numeral k)\ - by (fact shiftr_word_eq) - -lemma nth_mask [simp]: - \(mask n :: 'a::len word) !! i \ i < n \ i < size (mask n :: 'a word)\ - by (auto simp add: test_bit_word_eq word_size Word.bit_mask_iff) - -lemma slice_shiftr: "slice n w = ucast (w >> n)" - apply (rule bit_word_eqI) - apply (cases \n \ LENGTH('b)\) - apply (auto simp add: bit_slice_iff bit_ucast_iff bit_shiftr_word_iff ac_simps - dest: bit_imp_le_length) - done - -lemma nth_slice: "(slice n w :: 'a::len word) !! m = (w !! (m + n) \ m < LENGTH('a))" - by (simp add: slice_shiftr nth_ucast nth_shiftr) - -lemma revcast_down_uu [OF refl]: - "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (w >> n)" - for w :: "'a::len word" - apply (simp add: source_size_def target_size_def) - apply (rule bit_word_eqI) - apply (simp add: bit_revcast_iff bit_ucast_iff bit_shiftr_word_iff ac_simps) - done - -lemma revcast_down_su [OF refl]: - "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (w >> n)" - for w :: "'a::len word" - apply (simp add: source_size_def target_size_def) - apply (rule bit_word_eqI) - apply (simp add: bit_revcast_iff bit_word_scast_iff bit_shiftr_word_iff ac_simps) - done - -lemma cast_down_rev [OF refl]: - "uc = ucast \ source_size uc = target_size uc + n \ uc w = revcast (w << n)" - for w :: "'a::len word" - apply (simp add: source_size_def target_size_def) - apply (rule bit_word_eqI) - apply (simp add: bit_revcast_iff bit_word_ucast_iff bit_shiftl_word_iff) - done - -lemma revcast_up [OF refl]: - "rc = revcast \ source_size rc + n = target_size rc \ - rc w = (ucast w :: 'a::len word) << n" - apply (simp add: source_size_def target_size_def) - apply (rule bit_word_eqI) - apply (simp add: bit_revcast_iff bit_word_ucast_iff bit_shiftl_word_iff) - apply auto - apply (metis add.commute add_diff_cancel_right) - apply (metis diff_add_inverse2 diff_diff_add) - done - -lemmas rc1 = revcast_up [THEN - revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] -lemmas rc2 = revcast_down_uu [THEN - revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] - -lemmas ucast_up = - rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]] -lemmas ucast_down = - rc2 [simplified rev_shiftr revcast_ucast [symmetric]] - -\ \problem posed by TPHOLs referee: - criterion for overflow of addition of signed integers\ - -lemma sofl_test: - \sint x + sint y = sint (x + y) \ - (x + y XOR x) AND (x + y XOR y) >> (size x - 1) = 0\ - for x y :: \'a::len word\ -proof - - obtain n where n: \LENGTH('a) = Suc n\ - by (cases \LENGTH('a)\) simp_all - have *: \sint x + sint y + 2 ^ Suc n > signed_take_bit n (sint x + sint y) \ sint x + sint y \ - (2 ^ n)\ - \signed_take_bit n (sint x + sint y) > sint x + sint y - 2 ^ Suc n \ 2 ^ n > sint x + sint y\ - using signed_take_bit_int_greater_eq [of \sint x + sint y\ n] signed_take_bit_int_less_eq [of n \sint x + sint y\] - by (auto intro: ccontr) - have \sint x + sint y = sint (x + y) \ - (sint (x + y) < 0 \ sint x < 0) \ - (sint (x + y) < 0 \ sint y < 0)\ - using sint_less [of x] sint_greater_eq [of x] sint_less [of y] sint_greater_eq [of y] - signed_take_bit_int_eq_self [of \LENGTH('a) - 1\ \sint x + sint y\] - apply (auto simp add: not_less) - apply (unfold sint_word_ariths) - apply (subst signed_take_bit_int_eq_self) - prefer 4 - apply (subst signed_take_bit_int_eq_self) - prefer 7 - apply (subst signed_take_bit_int_eq_self) - prefer 10 - apply (subst signed_take_bit_int_eq_self) - apply (auto simp add: signed_take_bit_int_eq_self signed_take_bit_eq_take_bit_minus take_bit_Suc_from_most n not_less intro!: *) - done - then show ?thesis - apply (simp only: One_nat_def word_size shiftr_word_eq drop_bit_eq_zero_iff_not_bit_last bit_and_iff bit_xor_iff) - apply (simp add: bit_last_iff) - done -qed - -lemma shiftr_zero_size: "size x \ n \ x >> n = 0" - for x :: "'a :: len word" - by (rule word_eqI) (auto simp add: nth_shiftr dest: test_bit_size) - -lemma test_bit_cat [OF refl]: - "wc = word_cat a b \ wc !! n = (n < size wc \ - (if n < size b then b !! n else a !! (n - size b)))" - apply (simp add: word_size not_less; transfer) - apply (auto simp add: bit_concat_bit_iff bit_take_bit_iff) - done - -\ \keep quantifiers for use in simplification\ -lemma test_bit_split': - "word_split c = (a, b) \ - (\n m. - b !! n = (n < size b \ c !! n) \ - a !! m = (m < size a \ c !! (m + size b)))" - by (auto simp add: word_split_bin' test_bit_bin bit_unsigned_iff word_size - bit_drop_bit_eq ac_simps exp_eq_zero_iff - dest: bit_imp_le_length) - -lemma test_bit_split: - "word_split c = (a, b) \ - (\n::nat. b !! n \ n < size b \ c !! n) \ - (\m::nat. a !! m \ m < size a \ c !! (m + size b))" - by (simp add: test_bit_split') - -lemma test_bit_split_eq: - "word_split c = (a, b) \ - ((\n::nat. b !! n = (n < size b \ c !! n)) \ - (\m::nat. a !! m = (m < size a \ c !! (m + size b))))" - apply (rule_tac iffI) - apply (rule_tac conjI) - apply (erule test_bit_split [THEN conjunct1]) - apply (erule test_bit_split [THEN conjunct2]) - apply (case_tac "word_split c") - apply (frule test_bit_split) - apply (erule trans) - apply (fastforce intro!: word_eqI simp add: word_size) - done - -lemma test_bit_rcat: - "sw = size (hd wl) \ rc = word_rcat wl \ rc !! n = - (n < size rc \ n div sw < size wl \ (rev wl) ! (n div sw) !! (n mod sw))" - for wl :: "'a::len word list" - by (simp add: word_size word_rcat_def foldl_map rev_map bit_horner_sum_uint_exp_iff) - (simp add: test_bit_eq_bit) - -lemmas test_bit_cong = arg_cong [where f = "test_bit", THEN fun_cong] - -lemma max_test_bit: "(max_word::'a::len word) !! n \ n < LENGTH('a)" - by (fact nth_minus1) - -lemma shiftr_x_0 [iff]: "x >> 0 = x" - for x :: "'a::len word" - by transfer simp - -lemma shiftl_x_0 [simp]: "x << 0 = x" - for x :: "'a::len word" - by (simp add: shiftl_t2n) - -lemma shiftl_1 [simp]: "(1::'a::len word) << n = 2^n" - by (simp add: shiftl_t2n) - -lemma shiftr_1[simp]: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" - by (induct n) (auto simp: shiftr_def) - -lemma map_nth_0 [simp]: "map ((!!) (0::'a::len word)) xs = replicate (length xs) False" - by (induct xs) auto - -lemma word_and_1: - "n AND 1 = (if n !! 0 then 1 else 0)" for n :: "_ word" - by (rule bit_word_eqI) (auto simp add: bit_and_iff test_bit_eq_bit bit_1_iff intro: gr0I) - -lemma test_bit_1' [simp]: - "(1 :: 'a :: len word) !! n \ 0 < LENGTH('a) \ n = 0" - by simp - -lemma shiftl0: - "x << 0 = (x :: 'a :: len word)" - by (fact shiftl_x_0) - -end