diff -r a398b2a47aec -r c89d8e8bd8c7 src/HOL/Word/Word.thy --- a/src/HOL/Word/Word.thy Sun Oct 25 22:46:17 2020 +0000 +++ b/src/HOL/Word/Word.thy Mon Oct 26 11:28:43 2020 +0000 @@ -9,7 +9,6 @@ "HOL-Library.Type_Length" "HOL-Library.Boolean_Algebra" "HOL-Library.Bit_Operations" - Traditional_Syntax begin subsection \Preliminaries\ @@ -957,6 +956,14 @@ end lemma [code]: + \push_bit n w = w * 2 ^ n\ for w :: \'a::len word\ + by (fact push_bit_eq_mult) + +lemma [code]: + \Word.the_int (drop_bit n w) = drop_bit n (Word.the_int w)\ + by transfer (simp add: drop_bit_take_bit min_def le_less less_diff_conv) + +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) @@ -1784,138 +1791,10 @@ \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 \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_word_eq) -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)" @@ -2033,6 +1912,11 @@ apply (metis le_antisym less_eq_decr_length_iff) done +lemma [code]: + \Word.the_int (signed_drop_bit n w) = take_bit LENGTH('a) (drop_bit n (Word.the_signed_int w))\ + for w :: \'a::len word\ + by transfer simp + lemma signed_drop_bit_signed_drop_bit [simp]: \signed_drop_bit m (signed_drop_bit n w) = signed_drop_bit (m + n) w\ for w :: \'a::len word\ @@ -2054,10 +1938,6 @@ apply (auto simp add: bit_sint_iff bit_drop_bit_eq bit_signed_drop_bit_iff dest: bit_imp_le_length) done -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) - lift_definition sshiftr1 :: \'a::len word \ 'a word\ is \\k. take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - Suc 0) k div 2)\ by (simp flip: signed_take_bit_decr_length_iff) @@ -2066,10 +1946,6 @@ is \\b k. take_bit LENGTH('a) k div 2 + of_bool b * 2 ^ (LENGTH('a) - Suc 0)\ by (fact arg_cong) -lemma sshiftr_eq: - \w >>> n = signed_drop_bit n w\ - by transfer simp - lemma sshiftr1_eq_signed_drop_bit_Suc_0: \sshiftr1 = signed_drop_bit (Suc 0)\ by (rule ext) (transfer, simp add: drop_bit_Suc) @@ -2078,32 +1954,6 @@ \sshiftr1 w = word_of_int (sint w div 2)\ 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 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\ @@ -2445,31 +2295,6 @@ 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 \ bit (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: "bit (uint w) n \ n < LENGTH('a)" for w :: "'a::len word" by transfer (simp add: bit_take_bit_iff) @@ -2527,9 +2352,6 @@ lemma scast_id [simp]: "scast w = w" by transfer (simp add: take_bit_signed_take_bit) -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) @@ -2644,8 +2466,6 @@ 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: @@ -3686,52 +3506,6 @@ \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) \ 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 - \ \get from commutativity, associativity etc of \int_and\ etc to same for \word_and etc\\ lemmas bwsimps = wi_hom_add @@ -3742,21 +3516,21 @@ "(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]) + by (fact ac_simps)+ 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]) + by (fact ac_simps)+ 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]) + by (fact ac_simps)+ lemma word_log_esimps: "x AND 0 = 0" @@ -3797,24 +3571,24 @@ "(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]) + by (auto intro: bit_eqI simp add: bit_and_iff bit_or_iff) lemma word_not_not [simp]: "NOT (NOT x) = x" for x :: "'a::len word" - by simp + by (fact bit.double_compl) 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]) + by (fact bit.conj_disj_distrib2) 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]) - + by (fact bit.disj_conj_distrib2) + lemma word_add_not [simp]: "x + NOT x = -1" for x :: "'a::len word" - by transfer (simp add: not_int_def) - + by (simp add: not_eq_complement) + 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) @@ -3863,21 +3637,6 @@ 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 "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 uint_2p: "(0::'a::len word) < 2 ^ n \ uint (2 ^ n::'a::len word) = 2 ^ n" apply (cases \n < LENGTH('a)\; transfer) apply auto @@ -3886,14 +3645,6 @@ 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\ @@ -3924,37 +3675,6 @@ 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 \_ div 2\) takes normal arguments to normal results, @@ -3962,8 +3682,11 @@ \ lemma uint_shiftr1: "uint (shiftr1 w) = uint w div 2" - using uint_shiftr_eq [of w 1] - by (simp add: shiftr1_code) + using drop_bit_eq_div [of 1 \uint w\, symmetric] + apply simp + apply transfer + apply (simp add: drop_bit_take_bit min_def) + done lemma bit_sshiftr1_iff: \bit (sshiftr1 w) n \ bit w (if n = LENGTH('a) - 1 then LENGTH('a) - 1 else Suc n)\ @@ -3974,31 +3697,6 @@ 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 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) @@ -4006,13 +3704,6 @@ using sint_signed_drop_bit_eq [of 1 w] by (simp add: drop_bit_Suc sshiftr1_eq_signed_drop_bit_Suc_0) -lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" - by (fact uint_shiftr_eq) - -lemma sshiftr_div_2n: "sint (sshiftr 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) - lemma bit_bshiftr1_iff: \bit (bshiftr1 b w) n \ b \ n = LENGTH('a) - 1 \ bit w (Suc n)\ for w :: \'a::len word\ @@ -4030,28 +3721,6 @@ 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" @@ -4062,10 +3731,6 @@ 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 (take_bit LENGTH('a) (numeral w) div 2)" @@ -4083,18 +3748,6 @@ (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 bit_signed_take_bit_iff min_def 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 @@ -4156,10 +3809,6 @@ 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 transfer (simp add: take_bit_minus_one_eq_mask) @@ -4307,16 +3956,6 @@ 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 @@ -4394,153 +4033,15 @@ 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 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 - \ \this odd result is analogous to \ucast_id\, result to the length given by the result type\ @@ -4548,11 +4049,8 @@ by transfer (simp add: take_bit_concat_bit_eq) 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 + apply (rule bit_word_eqI) + apply (auto simp add: bit_word_cat_iff not_less word_size word_split_def bit_ucast_iff bit_drop_bit_eq) done lemmas word_cat_split_size = sym [THEN [2] word_cat_split_alt [symmetric]] @@ -4561,16 +4059,19 @@ 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)+ + apply (auto simp add: word_split_def word_size) + apply (rule bit_word_eqI) + apply (simp add: bit_slice_iff bit_ucast_iff bit_drop_bit_eq) + apply (cases \LENGTH('c) \ LENGTH('b)\) + apply (auto simp add: ac_simps dest: bit_imp_le_length) + apply (rule bit_word_eqI) + apply (auto simp add: bit_slice_iff bit_ucast_iff dest: bit_imp_le_length) 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) + apply (rule bit_word_eqI) + apply (auto simp add: bit_slice_iff bit_word_cat_iff word_size) done lemmas slice_cat2 = trans [OF slice_id word_cat_id] @@ -4578,20 +4079,17 @@ 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 + apply (rule bit_word_eqI) + apply (auto simp add: bit_slice_iff bit_word_cat_iff word_size) 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)+ + apply (auto simp add: word_split_def word_size) + apply (rule bit_eqI) + apply (auto simp add: bit_ucast_iff bit_drop_bit_eq bit_word_cat_iff dest: bit_imp_le_length) + apply (rule bit_eqI) + apply (auto simp add: bit_ucast_iff bit_drop_bit_eq bit_word_cat_iff dest: bit_imp_le_length) done lemma horner_sum_uint_exp_Cons_eq: @@ -4618,15 +4116,6 @@ (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 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] - subsection \Rotation\ @@ -4776,9 +4265,6 @@ 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) @@ -4787,33 +4273,22 @@ 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) + by (fact bit.conj_disj_distrib) 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) + by (fact bit.disj_conj_distrib) 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) + by (fact bit.conj_cancel_right) lemma word_or_not [simp]: "x OR NOT x = max_word" - by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) + by (fact bit.disj_cancel_right) 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) + by (fact bit.xor_def) lemma uint_lt_0 [simp]: "uint x < 0 = False" by (simp add: linorder_not_less) @@ -4821,16 +4296,10 @@ 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 @@ -5078,14 +4547,6 @@ 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) @@ -5095,10 +4556,6 @@ lemma bin_last_bintrunc: "odd (take_bit l n) \ l > 0 \ odd 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) - subsection \SMT support\