merge
authorpanny
Mon, 23 Dec 2013 16:59:56 +0100
changeset 54852 7137303f9f88
parent 54851 48a24d371ebb (current diff)
parent 54850 980817309b78 (diff)
child 54853 a435932a9f12
merge
--- a/NEWS	Mon Dec 23 15:30:31 2013 +0100
+++ b/NEWS	Mon Dec 23 16:59:56 2013 +0100
@@ -28,8 +28,11 @@
 
 *** HOL ***
 
+* Word library: bit representations prefer type bool over type bit.
+INCOMPATIBILITY.
+
 * Theorem disambiguation Inf_le_Sup (on finite sets) ~> Inf_fin_le_Sup_fin.
-INCOMPATBILITY.
+INCOMPATIBILITY.
 
 * Code generations are provided for make, fields, extend and truncate
 operations on records.
--- a/src/HOL/Word/Bit_Comparison.thy	Mon Dec 23 15:30:31 2013 +0100
+++ b/src/HOL/Word/Bit_Comparison.thy	Mon Dec 23 16:59:56 2013 +0100
@@ -20,10 +20,10 @@
   proof (cases y rule: bin_exhaust)
     case (1 bin' bit')
     from 3 have "0 \<le> bin"
-      by (simp add: Bit_def bitval_def split add: bit.split_asm)
+      by (cases bit) (simp_all add: Bit_def)
     then have "0 \<le> bin AND bin'" by (rule 3)
     with 1 show ?thesis
-      by simp (simp add: Bit_def bitval_def split add: bit.split)
+      by simp (simp add: Bit_def)
   qed
 next
   case 2
@@ -41,12 +41,12 @@
   proof (cases y rule: bin_exhaust)
     case (1 bin' bit')
     from 3 have "0 \<le> bin"
-      by (simp add: Bit_def bitval_def split add: bit.split_asm)
+      by (cases bit) (simp_all add: Bit_def)
     moreover from 1 3 have "0 \<le> bin'"
-      by (simp add: Bit_def bitval_def split add: bit.split_asm)
+      by (cases bit') (simp_all add: Bit_def)
     ultimately have "0 \<le> bin OR bin'" by (rule 3)
     with 1 show ?thesis
-      by simp (simp add: Bit_def bitval_def split add: bit.split)
+      by simp (simp add: Bit_def)
   qed
 qed simp_all
 
@@ -61,12 +61,12 @@
   proof (cases y rule: bin_exhaust)
     case (1 bin' bit')
     from 3 have "0 \<le> bin"
-      by (simp add: Bit_def bitval_def split add: bit.split_asm)
+      by (cases bit) (simp_all add: Bit_def)
     moreover from 1 3 have "0 \<le> bin'"
-      by (simp add: Bit_def bitval_def split add: bit.split_asm)
+      by (cases bit') (simp_all add: Bit_def)
     ultimately have "0 \<le> bin XOR bin'" by (rule 3)
     with 1 show ?thesis
-      by simp (simp add: Bit_def bitval_def split add: bit.split)
+      by simp (simp add: Bit_def)
   qed
 next
   case 2
@@ -84,10 +84,10 @@
   proof (cases y rule: bin_exhaust)
     case (1 bin' bit')
     from 3 have "0 \<le> bin"
-      by (simp add: Bit_def bitval_def split add: bit.split_asm)
+      by (cases bit) (simp_all add: Bit_def)
     then have "bin AND bin' \<le> bin" by (rule 3)
     with 1 show ?thesis
-      by simp (simp add: Bit_def bitval_def split add: bit.split)
+      by simp (simp add: Bit_def)
   qed
 next
   case 2
@@ -108,10 +108,10 @@
   proof (cases x rule: bin_exhaust)
     case (1 bin' bit')
     from 3 have "0 \<le> bin"
-      by (simp add: Bit_def bitval_def split add: bit.split_asm)
+      by (cases bit) (simp_all add: Bit_def)
     then have "bin' AND bin \<le> bin" by (rule 3)
     with 1 show ?thesis
-      by simp (simp add: Bit_def bitval_def split add: bit.split)
+      by simp (simp add: Bit_def)
   qed
 next
   case 2
@@ -135,21 +135,21 @@
     proof (cases n)
       case 0
       with 3 have "bin BIT bit = 0" by simp
-      then have "bin = 0" "bit = 0"
-        by (auto simp add: Bit_def bitval_def split add: bit.split_asm) arith
+      then have "bin = 0" and "\<not> bit"
+        by (auto simp add: Bit_def split: if_splits) arith
       then show ?thesis using 0 1 `y < 2 ^ n`
         by simp
     next
       case (Suc m)
       from 3 have "0 \<le> bin"
-        by (simp add: Bit_def bitval_def split add: bit.split_asm)
+        by (cases bit) (simp_all add: Bit_def)
       moreover from 3 Suc have "bin < 2 ^ m"
-        by (simp add: Bit_def bitval_def split add: bit.split_asm)
+        by (cases bit) (simp_all add: Bit_def)
       moreover from 1 3 Suc have "bin' < 2 ^ m"
-        by (simp add: Bit_def bitval_def split add: bit.split_asm)
+        by (cases bit') (simp_all add: Bit_def)
       ultimately have "bin OR bin' < 2 ^ m" by (rule 3)
       with 1 Suc show ?thesis
-        by simp (simp add: Bit_def bitval_def split add: bit.split)
+        by simp (simp add: Bit_def)
     qed
   qed
 qed simp_all
@@ -168,21 +168,21 @@
     proof (cases n)
       case 0
       with 3 have "bin BIT bit = 0" by simp
-      then have "bin = 0" "bit = 0"
-        by (auto simp add: Bit_def bitval_def split add: bit.split_asm) arith
+      then have "bin = 0" and "\<not> bit"
+        by (auto simp add: Bit_def split: if_splits) arith
       then show ?thesis using 0 1 `y < 2 ^ n`
         by simp
     next
       case (Suc m)
       from 3 have "0 \<le> bin"
-        by (simp add: Bit_def bitval_def split add: bit.split_asm)
+        by (cases bit) (simp_all add: Bit_def)
       moreover from 3 Suc have "bin < 2 ^ m"
-        by (simp add: Bit_def bitval_def split add: bit.split_asm)
+        by (cases bit) (simp_all add: Bit_def)
       moreover from 1 3 Suc have "bin' < 2 ^ m"
-        by (simp add: Bit_def bitval_def split add: bit.split_asm)
+        by (cases bit') (simp_all add: Bit_def)
       ultimately have "bin XOR bin' < 2 ^ m" by (rule 3)
       with 1 Suc show ?thesis
-        by simp (simp add: Bit_def bitval_def split add: bit.split)
+        by simp (simp add: Bit_def)
     qed
   qed
 next
--- a/src/HOL/Word/Bit_Int.thy	Mon Dec 23 15:30:31 2013 +0100
+++ b/src/HOL/Word/Bit_Int.thy	Mon Dec 23 16:59:56 2013 +0100
@@ -9,7 +9,7 @@
 header {* Bitwise Operations on Binary Integers *}
 
 theory Bit_Int
-imports Bit_Representation Bit_Bit
+imports Bit_Representation Bit_Operations
 begin
 
 subsection {* Logical operations *}
@@ -25,7 +25,7 @@
 function bitAND_int where
   "bitAND_int x y =
     (if x = 0 then 0 else if x = -1 then y else
-      (bin_rest x AND bin_rest y) BIT (bin_last x AND bin_last y))"
+      (bin_rest x AND bin_rest y) BIT (bin_last x \<and> bin_last y))"
   by pat_completeness simp
 
 termination
@@ -46,7 +46,7 @@
 subsubsection {* Basic simplification rules *}
 
 lemma int_not_BIT [simp]:
-  "NOT (w BIT b) = (NOT w) BIT (NOT b)"
+  "NOT (w BIT b) = (NOT w) BIT (\<not> b)"
   unfolding int_not_def Bit_def by (cases b, simp_all)
 
 lemma int_not_simps [simp]:
@@ -68,7 +68,7 @@
   by (simp add: bitAND_int.simps)
 
 lemma int_and_Bits [simp]: 
-  "(x BIT b) AND (y BIT c) = (x AND y) BIT (b AND c)" 
+  "(x BIT b) AND (y BIT c) = (x AND y) BIT (b \<and> c)" 
   by (subst bitAND_int.simps, simp add: Bit_eq_0_iff Bit_eq_m1_iff)
 
 lemma int_or_zero [simp]: "(0::int) OR x = x"
@@ -78,40 +78,40 @@
   unfolding int_or_def by simp
 
 lemma int_or_Bits [simp]: 
-  "(x BIT b) OR (y BIT c) = (x OR y) BIT (b OR c)"
-  unfolding int_or_def bit_or_def by simp
+  "(x BIT b) OR (y BIT c) = (x OR y) BIT (b \<or> c)"
+  unfolding int_or_def by simp
 
 lemma int_xor_zero [simp]: "(0::int) XOR x = x"
   unfolding int_xor_def by simp
 
 lemma int_xor_Bits [simp]: 
-  "(x BIT b) XOR (y BIT c) = (x XOR y) BIT (b XOR c)"
-  unfolding int_xor_def bit_xor_def by simp
+  "(x BIT b) XOR (y BIT c) = (x XOR y) BIT ((b \<or> c) \<and> \<not> (b \<and> c))"
+  unfolding int_xor_def by auto
 
 subsubsection {* Binary destructors *}
 
 lemma bin_rest_NOT [simp]: "bin_rest (NOT x) = NOT (bin_rest x)"
   by (cases x rule: bin_exhaust, simp)
 
-lemma bin_last_NOT [simp]: "bin_last (NOT x) = NOT (bin_last x)"
+lemma bin_last_NOT [simp]: "bin_last (NOT x) \<longleftrightarrow> \<not> bin_last x"
   by (cases x rule: bin_exhaust, simp)
 
 lemma bin_rest_AND [simp]: "bin_rest (x AND y) = bin_rest x AND bin_rest y"
   by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp)
 
-lemma bin_last_AND [simp]: "bin_last (x AND y) = bin_last x AND bin_last y"
+lemma bin_last_AND [simp]: "bin_last (x AND y) \<longleftrightarrow> bin_last x \<and> bin_last y"
   by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp)
 
 lemma bin_rest_OR [simp]: "bin_rest (x OR y) = bin_rest x OR bin_rest y"
   by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp)
 
