src/HOL/Library/Bit_Operations.thy
changeset 74101 d804e93ae9ff
parent 74100 fb9c119e5b49
child 74102 0572c733d12d
equal deleted inserted replaced
74100:fb9c119e5b49 74101:d804e93ae9ff
     1 (*  Author:  Florian Haftmann, TUM
       
     2 *)
       
     3 
       
     4 section \<open>Bit operations in suitable algebraic structures\<close>
       
     5 
       
     6 theory Bit_Operations
       
     7   imports
       
     8     Main
       
     9     "HOL-Library.Boolean_Algebra"
       
    10 begin
       
    11 
       
    12 subsection \<open>Bit operations\<close>
       
    13 
       
    14 class semiring_bit_operations = semiring_bit_shifts +
       
    15   fixes "and" :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr \<open>AND\<close> 64)
       
    16     and or :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr \<open>OR\<close>  59)
       
    17     and xor :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr \<open>XOR\<close> 59)
       
    18     and mask :: \<open>nat \<Rightarrow> 'a\<close>
       
    19     and set_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
       
    20     and unset_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
       
    21     and flip_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
       
    22   assumes bit_and_iff [bit_simps]: \<open>bit (a AND b) n \<longleftrightarrow> bit a n \<and> bit b n\<close>
       
    23     and bit_or_iff [bit_simps]: \<open>bit (a OR b) n \<longleftrightarrow> bit a n \<or> bit b n\<close>
       
    24     and bit_xor_iff [bit_simps]: \<open>bit (a XOR b) n \<longleftrightarrow> bit a n \<noteq> bit b n\<close>
       
    25     and mask_eq_exp_minus_1: \<open>mask n = 2 ^ n - 1\<close>
       
    26     and set_bit_eq_or: \<open>set_bit n a = a OR push_bit n 1\<close>
       
    27     and bit_unset_bit_iff [bit_simps]: \<open>bit (unset_bit m a) n \<longleftrightarrow> bit a n \<and> m \<noteq> n\<close>
       
    28     and flip_bit_eq_xor: \<open>flip_bit n a = a XOR push_bit n 1\<close>
       
    29 begin
       
    30 
       
    31 text \<open>
       
    32   We want the bitwise operations to bind slightly weaker
       
    33   than \<open>+\<close> and \<open>-\<close>.
       
    34   For the sake of code generation
       
    35   the operations \<^const>\<open>and\<close>, \<^const>\<open>or\<close> and \<^const>\<open>xor\<close>
       
    36   are specified as definitional class operations.
       
    37 \<close>
       
    38 
       
    39 sublocale "and": semilattice \<open>(AND)\<close>
       
    40   by standard (auto simp add: bit_eq_iff bit_and_iff)
       
    41 
       
    42 sublocale or: semilattice_neutr \<open>(OR)\<close> 0
       
    43   by standard (auto simp add: bit_eq_iff bit_or_iff)
       
    44 
       
    45 sublocale xor: comm_monoid \<open>(XOR)\<close> 0
       
    46   by standard (auto simp add: bit_eq_iff bit_xor_iff)
       
    47 
       
    48 lemma even_and_iff:
       
    49   \<open>even (a AND b) \<longleftrightarrow> even a \<or> even b\<close>
       
    50   using bit_and_iff [of a b 0] by auto
       
    51 
       
    52 lemma even_or_iff:
       
    53   \<open>even (a OR b) \<longleftrightarrow> even a \<and> even b\<close>
       
    54   using bit_or_iff [of a b 0] by auto
       
    55 
       
    56 lemma even_xor_iff:
       
    57   \<open>even (a XOR b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)\<close>
       
    58   using bit_xor_iff [of a b 0] by auto
       
    59 
       
    60 lemma zero_and_eq [simp]:
       
    61   \<open>0 AND a = 0\<close>
       
    62   by (simp add: bit_eq_iff bit_and_iff)
       
    63 
       
    64 lemma and_zero_eq [simp]:
       
    65   \<open>a AND 0 = 0\<close>
       
    66   by (simp add: bit_eq_iff bit_and_iff)
       
    67 
       
    68 lemma one_and_eq:
       
    69   \<open>1 AND a = a mod 2\<close>
       
    70   by (simp add: bit_eq_iff bit_and_iff) (auto simp add: bit_1_iff)
       
    71 
       
    72 lemma and_one_eq:
       
    73   \<open>a AND 1 = a mod 2\<close>
       
    74   using one_and_eq [of a] by (simp add: ac_simps)
       
    75 
       
    76 lemma one_or_eq:
       
    77   \<open>1 OR a = a + of_bool (even a)\<close>
       
    78   by (simp add: bit_eq_iff bit_or_iff add.commute [of _ 1] even_bit_succ_iff) (auto simp add: bit_1_iff)
       
    79 
       
    80 lemma or_one_eq:
       
    81   \<open>a OR 1 = a + of_bool (even a)\<close>
       
    82   using one_or_eq [of a] by (simp add: ac_simps)
       
    83 
       
    84 lemma one_xor_eq:
       
    85   \<open>1 XOR a = a + of_bool (even a) - of_bool (odd a)\<close>
       
    86   by (simp add: bit_eq_iff bit_xor_iff add.commute [of _ 1] even_bit_succ_iff) (auto simp add: bit_1_iff odd_bit_iff_bit_pred elim: oddE)
       
    87 
       
    88 lemma xor_one_eq:
       
    89   \<open>a XOR 1 = a + of_bool (even a) - of_bool (odd a)\<close>
       
    90   using one_xor_eq [of a] by (simp add: ac_simps)
       
    91 
       
    92 lemma take_bit_and [simp]:
       
    93   \<open>take_bit n (a AND b) = take_bit n a AND take_bit n b\<close>
       
    94   by (auto simp add: bit_eq_iff bit_take_bit_iff bit_and_iff)
       
    95 
       
    96 lemma take_bit_or [simp]:
       
    97   \<open>take_bit n (a OR b) = take_bit n a OR take_bit n b\<close>
       
    98   by (auto simp add: bit_eq_iff bit_take_bit_iff bit_or_iff)
       
    99 
       
   100 lemma take_bit_xor [simp]:
       
   101   \<open>take_bit n (a XOR b) = take_bit n a XOR take_bit n b\<close>
       
   102   by (auto simp add: bit_eq_iff bit_take_bit_iff bit_xor_iff)
       
   103 
       
   104 lemma push_bit_and [simp]:
       
   105   \<open>push_bit n (a AND b) = push_bit n a AND push_bit n b\<close>
       
   106   by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_and_iff)
       
   107 
       
   108 lemma push_bit_or [simp]:
       
   109   \<open>push_bit n (a OR b) = push_bit n a OR push_bit n b\<close>
       
   110   by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_or_iff)
       
   111 
       
   112 lemma push_bit_xor [simp]:
       
   113   \<open>push_bit n (a XOR b) = push_bit n a XOR push_bit n b\<close>
       
   114   by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_xor_iff)
       
   115 
       
   116 lemma drop_bit_and [simp]:
       
   117   \<open>drop_bit n (a AND b) = drop_bit n a AND drop_bit n b\<close>
       
   118   by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_and_iff)
       
   119 
       
   120 lemma drop_bit_or [simp]:
       
   121   \<open>drop_bit n (a OR b) = drop_bit n a OR drop_bit n b\<close>
       
   122   by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_or_iff)
       
   123 
       
   124 lemma drop_bit_xor [simp]:
       
   125   \<open>drop_bit n (a XOR b) = drop_bit n a XOR drop_bit n b\<close>
       
   126   by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_xor_iff)
       
   127 
       
   128 lemma bit_mask_iff [bit_simps]:
       
   129   \<open>bit (mask m) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n < m\<close>
       
   130   by (simp add: mask_eq_exp_minus_1 bit_mask_iff)
       
   131 
       
   132 lemma even_mask_iff:
       
   133   \<open>even (mask n) \<longleftrightarrow> n = 0\<close>
       
   134   using bit_mask_iff [of n 0] by auto
       
   135 
       
   136 lemma mask_0 [simp]:
       
   137   \<open>mask 0 = 0\<close>
       
   138   by (simp add: mask_eq_exp_minus_1)
       
   139 
       
   140 lemma mask_Suc_0 [simp]:
       
   141   \<open>mask (Suc 0) = 1\<close>
       
   142   by (simp add: mask_eq_exp_minus_1 add_implies_diff sym)
       
   143 
       
   144 lemma mask_Suc_exp:
       
   145   \<open>mask (Suc n) = 2 ^ n OR mask n\<close>
       
   146   by (rule bit_eqI)
       
   147     (auto simp add: bit_or_iff bit_mask_iff bit_exp_iff not_less le_less_Suc_eq)
       
   148 
       
   149 lemma mask_Suc_double:
       
   150   \<open>mask (Suc n) = 1 OR 2 * mask n\<close>
       
   151 proof (rule bit_eqI)
       
   152   fix q
       
   153   assume \<open>2 ^ q \<noteq> 0\<close>
       
   154   show \<open>bit (mask (Suc n)) q \<longleftrightarrow> bit (1 OR 2 * mask n) q\<close>
       
   155     by (cases q)
       
   156       (simp_all add: even_mask_iff even_or_iff bit_or_iff bit_mask_iff bit_exp_iff bit_double_iff not_less le_less_Suc_eq bit_1_iff, auto simp add: mult_2)
       
   157 qed
       
   158 
       
   159 lemma mask_numeral:
       
   160   \<open>mask (numeral n) = 1 + 2 * mask (pred_numeral n)\<close>
       
   161   by (simp add: numeral_eq_Suc mask_Suc_double one_or_eq ac_simps)
       
   162 
       
   163 lemma take_bit_mask [simp]:
       
   164   \<open>take_bit m (mask n) = mask (min m n)\<close>
       
   165   by (rule bit_eqI) (simp add: bit_simps)
       
   166 
       
   167 lemma take_bit_eq_mask:
       
   168   \<open>take_bit n a = a AND mask n\<close>
       
   169   by (rule bit_eqI)
       
   170     (auto simp add: bit_take_bit_iff bit_and_iff bit_mask_iff)
       
   171 
       
   172 lemma or_eq_0_iff:
       
   173   \<open>a OR b = 0 \<longleftrightarrow> a = 0 \<and> b = 0\<close>
       
   174   by (auto simp add: bit_eq_iff bit_or_iff)
       
   175 
       
   176 lemma disjunctive_add:
       
   177   \<open>a + b = a OR b\<close> if \<open>\<And>n. \<not> bit a n \<or> \<not> bit b n\<close>
       
   178   by (rule bit_eqI) (use that in \<open>simp add: bit_disjunctive_add_iff bit_or_iff\<close>)
       
   179 
       
   180 lemma bit_iff_and_drop_bit_eq_1:
       
   181   \<open>bit a n \<longleftrightarrow> drop_bit n a AND 1 = 1\<close>
       
   182   by (simp add: bit_iff_odd_drop_bit and_one_eq odd_iff_mod_2_eq_one)
       
   183 
       
   184 lemma bit_iff_and_push_bit_not_eq_0:
       
   185   \<open>bit a n \<longleftrightarrow> a AND push_bit n 1 \<noteq> 0\<close>
       
   186   apply (cases \<open>2 ^ n = 0\<close>)
       
   187   apply (simp_all add: push_bit_of_1 bit_eq_iff bit_and_iff bit_push_bit_iff exp_eq_0_imp_not_bit)
       
   188   apply (simp_all add: bit_exp_iff)
       
   189   done
       
   190 
       
   191 lemmas set_bit_def = set_bit_eq_or
       
   192 
       
   193 lemma bit_set_bit_iff [bit_simps]:
       
   194   \<open>bit (set_bit m a) n \<longleftrightarrow> bit a n \<or> (m = n \<and> 2 ^ n \<noteq> 0)\<close>
       
   195   by (auto simp add: set_bit_def push_bit_of_1 bit_or_iff bit_exp_iff)
       
   196 
       
   197 lemma even_set_bit_iff:
       
   198   \<open>even (set_bit m a) \<longleftrightarrow> even a \<and> m \<noteq> 0\<close>
       
   199   using bit_set_bit_iff [of m a 0] by auto
       
   200 
       
   201 lemma even_unset_bit_iff:
       
   202   \<open>even (unset_bit m a) \<longleftrightarrow> even a \<or> m = 0\<close>
       
   203   using bit_unset_bit_iff [of m a 0] by auto
       
   204 
       
   205 lemma and_exp_eq_0_iff_not_bit:
       
   206   \<open>a AND 2 ^ n = 0 \<longleftrightarrow> \<not> bit a n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
       
   207 proof
       
   208   assume ?Q
       
   209   then show ?P
       
   210     by (auto intro: bit_eqI simp add: bit_simps)
       
   211 next
       
   212   assume ?P
       
   213   show ?Q
       
   214   proof (rule notI)
       
   215     assume \<open>bit a n\<close>
       
   216     then have \<open>a AND 2 ^ n = 2 ^ n\<close>
       
   217       by (auto intro: bit_eqI simp add: bit_simps)
       
   218     with \<open>?P\<close> show False
       
   219       using \<open>bit a n\<close> exp_eq_0_imp_not_bit by auto
       
   220   qed
       
   221 qed
       
   222 
       
   223 lemmas flip_bit_def = flip_bit_eq_xor
       
   224 
       
   225 lemma bit_flip_bit_iff [bit_simps]:
       
   226   \<open>bit (flip_bit m a) n \<longleftrightarrow> (m = n \<longleftrightarrow> \<not> bit a n) \<and> 2 ^ n \<noteq> 0\<close>
       
   227   by (auto simp add: flip_bit_def push_bit_of_1 bit_xor_iff bit_exp_iff exp_eq_0_imp_not_bit)
       
   228 
       
   229 lemma even_flip_bit_iff:
       
   230   \<open>even (flip_bit m a) \<longleftrightarrow> \<not> (even a \<longleftrightarrow> m = 0)\<close>
       
   231   using bit_flip_bit_iff [of m a 0] by auto
       
   232 
       
   233 lemma set_bit_0 [simp]:
       
   234   \<open>set_bit 0 a = 1 + 2 * (a div 2)\<close>
       
   235 proof (rule bit_eqI)
       
   236   fix m
       
   237   assume *: \<open>2 ^ m \<noteq> 0\<close>
       
   238   then show \<open>bit (set_bit 0 a) m = bit (1 + 2 * (a div 2)) m\<close>
       
   239     by (simp add: bit_set_bit_iff bit_double_iff even_bit_succ_iff)
       
   240       (cases m, simp_all add: bit_Suc)
       
   241 qed
       
   242 
       
   243 lemma set_bit_Suc:
       
   244   \<open>set_bit (Suc n) a = a mod 2 + 2 * set_bit n (a div 2)\<close>
       
   245 proof (rule bit_eqI)
       
   246   fix m
       
   247   assume *: \<open>2 ^ m \<noteq> 0\<close>
       
   248   show \<open>bit (set_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * set_bit n (a div 2)) m\<close>
       
   249   proof (cases m)
       
   250     case 0
       
   251     then show ?thesis
       
   252       by (simp add: even_set_bit_iff)
       
   253   next
       
   254     case (Suc m)
       
   255     with * have \<open>2 ^ m \<noteq> 0\<close>
       
   256       using mult_2 by auto
       
   257     show ?thesis
       
   258       by (cases a rule: parity_cases)
       
   259         (simp_all add: bit_set_bit_iff bit_double_iff even_bit_succ_iff *,
       
   260         simp_all add: Suc \<open>2 ^ m \<noteq> 0\<close> bit_Suc)
       
   261   qed
       
   262 qed
       
   263 
       
   264 lemma unset_bit_0 [simp]:
       
   265   \<open>unset_bit 0 a = 2 * (a div 2)\<close>
       
   266 proof (rule bit_eqI)
       
   267   fix m
       
   268   assume *: \<open>2 ^ m \<noteq> 0\<close>
       
   269   then show \<open>bit (unset_bit 0 a) m = bit (2 * (a div 2)) m\<close>
       
   270     by (simp add: bit_unset_bit_iff bit_double_iff)
       
   271       (cases m, simp_all add: bit_Suc)
       
   272 qed
       
   273 
       
   274 lemma unset_bit_Suc:
       
   275   \<open>unset_bit (Suc n) a = a mod 2 + 2 * unset_bit n (a div 2)\<close>
       
   276 proof (rule bit_eqI)
       
   277   fix m
       
   278   assume *: \<open>2 ^ m \<noteq> 0\<close>
       
   279   then show \<open>bit (unset_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * unset_bit n (a div 2)) m\<close>
       
   280   proof (cases m)
       
   281     case 0
       
   282     then show ?thesis
       
   283       by (simp add: even_unset_bit_iff)
       
   284   next
       
   285     case (Suc m)
       
   286     show ?thesis
       
   287       by (cases a rule: parity_cases)
       
   288         (simp_all add: bit_unset_bit_iff bit_double_iff even_bit_succ_iff *,
       
   289          simp_all add: Suc bit_Suc)
       
   290   qed
       
   291 qed
       
   292 
       
   293 lemma flip_bit_0 [simp]:
       
   294   \<open>flip_bit 0 a = of_bool (even a) + 2 * (a div 2)\<close>
       
   295 proof (rule bit_eqI)
       
   296   fix m
       
   297   assume *: \<open>2 ^ m \<noteq> 0\<close>
       
   298   then show \<open>bit (flip_bit 0 a) m = bit (of_bool (even a) + 2 * (a div 2)) m\<close>
       
   299     by (simp add: bit_flip_bit_iff bit_double_iff even_bit_succ_iff)
       
   300       (cases m, simp_all add: bit_Suc)
       
   301 qed
       
   302 
       
   303 lemma flip_bit_Suc:
       
   304   \<open>flip_bit (Suc n) a = a mod 2 + 2 * flip_bit n (a div 2)\<close>
       
   305 proof (rule bit_eqI)
       
   306   fix m
       
   307   assume *: \<open>2 ^ m \<noteq> 0\<close>
       
   308   show \<open>bit (flip_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * flip_bit n (a div 2)) m\<close>
       
   309   proof (cases m)
       
   310     case 0
       
   311     then show ?thesis
       
   312       by (simp add: even_flip_bit_iff)
       
   313   next
       
   314     case (Suc m)
       
   315     with * have \<open>2 ^ m \<noteq> 0\<close>
       
   316       using mult_2 by auto
       
   317     show ?thesis
       
   318       by (cases a rule: parity_cases)
       
   319         (simp_all add: bit_flip_bit_iff bit_double_iff even_bit_succ_iff,
       
   320         simp_all add: Suc \<open>2 ^ m \<noteq> 0\<close> bit_Suc)
       
   321   qed
       
   322 qed
       
   323 
       
   324 lemma flip_bit_eq_if:
       
   325   \<open>flip_bit n a = (if bit a n then unset_bit else set_bit) n a\<close>
       
   326   by (rule bit_eqI) (auto simp add: bit_set_bit_iff bit_unset_bit_iff bit_flip_bit_iff)
       
   327 
       
   328 lemma take_bit_set_bit_eq:
       
   329   \<open>take_bit n (set_bit m a) = (if n \<le> m then take_bit n a else set_bit m (take_bit n a))\<close>
       
   330   by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_set_bit_iff)
       
   331 
       
   332 lemma take_bit_unset_bit_eq:
       
   333   \<open>take_bit n (unset_bit m a) = (if n \<le> m then take_bit n a else unset_bit m (take_bit n a))\<close>
       
   334   by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_unset_bit_iff)
       
   335 
       
   336 lemma take_bit_flip_bit_eq:
       
   337   \<open>take_bit n (flip_bit m a) = (if n \<le> m then take_bit n a else flip_bit m (take_bit n a))\<close>
       
   338   by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_flip_bit_iff)
       
   339 
       
   340 
       
   341 end
       
   342 
       
   343 class ring_bit_operations = semiring_bit_operations + ring_parity +
       
   344   fixes not :: \<open>'a \<Rightarrow> 'a\<close>  (\<open>NOT\<close>)
       
   345   assumes bit_not_iff [bit_simps]: \<open>\<And>n. bit (NOT a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> \<not> bit a n\<close>
       
   346   assumes minus_eq_not_minus_1: \<open>- a = NOT (a - 1)\<close>
       
   347 begin
       
   348 
       
   349 text \<open>
       
   350   For the sake of code generation \<^const>\<open>not\<close> is specified as
       
   351   definitional class operation.  Note that \<^const>\<open>not\<close> has no
       
   352   sensible definition for unlimited but only positive bit strings
       
   353   (type \<^typ>\<open>nat\<close>).
       
   354 \<close>
       
   355 
       
   356 lemma bits_minus_1_mod_2_eq [simp]:
       
   357   \<open>(- 1) mod 2 = 1\<close>
       
   358   by (simp add: mod_2_eq_odd)
       
   359 
       
   360 lemma not_eq_complement:
       
   361   \<open>NOT a = - a - 1\<close>
       
   362   using minus_eq_not_minus_1 [of \<open>a + 1\<close>] by simp
       
   363 
       
   364 lemma minus_eq_not_plus_1:
       
   365   \<open>- a = NOT a + 1\<close>
       
   366   using not_eq_complement [of a] by simp
       
   367 
       
   368 lemma bit_minus_iff [bit_simps]:
       
   369   \<open>bit (- a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> \<not> bit (a - 1) n\<close>
       
   370   by (simp add: minus_eq_not_minus_1 bit_not_iff)
       
   371 
       
   372 lemma even_not_iff [simp]:
       
   373   \<open>even (NOT a) \<longleftrightarrow> odd a\<close>
       
   374   using bit_not_iff [of a 0] by auto
       
   375 
       
   376 lemma bit_not_exp_iff [bit_simps]:
       
   377   \<open>bit (NOT (2 ^ m)) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n \<noteq> m\<close>
       
   378   by (auto simp add: bit_not_iff bit_exp_iff)
       
   379 
       
   380 lemma bit_minus_1_iff [simp]:
       
   381   \<open>bit (- 1) n \<longleftrightarrow> 2 ^ n \<noteq> 0\<close>
       
   382   by (simp add: bit_minus_iff)
       
   383 
       
   384 lemma bit_minus_exp_iff [bit_simps]:
       
   385   \<open>bit (- (2 ^ m)) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n \<ge> m\<close>
       
   386   by (auto simp add: bit_simps simp flip: mask_eq_exp_minus_1)
       
   387 
       
   388 lemma bit_minus_2_iff [simp]:
       
   389   \<open>bit (- 2) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n > 0\<close>
       
   390   by (simp add: bit_minus_iff bit_1_iff)
       
   391 
       
   392 lemma not_one [simp]:
       
   393   \<open>NOT 1 = - 2\<close>
       
   394   by (simp add: bit_eq_iff bit_not_iff) (simp add: bit_1_iff)
       
   395 
       
   396 sublocale "and": semilattice_neutr \<open>(AND)\<close> \<open>- 1\<close>
       
   397   by standard (rule bit_eqI, simp add: bit_and_iff)
       
   398 
       
   399 sublocale bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
       
   400   rewrites \<open>bit.xor = (XOR)\<close>
       
   401 proof -
       
   402   interpret bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
       
   403     by standard (auto simp add: bit_and_iff bit_or_iff bit_not_iff intro: bit_eqI)
       
   404   show \<open>boolean_algebra (AND) (OR) NOT 0 (- 1)\<close>
       
   405     by standard
       
   406   show \<open>boolean_algebra.xor (AND) (OR) NOT = (XOR)\<close>
       
   407     by (rule ext, rule ext, rule bit_eqI)
       
   408       (auto simp add: bit.xor_def bit_and_iff bit_or_iff bit_xor_iff bit_not_iff)
       
   409 qed
       
   410 
       
   411 lemma and_eq_not_not_or:
       
   412   \<open>a AND b = NOT (NOT a OR NOT b)\<close>
       
   413   by simp
       
   414 
       
   415 lemma or_eq_not_not_and:
       
   416   \<open>a OR b = NOT (NOT a AND NOT b)\<close>
       
   417   by simp
       
   418 
       
   419 lemma not_add_distrib:
       
   420   \<open>NOT (a + b) = NOT a - b\<close>
       
   421   by (simp add: not_eq_complement algebra_simps)
       
   422 
       
   423 lemma not_diff_distrib:
       
   424   \<open>NOT (a - b) = NOT a + b\<close>
       
   425   using not_add_distrib [of a \<open>- b\<close>] by simp
       
   426 
       
   427 lemma (in ring_bit_operations) and_eq_minus_1_iff:
       
   428   \<open>a AND b = - 1 \<longleftrightarrow> a = - 1 \<and> b = - 1\<close>
       
   429 proof
       
   430   assume \<open>a = - 1 \<and> b = - 1\<close>
       
   431   then show \<open>a AND b = - 1\<close>
       
   432     by simp
       
   433 next
       
   434   assume \<open>a AND b = - 1\<close>
       
   435   have *: \<open>bit a n\<close> \<open>bit b n\<close> if \<open>2 ^ n \<noteq> 0\<close> for n
       
   436   proof -
       
   437     from \<open>a AND b = - 1\<close>
       
   438     have \<open>bit (a AND b) n = bit (- 1) n\<close>
       
   439       by (simp add: bit_eq_iff)
       
   440     then show \<open>bit a n\<close> \<open>bit b n\<close>
       
   441       using that by (simp_all add: bit_and_iff)
       
   442   qed
       
   443   have \<open>a = - 1\<close>
       
   444     by (rule bit_eqI) (simp add: *)
       
   445   moreover have \<open>b = - 1\<close>
       
   446     by (rule bit_eqI) (simp add: *)
       
   447   ultimately show \<open>a = - 1 \<and> b = - 1\<close>
       
   448     by simp
       
   449 qed
       
   450 
       
   451 lemma disjunctive_diff:
       
   452   \<open>a - b = a AND NOT b\<close> if \<open>\<And>n. bit b n \<Longrightarrow> bit a n\<close>
       
   453 proof -
       
   454   have \<open>NOT a + b = NOT a OR b\<close>
       
   455     by (rule disjunctive_add) (auto simp add: bit_not_iff dest: that)
       
   456   then have \<open>NOT (NOT a + b) = NOT (NOT a OR b)\<close>
       
   457     by simp
       
   458   then show ?thesis
       
   459     by (simp add: not_add_distrib)
       
   460 qed
       
   461 
       
   462 lemma push_bit_minus:
       
   463   \<open>push_bit n (- a) = - push_bit n a\<close>
       
   464   by (simp add: push_bit_eq_mult)
       
   465 
       
   466 lemma take_bit_not_take_bit:
       
   467   \<open>take_bit n (NOT (take_bit n a)) = take_bit n (NOT a)\<close>
       
   468   by (auto simp add: bit_eq_iff bit_take_bit_iff bit_not_iff)
       
   469 
       
   470 lemma take_bit_not_iff:
       
   471   \<open>take_bit n (NOT a) = take_bit n (NOT b) \<longleftrightarrow> take_bit n a = take_bit n b\<close>
       
   472   apply (simp add: bit_eq_iff)
       
   473   apply (simp add: bit_not_iff bit_take_bit_iff bit_exp_iff)
       
   474   apply (use exp_eq_0_imp_not_bit in blast)
       
   475   done
       
   476 
       
   477 lemma take_bit_not_eq_mask_diff:
       
   478   \<open>take_bit n (NOT a) = mask n - take_bit n a\<close>
       
   479 proof -
       
   480   have \<open>take_bit n (NOT a) = take_bit n (NOT (take_bit n a))\<close>
       
   481     by (simp add: take_bit_not_take_bit)
       
   482   also have \<open>\<dots> = mask n AND NOT (take_bit n a)\<close>
       
   483     by (simp add: take_bit_eq_mask ac_simps)
       
   484   also have \<open>\<dots> = mask n - take_bit n a\<close>
       
   485     by (subst disjunctive_diff)
       
   486       (auto simp add: bit_take_bit_iff bit_mask_iff exp_eq_0_imp_not_bit)
       
   487   finally show ?thesis
       
   488     by simp
       
   489 qed
       
   490 
       
   491 lemma mask_eq_take_bit_minus_one:
       
   492   \<open>mask n = take_bit n (- 1)\<close>
       
   493   by (simp add: bit_eq_iff bit_mask_iff bit_take_bit_iff conj_commute)
       
   494 
       
   495 lemma take_bit_minus_one_eq_mask:
       
   496   \<open>take_bit n (- 1) = mask n\<close>
       
   497   by (simp add: mask_eq_take_bit_minus_one)
       
   498 
       
   499 lemma minus_exp_eq_not_mask:
       
   500   \<open>- (2 ^ n) = NOT (mask n)\<close>
       
   501   by (rule bit_eqI) (simp add: bit_minus_iff bit_not_iff flip: mask_eq_exp_minus_1)
       
   502 
       
   503 lemma push_bit_minus_one_eq_not_mask:
       
   504   \<open>push_bit n (- 1) = NOT (mask n)\<close>
       
   505   by (simp add: push_bit_eq_mult minus_exp_eq_not_mask)
       
   506 
       
   507 lemma take_bit_not_mask_eq_0:
       
   508   \<open>take_bit m (NOT (mask n)) = 0\<close> if \<open>n \<ge> m\<close>
       
   509   by (rule bit_eqI) (use that in \<open>simp add: bit_take_bit_iff bit_not_iff bit_mask_iff\<close>)
       
   510 
       
   511 lemma unset_bit_eq_and_not:
       
   512   \<open>unset_bit n a = a AND NOT (push_bit n 1)\<close>
       
   513   by (rule bit_eqI) (auto simp add: bit_simps)
       
   514 
       
   515 lemmas unset_bit_def = unset_bit_eq_and_not
       
   516 
       
   517 end
       
   518 
       
   519 
       
   520 subsection \<open>Instance \<^typ>\<open>int\<close>\<close>
       
   521 
       
   522 lemma int_bit_bound:
       
   523   fixes k :: int
       
   524   obtains n where \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close>
       
   525     and \<open>n > 0 \<Longrightarrow> bit k (n - 1) \<noteq> bit k n\<close>
       
   526 proof -
       
   527   obtain q where *: \<open>\<And>m. q \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k q\<close>
       
   528   proof (cases \<open>k \<ge> 0\<close>)
       
   529     case True
       
   530     moreover from power_gt_expt [of 2 \<open>nat k\<close>]
       
   531     have \<open>nat k < 2 ^ nat k\<close>
       
   532       by simp
       
   533     then have \<open>int (nat k) < int (2 ^ nat k)\<close>
       
   534       by (simp only: of_nat_less_iff)
       
   535     ultimately have *: \<open>k div 2 ^ nat k = 0\<close>
       
   536       by simp
       
   537     show thesis
       
   538     proof (rule that [of \<open>nat k\<close>])
       
   539       fix m
       
   540       assume \<open>nat k \<le> m\<close>
       
   541       then show \<open>bit k m \<longleftrightarrow> bit k (nat k)\<close>
       
   542         by (auto simp add: * bit_iff_odd power_add zdiv_zmult2_eq dest!: le_Suc_ex)
       
   543     qed
       
   544   next
       
   545     case False
       
   546     moreover from power_gt_expt [of 2 \<open>nat (- k)\<close>]
       
   547     have \<open>nat (- k) < 2 ^ nat (- k)\<close>
       
   548       by simp
       
   549     then have \<open>int (nat (- k)) < int (2 ^ nat (- k))\<close>
       
   550       by (simp only: of_nat_less_iff)
       
   551     ultimately have \<open>- k div - (2 ^ nat (- k)) = - 1\<close>
       
   552       by (subst div_pos_neg_trivial) simp_all
       
   553     then have *: \<open>k div 2 ^ nat (- k) = - 1\<close>
       
   554       by simp
       
   555     show thesis
       
   556     proof (rule that [of \<open>nat (- k)\<close>])
       
   557       fix m
       
   558       assume \<open>nat (- k) \<le> m\<close>
       
   559       then show \<open>bit k m \<longleftrightarrow> bit k (nat (- k))\<close>
       
   560         by (auto simp add: * bit_iff_odd power_add zdiv_zmult2_eq minus_1_div_exp_eq_int dest!: le_Suc_ex)
       
   561     qed
       
   562   qed
       
   563   show thesis
       
   564   proof (cases \<open>\<forall>m. bit k m \<longleftrightarrow> bit k q\<close>)
       
   565     case True
       
   566     then have \<open>bit k 0 \<longleftrightarrow> bit k q\<close>
       
   567       by blast
       
   568     with True that [of 0] show thesis
       
   569       by simp
       
   570   next
       
   571     case False
       
   572     then obtain r where **: \<open>bit k r \<noteq> bit k q\<close>
       
   573       by blast
       
   574     have \<open>r < q\<close>
       
   575       by (rule ccontr) (use * [of r] ** in simp)
       
   576     define N where \<open>N = {n. n < q \<and> bit k n \<noteq> bit k q}\<close>
       
   577     moreover have \<open>finite N\<close> \<open>r \<in> N\<close>
       
   578       using ** N_def \<open>r < q\<close> by auto
       
   579     moreover define n where \<open>n = Suc (Max N)\<close>
       
   580     ultimately have \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close>
       
   581       apply auto
       
   582          apply (metis (full_types, lifting) "*" Max_ge_iff Suc_n_not_le_n \<open>finite N\<close> all_not_in_conv mem_Collect_eq not_le)
       
   583         apply (metis "*" Max_ge Suc_n_not_le_n \<open>finite N\<close> linorder_not_less mem_Collect_eq)
       
   584         apply (metis "*" Max_ge Suc_n_not_le_n \<open>finite N\<close> linorder_not_less mem_Collect_eq)
       
   585       apply (metis (full_types, lifting) "*" Max_ge_iff Suc_n_not_le_n \<open>finite N\<close> all_not_in_conv mem_Collect_eq not_le)
       
   586       done
       
   587     have \<open>bit k (Max N) \<noteq> bit k n\<close>
       
   588       by (metis (mono_tags, lifting) "*" Max_in N_def \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m = bit k n\<close> \<open>finite N\<close> \<open>r \<in> N\<close> empty_iff le_cases mem_Collect_eq)
       
   589     show thesis apply (rule that [of n])
       
   590       using \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m = bit k n\<close> apply blast
       
   591       using \<open>bit k (Max N) \<noteq> bit k n\<close> n_def by auto
       
   592   qed
       
   593 qed
       
   594 
       
   595 instantiation int :: ring_bit_operations
       
   596 begin
       
   597 
       
   598 definition not_int :: \<open>int \<Rightarrow> int\<close>
       
   599   where \<open>not_int k = - k - 1\<close>
       
   600 
       
   601 lemma not_int_rec:
       
   602   \<open>NOT k = of_bool (even k) + 2 * NOT (k div 2)\<close> for k :: int
       
   603   by (auto simp add: not_int_def elim: oddE)
       
   604 
       
   605 lemma even_not_iff_int:
       
   606   \<open>even (NOT k) \<longleftrightarrow> odd k\<close> for k :: int
       
   607   by (simp add: not_int_def)
       
   608 
       
   609 lemma not_int_div_2:
       
   610   \<open>NOT k div 2 = NOT (k div 2)\<close> for k :: int
       
   611   by (simp add: not_int_def)
       
   612 
       
   613 lemma bit_not_int_iff [bit_simps]:
       
   614   \<open>bit (NOT k) n \<longleftrightarrow> \<not> bit k n\<close>
       
   615   for k :: int
       
   616   by (simp add: bit_not_int_iff' not_int_def)
       
   617 
       
   618 function and_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
       
   619   where \<open>(k::int) AND l = (if k \<in> {0, - 1} \<and> l \<in> {0, - 1}
       
   620     then - of_bool (odd k \<and> odd l)
       
   621     else of_bool (odd k \<and> odd l) + 2 * ((k div 2) AND (l div 2)))\<close>
       
   622   by auto
       
   623 
       
   624 termination
       
   625   by (relation \<open>measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))\<close>) auto
       
   626 
       
   627 declare and_int.simps [simp del]
       
   628 
       
   629 lemma and_int_rec:
       
   630   \<open>k AND l = of_bool (odd k \<and> odd l) + 2 * ((k div 2) AND (l div 2))\<close>
       
   631     for k l :: int
       
   632 proof (cases \<open>k \<in> {0, - 1} \<and> l \<in> {0, - 1}\<close>)
       
   633   case True
       
   634   then show ?thesis
       
   635     by auto (simp_all add: and_int.simps)
       
   636 next
       
   637   case False
       
   638   then show ?thesis
       
   639     by (auto simp add: ac_simps and_int.simps [of k l])
       
   640 qed
       
   641 
       
   642 lemma bit_and_int_iff:
       
   643   \<open>bit (k AND l) n \<longleftrightarrow> bit k n \<and> bit l n\<close> for k l :: int
       
   644 proof (induction n arbitrary: k l)
       
   645   case 0
       
   646   then show ?case
       
   647     by (simp add: and_int_rec [of k l])
       
   648 next
       
   649   case (Suc n)
       
   650   then show ?case
       
   651     by (simp add: and_int_rec [of k l] bit_Suc)
       
   652 qed
       
   653 
       
   654 lemma even_and_iff_int:
       
   655   \<open>even (k AND l) \<longleftrightarrow> even k \<or> even l\<close> for k l :: int
       
   656   using bit_and_int_iff [of k l 0] by auto
       
   657 
       
   658 definition or_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
       
   659   where \<open>k OR l = NOT (NOT k AND NOT l)\<close> for k l :: int
       
   660 
       
   661 lemma or_int_rec:
       
   662   \<open>k OR l = of_bool (odd k \<or> odd l) + 2 * ((k div 2) OR (l div 2))\<close>
       
   663   for k l :: int
       
   664   using and_int_rec [of \<open>NOT k\<close> \<open>NOT l\<close>]
       
   665   by (simp add: or_int_def even_not_iff_int not_int_div_2)
       
   666     (simp_all add: not_int_def)
       
   667 
       
   668 lemma bit_or_int_iff:
       
   669   \<open>bit (k OR l) n \<longleftrightarrow> bit k n \<or> bit l n\<close> for k l :: int
       
   670   by (simp add: or_int_def bit_not_int_iff bit_and_int_iff)
       
   671 
       
   672 definition xor_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
       
   673   where \<open>k XOR l = k AND NOT l OR NOT k AND l\<close> for k l :: int
       
   674 
       
   675 lemma xor_int_rec:
       
   676   \<open>k XOR l = of_bool (odd k \<noteq> odd l) + 2 * ((k div 2) XOR (l div 2))\<close>
       
   677   for k l :: int
       
   678   by (simp add: xor_int_def or_int_rec [of \<open>k AND NOT l\<close> \<open>NOT k AND l\<close>] even_and_iff_int even_not_iff_int)
       
   679     (simp add: and_int_rec [of \<open>NOT k\<close> \<open>l\<close>] and_int_rec [of \<open>k\<close> \<open>NOT l\<close>] not_int_div_2)
       
   680 
       
   681 lemma bit_xor_int_iff:
       
   682   \<open>bit (k XOR l) n \<longleftrightarrow> bit k n \<noteq> bit l n\<close> for k l :: int
       
   683   by (auto simp add: xor_int_def bit_or_int_iff bit_and_int_iff bit_not_int_iff)
       
   684 
       
   685 definition mask_int :: \<open>nat \<Rightarrow> int\<close>
       
   686   where \<open>mask n = (2 :: int) ^ n - 1\<close>
       
   687 
       
   688 definition set_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
       
   689   where \<open>set_bit n k = k OR push_bit n 1\<close> for k :: int
       
   690 
       
   691 definition unset_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
       
   692   where \<open>unset_bit n k = k AND NOT (push_bit n 1)\<close> for k :: int
       
   693 
       
   694 definition flip_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
       
   695   where \<open>flip_bit n k = k XOR push_bit n 1\<close> for k :: int
       
   696 
       
   697 instance proof
       
   698   fix k l :: int and m n :: nat
       
   699   show \<open>- k = NOT (k - 1)\<close>
       
   700     by (simp add: not_int_def)
       
   701   show \<open>bit (k AND l) n \<longleftrightarrow> bit k n \<and> bit l n\<close>
       
   702     by (fact bit_and_int_iff)
       
   703   show \<open>bit (k OR l) n \<longleftrightarrow> bit k n \<or> bit l n\<close>
       
   704     by (fact bit_or_int_iff)
       
   705   show \<open>bit (k XOR l) n \<longleftrightarrow> bit k n \<noteq> bit l n\<close>
       
   706     by (fact bit_xor_int_iff)
       
   707   show \<open>bit (unset_bit m k) n \<longleftrightarrow> bit k n \<and> m \<noteq> n\<close>
       
   708   proof -
       
   709     have \<open>unset_bit m k = k AND NOT (push_bit m 1)\<close>
       
   710       by (simp add: unset_bit_int_def)
       
   711     also have \<open>NOT (push_bit m 1 :: int) = - (push_bit m 1 + 1)\<close>
       
   712       by (simp add: not_int_def)
       
   713     finally show ?thesis by (simp only: bit_simps bit_and_int_iff) (auto simp add: bit_simps)
       
   714   qed
       
   715 qed (simp_all add: bit_not_int_iff mask_int_def set_bit_int_def flip_bit_int_def)
       
   716 
       
   717 end
       
   718 
       
   719 
       
   720 lemma mask_half_int:
       
   721   \<open>mask n div 2 = (mask (n - 1) :: int)\<close>
       
   722   by (cases n) (simp_all add: mask_eq_exp_minus_1 algebra_simps)
       
   723 
       
   724 lemma mask_nonnegative_int [simp]:
       
   725   \<open>mask n \<ge> (0::int)\<close>
       
   726   by (simp add: mask_eq_exp_minus_1)
       
   727 
       
   728 lemma not_mask_negative_int [simp]:
       
   729   \<open>\<not> mask n < (0::int)\<close>
       
   730   by (simp add: not_less)
       
   731 
       
   732 lemma not_nonnegative_int_iff [simp]:
       
   733   \<open>NOT k \<ge> 0 \<longleftrightarrow> k < 0\<close> for k :: int
       
   734   by (simp add: not_int_def)
       
   735 
       
   736 lemma not_negative_int_iff [simp]:
       
   737   \<open>NOT k < 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
       
   738   by (subst Not_eq_iff [symmetric]) (simp add: not_less not_le)
       
   739 
       
   740 lemma and_nonnegative_int_iff [simp]:
       
   741   \<open>k AND l \<ge> 0 \<longleftrightarrow> k \<ge> 0 \<or> l \<ge> 0\<close> for k l :: int
       
   742 proof (induction k arbitrary: l rule: int_bit_induct)
       
   743   case zero
       
   744   then show ?case
       
   745     by simp
       
   746 next
       
   747   case minus
       
   748   then show ?case
       
   749     by simp
       
   750 next
       
   751   case (even k)
       
   752   then show ?case
       
   753     using and_int_rec [of \<open>k * 2\<close> l] by (simp add: pos_imp_zdiv_nonneg_iff)
       
   754 next
       
   755   case (odd k)
       
   756   from odd have \<open>0 \<le> k AND l div 2 \<longleftrightarrow> 0 \<le> k \<or> 0 \<le> l div 2\<close>
       
   757     by simp
       
   758   then have \<open>0 \<le> (1 + k * 2) div 2 AND l div 2 \<longleftrightarrow> 0 \<le> (1 + k * 2) div 2\<or> 0 \<le> l div 2\<close>
       
   759     by simp
       
   760   with and_int_rec [of \<open>1 + k * 2\<close> l]
       
   761   show ?case
       
   762     by auto
       
   763 qed
       
   764 
       
   765 lemma and_negative_int_iff [simp]:
       
   766   \<open>k AND l < 0 \<longleftrightarrow> k < 0 \<and> l < 0\<close> for k l :: int
       
   767   by (subst Not_eq_iff [symmetric]) (simp add: not_less)
       
   768 
       
   769 lemma and_less_eq:
       
   770   \<open>k AND l \<le> k\<close> if \<open>l < 0\<close> for k l :: int
       
   771 using that proof (induction k arbitrary: l rule: int_bit_induct)
       
   772   case zero
       
   773   then show ?case
       
   774     by simp
       
   775 next
       
   776   case minus
       
   777   then show ?case
       
   778     by simp
       
   779 next
       
   780   case (even k)
       
   781   from even.IH [of \<open>l div 2\<close>] even.hyps even.prems
       
   782   show ?case
       
   783     by (simp add: and_int_rec [of _ l])
       
   784 next
       
   785   case (odd k)
       
   786   from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems
       
   787   show ?case
       
   788     by (simp add: and_int_rec [of _ l])
       
   789 qed
       
   790 
       
   791 lemma or_nonnegative_int_iff [simp]:
       
   792   \<open>k OR l \<ge> 0 \<longleftrightarrow> k \<ge> 0 \<and> l \<ge> 0\<close> for k l :: int
       
   793   by (simp only: or_eq_not_not_and not_nonnegative_int_iff) simp
       
   794 
       
   795 lemma or_negative_int_iff [simp]:
       
   796   \<open>k OR l < 0 \<longleftrightarrow> k < 0 \<or> l < 0\<close> for k l :: int
       
   797   by (subst Not_eq_iff [symmetric]) (simp add: not_less)
       
   798 
       
   799 lemma or_greater_eq:
       
   800   \<open>k OR l \<ge> k\<close> if \<open>l \<ge> 0\<close> for k l :: int
       
   801 using that proof (induction k arbitrary: l rule: int_bit_induct)
       
   802   case zero
       
   803   then show ?case
       
   804     by simp
       
   805 next
       
   806   case minus
       
   807   then show ?case
       
   808     by simp
       
   809 next
       
   810   case (even k)
       
   811   from even.IH [of \<open>l div 2\<close>] even.hyps even.prems
       
   812   show ?case
       
   813     by (simp add: or_int_rec [of _ l])
       
   814 next
       
   815   case (odd k)
       
   816   from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems
       
   817   show ?case
       
   818     by (simp add: or_int_rec [of _ l])
       
   819 qed
       
   820 
       
   821 lemma xor_nonnegative_int_iff [simp]:
       
   822   \<open>k XOR l \<ge> 0 \<longleftrightarrow> (k \<ge> 0 \<longleftrightarrow> l \<ge> 0)\<close> for k l :: int
       
   823   by (simp only: bit.xor_def or_nonnegative_int_iff) auto
       
   824 
       
   825 lemma xor_negative_int_iff [simp]:
       
   826   \<open>k XOR l < 0 \<longleftrightarrow> (k < 0) \<noteq> (l < 0)\<close> for k l :: int
       
   827   by (subst Not_eq_iff [symmetric]) (auto simp add: not_less)
       
   828 
       
   829 lemma OR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
       
   830   fixes x y :: int
       
   831   assumes \<open>0 \<le> x\<close> \<open>x < 2 ^ n\<close> \<open>y < 2 ^ n\<close>
       
   832   shows \<open>x OR y < 2 ^ n\<close>
       
   833 using assms proof (induction x arbitrary: y n rule: int_bit_induct)
       
   834   case zero
       
   835   then show ?case
       
   836     by simp
       
   837 next
       
   838   case minus
       
   839   then show ?case
       
   840     by simp
       
   841 next
       
   842   case (even x)
       
   843   from even.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] even.prems even.hyps
       
   844   show ?case 
       
   845     by (cases n) (auto simp add: or_int_rec [of \<open>_ * 2\<close>] elim: oddE)
       
   846 next
       
   847   case (odd x)
       
   848   from odd.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] odd.prems odd.hyps
       
   849   show ?case
       
   850     by (cases n) (auto simp add: or_int_rec [of \<open>1 + _ * 2\<close>], linarith)
       
   851 qed
       
   852 
       
   853 lemma XOR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
       
   854   fixes x y :: int
       
   855   assumes \<open>0 \<le> x\<close> \<open>x < 2 ^ n\<close> \<open>y < 2 ^ n\<close>
       
   856   shows \<open>x XOR y < 2 ^ n\<close>
       
   857 using assms proof (induction x arbitrary: y n rule: int_bit_induct)
       
   858   case zero
       
   859   then show ?case
       
   860     by simp
       
   861 next
       
   862   case minus
       
   863   then show ?case
       
   864     by simp
       
   865 next
       
   866   case (even x)
       
   867   from even.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] even.prems even.hyps
       
   868   show ?case 
       
   869     by (cases n) (auto simp add: xor_int_rec [of \<open>_ * 2\<close>] elim: oddE)
       
   870 next
       
   871   case (odd x)
       
   872   from odd.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] odd.prems odd.hyps
       
   873   show ?case
       
   874     by (cases n) (auto simp add: xor_int_rec [of \<open>1 + _ * 2\<close>])
       
   875 qed
       
   876 
       
   877 lemma AND_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
       
   878   fixes x y :: int
       
   879   assumes \<open>0 \<le> x\<close>
       
   880   shows \<open>0 \<le> x AND y\<close>
       
   881   using assms by simp
       
   882 
       
   883 lemma OR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
       
   884   fixes x y :: int
       
   885   assumes \<open>0 \<le> x\<close> \<open>0 \<le> y\<close>
       
   886   shows \<open>0 \<le> x OR y\<close>
       
   887   using assms by simp
       
   888 
       
   889 lemma XOR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
       
   890   fixes x y :: int
       
   891   assumes \<open>0 \<le> x\<close> \<open>0 \<le> y\<close>
       
   892   shows \<open>0 \<le> x XOR y\<close>
       
   893   using assms by simp
       
   894 
       
   895 lemma AND_upper1 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
       
   896   fixes x y :: int
       
   897   assumes \<open>0 \<le> x\<close>
       
   898   shows \<open>x AND y \<le> x\<close>
       
   899 using assms proof (induction x arbitrary: y rule: int_bit_induct)
       
   900   case (odd k)
       
   901   then have \<open>k AND y div 2 \<le> k\<close>
       
   902     by simp
       
   903   then show ?case 
       
   904     by (simp add: and_int_rec [of \<open>1 + _ * 2\<close>])
       
   905 qed (simp_all add: and_int_rec [of \<open>_ * 2\<close>])
       
   906 
       
   907 lemmas AND_upper1' [simp] = order_trans [OF AND_upper1] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
       
   908 lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
       
   909 
       
   910 lemma AND_upper2 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
       
   911   fixes x y :: int
       
   912   assumes \<open>0 \<le> y\<close>
       
   913   shows \<open>x AND y \<le> y\<close>
       
   914   using assms AND_upper1 [of y x] by (simp add: ac_simps)
       
   915 
       
   916 lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
       
   917 lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
       
   918 
       
   919 lemma plus_and_or: \<open>(x AND y) + (x OR y) = x + y\<close> for x y :: int
       
   920 proof (induction x arbitrary: y rule: int_bit_induct)
       
   921   case zero
       
   922   then show ?case
       
   923     by simp
       
   924 next
       
   925   case minus
       
   926   then show ?case
       
   927     by simp
       
   928 next
       
   929   case (even x)
       
   930   from even.IH [of \<open>y div 2\<close>]
       
   931   show ?case
       
   932     by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE)
       
   933 next
       
   934   case (odd x)
       
   935   from odd.IH [of \<open>y div 2\<close>]
       
   936   show ?case
       
   937     by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE)
       
   938 qed
       
   939 
       
   940 lemma set_bit_nonnegative_int_iff [simp]:
       
   941   \<open>set_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
       
   942   by (simp add: set_bit_def)
       
   943 
       
   944 lemma set_bit_negative_int_iff [simp]:
       
   945   \<open>set_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
       
   946   by (simp add: set_bit_def)
       
   947 
       
   948 lemma unset_bit_nonnegative_int_iff [simp]:
       
   949   \<open>unset_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
       
   950   by (simp add: unset_bit_def)
       
   951 
       
   952 lemma unset_bit_negative_int_iff [simp]:
       
   953   \<open>unset_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
       
   954   by (simp add: unset_bit_def)
       
   955 
       
   956 lemma flip_bit_nonnegative_int_iff [simp]:
       
   957   \<open>flip_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
       
   958   by (simp add: flip_bit_def)
       
   959 
       
   960 lemma flip_bit_negative_int_iff [simp]:
       
   961   \<open>flip_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
       
   962   by (simp add: flip_bit_def)
       
   963 
       
   964 lemma set_bit_greater_eq:
       
   965   \<open>set_bit n k \<ge> k\<close> for k :: int
       
   966   by (simp add: set_bit_def or_greater_eq)
       
   967 
       
   968 lemma unset_bit_less_eq:
       
   969   \<open>unset_bit n k \<le> k\<close> for k :: int
       
   970   by (simp add: unset_bit_def and_less_eq)
       
   971 
       
   972 lemma set_bit_eq:
       
   973   \<open>set_bit n k = k + of_bool (\<not> bit k n) * 2 ^ n\<close> for k :: int
       
   974 proof (rule bit_eqI)
       
   975   fix m
       
   976   show \<open>bit (set_bit n k) m \<longleftrightarrow> bit (k + of_bool (\<not> bit k n) * 2 ^ n) m\<close>
       
   977   proof (cases \<open>m = n\<close>)
       
   978     case True
       
   979     then show ?thesis
       
   980       apply (simp add: bit_set_bit_iff)
       
   981       apply (simp add: bit_iff_odd div_plus_div_distrib_dvd_right)
       
   982       done
       
   983   next
       
   984     case False
       
   985     then show ?thesis
       
   986       apply (clarsimp simp add: bit_set_bit_iff)
       
   987       apply (subst disjunctive_add)
       
   988       apply (clarsimp simp add: bit_exp_iff)
       
   989       apply (clarsimp simp add: bit_or_iff bit_exp_iff)
       
   990       done
       
   991   qed
       
   992 qed
       
   993 
       
   994 lemma unset_bit_eq:
       
   995   \<open>unset_bit n k = k - of_bool (bit k n) * 2 ^ n\<close> for k :: int
       
   996 proof (rule bit_eqI)
       
   997   fix m
       
   998   show \<open>bit (unset_bit n k) m \<longleftrightarrow> bit (k - of_bool (bit k n) * 2 ^ n) m\<close>
       
   999   proof (cases \<open>m = n\<close>)
       
  1000     case True
       
  1001     then show ?thesis
       
  1002       apply (simp add: bit_unset_bit_iff)
       
  1003       apply (simp add: bit_iff_odd)
       
  1004       using div_plus_div_distrib_dvd_right [of \<open>2 ^ n\<close> \<open>- (2 ^ n)\<close> k]
       
  1005       apply (simp add: dvd_neg_div)
       
  1006       done
       
  1007   next
       
  1008     case False
       
  1009     then show ?thesis
       
  1010       apply (clarsimp simp add: bit_unset_bit_iff)
       
  1011       apply (subst disjunctive_diff)
       
  1012       apply (clarsimp simp add: bit_exp_iff)
       
  1013       apply (clarsimp simp add: bit_and_iff bit_not_iff bit_exp_iff)
       
  1014       done
       
  1015   qed
       
  1016 qed
       
  1017 
       
  1018 lemma take_bit_eq_mask_iff:
       
  1019   \<open>take_bit n k = mask n \<longleftrightarrow> take_bit n (k + 1) = 0\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
       
  1020   for k :: int
       
  1021 proof
       
  1022   assume ?P
       
  1023   then have \<open>take_bit n (take_bit n k + take_bit n 1) = 0\<close>
       
  1024     by (simp add: mask_eq_exp_minus_1)
       
  1025   then show ?Q
       
  1026     by (simp only: take_bit_add)
       
  1027 next
       
  1028   assume ?Q
       
  1029   then have \<open>take_bit n (k + 1) - 1 = - 1\<close>
       
  1030     by simp
       
  1031   then have \<open>take_bit n (take_bit n (k + 1) - 1) = take_bit n (- 1)\<close>
       
  1032     by simp
       
  1033   moreover have \<open>take_bit n (take_bit n (k + 1) - 1) = take_bit n k\<close>
       
  1034     by (simp add: take_bit_eq_mod mod_simps)
       
  1035   ultimately show ?P
       
  1036     by (simp add: take_bit_minus_one_eq_mask)
       
  1037 qed
       
  1038 
       
  1039 lemma take_bit_eq_mask_iff_exp_dvd:
       
  1040   \<open>take_bit n k = mask n \<longleftrightarrow> 2 ^ n dvd k + 1\<close>
       
  1041   for k :: int
       
  1042   by (simp add: take_bit_eq_mask_iff flip: take_bit_eq_0_iff)
       
  1043 
       
  1044 context ring_bit_operations
       
  1045 begin
       
  1046 
       
  1047 lemma even_of_int_iff:
       
  1048   \<open>even (of_int k) \<longleftrightarrow> even k\<close>
       
  1049   by (induction k rule: int_bit_induct) simp_all
       
  1050 
       
  1051 lemma bit_of_int_iff [bit_simps]:
       
  1052   \<open>bit (of_int k) n \<longleftrightarrow> (2::'a) ^ n \<noteq> 0 \<and> bit k n\<close>
       
  1053 proof (cases \<open>(2::'a) ^ n = 0\<close>)
       
  1054   case True
       
  1055   then show ?thesis
       
  1056     by (simp add: exp_eq_0_imp_not_bit)
       
  1057 next
       
  1058   case False
       
  1059   then have \<open>bit (of_int k) n \<longleftrightarrow> bit k n\<close>
       
  1060   proof (induction k arbitrary: n rule: int_bit_induct)
       
  1061     case zero
       
  1062     then show ?case
       
  1063       by simp
       
  1064   next
       
  1065     case minus
       
  1066     then show ?case
       
  1067       by simp
       
  1068   next
       
  1069     case (even k)
       
  1070     then show ?case
       
  1071       using bit_double_iff [of \<open>of_int k\<close> n] Parity.bit_double_iff [of k n]
       
  1072       by (cases n) (auto simp add: ac_simps dest: mult_not_zero)
       
  1073   next
       
  1074     case (odd k)
       
  1075     then show ?case
       
  1076       using bit_double_iff [of \<open>of_int k\<close> n]
       
  1077       by (cases n) (auto simp add: ac_simps bit_double_iff even_bit_succ_iff Parity.bit_Suc dest: mult_not_zero)
       
  1078   qed
       
  1079   with False show ?thesis
       
  1080     by simp
       
  1081 qed
       
  1082 
       
  1083 lemma push_bit_of_int:
       
  1084   \<open>push_bit n (of_int k) = of_int (push_bit n k)\<close>
       
  1085   by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult)
       
  1086 
       
  1087 lemma of_int_push_bit:
       
  1088   \<open>of_int (push_bit n k) = push_bit n (of_int k)\<close>
       
  1089   by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult)
       
  1090 
       
  1091 lemma take_bit_of_int:
       
  1092   \<open>take_bit n (of_int k) = of_int (take_bit n k)\<close>
       
  1093   by (rule bit_eqI) (simp add: bit_take_bit_iff Parity.bit_take_bit_iff bit_of_int_iff)
       
  1094 
       
  1095 lemma of_int_take_bit:
       
  1096   \<open>of_int (take_bit n k) = take_bit n (of_int k)\<close>
       
  1097   by (rule bit_eqI) (simp add: bit_take_bit_iff Parity.bit_take_bit_iff bit_of_int_iff)
       
  1098 
       
  1099 lemma of_int_not_eq:
       
  1100   \<open>of_int (NOT k) = NOT (of_int k)\<close>
       
  1101   by (rule bit_eqI) (simp add: bit_not_iff Bit_Operations.bit_not_iff bit_of_int_iff)
       
  1102 
       
  1103 lemma of_int_and_eq:
       
  1104   \<open>of_int (k AND l) = of_int k AND of_int l\<close>
       
  1105   by (rule bit_eqI) (simp add: bit_of_int_iff bit_and_iff Bit_Operations.bit_and_iff)
       
  1106 
       
  1107 lemma of_int_or_eq:
       
  1108   \<open>of_int (k OR l) = of_int k OR of_int l\<close>
       
  1109   by (rule bit_eqI) (simp add: bit_of_int_iff bit_or_iff Bit_Operations.bit_or_iff)
       
  1110 
       
  1111 lemma of_int_xor_eq:
       
  1112   \<open>of_int (k XOR l) = of_int k XOR of_int l\<close>
       
  1113   by (rule bit_eqI) (simp add: bit_of_int_iff bit_xor_iff Bit_Operations.bit_xor_iff)
       
  1114 
       
  1115 lemma of_int_mask_eq:
       
  1116   \<open>of_int (mask n) = mask n\<close>
       
  1117   by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_int_or_eq)
       
  1118 
       
  1119 end
       
  1120 
       
  1121 lemma minus_numeral_inc_eq:
       
  1122   \<open>- numeral (Num.inc n) = NOT (numeral n :: int)\<close>
       
  1123   by (simp add: not_int_def sub_inc_One_eq add_One)
       
  1124 
       
  1125 lemma sub_one_eq_not_neg:
       
  1126   \<open>Num.sub n num.One = NOT (- numeral n :: int)\<close>
       
  1127   by (simp add: not_int_def)
       
  1128 
       
  1129 lemma int_not_numerals [simp]:
       
  1130   \<open>NOT (numeral (Num.Bit0 n) :: int) = - numeral (Num.Bit1 n)\<close>
       
  1131   \<open>NOT (numeral (Num.Bit1 n) :: int) = - numeral (Num.inc (num.Bit1 n))\<close>
       
  1132   \<open>NOT (numeral (Num.BitM n) :: int) = - numeral (num.Bit0 n)\<close>
       
  1133   \<open>NOT (- numeral (Num.Bit0 n) :: int) = numeral (Num.BitM n)\<close>
       
  1134   \<open>NOT (- numeral (Num.Bit1 n) :: int) = numeral (Num.Bit0 n)\<close>
       
  1135   by (simp_all add: not_int_def add_One inc_BitM_eq) 
       
  1136 
       
  1137 text \<open>FIXME: The rule sets below are very large (24 rules for each
       
  1138   operator). Is there a simpler way to do this?\<close>
       
  1139 
       
  1140 context
       
  1141 begin
       
  1142 
       
  1143 private lemma eqI:
       
  1144   \<open>k = l\<close>
       
  1145   if num: \<open>\<And>n. bit k (numeral n) \<longleftrightarrow> bit l (numeral n)\<close>
       
  1146     and even: \<open>even k \<longleftrightarrow> even l\<close>
       
  1147   for k l :: int
       
  1148 proof (rule bit_eqI)
       
  1149   fix n
       
  1150   show \<open>bit k n \<longleftrightarrow> bit l n\<close>
       
  1151   proof (cases n)
       
  1152     case 0
       
  1153     with even show ?thesis
       
  1154       by simp
       
  1155   next
       
  1156     case (Suc n)
       
  1157     with num [of \<open>num_of_nat (Suc n)\<close>] show ?thesis
       
  1158       by (simp only: numeral_num_of_nat)
       
  1159   qed
       
  1160 qed
       
  1161 
       
  1162 lemma int_and_numerals [simp]:
       
  1163   \<open>numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)\<close>
       
  1164   \<open>numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND numeral y)\<close>
       
  1165   \<open>numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)\<close>
       
  1166   \<open>numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND numeral y)\<close>
       
  1167   \<open>numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)\<close>
       
  1168   \<open>numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND - numeral (y + Num.One))\<close>
       
  1169   \<open>numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)\<close>
       
  1170   \<open>numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND - numeral (y + Num.One))\<close>
       
  1171   \<open>- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND numeral y)\<close>
       
  1172   \<open>- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND numeral y)\<close>
       
  1173   \<open>- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND numeral y)\<close>
       
  1174   \<open>- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND numeral y)\<close>
       
  1175   \<open>- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND - numeral y)\<close>
       
  1176   \<open>- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND - numeral (y + Num.One))\<close>
       
  1177   \<open>- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND - numeral y)\<close>
       
  1178   \<open>- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND - numeral (y + Num.One))\<close>
       
  1179   \<open>(1::int) AND numeral (Num.Bit0 y) = 0\<close>
       
  1180   \<open>(1::int) AND numeral (Num.Bit1 y) = 1\<close>
       
  1181   \<open>(1::int) AND - numeral (Num.Bit0 y) = 0\<close>
       
  1182   \<open>(1::int) AND - numeral (Num.Bit1 y) = 1\<close>
       
  1183   \<open>numeral (Num.Bit0 x) AND (1::int) = 0\<close>
       
  1184   \<open>numeral (Num.Bit1 x) AND (1::int) = 1\<close>
       
  1185   \<open>- numeral (Num.Bit0 x) AND (1::int) = 0\<close>
       
  1186   \<open>- numeral (Num.Bit1 x) AND (1::int) = 1\<close>
       
  1187   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)
       
  1188 
       
  1189 lemma int_or_numerals [simp]:
       
  1190   \<open>numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR numeral y)\<close>
       
  1191   \<open>numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)\<close>
       
  1192   \<open>numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR numeral y)\<close>
       
  1193   \<open>numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)\<close>
       
  1194   \<open>numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR - numeral y)\<close>
       
  1195   \<open>numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))\<close>
       
  1196   \<open>numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR - numeral y)\<close>
       
  1197   \<open>numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))\<close>
       
  1198   \<open>- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR numeral y)\<close>
       
  1199   \<open>- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR numeral y)\<close>
       
  1200   \<open>- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)\<close>
       
  1201   \<open>- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)\<close>
       
  1202   \<open>- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR - numeral y)\<close>
       
  1203   \<open>- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR - numeral (y + Num.One))\<close>
       
  1204   \<open>- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral y)\<close>
       
  1205   \<open>- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral (y + Num.One))\<close>
       
  1206   \<open>(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close>
       
  1207   \<open>(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)\<close>
       
  1208   \<open>(1::int) OR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)\<close>
       
  1209   \<open>(1::int) OR - numeral (Num.Bit1 y) = - numeral (Num.Bit1 y)\<close>
       
  1210   \<open>numeral (Num.Bit0 x) OR (1::int) = numeral (Num.Bit1 x)\<close>
       
  1211   \<open>numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)\<close>
       
  1212   \<open>- numeral (Num.Bit0 x) OR (1::int) = - numeral (Num.BitM x)\<close>
       
  1213   \<open>- numeral (Num.Bit1 x) OR (1::int) = - numeral (Num.Bit1 x)\<close>
       
  1214   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)
       
  1215 
       
  1216 lemma int_xor_numerals [simp]:
       
  1217   \<open>numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR numeral y)\<close>
       
  1218   \<open>numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR numeral y)\<close>
       
  1219   \<open>numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR numeral y)\<close>
       
  1220   \<open>numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR numeral y)\<close>
       
  1221   \<open>numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR - numeral y)\<close>
       
  1222   \<open>numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR - numeral (y + Num.One))\<close>
       
  1223   \<open>numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR - numeral y)\<close>
       
  1224   \<open>numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR - numeral (y + Num.One))\<close>
       
  1225   \<open>- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR numeral y)\<close>
       
  1226   \<open>- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR numeral y)\<close>
       
  1227   \<open>- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR numeral y)\<close>
       
  1228   \<open>- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR numeral y)\<close>
       
  1229   \<open>- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR - numeral y)\<close>
       
  1230   \<open>- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR - numeral (y + Num.One))\<close>
       
  1231   \<open>- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR - numeral y)\<close>
       
  1232   \<open>- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR - numeral (y + Num.One))\<close>
       
  1233   \<open>(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close>
       
  1234   \<open>(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)\<close>
       
  1235   \<open>(1::int) XOR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)\<close>
       
  1236   \<open>(1::int) XOR - numeral (Num.Bit1 y) = - numeral (Num.Bit0 (y + Num.One))\<close>
       
  1237   \<open>numeral (Num.Bit0 x) XOR (1::int) = numeral (Num.Bit1 x)\<close>
       
  1238   \<open>numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)\<close>
       
  1239   \<open>- numeral (Num.Bit0 x) XOR (1::int) = - numeral (Num.BitM x)\<close>
       
  1240   \<open>- numeral (Num.Bit1 x) XOR (1::int) = - numeral (Num.Bit0 (x + Num.One))\<close>
       
  1241   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)
       
  1242 
       
  1243 end
       
  1244 
       
  1245 
       
  1246 subsection \<open>Bit concatenation\<close>
       
  1247 
       
  1248 definition concat_bit :: \<open>nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int\<close>
       
  1249   where \<open>concat_bit n k l = take_bit n k OR push_bit n l\<close>
       
  1250 
       
  1251 lemma bit_concat_bit_iff [bit_simps]:
       
  1252   \<open>bit (concat_bit m k l) n \<longleftrightarrow> n < m \<and> bit k n \<or> m \<le> n \<and> bit l (n - m)\<close>
       
  1253   by (simp add: concat_bit_def bit_or_iff bit_and_iff bit_take_bit_iff bit_push_bit_iff ac_simps)
       
  1254 
       
  1255 lemma concat_bit_eq:
       
  1256   \<open>concat_bit n k l = take_bit n k + push_bit n l\<close>
       
  1257   by (simp add: concat_bit_def take_bit_eq_mask
       
  1258     bit_and_iff bit_mask_iff bit_push_bit_iff disjunctive_add)
       
  1259 
       
  1260 lemma concat_bit_0 [simp]:
       
  1261   \<open>concat_bit 0 k l = l\<close>
       
  1262   by (simp add: concat_bit_def)
       
  1263 
       
  1264 lemma concat_bit_Suc:
       
  1265   \<open>concat_bit (Suc n) k l = k mod 2 + 2 * concat_bit n (k div 2) l\<close>
       
  1266   by (simp add: concat_bit_eq take_bit_Suc push_bit_double)
       
  1267 
       
  1268 lemma concat_bit_of_zero_1 [simp]:
       
  1269   \<open>concat_bit n 0 l = push_bit n l\<close>
       
  1270   by (simp add: concat_bit_def)
       
  1271 
       
  1272 lemma concat_bit_of_zero_2 [simp]:
       
  1273   \<open>concat_bit n k 0 = take_bit n k\<close>
       
  1274   by (simp add: concat_bit_def take_bit_eq_mask)
       
  1275 
       
  1276 lemma concat_bit_nonnegative_iff [simp]:
       
  1277   \<open>concat_bit n k l \<ge> 0 \<longleftrightarrow> l \<ge> 0\<close>
       
  1278   by (simp add: concat_bit_def)
       
  1279 
       
  1280 lemma concat_bit_negative_iff [simp]:
       
  1281   \<open>concat_bit n k l < 0 \<longleftrightarrow> l < 0\<close>
       
  1282   by (simp add: concat_bit_def)
       
  1283 
       
  1284 lemma concat_bit_assoc:
       
  1285   \<open>concat_bit n k (concat_bit m l r) = concat_bit (m + n) (concat_bit n k l) r\<close>
       
  1286   by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps)
       
  1287 
       
  1288 lemma concat_bit_assoc_sym:
       
  1289   \<open>concat_bit m (concat_bit n k l) r = concat_bit (min m n) k (concat_bit (m - n) l r)\<close>
       
  1290   by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps min_def)
       
  1291 
       
  1292 lemma concat_bit_eq_iff:
       
  1293   \<open>concat_bit n k l = concat_bit n r s
       
  1294     \<longleftrightarrow> take_bit n k = take_bit n r \<and> l = s\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
       
  1295 proof
       
  1296   assume ?Q
       
  1297   then show ?P
       
  1298     by (simp add: concat_bit_def)
       
  1299 next
       
  1300   assume ?P
       
  1301   then have *: \<open>bit (concat_bit n k l) m = bit (concat_bit n r s) m\<close> for m
       
  1302     by (simp add: bit_eq_iff)
       
  1303   have \<open>take_bit n k = take_bit n r\<close>
       
  1304   proof (rule bit_eqI)
       
  1305     fix m
       
  1306     from * [of m]
       
  1307     show \<open>bit (take_bit n k) m \<longleftrightarrow> bit (take_bit n r) m\<close>
       
  1308       by (auto simp add: bit_take_bit_iff bit_concat_bit_iff)
       
  1309   qed
       
  1310   moreover have \<open>push_bit n l = push_bit n s\<close>
       
  1311   proof (rule bit_eqI)
       
  1312     fix m
       
  1313     from * [of m]
       
  1314     show \<open>bit (push_bit n l) m \<longleftrightarrow> bit (push_bit n s) m\<close>
       
  1315       by (auto simp add: bit_push_bit_iff bit_concat_bit_iff)
       
  1316   qed
       
  1317   then have \<open>l = s\<close>
       
  1318     by (simp add: push_bit_eq_mult)
       
  1319   ultimately show ?Q
       
  1320     by (simp add: concat_bit_def)
       
  1321 qed
       
  1322 
       
  1323 lemma take_bit_concat_bit_eq:
       
  1324   \<open>take_bit m (concat_bit n k l) = concat_bit (min m n) k (take_bit (m - n) l)\<close>
       
  1325   by (rule bit_eqI)
       
  1326     (auto simp add: bit_take_bit_iff bit_concat_bit_iff min_def)  
       
  1327 
       
  1328 lemma concat_bit_take_bit_eq:
       
  1329   \<open>concat_bit n (take_bit n b) = concat_bit n b\<close>
       
  1330   by (simp add: concat_bit_def [abs_def])
       
  1331 
       
  1332 
       
  1333 subsection \<open>Taking bits with sign propagation\<close>
       
  1334 
       
  1335 context ring_bit_operations
       
  1336 begin
       
  1337 
       
  1338 definition signed_take_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
       
  1339   where \<open>signed_take_bit n a = take_bit n a OR (of_bool (bit a n) * NOT (mask n))\<close>
       
  1340 
       
  1341 lemma signed_take_bit_eq_if_positive:
       
  1342   \<open>signed_take_bit n a = take_bit n a\<close> if \<open>\<not> bit a n\<close>
       
  1343   using that by (simp add: signed_take_bit_def)
       
  1344 
       
  1345 lemma signed_take_bit_eq_if_negative:
       
  1346   \<open>signed_take_bit n a = take_bit n a OR NOT (mask n)\<close> if \<open>bit a n\<close>
       
  1347   using that by (simp add: signed_take_bit_def)
       
  1348 
       
  1349 lemma even_signed_take_bit_iff:
       
  1350   \<open>even (signed_take_bit m a) \<longleftrightarrow> even a\<close>
       
  1351   by (auto simp add: signed_take_bit_def even_or_iff even_mask_iff bit_double_iff)
       
  1352 
       
  1353 lemma bit_signed_take_bit_iff [bit_simps]:
       
  1354   \<open>bit (signed_take_bit m a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> bit a (min m n)\<close>
       
  1355   by (simp add: signed_take_bit_def bit_take_bit_iff bit_or_iff bit_not_iff bit_mask_iff min_def not_le)
       
  1356     (use exp_eq_0_imp_not_bit in blast)
       
  1357 
       
  1358 lemma signed_take_bit_0 [simp]:
       
  1359   \<open>signed_take_bit 0 a = - (a mod 2)\<close>
       
  1360   by (simp add: signed_take_bit_def odd_iff_mod_2_eq_one)
       
  1361 
       
  1362 lemma signed_take_bit_Suc:
       
  1363   \<open>signed_take_bit (Suc n) a = a mod 2 + 2 * signed_take_bit n (a div 2)\<close>
       
  1364 proof (rule bit_eqI)
       
  1365   fix m
       
  1366   assume *: \<open>2 ^ m \<noteq> 0\<close>
       
  1367   show \<open>bit (signed_take_bit (Suc n) a) m \<longleftrightarrow>
       
  1368     bit (a mod 2 + 2 * signed_take_bit n (a div 2)) m\<close>
       
  1369   proof (cases m)
       
  1370     case 0
       
  1371     then show ?thesis
       
  1372       by (simp add: even_signed_take_bit_iff)
       
  1373   next
       
  1374     case (Suc m)
       
  1375     with * have \<open>2 ^ m \<noteq> 0\<close>
       
  1376       by (metis mult_not_zero power_Suc)
       
  1377     with Suc show ?thesis
       
  1378       by (simp add: bit_signed_take_bit_iff mod2_eq_if bit_double_iff even_bit_succ_iff
       
  1379         ac_simps flip: bit_Suc)
       
  1380   qed
       
  1381 qed
       
  1382 
       
  1383 lemma signed_take_bit_of_0 [simp]:
       
  1384   \<open>signed_take_bit n 0 = 0\<close>
       
  1385   by (simp add: signed_take_bit_def)
       
  1386 
       
  1387 lemma signed_take_bit_of_minus_1 [simp]:
       
  1388   \<open>signed_take_bit n (- 1) = - 1\<close>
       
  1389   by (simp add: signed_take_bit_def take_bit_minus_one_eq_mask mask_eq_exp_minus_1)
       
  1390 
       
  1391 lemma signed_take_bit_Suc_1 [simp]:
       
  1392   \<open>signed_take_bit (Suc n) 1 = 1\<close>
       
  1393   by (simp add: signed_take_bit_Suc)
       
  1394 
       
  1395 lemma signed_take_bit_rec:
       
  1396   \<open>signed_take_bit n a = (if n = 0 then - (a mod 2) else a mod 2 + 2 * signed_take_bit (n - 1) (a div 2))\<close>
       
  1397   by (cases n) (simp_all add: signed_take_bit_Suc)
       
  1398 
       
  1399 lemma signed_take_bit_eq_iff_take_bit_eq:
       
  1400   \<open>signed_take_bit n a = signed_take_bit n b \<longleftrightarrow> take_bit (Suc n) a = take_bit (Suc n) b\<close>
       
  1401 proof -
       
  1402   have \<open>bit (signed_take_bit n a) = bit (signed_take_bit n b) \<longleftrightarrow> bit (take_bit (Suc n) a) = bit (take_bit (Suc n) b)\<close>
       
  1403     by (simp add: fun_eq_iff bit_signed_take_bit_iff bit_take_bit_iff not_le less_Suc_eq_le min_def)
       
  1404       (use exp_eq_0_imp_not_bit in fastforce)
       
  1405   then show ?thesis
       
  1406     by (simp add: bit_eq_iff fun_eq_iff)
       
  1407 qed
       
  1408 
       
  1409 lemma signed_take_bit_signed_take_bit [simp]:
       
  1410   \<open>signed_take_bit m (signed_take_bit n a) = signed_take_bit (min m n) a\<close>
       
  1411 proof (rule bit_eqI)
       
  1412   fix q
       
  1413   show \<open>bit (signed_take_bit m (signed_take_bit n a)) q \<longleftrightarrow>
       
  1414     bit (signed_take_bit (min m n) a) q\<close>
       
  1415     by (simp add: bit_signed_take_bit_iff min_def bit_or_iff bit_not_iff bit_mask_iff bit_take_bit_iff)
       
  1416       (use le_Suc_ex exp_add_not_zero_imp in blast)
       
  1417 qed
       
  1418 
       
  1419 lemma signed_take_bit_take_bit:
       
  1420   \<open>signed_take_bit m (take_bit n a) = (if n \<le> m then take_bit n else signed_take_bit m) a\<close>
       
  1421   by (rule bit_eqI) (auto simp add: bit_signed_take_bit_iff min_def bit_take_bit_iff)
       
  1422 
       
  1423 lemma take_bit_signed_take_bit:
       
  1424   \<open>take_bit m (signed_take_bit n a) = take_bit m a\<close> if \<open>m \<le> Suc n\<close>
       
  1425   using that by (rule le_SucE; intro bit_eqI)
       
  1426    (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def less_Suc_eq)
       
  1427 
       
  1428 end
       
  1429 
       
  1430 text \<open>Modulus centered around 0\<close>
       
  1431 
       
  1432 lemma signed_take_bit_eq_concat_bit:
       
  1433   \<open>signed_take_bit n k = concat_bit n k (- of_bool (bit k n))\<close>
       
  1434   by (simp add: concat_bit_def signed_take_bit_def push_bit_minus_one_eq_not_mask)
       
  1435 
       
  1436 lemma signed_take_bit_add:
       
  1437   \<open>signed_take_bit n (signed_take_bit n k + signed_take_bit n l) = signed_take_bit n (k + l)\<close>
       
  1438   for k l :: int
       
  1439 proof -
       
  1440   have \<open>take_bit (Suc n)
       
  1441      (take_bit (Suc n) (signed_take_bit n k) +
       
  1442       take_bit (Suc n) (signed_take_bit n l)) =
       
  1443     take_bit (Suc n) (k + l)\<close>
       
  1444     by (simp add: take_bit_signed_take_bit take_bit_add)
       
  1445   then show ?thesis
       
  1446     by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_add)
       
  1447 qed
       
  1448 
       
  1449 lemma signed_take_bit_diff:
       
  1450   \<open>signed_take_bit n (signed_take_bit n k - signed_take_bit n l) = signed_take_bit n (k - l)\<close>
       
  1451   for k l :: int
       
  1452 proof -
       
  1453   have \<open>take_bit (Suc n)
       
  1454      (take_bit (Suc n) (signed_take_bit n k) -
       
  1455       take_bit (Suc n) (signed_take_bit n l)) =
       
  1456     take_bit (Suc n) (k - l)\<close>
       
  1457     by (simp add: take_bit_signed_take_bit take_bit_diff)
       
  1458   then show ?thesis
       
  1459     by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_diff)
       
  1460 qed
       
  1461 
       
  1462 lemma signed_take_bit_minus:
       
  1463   \<open>signed_take_bit n (- signed_take_bit n k) = signed_take_bit n (- k)\<close>
       
  1464   for k :: int
       
  1465 proof -
       
  1466   have \<open>take_bit (Suc n)
       
  1467      (- take_bit (Suc n) (signed_take_bit n k)) =
       
  1468     take_bit (Suc n) (- k)\<close>
       
  1469     by (simp add: take_bit_signed_take_bit take_bit_minus)
       
  1470   then show ?thesis
       
  1471     by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_minus)
       
  1472 qed
       
  1473 
       
  1474 lemma signed_take_bit_mult:
       
  1475   \<open>signed_take_bit n (signed_take_bit n k * signed_take_bit n l) = signed_take_bit n (k * l)\<close>
       
  1476   for k l :: int
       
  1477 proof -
       
  1478   have \<open>take_bit (Suc n)
       
  1479      (take_bit (Suc n) (signed_take_bit n k) *
       
  1480       take_bit (Suc n) (signed_take_bit n l)) =
       
  1481     take_bit (Suc n) (k * l)\<close>
       
  1482     by (simp add: take_bit_signed_take_bit take_bit_mult)
       
  1483   then show ?thesis
       
  1484     by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_mult)
       
  1485 qed
       
  1486 
       
  1487 lemma signed_take_bit_eq_take_bit_minus:
       
  1488   \<open>signed_take_bit n k = take_bit (Suc n) k - 2 ^ Suc n * of_bool (bit k n)\<close>
       
  1489   for k :: int
       
  1490 proof (cases \<open>bit k n\<close>)
       
  1491   case True
       
  1492   have \<open>signed_take_bit n k = take_bit (Suc n) k OR NOT (mask (Suc n))\<close>
       
  1493     by (rule bit_eqI) (auto simp add: bit_signed_take_bit_iff min_def bit_take_bit_iff bit_or_iff bit_not_iff bit_mask_iff less_Suc_eq True)
       
  1494   then have \<open>signed_take_bit n k = take_bit (Suc n) k + NOT (mask (Suc n))\<close>
       
  1495     by (simp add: disjunctive_add bit_take_bit_iff bit_not_iff bit_mask_iff)
       
  1496   with True show ?thesis
       
  1497     by (simp flip: minus_exp_eq_not_mask)
       
  1498 next
       
  1499   case False
       
  1500   show ?thesis
       
  1501     by (rule bit_eqI) (simp add: False bit_signed_take_bit_iff bit_take_bit_iff min_def less_Suc_eq)
       
  1502 qed
       
  1503 
       
  1504 lemma signed_take_bit_eq_take_bit_shift:
       
  1505   \<open>signed_take_bit n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\<close>
       
  1506   for k :: int
       
  1507 proof -
       
  1508   have *: \<open>take_bit n k OR 2 ^ n = take_bit n k + 2 ^ n\<close>
       
  1509     by (simp add: disjunctive_add bit_exp_iff bit_take_bit_iff)
       
  1510   have \<open>take_bit n k - 2 ^ n = take_bit n k + NOT (mask n)\<close>
       
  1511     by (simp add: minus_exp_eq_not_mask)
       
  1512   also have \<open>\<dots> = take_bit n k OR NOT (mask n)\<close>
       
  1513     by (rule disjunctive_add)
       
  1514       (simp add: bit_exp_iff bit_take_bit_iff bit_not_iff bit_mask_iff)
       
  1515   finally have **: \<open>take_bit n k - 2 ^ n = take_bit n k OR NOT (mask n)\<close> .
       
  1516   have \<open>take_bit (Suc n) (k + 2 ^ n) = take_bit (Suc n) (take_bit (Suc n) k + take_bit (Suc n) (2 ^ n))\<close>
       
  1517     by (simp only: take_bit_add)
       
  1518   also have \<open>take_bit (Suc n) k = 2 ^ n * of_bool (bit k n) + take_bit n k\<close>
       
  1519     by (simp add: take_bit_Suc_from_most)
       
  1520   finally have \<open>take_bit (Suc n) (k + 2 ^ n) = take_bit (Suc n) (2 ^ (n + of_bool (bit k n)) + take_bit n k)\<close>
       
  1521     by (simp add: ac_simps)
       
  1522   also have \<open>2 ^ (n + of_bool (bit k n)) + take_bit n k = 2 ^ (n + of_bool (bit k n)) OR take_bit n k\<close>
       
  1523     by (rule disjunctive_add)
       
  1524       (auto simp add: disjunctive_add bit_take_bit_iff bit_double_iff bit_exp_iff)
       
  1525   finally show ?thesis
       
  1526     using * ** by (simp add: signed_take_bit_def concat_bit_Suc min_def ac_simps)
       
  1527 qed
       
  1528 
       
  1529 lemma signed_take_bit_nonnegative_iff [simp]:
       
  1530   \<open>0 \<le> signed_take_bit n k \<longleftrightarrow> \<not> bit k n\<close>
       
  1531   for k :: int
       
  1532   by (simp add: signed_take_bit_def not_less concat_bit_def)
       
  1533 
       
  1534 lemma signed_take_bit_negative_iff [simp]:
       
  1535   \<open>signed_take_bit n k < 0 \<longleftrightarrow> bit k n\<close>
       
  1536   for k :: int
       
  1537   by (simp add: signed_take_bit_def not_less concat_bit_def)
       
  1538 
       
  1539 lemma signed_take_bit_int_greater_eq_minus_exp [simp]:
       
  1540   \<open>- (2 ^ n) \<le> signed_take_bit n k\<close>
       
  1541   for k :: int
       
  1542   by (simp add: signed_take_bit_eq_take_bit_shift)
       
  1543 
       
  1544 lemma signed_take_bit_int_less_exp [simp]:
       
  1545   \<open>signed_take_bit n k < 2 ^ n\<close>
       
  1546   for k :: int
       
  1547   using take_bit_int_less_exp [of \<open>Suc n\<close>]
       
  1548   by (simp add: signed_take_bit_eq_take_bit_shift)
       
  1549 
       
  1550 lemma signed_take_bit_int_eq_self_iff:
       
  1551   \<open>signed_take_bit n k = k \<longleftrightarrow> - (2 ^ n) \<le> k \<and> k < 2 ^ n\<close>
       
  1552   for k :: int
       
  1553   by (auto simp add: signed_take_bit_eq_take_bit_shift take_bit_int_eq_self_iff algebra_simps)
       
  1554 
       
  1555 lemma signed_take_bit_int_eq_self:
       
  1556   \<open>signed_take_bit n k = k\<close> if \<open>- (2 ^ n) \<le> k\<close> \<open>k < 2 ^ n\<close>
       
  1557   for k :: int
       
  1558   using that by (simp add: signed_take_bit_int_eq_self_iff)
       
  1559 
       
  1560 lemma signed_take_bit_int_less_eq_self_iff:
       
  1561   \<open>signed_take_bit n k \<le> k \<longleftrightarrow> - (2 ^ n) \<le> k\<close>
       
  1562   for k :: int
       
  1563   by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_less_eq_self_iff algebra_simps)
       
  1564     linarith
       
  1565 
       
  1566 lemma signed_take_bit_int_less_self_iff:
       
  1567   \<open>signed_take_bit n k < k \<longleftrightarrow> 2 ^ n \<le> k\<close>
       
  1568   for k :: int
       
  1569   by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_less_self_iff algebra_simps)
       
  1570 
       
  1571 lemma signed_take_bit_int_greater_self_iff:
       
  1572   \<open>k < signed_take_bit n k \<longleftrightarrow> k < - (2 ^ n)\<close>
       
  1573   for k :: int
       
  1574   by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_greater_self_iff algebra_simps)
       
  1575     linarith
       
  1576 
       
  1577 lemma signed_take_bit_int_greater_eq_self_iff:
       
  1578   \<open>k \<le> signed_take_bit n k \<longleftrightarrow> k < 2 ^ n\<close>
       
  1579   for k :: int
       
  1580   by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_greater_eq_self_iff algebra_simps)
       
  1581 
       
  1582 lemma signed_take_bit_int_greater_eq:
       
  1583   \<open>k + 2 ^ Suc n \<le> signed_take_bit n k\<close> if \<open>k < - (2 ^ n)\<close>
       
  1584   for k :: int
       
  1585   using that take_bit_int_greater_eq [of \<open>k + 2 ^ n\<close> \<open>Suc n\<close>]
       
  1586   by (simp add: signed_take_bit_eq_take_bit_shift)
       
  1587 
       
  1588 lemma signed_take_bit_int_less_eq:
       
  1589   \<open>signed_take_bit n k \<le> k - 2 ^ Suc n\<close> if \<open>k \<ge> 2 ^ n\<close>
       
  1590   for k :: int
       
  1591   using that take_bit_int_less_eq [of \<open>Suc n\<close> \<open>k + 2 ^ n\<close>]
       
  1592   by (simp add: signed_take_bit_eq_take_bit_shift)
       
  1593 
       
  1594 lemma signed_take_bit_Suc_bit0 [simp]:
       
  1595   \<open>signed_take_bit (Suc n) (numeral (Num.Bit0 k)) = signed_take_bit n (numeral k) * (2 :: int)\<close>
       
  1596   by (simp add: signed_take_bit_Suc)
       
  1597 
       
  1598 lemma signed_take_bit_Suc_bit1 [simp]:
       
  1599   \<open>signed_take_bit (Suc n) (numeral (Num.Bit1 k)) = signed_take_bit n (numeral k) * 2 + (1 :: int)\<close>
       
  1600   by (simp add: signed_take_bit_Suc)
       
  1601 
       
  1602 lemma signed_take_bit_Suc_minus_bit0 [simp]:
       
  1603   \<open>signed_take_bit (Suc n) (- numeral (Num.Bit0 k)) = signed_take_bit n (- numeral k) * (2 :: int)\<close>
       
  1604   by (simp add: signed_take_bit_Suc)
       
  1605 
       
  1606 lemma signed_take_bit_Suc_minus_bit1 [simp]:
       
  1607   \<open>signed_take_bit (Suc n) (- numeral (Num.Bit1 k)) = signed_take_bit n (- numeral k - 1) * 2 + (1 :: int)\<close>
       
  1608   by (simp add: signed_take_bit_Suc)
       
  1609 
       
  1610 lemma signed_take_bit_numeral_bit0 [simp]:
       
  1611   \<open>signed_take_bit (numeral l) (numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (numeral k) * (2 :: int)\<close>
       
  1612   by (simp add: signed_take_bit_rec)
       
  1613 
       
  1614 lemma signed_take_bit_numeral_bit1 [simp]:
       
  1615   \<open>signed_take_bit (numeral l) (numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (numeral k) * 2 + (1 :: int)\<close>
       
  1616   by (simp add: signed_take_bit_rec)
       
  1617 
       
  1618 lemma signed_take_bit_numeral_minus_bit0 [simp]:
       
  1619   \<open>signed_take_bit (numeral l) (- numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (- numeral k) * (2 :: int)\<close>
       
  1620   by (simp add: signed_take_bit_rec)
       
  1621 
       
  1622 lemma signed_take_bit_numeral_minus_bit1 [simp]:
       
  1623   \<open>signed_take_bit (numeral l) (- numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (- numeral k - 1) * 2 + (1 :: int)\<close>
       
  1624   by (simp add: signed_take_bit_rec)
       
  1625 
       
  1626 lemma signed_take_bit_code [code]:
       
  1627   \<open>signed_take_bit n a =
       
  1628   (let l = take_bit (Suc n) a
       
  1629    in if bit l n then l + push_bit (Suc n) (- 1) else l)\<close>
       
  1630 proof -
       
  1631   have *: \<open>take_bit (Suc n) a + push_bit n (- 2) =
       
  1632     take_bit (Suc n) a OR NOT (mask (Suc n))\<close>
       
  1633     by (auto simp add: bit_take_bit_iff bit_push_bit_iff bit_not_iff bit_mask_iff disjunctive_add
       
  1634        simp flip: push_bit_minus_one_eq_not_mask)
       
  1635   show ?thesis
       
  1636     by (rule bit_eqI)
       
  1637       (auto simp add: Let_def * bit_signed_take_bit_iff bit_take_bit_iff min_def less_Suc_eq bit_not_iff bit_mask_iff bit_or_iff)
       
  1638 qed
       
  1639 
       
  1640 
       
  1641 subsection \<open>Instance \<^typ>\<open>nat\<close>\<close>
       
  1642 
       
  1643 instantiation nat :: semiring_bit_operations
       
  1644 begin
       
  1645 
       
  1646 definition and_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
       
  1647   where \<open>m AND n = nat (int m AND int n)\<close> for m n :: nat
       
  1648 
       
  1649 definition or_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
       
  1650   where \<open>m OR n = nat (int m OR int n)\<close> for m n :: nat
       
  1651 
       
  1652 definition xor_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
       
  1653   where \<open>m XOR n = nat (int m XOR int n)\<close> for m n :: nat
       
  1654 
       
  1655 definition mask_nat :: \<open>nat \<Rightarrow> nat\<close>
       
  1656   where \<open>mask n = (2 :: nat) ^ n - 1\<close>
       
  1657 
       
  1658 definition set_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
       
  1659   where \<open>set_bit m n = n OR push_bit m 1\<close> for m n :: nat
       
  1660 
       
  1661 definition unset_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
       
  1662   where \<open>unset_bit m n = (if bit n m then n - push_bit m 1 else n)\<close> for m n :: nat
       
  1663 
       
  1664 definition flip_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
       
  1665   where \<open>flip_bit m n = n XOR push_bit m 1\<close> for m n :: nat
       
  1666 
       
  1667 instance proof
       
  1668   fix m n q :: nat
       
  1669   show \<open>bit (m AND n) q \<longleftrightarrow> bit m q \<and> bit n q\<close>
       
  1670     by (simp add: and_nat_def bit_simps)
       
  1671   show \<open>bit (m OR n) q \<longleftrightarrow> bit m q \<or> bit n q\<close>
       
  1672     by (simp add: or_nat_def bit_simps)
       
  1673   show \<open>bit (m XOR n) q \<longleftrightarrow> bit m q \<noteq> bit n q\<close>
       
  1674     by (simp add: xor_nat_def bit_simps)
       
  1675   show \<open>bit (unset_bit m n) q \<longleftrightarrow> bit n q \<and> m \<noteq> q\<close>
       
  1676   proof (cases \<open>bit n m\<close>)
       
  1677     case False
       
  1678     then show ?thesis by (auto simp add: unset_bit_nat_def)
       
  1679   next
       
  1680     case True
       
  1681     have \<open>push_bit m (drop_bit m n) + take_bit m n = n\<close>
       
  1682       by (fact bits_ident)
       
  1683     also from \<open>bit n m\<close> have \<open>drop_bit m n = 2 * drop_bit (Suc m) n  + 1\<close>
       
  1684       by (simp add: drop_bit_Suc drop_bit_half even_drop_bit_iff_not_bit ac_simps)
       
  1685     finally have \<open>push_bit m (2 * drop_bit (Suc m) n) + take_bit m n + push_bit m 1 = n\<close>
       
  1686       by (simp only: push_bit_add ac_simps)
       
  1687     then have \<open>n - push_bit m 1 = push_bit m (2 * drop_bit (Suc m) n) + take_bit m n\<close>
       
  1688       by simp
       
  1689     then have \<open>n - push_bit m 1 = push_bit m (2 * drop_bit (Suc m) n) OR take_bit m n\<close>
       
  1690       by (simp add: or_nat_def bit_simps flip: disjunctive_add)
       
  1691     with \<open>bit n m\<close> show ?thesis
       
  1692       by (auto simp add: unset_bit_nat_def or_nat_def bit_simps)
       
  1693   qed
       
  1694 qed (simp_all add: mask_nat_def set_bit_nat_def flip_bit_nat_def)
       
  1695 
       
  1696 end
       
  1697 
       
  1698 lemma and_nat_rec:
       
  1699   \<open>m AND n = of_bool (odd m \<and> odd n) + 2 * ((m div 2) AND (n div 2))\<close> for m n :: nat
       
  1700   by (simp add: and_nat_def and_int_rec [of \<open>int m\<close> \<open>int n\<close>] zdiv_int nat_add_distrib nat_mult_distrib)
       
  1701 
       
  1702 lemma or_nat_rec:
       
  1703   \<open>m OR n = of_bool (odd m \<or> odd n) + 2 * ((m div 2) OR (n div 2))\<close> for m n :: nat
       
  1704   by (simp add: or_nat_def or_int_rec [of \<open>int m\<close> \<open>int n\<close>] zdiv_int nat_add_distrib nat_mult_distrib)
       
  1705 
       
  1706 lemma xor_nat_rec:
       
  1707   \<open>m XOR n = of_bool (odd m \<noteq> odd n) + 2 * ((m div 2) XOR (n div 2))\<close> for m n :: nat
       
  1708   by (simp add: xor_nat_def xor_int_rec [of \<open>int m\<close> \<open>int n\<close>] zdiv_int nat_add_distrib nat_mult_distrib)
       
  1709 
       
  1710 lemma Suc_0_and_eq [simp]:
       
  1711   \<open>Suc 0 AND n = n mod 2\<close>
       
  1712   using one_and_eq [of n] by simp
       
  1713 
       
  1714 lemma and_Suc_0_eq [simp]:
       
  1715   \<open>n AND Suc 0 = n mod 2\<close>
       
  1716   using and_one_eq [of n] by simp
       
  1717 
       
  1718 lemma Suc_0_or_eq:
       
  1719   \<open>Suc 0 OR n = n + of_bool (even n)\<close>
       
  1720   using one_or_eq [of n] by simp
       
  1721 
       
  1722 lemma or_Suc_0_eq:
       
  1723   \<open>n OR Suc 0 = n + of_bool (even n)\<close>
       
  1724   using or_one_eq [of n] by simp
       
  1725 
       
  1726 lemma Suc_0_xor_eq:
       
  1727   \<open>Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)\<close>
       
  1728   using one_xor_eq [of n] by simp
       
  1729 
       
  1730 lemma xor_Suc_0_eq:
       
  1731   \<open>n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)\<close>
       
  1732   using xor_one_eq [of n] by simp
       
  1733 
       
  1734 context semiring_bit_operations
       
  1735 begin
       
  1736 
       
  1737 lemma of_nat_and_eq:
       
  1738   \<open>of_nat (m AND n) = of_nat m AND of_nat n\<close>
       
  1739   by (rule bit_eqI) (simp add: bit_of_nat_iff bit_and_iff Bit_Operations.bit_and_iff)
       
  1740 
       
  1741 lemma of_nat_or_eq:
       
  1742   \<open>of_nat (m OR n) = of_nat m OR of_nat n\<close>
       
  1743   by (rule bit_eqI) (simp add: bit_of_nat_iff bit_or_iff Bit_Operations.bit_or_iff)
       
  1744 
       
  1745 lemma of_nat_xor_eq:
       
  1746   \<open>of_nat (m XOR n) = of_nat m XOR of_nat n\<close>
       
  1747   by (rule bit_eqI) (simp add: bit_of_nat_iff bit_xor_iff Bit_Operations.bit_xor_iff)
       
  1748 
       
  1749 end
       
  1750 
       
  1751 context ring_bit_operations
       
  1752 begin
       
  1753 
       
  1754 lemma of_nat_mask_eq:
       
  1755   \<open>of_nat (mask n) = mask n\<close>
       
  1756   by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_nat_or_eq)
       
  1757 
       
  1758 end
       
  1759 
       
  1760 lemma Suc_mask_eq_exp:
       
  1761   \<open>Suc (mask n) = 2 ^ n\<close>
       
  1762   by (simp add: mask_eq_exp_minus_1)
       
  1763 
       
  1764 lemma less_eq_mask:
       
  1765   \<open>n \<le> mask n\<close>
       
  1766   by (simp add: mask_eq_exp_minus_1 le_diff_conv2)
       
  1767     (metis Suc_mask_eq_exp diff_Suc_1 diff_le_diff_pow diff_zero le_refl not_less_eq_eq power_0)
       
  1768 
       
  1769 lemma less_mask:
       
  1770   \<open>n < mask n\<close> if \<open>Suc 0 < n\<close>
       
  1771 proof -
       
  1772   define m where \<open>m = n - 2\<close>
       
  1773   with that have *: \<open>n = m + 2\<close>
       
  1774     by simp
       
  1775   have \<open>Suc (Suc (Suc m)) < 4 * 2 ^ m\<close>
       
  1776     by (induction m) simp_all
       
  1777   then have \<open>Suc (m + 2) < Suc (mask (m + 2))\<close>
       
  1778     by (simp add: Suc_mask_eq_exp)
       
  1779   then have \<open>m + 2 < mask (m + 2)\<close>
       
  1780     by (simp add: less_le)
       
  1781   with * show ?thesis
       
  1782     by simp
       
  1783 qed
       
  1784 
       
  1785 
       
  1786 subsection \<open>Symbolic computations on numeral expressions\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close>
       
  1787 
       
  1788 fun and_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close>
       
  1789 where
       
  1790   \<open>and_num num.One num.One = Some num.One\<close>
       
  1791 | \<open>and_num num.One (num.Bit0 n) = None\<close>
       
  1792 | \<open>and_num num.One (num.Bit1 n) = Some num.One\<close>
       
  1793 | \<open>and_num (num.Bit0 m) num.One = None\<close>
       
  1794 | \<open>and_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (and_num m n)\<close>
       
  1795 | \<open>and_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (and_num m n)\<close>
       
  1796 | \<open>and_num (num.Bit1 m) num.One = Some num.One\<close>
       
  1797 | \<open>and_num (num.Bit1 m) (num.Bit0 n) = map_option num.Bit0 (and_num m n)\<close>
       
  1798 | \<open>and_num (num.Bit1 m) (num.Bit1 n) = (case and_num m n of None \<Rightarrow> Some num.One | Some n' \<Rightarrow> Some (num.Bit1 n'))\<close>
       
  1799 
       
  1800 fun and_not_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close>
       
  1801 where
       
  1802   \<open>and_not_num num.One num.One = None\<close>
       
  1803 | \<open>and_not_num num.One (num.Bit0 n) = Some num.One\<close>
       
  1804 | \<open>and_not_num num.One (num.Bit1 n) = None\<close>
       
  1805 | \<open>and_not_num (num.Bit0 m) num.One = Some (num.Bit0 m)\<close>
       
  1806 | \<open>and_not_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (and_not_num m n)\<close>
       
  1807 | \<open>and_not_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (and_not_num m n)\<close>
       
  1808 | \<open>and_not_num (num.Bit1 m) num.One = Some (num.Bit0 m)\<close>
       
  1809 | \<open>and_not_num (num.Bit1 m) (num.Bit0 n) = (case and_not_num m n of None \<Rightarrow> Some num.One | Some n' \<Rightarrow> Some (num.Bit1 n'))\<close>
       
  1810 | \<open>and_not_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (and_not_num m n)\<close>
       
  1811 
       
  1812 fun or_num :: \<open>num \<Rightarrow> num \<Rightarrow> num\<close>
       
  1813 where
       
  1814   \<open>or_num num.One num.One = num.One\<close>
       
  1815 | \<open>or_num num.One (num.Bit0 n) = num.Bit1 n\<close>
       
  1816 | \<open>or_num num.One (num.Bit1 n) = num.Bit1 n\<close>
       
  1817 | \<open>or_num (num.Bit0 m) num.One = num.Bit1 m\<close>
       
  1818 | \<open>or_num (num.Bit0 m) (num.Bit0 n) = num.Bit0 (or_num m n)\<close>
       
  1819 | \<open>or_num (num.Bit0 m) (num.Bit1 n) = num.Bit1 (or_num m n)\<close>
       
  1820 | \<open>or_num (num.Bit1 m) num.One = num.Bit1 m\<close>
       
  1821 | \<open>or_num (num.Bit1 m) (num.Bit0 n) = num.Bit1 (or_num m n)\<close>
       
  1822 | \<open>or_num (num.Bit1 m) (num.Bit1 n) = num.Bit1 (or_num m n)\<close>
       
  1823 
       
  1824 fun or_not_num_neg :: \<open>num \<Rightarrow> num \<Rightarrow> num\<close>
       
  1825 where
       
  1826   \<open>or_not_num_neg num.One num.One = num.One\<close>
       
  1827 | \<open>or_not_num_neg num.One (num.Bit0 m) = num.Bit1 m\<close>
       
  1828 | \<open>or_not_num_neg num.One (num.Bit1 m) = num.Bit1 m\<close>
       
  1829 | \<open>or_not_num_neg (num.Bit0 n) num.One = num.Bit0 num.One\<close>
       
  1830 | \<open>or_not_num_neg (num.Bit0 n) (num.Bit0 m) = Num.BitM (or_not_num_neg n m)\<close>
       
  1831 | \<open>or_not_num_neg (num.Bit0 n) (num.Bit1 m) = num.Bit0 (or_not_num_neg n m)\<close>
       
  1832 | \<open>or_not_num_neg (num.Bit1 n) num.One = num.One\<close>
       
  1833 | \<open>or_not_num_neg (num.Bit1 n) (num.Bit0 m) = Num.BitM (or_not_num_neg n m)\<close>
       
  1834 | \<open>or_not_num_neg (num.Bit1 n) (num.Bit1 m) = Num.BitM (or_not_num_neg n m)\<close>
       
  1835 
       
  1836 fun xor_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close>
       
  1837 where
       
  1838   \<open>xor_num num.One num.One = None\<close>
       
  1839 | \<open>xor_num num.One (num.Bit0 n) = Some (num.Bit1 n)\<close>
       
  1840 | \<open>xor_num num.One (num.Bit1 n) = Some (num.Bit0 n)\<close>
       
  1841 | \<open>xor_num (num.Bit0 m) num.One = Some (num.Bit1 m)\<close>
       
  1842 | \<open>xor_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (xor_num m n)\<close>
       
  1843 | \<open>xor_num (num.Bit0 m) (num.Bit1 n) = Some (case xor_num m n of None \<Rightarrow> num.One | Some n' \<Rightarrow> num.Bit1 n')\<close>
       
  1844 | \<open>xor_num (num.Bit1 m) num.One = Some (num.Bit0 m)\<close>
       
  1845 | \<open>xor_num (num.Bit1 m) (num.Bit0 n) = Some (case xor_num m n of None \<Rightarrow> num.One | Some n' \<Rightarrow> num.Bit1 n')\<close>
       
  1846 | \<open>xor_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (xor_num m n)\<close>
       
  1847 
       
  1848 lemma int_numeral_and_num:
       
  1849   \<open>numeral m AND numeral n = (case and_num m n of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
       
  1850   by (induction m n rule: and_num.induct) (simp_all split: option.split)
       
  1851 
       
  1852 lemma and_num_eq_None_iff:
       
  1853   \<open>and_num m n = None \<longleftrightarrow> numeral m AND numeral n = (0::int)\<close>
       
  1854   by (simp add: int_numeral_and_num split: option.split)
       
  1855 
       
  1856 lemma and_num_eq_Some_iff:
       
  1857   \<open>and_num m n = Some q \<longleftrightarrow> numeral m AND numeral n = (numeral q :: int)\<close>
       
  1858   by (simp add: int_numeral_and_num split: option.split)
       
  1859 
       
  1860 lemma int_numeral_and_not_num:
       
  1861   \<open>numeral m AND NOT (numeral n) = (case and_not_num m n of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
       
  1862   by (induction m n rule: and_not_num.induct) (simp_all add: add_One BitM_inc_eq not_int_def split: option.split)
       
  1863 
       
  1864 lemma int_numeral_not_and_num:
       
  1865   \<open>NOT (numeral m) AND numeral n = (case and_not_num n m of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
       
  1866   using int_numeral_and_not_num [of n m] by (simp add: ac_simps)
       
  1867 
       
  1868 lemma and_not_num_eq_None_iff:
       
  1869   \<open>and_not_num m n = None \<longleftrightarrow> numeral m AND NOT (numeral n) = (0::int)\<close>
       
  1870   by (simp add: int_numeral_and_not_num split: option.split)
       
  1871 
       
  1872 lemma and_not_num_eq_Some_iff:
       
  1873   \<open>and_not_num m n = Some q \<longleftrightarrow> numeral m AND NOT (numeral n) = (numeral q :: int)\<close>
       
  1874   by (simp add: int_numeral_and_not_num split: option.split)
       
  1875 
       
  1876 lemma int_numeral_or_num:
       
  1877   \<open>numeral m OR numeral n = (numeral (or_num m n) :: int)\<close>
       
  1878   by (induction m n rule: or_num.induct) simp_all
       
  1879 
       
  1880 lemma numeral_or_num_eq:
       
  1881   \<open>numeral (or_num m n) = (numeral m OR numeral n :: int)\<close>
       
  1882   by (simp add: int_numeral_or_num)
       
  1883 
       
  1884 lemma int_numeral_or_not_num_neg:
       
  1885   \<open>numeral m OR NOT (numeral n :: int) = - numeral (or_not_num_neg m n)\<close>
       
  1886   by (induction m n rule: or_not_num_neg.induct) (simp_all add: add_One BitM_inc_eq not_int_def)
       
  1887 
       
  1888 lemma int_numeral_not_or_num_neg:
       
  1889   \<open>NOT (numeral m) OR (numeral n :: int) = - numeral (or_not_num_neg n m)\<close>
       
  1890   using int_numeral_or_not_num_neg [of n m] by (simp add: ac_simps)
       
  1891 
       
  1892 lemma numeral_or_not_num_eq:
       
  1893   \<open>numeral (or_not_num_neg m n) = - (numeral m OR NOT (numeral n :: int))\<close>
       
  1894   using int_numeral_or_not_num_neg [of m n] by simp
       
  1895 
       
  1896 lemma int_numeral_xor_num:
       
  1897   \<open>numeral m XOR numeral n = (case xor_num m n of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
       
  1898   by (induction m n rule: xor_num.induct) (simp_all split: option.split)
       
  1899 
       
  1900 lemma xor_num_eq_None_iff:
       
  1901   \<open>xor_num m n = None \<longleftrightarrow> numeral m XOR numeral n = (0::int)\<close>
       
  1902   by (simp add: int_numeral_xor_num split: option.split)
       
  1903 
       
  1904 lemma xor_num_eq_Some_iff:
       
  1905   \<open>xor_num m n = Some q \<longleftrightarrow> numeral m XOR numeral n = (numeral q :: int)\<close>
       
  1906   by (simp add: int_numeral_xor_num split: option.split)
       
  1907 
       
  1908 
       
  1909 subsection \<open>Instances for \<^typ>\<open>integer\<close> and \<^typ>\<open>natural\<close>\<close>
       
  1910 
       
  1911 unbundle integer.lifting natural.lifting
       
  1912 
       
  1913 instantiation integer :: ring_bit_operations
       
  1914 begin
       
  1915 
       
  1916 lift_definition not_integer :: \<open>integer \<Rightarrow> integer\<close>
       
  1917   is not .
       
  1918 
       
  1919 lift_definition and_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
       
  1920   is \<open>and\<close> .
       
  1921 
       
  1922 lift_definition or_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
       
  1923   is or .
       
  1924 
       
  1925 lift_definition xor_integer ::  \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
       
  1926   is xor .
       
  1927 
       
  1928 lift_definition mask_integer :: \<open>nat \<Rightarrow> integer\<close>
       
  1929   is mask .
       
  1930 
       
  1931 lift_definition set_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>
       
  1932   is set_bit .
       
  1933 
       
  1934 lift_definition unset_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>
       
  1935   is unset_bit .
       
  1936 
       
  1937 lift_definition flip_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>
       
  1938   is flip_bit .
       
  1939 
       
  1940 instance by (standard; transfer)
       
  1941   (simp_all add: minus_eq_not_minus_1 mask_eq_exp_minus_1
       
  1942     bit_not_iff bit_and_iff bit_or_iff bit_xor_iff
       
  1943     set_bit_def bit_unset_bit_iff flip_bit_def)
       
  1944 
       
  1945 end
       
  1946 
       
  1947 lemma [code]:
       
  1948   \<open>mask n = 2 ^ n - (1::integer)\<close>
       
  1949   by (simp add: mask_eq_exp_minus_1)
       
  1950 
       
  1951 instantiation natural :: semiring_bit_operations
       
  1952 begin
       
  1953 
       
  1954 lift_definition and_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>
       
  1955   is \<open>and\<close> .
       
  1956 
       
  1957 lift_definition or_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>
       
  1958   is or .
       
  1959 
       
  1960 lift_definition xor_natural ::  \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>
       
  1961   is xor .
       
  1962 
       
  1963 lift_definition mask_natural :: \<open>nat \<Rightarrow> natural\<close>
       
  1964   is mask .
       
  1965 
       
  1966 lift_definition set_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>
       
  1967   is set_bit .
       
  1968 
       
  1969 lift_definition unset_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>
       
  1970   is unset_bit .
       
  1971 
       
  1972 lift_definition flip_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>
       
  1973   is flip_bit .
       
  1974 
       
  1975 instance by (standard; transfer)
       
  1976   (simp_all add: mask_eq_exp_minus_1
       
  1977     bit_and_iff bit_or_iff bit_xor_iff
       
  1978     set_bit_def bit_unset_bit_iff flip_bit_def)
       
  1979 
       
  1980 end
       
  1981 
       
  1982 lemma [code]:
       
  1983   \<open>integer_of_natural (mask n) = mask n\<close>
       
  1984   by transfer (simp add: mask_eq_exp_minus_1 of_nat_diff)
       
  1985 
       
  1986 lifting_update integer.lifting
       
  1987 lifting_forget integer.lifting
       
  1988 
       
  1989 lifting_update natural.lifting
       
  1990 lifting_forget natural.lifting
       
  1991 
       
  1992 
       
  1993 subsection \<open>Key ideas of bit operations\<close>
       
  1994 
       
  1995 text \<open>
       
  1996   When formalizing bit operations, it is tempting to represent
       
  1997   bit values as explicit lists over a binary type. This however
       
  1998   is a bad idea, mainly due to the inherent ambiguities in
       
  1999   representation concerning repeating leading bits.
       
  2000 
       
  2001   Hence this approach avoids such explicit lists altogether
       
  2002   following an algebraic path:
       
  2003 
       
  2004   \<^item> Bit values are represented by numeric types: idealized
       
  2005     unbounded bit values can be represented by type \<^typ>\<open>int\<close>,
       
  2006     bounded bit values by quotient types over \<^typ>\<open>int\<close>.
       
  2007 
       
  2008   \<^item> (A special case are idealized unbounded bit values ending
       
  2009     in @{term [source] 0} which can be represented by type \<^typ>\<open>nat\<close> but
       
  2010     only support a restricted set of operations).
       
  2011 
       
  2012   \<^item> From this idea follows that
       
  2013 
       
  2014       \<^item> multiplication by \<^term>\<open>2 :: int\<close> is a bit shift to the left and
       
  2015 
       
  2016       \<^item> division by \<^term>\<open>2 :: int\<close> is a bit shift to the right.
       
  2017 
       
  2018   \<^item> Concerning bounded bit values, iterated shifts to the left
       
  2019     may result in eliminating all bits by shifting them all
       
  2020     beyond the boundary.  The property \<^prop>\<open>(2 :: int) ^ n \<noteq> 0\<close>
       
  2021     represents that \<^term>\<open>n\<close> is \<^emph>\<open>not\<close> beyond that boundary.
       
  2022 
       
  2023   \<^item> The projection on a single bit is then @{thm bit_iff_odd [where ?'a = int, no_vars]}.
       
  2024 
       
  2025   \<^item> This leads to the most fundamental properties of bit values:
       
  2026 
       
  2027       \<^item> Equality rule: @{thm bit_eqI [where ?'a = int, no_vars]}
       
  2028 
       
  2029       \<^item> Induction rule: @{thm bits_induct [where ?'a = int, no_vars]}
       
  2030 
       
  2031   \<^item> Typical operations are characterized as follows:
       
  2032 
       
  2033       \<^item> Singleton \<^term>\<open>n\<close>th bit: \<^term>\<open>(2 :: int) ^ n\<close>
       
  2034 
       
  2035       \<^item> Bit mask upto bit \<^term>\<open>n\<close>: @{thm mask_eq_exp_minus_1 [where ?'a = int, no_vars]}
       
  2036 
       
  2037       \<^item> Left shift: @{thm push_bit_eq_mult [where ?'a = int, no_vars]}
       
  2038 
       
  2039       \<^item> Right shift: @{thm drop_bit_eq_div [where ?'a = int, no_vars]}
       
  2040 
       
  2041       \<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]}
       
  2042 
       
  2043       \<^item> Negation: @{thm bit_not_iff [where ?'a = int, no_vars]}
       
  2044 
       
  2045       \<^item> And: @{thm bit_and_iff [where ?'a = int, no_vars]}
       
  2046 
       
  2047       \<^item> Or: @{thm bit_or_iff [where ?'a = int, no_vars]}
       
  2048 
       
  2049       \<^item> Xor: @{thm bit_xor_iff [where ?'a = int, no_vars]}
       
  2050 
       
  2051       \<^item> Set a single bit: @{thm set_bit_def [where ?'a = int, no_vars]}
       
  2052 
       
  2053       \<^item> Unset a single bit: @{thm unset_bit_def [where ?'a = int, no_vars]}
       
  2054 
       
  2055       \<^item> Flip a single bit: @{thm flip_bit_def [where ?'a = int, no_vars]}
       
  2056 
       
  2057       \<^item> Signed truncation, or modulus centered around \<^term>\<open>0::int\<close>: @{thm signed_take_bit_def [no_vars]}
       
  2058 
       
  2059       \<^item> Bit concatenation: @{thm concat_bit_def [no_vars]}
       
  2060 
       
  2061       \<^item> (Bounded) conversion from and to a list of bits: @{thm horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]}
       
  2062 \<close>
       
  2063 
       
  2064 code_identifier
       
  2065   type_class semiring_bits \<rightharpoonup>
       
  2066   (SML) Bit_Operations.semiring_bits and (OCaml) Bit_Operations.semiring_bits and (Haskell) Bit_Operations.semiring_bits and (Scala) Bit_Operations.semiring_bits
       
  2067 | class_relation semiring_bits < semiring_parity \<rightharpoonup>
       
  2068   (SML) Bit_Operations.semiring_parity_semiring_bits and (OCaml) Bit_Operations.semiring_parity_semiring_bits and (Haskell) Bit_Operations.semiring_parity_semiring_bits and (Scala) Bit_Operations.semiring_parity_semiring_bits
       
  2069 | constant bit \<rightharpoonup>
       
  2070   (SML) Bit_Operations.bit and (OCaml) Bit_Operations.bit and (Haskell) Bit_Operations.bit and (Scala) Bit_Operations.bit
       
  2071 | class_instance nat :: semiring_bits \<rightharpoonup>
       
  2072   (SML) Bit_Operations.semiring_bits_nat and (OCaml) Bit_Operations.semiring_bits_nat and (Haskell) Bit_Operations.semiring_bits_nat and (Scala) Bit_Operations.semiring_bits_nat
       
  2073 | class_instance int :: semiring_bits \<rightharpoonup>
       
  2074   (SML) Bit_Operations.semiring_bits_int and (OCaml) Bit_Operations.semiring_bits_int and (Haskell) Bit_Operations.semiring_bits_int and (Scala) Bit_Operations.semiring_bits_int
       
  2075 | type_class semiring_bit_shifts \<rightharpoonup>
       
  2076   (SML) Bit_Operations.semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bits and (Scala) Bit_Operations.semiring_bit_shifts
       
  2077 | class_relation semiring_bit_shifts < semiring_bits \<rightharpoonup>
       
  2078   (SML) Bit_Operations.semiring_bits_semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bits_semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bits_semiring_bit_shifts and (Scala) Bit_Operations.semiring_bits_semiring_bit_shifts
       
  2079 | constant push_bit \<rightharpoonup>
       
  2080   (SML) Bit_Operations.push_bit and (OCaml) Bit_Operations.push_bit and (Haskell) Bit_Operations.push_bit and (Scala) Bit_Operations.push_bit
       
  2081 | constant drop_bit \<rightharpoonup>
       
  2082   (SML) Bit_Operations.drop_bit and (OCaml) Bit_Operations.drop_bit and (Haskell) Bit_Operations.drop_bit and (Scala) Bit_Operations.drop_bit
       
  2083 | constant take_bit \<rightharpoonup>
       
  2084   (SML) Bit_Operations.take_bit and (OCaml) Bit_Operations.take_bit and (Haskell) Bit_Operations.take_bit and (Scala) Bit_Operations.take_bit
       
  2085 | class_instance nat :: semiring_bit_shifts \<rightharpoonup>
       
  2086   (SML) Bit_Operations.semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bit_shifts and (Scala) Bit_Operations.semiring_bit_shifts
       
  2087 | class_instance int :: semiring_bit_shifts \<rightharpoonup>
       
  2088   (SML) Bit_Operations.semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bit_shifts and (Scala) Bit_Operations.semiring_bit_shifts
       
  2089 
       
  2090 no_notation
       
  2091   "and"  (infixr \<open>AND\<close> 64)
       
  2092     and or  (infixr \<open>OR\<close>  59)
       
  2093     and xor  (infixr \<open>XOR\<close> 59)
       
  2094 
       
  2095 bundle bit_operations_syntax
       
  2096 begin      
       
  2097 
       
  2098 notation
       
  2099   "and"  (infixr \<open>AND\<close> 64)
       
  2100     and or  (infixr \<open>OR\<close>  59)
       
  2101     and xor  (infixr \<open>XOR\<close> 59)
       
  2102 
       
  2103 end
       
  2104 
       
  2105 end