From 1b15714cbf9db54fe9351449f2642e216b3fafd5 Mon Sep 17 00:00:00 2001 From: Gerwin Klein Date: Sun, 14 Nov 2021 21:05:54 +1100 Subject: [PATCH] isabelle2021-1: sync Word_Lib from afp Signed-off-by: Gerwin Klein --- lib/Word_Lib/Aligned.thy | 224 ++- lib/Word_Lib/Ancient_Numeral.thy | 18 +- lib/Word_Lib/Bit_Comprehension.thy | 7 +- lib/Word_Lib/Bit_Shifts_Infix_Syntax.thy | 201 +++ lib/Word_Lib/Bits_Int.thy | 712 +++++---- lib/Word_Lib/Bitwise.thy | 33 +- lib/Word_Lib/Enumeration_Word.thy | 33 +- lib/Word_Lib/Examples.thy | 1048 +++++++++++++- lib/Word_Lib/Generic_set_bit.thy | 130 +- lib/Word_Lib/Guide.thy | 114 +- lib/Word_Lib/Least_significant_bit.thy | 18 +- lib/Word_Lib/Legacy_Aliases.thy | 125 +- lib/Word_Lib/Machine_Word_32.thy | 136 ++ lib/Word_Lib/Machine_Word_32_Basics.thy | 66 + lib/Word_Lib/Machine_Word_64.thy | 136 ++ lib/Word_Lib/Machine_Word_64_Basics.thy | 66 + lib/Word_Lib/Many_More.thy | 6 + lib/Word_Lib/More_Arithmetic.thy | 14 +- lib/Word_Lib/More_Divides.thy | 6 +- lib/Word_Lib/More_Word.thy | 874 +++++++++++- lib/Word_Lib/More_Word_Operations.thy | 212 +-- lib/Word_Lib/Most_significant_bit.thy | 101 +- lib/Word_Lib/Next_and_Prev.thy | 4 +- lib/Word_Lib/Norm_Words.thy | 54 +- lib/Word_Lib/ROOT | 11 +- lib/Word_Lib/Reversed_Bit_Lists.thy | 400 +++--- lib/Word_Lib/Rsplit.thy | 33 +- lib/Word_Lib/Signed_Division_Word.thy | 119 +- lib/Word_Lib/Signed_Words.thy | 17 +- lib/Word_Lib/Singleton_Bit_Shifts.thy | 148 ++ lib/Word_Lib/Strict_part_mono.thy | 2 +- lib/Word_Lib/Syntax_Bundles.thy | 20 + lib/Word_Lib/Traditional_Infix_Syntax.thy | 1085 -------------- lib/Word_Lib/Type_Syntax.thy | 2 +- lib/Word_Lib/Typedef_Morphisms.thy | 38 +- lib/Word_Lib/Word_16.thy | 6 + lib/Word_Lib/Word_32.thy | 180 +-- lib/Word_Lib/Word_64.thy | 175 +-- lib/Word_Lib/Word_8.thy | 8 +- lib/Word_Lib/Word_EqI.thy | 12 +- lib/Word_Lib/Word_Lemmas.thy | 1584 +++++++++++---------- lib/Word_Lib/Word_Lib_Sumo.thy | 29 +- lib/Word_Lib/Word_Syntax.thy | 6 + lib/Word_Lib/Word_Type_Syntax.thy | 58 - lib/Word_Lib/document/root.tex | 1 + 45 files changed, 4943 insertions(+), 3329 deletions(-) create mode 100644 lib/Word_Lib/Bit_Shifts_Infix_Syntax.thy create mode 100644 lib/Word_Lib/Machine_Word_32.thy create mode 100644 lib/Word_Lib/Machine_Word_32_Basics.thy create mode 100644 lib/Word_Lib/Machine_Word_64.thy create mode 100644 lib/Word_Lib/Machine_Word_64_Basics.thy create mode 100644 lib/Word_Lib/Singleton_Bit_Shifts.thy create mode 100644 lib/Word_Lib/Syntax_Bundles.thy delete mode 100644 lib/Word_Lib/Traditional_Infix_Syntax.thy delete mode 100644 lib/Word_Lib/Word_Type_Syntax.thy diff --git a/lib/Word_Lib/Aligned.thy b/lib/Word_Lib/Aligned.thy index b943eec61..d0de3633c 100644 --- a/lib/Word_Lib/Aligned.thy +++ b/lib/Word_Lib/Aligned.thy @@ -10,8 +10,11 @@ theory Aligned imports "HOL-Library.Word" More_Word - Word_EqI - Typedef_Morphisms + Bit_Shifts_Infix_Syntax +begin + +context + includes bit_operations_syntax begin lift_definition is_aligned :: \'a::len word \ nat \ bool\ @@ -185,11 +188,9 @@ proof cases ultimately have upls: "unat x + unat y = 2 ^ m * (2 ^ k * q1 + q2)" proof - have f1: "unat x = 2 ^ (m + k) * q1" - by (metis (no_types) \x = of_nat (2 ^ (m + k) * q1)\ l1 nat_mod_lem word_unat.inverse_norm - zero_less_numeral zero_less_power) + using q1v unat_mult_power_lem xv by blast have "unat y = 2 ^ m * q2" - by (metis (no_types) \y = of_nat (2 ^ m * q2)\ l2 nat_mod_lem word_unat.inverse_norm - zero_less_numeral zero_less_power) + using q2v unat_mult_power_lem yv by blast then show ?thesis using f1 by (simp add: power_add semiring_normalization_rules(34)) qed @@ -249,29 +250,26 @@ proof cases from mv obtain q where mq: "m + q = LENGTH('a)" and "0 < q" by (auto dest: less_imp_add_positive) - have "(2::nat) ^ m dvd unat (k << m)" + have "(2::nat) ^ m dvd unat (push_bit m k)" proof have kv: "(unat k div 2 ^ q) * 2 ^ q + unat k mod 2 ^ q = unat k" by (rule div_mult_mod_eq) - have "unat (k << m) = unat (2 ^ m * k)" by (simp add: shiftl_t2n) + have "unat (push_bit m k) = unat (2 ^ m * k)" + by (simp add: push_bit_eq_mult ac_simps) also have "\ = (2 ^ m * unat k) mod (2 ^ LENGTH('a))" using mv by (simp add: unat_word_ariths(2)) also have "\ = 2 ^ m * (unat k mod 2 ^ q)" by (subst mq [symmetric], subst power_add, subst mod_mult2_eq) simp - finally show "unat (k << m) = 2 ^ m * (unat k mod 2 ^ q)" . + finally show "unat (push_bit m k) = 2 ^ m * (unat k mod 2 ^ q)" . qed - - then show ?thesis by (unfold is_aligned_iff_dvd_nat) + then show ?thesis by (unfold is_aligned_iff_dvd_nat shiftl_def) next assume "\ m < LENGTH('a)" then show ?thesis - by (simp add: not_less power_overflow is_aligned_mask shiftl_zero_size word_size) + by (simp add: not_less power_overflow is_aligned_mask word_size shiftl_def) qed -lemma word_mod_by_0: "k mod (0::'a::len word) = k" - by (simp add: word_arith_nat_mod) - lemma aligned_mod_eq_0: fixes p::"'a::len word" assumes al: "is_aligned p sz" @@ -454,7 +452,7 @@ lemma is_aligned_no_wrap''': apply (drule is_aligned_no_wrap[where off="of_nat off"]) apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) - apply (simp add: take_bit_eq_mod) + apply (simp add: take_bit_eq_mod unsigned_of_nat) apply (subst(asm) unat_of_nat_len) apply (erule order_less_trans) apply (erule power_strict_increasing) @@ -505,16 +503,19 @@ lemma is_aligned_no_overflow'': apply (erule word_sub_1_le) done -lemma is_aligned_nth [word_eqI_simps]: - "is_aligned p m = (\n < m. \p !! n)" - apply (clarsimp simp: is_aligned_mask bang_eq word_size) - apply (rule iffI) - apply clarsimp - apply (case_tac "n < size p") - apply (simp add: word_size) - apply (drule test_bit_size) - apply simp - apply clarsimp +lemma is_aligned_bitI: + \is_aligned p m\ if \\n. n < m \ \ bit p n\ + apply (simp add: is_aligned_mask) + apply (rule bit_word_eqI) + using that + apply (auto simp add: bit_simps) + done + +lemma is_aligned_nth: + "is_aligned p m = (\n < m. \ bit p n)" + apply (auto intro: is_aligned_bitI simp add: is_aligned_mask bit_eq_iff) + apply (auto simp: bit_simps) + using bit_imp_le_length not_less apply blast done lemma range_inter: @@ -612,7 +613,7 @@ qed lemma is_aligned_neg_mask: "m \ n \ is_aligned (x AND NOT (mask n)) m" - by (metis and_not_mask is_aligned_shift is_aligned_weaken) + by (rule is_aligned_bitI) (simp add: bit_simps) lemma unat_minus: "unat (- (x :: 'a :: len word)) = (if x = 0 then 0 else 2 ^ size x - unat x)" @@ -632,32 +633,32 @@ lemma is_aligned_minus: lemma add_mask_lower_bits: "\is_aligned (x :: 'a :: len word) n; - \n' \ n. n' < LENGTH('a) \ \ p !! n'\ \ x + p AND NOT (mask n) = x" + \n' \ n. n' < LENGTH('a) \ \ bit p n'\ \ x + p AND NOT (mask n) = x" apply (subst word_plus_and_or_coroll) apply (rule word_eqI) apply (clarsimp simp: word_size is_aligned_nth) apply (erule_tac x=na in allE)+ - apply simp + apply (simp add: bit_simps) apply (rule bit_word_eqI) - apply (auto simp add: bit_simps not_less test_bit_eq_bit) - apply (metis is_aligned_nth not_le test_bit_eq_bit) + apply (auto simp add: bit_simps not_less word_size) + apply (metis is_aligned_nth not_le) done lemma is_aligned_andI1: "is_aligned x n \ is_aligned (x AND y) n" - by (simp add: is_aligned_nth) + by (simp add: is_aligned_nth bit_simps) lemma is_aligned_andI2: "is_aligned y n \ is_aligned (x AND y) n" - by (simp add: is_aligned_nth) + by (simp add: is_aligned_nth bit_simps) lemma is_aligned_shiftl: "is_aligned w (n - m) \ is_aligned (w << m) n" - by (simp add: is_aligned_nth nth_shiftl) + by (simp add: is_aligned_nth bit_simps) lemma is_aligned_shiftr: "is_aligned w (n + m) \ is_aligned (w >> m) n" - by (simp add: is_aligned_nth nth_shiftr) + by (simp add: is_aligned_nth bit_simps) lemma is_aligned_shiftl_self: "is_aligned (p << n) n" @@ -665,22 +666,28 @@ lemma is_aligned_shiftl_self: lemma is_aligned_neg_mask_eq: "is_aligned p n \ p AND NOT (mask n) = p" - by (metis add.left_neutral is_aligned_mask word_plus_and_or_coroll2) + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps is_aligned_nth) + done lemma is_aligned_shiftr_shiftl: "is_aligned w n \ w >> n << n = w" - by (metis and_not_mask is_aligned_neg_mask_eq) + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps is_aligned_nth intro: ccontr) + apply (subst add_diff_inverse_nat) + apply (auto intro: ccontr) + done lemma aligned_shiftr_mask_shiftl: "is_aligned x n \ ((x >> n) AND mask v) << n = x AND mask (v + n)" apply (rule word_eqI) - apply (simp add: word_size nth_shiftl nth_shiftr) - apply (subgoal_tac "\m. x !! m \ m \ n") + apply (simp add: word_size bit_simps) + apply (subgoal_tac "\m. bit x m \ m \ n") apply auto[1] apply (clarsimp simp: is_aligned_mask) apply (drule_tac x=m in word_eqD) apply (frule test_bit_size) - apply (simp add: word_size) + apply (simp add: word_size bit_simps) done lemma mask_zero: @@ -715,7 +722,7 @@ lemma is_aligned_neg_mask_weaken: lemma is_aligned_neg_mask2 [simp]: "is_aligned (a AND NOT (mask n)) n" - by (simp add: and_not_mask is_aligned_shift) + by (rule is_aligned_bitI) (simp add: bit_simps) lemma is_aligned_0': "is_aligned 0 n" @@ -759,16 +766,12 @@ proof cases by (rule aligned_add_offset_no_wrap) fact+ show ?thesis using al szv - apply - - apply (erule is_alignedE) - apply (subst word_unat.Rep_inject [symmetric]) - apply (subst unat_mod) - apply (subst iffD1 [OF unat_add_lem], rule ux) - apply simp - apply (subst unat_mult_power_lem, assumption+) - apply (simp) - apply (rule mod_less[OF less_le_trans[OF unat_mono], OF kv]) - apply (erule eq_imp_le[OF unat_power_lower]) + apply (simp flip: take_bit_eq_mod) + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps) + apply (metis assms(2) bit_or_iff is_aligned_mask is_aligned_nth leD less_mask_eq word_and_le1 word_bw_lcs(1) word_neq_0_conv word_plus_and_or_coroll) + apply (meson assms(2) leI less_2p_is_upper_bits_unset) + apply (metis assms(2) bit_disjunctive_add_iff bit_imp_le_length bit_push_bit_iff is_alignedE' less_2p_is_upper_bits_unset) done next assume "\ sz < LENGTH('a)" @@ -868,14 +871,11 @@ lemma mask_out_add_aligned: lemma is_aligned_add_or: "\is_aligned p n; d < 2 ^ n\ \ p + d = p OR d" - apply (subst disjunctive_add) - apply (simp_all add: is_aligned_iff_take_bit_eq_0) - apply (simp add: bit_eq_iff) - apply (auto simp add: bit_simps) + apply (subst disjunctive_add, simp_all) + apply (clarsimp simp: is_aligned_nth less_2p_is_upper_bits_unset) subgoal for m apply (cases \m < n\) - apply (auto simp add: not_less) - apply (metis bit_take_bit_iff less_mask_eq take_bit_eq_mask) + apply (auto simp add: not_less dest: bit_imp_possible_bit) done done @@ -921,10 +921,29 @@ lemma and_neg_mask_eq_iff_not_mask_le: for w :: \'a::len word\ by (metis eq_iff neg_mask_mono_le word_and_le1 word_and_le2 word_bw_same(1)) -lemma neg_mask_le_high_bits [word_eqI_simps]: - "NOT(mask n) \ w \ (\i \ {n ..< size w}. w !! i)" +lemma neg_mask_le_high_bits: + \NOT (mask n) \ w \ (\i \ {n ..< size w}. bit w i)\ (is \?P \ ?Q\) for w :: \'a::len word\ - by (auto simp: word_size and_neg_mask_eq_iff_not_mask_le[symmetric] word_eq_iff neg_mask_test_bit) +proof + assume ?Q + then have \w AND NOT (mask n) = NOT (mask n)\ + by (auto simp add: bit_simps word_size intro: bit_word_eqI) + then show ?P + by (simp add: and_neg_mask_eq_iff_not_mask_le) +next + assume ?P + then have *: \w AND NOT (mask n) = NOT (mask n)\ + by (simp add: and_neg_mask_eq_iff_not_mask_le) + show \?Q\ + proof (rule ccontr) + assume \\ (\i\{n.. + then obtain m where m: \\ bit w m\ \n \ m\ \m < LENGTH('a)\ + by (auto simp add: word_size) + from * have \bit (w AND NOT (mask n)) m \ bit (NOT (mask n :: 'a word)) m\ + by auto + with m show False by (auto simp add: bit_simps) + qed +qed lemma is_aligned_add_less_t2n: "\is_aligned (p::'a::len word) n; d < 2^n; n \ m; p < 2^m\ \ p + d < 2^m" @@ -1016,8 +1035,6 @@ lemma aligned_less_plus_1: apply (clarsimp simp: field_simps) apply (drule arg_cong[where f="\x. x - 1"]) apply (clarsimp simp: is_aligned_mask) - apply (drule word_eqD[where x=0]) - apply simp done lemma aligned_add_offset_less: @@ -1134,12 +1151,12 @@ proof - by (rule diff_less, simp) ultimately show "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (LENGTH('a) - s1)" using take_bit_nat_less_self_iff [of \LENGTH('a)\ \mq - 2 ^ sq * nq\] - apply (auto simp add: word_less_nat_alt not_le not_less) + apply (auto simp add: word_less_nat_alt not_le not_less unsigned_of_nat) apply (metis take_bit_nat_eq_self_iff) done qed then have "mq - 2 ^ sq * nq < 2 ^ (s2 - s1)" using mq s2wb - apply (simp add: word_less_nat_alt take_bit_eq_mod) + apply (simp add: word_less_nat_alt take_bit_eq_mod unsigned_of_nat) apply (subst (asm) mod_less) apply auto apply (rule order_le_less_trans) @@ -1181,20 +1198,20 @@ lemma is_aligned_add: lemma aligned_shift: "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ - \ x + y >> n = y >> n" + \ (x + y) >> n = y >> n" apply (subst word_plus_and_or_coroll; rule bit_word_eqI) - apply (auto simp add: bit_simps is_aligned_nth test_bit_eq_bit) - apply (metis less_2p_is_upper_bits_unset not_le test_bit_word_eq) - apply (metis le_add1 less_2p_is_upper_bits_unset test_bit_bin test_bit_word_eq) + apply (auto simp add: bit_simps is_aligned_nth) + apply (metis less_2p_is_upper_bits_unset not_le) + apply (metis le_add1 less_2p_is_upper_bits_unset test_bit_bin) done lemma aligned_shift': "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ - \ y + x >> n = y >> n" + \ (y + x) >> n = y >> n" apply (subst word_plus_and_or_coroll; rule bit_word_eqI) - apply (auto simp add: bit_simps is_aligned_nth test_bit_eq_bit) - apply (metis less_2p_is_upper_bits_unset not_le test_bit_eq_bit) - apply (metis bit_imp_le_length le_add1 less_2p_is_upper_bits_unset test_bit_eq_bit) + apply (auto simp add: bit_simps is_aligned_nth) + apply (metis less_2p_is_upper_bits_unset not_le) + apply (metis bit_imp_le_length le_add1 less_2p_is_upper_bits_unset) done lemma and_neg_mask_plus_mask_mono: "(p AND NOT (mask n)) + mask n \ p" @@ -1242,4 +1259,67 @@ lemma and_mask_plus: apply (simp add:word_bw_comms word_bw_lcs) done +lemma is_aligned_add_not_aligned: + "\is_aligned (p::'a::len word) n; \ is_aligned (q::'a::len word) n\ \ \ is_aligned (p + q) n" + by (metis is_aligned_addD1) + +lemma neg_mask_add_aligned: + "\ is_aligned p n; q < 2 ^ n \ \ (p + q) AND NOT (mask n) = p AND NOT (mask n)" + by (metis is_aligned_add_helper is_aligned_neg_mask_eq) + +lemma word_add_power_off: + fixes a :: "'a :: len word" + assumes ak: "a < k" + and kw: "k < 2 ^ (LENGTH('a) - m)" + and mw: "m < LENGTH('a)" + and off: "off < 2 ^ m" +shows "(a * 2 ^ m) + off < k * 2 ^ m" +proof (cases "m = 0") + case True + then show ?thesis using off ak by simp +next + case False + + from ak have ak1: "a + 1 \ k" by (rule inc_le) + then have "(a + 1) * 2 ^ m \ 0" + apply - + apply (rule word_power_nonzero) + apply (erule order_le_less_trans [OF _ kw]) + apply (rule mw) + apply (rule less_is_non_zero_p1 [OF ak]) + done + then have "(a * 2 ^ m) + off < ((a + 1) * 2 ^ m)" using kw mw + apply - + apply (simp add: distrib_right) + apply (rule word_plus_strict_mono_right [OF off]) + apply (rule is_aligned_no_overflow'') + apply (rule is_aligned_mult_triv2) + apply assumption + done + also have "\ \ k * 2 ^ m" using ak1 mw kw False + apply - + apply (erule word_mult_le_mono1) + apply (simp add: p2_gt_0) + apply (simp add: word_less_nat_alt) + apply (meson nat_mult_power_less_eq zero_less_numeral) + done + finally show ?thesis . +qed + +lemma offset_not_aligned: + "\ is_aligned (p::'a::len word) n; i > 0; i < 2 ^ n; n < LENGTH('a)\ \ + \ is_aligned (p + of_nat i) n" + apply (erule is_aligned_add_not_aligned) + apply transfer + apply (auto simp add: is_aligned_iff_udvd) + apply (metis (no_types, lifting) le_less_trans less_not_refl2 less_or_eq_imp_le of_nat_eq_0_iff take_bit_eq_0_iff take_bit_nat_eq_self_iff take_bit_of_nat unat_lt2p unat_power_lower) + done + +lemma le_or_mask: + "w \ w' \ w OR mask x \ w' OR mask x" + for w w' :: \'a::len word\ + by (metis neg_mask_add_mask add.commute le_word_or1 mask_2pm1 neg_mask_mono_le word_plus_mono_left) + +end + end diff --git a/lib/Word_Lib/Ancient_Numeral.thy b/lib/Word_Lib/Ancient_Numeral.thy index 2224df6ae..88d054072 100644 --- a/lib/Word_Lib/Ancient_Numeral.thy +++ b/lib/Word_Lib/Ancient_Numeral.thy @@ -5,7 +5,7 @@ *) theory Ancient_Numeral - imports Main Reversed_Bit_Lists + imports Main Reversed_Bit_Lists Legacy_Aliases begin definition Bit :: "int \ bool \ int" (infixl "BIT" 90) @@ -188,6 +188,10 @@ lemma sbintrunc_Suc_minus_Is: lemma bin_cat_Suc_Bit: "bin_cat w (Suc n) (v BIT b) = bin_cat w n v BIT b" by (auto simp add: Bit_def concat_bit_Suc) +context + includes bit_operations_syntax +begin + lemma int_not_BIT [simp]: "NOT (w BIT b) = (NOT w) BIT (\ b)" by (simp add: not_int_def Bit_def) @@ -200,6 +204,8 @@ lemma int_or_Bits [simp]: "(x BIT b) OR (y BIT c) = (x OR y) BIT (b \ c)" lemma int_xor_Bits [simp]: "(x BIT b) XOR (y BIT c) = (x XOR y) BIT ((b \ c) \ \ (b \ c))" using xor_int_rec [of \x BIT b\ \y BIT c\] by (auto simp add: Bit_B0_2t Bit_B1_2t) +end + lemma mod_BIT: "bin BIT bit mod 2 ^ Suc n = (bin mod 2 ^ n) BIT bit" for bit proof - @@ -221,13 +227,9 @@ lemma int_lsb_BIT [simp]: fixes x :: int shows by(simp add: lsb_int_def) lemma int_shiftr_BIT [simp]: fixes x :: int - shows int_shiftr0: "x >> 0 = x" - and int_shiftr_Suc: "x BIT b >> Suc n = x >> n" -proof - - show "x >> 0 = x" by (simp add: shiftr_int_def) - show "x BIT b >> Suc n = x >> n" by (cases b) - (simp_all add: shiftr_int_def Bit_def add.commute pos_zdiv_mult_2) -qed + shows int_shiftr0: "drop_bit 0 x = x" + and int_shiftr_Suc: "drop_bit (Suc n) (x BIT b) = drop_bit n x" + by (simp_all add: drop_bit_Suc) lemma msb_BIT [simp]: "msb (x BIT b) = msb x" by(simp add: msb_int_def) diff --git a/lib/Word_Lib/Bit_Comprehension.thy b/lib/Word_Lib/Bit_Comprehension.thy index 05e763f43..647828efd 100644 --- a/lib/Word_Lib/Bit_Comprehension.thy +++ b/lib/Word_Lib/Bit_Comprehension.thy @@ -7,7 +7,8 @@ section \Comprehension syntax for bit expressions\ theory Bit_Comprehension - imports "HOL-Library.Word" + imports + "HOL-Library.Word" begin class bit_comprehension = ring_bit_operations + @@ -72,7 +73,7 @@ instance by standard end -lemma bit_set_bits_word_iff: +lemma bit_set_bits_word_iff [bit_simps]: \bit (set_bits P :: 'a::len word) n \ n < LENGTH('a) \ P n\ by (auto simp add: word_set_bits_def bit_horner_sum_bit_word_iff) @@ -100,7 +101,7 @@ proof (cases \\n. \m\n. f m \ f n\ then have n: \\m. n \ m \ f m \ f n\ by blast from n_def have n_eq: \(LEAST q. \m\q. f m \ f n) = n\ - by (smt Least_equality Least_le \\m\n. f m = f n\ dual_order.refl le_refl n order_refl) + by (smt (verit, best) Least_le \\m\n. f m = f n\ dual_order.antisym wellorder_Least_lemma(1)) show ?thesis proof (cases \f n\) case False diff --git a/lib/Word_Lib/Bit_Shifts_Infix_Syntax.thy b/lib/Word_Lib/Bit_Shifts_Infix_Syntax.thy new file mode 100644 index 000000000..4fa91fbec --- /dev/null +++ b/lib/Word_Lib/Bit_Shifts_Infix_Syntax.thy @@ -0,0 +1,201 @@ +(* + * Copyright Data61, CSIRO (ABN 41 687 119 230) + * + * SPDX-License-Identifier: BSD-2-Clause + *) + +(* Author: Jeremy Dawson, NICTA *) + +section \Shift operations with infix syntax\ + +theory Bit_Shifts_Infix_Syntax + imports "HOL-Library.Word" More_Word +begin + +context semiring_bit_operations +begin + +definition shiftl :: \'a \ nat \ 'a\ (infixl "<<" 55) + where [code_unfold]: \a << n = push_bit n a\ + +lemma bit_shiftl_iff [bit_simps]: + \bit (a << m) n \ m \ n \ possible_bit TYPE('a) n \ bit a (n - m)\ + by (simp add: shiftl_def bit_simps) + +definition shiftr :: \'a \ nat \ 'a\ (infixl ">>" 55) + where [code_unfold]: \a >> n = drop_bit n a\ + +lemma bit_shiftr_eq [bit_simps]: + \bit (a >> n) = bit a \ (+) n\ + by (simp add: shiftr_def bit_simps) + +end + +definition sshiftr :: \'a::len word \ nat \ 'a word\ (infixl \>>>\ 55) + where [code_unfold]: \w >>> n = signed_drop_bit n w\ + +lemma bit_sshiftr_iff [bit_simps]: + \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\ + by (simp add: sshiftr_def bit_simps) + +context + includes lifting_syntax +begin + +lemma shiftl_word_transfer [transfer_rule]: + \(pcr_word ===> (=) ===> pcr_word) (\k n. push_bit n k) (<<)\ + apply (unfold shiftl_def) + apply transfer_prover + done + +lemma shiftr_word_transfer [transfer_rule]: + \((pcr_word :: int \ 'a::len word \ _) ===> (=) ===> pcr_word) + (\k n. drop_bit n (take_bit LENGTH('a) k)) + (>>)\ +proof - + have \((pcr_word :: int \ 'a::len word \ _) ===> (=) ===> pcr_word) + (\k n. (drop_bit n \ take_bit LENGTH('a)) k) + (>>)\ + by (unfold shiftr_def) transfer_prover + then show ?thesis + by simp +qed + +lemma sshiftr_transfer [transfer_rule]: + \((pcr_word :: int \ 'a::len word \ _) ===> (=) ===> pcr_word) + (\k n. drop_bit n (signed_take_bit (LENGTH('a) - Suc 0) k)) + (>>>)\ +proof - + have \((pcr_word :: int \ 'a::len word \ _) ===> (=) ===> pcr_word) + (\k n. (drop_bit n \ signed_take_bit (LENGTH('a) - Suc 0)) k) + (>>>)\ + by (unfold sshiftr_def) transfer_prover + then show ?thesis + by simp +qed + +end + +context semiring_bit_operations +begin + +lemma shiftl_0 [simp]: + \0 << n = 0\ + by (simp add: shiftl_def) + +lemma shiftl_of_0 [simp]: + \a << 0 = a\ + by (simp add: shiftl_def) + +lemma shiftl_of_Suc [simp]: + \a << Suc n = (a * 2) << n\ + by (simp add: shiftl_def) + +lemma shiftl_1 [simp]: + \1 << n = 2 ^ n\ + by (simp add: shiftl_def) + +lemma shiftl_numeral_Suc [simp]: + \numeral m << Suc n = push_bit (Suc n) (numeral m)\ + by (fact shiftl_def) + +lemma shiftl_numeral_numeral [simp]: + \numeral m << numeral n = push_bit (numeral n) (numeral m)\ + by (fact shiftl_def) + +lemma shiftr_0 [simp]: + \0 >> n = 0\ + by (simp add: shiftr_def) + +lemma shiftr_of_0 [simp]: + \a >> 0 = a\ + by (simp add: shiftr_def) + +lemma shiftr_1 [simp]: + \1 >> n = of_bool (n = 0)\ + by (simp add: shiftr_def) + +lemma shiftr_numeral_Suc [simp]: + \numeral m >> Suc n = drop_bit (Suc n) (numeral m)\ + by (fact shiftr_def) + +lemma shiftr_numeral_numeral [simp]: + \numeral m >> numeral n = drop_bit (numeral n) (numeral m)\ + by (fact shiftr_def) + +end + +context ring_bit_operations +begin + +context + includes bit_operations_syntax +begin + +lemma shiftl_minus_1_numeral [simp]: + \- 1 << numeral n = NOT (mask (numeral n))\ + by (simp add: shiftl_def) + +end + +end + +lemma shiftl_Suc_0 [simp]: + \Suc 0 << n = 2 ^ n\ + by (simp add: shiftl_def) + +lemma shiftr_Suc_0 [simp]: + \Suc 0 >> n = of_bool (n = 0)\ + by (simp add: shiftr_def) + +lemma sshiftr_numeral_Suc [simp]: + \numeral m >>> Suc n = signed_drop_bit (Suc n) (numeral m)\ + by (fact sshiftr_def) + +lemma sshiftr_numeral_numeral [simp]: + \numeral m >>> numeral n = signed_drop_bit (numeral n) (numeral m)\ + by (fact sshiftr_def) + +context ring_bit_operations +begin + +lemma shiftl_minus_numeral_Suc [simp]: + \- numeral m << Suc n = push_bit (Suc n) (- numeral m)\ + by (fact shiftl_def) + +lemma shiftl_minus_numeral_numeral [simp]: + \- numeral m << numeral n = push_bit (numeral n) (- numeral m)\ + by (fact shiftl_def) + +lemma shiftr_minus_numeral_Suc [simp]: + \- numeral m >> Suc n = drop_bit (Suc n) (- numeral m)\ + by (fact shiftr_def) + +lemma shiftr_minus_numeral_numeral [simp]: + \- numeral m >> numeral n = drop_bit (numeral n) (- numeral m)\ + by (fact shiftr_def) + +end + +lemma sshiftr_0 [simp]: + \0 >>> n = 0\ + by (simp add: sshiftr_def) + +lemma sshiftr_of_0 [simp]: + \w >>> 0 = w\ + by (simp add: sshiftr_def) + +lemma sshiftr_1 [simp]: + \(1 :: 'a::len word) >>> n = of_bool (LENGTH('a) = 1 \ n = 0)\ + by (simp add: sshiftr_def) + +lemma sshiftr_minus_numeral_Suc [simp]: + \- numeral m >>> Suc n = signed_drop_bit (Suc n) (- numeral m)\ + by (fact sshiftr_def) + +lemma sshiftr_minus_numeral_numeral [simp]: + \- numeral m >>> numeral n = signed_drop_bit (numeral n) (- numeral m)\ + by (fact sshiftr_def) + +end diff --git a/lib/Word_Lib/Bits_Int.thy b/lib/Word_Lib/Bits_Int.thy index 9962de53b..45c1a10cd 100644 --- a/lib/Word_Lib/Bits_Int.thy +++ b/lib/Word_Lib/Bits_Int.thy @@ -8,22 +8,18 @@ section \Bitwise Operations on integers\ theory Bits_Int imports - "HOL-Library.Word" - Traditional_Infix_Syntax + "Word_Lib.Most_significant_bit" + "Word_Lib.Least_significant_bit" + "Word_Lib.Generic_set_bit" + "Word_Lib.Bit_Comprehension" 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" + "(odd :: int \ bool) 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]: "\ odd (0 :: int)" "odd (1 :: int)" @@ -36,53 +32,50 @@ lemma bin_last_numeral_simps [simp]: 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)" + "(\k::int. k div 2) 0 = 0" + "(\k::int. k div 2) 1 = 0" + "(\k::int. k div 2) (- 1) = - 1" + "(\k::int. k div 2) Numeral1 = 0" + "(\k::int. k div 2) (numeral (Num.Bit0 w)) = numeral w" + "(\k::int. k div 2) (numeral (Num.Bit1 w)) = numeral w" + "(\k::int. k div 2) (- numeral (Num.Bit0 w)) = - numeral w" + "(\k::int. k div 2) (- numeral (Num.Bit1 w)) = - numeral (w + Num.One)" by simp_all -lemma bin_rl_eqI: "\bin_rest x = bin_rest y; odd x = odd y\ \ x = y" +lemma bin_rl_eqI: "\(\k::int. k div 2) x = (\k::int. k div 2) y; odd x = odd 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" + shows bin_rest_lt0: "(\k::int. k div 2) i < 0 \ i < 0" + and bin_rest_ge_0: "(\k::int. k div 2) i \ 0 \ i \ 0" by auto -lemma bin_rest_gt_0 [simp]: "bin_rest x > 0 \ x > 1" +lemma bin_rest_gt_0 [simp]: "(\k::int. k div 2) 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" +lemma bin_nth_eq_iff: "(bit :: int \ nat \ bool) x = (bit :: int \ nat \ bool) 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) + "x = y" if "\n. (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" + using that by (rule bit_eqI) -lemma bin_eq_iff: "x = y \ (\n. bin_nth x n = bin_nth y n)" - by (fact bit_eq_iff) +lemma bin_eq_iff: "x = y \ (\n. (bit :: int \ nat \ bool) x n = (bit :: int \ nat \ bool) y n)" + by (metis bit_eq_iff) -lemma bin_nth_zero [simp]: "\ bin_nth 0 n" +lemma bin_nth_zero [simp]: "\ (bit :: int \ nat \ bool) 0 n" by simp -lemma bin_nth_1 [simp]: "bin_nth 1 n \ n = 0" +lemma bin_nth_1 [simp]: "(bit :: int \ nat \ bool) 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_minus1 [simp]: "(bit :: int \ nat \ bool) (- 1) n" + by simp -lemma bin_nth_numeral: "bin_rest x = y \ bin_nth x (numeral n) = bin_nth y (pred_numeral n)" +lemma bin_nth_numeral: "(\k::int. k div 2) x = y \ (bit :: int \ nat \ bool) x (numeral n) = (bit :: int \ nat \ bool) y (pred_numeral n)" by (simp add: numeral_eq_Suc bit_Suc) lemmas bin_nth_numeral_simps [simp] = @@ -92,10 +85,10 @@ 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\\ +lemma nth_2p_bin: "(bit :: int \ nat \ bool) (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)" +lemma nth_rest_power_bin: "(bit :: int \ nat \ bool) (((\k::int. k div 2) ^^ k) w) n = (bit :: int \ nat \ bool) w (n + k)" apply (induct k arbitrary: n) apply clarsimp apply clarsimp @@ -103,8 +96,8 @@ lemma nth_rest_power_bin: "bin_nth ((bin_rest ^^ k) w) n = bin_nth w (n + k)" 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))" + "(bit :: int \ nat \ bool) (numeral (num.Bit0 x)) n \ n > 0 \ (bit :: int \ nat \ bool) (numeral x) (n - 1)" + "(bit :: int \ nat \ bool) (numeral (num.Bit1 x)) n \ (n > 0 \ (bit :: int \ nat \ bool) (numeral x) (n - 1))" by (cases n; simp)+ @@ -121,100 +114,91 @@ lemma bin_sign_simps [simp]: "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" +lemma bin_sign_rest [simp]: "bin_sign ((\k::int. k div 2) 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" +lemma bintrunc_mod2p: "(take_bit :: nat \ int \ int) 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" +lemma sbintrunc_mod2p: "(signed_take_bit :: nat \ int \ int) 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\ + \(signed_take_bit :: nat \ int \ int) 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" +lemma sign_bintr: "bin_sign ((take_bit :: nat \ int \ int) n w) = 0" by (simp add: bin_sign_def) -lemma bintrunc_n_0: "bintrunc n 0 = 0" +lemma bintrunc_n_0: "(take_bit :: nat \ int \ int) n 0 = 0" by (fact take_bit_of_0) -lemma sbintrunc_n_0: "sbintrunc n 0 = 0" +lemma sbintrunc_n_0: "(signed_take_bit :: nat \ int \ int) n 0 = 0" by (fact signed_take_bit_of_0) -lemma sbintrunc_n_minus1: "sbintrunc n (- 1) = -1" +lemma sbintrunc_n_minus1: "(signed_take_bit :: nat \ int \ int) 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) + "(take_bit :: nat \ int \ int) (Suc n) 1 = 1" + "(take_bit :: nat \ int \ int) (Suc n) (- 1) = 1 + 2 * (take_bit :: nat \ int \ int) n (- 1)" + "(take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) n (numeral w)" + "(take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) n (numeral w)" + "(take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) n (- numeral w)" + "(take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) n (- numeral (w + Num.One))" + by (simp_all add: take_bit_Suc del: take_bit_minus_one_eq_mask) 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" + "(signed_take_bit :: nat \ int \ int) 0 1 = -1" + "(signed_take_bit :: nat \ int \ int) 0 (numeral (Num.Bit0 w)) = 0" + "(signed_take_bit :: nat \ int \ int) 0 (numeral (Num.Bit1 w)) = -1" + "(signed_take_bit :: nat \ int \ int) 0 (- numeral (Num.Bit0 w)) = 0" + "(signed_take_bit :: nat \ int \ int) 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))" + "(signed_take_bit :: nat \ int \ int) (Suc n) 1 = 1" + "(signed_take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) n (numeral w)" + "(signed_take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) n (numeral w)" + "(signed_take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) n (- numeral w)" + "(signed_take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) 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" +lemma bin_sign_lem: "(bin_sign ((signed_take_bit :: nat \ int \ int) 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" +lemma nth_bintr: "(bit :: int \ nat \ bool) ((take_bit :: nat \ int \ int) m w) n \ n < m \ (bit :: int \ nat \ bool) 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)" +lemma nth_sbintr: "(bit :: int \ nat \ bool) ((signed_take_bit :: nat \ int \ int) m w) n = (if n < m then (bit :: int \ nat \ bool) w n else (bit :: int \ nat \ bool) 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)" + "(bit :: int \ nat \ bool) (numeral (Num.Bit0 w)) n \ + (\m. n = Suc m \ (bit :: int \ nat \ bool) (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)" + "(bit :: int \ nat \ bool) (numeral (Num.Bit1 w)) n \ + n = 0 \ (\m. n = Suc m \ (bit :: int \ nat \ bool) (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" +lemma bintrunc_bintrunc_l: "n \ m \ (take_bit :: nat \ int \ int) m ((take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) n w" by simp -lemma sbintrunc_sbintrunc_l: "n \ m \ sbintrunc m (sbintrunc n w) = sbintrunc n w" - by (simp add: min_def) +lemma sbintrunc_sbintrunc_l: "n \ m \ (signed_take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) n w" + by simp -lemma bintrunc_bintrunc_ge: "n \ m \ bintrunc n (bintrunc m w) = bintrunc n w" +lemma bintrunc_bintrunc_ge: "n \ m \ (take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) m w) = (take_bit :: nat \ int \ int) 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" +lemma bintrunc_bintrunc_min [simp]: "(take_bit :: nat \ int \ int) m ((take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) (min m n) w" by (rule take_bit_take_bit) -lemma sbintrunc_sbintrunc_min [simp]: "sbintrunc m (sbintrunc n w) = sbintrunc (min m n) w" +lemma sbintrunc_sbintrunc_min [simp]: "(signed_take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) (min m n) w" by (rule signed_take_bit_signed_take_bit) lemmas sbintrunc_Suc_Pls = @@ -237,10 +221,10 @@ lemmas sbintrunc_0_simps = lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs -lemma bintrunc_minus: "0 < n \ bintrunc (Suc (n - 1)) w = bintrunc n w" +lemma bintrunc_minus: "0 < n \ (take_bit :: nat \ int \ int) (Suc (n - 1)) w = (take_bit :: nat \ int \ int) n w" by auto -lemma sbintrunc_minus: "0 < n \ sbintrunc (Suc (n - 1)) w = sbintrunc n w" +lemma sbintrunc_minus: "0 < n \ (signed_take_bit :: nat \ int \ int) (Suc (n - 1)) w = (signed_take_bit :: nat \ int \ int) n w" by auto lemmas sbintrunc_minus_simps = @@ -248,25 +232,25 @@ lemmas sbintrunc_minus_simps = lemma sbintrunc_BIT_I: \0 < n \ - sbintrunc (n - 1) 0 = y \ - sbintrunc n 0 = 2 * y\ + (signed_take_bit :: nat \ int \ int) (n - 1) 0 = y \ + (signed_take_bit :: nat \ int \ int) n 0 = 2 * y\ by simp lemma sbintrunc_Suc_Is: - \sbintrunc n (- 1) = y \ - sbintrunc (Suc n) (- 1) = 1 + 2 * y\ + \(signed_take_bit :: nat \ int \ int) n (- 1) = y \ + (signed_take_bit :: nat \ int \ int) (Suc n) (- 1) = 1 + 2 * y\ by auto -lemma sbintrunc_Suc_lem: "sbintrunc (Suc n) x = y \ m = Suc n \ sbintrunc m x = y" +lemma sbintrunc_Suc_lem: "(signed_take_bit :: nat \ int \ int) (Suc n) x = y \ m = Suc n \ (signed_take_bit :: nat \ int \ int) m x = y" by (rule ssubst) 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" +lemma sbintrunc_bintrunc_lt: "m > n \ (signed_take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) m w) = (signed_take_bit :: nat \ int \ int) 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" +lemma bintrunc_sbintrunc_le: "m \ Suc n \ (take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) m w" by (rule take_bit_signed_take_bit) lemmas bintrunc_sbintrunc [simp] = order_refl [THEN bintrunc_sbintrunc_le] @@ -274,13 +258,13 @@ 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" +lemma bintrunc_sbintrunc' [simp]: "0 < n \ (take_bit :: nat \ int \ int) n ((signed_take_bit :: nat \ int \ int) (n - 1) w) = (take_bit :: nat \ int \ int) n w" by (cases n) simp_all -lemma sbintrunc_bintrunc' [simp]: "0 < n \ sbintrunc (n - 1) (bintrunc n w) = sbintrunc (n - 1) w" +lemma sbintrunc_bintrunc' [simp]: "0 < n \ (signed_take_bit :: nat \ int \ int) (n - 1) ((take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) (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" +lemma bin_sbin_eq_iff: "(take_bit :: nat \ int \ int) (Suc n) x = (take_bit :: nat \ int \ int) (Suc n) y \ (signed_take_bit :: nat \ int \ int) n x = (signed_take_bit :: nat \ int \ int) n y" apply (rule iffI) apply (rule box_equals [OF _ sbintrunc_bintrunc sbintrunc_bintrunc]) apply simp @@ -289,7 +273,7 @@ lemma bin_sbin_eq_iff: "bintrunc (Suc n) x = bintrunc (Suc n) y \ bintrunc n x = bintrunc n y \ sbintrunc (n - 1) x = sbintrunc (n - 1) y" + "0 < n \ (take_bit :: nat \ int \ int) n x = (take_bit :: nat \ int \ int) n y \ (signed_take_bit :: nat \ int \ int) (n - 1) x = (signed_take_bit :: nat \ int \ int) (n - 1) y" by (cases n) (simp_all add: bin_sbin_eq_iff) lemmas bintrunc_sbintruncS0 [simp] = bintrunc_sbintrunc' [unfolded One_nat_def] @@ -306,51 +290,51 @@ 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)" + "(take_bit :: nat \ int \ int) (numeral k) x = of_bool (odd x) + 2 * (take_bit :: nat \ int \ int) (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)" + "(signed_take_bit :: nat \ int \ int) (numeral k) x = of_bool (odd x) + 2 * (signed_take_bit :: nat \ int \ int) (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" + "(take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit0 w)) = + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" + "(take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit1 w)) = + 1 + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" + "(take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit0 w)) = + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (- numeral w)" + "(take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit1 w)) = + 1 + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (- numeral (w + Num.One))" + "(take_bit :: nat \ int \ int) (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" + "(signed_take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit0 w)) = + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" + "(signed_take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit1 w)) = + 1 + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" + "(signed_take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit0 w)) = + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (- numeral w)" + "(signed_take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit1 w)) = + 1 + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (- numeral (w + Num.One))" + "(signed_take_bit :: nat \ int \ int) (numeral k) 1 = 1" by (simp_all add: sbintrunc_numeral) -lemma no_bintr_alt1: "bintrunc n = (\w. w mod 2 ^ n :: int)" +lemma no_bintr_alt1: "(take_bit :: nat \ int \ int) 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}" +lemma range_bintrunc: "range ((take_bit :: nat \ int \ int) 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)" +lemma no_sbintr_alt2: "(signed_take_bit :: nat \ int \ int) 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}" +lemma range_sbintrunc: "range ((signed_take_bit :: nat \ int \ int) 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))\ + moreover have \(signed_take_bit :: nat \ int \ int) 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) @@ -360,27 +344,27 @@ proof - qed lemma sbintrunc_inc: - \k + 2 ^ Suc n \ sbintrunc n k\ if \k < - (2 ^ n)\ + \k + 2 ^ Suc n \ (signed_take_bit :: nat \ int \ int) 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\ + \(signed_take_bit :: nat \ int \ int) 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" +lemma bintr_ge0: "0 \ (take_bit :: nat \ int \ int) n w" by (simp add: bintrunc_mod2p) -lemma bintr_lt2p: "bintrunc n w < 2 ^ n" +lemma bintr_lt2p: "(take_bit :: nat \ int \ int) 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 bintr_Min: "(take_bit :: nat \ int \ int) n (- 1) = 2 ^ n - 1" + by (simp add: stable_imp_take_bit_eq mask_eq_exp_minus_1) -lemma sbintr_ge: "- (2 ^ n) \ sbintrunc n w" - by (simp add: sbintrunc_mod2p) +lemma sbintr_ge: "- (2 ^ n) \ (signed_take_bit :: nat \ int \ int) n w" + by (fact signed_take_bit_int_greater_eq_minus_exp) -lemma sbintr_lt: "sbintrunc n w < 2 ^ n" - by (simp add: sbintrunc_mod2p) +lemma sbintr_lt: "(signed_take_bit :: nat \ int \ int) n w < 2 ^ n" + by (fact signed_take_bit_int_less_exp) lemma sign_Pls_ge_0: "bin_sign bin = 0 \ bin \ 0" for bin :: int @@ -390,29 +374,29 @@ 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)" +lemma bin_rest_trunc: "(\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin) = (take_bit :: nat \ int \ int) (n - 1) ((\k::int. k div 2) 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)" + "((\k::int. k div 2) ^^ k) ((take_bit :: nat \ int \ int) n bin) = (take_bit :: nat \ int \ int) (n - k) (((\k::int. k div 2) ^^ 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)" +lemma bin_rest_trunc_i: "(take_bit :: nat \ int \ int) n ((\k::int. k div 2) bin) = (\k::int. k div 2) ((take_bit :: nat \ int \ int) (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)" +lemma bin_rest_strunc: "(\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) (Suc n) bin) = (signed_take_bit :: nat \ int \ int) n ((\k::int. k div 2) bin)" by (simp add: signed_take_bit_Suc) -lemma bintrunc_rest [simp]: "bintrunc n (bin_rest (bintrunc n bin)) = bin_rest (bintrunc n bin)" +lemma bintrunc_rest [simp]: "(take_bit :: nat \ int \ int) n ((\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin)) = (\k::int. k div 2) ((take_bit :: nat \ int \ int) 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)" +lemma sbintrunc_rest [simp]: "(signed_take_bit :: nat \ int \ int) n ((\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) n bin)) = (\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) 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" +lemma bintrunc_rest': "(take_bit :: nat \ int \ int) n \ (\k::int. k div 2) \ (take_bit :: nat \ int \ int) n = (\k::int. k div 2) \ (take_bit :: nat \ int \ int) n" by (rule ext) auto -lemma sbintrunc_rest': "sbintrunc n \ bin_rest \ sbintrunc n = bin_rest \ sbintrunc n" +lemma sbintrunc_rest': "(signed_take_bit :: nat \ int \ int) n \ (\k::int. k div 2) \ (signed_take_bit :: nat \ int \ int) n = (\k::int. k div 2) \ (signed_take_bit :: nat \ int \ int) n" by (rule ext) auto lemma rco_lem: "f \ g \ f = g \ f \ f \ (g \ f) ^^ n = g ^^ n \ f" @@ -441,14 +425,11 @@ lemma [code]: "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\ + \concat_bit n l k = 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" +lemma bin_sign_cat: "bin_sign ((\k n l. concat_bit n l k) x n y) = bin_sign x" proof - have \0 \ x\ if \0 \ x * 2 ^ n + y mod 2 ^ n\ proof - @@ -472,20 +453,20 @@ proof - 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)" +lemma bin_cat_assoc: "(\k n l. concat_bit n l k) ((\k n l. concat_bit n l k) x m y) n z = (\k n l. concat_bit n l k) x (m + n) ((\k n l. concat_bit n l k) 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" +lemma bin_cat_assoc_sym: "(\k n l. concat_bit n l k) x m ((\k n l. concat_bit n l k) y n z) = (\k n l. concat_bit n l k) ((\k n l. concat_bit n l k) 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\ + \bin_rcat n = foldl (\u v. (\k n l. concat_bit n l k) u n v) 0\ proof fix ks :: \int list\ - show \bin_rcat n ks = foldl (\u v. bin_cat u n v) 0 ks\ + show \bin_rcat n ks = foldl (\u v. (\k n l. concat_bit n l k) 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 @@ -500,8 +481,6 @@ fun bin_rsplit_aux :: "nat \ nat \ int \ int 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 @@ -516,68 +495,67 @@ 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))" + "(bit :: int \ nat \ bool) ((\k n l. concat_bit n l k) x k y) n = + (if n < k then (bit :: int \ nat \ bool) y n else (bit :: int \ nat \ bool) 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)\ + \(bit :: int \ nat \ bool) (drop_bit n c) k \ (bit :: int \ nat \ bool) 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\ + \(bit :: int \ nat \ bool) (take_bit n c) k \ k < n \ (bit :: int \ nat \ bool) 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))" + (\k. (bit :: int \ nat \ bool) a k = (bit :: int \ nat \ bool) c (n + k)) \ + (\k. (bit :: int \ nat \ bool) b k = (k < n \ (bit :: int \ nat \ bool) 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" +lemma bin_cat_zero [simp]: "(\k n l. concat_bit n l k) 0 n w = (take_bit :: nat \ int \ int) 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" +lemma bintr_cat1: "(take_bit :: nat \ int \ int) (k + n) ((\k n l. concat_bit n l k) a n b) = (\k n l. concat_bit n l k) ((take_bit :: nat \ int \ int) 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)" - +lemma bintr_cat: "(take_bit :: nat \ int \ int) m ((\k n l. concat_bit n l k) a n b) = + (\k n l. concat_bit n l k) ((take_bit :: nat \ int \ int) (m - n) a) n ((take_bit :: nat \ int \ int) (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" +lemma bintr_cat_same [simp]: "(take_bit :: nat \ int \ int) n ((\k n l. concat_bit n l k) a n b) = (take_bit :: nat \ int \ int) n b" by (auto simp add : bintr_cat) -lemma cat_bintr [simp]: "bin_cat a n (bintrunc n b) = bin_cat a n b" +lemma cat_bintr [simp]: "(\k n l. concat_bit n l k) a n ((take_bit :: nat \ int \ int) n b) = (\k n l. concat_bit n l k) 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" +lemma split_bintrunc: "bin_split n c = (a, b) \ b = (take_bit :: nat \ int \ int) n c" by simp -lemma bin_cat_split: "bin_split n w = (u, v) \ w = bin_cat u n v" +lemma bin_cat_split: "bin_split n w = (u, v) \ w = (\k n l. concat_bit n l k) 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\ + \drop_bit n ((\k n l. concat_bit n l k) 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\ + \take_bit n ((\k n l. concat_bit n l k) 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)" +lemma bin_split_cat: "bin_split n ((\k n l. concat_bit n l k) v n w) = (v, (take_bit :: nat \ int \ int) 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))" + "bin_split n (- 1) = (- 1, (take_bit :: nat \ int \ int) 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)" + bin_split n ((take_bit :: nat \ int \ int) m c) = ((take_bit :: nat \ int \ int) (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) @@ -586,14 +564,14 @@ lemma bin_split_trunc: lemma bin_split_trunc1: "bin_split n c = (a, b) \ - bin_split n (bintrunc m c) = (bintrunc (m - n) a, bintrunc m b)" + bin_split n ((take_bit :: nat \ int \ int) m c) = ((take_bit :: nat \ int \ int) (m - n) a, (take_bit :: nat \ int \ int) 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" +lemma bin_cat_num: "(\k n l. concat_bit n l k) a n b = a * 2 ^ n + (take_bit :: nat \ int \ int) 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)" @@ -643,7 +621,7 @@ lemma bin_split_minus: "0 < n \ bin_split (Suc (n - 1)) w = bin_ 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) + (let (w1, w2) = bin_split (numeral bin - 1) ((\k::int. k div 2) w) in (w1, of_bool (odd w) + 2 * w2))" by (simp add: take_bit_rec drop_bit_rec mod_2_eq_odd) @@ -662,7 +640,7 @@ lemmas bin_rsplit_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" + "n > 0 \ rev sw = bin_rsplit n (nw, w) \ \v\set sw. (take_bit :: nat \ int \ int) n v = v" apply (induct sw arbitrary: nw w) apply clarsimp apply clarsimp @@ -679,7 +657,7 @@ 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)" + k < size sw \ (bit :: int \ nat \ bool) (sw ! k) m = (bit :: int \ nat \ bool) w (k * n + m)" apply (induct sw) apply clarsimp apply clarsimp @@ -696,11 +674,11 @@ lemma bin_nth_rsplit [rule_format] : 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]" +lemma bin_rsplit_all: "0 < nw \ nw \ n \ bin_rsplit n (nw, w) = [(take_bit :: nat \ int \ int) 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)" + "\bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, (take_bit :: nat \ int \ int) 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) @@ -718,7 +696,7 @@ lemma bin_rsplit_l [rule_format]: done lemma bin_rsplit_rcat [rule_format]: - "n > 0 \ bin_rsplit n (n * size ws, bin_rcat n ws) = map (bintrunc n) ws" + "n > 0 \ bin_rsplit n (n * size ws, bin_rcat n ws) = map ((take_bit :: nat \ int \ int) n) ws" apply (unfold bin_rsplit_def bin_rcat_eq_foldl) apply (rule_tac xs = ws in rev_induct) apply clarsimp @@ -810,13 +788,19 @@ lemma bin_rsplit_len_indep: 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)" +abbreviation (input) bin_sc :: \nat \ bool \ int \ int\ + where \bin_sc n b i \ set_bit i n b\ -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_0 [simp]: + "bin_sc 0 b w = of_bool b + 2 * (\k::int. k div 2) w" + by (simp add: set_bit_int_def) + +lemma bin_sc_Suc [simp]: + "bin_sc (Suc n) b w = of_bool (odd w) + 2 * bin_sc n b (w div 2)" + by (simp add: set_bit_int_def set_bit_Suc unset_bit_Suc bin_last_def) + +lemma bin_nth_sc [bit_simps]: "bit (bin_sc n b w) n \ b" + by (simp add: bit_simps) 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) @@ -827,7 +811,7 @@ lemma bin_sc_sc_diff: "m \ n \ bin_sc m c (bin_sc n b w) 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)" +lemma bin_nth_sc_gen: "(bit :: int \ nat \ bool) (bin_sc n b w) m = (if m = n then b else (bit :: int \ nat \ bool) w m)" apply (induct n arbitrary: w m) apply (case_tac m; simp add: bit_Suc) apply (case_tac m; simp add: bit_Suc) @@ -836,10 +820,11 @@ lemma bin_nth_sc_gen: "bin_nth (bin_sc n b w) m = (if m = n then b else bin_nth 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) + apply (simp_all add: fun_eq_iff bit_eq_iff) + apply (simp_all add: bit_simps bin_nth_sc_gen) + done -lemma bin_sc_nth [simp]: "bin_sc n (bin_nth w n) w = w" +lemma bin_sc_nth [simp]: "bin_sc n ((bit :: int \ nat \ bool) 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" @@ -854,23 +839,23 @@ next qed lemma bin_sc_bintr [simp]: - "bintrunc m (bin_sc n x (bintrunc m w)) = bintrunc m (bin_sc n x w)" + "(take_bit :: nat \ int \ int) m (bin_sc n x ((take_bit :: nat \ int \ int) m w)) = (take_bit :: nat \ int \ int) m (bin_sc n x w)" + apply (rule bit_eqI) 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) + apply (auto simp add: bit_simps bin_sc_eq) done lemma bin_clr_le: "bin_sc n False w \ w" - by (simp add: bin_sc_eq unset_bit_less_eq) + by (simp add: set_bit_int_def unset_bit_less_eq) lemma bin_set_ge: "bin_sc n True w \ w" - by (simp add: bin_sc_eq set_bit_greater_eq) + by (simp add: set_bit_int_def 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_clr_le: "(take_bit :: nat \ int \ int) n (bin_sc m False w) \ (take_bit :: nat \ int \ int) n w" + by (simp add: set_bit_int_def 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 bintr_bin_set_ge: "(take_bit :: nat \ int \ int) n (bin_sc m True w) \ (take_bit :: nat \ int \ int) n w" + by (simp add: set_bit_int_def take_bit_set_bit_eq set_bit_greater_eq) lemma bin_sc_FP [simp]: "bin_sc n False 0 = 0" by (induct n) auto @@ -878,13 +863,13 @@ lemma bin_sc_FP [simp]: "bin_sc n False 0 = 0" 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 +lemmas bin_sc_simps = bin_sc_0 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] + trans [OF bin_sc_minus [symmetric] bin_sc_Suc] lemma bin_sc_numeral [simp]: "bin_sc (numeral k) b w = @@ -894,26 +879,89 @@ lemma bin_sc_numeral [simp]: lemmas bin_sc_minus_simps = bin_sc_simps (2,3,4) [THEN [2] trans, OF bin_sc_minus [THEN sym]] -instance int :: semiring_bit_syntax .. +lemma int_set_bit_0 [simp]: fixes x :: int shows + "set_bit x 0 b = of_bool b + 2 * (x div 2)" + by (fact bin_sc_0) -lemma test_bit_int_def [iff]: - "i !! n \ bin_nth i n" - by (simp add: test_bit_eq_bit) +lemma int_set_bit_Suc: fixes x :: int shows + "set_bit x (Suc n) b = of_bool (odd x) + 2 * set_bit (x div 2) n b" + by (fact bin_sc_Suc) + +lemma bin_last_set_bit: + "odd (set_bit x n b :: int) = (if n > 0 then odd x else b)" + by (cases n) (simp_all add: int_set_bit_Suc) + +lemma bin_rest_set_bit: + "(set_bit x n b :: int) div 2 = (if n > 0 then set_bit (x div 2) (n - 1) b else x div 2)" + by (cases n) (simp_all add: int_set_bit_Suc) + +lemma int_set_bit_numeral: fixes x :: int shows + "set_bit x (numeral w) b = of_bool (odd x) + 2 * set_bit (x div 2) (pred_numeral w) b" + by (fact bin_sc_numeral) + +lemmas int_set_bit_numerals [simp] = + int_set_bit_numeral[where x="numeral w'"] + int_set_bit_numeral[where x="- numeral w'"] + int_set_bit_numeral[where x="Numeral1"] + int_set_bit_numeral[where x="1"] + int_set_bit_numeral[where x="0"] + int_set_bit_Suc[where x="numeral w'"] + int_set_bit_Suc[where x="- numeral w'"] + int_set_bit_Suc[where x="Numeral1"] + int_set_bit_Suc[where x="1"] + int_set_bit_Suc[where x="0"] + for w' + +lemma msb_set_bit [simp]: + "msb (set_bit (x :: int) n b) \ msb x" + by (simp add: msb_int_def set_bit_int_def) + +lemma word_set_bit_def: + \set_bit a n x = word_of_int (bin_sc n x (uint a))\ + apply (rule bit_word_eqI) + apply (cases x) + apply (simp_all add: bit_simps bin_sc_eq) + done + +lemma set_bit_word_of_int: + "set_bit (word_of_int x) n b = word_of_int (bin_sc n b x)" + unfolding word_set_bit_def + by (rule word_eqI) (simp add: word_size bin_nth_sc_gen nth_bintr bit_simps) + +lemma word_set_numeral [simp]: + "set_bit (numeral bin::'a::len word) n b = + word_of_int (bin_sc n b (numeral bin))" + unfolding word_numeral_alt by (rule set_bit_word_of_int) + +lemma word_set_neg_numeral [simp]: + "set_bit (- numeral bin::'a::len word) n b = + word_of_int (bin_sc n b (- numeral bin))" + unfolding word_neg_numeral_alt by (rule set_bit_word_of_int) + +lemma word_set_bit_0 [simp]: "set_bit 0 n b = word_of_int (bin_sc n b 0)" + unfolding word_0_wi by (rule set_bit_word_of_int) + +lemma word_set_bit_1 [simp]: "set_bit 1 n b = word_of_int (bin_sc n b 1)" + unfolding word_1_wi by (rule set_bit_word_of_int) lemma shiftl_int_def: - "shiftl x n = x * 2 ^ n" for x :: int - by (simp add: push_bit_int_def shiftl_eq_push_bit) + "push_bit n x = x * 2 ^ n" for x :: int + by (fact push_bit_eq_mult) 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) + "drop_bit n x = x div 2 ^ n" for x :: int + by (fact drop_bit_eq_div) subsubsection \Basic simplification rules\ +context + includes bit_operations_syntax +begin + lemmas int_not_def = not_int_def -lemma int_not_simps [simp]: +lemma int_not_simps: "NOT (0::int) = -1" "NOT (1::int) = -2" "NOT (- 1::int) = 0" @@ -932,11 +980,11 @@ lemma int_and_0 [simp]: "0 AND x = 0" lemma int_and_m1 [simp]: "-1 AND x = x" for x :: int - by (fact bit.conj_one_left) + by (fact and.left_neutral) lemma int_or_zero [simp]: "0 OR x = x" for x :: int - by (fact bit.disj_zero_left) + by (fact or.left_neutral) lemma int_or_minus1 [simp]: "-1 OR x = -1" for x :: int @@ -944,40 +992,40 @@ lemma int_or_minus1 [simp]: "-1 OR x = -1" lemma int_xor_zero [simp]: "0 XOR x = x" for x :: int - by (fact bit.xor_zero_left) + by (fact xor.left_neutral) subsubsection \Binary destructors\ -lemma bin_rest_NOT [simp]: "bin_rest (NOT x) = NOT (bin_rest x)" +lemma bin_rest_NOT [simp]: "(\k::int. k div 2) (NOT x) = NOT ((\k::int. k div 2) x)" by (fact not_int_div_2) -lemma bin_last_NOT [simp]: "bin_last (NOT x) \ \ bin_last x" +lemma bin_last_NOT [simp]: "(odd :: int \ bool) (NOT x) \ \ (odd :: int \ bool) x" by simp -lemma bin_rest_AND [simp]: "bin_rest (x AND y) = bin_rest x AND bin_rest y" +lemma bin_rest_AND [simp]: "(\k::int. k div 2) (x AND y) = (\k::int. k div 2) x AND (\k::int. k div 2) y" by (subst and_int_rec) auto -lemma bin_last_AND [simp]: "bin_last (x AND y) \ bin_last x \ bin_last y" +lemma bin_last_AND [simp]: "(odd :: int \ bool) (x AND y) \ (odd :: int \ bool) x \ (odd :: int \ bool) y" by (subst and_int_rec) auto -lemma bin_rest_OR [simp]: "bin_rest (x OR y) = bin_rest x OR bin_rest y" +lemma bin_rest_OR [simp]: "(\k::int. k div 2) (x OR y) = (\k::int. k div 2) x OR (\k::int. k div 2) y" by (subst or_int_rec) auto -lemma bin_last_OR [simp]: "bin_last (x OR y) \ bin_last x \ bin_last y" +lemma bin_last_OR [simp]: "(odd :: int \ bool) (x OR y) \ (odd :: int \ bool) x \ (odd :: int \ bool) y" by (subst or_int_rec) auto -lemma bin_rest_XOR [simp]: "bin_rest (x XOR y) = bin_rest x XOR bin_rest y" +lemma bin_rest_XOR [simp]: "(\k::int. k div 2) (x XOR y) = (\k::int. k div 2) x XOR (\k::int. k div 2) 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)" +lemma bin_last_XOR [simp]: "(odd :: int \ bool) (x XOR y) \ ((odd :: int \ bool) x \ (odd :: int \ bool) y) \ \ ((odd :: int \ bool) x \ (odd :: int \ bool) 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" + "\x y. (bit :: int \ nat \ bool) (x AND y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" + "\x y. (bit :: int \ nat \ bool) (x OR y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" + "\x y. (bit :: int \ nat \ bool) (x XOR y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" + "\x. (bit :: int \ nat \ bool) (NOT x) n \ \ (bit :: int \ nat \ bool) x n" by (simp_all add: bit_and_iff bit_or_iff bit_xor_iff bit_not_iff) @@ -1080,22 +1128,17 @@ 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" + "(\k::int. k div 2) (- numeral (Num.BitM w)) = - numeral w" by simp lemma bin_last_neg_numeral_BitM [simp]: - "bin_last (- numeral (Num.BitM w))" + "(odd :: int \ bool) (- numeral (Num.BitM w))" by simp @@ -1120,18 +1163,18 @@ lemma AND_mod: "x AND (2 ^ n - 1) = x mod 2 ^ n" 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)" + "(take_bit :: nat \ int \ int) n x AND (take_bit :: nat \ int \ int) n y = (take_bit :: nat \ int \ int) n (x AND y)" + "(take_bit :: nat \ int \ int) n x OR (take_bit :: nat \ int \ int) n y = (take_bit :: nat \ int \ int) 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)" +lemma bin_trunc_xor: "(take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) n x XOR (take_bit :: nat \ int \ int) n y) = (take_bit :: nat \ int \ int) n (x XOR y)" by simp -lemma bin_trunc_not: "bintrunc n (NOT (bintrunc n x)) = bintrunc n (NOT x)" +lemma bin_trunc_not: "(take_bit :: nat \ int \ int) n (NOT ((take_bit :: nat \ int \ int) n x)) = (take_bit :: nat \ int \ int) 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" +lemma bintr_bintr_i: "x = (take_bit :: nat \ int \ int) n y \ (take_bit :: nat \ int \ int) n x = (take_bit :: nat \ int \ int) n y" by auto lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i] @@ -1198,11 +1241,11 @@ lemma even_conv_AND: by (simp add: and_one_eq mod2_eq_if) lemma bin_last_conv_AND: - "bin_last i \ i AND 1 \ 0" + "(odd :: int \ bool) 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" + "of_bool ((odd :: int \ bool) i) = i AND 1" by (simp add: and_one_eq mod2_eq_if) lemma bin_sign_and: @@ -1219,95 +1262,95 @@ 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)" + shows int_shiftl0: "push_bit 0 x = x" + and int_shiftl_Suc: "push_bit (Suc n) x = 2 * push_bit n x" by (auto simp add: shiftl_int_def) -lemma int_0_shiftl [simp]: "0 << n = (0 :: int)" -by(induct n) simp_all +lemma int_0_shiftl: "push_bit n 0 = (0 :: int)" + by (fact push_bit_of_0) -lemma bin_last_shiftl: "bin_last (x << n) \ n = 0 \ bin_last x" -by(cases n)(simp_all) +lemma bin_last_shiftl: "odd (push_bit n x) \ n = 0 \ (odd :: int \ bool) x" + by simp -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_rest_shiftl: "(\k::int. k div 2) (push_bit n x) = (if n > 0 then push_bit (n - 1) x else (\k::int. k div 2) x)" + by (cases n) (simp_all add: push_bit_eq_mult) -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_nth_shiftl: "(bit :: int \ nat \ bool) (push_bit n x) m \ n \ m \ (bit :: int \ nat \ bool) x (m - n)" + by (fact bit_push_bit_iff_int) -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_last_shiftr: "odd (drop_bit n x) \ bit x n" for x :: int + by (simp add: 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_rest_shiftr: "(\k::int. k div 2) (drop_bit n x) = drop_bit (Suc n) x" + by (simp add: drop_bit_Suc 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_shiftr: "(bit :: int \ nat \ bool) (drop_bit n x) m = (bit :: int \ nat \ bool) x (n + m)" + by (simp add: bit_simps) 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) + "(bit :: int \ nat \ bool) x n \ x AND (push_bit n 1) \ 0" + by (fact bit_iff_and_push_bit_not_eq_0) 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'" + "push_bit (numeral w') (numeral w :: int) = push_bit (pred_numeral w') (numeral (num.Bit0 w))" + "push_bit (numeral w') (- numeral w :: int) = push_bit (pred_numeral w') (- numeral (num.Bit0 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 + "push_bit (numeral w) (1::int) = push_bit (pred_numeral w) 2" + using int_shiftl_numeral [of Num.One w] + by (simp only: numeral_eq_Suc push_bit_Suc) simp -lemma shiftl_ge_0 [simp]: fixes i :: int shows "i << n \ 0 \ i \ 0" -by(induct n) simp_all +lemma shiftl_ge_0: fixes i :: int shows "push_bit n i \ 0 \ i \ 0" + by (fact push_bit_nonnegative_int_iff) -lemma shiftl_lt_0 [simp]: fixes i :: int shows "i << n < 0 \ i < 0" -by (metis not_le shiftl_ge_0) +lemma shiftl_lt_0: fixes i :: int shows "push_bit n i < 0 \ i < 0" + by (fact push_bit_negative_int_iff) -lemma int_shiftl_test_bit: "(n << i :: int) !! m \ m \ i \ n !! (m - i)" - by simp +lemma int_shiftl_test_bit: "bit (push_bit i n :: int) m \ m \ i \ bit n (m - i)" + by (fact bit_push_bit_iff_int) -lemma int_0shiftr [simp]: "(0 :: int) >> x = 0" -by(simp add: shiftr_int_def) +lemma int_0shiftr: "drop_bit x (0 :: int) = 0" + by (fact drop_bit_of_0) -lemma int_minus1_shiftr [simp]: "(-1 :: int) >> x = -1" -by(simp add: shiftr_int_def div_eq_minus1) +lemma int_minus1_shiftr: "drop_bit x (-1 :: int) = -1" + by (fact drop_bit_minus_one) -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_ge_0: fixes i :: int shows "drop_bit n i \ 0 \ i \ 0" + by (fact drop_bit_nonnegative_int_iff) -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_lt_0 [simp]: fixes i :: int shows "drop_bit n i < 0 \ i < 0" + by (fact drop_bit_negative_int_iff) 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) + "drop_bit (numeral w') (1 :: int) = 0" + "drop_bit (numeral w') (numeral num.One :: int) = 0" + "drop_bit (numeral w') (numeral (num.Bit0 w) :: int) = drop_bit (pred_numeral w') (numeral w)" + "drop_bit (numeral w') (numeral (num.Bit1 w) :: int) = drop_bit (pred_numeral w') (numeral w)" + "drop_bit (numeral w') (- numeral (num.Bit0 w) :: int) = drop_bit (pred_numeral w') (- numeral w)" + "drop_bit (numeral w') (- numeral (num.Bit1 w) :: int) = drop_bit (pred_numeral w') (- numeral (Num.inc w))" + by (simp_all add: 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) + "drop_bit (Suc 0) (1 :: int) = 0" + "drop_bit (Suc 0) (numeral num.One :: int) = 0" + "drop_bit (Suc 0) (numeral (num.Bit0 w) :: int) = numeral w" + "drop_bit (Suc 0) (numeral (num.Bit1 w) :: int) = numeral w" + "drop_bit (Suc 0) (- numeral (num.Bit0 w) :: int) = - numeral w" + "drop_bit (Suc 0) (- numeral (num.Bit1 w) :: int) = - numeral (Num.inc w)" + by (simp_all add: drop_bit_Suc add_One) lemma bin_nth_minus_p2: assumes sign: "bin_sign x = 0" - and y: "y = 1 << n" + and y: "y = push_bit n 1" and m: "m < n" and x: "x < y" - shows "bin_nth (x - y) m = bin_nth x m" + shows "bit (x - y) m = bit 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) + by (simp_all add: bin_sign_def 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 @@ -1331,16 +1374,35 @@ proof - 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)+ + "bin_sc n False i = i AND NOT (push_bit n 1)" + by (rule bit_eqI) (auto simp add: bin_sc_eq bit_simps) lemma bin_set_conv_OR: - "bin_sc n True i = i OR (1 << n)" - by (induct n arbitrary: i) (rule bin_rl_eqI; simp)+ + "bin_sc n True i = i OR (push_bit n 1)" + by (rule bit_eqI) (auto simp add: bin_sc_eq bit_simps) + +end subsection \More lemmas on words\ +lemma msb_conv_bin_sign: + "msb x \ bin_sign x = -1" + by (simp add: bin_sign_def not_le msb_int_def) + +lemma msb_bin_sc: + "msb (bin_sc n b x) \ msb x" + by (simp add: msb_conv_bin_sign) + +lemma msb_word_def: + \msb a \ bin_sign (signed_take_bit (LENGTH('a) - 1) (uint a)) = - 1\ + for a :: \'a::len word\ + by (simp add: bin_sign_def bit_simps msb_word_iff_bit) + +lemma word_msb_def: + "msb a \ bin_sign (sint a) = - 1" + by (simp add: msb_word_def sint_uint) + lemma word_rcat_eq: \word_rcat ws = word_of_int (bin_rcat (LENGTH('a::len)) (map uint ws))\ for ws :: \'a::len word list\ @@ -1364,12 +1426,12 @@ lemmas wils1 = bin_log_bintrs [THEN word_of_int_eq_iff [THEN iffD2], 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 setBit_no: "Bit_Operations.set_bit n (numeral bin) = word_of_int (bin_sc n True (numeral bin))" + by (rule bit_word_eqI) (simp add: bit_simps) -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 clearBit_no: + "unset_bit n (numeral bin) = word_of_int (bin_sc n False (numeral bin))" + by (rule bit_word_eqI) (simp add: bit_simps) lemma eq_mod_iff: "0 < n \ b = b mod n \ 0 \ b \ b < n" for b n :: int @@ -1384,17 +1446,17 @@ lemma split_uint_lem: "bin_split n (uint w) = (a, b) \ 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))" + word_of_int ((\k n l. concat_bit n l k) 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" + "take_bit n (push_bit i m) = push_bit i (take_bit (n - i) m)" for m :: int - by (rule bit_eqI) (auto simp add: bit_take_bit_iff) + by (fact take_bit_push_bit) 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) + "uint (push_bit i n) = take_bit (size n) (push_bit i (uint n))" + by (simp add: unsigned_push_bit_eq word_size) lemma bin_mask_conv_pow2: "mask n = 2 ^ n - (1 :: int)" @@ -1403,33 +1465,39 @@ lemma bin_mask_conv_pow2: lemma bin_mask_ge0: "mask n \ (0 :: int)" by (fact mask_nonnegative_int) +context + includes bit_operations_syntax +begin + lemma and_bin_mask_conv_mod: "x AND mask n = x mod 2 ^ n" for x :: int by (simp flip: take_bit_eq_mod add: take_bit_eq_mask) +end + lemma bin_mask_numeral: "mask (numeral n) = (1 :: int) + 2 * mask (pred_numeral n)" by (fact mask_numeral) -lemma bin_nth_mask [simp]: "bit (mask n :: int) i \ i < n" +lemma bin_nth_mask: "bit (mask n :: int) i \ i < n" by (simp add: bit_mask_iff) lemma bin_sign_mask [simp]: "bin_sign (mask n) = 0" by (simp add: bin_sign_def bin_mask_conv_pow2) -lemma bin_mask_p1_conv_shift: "mask n + 1 = (1 :: int) << n" +lemma bin_mask_p1_conv_shift: "mask n + 1 = push_bit n (1 :: int)" by (simp add: bin_mask_conv_pow2 shiftl_int_def) lemma sbintrunc_eq_in_range: - "(sbintrunc n x = x) = (x \ range (sbintrunc n))" - "(x = sbintrunc n x) = (x \ range (sbintrunc n))" + "((signed_take_bit :: nat \ int \ int) n x = x) = (x \ range ((signed_take_bit :: nat \ int \ int) n))" + "(x = (signed_take_bit :: nat \ int \ int) n x) = (x \ range ((signed_take_bit :: nat \ int \ int) n))" apply (simp_all add: image_def) apply (metis sbintrunc_sbintrunc)+ done lemma sbintrunc_If: "- 3 * (2 ^ n) \ x \ x < 3 * (2 ^ n) - \ sbintrunc n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n) + \ (signed_take_bit :: nat \ int \ int) n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n) else if x \ 2 ^ n then x - 2 * (2 ^ n) else x)" apply (simp add: no_sbintr_alt2, safe) apply (simp add: mod_pos_geq) @@ -1458,7 +1526,7 @@ lemma signed_mult_eq_checks_double_size: shows "(sint (a :: 'a :: len word) * sint b = sint (a * b)) = (scast a * scast b = (scast (a * b) :: 'b :: len word))" proof - - have P: "sbintrunc (size a - 1) (sint a * sint b) \ range (sbintrunc (size a - 1))" + have P: "(signed_take_bit :: nat \ int \ int) (size a - 1) (sint a * sint b) \ range ((signed_take_bit :: nat \ int \ int) (size a - 1))" by simp have abs: "!! x :: 'a word. abs (sint x) < 2 ^ (size a - 1) + 1" @@ -1480,6 +1548,24 @@ proof - done qed +lemma bintrunc_id: + "\m \ int n; 0 < m\ \ take_bit n m = m" + by (simp add: take_bit_int_eq_self_iff le_less_trans) + +lemma bin_cat_cong: "concat_bit n b a = concat_bit m d c" + if "n = m" "a = c" "take_bit m b = take_bit m d" + using that(3) unfolding that(1,2) + by (simp add: bin_cat_eq_push_bit_add_take_bit) + +lemma bin_cat_eqD1: "concat_bit n b a = concat_bit n d c \ a = c" + by (metis drop_bit_bin_cat_eq) + +lemma bin_cat_eqD2: "concat_bit n b a = concat_bit n d c \ take_bit n b = take_bit n d" + by (metis take_bit_bin_cat_eq) + +lemma bin_cat_inj: "(concat_bit n b a) = concat_bit n d c \ a = c \ take_bit n b = take_bit n d" + by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) + lemma bin_sc_pos: "0 \ i \ 0 \ bin_sc n b i" by (metis bin_sign_sc sign_Pls_ge_0) diff --git a/lib/Word_Lib/Bitwise.thy b/lib/Word_Lib/Bitwise.thy index 29de7a9f3..54f7088a6 100644 --- a/lib/Word_Lib/Bitwise.thy +++ b/lib/Word_Lib/Bitwise.thy @@ -9,6 +9,7 @@ theory Bitwise "HOL-Library.Word" More_Arithmetic Reversed_Bit_Lists + Bit_Shifts_Infix_Syntax begin text \Helper constants used in defining addition\ @@ -101,7 +102,7 @@ lemma rbl_succ2_simps: "rbl_succ2 b (x # xs) = (b \ x) # rbl_succ2 (x \ b) xs" by (simp_all add: rbl_succ2_def) -lemma twos_complement: "- x = word_succ (NOT x)" +lemma twos_complement: "- x = word_succ (not x)" using arg_cong[OF word_add_not[where x=x], where f="\a. a - x + 1"] by (simp add: word_succ_p1 word_sp_01[unfolded word_succ_p1] del: word_add_not) @@ -165,22 +166,22 @@ lemma rbl_sshiftr: drop_nonempty_def take_Cons') apply (case_tac "LENGTH('a)", simp_all) apply (rule word_eqI) - apply (simp add: nth_sshiftr word_size test_bit_of_bl + apply (simp add: bit_simps word_size test_bit_of_bl msb_nth) done lemma nth_word_of_int: - "(word_of_int x :: 'a::len word) !! n = (n < LENGTH('a) \ bin_nth x n)" + "bit (word_of_int x :: 'a::len word) n = (n < LENGTH('a) \ bit x n)" apply (simp add: test_bit_bl word_size to_bl_of_bin) apply (subst conj_cong[OF refl], erule bin_nth_bl) apply auto done lemma nth_scast: - "(scast (x :: 'a::len word) :: 'b::len word) !! n = + "bit (scast (x :: 'a::len word) :: 'b::len word) n = (n < LENGTH('b) \ - (if n < LENGTH('a) - 1 then x !! n - else x !! (LENGTH('a) - 1)))" + (if n < LENGTH('a) - 1 then bit x n + else bit x (LENGTH('a) - 1)))" apply transfer apply (auto simp add: bit_signed_take_bit_iff min_def) done @@ -235,7 +236,7 @@ next by (cases z) (simp cong: map_cong, simp add: map_replicate_const cong: map_cong) have shiftl: "of_bl xs * 2 * y = (of_bl xs * y) << 1" for xs - by (simp add: shiftl_t2n) + by (simp add: push_bit_eq_mult shiftl_def) have zip_take_triv: "\xs ys n. n = length ys \ zip (take n xs) ys = zip xs ys" by (rule nth_equalityI) simp_all @@ -244,7 +245,10 @@ next apply (simp add: trans [OF of_bl_append add.commute] rbl_mul_simps rbl_word_plus' distrib_right mult_bit shiftl rbl_shiftl) apply (simp add: takefill_alt word_size rev_map take_rbl_plus min_def) - apply (simp add: rbl_plus_def zip_take_triv) + apply (simp add: rbl_plus_def) + apply (simp add: zip_take_triv) + apply (simp only: mult.commute [of _ 2] to_bl_double_eq) + apply (simp flip: butlast_rev add: take_butlast) done qed @@ -349,13 +353,10 @@ lemma rev_bin_to_bl_simps: False # rev (bin_to_bl n (- numeral num.One))" by (simp_all add: bin_to_bl_aux_append bin_to_bl_zero_aux bin_to_bl_minus1_aux replicate_append_same) -lemma to_bl_upt: "to_bl x = rev (map ((!!) x) [0 ..< size x])" - apply (rule nth_equalityI) - apply (simp add: word_size) - apply (auto simp: to_bl_nth word_size rev_nth) - done +lemma to_bl_upt: "to_bl x = rev (map (bit x) [0 ..< size x])" + by (simp add: to_bl_eq_rev word_size rev_map) -lemma rev_to_bl_upt: "rev (to_bl x) = map ((!!) x) [0 ..< size x]" +lemma rev_to_bl_upt: "rev (to_bl x) = map (bit x) [0 ..< size x]" by (simp add: to_bl_upt) lemma upt_eq_list_intros: @@ -383,7 +384,7 @@ fun mk_nat_clist ns = fun upt_conv ctxt ct = case Thm.term_of ct of - (\<^const>\upt\ $ n $ m) => + \<^Const_>\upt for n m\ => let val (i, j) = apply2 (snd o HOLogic.dest_number) (n, m); val ns = map (Numeral.mk_cnumber \<^ctyp>\nat\) (i upto (j - 1)) @@ -405,7 +406,7 @@ val expand_upt_simproc = fun word_len_simproc_fn ctxt ct = (case Thm.term_of ct of - Const (\<^const_name>\len_of\, _) $ t => + \<^Const_>\len_of _ for t\ => (let val T = fastype_of t |> dest_Type |> snd |> the_single val n = Numeral.mk_cnumber \<^ctyp>\nat\ (Word_Lib.dest_binT T); diff --git a/lib/Word_Lib/Enumeration_Word.thy b/lib/Word_Lib/Enumeration_Word.thy index b29b49b7c..8550ac4d5 100644 --- a/lib/Word_Lib/Enumeration_Word.thy +++ b/lib/Word_Lib/Enumeration_Word.thy @@ -63,7 +63,7 @@ lemma minBound_word: by (simp add: minBound_def enum_word_def upt_conv_Cons) lemma maxBound_max_word: - "(maxBound::'a::len word) = max_word" + "(maxBound::'a::len word) = - 1" by (fact maxBound_word) lemma leq_maxBound [simp]: @@ -295,4 +295,35 @@ lemma enum_word_div: apply simp done +lemma remdups_enum_upto: + fixes s::"'a::len word" + shows "remdups [s .e. e] = [s .e. e]" + by simp + +lemma card_enum_upto: + fixes s::"'a::len word" + shows "card (set [s .e. e]) = Suc (unat e) - unat s" + by (subst List.card_set) (simp add: remdups_enum_upto) + +lemma length_upto_enum_one: + fixes x :: "'a :: len word" + assumes lt1: "x < y" and lt2: "z < y" and lt3: "x \ z" + shows "[x , y .e. z] = [x]" + unfolding upto_enum_step_def +proof (subst upto_enum_red, subst if_not_P [OF leD [OF lt3]], clarsimp, rule conjI) + show "unat ((z - x) div (y - x)) = 0" + proof (subst unat_div, rule div_less) + have syx: "unat (y - x) = unat y - unat x" + by (rule unat_sub [OF order_less_imp_le]) fact + moreover have "unat (z - x) = unat z - unat x" + by (rule unat_sub) fact + + ultimately show "unat (z - x) < unat (y - x)" + using lt2 lt3 unat_mono word_less_minus_mono_left by blast + qed + + then show "(z - x) div (y - x) * (y - x) = 0" + by (simp add: unat_div) (simp add: word_arith_nat_defs(6)) +qed + end diff --git a/lib/Word_Lib/Examples.thy b/lib/Word_Lib/Examples.thy index 4f8c62d3c..39e4ad9e7 100644 --- a/lib/Word_Lib/Examples.thy +++ b/lib/Word_Lib/Examples.thy @@ -5,9 +5,970 @@ *) theory Examples - imports Bitwise Next_and_Prev Generic_set_bit Word_Syntax Signed_Division_Word + imports + Bit_Shifts_Infix_Syntax Next_and_Prev Signed_Division_Word Bitwise begin +context + includes bit_operations_syntax +begin + +section \\<^typ>\nat\\ + +lemma + \bit (1705 :: nat) (Suc (Suc (Suc 0)))\ + by simp + +lemma + \bit (1705 :: nat) 3\ + by simp + +lemma + \\ bit (1 :: nat) (Suc (Suc (Suc 0)))\ + by simp + +lemma + \\ bit (1 :: nat) 3\ + by simp + +lemma + \(1705 :: nat) AND 42 = 40\ + by simp + +lemma + \(1705 :: nat) AND Suc 0 = 1\ + by simp + +lemma + \(1705 :: nat) OR 42 = 1707\ + by simp + +lemma + \(1705 :: nat) OR Suc 0 = 1705\ + by simp + +lemma + \(1705 :: nat) XOR 42 = 1667\ + by simp + +lemma + \(1705 :: nat) XOR 1 = 1704\ + by simp + +lemma + \push_bit 3 (1705 :: nat) = 13640\ + by simp + +lemma + \push_bit (Suc (Suc (Suc 0))) (1705 :: nat) = 13640\ + by simp + +lemma + \push_bit 3 (Suc 0) = 8\ + by simp + +lemma + \push_bit (Suc (Suc (Suc 0))) (Suc 0) = 8\ + by simp + +lemma + \(1705 :: nat) << 3 = 13640\ + by simp + +lemma + \(1705 :: nat) << Suc (Suc (Suc 0)) = 13640\ + by simp + +lemma + \Suc 0 << 3 = 8\ + by simp + +lemma + \Suc 0 << Suc (Suc (Suc 0)) = 8\ + by simp + +lemma + \drop_bit 3 (1705 :: nat) = 213\ + by simp + +lemma + \drop_bit (Suc (Suc (Suc 0))) (1705 :: nat) = 213\ + by simp + +lemma + \drop_bit 3 (Suc 0) = 0\ + by simp + +lemma + \drop_bit (Suc (Suc (Suc 0))) (Suc 0) = 0\ + by simp + +lemma + \(1705 :: nat) >> 3 = 213\ + by simp + +lemma + \(1705 :: nat) >> Suc (Suc (Suc 0)) = 213\ + by simp + +lemma + \Suc 0 >> 3 = 0\ + by simp + +lemma + \Suc 0 >> Suc (Suc (Suc 0)) = 0\ + by simp + +lemma + \take_bit 3 (1705 :: nat) = 1\ + by simp + +lemma + \take_bit (Suc (Suc (Suc 0))) (1705 :: nat) = 1\ + by (simp flip: add_2_eq_Suc) + +lemma + \take_bit 3 (Suc 0) = 1\ + by simp + +lemma + \take_bit (Suc (Suc (Suc 0))) (Suc 0) = 1\ + by simp + + +section \\<^typ>\int\\ + +lemma + \bit (1705 :: int) (Suc (Suc (Suc 0)))\ + by simp + +lemma + \bit (1705 :: int) 3\ + by simp + +lemma + \\ bit (- 1705 :: int) (Suc (Suc (Suc 0)))\ + by simp + +lemma + \\ bit (- 1705 :: int) 3\ + by simp + +lemma + \\ bit (1 :: int) (Suc (Suc (Suc 0)))\ + by simp + +lemma + \\ bit (1 :: int) 3\ + by simp + +lemma + \(NOT 1705 :: int) = - 1706\ + by simp + +lemma + \(NOT (- 42 :: int)) = 41\ + by simp + +lemma + \(NOT 1 :: int) = - 2\ + by simp + +lemma + \(1705 :: int) AND 42 = 40\ + by simp + +lemma + \(1705 :: int) AND - 42 = 1664\ + by simp + +lemma + \(1705 :: int) AND 1 = 1\ + by simp + +lemma + \- (1705 :: int) AND 42 = 2\ + by simp + +lemma + \- (1705 :: int) AND - 42 = - 1706\ + by simp + +lemma + \- (1705 :: int) AND 1 = 1\ + by simp + +lemma + \(1705 :: int) OR 42 = 1707\ + by simp + +lemma + \(1705 :: int) OR - 42 = - 1\ + by simp + +lemma + \(1705 :: int) OR 1 = 1705\ + by simp + +lemma + \- (1705 :: int) OR 42 = - 1665\ + by simp + +lemma + \- (1705 :: int) OR - 42 = - 41\ + by simp + +lemma + \- (1705 :: int) OR 1 = - 1705\ + by simp + +lemma + \(1705 :: int) XOR 42 = 1667\ + by simp + +lemma + \(1705 :: int) XOR - 42 = - 1665\ + by simp + +lemma + \(1705 :: int) XOR 1 = 1704\ + by simp + +lemma + \- (1705 :: int) XOR 42 = - 1667\ + by simp + +lemma + \- (1705 :: int) XOR - 42 = 1665\ + by simp + +lemma + \- (1705 :: int) XOR 1 = - 1706\ + by simp + +lemma + \push_bit 3 (1705 :: int) = 13640\ + by simp + +lemma + \push_bit (Suc (Suc (Suc 0))) (1705 :: int) = 13640\ + by simp + +lemma + \push_bit 3 (- 1705 :: int) = - 13640\ + by simp + +lemma + \push_bit (Suc (Suc (Suc 0))) (- 1705 :: int) = - 13640\ + by simp + +lemma + \push_bit 3 (1 :: int) = 8\ + by simp + +lemma + \push_bit (Suc (Suc (Suc 0))) (1 :: int) = 8\ + by simp + +lemma + \push_bit 3 (- 1 :: int) = - 8\ + by (simp add: mask_eq_exp_minus_1) + +lemma + \push_bit (Suc (Suc (Suc 0))) (- 1 :: int) = - 8\ + by (simp add: mask_eq_exp_minus_1) + +lemma + \(1705 :: int) << 3 = 13640\ + by simp + +lemma + \(1705 :: int) << Suc (Suc (Suc 0)) = 13640\ + by simp + +lemma + \(- 1705 :: int) << 3 = - 13640\ + by simp + +lemma + \(- 1705 :: int) << Suc (Suc (Suc 0)) = - 13640\ + by simp + +lemma + \(1 :: int) << 3 = 8\ + by simp + +lemma + \(1 :: int) << Suc (Suc (Suc 0)) = 8\ + by simp + +lemma + \(- 1 :: int) << 3 = - 8\ + by (simp add: mask_eq_exp_minus_1) + +lemma + \(- 1 :: int) << Suc (Suc (Suc 0)) = - 8\ + by simp + +lemma + \drop_bit 3 (1705 :: int) = 213\ + by simp + +lemma + \drop_bit (Suc (Suc (Suc 0))) (1705 :: int) = 213\ + by simp + +lemma + \drop_bit 3 (- 1705 :: int) = - 214\ + by simp + +lemma + \drop_bit (Suc (Suc (Suc 0))) (- 1705 :: int) = - 214\ + by simp + +lemma + \drop_bit 3 (1 :: int) = 0\ + by simp + +lemma + \drop_bit (Suc (Suc (Suc 0))) (1 :: int) = 0\ + by simp + +lemma + \(1705 :: int) >> 3 = 213\ + by simp + +lemma + \(1705 :: int) >> Suc (Suc (Suc 0)) = 213\ + by simp + +lemma + \(- 1705 :: int) >> 3 = - 214\ + by simp + +lemma + \(- 1705 :: int) >> Suc (Suc (Suc 0)) = - 214\ + by simp + +lemma + \(1 :: int) >> 3 = 0\ + by simp + +lemma + \(1 :: int) >> Suc (Suc (Suc 0)) = 0\ + by simp + +lemma + \take_bit 3 (1705 :: int) = 1\ + by simp + +lemma + \take_bit (Suc (Suc (Suc 0))) (1705 :: int) = 1\ + by (simp flip: add_2_eq_Suc) + +lemma + \take_bit 3 (- 1705 :: int) = 7\ + by simp + +lemma + \take_bit (Suc (Suc (Suc 0))) (- 1705 :: int) = 7\ + by (simp flip: add_2_eq_Suc) + +lemma + \take_bit 3 (1 :: int) = 1\ + by simp + +lemma + \take_bit (Suc (Suc (Suc 0))) (1 :: int) = 1\ + by simp + +lemma + \take_bit 3 (- 1 :: int) = 7\ + by (simp add: mask_eq_exp_minus_1) + +lemma + \take_bit (Suc (Suc (Suc 0))) (- 1 :: int) = 7\ + by (simp add: mask_eq_exp_minus_1) + +lemma + \signed_take_bit 3 (1705 :: int) = - 7\ + by simp + +lemma + \signed_take_bit (Suc (Suc (Suc 0))) (1705 :: int) = - 7\ + by simp + +lemma + \signed_take_bit 3 (- 1705 :: int) = 7\ + by simp + +lemma + \signed_take_bit (Suc (Suc (Suc 0))) (- 1705 :: int) = 7\ + by simp + +lemma + \signed_take_bit 3 (1 :: int) = 1\ + by simp + +lemma + \signed_take_bit (Suc (Suc (Suc 0))) (1 :: int) = 1\ + by simp + + +section \\<^typ>\'a word\\ + +lemma + \(1705 :: 8 word) = 169\ + by simp + +lemma + \(- 1705 :: 8 word) = 87\ + by simp + +lemma + \(257 :: 8 word) = 1\ + by simp + +lemma + \(42 :: 8 word) \ 1705\ + by simp + +lemma + \(- 42 :: 8 word) \ 230\ + by simp + +lemma + \(42 :: 8 word) \ - 1705\ + by simp + +lemma + \- (42 :: 8 word) \ 235\ + by simp + +lemma + \(1 :: 8 word) \ 1705\ + by simp + +lemma + \(- 1 :: 8 word) \ 65535\ + by simp + +lemma + \(42 :: 8 word) < 1705\ + by simp + +lemma + \(- 42 :: 8 word) < 230\ + by simp + +lemma + \(42 :: 8 word) < - 1705\ + by simp + +lemma + \- (42 :: 8 word) < 230\ + by simp + +lemma + \(1 :: 8 word) < 1705\ + by simp + +lemma + \(1705 :: 8 word) < - 1\ + by simp + +lemma + \(42 :: 8 word) \s 1333\ + by simp + +lemma + \(- 42 :: 8 word) \s 230\ + by simp + +lemma + \(42 :: 8 word) \s - 1705\ + by simp + +lemma + \- (42 :: 8 word) \s - 1705\ + by simp + +lemma + \(1 :: 8 word) \s 42\ + by simp + +lemma + \(42 :: 8 word) + by simp + +lemma + \(- 42 :: 8 word) + by simp + +lemma + \(42 :: 8 word) + by simp + +lemma + \- (42 :: 8 word) + by simp + +lemma + \(1 :: 8 word) + by simp + +lemma + \bit (1705 :: 16 word) (Suc (Suc (Suc 0)))\ + by simp + +lemma + \bit (1705 :: 16 word) 3\ + by simp + +lemma + \\ bit (- 1705 :: 16 word) (Suc (Suc (Suc 0)))\ + by simp + +lemma + \\ bit (- 1705 :: 16 word) 3\ + by simp + +lemma + \\ bit (1 :: 'a::len word) (Suc (Suc (Suc 0)))\ + by simp + +lemma + \\ bit (1 :: 'a::len word) 3\ + by simp + +lemma + \(NOT 1705 :: 'a::len word) = - 1706\ + by simp + +lemma + \(NOT (- 42 :: 'a::len word)) = 41\ + by simp + +lemma + \(NOT 1 :: 'a::len word) = - 2\ + by simp + +lemma + \(1705 :: 'a::len word) AND 42 = 40\ + by simp + +lemma + \(1705 :: 'a::len word) AND - 42 = 1664\ + by simp + +lemma + \(1705 :: 'a::len word) AND 1 = 1\ + by simp + +lemma + \- (1705 :: 'a::len word) AND 42 = 2\ + by simp + +lemma + \- (1705 :: 'a::len word) AND - 42 = - 1706\ + by simp + +lemma + \- (1705 :: 'a::len word) AND 1 = 1\ + by simp + +lemma + \(1705 :: 'a::len word) OR 42 = 1707\ + by simp + +lemma + \(1705 :: 'a::len word) OR - 42 = - 1\ + by simp + +lemma + \(1705 :: 'a::len word) OR 1 = 1705\ + by simp + +lemma + \- (1705 :: 'a::len word) OR 42 = - 1665\ + by simp + +lemma + \- (1705 :: 'a::len word) OR - 42 = - 41\ + by simp + +lemma + \- (1705 :: 'a::len word) OR 1 = - 1705\ + by simp + +lemma + \(1705 :: 'a::len word) XOR 42 = 1667\ + by simp + +lemma + \(1705 :: 'a::len word) XOR - 42 = - 1665\ + by simp + +lemma + \(1705 :: 'a::len word) XOR 1 = 1704\ + by simp + +lemma + \- (1705 :: 'a::len word) XOR 42 = - 1667\ + by simp + +lemma + \- (1705 :: 'a::len word) XOR - 42 = 1665\ + by simp + +lemma + \- (1705 :: 'a::len word) XOR 1 = - 1706\ + by simp + +lemma + \push_bit 3 (1705 :: 'a::len word) = 13640\ + by simp + +lemma + \push_bit (Suc (Suc (Suc 0))) (1705 :: 'a::len word) = 13640\ + by simp + +lemma + \push_bit 3 (- 1705 :: 'a::len word) = - 13640\ + by simp + +lemma + \push_bit (Suc (Suc (Suc 0))) (- 1705 :: 'a::len word) = - 13640\ + by simp + +lemma + \push_bit 3 (1 :: 'a::len word) = 8\ + by simp + +lemma + \push_bit (Suc (Suc (Suc 0))) (1 :: 'a::len word) = 8\ + by simp + +lemma + \push_bit 3 (- 1 :: 'a::len word) = - 8\ + by (simp add: mask_eq_exp_minus_1) + +lemma + \push_bit (Suc (Suc (Suc 0))) (- 1 :: 'a::len word) = - 8\ + by (simp add: mask_eq_exp_minus_1) + +lemma + \(1705 :: 'a::len word) << 3 = 13640\ + by simp + +lemma + \(1705 :: 'a::len word) << Suc (Suc (Suc 0)) = 13640\ + by simp + +lemma + \(- 1705 :: 'a::len word) << 3 = - 13640\ + by simp + +lemma + \(- 1705 :: 'a::len word) << Suc (Suc (Suc 0)) = - 13640\ + by simp + +lemma + \(1 :: 'a::len word) << 3 = 8\ + by simp + +lemma + \(1 :: 'a::len word) << Suc (Suc (Suc 0)) = 8\ + by simp + +lemma + \(- 1 :: 'a::len word) << 3 = - 8\ + by (simp add: mask_eq_exp_minus_1) + +lemma + \(- 1 :: 'a::len word) << Suc (Suc (Suc 0)) = - 8\ + by simp + +lemma + \drop_bit 3 (1705 :: 16 word) = 213\ + by simp + +lemma + \drop_bit (Suc (Suc (Suc 0))) (1705 :: 16 word) = 213\ + by simp + +lemma + \drop_bit 3 (- 1705 :: 16 word) = 7978\ + by simp + +lemma + \drop_bit (Suc (Suc (Suc 0))) (- 1705 :: 16 word) = 7978\ + by simp + +lemma + \drop_bit 3 (1 :: 16 word) = 0\ + by simp + +lemma + \drop_bit (Suc (Suc (Suc 0))) (1 :: 16 word) = 0\ + by simp + +lemma + \(1705 :: 16 word) >> 3 = 213\ + by simp + +lemma + \(1705 :: 16 word) >> Suc (Suc (Suc 0)) = 213\ + by simp + +lemma + \(- 1705 :: 16 word) >> 3 = 7978\ + by simp + +lemma + \(- 1705 :: 16 word) >> Suc (Suc (Suc 0)) = 7978\ + by simp + +lemma + \(1 :: 16 word) >> 3 = 0\ + by simp + +lemma + \(1 :: 16 word) >> Suc (Suc (Suc 0)) = 0\ + by simp + +lemma + \signed_drop_bit 3 (1705 :: 16 word) = 213\ + by simp + +lemma + \signed_drop_bit (Suc (Suc (Suc 0))) (1705 :: 16 word) = 213\ + by simp + +lemma + \signed_drop_bit 3 (- 1705 :: 16 word) = - 214\ + by simp + +lemma + \signed_drop_bit (Suc (Suc (Suc 0))) (- 1705 :: 16 word) = - 214\ + by simp + +lemma + \signed_drop_bit 3 (1 :: 16 word) = 0\ + by simp + +lemma + \signed_drop_bit (Suc (Suc (Suc 0))) (1 :: 16 word) = 0\ + by simp + +lemma + \(1705 :: 16 word) >>> 3 = 213\ + by simp + +lemma + \(1705 :: 16 word) >>> Suc (Suc (Suc 0)) = 213\ + by simp + +lemma + \(- 1705 :: 16 word) >>> 3 = - 214\ + by simp + +lemma + \(- 1705 :: 16 word) >>> Suc (Suc (Suc 0)) = - 214\ + by simp + +lemma + \(1 :: 16 word) >>> 3 = 0\ + by simp + +lemma + \(1 :: 16 word) >>> Suc (Suc (Suc 0)) = 0\ + by simp + +lemma + \take_bit 3 (1705 :: 16 word) = 1\ + by simp + +lemma + \take_bit (Suc (Suc (Suc 0))) (1705 :: 16 word) = 1\ + by (simp flip: add_2_eq_Suc) + +lemma + \take_bit 3 (- 1705 :: 16 word) = 7\ + by simp + +lemma + \take_bit (Suc (Suc (Suc 0))) (- 1705 :: 16 word) = 7\ + by (simp flip: add_2_eq_Suc) + +lemma + \take_bit 3 (1 :: 16 word) = 1\ + by simp + +lemma + \take_bit (Suc (Suc (Suc 0))) (1 :: 16 word) = 1\ + by simp + +lemma + \take_bit 3 (- 1 :: 16 word) = 7\ + by (simp add: mask_eq_exp_minus_1) + +lemma + \take_bit (Suc (Suc (Suc 0))) (- 1 :: 16 word) = 7\ + by (simp add: mask_eq_exp_minus_1) + +lemma + \signed_take_bit 3 (1705 :: 16 word) = - 7\ + by simp + +lemma + \signed_take_bit (Suc (Suc (Suc 0))) (1705 :: 16 word) = - 7\ + by simp + +lemma + \signed_take_bit 3 (- 1705 :: 16 word) = 7\ + by simp + +lemma + \signed_take_bit (Suc (Suc (Suc 0))) (- 1705 :: 16 word) = 7\ + by simp + +lemma + \signed_take_bit 3 (1 :: 16 word) = 1\ + by simp + +lemma + \signed_take_bit (Suc (Suc (Suc 0))) (1 :: 16 word) = 1\ + by simp + +lemma + \(1705 :: 16 word) div 42 = 40\ + by simp + +lemma + \(- 1705 :: 16 word) div 42 = 1519\ + by simp + +lemma + \(1705 :: 16 word) div - 42 = 0\ + by simp + +lemma + \(- 1705 :: 16 word) div - 42 = 0\ + by simp + +lemma + \(1705 :: 16 word) div 1 = 1705\ + by simp + +lemma + \(1705 :: 16 word) div - 1 = 0\ + by simp + +lemma + \(1 :: 16 word) div 42 = 0\ + by simp + +lemma + \(- 1 :: 16 word) div 42 = 1560\ + by simp + +lemma + \(1705 :: 16 word) mod 42 = 25\ + by simp + +lemma + \(- 1705 :: 16 word) mod 42 = 33\ + by simp + +lemma + \(1705 :: 16 word) mod - 42 = 1705\ + by simp + +lemma + \(- 1705 :: 16 word) mod - 42 = 63831\ + by simp + +lemma + \(1705 :: 16 word) mod 1 = 0\ + by simp + +lemma + \(1705 :: 16 word) mod - 1 = 1705\ + by simp + +lemma + \(1 :: 16 word) mod 42 = 1\ + by simp + +lemma + \(- 1 :: 16 word) mod 42 = 15\ + by simp + +lemma + \(1705 :: 16 word) sdiv 42 = 40\ + by simp + +lemma + \(- 1705 :: 16 word) sdiv 42 = 65496\ + by simp + +lemma + \(1705 :: 16 word) sdiv - 42 = 65496\ + by simp + +lemma + \(- 1705 :: 16 word) sdiv - 42 = 40\ + by simp + +lemma + \(1705 :: 16 word) sdiv 1 = 1705\ + by simp + +lemma + \(1705 :: 16 word) sdiv - 1 = 63831\ + by simp + +lemma + \(1 :: 16 word) sdiv 42 = 0\ + by simp + +lemma + \(- 1 :: 16 word) sdiv 42 = 0\ + by simp + +lemma + \(1705 :: 16 word) smod 42 = 25\ + by simp + +lemma + \(- 1705 :: 16 word) smod 42 = 65511\ + by simp + +lemma + \(1705 :: 16 word) smod - 42 = 25\ + by simp + +lemma + \(- 1705 :: 16 word) smod - 42 = 65511\ + by simp + +lemma + \(1705 :: 16 word) smod 1 = 0\ + by simp + +lemma + \(1705 :: 16 word) smod - 1 = 0\ + by simp + +lemma + \(1 :: 16 word) smod 42 = 1\ + by simp + +lemma + \(- 1 :: 16 word) smod 42 = 65535\ + by simp + text "modulus" lemma "(27 :: 4 word) = -5" by simp @@ -68,15 +1029,6 @@ lemma "i < x \ i < i + 1" for i x :: "'a::len word" lemma "i < x \ i < i + 1" for i x :: "'a::len word" by unat_arith -text "bool lists" - -lemma "of_bl [True, False, True, True] = (0b1011::'a::len word)" by simp - -lemma "to_bl (0b110::4 word) = [False, True, True, False]" by simp - -lemma "of_bl (replicate 32 True) = (0xFFFFFFFF::32 word)" - by (simp add: numeral_eq_Suc) - text "bit operations" lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by simp @@ -96,13 +1048,13 @@ lemma "NOT 0 = (255 :: 8 word)" by simp lemma "(-1 :: 32 word) = 0xFFFFFFFF" by simp -lemma "(0b0010 :: 4 word) !! 1" by simp -lemma "\ (0b0010 :: 4 word) !! 0" by simp -lemma "\ (0b1000 :: 3 word) !! 4" by simp -lemma "\ (1 :: 3 word) !! 2" by simp +lemma "bit (0b0010 :: 4 word) 1" by simp +lemma "\ bit (0b0010 :: 4 word) 0" by simp +lemma "\ bit (0b1000 :: 3 word) 4" by simp +lemma "\ bit (1 :: 3 word) 2" by simp -lemma "(0b11000 :: 10 word) !! n = (n = 4 \ n = 3)" - by (auto simp add: bin_nth_Bit0 bin_nth_Bit1) +lemma "bit (0b11000 :: 10 word) n = (n = 4 \ n = 3)" + by (auto simp add: bit_numeral_rec bit_1_iff split: nat.splits) lemma "set_bit 55 7 True = (183::'a::len word)" by simp lemma "set_bit 0b0010 7 True = (0b10000010::'a::len word)" by simp @@ -122,7 +1074,8 @@ lemma "msb (0b1000::4 word)" by simp lemma "\ msb (1::4 word)" by simp lemma "\ msb (0::4 word)" by simp -lemma "word_cat (27::4 word) (27::8 word) = (2843::'a::len word)" by simp +lemma "word_cat (27::4 word) (27::8 word) = (2843::'a::len word)" + by simp lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)" by simp @@ -159,6 +1112,43 @@ lemma "word_next (255:: 8 word) = 255" by eval lemma "word_prev (2:: 8 word) = 1" by eval lemma "word_prev (0:: 8 word) = 0" by eval +text \singed division\ + +lemma + "( 4 :: 32 word) sdiv 4 = 1" + "(-4 :: 32 word) sdiv 4 = -1" + "(-3 :: 32 word) sdiv 4 = 0" + "( 3 :: 32 word) sdiv -4 = 0" + "(-3 :: 32 word) sdiv -4 = 0" + "(-5 :: 32 word) sdiv -4 = 1" + "( 5 :: 32 word) sdiv -4 = -1" + by (simp_all add: sdiv_word_def signed_divide_int_def) + +lemma + "( 4 :: 32 word) smod 4 = 0" + "( 3 :: 32 word) smod 4 = 3" + "(-3 :: 32 word) smod 4 = -3" + "( 3 :: 32 word) smod -4 = 3" + "(-3 :: 32 word) smod -4 = -3" + "(-5 :: 32 word) smod -4 = -1" + "( 5 :: 32 word) smod -4 = 1" + by (simp_all add: smod_word_def signed_modulo_int_def signed_divide_int_def) + + +text \comparison\ + +lemma "1 < (1024::32 word) \ 1 \ (1024::32 word)" + by simp + +text "bool lists" + +lemma "of_bl [True, False, True, True] = (0b1011::'a::len word)" by simp + +lemma "to_bl (0b110::4 word) = [False, True, True, False]" by simp + +lemma "of_bl (replicate 32 True) = (0xFFFFFFFF::32 word)" + by (simp add: numeral_eq_Suc) + text "proofs using bitwise expansion" lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" @@ -199,28 +1189,6 @@ lemma shiftr_overflow: "32 \ a \ b >> a = 0" lemma "((x :: 32 word) >> 3) AND 7 = (x AND 56) >> 3" by word_bitwise -(* Tests *) -lemma - "( 4 :: 32 word) sdiv 4 = 1" - "(-4 :: 32 word) sdiv 4 = -1" - "(-3 :: 32 word) sdiv 4 = 0" - "( 3 :: 32 word) sdiv -4 = 0" - "(-3 :: 32 word) sdiv -4 = 0" - "(-5 :: 32 word) sdiv -4 = 1" - "( 5 :: 32 word) sdiv -4 = -1" - by (simp_all add: sdiv_word_def signed_divide_int_def) - -lemma - "( 4 :: 32 word) smod 4 = 0" - "( 3 :: 32 word) smod 4 = 3" - "(-3 :: 32 word) smod 4 = -3" - "( 3 :: 32 word) smod -4 = 3" - "(-3 :: 32 word) smod -4 = -3" - "(-5 :: 32 word) smod -4 = -1" - "( 5 :: 32 word) smod -4 = 1" - by (simp_all add: smod_word_def signed_modulo_int_def signed_divide_int_def) - -lemma "1 < (1024::32 word) \ 1 \ (1024::32 word)" - by simp +end end diff --git a/lib/Word_Lib/Generic_set_bit.thy b/lib/Word_Lib/Generic_set_bit.thy index cad22f94b..f3641eb8c 100644 --- a/lib/Word_Lib/Generic_set_bit.thy +++ b/lib/Word_Lib/Generic_set_bit.thy @@ -11,16 +11,17 @@ section \Operation variant for setting and unsetting bits\ theory Generic_set_bit imports "HOL-Library.Word" - Bits_Int Most_significant_bit begin class set_bit = semiring_bits + fixes set_bit :: \'a \ nat \ bool \ 'a\ - assumes bit_set_bit_iff [bit_simps]: + assumes bit_set_bit_iff_2n: \bit (set_bit a m b) n \ (if m = n then b else bit a n) \ 2 ^ n \ 0\ +lemmas bit_set_bit_iff[bit_simps] = bit_set_bit_iff_2n[simplified fold_possible_bit simp_thms] + lemma set_bit_eq: \set_bit a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ for a :: \'a::{ring_bit_operations, set_bit}\ @@ -30,88 +31,43 @@ instantiation int :: set_bit begin definition set_bit_int :: \int \ nat \ bool \ int\ - where \set_bit i n b = bin_sc n b i\ + where \set_bit_int i n b = (if b then Bit_Operations.set_bit else Bit_Operations.unset_bit) n i\ instance - by standard - (simp_all add: set_bit_int_def bin_nth_sc_gen bit_simps) + by standard (simp_all add: set_bit_int_def bit_simps) end -lemma int_set_bit_0 [simp]: fixes x :: int shows - "set_bit x 0 b = of_bool b + 2 * (x div 2)" - by (auto simp add: set_bit_int_def intro: bin_rl_eqI) - -lemma int_set_bit_Suc: fixes x :: int shows - "set_bit x (Suc n) b = of_bool (odd x) + 2 * set_bit (x div 2) n b" - by (auto simp add: set_bit_int_def intro: bin_rl_eqI) - -lemma bin_last_set_bit: - "bin_last (set_bit x n b) = (if n > 0 then bin_last x else b)" - by (cases n) (simp_all add: int_set_bit_Suc) - -lemma bin_rest_set_bit: - "bin_rest (set_bit x n b) = (if n > 0 then set_bit (x div 2) (n - 1) b else x div 2)" - by (cases n) (simp_all add: int_set_bit_Suc) - -lemma int_set_bit_numeral: fixes x :: int shows - "set_bit x (numeral w) b = of_bool (odd x) + 2 * set_bit (x div 2) (pred_numeral w) b" - by (simp add: set_bit_int_def) - -lemmas int_set_bit_numerals [simp] = - int_set_bit_numeral[where x="numeral w'"] - int_set_bit_numeral[where x="- numeral w'"] - int_set_bit_numeral[where x="Numeral1"] - int_set_bit_numeral[where x="1"] - int_set_bit_numeral[where x="0"] - int_set_bit_Suc[where x="numeral w'"] - int_set_bit_Suc[where x="- numeral w'"] - int_set_bit_Suc[where x="Numeral1"] - int_set_bit_Suc[where x="1"] - int_set_bit_Suc[where x="0"] - for w' - -lemma msb_set_bit [simp]: "msb (set_bit (x :: int) n b) \ msb x" -by(simp add: msb_conv_bin_sign set_bit_int_def) - instantiation word :: (len) set_bit begin definition set_bit_word :: \'a word \ nat \ bool \ 'a word\ - where word_set_bit_def: \set_bit a n x = word_of_int (bin_sc n x (uint a))\ + where set_bit_unfold: \set_bit w n b = (if b then Bit_Operations.set_bit n w else unset_bit n w)\ + for w :: \'a::len word\ instance - by standard - (auto simp add: word_set_bit_def bin_nth_sc_gen bit_simps) + by standard (auto simp add: set_bit_unfold bit_simps dest: bit_imp_le_length) end -lemma set_bit_unfold: - \set_bit w n b = (if b then Bit_Operations.set_bit n w else unset_bit n w)\ - for w :: \'a::len word\ - by (simp add: set_bit_eq) - lemma bit_set_bit_word_iff [bit_simps]: \bit (set_bit w m b) n \ (if m = n then n < LENGTH('a) \ b else bit w n)\ for w :: \'a::len word\ by (auto simp add: bit_simps dest: bit_imp_le_length) -lemma word_set_nth [simp]: "set_bit w n (test_bit w n) = w" - for w :: "'a::len word" - by (auto simp: word_test_bit_def word_set_bit_def) - -lemma test_bit_set: "(set_bit w n x) !! n \ n < size w \ x" - for w :: "'a::len word" - by (auto simp: word_size word_test_bit_def word_set_bit_def nth_bintr) - lemma test_bit_set_gen: - "test_bit (set_bit w n x) m = (if m = n then n < size w \ x else test_bit w m)" + "bit (set_bit w n x) m \ (if m = n then n < size w \ x else bit w m)" for w :: "'a::len word" - apply (unfold word_size word_test_bit_def word_set_bit_def) - apply (clarsimp simp add: nth_bintr bin_nth_sc_gen) - apply (auto elim!: test_bit_size [unfolded word_size] - simp add: word_test_bit_def [symmetric]) - done + by (simp add: bit_set_bit_word_iff word_size) + +lemma test_bit_set: + "bit (set_bit w n x) n \ n < size w \ x" + for w :: "'a::len word" + by (auto simp add: bit_simps word_size) + +lemma word_set_nth: "set_bit w n (bit w n) = w" + for w :: "'a::len word" + by (rule bit_word_eqI) (simp add: bit_simps) lemma word_set_set_same [simp]: "set_bit (set_bit w n x) n y = set_bit w n y" for w :: "'a::len word" @@ -123,27 +79,7 @@ lemma word_set_set_diff: shows "set_bit (set_bit w m x) n y = set_bit (set_bit w n y) m x" by (rule word_eqI) (auto simp: test_bit_set_gen word_size assms) -lemma set_bit_word_of_int: "set_bit (word_of_int x) n b = word_of_int (bin_sc n b x)" - unfolding word_set_bit_def - by (rule word_eqI)(simp add: word_size bin_nth_sc_gen nth_bintr) - -lemma word_set_numeral [simp]: - "set_bit (numeral bin::'a::len word) n b = - word_of_int (bin_sc n b (numeral bin))" - unfolding word_numeral_alt by (rule set_bit_word_of_int) - -lemma word_set_neg_numeral [simp]: - "set_bit (- numeral bin::'a::len word) n b = - word_of_int (bin_sc n b (- numeral bin))" - unfolding word_neg_numeral_alt by (rule set_bit_word_of_int) - -lemma word_set_bit_0 [simp]: "set_bit 0 n b = word_of_int (bin_sc n b 0)" - unfolding word_0_wi by (rule set_bit_word_of_int) - -lemma word_set_bit_1 [simp]: "set_bit 1 n b = word_of_int (bin_sc n b 1)" - unfolding word_1_wi by (rule set_bit_word_of_int) - -lemma word_set_nth_iff: "set_bit w n b = w \ w !! n = b \ n \ size w" +lemma word_set_nth_iff: "set_bit w n b = w \ bit w n = b \ n \ size w" for w :: "'a::len word" apply (rule iffI) apply (rule disjCI) @@ -153,38 +89,36 @@ lemma word_set_nth_iff: "set_bit w n b = w \ w !! n = b \ set_bit w n False" for w :: "'a::len word" - apply (simp add: word_set_bit_def word_le_def) + apply (simp add: set_bit_unfold) apply transfer - apply (rule order_trans) - apply (rule bintr_bin_clr_le) - apply simp + apply (simp add: take_bit_unset_bit_eq unset_bit_less_eq) done lemma word_set_ge: "w \ set_bit w n True" for w :: "'a::len word" - apply (simp add: word_set_bit_def word_le_def) + apply (simp add: set_bit_unfold) apply transfer - apply (rule order_trans [OF _ bintr_bin_set_ge]) - apply simp + apply (simp add: take_bit_set_bit_eq set_bit_greater_eq) done lemma set_bit_beyond: "size x \ n \ set_bit x n b = x" for x :: "'a :: len word" - by (auto intro: word_eqI simp add: test_bit_set_gen word_size) + by (simp add: word_set_nth_iff) lemma one_bit_shiftl: "set_bit 0 n True = (1 :: 'a :: len word) << n" apply (rule word_eqI) - apply (auto simp add: test_bit_set_gen nth_shiftl word_size - simp del: word_set_bit_0 shiftl_1) + apply (auto simp add: word_size bit_simps) done -lemmas one_bit_pow = trans [OF one_bit_shiftl shiftl_1] +lemma one_bit_pow: "set_bit 0 n True = (2 :: 'a :: len word) ^ n" + by (simp add: one_bit_shiftl shiftl_def) end diff --git a/lib/Word_Lib/Guide.thy b/lib/Word_Lib/Guide.thy index 359699f2f..ab5220883 100644 --- a/lib/Word_Lib/Guide.thy +++ b/lib/Word_Lib/Guide.thy @@ -6,15 +6,28 @@ (*<*) theory Guide - imports Word_Lib_Sumo Word_64 + imports Word_Lib_Sumo Machine_Word_32 Machine_Word_64 Ancient_Numeral begin +context semiring_bit_operations +begin + +lemma bit_eq_iff: + \a = b \ (\n. 2 ^ n \ 0 \ bit a n \ bit b n)\ + using bit_eq_iff [of a b] by (simp add: possible_bit_def) + +end + +notation (output) Generic_set_bit.set_bit (\Generic'_set'_bit.set'_bit\) + hide_const (open) Generic_set_bit.set_bit +no_notation bit (infixl \!!\ 100) + (*>*) section \A short overview over bit operations and word types\ -subsection \Basic theories and key ideas\ +subsection \Key principles\ text \ When formalizing bit operations, it is tempting to represent @@ -33,7 +46,7 @@ text \ in @{term [source] 0} which can be represented by type \<^typ>\nat\ but only support a restricted set of operations). - The most fundamental ideas are developed in theory \<^theory>\HOL.Parity\ + The fundamental principles are developed in theory \<^theory>\HOL.Bit_Operations\ (which is part of \<^theory>\Main\): \<^item> Multiplication by \<^term>\2 :: int\ is a bit shift to the left and @@ -53,11 +66,10 @@ text \ \<^item> Induction rule: @{thm [display, mode=iff] bits_induct [where ?'a = int, no_vars]} - \<^item> Characteristic properties \<^prop>\bit (f x) n \ P x n\ + \<^item> Characteristic properties @{prop [source] \bit (f x) n \ P x n\} are available in fact collection \<^text>\bit_simps\. - On top of this, the following generic operations are provided - after import of theory \<^theory>\HOL-Library.Bit_Operations\: + On top of this, the following generic operations are provided: \<^item> Singleton \<^term>\n\th bit: \<^term>\(2 :: int) ^ n\ @@ -69,19 +81,19 @@ text \ \<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]} - \<^item> Negation: @{thm [mode=iff] bit_not_iff [where ?'a = int, no_vars]} + \<^item> Bitwise negation: @{thm [mode=iff] bit_not_iff_eq [where ?'a = int, no_vars]} - \<^item> And: @{thm [mode=iff] bit_and_iff [where ?'a = int, no_vars]} + \<^item> Bitwise conjunction: @{thm [mode=iff] bit_and_iff [where ?'a = int, no_vars]} - \<^item> Or: @{thm [mode=iff] bit_or_iff [where ?'a = int, no_vars]} + \<^item> Bitwise disjunction: @{thm [mode=iff] bit_or_iff [where ?'a = int, no_vars]} - \<^item> Xor: @{thm [mode=iff] bit_xor_iff [where ?'a = int, no_vars]} + \<^item> Bitwise exclusive disjunction: @{thm [mode=iff] bit_xor_iff [where ?'a = int, no_vars]} - \<^item> Set a single bit: @{thm set_bit_def [where ?'a = int, no_vars]} + \<^item> Setting a single bit: @{thm set_bit_def [where ?'a = int, no_vars]} - \<^item> Unset a single bit: @{thm unset_bit_def [where ?'a = int, no_vars]} + \<^item> Unsetting a single bit: @{thm unset_bit_def [where ?'a = int, no_vars]} - \<^item> Flip a single bit: @{thm flip_bit_def [where ?'a = int, no_vars]} + \<^item> Flipping a single bit: @{thm flip_bit_def [where ?'a = int, no_vars]} \<^item> Signed truncation, or modulus centered around \<^term>\0::int\: @{thm [display] signed_take_bit_def [where ?'a = int, no_vars]} @@ -89,6 +101,16 @@ text \ \<^item> (Bounded) conversion from and to a list of bits: @{thm [display] horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]} + Bit concatenation on \<^typ>\int\ as given by + @{thm [display] concat_bit_def [no_vars]} + appears quite + technical but is the logical foundation for the quite natural bit concatenation + on \<^typ>\'a word\ (see below). +\ + +subsection \Core word theory\ + +text \ Proper word types are introduced in theory \<^theory>\HOL-Library.Word\, with the following specific operations: @@ -159,22 +181,26 @@ text \ For proofs about words the following default strategies are applicable: - \<^item> Using bit extensionality (facts \<^text>\bit_eq_iff\, \<^text>\bit_eqI\; fact + \<^item> Using bit extensionality (facts \<^text>\bit_eq_iff\, \<^text>\bit_word_eqI\; fact collection \<^text>\bit_simps\). \<^item> Using the @{method transfer} method. \ + subsection \More library theories\ text \ - Note: currently, the theories listed here are hardly separate + Note: currently, most theories listed here are hardly separate entities since they import each other in various ways. Always inspect them to understand what you pull in if you want to import one. \<^descr>[Syntax] + \<^descr>[\<^theory>\Word_Lib.Syntax_Bundles\] + Bundles to provide alternative syntax for various bit operations. + \<^descr>[\<^theory>\Word_Lib.Hex_Words\] Printing word numerals as hexadecimal numerals. @@ -192,6 +218,9 @@ text \ \<^descr>[\<^theory>\Word_Lib.Bitwise\] Method @{method word_bitwise} decomposes word equalities and inequalities into bit propositions. + \<^descr>[\<^theory>\Word_Lib.Bitwise_Signed\] + Method @{method word_bitwise_signed} decomposes word equalities and inequalities into bit propositions. + \<^descr>[\<^theory>\Word_Lib.Word_EqI\] Method @{method word_eqI_solve} decomposes word equalities and inequalities into bit propositions. @@ -226,18 +255,15 @@ text \ \<^item> @{thm [mode=iff] msb_word_iff_bit [no_vars]} - \<^descr>[\<^theory>\Word_Lib.Traditional_Infix_Syntax\] + \<^descr>[\<^theory>\Word_Lib.Bit_Shifts_Infix_Syntax\] - Clones of existing operations decorated with - traditional syntax: + Bit shifts decorated with infix syntax: - \<^item> @{thm test_bit_eq_bit [no_vars]} + \<^item> @{thm Bit_Shifts_Infix_Syntax.shiftl_def [no_vars]} - \<^item> @{thm shiftl_eq_push_bit [no_vars]} + \<^item> @{thm Bit_Shifts_Infix_Syntax.shiftr_def [no_vars]} - \<^item> @{thm shiftr_eq_drop_bit [no_vars]} - - \<^item> @{thm sshiftr_eq [no_vars]} + \<^item> @{thm Bit_Shifts_Infix_Syntax.sshiftr_def [no_vars]} \<^descr>[\<^theory>\Word_Lib.Next_and_Prev\] \ @@ -281,15 +307,19 @@ text \ \<^descr>[\<^theory>\Word_Lib.Word_32\] - for 32-bit words. This theory is not part of \<^text>\Word_Lib_Sumo\, because it shadows - names from \<^theory>\Word_Lib.Word_64\. They can be used together, but then will have - to use qualified names in applications. + for 32-bit words. \<^descr>[\<^theory>\Word_Lib.Word_64\] for 64-bit words. This theory is not part of \<^text>\Word_Lib_Sumo\, because it shadows names from \<^theory>\Word_Lib.Word_32\. They can be used together, but then will have to use qualified names in applications. + + \<^descr>[\<^theory>\Word_Lib.Machine_Word_32\ and \<^theory>\Word_Lib.Machine_Word_64\] + + provide lemmas for 32-bit words and 64-bit words under the same name, + which can help to organize applications relying on some form + of genericity. \ @@ -319,7 +349,7 @@ text \ for backward compatibility: start importing this theory when migrating applications to Isabelle2021, and later sort out what you really need. You may need to include - \<^theory>\Word_Lib.Word_32\ or \<^theory>\Word_Lib.Word_64\ separately. + \<^theory>\Word_Lib.Word_64\ separately. \<^descr>[\<^theory>\Word_Lib.Generic_set_bit\] @@ -355,6 +385,36 @@ text \ They are used in applications of \<^text>\Word_Lib\, but should be migrated to there. \ +section \Changelog\ + +text \ + \<^descr>[Changes since AFP 2021] ~ + + \<^item> Theory \<^theory>\Word_Lib.Ancient_Numeral\ is not part of \<^theory>\Word_Lib.Word_Lib_Sumo\ + any longer. + + \<^item> Infix syntax for \<^term>\(AND)\, \<^term>\(OR)\, \<^term>\(XOR)\ organized in + syntax bundle \<^bundle>\bit_operations_syntax\. + + \<^item> Abbreviation \<^abbrev>\max_word\ moved from distribution into theory + \<^theory>\Word_Lib.Legacy_Aliases\. + + \<^item> Operation \<^const>\test_bit\ replaced by input abbreviation \<^abbrev>\test_bit\. + + \<^item> Abbreviations \<^abbrev>\bin_nth\, \<^abbrev>\bin_last\, \<^abbrev>\bin_rest\, + \<^abbrev>\bintrunc\, \<^abbrev>\sbintrunc\, \<^abbrev>\norm_sint\, + \<^abbrev>\bin_cat\ moved into theory \<^theory>\Word_Lib.Legacy_Aliases\. + + \<^item> Operations \<^abbrev>\bshiftr1\, + \<^abbrev>\setBit\, \<^abbrev>\clearBit\ moved from distribution into theory + \<^theory>\Word_Lib.Legacy_Aliases\ and replaced by input abbreviations. + + \<^item> Operations \<^const>\shiftl1\, \<^const>\shiftr1\, \<^const>\sshiftr1\ + moved here from distribution. + + \<^item> Operation \<^const>\complement\ replaced by input abbreviation \<^abbrev>\complement\. +\ + (*<*) end (*>*) diff --git a/lib/Word_Lib/Least_significant_bit.thy b/lib/Word_Lib/Least_significant_bit.thy index a91d8be3e..587b0318c 100644 --- a/lib/Word_Lib/Least_significant_bit.thy +++ b/lib/Word_Lib/Least_significant_bit.thy @@ -11,7 +11,7 @@ section \Operation variant for the least significant bit\ theory Least_significant_bit imports "HOL-Library.Word" - Bits_Int + More_Word begin class lsb = semiring_bits + @@ -22,14 +22,14 @@ instantiation int :: lsb begin definition lsb_int :: \int \ bool\ - where \lsb i = i !! 0\ for i :: int + where \lsb i = bit i 0\ for i :: int instance by standard (simp add: fun_eq_iff lsb_int_def) end -lemma bin_last_conv_lsb: "bin_last = lsb" +lemma bin_last_conv_lsb: "odd = (lsb :: int \ bool)" by (simp add: lsb_odd) lemma int_lsb_numeral [simp]: @@ -62,9 +62,9 @@ lemma lsb_word_eq: \lsb = (odd :: 'a word \ bool)\ for w :: \'a::len word\ by (fact lsb_odd) -lemma word_lsb_alt: "lsb w = test_bit w 0" +lemma word_lsb_alt: "lsb w = bit w 0" for w :: "'a::len word" - by (auto simp: word_test_bit_def word_lsb_def) + by (simp add: lsb_word_eq) lemma word_lsb_1_0 [simp]: "lsb (1::'a::len word) \ \ lsb (0::'b::len word)" unfolding word_lsb_def by simp @@ -78,12 +78,12 @@ lemma word_lsb_int: "lsb w \ uint w mod 2 = 1" lemmas word_ops_lsb = lsb0 [unfolded word_lsb_alt] lemma word_lsb_numeral [simp]: - "lsb (numeral bin :: 'a::len word) \ bin_last (numeral bin)" - unfolding word_lsb_alt test_bit_numeral by simp + "lsb (numeral bin :: 'a::len word) \ odd (numeral bin :: int)" + by (simp only: lsb_odd, transfer) rule lemma word_lsb_neg_numeral [simp]: - "lsb (- numeral bin :: 'a::len word) \ bin_last (- numeral bin)" - by (simp add: word_lsb_alt) + "lsb (- numeral bin :: 'a::len word) \ odd (- numeral bin :: int)" + by (simp only: lsb_odd, transfer) rule lemma word_lsb_nat:"lsb w = (unat w mod 2 = 1)" apply (simp add: word_lsb_def Groebner_Basis.algebra(31)) diff --git a/lib/Word_Lib/Legacy_Aliases.thy b/lib/Word_Lib/Legacy_Aliases.thy index 2223dc973..4d50a3d92 100644 --- a/lib/Word_Lib/Legacy_Aliases.thy +++ b/lib/Word_Lib/Legacy_Aliases.thy @@ -10,13 +10,128 @@ theory Legacy_Aliases imports "HOL-Library.Word" begin -definition - complement :: "'a :: len word \ 'a word" where - "complement x \ NOT x" +context abstract_boolean_algebra +begin + +lemma conj_assoc: "(x \<^bold>\ y) \<^bold>\ z = x \<^bold>\ (y \<^bold>\ z)" + by (fact conj.assoc) + +lemma conj_commute: "x \<^bold>\ y = y \<^bold>\ x" + by (fact conj.commute) + +lemmas conj_left_commute = conj.left_commute +lemmas conj_ac = conj.assoc conj.commute conj.left_commute + +lemma conj_one_left: "\<^bold>1 \<^bold>\ x = x" + by (fact conj.left_neutral) + +lemma conj_left_absorb: "x \<^bold>\ (x \<^bold>\ y) = x \<^bold>\ y" + by (fact conj.left_idem) + +lemma conj_absorb: "x \<^bold>\ x = x" + by (fact conj.idem) + +lemma disj_assoc: "(x \<^bold>\ y) \<^bold>\ z = x \<^bold>\ (y \<^bold>\ z)" + by (fact disj.assoc) + +lemma disj_commute: "x \<^bold>\ y = y \<^bold>\ x" + by (fact disj.commute) + +lemmas disj_left_commute = disj.left_commute + +lemmas disj_ac = disj.assoc disj.commute disj.left_commute + +lemma disj_zero_left: "\<^bold>0 \<^bold>\ x = x" + by (fact disj.left_neutral) + +lemma disj_left_absorb: "x \<^bold>\ (x \<^bold>\ y) = x \<^bold>\ y" + by (fact disj.left_idem) + +lemma disj_absorb: "x \<^bold>\ x = x" + by (fact disj.idem) + +end + +context abstract_boolean_algebra_sym_diff +begin + +lemmas xor_assoc = xor.assoc +lemmas xor_commute = xor.commute +lemmas xor_left_commute = xor.left_commute + +lemmas xor_ac = xor.assoc xor.commute xor.left_commute + +lemma xor_zero_right: "x \<^bold>\ \<^bold>0 = x" + by (fact xor.comm_neutral) + +lemma xor_zero_left: "\<^bold>0 \<^bold>\ x = x" + by (fact xor.left_neutral) + +end + +abbreviation (input) test_bit :: \'a::semiring_bits \ nat \ bool\ + where \test_bit \ bit\ + +abbreviation (input) bin_nth :: \int \ nat \ bool\ + where \bin_nth \ bit\ + +abbreviation (input) bin_last :: \int \ bool\ + where \bin_last \ odd\ + +abbreviation (input) bin_rest :: \int \ int\ + where \bin_rest w \ w div 2\ + +abbreviation (input) bintrunc :: \nat \ int \ int\ + where \bintrunc \ take_bit\ + +abbreviation (input) sbintrunc :: \nat \ int \ int\ + where \sbintrunc \ signed_take_bit\ + +abbreviation (input) bin_cat :: \int \ nat \ int \ int\ + where \bin_cat k n l \ concat_bit n l k\ + +abbreviation (input) norm_sint :: \nat \ int \ int\ + where \norm_sint n \ signed_take_bit (n - 1)\ + +abbreviation (input) max_word :: \'a::len word\ + where \max_word \ - 1\ + +abbreviation (input) complement :: \'a::len word \ 'a word\ + where \complement \ not\ lemma complement_mask: - "complement (2 ^ n - 1) = NOT (mask n)" - unfolding complement_def mask_eq_decr_exp by simp + "complement (2 ^ n - 1) = not (mask n)" + unfolding mask_eq_decr_exp by simp + +context + includes bit_operations_syntax +begin + +abbreviation (input) bshiftr1 :: \bool \ 'a::len word \ 'a word\ + where \bshiftr1 b w \ w div 2 OR push_bit (LENGTH('a) - Suc 0) (of_bool b) \ + +end + +lemma bit_bshiftr1_iff: + \bit (bshiftr1 b w) n \ b \ n = LENGTH('a) - 1 \ bit w (Suc n)\ + for w :: \'a::len word\ + by (auto simp add: bit_simps simp flip: bit_Suc) + +abbreviation (input) setBit :: \'a::len word \ nat \ 'a word\ + where \setBit w n \ set_bit n w\ + +abbreviation (input) clearBit :: \'a::len word \ nat \ 'a word\ + where \clearBit w n \ unset_bit n w\ + +lemma bit_setBit_iff: + \bit (setBit w m) n \ (m = n \ n < LENGTH('a) \ bit w n)\ + for w :: \'a::len word\ + by (auto simp add: bit_simps) + +lemma bit_clearBit_iff: + \bit (clearBit w m) n \ m \ n \ bit w n\ + for w :: \'a::len word\ + by (auto simp add: bit_simps) lemmas less_def = less_eq [symmetric] lemmas le_def = not_less [symmetric, where ?'a = nat] diff --git a/lib/Word_Lib/Machine_Word_32.thy b/lib/Word_Lib/Machine_Word_32.thy new file mode 100644 index 000000000..fb3f5fb90 --- /dev/null +++ b/lib/Word_Lib/Machine_Word_32.thy @@ -0,0 +1,136 @@ +(* + * Copyright 2014, NICTA + * + * This software may be distributed and modified according to the terms of + * the BSD 2-Clause license. Note that NO WARRANTY is provided. + * See "LICENSE_BSD2.txt" for details. + * + * @TAG(NICTA_BSD) + *) + +section "32-Bit Machine Word Setup" + +theory Machine_Word_32 + imports Machine_Word_32_Basics More_Word Bit_Shifts_Infix_Syntax Rsplit +begin + +context + includes bit_operations_syntax +begin + +type_synonym machine_word = \machine_word_len word\ + +lemma word_bits_len_of: + \LENGTH(machine_word_len) = word_bits\ + by (simp add: word_bits_conv) + +lemma word_bits_size: + "size (w :: machine_word) = word_bits" + by (simp add: word_bits_def word_size) + +lemma word_bits_word_size_conv: + \word_bits = word_size * 8\ + by (simp add: word_bits_def word_size_def) + +lemma word_size_word_size_bits: + \word_size = (2 :: 'a :: semiring_1) ^ word_size_bits\ + by (simp add: word_size_def word_size_bits_def) + +lemma lt_word_bits_lt_pow: + "sz < word_bits \ sz < 2 ^ word_bits" + by (simp add: word_bits_conv) + +lemma if_then_1_else_0: + "((if P then 1 else 0) = (0 :: machine_word)) = (\ P)" + by simp + +lemma if_then_0_else_1: + "((if P then 0 else 1) = (0 :: machine_word)) = (P)" + by simp + +lemmas if_then_simps = if_then_0_else_1 if_then_1_else_0 + +lemma bool_mask [simp]: + \0 < x AND 1 \ x AND 1 = 1\ for x :: machine_word + by (rule bool_mask') auto + +lemma in_16_range: + "0 \ S \ r \ (\x. r + x * (16 :: machine_word)) ` S" + "n - 1 \ S \ (r + (16 * n - 16)) \ (\x :: machine_word. r + x * 16) ` S" + by (clarsimp simp: image_def elim!: bexI[rotated])+ + +lemma le_step_down_word_3: + fixes x :: machine_word + shows "\x \ y; x \ y\ \ x \ y - 1" + by (fact le_step_down_word_2) + +lemma shiftr_1: + "(x::machine_word) >> 1 = 0 \ x < 2" + apply transfer + apply (simp add: take_bit_drop_bit) + apply (simp add: drop_bit_Suc) + done + +lemma Suc_unat_mask_div: + "Suc (unat (mask sz div word_size :: machine_word)) = 2 ^ (min sz word_bits - word_size_bits)" + by (simp add: word_size_word_size_bits unat_drop_bit_eq unat_mask_eq drop_bit_mask_eq Suc_mask_eq_exp + flip: drop_bit_eq_div word_bits_conv) + +lemma ucast_not_helper: + fixes a::"8 word" + assumes a: "a \ 0xFF" + shows "ucast a \ (0xFF::machine_word)" +proof + assume "ucast a = (0xFF::machine_word)" + also + have "(0xFF::machine_word) = ucast (0xFF::8 word)" by simp + finally + show False using a + apply - + apply (drule up_ucast_inj, simp) + apply simp + done +qed + +lemma unat_less_2p_word_bits: + "unat (x :: machine_word) < 2 ^ word_bits" + apply (simp only: word_bits_def) + apply (rule unat_lt2p) + done + +lemma unat_less_word_bits: + fixes y :: machine_word + shows "x < unat y \ x < 2 ^ word_bits" + unfolding word_bits_def + by (rule order_less_trans [OF _ unat_lt2p]) + +lemma unat_mask_2_less_4: + "unat (p AND mask 2 :: machine_word) < 4" + by (rule unat_less_helper) (simp only: take_bit_eq_mod word_mod_less_divisor flip: take_bit_eq_mask, simp add: word_mod_less_divisor) + +lemma unat_mult_simple: + \unat (x * y) = unat x * unat y\ + if \unat x * unat y < 2 ^ LENGTH(machine_word_len)\ + for x y :: machine_word + using that by (simp flip: unat_mult_lem) + +lemma upto_2_helper: + "{0..<2 :: machine_word} = {0, 1}" + by (safe; simp) unat_arith + +lemma word_ge_min: + \- (2 ^ (word_bits - 1)) \ sint x\ for x :: machine_word + using sint_ge [of x] by (simp add: word_bits_def) + +lemma word_rsplit_0: + "word_rsplit (0 :: machine_word) = replicate (word_bits div 8) (0 :: 8 word)" + by (simp add: word_rsplit_def bin_rsplit_def word_bits_def word_size_def Cons_replicate_eq) + +lemma x_less_2_0_1: + fixes x :: machine_word + shows "x < 2 \ x = 0 \ x = 1" + by (rule x_less_2_0_1') auto + +end + +end diff --git a/lib/Word_Lib/Machine_Word_32_Basics.thy b/lib/Word_Lib/Machine_Word_32_Basics.thy new file mode 100644 index 000000000..51de6baa8 --- /dev/null +++ b/lib/Word_Lib/Machine_Word_32_Basics.thy @@ -0,0 +1,66 @@ +(* + * Copyright 2014, NICTA + * + * This software may be distributed and modified according to the terms of + * the BSD 2-Clause license. Note that NO WARRANTY is provided. + * See "LICENSE_BSD2.txt" for details. + * + * @TAG(NICTA_BSD) + *) + +section \32 bit standard platform-specific word size and alignment.\ + +theory Machine_Word_32_Basics +imports "HOL-Library.Word" Word_32 +begin + +type_synonym machine_word_len = 32 + +definition word_bits :: nat +where + \word_bits = LENGTH(machine_word_len)\ + +lemma word_bits_conv [code]: + \word_bits = 32\ + by (simp add: word_bits_def) + +text \The following two are numerals so they can be used as nats and words.\ + +definition word_size_bits :: \'a :: numeral\ +where + \word_size_bits = 2\ + +definition word_size :: \'a :: numeral\ +where + \word_size = 4\ + +lemma n_less_word_bits: + "(n < word_bits) = (n < 32)" + by (simp add: word_bits_def word_size_def) + +lemmas upper_bits_unset_is_l2p_32 = upper_bits_unset_is_l2p [where 'a=32, folded word_bits_def] + +lemmas le_2p_upper_bits_32 = le_2p_upper_bits [where 'a=32, folded word_bits_def] +lemmas le2p_bits_unset_32 = le2p_bits_unset[where 'a=32, folded word_bits_def] + +lemmas unat_power_lower32 [simp] = unat_power_lower32'[folded word_bits_def] + +lemmas word32_less_sub_le[simp] = word32_less_sub_le' [folded word_bits_def] + +lemmas word32_power_less_1[simp] = word32_power_less_1'[folded word_bits_def] + +lemma of_nat32_0: + "\of_nat n = (0::word32); n < 2 ^ word_bits\ \ n = 0" + by (erule of_nat_0, simp add: word_bits_def) + +lemmas unat_of_nat32 = unat_of_nat32'[folded word_bits_def] + +lemmas word_power_nonzero_32 = word_power_nonzero [where 'a=32, folded word_bits_def] + +lemmas div_power_helper_32 = div_power_helper [where 'a=32, folded word_bits_def] + +lemmas of_nat_less_pow_32 = of_nat_power [where 'a=32, folded word_bits_def] + +lemmas unat_mask_word32 = unat_mask_word32'[folded word_bits_def] + +end diff --git a/lib/Word_Lib/Machine_Word_64.thy b/lib/Word_Lib/Machine_Word_64.thy new file mode 100644 index 000000000..f5f5c62c8 --- /dev/null +++ b/lib/Word_Lib/Machine_Word_64.thy @@ -0,0 +1,136 @@ +(* + * Copyright 2014, NICTA + * + * This software may be distributed and modified according to the terms of + * the BSD 2-Clause license. Note that NO WARRANTY is provided. + * See "LICENSE_BSD2.txt" for details. + * + * @TAG(NICTA_BSD) + *) + +section "64-Bit Machine Word Setup" + +theory Machine_Word_64 +imports Machine_Word_64_Basics More_Word Bit_Shifts_Infix_Syntax Rsplit +begin + +context + includes bit_operations_syntax +begin + +type_synonym machine_word = \machine_word_len word\ + +lemma word_bits_len_of: + \LENGTH(machine_word_len) = word_bits\ + by (simp add: word_bits_conv) + +lemma word_bits_size: + "size (w :: machine_word) = word_bits" + by (simp add: word_bits_def word_size) + +lemma word_bits_word_size_conv: + \word_bits = word_size * 8\ + by (simp add: word_bits_def word_size_def) + +lemma word_size_word_size_bits: + \word_size = (2 :: 'a :: semiring_1) ^ word_size_bits\ + by (simp add: word_size_def word_size_bits_def) + +lemma lt_word_bits_lt_pow: + "sz < word_bits \ sz < 2 ^ word_bits" + by (simp add: word_bits_conv) + +lemma if_then_1_else_0: + "((if P then 1 else 0) = (0 :: machine_word)) = (\ P)" + by simp + +lemma if_then_0_else_1: + "((if P then 0 else 1) = (0 :: machine_word)) = (P)" + by simp + +lemmas if_then_simps = if_then_0_else_1 if_then_1_else_0 + +lemma bool_mask [simp]: + \0 < x AND 1 \ x AND 1 = 1\ for x :: machine_word + by (rule bool_mask') auto + +lemma in_16_range: + "0 \ S \ r \ (\x. r + x * (16 :: machine_word)) ` S" + "n - 1 \ S \ (r + (16 * n - 16)) \ (\x :: machine_word. r + x * 16) ` S" + by (clarsimp simp: image_def elim!: bexI[rotated])+ + +lemma le_step_down_word_3: + fixes x :: machine_word + shows "\x \ y; x \ y\ \ x \ y - 1" + by (fact le_step_down_word_2) + +lemma shiftr_1: + "(x::machine_word) >> 1 = 0 \ x < 2" + apply transfer + apply (simp add: take_bit_drop_bit) + apply (simp add: drop_bit_Suc) + done + +lemma Suc_unat_mask_div: + "Suc (unat (mask sz div word_size :: machine_word)) = 2 ^ (min sz word_bits - word_size_bits)" + by (simp add: word_size_word_size_bits unat_drop_bit_eq unat_mask_eq drop_bit_mask_eq Suc_mask_eq_exp + flip: drop_bit_eq_div word_bits_conv) + +lemma ucast_not_helper: + fixes a::"8 word" + assumes a: "a \ 0xFF" + shows "ucast a \ (0xFF::machine_word)" +proof + assume "ucast a = (0xFF::machine_word)" + also + have "(0xFF::machine_word) = ucast (0xFF::8 word)" by simp + finally + show False using a + apply - + apply (drule up_ucast_inj, simp) + apply simp + done +qed + +lemma unat_less_2p_word_bits: + "unat (x :: machine_word) < 2 ^ word_bits" + apply (simp only: word_bits_def) + apply (rule unat_lt2p) + done + +lemma unat_less_word_bits: + fixes y :: machine_word + shows "x < unat y \ x < 2 ^ word_bits" + unfolding word_bits_def + by (rule order_less_trans [OF _ unat_lt2p]) + +lemma unat_mask_2_less_4: + "unat (p AND mask 2 :: machine_word) < 4" + by (rule unat_less_helper) (simp only: take_bit_eq_mod word_mod_less_divisor flip: take_bit_eq_mask, simp add: word_mod_less_divisor) + +lemma unat_mult_simple: + \unat (x * y) = unat x * unat y\ + if \unat x * unat y < 2 ^ LENGTH(machine_word_len)\ + for x y :: machine_word + using that by (simp flip: unat_mult_lem) + +lemma upto_2_helper: + "{0..<2 :: machine_word} = {0, 1}" + by (safe; simp) unat_arith + +lemma word_ge_min: + \- (2 ^ (word_bits - 1)) \ sint x\ for x :: machine_word + using sint_ge [of x] by (simp add: word_bits_def) + +lemma word_rsplit_0: + "word_rsplit (0 :: machine_word) = replicate (word_bits div 8) (0 :: 8 word)" + by (simp add: word_rsplit_def bin_rsplit_def word_bits_def word_size_def Cons_replicate_eq) + +lemma x_less_2_0_1: + fixes x :: machine_word + shows "x < 2 \ x = 0 \ x = 1" + by (rule x_less_2_0_1') auto + +end + +end diff --git a/lib/Word_Lib/Machine_Word_64_Basics.thy b/lib/Word_Lib/Machine_Word_64_Basics.thy new file mode 100644 index 000000000..d37ef5d7c --- /dev/null +++ b/lib/Word_Lib/Machine_Word_64_Basics.thy @@ -0,0 +1,66 @@ +(* + * Copyright 2014, NICTA + * + * This software may be distributed and modified according to the terms of + * the BSD 2-Clause license. Note that NO WARRANTY is provided. + * See "LICENSE_BSD2.txt" for details. + * + * @TAG(NICTA_BSD) + *) + +section \64 bit standard platform-specific word size and alignment.\ + +theory Machine_Word_64_Basics +imports "HOL-Library.Word" Word_64 +begin + +type_synonym machine_word_len = 64 + +definition word_bits :: nat +where + \word_bits = LENGTH(machine_word_len)\ + +lemma word_bits_conv [code]: + \word_bits = 64\ + by (simp add: word_bits_def) + +text \The following two are numerals so they can be used as nats and words.\ + +definition word_size_bits :: \'a :: numeral\ +where + \word_size_bits = 3\ + +definition word_size :: \'a :: numeral\ +where + \word_size = 8\ + +lemma n_less_word_bits: + "(n < word_bits) = (n < 64)" + by (simp add: word_bits_def word_size_def) + +lemmas upper_bits_unset_is_l2p_64 = upper_bits_unset_is_l2p [where 'a=64, folded word_bits_def] + +lemmas le_2p_upper_bits_64 = le_2p_upper_bits [where 'a=64, folded word_bits_def] +lemmas le2p_bits_unset_64 = le2p_bits_unset[where 'a=64, folded word_bits_def] + +lemmas unat_power_lower64 [simp] = unat_power_lower64'[folded word_bits_def] + +lemmas word64_less_sub_le[simp] = word64_less_sub_le' [folded word_bits_def] + +lemmas word64_power_less_1[simp] = word64_power_less_1'[folded word_bits_def] + +lemma of_nat64_0: + "\of_nat n = (0::word64); n < 2 ^ word_bits\ \ n = 0" + by (erule of_nat_0, simp add: word_bits_def) + +lemmas unat_of_nat64 = unat_of_nat64'[folded word_bits_def] + +lemmas word_power_nonzero_64 = word_power_nonzero [where 'a=64, folded word_bits_def] + +lemmas div_power_helper_64 = div_power_helper [where 'a=64, folded word_bits_def] + +lemmas of_nat_less_pow_64 = of_nat_power [where 'a=64, folded word_bits_def] + +lemmas unat_mask_word64 = unat_mask_word64'[folded word_bits_def] + +end diff --git a/lib/Word_Lib/Many_More.thy b/lib/Word_Lib/Many_More.thy index e24e64f90..c15241ec4 100644 --- a/lib/Word_Lib/Many_More.thy +++ b/lib/Word_Lib/Many_More.thy @@ -18,10 +18,16 @@ lemma nat_less_mult_monoish: "\ a < b; c < (d :: nat) \ \ (a = x) = (b = x)" by simp diff --git a/lib/Word_Lib/More_Arithmetic.thy b/lib/Word_Lib/More_Arithmetic.thy index b51d1fed1..623dfdaa4 100644 --- a/lib/Word_Lib/More_Arithmetic.thy +++ b/lib/Word_Lib/More_Arithmetic.thy @@ -7,14 +7,10 @@ section \Arithmetic lemmas\ theory More_Arithmetic - imports Main "HOL-Library.Type_Length" "HOL-Library.Bit_Operations" + imports Main "HOL-Library.Type_Length" begin -declare iszero_0 [intro] - -declare min.absorb1 [simp] min.absorb2 [simp] - -lemma n_less_equal_power_2 [simp]: +lemma n_less_equal_power_2: "n < 2 ^ n" by (fact less_exp) @@ -138,4 +134,10 @@ proof - by simp qed +lemma msrevs: + "0 < n \ (k * n + m) div n = m div n + k" + "(k * n + m) mod n = m mod n" + for n :: nat + by simp_all + end diff --git a/lib/Word_Lib/More_Divides.thy b/lib/Word_Lib/More_Divides.thy index 45acf717f..670615546 100644 --- a/lib/Word_Lib/More_Divides.thy +++ b/lib/Word_Lib/More_Divides.thy @@ -403,13 +403,11 @@ lemma two_pow_div_gt_le: lemma td_gal_lt: \0 < c \ a < b * c \ a div c < b\ for a b c :: nat - apply (auto dest: less_mult_imp_div_less) - apply (metis div_le_mono div_mult_self_is_m leD leI) - done + by (simp add: div_less_iff_less_mult) lemma td_gal: \0 < c \ b * c \ a \ b \ a div c\ for a b c :: nat - by (meson not_le td_gal_lt) + by (simp add: less_eq_div_iff_mult_less_eq) end diff --git a/lib/Word_Lib/More_Word.thy b/lib/Word_Lib/More_Word.thy index 2d83eef81..4623d21af 100644 --- a/lib/Word_Lib/More_Word.thy +++ b/lib/Word_Lib/More_Word.thy @@ -13,6 +13,63 @@ theory More_Word More_Divides begin +context unique_euclidean_semiring_with_bit_operations \\TODO: move\ +begin + +lemma possible_bit [simp]: + \possible_bit TYPE('a) n\ + by (simp add: possible_bit_def) + +lemma drop_bit_mask_eq: + \drop_bit m (mask n) = mask (n - m)\ + by (rule bit_eqI) (auto simp add: bit_simps possible_bit_def) + +end + +context + includes bit_operations_syntax +begin + +\ \problem posed by TPHOLs referee: + criterion for overflow of addition of signed integers\ + +lemma sofl_test: + \sint x + sint y = sint (x + y) \ + drop_bit (size x - 1) ((x + y XOR x) AND (x + y XOR y)) = 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!: *) + apply (smt (z3) take_bit_nonnegative) + apply (smt (z3) take_bit_int_less_exp) + apply (smt (z3) take_bit_nonnegative) + apply (smt (z3) take_bit_int_less_exp) + done + then show ?thesis + apply (simp only: One_nat_def word_size drop_bit_eq_zero_iff_not_bit_last bit_and_iff bit_xor_iff) + apply (simp add: bit_last_iff) + done +qed + lemma unat_power_lower [simp]: "unat ((2::'a::len word) ^ n) = 2 ^ n" if "n < LENGTH('a::len)" using that by transfer simp @@ -29,13 +86,6 @@ lemma word_div_eq_1_iff: "n div m = 1 \ n \ m \ una apply (simp flip: unat_div unsigned_take_bit_eq) done -lemma shiftl_power: - "(shiftl1 ^^ x) (y::'a::len word) = 2 ^ x * y" - apply (induct x) - apply simp - apply (simp add: shiftl1_2t) - done - lemma AND_twice [simp]: "(w AND m) AND m = w AND m" by (fact and.right_idem) @@ -104,21 +154,6 @@ lemma is_aligned_AND_less_0: apply (auto simp add: bit_simps) done -lemma le_shiftr1: - \shiftr1 u \ shiftr1 v\ if \u \ v\ -using that proof transfer - fix k l :: int - assume \take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ - then have \take_bit LENGTH('a) (drop_bit 1 (take_bit LENGTH('a) k)) - \ take_bit LENGTH('a) (drop_bit 1 (take_bit LENGTH('a) l))\ - apply (simp add: take_bit_drop_bit min_def) - apply (simp add: drop_bit_eq_div) - done - then show \take_bit LENGTH('a) (take_bit LENGTH('a) k div 2) - \ take_bit LENGTH('a) (take_bit LENGTH('a) l div 2)\ - by (simp add: drop_bit_eq_div) -qed - lemma and_mask_eq_iff_le_mask: \w AND mask n = w \ w \ mask n\ for w :: \'a::len word\ @@ -183,19 +218,12 @@ lemma of_nat_mono_maybe: lemma word_and_max_word: fixes a::"'a::len word" - shows "x = max_word \ a AND x = a" + shows "x = - 1 \ a AND x = a" by simp lemma word_and_full_mask_simp: \x AND mask LENGTH('a) = x\ for x :: \'a::len word\ -proof (rule bit_eqI) - fix n - assume \2 ^ n \ (0 :: 'a word)\ - then have \n < LENGTH('a)\ - by simp - then show \bit (x AND Bit_Operations.mask LENGTH('a)) n \ bit x n\ - by (simp add: bit_and_iff bit_mask_iff) -qed + by (simp add: bit_eq_iff bit_simps) lemma of_int_uint: "of_int (uint x) = x" @@ -215,10 +243,6 @@ corollary word_plus_and_or_coroll2: apply (simp flip: bit.conj_disj_distrib) done -lemma nat_mask_eq: - \nat (mask n) = mask n\ - by (simp add: nat_eq_iff of_nat_mask_eq) - lemma unat_mask_eq: \unat (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ by transfer (simp add: nat_mask_eq) @@ -650,7 +674,7 @@ lemma unat_mult_power_lem: proof (cases \sz < LENGTH('a)\) case True with assms show ?thesis - by (simp add: unat_word_ariths take_bit_eq_mod mod_simps) + by (simp add: unat_word_ariths take_bit_eq_mod mod_simps unsigned_of_nat) (simp add: take_bit_nat_eq_self_iff nat_less_power_trans flip: take_bit_eq_mod) next case False @@ -678,7 +702,7 @@ lemma word_less_sub_le[simp]: lemma unat_of_nat_len: "x < 2 ^ LENGTH('a) \ unat (of_nat x :: 'a::len word) = x" - by (simp add: take_bit_nat_eq_self_iff) + by (simp add: unsigned_of_nat take_bit_nat_eq_self_iff) lemma unat_of_nat_eq: "x < 2 ^ LENGTH('a) \ unat (of_nat x ::'a::len word) = x" @@ -781,7 +805,7 @@ lemma plus_one_helper2: lemma less_x_plus_1: fixes x :: "'a :: len word" shows - "x \ max_word \ (y < (x + 1)) = (y < x \ y = x)" + "x \ - 1 \ (y < (x + 1)) = (y < x \ y = x)" apply (rule iffI) apply (rule disjCI) apply (drule plus_one_helper) @@ -794,11 +818,11 @@ lemma less_x_plus_1: done lemma word_Suc_leq: - fixes k::"'a::len word" shows "k \ max_word \ x < k + 1 \ x \ k" + fixes k::"'a::len word" shows "k \ - 1 \ x < k + 1 \ x \ k" using less_x_plus_1 word_le_less_eq by auto lemma word_Suc_le: - fixes k::"'a::len word" shows "x \ max_word \ x + 1 \ k \ x < k" + fixes k::"'a::len word" shows "x \ - 1 \ x + 1 \ k \ x < k" by (meson not_less word_Suc_leq) lemma word_lessThan_Suc_atMost: @@ -815,7 +839,7 @@ lemma word_atLeastAtMost_Suc_greaterThanAtMost: lemma word_atLeastLessThan_Suc_atLeastAtMost_union: fixes l::"'a::len word" - assumes "m \ max_word" and "l \ m" and "m \ u" + assumes "m \ - 1" and "l \ m" and "m \ u" shows "{l..m} \ {m+1..u} = {l..u}" proof - from ivl_disj_un_two(8)[OF assms(2) assms(3)] have "{l..u} = {l..m} \ {m<..u}" by blast @@ -877,10 +901,7 @@ lemma scast_1: lemma unat_minus_one_word: "unat (-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" - apply (simp only: flip: mask_eq_exp_minus_1) - apply transfer - apply (simp add: take_bit_minus_one_eq_mask nat_mask_eq) - done + by (simp add: mask_eq_exp_minus_1 unsigned_minus_1_eq_mask) lemmas word_diff_ls'' = word_diff_ls [where xa=x and x=x for x] lemmas word_diff_ls' = word_diff_ls'' [simplified] @@ -923,7 +944,7 @@ lemma word_of_nat_less: "\ n < unat x \ \ of_nat n < x" apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) - apply (simp add: take_bit_eq_mod) + apply (simp add: unsigned_of_nat take_bit_eq_mod) done lemma unat_mask: @@ -946,7 +967,7 @@ lemma Suc_2p_unat_mask: lemma sint_of_nat_ge_zero: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) \ 0" - by (simp add: bit_iff_odd) + by (simp add: bit_iff_odd signed_of_nat) lemma int_eq_sint: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) = int x" @@ -1048,7 +1069,7 @@ lemma ucast_range_less: apply (rule_tac x="ucast x" in exI) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) - apply (metis bit_take_bit_iff less_mask_eq not_less take_bit_eq_mask) + apply (metis bit_take_bit_iff take_bit_word_eq_self_iff) done lemma word_power_less_diff: @@ -1130,7 +1151,7 @@ lemma unat_less_helper: "x < of_nat n \ unat x < n" apply (simp add: word_less_nat_alt) apply (erule order_less_le_trans) - apply (simp add: take_bit_eq_mod) + apply (simp add: take_bit_eq_mod unsigned_of_nat) done lemma nat_uint_less_helper: @@ -1142,7 +1163,7 @@ lemma nat_uint_less_helper: lemma of_nat_0: "\of_nat n = (0::'a::len word); n < 2 ^ LENGTH('a)\ \ n = 0" - by transfer (simp add: take_bit_eq_mod) + by (auto simp add: word_of_nat_eq_0_iff) lemma of_nat_inj: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ @@ -1285,7 +1306,7 @@ lemma mask_eq_x_eq_0: by auto lemma compl_of_1: "NOT 1 = (-2 :: 'a :: len word)" - by (fact not_one) + by (fact not_one_eq) lemma split_word_eq_on_mask: "(x = y) = (x AND m = y AND m \ x AND NOT m = y AND NOT m)" @@ -1353,10 +1374,6 @@ lemma word_minus_one_le_leq: apply arith done -lemma word_div_less: - "m < n \ m div n = 0" for m :: "'a :: len word" - by (simp add: unat_mono word_arith_nat_defs(6)) - lemma word_must_wrap: "\ x \ n - 1; n \ x \ \ n = (0 :: 'a :: len word)" using dual_order.trans sub_wrap word_less_1 by blast @@ -1448,7 +1465,10 @@ lemma word_less_power_trans_ofnat: \ of_nat n * 2 ^ k < (2::'a::len word) ^ m" apply (subst mult.commute) apply (rule word_less_power_trans) - apply (simp_all add: word_less_nat_alt less_le_trans take_bit_eq_mod) + apply (simp_all add: word_less_nat_alt unsigned_of_nat) + using take_bit_nat_less_eq_self + apply (rule le_less_trans) + apply assumption done lemma word_1_le_power: @@ -1664,7 +1684,8 @@ lemma sint_minus1 [simp]: "(sint x = -1) = (x = -1)" apply (cases \LENGTH('a)\) apply simp_all apply transfer - apply (simp flip: signed_take_bit_eq_iff_take_bit_eq) + apply (simp only: flip: signed_take_bit_eq_iff_take_bit_eq) + apply simp done lemma sint_0 [simp]: "(sint x = 0) = (x = 0)" @@ -1712,7 +1733,7 @@ lemma uint_range': lemma sint_of_int_eq: "\ - (2 ^ (LENGTH('a) - 1)) \ x; x < 2 ^ (LENGTH('a) - 1) \ \ sint (of_int x :: ('a::len) word) = x" - by (simp add: signed_take_bit_int_eq_self) + by (simp add: signed_take_bit_int_eq_self signed_of_int) lemma of_int_sint: "of_int (sint a) = a" @@ -1788,7 +1809,7 @@ lemma le_step_down_word_2: simp add: word_le_minus_one_leq) lemma NOT_mask_AND_mask[simp]: "(w AND mask n) AND NOT (mask n) = 0" - by (clarsimp simp add: mask_eq_decr_exp Parity.bit_eq_iff bit_and_iff bit_not_iff bit_mask_iff) + by (rule bit_eqI) (simp add: bit_simps) lemma and_and_not[simp]:"(a AND b) AND NOT b = 0" for a b :: \'a::len word\ @@ -1804,4 +1825,739 @@ lemma ex_mask_1[simp]: "(\x. mask x = (1 :: 'a::len word))" lemma not_switch:"NOT a = x \ a = NOT x" by auto +lemma test_bit_eq_iff: "bit u = bit v \ u = v" + for u v :: "'a::len word" + by (auto intro: bit_eqI simp add: fun_eq_iff) + +lemma test_bit_size: "bit 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 \ bit u n = bit v n) \ u = v" + for u :: "'a::len word" + by (simp add: word_size word_eq_iff) + +lemma word_eqD: "u = v \ bit u x = bit v x" + for u v :: "'a::len word" + by simp + +lemma test_bit_bin': "bit 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: + \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) (bit :: 'a::len word \ _)" + by transfer_prover + +lemma test_bit_wi: + "bit (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 \ + bit (x OR y) n = (bit x n | bit y n) \ + bit (x AND y) n = (bit x n \ bit y n) \ + bit (x XOR y) n = (bit x n \ bit y n) \ + bit (NOT x) n = (\ bit 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: + "bit (x OR y) n = (bit x n | bit y n) \ + bit (x AND y) n = (bit x n \ bit y n)" + for x :: "'a::len word" + by transfer (auto simp add: bit_or_iff bit_and_iff) + +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 bit w n else bit 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: "bit (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: "bit ((2::'a::len word) ^ n) m \ m = n \ m < LENGTH('a::len)" + by transfer (auto simp add: bit_exp_iff) + +lemma bang_is_le: "bit 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 + +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_1 [iff]: "bit (1 :: 'a::len word) n \ n = 0" + by transfer (auto simp add: bit_1_iff) + +lemma nth_0: "\ bit (0 :: 'a::len word) n" + by transfer simp + +lemma nth_minus1: "bit (-1 :: 'a::len word) n \ n < LENGTH('a)" + by transfer simp + +lemma nth_ucast: + "bit (ucast w::'a::len word) n = (bit w n \ n < LENGTH('a))" + by transfer (simp add: bit_take_bit_iff ac_simps) + +lemma drop_bit_numeral_bit0_1 [simp]: + \drop_bit (Suc 0) (numeral k) = + (word_of_int (drop_bit (Suc 0) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\ + by (metis Word_eq_word_of_int drop_bit_word.abs_eq of_int_numeral) + +lemma nth_mask: + \bit (mask n :: 'a::len word) i \ i < n \ i < size (mask n :: 'a word)\ + by (auto simp add: word_size Word.bit_mask_iff) + +lemma nth_slice: "bit (slice n w :: 'a::len word) m = (bit w (m + n) \ m < LENGTH('a))" + apply (auto simp add: bit_simps less_diff_conv dest: bit_imp_le_length) + using bit_imp_le_length + apply fastforce + done + +lemma test_bit_cat [OF refl]: + "wc = word_cat a b \ bit wc n = (n < size wc \ + (if n < size b then bit b n else bit 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. + bit b n = (n < size b \ bit c n) \ + bit a m = (m < size a \ bit c (m + size b)))" + by (auto simp add: word_split_bin' bit_unsigned_iff word_size bit_drop_bit_eq ac_simps + dest: bit_imp_le_length) + +lemma test_bit_split: + "word_split c = (a, b) \ + (\n::nat. bit b n \ n < size b \ bit c n) \ + (\m::nat. bit a m \ m < size a \ bit c (m + size b))" + by (simp add: test_bit_split') + +lemma test_bit_split_eq: + "word_split c = (a, b) \ + ((\n::nat. bit b n = (n < size b \ bit c n)) \ + (\m::nat. bit a m = (m < size a \ bit 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 \ bit rc n = + (n < size rc \ n div sw < size wl \ bit ((rev wl) ! (n div sw)) (n mod sw))" + for wl :: "'a::len word list" + by (simp add: word_size word_rcat_def rev_map bit_horner_sum_uint_exp_iff bit_simps not_le) + +lemmas test_bit_cong = arg_cong [where f = "bit", THEN fun_cong] + +lemma max_test_bit: "bit (- 1::'a::len word) n \ n < LENGTH('a)" + by (fact nth_minus1) + +lemma map_nth_0 [simp]: "map (bit (0::'a::len word)) xs = replicate (length xs) False" + by (simp flip: map_replicate_const) + +lemma word_and_1: + "n AND 1 = (if bit n 0 then 1 else 0)" for n :: "_ word" + by (rule bit_word_eqI) (auto simp add: bit_and_iff bit_1_iff intro: gr0I) + +lemma test_bit_1': + "bit (1 :: 'a :: len word) n \ 0 < LENGTH('a) \ n = 0" + by simp + +lemma nth_w2p_same: + "bit (2^n :: 'a :: len word) n = (n < LENGTH('a))" + by (simp add: nth_w2p) + +lemma word_leI: + "(\n. \n < size (u::'a::len word); bit u n \ \ bit (v::'a::len word) n) \ u <= v" + apply (rule order_trans [of u \u AND v\ v]) + apply (rule eq_refl) + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps word_and_le1 word_size) + done + +lemma bang_eq: + fixes x :: "'a::len word" + shows "(x = y) = (\n. bit x n = bit y n)" + by (auto intro!: bit_eqI) + +lemma neg_mask_test_bit: + "bit (NOT(mask n) :: 'a :: len word) m = (n \ m \ m < LENGTH('a))" + by (auto simp add: bit_simps) + +lemma upper_bits_unset_is_l2p: + \(\n' \ n. n' < LENGTH('a) \ \ bit p n') \ (p < 2 ^ n)\ (is \?P \ ?Q\) + if \n < LENGTH('a)\ + for p :: "'a :: len word" +proof + assume ?Q + then show ?P + by (meson bang_is_le le_less_trans not_le word_power_increasing) +next + assume ?P + have \take_bit n p = p\ + proof (rule bit_word_eqI) + fix q + assume \q < LENGTH('a)\ + show \bit (take_bit n p) q \ bit p q\ + proof (cases \q < n\) + case True + then show ?thesis + by (auto simp add: bit_simps) + next + case False + then have \n \ q\ + by simp + with \?P\ \q < LENGTH('a)\ have \\ bit p q\ + by simp + then show ?thesis + by (simp add: bit_simps) + qed + qed + with that show ?Q + using take_bit_word_eq_self_iff [of n p] by auto +qed + +lemma less_2p_is_upper_bits_unset: + "p < 2 ^ n \ n < LENGTH('a) \ (\n' \ n. n' < LENGTH('a) \ \ bit p n')" for p :: "'a :: len word" + by (meson le_less_trans le_mask_iff_lt_2n upper_bits_unset_is_l2p word_zero_le) + +lemma test_bit_over: + "n \ size (x::'a::len word) \ (bit x n) = False" + by transfer auto + +lemma le_mask_high_bits: + "w \ mask n \ (\i \ {n ..< size w}. \ bit w i)" + for w :: \'a::len word\ + apply (auto simp add: bit_simps word_size less_eq_mask_iff_take_bit_eq_self) + apply (metis bit_take_bit_iff leD) + apply (metis atLeastLessThan_iff leI take_bit_word_eq_self_iff upper_bits_unset_is_l2p) + done + +lemma test_bit_conj_lt: + "(bit x m \ m < LENGTH('a)) = bit x m" for x :: "'a :: len word" + using test_bit_bin by blast + +lemma neg_test_bit: + "bit (NOT x) n = (\ bit x n \ n < LENGTH('a))" for x :: "'a::len word" + by (cases "n < LENGTH('a)") (auto simp add: test_bit_over word_ops_nth_size word_size) + +lemma nth_bounded: + "\bit (x :: 'a :: len word) n; x < 2 ^ m; m \ len_of TYPE ('a)\ \ n < m" + apply (rule ccontr) + apply (auto simp add: not_less) + apply (meson bit_imp_le_length bit_uint_iff less_2p_is_upper_bits_unset test_bit_bin) + done + +lemma and_neq_0_is_nth: + \x AND y \ 0 \ bit x n\ if \y = 2 ^ n\ for x y :: \'a::len word\ + apply (simp add: bit_eq_iff bit_simps) + using that apply (simp add: bit_simps not_le) + apply transfer + apply auto + done + +lemma nth_is_and_neq_0: + "bit (x::'a::len word) n = (x AND 2 ^ n \ 0)" + by (subst and_neq_0_is_nth; rule refl) + +lemma max_word_not_less [simp]: + "\ - 1 < x" for x :: \'a::len word\ + by (fact word_order.extremum_strict) + +lemma bit_twiddle_min: + "(y::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = min x y" + by (rule bit_eqI) (auto simp add: bit_simps) + +lemma bit_twiddle_max: + "(x::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = max x y" + by (rule bit_eqI) (auto simp add: bit_simps max_def) + +lemma swap_with_xor: + "\(x::'a::len word) = a XOR b; y = b XOR x; z = x XOR y\ \ z = b \ y = a" + by (auto intro: bit_word_eqI simp add: bit_simps) + +lemma le_mask_imp_and_mask: + "(x::'a::len word) \ mask n \ x AND mask n = x" + by (metis and_mask_eq_iff_le_mask) + +lemma or_not_mask_nop: + "((x::'a::len word) OR NOT (mask n)) AND mask n = x AND mask n" + by (metis word_and_not word_ao_dist2 word_bw_comms(1) word_log_esimps(3)) + +lemma mask_subsume: + "\n \ m\ \ ((x::'a::len word) OR y AND mask n) AND NOT (mask m) = x AND NOT (mask m)" + by (rule bit_word_eqI) (auto simp add: bit_simps word_size) + +lemma and_mask_0_iff_le_mask: + fixes w :: "'a::len word" + shows "(w AND NOT(mask n) = 0) = (w \ mask n)" + by (simp add: mask_eq_0_eq_x le_mask_imp_and_mask and_mask_eq_iff_le_mask) + +lemma mask_twice2: + "n \ m \ ((x::'a::len word) AND mask m) AND mask n = x AND mask n" + by (metis mask_twice min_def) + +lemma uint_2_id: + "LENGTH('a) \ 2 \ uint (2::('a::len) word) = 2" + by simp + +lemma div_of_0_id[simp]:"(0::('a::len) word) div n = 0" + by (simp add: word_div_def) + +lemma degenerate_word:"LENGTH('a) = 1 \ (x::('a::len) word) = 0 \ x = 1" + by (metis One_nat_def less_irrefl_nat sint_1_cases) + +lemma div_by_0_word:"(x::('a::len) word) div 0 = 0" + by (metis div_0 div_by_0 unat_0 word_arith_nat_defs(6) word_div_1) + +lemma div_less_dividend_word:"\x \ 0; n \ 1\ \ (x::('a::len) word) div n < x" + apply (cases \n = 0\) + apply clarsimp + apply (simp add:word_neq_0_conv) + apply (subst word_arith_nat_div) + apply (rule word_of_nat_less) + apply (rule div_less_dividend) + using unat_eq_zero word_unat_Rep_inject1 apply force + apply (simp add:unat_gt_0) + done + +lemma word_less_div: + fixes x :: "('a::len) word" + and y :: "('a::len) word" + shows "x div y = 0 \ y = 0 \ x < y" + apply (case_tac "y = 0", clarsimp+) + by (metis One_nat_def Suc_le_mono le0 le_div_geq not_less unat_0 unat_div unat_gt_0 word_less_nat_alt zero_less_one) + +lemma not_degenerate_imp_2_neq_0:"LENGTH('a) > 1 \ (2::('a::len) word) \ 0" + by (metis numerals(1) power_not_zero power_zero_numeral) + +lemma word_overflow:"(x::('a::len) word) + 1 > x \ x + 1 = 0" + apply clarsimp + by (metis diff_0 eq_diff_eq less_x_plus_1) + +lemma word_overflow_unat:"unat ((x::('a::len) word) + 1) = unat x + 1 \ x + 1 = 0" + by (metis Suc_eq_plus1 add.commute unatSuc) + +lemma even_word_imp_odd_next:"even (unat (x::('a::len) word)) \ x + 1 = 0 \ odd (unat (x + 1))" + apply (cut_tac x=x in word_overflow_unat) + apply clarsimp + done + +lemma odd_word_imp_even_next:"odd (unat (x::('a::len) word)) \ x + 1 = 0 \ even (unat (x + 1))" + apply (cut_tac x=x in word_overflow_unat) + apply clarsimp + done + +lemma overflow_imp_lsb:"(x::('a::len) word) + 1 = 0 \ bit x 0" + using even_plus_one_iff [of x] by simp + +lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = bit x 0" + by transfer (simp add: even_nat_iff) + +lemma of_nat_neq_iff_word: + "x mod 2 ^ LENGTH('a) \ y mod 2 ^ LENGTH('a) \ + (((of_nat x)::('a::len) word) \ of_nat y) = (x \ y)" + apply (rule iffI) + apply (case_tac "x = y") + apply (subst (asm) of_nat_eq_iff[symmetric]) + apply auto + apply (case_tac "((of_nat x)::('a::len) word) = of_nat y") + apply auto + apply (metis unat_of_nat) + done + +lemma lsb_this_or_next: "\ (bit ((x::('a::len) word) + 1) 0) \ bit x 0" + by simp + +lemma mask_or_not_mask: + "x AND mask n OR x AND NOT (mask n) = x" + for x :: \'a::len word\ + apply (subst word_oa_dist, simp) + apply (subst word_oa_dist2, simp) + done + +lemma word_gr0_conv_Suc: "(m::'a::len word) > 0 \ \n. m = n + 1" + by (metis add.commute add_minus_cancel) + +lemma revcast_down_us [OF refl]: + "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (signed_drop_bit n w)" + for w :: "'a::len word" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_simps ac_simps) + done + +lemma revcast_down_ss [OF refl]: + "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (signed_drop_bit n w)" + for w :: "'a::len word" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_simps ac_simps) + done + +lemma revcast_down_uu [OF refl]: + "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (drop_bit n w)" + for w :: "'a::len word" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_simps ac_simps) + done + +lemma revcast_down_su [OF refl]: + "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (drop_bit n w)" + for w :: "'a::len word" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_simps ac_simps) + done + +lemma cast_down_rev [OF refl]: + "uc = ucast \ source_size uc = target_size uc + n \ uc w = revcast (push_bit n w)" + for w :: "'a::len word" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_simps) + done + +lemma revcast_up [OF refl]: + "rc = revcast \ source_size rc + n = target_size rc \ + rc w = push_bit n (ucast w :: 'a::len word)" + apply (simp add: source_size_def target_size_def) + apply (rule bit_word_eqI) + apply (simp add: bit_simps) + 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]] + +lemma word_ops_nth: + fixes x y :: \'a::len word\ + shows + word_or_nth: "bit (x OR y) n = (bit x n \ bit y n)" and + word_and_nth: "bit (x AND y) n = (bit x n \ bit y n)" and + word_xor_nth: "bit (x XOR y) n = (bit x n \ bit y n)" + by (simp_all add: bit_simps) + +lemma word_power_nonzero: + "\ (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x \ 0 \ + \ x * 2 ^ n \ 0" + by (metis gr_implies_not0 mult_eq_0_iff nat_mult_power_less_eq numeral_2_eq_2 + p2_gt_0 unat_eq_zero unat_less_power unat_mult_lem unat_power_lower word_gt_a_gt_0 zero_less_Suc) + +lemma less_1_helper: + "n \ m \ (n - 1 :: int) < m" + by arith + +lemma div_power_helper: + "\ x \ y; y < LENGTH('a) \ \ (2 ^ y - 1) div (2 ^ x :: 'a::len word) = 2 ^ (y - x) - 1" + apply (simp flip: mask_eq_exp_minus_1 drop_bit_eq_div) + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps not_le) + done + +lemma max_word_mask: + "(- 1 :: 'a::len word) = mask LENGTH('a)" + by (fact minus_1_eq_mask) + +lemmas mask_len_max = max_word_mask[symmetric] + +lemma mask_out_first_mask_some: + "\ x AND NOT (mask n) = y; n \ m \ \ x AND NOT (mask m) = y AND NOT (mask m)" + for x y :: \'a::len word\ + by (rule bit_word_eqI) (auto simp add: bit_simps word_size) + +lemma mask_lower_twice: + "n \ m \ (x AND NOT (mask n)) AND NOT (mask m) = x AND NOT (mask m)" + for x :: \'a::len word\ + by (rule bit_word_eqI) (auto simp add: bit_simps word_size) + +lemma mask_lower_twice2: + "(a AND NOT (mask n)) AND NOT (mask m) = a AND NOT (mask (max n m))" + for a :: \'a::len word\ + by (rule bit_word_eqI) (auto simp add: bit_simps) + +lemma ucast_and_neg_mask: + "ucast (x AND NOT (mask n)) = ucast x AND NOT (mask n)" + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps dest: bit_imp_le_length) + done + +lemma ucast_and_mask: + "ucast (x AND mask n) = ucast x AND mask n" + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps dest: bit_imp_le_length) + done + +lemma ucast_mask_drop: + "LENGTH('a :: len) \ n \ (ucast (x AND mask n) :: 'a word) = ucast x" + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps dest: bit_imp_le_length) + done + +lemma mask_exceed: + "n \ LENGTH('a) \ (x::'a::len word) AND NOT (mask n) = 0" + by (rule bit_word_eqI) (simp add: bit_simps) + +lemma word_add_no_overflow:"(x::'a::len word) < - 1 \ x < x + 1" + using less_x_plus_1 order_less_le by blast + +lemma lt_plus_1_le_word: + fixes x :: "'a::len word" + assumes bound:"n < unat (maxBound::'a word)" + shows "x < 1 + of_nat n = (x \ of_nat n)" + by (metis add.commute bound max_word_max word_Suc_leq word_not_le word_of_nat_less) + +lemma unat_ucast_up_simp: + fixes x :: "'a::len word" + assumes "LENGTH('a) \ LENGTH('b)" + shows "unat (ucast x :: 'b::len word) = unat x" + apply (rule bit_eqI) + using assms apply (auto simp add: bit_simps dest: bit_imp_le_length) + done + +lemma unat_ucast_less_no_overflow: + "\n < 2 ^ LENGTH('a); unat f < n\ \ (f::('a::len) word) < of_nat n" + by (erule (1) order_le_less_trans[OF _ of_nat_mono_maybe,rotated]) simp + +lemma unat_ucast_less_no_overflow_simp: + "n < 2 ^ LENGTH('a) \ (unat f < n) = ((f::('a::len) word) < of_nat n)" + using unat_less_helper unat_ucast_less_no_overflow by blast + +lemma unat_ucast_no_overflow_le: + assumes no_overflow: "unat b < (2 :: nat) ^ LENGTH('a)" + and upward_cast: "LENGTH('a) < LENGTH('b)" + shows "(ucast (f::'a::len word) < (b :: 'b :: len word)) = (unat f < unat b)" +proof - + have LR: "ucast f < b \ unat f < unat b" + apply (rule unat_less_helper) + apply (simp add:ucast_nat_def) + apply (rule_tac 'b1 = 'b in ucast_less_ucast[OF order.strict_implies_order, THEN iffD1]) + apply (rule upward_cast) + apply (simp add: ucast_ucast_mask less_mask_eq word_less_nat_alt + unat_power_lower[OF upward_cast] no_overflow) + done + have RL: "unat f < unat b \ ucast f < b" + proof- + assume ineq: "unat f < unat b" + have "ucast (f::'a::len word) < ((ucast (ucast b ::'a::len word)) :: 'b :: len word)" + apply (simp add: ucast_less_ucast[OF order.strict_implies_order] upward_cast) + apply (simp only: flip: ucast_nat_def) + apply (rule unat_ucast_less_no_overflow[OF no_overflow ineq]) + done + then show ?thesis + apply (rule order_less_le_trans) + apply (simp add:ucast_ucast_mask word_and_le2) + done + qed + then show ?thesis by (simp add:RL LR iffI) +qed + +lemmas ucast_up_mono = ucast_less_ucast[THEN iffD2] + +lemma minus_one_word: + "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" + by simp + +lemma le_2p_upper_bits: + "\ (p::'a::len word) \ 2^n - 1; n < LENGTH('a) \ \ + \n'\n. n' < LENGTH('a) \ \ bit p n'" + by (subst upper_bits_unset_is_l2p; simp) + +lemma le2p_bits_unset: + "p \ 2 ^ n - 1 \ \n'\n. n' < LENGTH('a) \ \ bit (p::'a::len word) n'" + using upper_bits_unset_is_l2p [where p=p] + by (cases "n < LENGTH('a)") auto + +lemma complement_nth_w2p: + shows "n' < LENGTH('a) \ bit (NOT (2 ^ n :: 'a::len word)) n' = (n' \ n)" + by (fastforce simp: word_ops_nth_size word_size nth_w2p) + +lemma word_unat_and_lt: + "unat x < n \ unat y < n \ unat (x AND y) < n" + by (meson le_less_trans word_and_le1 word_and_le2 word_le_nat_alt) + +lemma word_unat_mask_lt: + "m \ size w \ unat ((w::'a::len word) AND mask m) < 2 ^ m" + by (rule word_unat_and_lt) (simp add: unat_mask word_size) + +lemma word_sless_sint_le:"x sint x \ sint y - 1" + by (metis word_sless_alt zle_diff1_eq) + +lemma upper_trivial: + fixes x :: "'a::len word" + shows "x \ 2 ^ LENGTH('a) - 1 \ x < 2 ^ LENGTH('a) - 1" + by (simp add: less_le) + +lemma constraint_expand: + fixes x :: "'a::len word" + shows "x \ {y. lower \ y \ y \ upper} = (lower \ x \ x \ upper)" + by (rule mem_Collect_eq) + +lemma card_map_elide: + "card ((of_nat :: nat \ 'a::len word) ` {0.. CARD('a::len word)" +proof - + let ?of_nat = "of_nat :: nat \ 'a word" + have "inj_on ?of_nat {i. i < CARD('a word)}" + by (rule inj_onI) (simp add: card_word of_nat_inj) + moreover have "{0.. {i. i < CARD('a word)}" + using that by auto + ultimately have "inj_on ?of_nat {0.. CARD('a::len word) \ card ((of_nat::nat \ 'a::len word) ` {0.. LENGTH('a) \ x = ucast y \ ucast x = y" + for x :: "'a::len word" and y :: "'b::len word" + by transfer simp + +lemma le_ucast_ucast_le: + "x \ ucast y \ ucast x \ y" + for x :: "'a::len word" and y :: "'b::len word" + by (smt le_unat_uoi linorder_not_less order_less_imp_le ucast_nat_def unat_arith_simps(1)) + +lemma less_ucast_ucast_less: + "LENGTH('b) \ LENGTH('a) \ x < ucast y \ ucast x < y" + for x :: "'a::len word" and y :: "'b::len word" + by (metis ucast_nat_def unat_mono unat_ucast_up_simp word_of_nat_less) + +lemma ucast_le_ucast: + "LENGTH('a) \ LENGTH('b) \ (ucast x \ (ucast y::'b::len word)) = (x \ y)" + for x :: "'a::len word" + by (simp add: unat_arith_simps(1) unat_ucast_up_simp) + +lemmas ucast_up_mono_le = ucast_le_ucast[THEN iffD2] + +lemma ucast_or_distrib: + fixes x :: "'a::len word" + fixes y :: "'a::len word" + shows "(ucast (x OR y) :: ('b::len) word) = ucast x OR ucast y" + by (fact unsigned_or_eq) + +lemma word_exists_nth: + "(w::'a::len word) \ 0 \ \i. bit w i" + by (auto simp add: bit_eq_iff) + +lemma max_word_not_0 [simp]: + "- 1 \ (0 :: 'a::len word)" + by simp + +lemma unat_max_word_pos[simp]: "0 < unat (- 1 :: 'a::len word)" + using unat_gt_0 [of \- 1 :: 'a::len word\] by simp + + +(* Miscellaneous conditional injectivity rules. *) + +lemma mult_pow2_inj: + assumes ws: "m + n \ LENGTH('a)" + assumes le: "x \ mask m" "y \ mask m" + assumes eq: "x * 2 ^ n = y * (2 ^ n::'a::len word)" + shows "x = y" +proof (rule bit_word_eqI) + fix q + assume \q < LENGTH('a)\ + from eq have \push_bit n x = push_bit n y\ + by (simp add: push_bit_eq_mult) + moreover from le have \take_bit m x = x\ \take_bit m y = y\ + by (simp_all add: less_eq_mask_iff_take_bit_eq_self) + ultimately have \push_bit n (take_bit m x) = push_bit n (take_bit m y)\ + by simp_all + with \q < LENGTH('a)\ ws show \bit x q \ bit y q\ + apply (simp add: push_bit_take_bit) + unfolding bit_eq_iff + apply (simp add: bit_simps not_le) + apply (metis (full_types) \take_bit m x = x\ \take_bit m y = y\ add.commute add_diff_cancel_right' add_less_cancel_right bit_take_bit_iff le_add2 less_le_trans) + done +qed + +lemma word_of_nat_inj: + assumes bounded: "x < 2 ^ LENGTH('a)" "y < 2 ^ LENGTH('a)" + assumes of_nats: "of_nat x = (of_nat y :: 'a::len word)" + shows "x = y" + by (rule contrapos_pp[OF of_nats]; cases "x < y"; cases "y < x") + (auto dest: bounded[THEN of_nat_mono_maybe]) + +lemma word_of_int_bin_cat_eq_iff: + "(word_of_int (concat_bit LENGTH('b) (uint b) (uint a))::'c::len word) = + word_of_int (concat_bit LENGTH('b) (uint d) (uint c)) \ b = d \ a = c" + if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" + for a::"'a::len word" and b::"'b::len word" +proof - + from that show ?thesis + using that concat_bit_eq_iff [of \LENGTH('b)\ \uint b\ \uint a\ \uint d\ \uint c\] + apply (simp add: word_of_int_eq_iff take_bit_int_eq_self flip: word_eq_iff_unsigned) + apply (simp add: concat_bit_def take_bit_int_eq_self bintr_uint take_bit_push_bit) + done +qed + +lemma word_cat_inj: "(word_cat a b::'c::len word) = word_cat c d \ a = c \ b = d" + if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" + for a::"'a::len word" and b::"'b::len word" + using word_of_int_bin_cat_eq_iff [OF that, of b a d c] + by (simp add: word_cat_eq' ac_simps) + +lemma p2_eq_1: "2 ^ n = (1::'a::len word) \ n = 0" +proof - + have "2 ^ n = (1::'a word) \ n = 0" + by (metis One_nat_def not_less one_less_numeral_iff p2_eq_0 p2_gt_0 power_0 power_0 + power_inject_exp semiring_norm(76) unat_power_lower zero_neq_one) + then show ?thesis by auto +qed + +end + +lemmas word_div_less = div_word_less + +(* FIXME: move to Word distribution? *) +lemma bin_nth_minus_Bit0[simp]: + "0 < n \ bit (numeral (num.Bit0 w) :: int) n = bit (numeral w :: int) (n - 1)" + by (cases n; simp) + +lemma bin_nth_minus_Bit1[simp]: + "0 < n \ bit (numeral (num.Bit1 w) :: int) n = bit (numeral w :: int) (n - 1)" + by (cases n; simp) + +lemma word_mod_by_0: "k mod (0::'a::len word) = k" + by (simp add: word_arith_nat_mod) + end diff --git a/lib/Word_Lib/More_Word_Operations.thy b/lib/Word_Lib/More_Word_Operations.thy index 9d1798509..99a592a9c 100644 --- a/lib/Word_Lib/More_Word_Operations.thy +++ b/lib/Word_Lib/More_Word_Operations.thy @@ -13,6 +13,12 @@ theory More_Word_Operations Reversed_Bit_Lists More_Misc Signed_Words + Word_Lemmas + Word_EqI +begin + +context + includes bit_operations_syntax begin definition @@ -48,6 +54,36 @@ definition where "word_ctz w \ length (takeWhile Not (rev (to_bl w)))" +lemma word_ctz_unfold: + \word_ctz w = length (takeWhile (Not \ bit w) [0.. for w :: \'a::len word\ + by (simp add: word_ctz_def rev_to_bl_eq takeWhile_map) + +lemma word_ctz_unfold': + \word_ctz w = Min (insert LENGTH('a) {n. bit w n})\ for w :: \'a::len word\ +proof (cases \\n. bit w n\) + case True + then obtain n where \bit w n\ .. + from \bit w n\ show ?thesis + apply (simp add: word_ctz_unfold) + apply (subst Min_eq_length_takeWhile [symmetric]) + apply (auto simp add: bit_imp_le_length) + apply (subst Min_insert) + apply auto + apply (subst min.absorb2) + apply (subst Min_le_iff) + apply auto + apply (meson bit_imp_le_length order_less_le) + done +next + case False + then have \bit w = bot\ + by auto + then have \word_ctz w = LENGTH('a)\ + by (simp add: word_ctz_def rev_to_bl_eq bot_fun_def map_replicate_const) + with \bit w = bot\ show ?thesis + by simp +qed + lemma word_ctz_le: "word_ctz (w :: ('a::len word)) \ LENGTH('a)" apply (clarsimp simp: word_ctz_def) @@ -87,11 +123,11 @@ qed lemma unat_of_nat_ctz_mw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len word) = word_ctz w" - by simp + by (simp add: unsigned_of_nat) lemma unat_of_nat_ctz_smw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len signed word) = word_ctz w" - by simp + by (simp add: unsigned_of_nat) definition word_log2 :: "'a::len word \ nat" @@ -108,7 +144,7 @@ where definition sign_extend :: "nat \ 'a::len word \ 'a word" where - "sign_extend n w \ if w !! n then w OR NOT (mask n) else w AND mask n" + "sign_extend n w \ if bit w n then w OR NOT (mask n) else w AND mask n" lemma sign_extend_eq_signed_take_bit: \sign_extend = signed_take_bit\ @@ -119,7 +155,7 @@ proof (rule ext)+ fix q assume \q < LENGTH('a)\ then show \bit (sign_extend n w) q \ bit (signed_take_bit n w) q\ - by (auto simp add: test_bit_eq_bit bit_signed_take_bit_iff + by (auto simp add: bit_signed_take_bit_iff sign_extend_def bit_and_iff bit_or_iff bit_not_iff bit_mask_iff not_less exp_eq_0_imp_not_bit not_le min_def) qed @@ -128,7 +164,7 @@ qed definition sign_extended :: "nat \ 'a::len word \ bool" where - "sign_extended n w \ \i. n < i \ i < size w \ w !! i = w !! n" + "sign_extended n w \ \i. n < i \ i < size w \ bit w i = bit w n" lemma ptr_add_0 [simp]: "ptr_add ref 0 = ref " @@ -193,7 +229,7 @@ lemma bit_word_log2: \bit w (word_log2 w)\ if \w \ 0\ proof - from \w \ 0\ have \\r. bit w r\ - by (simp add: bit_eq_iff) + by (auto intro: bit_eqI) then obtain r where \bit w r\ .. from \w \ 0\ have \word_log2 w = Max {n. bit w n}\ by (simp add: word_log2_unfold) @@ -216,17 +252,17 @@ proof - qed lemma word_log2_nth_same: - "w \ 0 \ w !! word_log2 w" - by (drule bit_word_log2) (simp add: test_bit_eq_bit) + "w \ 0 \ bit w (word_log2 w)" + by (drule bit_word_log2) simp lemma word_log2_nth_not_set: - "\ word_log2 w < i ; i < size w \ \ \ w !! i" - using word_log2_maximum [of w i] by (auto simp add: test_bit_eq_bit) + "\ word_log2 w < i ; i < size w \ \ \ bit w i" + using word_log2_maximum [of w i] by auto lemma word_log2_highest: - assumes a: "w !! i" + assumes a: "bit w i" shows "i \ word_log2 w" - using a by (simp add: test_bit_eq_bit word_log2_maximum) + using a by (simp add: word_log2_maximum) lemma word_log2_max: "word_log2 w < size w" @@ -319,7 +355,7 @@ next case False have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz - by (metis shiftr_div_2n' word_shiftr_lt) + by (metis le_add_diff_inverse2 less_mult_imp_div_less order_less_imp_le power_add unsigned_less) have"2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans nat_less_le) @@ -413,7 +449,7 @@ next assume asm: "alignUp a n = 0" have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz - by (metis shiftr_div_2n' word_shiftr_lt) + by (metis le_add_diff_inverse2 less_mult_imp_div_less order_less_imp_le power_add unsigned_less) have leq: "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans @@ -438,7 +474,7 @@ next with leq have "2 ^ n * (unat a div 2 ^ n + 1) = 2 ^ LENGTH('a)" by (force elim!: le_SucE) then have "unat a div 2 ^ n = 2 ^ LENGTH('a) div 2 ^ n - 1" - by (metis (no_types, hide_lams) Groups.add_ac(2) add.right_neutral + by (metis (no_types, opaque_lifting) Groups.add_ac(2) add.right_neutral add_diff_cancel_left' div_le_dividend div_mult_self4 gr_implies_not0 le_neq_implies_less power_eq_0_iff zero_neq_numeral) then have "unat a div 2 ^ n = 2 ^ (LENGTH('a) - n) - 1" @@ -536,21 +572,26 @@ qed (* Sign extension from bit n. *) +lemma bin_sign_extend_iff [bit_simps]: + \bit (sign_extend e w) i \ bit w (min e i)\ + if \i < LENGTH('a)\ for w :: \'a::len word\ + using that by (simp add: sign_extend_def bit_simps min_def) + lemma sign_extend_bitwise_if: - "i < size w \ sign_extend e w !! i \ (if i < e then w !! i else w !! e)" - by (simp add: sign_extend_def neg_mask_test_bit word_size) + "i < size w \ bit (sign_extend e w) i \ (if i < e then bit w i else bit w e)" + by (simp add: word_size bit_simps) lemma sign_extend_bitwise_if' [word_eqI_simps]: - \i < LENGTH('a) \ sign_extend e w !! i \ (if i < e then w !! i else w !! e)\ + \i < LENGTH('a) \ bit (sign_extend e w) i \ (if i < e then bit w i else bit w e)\ for w :: \'a::len word\ using sign_extend_bitwise_if [of i w e] by (simp add: word_size) lemma sign_extend_bitwise_disj: - "i < size w \ sign_extend e w !! i \ i \ e \ w !! i \ e \ i \ w !! e" + "i < size w \ bit (sign_extend e w) i \ i \ e \ bit w i \ e \ i \ bit w e" by (auto simp: sign_extend_bitwise_if) lemma sign_extend_bitwise_cases: - "i < size w \ sign_extend e w !! i \ (i \ e \ w !! i) \ (e \ i \ w !! e)" + "i < size w \ bit (sign_extend e w) i \ (i \ e \ bit w i) \ (e \ i \ bit w e)" by (auto simp: sign_extend_bitwise_if) lemmas sign_extend_bitwise_disj' = sign_extend_bitwise_disj[simplified word_size] @@ -559,8 +600,8 @@ lemmas sign_extend_bitwise_cases' = sign_extend_bitwise_cases[simplified word_si (* Often, it is easier to reason about an operation which does not overwrite the bit which determines which mask operation to apply. *) lemma sign_extend_def': - "sign_extend n w = (if w !! n then w OR NOT (mask (Suc n)) else w AND mask (Suc n))" - by (rule bit_word_eqI) (auto simp add: bit_simps sign_extend_eq_signed_take_bit min_def test_bit_eq_bit less_Suc_eq_le) + "sign_extend n w = (if bit w n then w OR NOT (mask (Suc n)) else w AND mask (Suc n))" + by (rule bit_word_eqI) (auto simp add: bit_simps sign_extend_eq_signed_take_bit min_def less_Suc_eq_le) lemma sign_extended_sign_extend: "sign_extended n (sign_extend n w)" @@ -570,10 +611,8 @@ lemma sign_extended_iff_sign_extend: "sign_extended n w \ sign_extend n w = w" apply auto apply (auto simp add: bit_eq_iff) - apply (simp_all add: bit_simps sign_extend_eq_signed_take_bit not_le min_def sign_extended_def test_bit_eq_bit word_size split: if_splits) - using le_imp_less_or_eq apply auto[1] - apply (metis bit_imp_le_length nat_less_le) - apply (metis Suc_leI Suc_n_not_le_n le_trans nat_less_le) + apply (simp_all add: bit_simps sign_extend_eq_signed_take_bit not_le min_def sign_extended_def word_size split: if_splits) + using le_imp_less_or_eq apply auto done lemma sign_extended_weaken: @@ -585,7 +624,7 @@ lemma sign_extend_sign_extend_eq: by (rule bit_word_eqI) (simp add: sign_extend_eq_signed_take_bit bit_simps) lemma sign_extended_high_bits: - "\ sign_extended e p; j < size p; e \ i; i < j \ \ p !! i = p !! j" + "\ sign_extended e p; j < size p; e \ i; i < j \ \ bit p i = bit p j" by (drule (1) sign_extended_weaken; simp add: sign_extended_def) lemma sign_extend_eq: @@ -601,11 +640,11 @@ lemma sign_extended_add: proof (cases "e < size p") case True note and_or = is_aligned_add_or[OF p f] - have "\ f !! e" + have "\ bit f e" using True e less_2p_is_upper_bits_unset[THEN iffD1, OF f] by (fastforce simp: word_size) - hence i: "(p + f) !! e = p !! e" - by (simp add: and_or) + hence i: "bit (p + f) e = bit p e" + by (simp add: and_or bit_simps) have fm: "f AND mask e = f" by (fastforce intro: subst[where P="\f. f AND mask e = f", OF less_mask_eq[OF f]] simp: mask_twice e) @@ -622,7 +661,7 @@ qed lemma sign_extended_neq_mask: "\sign_extended n ptr; m \ n\ \ sign_extended n (ptr AND NOT (mask m))" - by (fastforce simp: sign_extended_def word_size neg_mask_test_bit) + by (fastforce simp: sign_extended_def word_size neg_mask_test_bit bit_simps) definition "limited_and (x :: 'a :: len word) y \ (x AND y = x)" @@ -642,13 +681,11 @@ lemma limited_and_eq_id: lemma lshift_limited_and: "limited_and x z \ limited_and (x << n) (z << n)" - unfolding limited_and_def - by (simp add: shiftl_over_and_dist[symmetric]) + using push_bit_and [of n x z] by (simp add: limited_and_def shiftl_def) lemma rshift_limited_and: "limited_and x z \ limited_and (x >> n) (z >> n)" - unfolding limited_and_def - by (simp add: shiftr_over_and_dist[symmetric]) + using drop_bit_and [of n x z] by (simp add: limited_and_def shiftr_def) lemmas limited_and_simps1 = limited_and_eq_0 limited_and_eq_id @@ -660,8 +697,7 @@ lemmas limited_and_simps = limited_and_simps1 limited_and_simps1[OF lshift_limited_and] limited_and_simps1[OF rshift_limited_and] limited_and_simps1[OF rshift_limited_and, OF is_aligned_limited_and] - not_one shiftl_shiftr1[unfolded word_size mask_eq_decr_exp] - shiftl_shiftr2[unfolded word_size mask_eq_decr_exp] + not_one_eq definition from_bool :: "bool \ 'a::len word" where @@ -685,8 +721,8 @@ definition "to_bool \ (\) 0" lemma to_bool_and_1: - "to_bool (x AND 1) = (x !! 0)" - by (simp add: test_bit_word_eq to_bool_def and_one_eq mod_2_eq_odd) + "to_bool (x AND 1) \ bit x 0" + by (simp add: to_bool_def and_one_eq mod_2_eq_odd) lemma to_bool_from_bool [simp]: "to_bool (from_bool r) = r" @@ -768,7 +804,7 @@ lemma aligned_offset_in_range: apply (simp add: bit_simps) apply (erule is_alignedE') apply (auto simp add: bit_simps not_le)[1] - apply (metis less_2p_is_upper_bits_unset test_bit_eq_bit) + apply (metis less_2p_is_upper_bits_unset) apply (simp only: is_aligned_add_or word_ao_dist flip: neg_mask_in_mask_range) apply (subgoal_tac \y AND NOT (mask n) = 0\) apply simp @@ -880,14 +916,19 @@ lemma aligned_mask_ranges_disjoint2: apply safe apply (simp only: flip: neg_mask_in_mask_range) apply (drule_tac x="x AND mask n >> m" in spec) - apply (clarsimp simp: and_mask_less_size wsst_TYs shiftr_less_t2n multiple_mask_trivia neg_mask_twice - word_bw_assocs max_absorb2 shiftr_shiftl1) + apply (erule notE[OF mp]) + apply (simp flip: take_bit_eq_mask add: shiftr_def drop_bit_take_bit) + apply transfer + apply simp + apply (simp add: word_size and_mask_less_size) + apply (subst disjunctive_add) + apply (auto simp add: bit_simps word_size intro!: bit_eqI) done lemma word_clz_sint_upper[simp]: "LENGTH('a) \ 3 \ sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a sword) \ int (LENGTH('a))" using word_clz_max [of w] - apply (simp add: word_size) + apply (simp add: word_size signed_of_nat) apply (subst signed_take_bit_int_eq_self) apply simp_all apply (metis negative_zle of_nat_numeral semiring_1_class.of_nat_power) @@ -901,9 +942,9 @@ lemma word_clz_sint_lower[simp]: \ - sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a signed word) \ int (LENGTH('a))" apply (subst sint_eq_uint) using word_clz_max [of w] - apply (simp_all add: word_size) + apply (simp_all add: word_size unsigned_of_nat) apply (rule not_msb_from_less) - apply (simp add: word_less_nat_alt) + apply (simp add: word_less_nat_alt unsigned_of_nat) apply (subst take_bit_nat_eq_self) apply (simp add: le_less_trans) apply (drule small_powers_of_2) @@ -916,56 +957,6 @@ lemma mask_range_subsetD: x' \ mask_range p n" using aligned_mask_step by fastforce -lemma nasty_split_lt: - "\ (x :: 'a:: len word) < 2 ^ (m - n); n \ m; m < LENGTH('a::len) \ - \ x * 2 ^ n + (2 ^ n - 1) \ 2 ^ m - 1" - apply (simp only: add_diff_eq) - apply (subst mult_1[symmetric], subst distrib_right[symmetric]) - apply (rule word_sub_mono) - apply (rule order_trans) - apply (rule word_mult_le_mono1) - apply (rule inc_le) - apply assumption - apply (subst word_neq_0_conv[symmetric]) - apply (rule power_not_zero) - apply simp - apply (subst unat_power_lower, simp)+ - apply (subst power_add[symmetric]) - apply (rule power_strict_increasing) - apply simp - apply simp - apply (subst power_add[symmetric]) - apply simp - apply simp - apply (rule word_sub_1_le) - apply (subst mult.commute) - apply (subst shiftl_t2n[symmetric]) - apply (rule word_shift_nonzero) - apply (erule inc_le) - apply simp - apply (unat_arith) - apply (drule word_power_less_1) - apply simp - done - -lemma nasty_split_less: - "\m \ n; n \ nm; nm < LENGTH('a::len); x < 2 ^ (nm - n)\ - \ (x :: 'a word) * 2 ^ n + (2 ^ m - 1) < 2 ^ nm" - apply (simp only: word_less_sub_le[symmetric]) - apply (rule order_trans [OF _ nasty_split_lt]) - apply (rule word_plus_mono_right) - apply (rule word_sub_mono) - apply (simp add: word_le_nat_alt) - apply simp - apply (simp add: word_sub_1_le[OF power_not_zero]) - apply (simp add: word_sub_1_le[OF power_not_zero]) - apply (rule is_aligned_no_wrap') - apply (rule is_aligned_mult_triv2) - apply simp - apply (erule order_le_less_trans, simp) - apply simp+ - done - lemma add_mult_in_mask_range: "\ is_aligned (base :: 'a :: len word) n; n < LENGTH('a); bits \ n; x < 2 ^ (n - bits) \ \ base + x * 2^bits \ mask_range base n" @@ -977,14 +968,27 @@ lemma from_to_bool_last_bit: by (metis from_bool_to_bool_iff word_and_1) lemma sint_ctz: - "LENGTH('a) > 2 - \ 0 \ sint (of_nat (word_ctz (x :: 'a :: len word)) :: 'a signed word) - \ sint (of_nat (word_ctz x) :: 'a signed word) \ int (LENGTH('a))" - apply (subgoal_tac "LENGTH('a) < 2 ^ (LENGTH('a) - 1)") - apply (rule conjI) - apply (metis len_signed order_le_less_trans sint_of_nat_ge_zero word_ctz_le) - apply (metis int_eq_sint len_signed sint_of_nat_le word_ctz_le) - using small_powers_of_2 [of \LENGTH('a)\] by simp + \0 \ sint (of_nat (word_ctz (x :: 'a :: len word)) :: 'a signed word) + \ sint (of_nat (word_ctz x) :: 'a signed word) \ int (LENGTH('a))\ (is \?P \ ?Q\) + if \LENGTH('a) > 2\ +proof + have *: \word_ctz x < 2 ^ (LENGTH('a) - Suc 0)\ + using word_ctz_le apply (rule le_less_trans) + using that small_powers_of_2 [of \LENGTH('a)\] apply simp + done + have \int (word_ctz x) div 2 ^ (LENGTH('a) - Suc 0) = 0\ + apply (rule div_pos_pos_trivial) + apply (simp_all add: *) + done + then show ?P by (simp add: signed_of_nat bit_iff_odd) + show ?Q + apply (auto simp add: signed_of_nat) + apply (subst signed_take_bit_int_eq_self) + apply (auto simp add: word_ctz_le * minus_le_iff [of _ \int (word_ctz x)\]) + apply (rule order.trans [of _ 0]) + apply simp_all + done +qed lemma unat_of_nat_word_log2: "LENGTH('a) < 2 ^ LENGTH('b) @@ -1006,4 +1010,6 @@ lemma aligned_mask_diff: apply (meson aligned_add_mask_less_eq is_aligned_weaken le_less_trans) done +end + end \ No newline at end of file diff --git a/lib/Word_Lib/Most_significant_bit.thy b/lib/Word_Lib/Most_significant_bit.thy index 18a55d14a..e3f6f9d12 100644 --- a/lib/Word_Lib/Most_significant_bit.thy +++ b/lib/Word_Lib/Most_significant_bit.thy @@ -11,8 +11,8 @@ section \Dedicated operation for the most significant bit\ theory Most_significant_bit imports "HOL-Library.Word" - Bits_Int - Traditional_Infix_Syntax + Bit_Shifts_Infix_Syntax + More_Word More_Arithmetic begin @@ -28,13 +28,14 @@ instance .. end -lemma msb_conv_bin_sign: "msb x \ bin_sign x = -1" -by(simp add: bin_sign_def not_le msb_int_def) - lemma msb_bin_rest [simp]: "msb (x div 2) = msb x" for x :: int by (simp add: msb_int_def) +context + includes bit_operations_syntax +begin + lemma int_msb_and [simp]: "msb ((x :: int) AND y) \ msb x \ msb y" by(simp add: msb_int_def) @@ -47,14 +48,13 @@ by(simp add: msb_int_def) lemma int_msb_not [simp]: "msb (NOT (x :: int)) \ \ msb x" by(simp add: msb_int_def not_less) +end + lemma msb_shiftl [simp]: "msb ((x :: int) << n) \ msb x" -by(simp add: msb_int_def) + by (simp add: msb_int_def shiftl_def) lemma msb_shiftr [simp]: "msb ((x :: int) >> r) \ msb x" -by(simp add: msb_int_def) - -lemma msb_bin_sc [simp]: "msb (bin_sc n b x) \ msb x" -by(simp add: msb_conv_bin_sign) + by (simp add: msb_int_def shiftr_def) lemma msb_0 [simp]: "msb (0 :: int) = False" by(simp add: msb_int_def) @@ -71,24 +71,15 @@ instantiation word :: (len) msb begin definition msb_word :: \'a word \ bool\ - where \msb a \ bin_sign (sbintrunc (LENGTH('a) - 1) (uint a)) = - 1\ - -lemma msb_word_eq: - \msb w \ bit w (LENGTH('a) - 1)\ for w :: \'a::len word\ - by (simp add: msb_word_def bin_sign_lem bit_uint_iff) + where msb_word_iff_bit: \msb w \ bit w (LENGTH('a) - Suc 0)\ for w :: \'a::len word\ instance .. end -lemma msb_word_iff_bit: - \msb w \ bit w (LENGTH('a) - Suc 0)\ - for w :: \'a::len word\ - by (simp add: msb_word_def bin_sign_def bit_uint_iff) - -lemma word_msb_def: - "msb a \ bin_sign (sint a) = - 1" - by (simp add: msb_word_def sint_uint) +lemma msb_word_eq: + \msb w \ bit w (LENGTH('a) - 1)\ for w :: \'a::len word\ + by (simp add: msb_word_iff_bit) lemma word_msb_sint: "msb w \ sint w < 0" by (simp add: msb_word_eq bit_last_iff) @@ -97,48 +88,46 @@ lemma msb_word_iff_sless_0: \msb w \ w by (simp add: word_msb_sint word_sless_alt) -lemma msb_word_of_int: "msb (word_of_int x::'a::len word) = bin_nth x (LENGTH('a) - 1)" - by (simp add: word_msb_def bin_sign_lem) +lemma msb_word_of_int: + "msb (word_of_int x::'a::len word) = bit x (LENGTH('a) - 1)" + by (simp add: msb_word_iff_bit bit_simps) lemma word_msb_numeral [simp]: - "msb (numeral w::'a::len word) = bin_nth (numeral w) (LENGTH('a) - 1)" + "msb (numeral w::'a::len word) = bit (numeral w :: int) (LENGTH('a) - 1)" unfolding word_numeral_alt by (rule msb_word_of_int) lemma word_msb_neg_numeral [simp]: - "msb (- numeral w::'a::len word) = bin_nth (- numeral w) (LENGTH('a) - 1)" + "msb (- numeral w::'a::len word) = bit (- numeral w :: int) (LENGTH('a) - 1)" unfolding word_neg_numeral_alt by (rule msb_word_of_int) lemma word_msb_0 [simp]: "\ msb (0::'a::len word)" - by (simp add: word_msb_def bin_sign_def sint_uint sbintrunc_eq_take_bit) + by (simp add: msb_word_iff_bit) lemma word_msb_1 [simp]: "msb (1::'a::len word) \ LENGTH('a) = 1" - unfolding word_1_wi msb_word_of_int eq_iff [where 'a=nat] - by (simp add: Suc_le_eq) + by (simp add: msb_word_iff_bit le_Suc_eq) -lemma word_msb_nth: "msb w = bin_nth (uint w) (LENGTH('a) - 1)" +lemma word_msb_nth: "msb w = bit (uint w) (LENGTH('a) - 1)" for w :: "'a::len word" - by (simp add: word_msb_def sint_uint bin_sign_lem) + by (simp add: msb_word_iff_bit bit_simps) -lemma msb_nth: "msb w = w !! (LENGTH('a) - 1)" +lemma msb_nth: "msb w = bit w (LENGTH('a) - 1)" for w :: "'a::len word" - by (simp add: word_msb_nth word_test_bit_def) + by (fact msb_word_eq) lemma word_msb_n1 [simp]: "msb (-1::'a::len word)" by (simp add: msb_word_eq not_le) -lemma msb_shift: "msb w \ w >> (LENGTH('a) - 1) \ 0" +lemma msb_shift: "msb w \ w >> LENGTH('a) - 1 \ 0" for w :: "'a::len word" - by (simp add: msb_word_eq shiftr_word_eq bit_iff_odd_drop_bit drop_bit_eq_zero_iff_not_bit_last) + by (simp add: drop_bit_eq_zero_iff_not_bit_last msb_word_eq shiftr_def) lemmas word_ops_msb = msb1 [unfolded msb_nth [symmetric, unfolded One_nat_def]] lemma word_sint_msb_eq: "sint x = uint x - (if msb x then 2 ^ size x else 0)" apply (cases \LENGTH('a)\) - apply (simp_all add: msb_word_def bin_sign_def bit_simps word_size) + apply (simp_all add: msb_word_iff_bit word_size) apply transfer - apply (auto simp add: take_bit_Suc_from_most signed_take_bit_eq_if_positive signed_take_bit_eq_if_negative minus_exp_eq_not_mask ac_simps) - apply (subst disjunctive_add) - apply (simp_all add: bit_simps) + apply (simp add: signed_take_bit_eq_take_bit_minus) done lemma word_sle_msb_le: "x <=s y \ (msb y \ msb x) \ ((msb x \ \ msb y) \ x \ y)" @@ -159,15 +148,15 @@ lemma not_msb_from_less: apply (clarsimp simp add: msb_nth) apply (drule less_mask_eq) apply (drule word_eqD, drule(1) iffD2) - apply simp + apply (simp add: bit_simps) done lemma sint_eq_uint: "\ msb x \ sint x = uint x" - apply (simp add: msb_word_eq) + apply (cases \LENGTH('a)\) + apply (simp_all add: msb_word_iff_bit) apply transfer - apply auto - apply (smt One_nat_def bintrunc_bintrunc_l bintrunc_sbintrunc' diff_le_self len_gt_0 signed_take_bit_eq_if_positive) + apply (simp add: signed_take_bit_eq_take_bit_minus) done lemma scast_eq_ucast: @@ -177,7 +166,7 @@ lemma scast_eq_ucast: apply (rule bit_word_eqI) apply (auto simp add: bit_signed_iff bit_unsigned_iff min_def msb_word_eq) apply (erule notE) - apply (metis le_less_Suc_eq test_bit_bin test_bit_word_eq) + apply (metis le_less_Suc_eq test_bit_bin) done lemma msb_ucast_eq: @@ -186,18 +175,20 @@ lemma msb_ucast_eq: by (simp add: msb_word_eq bit_simps) lemma msb_big: - "msb (a :: ('a::len) word) = (a \ 2 ^ (LENGTH('a) - Suc 0))" - apply (rule iffI) - apply (clarsimp simp: msb_nth) - apply (drule bang_is_le) - apply simp + \msb a \ 2 ^ (LENGTH('a) - Suc 0) \ a\ + for a :: \'a::len word\ + using bang_is_le [of a \LENGTH('a) - Suc 0\] + apply (auto simp add: msb_nth word_le_not_less) apply (rule ccontr) - apply (subgoal_tac "a = a AND mask (LENGTH('a) - Suc 0)") + apply (erule notE) + apply (rule ccontr) + apply (clarsimp simp: not_less) + apply (subgoal_tac "a = take_bit (LENGTH('a) - Suc 0) a") apply (cut_tac and_mask_less' [where w=a and n="LENGTH('a) - Suc 0"]) - apply (clarsimp simp: word_not_le [symmetric]) - apply clarsimp - apply (rule sym, subst and_mask_eq_iff_shiftr_0) - apply (clarsimp simp: msb_shift) + apply auto + apply (simp flip: take_bit_eq_mask) + apply (rule sym) + apply (simp add: take_bit_eq_self_iff_drop_bit_eq_0 drop_bit_eq_zero_iff_not_bit_last) done end diff --git a/lib/Word_Lib/Next_and_Prev.thy b/lib/Word_Lib/Next_and_Prev.thy index cadd4e4dd..1de65546b 100644 --- a/lib/Word_Lib/Next_and_Prev.thy +++ b/lib/Word_Lib/Next_and_Prev.thy @@ -22,7 +22,7 @@ lift_definition word_prev :: \'a::len word \ 'a word\ lemma word_next_unfold: \word_next w = (if w = - 1 then - 1 else w + 1)\ - by transfer (simp add: take_bit_minus_one_eq_mask flip: take_bit_eq_mask_iff_exp_dvd) + by transfer (simp flip: take_bit_eq_mask_iff_exp_dvd) lemma word_prev_unfold: \word_prev w = (if w = 0 then 0 else w - 1)\ @@ -32,7 +32,7 @@ lemma [code]: \Word.the_int (word_next w :: 'a::len word) = (if w = - 1 then Word.the_int w else Word.the_int w + 1)\ by transfer - (simp add: take_bit_minus_one_eq_mask mask_eq_exp_minus_1 take_bit_incr_eq flip: take_bit_eq_mask_iff_exp_dvd) + (simp add: mask_eq_exp_minus_1 take_bit_incr_eq flip: take_bit_eq_mask_iff_exp_dvd) lemma [code]: \Word.the_int (word_prev w :: 'a::len word) = diff --git a/lib/Word_Lib/Norm_Words.thy b/lib/Word_Lib/Norm_Words.thy index c574add6b..eeae2ba11 100644 --- a/lib/Word_Lib/Norm_Words.thy +++ b/lib/Word_Lib/Norm_Words.thy @@ -7,7 +7,7 @@ section "Normalising Word Numerals" theory Norm_Words - imports Bits_Int Signed_Words + imports Signed_Words begin text \ @@ -15,25 +15,28 @@ text \ interval \[0..2^len_of 'a)\. Only for concrete word lengths. \ +lemma bintrunc_numeral: + "(take_bit :: nat \ int \ int) (numeral k) x = of_bool (odd x) + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (x div 2)" + by (simp add: numeral_eq_Suc take_bit_Suc mod_2_eq_odd) + lemma neg_num_bintr: - "(- numeral x :: 'a::len word) = - word_of_int (bintrunc (LENGTH('a)) (-numeral x))" + "(- numeral x :: 'a::len word) = word_of_int (take_bit LENGTH('a) (- numeral x))" by transfer simp ML \ - fun is_refl (Const (@{const_name Pure.eq}, _) $ x $ y) = (x = y) + fun is_refl \<^Const_>\Pure.eq _ for x y\ = (x = y) | is_refl _ = false; - fun signed_dest_wordT (Type (@{type_name word}, [Type (@{type_name signed}, [T])])) = Word_Lib.dest_binT T + fun signed_dest_wordT \<^Type>\word \<^Type>\signed T\\ = Word_Lib.dest_binT T | signed_dest_wordT T = Word_Lib.dest_wordT T fun typ_size_of t = signed_dest_wordT (type_of (Thm.term_of t)); - fun num_len (Const (@{const_name Num.Bit0}, _) $ n) = num_len n + 1 - | num_len (Const (@{const_name Num.Bit1}, _) $ n) = num_len n + 1 - | num_len (Const (@{const_name Num.One}, _)) = 1 - | num_len (Const (@{const_name numeral}, _) $ t) = num_len t - | num_len (Const (@{const_name uminus}, _) $ t) = num_len t + fun num_len \<^Const_>\Num.Bit0 for n\ = num_len n + 1 + | num_len \<^Const_>\Num.Bit1 for n\ = num_len n + 1 + | num_len \<^Const_>\Num.One\ = 1 + | num_len \<^Const_>\numeral _ for t\ = num_len t + | num_len \<^Const_>\uminus _ for t\ = num_len t | num_len t = raise TERM ("num_len", [t]) fun unsigned_norm is_neg _ ctxt ct = @@ -43,7 +46,8 @@ ML \ val th = [Thm.reflexive ct, mk_eq btr] MRS transitive_thm (* will work in context of theory Word as well *) - val ss = simpset_of (@{context} addsimps @{thms bintrunc_numeral}) + val ss = simpset_of (@{context} addsimps @{thms bintrunc_numeral} delsimps @{thms take_bit_minus_one_eq_mask}) + (* TODO: completely explicitly determined simpset *) val cnv = simplify (put_simpset ss ctxt) th in if is_refl (Thm.prop_of cnv) then NONE else SOME cnv end else NONE) @@ -67,22 +71,24 @@ lemma minus_one_norm: lemmas minus_one_norm_num = minus_one_norm [where 'a="'b::len bit0"] minus_one_norm [where 'a="'b::len0 bit1"] -lemma "f (7 :: 2 word) = f 3" by simp +context +begin -lemma "f 7 = f (3 :: 2 word)" by simp +private lemma "f (7 :: 2 word) = f 3" by simp -lemma "f (-2) = f (21 + 1 :: 3 word)" by simp +private lemma "f 7 = f (3 :: 2 word)" by simp -lemma "f (-2) = f (13 + 1 :: 'a::len word)" +private lemma "f (-2) = f (21 + 1 :: 3 word)" by simp + +private lemma "f (-2) = f (13 + 1 :: 'a::len word)" apply simp (* does not touch generic word length *) oops -lemma "f (-2) = f (0xFFFFFFFE :: 32 word)" by simp +private lemma "f (-2) = f (0xFFFFFFFE :: 32 word)" by simp -lemma "(-1 :: 2 word) = 3" by simp - -lemma "f (-2) = f (0xFFFFFFFE :: 32 signed word)" by simp +private lemma "(-1 :: 2 word) = 3" by simp +private lemma "f (-2) = f (0xFFFFFFFE :: 32 signed word)" by simp text \ We leave @{term "-1"} untouched by default, because it is often useful @@ -95,16 +101,18 @@ context notes minus_one_norm_num [simp] begin -lemma "f (-1) = f (15 :: 4 word)" by simp +private lemma "f (-1) = f (15 :: 4 word)" by simp -lemma "f (-1) = f (7 :: 3 word)" by simp +private lemma "f (-1) = f (7 :: 3 word)" by simp -lemma "f (-1) = f (0xFFFF :: 16 word)" by simp +private lemma "f (-1) = f (0xFFFF :: 16 word)" by simp -lemma "f (-1) = f (0xFFFF + 1 :: 'a::len word)" +private lemma "f (-1) = f (0xFFFF + 1 :: 'a::len word)" apply simp (* does not touch generic -1 *) oops end end + +end diff --git a/lib/Word_Lib/ROOT b/lib/Word_Lib/ROOT index 91272ddae..0c3c2d429 100644 --- a/lib/Word_Lib/ROOT +++ b/lib/Word_Lib/ROOT @@ -9,14 +9,23 @@ chapter Lib session Word_Lib (lib) = HOL + options [timeout = 300, document=pdf] sessions + "HOL-Library" "HOL-Eisbach" directories (* not in the AFP: *) "$L4V_ARCH" theories (* not in the AFP: *) "Distinct_Prop" "$L4V_ARCH/WordSetup" + theories [document=false] + More_Arithmetic + Even_More_List + More_Sublist + More_Misc + Strict_part_mono + Many_More + Ancient_Numeral + Examples theories Guide - Examples document_files "root.tex" diff --git a/lib/Word_Lib/Reversed_Bit_Lists.thy b/lib/Word_Lib/Reversed_Bit_Lists.thy index 20447c1c9..2d4484cb2 100644 --- a/lib/Word_Lib/Reversed_Bit_Lists.thy +++ b/lib/Word_Lib/Reversed_Bit_Lists.thy @@ -17,6 +17,12 @@ theory Reversed_Bit_Lists Even_More_List "HOL-Library.Sublist" Aligned + Singleton_Bit_Shifts + Legacy_Aliases +begin + +context + includes bit_operations_syntax begin lemma horner_sum_of_bool_2_concat: @@ -336,7 +342,7 @@ definition bl_to_bin :: "bool list \ int" primrec bin_to_bl_aux :: "nat \ int \ bool list \ bool list" where Z: "bin_to_bl_aux 0 w bl = bl" - | Suc: "bin_to_bl_aux (Suc n) w bl = bin_to_bl_aux n (bin_rest w) ((bin_last w) # bl)" + | Suc: "bin_to_bl_aux (Suc n) w bl = bin_to_bl_aux n (w div 2) (odd w # bl)" definition bin_to_bl :: "nat \ int \ bool list" where "bin_to_bl n w = bin_to_bl_aux n w []" @@ -429,20 +435,20 @@ lemma bin_to_bl_minus1: "bin_to_bl n (- 1) = replicate n True" subsection \Semantic interpretation of \<^typ>\bool list\ as \<^typ>\int\\ -lemma bin_bl_bin': "bl_to_bin (bin_to_bl_aux n w bs) = bl_to_bin_aux bs (bintrunc n w)" +lemma bin_bl_bin': "bl_to_bin (bin_to_bl_aux n w bs) = bl_to_bin_aux bs (take_bit n w)" by (induct n arbitrary: w bs) (auto simp: bl_to_bin_def take_bit_Suc ac_simps mod_2_eq_odd) -lemma bin_bl_bin [simp]: "bl_to_bin (bin_to_bl n w) = bintrunc n w" +lemma bin_bl_bin [simp]: "bl_to_bin (bin_to_bl n w) = take_bit n w" by (auto simp: bin_to_bl_def bin_bl_bin') lemma bl_to_bin_rep_F: "bl_to_bin (replicate n False @ bl) = bl_to_bin bl" by (simp add: bin_to_bl_zero_aux [symmetric] bin_bl_bin') (simp add: bl_to_bin_def) -lemma bin_to_bl_trunc [simp]: "n \ m \ bin_to_bl n (bintrunc m w) = bin_to_bl n w" +lemma bin_to_bl_trunc [simp]: "n \ m \ bin_to_bl n (take_bit m w) = bin_to_bl n w" by (auto intro: bl_to_bin_inj) lemma bin_to_bl_aux_bintr: - "bin_to_bl_aux n (bintrunc m bin) bl = + "bin_to_bl_aux n (take_bit m bin) bl = replicate (n - m) False @ bin_to_bl_aux (min n m) bin bl" apply (induct n arbitrary: m bin bl) apply clarsimp @@ -455,7 +461,7 @@ lemma bin_to_bl_aux_bintr: done lemma bin_to_bl_bintr: - "bin_to_bl n (bintrunc m bin) = replicate (n - m) False @ bin_to_bl (min n m) bin" + "bin_to_bl n (take_bit m bin) = replicate (n - m) False @ bin_to_bl (min n m) bin" unfolding bin_to_bl_def by (rule bin_to_bl_aux_bintr) lemma bl_to_bin_rep_False: "bl_to_bin (replicate n False) = 0" @@ -473,25 +479,25 @@ lemma sign_bl_bin': "bin_sign (bl_to_bin_aux bs w) = bin_sign w" lemma sign_bl_bin: "bin_sign (bl_to_bin bs) = 0" by (simp add: bl_to_bin_def sign_bl_bin') -lemma bl_sbin_sign_aux: "hd (bin_to_bl_aux (Suc n) w bs) = (bin_sign (sbintrunc n w) = -1)" +lemma bl_sbin_sign_aux: "hd (bin_to_bl_aux (Suc n) w bs) = (bin_sign (signed_take_bit n w) = -1)" by (induction n arbitrary: w bs) (auto simp add: bin_sign_def even_iff_mod_2_eq_zero bit_Suc) -lemma bl_sbin_sign: "hd (bin_to_bl (Suc n) w) = (bin_sign (sbintrunc n w) = -1)" +lemma bl_sbin_sign: "hd (bin_to_bl (Suc n) w) = (bin_sign (signed_take_bit n w) = -1)" unfolding bin_to_bl_def by (rule bl_sbin_sign_aux) lemma bin_nth_of_bl_aux: - "bin_nth (bl_to_bin_aux bl w) n = - (n < size bl \ rev bl ! n \ n \ length bl \ bin_nth w (n - size bl))" + "bit (bl_to_bin_aux bl w) n = + (n < size bl \ rev bl ! n \ n \ length bl \ bit w (n - size bl))" apply (induction bl arbitrary: w) apply simp_all apply safe apply (simp_all add: not_le nth_append bit_double_iff even_bit_succ_iff split: if_splits) done -lemma bin_nth_of_bl: "bin_nth (bl_to_bin bl) n = (n < length bl \ rev bl ! n)" +lemma bin_nth_of_bl: "bit (bl_to_bin bl) n = (n < length bl \ rev bl ! n)" by (simp add: bl_to_bin_def bin_nth_of_bl_aux) -lemma bin_nth_bl: "n < m \ bin_nth w n = nth (rev (bin_to_bl m w)) n" +lemma bin_nth_bl: "n < m \ bit w n = nth (rev (bin_to_bl m w)) n" apply (induct n arbitrary: m w) apply clarsimp apply (case_tac m, clarsimp) @@ -514,7 +520,7 @@ lemma nth_bin_to_bl_aux: less_antisym less_imp_Suc_add list.size(3) nat_less_le nth_append size_bin_to_bl_aux) done -lemma nth_bin_to_bl: "n < m \ (bin_to_bl m w) ! n = bin_nth w (m - Suc n)" +lemma nth_bin_to_bl: "n < m \ (bin_to_bl m w) ! n = bit w (m - Suc n)" by (simp add: bin_to_bl_def nth_bin_to_bl_aux) lemma takefill_bintrunc: "takefill False n bl = rev (bin_to_bl n (bl_to_bin (rev bl)))" @@ -579,26 +585,26 @@ lemma bl_to_bin_ge0: "bl_to_bin bs \ 0" apply simp done -lemma butlast_rest_bin: "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (bin_rest w)" +lemma butlast_rest_bin: "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (w div 2)" apply (unfold bin_to_bl_def) apply (cases n, clarsimp) apply clarsimp apply (auto simp add: bin_to_bl_aux_alt) done -lemma butlast_bin_rest: "butlast bl = bin_to_bl (length bl - Suc 0) (bin_rest (bl_to_bin bl))" +lemma butlast_bin_rest: "butlast bl = bin_to_bl (length bl - Suc 0) (bl_to_bin bl div 2)" using butlast_rest_bin [where w="bl_to_bin bl" and n="length bl"] by simp lemma butlast_rest_bl2bin_aux: - "bl \ [] \ bl_to_bin_aux (butlast bl) w = bin_rest (bl_to_bin_aux bl w)" + "bl \ [] \ bl_to_bin_aux (butlast bl) w = bl_to_bin_aux bl w div 2" by (induct bl arbitrary: w) auto -lemma butlast_rest_bl2bin: "bl_to_bin (butlast bl) = bin_rest (bl_to_bin bl)" +lemma butlast_rest_bl2bin: "bl_to_bin (butlast bl) = bl_to_bin bl div 2" by (cases bl) (auto simp: bl_to_bin_def butlast_rest_bl2bin_aux) lemma trunc_bl2bin_aux: - "bintrunc m (bl_to_bin_aux bl w) = - bl_to_bin_aux (drop (length bl - m) bl) (bintrunc (m - length bl) w)" + "take_bit m (bl_to_bin_aux bl w) = + bl_to_bin_aux (drop (length bl - m) bl) (take_bit (m - length bl) w)" proof (induct bl arbitrary: w) case Nil show ?case by simp @@ -616,13 +622,13 @@ next qed qed -lemma trunc_bl2bin: "bintrunc m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)" +lemma trunc_bl2bin: "take_bit m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)" by (simp add: bl_to_bin_def trunc_bl2bin_aux) -lemma trunc_bl2bin_len [simp]: "bintrunc (length bl) (bl_to_bin bl) = bl_to_bin bl" +lemma trunc_bl2bin_len [simp]: "take_bit (length bl) (bl_to_bin bl) = bl_to_bin bl" by (simp add: trunc_bl2bin) -lemma bl2bin_drop: "bl_to_bin (drop k bl) = bintrunc (length bl - k) (bl_to_bin bl)" +lemma bl2bin_drop: "bl_to_bin (drop k bl) = take_bit (length bl - k) (bl_to_bin bl)" apply (rule trans) prefer 2 apply (rule trunc_bl2bin [symmetric]) @@ -630,19 +636,19 @@ lemma bl2bin_drop: "bl_to_bin (drop k bl) = bintrunc (length bl - k) (bl_to_bin apply auto done -lemma take_rest_power_bin: "m \ n \ take m (bin_to_bl n w) = bin_to_bl m ((bin_rest ^^ (n - m)) w)" +lemma take_rest_power_bin: "m \ n \ take m (bin_to_bl n w) = bin_to_bl m (((\w. w div 2) ^^ (n - m)) w)" apply (rule nth_equalityI) apply simp apply (clarsimp simp add: nth_bin_to_bl nth_rest_power_bin) done -lemma last_bin_last': "size xs > 0 \ last xs \ bin_last (bl_to_bin_aux xs w)" +lemma last_bin_last': "size xs > 0 \ last xs \ odd (bl_to_bin_aux xs w)" by (induct xs arbitrary: w) auto -lemma last_bin_last: "size xs > 0 \ last xs \ bin_last (bl_to_bin xs)" +lemma last_bin_last: "size xs > 0 \ last xs \ odd (bl_to_bin xs)" unfolding bl_to_bin_def by (erule last_bin_last') -lemma bin_last_last: "bin_last w \ last (bin_to_bl (Suc n) w)" +lemma bin_last_last: "odd w \ last (bin_to_bl (Suc n) w)" by (simp add: bin_to_bl_def) (auto simp: bin_to_bl_aux_alt) lemma drop_bin2bl_aux: @@ -694,22 +700,22 @@ lemma bl_bin_bl_rep_drop: by (simp add: bl_to_bin_inj bl_to_bin_rep_F trunc_bl2bin) lemma bl_to_bin_aux_cat: - "bl_to_bin_aux bs (bin_cat w nv v) = - bin_cat w (nv + length bs) (bl_to_bin_aux bs v)" + "bl_to_bin_aux bs (concat_bit nv v w) = + concat_bit (nv + length bs) (bl_to_bin_aux bs v) w" by (rule bit_eqI) (auto simp add: bin_nth_of_bl_aux bin_nth_cat algebra_simps) lemma bin_to_bl_aux_cat: - "bin_to_bl_aux (nv + nw) (bin_cat v nw w) bs = + "bin_to_bl_aux (nv + nw) (concat_bit nw w v) bs = bin_to_bl_aux nv v (bin_to_bl_aux nw w bs)" by (induction nw arbitrary: w bs) (simp_all add: concat_bit_Suc) -lemma bl_to_bin_aux_alt: "bl_to_bin_aux bs w = bin_cat w (length bs) (bl_to_bin bs)" +lemma bl_to_bin_aux_alt: "bl_to_bin_aux bs w = concat_bit (length bs) (bl_to_bin bs) w" using bl_to_bin_aux_cat [where nv = "0" and v = "0"] by (simp add: bl_to_bin_def [symmetric]) lemma bin_to_bl_cat: - "bin_to_bl (nv + nw) (bin_cat v nw w) = + "bin_to_bl (nv + nw) (concat_bit nw w v) = bin_to_bl_aux nv v (bin_to_bl nw w)" by (simp add: bin_to_bl_def bin_to_bl_aux_cat) @@ -720,15 +726,15 @@ lemmas bin_to_bl_aux_cat_app = trans [OF bin_to_bl_aux_cat bin_to_bl_aux_alt] lemma bl_to_bin_app_cat: - "bl_to_bin (bsa @ bs) = bin_cat (bl_to_bin bsa) (length bs) (bl_to_bin bs)" + "bl_to_bin (bsa @ bs) = concat_bit (length bs) (bl_to_bin bs) (bl_to_bin bsa)" by (simp only: bl_to_bin_aux_app_cat bl_to_bin_def) lemma bin_to_bl_cat_app: - "bin_to_bl (n + nw) (bin_cat w nw wa) = bin_to_bl n w @ bin_to_bl nw wa" + "bin_to_bl (n + nw) (concat_bit nw wa w) = bin_to_bl n w @ bin_to_bl nw wa" by (simp only: bin_to_bl_def bin_to_bl_aux_cat_app) text \\bl_to_bin_app_cat_alt\ and \bl_to_bin_app_cat\ are easily interderivable.\ -lemma bl_to_bin_app_cat_alt: "bin_cat (bl_to_bin cs) n w = bl_to_bin (cs @ bin_to_bl n w)" +lemma bl_to_bin_app_cat_alt: "concat_bit n w (bl_to_bin cs) = bl_to_bin (cs @ bin_to_bl n w)" by (simp add: bl_to_bin_app_cat) lemma mask_lem: "(bl_to_bin (True # replicate n False)) = bl_to_bin (replicate n True) + 1" @@ -850,7 +856,7 @@ next by simp ultimately show ?case using Suc [of \bin div 2\] - by simp (simp add: bin_to_bl_aux_alt) + by simp (auto simp add: bin_to_bl_aux_alt) qed lemma rbl_succ: "rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin + 1))" @@ -858,8 +864,7 @@ lemma rbl_succ: "rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin + 1))" apply (induction n arbitrary: bin) apply simp_all apply (case_tac bin rule: bin_exhaust) - apply simp - apply (simp add: bin_to_bl_aux_alt ac_simps) + apply (simp_all add: bin_to_bl_aux_alt ac_simps) done lemma rbl_add: @@ -918,24 +923,23 @@ lemma rbl_mult: apply clarsimp apply (case_tac bina rule: bin_exhaust) apply (case_tac binb rule: bin_exhaust) - apply simp - apply (simp add: bin_to_bl_aux_alt) - apply (simp add: rbbl_Cons rbl_mult_Suc rbl_add algebra_simps) + apply (simp_all add: bin_to_bl_aux_alt) + apply (simp_all add: rbbl_Cons rbl_mult_Suc rbl_add algebra_simps) done lemma sclem: "size (concat (map (bin_to_bl n) xs)) = length xs * n" - by (induct xs) auto + by (simp add: length_concat comp_def sum_list_triv) lemma bin_cat_foldl_lem: - "foldl (\u. bin_cat u n) x xs = - bin_cat x (size xs * n) (foldl (\u. bin_cat u n) y xs)" + "foldl (\u k. concat_bit n k u) x xs = + concat_bit (size xs * n) (foldl (\u k. concat_bit n k u) y xs) x" apply (induct xs arbitrary: x) apply simp apply (simp (no_asm)) apply (frule asm_rl) apply (drule meta_spec) apply (erule trans) - apply (drule_tac x = "bin_cat y n a" in meta_spec) + apply (drule_tac x = "concat_bit n a y" in meta_spec) apply (simp add: bin_cat_assoc_sym) done @@ -948,10 +952,10 @@ lemma bin_rcat_bl: "bin_rcat n wl = bl_to_bin (concat (map (bin_to_bl n) wl))" apply (simp add: bin_cat_foldl_lem [symmetric]) done -lemma bin_last_bl_to_bin: "bin_last (bl_to_bin bs) \ bs \ [] \ last bs" +lemma bin_last_bl_to_bin: "odd (bl_to_bin bs) \ bs \ [] \ last bs" by(cases "bs = []")(auto simp add: bl_to_bin_def last_bin_last'[where w=0]) -lemma bin_rest_bl_to_bin: "bin_rest (bl_to_bin bs) = bl_to_bin (butlast bs)" +lemma bin_rest_bl_to_bin: "bl_to_bin bs div 2 = bl_to_bin (butlast bs)" by(cases "bs = []")(simp_all add: bl_to_bin_def butlast_rest_bl2bin_aux) lemma bl_xor_aux_bin: @@ -1038,12 +1042,9 @@ lemma of_bl_eq: lemma bshiftr1_eq: \bshiftr1 b w = of_bl (b # butlast (to_bl w))\ - apply transfer - apply auto - apply (subst bl_to_bin_app_cat [of \[True]\, simplified]) - apply simp - apply (metis One_nat_def add.commute bin_bl_bin bin_last_bl_to_bin bin_rest_bl_to_bin butlast_bin_rest concat_bit_eq last.simps list.distinct(1) list.size(3) list.size(4) odd_iff_mod_2_eq_one plus_1_eq_Suc power_Suc0_right push_bit_of_1 size_bin_to_bl take_bit_eq_mod trunc_bl2bin_len) - apply (simp add: butlast_rest_bl2bin) + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps to_bl_eq_rev nth_append rev_nth nth_butlast not_less simp flip: bit_Suc) + apply (metis Suc_pred len_gt_0 less_eq_decr_length_iff not_bit_length verit_la_disequality) done lemma length_to_bl_eq: @@ -1088,7 +1089,7 @@ lemma td_bl: apply (auto dest: sym) done -interpretation word_bl: +global_interpretation word_bl: type_definition "to_bl :: 'a::len word \ bool list" of_bl @@ -1115,12 +1116,13 @@ lemma length_bl_neq_0 [iff]: "length (to_bl x) \ 0" for x :: "'a::len word" by (fact length_bl_gt_0 [THEN gr_implies_not0]) +lemma hd_to_bl_iff: + \hd (to_bl w) \ bit w (LENGTH('a) - 1)\ + for w :: \'a::len word\ + by (simp add: to_bl_eq_rev hd_map hd_rev) + lemma hd_bl_sign_sint: "hd (to_bl w) = (bin_sign (sint w) = -1)" - apply transfer - apply (auto simp add: bin_sign_def) - using bin_sign_lem bl_sbin_sign apply fastforce - using bin_sign_lem bl_sbin_sign apply force - done + by (simp add: hd_to_bl_iff bit_last_iff bin_sign_def) lemma of_bl_drop': "lend = length bl - LENGTH('a::len) \ @@ -1128,7 +1130,7 @@ lemma of_bl_drop': by transfer (simp flip: trunc_bl2bin) lemma test_bit_of_bl: - "(of_bl bl::'a::len word) !! n = (rev bl ! n \ n < LENGTH('a) \ n < length bl)" + "bit (of_bl bl::'a::len word) n = (rev bl ! n \ n < LENGTH('a) \ n < length bl)" by transfer (simp add: bin_nth_of_bl ac_simps) lemma no_of_bl: "(numeral bin ::'a::len word) = of_bl (bin_to_bl (LENGTH('a)) (numeral bin))" @@ -1302,7 +1304,7 @@ lemma bl_word_or: "to_bl (v OR w) = map2 (\) (to_bl v) (to_bl w)" lemma bl_word_and: "to_bl (v AND w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_and_bin) -lemma bin_nth_uint': "bin_nth (uint w) n \ rev (bin_to_bl (size w) (uint w)) ! n \ n < size w" +lemma bin_nth_uint': "bit (uint w) n \ rev (bin_to_bl (size w) (uint w)) ! n \ n < size w" apply (unfold word_size) apply (safe elim!: bin_nth_uint_imp) apply (frule bin_nth_uint_imp) @@ -1311,10 +1313,10 @@ lemma bin_nth_uint': "bin_nth (uint w) n \ rev (bin_to_bl (s lemmas bin_nth_uint = bin_nth_uint' [unfolded word_size] -lemma test_bit_bl: "w !! n \ rev (to_bl w) ! n \ n < size w" +lemma test_bit_bl: "bit w n \ rev (to_bl w) ! n \ n < size w" by transfer (auto simp add: bin_nth_bl) -lemma to_bl_nth: "n < size w \ to_bl w ! n = w !! (size w - Suc n)" +lemma to_bl_nth: "n < size w \ to_bl w ! n = bit w (size w - Suc n)" by (simp add: word_size rev_nth test_bit_bl) lemma map_bit_interval_eq: @@ -1328,7 +1330,7 @@ proof (rule nth_equalityI) then have \m < n\ by simp then have \bit w m \ takefill False n (rev (to_bl w)) ! m\ - by (auto simp add: nth_takefill not_less rev_nth to_bl_nth word_size test_bit_word_eq dest: bit_imp_le_length) + by (auto simp add: nth_takefill not_less rev_nth to_bl_nth word_size dest: bit_imp_le_length) with \m < n \show \map (bit w) [0.. takefill False n (rev (to_bl w)) ! m\ by simp qed @@ -1395,28 +1397,37 @@ lemma rbl_word_not: "rev (to_bl (NOT x)) = map Not (rev (to_bl x))" lemma bshiftr1_numeral [simp]: \bshiftr1 b (numeral w :: 'a word) = of_bl (b # butlast (bin_to_bl LENGTH('a::len) (numeral w)))\ - by (simp add: bshiftr1_eq) + by (rule bit_word_eqI) (auto simp add: bit_simps rev_nth nth_append nth_butlast nth_bin_to_bl simp flip: bit_Suc) lemma bshiftr1_bl: "to_bl (bshiftr1 b w) = b # butlast (to_bl w)" unfolding bshiftr1_eq by (rule word_bl.Abs_inverse) simp lemma shiftl1_of_bl: "shiftl1 (of_bl bl) = of_bl (bl @ [False])" - by transfer (simp add: bl_to_bin_append) + apply (rule bit_word_eqI) + apply (simp add: bit_simps) + subgoal for n + apply (cases n) + apply simp_all + done + done lemma shiftl1_bl: "shiftl1 w = of_bl (to_bl w @ [False])" - for w :: "'a::len word" -proof - - have "shiftl1 w = shiftl1 (of_bl (to_bl w))" - by simp - also have "\ = of_bl (to_bl w @ [False])" - by (rule shiftl1_of_bl) - finally show ?thesis . -qed + apply (rule bit_word_eqI) + apply (simp add: bit_simps) + subgoal for n + apply (cases n) + apply (simp_all add: nth_rev_to_bl) + done + done lemma bl_shiftl1: "to_bl (shiftl1 w) = tl (to_bl w) @ [False]" for w :: "'a::len word" by (simp add: shiftl1_bl word_rep_drop drop_Suc drop_Cons') (fast intro!: Suc_leI) +lemma to_bl_double_eq: + \to_bl (2 * w) = tl (to_bl w) @ [False]\ + using bl_shiftl1 [of w] by (simp add: shiftl1_def ac_simps) + \ \Generalized version of \bl_shiftl1\. Maybe this one should replace it?\ lemma bl_shiftl1': "to_bl (shiftl1 w) = tl (to_bl w @ [False])" by (simp add: shiftl1_bl word_rep_drop drop_Suc del: drop_append) @@ -1438,7 +1449,7 @@ proof (rule bit_word_eqI) by simp with \n < LENGTH('a)\ show ?thesis by (simp add: bit_shiftr1_iff bit_of_bl_iff rev_nth nth_butlast - word_size test_bit_word_eq to_bl_nth) + word_size to_bl_nth) qed qed @@ -1462,45 +1473,35 @@ proof (rule nth_equalityI) by simp then show \to_bl (sshiftr1 w) ! n \ (hd (to_bl w) # butlast (to_bl w)) ! n\ apply (cases n) - apply (simp_all add: to_bl_nth word_size hd_conv_nth test_bit_eq_bit bit_sshiftr1_iff nth_butlast Suc_diff_Suc nth_to_bl) + apply (simp_all add: to_bl_nth word_size hd_conv_nth bit_sshiftr1_iff nth_butlast Suc_diff_Suc nth_to_bl) done qed simp lemma drop_shiftr: "drop n (to_bl (w >> n)) = take (size w - n) (to_bl w)" for w :: "'a::len word" - apply (unfold shiftr_def) - apply (induct n) - prefer 2 - apply (simp add: drop_Suc bl_shiftr1 butlast_drop [symmetric]) - apply (rule butlast_take [THEN trans]) - apply (auto simp: word_size) + apply (rule nth_equalityI) + apply (simp_all add: word_size to_bl_nth bit_simps) done lemma drop_sshiftr: "drop n (to_bl (w >>> n)) = take (size w - n) (to_bl w)" for w :: "'a::len word" - apply (simp_all add: word_size sshiftr_eq) apply (rule nth_equalityI) - apply (simp_all add: word_size nth_to_bl bit_signed_drop_bit_iff) + apply (simp_all add: word_size nth_to_bl bit_simps) done lemma take_shiftr: "n \ size w \ take n (to_bl (w >> n)) = replicate n False" - apply (unfold shiftr_def) - apply (induct n) - prefer 2 - apply (simp add: bl_shiftr1' length_0_conv [symmetric] word_size) - apply (rule take_butlast [THEN trans]) - apply (auto simp: word_size) + apply (rule nth_equalityI) + apply (auto simp add: word_size to_bl_nth bit_simps dest: bit_imp_le_length) done lemma take_sshiftr': "n \ size w \ hd (to_bl (w >>> n)) = hd (to_bl w) \ take n (to_bl (w >>> n)) = replicate n (hd (to_bl w))" for w :: "'a::len word" - apply (auto simp add: sshiftr_eq hd_bl_sign_sint bin_sign_def not_le word_size sint_signed_drop_bit_eq) + apply (cases n) + apply (auto simp add: hd_to_bl_iff bit_simps not_less word_size) apply (rule nth_equalityI) - apply (auto simp add: nth_to_bl bit_signed_drop_bit_iff bit_last_iff) - apply (rule nth_equalityI) - apply (auto simp add: nth_to_bl bit_signed_drop_bit_iff bit_last_iff) + apply (auto simp add: nth_to_bl bit_simps nth_Cons split: nat.split) done lemmas hd_sshiftr = take_sshiftr' [THEN conjunct1] @@ -1513,17 +1514,13 @@ lemmas bl_shiftr = atd_lem [OF take_shiftr drop_shiftr] lemmas bl_sshiftr = atd_lem [OF take_sshiftr drop_sshiftr] lemma shiftl_of_bl: "of_bl bl << n = of_bl (bl @ replicate n False)" - by (induct n) (auto simp: shiftl_def shiftl1_of_bl replicate_app_Cons_same) + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps nth_append) + done lemma shiftl_bl: "w << n = of_bl (to_bl w @ replicate n False)" for w :: "'a::len word" -proof - - have "w << n = of_bl (to_bl w) << n" - by simp - also have "\ = of_bl (to_bl w @ replicate n False)" - by (rule shiftl_of_bl) - finally show ?thesis . -qed + by (simp flip: shiftl_of_bl) lemma bl_shiftl: "to_bl (w << n) = drop n (to_bl w) @ replicate (min (size w) n) False" by (simp add: shiftl_bl word_rep_drop word_size) @@ -1531,19 +1528,16 @@ lemma bl_shiftl: "to_bl (w << n) = drop n (to_bl w) @ replicate (min (size w) n) lemma shiftr1_bl_of: "length bl \ LENGTH('a) \ shiftr1 (of_bl bl::'a::len word) = of_bl (butlast bl)" - by transfer (simp add: butlast_rest_bl2bin trunc_bl2bin) + apply (rule bit_word_eqI) + apply (simp add: bit_simps) + apply (cases bl rule: rev_cases) + apply auto + done lemma shiftr_bl_of: "length bl \ LENGTH('a) \ - (of_bl bl::'a::len word) >> n = of_bl (take (length bl - n) bl)" - apply (unfold shiftr_def) - apply (induct n) - apply clarsimp - apply clarsimp - apply (subst shiftr1_bl_of) - apply simp - apply (simp add: butlast_take) - done + (of_bl bl::'a::len word) >> n = of_bl (take (length bl - n) bl)" + by (rule bit_word_eqI) (auto simp add: bit_simps rev_nth) lemma shiftr_bl: "x >> n \ of_bl (take (LENGTH('a) - n) (to_bl x))" for x :: "'a::len word" @@ -1567,7 +1561,7 @@ lemma aligned_bl_add_size [OF refl]: done lemma mask_bl: "mask n = of_bl (replicate n True)" - by (auto simp add : test_bit_of_bl word_size intro: word_eqI) + by (auto simp add: bit_simps intro!: word_eqI) lemma bl_and_mask': "to_bl (w AND mask n :: 'a::len word) = @@ -1575,7 +1569,7 @@ lemma bl_and_mask': drop (LENGTH('a) - n) (to_bl w)" apply (rule nth_equalityI) apply simp - apply (clarsimp simp add: to_bl_nth word_size) + apply (clarsimp simp add: to_bl_nth word_size bit_simps) apply (auto simp add: word_size test_bit_bl nth_append rev_nth) done @@ -1771,6 +1765,8 @@ declare word_rotl_eqs (1) [simp] lemmas abl_cong = arg_cong [where f = "of_bl"] +end + locale word_rotate begin @@ -1823,7 +1819,7 @@ lemmas word_rotl_dt_no_bin' [simp] = word_rotl_dt [where w="numeral w", unfolded to_bl_numeral] for w (* FIXME: negative numerals, 0 and 1 *) -lemma max_word_bl: "to_bl (max_word::'a::len word) = replicate LENGTH('a) True" +lemma max_word_bl: "to_bl (- 1::'a::len word) = replicate LENGTH('a) True" by (fact to_bl_n1) lemma to_bl_mask: @@ -1842,6 +1838,10 @@ lemma map_replicate_False: (zip xs (replicate n False)) = replicate n False" by (induct xs arbitrary: n) auto +context + includes bit_operations_syntax +begin + lemma bl_and_mask: fixes w :: "'a::len word" and n :: nat @@ -1890,18 +1890,18 @@ lemma bin_to_bl_or: lemma word_and_1_bl: fixes x::"'a::len word" - shows "(x AND 1) = of_bl [x !! 0]" - by (simp add: mod_2_eq_odd test_bit_word_eq and_one_eq) + shows "(x AND 1) = of_bl [bit x 0]" + by (simp add: mod_2_eq_odd and_one_eq) lemma word_1_and_bl: fixes x::"'a::len word" - shows "(1 AND x) = of_bl [x !! 0]" - by (simp add: mod_2_eq_odd test_bit_word_eq one_and_eq) + shows "(1 AND x) = of_bl [bit x 0]" + by (simp add: mod_2_eq_odd one_and_eq) lemma of_bl_drop: "of_bl (drop n xs) = (of_bl xs AND mask (length xs - n))" - apply (clarsimp simp: bang_eq test_bit_of_bl rev_nth cong: rev_conj_cong) - apply (safe; simp add: word_size to_bl_nth) + apply (rule bit_word_eqI) + apply (auto simp: rev_nth bit_simps cong: rev_conj_cong) done lemma to_bl_1: @@ -1959,41 +1959,16 @@ lemma word_lsb_last: lemma is_aligned_to_bl: "is_aligned (w :: 'a :: len word) n = (True \ set (drop (size w - n) (to_bl w)))" - apply (simp add: is_aligned_mask eq_zero_set_bl) - apply (clarsimp simp: in_set_conv_nth word_size) - apply (simp add: to_bl_nth word_size cong: conj_cong) - apply (simp add: diff_diff_less) - apply safe - apply (case_tac "n \ LENGTH('a)") - prefer 2 - apply (rule_tac x=i in exI) - apply clarsimp - apply (subgoal_tac "\j < LENGTH('a). j < n \ LENGTH('a) - n + j = i") - apply (erule exE) - apply (rule_tac x=j in exI) - apply clarsimp - apply (thin_tac "w !! t" for t) - apply (rule_tac x="i + n - LENGTH('a)" in exI) - apply clarsimp - apply arith - apply (rule_tac x="LENGTH('a) - n + i" in exI) - apply clarsimp - apply arith - done + by (simp add: is_aligned_mask eq_zero_set_bl bl_and_mask word_size) lemma is_aligned_replicate: fixes w::"'a::len word" assumes aligned: "is_aligned w n" and nv: "n \ LENGTH('a)" shows "to_bl w = (take (LENGTH('a) - n) (to_bl w)) @ replicate n False" -proof - - from nv have rl: "\q. q < 2 ^ (LENGTH('a) - n) \ - to_bl (2 ^ n * (of_nat q :: 'a word)) = - drop n (to_bl (of_nat q :: 'a word)) @ replicate n False" - by (metis bl_shiftl le_antisym min_def shiftl_t2n wsst_TYs(3)) - show ?thesis using aligned - by (auto simp: rl elim: is_alignedE) -qed + apply (rule nth_equalityI) + using assms apply (simp_all add: nth_append not_less word_size to_bl_nth is_aligned_imp_not_bit) + done lemma is_aligned_drop: fixes w::"'a::len word" @@ -2048,9 +2023,12 @@ lemma to_bl_2p: "n < LENGTH('a) \ to_bl ((2::'a::len word) ^ n) = replicate (LENGTH('a) - Suc n) False @ True # replicate n False" - apply (subst shiftl_1 [symmetric]) - apply (subst bl_shiftl) - apply (simp add: to_bl_1 min_def word_size) + apply (rule nth_equalityI) + apply (auto simp add: nth_append to_bl_nth word_size bit_simps not_less nth_Cons le_diff_conv) + subgoal for i + apply (cases \Suc (i + n) - LENGTH('a)\) + apply simp_all + done done lemma xor_2p_to_bl: @@ -2059,22 +2037,10 @@ lemma xor_2p_to_bl: (if n < LENGTH('a) then take (LENGTH('a)-Suc n) (to_bl x) @ (\rev (to_bl x)!n) # drop (LENGTH('a)-n) (to_bl x) else to_bl x)" -proof - - have x: "to_bl x = take (LENGTH('a)-Suc n) (to_bl x) @ drop (LENGTH('a)-Suc n) (to_bl x)" - by simp - - show ?thesis - apply simp - apply (rule conjI) - apply (clarsimp simp: word_size) - apply (simp add: bl_word_xor to_bl_2p) - apply (subst x) - apply (subst zip_append) - apply simp - apply (simp add: map_zip_replicate_False_xor drop_minus) - apply (auto simp add: word_size nth_w2p intro!: word_eqI) + apply (auto simp add: to_bl_eq_rev take_map drop_map take_rev drop_rev bit_simps) + apply (rule nth_equalityI) + apply (auto simp add: bit_simps rev_nth nth_append Suc_diff_Suc) done -qed lemma is_aligned_replicateD: "\ is_aligned (w::'a::len word) n; n \ LENGTH('a) \ @@ -2120,7 +2086,7 @@ lemma of_bl_mult_and_not_mask_eq: lemma bin_to_bl_of_bl_eq: "\is_aligned (a::'a::len word) n; length b + c \ n; length b + c < LENGTH('a)\ \ bin_to_bl (length b) (uint ((a + of_bl b * 2^c) >> c)) = b" - apply (simp flip: push_bit_eq_mult take_bit_eq_mask add: shiftr_eq_drop_bit) + apply (simp flip: push_bit_eq_mult take_bit_eq_mask) apply (subst disjunctive_add) apply (auto simp add: bit_simps not_le not_less unsigned_or_eq unsigned_drop_bit_eq unsigned_push_bit_eq bin_to_bl_or simp flip: bin_to_bl_def) @@ -2130,15 +2096,6 @@ lemma bin_to_bl_of_bl_eq: apply (auto simp add: nth_bin_to_bl bit_simps rev_nth simp flip: bin_to_bl_def) done -(* FIXME: move to Word distribution *) -lemma bin_nth_minus_Bit0[simp]: - "0 < n \ bin_nth (numeral (num.Bit0 w)) n = bin_nth (numeral w) (n - 1)" - by (cases n; simp) - -lemma bin_nth_minus_Bit1[simp]: - "0 < n \ bin_nth (numeral (num.Bit1 w)) n = bin_nth (numeral w) (n - 1)" - by (cases n; simp) - (* casting a long word to a shorter word and casting back to the long word is equal to the original long word -- if the word is small enough. 'l is the longer word. @@ -2174,7 +2131,7 @@ proof - qed lemma map_bits_rev_to_bl: - "map ((!!) x) [0..Some auxiliaries for sign-shifting by the entire word length or more\ + +lemma sshiftr_clamp_pos: + assumes + "LENGTH('a) \ n" + "0 \ sint x" + shows "(x::'a::len word) >>> n = 0" + apply (rule bit_word_eqI) + using assms + apply (auto simp add: bit_simps bit_last_iff) + done + +lemma sshiftr_clamp_neg: + assumes + "LENGTH('a) \ n" + "sint x < 0" + shows "(x::'a::len word) >>> n = -1" + apply (rule bit_word_eqI) + using assms + apply (auto simp add: bit_simps bit_last_iff) + done + +lemma sshiftr_clamp: + assumes "LENGTH('a) \ n" + shows "(x::'a::len word) >>> n = x >>> LENGTH('a)" + apply (rule bit_word_eqI) + using assms + apply (auto simp add: bit_simps bit_last_iff) + done + +text\ +Like @{thm shiftr1_bl_of}, but the precondition is stronger because we need to pick the msb out of +the list. +\ +lemma sshiftr1_bl_of: + "length bl = LENGTH('a) \ + sshiftr1 (of_bl bl::'a::len word) = of_bl (hd bl # butlast bl)" + apply (rule word_bl.Rep_eqD) + apply (subst bl_sshiftr1[of "of_bl bl :: 'a word"]) + by (simp add: word_bl.Abs_inverse) + +text\ +Like @{thm sshiftr1_bl_of}, with a weaker precondition. +We still get a direct equation for @{term \sshiftr1 (of_bl bl)\}, it's just uglier. +\ +lemma sshiftr1_bl_of': + "LENGTH('a) \ length bl \ + sshiftr1 (of_bl bl::'a::len word) = + of_bl (hd (drop (length bl - LENGTH('a)) bl) # butlast (drop (length bl - LENGTH('a)) bl))" + apply (subst of_bl_drop'[symmetric, of "length bl - LENGTH('a)"]) + using sshiftr1_bl_of[of "drop (length bl - LENGTH('a)) bl"] + by auto + +text\ +Like @{thm shiftr_bl_of}. +\ +lemma sshiftr_bl_of: + assumes "length bl = LENGTH('a)" + shows "(of_bl bl::'a::len word) >>> n = of_bl (replicate n (hd bl) @ take (length bl - n) bl)" +proof - + from assms obtain b bs where \bl = b # bs\ + by (cases bl) simp_all + then have *: \bl ! 0 \ b\ \hd bl \ b\ + by simp_all + show ?thesis + apply (rule bit_word_eqI) + using assms * by (auto simp add: bit_simps nth_append rev_nth not_less) +qed + +text\Like @{thm shiftr_bl}\ +lemma sshiftr_bl: "x >>> n \ of_bl (replicate n (msb x) @ take (LENGTH('a) - n) (to_bl x))" + for x :: "'a::len word" + unfolding word_msb_alt + by (smt (z3) length_to_bl_eq sshiftr_bl_of word_bl.Rep_inverse) + +end + +lemma of_bl_drop_eq_take_bit: + \of_bl (drop n xs) = take_bit (length xs - n) (of_bl xs)\ + by (simp add: of_bl_drop take_bit_eq_mask) + +lemma of_bl_take_to_bl_eq_drop_bit: + \of_bl (take n (to_bl w)) = drop_bit (LENGTH('a) - n) w\ + if \n \ LENGTH('a)\ + for w :: \'a::len word\ + using that shiftr_bl [of w \LENGTH('a) - n\] by (simp add: shiftr_def) + end diff --git a/lib/Word_Lib/Rsplit.thy b/lib/Word_Lib/Rsplit.thy index 405361488..58a6f7b87 100644 --- a/lib/Word_Lib/Rsplit.thy +++ b/lib/Word_Lib/Rsplit.thy @@ -6,12 +6,14 @@ (* Author: Jeremy Dawson and Gerwin Klein, NICTA *) +section \Splitting words into lists\ + theory Rsplit - imports "HOL-Library.Word" Bits_Int + imports "HOL-Library.Word" More_Word Bits_Int begin definition word_rsplit :: "'a::len word \ 'b::len word list" - where "word_rsplit w = map word_of_int (bin_rsplit (LENGTH('b)) (LENGTH('a), uint w))" + where "word_rsplit w = map word_of_int (bin_rsplit LENGTH('b) (LENGTH('a), uint w))" lemma word_rsplit_no: "(word_rsplit (numeral bin :: 'b::len word) :: 'a word list) = @@ -39,11 +41,11 @@ lemma word_rsplit_empty_iff_size: "word_rsplit w = [] \ size lemma test_bit_rsplit: "sw = word_rsplit w \ m < size (hd sw) \ - k < length sw \ (rev sw ! k) !! m = w !! (k * size (hd sw) + m)" + k < length sw \ bit (rev sw ! k) m = bit w (k * size (hd sw) + m)" for sw :: "'a::len word list" apply (unfold word_rsplit_def word_test_bit_def) apply (rule trans) - apply (rule_tac f = "\x. bin_nth x m" in arg_cong) + apply (rule_tac f = "\x. bit x m" in arg_cong) apply (rule nth_map [symmetric]) apply simp apply (rule bin_nth_rsplit) @@ -53,12 +55,12 @@ lemma test_bit_rsplit: defer apply (rule map_ident [THEN fun_cong]) apply (rule refl [THEN map_cong]) - apply simp + apply (simp add: unsigned_of_int take_bit_int_eq_self_iff) using bin_rsplit_size_sign take_bit_int_eq_self_iff by blast lemma test_bit_rsplit_alt: - \(word_rsplit w :: 'b::len word list) ! i !! m \ - w !! ((length (word_rsplit w :: 'b::len word list) - Suc i) * size (hd (word_rsplit w :: 'b::len word list)) + m)\ + \bit ((word_rsplit w :: 'b::len word list) ! i) m \ + bit w ((length (word_rsplit w :: 'b::len word list) - Suc i) * size (hd (word_rsplit w :: 'b::len word list)) + m)\ if \i < length (word_rsplit w :: 'b::len word list)\ \m < size (hd (word_rsplit w :: 'b::len word list))\ \0 < length (word_rsplit w :: 'b::len word list)\ for w :: \'a::len word\ apply (rule trans) @@ -120,12 +122,6 @@ lemma size_word_rsplit_rcat_size: for ws :: "'a::len word list" and frcw :: "'b::len word" by (cases \LENGTH('a)\) (simp_all add: word_size length_word_rsplit_exp_size' div_nat_eqI) -lemma msrevs: - "0 < n \ (k * n + m) div n = m div n + k" - "(k * n + m) mod n = m mod n" - for n :: nat - by (auto simp: add.commute) - lemma word_rsplit_rcat_size [OF refl]: "word_rcat ws = frcw \ size frcw = length ws * LENGTH('a) \ word_rsplit frcw = ws" @@ -152,15 +148,16 @@ lemma word_rsplit_rcat_size [OF refl]: lemma word_rsplit_upt: "\ size x = LENGTH('a :: len) * n; n \ 0 \ - \ word_rsplit x = map (\i. ucast (x >> i * len_of TYPE ('a)) :: 'a word) (rev [0 ..< n])" + \ word_rsplit x = map (\i. ucast (x >> (i * LENGTH('a))) :: 'a word) (rev [0 ..< n])" apply (subgoal_tac "length (word_rsplit x :: 'a word list) = n") apply (rule nth_equalityI, simp) apply (intro allI word_eqI impI) apply (simp add: test_bit_rsplit_alt word_size) - apply (simp add: nth_ucast nth_shiftr rev_nth field_simps) - apply (simp add: length_word_rsplit_exp_size) - apply transfer - apply (metis (no_types, lifting) Nat.add_diff_assoc Suc_leI add_0_left diff_Suc_less div_less len_gt_0 msrevs(1) mult.commute) + apply (simp add: nth_ucast bit_simps rev_nth field_simps) + apply (simp add: length_word_rsplit_exp_size word_size) + apply (subst diff_add_assoc) + apply (simp flip: less_eq_Suc_le) + apply simp done end \ No newline at end of file diff --git a/lib/Word_Lib/Signed_Division_Word.thy b/lib/Word_Lib/Signed_Division_Word.thy index 288cb02b4..7aeaf4aa8 100644 --- a/lib/Word_Lib/Signed_Division_Word.thy +++ b/lib/Word_Lib/Signed_Division_Word.thy @@ -52,6 +52,15 @@ lemma signed_mod_arith: apply simp done +lemma word_sdiv_div0 [simp]: + "(a :: ('a::len) word) sdiv 0 = 0" + apply (auto simp: sdiv_word_def signed_divide_int_def sgn_if) + done + +lemma smod_word_zero [simp]: + \w smod 0 = w\ for w :: \'a::len word\ + by (simp add: smod_word_def signed_modulo_int_def) + lemma word_sdiv_div1 [simp]: "(a :: ('a::len) word) sdiv 1 = a" apply (cases \LENGTH('a)\) @@ -63,10 +72,9 @@ lemma word_sdiv_div1 [simp]: apply (simp add: take_bit_signed_take_bit) done -lemma word_sdiv_div0 [simp]: - "(a :: ('a::len) word) sdiv 0 = 0" - apply (auto simp: sdiv_word_def signed_divide_int_def sgn_if) - done +lemma smod_word_one [simp]: + \w smod 1 = 0\ for w :: \'a::len word\ + by (simp add: smod_word_def signed_modulo_int_def) lemma word_sdiv_div_minus1 [simp]: "(a :: ('a::len) word) sdiv -1 = -a" @@ -76,6 +84,78 @@ lemma word_sdiv_div_minus1 [simp]: apply (metis Suc_pred len_gt_0 signed_take_bit_eq_iff_take_bit_eq signed_take_bit_of_0 take_bit_of_0) done +lemma smod_word_minus_one [simp]: + \w smod - 1 = 0\ for w :: \'a::len word\ + by (simp add: smod_word_def signed_modulo_int_def) + +lemma one_sdiv_word_eq [simp]: + \1 sdiv w = of_bool (w = 1 \ w = - 1) * w\ for w :: \'a::len word\ +proof (cases \1 < \sint w\\) + case True + then show ?thesis + by (auto simp add: sdiv_word_def signed_divide_int_def split: if_splits) +next + case False + then have \\sint w\ \ 1\ + by simp + then have \sint w \ {- 1, 0, 1}\ + by auto + then have \(word_of_int (sint w) :: 'a::len word) \ word_of_int ` {- 1, 0, 1}\ + by blast + then have \w \ {- 1, 0, 1}\ + by simp + then show ?thesis by auto +qed + +lemma one_smod_word_eq [simp]: + \1 smod w = 1 - of_bool (w = 1 \ w = - 1)\ for w :: \'a::len word\ + using sdiv_smod_id [of 1 w] by auto + +lemma minus_one_sdiv_word_eq [simp]: + \- 1 sdiv w = - (1 sdiv w)\ for w :: \'a::len word\ + apply (auto simp add: sdiv_word_def signed_divide_int_def) + apply transfer + apply simp + done + +lemma minus_one_smod_word_eq [simp]: + \- 1 smod w = - (1 smod w)\ for w :: \'a::len word\ + using sdiv_smod_id [of \- 1\ w] by auto + +lemma smod_word_alt_def: + "(a :: ('a::len) word) smod b = a - (a sdiv b) * b" + apply (cases \a \ - (2 ^ (LENGTH('a) - 1)) \ b \ - 1\) + apply (clarsimp simp: smod_word_def sdiv_word_def signed_modulo_int_def + simp flip: wi_hom_sub wi_hom_mult) + apply (clarsimp simp: smod_word_def signed_modulo_int_def) + done + +lemmas sdiv_word_numeral_numeral [simp] = + sdiv_word_def [of \numeral a\ \numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] + for a b +lemmas sdiv_word_minus_numeral_numeral [simp] = + sdiv_word_def [of \- numeral a\ \numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] + for a b +lemmas sdiv_word_numeral_minus_numeral [simp] = + sdiv_word_def [of \numeral a\ \- numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] + for a b +lemmas sdiv_word_minus_numeral_minus_numeral [simp] = + sdiv_word_def [of \- numeral a\ \- numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] + for a b + +lemmas smod_word_numeral_numeral [simp] = + smod_word_def [of \numeral a\ \numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] + for a b +lemmas smod_word_minus_numeral_numeral [simp] = + smod_word_def [of \- numeral a\ \numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] + for a b +lemmas smod_word_numeral_minus_numeral [simp] = + smod_word_def [of \numeral a\ \- numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] + for a b +lemmas smod_word_minus_numeral_minus_numeral [simp] = + smod_word_def [of \- numeral a\ \- numeral b\, simplified sint_sbintrunc sint_sbintrunc_neg] + for a b + lemmas word_sdiv_0 = word_sdiv_div0 lemma sdiv_word_min: @@ -93,15 +173,34 @@ lemma sdiv_word_min: apply (metis abs_le_iff add.inverse_inverse le_cases le_minus_iff not_le signed_take_bit_int_eq_self_iff signed_take_bit_minus) done +lemma sdiv_word_max: + \sint a sdiv sint b \ 2 ^ (size a - Suc 0)\ + for a b :: \'a::len word\ +proof (cases \sint a = 0 \ sint b = 0 \ sgn (sint a) \ sgn (sint b)\) + case True then show ?thesis + apply (auto simp add: sgn_if not_less signed_divide_int_def split: if_splits) + apply (smt (z3) pos_imp_zdiv_neg_iff zero_le_power) + apply (smt (z3) not_exp_less_eq_0_int pos_imp_zdiv_neg_iff) + done +next + case False + then have \\sint a\ div \sint b\ \ \sint a\\ + by (subst nat_le_eq_zle [symmetric]) (simp_all add: div_abs_eq_div_nat) + also have \\sint a\ \ 2 ^ (size a - Suc 0)\ + using sint_range_size [of a] by auto + finally show ?thesis + using False by (simp add: signed_divide_int_def) +qed + lemmas word_sdiv_numerals_lhs = sdiv_word_def[where v="numeral x" for x] sdiv_word_def[where v=0] sdiv_word_def[where v=1] lemmas word_sdiv_numerals = word_sdiv_numerals_lhs[where w="numeral y" for y] word_sdiv_numerals_lhs[where w=0] word_sdiv_numerals_lhs[where w=1] -lemma smod_word_mod_0 [simp]: +lemma smod_word_mod_0: "x smod (0 :: ('a::len) word) = x" - by (clarsimp simp: smod_word_def) + by (fact smod_word_zero) lemma smod_word_0_mod [simp]: "0 smod (x :: ('a::len) word) = 0" @@ -139,14 +238,6 @@ lemma smod_word_min: apply (metis abs_zero add.inverse_inverse add.left_neutral add_mono_thms_linordered_semiring(1) diff_Suc_1 le_cases le_minus_iff linorder_not_less sint_ge zabs_less_one_iff) done -lemma smod_word_alt_def: - "(a :: ('a::len) word) smod b = a - (a sdiv b) * b" - apply (cases \a \ - (2 ^ (LENGTH('a) - 1)) \ b \ - 1\) - apply (clarsimp simp: smod_word_def sdiv_word_def signed_modulo_int_def - simp flip: wi_hom_sub wi_hom_mult) - apply (clarsimp simp: smod_word_def signed_modulo_int_def) - done - lemmas word_smod_numerals_lhs = smod_word_def[where v="numeral x" for x] smod_word_def[where v=0] smod_word_def[where v=1] diff --git a/lib/Word_Lib/Signed_Words.thy b/lib/Word_Lib/Signed_Words.thy index 0ce30a19c..4aa26f75f 100644 --- a/lib/Word_Lib/Signed_Words.thy +++ b/lib/Word_Lib/Signed_Words.thy @@ -81,7 +81,7 @@ lemma word_le_ucast_sless: apply simp_all apply transfer apply (simp add: signed_take_bit_take_bit) - apply (metis add.commute mask_eq_exp_minus_1 mask_eq_take_bit_minus_one take_bit_incr_eq zle_add1_eq_le) + apply (metis add.commute mask_eq_exp_minus_1 take_bit_incr_eq zle_add1_eq_le) done lemma zero_sle_ucast: @@ -93,6 +93,21 @@ lemma zero_sle_ucast: apply (simp_all add: bit_simps disjunctive_add) done +lemma nth_w2p_scast: + "(bit (scast ((2::'a::len signed word) ^ n) :: 'a word) m) + \ (bit (((2::'a::len word) ^ n) :: 'a word) m)" + by (simp add: bit_simps) + +lemma scast_nop1 [simp]: + "((scast ((of_int x)::('a::len) word))::'a signed word) = of_int x" + by transfer (simp add: take_bit_signed_take_bit) + +lemma scast_nop2 [simp]: + "((scast ((of_int x)::('a::len) signed word))::'a word) = of_int x" + by transfer (simp add: take_bit_signed_take_bit) + +lemmas scast_nop = scast_nop1 scast_nop2 scast_id + type_synonym 'a sword = "'a signed word" end diff --git a/lib/Word_Lib/Singleton_Bit_Shifts.thy b/lib/Word_Lib/Singleton_Bit_Shifts.thy new file mode 100644 index 000000000..05a20040a --- /dev/null +++ b/lib/Word_Lib/Singleton_Bit_Shifts.thy @@ -0,0 +1,148 @@ +theory Singleton_Bit_Shifts + imports + "HOL-Library.Word" + Bit_Shifts_Infix_Syntax +begin + +definition shiftl1 :: \'a::len word \ 'a word\ + where \shiftl1 = push_bit 1\ + +lemma bit_shiftl1_iff [bit_simps]: + \bit (shiftl1 w) n \ 0 < n \ n < LENGTH('a) \ bit w (n - 1)\ + for w :: \'a::len word\ + by (simp only: shiftl1_def bit_push_bit_iff) auto + +definition shiftr1 :: \'a::len word \ 'a word\ + where \shiftr1 = drop_bit 1\ + +lemma bit_shiftr1_iff [bit_simps]: + \bit (shiftr1 w) n \ bit w (Suc n)\ + for w :: \'a::len word\ + by (simp add: shiftr1_def bit_drop_bit_eq) + +definition sshiftr1 :: \'a::len word \ 'a word\ + where \sshiftr1 \ signed_drop_bit 1\ + +lemma bit_sshiftr1_iff [bit_simps]: + \bit (sshiftr1 w) n \ bit w (if n = LENGTH('a) - 1 then LENGTH('a) - 1 else Suc n)\ + for w :: \'a::len word\ + by (auto simp add: sshiftr1_def bit_signed_drop_bit_iff) + +lemma shiftr1_1: "shiftr1 (1::'a::len word) = 0" + by (simp add: shiftr1_def) + +lemma sshiftr1_eq: + \sshiftr1 w = word_of_int (sint w div 2)\ + by (rule bit_word_eqI) (auto simp add: bit_simps min_def simp flip: bit_Suc elim: le_SucE) + +lemma shiftl1_eq: + \shiftl1 w = word_of_int (2 * uint w)\ + by (rule bit_word_eqI) (auto simp add: bit_simps) + +lemma shiftr1_eq: + \shiftr1 w = word_of_int (uint w div 2)\ + by (rule bit_word_eqI) (simp add: bit_simps flip: bit_Suc) + +lemma shiftl1_rev: + "shiftl1 w = word_reverse (shiftr1 (word_reverse w))" + by (rule bit_word_eqI) (auto simp add: bit_simps Suc_diff_Suc simp flip: bit_Suc) + +lemma shiftl1_p: + "shiftl1 w = w + w" + for w :: "'a::len word" + by (simp add: shiftl1_def) + +lemma shiftr1_bintr: + "(shiftr1 (numeral w) :: 'a::len word) = + word_of_int (take_bit LENGTH('a) (numeral w) div 2)" + by (rule bit_word_eqI) (simp add: bit_simps bit_numeral_iff [where ?'a = int] flip: bit_Suc) + +lemma sshiftr1_sbintr: + "(sshiftr1 (numeral w) :: 'a::len word) = + word_of_int (signed_take_bit (LENGTH('a) - 1) (numeral w) div 2)" + apply (cases \LENGTH('a)\) + apply simp_all + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps min_def simp flip: bit_Suc elim: le_SucE) + done + +lemma shiftl1_wi: + "shiftl1 (word_of_int w) = word_of_int (2 * w)" + by (rule bit_word_eqI) (auto simp add: bit_simps) + +lemma shiftl1_numeral: + "shiftl1 (numeral w) = numeral (Num.Bit0 w)" + unfolding word_numeral_alt shiftl1_wi by simp + +lemma shiftl1_neg_numeral: + "shiftl1 (- numeral w) = - numeral (Num.Bit0 w)" + unfolding word_neg_numeral_alt shiftl1_wi by simp + +lemma shiftl1_0: + "shiftl1 0 = 0" + by (simp add: shiftl1_def) + +lemma shiftl1_def_u: + "shiftl1 w = word_of_int (2 * uint w)" + by (fact shiftl1_eq) + +lemma shiftl1_def_s: + "shiftl1 w = word_of_int (2 * sint w)" + by (simp add: shiftl1_def) + +lemma shiftr1_0: + "shiftr1 0 = 0" + by (simp add: shiftr1_def) + +lemma sshiftr1_0: + "sshiftr1 0 = 0" + by (simp add: sshiftr1_def) + +lemma sshiftr1_n1: + "sshiftr1 (- 1) = - 1" + by (simp add: sshiftr1_def) + +lemma uint_shiftr1: + "uint (shiftr1 w) = uint w div 2" + by (rule bit_eqI) (simp add: bit_simps flip: bit_Suc) + +lemma shiftr1_div_2: + "uint (shiftr1 w) = uint w div 2" + by (fact uint_shiftr1) + +lemma sshiftr1_div_2: + "sint (sshiftr1 w) = sint w div 2" + by (rule bit_eqI) (auto simp add: bit_simps ac_simps min_def simp flip: bit_Suc elim: le_SucE) + +lemma nth_shiftl1: + "bit (shiftl1 w) n \ n < size w \ n > 0 \ bit w (n - 1)" + by (auto simp add: word_size bit_simps) + +lemma nth_shiftr1: + "bit (shiftr1 w) n = bit w (Suc n)" + by (fact bit_shiftr1_iff) + +lemma nth_sshiftr1: "bit (sshiftr1 w) n = (if n = size w - 1 then bit w n else bit w (Suc n))" + by (auto simp add: word_size bit_simps) + +lemma shiftl_power: + "(shiftl1 ^^ x) (y::'a::len word) = 2 ^ x * y" + by (induction x) (simp_all add: shiftl1_def) + +lemma le_shiftr1: + \shiftr1 u \ shiftr1 v\ if \u \ v\ + using that by (simp add: word_le_nat_alt unat_div div_le_mono shiftr1_def drop_bit_Suc) + +lemma le_shiftr1': + "\ shiftr1 u \ shiftr1 v ; shiftr1 u \ shiftr1 v \ \ u \ v" + by (meson dual_order.antisym le_cases le_shiftr1) + +lemma sshiftr_eq_funpow_sshiftr1: + \w >>> n = (sshiftr1 ^^ n) w\ + apply (rule sym) + apply (simp add: sshiftr1_def sshiftr_def) + apply (induction n) + apply simp_all + done + +end diff --git a/lib/Word_Lib/Strict_part_mono.thy b/lib/Word_Lib/Strict_part_mono.thy index 937e8c67f..b4bd3dd6f 100644 --- a/lib/Word_Lib/Strict_part_mono.thy +++ b/lib/Word_Lib/Strict_part_mono.thy @@ -28,7 +28,7 @@ lemma strict_part_mono_singleton[simp]: lemma strict_part_mono_lt: "\ x < f 0; strict_part_mono {.. n :: nat} f \ \ \m \ n. x < f m" - by (metis atMost_iff le_0_eq le_cases neq0_conv order.strict_trans strict_part_mono_def) + by (auto simp add: strict_part_mono_def Ball_def intro: order.strict_trans) lemma strict_part_mono_reverseE: "\ f n \ f m; strict_part_mono {.. N :: nat} f; n \ N \ \ n \ m" diff --git a/lib/Word_Lib/Syntax_Bundles.thy b/lib/Word_Lib/Syntax_Bundles.thy new file mode 100644 index 000000000..6507af8ee --- /dev/null +++ b/lib/Word_Lib/Syntax_Bundles.thy @@ -0,0 +1,20 @@ +(* + * Copyright Florian Haftmann + * + * SPDX-License-Identifier: BSD-2-Clause + *) + +section \Syntax bundles for traditional infix syntax\ + +theory Syntax_Bundles + imports "HOL-Library.Word" +begin + +bundle bit_projection_infix_syntax +begin + +notation bit (infixl \!!\ 100) + +end + +end diff --git a/lib/Word_Lib/Traditional_Infix_Syntax.thy b/lib/Word_Lib/Traditional_Infix_Syntax.thy deleted file mode 100644 index 9c94ca71e..000000000 --- a/lib/Word_Lib/Traditional_Infix_Syntax.thy +++ /dev/null @@ -1,1085 +0,0 @@ -(* - * Copyright Data61, CSIRO (ABN 41 687 119 230) - * - * SPDX-License-Identifier: BSD-2-Clause - *) - -(* Author: Jeremy Dawson, NICTA *) - -section \Operation variants with traditional syntax\ - -theory Traditional_Infix_Syntax - imports "HOL-Library.Word" More_Word Signed_Words -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 [iff]: "(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_simps]: - \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 not_le) - -lemma bit_shiftr_word_iff [bit_simps]: - \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_simps]: - \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 sshiftr_numeral_Suc_0 [simp]: - \(numeral k >>> Suc 0 :: 'a::len word) = - word_of_int (drop_bit (Suc 0) (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) - by (metis dual_order.antisym le_Suc_eq) - -setup \ - Context.theory_map (fold SMT_Word.add_word_shift' [ - (\<^term>\shiftl :: 'a::len word \ _\, "bvshl"), - (\<^term>\shiftr :: 'a::len word \ _\, "bvlshr"), - (\<^term>\sshiftr :: 'a::len word \ _\, "bvashr") - ]) -\ - -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_numeral_Suc_0 [simp]: - \numeral k << Suc 0 = (numeral (k * num.Bit0 num.One) :: 'a::len word)\ - by (subst shiftl_word_eq) simp - -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 shiftr_numeral_Suc [simp]: - \(numeral k >> Suc 0 :: 'a::len word) = drop_bit (Suc 0) (numeral k)\ - by (fact shiftr_word_eq) - -lemma drop_bit_numeral_bit0_1 [simp]: - \drop_bit (Suc 0) (numeral k) = - (word_of_int (drop_bit (Suc 0) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\ - by (metis Word_eq_word_of_int drop_bit_word.abs_eq of_int_numeral) - -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 - 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) - -lemma word_ops_nth [simp]: - fixes x y :: \'a::len word\ - shows - word_or_nth: "(x OR y) !! n = (x !! n \ y !! n)" and - word_and_nth: "(x AND y) !! n = (x !! n \ y !! n)" and - word_xor_nth: "(x XOR y) !! n = (x !! n \ y !! n)" - by ((cases "n < size x", - auto dest: test_bit_size simp: word_ops_nth_size word_size)[1])+ - -lemma and_not_mask: - "w AND NOT (mask n) = (w >> n) << n" - for w :: \'a::len word\ - apply (rule word_eqI) - apply (simp add : word_ops_nth_size word_size) - apply (simp add : nth_shiftr nth_shiftl) - by auto - -lemma and_mask: - "w AND mask n = (w << (size w - n)) >> (size w - n)" - for w :: \'a::len word\ - apply (rule word_eqI) - apply (simp add : word_ops_nth_size word_size) - apply (simp add : nth_shiftr nth_shiftl) - by auto - -lemma nth_w2p_same: - "(2^n :: 'a :: len word) !! n = (n < LENGTH('a))" - by (simp add : nth_w2p) - -lemma shiftr_div_2n_w: "n < size w \ w >> n = w div (2^n :: 'a :: len word)" - apply (unfold word_div_def) - apply (simp add: uint_2p_alt word_size) - apply (metis uint_shiftr_eq word_of_int_uint) - done - -lemma le_shiftr: - "u \ v \ u >> (n :: nat) \ (v :: 'a :: len word) >> n" - apply (unfold shiftr_def) - apply (induct_tac "n") - apply auto - apply (erule le_shiftr1) - done - -lemma shiftr_mask_le: - "n <= m \ mask n >> m = (0 :: 'a::len word)" - apply (rule word_eqI) - apply (simp add: word_size nth_shiftr) - done - -lemma shiftr_mask [simp]: - \mask m >> m = (0::'a::len word)\ - by (rule shiftr_mask_le) simp - -lemma word_leI: - "(\n. \n < size (u::'a::len word); u !! n \ \ (v::'a::len word) !! n) \ u <= v" - apply (rule xtrans(4)) - apply (rule word_and_le2) - apply (rule word_eqI) - apply (simp add: word_ao_nth) - apply safe - apply assumption - apply (erule_tac [2] asm_rl) - apply (unfold word_size) - by auto - -lemma le_mask_iff: - "(w \ mask n) = (w >> n = 0)" - for w :: \'a::len word\ - apply safe - apply (rule word_le_0_iff [THEN iffD1]) - apply (rule xtrans(3)) - apply (erule_tac [2] le_shiftr) - apply simp - apply (rule word_leI) - apply (rename_tac n') - apply (drule_tac x = "n' - n" in word_eqD) - apply (simp add : nth_shiftr word_size) - apply (case_tac "n <= n'") - by auto - -lemma and_mask_eq_iff_shiftr_0: - "(w AND mask n = w) = (w >> n = 0)" - for w :: \'a::len word\ - apply (unfold test_bit_eq_iff [THEN sym]) - apply (rule iffI) - apply (rule ext) - apply (rule_tac [2] ext) - apply (auto simp add : word_ao_nth nth_shiftr) - apply (drule arg_cong) - apply (drule iffD2) - apply assumption - apply (simp add : word_ao_nth) - prefer 2 - apply (simp add : word_size test_bit_bin) - apply transfer - apply (auto simp add: fun_eq_iff bit_simps) - apply (metis add_diff_inverse_nat) - done - -lemma mask_shiftl_decompose: - "mask m << n = mask (m + n) AND NOT (mask n :: 'a::len word)" - by (auto intro!: word_eqI simp: and_not_mask nth_shiftl nth_shiftr word_size) - -lemma bang_eq: - fixes x :: "'a::len word" - shows "(x = y) = (\n. x !! n = y !! n)" - by (subst test_bit_eq_iff[symmetric]) fastforce - -lemma shiftl_over_and_dist: - fixes a::"'a::len word" - shows "(a AND b) << c = (a << c) AND (b << c)" - apply(rule word_eqI) - apply(simp add: word_ao_nth nth_shiftl, safe) - done - -lemma shiftr_over_and_dist: - fixes a::"'a::len word" - shows "a AND b >> c = (a >> c) AND (b >> c)" - apply(rule word_eqI) - apply(simp add:nth_shiftr word_ao_nth) - done - -lemma sshiftr_over_and_dist: - fixes a::"'a::len word" - shows "a AND b >>> c = (a >>> c) AND (b >>> c)" - apply(rule word_eqI) - apply(simp add:nth_sshiftr word_ao_nth word_size) - done - -lemma shiftl_over_or_dist: - fixes a::"'a::len word" - shows "a OR b << c = (a << c) OR (b << c)" - apply(rule word_eqI) - apply(simp add:nth_shiftl word_ao_nth, safe) - done - -lemma shiftr_over_or_dist: - fixes a::"'a::len word" - shows "a OR b >> c = (a >> c) OR (b >> c)" - apply(rule word_eqI) - apply(simp add:nth_shiftr word_ao_nth) - done - -lemma sshiftr_over_or_dist: - fixes a::"'a::len word" - shows "a OR b >>> c = (a >>> c) OR (b >>> c)" - apply(rule word_eqI) - apply(simp add:nth_sshiftr word_ao_nth word_size) - done - -lemmas shift_over_ao_dists = - shiftl_over_or_dist shiftr_over_or_dist - sshiftr_over_or_dist shiftl_over_and_dist - shiftr_over_and_dist sshiftr_over_and_dist - -lemma shiftl_shiftl: - fixes a::"'a::len word" - shows "a << b << c = a << (b + c)" - apply(rule word_eqI) - apply(auto simp:word_size nth_shiftl add.commute add.left_commute) - done - -lemma shiftr_shiftr: - fixes a::"'a::len word" - shows "a >> b >> c = a >> (b + c)" - apply(rule word_eqI) - apply(simp add:word_size nth_shiftr add.left_commute add.commute) - done - -lemma shiftl_shiftr1: - fixes a::"'a::len word" - shows "c \ b \ a << b >> c = a AND (mask (size a - b)) << (b - c)" - apply(rule word_eqI) - apply(auto simp:nth_shiftr nth_shiftl word_size word_ao_nth) - done - -lemma shiftl_shiftr2: - fixes a::"'a::len word" - shows "b < c \ a << b >> c = (a >> (c - b)) AND (mask (size a - c))" - apply(rule word_eqI) - apply(auto simp:nth_shiftr nth_shiftl word_size word_ao_nth) - done - -lemma shiftr_shiftl1: - fixes a::"'a::len word" - shows "c \ b \ a >> b << c = (a >> (b - c)) AND (NOT (mask c))" - apply(rule word_eqI) - apply(auto simp:nth_shiftr nth_shiftl word_size word_ops_nth_size) - done - -lemma shiftr_shiftl2: - fixes a::"'a::len word" - shows "b < c \ a >> b << c = (a << (c - b)) AND (NOT (mask c))" - apply(rule word_eqI) - apply(auto simp:nth_shiftr nth_shiftl word_size word_ops_nth_size) - done - -lemmas multi_shift_simps = - shiftl_shiftl shiftr_shiftr - shiftl_shiftr1 shiftl_shiftr2 - shiftr_shiftl1 shiftr_shiftl2 - -lemma shiftr_mask2: - "n \ LENGTH('a) \ (mask n >> m :: ('a :: len) word) = mask (n - m)" - apply (rule word_eqI) - apply (simp add: nth_shiftr word_size) - apply arith - done - -lemma word_shiftl_add_distrib: - fixes x :: "'a :: len word" - shows "(x + y) << n = (x << n) + (y << n)" - by (simp add: shiftl_t2n ring_distribs) - -lemma mask_shift: - "(x AND NOT (mask y)) >> y = x >> y" - for x :: \'a::len word\ - apply (rule bit_eqI) - apply (simp add: bit_and_iff bit_not_iff bit_shiftr_word_iff bit_mask_iff not_le) - using bit_imp_le_length apply auto - done - -lemma shiftr_div_2n': - "unat (w >> n) = unat w div 2 ^ n" - apply (unfold unat_eq_nat_uint) - apply (subst shiftr_div_2n) - apply (subst nat_div_distrib) - apply simp - apply (simp add: nat_power_eq) - done - -lemma shiftl_shiftr_id: - assumes nv: "n < LENGTH('a)" - and xv: "x < 2 ^ (LENGTH('a) - n)" - shows "x << n >> n = (x::'a::len word)" - apply (simp add: shiftl_t2n) - apply (rule word_eq_unatI) - apply (subst shiftr_div_2n') - apply (cases n) - apply simp - apply (subst iffD1 [OF unat_mult_lem])+ - apply (subst unat_power_lower[OF nv]) - apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) - apply (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl]) - apply (rule unat_power_lower) - apply simp - apply (subst unat_power_lower[OF nv]) - apply simp - done - -lemma ucast_shiftl_eq_0: - fixes w :: "'a :: len word" - shows "\ n \ LENGTH('b) \ \ ucast (w << n) = (0 :: 'b :: len word)" - by transfer (simp add: take_bit_push_bit) - -lemma word_shift_nonzero: - "\ (x::'a::len word) \ 2 ^ m; m + n < LENGTH('a::len); x \ 0\ - \ x << n \ 0" - apply (simp only: word_neq_0_conv word_less_nat_alt - shiftl_t2n mod_0 unat_word_ariths - unat_power_lower word_le_nat_alt) - apply (subst mod_less) - apply (rule order_le_less_trans) - apply (erule mult_le_mono2) - apply (subst power_add[symmetric]) - apply (rule power_strict_increasing) - apply simp - apply simp - apply simp - done - -lemma word_shiftr_lt: - fixes w :: "'a::len word" - shows "unat (w >> n) < (2 ^ (LENGTH('a) - n))" - apply (subst shiftr_div_2n') - apply transfer - apply (simp flip: drop_bit_eq_div add: drop_bit_nat_eq drop_bit_take_bit) - done - -lemma neg_mask_test_bit: - "(NOT(mask n) :: 'a :: len word) !! m = (n \ m \ m < LENGTH('a))" - by (metis not_le nth_mask test_bit_bin word_ops_nth_size word_size) - -lemma upper_bits_unset_is_l2p: - \(\n' \ n. n' < LENGTH('a) \ \ p !! n') \ (p < 2 ^ n)\ (is \?P \ ?Q\) - if \n < LENGTH('a)\ - for p :: "'a :: len word" -proof - assume ?Q - then show ?P - by (meson bang_is_le le_less_trans not_le word_power_increasing) -next - assume ?P - have \take_bit n p = p\ - proof (rule bit_word_eqI) - fix q - assume \q < LENGTH('a)\ - show \bit (take_bit n p) q \ bit p q\ - proof (cases \q < n\) - case True - then show ?thesis - by (auto simp add: bit_simps) - next - case False - then have \n \ q\ - by simp - with \?P\ \q < LENGTH('a)\ have \\ bit p q\ - by (simp add: test_bit_eq_bit) - then show ?thesis - by (simp add: bit_simps) - qed - qed - with that show ?Q - using take_bit_word_eq_self_iff [of n p] by auto -qed - -lemma less_2p_is_upper_bits_unset: - "p < 2 ^ n \ n < LENGTH('a) \ (\n' \ n. n' < LENGTH('a) \ \ p !! n')" for p :: "'a :: len word" - by (meson le_less_trans le_mask_iff_lt_2n upper_bits_unset_is_l2p word_zero_le) - -lemma test_bit_over: - "n \ size (x::'a::len word) \ (x !! n) = False" - by transfer auto - -lemma le_mask_high_bits: - "w \ mask n \ (\i \ {n ..< size w}. \ w !! i)" - for w :: \'a::len word\ - by (auto simp: word_size and_mask_eq_iff_le_mask[symmetric] word_eq_iff) - -lemma test_bit_conj_lt: - "(x !! m \ m < LENGTH('a)) = x !! m" for x :: "'a :: len word" - using test_bit_bin by blast - -lemma neg_test_bit: - "(NOT x) !! n = (\ x !! n \ n < LENGTH('a))" for x :: "'a::len word" - by (cases "n < LENGTH('a)") (auto simp add: test_bit_over word_ops_nth_size word_size) - -lemma shiftr_less_t2n': - "\ x AND mask (n + m) = x; m < LENGTH('a) \ \ x >> n < 2 ^ m" for x :: "'a :: len word" - apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) - apply transfer - apply (simp add: take_bit_drop_bit ac_simps) - done - -lemma shiftr_less_t2n: - "x < 2 ^ (n + m) \ x >> n < 2 ^ m" for x :: "'a :: len word" - apply (rule shiftr_less_t2n') - apply (erule less_mask_eq) - apply (rule ccontr) - apply (simp add: not_less) - apply (subst (asm) p2_eq_0[symmetric]) - apply (simp add: power_add) - done - -lemma shiftr_eq_0: - "n \ LENGTH('a) \ ((w::'a::len word) >> n) = 0" - apply (cut_tac shiftr_less_t2n'[of w n 0], simp) - apply (simp add: mask_eq_iff) - apply (simp add: lt2p_lem) - apply simp - done - -lemma shiftr_not_mask_0: - "n+m \ LENGTH('a :: len) \ ((w::'a::len word) >> n) AND NOT (mask m) = 0" - by (rule bit_word_eqI) (auto simp add: bit_simps dest: bit_imp_le_length) - -lemma shiftl_less_t2n: - fixes x :: "'a :: len word" - shows "\ x < (2 ^ (m - n)); m < LENGTH('a) \ \ (x << n) < 2 ^ m" - apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) - apply transfer - apply (simp add: take_bit_push_bit) - done - -lemma shiftl_less_t2n': - "(x::'a::len word) < 2 ^ m \ m+n < LENGTH('a) \ x << n < 2 ^ (m + n)" - by (rule shiftl_less_t2n) simp_all - -lemma nth_w2p_scast [simp]: - "((scast ((2::'a::len signed word) ^ n) :: 'a word) !! m) - \ ((((2::'a::len word) ^ n) :: 'a word) !! m)" - by transfer (auto simp add: bit_simps) - -lemma scast_bit_test [simp]: - "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" - by (clarsimp simp: word_eq_iff) - -lemma signed_shift_guard_to_word: - "\ n < len_of TYPE ('a); n > 0 \ - \ (unat (x :: 'a :: len word) * 2 ^ y < 2 ^ n) - = (x = 0 \ x < (1 << n >> y))" - apply (simp only: nat_mult_power_less_eq) - apply (cases "y \ n") - apply (simp only: shiftl_shiftr1) - apply (subst less_mask_eq) - apply (simp add: word_less_nat_alt word_size) - apply (rule order_less_le_trans[rotated], rule power_increasing[where n=1]) - apply simp - apply simp - apply simp - apply (simp add: nat_mult_power_less_eq word_less_nat_alt word_size) - apply auto[1] - apply (simp only: shiftl_shiftr2, simp add: unat_eq_0) - done - -lemma nth_bounded: - "\(x :: 'a :: len word) !! n; x < 2 ^ m; m \ len_of TYPE ('a)\ \ n < m" - apply (rule ccontr) - apply (auto simp add: not_less test_bit_word_eq) - apply (meson bit_imp_le_length bit_uint_iff less_2p_is_upper_bits_unset test_bit_bin) - done - -lemma shiftl_mask_is_0[simp]: - "(x << n) AND mask n = 0" - for x :: \'a::len word\ - by (simp flip: take_bit_eq_mask add: shiftl_eq_push_bit take_bit_push_bit) - -lemma rshift_sub_mask_eq: - "(a >> (size a - b)) AND mask b = a >> (size a - b)" - for a :: \'a::len word\ - using shiftl_shiftr2[where a=a and b=0 and c="size a - b"] - apply (cases "b < size a") - apply simp - apply (simp add: linorder_not_less mask_eq_decr_exp word_size - p2_eq_0[THEN iffD2]) - done - -lemma shiftl_shiftr3: - "b \ c \ a << b >> c = (a >> c - b) AND mask (size a - c)" - for a :: \'a::len word\ - apply (cases "b = c") - apply (simp add: shiftl_shiftr1) - apply (simp add: shiftl_shiftr2) - done - -lemma and_mask_shiftr_comm: - "m \ size w \ (w AND mask m) >> n = (w >> n) AND mask (m-n)" - for w :: \'a::len word\ - by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3) - -lemma and_mask_shiftl_comm: - "m+n \ size w \ (w AND mask m) << n = (w << n) AND mask (m+n)" - for w :: \'a::len word\ - by (simp add: and_mask word_size shiftl_shiftl) (simp add: shiftl_shiftr1) - -lemma le_mask_shiftl_le_mask: "s = m + n \ x \ mask n \ x << m \ mask s" - for x :: \'a::len word\ - by (simp add: le_mask_iff shiftl_shiftr3) - -lemma word_and_1_shiftl: - "x AND (1 << n) = (if x !! n then (1 << n) else 0)" for x :: "'a :: len word" - apply (rule bit_word_eqI; transfer) - apply (auto simp add: bit_simps not_le ac_simps) - done - -lemmas word_and_1_shiftls' - = word_and_1_shiftl[where n=0] - word_and_1_shiftl[where n=1] - word_and_1_shiftl[where n=2] - -lemmas word_and_1_shiftls = word_and_1_shiftls' [simplified] - -lemma word_and_mask_shiftl: - "x AND (mask n << m) = ((x >> m) AND mask n) << m" - for x :: \'a::len word\ - apply (rule bit_word_eqI; transfer) - apply (auto simp add: bit_simps not_le ac_simps) - done - -lemma shift_times_fold: - "(x :: 'a :: len word) * (2 ^ n) << m = x << (m + n)" - by (simp add: shiftl_t2n ac_simps power_add) - -lemma of_bool_nth: - "of_bool (x !! v) = (x >> v) AND 1" - for x :: \'a::len word\ - by (simp add: test_bit_word_eq shiftr_word_eq bit_eq_iff) - (auto simp add: bit_1_iff bit_and_iff bit_drop_bit_eq intro: ccontr) - -lemma shiftr_mask_eq: - "(x >> n) AND mask (size x - n) = x >> n" for x :: "'a :: len word" - apply (simp flip: take_bit_eq_mask) - apply transfer - apply (simp add: take_bit_drop_bit) - done - -lemma shiftr_mask_eq': - "m = (size x - n) \ (x >> n) AND mask m = x >> n" for x :: "'a :: len word" - by (simp add: shiftr_mask_eq) - -lemma and_eq_0_is_nth: - fixes x :: "'a :: len word" - shows "y = 1 << n \ ((x AND y) = 0) = (\ (x !! n))" - apply safe - apply (drule_tac u="(x AND (1 << n))" and x=n in word_eqD) - apply (simp add: nth_w2p) - apply (simp add: test_bit_bin) - apply (rule bit_word_eqI) - apply (auto simp add: bit_simps test_bit_eq_bit) - done - -lemma and_neq_0_is_nth: - \x AND y \ 0 \ x !! n\ if \y = 2 ^ n\ for x y :: \'a::len word\ - apply (simp add: bit_eq_iff bit_simps) - using that apply (simp add: bit_simps not_le) - apply transfer - apply auto - done - -lemma nth_is_and_neq_0: - "(x::'a::len word) !! n = (x AND 2 ^ n \ 0)" - by (subst and_neq_0_is_nth; rule refl) - -lemma word_shift_zero: - "\ x << n = 0; x \ 2^m; m + n < LENGTH('a)\ \ (x::'a::len word) = 0" - apply (rule ccontr) - apply (drule (2) word_shift_nonzero) - apply simp - done - -lemma mask_shift_and_negate[simp]:"(w AND mask n << m) AND NOT (mask n << m) = 0" - for w :: \'a::len word\ - by (clarsimp simp add: mask_eq_decr_exp Parity.bit_eq_iff bit_and_iff bit_not_iff shiftl_word_eq bit_push_bit_iff) - -end diff --git a/lib/Word_Lib/Type_Syntax.thy b/lib/Word_Lib/Type_Syntax.thy index 6f1c8014a..e99104d6e 100644 --- a/lib/Word_Lib/Type_Syntax.thy +++ b/lib/Word_Lib/Type_Syntax.thy @@ -18,7 +18,7 @@ struct fun tr' cnst ctxt typ ts = if Config.get ctxt show_word_types then case (Term.binder_types typ, Term.body_type typ) of - ([Type (@{type_name "word"}, [S])], Type (@{type_name "word"}, [T])) => + ([\<^Type>\word S\], \<^Type>\word T\) => list_comb (Syntax.const cnst $ Syntax_Phases.term_of_typ ctxt S $ Syntax_Phases.term_of_typ ctxt T , ts) diff --git a/lib/Word_Lib/Typedef_Morphisms.thy b/lib/Word_Lib/Typedef_Morphisms.thy index d2f9cfba8..b6e069ab9 100644 --- a/lib/Word_Lib/Typedef_Morphisms.thy +++ b/lib/Word_Lib/Typedef_Morphisms.thy @@ -292,7 +292,7 @@ lemma unat_le: "y \ unat z \ y \ unats (LENGTH('a))" for z :: "'a::len word" apply (unfold unats_def) apply clarsimp - apply (rule xtrans, rule unat_lt2p, assumption) + apply (metis le_unat_uoi unsigned_less) done lemma td_ext_sbin: @@ -346,11 +346,11 @@ lemmas uint_mod_alt = word_mod_def [THEN trans [OF uint_cong int_word_uint]] interpretation test_bit: td_ext - "(!!) :: 'a::len word \ nat \ bool" + "bit :: 'a::len word \ nat \ bool" set_bits "{f. \i. f i \ i < LENGTH('a::len)}" "(\h i. h i \ i < LENGTH('a::len))" - by standard (auto simp add: test_bit_word_eq bit_imp_le_length bit_set_bits_word_iff set_bits_bit_eq) + by standard (auto simp add: bit_imp_le_length bit_set_bits_word_iff set_bits_bit_eq) lemmas td_nth = test_bit.td_thm @@ -365,4 +365,36 @@ lemma sints_subset: apply simp done +lemma uints_mono_iff: "uints l \ uints m \ l \ m" + using power_increasing_iff[of "2::int" l m] + apply (auto simp: uints_num subset_iff simp del: power_increasing_iff) + apply (meson less_irrefl not_le zero_le_numeral zero_le_power) + done + +lemmas uints_monoI = uints_mono_iff[THEN iffD2] + +lemma Bit_in_uints_Suc: "of_bool c + 2 * w \ uints (Suc m)" if "w \ uints m" + using that + by (auto simp: uints_num) + +lemma Bit_in_uintsI: "of_bool c + 2 * w \ uints m" if "w \ uints (m - 1)" "m > 0" + using Bit_in_uints_Suc[OF that(1)] that(2) + by auto + +lemma bin_cat_in_uintsI: + \concat_bit n b a \ uints m\ if \a \ uints l\ \m \ l + n\ +proof - + from \m \ l + n\ obtain q where \m = l + n + q\ + using le_Suc_ex by blast + then have \(2::int) ^ m = 2 ^ n * 2 ^ (l + q)\ + by (simp add: ac_simps power_add) + moreover have \a mod 2 ^ (l + q) = a\ + using \a \ uints l\ + by (auto simp add: uints_def take_bit_eq_mod power_add Divides.mod_mult2_eq) + ultimately have \concat_bit n b a = take_bit m (concat_bit n b a)\ + by (simp add: concat_bit_eq take_bit_eq_mod push_bit_eq_mult Divides.mod_mult2_eq) + then show ?thesis + by (simp add: uints_def) +qed + end diff --git a/lib/Word_Lib/Word_16.thy b/lib/Word_Lib/Word_16.thy index a91c5c424..58366290b 100644 --- a/lib/Word_Lib/Word_16.thy +++ b/lib/Word_Lib/Word_16.thy @@ -14,9 +14,15 @@ begin lemma len16: "len_of (x :: 16 itself) = 16" by simp +context + includes bit_operations_syntax +begin + lemma word16_and_max_simp: \x AND 0xFFFF = x\ for x :: \16 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) end + +end diff --git a/lib/Word_Lib/Word_32.thy b/lib/Word_Lib/Word_32.thy index d2d4651e9..5d187cd9d 100644 --- a/lib/Word_Lib/Word_32.thy +++ b/lib/Word_Lib/Word_32.thy @@ -16,207 +16,56 @@ theory Word_32 Bitwise begin +context + includes bit_operations_syntax +begin + type_synonym word32 = "32 word" lemma len32: "len_of (x :: 32 itself) = 32" by simp type_synonym sword32 = "32 sword" -type_synonym machine_word_len = 32 -type_synonym machine_word = "machine_word_len word" - -definition word_bits :: nat -where - "word_bits = LENGTH(machine_word_len)" - -text \The following two are numerals so they can be used as nats and words.\ -definition word_size_bits :: "'a :: numeral" -where - "word_size_bits = 2" - -definition word_size :: "'a :: numeral" -where - "word_size = 4" - -lemma word_bits_conv[code]: - "word_bits = 32" - unfolding word_bits_def by simp - -lemma word_size_word_size_bits: - "(word_size::nat) = 2 ^ word_size_bits" - unfolding word_size_def word_size_bits_def by simp - -lemma word_bits_word_size_conv: - "word_bits = word_size * 8" - unfolding word_bits_def word_size_def by simp - lemma ucast_8_32_inj: "inj (ucast :: 8 word \ 32 word)" by (rule down_ucast_inj) (clarsimp simp: is_down_def target_size source_size) -lemma upto_2_helper: - "{0..<2 :: 32 word} = {0, 1}" - by (safe; simp) unat_arith - -lemmas upper_bits_unset_is_l2p_32 = upper_bits_unset_is_l2p [where 'a=32, folded word_bits_def] -lemmas le_2p_upper_bits_32 = le_2p_upper_bits [where 'a=32, folded word_bits_def] -lemmas le2p_bits_unset_32 = le2p_bits_unset[where 'a=32, folded word_bits_def] - -lemma word_bits_len_of: - "len_of TYPE (32) = word_bits" - by (simp add: word_bits_conv) - lemmas unat_power_lower32' = unat_power_lower[where 'a=32] -lemmas unat_power_lower32 [simp] = unat_power_lower32'[unfolded word_bits_len_of] lemmas word32_less_sub_le' = word_less_sub_le[where 'a = 32] -lemmas word32_less_sub_le[simp] = word32_less_sub_le' [folded word_bits_def] - -lemma word_bits_size: - "size (w::word32) = word_bits" - by (simp add: word_bits_def word_size) lemmas word32_power_less_1' = word_power_less_1[where 'a = 32] -lemmas word32_power_less_1[simp] = word32_power_less_1'[folded word_bits_def] - -lemma of_nat32_0: - "\of_nat n = (0::word32); n < 2 ^ word_bits\ \ n = 0" - by (erule of_nat_0, simp add: word_bits_def) - -lemma unat_mask_2_less_4: - "unat (p && mask 2 :: word32) < 4" - apply (rule unat_less_helper) - apply (rule order_le_less_trans, rule word_and_le1) - apply (simp add: mask_eq) - done lemmas unat_of_nat32' = unat_of_nat_eq[where 'a=32] -lemmas unat_of_nat32 = unat_of_nat32'[unfolded word_bits_len_of] - -lemmas word_power_nonzero_32 = word_power_nonzero [where 'a=32, folded word_bits_def] - -lemmas unat_mult_simple = iffD1 [OF unat_mult_lem [where 'a = 32, unfolded word_bits_len_of]] - -lemmas div_power_helper_32 = div_power_helper [where 'a=32, folded word_bits_def] - -lemma n_less_word_bits: - "(n < word_bits) = (n < 32)" - by (simp add: word_bits_def) - -lemmas of_nat_less_pow_32 = of_nat_power [where 'a=32, folded word_bits_def] - -lemma lt_word_bits_lt_pow: - "sz < word_bits \ sz < 2 ^ word_bits" - by (simp add: word_bits_conv) - -lemma unat_less_word_bits: - fixes y :: word32 - shows "x < unat y \ x < 2 ^ word_bits" - unfolding word_bits_def - by (rule order_less_trans [OF _ unat_lt2p]) lemmas unat_mask_word32' = unat_mask[where 'a=32] -lemmas unat_mask_word32 = unat_mask_word32'[folded word_bits_def] - -lemma unat_less_2p_word_bits: - "unat (x :: 32 word) < 2 ^ word_bits" - apply (simp only: word_bits_def) - apply (rule unat_lt2p) - done - -lemma Suc_unat_mask_div: - "Suc (unat (mask sz div word_size::word32)) = 2 ^ (min sz word_bits - 2)" - apply (case_tac "sz < word_bits") - apply (case_tac "2 \ sz") - apply (clarsimp simp: word_size_def word_bits_def min_def mask_eq) - apply (drule (2) Suc_div_unat_helper - [where 'a=32 and sz=sz and us=2, simplified, symmetric]) - apply (simp add: not_le word_size_def word_bits_def) - apply (case_tac sz, simp add: unat_word_ariths) - apply (case_tac nat, simp add: unat_word_ariths - unat_mask_word32 min_def word_bits_def) - apply simp - apply (simp add: unat_word_ariths - unat_mask_word32 min_def word_bits_def word_size_def) - done lemmas word32_minus_one_le' = word_minus_one_le[where 'a=32] lemmas word32_minus_one_le = word32_minus_one_le'[simplified] -lemma ucast_not_helper: - fixes a::"8 word" - assumes a: "a \ 0xFF" - shows "ucast a \ (0xFF::word32)" -proof - assume "ucast a = (0xFF::word32)" - also - have "(0xFF::word32) = ucast (0xFF::8 word)" by simp - finally - show False using a - apply - - apply (drule up_ucast_inj, simp) - apply simp - done -qed - -lemma less_4_cases: - "(x::word32) < 4 \ x=0 \ x=1 \ x=2 \ x=3" - apply clarsimp - apply (drule word_less_cases, erule disjE, simp, simp)+ - done - lemma unat_ucast_8_32: fixes x :: "8 word" shows "unat (ucast x :: word32) = unat x" by transfer simp -lemma if_then_1_else_0: - "((if P then 1 else 0) = (0 :: word32)) = (\ P)" - by simp - -lemma if_then_0_else_1: - "((if P then 0 else 1) = (0 :: word32)) = (P)" - by simp - -lemmas if_then_simps = if_then_0_else_1 if_then_1_else_0 - lemma ucast_le_ucast_8_32: "(ucast x \ (ucast y :: word32)) = (x \ (y :: 8 word))" by (simp add: ucast_le_ucast) -lemma in_16_range: - "0 \ S \ r \ (\x. r + x * (16 :: word32)) ` S" - "n - 1 \ S \ (r + (16 * n - 16)) \ (\x :: word32. r + x * 16) ` S" - by (clarsimp simp: image_def elim!: bexI[rotated])+ - lemma eq_2_32_0: "(2 ^ 32 :: word32) = 0" by simp -lemma x_less_2_0_1: - fixes x :: word32 shows - "x < 2 \ x = 0 \ x = 1" - by (rule x_less_2_0_1') auto - lemmas mask_32_max_word = max_word_mask [symmetric, where 'a=32, simplified] lemma of_nat32_n_less_equal_power_2: "n < 32 \ ((of_nat n)::32 word) < 2 ^ n" by (rule of_nat_n_less_equal_power_2, clarsimp simp: word_size) -lemma word_rsplit_0: - "word_rsplit (0 :: word32) = [0, 0, 0, 0 :: 8 word]" - by (simp add: word_rsplit_def bin_rsplit_def) - lemma unat_ucast_10_32 : fixes x :: "10 word" shows "unat (ucast x :: word32) = unat x" by transfer simp -lemma bool_mask [simp]: - fixes x :: word32 - shows "(0 < x && 1) = (x && 1 = 1)" - by (rule bool_mask') auto - lemma word32_bounds: "- (2 ^ (size (x :: word32) - 1)) = (-2147483648 :: int)" "((2 ^ (size (x :: word32) - 1)) - 1) = (2147483647 :: int)" @@ -224,9 +73,6 @@ lemma word32_bounds: "((2 ^ (size (y :: 32 signed word) - 1)) - 1) = (2147483647 :: int)" by (simp_all add: word_size) -lemma word_ge_min:"sint (x::32 word) \ -2147483648" - by (metis sint_ge word32_bounds(1) word_size) - lemmas signed_arith_ineq_checks_to_eq_word32' = signed_arith_ineq_checks_to_eq[where 'a=32] signed_arith_ineq_checks_to_eq[where 'a="32 signed"] @@ -256,7 +102,7 @@ lemma ucast_of_nats [simp]: "(ucast (of_nat x :: 16 word) :: 16 sword) = (of_nat x)" "(ucast (of_nat x :: 16 word) :: 8 sword) = (of_nat x)" "(ucast (of_nat x :: 8 word) :: 8 sword) = (of_nat x)" - by (simp_all add: of_nat_take_bit take_bit_word_eq_self) + by (simp_all add: of_nat_take_bit take_bit_word_eq_self unsigned_of_nat) lemmas signed_shift_guard_simpler_32' = power_strict_increasing_iff[where b="2 :: nat" and y=31] @@ -271,15 +117,6 @@ lemmas signed_shift_guard_to_word_32 = signed_shift_guard_to_word[OF word32_31_less(1-2)] signed_shift_guard_to_word[OF word32_31_less(3-4)] -lemma le_step_down_word_3: - fixes x :: "32 word" - shows "\x \ y; x \ y; y < 2 ^ 32 - 1\ \ x \ y - 1" - by (rule le_step_down_word_2, assumption+) - -lemma shiftr_1: - "(x::word32) >> 1 = 0 \ x < 2" - by transfer (simp add: take_bit_drop_bit drop_bit_Suc) - lemma has_zero_byte: "~~ (((((v::word32) && 0x7f7f7f7f) + 0x7f7f7f7f) || v) || 0x7f7f7f7f) \ 0 \ v && 0xff000000 = 0 \ v && 0xff0000 = 0 \ v && 0xff00 = 0 \ v && 0xff = 0" @@ -309,10 +146,7 @@ qed lemma unat_of_int_32: "\i \ 0; i \2 ^ 31\ \ (unat ((of_int i)::sword32)) = nat i" - unfolding unat_eq_nat_uint - apply (subst eq_nat_nat_iff) - apply (auto simp add: take_bit_int_eq_self) - done + by (simp add: unsigned_of_int nat_take_bit_eq take_bit_nat_eq_self) lemmas word_ctz_not_minus_1_32 = word_ctz_not_minus_1[where 'a=32, simplified] @@ -339,3 +173,5 @@ lemma word32_and_max_simp: by (simp add: numeral_eq_Suc mask_Suc_exp) end + +end diff --git a/lib/Word_Lib/Word_64.thy b/lib/Word_Lib/Word_64.thy index fc8340aa0..9c63a998c 100644 --- a/lib/Word_Lib/Word_64.thy +++ b/lib/Word_Lib/Word_64.thy @@ -15,200 +15,54 @@ theory Word_64 More_Word_Operations begin +context + includes bit_operations_syntax +begin + lemma len64: "len_of (x :: 64 itself) = 64" by simp -type_synonym machine_word_len = 64 -type_synonym machine_word = "machine_word_len word" - -definition word_bits :: nat -where - "word_bits = LENGTH(machine_word_len)" - -text \The following two are numerals so they can be used as nats and words.\ -definition word_size_bits :: "'a :: numeral" -where - "word_size_bits = 3" - -definition word_size :: "'a :: numeral" -where - "word_size = 8" - -lemma word_bits_conv[code]: - "word_bits = 64" - unfolding word_bits_def by simp - -lemma word_size_word_size_bits: - "(word_size::nat) = 2 ^ word_size_bits" - unfolding word_size_def word_size_bits_def by simp - -lemma word_bits_word_size_conv: - "word_bits = word_size * 8" - unfolding word_bits_def word_size_def by simp - lemma ucast_8_64_inj: "inj (ucast :: 8 word \ 64 word)" by (rule down_ucast_inj) (clarsimp simp: is_down_def target_size source_size) -lemma upto_2_helper: - "{0..<2 :: 64 word} = {0, 1}" - by (safe; simp) unat_arith - -lemmas upper_bits_unset_is_l2p_64 = upper_bits_unset_is_l2p [where 'a=64, folded word_bits_def] -lemmas le_2p_upper_bits_64 = le_2p_upper_bits [where 'a=64, folded word_bits_def] -lemmas le2p_bits_unset_64 = le2p_bits_unset[where 'a=64, folded word_bits_def] - -lemma word_bits_len_of: - "len_of TYPE (64) = word_bits" - by (simp add: word_bits_conv) - lemmas unat_power_lower64' = unat_power_lower[where 'a=64] -lemmas unat_power_lower64 [simp] = unat_power_lower64'[unfolded word_bits_len_of] lemmas word64_less_sub_le' = word_less_sub_le[where 'a = 64] -lemmas word64_less_sub_le[simp] = word64_less_sub_le' [folded word_bits_def] - -lemma word_bits_size: - "size (w::word64) = word_bits" - by (simp add: word_bits_def word_size) lemmas word64_power_less_1' = word_power_less_1[where 'a = 64] -lemmas word64_power_less_1[simp] = word64_power_less_1'[folded word_bits_def] - -lemma of_nat64_0: - "\of_nat n = (0::word64); n < 2 ^ word_bits\ \ n = 0" - by (erule of_nat_0, simp add: word_bits_def) - -lemma unat_mask_2_less_4: - "unat (p && mask 2 :: word64) < 4" - apply (rule unat_less_helper) - apply (rule order_le_less_trans, rule word_and_le1) - apply (simp add: mask_eq) - done lemmas unat_of_nat64' = unat_of_nat_eq[where 'a=64] -lemmas unat_of_nat64 = unat_of_nat64'[unfolded word_bits_len_of] - -lemmas word_power_nonzero_64 = word_power_nonzero [where 'a=64, folded word_bits_def] - -lemmas unat_mult_simple = iffD1 [OF unat_mult_lem [where 'a = 64, unfolded word_bits_len_of]] - -lemmas div_power_helper_64 = div_power_helper [where 'a=64, folded word_bits_def] - -lemma n_less_word_bits: - "(n < word_bits) = (n < 64)" - by (simp add: word_bits_def) - -lemmas of_nat_less_pow_64 = of_nat_power [where 'a=64, folded word_bits_def] - -lemma lt_word_bits_lt_pow: - "sz < word_bits \ sz < 2 ^ word_bits" - by (simp add: word_bits_conv) - -lemma unat_less_word_bits: - fixes y :: word64 - shows "x < unat y \ x < 2 ^ word_bits" - unfolding word_bits_def - by (rule order_less_trans [OF _ unat_lt2p]) lemmas unat_mask_word64' = unat_mask[where 'a=64] -lemmas unat_mask_word64 = unat_mask_word64'[folded word_bits_def] - -lemma unat_less_2p_word_bits: - "unat (x :: 64 word) < 2 ^ word_bits" - apply (simp only: word_bits_def) - apply (rule unat_lt2p) - done - -lemma Suc_unat_mask_div: - "Suc (unat (mask sz div word_size::word64)) = 2 ^ (min sz word_bits - 3)" - apply (case_tac "sz < word_bits") - apply (case_tac "3 \ sz") - apply (clarsimp simp: word_size_def word_bits_def min_def mask_eq) - apply (drule (2) Suc_div_unat_helper - [where 'a=64 and sz=sz and us=3, simplified, symmetric]) - apply (simp add: not_le word_size_def word_bits_def) - apply (case_tac sz, simp add: unat_word_ariths) - apply (case_tac nat, simp add: unat_word_ariths - unat_mask_word64 min_def word_bits_def) - apply (case_tac nata, simp add: unat_word_ariths unat_mask_word64 word_bits_def) - apply simp - apply (simp add: unat_word_ariths - unat_mask_word64 min_def word_bits_def word_size_def) - done lemmas word64_minus_one_le' = word_minus_one_le[where 'a=64] lemmas word64_minus_one_le = word64_minus_one_le'[simplified] -lemma ucast_not_helper: - fixes a::"8 word" - assumes a: "a \ 0xFF" - shows "ucast a \ (0xFF::word64)" -proof - assume "ucast a = (0xFF::word64)" - also - have "(0xFF::word64) = ucast (0xFF::8 word)" by simp - finally - show False using a - apply - - apply (drule up_ucast_inj, simp) - apply simp - done -qed - lemma less_4_cases: "(x::word64) < 4 \ x=0 \ x=1 \ x=2 \ x=3" apply clarsimp apply (drule word_less_cases, erule disjE, simp, simp)+ done -lemma if_then_1_else_0: - "((if P then 1 else 0) = (0 :: word64)) = (\ P)" - by simp - -lemma if_then_0_else_1: - "((if P then 0 else 1) = (0 :: word64)) = (P)" - by simp - -lemmas if_then_simps = if_then_0_else_1 if_then_1_else_0 - lemma ucast_le_ucast_8_64: "(ucast x \ (ucast y :: word64)) = (x \ (y :: 8 word))" by (simp add: ucast_le_ucast) -lemma in_16_range: - "0 \ S \ r \ (\x. r + x * (16 :: word64)) ` S" - "n - 1 \ S \ (r + (16 * n - 16)) \ (\x :: word64. r + x * 16) ` S" - by (clarsimp simp: image_def elim!: bexI[rotated])+ - lemma eq_2_64_0: "(2 ^ 64 :: word64) = 0" by simp -lemma x_less_2_0_1: - fixes x :: word64 shows - "x < 2 \ x = 0 \ x = 1" - by (rule x_less_2_0_1') auto - lemmas mask_64_max_word = max_word_mask [symmetric, where 'a=64, simplified] lemma of_nat64_n_less_equal_power_2: "n < 64 \ ((of_nat n)::64 word) < 2 ^ n" by (rule of_nat_n_less_equal_power_2, clarsimp simp: word_size) -lemma word_rsplit_0: - "word_rsplit (0 :: word64) = [0, 0, 0, 0, 0, 0, 0, 0 :: 8 word]" - by (simp add: word_rsplit_def bin_rsplit_def) - lemma unat_ucast_10_64 : fixes x :: "10 word" shows "unat (ucast x :: word64) = unat x" by transfer simp -lemma bool_mask [simp]: - fixes x :: word64 - shows "(0 < x && 1) = (x && 1 = 1)" - by (rule bool_mask') auto - lemma word64_bounds: "- (2 ^ (size (x :: word64) - 1)) = (-9223372036854775808 :: int)" "((2 ^ (size (x :: word64) - 1)) - 1) = (9223372036854775807 :: int)" @@ -216,9 +70,6 @@ lemma word64_bounds: "((2 ^ (size (y :: 64 signed word) - 1)) - 1) = (9223372036854775807 :: int)" by (simp_all add: word_size) -lemma word_ge_min:"sint (x::64 word) \ -9223372036854775808" - by (metis sint_ge word64_bounds(1) word_size) - lemmas signed_arith_ineq_checks_to_eq_word64' = signed_arith_ineq_checks_to_eq[where 'a=64] signed_arith_ineq_checks_to_eq[where 'a="64 signed"] @@ -245,7 +96,7 @@ lemma ucast_of_nats [simp]: "(ucast (of_nat x :: word64) :: sword64) = (of_nat x)" "(ucast (of_nat x :: word64) :: 16 sword) = (of_nat x)" "(ucast (of_nat x :: word64) :: 8 sword) = (of_nat x)" - by (simp_all add: of_nat_take_bit take_bit_word_eq_self) + by (simp_all add: of_nat_take_bit take_bit_word_eq_self unsigned_of_nat) lemmas signed_shift_guard_simpler_64' = power_strict_increasing_iff[where b="2 :: nat" and y=31] @@ -260,15 +111,6 @@ lemmas signed_shift_guard_to_word_64 = signed_shift_guard_to_word[OF word64_31_less(1-2)] signed_shift_guard_to_word[OF word64_31_less(3-4)] -lemma le_step_down_word_3: - fixes x :: "64 word" - shows "\x \ y; x \ y; y < 2 ^ 64 - 1\ \ x \ y - 1" - by (rule le_step_down_word_2, assumption+) - -lemma shiftr_1: - "(x::word64) >> 1 = 0 \ x < 2" - by transfer (simp add: take_bit_drop_bit drop_bit_Suc) - lemma mask_step_down_64: \\x. mask x = b\ if \b && 1 = 1\ and \\x. x < 64 \ mask x = b >> 1\ for b :: \64word\ @@ -293,10 +135,7 @@ qed lemma unat_of_int_64: "\i \ 0; i \ 2 ^ 63\ \ (unat ((of_int i)::sword64)) = nat i" - unfolding unat_eq_nat_uint - apply (subst eq_nat_nat_iff) - apply (simp_all add: take_bit_int_eq_self) - done + by (simp add: unsigned_of_int nat_take_bit_eq take_bit_nat_eq_self) lemmas word_ctz_not_minus_1_64 = word_ctz_not_minus_1[where 'a=64, simplified] @@ -306,3 +145,5 @@ lemma word64_and_max_simp: by (simp add: numeral_eq_Suc mask_Suc_exp) end + +end diff --git a/lib/Word_Lib/Word_8.thy b/lib/Word_Lib/Word_8.thy index c9ac1e6b2..09b57ae51 100644 --- a/lib/Word_Lib/Word_8.thy +++ b/lib/Word_Lib/Word_8.thy @@ -15,6 +15,10 @@ imports Word_Lemmas begin +context + includes bit_operations_syntax +begin + lemma len8: "len_of (x :: 8 itself) = 8" by simp lemma word8_and_max_simp: @@ -42,7 +46,7 @@ lemma enum_word8_eq: 244, 245, 246, 247, 248, 249, 250, 251, 252, 253, 254, 255]\ (is \?lhs = ?rhs\) proof - have \map unat ?lhs = [0..<256]\ - by (simp add: enum_word_def comp_def take_bit_nat_eq_self map_idem_upt_eq) + by (simp add: enum_word_def comp_def take_bit_nat_eq_self map_idem_upt_eq unsigned_of_nat) also have \\ = map unat ?rhs\ by (simp add: upt_zero_numeral_unfold) finally show ?thesis @@ -106,3 +110,5 @@ lemma word8_exhaust: done end + +end diff --git a/lib/Word_Lib/Word_EqI.thy b/lib/Word_Lib/Word_EqI.thy index 573c9b65b..d0c8cac3a 100644 --- a/lib/Word_Lib/Word_EqI.thy +++ b/lib/Word_Lib/Word_EqI.thy @@ -9,7 +9,7 @@ section "Solving Word Equalities" theory Word_EqI imports More_Word - Traditional_Infix_Syntax + Aligned "HOL-Eisbach.Eisbach_Tools" begin @@ -23,23 +23,27 @@ named_theorems word_eqI_simps lemmas [word_eqI_simps] = word_ops_nth_size + word_ao_nth + bit_mask_iff word_size word_or_zero neg_mask_test_bit nth_ucast - nth_w2p nth_shiftl - nth_shiftr + nth_w2p bit_push_bit_iff + bit_drop_bit_eq less_2p_is_upper_bits_unset le_mask_high_bits bang_eq neg_test_bit is_up is_down + is_aligned_nth + neg_mask_le_high_bits lemmas word_eqI_rule = word_eqI [rule_format] lemma test_bit_lenD: - "x !! n \ n < LENGTH('a) \ x !! n" for x :: "'a :: len word" + "bit x n \ n < LENGTH('a) \ bit x n" for x :: "'a :: len word" by (fastforce dest: test_bit_size simp: word_size) method word_eqI uses simp simp_del split split_del cong flip = diff --git a/lib/Word_Lib/Word_Lemmas.thy b/lib/Word_Lib/Word_Lemmas.thy index 81d225265..78efa5e96 100644 --- a/lib/Word_Lib/Word_Lemmas.thy +++ b/lib/Word_Lib/Word_Lemmas.thy @@ -15,268 +15,47 @@ theory Word_Lemmas Most_significant_bit Enumeration_Word Aligned + Bit_Shifts_Infix_Syntax + Word_EqI begin -(* The seL4 bitfield generator produces functions containing mask and shift operations, such that - * invoking two of them consecutively can produce something like the following. - *) -lemma bitfield_op_twice: - "(x AND NOT (mask n << m) OR ((y AND mask n) << m)) AND NOT (mask n << m) = x AND NOT (mask n << m)" - for x :: \'a::len word\ - by (induct n arbitrary: m) (auto simp: word_ao_dist) +context + includes bit_operations_syntax +begin -lemma bitfield_op_twice'': - "\NOT a = b << c; \x. b = mask x\ \ (x AND a OR (y AND b << c)) AND a = x AND a" - for a b :: \'a::len word\ - apply clarsimp - apply (cut_tac n=xa and m=c and x=x and y=y in bitfield_op_twice) - apply (clarsimp simp:mask_eq_decr_exp) - apply (drule not_switch) - apply clarsimp - done - -lemma bit_twiddle_min: - "(y::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = min x y" - by (auto simp add: Parity.bit_eq_iff bit_xor_iff min_def) - -lemma bit_twiddle_max: - "(x::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = max x y" - by (auto simp add: Parity.bit_eq_iff bit_xor_iff max_def) - -lemma swap_with_xor: - "\(x::'a::len word) = a XOR b; y = b XOR x; z = x XOR y\ \ z = b \ y = a" - by (auto simp add: Parity.bit_eq_iff bit_xor_iff max_def) - -lemma scast_nop1 [simp]: - "((scast ((of_int x)::('a::len) word))::'a sword) = of_int x" - apply (simp only: scast_eq) - by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) - -lemma scast_nop2 [simp]: - "((scast ((of_int x)::('a::len) sword))::'a word) = of_int x" - apply (simp only: scast_eq) - by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) - -lemmas scast_nop = scast_nop1 scast_nop2 scast_id - -lemma le_mask_imp_and_mask: - "(x::'a::len word) \ mask n \ x AND mask n = x" - by (metis and_mask_eq_iff_le_mask) - -lemma or_not_mask_nop: - "((x::'a::len word) OR NOT (mask n)) AND mask n = x AND mask n" - by (metis word_and_not word_ao_dist2 word_bw_comms(1) word_log_esimps(3)) - -lemma mask_subsume: - "\n \ m\ \ ((x::'a::len word) OR y AND mask n) AND NOT (mask m) = x AND NOT (mask m)" - by (auto simp add: Parity.bit_eq_iff bit_not_iff bit_or_iff bit_and_iff bit_mask_iff) - -lemma and_mask_0_iff_le_mask: - fixes w :: "'a::len word" - shows "(w AND NOT(mask n) = 0) = (w \ mask n)" - by (simp add: mask_eq_0_eq_x le_mask_imp_and_mask and_mask_eq_iff_le_mask) - -lemma mask_twice2: - "n \ m \ ((x::'a::len word) AND mask m) AND mask n = x AND mask n" - by (metis mask_twice min_def) - -lemma uint_2_id: - "LENGTH('a) \ 2 \ uint (2::('a::len) word) = 2" - by simp - -lemma bintrunc_id: - "\m \ of_nat n; 0 < m\ \ bintrunc n m = m" - by (simp add: bintrunc_mod2p le_less_trans) - -lemma shiftr1_unfold: "shiftr1 x = x >> 1" - by (metis One_nat_def comp_apply funpow.simps(1) funpow.simps(2) id_apply shiftr_def) - -lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2" - by transfer (simp add: drop_bit_Suc) - -lemma shiftl1_is_mult: "(x << 1) = (x :: 'a::len word) * 2" - by (metis One_nat_def mult_2 mult_2_right one_add_one - power_0 power_Suc shiftl_t2n) - -lemma div_of_0_id[simp]:"(0::('a::len) word) div n = 0" - by (simp add: word_div_def) - -lemma degenerate_word:"LENGTH('a) = 1 \ (x::('a::len) word) = 0 \ x = 1" - by (metis One_nat_def less_irrefl_nat sint_1_cases) - -lemma div_by_0_word:"(x::('a::len) word) div 0 = 0" - by (metis div_0 div_by_0 unat_0 word_arith_nat_defs(6) word_div_1) - -lemma div_less_dividend_word:"\x \ 0; n \ 1\ \ (x::('a::len) word) div n < x" - apply (cases \n = 0\) - apply clarsimp - apply (simp add:word_neq_0_conv) - apply (subst word_arith_nat_div) - apply (rule word_of_nat_less) - apply (rule div_less_dividend) - using unat_eq_zero word_unat_Rep_inject1 apply force - apply (simp add:unat_gt_0) - done - -lemma shiftr1_lt:"x \ 0 \ (x::('a::len) word) >> 1 < x" - apply (subst shiftr1_is_div_2) - apply (rule div_less_dividend_word) - apply simp+ - done - -lemma word_less_div: - fixes x :: "('a::len) word" - and y :: "('a::len) word" - shows "x div y = 0 \ y = 0 \ x < y" - apply (case_tac "y = 0", clarsimp+) - by (metis One_nat_def Suc_le_mono le0 le_div_geq not_less unat_0 unat_div unat_gt_0 word_less_nat_alt zero_less_one) - -lemma not_degenerate_imp_2_neq_0:"LENGTH('a) > 1 \ (2::('a::len) word) \ 0" - by (metis numerals(1) power_not_zero power_zero_numeral) - -lemma shiftr1_0_or_1:"(x::('a::len) word) >> 1 = 0 \ x = 0 \ x = 1" - apply (subst (asm) shiftr1_is_div_2) - apply (drule word_less_div) - apply (case_tac "LENGTH('a) = 1") - apply (simp add:degenerate_word) - apply (erule disjE) - apply (subgoal_tac "(2::'a word) \ 0") - apply simp - apply (rule not_degenerate_imp_2_neq_0) - apply (subgoal_tac "LENGTH('a) \ 0") - apply arith - apply simp - apply (rule x_less_2_0_1', simp+) - done - -lemma word_overflow:"(x::('a::len) word) + 1 > x \ x + 1 = 0" - apply clarsimp - by (metis diff_0 eq_diff_eq less_x_plus_1) - -lemma word_overflow_unat:"unat ((x::('a::len) word) + 1) = unat x + 1 \ x + 1 = 0" - by (metis Suc_eq_plus1 add.commute unatSuc) - -lemma even_word_imp_odd_next:"even (unat (x::('a::len) word)) \ x + 1 = 0 \ odd (unat (x + 1))" - apply (cut_tac x=x in word_overflow_unat) - apply clarsimp - done - -lemma odd_word_imp_even_next:"odd (unat (x::('a::len) word)) \ x + 1 = 0 \ even (unat (x + 1))" - apply (cut_tac x=x in word_overflow_unat) - apply clarsimp - done - -lemma overflow_imp_lsb:"(x::('a::len) word) + 1 = 0 \ x !! 0" - using even_plus_one_iff [of x] by (simp add: test_bit_word_eq) - -lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = x !! 0" - by transfer (simp add: even_nat_iff) - -lemma of_nat_neq_iff_word: - "x mod 2 ^ LENGTH('a) \ y mod 2 ^ LENGTH('a) \ - (((of_nat x)::('a::len) word) \ of_nat y) = (x \ y)" +lemma ucast_le_ucast_eq: + fixes x y :: "'a::len word" + assumes x: "x < 2 ^ n" + assumes y: "y < 2 ^ n" + assumes n: "n = LENGTH('b::len)" + shows "(UCAST('a \ 'b) x \ UCAST('a \ 'b) y) = (x \ y)" apply (rule iffI) - apply (case_tac "x = y") - apply (subst (asm) of_nat_eq_iff[symmetric]) - apply simp+ - apply (case_tac "((of_nat x)::('a::len) word) = of_nat y") - apply (subst (asm) word_unat.norm_eq_iff[symmetric]) - apply simp+ + apply (cases "LENGTH('b) < LENGTH('a)") + apply (subst less_mask_eq[OF x, symmetric]) + apply (subst less_mask_eq[OF y, symmetric]) + apply (unfold n) + apply (subst ucast_ucast_mask[symmetric])+ + apply (simp add: ucast_le_ucast)+ + apply (erule ucast_mono_le[OF _ y[unfolded n]]) done -lemma shiftr1_irrelevant_lsb:"(x::('a::len) word) !! 0 \ x >> 1 = (x + 1) >> 1" - apply (cases \LENGTH('a)\; transfer) - apply (simp_all add: take_bit_drop_bit) - apply (simp add: drop_bit_take_bit drop_bit_Suc) - done - -lemma shiftr1_0_imp_only_lsb:"((x::('a::len) word) + 1) >> 1 = 0 \ x = 0 \ x + 1 = 0" - by (metis One_nat_def shiftr1_0_or_1 word_less_1 word_overflow) - -lemma shiftr1_irrelevant_lsb':"\((x::('a::len) word) !! 0) \ x >> 1 = (x + 1) >> 1" - by (metis shiftr1_irrelevant_lsb) - -lemma lsb_this_or_next:"\(((x::('a::len) word) + 1) !! 0) \ x !! 0" - by (metis (poly_guards_query) even_word_imp_odd_next odd_iff_lsb overflow_imp_lsb) - -(* Perhaps this one should be a simp lemma, but it seems a little dangerous. *) -lemma cast_chunk_assemble_id: - "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ - (((ucast ((ucast (x::'b word))::'a word))::'b word) OR (((ucast ((ucast (x >> n))::'a word))::'b word) << n)) = x" - apply (subgoal_tac "((ucast ((ucast (x >> n))::'a word))::'b word) = x >> n") - apply clarsimp - apply (subst and_not_mask[symmetric]) - apply (subst ucast_ucast_mask) - apply (subst word_ao_dist2[symmetric]) - apply clarsimp - apply (rule ucast_ucast_len) - apply (rule shiftr_less_t2n') - apply (subst and_mask_eq_iff_le_mask) - apply (simp_all add: mask_eq_decr_exp flip: mult_2_right) - apply (metis add_diff_cancel_left' len_gt_0 mult_2_right zero_less_diff) - done - -lemma cast_chunk_scast_assemble_id: - "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ - (((ucast ((scast (x::'b word))::'a word))::'b word) OR - (((ucast ((scast (x >> n))::'a word))::'b word) << n)) = x" - apply (subgoal_tac "((scast x)::'a word) = ((ucast x)::'a word)") - apply (subgoal_tac "((scast (x >> n))::'a word) = ((ucast (x >> n))::'a word)") - apply (simp add:cast_chunk_assemble_id) - apply (subst down_cast_same[symmetric], subst is_down, arith, simp)+ - done - -lemma mask_or_not_mask: - "x AND mask n OR x AND NOT (mask n) = x" - for x :: \'a::len word\ - apply (subst word_oa_dist, simp) - apply (subst word_oa_dist2, simp) - done - -lemma is_aligned_add_not_aligned: - "\is_aligned (p::'a::len word) n; \ is_aligned (q::'a::len word) n\ \ \ is_aligned (p + q) n" - by (metis is_aligned_addD1) - -lemma word_gr0_conv_Suc: "(m::'a::len word) > 0 \ \n. m = n + 1" - by (metis add.commute add_minus_cancel) - -lemma neg_mask_add_aligned: - "\ is_aligned p n; q < 2 ^ n \ \ (p + q) AND NOT (mask n) = p AND NOT (mask n)" - by (metis is_aligned_add_helper is_aligned_neg_mask_eq) - -lemma word_sless_sint_le:"x sint x \ sint y - 1" - by (metis word_sless_alt zle_diff1_eq) - - -lemma upper_trivial: - fixes x :: "'a::len word" - shows "x \ 2 ^ LENGTH('a) - 1 \ x < 2 ^ LENGTH('a) - 1" - by (simp add: less_le) - -lemma constraint_expand: - fixes x :: "'a::len word" - shows "x \ {y. lower \ y \ y \ upper} = (lower \ x \ x \ upper)" - by (rule mem_Collect_eq) - -lemma card_map_elide: - "card ((of_nat :: nat \ 'a::len word) ` {0.. CARD('a::len word)" -proof - - let ?of_nat = "of_nat :: nat \ 'a word" - from word_unat.Abs_inj_on - have "inj_on ?of_nat {i. i < CARD('a word)}" - by (simp add: unats_def card_word) - moreover have "{0.. {i. i < CARD('a word)}" - using that by auto - ultimately have "inj_on ?of_nat {0..is_aligned w n\ if \UCAST('a::len \ 'b::len) w = 0\ \n \ LENGTH('b)\ +proof (rule is_aligned_bitI) + fix q + assume \q < n\ + moreover have \bit (UCAST('a::len \ 'b::len) w) q = bit 0 q\ + using that by simp + with \q < n\ \n \ LENGTH('b)\ show \\ bit w q\ + by (simp add: bit_simps) qed -lemma card_map_elide2: - "n \ CARD('a::len word) \ card ((of_nat::nat \ 'a::len word) ` {0.. 'a::len) w) = unat (w AND mask LENGTH('a))" + apply (simp flip: take_bit_eq_mask) + apply transfer + apply (simp add: ac_simps) + done lemma le_max_word_ucast_id: \UCAST('b \ 'a) (UCAST('a \ 'b) x) = x\ @@ -284,7 +63,7 @@ lemma le_max_word_ucast_id: for x :: \'a::len word\ proof - from that have a1: \x \ word_of_int (uint (word_of_int (2 ^ LENGTH('b) - 1) :: 'b word))\ - by simp + by (simp add: of_int_mask_eq) have f2: "((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (0::int) \ - 1 + 2 ^ LENGTH('b) \ (0::int) \ - 1 + 2 ^ LENGTH('b) + - 1 * 2 ^ LENGTH('b) \ (- (1::int) + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b) = @@ -316,122 +95,100 @@ proof - by (simp add: uint_word_ariths) (metis \1 \ 2 ^ LENGTH('b)\ \uint x \ uint (2 ^ LENGTH('b) - 1)\ linorder_not_less lt2p_lem uint_1 uint_minus_simple_alt uint_power_lower word_le_def zle_diff1_eq) then show ?thesis - apply (simp add: word_ubin.eq_norm bintrunc_mod2p unsigned_ucast_eq) - apply (metis \x \ 2 ^ LENGTH('b) - 1\ and_mask_eq_iff_le_mask mask_eq_decr_exp take_bit_eq_mask) + apply (simp add: unsigned_ucast_eq take_bit_word_eq_self_iff) + apply (meson \x \ 2 ^ LENGTH('b) - 1\ not_le word_less_sub_le) done qed -lemma remdups_enum_upto: - fixes s::"'a::len word" - shows "remdups [s .e. e] = [s .e. e]" - by simp +lemma uint_shiftr_eq: + \uint (w >> n) = uint w div 2 ^ n\ + by (rule bit_eqI) (simp flip: drop_bit_eq_div add: bit_simps) -lemma card_enum_upto: - fixes s::"'a::len word" - shows "card (set [s .e. e]) = Suc (unat e) - unat s" - by (subst List.card_set) (simp add: remdups_enum_upto) +lemma bit_shiftl_word_iff [bit_simps]: + \bit (w << m) n \ m \ n \ n < LENGTH('a) \ bit w (n - m)\ + for w :: \'a::len word\ + by (simp add: bit_simps) -lemma complement_nth_w2p: - shows "n' < LENGTH('a) \ (NOT (2 ^ n :: 'a::len word)) !! n' = (n' \ n)" - by (fastforce simp: word_ops_nth_size word_size nth_w2p) +lemma bit_shiftr_word_iff: + \bit (w >> m) n \ bit w (m + n)\ + for w :: \'a::len word\ + by (simp add: bit_simps) -lemma word_unat_and_lt: - "unat x < n \ unat y < n \ unat (x AND y) < n" - by (meson le_less_trans word_and_le1 word_and_le2 word_le_nat_alt) +lemma uint_sshiftr_eq: + \uint (w >>> n) = take_bit LENGTH('a) (sint w div 2 ^ n)\ + for w :: \'a::len word\ + by (rule bit_eqI) (auto simp add: bit_simps not_less simp flip: drop_bit_eq_div dest: bit_imp_le_length) -lemma word_unat_mask_lt: - "m \ size w \ unat ((w::'a::len word) AND mask m) < 2 ^ m" - by (rule word_unat_and_lt) (simp add: unat_mask word_size) +lemma sshiftr_0: "0 >>> n = 0" + by (simp add: sshiftr_def) -lemma unat_shiftr_less_t2n: - fixes x :: "'a :: len word" - shows "unat x < 2 ^ (n + m) \ unat (x >> n) < 2 ^ m" - by (simp add: shiftr_div_2n' power_add mult.commute less_mult_imp_div_less) +lemma sshiftr_n1: "-1 >>> n = -1" + by (simp add: sshiftr_def) -lemma le_or_mask: - "w \ w' \ w OR mask x \ w' OR mask x" - for w w' :: \'a::len word\ - by (metis neg_mask_add_mask add.commute le_word_or1 mask_2pm1 neg_mask_mono_le word_plus_mono_left) +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\ + by (fact bit_sshiftr_iff) -lemma le_shiftr1': - "\ shiftr1 u \ shiftr1 v ; shiftr1 u \ shiftr1 v \ \ u \ v" - apply transfer - apply simp +lemma nth_sshiftr : + "bit (w >>> m) n = + (n < size w \ (if n + m \ size w then bit w (size w - 1) else bit w (n + m)))" + apply (auto simp add: bit_simps word_size ac_simps not_less) + apply (meson bit_imp_le_length bit_shiftr_word_iff leD) done -lemma le_shiftr': - "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len word) \ v" - apply (induct n; simp add: shiftr_def) - apply (case_tac "(shiftr1 ^^ n) u = (shiftr1 ^^ n) v", simp) - apply (fastforce dest: le_shiftr1') +lemma sshiftr_numeral: + \(numeral k >>> numeral n :: 'a::len word) = + word_of_int (signed_take_bit (LENGTH('a) - 1) (numeral k) >> numeral n)\ + using signed_drop_bit_word_numeral [of n k] by (simp add: sshiftr_def shiftr_def) + +lemma sshiftr_div_2n: "sint (w >>> n) = sint w div 2 ^ n" + apply (rule bit_eqI) + apply (cases \n < LENGTH('a)\) + apply (auto simp add: bit_simps not_less le_diff_conv2 simp flip: drop_bit_eq_div) done -lemma word_add_no_overflow:"(x::'a::len word) < max_word \ x < x + 1" - using less_x_plus_1 order_less_le by blast +lemma mask_eq: + \mask n = (1 << n) - (1 :: 'a::len word)\ + by (simp add: mask_eq_exp_minus_1 shiftl_def) -lemma lt_plus_1_le_word: - fixes x :: "'a::len word" - assumes bound:"n < unat (maxBound::'a word)" - shows "x < 1 + of_nat n = (x \ of_nat n)" - by (metis add.commute bound max_word_max word_Suc_leq word_not_le word_of_nat_less) +lemma nth_shiftl': "bit (w << m) n \ n < size w \ n >= m \ bit w (n - m)" + for w :: "'a::len word" + by (simp add: bit_simps word_size ac_simps) -lemma unat_ucast_up_simp: - fixes x :: "'a::len word" - assumes "LENGTH('a) \ LENGTH('b)" - shows "unat (ucast x :: 'b::len word) = unat x" - unfolding ucast_eq unat_eq_nat_uint - apply (subst int_word_uint) - apply (subst mod_pos_pos_trivial; simp?) - apply (rule lt2p_lem) - apply (simp add: assms) - done +lemmas nth_shiftl = nth_shiftl' [unfolded word_size] -lemma unat_ucast_less_no_overflow: - "\n < 2 ^ LENGTH('a); unat f < n\ \ (f::('a::len) word) < of_nat n" - by (erule (1) order_le_less_trans[OF _ of_nat_mono_maybe,rotated]) simp +lemma nth_shiftr: "bit (w >> m) n = bit w (n + m)" + for w :: "'a::len word" + by (simp add: bit_simps ac_simps) -lemma unat_ucast_less_no_overflow_simp: - "n < 2 ^ LENGTH('a) \ (unat f < n) = ((f::('a::len) word) < of_nat n)" - using unat_less_helper unat_ucast_less_no_overflow by blast +lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" + by (fact uint_shiftr_eq) -lemma unat_ucast_no_overflow_le: - assumes no_overflow: "unat b < (2 :: nat) ^ LENGTH('a)" - and upward_cast: "LENGTH('a) < LENGTH('b)" - shows "(ucast (f::'a::len word) < (b :: 'b :: len word)) = (unat f < unat b)" -proof - - have LR: "ucast f < b \ unat f < unat b" - apply (rule unat_less_helper) - apply (simp add:ucast_nat_def) - apply (rule_tac 'b1 = 'b in ucast_less_ucast[OF order.strict_implies_order, THEN iffD1]) - apply (rule upward_cast) - apply (simp add: ucast_ucast_mask less_mask_eq word_less_nat_alt - unat_power_lower[OF upward_cast] no_overflow) - done - have RL: "unat f < unat b \ ucast f < b" - proof- - assume ineq: "unat f < unat b" - have "ucast (f::'a::len word) < ((ucast (ucast b ::'a::len word)) :: 'b :: len word)" - apply (simp add: ucast_less_ucast[OF order.strict_implies_order] upward_cast) - apply (simp only: flip: ucast_nat_def) - apply (rule unat_ucast_less_no_overflow[OF no_overflow ineq]) - done - then show ?thesis - apply (rule order_less_le_trans) - apply (simp add:ucast_ucast_mask word_and_le2) - done - qed - then show ?thesis by (simp add:RL LR iffI) -qed +lemma shiftl_rev: "shiftl w n = word_reverse (shiftr (word_reverse w) n)" + by (rule bit_word_eqI) (auto simp add: bit_simps) -lemmas ucast_up_mono = ucast_less_ucast[THEN iffD2] +lemma rev_shiftl: "word_reverse w << n = word_reverse (w >> n)" + by (simp add: shiftl_rev) -lemma minus_one_word: - "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" - by simp +lemma shiftr_rev: "w >> n = word_reverse (word_reverse w << n)" + by (simp add: rev_shiftl) -lemma mask_exceed: - "n \ LENGTH('a) \ (x::'a::len word) AND NOT (mask n) = 0" - by (simp add: and_not_mask shiftr_eq_0) +lemma rev_shiftr: "word_reverse w >> n = word_reverse (w << n)" + by (simp add: shiftr_rev) + +lemmas ucast_up = + rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]] +lemmas ucast_down = + rc2 [simplified rev_shiftr revcast_ucast [symmetric]] + +lemma shiftl_zero_size: "size x \ n \ x << n = 0" + for x :: "'a::len word" + by (simp add: shiftl_def word_size) + +lemma shiftl_t2n: "shiftl w n = 2 ^ n * w" + for w :: "'a::len word" + by (simp add: shiftl_def push_bit_eq_mult) lemma word_shift_by_2: "x * 4 = (x::'a::len word) << 2" @@ -441,15 +198,542 @@ lemma word_shift_by_3: "x * 8 = (x::'a::len word) << 3" by (simp add: shiftl_t2n) -lemma le_2p_upper_bits: - "\ (p::'a::len word) \ 2^n - 1; n < LENGTH('a) \ \ - \n'\n. n' < LENGTH('a) \ \ p !! n'" - by (subst upper_bits_unset_is_l2p; simp) +lemma 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 le2p_bits_unset: - "p \ 2 ^ n - 1 \ \n'\n. n' < LENGTH('a) \ \ (p::'a::len word) !! n'" - using upper_bits_unset_is_l2p [where p=p] - by (cases "n < LENGTH('a)") auto +lemma 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 shiftr_x_0 [simp]: "x >> 0 = x" + for x :: "'a::len word" + by (simp add: shiftr_def) + +lemma shiftl_x_0 [simp]: "x << 0 = x" + for x :: "'a::len word" + by (simp add: shiftl_def) + +lemma shiftl_1: "(1::'a::len word) << n = 2^n" + by (simp add: shiftl_def) + +lemma shiftr_1 [simp]: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" + by (simp add: shiftr_def) + +lemma shiftl0: + "x << 0 = (x :: 'a :: len word)" + by (fact shiftl_x_0) + +lemma and_not_mask: + "w AND NOT (mask n) = (w >> n) << n" + for w :: \'a::len word\ + by (rule bit_word_eqI) (auto simp add: bit_simps) + +lemma and_mask: + "w AND mask n = (w << (size w - n)) >> (size w - n)" + for w :: \'a::len word\ + by (rule bit_word_eqI) (auto simp add: bit_simps word_size) + +lemma shiftr_div_2n_w: "n < size w \ w >> n = w div (2^n :: 'a :: len word)" + apply (unfold word_div_def) + apply (simp add: uint_2p_alt word_size) + apply (metis uint_shiftr_eq word_of_int_uint) + done + +lemma le_shiftr: + "u \ v \ u >> (n :: nat) \ (v :: 'a :: len word) >> n" + apply (unfold shiftr_def) + apply transfer + apply (simp add: take_bit_drop_bit) + apply (simp add: drop_bit_eq_div zdiv_mono1) + done + +lemma le_shiftr': + "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len word) \ v" + apply (metis le_cases le_shiftr verit_la_disequality) + done + +lemma shiftr_mask_le: + "n <= m \ mask n >> m = (0 :: 'a::len word)" + by (rule bit_word_eqI) (auto simp add: bit_simps) + +lemma shiftr_mask [simp]: + \mask m >> m = (0::'a::len word)\ + by (rule shiftr_mask_le) simp + +lemma le_mask_iff: + "(w \ mask n) = (w >> n = 0)" + for w :: \'a::len word\ + apply safe + apply (rule word_le_0_iff [THEN iffD1]) + apply (rule xtrans(3)) + apply (erule_tac [2] le_shiftr) + apply simp + apply (rule word_leI) + apply (rename_tac n') + apply (drule_tac x = "n' - n" in word_eqD) + apply (simp add : nth_shiftr word_size bit_simps) + apply (case_tac "n <= n'") + by auto + +lemma and_mask_eq_iff_shiftr_0: + "(w AND mask n = w) = (w >> n = 0)" + for w :: \'a::len word\ + by (simp flip: take_bit_eq_mask add: shiftr_def take_bit_eq_self_iff_drop_bit_eq_0) + +lemma mask_shiftl_decompose: + "mask m << n = mask (m + n) AND NOT (mask n :: 'a::len word)" + by (rule bit_word_eqI) (auto simp add: bit_simps) + +lemma shiftl_over_and_dist: + fixes a::"'a::len word" + shows "(a AND b) << c = (a << c) AND (b << c)" + by (unfold shiftl_def) (fact push_bit_and) + +lemma shiftr_over_and_dist: + fixes a::"'a::len word" + shows "a AND b >> c = (a >> c) AND (b >> c)" + by (unfold shiftr_def) (fact drop_bit_and) + +lemma sshiftr_over_and_dist: + fixes a::"'a::len word" + shows "a AND b >>> c = (a >>> c) AND (b >>> c)" + apply(rule word_eqI) + apply(simp add:nth_sshiftr word_ao_nth word_size) + done + +lemma shiftl_over_or_dist: + fixes a::"'a::len word" + shows "a OR b << c = (a << c) OR (b << c)" + by (unfold shiftl_def) (fact push_bit_or) + +lemma shiftr_over_or_dist: + fixes a::"'a::len word" + shows "a OR b >> c = (a >> c) OR (b >> c)" + by (unfold shiftr_def) (fact drop_bit_or) + +lemma sshiftr_over_or_dist: + fixes a::"'a::len word" + shows "a OR b >>> c = (a >>> c) OR (b >>> c)" + by (rule bit_word_eqI) (simp add: bit_simps) + +lemmas shift_over_ao_dists = + shiftl_over_or_dist shiftr_over_or_dist + sshiftr_over_or_dist shiftl_over_and_dist + shiftr_over_and_dist sshiftr_over_and_dist + +lemma shiftl_shiftl: + fixes a::"'a::len word" + shows "a << b << c = a << (b + c)" + apply(rule word_eqI) + apply(auto simp:word_size nth_shiftl add.commute add.left_commute) + done + +lemma shiftr_shiftr: + fixes a::"'a::len word" + shows "a >> b >> c = a >> (b + c)" + apply(rule word_eqI) + apply(simp add:word_size nth_shiftr add.left_commute add.commute) + done + +lemma shiftl_shiftr1: + fixes a::"'a::len word" + shows "c \ b \ a << b >> c = a AND (mask (size a - b)) << (b - c)" + apply (rule word_eqI) + apply (auto simp add: bit_simps not_le word_size ac_simps) + done + +lemma shiftl_shiftr2: + fixes a::"'a::len word" + shows "b < c \ a << b >> c = (a >> (c - b)) AND (mask (size a - c))" + apply(rule word_eqI) + apply(auto simp:nth_shiftr nth_shiftl word_size word_ao_nth bit_simps) + done + +lemma shiftr_shiftl1: + fixes a::"'a::len word" + shows "c \ b \ a >> b << c = (a >> (b - c)) AND (NOT (mask c))" + by (rule bit_word_eqI) (auto simp add: bit_simps) + +lemma shiftr_shiftl2: + fixes a::"'a::len word" + shows "b < c \ a >> b << c = (a << (c - b)) AND (NOT (mask c))" + apply (rule word_eqI) + apply (auto simp add: bit_simps not_le word_size ac_simps) + done + +lemmas multi_shift_simps = + shiftl_shiftl shiftr_shiftr + shiftl_shiftr1 shiftl_shiftr2 + shiftr_shiftl1 shiftr_shiftl2 + +lemma shiftr_mask2: + "n \ LENGTH('a) \ (mask n >> m :: ('a :: len) word) = mask (n - m)" + by (rule bit_word_eqI) (auto simp add: bit_simps) + +lemma word_shiftl_add_distrib: + fixes x :: "'a :: len word" + shows "(x + y) << n = (x << n) + (y << n)" + by (simp add: shiftl_t2n ring_distribs) + +lemma mask_shift: + "(x AND NOT (mask y)) >> y = x >> y" + for x :: \'a::len word\ + apply (rule bit_eqI) + apply (simp add: bit_and_iff bit_not_iff bit_shiftr_word_iff bit_mask_iff not_le) + using bit_imp_le_length apply auto + done + +lemma shiftr_div_2n': + "unat (w >> n) = unat w div 2 ^ n" + apply (unfold unat_eq_nat_uint) + apply (subst shiftr_div_2n) + apply (subst nat_div_distrib) + apply simp + apply (simp add: nat_power_eq) + done + +lemma shiftl_shiftr_id: + assumes nv: "n < LENGTH('a)" + and xv: "x < 2 ^ (LENGTH('a) - n)" + shows "x << n >> n = (x::'a::len word)" + apply (simp add: shiftl_t2n) + apply (rule word_eq_unatI) + apply (subst shiftr_div_2n') + apply (cases n) + apply simp + apply (subst iffD1 [OF unat_mult_lem])+ + apply (subst unat_power_lower[OF nv]) + apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) + apply (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl]) + apply (rule unat_power_lower) + apply simp + apply (subst unat_power_lower[OF nv]) + apply simp + done + +lemma ucast_shiftl_eq_0: + fixes w :: "'a :: len word" + shows "\ n \ LENGTH('b) \ \ ucast (w << n) = (0 :: 'b :: len word)" + by (transfer fixing: n) (simp add: take_bit_push_bit) + +lemma word_shift_nonzero: + "\ (x::'a::len word) \ 2 ^ m; m + n < LENGTH('a::len); x \ 0\ + \ x << n \ 0" + apply (simp only: word_neq_0_conv word_less_nat_alt + shiftl_t2n mod_0 unat_word_ariths + unat_power_lower word_le_nat_alt) + apply (subst mod_less) + apply (rule order_le_less_trans) + apply (erule mult_le_mono2) + apply (subst power_add[symmetric]) + apply (rule power_strict_increasing) + apply simp + apply simp + apply simp + done + +lemma word_shiftr_lt: + fixes w :: "'a::len word" + shows "unat (w >> n) < (2 ^ (LENGTH('a) - n))" + apply (subst shiftr_div_2n') + apply transfer + apply (simp flip: drop_bit_eq_div add: drop_bit_nat_eq drop_bit_take_bit) + done + +lemma shiftr_less_t2n': + "\ x AND mask (n + m) = x; m < LENGTH('a) \ \ x >> n < 2 ^ m" for x :: "'a :: len word" + apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) + apply transfer + apply (simp add: take_bit_drop_bit ac_simps) + done + +lemma shiftr_less_t2n: + "x < 2 ^ (n + m) \ x >> n < 2 ^ m" for x :: "'a :: len word" + apply (rule shiftr_less_t2n') + apply (erule less_mask_eq) + apply (rule ccontr) + apply (simp add: not_less) + apply (subst (asm) p2_eq_0[symmetric]) + apply (simp add: power_add) + done + +lemma shiftr_eq_0: + "n \ LENGTH('a) \ ((w::'a::len word) >> n) = 0" + apply (cut_tac shiftr_less_t2n'[of w n 0], simp) + apply (simp add: mask_eq_iff) + apply (simp add: lt2p_lem) + apply simp + done + +lemma shiftl_less_t2n: + fixes x :: "'a :: len word" + shows "\ x < (2 ^ (m - n)); m < LENGTH('a) \ \ (x << n) < 2 ^ m" + apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) + apply transfer + apply (simp add: take_bit_push_bit) + done + +lemma shiftl_less_t2n': + "(x::'a::len word) < 2 ^ m \ m+n < LENGTH('a) \ x << n < 2 ^ (m + n)" + by (rule shiftl_less_t2n) simp_all + +lemma scast_bit_test [simp]: + "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" + by (rule bit_word_eqI) (simp add: bit_simps) + +lemma signed_shift_guard_to_word: + \unat x * 2 ^ y < 2 ^ n \ x = 0 \ x < 1 << n >> y\ + if \n < LENGTH('a)\ \0 < n\ + for x :: \'a::len word\ +proof (cases \x = 0\) + case True + then show ?thesis + by simp +next + case False + then have \unat x \ 0\ + by (simp add: unat_eq_0) + then have \unat x \ 1\ + by simp + show ?thesis + proof (cases \y < n\) + case False + then have \n \ y\ + by simp + then obtain q where \y = n + q\ + using le_Suc_ex by blast + moreover have \(2 :: nat) ^ n >> n + q \ 1\ + by (simp add: drop_bit_eq_div power_add shiftr_def) + ultimately show ?thesis + using \x \ 0\ \unat x \ 1\ \n < LENGTH('a)\ + by (simp add: power_add not_less word_le_nat_alt unat_drop_bit_eq shiftr_def shiftl_def) + next + case True + with that have \y < LENGTH('a)\ + by simp + show ?thesis + proof (cases \2 ^ n = unat x * 2 ^ y\) + case True + moreover have \unat x * 2 ^ y < 2 ^ LENGTH('a)\ + using \n < LENGTH('a)\ by (simp flip: True) + moreover have \(word_of_nat (2 ^ n) :: 'a word) = word_of_nat (unat x * 2 ^ y)\ + using True by simp + then have \2 ^ n = x * 2 ^ y\ + by simp + ultimately show ?thesis + using \y < LENGTH('a)\ + by (auto simp add: drop_bit_eq_div word_less_nat_alt unat_div unat_word_ariths + shiftr_def shiftl_def) + next + case False + with \y < n\ have *: \unat x \ 2 ^ n div 2 ^ y\ + by (auto simp flip: power_sub power_add) + have \unat x * 2 ^ y < 2 ^ n \ unat x * 2 ^ y \ 2 ^ n\ + using False by (simp add: less_le) + also have \\ \ unat x \ 2 ^ n div 2 ^ y\ + by (simp add: less_eq_div_iff_mult_less_eq) + also have \\ \ unat x < 2 ^ n div 2 ^ y\ + using * by (simp add: less_le) + finally show ?thesis + using that \x \ 0\ by (simp flip: push_bit_eq_mult drop_bit_eq_div + add: shiftr_def shiftl_def unat_drop_bit_eq word_less_iff_unsigned [where ?'a = nat]) + qed + qed +qed + +lemma shiftr_not_mask_0: + "n+m \ LENGTH('a :: len) \ ((w::'a::len word) >> n) AND NOT (mask m) = 0" + by (rule bit_word_eqI) (auto simp add: bit_simps word_size dest: bit_imp_le_length) + +lemma shiftl_mask_is_0[simp]: + "(x << n) AND mask n = 0" + for x :: \'a::len word\ + by (simp flip: take_bit_eq_mask add: take_bit_push_bit shiftl_def) + +lemma rshift_sub_mask_eq: + "(a >> (size a - b)) AND mask b = a >> (size a - b)" + for a :: \'a::len word\ + using shiftl_shiftr2[where a=a and b=0 and c="size a - b"] + apply (cases "b < size a") + apply simp + apply (simp add: linorder_not_less mask_eq_decr_exp word_size + p2_eq_0[THEN iffD2]) + done + +lemma shiftl_shiftr3: + "b \ c \ a << b >> c = (a >> c - b) AND mask (size a - c)" + for a :: \'a::len word\ + apply (cases "b = c") + apply (simp add: shiftl_shiftr1) + apply (simp add: shiftl_shiftr2) + done + +lemma and_mask_shiftr_comm: + "m \ size w \ (w AND mask m) >> n = (w >> n) AND mask (m-n)" + for w :: \'a::len word\ + by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3) + +lemma and_mask_shiftl_comm: + "m+n \ size w \ (w AND mask m) << n = (w << n) AND mask (m+n)" + for w :: \'a::len word\ + by (simp add: and_mask word_size shiftl_shiftl) (simp add: shiftl_shiftr1) + +lemma le_mask_shiftl_le_mask: "s = m + n \ x \ mask n \ x << m \ mask s" + for x :: \'a::len word\ + by (simp add: le_mask_iff shiftl_shiftr3) + +lemma word_and_1_shiftl: + "x AND (1 << n) = (if bit x n then (1 << n) else 0)" for x :: "'a :: len word" + apply (rule bit_word_eqI; transfer) + apply (auto simp add: bit_simps not_le ac_simps) + done + +lemmas word_and_1_shiftls' + = word_and_1_shiftl[where n=0] + word_and_1_shiftl[where n=1] + word_and_1_shiftl[where n=2] + +lemmas word_and_1_shiftls = word_and_1_shiftls' [simplified] + +lemma word_and_mask_shiftl: + "x AND (mask n << m) = ((x >> m) AND mask n) << m" + for x :: \'a::len word\ + apply (rule bit_word_eqI; transfer) + apply (auto simp add: bit_simps not_le ac_simps) + done + +lemma shift_times_fold: + "(x :: 'a :: len word) * (2 ^ n) << m = x << (m + n)" + by (simp add: shiftl_t2n ac_simps power_add) + +lemma of_bool_nth: + "of_bool (bit x v) = (x >> v) AND 1" + for x :: \'a::len word\ + by (simp add: bit_iff_odd_drop_bit word_and_1 shiftr_def) + +lemma shiftr_mask_eq: + "(x >> n) AND mask (size x - n) = x >> n" for x :: "'a :: len word" + apply (simp flip: take_bit_eq_mask) + apply transfer + apply (simp add: take_bit_drop_bit) + done + +lemma shiftr_mask_eq': + "m = (size x - n) \ (x >> n) AND mask m = x >> n" for x :: "'a :: len word" + by (simp add: shiftr_mask_eq) + +lemma and_eq_0_is_nth: + fixes x :: "'a :: len word" + shows "y = 1 << n \ ((x AND y) = 0) = (\ (bit x n))" + by (simp add: and_exp_eq_0_iff_not_bit shiftl_def) + +lemma word_shift_zero: + "\ x << n = 0; x \ 2^m; m + n < LENGTH('a)\ \ (x::'a::len word) = 0" + apply (rule ccontr) + apply (drule (2) word_shift_nonzero) + apply simp + done + +lemma mask_shift_and_negate[simp]:"(w AND mask n << m) AND NOT (mask n << m) = 0" + for w :: \'a::len word\ + by (rule bit_word_eqI) (simp add: bit_simps) + +(* The seL4 bitfield generator produces functions containing mask and shift operations, such that + * invoking two of them consecutively can produce something like the following. + *) +lemma bitfield_op_twice: + "(x AND NOT (mask n << m) OR ((y AND mask n) << m)) AND NOT (mask n << m) = x AND NOT (mask n << m)" + for x :: \'a::len word\ + by (induct n arbitrary: m) (auto simp: word_ao_dist) + +lemma bitfield_op_twice'': + "\NOT a = b << c; \x. b = mask x\ \ (x AND a OR (y AND b << c)) AND a = x AND a" + for a b :: \'a::len word\ + apply clarsimp + apply (cut_tac n=xa and m=c and x=x and y=y in bitfield_op_twice) + apply (clarsimp simp:mask_eq_decr_exp) + apply (drule not_switch) + apply clarsimp + done + +lemma shiftr1_unfold: "x div 2 = x >> 1" + by (simp add: drop_bit_eq_div shiftr_def) + +lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2" + by (simp add: drop_bit_eq_div shiftr_def) + +lemma shiftl1_is_mult: "(x << 1) = (x :: 'a::len word) * 2" + by (metis One_nat_def mult_2 mult_2_right one_add_one + power_0 power_Suc shiftl_t2n) + +lemma shiftr1_lt:"x \ 0 \ (x::('a::len) word) >> 1 < x" + apply (subst shiftr1_is_div_2) + apply (rule div_less_dividend_word) + apply simp+ + done + +lemma shiftr1_0_or_1:"(x::('a::len) word) >> 1 = 0 \ x = 0 \ x = 1" + apply (subst (asm) shiftr1_is_div_2) + apply (drule word_less_div) + apply (case_tac "LENGTH('a) = 1") + apply (simp add:degenerate_word) + apply (erule disjE) + apply (subgoal_tac "(2::'a word) \ 0") + apply simp + apply (rule not_degenerate_imp_2_neq_0) + apply (subgoal_tac "LENGTH('a) \ 0") + apply arith + apply simp + apply (rule x_less_2_0_1', simp+) + done + +lemma shiftr1_irrelevant_lsb: "bit (x::('a::len) word) 0 \ x >> 1 = (x + 1) >> 1" + apply (cases \LENGTH('a)\; transfer) + apply (simp_all add: take_bit_drop_bit) + apply (simp add: drop_bit_take_bit drop_bit_Suc) + done + +lemma shiftr1_0_imp_only_lsb:"((x::('a::len) word) + 1) >> 1 = 0 \ x = 0 \ x + 1 = 0" + by (metis One_nat_def shiftr1_0_or_1 word_less_1 word_overflow) + +lemma shiftr1_irrelevant_lsb': "\ (bit (x::('a::len) word) 0) \ x >> 1 = (x + 1) >> 1" + by (metis shiftr1_irrelevant_lsb) + +(* Perhaps this one should be a simp lemma, but it seems a little dangerous. *) +lemma cast_chunk_assemble_id: + "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ + (((ucast ((ucast (x::'b word))::'a word))::'b word) OR (((ucast ((ucast (x >> n))::'a word))::'b word) << n)) = x" + apply (subgoal_tac "((ucast ((ucast (x >> n))::'a word))::'b word) = x >> n") + apply clarsimp + apply (subst and_not_mask[symmetric]) + apply (subst ucast_ucast_mask) + apply (subst word_ao_dist2[symmetric]) + apply clarsimp + apply (rule ucast_ucast_len) + apply (rule shiftr_less_t2n') + apply (subst and_mask_eq_iff_le_mask) + apply (simp_all add: mask_eq_decr_exp flip: mult_2_right) + apply (metis add_diff_cancel_left' len_gt_0 mult_2_right zero_less_diff) + done + +lemma cast_chunk_scast_assemble_id: + "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ + (((ucast ((scast (x::'b word))::'a word))::'b word) OR + (((ucast ((scast (x >> n))::'a word))::'b word) << n)) = x" + apply (subgoal_tac "((scast x)::'a word) = ((ucast x)::'a word)") + apply (subgoal_tac "((scast (x >> n))::'a word) = ((ucast (x >> n))::'a word)") + apply (simp add:cast_chunk_assemble_id) + apply (subst down_cast_same[symmetric], subst is_down, arith, simp)+ + done + +lemma unat_shiftr_less_t2n: + fixes x :: "'a :: len word" + shows "unat x < 2 ^ (n + m) \ unat (x >> n) < 2 ^ m" + by (simp add: shiftr_div_2n' power_add mult.commute less_mult_imp_div_less) lemma ucast_less_shiftl_helper: "\ LENGTH('b) + 2 < LENGTH('a); 2 ^ (LENGTH('b) + 2) \ n\ @@ -460,194 +744,14 @@ lemma ucast_less_shiftl_helper: apply (rule word_less_power_trans2; simp) done -lemma word_power_nonzero: - "\ (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x \ 0 \ - \ x * 2 ^ n \ 0" - by (metis and_mask_eq_iff_shiftr_0 less_mask_eq p2_gt_0 semiring_normalization_rules(7) - shiftl_shiftr_id shiftl_t2n) - -lemma less_1_helper: - "n \ m \ (n - 1 :: int) < m" - by arith - -lemma div_power_helper: - "\ x \ y; y < LENGTH('a) \ \ (2 ^ y - 1) div (2 ^ x :: 'a::len word) = 2 ^ (y - x) - 1" - apply (rule word_uint.Rep_eqD) - apply (simp only: uint_word_ariths uint_div uint_power_lower) - apply (subst mod_pos_pos_trivial, fastforce, fastforce)+ - apply (subst mod_pos_pos_trivial) - apply (simp add: le_diff_eq uint_2p_alt) - apply (rule less_1_helper) - apply (rule power_increasing; simp) - apply (subst mod_pos_pos_trivial) - apply (simp add: uint_2p_alt) - apply (rule less_1_helper) - apply (rule power_increasing; simp) - apply (subst int_div_sub_1; simp add: uint_2p_alt) - apply (subst power_0[symmetric]) - apply (simp add: uint_2p_alt le_imp_power_dvd power_diff_power_eq) - done - -lemma word_add_power_off: - fixes a :: "'a :: len word" - assumes ak: "a < k" - and kw: "k < 2 ^ (LENGTH('a) - m)" - and mw: "m < LENGTH('a)" - and off: "off < 2 ^ m" - shows "(a * 2 ^ m) + off < k * 2 ^ m" -proof (cases "m = 0") - case True - then show ?thesis using off ak by simp -next - case False - - from ak have ak1: "a + 1 \ k" by (rule inc_le) - then have "(a + 1) * 2 ^ m \ 0" - apply - - apply (rule word_power_nonzero) - apply (erule order_le_less_trans [OF _ kw]) - apply (rule mw) - apply (rule less_is_non_zero_p1 [OF ak]) - done - then have "(a * 2 ^ m) + off < ((a + 1) * 2 ^ m)" using kw mw - apply - - apply (simp add: distrib_right) - apply (rule word_plus_strict_mono_right [OF off]) - apply (rule is_aligned_no_overflow'') - apply (rule is_aligned_mult_triv2) - apply assumption - done - also have "\ \ k * 2 ^ m" using ak1 mw kw False - apply - - apply (erule word_mult_le_mono1) - apply (simp add: p2_gt_0) - apply (simp add: word_less_nat_alt) - apply (meson nat_mult_power_less_eq zero_less_numeral) - done - finally show ?thesis . -qed - -lemma offset_not_aligned: - "\ is_aligned (p::'a::len word) n; i > 0; i < 2 ^ n; n < LENGTH('a)\ \ - \ is_aligned (p + of_nat i) n" - apply (erule is_aligned_add_not_aligned) - apply transfer - apply (auto simp add: is_aligned_iff_udvd) - apply (metis bintrunc_bintrunc_ge int_ops(1) nat_int_comparison(1) nat_less_le take_bit_eq_0_iff take_bit_nat_eq_self_iff take_bit_of_nat) - done - -lemma length_upto_enum_one: - fixes x :: "'a :: len word" - assumes lt1: "x < y" and lt2: "z < y" and lt3: "x \ z" - shows "[x , y .e. z] = [x]" - unfolding upto_enum_step_def -proof (subst upto_enum_red, subst if_not_P [OF leD [OF lt3]], clarsimp, rule conjI) - show "unat ((z - x) div (y - x)) = 0" - proof (subst unat_div, rule div_less) - have syx: "unat (y - x) = unat y - unat x" - by (rule unat_sub [OF order_less_imp_le]) fact - moreover have "unat (z - x) = unat z - unat x" - by (rule unat_sub) fact - - ultimately show "unat (z - x) < unat (y - x)" - using lt2 lt3 unat_mono word_less_minus_mono_left by blast - qed - - then show "(z - x) div (y - x) * (y - x) = 0" - by (metis mult_zero_left unat_0 word_unat.Rep_eqD) -qed - -lemma max_word_mask: - "(max_word :: 'a::len word) = mask LENGTH('a)" - unfolding mask_eq_decr_exp by simp - -lemmas mask_len_max = max_word_mask[symmetric] - -lemma mask_out_first_mask_some: - "\ x AND NOT (mask n) = y; n \ m \ \ x AND NOT (mask m) = y AND NOT (mask m)" - for x y :: \'a::len word\ - by (rule bit_word_eqI) (auto simp add: bit_simps) - -lemma mask_lower_twice: - "n \ m \ (x AND NOT (mask n)) AND NOT (mask m) = x AND NOT (mask m)" - for x :: \'a::len word\ - by (rule bit_word_eqI) (auto simp add: bit_simps) - -lemma mask_lower_twice2: - "(a AND NOT (mask n)) AND NOT (mask m) = a AND NOT (mask (max n m))" - for a :: \'a::len word\ - by (rule bit_word_eqI) (auto simp add: bit_simps) - -lemma ucast_and_neg_mask: - "ucast (x AND NOT (mask n)) = ucast x AND NOT (mask n)" - apply (rule bit_word_eqI) - apply (auto simp add: bit_simps dest: bit_imp_le_length) - done - -lemma ucast_and_mask: - "ucast (x AND mask n) = ucast x AND mask n" - apply (rule bit_word_eqI) - apply (auto simp add: bit_simps dest: bit_imp_le_length) - done - -lemma ucast_mask_drop: - "LENGTH('a :: len) \ n \ (ucast (x AND mask n) :: 'a word) = ucast x" - apply (rule bit_word_eqI) - apply (auto simp add: bit_simps dest: bit_imp_le_length) - done - (* negating a mask which has been shifted to the very left *) lemma NOT_mask_shifted_lenword: "NOT (mask len << (LENGTH('a) - len) ::'a::len word) = mask (LENGTH('a) - len)" by (rule bit_word_eqI) - (auto simp add: shiftl_word_eq word_size bit_not_iff bit_push_bit_iff bit_mask_iff) + (auto simp add: word_size bit_not_iff bit_push_bit_iff bit_mask_iff shiftl_def) (* Comparisons between different word sizes. *) -lemma eq_ucast_ucast_eq: - "LENGTH('b) \ LENGTH('a) \ x = ucast y \ ucast x = y" - for x :: "'a::len word" and y :: "'b::len word" - by transfer simp - -lemma le_ucast_ucast_le: - "x \ ucast y \ ucast x \ y" - for x :: "'a::len word" and y :: "'b::len word" - by (smt le_unat_uoi linorder_not_less order_less_imp_le ucast_nat_def unat_arith_simps(1)) - -lemma less_ucast_ucast_less: - "LENGTH('b) \ LENGTH('a) \ x < ucast y \ ucast x < y" - for x :: "'a::len word" and y :: "'b::len word" - by (metis ucast_nat_def unat_mono unat_ucast_up_simp word_of_nat_less) - -lemma ucast_le_ucast: - "LENGTH('a) \ LENGTH('b) \ (ucast x \ (ucast y::'b::len word)) = (x \ y)" - for x :: "'a::len word" - by (simp add: unat_arith_simps(1) unat_ucast_up_simp) - -lemmas ucast_up_mono_le = ucast_le_ucast[THEN iffD2] - -lemma ucast_le_ucast_eq: - fixes x y :: "'a::len word" - assumes x: "x < 2 ^ n" - assumes y: "y < 2 ^ n" - assumes n: "n = LENGTH('b::len)" - shows "(UCAST('a \ 'b) x \ UCAST('a \ 'b) y) = (x \ y)" - apply (rule iffI) - apply (cases "LENGTH('b) < LENGTH('a)") - apply (subst less_mask_eq[OF x, symmetric]) - apply (subst less_mask_eq[OF y, symmetric]) - apply (unfold n) - apply (subst ucast_ucast_mask[symmetric])+ - apply (simp add: ucast_le_ucast)+ - apply (erule ucast_mono_le[OF _ y[unfolded n]]) - done - -lemma ucast_or_distrib: - fixes x :: "'a::len word" - fixes y :: "'a::len word" - shows "(ucast (x OR y) :: ('b::len) word) = ucast x OR ucast y" - by transfer simp - lemma shiftr_less: "(w::'a::len word) < k \ w >> n < k" by (metis div_le_dividend le_less_trans shiftr_div_2n' unat_arith_simps(2)) @@ -656,13 +760,10 @@ lemma word_and_notzeroD: "w AND w' \ 0 \ w \ 0 \ w' \ 0" by auto -lemma word_exists_nth: - "(w::'a::len word) \ 0 \ \i. w !! i" - by (simp add: bit_eq_iff test_bit_word_eq) - lemma shiftr_le_0: "unat (w::'a::len word) < 2 ^ n \ w >> n = (0::'a::len word)" - by (rule word_unat.Rep_eqD) (simp add: shiftr_div_2n') + by (auto simp add: take_bit_word_eq_self_iff word_less_nat_alt shiftr_def + simp flip: take_bit_eq_self_iff_drop_bit_eq_0 intro: ccontr) lemma of_nat_shiftl: "(of_nat x << n) = (of_nat (x * 2 ^ n) :: ('a::len) word)" @@ -676,131 +777,13 @@ lemma shiftl_1_not_0: "n < LENGTH('a) \ (1::'a::len word) << n \ 0" by (simp add: shiftl_t2n) -lemma max_word_not_0 [simp]: - "- 1 \ (0 :: 'a::len word)" - by simp - -lemma ucast_zero_is_aligned: - "UCAST('a::len \ 'b::len) w = 0 \ n \ LENGTH('b) \ is_aligned w n" - by (clarsimp simp: is_aligned_mask word_eq_iff word_size nth_ucast) - -lemma unat_ucast_eq_unat_and_mask: - "unat (UCAST('b::len \ 'a::len) w) = unat (w AND mask LENGTH('a))" - apply (simp flip: take_bit_eq_mask) - apply transfer - apply (simp add: ac_simps) - done - -lemma unat_max_word_pos[simp]: "0 < unat (- 1 :: 'a::len word)" - using unat_gt_0 [of \- 1 :: 'a::len word\] by simp - - -(* Miscellaneous conditional injectivity rules. *) - -lemma mult_pow2_inj: - assumes ws: "m + n \ LENGTH('a)" - assumes le: "x \ mask m" "y \ mask m" - assumes eq: "x * 2 ^ n = y * (2 ^ n::'a::len word)" - shows "x = y" -proof (rule bit_word_eqI) - fix q - assume \q < LENGTH('a)\ - from eq have \push_bit n x = push_bit n y\ - by (simp add: push_bit_eq_mult) - moreover from le have \take_bit m x = x\ \take_bit m y = y\ - by (simp_all add: less_eq_mask_iff_take_bit_eq_self) - ultimately have \push_bit n (take_bit m x) = push_bit n (take_bit m y)\ - by simp_all - with \q < LENGTH('a)\ ws show \bit x q \ bit y q\ - apply (simp add: push_bit_take_bit) - unfolding bit_eq_iff - apply (simp add: bit_simps not_le) - apply (metis (full_types) \take_bit m x = x\ \take_bit m y = y\ add.commute add_diff_cancel_right' add_less_cancel_right bit_take_bit_iff le_add2 less_le_trans) - done -qed - -lemma word_of_nat_inj: - assumes bounded: "x < 2 ^ LENGTH('a)" "y < 2 ^ LENGTH('a)" - assumes of_nats: "of_nat x = (of_nat y :: 'a::len word)" - shows "x = y" - by (rule contrapos_pp[OF of_nats]; cases "x < y"; cases "y < x") - (auto dest: bounded[THEN of_nat_mono_maybe]) - -(* Uints *) - -lemma uints_mono_iff: "uints l \ uints m \ l \ m" - using power_increasing_iff[of "2::int" l m] - apply (auto simp: uints_num subset_iff simp del: power_increasing_iff) - apply (meson less_irrefl not_le zero_le_numeral zero_le_power) - done - -lemmas uints_monoI = uints_mono_iff[THEN iffD2] - -lemma Bit_in_uints_Suc: "of_bool c + 2 * w \ uints (Suc m)" if "w \ uints m" - using that - by (auto simp: uints_num) - -lemma Bit_in_uintsI: "of_bool c + 2 * w \ uints m" if "w \ uints (m - 1)" "m > 0" - using Bit_in_uints_Suc[OF that(1)] that(2) - by auto - -lemma bin_cat_in_uintsI: - \bin_cat a n b \ uints m\ if \a \ uints l\ \m \ l + n\ -proof - - from \m \ l + n\ obtain q where \m = l + n + q\ - using le_Suc_ex by blast - then have \(2::int) ^ m = 2 ^ n * 2 ^ (l + q)\ - by (simp add: ac_simps power_add) - moreover have \a mod 2 ^ (l + q) = a\ - using \a \ uints l\ - by (auto simp add: uints_def take_bit_eq_mod power_add Divides.mod_mult2_eq) - ultimately have \concat_bit n b a = take_bit m (concat_bit n b a)\ - by (simp add: concat_bit_eq take_bit_eq_mod push_bit_eq_mult Divides.mod_mult2_eq) - then show ?thesis - by (simp add: uints_def) -qed - -lemma bin_cat_cong: "bin_cat a n b = bin_cat c m d" - if "n = m" "a = c" "bintrunc m b = bintrunc m d" - using that(3) unfolding that(1,2) by (simp add: bin_cat_eq_push_bit_add_take_bit) - -lemma bin_cat_eqD1: "bin_cat a n b = bin_cat c n d \ a = c" - by (metis drop_bit_bin_cat_eq) - -lemma bin_cat_eqD2: "bin_cat a n b = bin_cat c n d \ bintrunc n b = bintrunc n d" - by (metis take_bit_bin_cat_eq) - -lemma bin_cat_inj: "(bin_cat a n b) = bin_cat c n d \ a = c \ bintrunc n b = bintrunc n d" - by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) - -lemma word_of_int_bin_cat_eq_iff: - "(word_of_int (bin_cat (uint a) LENGTH('b) (uint b))::'c::len word) = - word_of_int (bin_cat (uint c) LENGTH('b) (uint d)) \ b = d \ a = c" - if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" - for a::"'a::len word" and b::"'b::len word" - by (subst word_uint.Abs_inject) - (auto simp: bin_cat_inj intro!: that bin_cat_in_uintsI) - -lemma word_cat_inj: "(word_cat a b::'c::len word) = word_cat c d \ a = c \ b = d" - if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" - for a::"'a::len word" and b::"'b::len word" - using word_of_int_bin_cat_eq_iff [OF that, of b a d c] - by transfer auto - -lemma p2_eq_1: "2 ^ n = (1::'a::len word) \ n = 0" -proof - - have "2 ^ n = (1::'a word) \ n = 0" - by (metis One_nat_def not_less one_less_numeral_iff p2_eq_0 p2_gt_0 power_0 power_0 - power_inject_exp semiring_norm(76) unat_power_lower zero_neq_one) - then show ?thesis by auto -qed +(* continue sorting out from here *) (* usually: x,y = (len_of TYPE ('a)) *) lemma bitmagic_zeroLast_leq_or1Last: "(a::('a::len) word) AND (mask len << x - len) \ a OR mask (y - len)" by (meson le_word_or2 order_trans word_and_le2) - lemma zero_base_lsb_imp_set_eq_as_bit_operation: fixes base ::"'a::len word" assumes valid_prefix: "mask (LENGTH('a) - len) AND base = 0" @@ -870,10 +853,7 @@ lemma word_and_or_mask_aligned2: lemma is_aligned_ucastI: "is_aligned w n \ is_aligned (ucast w) n" - apply transfer - apply (auto simp add: min_def) - apply (metis bintrunc_bintrunc_ge bintrunc_n_0 nat_less_le not_le take_bit_eq_0_iff) - done + by (simp add: bit_ucast_iff is_aligned_nth) lemma ucast_le_maskI: "a \ mask n \ UCAST('a::len \ 'b::len) a \ mask n" @@ -885,7 +865,7 @@ lemma ucast_add_mask_aligned: lemma ucast_shiftl: "LENGTH('b) \ LENGTH ('a) \ UCAST ('a::len \ 'b::len) x << n = ucast (x << n)" - by word_eqI_solve + by (rule bit_word_eqI) (auto simp add: bit_simps) lemma ucast_leq_mask: "LENGTH('a) \ n \ ucast (x::'a::len word) \ mask n" @@ -895,26 +875,57 @@ lemma ucast_leq_mask: done lemma shiftl_inj: - "\ x << n = y << n; x \ mask (LENGTH('a)-n); y \ mask (LENGTH('a)-n) \ \ - x = (y :: 'a :: len word)" - apply word_eqI - apply (rename_tac n') - apply (case_tac "LENGTH('a) - n \ n'", simp) - by (metis add.commute add.right_neutral diff_add_inverse le_diff_conv linorder_not_less zero_order(1)) + \x = y\ + if \x << n = y << n\ \x \ mask (LENGTH('a) - n)\ \y \ mask (LENGTH('a) - n)\ + for x y :: \'a::len word\ +proof (cases \n < LENGTH('a)\) + case False + with that show ?thesis + by simp +next + case True + moreover from that have \take_bit (LENGTH('a) - n) x = x\ \take_bit (LENGTH('a) - n) y = y\ + by (simp_all add: less_eq_mask_iff_take_bit_eq_self) + ultimately show ?thesis + using \x << n = y << n\ by (metis diff_less gr_implies_not0 linorder_cases linorder_not_le shiftl_shiftr_id shiftl_x_0 take_bit_word_eq_self_iff) +qed lemma distinct_word_add_ucast_shift_inj: - "\ p + (UCAST('a::len \ 'b::len) off << n) = p' + (ucast off' << n); - is_aligned p n'; is_aligned p' n'; n' = n + LENGTH('a); n' < LENGTH('b) \ - \ p' = p \ off' = off" - apply (simp add: word_and_or_mask_aligned le_mask_shiftl_le_mask[where n="LENGTH('a)"] - ucast_leq_mask) - apply (simp add: is_aligned_nth) - apply (rule conjI; word_eqI) - apply (metis add.commute test_bit_conj_lt diff_add_inverse le_diff_conv nat_less_le) - apply (rename_tac i) - apply (erule_tac x="i+n" in allE) - apply simp - done + \p' = p \ off' = off\ + if *: \p + (UCAST('a::len \ 'b::len) off << n) = p' + (ucast off' << n)\ + and \is_aligned p n'\ \is_aligned p' n'\ \n' = n + LENGTH('a)\ \n' < LENGTH('b)\ +proof - + from \n' = n + LENGTH('a)\ + have [simp]: \n' - n = LENGTH('a)\ \n + LENGTH('a) = n'\ + by simp_all + from \is_aligned p n'\ obtain q + where p: \p = push_bit n' (word_of_nat q)\ \q < 2 ^ (LENGTH('b) - n')\ + by (rule is_alignedE') + from \is_aligned p' n'\ obtain q' + where p': \p' = push_bit n' (word_of_nat q')\ \q' < 2 ^ (LENGTH('b) - n')\ + by (rule is_alignedE') + define m :: nat where \m = unat off\ + then have off: \off = word_of_nat m\ + by simp + define m' :: nat where \m' = unat off'\ + then have off': \off' = word_of_nat m'\ + by simp + have \push_bit n' q + take_bit n' (push_bit n m) < 2 ^ LENGTH('b)\ + by (metis id_apply is_aligned_no_wrap''' of_nat_eq_id of_nat_push_bit p(1) p(2) take_bit_nat_eq_self_iff take_bit_nat_less_exp take_bit_push_bit that(2) that(5) unsigned_of_nat) + moreover have \push_bit n' q' + take_bit n' (push_bit n m') < 2 ^ LENGTH('b)\ + by (metis \n' - n = LENGTH('a)\ id_apply is_aligned_no_wrap''' m'_def of_nat_eq_id of_nat_push_bit off' p'(1) p'(2) take_bit_nat_eq_self_iff take_bit_push_bit that(3) that(5) unsigned_of_nat) + ultimately have \push_bit n' q + take_bit n' (push_bit n m) = push_bit n' q' + take_bit n' (push_bit n m')\ + using * by (simp add: p p' off off' push_bit_of_nat push_bit_take_bit word_of_nat_inj unsigned_of_nat shiftl_def flip: of_nat_add) + then have \int (push_bit n' q + take_bit n' (push_bit n m)) + = int (push_bit n' q' + take_bit n' (push_bit n m'))\ + by simp + then have \concat_bit n' (int (push_bit n m)) (int q) + = concat_bit n' (int (push_bit n m')) (int q')\ + by (simp add: of_nat_push_bit of_nat_take_bit concat_bit_eq) + then show ?thesis + by (simp add: p p' off off' take_bit_of_nat take_bit_push_bit word_of_nat_eq_iff concat_bit_eq_iff) + (simp add: push_bit_eq_mult) +qed lemma word_upto_Nil: "y < x \ [x .e. y ::'a::len word] = []" @@ -929,10 +940,6 @@ proof - then show ?thesis by auto qed -lemma max_word_not_less[simp]: - "\ max_word < x" - by (simp add: not_less) - lemma word_enum_prefix: "[x .e. (y ::'a::len word)] = as @ a # bs \ as = (if x < a then [x .e. a - 1] else [])" apply (induct as arbitrary: x; clarsimp) @@ -999,7 +1006,7 @@ lemma add_mult_aligned_neg_mask: \(x + y * m) AND NOT(mask n) = (x AND NOT(mask n)) + y * m\ if \m AND (2 ^ n - 1) = 0\ for x y m :: \'a::len word\ - by (metis (no_types, hide_lams) + by (metis (no_types, opaque_lifting) add.assoc add.commute add.right_neutral add_uminus_conv_diff mask_eq_decr_exp mask_eqs(2) mask_eqs(6) mult.commute mult_zero_left subtract_mask(1) that) @@ -1048,8 +1055,8 @@ lemma neg_mask_add: by (clarsimp simp: mask_out_sub_mask mask_eqs(7)[symmetric] mask_twice) lemma shiftr_shiftl_shiftr[simp]: - "(x :: 'a :: len word) >> a << a >> a = x >> a" - by word_eqI_solve + "(x :: 'a :: len word) >> a << a >> a = x >> a" + by (rule bit_word_eqI) (auto simp add: bit_simps dest: bit_imp_le_length) lemma add_right_shift: "\ x AND mask n = 0; y AND mask n = 0; x \ x + y \ @@ -1086,11 +1093,11 @@ lemma word_and_less': lemma shiftr_w2p: "x < LENGTH('a) \ 2 ^ x = (2 ^ (LENGTH('a) - 1) >> (LENGTH('a) - 1 - x) :: 'a :: len word)" - by word_eqI_solve + by (rule bit_word_eqI) (auto simp add: bit_simps) lemma t2p_shiftr: "\ b \ a; a < LENGTH('a) \ \ (2 :: 'a :: len word) ^ a >> b = 2 ^ (a - b)" - by word_eqI_solve + by (rule bit_word_eqI) (auto simp add: bit_simps) lemma scast_1[simp]: "scast (1 :: 'a :: len signed word) = (1 :: 'a word)" @@ -1098,16 +1105,16 @@ lemma scast_1[simp]: lemma unsigned_uminus1 [simp]: \(unsigned (-1::'b::len word)::'c::len word) = mask LENGTH('b)\ - by word_eqI + by (fact unsigned_minus_1_eq_mask) lemma ucast_ucast_mask_eq: "\ UCAST('a::len \ 'b::len) x = y; x AND mask LENGTH('b) = x \ \ x = ucast y" - by word_eqI_solve + by (drule sym) (simp flip: take_bit_eq_mask add: unsigned_ucast_eq) lemma ucast_up_eq: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ LENGTH ('b) \ \ ucast x = (ucast y::'a::len word)" - by word_eqI_solve + by (simp add: word_eq_iff bit_simps) lemma ucast_up_neq: "\ ucast x \ (ucast y::'b::len word); LENGTH('b) \ LENGTH ('a) \ @@ -1138,17 +1145,21 @@ lemma mask_eq1_nochoice: lemma shiftr_and_eq_shiftl: "(w >> n) AND x = y \ w AND (x << n) = (y << n)" for y :: "'a:: len word" - by (metis (no_types, lifting) and_not_mask bit.conj_ac(1) bit.conj_ac(2) mask_eq_0_eq_x shiftl_mask_is_0 shiftl_over_and_dist) + apply (drule sym) + apply simp + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps) + done lemma add_mask_lower_bits': "\ len = LENGTH('a); is_aligned (x :: 'a :: len word) n; - \n' \ n. n' < len \ \ p !! n' \ + \n' \ n. n' < len \ \ bit p n' \ \ x + p AND NOT(mask n) = x" using add_mask_lower_bits by auto lemma leq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + high_bits) \ (x >> low_bits) \ mask high_bits" - by (simp add: le_mask_iff shiftr_shiftr) + by (simp add: le_mask_iff shiftr_shiftr ac_simps) lemma ucast_ucast_eq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + LENGTH('b)) @@ -1158,11 +1169,7 @@ lemma ucast_ucast_eq_mask_shift: lemma const_le_unat: "\ b < 2 ^ LENGTH('a); of_nat b \ a \ \ b \ unat (a :: 'a :: len word)" - apply (simp add: word_le_def) - apply (simp only: uint_nat zle_int) - apply transfer - apply (simp add: take_bit_nat_eq_self) - done + by (simp add: word_le_nat_alt unsigned_of_nat take_bit_nat_eq_self) lemma upt_enum_offset_trivial: "\ x < 2 ^ LENGTH('a) - 1 ; n \ unat x \ @@ -1183,7 +1190,7 @@ lemma ucast_add: lemma ucast_minus: "ucast (a - (b :: 'a :: len word)) = ucast a - (ucast b :: ('a signed word))" apply (insert ucast_add[where a=a and b="-b"]) - apply (metis (no_types, hide_lams) add_diff_eq diff_add_cancel ucast_add) + apply (metis (no_types, opaque_lifting) add_diff_eq diff_add_cancel ucast_add) done lemma scast_ucast_add_one [simp]: @@ -1231,7 +1238,9 @@ lemma shiftr_eqD: lemma word_shiftr_shiftl_shiftr_eq_shiftr: "a \ b \ (x :: 'a :: len word) >> a << b >> b = x >> a" - by (simp add: mask_shift multi_shift_simps(5) shiftr_shiftr) + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps dest: bit_imp_le_length) + done lemma of_int_uint_ucast: "of_int (uint (x :: 'a::len word)) = (ucast x :: 'b::len word)" @@ -1302,9 +1311,30 @@ lemma unat_pow_le_intro: by (metis lt2p_lem not_le of_nat_le_iff of_nat_numeral semiring_1_class.of_nat_power uint_nat) lemma unat_shiftl_less_t2n: - "\ unat (x :: 'a :: len word) < 2 ^ (m - n); m < LENGTH('a) \ \ unat (x << n) < 2 ^ m" - by (metis (no_types) of_nat_power diff_le_self le_less_trans shiftl_less_t2n - unat_less_power word_unat.Rep_inverse) + \unat (x << n) < 2 ^ m\ + if \unat (x :: 'a :: len word) < 2 ^ (m - n)\ \m < LENGTH('a)\ +proof (cases \n \ m\) + case False + with that show ?thesis + apply (transfer fixing: m n) + apply (simp add: not_le take_bit_push_bit) + apply (metis diff_le_self order_le_less_trans push_bit_of_0 take_bit_0 take_bit_int_eq_self + take_bit_int_less_exp take_bit_nonnegative take_bit_tightened) + done +next + case True + moreover define q r where \q = m - n\ and \r = LENGTH('a) - n - q\ + ultimately have \m - n = q\ \m = n + q\ \LENGTH('a) = r + q + n\ + using that by simp_all + with that show ?thesis + apply (transfer fixing: m n q r) + apply (simp add: not_le take_bit_push_bit) + apply (simp add: push_bit_eq_mult power_add) + using take_bit_tightened_less_eq_int [of \r + q\ \r + q + n\] + apply (rule le_less_trans) + apply simp_all + done +qed lemma unat_is_aligned_add: "\ is_aligned p n; unat d < 2 ^ n \ @@ -1403,7 +1433,7 @@ lemma unat_2tp_if: lemma mask_of_mask: "mask (n::nat) AND mask (m::nat) = (mask (min m n) :: 'a::len word)" - by word_eqI_solve + by (rule bit_word_eqI) (auto simp add: bit_simps) lemma unat_signed_ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ unat (ucast (x :: 'a :: len word) :: 'b :: len signed word) = unat x" @@ -1415,22 +1445,21 @@ lemma toEnum_of_ucast: by (simp add: unat_pow_le_intro) lemma plus_mask_AND_NOT_mask_eq: - "x AND NOT(mask n) = x \ (x + mask n) AND NOT(mask n) = x" - for x::\'a::len word\ + "x AND NOT(mask n) = x \ (x + mask n) AND NOT(mask n) = x" for x::\'a::len word\ by (subst word_plus_and_or_coroll; word_eqI_solve) lemmas unat_ucast_mask = unat_ucast_eq_unat_and_mask[where w=a for a] lemma t2n_mask_eq_if: "2 ^ n AND mask m = (if n < m then 2 ^ n else (0 :: 'a::len word))" - by (rule word_eqI, auto simp add: word_size nth_w2p split: if_split) + by (rule word_eqI) (auto simp add: bit_simps) lemma unat_ucast_le: "unat (ucast (x :: 'a :: len word) :: 'b :: len word) \ unat x" by (simp add: ucast_nat_def word_unat_less_le) lemma ucast_le_up_down_iff: - "\ LENGTH('a) \ LENGTH('b); (x :: 'b :: len word) \ ucast (max_word :: 'a :: len word) \ + "\ LENGTH('a) \ LENGTH('b); (x :: 'b :: len word) \ ucast (- 1 :: 'a :: len word) \ \ (ucast x \ (y :: 'a word)) = (x \ ucast y)" using le_max_word_ucast_id ucast_le_ucast by metis @@ -1480,7 +1509,14 @@ lemma mask_shift_eq_mask_mask: lemma mask_shift_sum: "\ a \ b; unat n = unat (p AND mask b) \ \ (p AND NOT(mask a)) + (p AND mask a >> b) * (1 << b) + n = (p :: 'a :: len word)" - by (metis and_not_mask mask_rshift_mult_eq_rshift_lshift mask_split_sum_twice word_unat.Rep_eqD) + apply (simp add: shiftl_def shiftr_def flip: push_bit_eq_mult take_bit_eq_mask word_unat_eq_iff) + apply (subst disjunctive_add) + apply (auto simp add: bit_simps) + apply (subst disjunctive_add) + apply (auto simp add: bit_simps) + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps) + done lemma is_up_compose: "\ is_up uc; is_up uc' \ \ is_up (uc' \ uc)" @@ -1492,7 +1528,7 @@ lemma of_int_sint_scast: lemma scast_of_nat_to_signed [simp]: "scast (of_nat x :: 'a :: len word) = (of_nat x :: 'a signed word)" - by transfer simp + by (rule bit_word_eqI) (simp add: bit_simps) lemma scast_of_nat_signed_to_unsigned_add: "scast (of_nat x + of_nat y :: 'a :: len signed word) = (of_nat x + of_nat y :: 'a :: len word)" @@ -1521,8 +1557,7 @@ lemma sint_eq_uint_2pl: lemma pow_sub_less: "\ a + b \ LENGTH('a); unat (x :: 'a :: len word) = 2 ^ a \ \ unat (x * 2 ^ b - 1) < 2 ^ (a + b)" - by (metis (mono_tags) eq_or_less_helperD not_less of_nat_numeral power_add - semiring_1_class.of_nat_power unat_pow_le_intro word_unat.Rep_inverse) + by (smt (z3) eq_or_less_helperD le_add2 le_eq_less_or_eq le_trans power_add unat_mult_lem unat_pow_le_intro unat_power_lower word_eq_unatI) lemma sle_le_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a \ b \ \ a <=s b" @@ -1535,7 +1570,7 @@ lemma sless_less_2pl: lemma and_mask2: "w << n >> n = w AND mask (size w - n)" for w :: \'a::len word\ - by (cases "n \ size w"; clarsimp simp: word_and_le2 and_mask shiftl_zero_size) + by (rule bit_word_eqI) (auto simp add: bit_simps word_size) lemma aligned_sub_aligned_simple: "\ is_aligned a n; is_aligned b n \ \ is_aligned (a - b) n" @@ -1543,12 +1578,12 @@ lemma aligned_sub_aligned_simple: lemma minus_one_shift: "- (1 << n) = (-1 << n :: 'a::len word)" - by (simp add: mask_eq_decr_exp NOT_eq flip: mul_not_mask_eq_neg_shiftl) + by (simp add: shiftl_def minus_exp_eq_not_mask) lemma ucast_eq_mask: "(UCAST('a::len \ 'b::len) x = UCAST('a \ 'b) y) = (x AND mask LENGTH('b) = y AND mask LENGTH('b))" - by (rule iffI; word_eqI_solve) + by transfer (simp flip: take_bit_eq_mask add: ac_simps) context fixes w :: "'a::len word" @@ -1556,30 +1591,24 @@ begin private lemma sbintrunc_uint_ucast: assumes "Suc n = LENGTH('b::len)" - shows "sbintrunc n (uint (ucast w :: 'b word)) = sbintrunc n (uint w)" - by (metis assms sbintrunc_bintrunc ucast_eq word_ubin.eq_norm) + shows "signed_take_bit n (uint (ucast w :: 'b word)) = signed_take_bit n (uint w)" + by (rule bit_eqI) (use assms in \simp add: bit_simps\) private lemma test_bit_sbintrunc: assumes "i < LENGTH('a)" - shows "(word_of_int (sbintrunc n (uint w)) :: 'a word) !! i - = (if n < i then w !! n else w !! i)" - using assms by (simp add: nth_sbintr) - (simp add: test_bit_bin) + shows "bit (word_of_int (signed_take_bit n (uint w)) :: 'a word) i + = (if n < i then bit w n else bit w i)" + using assms by (simp add: bit_simps) private lemma test_bit_sbintrunc_ucast: assumes len_a: "i < LENGTH('a)" - shows "(word_of_int (sbintrunc (LENGTH('b) - 1) (uint (ucast w :: 'b word))) :: 'a word) !! i - = (if LENGTH('b::len) \ i then w !! (LENGTH('b) - 1) else w !! i)" - apply (subst sbintrunc_uint_ucast) - apply simp - apply (subst test_bit_sbintrunc) - apply (rule len_a) - apply (rule if_cong[OF _ refl refl]) - using leD less_linear by fastforce + shows "bit (word_of_int (signed_take_bit (LENGTH('b) - 1) (uint (ucast w :: 'b word))) :: 'a word) i + = (if LENGTH('b::len) \ i then bit w (LENGTH('b) - 1) else bit w i)" + using len_a by (auto simp add: sbintrunc_uint_ucast bit_simps) lemma scast_ucast_high_bits: \scast (ucast w :: 'b::len word) = w - \ (\ i \ {LENGTH('b) ..< size w}. w !! i = w !! (LENGTH('b) - 1))\ + \ (\ i \ {LENGTH('b) ..< size w}. bit w i = bit w (LENGTH('b) - 1))\ proof (cases \LENGTH('a) \ LENGTH('b)\) case True moreover define m where \m = LENGTH('b) - LENGTH('a)\ @@ -1633,15 +1662,18 @@ end lemma ucast_ucast_mask2: "is_down (UCAST ('a \ 'b)) \ UCAST ('b::len \ 'c::len) (UCAST ('a::len \ 'b::len) x) = UCAST ('a \ 'c) (x AND mask LENGTH('b))" - by word_eqI_solve + apply (simp flip: take_bit_eq_mask) + apply transfer + apply simp + done lemma ucast_NOT: "ucast (NOT x) = NOT(ucast x) AND mask (LENGTH('a))" for x::"'a::len word" - by word_eqI + by (rule bit_word_eqI) (auto simp add: bit_simps) lemma ucast_NOT_down: "is_down UCAST('a::len \ 'b::len) \ UCAST('a \ 'b) (NOT x) = NOT(UCAST('a \ 'b) x)" - by word_eqI + by (rule bit_word_eqI) (auto simp add: bit_simps is_down.rep_eq) lemma upto_enum_step_shift: "\ is_aligned p n \ \ @@ -1676,7 +1708,6 @@ lemma upto_enum_step_subset: apply simp done - lemma ucast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" @@ -1690,14 +1721,7 @@ lemma ucast_distrib: apply (simp only: ucast_eq) apply (subst lift_M) apply (subst of_int_uint [symmetric], subst lift_M') - apply (subst (1 2) int_word_uint) - apply (subst word_ubin.norm_eq_iff [symmetric]) - apply (subst (1 2) bintrunc_mod2p) - apply (insert is_down) - apply (unfold is_down_def) - apply (clarsimp simp: target_size source_size) - apply (clarsimp simp: mod_exp_eq min_def) - apply (rule distrib [symmetric]) + apply (metis local.distrib local.is_down take_bit_eq_mod ucast_down_wi uint_word_of_int_eq word_of_int_uint) done lemma ucast_down_add: @@ -1834,12 +1858,10 @@ proof (rule classical) using not_thesis by (clarsimp) - have result_range: "sint a sdiv sint b \ (sints (size a)) \ {2 ^ (size a - 1)}" - apply (cut_tac sdiv_int_range [where a="sint a" and b="sint b"]) - apply (erule rev_subsetD) - using sint_range' [where x=a] sint_range' [where x=b] - apply (auto simp: max_def abs_if word_size sints_num) - done + let ?range = \{- (2 ^ (size a - 1))..<2 ^ (size a - 1)} :: int set\ + + have result_range: "sint a sdiv sint b \ ?range \ {2 ^ (size a - 1)}" + using sdiv_word_min [of a b] sdiv_word_max [of a b] by auto have result_range_overflow: "(sint a sdiv sint b = 2 ^ (size a - 1)) = (?a_int_min \ ?b_minus1)" apply (rule iffI [rotated]) @@ -1854,14 +1876,15 @@ proof (rule classical) apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (clarsimp simp: word_size) apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] - apply (insert word_sint.Rep [where x="a"])[1] - apply (clarsimp simp: minus_le_iff word_size abs_if sints_num split: if_split_asm) - apply (metis minus_minus sint_int_min word_sint.Rep_inject) + apply auto + apply (cases \size a\) + apply simp_all + apply (smt (z3) One_nat_def diff_Suc_1 signed_word_eqI sint_int_min sint_range_size wsst_TYs(3)) done - have result_range_simple: "(sint a sdiv sint b \ (sints (size a))) \ ?thesis" + have result_range_simple: "(sint a sdiv sint b \ ?range) \ ?thesis" apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) - apply (clarsimp simp: word_size sints_num sint_int_min) + apply (clarsimp simp: word_size sint_int_min) done show ?thesis @@ -1889,8 +1912,7 @@ lemma signed_arith_ineq_checks_to_eq: = (sint a sdiv sint b = sint (a sdiv b))" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) = (sint a smod sint b = sint (a smod b))" - by (auto simp: sint_word_ariths word_size signed_div_arith signed_mod_arith - sbintrunc_eq_in_range range_sbintrunc) + by (auto simp: sint_word_ariths word_size signed_div_arith signed_mod_arith signed_take_bit_int_eq_self_iff intro: sym dest: sym) lemma signed_arith_sint: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) @@ -1907,4 +1929,42 @@ lemma signed_arith_sint: \ sint (a smod b) = (sint a smod sint b)" by (subst (asm) signed_arith_ineq_checks_to_eq; simp)+ +lemma nasty_split_lt: + \x * 2 ^ n + (2 ^ n - 1) \ 2 ^ m - 1\ + if \x < 2 ^ (m - n)\ \n \ m\ \m < LENGTH('a::len)\ + for x :: \'a::len word\ +proof - + define q where \q = m - n\ + with \n \ m\ have \m = q + n\ + by simp + with \x < 2 ^ (m - n)\ have *: \i < q\ if \bit x i\ for i + using that by simp (metis bit_take_bit_iff take_bit_word_eq_self_iff) + from \m = q + n\ have \push_bit n x OR mask n \ mask m\ + by (auto simp add: le_mask_high_bits word_size bit_simps dest!: *) + then have \push_bit n x + mask n \ mask m\ + by (simp add: disjunctive_add bit_simps) + then show ?thesis + by (simp add: mask_eq_exp_minus_1 push_bit_eq_mult) +qed + +lemma nasty_split_less: + "\m \ n; n \ nm; nm < LENGTH('a::len); x < 2 ^ (nm - n)\ + \ (x :: 'a word) * 2 ^ n + (2 ^ m - 1) < 2 ^ nm" + apply (simp only: word_less_sub_le[symmetric]) + apply (rule order_trans [OF _ nasty_split_lt]) + apply (rule word_plus_mono_right) + apply (rule word_sub_mono) + apply (simp add: word_le_nat_alt) + apply simp + apply (simp add: word_sub_1_le[OF power_not_zero]) + apply (simp add: word_sub_1_le[OF power_not_zero]) + apply (rule is_aligned_no_wrap') + apply (rule is_aligned_mult_triv2) + apply simp + apply (erule order_le_less_trans, simp) + apply simp+ + done + +end + end diff --git a/lib/Word_Lib/Word_Lib_Sumo.thy b/lib/Word_Lib/Word_Lib_Sumo.thy index 5d8c7b88b..7471754e9 100644 --- a/lib/Word_Lib/Word_Lib_Sumo.thy +++ b/lib/Word_Lib/Word_Lib_Sumo.thy @@ -10,8 +10,8 @@ theory Word_Lib_Sumo imports "HOL-Library.Word" Aligned - Ancient_Numeral Bit_Comprehension + Bit_Shifts_Infix_Syntax Bits_Int Bitwise_Signed Bitwise @@ -32,7 +32,7 @@ imports Reversed_Bit_Lists Rsplit Signed_Words - Traditional_Infix_Syntax + Syntax_Bundles Typedef_Morphisms Type_Syntax Word_EqI @@ -42,13 +42,15 @@ imports Word_32 Word_Syntax Signed_Division_Word + Singleton_Bit_Shifts More_Word_Operations Many_More - Word_Lemmas_Internal - Word_Lemmas_Prefix begin -declare word_induct2[case_names zero suc, induct type] +unbundle bit_operations_syntax +unbundle bit_projection_infix_syntax + +declare word_induct2[induct type] declare word_nat_cases[cases type] declare signed_take_bit_Suc [simp] @@ -58,13 +60,7 @@ lemmas of_int_and_nat = unsigned_of_nat unsigned_of_int signed_of_int signed_of_ bundle no_take_bit begin - declare of_int_and_nat[simp del] -end - -bundle no_0_dvd -begin - declare word_of_int_eq_0_iff[simp del] - declare word_of_nat_eq_0_iff[simp del] +declare of_int_and_nat[simp del] end lemmas bshiftr1_def = bshiftr1_eq @@ -103,11 +99,6 @@ lemmas word_prev_def = word_prev_unfold lemmas is_aligned_def = is_aligned_iff_dvd_nat -lemma shiftl_transfer [transfer_rule]: - includes lifting_syntax - shows "(pcr_word ===> (=) ===> pcr_word) (<<) (<<)" - by (unfold shiftl_eq_push_bit) transfer_prover - lemmas word_and_max_simps = word8_and_max_simp word16_and_max_simp @@ -132,7 +123,7 @@ declare of_nat_diff [simp] (* Haskellish names/syntax *) notation (input) - test_bit ("testBit") + bit ("testBit") lemmas cast_simps = cast_simps ucast_down_bl @@ -140,6 +131,6 @@ lemmas cast_simps = cast_simps ucast_down_bl lemma nth_ucast: "(ucast (w::'a::len word)::'b::len word) !! n = (w !! n \ n < min LENGTH('a) LENGTH('b))" - by transfer (simp add: bit_take_bit_iff ac_simps) + by (auto simp add: bit_simps not_le dest: bit_imp_le_length) end diff --git a/lib/Word_Lib/Word_Syntax.thy b/lib/Word_Lib/Word_Syntax.thy index 3edae8436..765cc5430 100644 --- a/lib/Word_Lib/Word_Syntax.thy +++ b/lib/Word_Lib/Word_Syntax.thy @@ -13,6 +13,10 @@ begin text \Additional bit and type syntax that forces word types.\ +context + includes bit_operations_syntax +begin + abbreviation wordNOT :: "'a::len word \ 'a word" ("~~ _" [70] 71) where @@ -34,3 +38,5 @@ where "a xor b == a XOR b" end + +end diff --git a/lib/Word_Lib/Word_Type_Syntax.thy b/lib/Word_Lib/Word_Type_Syntax.thy deleted file mode 100644 index 11cdde530..000000000 --- a/lib/Word_Lib/Word_Type_Syntax.thy +++ /dev/null @@ -1,58 +0,0 @@ -(* - * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) - * - * SPDX-License-Identifier: BSD-2-Clause - *) - -section "Displaying Phantom Types for Word Operations" - -theory Word_Type_Syntax -imports "HOL-Library.Word" -begin - -ML \ -structure Word_Syntax = -struct - - val show_word_types = Attrib.setup_config_bool @{binding show_word_types} (K true) - - fun tr' cnst ctxt typ ts = if Config.get ctxt show_word_types then - case (Term.binder_types typ, Term.body_type typ) of - ([Type (@{type_name "word"}, [S])], Type (@{type_name "word"}, [T])) => - list_comb - (Syntax.const cnst $ Syntax_Phases.term_of_typ ctxt S $ Syntax_Phases.term_of_typ ctxt T - , ts) - | _ => raise Match - else raise Match - -end -\ - - -syntax - "_Ucast" :: "type \ type \ logic" ("(1UCAST/(1'(_ \ _')))") -translations - "UCAST('s \ 't)" => "CONST ucast :: ('s word \ 't word)" -typed_print_translation - \ [(@{const_syntax ucast}, Word_Syntax.tr' @{syntax_const "_Ucast"})] \ - - -syntax - "_Scast" :: "type \ type \ logic" ("(1SCAST/(1'(_ \ _')))") -translations - "SCAST('s \ 't)" => "CONST scast :: ('s word \ 't word)" -typed_print_translation - \ [(@{const_syntax scast}, Word_Syntax.tr' @{syntax_const "_Scast"})] \ - - -syntax - "_Revcast" :: "type \ type \ logic" ("(1REVCAST/(1'(_ \ _')))") -translations - "REVCAST('s \ 't)" => "CONST revcast :: ('s word \ 't word)" -typed_print_translation - \ [(@{const_syntax revcast}, Word_Syntax.tr' @{syntax_const "_Revcast"})] \ - -(* Further candidates for showing word sizes, but with different arities: - slice, word_cat, word_split, word_rcat, word_rsplit *) - -end diff --git a/lib/Word_Lib/document/root.tex b/lib/Word_Lib/document/root.tex index 588c9c9b3..33f4eb24c 100644 --- a/lib/Word_Lib/document/root.tex +++ b/lib/Word_Lib/document/root.tex @@ -5,6 +5,7 @@ % \documentclass[11pt,a4paper]{article} +\usepackage[T1]{fontenc} \usepackage{isabelle,isabellesym} % this should be the last package used