-lemma bin_last_OR [simp]: "bin_last (x OR y) = bin_last x OR bin_last y"
+lemma bin_last_OR [simp]: "bin_last (x OR y) \<longleftrightarrow> bin_last x \<or> bin_last y"
   by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp)
 
 lemma bin_rest_XOR [simp]: "bin_rest (x XOR y) = bin_rest x XOR bin_rest y"
   by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp)
 
-lemma bin_last_XOR [simp]: "bin_last (x XOR y) = bin_last x XOR bin_last y"
+lemma bin_last_XOR [simp]: "bin_last (x XOR y) \<longleftrightarrow> (bin_last x \<or> bin_last y) \<and> \<not> (bin_last x \<and> bin_last y)"
   by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp)
 
 lemma bin_nth_ops:
@@ -225,36 +225,36 @@
   other simp rules. *}
 
 lemma bin_rl_eqI: "\<lbrakk>bin_rest x = bin_rest y; bin_last x = bin_last y\<rbrakk> \<Longrightarrow> x = y"
-  by (metis bin_rl_simp)
+  by (metis (mono_tags) BIT_eq_iff bin_ex_rl bin_last_BIT bin_rest_BIT)
 
 lemma bin_rest_neg_numeral_BitM [simp]:
   "bin_rest (- numeral (Num.BitM w)) = - numeral w"
   by (simp only: BIT_bin_simps [symmetric] bin_rest_BIT)
 
 lemma bin_last_neg_numeral_BitM [simp]:
-  "bin_last (-  numeral (Num.BitM w)) = 1"
+  "bin_last (- numeral (Num.BitM w))"
   by (simp only: BIT_bin_simps [symmetric] bin_last_BIT)
 
 text {* FIXME: The rule sets below are very large (24 rules for each
   operator). Is there a simpler way to do this? *}
 
 lemma int_and_numerals [simp]:
-  "numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (numeral x AND numeral y) BIT 0"
-  "numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (numeral x AND numeral y) BIT 0"
-  "numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (numeral x AND numeral y) BIT 0"
-  "numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = (numeral x AND numeral y) BIT 1"
-  "numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (numeral x AND - numeral y) BIT 0"
-  "numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (numeral x AND - numeral (y + Num.One)) BIT 0"
-  "numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (numeral x AND - numeral y) BIT 0"
-  "numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = (numeral x AND - numeral (y + Num.One)) BIT 1"
-  "- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (- numeral x AND numeral y) BIT 0"
-  "- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (- numeral x AND numeral y) BIT 0"
-  "- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (- numeral (x + Num.One) AND numeral y) BIT 0"
-  "- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = (- numeral (x + Num.One) AND numeral y) BIT 1"
-  "- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (- numeral x AND - numeral y) BIT 0"
-  "- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (- numeral x AND - numeral (y + Num.One)) BIT 0"
-  "- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (- numeral (x + Num.One) AND - numeral y) BIT 0"
-  "- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = (- numeral (x + Num.One) AND - numeral (y + Num.One)) BIT 1"
+  "numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (numeral x AND numeral y) BIT False"
+  "numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (numeral x AND numeral y) BIT False"
+  "numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (numeral x AND numeral y) BIT False"
+  "numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = (numeral x AND numeral y) BIT True"
+  "numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (numeral x AND - numeral y) BIT False"
+  "numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (numeral x AND - numeral (y + Num.One)) BIT False"
+  "numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (numeral x AND - numeral y) BIT False"
+  "numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = (numeral x AND - numeral (y + Num.One)) BIT True"
+  "- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (- numeral x AND numeral y) BIT False"
+  "- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (- numeral x AND numeral y) BIT False"
+  "- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (- numeral (x + Num.One) AND numeral y) BIT False"
+  "- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = (- numeral (x + Num.One) AND numeral y) BIT True"
+  "- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (- numeral x AND - numeral y) BIT False"
+  "- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (- numeral x AND - numeral (y + Num.One)) BIT False"
+  "- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (- numeral (x + Num.One) AND - numeral y) BIT False"
+  "- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = (- numeral (x + Num.One) AND - numeral (y + Num.One)) BIT True"
   "(1::int) AND numeral (Num.Bit0 y) = 0"
   "(1::int) AND numeral (Num.Bit1 y) = 1"
   "(1::int) AND - numeral (Num.Bit0 y) = 0"
@@ -266,22 +266,22 @@
   by (rule bin_rl_eqI, simp, simp)+
 
 lemma int_or_numerals [simp]:
-  "numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (numeral x OR numeral y) BIT 0"
-  "numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = (numeral x OR numeral y) BIT 1"
-  "numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = (numeral x OR numeral y) BIT 1"
-  "numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = (numeral x OR numeral y) BIT 1"
-  "numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (numeral x OR - numeral y) BIT 0"
-  "numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = (numeral x OR - numeral (y + Num.One)) BIT 1"
-  "numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = (numeral x OR - numeral y) BIT 1"
-  "numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = (numeral x OR - numeral (y + Num.One)) BIT 1"
-  "- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (- numeral x OR numeral y) BIT 0"
-  "- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = (- numeral x OR numeral y) BIT 1"
-  "- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = (- numeral (x + Num.One) OR numeral y) BIT 1"
-  "- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = (- numeral (x + Num.One) OR numeral y) BIT 1"
-  "- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (- numeral x OR - numeral y) BIT 0"
-  "- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = (- numeral x OR - numeral (y + Num.One)) BIT 1"
-  "- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = (- numeral (x + Num.One) OR - numeral y) BIT 1"
-  "- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = (- numeral (x + Num.One) OR - numeral (y + Num.One)) BIT 1"
+  "numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (numeral x OR numeral y) BIT False"
+  "numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = (numeral x OR numeral y) BIT True"
+  "numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = (numeral x OR numeral y) BIT True"
+  "numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = (numeral x OR numeral y) BIT True"
+  "numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (numeral x OR - numeral y) BIT False"
+  "numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = (numeral x OR - numeral (y + Num.One)) BIT True"
+  "numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = (numeral x OR - numeral y) BIT True"
+  "numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = (numeral x OR - numeral (y + Num.One)) BIT True"
+  "- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (- numeral x OR numeral y) BIT False"
+  "- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = (- numeral x OR numeral y) BIT True"
+  "- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = (- numeral (x + Num.One) OR numeral y) BIT True"
+  "- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = (- numeral (x + Num.One) OR numeral y) BIT True"
+  "- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (- numeral x OR - numeral y) BIT False"
+  "- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = (- numeral x OR - numeral (y + Num.One)) BIT True"
+  "- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = (- numeral (x + Num.One) OR - numeral y) BIT True"
+  "- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = (- numeral (x + Num.One) OR - numeral (y + Num.One)) BIT True"
   "(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)"
   "(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)"
   "(1::int) OR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)"
@@ -293,22 +293,22 @@
   by (rule bin_rl_eqI, simp, simp)+
 
 lemma int_xor_numerals [simp]:
-  "numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (numeral x XOR numeral y) BIT 0"
-  "numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = (numeral x XOR numeral y) BIT 1"
-  "numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = (numeral x XOR numeral y) BIT 1"
-  "numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (numeral x XOR numeral y) BIT 0"
-  "numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (numeral x XOR - numeral y) BIT 0"
-  "numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = (numeral x XOR - numeral (y + Num.One)) BIT 1"
-  "numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = (numeral x XOR - numeral y) BIT 1"
-  "numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (numeral x XOR - numeral (y + Num.One)) BIT 0"
-  "- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (- numeral x XOR numeral y) BIT 0"
-  "- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = (- numeral x XOR numeral y) BIT 1"
-  "- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = (- numeral (x + Num.One) XOR numeral y) BIT 1"
-  "- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (- numeral (x + Num.One) XOR numeral y) BIT 0"
-  "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (- numeral x XOR - numeral y) BIT 0"
-  "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = (- numeral x XOR - numeral (y + Num.One)) BIT 1"
-  "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = (- numeral (x + Num.One) XOR - numeral y) BIT 1"
-  "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (- numeral (x + Num.One) XOR - numeral (y + Num.One)) BIT 0"
+  "numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (numeral x XOR numeral y) BIT False"
+  "numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = (numeral x XOR numeral y) BIT True"
+  "numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = (numeral x XOR numeral y) BIT True"
+  "numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (numeral x XOR numeral y) BIT False"
+  "numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (numeral x XOR - numeral y) BIT False"
+  "numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = (numeral x XOR - numeral (y + Num.One)) BIT True"
+  "numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = (numeral x XOR - numeral y) BIT True"
+  "numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (numeral x XOR - numeral (y + Num.One)) BIT False"
+  "- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (- numeral x XOR numeral y) BIT False"
+  "- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = (- numeral x XOR numeral y) BIT True"
+  "- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = (- numeral (x + Num.One) XOR numeral y) BIT True"
+  "- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (- numeral (x + Num.One) XOR numeral y) BIT False"
+  "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (- numeral x XOR - numeral y) BIT False"
+  "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = (- numeral x XOR - numeral (y + Num.One)) BIT True"
+  "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = (- numeral (x + Num.One) XOR - numeral y) BIT True"
+  "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (- numeral (x + Num.One) XOR - numeral (y + Num.One)) BIT False"
   "(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)"
   "(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)"
   "(1::int) XOR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)"
