# HG changeset patch # User haftmann # Date 1602953796 -7200 # Node ID ee659bca8955baf931be1c9a9618e58e1f1aedb1 # Parent ab32922f139bc9f1f52952f71c5104129ed9d7c3 factored out theory Bits_Int diff -r ab32922f139b -r ee659bca8955 src/HOL/Library/Bit_Operations.thy --- a/src/HOL/Library/Bit_Operations.thy Thu Oct 15 14:55:19 2020 +0200 +++ b/src/HOL/Library/Bit_Operations.thy Sat Oct 17 18:56:36 2020 +0200 @@ -9,6 +9,41 @@ Main begin +lemma sub_BitM_One_eq: + \(Num.sub (Num.BitM n) num.One) = 2 * (Num.sub n Num.One :: int)\ + by (cases n) simp_all + +lemma bit_not_int_iff': + \bit (- k - 1) n \ \ bit k n\ + for k :: int +proof (induction n arbitrary: k) + case 0 + show ?case + by simp +next + case (Suc n) + have \(- k - 1) div 2 = - (k div 2) - 1\ + by simp + with Suc show ?case + by (simp add: bit_Suc) +qed + +lemma bit_minus_int_iff: + \bit (- k) n \ \ bit (k - 1) n\ + for k :: int + using bit_not_int_iff' [of \k - 1\] by simp + +lemma bit_numeral_int_simps [simp]: + \bit (1 :: int) (numeral n) \ bit (0 :: int) (pred_numeral n)\ + \bit (numeral (num.Bit0 w) :: int) (numeral n) \ bit (numeral w :: int) (pred_numeral n)\ + \bit (numeral (num.Bit1 w) :: int) (numeral n) \ bit (numeral w :: int) (pred_numeral n)\ + \bit (numeral (Num.BitM w) :: int) (numeral n) \ \ bit (- numeral w :: int) (pred_numeral n)\ + \bit (- numeral (num.Bit0 w) :: int) (numeral n) \ bit (- numeral w :: int) (pred_numeral n)\ + \bit (- numeral (num.Bit1 w) :: int) (numeral n) \ \ bit (numeral w :: int) (pred_numeral n)\ + \bit (- numeral (Num.BitM w) :: int) (numeral n) \ bit (- (numeral w) :: int) (pred_numeral n)\ + by (simp_all add: bit_1_iff numeral_eq_Suc bit_Suc add_One sub_inc_One_eq bit_minus_int_iff) + + subsection \Bit operations\ class semiring_bit_operations = semiring_bit_shifts + @@ -574,8 +609,8 @@ lemma bit_not_int_iff: \bit (NOT k) n \ \ bit k n\ - for k :: int - by (induction n arbitrary: k) (simp_all add: not_int_div_2 even_not_iff_int bit_Suc) + for k :: int + by (simp add: bit_not_int_iff' not_int_def) function and_int :: \int \ int \ int\ where \(k::int) AND l = (if k \ {0, - 1} \ l \ {0, - 1} @@ -771,6 +806,112 @@ \k XOR l < 0 \ (k < 0) \ (l < 0)\ for k l :: int by (subst Not_eq_iff [symmetric]) (auto simp add: not_less) +lemma OR_upper: \<^marker>\contributor \Stefan Berghofer\\ + fixes x y :: int + assumes "0 \ x" "x < 2 ^ n" "y < 2 ^ n" + shows "x OR y < 2 ^ n" +using assms proof (induction x arbitrary: y n rule: int_bit_induct) + case zero + then show ?case + by simp +next + case minus + then show ?case + by simp +next + case (even x) + from even.IH [of \n - 1\ \y div 2\] even.prems even.hyps + show ?case + by (cases n) (auto simp add: or_int_rec [of \_ * 2\] elim: oddE) +next + case (odd x) + from odd.IH [of \n - 1\ \y div 2\] odd.prems odd.hyps + show ?case + by (cases n) (auto simp add: or_int_rec [of \1 + _ * 2\], linarith) +qed + +lemma XOR_upper: \<^marker>\contributor \Stefan Berghofer\\ + fixes x y :: int + assumes "0 \ x" "x < 2 ^ n" "y < 2 ^ n" + shows "x XOR y < 2 ^ n" +using assms proof (induction x arbitrary: y n rule: int_bit_induct) + case zero + then show ?case + by simp +next + case minus + then show ?case + by simp +next + case (even x) + from even.IH [of \n - 1\ \y div 2\] even.prems even.hyps + show ?case + by (cases n) (auto simp add: xor_int_rec [of \_ * 2\] elim: oddE) +next + case (odd x) + from odd.IH [of \n - 1\ \y div 2\] odd.prems odd.hyps + show ?case + by (cases n) (auto simp add: xor_int_rec [of \1 + _ * 2\]) +qed + +lemma AND_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ + fixes x y :: int + assumes "0 \ x" + shows "0 \ x AND y" + using assms by simp + +lemma OR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ + fixes x y :: int + assumes "0 \ x" "0 \ y" + shows "0 \ x OR y" + using assms by simp + +lemma XOR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ + fixes x y :: int + assumes "0 \ x" "0 \ y" + shows "0 \ x XOR y" + using assms by simp + +lemma AND_upper1 [simp]: \<^marker>\contributor \Stefan Berghofer\\ + fixes x y :: int + assumes "0 \ x" + shows "x AND y \ x" + using assms by (induction x arbitrary: y rule: int_bit_induct) + (simp_all add: and_int_rec [of \_ * 2\] and_int_rec [of \1 + _ * 2\] add_increasing) + +lemmas AND_upper1' [simp] = order_trans [OF AND_upper1] \<^marker>\contributor \Stefan Berghofer\\ +lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] \<^marker>\contributor \Stefan Berghofer\\ + +lemma AND_upper2 [simp]: \<^marker>\contributor \Stefan Berghofer\\ + fixes x y :: int + assumes "0 \ y" + shows "x AND y \ y" + using assms AND_upper1 [of y x] by (simp add: ac_simps) + +lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] \<^marker>\contributor \Stefan Berghofer\\ +lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] \<^marker>\contributor \Stefan Berghofer\\ + +lemma plus_and_or: \(x AND y) + (x OR y) = x + y\ for x y :: int +proof (induction x arbitrary: y rule: int_bit_induct) + case zero + then show ?case + by simp +next + case minus + then show ?case + by simp +next + case (even x) + from even.IH [of \y div 2\] + show ?case + by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE) +next + case (odd x) + from odd.IH [of \y div 2\] + show ?case + by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE) +qed + lemma set_bit_nonnegative_int_iff [simp]: \set_bit n k \ 0 \ k \ 0\ for k :: int by (simp add: set_bit_def) @@ -926,6 +1067,114 @@ end +text \FIXME: The rule sets below are very large (24 rules for each + operator). Is there a simpler way to do this?\ + +context +begin + +private lemma eqI: + \k = l\ + if num: \\n. bit k (numeral n) \ bit l (numeral n)\ + and even: \even k \ even l\ + for k l :: int +proof (rule bit_eqI) + fix n + show \bit k n \ bit l n\ + proof (cases n) + case 0 + with even show ?thesis + by simp + next + case (Suc n) + with num [of \num_of_nat (Suc n)\] show ?thesis + by (simp only: numeral_num_of_nat) + qed +qed + +lemma int_and_numerals [simp]: + "numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)" + "numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND numeral y)" + "numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)" + "numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND numeral y)" + "numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)" + "numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND - numeral (y + Num.One))" + "numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)" + "numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND - numeral (y + Num.One))" + "- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND numeral y)" + "- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND numeral y)" + "- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND numeral y)" + "- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND numeral y)" + "- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND - numeral y)" + "- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND - numeral (y + Num.One))" + "- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND - numeral y)" + "- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND - numeral (y + Num.One))" + "(1::int) AND numeral (Num.Bit0 y) = 0" + "(1::int) AND numeral (Num.Bit1 y) = 1" + "(1::int) AND - numeral (Num.Bit0 y) = 0" + "(1::int) AND - numeral (Num.Bit1 y) = 1" + "numeral (Num.Bit0 x) AND (1::int) = 0" + "numeral (Num.Bit1 x) AND (1::int) = 1" + "- numeral (Num.Bit0 x) AND (1::int) = 0" + "- numeral (Num.Bit1 x) AND (1::int) = 1" + by (auto simp add: bit_and_iff bit_minus_iff even_and_iff bit_double_iff even_bit_succ_iff add_One sub_inc_One_eq intro: eqI) + +lemma int_or_numerals [simp]: + "numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR numeral y)" + "numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)" + "numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR numeral y)" + "numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)" + "numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR - numeral y)" + "numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))" + "numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR - numeral y)" + "numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))" + "- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR numeral y)" + "- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR numeral y)" + "- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)" + "- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)" + "- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR - numeral y)" + "- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR - numeral (y + Num.One))" + "- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral y)" + "- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral (y + Num.One))" + "(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" + "(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)" + "(1::int) OR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)" + "(1::int) OR - numeral (Num.Bit1 y) = - numeral (Num.Bit1 y)" + "numeral (Num.Bit0 x) OR (1::int) = numeral (Num.Bit1 x)" + "numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)" + "- numeral (Num.Bit0 x) OR (1::int) = - numeral (Num.BitM x)" + "- numeral (Num.Bit1 x) OR (1::int) = - numeral (Num.Bit1 x)" + by (auto simp add: bit_or_iff bit_minus_iff even_or_iff bit_double_iff even_bit_succ_iff add_One sub_inc_One_eq sub_BitM_One_eq intro: eqI) + +lemma int_xor_numerals [simp]: + "numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR numeral y)" + "numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR numeral y)" + "numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR numeral y)" + "numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR numeral y)" + "numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR - numeral y)" + "numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR - numeral (y + Num.One))" + "numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR - numeral y)" + "numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR - numeral (y + Num.One))" + "- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR numeral y)" + "- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR numeral y)" + "- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR numeral y)" + "- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR numeral y)" + "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR - numeral y)" + "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR - numeral (y + Num.One))" + "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR - numeral y)" + "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR - numeral (y + Num.One))" + "(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" + "(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)" + "(1::int) XOR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)" + "(1::int) XOR - numeral (Num.Bit1 y) = - numeral (Num.Bit0 (y + Num.One))" + "numeral (Num.Bit0 x) XOR (1::int) = numeral (Num.Bit1 x)" + "numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)" + "- numeral (Num.Bit0 x) XOR (1::int) = - numeral (Num.BitM x)" + "- numeral (Num.Bit1 x) XOR (1::int) = - numeral (Num.Bit0 (x + Num.One))" + by (auto simp add: bit_xor_iff bit_minus_iff even_xor_iff bit_double_iff even_bit_succ_iff add_One sub_inc_One_eq sub_BitM_One_eq intro: eqI) + +end + subsection \Bit concatenation\ @@ -1009,6 +1258,10 @@ by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_concat_bit_iff min_def) +lemma concat_bit_take_bit_eq: + \concat_bit n (take_bit n b) = concat_bit n b\ + by (simp add: concat_bit_def [abs_def]) + subsection \Taking bits with sign propagation\ diff -r ab32922f139b -r ee659bca8955 src/HOL/SPARK/Manual/Reference.thy --- a/src/HOL/SPARK/Manual/Reference.thy Thu Oct 15 14:55:19 2020 +0200 +++ b/src/HOL/SPARK/Manual/Reference.thy Sat Oct 17 18:56:36 2020 +0200 @@ -1,6 +1,6 @@ (*<*) theory Reference -imports "HOL-SPARK.SPARK" +imports "HOL-SPARK.SPARK" "HOL-Word.Bits_Int" begin syntax (my_constrain output) diff -r ab32922f139b -r ee659bca8955 src/HOL/SPARK/SPARK.thy --- a/src/HOL/SPARK/SPARK.thy Thu Oct 15 14:55:19 2020 +0200 +++ b/src/HOL/SPARK/SPARK.thy Sat Oct 17 18:56:36 2020 +0200 @@ -39,13 +39,7 @@ lemma bit_not_spark_eq: "NOT (word_of_int x :: ('a::len) word) = word_of_int (2 ^ LENGTH('a) - 1 - x)" -proof - - have "word_of_int x + NOT (word_of_int x) = - word_of_int x + (word_of_int (2 ^ LENGTH('a) - 1 - x)::'a word)" - by (simp only: bwsimps bin_add_not Min_def) - (simp add: word_of_int_hom_syms word_of_int_2p_len) - then show ?thesis by (rule add_left_imp_eq) -qed + by (simp flip: mask_eq_exp_minus_1 add: of_int_mask_eq not_eq_complement) lemmas [simp] = bit_not_spark_eq [where 'a=8, simplified] diff -r ab32922f139b -r ee659bca8955 src/HOL/Word/Bit_Comprehension.thy --- a/src/HOL/Word/Bit_Comprehension.thy Thu Oct 15 14:55:19 2020 +0200 +++ b/src/HOL/Word/Bit_Comprehension.thy Sat Oct 17 18:56:36 2020 +0200 @@ -47,7 +47,8 @@ apply simp apply (rule bit_eqI) apply (simp add: bit_signed_take_bit_iff min_def) - using "*" by blast + apply (auto simp add: not_le bit_take_bit_iff dest: *) + done qed end @@ -223,15 +224,15 @@ qed lemma bin_last_set_bits [simp]: - "bin_last (set_bits f) = f 0" + "odd (set_bits f :: int) = f 0" by (subst int_set_bits_unfold_BIT) simp_all lemma bin_rest_set_bits [simp]: - "bin_rest (set_bits f) = set_bits (f \ Suc)" + "set_bits f div (2 :: int) = set_bits (f \ Suc)" by (subst int_set_bits_unfold_BIT) simp_all lemma bin_nth_set_bits [simp]: - "bin_nth (set_bits f) m = f m" + "bit (set_bits f :: int) m \ f m" using wff proof (induction m arbitrary: f) case 0 then show ?case diff -r ab32922f139b -r ee659bca8955 src/HOL/Word/Bits_Int.thy --- a/src/HOL/Word/Bits_Int.thy Thu Oct 15 14:55:19 2020 +0200 +++ b/src/HOL/Word/Bits_Int.thy Sat Oct 17 18:56:36 2020 +0200 @@ -8,15 +8,9 @@ imports "HOL-Library.Bit_Operations" Traditional_Syntax + Word begin -subsection \Generic auxiliary\ - -lemma int_mod_ge: "a < n \ 0 < n \ a \ a mod n" - for a n :: int - by (metis dual_order.trans le_cases mod_pos_pos_trivial pos_mod_conj) - - subsection \Implicit bit representation of \<^typ>\int\\ abbreviation (input) bin_last :: "int \ bool" @@ -1114,114 +1108,9 @@ "bin_last (- numeral (Num.BitM w))" by simp -(* FIXME: The rule sets below are very large (24 rules for each - operator). Is there a simpler way to do this? *) - -lemma int_and_numerals [simp]: - "numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)" - "numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND numeral y)" - "numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)" - "numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND numeral y)" - "numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)" - "numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND - numeral (y + Num.One))" - "numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)" - "numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND - numeral (y + Num.One))" - "- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND numeral y)" - "- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND numeral y)" - "- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND numeral y)" - "- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND numeral y)" - "- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND - numeral y)" - "- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND - numeral (y + Num.One))" - "- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND - numeral y)" - "- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND - numeral (y + Num.One))" - "(1::int) AND numeral (Num.Bit0 y) = 0" - "(1::int) AND numeral (Num.Bit1 y) = 1" - "(1::int) AND - numeral (Num.Bit0 y) = 0" - "(1::int) AND - numeral (Num.Bit1 y) = 1" - "numeral (Num.Bit0 x) AND (1::int) = 0" - "numeral (Num.Bit1 x) AND (1::int) = 1" - "- numeral (Num.Bit0 x) AND (1::int) = 0" - "- numeral (Num.Bit1 x) AND (1::int) = 1" - by (rule bin_rl_eqI; simp)+ - -lemma int_or_numerals [simp]: - "numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR numeral y)" - "numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)" - "numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR numeral y)" - "numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)" - "numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR - numeral y)" - "numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))" - "numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR - numeral y)" - "numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))" - "- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR numeral y)" - "- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR numeral y)" - "- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)" - "- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)" - "- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR - numeral y)" - "- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR - numeral (y + Num.One))" - "- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral y)" - "- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral (y + Num.One))" - "(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" - "(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)" - "(1::int) OR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)" - "(1::int) OR - numeral (Num.Bit1 y) = - numeral (Num.Bit1 y)" - "numeral (Num.Bit0 x) OR (1::int) = numeral (Num.Bit1 x)" - "numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)" - "- numeral (Num.Bit0 x) OR (1::int) = - numeral (Num.BitM x)" - "- numeral (Num.Bit1 x) OR (1::int) = - numeral (Num.Bit1 x)" - by (rule bin_rl_eqI; simp)+ - -lemma int_xor_numerals [simp]: - "numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR numeral y)" - "numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR numeral y)" - "numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR numeral y)" - "numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR numeral y)" - "numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR - numeral y)" - "numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR - numeral (y + Num.One))" - "numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR - numeral y)" - "numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR - numeral (y + Num.One))" - "- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR numeral y)" - "- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR numeral y)" - "- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR numeral y)" - "- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR numeral y)" - "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR - numeral y)" - "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR - numeral (y + Num.One))" - "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR - numeral y)" - "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR - numeral (y + Num.One))" - "(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" - "(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)" - "(1::int) XOR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)" - "(1::int) XOR - numeral (Num.Bit1 y) = - numeral (Num.Bit0 (y + Num.One))" - "numeral (Num.Bit0 x) XOR (1::int) = numeral (Num.Bit1 x)" - "numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)" - "- numeral (Num.Bit0 x) XOR (1::int) = - numeral (Num.BitM x)" - "- numeral (Num.Bit1 x) XOR (1::int) = - numeral (Num.Bit0 (x + Num.One))" - by (rule bin_rl_eqI; simp)+ - subsubsection \Interactions with arithmetic\ -lemma plus_and_or: "(x AND y) + (x OR y) = x + y" for x y :: int -proof (induction x arbitrary: y rule: int_bit_induct) - case zero - then show ?case - by simp -next - case minus - then show ?case - by simp -next - case (even x) - from even.IH [of \y div 2\] - show ?case - by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE) -next - case (odd x) - from odd.IH [of \y div 2\] - show ?case - by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE) -qed - lemma le_int_or: "bin_sign y = 0 \ x \ x OR y" for x y :: int by (simp add: bin_sign_def or_greater_eq split: if_splits) @@ -1238,106 +1127,18 @@ by (simp flip: take_bit_eq_mod add: take_bit_eq_mask mask_eq_exp_minus_1) -subsubsection \Comparison\ - -lemma AND_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ - fixes x y :: int - assumes "0 \ x" - shows "0 \ x AND y" - using assms by simp - -lemma OR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ - fixes x y :: int - assumes "0 \ x" "0 \ y" - shows "0 \ x OR y" - using assms by simp - -lemma XOR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ - fixes x y :: int - assumes "0 \ x" "0 \ y" - shows "0 \ x XOR y" - using assms by simp - -lemma AND_upper1 [simp]: \<^marker>\contributor \Stefan Berghofer\\ - fixes x y :: int - assumes "0 \ x" - shows "x AND y \ x" - using assms by (induction x arbitrary: y rule: int_bit_induct) - (simp_all add: and_int_rec [of \_ * 2\] and_int_rec [of \1 + _ * 2\] add_increasing) - -lemmas AND_upper1' [simp] = order_trans [OF AND_upper1] \<^marker>\contributor \Stefan Berghofer\\ -lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] \<^marker>\contributor \Stefan Berghofer\\ - -lemma AND_upper2 [simp]: \<^marker>\contributor \Stefan Berghofer\\ - fixes x y :: int - assumes "0 \ y" - shows "x AND y \ y" - using assms AND_upper1 [of y x] by (simp add: ac_simps) - -lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] \<^marker>\contributor \Stefan Berghofer\\ -lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] \<^marker>\contributor \Stefan Berghofer\\ - -lemma OR_upper: \<^marker>\contributor \Stefan Berghofer\\ - fixes x y :: int - assumes "0 \ x" "x < 2 ^ n" "y < 2 ^ n" - shows "x OR y < 2 ^ n" -using assms proof (induction x arbitrary: y n rule: int_bit_induct) - case zero - then show ?case - by simp -next - case minus - then show ?case - by simp -next - case (even x) - from even.IH [of \n - 1\ \y div 2\] even.prems even.hyps - show ?case - by (cases n) (auto simp add: or_int_rec [of \_ * 2\] elim: oddE) -next - case (odd x) - from odd.IH [of \n - 1\ \y div 2\] odd.prems odd.hyps - show ?case - by (cases n) (auto simp add: or_int_rec [of \1 + _ * 2\], linarith) -qed - -lemma XOR_upper: \<^marker>\contributor \Stefan Berghofer\\ - fixes x y :: int - assumes "0 \ x" "x < 2 ^ n" "y < 2 ^ n" - shows "x XOR y < 2 ^ n" -using assms proof (induction x arbitrary: y n rule: int_bit_induct) - case zero - then show ?case - by simp -next - case minus - then show ?case - by simp -next - case (even x) - from even.IH [of \n - 1\ \y div 2\] even.prems even.hyps - show ?case - by (cases n) (auto simp add: xor_int_rec [of \_ * 2\] elim: oddE) -next - case (odd x) - from odd.IH [of \n - 1\ \y div 2\] odd.prems odd.hyps - show ?case - by (cases n) (auto simp add: xor_int_rec [of \1 + _ * 2\]) -qed - - subsubsection \Truncating results of bit-wise operations\ lemma bin_trunc_ao: "bintrunc n x AND bintrunc n y = bintrunc n (x AND y)" "bintrunc n x OR bintrunc n y = bintrunc n (x OR y)" - by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) + by simp_all lemma bin_trunc_xor: "bintrunc n (bintrunc n x XOR bintrunc n y) = bintrunc n (x XOR y)" - by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) + by simp lemma bin_trunc_not: "bintrunc n (NOT (bintrunc n x)) = bintrunc n (NOT x)" - by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) + 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" @@ -1547,4 +1348,62 @@ "bin_sc n True i = i OR (1 << n)" by (induct n arbitrary: i) (rule bin_rl_eqI; simp)+ + +subsection \More lemmas on words\ + +lemma word_rcat_eq: + \word_rcat ws = word_of_int (bin_rcat (LENGTH('a::len)) (map uint ws))\ + for ws :: \'a::len word list\ + apply (simp add: word_rcat_def bin_rcat_def rev_map) + apply transfer + apply (simp add: horner_sum_foldr foldr_map comp_def) + done + +lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" + by (simp add: sign_Pls_ge_0) + +lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or + +\ \following definitions require both arithmetic and bit-wise word operations\ + +\ \to get \word_no_log_defs\ from \word_log_defs\, using \bin_log_bintrs\\ +lemmas wils1 = bin_log_bintrs [THEN word_of_int_eq_iff [THEN iffD2], + folded uint_word_of_int_eq, THEN eq_reflection] + +\ \the binary operations only\ (* BH: why is this needed? *) +lemmas word_log_binary_defs = + word_and_def word_or_def word_xor_def + +lemma setBit_no [simp]: "setBit (numeral bin) n = word_of_int (bin_sc n True (numeral bin))" + by transfer (simp add: bin_sc_eq) + +lemma clearBit_no [simp]: + "clearBit (numeral bin) n = word_of_int (bin_sc n False (numeral bin))" + by transfer (simp add: bin_sc_eq) + +lemma eq_mod_iff: "0 < n \ b = b mod n \ 0 \ b \ b < n" + for b n :: int + by auto (metis pos_mod_conj)+ + +lemma split_uint_lem: "bin_split n (uint w) = (a, b) \ + a = take_bit (LENGTH('a) - n) a \ b = take_bit (LENGTH('a)) b" + for w :: "'a::len word" + by transfer (simp add: drop_bit_take_bit ac_simps) + +\ \limited hom result\ +lemma word_cat_hom: + "LENGTH('a::len) \ LENGTH('b::len) + LENGTH('c::len) \ + (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = + word_of_int (bin_cat w (size b) (uint b))" + by transfer (simp add: take_bit_concat_bit_eq) + +lemma bintrunc_shiftl: + "take_bit n (m << i) = take_bit (n - i) m << i" + for m :: int + by (rule bit_eqI) (auto simp add: bit_take_bit_iff) + +lemma uint_shiftl: + "uint (n << i) = take_bit (size n) (uint n << i)" + by transfer (simp add: push_bit_take_bit shiftl_eq_push_bit) + end diff -r ab32922f139b -r ee659bca8955 src/HOL/Word/Misc_Typedef.thy --- a/src/HOL/Word/Misc_Typedef.thy Thu Oct 15 14:55:19 2020 +0200 +++ b/src/HOL/Word/Misc_Typedef.thy Sat Oct 17 18:56:36 2020 +0200 @@ -7,7 +7,7 @@ section \Type Definition Theorems\ theory Misc_Typedef - imports Main Word Bit_Comprehension + imports Main Word Bit_Comprehension Bits_Int begin subsection "More lemmas about normal type definitions" diff -r ab32922f139b -r ee659bca8955 src/HOL/Word/Reversed_Bit_Lists.thy --- a/src/HOL/Word/Reversed_Bit_Lists.thy Thu Oct 15 14:55:19 2020 +0200 +++ b/src/HOL/Word/Reversed_Bit_Lists.thy Sat Oct 17 18:56:36 2020 +0200 @@ -1684,7 +1684,7 @@ apply (simp add: word_split_def) apply transfer apply (cases \LENGTH('b) \ LENGTH('a)\) - apply (auto simp add: drop_bit_take_bit drop_bin2bl bin_to_bl_drop_bit [symmetric, of \LENGTH('a)\ \LENGTH('a) - LENGTH('b)\ \LENGTH('b)\]) + apply (auto simp add: drop_bit_take_bit drop_bin2bl bin_to_bl_drop_bit [symmetric, of \LENGTH('a)\ \LENGTH('a) - LENGTH('b)\ \LENGTH('b)\] min_absorb2) done lemma word_split_bl: "std = size c - size b \ diff -r ab32922f139b -r ee659bca8955 src/HOL/Word/Word.thy --- a/src/HOL/Word/Word.thy Thu Oct 15 14:55:19 2020 +0200 +++ b/src/HOL/Word/Word.thy Sat Oct 17 18:56:36 2020 +0200 @@ -9,7 +9,6 @@ "HOL-Library.Type_Length" "HOL-Library.Boolean_Algebra" "HOL-Library.Bit_Operations" - Bits_Int Traditional_Syntax begin @@ -295,8 +294,7 @@ lemma signed_1 [simp]: \signed (1 :: 'b::len word) = (if LENGTH('b) = 1 then - 1 else 1)\ - by (transfer fixing: uminus; cases \LENGTH('b)\) - (simp_all add: sbintrunc_minus_simps) + by (transfer fixing: uminus; cases \LENGTH('b)\) (auto dest: gr0_implies_Suc) lemma signed_minus_1 [simp]: \signed (- 1 :: 'b::len word) = - 1\ @@ -476,7 +474,7 @@ lemma [code]: \Word.the_signed_int w = signed_take_bit (LENGTH('a) - Suc 0) (Word.the_int w)\ for w :: \'a::len word\ - by transfer simp + by transfer (simp add: signed_take_bit_take_bit) lemma [code_abbrev, simp]: \Word.the_signed_int = sint\ @@ -1658,7 +1656,7 @@ by (fact word_coorder.extremum) lemma word_m1_ge [simp] : "word_pred 0 \ y" (* FIXME: delete *) - by transfer (simp add: take_bit_minus_one_eq_mask mask_eq_exp_minus_1 bintr_lt2p) + by transfer (simp add: take_bit_minus_one_eq_mask mask_eq_exp_minus_1 ) lemma word_n1_ge [simp]: "y \ -1" for y :: "'a::len word" @@ -1766,7 +1764,8 @@ lift_definition shiftr1 :: \'a::len word \ 'a word\ \ \shift right as unsigned or as signed, ie logical or arithmetic\ - is \\k. take_bit LENGTH('a) k div 2\ by simp + is \\k. take_bit LENGTH('a) k div 2\ + by simp lemma shiftr1_eq_div_2: \shiftr1 w = w div 2\ @@ -1884,14 +1883,14 @@ lemma shiftr_def: \w >> n = (shiftr1 ^^ n) w\ proof - - have \drop_bit n = (((\k::int. k div 2) ^^ n))\ for n - by (rule sym, induction n) - (simp_all add: fun_eq_iff drop_bit_Suc flip: drop_bit_half) + 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 - apply transfer - apply simp - apply (metis bintrunc_bintrunc rco_bintr) - done + by (simp add: shiftr_word_eq) qed lemma bit_shiftl_word_iff: @@ -2029,9 +2028,18 @@ apply (metis le_antisym less_eq_decr_length_iff) done +lemma signed_drop_bit_signed_drop_bit [simp]: + \signed_drop_bit m (signed_drop_bit n w) = signed_drop_bit (m + n) w\ + for w :: \'a::len word\ + apply (cases \LENGTH('a)\) + apply simp_all + apply (rule bit_word_eqI) + apply (auto simp add: bit_signed_drop_bit_iff not_le less_diff_conv ac_simps) + done + lemma signed_drop_bit_0 [simp]: \signed_drop_bit 0 w = w\ - by transfer simp + by transfer (simp add: take_bit_signed_take_bit) lemma sint_signed_drop_bit_eq: \sint (signed_drop_bit n w) = drop_bit n (sint w)\ @@ -2041,42 +2049,37 @@ apply (auto simp add: bit_sint_iff bit_drop_bit_eq bit_signed_drop_bit_iff dest: bit_imp_le_length) done -lift_definition sshiftr1 :: \'a::len word \ 'a word\ - is \\k. take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - 1) k div 2)\ +lift_definition sshiftr :: \'a::len word \ nat \ 'a word\ (infixl \>>>\ 55) + is \\k n. take_bit LENGTH('a) (drop_bit n (signed_take_bit (LENGTH('a) - Suc 0) k))\ by (simp flip: signed_take_bit_decr_length_iff) - -lift_definition sshiftr :: \'a::len word \ nat \ 'a word\ (infixl \>>>\ 55) - is \\k n. take_bit LENGTH('a) (drop_bit n (signed_take_bit (LENGTH('a) - 1) k))\ + +lift_definition sshiftr1 :: \'a::len word \ 'a word\ + is \\k. take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - Suc 0) k div 2)\ by (simp flip: signed_take_bit_decr_length_iff) lift_definition bshiftr1 :: \bool \ 'a::len word \ 'a word\ is \\b k. take_bit LENGTH('a) k div 2 + of_bool b * 2 ^ (LENGTH('a) - Suc 0)\ by (fact arg_cong) +lemma sshiftr_eq: + \w >>> n = signed_drop_bit n w\ + by transfer simp + +lemma sshiftr1_eq_signed_drop_bit_Suc_0: + \sshiftr1 = signed_drop_bit (Suc 0)\ + by (rule ext) (transfer, simp add: drop_bit_Suc) + lemma sshiftr1_eq: \sshiftr1 w = word_of_int (sint w div 2)\ by transfer simp lemma sshiftr_eq_funpow_sshiftr1: \w >>> n = (sshiftr1 ^^ n) w\ -proof - - have *: \(\k::int. take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - Suc 0) k div 2)) ^^ Suc n = - take_bit LENGTH('a) \ drop_bit (Suc n) \ signed_take_bit (LENGTH('a) - Suc 0)\ - for n - apply (induction n) - apply (auto simp add: fun_eq_iff drop_bit_Suc) - apply (metis (no_types, lifting) Suc_pred funpow_swap1 len_gt_0 sbintrunc_bintrunc sbintrunc_rest) - done - show ?thesis - apply transfer - apply simp - subgoal for k n - apply (cases n) - apply (simp_all only: *) - apply simp_all - done - done -qed + apply (rule sym) + apply (simp add: sshiftr1_eq_signed_drop_bit_Suc_0 sshiftr_eq) + apply (induction n) + apply simp_all + done lemma mask_eq: \mask n = (1 << n) - (1 :: 'a::len word)\ @@ -2260,29 +2263,20 @@ lemma word_cat_eq' [code]: \word_cat a b = word_of_int (concat_bit LENGTH('b) (uint b) (uint a))\ for a :: \'a::len word\ and b :: \'b::len word\ - by transfer simp + by transfer (simp add: concat_bit_take_bit_eq) lemma bit_word_cat_iff: \bit (word_cat v w :: 'c::len word) n \ n < LENGTH('c) \ (if n < LENGTH('b) then bit w n else bit v (n - LENGTH('b)))\ for v :: \'a::len word\ and w :: \'b::len word\ by transfer (simp add: bit_concat_bit_iff bit_take_bit_iff) -definition word_split :: "'a::len word \ 'b::len word \ 'c::len word" - where "word_split a = - (case bin_split (LENGTH('c)) (uint a) of - (u, v) \ (word_of_int u, word_of_int v))" +definition word_split :: \'a::len word \ 'b::len word \ 'c::len word\ + where \word_split w = + (ucast (drop_bit LENGTH('c) w) :: 'b::len word, ucast w :: 'c::len word)\ definition word_rcat :: \'a::len word list \ 'b::len word\ where \word_rcat = word_of_int \ horner_sum uint (2 ^ LENGTH('a)) \ rev\ -lemma word_rcat_eq: - \word_rcat ws = word_of_int (bin_rcat (LENGTH('a::len)) (map uint ws))\ - for ws :: \'a::len word list\ - apply (simp add: word_rcat_def bin_rcat_def rev_map) - apply transfer - apply (simp add: horner_sum_foldr foldr_map comp_def) - done - abbreviation (input) max_word :: \'a::len word\ \ \Largest representable machine integer.\ where "max_word \ - 1" @@ -2292,14 +2286,14 @@ lemma int_word_sint: \sint (word_of_int x :: 'a::len word) = (x + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)\ - by transfer (simp add: no_sbintr_alt2) + by transfer (simp flip: take_bit_eq_mod add: signed_take_bit_eq_take_bit_shift) lemma sint_sbintrunc': "sint (word_of_int bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) bin" by simp -lemma uint_sint: "uint w = take_bit (LENGTH('a)) (sint w)" +lemma uint_sint: "uint w = take_bit LENGTH('a) (sint w)" for w :: "'a::len word" - by transfer simp + by transfer (simp add: take_bit_signed_take_bit) lemma bintr_uint: "LENGTH('a) \ n \ take_bit n (uint w) = uint w" for w :: "'a::len word" @@ -2364,9 +2358,6 @@ for x :: "'a::len word" using sint_less [of x] by simp -lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" - by (simp add: sign_Pls_ge_0) - lemma uint_m2p_neg: "uint x - 2 ^ LENGTH('a) < 0" for x :: "'a::len word" by (simp only: diff_less_0_iff_less uint_lt2p) @@ -2377,7 +2368,7 @@ lemma lt2p_lem: "LENGTH('a) \ n \ uint w < 2 ^ n" for w :: "'a::len word" - by (metis bintr_lt2p bintr_uint) + using uint_bounded [of w] by (rule less_le_trans) simp lemma uint_le_0_iff [simp]: "uint x \ 0 \ uint x = 0" by (fact uint_ge_0 [THEN leD, THEN antisym_conv1]) @@ -2436,7 +2427,7 @@ by transfer simp lemma sint_range_size: "- (2 ^ (size w - Suc 0)) \ sint w \ sint w < 2 ^ (size w - Suc 0)" - by transfer (use sbintr_ge sbintr_lt in blast) + by (simp add: word_size sint_greater_eq sint_less) lemma sint_above_size: "2 ^ (size w - 1) \ x \ sint w < x" for w :: "'a::len word" @@ -2469,18 +2460,18 @@ for u v :: "'a::len word" by simp -lemma test_bit_bin': "w !! n \ n < size w \ bin_nth (uint w) n" +lemma test_bit_bin': "w !! n \ n < size w \ bit (uint w) n" by transfer (simp add: bit_take_bit_iff) lemmas test_bit_bin = test_bit_bin' [unfolded word_size] -lemma bin_nth_uint_imp: "bin_nth (uint w) n \ n < LENGTH('a)" +lemma bin_nth_uint_imp: "bit (uint w) n \ n < LENGTH('a)" for w :: "'a::len word" by transfer (simp add: bit_take_bit_iff) lemma bin_nth_sint: "LENGTH('a) \ n \ - bin_nth (sint w) n = bin_nth (sint w) (LENGTH('a) - 1)" + bit (sint w) n = bit (sint w) (LENGTH('a) - 1)" for w :: "'a::len word" by (transfer fixing: n) (simp add: bit_signed_take_bit_iff le_diff_conv min_def) @@ -2503,7 +2494,7 @@ then have \take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - 1) (numeral a :: int)) = take_bit LENGTH('a) (numeral b)\ by simp then show \take_bit LENGTH('a) (numeral a :: int) = take_bit LENGTH('a) (numeral b)\ - by simp + by (simp add: take_bit_signed_take_bit) qed lemma num_abs_bintr: @@ -2514,7 +2505,7 @@ lemma num_abs_sbintr: "(numeral x :: 'a word) = word_of_int (signed_take_bit (LENGTH('a::len) - 1) (numeral x))" - by transfer simp + by transfer (simp add: take_bit_signed_take_bit) text \ \cast\ -- note, no arg for new length, as it's determined by type of result, @@ -2529,7 +2520,7 @@ by transfer simp lemma scast_id [simp]: "scast w = w" - by transfer simp + by transfer (simp add: take_bit_signed_take_bit) lemma nth_ucast: "(ucast w::'a::len word) !! n = (w !! n \ n < LENGTH('a))" by transfer (simp add: bit_take_bit_iff ac_simps) @@ -2743,16 +2734,13 @@ and "sint (word_pred a) = signed_take_bit (LENGTH('a) - 1) (sint a - 1)" and "sint (0 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 0" and "sint (1 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 1" - prefer 8 - apply (simp add: Suc_lessI sbintrunc_minus_simps(3)) - prefer 7 - apply simp - apply transfer apply (simp add: signed_take_bit_add) - apply transfer apply (simp add: signed_take_bit_diff) - apply transfer apply (simp add: signed_take_bit_mult) - apply transfer apply (simp add: signed_take_bit_minus) - apply (metis One_nat_def id_apply of_int_eq_id sbintrunc_sbintrunc signed.rep_eq signed_word_eqI sint_sbintrunc' wi_hom_succ) - apply (metis (no_types, lifting) One_nat_def signed_take_bit_decr_length_iff sint_uint uint_sint uint_word_of_int_eq wi_hom_pred word_of_int_uint) + apply transfer apply (simp add: signed_take_bit_add) + apply transfer apply (simp add: signed_take_bit_diff) + apply transfer apply (simp add: signed_take_bit_mult) + apply transfer apply (simp add: signed_take_bit_minus) + apply (metis of_int_sint scast_id sint_sbintrunc' wi_hom_succ) + apply (metis of_int_sint scast_id sint_sbintrunc' wi_hom_pred) + apply (simp_all add: sint_uint) done lemma word_pred_0_n1: "word_pred 0 = word_of_int (- 1)" @@ -2937,18 +2925,31 @@ unfolding uint_word_ariths by (simp add: zmod_le_nonneg_dividend) lemma uint_sub_ge: "uint (x - y) \ uint x - uint y" - unfolding uint_word_ariths by (simp add: int_mod_ge) - + unfolding uint_word_ariths + by (simp flip: take_bit_eq_mod add: take_bit_int_greater_eq_self_iff) + +lemma int_mod_ge: \a \ a mod n\ if \a < n\ \0 < n\ + for a n :: int +proof (cases \a < 0\) + case True + with \0 < n\ show ?thesis + by (metis less_trans not_less pos_mod_conj) + +next + case False + with \a < n\ show ?thesis + by simp +qed + lemma mod_add_if_z: "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ (x + y) mod z = (if x + y < z then x + y else x + y - z)" for x y z :: int apply (auto simp add: not_less) apply (rule antisym) - apply (metis diff_ge_0_iff_ge minus_mod_self2 zmod_le_nonneg_dividend) - apply (simp only: flip: minus_mod_self2 [of \x + y\ z]) - apply (rule int_mod_ge) - apply simp_all + apply (metis diff_ge_0_iff_ge minus_mod_self2 zmod_le_nonneg_dividend) + apply (simp only: flip: minus_mod_self2 [of \x + y\ z]) + apply (metis add.commute add_less_cancel_left add_mono_thms_linordered_field(5) diff_add_cancel diff_ge_0_iff_ge mod_pos_pos_trivial order_refl) done lemma uint_plus_if': @@ -3593,16 +3594,13 @@ by (fact inj_unsigned) lemma range_uint: \range (uint :: 'a word \ int) = {0..<2 ^ LENGTH('a::len)}\ - by transfer (auto simp add: bintr_lt2p range_bintrunc) + apply transfer + apply (auto simp add: image_iff) + apply (metis take_bit_int_eq_self_iff) + done lemma UNIV_eq: \(UNIV :: 'a word set) = word_of_int ` {0..<2 ^ LENGTH('a::len)}\ -proof - - have \uint ` (UNIV :: 'a word set) = uint ` (word_of_int :: int \ 'a word) ` {0..<2 ^ LENGTH('a::len)}\ - by transfer (simp add: range_bintrunc, auto simp add: take_bit_eq_mod) - then show ?thesis - using inj_image_eq_iff [of \uint :: 'a word \ int\ \UNIV :: 'a word set\ \word_of_int ` {0..<2 ^ LENGTH('a)} :: 'a word set\, OF inj_uint] - by simp -qed + by (auto simp add: image_iff) (metis atLeastLessThan_iff linorder_not_le uint_split) lemma card_word: "CARD('a word) = 2 ^ LENGTH('a::len)" by (simp add: UNIV_eq card_image inj_on_word_of_int) @@ -3617,18 +3615,6 @@ subsection \Bitwise Operations on Words\ -lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or - -\ \following definitions require both arithmetic and bit-wise word operations\ - -\ \to get \word_no_log_defs\ from \word_log_defs\, using \bin_log_bintrs\\ -lemmas wils1 = bin_log_bintrs [THEN word_of_int_eq_iff [THEN iffD2], - folded uint_word_of_int_eq, THEN eq_reflection] - -\ \the binary operations only\ (* BH: why is this needed? *) -lemmas word_log_binary_defs = - word_and_def word_or_def word_xor_def - lemma word_wi_log_defs: "NOT (word_of_int a) = word_of_int (NOT a)" "word_of_int a AND word_of_int b = word_of_int (a AND b)" @@ -3701,7 +3687,7 @@ by (simp only: test_bit_eq_bit) transfer_prover lemma test_bit_wi [simp]: - "(word_of_int x :: 'a::len word) !! n \ n < LENGTH('a) \ bin_nth x n" + "(word_of_int x :: 'a::len word) !! n \ n < LENGTH('a) \ bit x n" by transfer simp lemma word_ops_nth_size: @@ -3711,29 +3697,29 @@ (x XOR y) !! n = (x !! n \ y !! n) \ (NOT x) !! n = (\ x !! n)" for x :: "'a::len word" - unfolding word_size by transfer (simp add: bin_nth_ops) + 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: bin_nth_ops) + 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) \ bin_nth (numeral w) 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) \ bin_nth (- numeral w) n" + n < LENGTH('a) \ bit (- numeral w :: int) n" by transfer (rule refl) lemma test_bit_1 [simp]: "(1 :: 'a::len word) !! n \ n = 0" - by transfer auto + by transfer (auto simp add: bit_1_iff) lemma nth_0 [simp]: "\ (0 :: 'a::len word) !! n" by transfer simp @@ -3822,7 +3808,7 @@ lemma word_add_not [simp]: "x + NOT x = -1" for x :: "'a::len word" - by transfer (simp add: bin_add_not) + by transfer (simp add: not_int_def) lemma word_plus_and_or [simp]: "(x AND y) + (x OR y) = x + y" for x :: "'a::len word" @@ -3842,7 +3828,7 @@ lemma le_word_or2: "x \ x OR y" for x y :: "'a::len word" - by (auto simp: word_le_def uint_or intro: le_int_or) + by (simp add: or_greater_eq uint_or word_le_def) lemmas le_word_or1 = xtrans(3) [OF word_bw_comms (2) le_word_or2] lemmas word_and_le1 = xtrans(3) [OF word_ao_absorbs (4) [symmetric] le_word_or2] @@ -3877,16 +3863,9 @@ lemma nth_sint: fixes w :: "'a::len word" defines "l \ LENGTH('a)" - shows "bin_nth (sint w) n = (if n < l - 1 then w !! n else w !! (l - 1))" + shows "bit (sint w) n = (if n < l - 1 then w !! n else w !! (l - 1))" unfolding sint_uint l_def - by (auto simp: nth_sbintr word_test_bit_def [symmetric]) - -lemma setBit_no [simp]: "setBit (numeral bin) n = word_of_int (bin_sc n True (numeral bin))" - by transfer (simp add: bin_sc_eq) - -lemma clearBit_no [simp]: - "clearBit (numeral bin) n = word_of_int (bin_sc n False (numeral bin))" - by transfer (simp add: bin_sc_eq) + 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) @@ -3973,12 +3952,13 @@ text \ see paper page 10, (1), (2), \shiftr1_def\ is of the form of (1), - where \f\ (ie \bin_rest\) takes normal arguments to normal results, + where \f\ (ie \_ div 2\) takes normal arguments to normal results, thus we get (2) from (1) \ -lemma uint_shiftr1: "uint (shiftr1 w) = bin_rest (uint w)" - by transfer simp +lemma uint_shiftr1: "uint (shiftr1 w) = uint w div 2" + using uint_shiftr_eq [of w 1] + by (simp add: shiftr1_code) lemma bit_sshiftr1_iff: \bit (sshiftr1 w) n \ bit w (if n = LENGTH('a) - 1 then LENGTH('a) - 1 else Suc n)\ @@ -3998,10 +3978,6 @@ using le_less_Suc_eq apply fastforce done -lemma sshiftr_eq: - \w >>> m = signed_drop_bit m w\ - by (rule bit_eqI) (simp add: bit_signed_drop_bit_iff bit_sshiftr_word_iff) - lemma nth_sshiftr1: "sshiftr1 w !! n = (if n = size w - 1 then w !! n else w !! Suc n)" apply transfer apply (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def simp flip: bit_Suc) @@ -4022,20 +3998,15 @@ by (fact uint_shiftr1) lemma sshiftr1_div_2: "sint (sshiftr1 w) = sint w div 2" - by transfer simp + using sint_signed_drop_bit_eq [of 1 w] + by (simp add: drop_bit_Suc sshiftr1_eq_signed_drop_bit_Suc_0) lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" - apply (unfold shiftr_def) - apply (induct n) - apply simp - apply (simp add: shiftr1_div_2 mult.commute zdiv_zmult2_eq [symmetric]) - done + by (fact uint_shiftr_eq) lemma sshiftr_div_2n: "sint (sshiftr w n) = sint w div 2 ^ n" - apply transfer - apply (rule bit_eqI) - apply (simp add: bit_signed_take_bit_iff bit_drop_bit_eq min_def flip: drop_bit_eq_div) - done + using sint_signed_drop_bit_eq [of n w] + by (simp add: drop_bit_eq_div sshiftr_eq) lemma bit_bshiftr1_iff: \bit (bshiftr1 b w) n \ b \ n = LENGTH('a) - 1 \ bit w (Suc n)\ @@ -4092,12 +4063,12 @@ lemma shiftr1_bintr [simp]: "(shiftr1 (numeral w) :: 'a::len word) = - word_of_int (bin_rest (take_bit (LENGTH('a)) (numeral w)))" + word_of_int (take_bit LENGTH('a) (numeral w) div 2)" by transfer simp lemma sshiftr1_sbintr [simp]: "(sshiftr1 (numeral w) :: 'a::len word) = - word_of_int (bin_rest (signed_take_bit (LENGTH('a) - 1) (numeral w)))" + word_of_int (signed_take_bit (LENGTH('a) - 1) (numeral w) div 2)" by transfer simp text \TODO: rules for \<^term>\- (numeral n)\\ @@ -4116,7 +4087,7 @@ word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (numeral k)))\ apply (rule word_eqI) apply (cases \LENGTH('a)\) - apply (simp_all add: word_size bit_drop_bit_eq nth_sshiftr nth_sbintr not_le not_less less_Suc_eq_le ac_simps) + apply (simp_all add: word_size bit_drop_bit_eq nth_sshiftr bit_signed_take_bit_iff min_def not_le not_less less_Suc_eq_le ac_simps) done lemma zip_replicate: "n \ length ys \ zip (replicate n x) ys = map (\y. (x, y)) ys" @@ -4185,20 +4156,17 @@ by (auto simp add: test_bit_word_eq word_size Word.bit_mask_iff) lemma mask_bin: "mask n = word_of_int (take_bit n (- 1))" - by (auto simp add: nth_bintr word_size intro: word_eqI) + by transfer (simp add: take_bit_minus_one_eq_mask) lemma and_mask_bintr: "w AND mask n = word_of_int (take_bit n (uint w))" - apply (rule word_eqI) - apply (simp add: nth_bintr word_size word_ops_nth_size) - apply (auto simp add: test_bit_bin) - done + by transfer (simp add: ac_simps take_bit_eq_mask) lemma and_mask_wi: "word_of_int i AND mask n = word_of_int (take_bit n i)" - by (auto simp add: nth_bintr word_size word_ops_nth_size word_eq_iff) + by (auto simp add: and_mask_bintr min_def not_le wi_bintr) lemma and_mask_wi': "word_of_int i AND mask n = (word_of_int (take_bit (min LENGTH('a) n) i) :: 'a::len word)" - by (auto simp add: nth_bintr word_size word_ops_nth_size word_eq_iff) + by (auto simp add: and_mask_wi min_def wi_bintr) lemma and_mask_no: "numeral i AND mask n = word_of_int (take_bit n (numeral i))" unfolding word_numeral_alt by (rule and_mask_wi) @@ -4211,16 +4179,10 @@ by transfer simp lemma and_mask_lt_2p: "uint (w AND mask n) < 2 ^ n" - apply (simp add: uint_and uint_mask_eq) - apply (rule AND_upper2'') - apply simp - apply (auto simp add: mask_eq_exp_minus_1 min_def power_add algebra_simps dest!: le_Suc_ex) - apply (metis add_minus_cancel le_add2 one_le_numeral power_add power_increasing uminus_add_conv_diff zle_diff1_eq) - done - -lemma eq_mod_iff: "0 < n \ b = b mod n \ 0 \ b \ b < n" - for b n :: int - by auto (metis pos_mod_conj)+ + apply (simp flip: take_bit_eq_mask) + apply transfer + apply (auto simp add: min_def) + using antisym_conv take_bit_int_eq_self_iff by fastforce lemma mask_eq_iff: "w AND mask n = w \ uint w < 2 ^ n" apply (auto simp flip: take_bit_eq_mask) @@ -4240,11 +4202,11 @@ lemma less_mask_eq: "x < 2 ^ n \ x AND mask n = x" for x :: "'a::len word" - apply (simp add: and_mask_bintr) + apply (cases \n < LENGTH('a)\) + apply (simp_all add: not_less flip: take_bit_eq_mask exp_eq_zero_iff) apply transfer - apply (simp add: ac_simps) - apply (auto simp add: min_def) - apply (metis bintrunc_bintrunc_ge mod_pos_pos_trivial mult.commute mult.left_neutral mult_zero_left not_le of_bool_def take_bit_eq_mod take_bit_nonnegative) + apply (simp add: min_def) + apply (metis min_def nat_less_le take_bit_int_eq_self_iff take_bit_take_bit) done lemmas mask_eq_iff_w2p = trans [OF mask_eq_iff word_2p_lem [symmetric]] @@ -4544,27 +4506,15 @@ apply (auto simp add: bit_concat_bit_iff bit_take_bit_iff) done -lemma split_uint_lem: "bin_split n (uint w) = (a, b) \ - a = take_bit (LENGTH('a) - n) a \ b = take_bit (LENGTH('a)) b" - for w :: "'a::len word" - by transfer (simp add: drop_bit_take_bit ac_simps) - \ \keep quantifiers for use in simplification\ lemma test_bit_split': "word_split c = (a, b) \ (\n m. b !! n = (n < size b \ c !! n) \ a !! m = (m < size a \ c !! (m + size b)))" - apply (unfold word_split_bin' test_bit_bin) - apply (clarify) - apply simp - apply (erule conjE) - apply (drule sym) - apply (drule sym) - apply (simp add: bit_take_bit_iff bit_drop_bit_eq) - apply transfer - apply (simp add: bit_take_bit_iff ac_simps) - done + by (auto simp add: word_split_bin' test_bit_bin bit_unsigned_iff word_size + bit_drop_bit_eq ac_simps exp_eq_zero_iff + dest: bit_imp_le_length) lemma test_bit_split: "word_split c = (a, b) \ @@ -4590,15 +4540,7 @@ result to the length given by the result type\ lemma word_cat_id: "word_cat a b = b" - by transfer simp - -\ \limited hom result\ -lemma word_cat_hom: - "LENGTH('a::len) \ LENGTH('b::len) + LENGTH('c::len) \ - (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = - word_of_int (bin_cat w (size b) (uint b))" - apply transfer - using bintr_cat by auto + by transfer (simp add: take_bit_concat_bit_eq) lemma word_cat_split_alt: "size w \ size u + size v \ word_split w = (u, v) \ word_cat u v = w" apply (rule word_eqI) @@ -4675,7 +4617,7 @@ "sw = size (hd wl) \ rc = word_rcat wl \ rc !! n = (n < size rc \ n div sw < size wl \ (rev wl) ! (n div sw) !! (n mod sw))" for wl :: "'a::len word list" - by (simp add: word_size word_rcat_def bin_rcat_def foldl_map rev_map bit_horner_sum_uint_exp_iff) + 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] @@ -4932,11 +4874,11 @@ by transfer simp with that show ?thesis apply transfer - apply simp + apply simp apply (subst take_bit_diff [symmetric]) apply (subst nat_take_bit_eq) - apply (simp add: nat_le_eq_zle) - apply (simp add: nat_diff_distrib take_bit_nat_eq_self_iff less_imp_diff_less bintr_lt2p) + apply (simp add: nat_le_eq_zle) + apply (simp add: nat_diff_distrib take_bit_nat_eq_self_iff less_imp_diff_less) done qed @@ -5145,22 +5087,13 @@ lemma mask_Suc_0: "mask (Suc 0) = 1" using mask_1 by simp -lemma bin_last_bintrunc: "bin_last (take_bit l n) = (l > 0 \ bin_last n)" +lemma bin_last_bintrunc: "odd (take_bit l n) \ l > 0 \ odd n" by simp lemma word_and_1: "n AND 1 = (if n !! 0 then 1 else 0)" for n :: "_ word" by (rule bit_word_eqI) (auto simp add: bit_and_iff test_bit_eq_bit bit_1_iff intro: gr0I) -lemma bintrunc_shiftl: - "take_bit n (m << i) = take_bit (n - i) m << i" - for m :: int - by (rule bit_eqI) (auto simp add: bit_take_bit_iff) - -lemma uint_shiftl: - "uint (n << i) = take_bit (size n) (uint n << i)" - by transfer (simp add: push_bit_take_bit shiftl_eq_push_bit) - subsection \Misc\