src/HOL/Word/BinOperations.thy
changeset 24364 31e359126ab6
parent 24353 9a7a9b19e925
child 24366 08b116049547
--- a/src/HOL/Word/BinOperations.thy	Mon Aug 20 20:44:03 2007 +0200
+++ b/src/HOL/Word/BinOperations.thy	Mon Aug 20 21:31:10 2007 +0200
@@ -13,7 +13,7 @@
 
 begin
 
-subsection {* NOT, AND, OR, XOR *}
+subsection {* Logical operations *}
 
 text "bit-wise logical operations on the int type"
 
@@ -28,121 +28,6 @@
     (\<lambda>w b s y. s (bin_rest y) BIT (b XOR bin_last y))"
   ..
 
-consts
-(*
-  int_and :: "int => int => int"
-  int_or :: "int => int => int"
-  bit_not :: "bit => bit"
-  bit_and :: "bit => bit => bit"
-  bit_or :: "bit => bit => bit"
-  bit_xor :: "bit => bit => bit"
-  int_not :: "int => int"
-  int_xor :: "int => int => int"
-*)
-  bin_sc :: "nat => bit => int => int"
-
-(*
-primrec
-  B0 : "bit_not bit.B0 = bit.B1"
-  B1 : "bit_not bit.B1 = bit.B0"
-
-primrec
-  B1 : "bit_xor bit.B1 x = bit_not x"
-  B0 : "bit_xor bit.B0 x = x"
-
-primrec
-  B1 : "bit_or bit.B1 x = bit.B1"
-  B0 : "bit_or bit.B0 x = x"
-
-primrec
-  B0 : "bit_and bit.B0 x = bit.B0"
-  B1 : "bit_and bit.B1 x = x"
-*)
-
-primrec
-  Z : "bin_sc 0 b w = bin_rest w BIT b"
-  Suc :
-    "bin_sc (Suc n) b w = bin_sc n b (bin_rest w) BIT bin_last w"
-
-(*
-defs (overloaded)
-  int_not_def : "int_not == bin_rec Numeral.Min Numeral.Pls 
-    (%w b s. s BIT bit_not b)"
-    int_and_def : "int_and == bin_rec (%x. Numeral.Pls) (%y. y) 
-    (%w b s y. s (bin_rest y) BIT (bit_and b (bin_last y)))"
-  int_or_def : "int_or == bin_rec (%x. x) (%y. Numeral.Min) 
-    (%w b s y. s (bin_rest y) BIT (bit_or b (bin_last y)))"
-  int_xor_def : "int_xor == bin_rec (%x. x) int_not 
-    (%w b s y. s (bin_rest y) BIT (bit_xor b (bin_last y)))"
-*)
-
-consts
-  bin_to_bl :: "nat => int => bool list"
-  bin_to_bl_aux :: "nat => int => bool list => bool list"
-  bl_to_bin :: "bool list => int"
-  bl_to_bin_aux :: "int => bool list => int"
-
-  bl_of_nth :: "nat => (nat => bool) => bool list"
-
-primrec
-  Nil : "bl_to_bin_aux w [] = w"
-  Cons : "bl_to_bin_aux w (b # bs) = 
-      bl_to_bin_aux (w BIT (if b then bit.B1 else bit.B0)) bs"
-
-primrec
-  Z : "bin_to_bl_aux 0 w bl = bl"
-  Suc : "bin_to_bl_aux (Suc n) w bl =
-    bin_to_bl_aux n (bin_rest w) ((bin_last w = bit.B1) # bl)"
-
-defs
-  bin_to_bl_def : "bin_to_bl n w == bin_to_bl_aux n w []"
-  bl_to_bin_def : "bl_to_bin bs == bl_to_bin_aux Numeral.Pls bs"
-
-primrec
-  Suc : "bl_of_nth (Suc n) f = f n # bl_of_nth n f"
-  Z : "bl_of_nth 0 f = []"
-
-consts
-  takefill :: "'a => nat => 'a list => 'a list"
-  app2 :: "('a => 'b => 'c) => 'a list => 'b list => 'c list"
-
--- "takefill - like take but if argument list too short,"
--- "extends result to get requested length"
-primrec
-  Z : "takefill fill 0 xs = []"
-  Suc : "takefill fill (Suc n) xs = (
-    case xs of [] => fill # takefill fill n xs
-      | y # ys => y # takefill fill n ys)"
-
-defs
-  app2_def : "app2 f as bs == map (split f) (zip as bs)"
-
--- "rcat and rsplit"
-consts
-  bin_rcat :: "nat => int list => int"
-  bin_rsplit_aux :: "nat * int list * nat * int => int list"
-  bin_rsplit :: "nat => (nat * int) => int list"
-  bin_rsplitl_aux :: "nat * int list * nat * int => int list"
-  bin_rsplitl :: "nat => (nat * int) => int list"
-  
-recdef bin_rsplit_aux "measure (fst o snd o snd)"
-  "bin_rsplit_aux (n, bs, (m, c)) =
-    (if m = 0 | n = 0 then bs else
-      let (a, b) = bin_split n c 
-      in bin_rsplit_aux (n, b # bs, (m - n, a)))"
-
-recdef bin_rsplitl_aux "measure (fst o snd o snd)"
-  "bin_rsplitl_aux (n, bs, (m, c)) =
-    (if m = 0 | n = 0 then bs else
-      let (a, b) = bin_split (min m n) c 
-      in bin_rsplitl_aux (n, b # bs, (m - n, a)))"
-
-defs
-  bin_rcat_def : "bin_rcat n bs == foldl (%u v. bin_cat u n v) Numeral.Pls bs"
-  bin_rsplit_def : "bin_rsplit n w == bin_rsplit_aux (n, [], w)"
-  bin_rsplitl_def : "bin_rsplitl n w == bin_rsplitl_aux (n, [], w)"
-     
- 
 lemma int_not_simps [simp]:
   "NOT Numeral.Pls = Numeral.Min"
   "NOT Numeral.Min = Numeral.Pls"
@@ -274,124 +159,6 @@
   int_and_extra_simps  int_or_extra_simps  int_xor_extra_simps
   int_and_Pls int_and_Min  int_or_Pls int_or_Min  int_xor_Pls int_xor_Min
 
-(* potential for looping *)
-declare bin_rsplit_aux.simps [simp del]
-declare bin_rsplitl_aux.simps [simp del]
-
-
-lemma bin_sign_cat: 
-  "!!y. bin_sign (bin_cat x n y) = bin_sign x"
-  by (induct n) auto
-
-lemma bin_cat_Suc_Bit:
-  "bin_cat w (Suc n) (v BIT b) = bin_cat w n v BIT b"
-  by auto
-
-lemma bin_nth_cat: 
-  "!!n y. bin_nth (bin_cat x k y) n = 
-    (if n < k then bin_nth y n else bin_nth x (n - k))"
-  apply (induct k)
-   apply clarsimp
-  apply (case_tac n, auto)
-  done
-
-lemma bin_nth_split:
-  "!!b c. bin_split n c = (a, b) ==> 
-    (ALL k. bin_nth a k = bin_nth c (n + k)) & 
-    (ALL k. bin_nth b k = (k < n & bin_nth c k))"
-  apply (induct n)
-   apply clarsimp
-  apply (clarsimp simp: Let_def split: ls_splits)
-  apply (case_tac k)
-  apply auto
-  done
-
-lemma bin_cat_assoc: 
-  "!!z. bin_cat (bin_cat x m y) n z = bin_cat x (m + n) (bin_cat y n z)" 
-  by (induct n) auto
-
-lemma bin_cat_assoc_sym: "!!z m. 
-  bin_cat x m (bin_cat y n z) = bin_cat (bin_cat x (m - n) y) (min m n) z"
-  apply (induct n, clarsimp)
-  apply (case_tac m, auto)
-  done
-
-lemma bin_cat_Pls [simp]: 
-  "!!w. bin_cat Numeral.Pls n w = bintrunc n w"
-  by (induct n) auto
-
-lemma bintr_cat1: 
-  "!!b. bintrunc (k + n) (bin_cat a n b) = bin_cat (bintrunc k a) n b"
-  by (induct n) auto
-    
-lemma bintr_cat: "bintrunc m (bin_cat a n b) = 
-    bin_cat (bintrunc (m - n) a) n (bintrunc (min m n) b)"
-  by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr)
-    
-lemma bintr_cat_same [simp]: 
-  "bintrunc n (bin_cat a n b) = bintrunc n b"
-  by (auto simp add : bintr_cat)
-
-lemma cat_bintr [simp]: 
-  "!!b. bin_cat a n (bintrunc n b) = bin_cat a n b"
-  by (induct n) auto
-
-lemma split_bintrunc: 
-  "!!b c. bin_split n c = (a, b) ==> b = bintrunc n c"
-  by (induct n) (auto simp: Let_def split: ls_splits)
-
-lemma bin_cat_split:
-  "!!v w. bin_split n w = (u, v) ==> w = bin_cat u n v"
-  by (induct n) (auto simp: Let_def split: ls_splits)
-
-lemma bin_split_cat:
-  "!!w. bin_split n (bin_cat v n w) = (v, bintrunc n w)"
-  by (induct n) auto
-
-lemma bin_split_Pls [simp]:
-  "bin_split n Numeral.Pls = (Numeral.Pls, Numeral.Pls)"
-  by (induct n) (auto simp: Let_def split: ls_splits)
-
-lemma bin_split_Min [simp]:
-  "bin_split n Numeral.Min = (Numeral.Min, bintrunc n Numeral.Min)"
-  by (induct n) (auto simp: Let_def split: ls_splits)
-
-lemma bin_split_trunc:
-  "!!m b c. bin_split (min m n) c = (a, b) ==> 
-    bin_split n (bintrunc m c) = (bintrunc (m - n) a, b)"
-  apply (induct n, clarsimp)
-  apply (simp add: bin_rest_trunc Let_def split: ls_splits)
-  apply (case_tac m)
-   apply (auto simp: Let_def split: ls_splits)
-  done
-
-lemma bin_split_trunc1:
-  "!!m b c. bin_split n c = (a, b) ==> 
-    bin_split n (bintrunc m c) = (bintrunc (m - n) a, bintrunc m b)"
-  apply (induct n, clarsimp)
-  apply (simp add: bin_rest_trunc Let_def split: ls_splits)
-  apply (case_tac m)
-   apply (auto simp: Let_def split: ls_splits)
-  done
-
-lemma bin_cat_num:
-  "!!b. bin_cat a n b = a * 2 ^ n + bintrunc n b"
-  apply (induct n, clarsimp)
-  apply (simp add: Bit_def cong: number_of_False_cong)
-  done
-
-lemma bin_split_num:
-  "!!b. bin_split n b = (b div 2 ^ n, b mod 2 ^ n)"
-  apply (induct n, clarsimp)
-  apply (simp add: bin_rest_div zdiv_zmult2_eq)
-  apply (case_tac b rule: bin_exhaust)
-  apply simp
-  apply (simp add: Bit_def zmod_zmult_zmult1 p1mod22k
-              split: bit.split 
-              cong: number_of_False_cong)
-  done 
-
-
 (* basic properties of logical (bit-wise) operations *)
 
 lemma bbw_ao_absorb: 
@@ -523,6 +290,76 @@
 lemmas int_and_le =
   xtr3 [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] ;
 
+lemma bin_nth_ops:
+  "!!x y. bin_nth (x AND y) n = (bin_nth x n & bin_nth y n)" 
+  "!!x y. bin_nth (x OR y) n = (bin_nth x n | bin_nth y n)"
+  "!!x y. bin_nth (x XOR y) n = (bin_nth x n ~= bin_nth y n)" 
+  "!!x. bin_nth (NOT x) n = (~ bin_nth x n)"
+  apply (induct n)
+         apply safe
+                         apply (case_tac [!] x rule: bin_exhaust)
+                         apply simp_all
+                      apply (case_tac [!] y rule: bin_exhaust)
+                      apply simp_all
+        apply (auto dest: not_B1_is_B0 intro: B1_ass_B0)
+  done
+
+(* interaction between bit-wise and arithmetic *)
+(* good example of bin_induction *)
+lemma bin_add_not: "x + NOT x = Numeral.Min"
+  apply (induct x rule: bin_induct)
+    apply clarsimp
+   apply clarsimp
+  apply (case_tac bit, auto)
+  done
+
+(* truncating results of bit-wise operations *)
+lemma bin_trunc_ao: 
+  "!!x y. (bintrunc n x) AND (bintrunc n y) = bintrunc n (x AND y)" 
+  "!!x y. (bintrunc n x) OR (bintrunc n y) = bintrunc n (x OR y)"
+  apply (induct n)
+      apply auto
+      apply (case_tac [!] x rule: bin_exhaust)
+      apply (case_tac [!] y rule: bin_exhaust)
+      apply auto
+  done
+
+lemma bin_trunc_xor: 
+  "!!x y. bintrunc n (bintrunc n x XOR bintrunc n y) = 
+          bintrunc n (x XOR y)"
+  apply (induct n)
+   apply auto
+   apply (case_tac [!] x rule: bin_exhaust)
+   apply (case_tac [!] y rule: bin_exhaust)
+   apply auto
+  done
+
+lemma bin_trunc_not: 
+  "!!x. bintrunc n (NOT (bintrunc n x)) = bintrunc n (NOT x)"
+  apply (induct n)
+   apply auto
+   apply (case_tac [!] x rule: bin_exhaust)
+   apply auto
+  done
+
+(* want theorems of the form of bin_trunc_xor *)
+lemma bintr_bintr_i:
+  "x = bintrunc n y ==> bintrunc n x = bintrunc n y"
+  by auto
+
+lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i]
+lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i]
+
+subsection {* Setting and clearing bits *}
+
+consts
+  bin_sc :: "nat => bit => int => int"
+
+primrec
+  Z : "bin_sc 0 b w = bin_rest w BIT b"
+  Suc :
+    "bin_sc (Suc n) b w = bin_sc n b (bin_rest w) BIT bin_last w"
+
 (** nth bit, set/clear **)
 
 lemma bin_nth_sc [simp]: 
@@ -601,20 +438,6 @@
    apply (simp_all split: bit.split)
   done
 
-lemma bin_nth_ops:
-  "!!x y. bin_nth (x AND y) n = (bin_nth x n & bin_nth y n)" 
-  "!!x y. bin_nth (x OR y) n = (bin_nth x n | bin_nth y n)"
-  "!!x y. bin_nth (x XOR y) n = (bin_nth x n ~= bin_nth y n)" 
-  "!!x. bin_nth (NOT x) n = (~ bin_nth x n)"
-  apply (induct n)
-         apply safe
-                         apply (case_tac [!] x rule: bin_exhaust)
-                         apply simp_all
-                      apply (case_tac [!] y rule: bin_exhaust)
-                      apply simp_all
-        apply (auto dest: not_B1_is_B0 intro: B1_ass_B0)
-  done
-
 lemma bin_sc_FP [simp]: "bin_sc n bit.B0 Numeral.Pls = Numeral.Pls"
   by (induct n) auto
 
@@ -633,51 +456,194 @@
 lemmas bin_sc_Suc_pred [simp] = 
   bin_sc_Suc_minus [of "number_of bin", simplified nobm1, standard]
 
-(* interaction between bit-wise and arithmetic *)
-(* good example of bin_induction *)
-lemma bin_add_not: "x + NOT x = Numeral.Min"
-  apply (induct x rule: bin_induct)
-    apply clarsimp
-   apply clarsimp
-  apply (case_tac bit, auto)
-  done
+subsection {* Operations on lists of booleans *}
+
+consts
+  bin_to_bl :: "nat => int => bool list"
+  bin_to_bl_aux :: "nat => int => bool list => bool list"
+  bl_to_bin :: "bool list => int"
+  bl_to_bin_aux :: "int => bool list => int"
+
+  bl_of_nth :: "nat => (nat => bool) => bool list"
+
+primrec
+  Nil : "bl_to_bin_aux w [] = w"
+  Cons : "bl_to_bin_aux w (b # bs) = 
+      bl_to_bin_aux (w BIT (if b then bit.B1 else bit.B0)) bs"
+
+primrec
+  Z : "bin_to_bl_aux 0 w bl = bl"
+  Suc : "bin_to_bl_aux (Suc n) w bl =
+    bin_to_bl_aux n (bin_rest w) ((bin_last w = bit.B1) # bl)"
+
+defs
+  bin_to_bl_def : "bin_to_bl n w == bin_to_bl_aux n w []"
+  bl_to_bin_def : "bl_to_bin bs == bl_to_bin_aux Numeral.Pls bs"
+
+primrec
+  Suc : "bl_of_nth (Suc n) f = f n # bl_of_nth n f"
+  Z : "bl_of_nth 0 f = []"
+
+consts
+  takefill :: "'a => nat => 'a list => 'a list"
+  app2 :: "('a => 'b => 'c) => 'a list => 'b list => 'c list"
+
+-- "takefill - like take but if argument list too short,"
+-- "extends result to get requested length"
+primrec
+  Z : "takefill fill 0 xs = []"
+  Suc : "takefill fill (Suc n) xs = (
+    case xs of [] => fill # takefill fill n xs
+      | y # ys => y # takefill fill n ys)"
+
+defs
+  app2_def : "app2 f as bs == map (split f) (zip as bs)"
+
+subsection {* Splitting and concatenation *}
 
-(* truncating results of bit-wise operations *)
-lemma bin_trunc_ao: 
-  "!!x y. (bintrunc n x) AND (bintrunc n y) = bintrunc n (x AND y)" 
-  "!!x y. (bintrunc n x) OR (bintrunc n y) = bintrunc n (x OR y)"
-  apply (induct n)
-      apply auto
-      apply (case_tac [!] x rule: bin_exhaust)
-      apply (case_tac [!] y rule: bin_exhaust)
-      apply auto
+-- "rcat and rsplit"
+consts
+  bin_rcat :: "nat => int list => int"
+  bin_rsplit_aux :: "nat * int list * nat * int => int list"
+  bin_rsplit :: "nat => (nat * int) => int list"
+  bin_rsplitl_aux :: "nat * int list * nat * int => int list"
+  bin_rsplitl :: "nat => (nat * int) => int list"
+  
+recdef bin_rsplit_aux "measure (fst o snd o snd)"
+  "bin_rsplit_aux (n, bs, (m, c)) =
+    (if m = 0 | n = 0 then bs else
+      let (a, b) = bin_split n c 
+      in bin_rsplit_aux (n, b # bs, (m - n, a)))"
+
+recdef bin_rsplitl_aux "measure (fst o snd o snd)"
+  "bin_rsplitl_aux (n, bs, (m, c)) =
+    (if m = 0 | n = 0 then bs else
+      let (a, b) = bin_split (min m n) c 
+      in bin_rsplitl_aux (n, b # bs, (m - n, a)))"
+
+defs
+  bin_rcat_def : "bin_rcat n bs == foldl (%u v. bin_cat u n v) Numeral.Pls bs"
+  bin_rsplit_def : "bin_rsplit n w == bin_rsplit_aux (n, [], w)"
+  bin_rsplitl_def : "bin_rsplitl n w == bin_rsplitl_aux (n, [], w)"
+     
+ 
+(* potential for looping *)
+declare bin_rsplit_aux.simps [simp del]
+declare bin_rsplitl_aux.simps [simp del]
+
+lemma bin_sign_cat: 
+  "!!y. bin_sign (bin_cat x n y) = bin_sign x"
+  by (induct n) auto
+
+lemma bin_cat_Suc_Bit:
+  "bin_cat w (Suc n) (v BIT b) = bin_cat w n v BIT b"
+  by auto
+
+lemma bin_nth_cat: 
+  "!!n y. bin_nth (bin_cat x k y) n = 
+    (if n < k then bin_nth y n else bin_nth x (n - k))"
+  apply (induct k)
+   apply clarsimp
+  apply (case_tac n, auto)
   done
 
-lemma bin_trunc_xor: 
-  "!!x y. bintrunc n (bintrunc n x XOR bintrunc n y) = 
-          bintrunc n (x XOR y)"
+lemma bin_nth_split:
+  "!!b c. bin_split n c = (a, b) ==> 
+    (ALL k. bin_nth a k = bin_nth c (n + k)) & 
+    (ALL k. bin_nth b k = (k < n & bin_nth c k))"
   apply (induct n)
-   apply auto
-   apply (case_tac [!] x rule: bin_exhaust)
-   apply (case_tac [!] y rule: bin_exhaust)
-   apply auto
+   apply clarsimp
+  apply (clarsimp simp: Let_def split: ls_splits)
+  apply (case_tac k)
+  apply auto
+  done
+
+lemma bin_cat_assoc: 
+  "!!z. bin_cat (bin_cat x m y) n z = bin_cat x (m + n) (bin_cat y n z)" 
+  by (induct n) auto
+
+lemma bin_cat_assoc_sym: "!!z m. 
+  bin_cat x m (bin_cat y n z) = bin_cat (bin_cat x (m - n) y) (min m n) z"
+  apply (induct n, clarsimp)
+  apply (case_tac m, auto)
   done
 
-lemma bin_trunc_not: 
-  "!!x. bintrunc n (NOT (bintrunc n x)) = bintrunc n (NOT x)"
-  apply (induct n)
-   apply auto
-   apply (case_tac [!] x rule: bin_exhaust)
-   apply auto
+lemma bin_cat_Pls [simp]: 
+  "!!w. bin_cat Numeral.Pls n w = bintrunc n w"
+  by (induct n) auto
+
+lemma bintr_cat1: 
+  "!!b. bintrunc (k + n) (bin_cat a n b) = bin_cat (bintrunc k a) n b"
+  by (induct n) auto
+    
+lemma bintr_cat: "bintrunc m (bin_cat a n b) = 
+    bin_cat (bintrunc (m - n) a) n (bintrunc (min m n) b)"
+  by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr)
+    
+lemma bintr_cat_same [simp]: 
+  "bintrunc n (bin_cat a n b) = bintrunc n b"
+  by (auto simp add : bintr_cat)
+
+lemma cat_bintr [simp]: 
+  "!!b. bin_cat a n (bintrunc n b) = bin_cat a n b"
+  by (induct n) auto
+
+lemma split_bintrunc: 
+  "!!b c. bin_split n c = (a, b) ==> b = bintrunc n c"
+  by (induct n) (auto simp: Let_def split: ls_splits)
+
+lemma bin_cat_split:
+  "!!v w. bin_split n w = (u, v) ==> w = bin_cat u n v"
+  by (induct n) (auto simp: Let_def split: ls_splits)
+
+lemma bin_split_cat:
+  "!!w. bin_split n (bin_cat v n w) = (v, bintrunc n w)"
+  by (induct n) auto
+
+lemma bin_split_Pls [simp]:
+  "bin_split n Numeral.Pls = (Numeral.Pls, Numeral.Pls)"
+  by (induct n) (auto simp: Let_def split: ls_splits)
+
+lemma bin_split_Min [simp]:
+  "bin_split n Numeral.Min = (Numeral.Min, bintrunc n Numeral.Min)"
+  by (induct n) (auto simp: Let_def split: ls_splits)
+
+lemma bin_split_trunc:
+  "!!m b c. bin_split (min m n) c = (a, b) ==> 
+    bin_split n (bintrunc m c) = (bintrunc (m - n) a, b)"
+  apply (induct n, clarsimp)
+  apply (simp add: bin_rest_trunc Let_def split: ls_splits)
+  apply (case_tac m)
+   apply (auto simp: Let_def split: ls_splits)
   done
 
-(* want theorems of the form of bin_trunc_xor *)
-lemma bintr_bintr_i:
-  "x = bintrunc n y ==> bintrunc n x = bintrunc n y"
-  by auto
+lemma bin_split_trunc1:
+  "!!m b c. bin_split n c = (a, b) ==> 
+    bin_split n (bintrunc m c) = (bintrunc (m - n) a, bintrunc m b)"
+  apply (induct n, clarsimp)
+  apply (simp add: bin_rest_trunc Let_def split: ls_splits)
+  apply (case_tac m)
+   apply (auto simp: Let_def split: ls_splits)
+  done
 
-lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i]
-lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i]
+lemma bin_cat_num:
+  "!!b. bin_cat a n b = a * 2 ^ n + bintrunc n b"
+  apply (induct n, clarsimp)
+  apply (simp add: Bit_def cong: number_of_False_cong)
+  done
+
+lemma bin_split_num:
+  "!!b. bin_split n b = (b div 2 ^ n, b mod 2 ^ n)"
+  apply (induct n, clarsimp)
+  apply (simp add: bin_rest_div zdiv_zmult2_eq)
+  apply (case_tac b rule: bin_exhaust)
+  apply simp
+  apply (simp add: Bit_def zmod_zmult_zmult1 p1mod22k
+              split: bit.split 
+              cong: number_of_False_cong)
+  done 
+
+subsection {* Miscellaneous lemmas *}
 
 lemma nth_2p_bin: 
   "!!m. bin_nth (2 ^ n) m = (m = n)"