@@ -332,7 +332,7 @@
   apply (unfold Bit_def)
   apply clarsimp
   apply (erule_tac x = "x" in allE)
-  apply (simp add: bitval_def split: bit.split)
+  apply simp
   done
 
 lemma le_int_or:
@@ -385,7 +385,7 @@
 subsection {* Setting and clearing bits *}
 
 primrec
-  bin_sc :: "nat => bit => int => int"
+  bin_sc :: "nat => bool => int => int"
 where
   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"
@@ -393,7 +393,7 @@
 (** nth bit, set/clear **)
 
 lemma bin_nth_sc [simp]: 
-  "bin_nth (bin_sc n b w) n = (b = 1)"
+  "bin_nth (bin_sc n b w) n \<longleftrightarrow> b"
   by (induct n arbitrary: w) auto
 
 lemma bin_sc_sc_same [simp]: 
@@ -409,11 +409,11 @@
   done
 
 lemma bin_nth_sc_gen: 
-  "bin_nth (bin_sc n b w) m = (if m = n then b = 1 else bin_nth w m)"
+  "bin_nth (bin_sc n b w) m = (if m = n then b else bin_nth w m)"
   by (induct n arbitrary: w m) (case_tac [!] m, auto)
   
 lemma bin_sc_nth [simp]:
-  "(bin_sc n (If (bin_nth w n) 1 0) w) = w"
+  "(bin_sc n (bin_nth w n) w) = w"
   by (induct n arbitrary: w) auto
 
 lemma bin_sign_sc [simp]:
@@ -428,21 +428,21 @@
   done
 
 lemma bin_clr_le:
-  "bin_sc n 0 w <= w"
+  "bin_sc n False w <= w"
   apply (induct n arbitrary: w)
    apply (case_tac [!] w rule: bin_exhaust)
    apply (auto simp: le_Bits)
   done
 
 lemma bin_set_ge:
-  "bin_sc n 1 w >= w"
+  "bin_sc n True w >= w"
   apply (induct n arbitrary: w)
    apply (case_tac [!] w rule: bin_exhaust)
    apply (auto simp: le_Bits)
   done
 
 lemma bintr_bin_clr_le:
-  "bintrunc n (bin_sc m 0 w) <= bintrunc n w"
+  "bintrunc n (bin_sc m False w) <= bintrunc n w"
   apply (induct n arbitrary: w m)
    apply simp
   apply (case_tac w rule: bin_exhaust)
@@ -451,7 +451,7 @@
   done
 
 lemma bintr_bin_set_ge:
-  "bintrunc n (bin_sc m 1 w) >= bintrunc n w"
+  "bintrunc n (bin_sc m True w) >= bintrunc n w"
   apply (induct n arbitrary: w m)
    apply simp
   apply (case_tac w rule: bin_exhaust)
@@ -459,10 +459,10 @@
    apply (auto simp: le_Bits)
   done
 
-lemma bin_sc_FP [simp]: "bin_sc n 0 0 = 0"
+lemma bin_sc_FP [simp]: "bin_sc n False 0 = 0"
   by (induct n) auto
 
-lemma bin_sc_TM [simp]: "bin_sc n 1 -1 = -1"
+lemma bin_sc_TM [simp]: "bin_sc n True -1 = -1"
   by (induct n) auto
   
 lemmas bin_sc_simps = bin_sc.Z bin_sc.Suc bin_sc_TM bin_sc_FP
@@ -482,25 +482,30 @@
 
 subsection {* Splitting and concatenation *}
 
-definition bin_rcat :: "nat \<Rightarrow> int list \<Rightarrow> int" where
+definition bin_rcat :: "nat \<Rightarrow> int list \<Rightarrow> int"
+where
   "bin_rcat n = foldl (\<lambda>u v. bin_cat u n v) 0"
 
