src/HOL/Word/BinOperations.thy
changeset 24333 e77ea0ea7f2c
child 24350 4d74f37c6367
equal deleted inserted replaced
24332:e3a2b75b1cf9 24333:e77ea0ea7f2c
       
     1 (* 
       
     2   ID:     $Id$
       
     3   Author: Jeremy Dawson and Gerwin Klein, NICTA
       
     4 
       
     5   definition and basic theorems for bit-wise logical operations 
       
     6   for integers expressed using Pls, Min, BIT,
       
     7   and converting them to and from lists of bools
       
     8 *) 
       
     9 
       
    10 theory BinOperations imports BinGeneral
       
    11 
       
    12 begin
       
    13 
       
    14 -- "bit-wise logical operations on the int type"
       
    15 consts
       
    16   int_and :: "int => int => int"
       
    17   int_or :: "int => int => int"
       
    18   bit_not :: "bit => bit"
       
    19   bit_and :: "bit => bit => bit"
       
    20   bit_or :: "bit => bit => bit"
       
    21   bit_xor :: "bit => bit => bit"
       
    22   int_not :: "int => int"
       
    23   int_xor :: "int => int => int"
       
    24   bin_sc :: "nat => bit => int => int"
       
    25 
       
    26 primrec
       
    27   B0 : "bit_not bit.B0 = bit.B1"
       
    28   B1 : "bit_not bit.B1 = bit.B0"
       
    29 
       
    30 primrec
       
    31   B1 : "bit_xor bit.B1 x = bit_not x"
       
    32   B0 : "bit_xor bit.B0 x = x"
       
    33 
       
    34 primrec
       
    35   B1 : "bit_or bit.B1 x = bit.B1"
       
    36   B0 : "bit_or bit.B0 x = x"
       
    37 
       
    38 primrec
       
    39   B0 : "bit_and bit.B0 x = bit.B0"
       
    40   B1 : "bit_and bit.B1 x = x"
       
    41 
       
    42 primrec
       
    43   Z : "bin_sc 0 b w = bin_rest w BIT b"
       
    44   Suc :
       
    45     "bin_sc (Suc n) b w = bin_sc n b (bin_rest w) BIT bin_last w"
       
    46 
       
    47 defs
       
    48   int_not_def : "int_not == bin_rec Numeral.Min Numeral.Pls 
       
    49     (%w b s. s BIT bit_not b)"
       
    50     int_and_def : "int_and == bin_rec (%x. Numeral.Pls) (%y. y) 
       
    51     (%w b s y. s (bin_rest y) BIT (bit_and b (bin_last y)))"
       
    52   int_or_def : "int_or == bin_rec (%x. x) (%y. Numeral.Min) 
       
    53     (%w b s y. s (bin_rest y) BIT (bit_or b (bin_last y)))"
       
    54   int_xor_def : "int_xor == bin_rec (%x. x) int_not 
       
    55     (%w b s y. s (bin_rest y) BIT (bit_xor b (bin_last y)))"
       
    56 
       
    57 consts
       
    58   bin_to_bl :: "nat => int => bool list"
       
    59   bin_to_bl_aux :: "nat => int => bool list => bool list"
       
    60   bl_to_bin :: "bool list => int"
       
    61   bl_to_bin_aux :: "int => bool list => int"
       
    62 
       
    63   bl_of_nth :: "nat => (nat => bool) => bool list"
       
    64 
       
    65 primrec
       
    66   Nil : "bl_to_bin_aux w [] = w"
       
    67   Cons : "bl_to_bin_aux w (b # bs) = 
       
    68       bl_to_bin_aux (w BIT (if b then bit.B1 else bit.B0)) bs"
       
    69 
       
    70 primrec
       
    71   Z : "bin_to_bl_aux 0 w bl = bl"
       
    72   Suc : "bin_to_bl_aux (Suc n) w bl =
       
    73     bin_to_bl_aux n (bin_rest w) ((bin_last w = bit.B1) # bl)"
       
    74 
       
    75 defs
       
    76   bin_to_bl_def : "bin_to_bl n w == bin_to_bl_aux n w []"
       
    77   bl_to_bin_def : "bl_to_bin bs == bl_to_bin_aux Numeral.Pls bs"
       
    78 
       
    79 primrec
       
    80   Suc : "bl_of_nth (Suc n) f = f n # bl_of_nth n f"
       
    81   Z : "bl_of_nth 0 f = []"
       
    82 
       
    83 consts
       
    84   takefill :: "'a => nat => 'a list => 'a list"
       
    85   app2 :: "('a => 'b => 'c) => 'a list => 'b list => 'c list"
       
    86 
       
    87 -- "takefill - like take but if argument list too short,"
       
    88 -- "extends result to get requested length"
       
    89 primrec
       
    90   Z : "takefill fill 0 xs = []"
       
    91   Suc : "takefill fill (Suc n) xs = (
       
    92     case xs of [] => fill # takefill fill n xs
       
    93       | y # ys => y # takefill fill n ys)"
       
    94 
       
    95 defs
       
    96   app2_def : "app2 f as bs == map (split f) (zip as bs)"
       
    97 
       
    98 -- "rcat and rsplit"
       
    99 consts
       
   100   bin_rcat :: "nat => int list => int"
       
   101   bin_rsplit_aux :: "nat * int list * nat * int => int list"
       
   102   bin_rsplit :: "nat => (nat * int) => int list"
       
   103   bin_rsplitl_aux :: "nat * int list * nat * int => int list"
       
   104   bin_rsplitl :: "nat => (nat * int) => int list"
       
   105   
       
   106 recdef bin_rsplit_aux "measure (fst o snd o snd)"
       
   107   "bin_rsplit_aux (n, bs, (m, c)) =
       
   108     (if m = 0 | n = 0 then bs else
       
   109       let (a, b) = bin_split n c 
       
   110       in bin_rsplit_aux (n, b # bs, (m - n, a)))"
       
   111 
       
   112 recdef bin_rsplitl_aux "measure (fst o snd o snd)"
       
   113   "bin_rsplitl_aux (n, bs, (m, c)) =
       
   114     (if m = 0 | n = 0 then bs else
       
   115       let (a, b) = bin_split (min m n) c 
       
   116       in bin_rsplitl_aux (n, b # bs, (m - n, a)))"
       
   117 
       
   118 defs
       
   119   bin_rcat_def : "bin_rcat n bs == foldl (%u v. bin_cat u n v) Numeral.Pls bs"
       
   120   bin_rsplit_def : "bin_rsplit n w == bin_rsplit_aux (n, [], w)"
       
   121   bin_rsplitl_def : "bin_rsplitl n w == bin_rsplitl_aux (n, [], w)"
       
   122      
       
   123  
       
   124 lemma int_not_simps [simp]:
       
   125   "int_not Numeral.Pls = Numeral.Min"
       
   126   "int_not Numeral.Min = Numeral.Pls"
       
   127   "int_not (w BIT b) = int_not w BIT bit_not b"
       
   128   by (unfold int_not_def) (auto intro: bin_rec_simps)
       
   129 
       
   130 lemma bit_extra_simps [simp]: 
       
   131   "bit_and x bit.B0 = bit.B0"
       
   132   "bit_and x bit.B1 = x"
       
   133   "bit_or x bit.B1 = bit.B1"
       
   134   "bit_or x bit.B0 = x"
       
   135   "bit_xor x bit.B1 = bit_not x"
       
   136   "bit_xor x bit.B0 = x"
       
   137   by (cases x, auto)+
       
   138 
       
   139 lemma bit_ops_comm: 
       
   140   "bit_and x y = bit_and y x"
       
   141   "bit_or x y = bit_or y x"
       
   142   "bit_xor x y = bit_xor y x"
       
   143   by (cases y, auto)+
       
   144 
       
   145 lemma bit_ops_same [simp]: 
       
   146   "bit_and x x = x"
       
   147   "bit_or x x = x"
       
   148   "bit_xor x x = bit.B0"
       
   149   by (cases x, auto)+
       
   150 
       
   151 lemma bit_not_not [simp]: "bit_not (bit_not x) = x"
       
   152   by (cases x) auto
       
   153 
       
   154 lemma int_xor_Pls [simp]: 
       
   155   "int_xor Numeral.Pls x = x"
       
   156   unfolding int_xor_def by (simp add: bin_rec_PM)
       
   157 
       
   158 lemma int_xor_Min [simp]: 
       
   159   "int_xor Numeral.Min x = int_not x"
       
   160   unfolding int_xor_def by (simp add: bin_rec_PM)
       
   161 
       
   162 lemma int_xor_Bits [simp]: 
       
   163   "int_xor (x BIT b) (y BIT c) = int_xor x y BIT bit_xor b c"
       
   164   apply (unfold int_xor_def)
       
   165   apply (rule bin_rec_simps (1) [THEN fun_cong, THEN trans])
       
   166     apply (rule ext, simp)
       
   167    prefer 2
       
   168    apply simp
       
   169   apply (rule ext)
       
   170   apply (simp add: int_not_simps [symmetric])
       
   171   done
       
   172 
       
   173 lemma int_xor_x_simps':
       
   174   "int_xor w (Numeral.Pls BIT bit.B0) = w"
       
   175   "int_xor w (Numeral.Min BIT bit.B1) = int_not w"
       
   176   apply (induct w rule: bin_induct)
       
   177        apply simp_all[4]
       
   178    apply (unfold int_xor_Bits)
       
   179    apply clarsimp+
       
   180   done
       
   181 
       
   182 lemmas int_xor_extra_simps [simp] = int_xor_x_simps' [simplified arith_simps]
       
   183 
       
   184 lemma int_or_Pls [simp]: 
       
   185   "int_or Numeral.Pls x = x"
       
   186   by (unfold int_or_def) (simp add: bin_rec_PM)
       
   187   
       
   188 lemma int_or_Min [simp]:
       
   189   "int_or Numeral.Min x = Numeral.Min"
       
   190   by (unfold int_or_def) (simp add: bin_rec_PM)
       
   191 
       
   192 lemma int_or_Bits [simp]: 
       
   193   "int_or (x BIT b) (y BIT c) = int_or x y BIT bit_or b c"
       
   194   unfolding int_or_def by (simp add: bin_rec_simps)
       
   195 
       
   196 lemma int_or_x_simps': 
       
   197   "int_or w (Numeral.Pls BIT bit.B0) = w"
       
   198   "int_or w (Numeral.Min BIT bit.B1) = Numeral.Min"
       
   199   apply (induct w rule: bin_induct)
       
   200        apply simp_all[4]
       
   201    apply (unfold int_or_Bits)
       
   202    apply clarsimp+
       
   203   done
       
   204 
       
   205 lemmas int_or_extra_simps [simp] = int_or_x_simps' [simplified arith_simps]
       
   206 
       
   207 
       
   208 lemma int_and_Pls [simp]:
       
   209   "int_and Numeral.Pls x = Numeral.Pls"
       
   210   unfolding int_and_def by (simp add: bin_rec_PM)
       
   211 
       
   212 lemma  int_and_Min [simp]:
       
   213   "int_and Numeral.Min x = x"
       
   214   unfolding int_and_def by (simp add: bin_rec_PM)
       
   215 
       
   216 lemma int_and_Bits [simp]: 
       
   217   "int_and (x BIT b) (y BIT c) = int_and x y BIT bit_and b c" 
       
   218   unfolding int_and_def by (simp add: bin_rec_simps)
       
   219 
       
   220 lemma int_and_x_simps': 
       
   221   "int_and w (Numeral.Pls BIT bit.B0) = Numeral.Pls"
       
   222   "int_and w (Numeral.Min BIT bit.B1) = w"
       
   223   apply (induct w rule: bin_induct)
       
   224        apply simp_all[4]
       
   225    apply (unfold int_and_Bits)
       
   226    apply clarsimp+
       
   227   done
       
   228 
       
   229 lemmas int_and_extra_simps [simp] = int_and_x_simps' [simplified arith_simps]
       
   230 
       
   231 (* commutativity of the above *)
       
   232 lemma bin_ops_comm:
       
   233   shows
       
   234   int_and_comm: "!!y. int_and x y = int_and y x" and
       
   235   int_or_comm:  "!!y. int_or x y = int_or y x" and
       
   236   int_xor_comm: "!!y. int_xor x y = int_xor y x"
       
   237   apply (induct x rule: bin_induct)
       
   238           apply simp_all[6]
       
   239     apply (case_tac y rule: bin_exhaust, simp add: bit_ops_comm)+
       
   240   done
       
   241 
       
   242 lemma bin_ops_same [simp]:
       
   243   "int_and x x = x" 
       
   244   "int_or x x = x" 
       
   245   "int_xor x x = Numeral.Pls"
       
   246   by (induct x rule: bin_induct) auto
       
   247 
       
   248 lemma int_not_not [simp]: "int_not (int_not x) = x"
       
   249   by (induct x rule: bin_induct) auto
       
   250 
       
   251 lemmas bin_log_esimps = 
       
   252   int_and_extra_simps  int_or_extra_simps  int_xor_extra_simps
       
   253   int_and_Pls int_and_Min  int_or_Pls int_or_Min  int_xor_Pls int_xor_Min
       
   254 
       
   255 (* potential for looping *)
       
   256 declare bin_rsplit_aux.simps [simp del]
       
   257 declare bin_rsplitl_aux.simps [simp del]
       
   258 
       
   259 
       
   260 lemma bin_sign_cat: 
       
   261   "!!y. bin_sign (bin_cat x n y) = bin_sign x"
       
   262   by (induct n) auto
       
   263 
       
   264 lemma bin_cat_Suc_Bit:
       
   265   "bin_cat w (Suc n) (v BIT b) = bin_cat w n v BIT b"
       
   266   by auto
       
   267 
       
   268 lemma bin_nth_cat: 
       
   269   "!!n y. bin_nth (bin_cat x k y) n = 
       
   270     (if n < k then bin_nth y n else bin_nth x (n - k))"
       
   271   apply (induct k)
       
   272    apply clarsimp
       
   273   apply (case_tac n, auto)
       
   274   done
       
   275 
       
   276 lemma bin_nth_split:
       
   277   "!!b c. bin_split n c = (a, b) ==> 
       
   278     (ALL k. bin_nth a k = bin_nth c (n + k)) & 
       
   279     (ALL k. bin_nth b k = (k < n & bin_nth c k))"
       
   280   apply (induct n)
       
   281    apply clarsimp
       
   282   apply (clarsimp simp: Let_def split: ls_splits)
       
   283   apply (case_tac k)
       
   284   apply auto
       
   285   done
       
   286 
       
   287 lemma bin_cat_assoc: 
       
   288   "!!z. bin_cat (bin_cat x m y) n z = bin_cat x (m + n) (bin_cat y n z)" 
       
   289   by (induct n) auto
       
   290 
       
   291 lemma bin_cat_assoc_sym: "!!z m. 
       
   292   bin_cat x m (bin_cat y n z) = bin_cat (bin_cat x (m - n) y) (min m n) z"
       
   293   apply (induct n, clarsimp)
       
   294   apply (case_tac m, auto)
       
   295   done
       
   296 
       
   297 lemma bin_cat_Pls [simp]: 
       
   298   "!!w. bin_cat Numeral.Pls n w = bintrunc n w"
       
   299   by (induct n) auto
       
   300 
       
   301 lemma bintr_cat1: 
       
   302   "!!b. bintrunc (k + n) (bin_cat a n b) = bin_cat (bintrunc k a) n b"
       
   303   by (induct n) auto
       
   304     
       
   305 lemma bintr_cat: "bintrunc m (bin_cat a n b) = 
       
   306     bin_cat (bintrunc (m - n) a) n (bintrunc (min m n) b)"
       
   307   by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr)
       
   308     
       
   309 lemma bintr_cat_same [simp]: 
       
   310   "bintrunc n (bin_cat a n b) = bintrunc n b"
       
   311   by (auto simp add : bintr_cat)
       
   312 
       
   313 lemma cat_bintr [simp]: 
       
   314   "!!b. bin_cat a n (bintrunc n b) = bin_cat a n b"
       
   315   by (induct n) auto
       
   316 
       
   317 lemma split_bintrunc: 
       
   318   "!!b c. bin_split n c = (a, b) ==> b = bintrunc n c"
       
   319   by (induct n) (auto simp: Let_def split: ls_splits)
       
   320 
       
   321 lemma bin_cat_split:
       
   322   "!!v w. bin_split n w = (u, v) ==> w = bin_cat u n v"
       
   323   by (induct n) (auto simp: Let_def split: ls_splits)
       
   324 
       
   325 lemma bin_split_cat:
       
   326   "!!w. bin_split n (bin_cat v n w) = (v, bintrunc n w)"
       
   327   by (induct n) auto
       
   328 
       
   329 lemma bin_split_Pls [simp]:
       
   330   "bin_split n Numeral.Pls = (Numeral.Pls, Numeral.Pls)"
       
   331   by (induct n) (auto simp: Let_def split: ls_splits)
       
   332 
       
   333 lemma bin_split_Min [simp]:
       
   334   "bin_split n Numeral.Min = (Numeral.Min, bintrunc n Numeral.Min)"
       
   335   by (induct n) (auto simp: Let_def split: ls_splits)
       
   336 
       
   337 lemma bin_split_trunc:
       
   338   "!!m b c. bin_split (min m n) c = (a, b) ==> 
       
   339     bin_split n (bintrunc m c) = (bintrunc (m - n) a, b)"
       
   340   apply (induct n, clarsimp)
       
   341   apply (simp add: bin_rest_trunc Let_def split: ls_splits)
       
   342   apply (case_tac m)
       
   343    apply (auto simp: Let_def split: ls_splits)
       
   344   done
       
   345 
       
   346 lemma bin_split_trunc1:
       
   347   "!!m b c. bin_split n c = (a, b) ==> 
       
   348     bin_split n (bintrunc m c) = (bintrunc (m - n) a, bintrunc m b)"
       
   349   apply (induct n, clarsimp)
       
   350   apply (simp add: bin_rest_trunc Let_def split: ls_splits)
       
   351   apply (case_tac m)
       
   352    apply (auto simp: Let_def split: ls_splits)
       
   353   done
       
   354 
       
   355 lemma bin_cat_num:
       
   356   "!!b. bin_cat a n b = a * 2 ^ n + bintrunc n b"
       
   357   apply (induct n, clarsimp)
       
   358   apply (simp add: Bit_def cong: number_of_False_cong)
       
   359   done
       
   360 
       
   361 lemma bin_split_num:
       
   362   "!!b. bin_split n b = (b div 2 ^ n, b mod 2 ^ n)"
       
   363   apply (induct n, clarsimp)
       
   364   apply (simp add: bin_rest_div zdiv_zmult2_eq)
       
   365   apply (case_tac b rule: bin_exhaust)
       
   366   apply simp
       
   367   apply (simp add: Bit_def zmod_zmult_zmult1 p1mod22k
       
   368               split: bit.split 
       
   369               cong: number_of_False_cong)
       
   370   done 
       
   371 
       
   372 
       
   373 (* basic properties of logical (bit-wise) operations *)
       
   374 
       
   375 lemma bbw_ao_absorb: 
       
   376   "!!y. int_and x (int_or y x) = x & int_or x (int_and y x) = x"
       
   377   apply (induct x rule: bin_induct)
       
   378     apply auto 
       
   379    apply (case_tac [!] y rule: bin_exhaust)
       
   380    apply auto
       
   381    apply (case_tac [!] bit)
       
   382      apply auto
       
   383   done
       
   384 
       
   385 lemma bbw_ao_absorbs_other:
       
   386   "int_and x (int_or x y) = x \<and> int_or (int_and y x) x = x"
       
   387   "int_and (int_or y x) x = x \<and> int_or x (int_and x y) = x"
       
   388   "int_and (int_or x y) x = x \<and> int_or (int_and x y) x = x"
       
   389   apply (auto simp: bbw_ao_absorb int_or_comm)  
       
   390       apply (subst int_or_comm)
       
   391     apply (simp add: bbw_ao_absorb)
       
   392    apply (subst int_and_comm)
       
   393    apply (subst int_or_comm)
       
   394    apply (simp add: bbw_ao_absorb)
       
   395   apply (subst int_and_comm)
       
   396   apply (simp add: bbw_ao_absorb)
       
   397   done
       
   398   
       
   399 lemmas bbw_ao_absorbs [simp] = bbw_ao_absorb bbw_ao_absorbs_other
       
   400 
       
   401 lemma int_xor_not:
       
   402   "!!y. int_xor (int_not x) y = int_not (int_xor x y) & 
       
   403         int_xor x (int_not y) = int_not (int_xor x y)"
       
   404   apply (induct x rule: bin_induct)
       
   405     apply auto
       
   406    apply (case_tac y rule: bin_exhaust, auto, 
       
   407           case_tac b, auto)+
       
   408   done
       
   409 
       
   410 lemma bbw_assocs': 
       
   411   "!!y z. int_and (int_and x y) z = int_and x (int_and y z) & 
       
   412           int_or (int_or x y) z = int_or x (int_or y z) & 
       
   413           int_xor (int_xor x y) z = int_xor x (int_xor y z)"
       
   414   apply (induct x rule: bin_induct)
       
   415     apply (auto simp: int_xor_not)
       
   416     apply (case_tac [!] y rule: bin_exhaust)
       
   417     apply (case_tac [!] z rule: bin_exhaust)
       
   418     apply (case_tac [!] bit)
       
   419        apply (case_tac [!] b)
       
   420              apply auto
       
   421   done
       
   422 
       
   423 lemma int_and_assoc:
       
   424   "int_and (int_and x y) z = int_and x (int_and y z)"
       
   425   by (simp add: bbw_assocs')
       
   426 
       
   427 lemma int_or_assoc:
       
   428   "int_or (int_or x y) z = int_or x (int_or y z)"
       
   429   by (simp add: bbw_assocs')
       
   430 
       
   431 lemma int_xor_assoc:
       
   432   "int_xor (int_xor x y) z = int_xor x (int_xor y z)"
       
   433   by (simp add: bbw_assocs')
       
   434 
       
   435 lemmas bbw_assocs = int_and_assoc int_or_assoc int_xor_assoc
       
   436 
       
   437 lemma bbw_lcs [simp]: 
       
   438   "int_and y (int_and x z) = int_and x (int_and y z)"
       
   439   "int_or y (int_or x z) = int_or x (int_or y z)"
       
   440   "int_xor y (int_xor x z) = int_xor x (int_xor y z)" 
       
   441   apply (auto simp: bbw_assocs [symmetric])
       
   442   apply (auto simp: bin_ops_comm)
       
   443   done
       
   444 
       
   445 lemma bbw_not_dist: 
       
   446   "!!y. int_not (int_or x y) = int_and (int_not x) (int_not y)" 
       
   447   "!!y. int_not (int_and x y) = int_or (int_not x) (int_not y)"
       
   448   apply (induct x rule: bin_induct)
       
   449        apply auto
       
   450    apply (case_tac [!] y rule: bin_exhaust)
       
   451    apply (case_tac [!] bit, auto)
       
   452   done
       
   453 
       
   454 lemma bbw_oa_dist: 
       
   455   "!!y z. int_or (int_and x y) z = 
       
   456           int_and (int_or x z) (int_or y z)"
       
   457   apply (induct x rule: bin_induct)
       
   458     apply auto
       
   459   apply (case_tac y rule: bin_exhaust)
       
   460   apply (case_tac z rule: bin_exhaust)
       
   461   apply (case_tac ba, auto)
       
   462   done
       
   463 
       
   464 lemma bbw_ao_dist: 
       
   465   "!!y z. int_and (int_or x y) z = 
       
   466           int_or (int_and x z) (int_and y z)"
       
   467    apply (induct x rule: bin_induct)
       
   468     apply auto
       
   469   apply (case_tac y rule: bin_exhaust)
       
   470   apply (case_tac z rule: bin_exhaust)
       
   471   apply (case_tac ba, auto)
       
   472   done
       
   473 
       
   474 declare bin_ops_comm [simp] bbw_assocs [simp] 
       
   475 
       
   476 lemma plus_and_or [rule_format]:
       
   477   "ALL y. int_and x y + int_or x y = x + y"
       
   478   apply (induct x rule: bin_induct)
       
   479     apply clarsimp
       
   480    apply clarsimp
       
   481   apply clarsimp
       
   482   apply (case_tac y rule: bin_exhaust)
       
   483   apply clarsimp
       
   484   apply (unfold Bit_def)
       
   485   apply clarsimp
       
   486   apply (erule_tac x = "x" in allE)
       
   487   apply (simp split: bit.split)
       
   488   done
       
   489 
       
   490 lemma le_int_or:
       
   491   "!!x.  bin_sign y = Numeral.Pls ==> x <= int_or x y"
       
   492   apply (induct y rule: bin_induct)
       
   493     apply clarsimp
       
   494    apply clarsimp
       
   495   apply (case_tac x rule: bin_exhaust)
       
   496   apply (case_tac b)
       
   497    apply (case_tac [!] bit)
       
   498      apply (auto simp: less_eq_numeral_code)
       
   499   done
       
   500 
       
   501 lemmas int_and_le =
       
   502   xtr3 [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] ;
       
   503 
       
   504 (** nth bit, set/clear **)
       
   505 
       
   506 lemma bin_nth_sc [simp]: 
       
   507   "!!w. bin_nth (bin_sc n b w) n = (b = bit.B1)"
       
   508   by (induct n)  auto
       
   509 
       
   510 lemma bin_sc_sc_same [simp]: 
       
   511   "!!w. bin_sc n c (bin_sc n b w) = bin_sc n c w"
       
   512   by (induct n) auto
       
   513 
       
   514 lemma bin_sc_sc_diff:
       
   515   "!!w m. m ~= n ==> 
       
   516     bin_sc m c (bin_sc n b w) = bin_sc n b (bin_sc m c w)"
       
   517   apply (induct n)
       
   518    apply safe
       
   519    apply (case_tac [!] m)
       
   520      apply auto
       
   521   done
       
   522 
       
   523 lemma bin_nth_sc_gen: 
       
   524   "!!w m. bin_nth (bin_sc n b w) m = (if m = n then b = bit.B1 else bin_nth w m)"
       
   525   by (induct n) (case_tac [!] m, auto)
       
   526   
       
   527 lemma bin_sc_nth [simp]:
       
   528   "!!w. (bin_sc n (If (bin_nth w n) bit.B1 bit.B0) w) = w"
       
   529   by (induct n) auto
       
   530 
       
   531 lemma bin_sign_sc [simp]:
       
   532   "!!w. bin_sign (bin_sc n b w) = bin_sign w"
       
   533   by (induct n) auto
       
   534   
       
   535 lemma bin_sc_bintr [simp]: 
       
   536   "!!w m. bintrunc m (bin_sc n x (bintrunc m (w))) = bintrunc m (bin_sc n x w)"
       
   537   apply (induct n)
       
   538    apply (case_tac [!] w rule: bin_exhaust)
       
   539    apply (case_tac [!] m, auto)
       
   540   done
       
   541 
       
   542 lemma bin_clr_le:
       
   543   "!!w. bin_sc n bit.B0 w <= w"
       
   544   apply (induct n) 
       
   545    apply (case_tac [!] w rule: bin_exhaust)
       
   546    apply auto
       
   547    apply (unfold Bit_def)
       
   548    apply (simp_all split: bit.split)
       
   549   done
       
   550 
       
   551 lemma bin_set_ge:
       
   552   "!!w. bin_sc n bit.B1 w >= w"
       
   553   apply (induct n) 
       
   554    apply (case_tac [!] w rule: bin_exhaust)
       
   555    apply auto
       
   556    apply (unfold Bit_def)
       
   557    apply (simp_all split: bit.split)
       
   558   done
       
   559 
       
   560 lemma bintr_bin_clr_le:
       
   561   "!!w m. bintrunc n (bin_sc m bit.B0 w) <= bintrunc n w"
       
   562   apply (induct n)
       
   563    apply simp
       
   564   apply (case_tac w rule: bin_exhaust)
       
   565   apply (case_tac m)
       
   566    apply auto
       
   567    apply (unfold Bit_def)
       
   568    apply (simp_all split: bit.split)
       
   569   done
       
   570 
       
   571 lemma bintr_bin_set_ge:
       
   572   "!!w m. bintrunc n (bin_sc m bit.B1 w) >= bintrunc n w"
       
   573   apply (induct n)
       
   574    apply simp
       
   575   apply (case_tac w rule: bin_exhaust)
       
   576   apply (case_tac m)
       
   577    apply auto
       
   578    apply (unfold Bit_def)
       
   579    apply (simp_all split: bit.split)
       
   580   done
       
   581 
       
   582 lemma bin_nth_ops:
       
   583   "!!x y. bin_nth (int_and x y) n = (bin_nth x n & bin_nth y n)" 
       
   584   "!!x y. bin_nth (int_or x y) n = (bin_nth x n | bin_nth y n)"
       
   585   "!!x y. bin_nth (int_xor x y) n = (bin_nth x n ~= bin_nth y n)" 
       
   586   "!!x. bin_nth (int_not x) n = (~ bin_nth x n)"
       
   587   apply (induct n)
       
   588          apply safe
       
   589                          apply (case_tac [!] x rule: bin_exhaust)
       
   590                          apply simp_all
       
   591                       apply (case_tac [!] y rule: bin_exhaust)
       
   592                       apply simp_all
       
   593         apply (auto dest: not_B1_is_B0 intro: B1_ass_B0)
       
   594   done
       
   595 
       
   596 lemma bin_sc_FP [simp]: "bin_sc n bit.B0 Numeral.Pls = Numeral.Pls"
       
   597   by (induct n) auto
       
   598 
       
   599 lemma bin_sc_TM [simp]: "bin_sc n bit.B1 Numeral.Min = Numeral.Min"
       
   600   by (induct n) auto
       
   601   
       
   602 lemmas bin_sc_simps = bin_sc.Z bin_sc.Suc bin_sc_TM bin_sc_FP
       
   603 
       
   604 lemma bin_sc_minus:
       
   605   "0 < n ==> bin_sc (Suc (n - 1)) b w = bin_sc n b w"
       
   606   by auto
       
   607 
       
   608 lemmas bin_sc_Suc_minus = 
       
   609   trans [OF bin_sc_minus [symmetric] bin_sc.Suc, standard]
       
   610 
       
   611 lemmas bin_sc_Suc_pred [simp] = 
       
   612   bin_sc_Suc_minus [of "number_of bin", simplified nobm1, standard]
       
   613 
       
   614 (* interaction between bit-wise and arithmetic *)
       
   615 (* good example of bin_induction *)
       
   616 lemma bin_add_not: "x + int_not x = Numeral.Min"
       
   617   apply (induct x rule: bin_induct)
       
   618     apply clarsimp
       
   619    apply clarsimp
       
   620   apply (case_tac bit, auto)
       
   621   done
       
   622 
       
   623 (* truncating results of bit-wise operations *)
       
   624 lemma bin_trunc_ao: 
       
   625   "!!x y. int_and (bintrunc n x) (bintrunc n y) = bintrunc n (int_and x y)" 
       
   626   "!!x y. int_or (bintrunc n x) (bintrunc n y) = bintrunc n (int_or x y)"
       
   627   apply (induct n)
       
   628       apply auto
       
   629       apply (case_tac [!] x rule: bin_exhaust)
       
   630       apply (case_tac [!] y rule: bin_exhaust)
       
   631       apply auto
       
   632   done
       
   633 
       
   634 lemma bin_trunc_xor: 
       
   635   "!!x y. bintrunc n (int_xor (bintrunc n x) (bintrunc n y)) = 
       
   636           bintrunc n (int_xor x y)"
       
   637   apply (induct n)
       
   638    apply auto
       
   639    apply (case_tac [!] x rule: bin_exhaust)
       
   640    apply (case_tac [!] y rule: bin_exhaust)
       
   641    apply auto
       
   642   done
       
   643 
       
   644 lemma bin_trunc_not: 
       
   645   "!!x. bintrunc n (int_not (bintrunc n x)) = bintrunc n (int_not x)"
       
   646   apply (induct n)
       
   647    apply auto
       
   648    apply (case_tac [!] x rule: bin_exhaust)
       
   649    apply auto
       
   650   done
       
   651 
       
   652 (* want theorems of the form of bin_trunc_xor *)
       
   653 lemma bintr_bintr_i:
       
   654   "x = bintrunc n y ==> bintrunc n x = bintrunc n y"
       
   655   by auto
       
   656 
       
   657 lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i]
       
   658 lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i]
       
   659 
       
   660 lemma nth_2p_bin: 
       
   661   "!!m. bin_nth (2 ^ n) m = (m = n)"
       
   662   apply (induct n)
       
   663    apply clarsimp
       
   664    apply safe
       
   665      apply (case_tac m) 
       
   666       apply (auto simp: trans [OF numeral_1_eq_1 [symmetric] number_of_eq])
       
   667    apply (case_tac m) 
       
   668     apply (auto simp: Bit_B0_2t [symmetric])
       
   669   done
       
   670 
       
   671 (* for use when simplifying with bin_nth_Bit *)
       
   672 
       
   673 lemma ex_eq_or:
       
   674   "(EX m. n = Suc m & (m = k | P m)) = (n = Suc k | (EX m. n = Suc m & P m))"
       
   675   by auto
       
   676 
       
   677 end
       
   678