-fun bin_rsplit_aux :: "nat \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int list \<Rightarrow> int list" where
+fun bin_rsplit_aux :: "nat \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int list \<Rightarrow> int list"
+where
   "bin_rsplit_aux n m c bs =
     (if m = 0 | n = 0 then bs else
       let (a, b) = bin_split n c 
       in bin_rsplit_aux n (m - n) a (b # bs))"
 
-definition bin_rsplit :: "nat \<Rightarrow> nat \<times> int \<Rightarrow> int list" where
+definition bin_rsplit :: "nat \<Rightarrow> nat \<times> int \<Rightarrow> int list"
+where
   "bin_rsplit n w = bin_rsplit_aux n (fst w) (snd w) []"
 
-fun bin_rsplitl_aux :: "nat \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int list \<Rightarrow> int list" where
+fun bin_rsplitl_aux :: "nat \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int list \<Rightarrow> int list"
+where
   "bin_rsplitl_aux n m c bs =
     (if m = 0 | n = 0 then bs else
       let (a, b) = bin_split (min m n) c 
       in bin_rsplitl_aux n (m - n) a (b # bs))"
 
-definition bin_rsplitl :: "nat \<Rightarrow> nat \<times> int \<Rightarrow> int list" where
+definition bin_rsplitl :: "nat \<Rightarrow> nat \<times> int \<Rightarrow> int list"
+where
   "bin_rsplitl n w = bin_rsplitl_aux n (fst w) (snd w) []"
 
 declare bin_rsplit_aux.simps [simp del]
@@ -611,8 +616,7 @@
   apply (simp add: bin_rest_def zdiv_zmult2_eq)
   apply (case_tac b rule: bin_exhaust)
   apply simp
-  apply (simp add: Bit_def mod_mult_mult1 p1mod22k bitval_def
-              split: bit.split)
+  apply (simp add: Bit_def mod_mult_mult1 p1mod22k)
   done
 
 subsection {* Miscellaneous lemmas *}
@@ -632,7 +636,7 @@
   "(EX m. n = Suc m & (m = k | P m)) = (n = Suc k | (EX m. n = Suc m & P m))"
   by auto
 
-lemma power_BIT: "2 ^ (Suc n) - 1 = (2 ^ n - 1) BIT 1"
+lemma power_BIT: "2 ^ (Suc n) - 1 = (2 ^ n - 1) BIT True"
   unfolding Bit_B1
   by (induct n) simp_all
 
@@ -645,8 +649,8 @@
     by (rule mult_left_mono) simp
   then have "2 * (bin mod 2 ^ n) + 1 < 2 * 2 ^ n" by simp
   then show ?thesis
-    by (auto simp add: Bit_def bitval_def mod_mult_mult1 mod_add_left_eq [of "2 * bin"]
-      mod_pos_pos_trivial split add: bit.split)
+    by (auto simp add: Bit_def mod_mult_mult1 mod_add_left_eq [of "2 * bin"]
+      mod_pos_pos_trivial)
 qed
 
 lemma AND_mod:
@@ -665,7 +669,7 @@
   show ?case
   proof (cases n)
     case 0
-    then show ?thesis by (simp add: int_and_extra_simps)
+    then show ?thesis by simp
   next
     case (Suc m)
     with 3 show ?thesis
--- a/src/HOL/Word/Bit_Operations.thy	Mon Dec 23 15:30:31 2013 +0100
+++ b/src/HOL/Word/Bit_Operations.thy	Mon Dec 23 16:59:56 2013 +0100
@@ -5,7 +5,7 @@
 header {* Syntactic classes for bitwise operations *}
 
 theory Bit_Operations
-imports "~~/src/HOL/Library/Bit"
+imports Main
 begin
 
 class bit =
--- a/src/HOL/Word/Bit_Representation.thy	Mon Dec 23 15:30:31 2013 +0100
+++ b/src/HOL/Word/Bit_Representation.thy	Mon Dec 23 16:59:56 2013 +0100
@@ -5,40 +5,39 @@
 header {* Integers as implict bit strings *}
 
 theory Bit_Representation
-imports "~~/src/HOL/Library/Bit" Misc_Numeric
+imports Misc_Numeric
 begin
 
 subsection {* Constructors and destructors for binary integers *}
 
-definition bitval :: "bit \<Rightarrow> 'a\<Colon>zero_neq_one" where
-  "bitval = bit_case 0 1"
-
-lemma bitval_simps [simp]:
-  "bitval 0 = 0"
-  "bitval 1 = 1"
-  by (simp_all add: bitval_def)
-
-definition Bit :: "int \<Rightarrow> bit \<Rightarrow> int" (infixl "BIT" 90) where
-  "k BIT b = bitval b + k + k"
+definition Bit :: "int \<Rightarrow> bool \<Rightarrow> int" (infixl "BIT" 90)
+where
+  "k BIT b = (if b then 1 else 0) + k + k"
 
 lemma Bit_B0:
-  "k BIT (0::bit) = k + k"
+  "k BIT False = k + k"
    by (unfold Bit_def) simp
 
 lemma Bit_B1:
-  "k BIT (1::bit) = k + k + 1"
+  "k BIT True = k + k + 1"
    by (unfold Bit_def) simp
   
-lemma Bit_B0_2t: "k BIT (0::bit) = 2 * k"
+lemma Bit_B0_2t: "k BIT False = 2 * k"
   by (rule trans, rule Bit_B0) simp
 
-lemma Bit_B1_2t: "k BIT (1::bit) = 2 * k + 1"
+lemma Bit_B1_2t: "k BIT True = 2 * k + 1"
   by (rule trans, rule Bit_B1) simp
 
-definition bin_last :: "int \<Rightarrow> bit" where
-  "bin_last w = (if w mod 2 = 0 then (0::bit) else (1::bit))"
+definition bin_last :: "int \<Rightarrow> bool"
+where
+  "bin_last w \<longleftrightarrow> w mod 2 = 1"
 
-definition bin_rest :: "int \<Rightarrow> int" where
+lemma bin_last_odd:
+  "bin_last = odd"
+  by (rule ext) (simp add: bin_last_def even_def)
+
+definition bin_rest :: "int \<Rightarrow> int"
+where
   "bin_rest w = w div 2"
 
 lemma bin_rl_simp [simp]:
@@ -56,48 +55,55 @@
   by (cases b, simp_all add: z1pmod2)
 
 lemma BIT_eq_iff [iff]: "u BIT b = v BIT c \<longleftrightarrow> u = v \<and> b = c"
-  by (metis bin_rest_BIT bin_last_BIT)
+  apply (auto simp add: Bit_def)
+  apply arith
+  apply arith
+  done
 
 lemma BIT_bin_simps [simp]:
-  "numeral k BIT 0 = numeral (Num.Bit0 k)"
-  "numeral k BIT 1 = numeral (Num.Bit1 k)"
-  "(- numeral k) BIT 0 = - numeral (Num.Bit0 k)"
-  "(- numeral k) BIT 1 = - numeral (Num.BitM k)"
+  "numeral k BIT False = numeral (Num.Bit0 k)"
+  "numeral k BIT True = numeral (Num.Bit1 k)"
+  "(- numeral k) BIT False = - numeral (Num.Bit0 k)"
+  "(- numeral k) BIT True = - numeral (Num.BitM k)"
   unfolding numeral.simps numeral_BitM
-  unfolding Bit_def bitval_simps
+  unfolding Bit_def
   by (simp_all del: arith_simps add_numeral_special diff_numeral_special)
 
 lemma BIT_special_simps [simp]:
-  shows "0 BIT 0 = 0" and "0 BIT 1 = 1"
-  and "1 BIT 0 = 2" and "1 BIT 1 = 3"
-  and "(- 1) BIT 0 = - 2" and "(- 1) BIT 1 = - 1"
+  shows "0 BIT False = 0" and "0 BIT True = 1"
+  and "1 BIT False = 2" and "1 BIT True = 3"
+  and "(- 1) BIT False = - 2" and "(- 1) BIT True = - 1"
   unfolding Bit_def by simp_all
 
-lemma Bit_eq_0_iff: "w BIT b = 0 \<longleftrightarrow> w = 0 \<and> b = 0"
-  by (subst BIT_eq_iff [symmetric], simp)
+lemma Bit_eq_0_iff: "w BIT b = 0 \<longleftrightarrow> w = 0 \<and> \<not> b"
+  apply (auto simp add: Bit_def)
+  apply arith
+  done
 
-lemma Bit_eq_m1_iff: "w BIT b = - 1 \<longleftrightarrow> w = - 1 \<and> b = 1"
-  by (cases b) (auto simp add: Bit_def, arith)
+lemma Bit_eq_m1_iff: "w BIT b = -1 \<longleftrightarrow> w = -1 \<and> b"
+  apply (auto simp add: Bit_def)
+  apply arith
+  done
 
 lemma BitM_inc: "Num.BitM (Num.inc w) = Num.Bit1 w"
   by (induct w, simp_all)
 
 lemma expand_BIT:
-  "numeral (Num.Bit0 w) = numeral w BIT 0"
-  "numeral (Num.Bit1 w) = numeral w BIT 1"
-  "- numeral (Num.Bit0 w) = - numeral w BIT 0"
-  "- numeral (Num.Bit1 w) = (- numeral (w + Num.One)) BIT 1"
+  "numeral (Num.Bit0 w) = numeral w BIT False"
+  "numeral (Num.Bit1 w) = numeral w BIT True"
+  "- numeral (Num.Bit0 w) = (- numeral w) BIT False"
+  "- numeral (Num.Bit1 w) = (- numeral (w + Num.One)) BIT True"
   unfolding add_One by (simp_all add: BitM_inc)
 
 lemma bin_last_numeral_simps [simp]:
-  "bin_last 0 = 0"
-  "bin_last 1 = 1"
-  "bin_last -1 = 1"
-  "bin_last Numeral1 = 1"
-  "bin_last (numeral (Num.Bit0 w)) = 0"
-  "bin_last (numeral (Num.Bit1 w)) = 1"
-  "bin_last (- numeral (Num.Bit0 w)) = 0"
-  "bin_last (- numeral (Num.Bit1 w)) = 1"
+  "\<not> bin_last 0"
+  "bin_last 1"
+  "bin_last -1"
+  "bin_last Numeral1"
+  "\<not> bin_last (numeral (Num.Bit0 w))"
+  "bin_last (numeral (Num.Bit1 w))"
+  "\<not> bin_last (- numeral (Num.Bit0 w))"
+  "bin_last (- numeral (Num.Bit1 w))"
   unfolding expand_BIT bin_last_BIT by (simp_all add: bin_last_def zmod_zminus1_eq_if)
 
 lemma bin_rest_numeral_simps [simp]:
@@ -112,51 +118,42 @@
   unfolding expand_BIT bin_rest_BIT by (simp_all add: bin_rest_def zdiv_zminus1_eq_if)
 
 lemma less_Bits: 
-  "(v BIT b < w BIT c) = (v < w | v <= w & b = (0::bit) & c = (1::bit))"
-  unfolding Bit_def by (auto simp add: bitval_def split: bit.split)
+  "v BIT b < w BIT c \<longleftrightarrow> v < w \<or> v \<le> w \<and> \<not> b \<and> c"
+  unfolding Bit_def by auto
 
 lemma le_Bits: 
-  "(v BIT b <= w BIT c) = (v < w | v <= w & (b ~= (1::bit) | c ~= (0::bit)))" 
-  unfolding Bit_def by (auto simp add: bitval_def split: bit.split)
+  "v BIT b \<le> w BIT c \<longleftrightarrow> v < w \<or> v \<le> w \<and> (\<not> b \<or> c)" 
+  unfolding Bit_def by auto
 
 lemma pred_BIT_simps [simp]:
-  "x BIT 0 - 1 = (x - 1) BIT 1"
-  "x BIT 1 - 1 = x BIT 0"
+  "x BIT False - 1 = (x - 1) BIT True"
+  "x BIT True - 1 = x BIT False"
   by (simp_all add: Bit_B0_2t Bit_B1_2t)
 
 lemma succ_BIT_simps [simp]:
-  "x BIT 0 + 1 = x BIT 1"
-  "x BIT 1 + 1 = (x + 1) BIT 0"
+  "x BIT False + 1 = x BIT True"
+  "x BIT True + 1 = (x + 1) BIT False"
   by (simp_all add: Bit_B0_2t Bit_B1_2t)
 
 lemma add_BIT_simps [simp]:
-  "x BIT 0 + y BIT 0 = (x + y) BIT 0"
-  "x BIT 0 + y BIT 1 = (x + y) BIT 1"
-  "x BIT 1 + y BIT 0 = (x + y) BIT 1"
-  "x BIT 1 + y BIT 1 = (x + y + 1) BIT 0"
+  "x BIT False + y BIT False = (x + y) BIT False"
+  "x BIT False + y BIT True = (x + y) BIT True"
+  "x BIT True + y BIT False = (x + y) BIT True"
+  "x BIT True + y BIT True = (x + y + 1) BIT False"
   by (simp_all add: Bit_B0_2t Bit_B1_2t)
 
 lemma mult_BIT_simps [simp]:
-  "x BIT 0 * y = (x * y) BIT 0"
-  "x * y BIT 0 = (x * y) BIT 0"
-  "x BIT 1 * y = (x * y) BIT 0 + y"
+  "x BIT False * y = (x * y) BIT False"
+  "x * y BIT False = (x * y) BIT False"
+  "x BIT True * y = (x * y) BIT False + y"
   by (simp_all add: Bit_B0_2t Bit_B1_2t algebra_simps)
 
 lemma B_mod_2': 
-  "X = 2 ==> (w BIT (1::bit)) mod X = 1 & (w BIT (0::bit)) mod X = 0"
+  "X = 2 ==> (w BIT True) mod X = 1 & (w BIT False) mod X = 0"
   apply (simp (no_asm) only: Bit_B0 Bit_B1)
   apply (simp add: z1pmod2)
   done
 
-lemma neB1E [elim!]:
-  assumes ne: "y \<noteq> (1::bit)"
-  assumes y: "y = (0::bit) \<Longrightarrow> P"
-  shows "P"
-  apply (rule y)
-  apply (cases y rule: bit.exhaust, simp)
-  apply (simp add: ne)
-  done
-
 lemma bin_ex_rl: "EX w b. w BIT b = bin"
   by (metis bin_rl_simp)
 
@@ -170,8 +167,10 @@
   done
 
 primrec bin_nth where
-  Z: "bin_nth w 0 = (bin_last w = (1::bit))"
-  | Suc: "bin_nth w (Suc n) = bin_nth (bin_rest w) n"
+  Z: "bin_nth w 0 \<longleftrightarrow> bin_last w"
+  | Suc: "bin_nth w (Suc n) \<longleftrightarrow> bin_nth (bin_rest w) n"
+
+find_theorems "bin_rest _ = _"
 
 lemma bin_abs_lem:
   "bin = (w BIT b) ==> bin ~= -1 --> bin ~= 0 -->
@@ -248,7 +247,7 @@
 lemma bin_nth_minus1 [simp]: "bin_nth -1 n"
   by (induct n) auto
 
-lemma bin_nth_0_BIT: "bin_nth (w BIT b) 0 = (b = (1::bit))"
+lemma bin_nth_0_BIT: "bin_nth (w BIT b) 0 \<longleftrightarrow> b"
   by auto
 
 lemma bin_nth_Suc_BIT: "bin_nth (w BIT b) (Suc n) = bin_nth w n"
@@ -275,7 +274,8 @@
 
 subsection {* Truncating binary integers *}
 
-definition bin_sign :: "int \<Rightarrow> int" where
+definition bin_sign :: "int \<Rightarrow> int"
+where
   bin_sign_def: "bin_sign k = (if k \<ge> 0 then 0 else - 1)"
 
 lemma bin_sign_simps [simp]:
@@ -285,8 +285,8 @@
   "bin_sign (numeral k) = 0"
   "bin_sign (- numeral k) = -1"
   "bin_sign (w BIT b) = bin_sign w"
-  unfolding bin_sign_def Bit_def bitval_def
-  by (simp_all split: bit.split)
+  unfolding bin_sign_def Bit_def
+  by simp_all
 
 lemma bin_sign_rest [simp]: 
   "bin_sign (bin_rest w) = bin_sign w"
@@ -297,7 +297,7 @@
 | Suc : "bintrunc (Suc n) bin = bintrunc n (bin_rest bin) BIT (bin_last bin)"
 
 primrec sbintrunc :: "nat => int => int" where
-  Z : "sbintrunc 0 bin = (case bin_last bin of (1::bit) \<Rightarrow> -1 | (0::bit) \<Rightarrow> 0)"
+  Z : "sbintrunc 0 bin = (if bin_last bin then -1 else 0)"
 | Suc : "sbintrunc (Suc n) bin = sbintrunc n (bin_rest bin) BIT (bin_last bin)"
 
 lemma sign_bintr: "bin_sign (bintrunc n w) = 0"
@@ -313,7 +313,8 @@
    apply simp
    apply (subst mod_add_left_eq)
    apply (simp add: bin_last_def)
-  apply (simp add: bin_last_def bin_rest_def Bit_def)
+   apply arith
+  apply (simp add: bin_last_def bin_rest_def Bit_def mod_2_neq_1_eq_eq_0)
   apply (clarsimp simp: mod_mult_mult1 [symmetric] 
          zmod_zdiv_equality [THEN diff_eq_eq [THEN iffD2 [THEN sym]]])
   apply (rule trans [symmetric, OF _ emep1])
@@ -334,13 +335,13 @@
 
 lemma bintrunc_Suc_numeral:
   "bintrunc (Suc n) 1 = 1"
-  "bintrunc (Suc n) -1 = bintrunc n -1 BIT 1"
-  "bintrunc (Suc n) (numeral (Num.Bit0 w)) = bintrunc n (numeral w) BIT 0"
-  "bintrunc (Suc n) (numeral (Num.Bit1 w)) = bintrunc n (numeral w) BIT 1"
+  "bintrunc (Suc n) -1 = bintrunc n -1 BIT True"
+  "bintrunc (Suc n) (numeral (Num.Bit0 w)) = bintrunc n (numeral w) BIT False"
+  "bintrunc (Suc n) (numeral (Num.Bit1 w)) = bintrunc n (numeral w) BIT True"
   "bintrunc (Suc n) (- numeral (Num.Bit0 w)) =
-    bintrunc n (- numeral w) BIT 0"
+    bintrunc n (- numeral w) BIT False"
   "bintrunc (Suc n) (- numeral (Num.Bit1 w)) =
-    bintrunc n (- numeral (w + Num.One)) BIT 1"
+    bintrunc n (- numeral (w + Num.One)) BIT True"
   by simp_all
 
 lemma sbintrunc_0_numeral [simp]:
@@ -354,21 +355,15 @@
 lemma sbintrunc_Suc_numeral:
   "sbintrunc (Suc n) 1 = 1"
   "sbintrunc (Suc n) (numeral (Num.Bit0 w)) =
-    sbintrunc n (numeral w) BIT 0"
+    sbintrunc n (numeral w) BIT False"
   "sbintrunc (Suc n) (numeral (Num.Bit1 w)) =
-    sbintrunc n (numeral w) BIT 1"
+    sbintrunc n (numeral w) BIT True"
   "sbintrunc (Suc n) (- numeral (Num.Bit0 w)) =
-    sbintrunc n (- numeral w) BIT 0"
+    sbintrunc n (- numeral w) BIT False"
   "sbintrunc (Suc n) (- numeral (Num.Bit1 w)) =
-    sbintrunc n (- numeral (w + Num.One)) BIT 1"
+    sbintrunc n (- numeral (w + Num.One)) BIT True"
   by simp_all
 
-lemma bit_bool:
-  "(b = (b' = (1::bit))) = (b' = (if b then (1::bit) else (0::bit)))"
-  by (cases b') auto
-
-lemmas bit_bool1 [simp] = refl [THEN bit_bool [THEN iffD1], symmetric]
-
 lemma bin_sign_lem: "(bin_sign (sbintrunc n bin) = -1) = bin_nth bin n"
   apply (induct n arbitrary: bin)
   apply (case_tac bin rule: bin_exhaust, case_tac b, auto)
@@ -384,23 +379,25 @@
   "bin_nth (sbintrunc m w) n = 
           (if n < m then bin_nth w n else bin_nth w m)"
   apply (induct n arbitrary: w m)
-   apply (case_tac m, simp_all split: bit.splits)[1]
-  apply (case_tac m, simp_all split: bit.splits)[1]
+  apply (case_tac m)
+  apply simp_all
+  apply (case_tac m)
+  apply simp_all
   done
 
 lemma bin_nth_Bit:
-  "bin_nth (w BIT b) n = (n = 0 & b = (1::bit) | (EX m. n = Suc m & bin_nth w m))"
+  "bin_nth (w BIT b) n = (n = 0 & b | (EX m. n = Suc m & bin_nth w m))"
   by (cases n) auto
 
 lemma bin_nth_Bit0:
   "bin_nth (numeral (Num.Bit0 w)) n \<longleftrightarrow>
     (\<exists>m. n = Suc m \<and> bin_nth (numeral w) m)"
-  using bin_nth_Bit [where w="numeral w" and b="(0::bit)"] by simp
+  using bin_nth_Bit [where w="numeral w" and b="False"] by simp
 
 lemma bin_nth_Bit1:
   "bin_nth (numeral (Num.Bit1 w)) n \<longleftrightarrow>
     n = 0 \<or> (\<exists>m. n = Suc m \<and> bin_nth (numeral w) m)"
-  using bin_nth_Bit [where w="numeral w" and b="(1::bit)"] by simp
+  using bin_nth_Bit [where w="numeral w" and b="True"] by simp
 
 lemma bintrunc_bintrunc_l:
   "n <= m ==> (bintrunc m (bintrunc n w) = bintrunc n w)"
@@ -452,19 +449,19 @@
 
 lemmas sbintrunc_Pls = 
   sbintrunc.Z [where bin="0", 
-               simplified bin_last_numeral_simps bin_rest_numeral_simps bit.simps]
+               simplified bin_last_numeral_simps bin_rest_numeral_simps]
 
 lemmas sbintrunc_Min = 
   sbintrunc.Z [where bin="-1",
-               simplified bin_last_numeral_simps bin_rest_numeral_simps bit.simps]
+               simplified bin_last_numeral_simps bin_rest_numeral_simps]
 
 lemmas sbintrunc_0_BIT_B0 [simp] = 
-  sbintrunc.Z [where bin="w BIT (0::bit)", 
-               simplified bin_last_numeral_simps bin_rest_numeral_simps bit.simps] for w
+  sbintrunc.Z [where bin="w BIT False", 
+               simplified bin_last_numeral_simps bin_rest_numeral_simps] for w
 
 lemmas sbintrunc_0_BIT_B1 [simp] = 
-  sbintrunc.Z [where bin="w BIT (1::bit)", 
-               simplified bin_last_BIT bin_rest_numeral_simps bit.simps] for w
+  sbintrunc.Z [where bin="w BIT True", 
+               simplified bin_last_BIT bin_rest_numeral_simps] for w
 
 lemmas sbintrunc_0_simps =
   sbintrunc_Pls sbintrunc_Min sbintrunc_0_BIT_B0 sbintrunc_0_BIT_B1
@@ -583,25 +580,25 @@
 
 lemma bintrunc_numeral_simps [simp]:
   "bintrunc (numeral k) (numeral (Num.Bit0 w)) =
-    bintrunc (pred_numeral k) (numeral w) BIT 0"
+    bintrunc (pred_numeral k) (numeral w) BIT False"
   "bintrunc (numeral k) (numeral (Num.Bit1 w)) =
-    bintrunc (pred_numeral k) (numeral w) BIT 1"
+    bintrunc (pred_numeral k) (numeral w) BIT True"
   "bintrunc (numeral k) (- numeral (Num.Bit0 w)) =
-    bintrunc (pred_numeral k) (- numeral w) BIT 0"
+    bintrunc (pred_numeral k) (- numeral w) BIT False"
   "bintrunc (numeral k) (- numeral (Num.Bit1 w)) =
-    bintrunc (pred_numeral k) (- numeral (w + Num.One)) BIT 1"
+    bintrunc (pred_numeral k) (- numeral (w + Num.One)) BIT True"
   "bintrunc (numeral k) 1 = 1"
   by (simp_all add: bintrunc_numeral)
 
 lemma sbintrunc_numeral_simps [simp]:
   "sbintrunc (numeral k) (numeral (Num.Bit0 w)) =
-    sbintrunc (pred_numeral k) (numeral w) BIT 0"
+    sbintrunc (pred_numeral k) (numeral w) BIT False"
   "sbintrunc (numeral k) (numeral (Num.Bit1 w)) =
-    sbintrunc (pred_numeral k) (numeral w) BIT 1"
+    sbintrunc (pred_numeral k) (numeral w) BIT True"
   "sbintrunc (numeral k) (- numeral (Num.Bit0 w)) =
-    sbintrunc (pred_numeral k) (- numeral w) BIT 0"
+    sbintrunc (pred_numeral k) (- numeral w) BIT False"
   "sbintrunc (numeral k) (- numeral (Num.Bit1 w)) =
-    sbintrunc (pred_numeral k) (- numeral (w + Num.One)) BIT 1"
+    sbintrunc (pred_numeral k) (- numeral (w + Num.One)) BIT True"
   "sbintrunc (numeral k) 1 = 1"
   by (simp_all add: sbintrunc_numeral)
 
@@ -728,7 +725,7 @@
   "sbintrunc n (bin_rest (sbintrunc n bin)) = bin_rest (sbintrunc n bin)"
   apply (induct n arbitrary: bin, simp)
   apply (case_tac bin rule: bin_exhaust)
-  apply (auto simp: bintrunc_bintrunc_l split: bit.splits)
+  apply (auto simp: bintrunc_bintrunc_l split: bool.splits)
   done
 
 lemma bintrunc_rest':
--- a/src/HOL/Word/Bool_List_Representation.thy	Mon Dec 23 15:30:31 2013 +0100
+++ b/src/HOL/Word/Bool_List_Representation.thy	Mon Dec 23 16:59:56 2013 +0100
@@ -31,27 +31,33 @@
 
 subsection {* Operations on lists of booleans *}
 
-primrec bl_to_bin_aux :: "bool list \<Rightarrow> int \<Rightarrow> int" where
+primrec bl_to_bin_aux :: "bool list \<Rightarrow> int \<Rightarrow> int"
+where
   Nil: "bl_to_bin_aux [] w = w"
   | Cons: "bl_to_bin_aux (b # bs) w = 
-      bl_to_bin_aux bs (w BIT (if b then 1 else 0))"
+      bl_to_bin_aux bs (w BIT b)"
 
-definition bl_to_bin :: "bool list \<Rightarrow> int" where
+definition bl_to_bin :: "bool list \<Rightarrow> int"
+where
   bl_to_bin_def: "bl_to_bin bs = bl_to_bin_aux bs 0"
 
-primrec bin_to_bl_aux :: "nat \<Rightarrow> int \<Rightarrow> bool list \<Rightarrow> bool list" where
+primrec bin_to_bl_aux :: "nat \<Rightarrow> int \<Rightarrow> bool list \<Rightarrow> bool list"
+where
   Z: "bin_to_bl_aux 0 w bl = bl"
   | Suc: "bin_to_bl_aux (Suc n) w bl =
-      bin_to_bl_aux n (bin_rest w) ((bin_last w = 1) # bl)"
+      bin_to_bl_aux n (bin_rest w) ((bin_last w) # bl)"
 
-definition bin_to_bl :: "nat \<Rightarrow> int \<Rightarrow> bool list" where
+definition bin_to_bl :: "nat \<Rightarrow> int \<Rightarrow> bool list"
+where
   bin_to_bl_def : "bin_to_bl n w = bin_to_bl_aux n w []"
 
-primrec bl_of_nth :: "nat \<Rightarrow> (nat \<Rightarrow> bool) \<Rightarrow> bool list" where
+primrec bl_of_nth :: "nat \<Rightarrow> (nat \<Rightarrow> bool) \<Rightarrow> bool list"
+where
   Suc: "bl_of_nth (Suc n) f = f n # bl_of_nth n f"
   | Z: "bl_of_nth 0 f = []"
 
-primrec takefill :: "'a \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+primrec takefill :: "'a \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
+where
   Z: "takefill fill 0 xs = []"
   | Suc: "takefill fill (Suc n) xs = (
       case xs of [] => fill # takefill fill n xs
@@ -65,21 +71,25 @@
   assuming input list(s) the same length, and don't extend them. 
 *}
 
-primrec rbl_succ :: "bool list => bool list" where
+primrec rbl_succ :: "bool list => bool list"
+where
   Nil: "rbl_succ Nil = Nil"
   | Cons: "rbl_succ (x # xs) = (if x then False # rbl_succ xs else True # xs)"
 
-primrec rbl_pred :: "bool list => bool list" where
+primrec rbl_pred :: "bool list => bool list"
+where
   Nil: "rbl_pred Nil = Nil"
   | Cons: "rbl_pred (x # xs) = (if x then False # xs else True # rbl_pred xs)"
 
-primrec rbl_add :: "bool list => bool list => bool list" where
+primrec rbl_add :: "bool list => bool list => bool list"
+where
   -- "result is length of first arg, second arg may be longer"
   Nil: "rbl_add Nil x = Nil"
   | Cons: "rbl_add (y # ys) x = (let ws = rbl_add ys (tl x) in 
     (y ~= hd x) # (if hd x & y then rbl_succ ws else ws))"
 
-primrec rbl_mult :: "bool list => bool list => bool list" where
+primrec rbl_mult :: "bool list => bool list => bool list"
+where
   -- "result is length of first arg, second arg may be longer"
   Nil: "rbl_mult Nil x = Nil"
   | Cons: "rbl_mult (y # ys) x = (let ws = False # rbl_mult ys x in 
@@ -106,7 +116,7 @@
 
 lemma bin_to_bl_aux_Bit_minus_simp [simp]:
   "0 < n ==> bin_to_bl_aux n (w BIT b) bl = 
-    bin_to_bl_aux (n - 1) w ((b = 1) # bl)"
+    bin_to_bl_aux (n - 1) w (b # bl)"
   by (cases n) auto
 
 lemma bin_to_bl_aux_Bit0_minus_simp [simp]:
@@ -253,8 +263,7 @@
   apply (induct n arbitrary: w bs)
    apply clarsimp
    apply (cases w rule: bin_exhaust)
-   apply (simp split add : bit.split)
-  apply clarsimp
+   apply simp
   done
     
 lemma bl_sbin_sign: 
@@ -317,7 +326,6 @@
   apply (induct bs arbitrary: w)
    apply clarsimp
   apply clarsimp
-  apply safe
   apply (drule meta_spec, erule xtrans(8) [rotated], simp add: Bit_def)+
   done
 
@@ -334,9 +342,9 @@
   apply (induct bs arbitrary: w)
    apply clarsimp
   apply clarsimp
-  apply safe
    apply (drule meta_spec, erule order_trans [rotated],
           simp add: Bit_B0_2t Bit_B1_2t algebra_simps)+
+   apply (simp add: Bit_def)
   done
 
 lemma bl_to_bin_ge0: "bl_to_bin bs >= 0"
@@ -422,15 +430,15 @@
   by (cases xs) auto
 
 lemma last_bin_last': 
-  "size xs > 0 \<Longrightarrow> last xs = (bin_last (bl_to_bin_aux xs w) = 1)" 
+  "size xs > 0 \<Longrightarrow> last xs \<longleftrightarrow> bin_last (bl_to_bin_aux xs w)" 
   by (induct xs arbitrary: w) auto
 
 lemma last_bin_last: 
-  "size xs > 0 ==> last xs = (bin_last (bl_to_bin xs) = 1)" 
+  "size xs > 0 ==> last xs \<longleftrightarrow> bin_last (bl_to_bin xs)" 
   unfolding bl_to_bin_def by (erule last_bin_last')
   
 lemma bin_last_last: 
-  "bin_last w = (if last (bin_to_bl (Suc n) w) then 1 else 0)" 
+  "bin_last w \<longleftrightarrow> last (bin_to_bl (Suc n) w)" 
   apply (unfold bin_to_bl_def)
   apply simp
   apply (auto simp add: bin_to_bl_aux_alt)
@@ -447,7 +455,7 @@
   apply (case_tac w rule: bin_exhaust)
   apply clarsimp
   apply (case_tac b)
-  apply (case_tac ba, safe, simp_all)+
+  apply auto
   done
 
 lemma bl_or_aux_bin:
@@ -458,8 +466,6 @@
   apply (case_tac v rule: bin_exhaust)
   apply (case_tac w rule: bin_exhaust)
   apply clarsimp
-  apply (case_tac b)
-  apply (case_tac ba, safe, simp_all)+
   done
     
 lemma bl_and_aux_bin:
@@ -846,7 +852,7 @@
 lemmas rbl_mult_Suc = lessI [THEN rbl_mult_gt]
 
 lemma rbbl_Cons: 
-  "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (x BIT If b 1 0))"
+  "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (x BIT b))"
   apply (unfold bin_to_bl_def)
   apply simp
   apply (simp add: bin_to_bl_aux_alt)
--- a/src/HOL/Word/Misc_Numeric.thy	Mon Dec 23 15:30:31 2013 +0100
+++ b/src/HOL/Word/Misc_Numeric.thy	Mon Dec 23 16:59:56 2013 +0100
@@ -22,6 +22,11 @@
   
 lemmas min_minus' [simp] = trans [OF min_max.inf_commute min_minus]
 
+lemma mod_2_neq_1_eq_eq_0:
+  fixes k :: int
+  shows "k mod 2 \<noteq> 1 \<longleftrightarrow> k mod 2 = 0"
+  by arith
+
 lemma z1pmod2:
   "(2 * b + 1) mod 2 = (1::int)" by arith
 
--- a/src/HOL/Word/Word.thy	Mon Dec 23 15:30:31 2013 +0100
+++ b/src/HOL/Word/Word.thy	Mon Dec 23 16:59:56 2013 +0100
@@ -35,7 +35,8 @@
   shows "uint w mod 2 ^ len_of TYPE('a) = uint w"
   using uint_nonnegative uint_bounded by (rule mod_pos_pos_trivial)
 
-definition word_of_int :: "int \<Rightarrow> 'a\<Colon>len0 word" where
+definition word_of_int :: "int \<Rightarrow> 'a\<Colon>len0 word"
+where
   -- {* representation of words using unsigned or signed bins, 
         only difference in these is the type class *}
   "word_of_int k = Abs_word (k mod 2 ^ len_of TYPE('a))" 
@@ -72,7 +73,8 @@
 instantiation word :: (len0) equal
 begin
 
-definition equal_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool" where
+definition equal_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"
+where
   "equal_word k l \<longleftrightarrow> HOL.equal (uint k) (uint l)"
 
 instance proof
@@ -101,31 +103,39 @@
 
 subsection {* Type conversions and casting *}
 
-definition sint :: "'a :: len word => int" where
+definition sint :: "'a :: len word => int"
+where
   -- {* treats the most-significant-bit as a sign bit *}
   sint_uint: "sint w = sbintrunc (len_of TYPE ('a) - 1) (uint w)"
 
-definition unat :: "'a :: len0 word => nat" where
+definition unat :: "'a :: len0 word => nat"
+where
   "unat w = nat (uint w)"
 
-definition uints :: "nat => int set" where
+definition uints :: "nat => int set"
+where
   -- "the sets of integers representing the words"
   "uints n = range (bintrunc n)"
 
-definition sints :: "nat => int set" where
+definition sints :: "nat => int set"
+where
   "sints n = range (sbintrunc (n - 1))"
 
-definition unats :: "nat => nat set" where
+definition unats :: "nat => nat set"
+where
   "unats n = {i. i < 2 ^ n}"
 
-definition norm_sint :: "nat => int => int" where
+definition norm_sint :: "nat => int => int"
+where
   "norm_sint n w = (w + 2 ^ (n - 1)) mod 2 ^ n - 2 ^ (n - 1)"
 
-definition scast :: "'a :: len word => 'b :: len word" where
+definition scast :: "'a :: len word => 'b :: len word"
+where
   -- "cast a word to a different length"
   "scast w = word_of_int (sint w)"
 
-definition ucast :: "'a :: len0 word => 'b :: len0 word" where
+definition ucast :: "'a :: len0 word => 'b :: len0 word"
+where
   "ucast w = word_of_int (uint w)"
 
 instantiation word :: (len0) size
@@ -138,29 +148,37 @@
 
 end
 
-definition source_size :: "('a :: len0 word => 'b) => nat" where
+definition source_size :: "('a :: len0 word => 'b) => nat"
+where
   -- "whether a cast (or other) function is to a longer or shorter length"
   "source_size c = (let arb = undefined ; x = c arb in size arb)"  
 
-definition target_size :: "('a => 'b :: len0 word) => nat" where
+definition target_size :: "('a => 'b :: len0 word) => nat"
+where
   "target_size c = size (c undefined)"
 
-definition is_up :: "('a :: len0 word => 'b :: len0 word) => bool" where
+definition is_up :: "('a :: len0 word => 'b :: len0 word) => bool"
+where
   "is_up c \<longleftrightarrow> source_size c <= target_size c"
 
-definition is_down :: "('a :: len0 word => 'b :: len0 word) => bool" where
+definition is_down :: "('a :: len0 word => 'b :: len0 word) => bool"
+where
   "is_down c \<longleftrightarrow> target_size c <= source_size c"
 
-definition of_bl :: "bool list => 'a :: len0 word" where
+definition of_bl :: "bool list => 'a :: len0 word"
+where
   "of_bl bl = word_of_int (bl_to_bin bl)"
 
-definition to_bl :: "'a :: len0 word => bool list" where
+definition to_bl :: "'a :: len0 word => bool list"
+where
   "to_bl w = bin_to_bl (len_of TYPE ('a)) (uint w)"
 
-definition word_reverse :: "'a :: len0 word => 'a word" where
+definition word_reverse :: "'a :: len0 word => 'a word"
+where
   "word_reverse w = of_bl (rev (to_bl w))"
 
-definition word_int_case :: "(int => 'b) => ('a :: len0 word) => 'b" where
+definition word_int_case :: "(int => 'b) => ('a :: len0 word) => 'b" 
+where
   "word_int_case f w = f (uint w)"
 
 translations
@@ -232,7 +250,8 @@
 subsection {* Correspondence relation for theorem transfer *}
 
 definition cr_word :: "int \<Rightarrow> 'a::len0 word \<Rightarrow> bool"
-  where "cr_word \<equiv> (\<lambda>x y. word_of_int x = y)"
+where
+  "cr_word = (\<lambda>x y. word_of_int x = y)"
 
 lemma Quotient_word:
   "Quotient (\<lambda>x y. bintrunc (len_of TYPE('a)) x = bintrunc (len_of TYPE('a)) y)
@@ -341,7 +360,8 @@
   apply (simp add: word_of_nat wi_hom_sub)
   done
 
-definition udvd :: "'a::len word => 'a::len word => bool" (infixl "udvd" 50) where
+definition udvd :: "'a::len word => 'a::len word => bool" (infixl "udvd" 50)
+where
   "a udvd b = (EX n>=0. uint b = n * uint a)"
 
 
@@ -361,10 +381,12 @@
 
 end
 
-definition word_sle :: "'a :: len word => 'a word => bool" ("(_/ <=s _)" [50, 51] 50) where
+definition word_sle :: "'a :: len word => 'a word => bool" ("(_/ <=s _)" [50, 51] 50)
+where
   "a <=s b = (sint a <= sint b)"
 
-definition word_sless :: "'a :: len word => 'a word => bool" ("(_/ <s _)" [50, 51] 50) where
+definition word_sless :: "'a :: len word => 'a word => bool" ("(_/ <s _)" [50, 51] 50)
+where
   "(x <s y) = (x <=s y & x ~= y)"
 
 
@@ -390,18 +412,20 @@
 
 definition
   word_set_bit_def: "set_bit a n x =
-   word_of_int (bin_sc n (If x 1 0) (uint a))"
+   word_of_int (bin_sc n x (uint a))"
 
 definition
   word_set_bits_def: "(BITS n. f n) = of_bl (bl_of_nth (len_of TYPE ('a)) f)"
 
 definition
-  word_lsb_def: "lsb a \<longleftrightarrow> bin_last (uint a) = 1"
-
-definition shiftl1 :: "'a word \<Rightarrow> 'a word" where
-  "shiftl1 w = word_of_int (uint w BIT 0)"
-
-definition shiftr1 :: "'a word \<Rightarrow> 'a word" where
+  word_lsb_def: "lsb a \<longleftrightarrow> bin_last (uint a)"
+
+definition shiftl1 :: "'a word \<Rightarrow> 'a word"
+where
+  "shiftl1 w = word_of_int (uint w BIT False)"
+
+definition shiftr1 :: "'a word \<Rightarrow> 'a word"
+where
   -- "shift right as unsigned or as signed, ie logical or arithmetic"
   "shiftr1 w = word_of_int (bin_rest (uint w))"
 
@@ -434,76 +458,95 @@
 
 end
 
-definition setBit :: "'a :: len0 word => nat => 'a word" where 
+definition setBit :: "'a :: len0 word => nat => 'a word"
+where 
   "setBit w n = set_bit w n True"
 
-definition clearBit :: "'a :: len0 word => nat => 'a word" where
+definition clearBit :: "'a :: len0 word => nat => 'a word"
+where
   "clearBit w n = set_bit w n False"
 
 
 subsection "Shift operations"
 
-definition sshiftr1 :: "'a :: len word => 'a word" where 
+definition sshiftr1 :: "'a :: len word => 'a word"
+where 
   "sshiftr1 w = word_of_int (bin_rest (sint w))"
 
-definition bshiftr1 :: "bool => 'a :: len word => 'a word" where
+definition bshiftr1 :: "bool => 'a :: len word => 'a word"
+where
   "bshiftr1 b w = of_bl (b # butlast (to_bl w))"
 
-definition sshiftr :: "'a :: len word => nat => 'a word" (infixl ">>>" 55) where
+definition sshiftr :: "'a :: len word => nat => 'a word" (infixl ">>>" 55)
+where
   "w >>> n = (sshiftr1 ^^ n) w"
 
-definition mask :: "nat => 'a::len word" where
+definition mask :: "nat => 'a::len word"
+where
   "mask n = (1 << n) - 1"
 
-definition revcast :: "'a :: len0 word => 'b :: len0 word" where
+definition revcast :: "'a :: len0 word => 'b :: len0 word"
+where
   "revcast w =  of_bl (takefill False (len_of TYPE('b)) (to_bl w))"
 
-definition slice1 :: "nat => 'a :: len0 word => 'b :: len0 word" where
+definition slice1 :: "nat => 'a :: len0 word => 'b :: len0 word"
+where
   "slice1 n w = of_bl (takefill False n (to_bl w))"
 
-definition slice :: "nat => 'a :: len0 word => 'b :: len0 word" where
+definition slice :: "nat => 'a :: len0 word => 'b :: len0 word"
+where
   "slice n w = slice1 (size w - n) w"
 
 
 subsection "Rotation"
 
-definition rotater1 :: "'a list => 'a list" where
+definition rotater1 :: "'a list => 'a list"
+where
   "rotater1 ys = 
     (case ys of [] => [] | x # xs => last ys # butlast ys)"
 
-definition rotater :: "nat => 'a list => 'a list" where
+definition rotater :: "nat => 'a list => 'a list"
+where
   "rotater n = rotater1 ^^ n"
 
-definition word_rotr :: "nat => 'a :: len0 word => 'a :: len0 word" where
+definition word_rotr :: "nat => 'a :: len0 word => 'a :: len0 word"
+where
   "word_rotr n w = of_bl (rotater n (to_bl w))"
 
-definition word_rotl :: "nat => 'a :: len0 word => 'a :: len0 word" where
+definition word_rotl :: "nat => 'a :: len0 word => 'a :: len0 word"
+where
   "word_rotl n w = of_bl (rotate n (to_bl w))"
 
-definition word_roti :: "int => 'a :: len0 word => 'a :: len0 word" where
+definition word_roti :: "int => 'a :: len0 word => 'a :: len0 word"
+where
   "word_roti i w = (if i >= 0 then word_rotr (nat i) w
                     else word_rotl (nat (- i)) w)"
 
 
 subsection "Split and cat operations"
 
-definition word_cat :: "'a :: len0 word => 'b :: len0 word => 'c :: len0 word" where
+definition word_cat :: "'a :: len0 word => 'b :: len0 word => 'c :: len0 word"
+where
   "word_cat a b = word_of_int (bin_cat (uint a) (len_of TYPE ('b)) (uint b))"
 
-definition word_split :: "'a :: len0 word => ('b :: len0 word) * ('c :: len0 word)" where
+definition word_split :: "'a :: len0 word => ('b :: len0 word) * ('c :: len0 word)"
+where
   "word_split a = 
    (case bin_split (len_of TYPE ('c)) (uint a) of 
      (u, v) => (word_of_int u, word_of_int v))"
 
-definition word_rcat :: "'a :: len0 word list => 'b :: len0 word" where
+definition word_rcat :: "'a :: len0 word list => 'b :: len0 word"
+where
   "word_rcat ws = 
   word_of_int (bin_rcat (len_of TYPE ('a)) (map uint ws))"
 
-definition word_rsplit :: "'a :: len0 word => 'b :: len word list" where
+definition word_rsplit :: "'a :: len0 word => 'b :: len word list"
+where
   "word_rsplit w = 
   map word_of_int (bin_rsplit (len_of TYPE ('b)) (len_of TYPE ('a), uint w))"
 
-definition max_word :: "'a::len word" -- "Largest representable machine integer." where
+definition max_word :: "'a::len word" -- "Largest representable machine integer."
+where
   "max_word = word_of_int (2 ^ len_of TYPE('a) - 1)"
 
 (* FIXME: only provide one theorem name *)
@@ -2524,15 +2567,15 @@
   by (clarsimp simp add: nth_sbintr word_test_bit_def [symmetric])
 
 lemma word_lsb_numeral [simp]:
-  "lsb (numeral bin :: 'a :: len word) = (bin_last (numeral bin) = 1)"
+  "lsb (numeral bin :: 'a :: len word) \<longleftrightarrow> bin_last (numeral bin)"
   unfolding word_lsb_alt test_bit_numeral by simp
 
 lemma word_lsb_neg_numeral [simp]:
-  "lsb (- numeral bin :: 'a :: len word) = (bin_last (- numeral bin) = 1)"
+  "lsb (- numeral bin :: 'a :: len word) \<longleftrightarrow> bin_last (- numeral bin)"
   unfolding word_lsb_alt test_bit_neg_numeral by simp
 
 lemma set_bit_word_of_int:
-  "set_bit (word_of_int x) n b = word_of_int (bin_sc n (if b then 1 else 0) x)"
+  "set_bit (word_of_int x) n b = word_of_int (bin_sc n b x)"
   unfolding word_set_bit_def
   apply (rule word_eqI)
   apply (simp add: word_size bin_nth_sc_gen word_ubin.eq_norm nth_bintr)
@@ -2540,28 +2583,28 @@
 
 lemma word_set_numeral [simp]:
   "set_bit (numeral bin::'a::len0 word) n b = 
-    word_of_int (bin_sc n (if b then 1 else 0) (numeral bin))"
+    word_of_int (bin_sc n b (numeral bin))"
   unfolding word_numeral_alt by (rule set_bit_word_of_int)
 
 lemma word_set_neg_numeral [simp]:
   "set_bit (- numeral bin::'a::len0 word) n b = 
-    word_of_int (bin_sc n (if b then 1 else 0) (- numeral bin))"
+    word_of_int (bin_sc n b (- numeral bin))"
   unfolding word_neg_numeral_alt by (rule set_bit_word_of_int)
 
 lemma word_set_bit_0 [simp]:
-  "set_bit 0 n b = word_of_int (bin_sc n (if b then 1 else 0) 0)"
+  "set_bit 0 n b = word_of_int (bin_sc n b 0)"
   unfolding word_0_wi by (rule set_bit_word_of_int)
 
 lemma word_set_bit_1 [simp]:
-  "set_bit 1 n b = word_of_int (bin_sc n (if b then 1 else 0) 1)"
+  "set_bit 1 n b = word_of_int (bin_sc n b 1)"
   unfolding word_1_wi by (rule set_bit_word_of_int)
 
 lemma setBit_no [simp]:
-  "setBit (numeral bin) n = word_of_int (bin_sc n 1 (numeral bin))"
+  "setBit (numeral bin) n = word_of_int (bin_sc n True (numeral bin))"
   by (simp add: setBit_def)
 
 lemma clearBit_no [simp]:
-  "clearBit (numeral bin) n = word_of_int (bin_sc n 0 (numeral bin))"
+  "clearBit (numeral bin) n = word_of_int (bin_sc n False (numeral bin))"
   by (simp add: clearBit_def)
 
 lemma to_bl_n1: 
@@ -2645,7 +2688,6 @@
   fixes w :: "'a::len0 word"
   shows "w >= set_bit w n False"
   apply (unfold word_set_bit_def word_le_def word_ubin.eq_norm)
-  apply simp
   apply (rule order_trans)
    apply (rule bintr_bin_clr_le)
   apply simp
@@ -2655,7 +2697,6 @@
   fixes w :: "'a::len word"
   shows "w <= set_bit w n True"
   apply (unfold word_set_bit_def word_le_def word_ubin.eq_norm)
-  apply simp
   apply (rule order_trans [OF _ bintr_bin_set_ge])
   apply simp
   done
@@ -2663,7 +2704,7 @@
 
 subsection {* Shifting, Rotating, and Splitting Words *}
 
-lemma shiftl1_wi [simp]: "shiftl1 (word_of_int w) = word_of_int (w BIT 0)"
+lemma shiftl1_wi [simp]: "shiftl1 (word_of_int w) = word_of_int (w BIT False)"
   unfolding shiftl1_def
   apply (simp add: word_ubin.norm_eq_iff [symmetric] word_ubin.eq_norm)
   apply (subst refl [THEN bintrunc_BIT_I, symmetric])
@@ -2682,10 +2723,10 @@
 lemma shiftl1_0 [simp] : "shiftl1 0 = 0"
   unfolding shiftl1_def by simp
 
-lemma shiftl1_def_u: "shiftl1 w = word_of_int (uint w BIT 0)"
+lemma shiftl1_def_u: "shiftl1 w = word_of_int (uint w BIT False)"
   by (simp only: shiftl1_def) (* FIXME: duplicate *)
 
-lemma shiftl1_def_s: "shiftl1 w = word_of_int (sint w BIT 0)"
+lemma shiftl1_def_s: "shiftl1 w = word_of_int (sint w BIT False)"
   unfolding shiftl1_def Bit_B0 wi_hom_syms by simp
 
 lemma shiftr1_0 [simp]: "shiftr1 0 = 0"
@@ -3969,9 +4010,6 @@
 
 subsubsection "map, map2, commuting with rotate(r)"
 
-lemma last_map: "xs ~= [] \<Longrightarrow> last (map f xs) = f (last xs)"
-  by (induct xs) auto
-
 lemma butlast_map:
   "xs ~= [] \<Longrightarrow> butlast (map f xs) = map f (butlast xs)"
   by (induct xs) auto
@@ -4538,7 +4576,8 @@
 
 subsection {* Recursion combinator for words *}
 
-definition word_rec :: "'a \<Rightarrow> ('b::len word \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b word \<Rightarrow> 'a" where
+definition word_rec :: "'a \<Rightarrow> ('b::len word \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b word \<Rightarrow> 'a"
+where
   "word_rec forZero forSuc n = nat_rec forZero (forSuc \<circ> of_nat) (unat n)"
 
 lemma word_rec_0: "word_rec z s 0 = z"
--- a/src/HOL/Word/WordBitwise.thy	Mon Dec 23 15:30:31 2013 +0100
+++ b/src/HOL/Word/WordBitwise.thy	Mon Dec 23 16:59:56 2013 +0100
@@ -364,8 +364,7 @@
             = (bl_to_bin (rev xs) < bl_to_bin (rev ys))"
   apply (induct xs ys rule: list_induct2)
    apply (simp_all add: rev_bl_order_simps bl_to_bin_app_cat)
-  apply (simp add: bl_to_bin_def Bit_B0 Bit_B1 add1_zle_eq)
-  apply arith?
+  apply (auto simp add: bl_to_bin_def Bit_B0 Bit_B1 add1_zle_eq Bit_def)
   done
 
 lemma word_le_rbl: