merged
authorwenzelm
Sun, 18 Aug 2013 17:00:10 +0200
changeset 53075 98fc47533342
parent 53070 6a3410845bb2 (diff)
parent 53074 e62c7a4b6697 (current diff)
child 53076 47c9aff07725
merged
--- a/src/HOL/Code_Numeral.thy	Sun Aug 18 15:59:48 2013 +0200
+++ b/src/HOL/Code_Numeral.thy	Sun Aug 18 17:00:10 2013 +0200
@@ -223,6 +223,21 @@
   "of_nat (nat_of_integer k) = max 0 k"
   by transfer auto
 
+instance integer :: semiring_numeral_div
+  by intro_classes (transfer,
+    fact semiring_numeral_div_class.diff_invert_add1
+    semiring_numeral_div_class.le_add_diff_inverse2
+    semiring_numeral_div_class.mult_div_cancel
+    semiring_numeral_div_class.div_less
+    semiring_numeral_div_class.mod_less
+    semiring_numeral_div_class.div_positive
+    semiring_numeral_div_class.mod_less_eq_dividend
+    semiring_numeral_div_class.pos_mod_bound
+    semiring_numeral_div_class.pos_mod_sign
+    semiring_numeral_div_class.mod_mult2_eq
+    semiring_numeral_div_class.div_mult2_eq
+    semiring_numeral_div_class.discrete)+
+
 
 subsection {* Code theorems for target language integers *}
 
@@ -347,24 +362,15 @@
   "snd (divmod_abs k l) = \<bar>k\<bar> mod \<bar>l\<bar>"
   by (simp add: divmod_abs_def)
 
-lemma divmod_abs_terminate_code [code]:
-  "divmod_abs (Neg k) (Neg l) = divmod_abs (Pos k) (Pos l)"
-  "divmod_abs (Neg k) (Pos l) = divmod_abs (Pos k) (Pos l)"
-  "divmod_abs (Pos k) (Neg l) = divmod_abs (Pos k) (Pos l)"
+lemma divmod_abs_code [code]:
+  "divmod_abs (Pos k) (Pos l) = divmod k l"
+  "divmod_abs (Neg k) (Neg l) = divmod k l"
+  "divmod_abs (Neg k) (Pos l) = divmod k l"
+  "divmod_abs (Pos k) (Neg l) = divmod k l"
   "divmod_abs j 0 = (0, \<bar>j\<bar>)"
   "divmod_abs 0 j = (0, 0)"
   by (simp_all add: prod_eq_iff)
 
-lemma divmod_abs_rec_code [code]:
-  "divmod_abs (Pos k) (Pos l) =
-    (let j = sub k l in
-       if j < 0 then (0, Pos k)
-       else let (q, r) = divmod_abs j (Pos l) in (q + 1, r))"
-  apply (simp add: prod_eq_iff Let_def prod_case_beta)
-  apply transfer
-  apply (simp add: sub_non_negative sub_negative div_pos_pos_trivial mod_pos_pos_trivial div_pos_geq mod_pos_geq)
-  done
-
 lemma divmod_integer_code [code]:
   "divmod_integer k l =
     (if k = 0 then (0, 0) else if l = 0 then (0, k) else
--- a/src/HOL/Divides.thy	Sun Aug 18 15:59:48 2013 +0200
+++ b/src/HOL/Divides.thy	Sun Aug 18 17:00:10 2013 +0200
@@ -480,6 +480,205 @@
 end
 
 
+subsection {* Generic numeral division with a pragmatic type class *}
+
+text {*
+  The following type class contains everything necessary to formulate
+  a division algorithm in ring structures with numerals, restricted
+  to its positive segments.  This is its primary motiviation, and it
+  could surely be formulated using a more fine-grained, more algebraic
+  and less technical class hierarchy.
+*}
+
+
+class semiring_numeral_div = linordered_semidom + minus + semiring_div +
+  assumes diff_invert_add1: "a + b = c \<Longrightarrow> a = c - b"
+    and le_add_diff_inverse2: "b \<le> a \<Longrightarrow> a - b + b = a"
+  assumes mult_div_cancel: "b * (a div b) = a - a mod b"
+    and div_less: "0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a div b = 0"
+    and mod_less: " 0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a mod b = a"
+    and div_positive: "0 < b \<Longrightarrow> b \<le> a \<Longrightarrow> a div b > 0"
+    and mod_less_eq_dividend: "0 \<le> a \<Longrightarrow> a mod b \<le> a"
+    and pos_mod_bound: "0 < b \<Longrightarrow> a mod b < b"
+    and pos_mod_sign: "0 < b \<Longrightarrow> 0 \<le> a mod b"
+    and mod_mult2_eq: "0 \<le> c \<Longrightarrow> a mod (b * c) = b * (a div b mod c) + a mod b"
+    and div_mult2_eq: "0 \<le> c \<Longrightarrow> a div (b * c) = a div b div c"
+  assumes discrete: "a < b \<longleftrightarrow> a + 1 \<le> b"
+begin
+
+lemma diff_zero [simp]:
+  "a - 0 = a"
+  by (rule diff_invert_add1 [symmetric]) simp
+
+lemma parity:
+  "a mod 2 = 0 \<or> a mod 2 = 1"
+proof (rule ccontr)
+  assume "\<not> (a mod 2 = 0 \<or> a mod 2 = 1)"
+  then have "a mod 2 \<noteq> 0" and "a mod 2 \<noteq> 1" by simp_all
+  have "0 < 2" by simp
+  with pos_mod_bound pos_mod_sign have "0 \<le> a mod 2" "a mod 2 < 2" by simp_all
+  with `a mod 2 \<noteq> 0` have "0 < a mod 2" by simp
+  with discrete have "1 \<le> a mod 2" by simp
+  with `a mod 2 \<noteq> 1` have "1 < a mod 2" by simp
+  with discrete have "2 \<le> a mod 2" by simp
+  with `a mod 2 < 2` show False by simp
+qed
+
+lemma divmod_digit_1:
+  assumes "0 \<le> a" "0 < b" and "b \<le> a mod (2 * b)"
+  shows "2 * (a div (2 * b)) + 1 = a div b" (is "?P")
+    and "a mod (2 * b) - b = a mod b" (is "?Q")
+proof -
+  from assms mod_less_eq_dividend [of a "2 * b"] have "b \<le> a"
+    by (auto intro: trans)
+  with `0 < b` have "0 < a div b" by (auto intro: div_positive)
+  then have [simp]: "1 \<le> a div b" by (simp add: discrete)
+  with `0 < b` have mod_less: "a mod b < b" by (simp add: pos_mod_bound)
+  def w \<equiv> "a div b mod 2" with parity have w_exhaust: "w = 0 \<or> w = 1" by auto
+  have mod_w: "a mod (2 * b) = a mod b + b * w"
+    by (simp add: w_def mod_mult2_eq ac_simps)
+  from assms w_exhaust have "w = 1"
+    by (auto simp add: mod_w) (insert mod_less, auto)
+  with mod_w have mod: "a mod (2 * b) = a mod b + b" by simp
+  have "2 * (a div (2 * b)) = a div b - w"
+    by (simp add: w_def div_mult2_eq mult_div_cancel ac_simps)
+  with `w = 1` have div: "2 * (a div (2 * b)) = a div b - 1" by simp
+  then show ?P and ?Q
+    by (simp_all add: div mod diff_invert_add1 [symmetric] le_add_diff_inverse2)
+qed
+
+lemma divmod_digit_0:
+  assumes "0 < b" and "a mod (2 * b) < b"
+  shows "2 * (a div (2 * b)) = a div b" (is "?P")
+    and "a mod (2 * b) = a mod b" (is "?Q")
+proof -
+  def w \<equiv> "a div b mod 2" with parity have w_exhaust: "w = 0 \<or> w = 1" by auto
+  have mod_w: "a mod (2 * b) = a mod b + b * w"
+    by (simp add: w_def mod_mult2_eq ac_simps)
+  moreover have "b \<le> a mod b + b"
+  proof -
+    from `0 < b` pos_mod_sign have "0 \<le> a mod b" by blast
+    then have "0 + b \<le> a mod b + b" by (rule add_right_mono)
+    then show ?thesis by simp
+  qed
+  moreover note assms w_exhaust
+  ultimately have "w = 0" by auto
+  find_theorems "_ + ?b < _ + ?b"
+  with mod_w have mod: "a mod (2 * b) = a mod b" by simp
+  have "2 * (a div (2 * b)) = a div b - w"
+    by (simp add: w_def div_mult2_eq mult_div_cancel ac_simps)
+  with `w = 0` have div: "2 * (a div (2 * b)) = a div b" by simp
+  then show ?P and ?Q
+    by (simp_all add: div mod)
+qed
+
+definition divmod :: "num \<Rightarrow> num \<Rightarrow> 'a \<times> 'a"
+where
+  "divmod m n = (numeral m div numeral n, numeral m mod numeral n)"
+
+lemma fst_divmod [simp]:
+  "fst (divmod m n) = numeral m div numeral n"
+  by (simp add: divmod_def)
+
+lemma snd_divmod [simp]:
+  "snd (divmod m n) = numeral m mod numeral n"
+  by (simp add: divmod_def)
+
+definition divmod_step :: "num \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<times> 'a"
+where
+  "divmod_step l qr = (let (q, r) = qr
+    in if r \<ge> numeral l then (2 * q + 1, r - numeral l)
+    else (2 * q, r))"
+
+text {*
+  This is a formulation of one step (referring to one digit position)
+  in school-method division: compare the dividend at the current
+  digit position with the remainder from previous division steps
+  and evaluate accordingly.
+*}
+
+lemma divmod_step_eq [code]:
+  "divmod_step l (q, r) = (if numeral l \<le> r
+    then (2 * q + 1, r - numeral l) else (2 * q, r))"
+  by (simp add: divmod_step_def)
+
+lemma divmod_step_simps [simp]:
+  "r < numeral l \<Longrightarrow> divmod_step l (q, r) = (2 * q, r)"
+  "numeral l \<le> r \<Longrightarrow> divmod_step l (q, r) = (2 * q + 1, r - numeral l)"
+  by (auto simp add: divmod_step_eq not_le)
+
+text {*
+  This is a formulation of school-method division.
+  If the divisor is smaller than the dividend, terminate.
+  If not, shift the dividend to the right until termination
+  occurs and then reiterate single division steps in the
+  opposite direction.
+*}
+
+lemma divmod_divmod_step [code]:
+  "divmod m n = (if m < n then (0, numeral m)
+    else divmod_step n (divmod m (Num.Bit0 n)))"
+proof (cases "m < n")
+  case True then have "numeral m < numeral n" by simp
+  then show ?thesis
+    by (simp add: prod_eq_iff div_less mod_less)
+next
+  case False
+  have "divmod m n =
+    divmod_step n (numeral m div (2 * numeral n),
+      numeral m mod (2 * numeral n))"
+  proof (cases "numeral n \<le> numeral m mod (2 * numeral n)")
+    case True
+    with divmod_step_simps
+      have "divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) =
+        (2 * (numeral m div (2 * numeral n)) + 1, numeral m mod (2 * numeral n) - numeral n)"
+        by blast
+    moreover from True divmod_digit_1 [of "numeral m" "numeral n"]
+      have "2 * (numeral m div (2 * numeral n)) + 1 = numeral m div numeral n"
+      and "numeral m mod (2 * numeral n) - numeral n = numeral m mod numeral n"
+      by simp_all
+    ultimately show ?thesis by (simp only: divmod_def)
+  next
+    case False then have *: "numeral m mod (2 * numeral n) < numeral n"
+      by (simp add: not_le)
+    with divmod_step_simps
+      have "divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) =
+        (2 * (numeral m div (2 * numeral n)), numeral m mod (2 * numeral n))"
+        by blast
+    moreover from * divmod_digit_0 [of "numeral n" "numeral m"]
+      have "2 * (numeral m div (2 * numeral n)) = numeral m div numeral n"
+      and "numeral m mod (2 * numeral n) = numeral m mod numeral n"
+      by (simp_all only: zero_less_numeral)
+    ultimately show ?thesis by (simp only: divmod_def)
+  qed
+  then have "divmod m n =
+    divmod_step n (numeral m div numeral (Num.Bit0 n),
+      numeral m mod numeral (Num.Bit0 n))"
+    by (simp only: numeral.simps distrib mult_1) 
+  then have "divmod m n = divmod_step n (divmod m (Num.Bit0 n))"
+    by (simp add: divmod_def)
+  with False show ?thesis by simp
+qed
+
+lemma divmod_cancel [code]:
+  "divmod (Num.Bit0 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r))" (is ?P)
+  "divmod (Num.Bit1 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r + 1))" (is ?Q)
+proof -
+  have *: "\<And>q. numeral (Num.Bit0 q) = 2 * numeral q"
+    "\<And>q. numeral (Num.Bit1 q) = 2 * numeral q + 1"
+    by (simp_all only: numeral_mult numeral.simps distrib) simp_all
+  have "1 div 2 = 0" "1 mod 2 = 1" by (auto intro: div_less mod_less)
+  then show ?P and ?Q
+    by (simp_all add: prod_eq_iff split_def * [of m] * [of n] mod_mult_mult1
+      div_mult2_eq [of _ _ 2] mod_mult2_eq [of _ _ 2] add.commute del: numeral_times_numeral)
+ qed
+
+end
+
+hide_fact (open) diff_invert_add1 le_add_diff_inverse2 diff_zero
+  -- {* restore simple accesses for more general variants of theorems *}
+
+  
 subsection {* Division on @{typ nat} *}
 
 text {*
@@ -733,6 +932,17 @@
 lemma div_mult_self1_is_m [simp]: "0<n ==> (n*m) div n = (m::nat)"
 by simp
 
+lemma div_positive:
+  fixes m n :: nat
+  assumes "n > 0"
+  assumes "m \<ge> n"
+  shows "m div n > 0"
+proof -
+  from `m \<ge> n` obtain q where "m = n + q"
+    by (auto simp add: le_iff_add)
+  with `n > 0` show ?thesis by simp
+qed
+
 
 subsubsection {* Remainder *}
 
@@ -1180,6 +1390,9 @@
   shows "n mod 2 \<noteq> 0 \<longleftrightarrow> n mod 2 = 1"
   by simp
 
+instance nat :: semiring_numeral_div
+  by intro_classes (auto intro: div_positive simp add: mult_div_cancel mod_mult2_eq div_mult2_eq)
+
 
 subsection {* Division on @{typ int} *}
 
@@ -1188,6 +1401,14 @@
   "divmod_int_rel a b = (\<lambda>(q, r). a = b * q + r \<and>
     (if 0 < b then 0 \<le> r \<and> r < b else if b < 0 then b < r \<and> r \<le> 0 else q = 0))"
 
+text {*
+  The following algorithmic devlopment actually echos what has already
+  been developed in class @{class semiring_numeral_div}.  In the long
+  run it seems better to derive division on @{typ int} just from
+  division on @{typ nat} and instantiate @{class semiring_numeral_div}
+  accordingly.
+*}
+
 definition adjust :: "int \<Rightarrow> int \<times> int \<Rightarrow> int \<times> int" where
     --{*for the division algorithm*}
     "adjust b = (\<lambda>(q, r). if 0 \<le> r - b then (2 * q + 1, r - b)
@@ -1638,7 +1859,11 @@
   by (rule mod_int_unique [of a b q r],
     simp add: divmod_int_rel_def)
 
-(* simprocs adapted from HOL/ex/Binary.thy *)
+text {*
+  numeral simprocs -- high chance that these can be replaced
+  by divmod algorithm from @{class semiring_numeral_div}
+*}
+
 ML {*
 local
   val mk_number = HOLogic.mk_number HOLogic.intT
@@ -1914,15 +2139,18 @@
                    zero_less_mult_iff distrib_left [symmetric] 
                    zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4 mult_less_0_iff split: split_if_asm)
 
-lemma zdiv_zmult2_eq: "(0::int) < c ==> a div (b*c) = (a div b) div c"
+lemma zdiv_zmult2_eq:
+  fixes a b c :: int
+  shows "0 \<le> c \<Longrightarrow> a div (b * c) = (a div b) div c"
 apply (case_tac "b = 0", simp)
-apply (force simp add: divmod_int_rel_div_mod [THEN zmult2_lemma, THEN div_int_unique])
+apply (force simp add: le_less divmod_int_rel_div_mod [THEN zmult2_lemma, THEN div_int_unique])
 done
 
 lemma zmod_zmult2_eq:
-     "(0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b"
+  fixes a b c :: int
+  shows "0 \<le> c \<Longrightarrow> a mod (b * c) = b * (a div b mod c) + a mod b"
 apply (case_tac "b = 0", simp)
-apply (force simp add: divmod_int_rel_div_mod [THEN zmult2_lemma, THEN mod_int_unique])
+apply (force simp add: le_less divmod_int_rel_div_mod [THEN zmult2_lemma, THEN mod_int_unique])
 done
 
 lemma div_pos_geq:
@@ -2319,6 +2547,12 @@
   thus  ?lhs by simp
 qed
 
+text {*
+  This re-embedding of natural division on integers goes back to the
+  time when numerals had been signed numerals.  It should
+  now be replaced by the algorithm developed in @{class semiring_numeral_div}.  
+*}
+
 lemma div_nat_numeral [simp]:
   "(numeral v :: nat) div numeral v' = nat (numeral v div numeral v')"
   by (simp add: nat_div_distrib)
@@ -2340,6 +2574,12 @@
   shows "k mod 2 \<noteq> 0 \<longleftrightarrow> k mod 2 = 1"
   by auto
 
+instance int :: semiring_numeral_div
+  by intro_classes (auto intro: zmod_le_nonneg_dividend
+    simp add: zmult_div_cancel
+    pos_imp_zdiv_pos_iff div_pos_pos_trivial mod_pos_pos_trivial
+    zmod_zmult2_eq zdiv_zmult2_eq)
+
 
 subsubsection {* Tools setup *}
 
@@ -2350,37 +2590,55 @@
 
 subsubsection {* Code generation *}
 
-definition pdivmod :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
-  "pdivmod k l = (\<bar>k\<bar> div \<bar>l\<bar>, \<bar>k\<bar> mod \<bar>l\<bar>)"
-
-lemma pdivmod_posDivAlg [code]:
-  "pdivmod k l = (if l = 0 then (0, \<bar>k\<bar>) else posDivAlg \<bar>k\<bar> \<bar>l\<bar>)"
-by (subst posDivAlg_div_mod) (simp_all add: pdivmod_def)
-
-lemma divmod_int_pdivmod: "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else
+definition divmod_abs :: "int \<Rightarrow> int \<Rightarrow> int \<times> int"
+where
+  "divmod_abs k l = (\<bar>k\<bar> div \<bar>l\<bar>, \<bar>k\<bar> mod \<bar>l\<bar>)"
+
+lemma fst_divmod_abs [simp]:
+  "fst (divmod_abs k l) = \<bar>k\<bar> div \<bar>l\<bar>"
+  by (simp add: divmod_abs_def)
+
+lemma snd_divmod_abs [simp]:
+  "snd (divmod_abs k l) = \<bar>k\<bar> mod \<bar>l\<bar>"
+  by (simp add: divmod_abs_def)
+
+lemma divmod_abs_code [code]:
+  "divmod_abs (Int.Pos k) (Int.Pos l) = divmod k l"
+  "divmod_abs (Int.Neg k) (Int.Neg l) = divmod k l"
+  "divmod_abs (Int.Neg k) (Int.Pos l) = divmod k l"
+  "divmod_abs (Int.Pos k) (Int.Neg l) = divmod k l"
+  "divmod_abs j 0 = (0, \<bar>j\<bar>)"
+  "divmod_abs 0 j = (0, 0)"
+  by (simp_all add: prod_eq_iff)
+
+lemma divmod_int_divmod_abs:
+  "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else
   apsnd ((op *) (sgn l)) (if 0 < l \<and> 0 \<le> k \<or> l < 0 \<and> k < 0
-    then pdivmod k l
-    else (let (r, s) = pdivmod k l in
+    then divmod_abs k l
+    else (let (r, s) = divmod_abs k l in
        if s = 0 then (- r, 0) else (- r - 1, \<bar>l\<bar> - s))))"
 proof -
   have aux: "\<And>q::int. - k = l * q \<longleftrightarrow> k = l * - q" by auto
   show ?thesis
-    by (simp add: divmod_int_mod_div pdivmod_def)
+    by (simp add: prod_eq_iff split_def Let_def)
       (auto simp add: aux not_less not_le zdiv_zminus1_eq_if
       zmod_zminus1_eq_if zdiv_zminus2_eq_if zmod_zminus2_eq_if)
 qed
 
-lemma divmod_int_code [code]: "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else
+lemma divmod_int_code [code]:
+  "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else
   apsnd ((op *) (sgn l)) (if sgn k = sgn l
-    then pdivmod k l
-    else (let (r, s) = pdivmod k l in
+    then divmod_abs k l
+    else (let (r, s) = divmod_abs k l in
       if s = 0 then (- r, 0) else (- r - 1, \<bar>l\<bar> - s))))"
 proof -
   have "k \<noteq> 0 \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> 0 < l \<and> 0 \<le> k \<or> l < 0 \<and> k < 0 \<longleftrightarrow> sgn k = sgn l"
     by (auto simp add: not_less sgn_if)
-  then show ?thesis by (simp add: divmod_int_pdivmod)
+  then show ?thesis by (simp add: divmod_int_divmod_abs)
 qed
 
+hide_const (open) divmod_abs
+
 code_identifier
   code_module Divides \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
 
--- a/src/HOL/Int.thy	Sun Aug 18 15:59:48 2013 +0200
+++ b/src/HOL/Int.thy	Sun Aug 18 17:00:10 2013 +0200
@@ -384,6 +384,10 @@
 lemma nat_zminus_int [simp]: "nat (- int n) = 0"
   by transfer simp
 
+lemma le_nat_iff:
+  "k \<ge> 0 \<Longrightarrow> n \<le> nat k \<longleftrightarrow> int n \<le> k"
+  by transfer auto
+  
 lemma zless_nat_eq_int_zless: "(m < nat z) = (int m < z)"
   by transfer (clarsimp simp add: less_diff_conv)
 
--- a/src/HOL/Library/Bit.thy	Sun Aug 18 15:59:48 2013 +0200
+++ b/src/HOL/Library/Bit.thy	Sun Aug 18 17:00:10 2013 +0200
@@ -10,16 +10,18 @@
 
 subsection {* Bits as a datatype *}
 
-typedef bit = "UNIV :: bool set" ..
+typedef bit = "UNIV :: bool set"
+  morphisms set Bit
+  ..
 
 instantiation bit :: "{zero, one}"
 begin
 
 definition zero_bit_def:
-  "0 = Abs_bit False"
+  "0 = Bit False"
 
 definition one_bit_def:
-  "1 = Abs_bit True"
+  "1 = Bit True"
 
 instance ..
 
@@ -29,7 +31,7 @@
 proof -
   fix P and x :: bit
   assume "P (0::bit)" and "P (1::bit)"
-  then have "\<forall>b. P (Abs_bit b)"
+  then have "\<forall>b. P (Bit b)"
     unfolding zero_bit_def one_bit_def
     by (simp add: all_bool_eq)
   then show "P x"
@@ -37,16 +39,63 @@
 next
   show "(0::bit) \<noteq> (1::bit)"
     unfolding zero_bit_def one_bit_def
-    by (simp add: Abs_bit_inject)
+    by (simp add: Bit_inject)
 qed
 
-lemma bit_not_0_iff [iff]: "(x::bit) \<noteq> 0 \<longleftrightarrow> x = 1"
-  by (induct x) simp_all
+lemma Bit_set_eq [simp]:
+  "Bit (set b) = b"
+  by (fact set_inverse)
+  
+lemma set_Bit_eq [simp]:
+  "set (Bit P) = P"
+  by (rule Bit_inverse) rule
+
+lemma bit_eq_iff:
+  "x = y \<longleftrightarrow> (set x \<longleftrightarrow> set y)"
+  by (auto simp add: set_inject)
+
+lemma Bit_inject [simp]:
+  "Bit P = Bit Q \<longleftrightarrow> (P \<longleftrightarrow> Q)"
+  by (auto simp add: Bit_inject)  
+
+lemma set [iff]:
+  "\<not> set 0"
+  "set 1"
+  by (simp_all add: zero_bit_def one_bit_def Bit_inverse)
+
+lemma [code]:
+  "set 0 \<longleftrightarrow> False"
+  "set 1 \<longleftrightarrow> True"
+  by simp_all
 
-lemma bit_not_1_iff [iff]: "(x::bit) \<noteq> 1 \<longleftrightarrow> x = 0"
-  by (induct x) simp_all
+lemma set_iff:
+  "set b \<longleftrightarrow> b = 1"
+  by (cases b) simp_all
+
+lemma bit_eq_iff_set:
+  "b = 0 \<longleftrightarrow> \<not> set b"
+  "b = 1 \<longleftrightarrow> set b"
+  by (simp_all add: bit_eq_iff)
+
+lemma Bit [simp, code]:
+  "Bit False = 0"
+  "Bit True = 1"
+  by (simp_all add: zero_bit_def one_bit_def)
 
+lemma bit_not_0_iff [iff]:
+  "(x::bit) \<noteq> 0 \<longleftrightarrow> x = 1"
+  by (simp add: bit_eq_iff)
 
+lemma bit_not_1_iff [iff]:
+  "(x::bit) \<noteq> 1 \<longleftrightarrow> x = 0"
+  by (simp add: bit_eq_iff)
+
+lemma [code]:
+  "HOL.equal 0 b \<longleftrightarrow> \<not> set b"
+  "HOL.equal 1 b \<longleftrightarrow> set b"
+  by (simp_all add: equal set_iff)  
+
+  
 subsection {* Type @{typ bit} forms a field *}
 
 instantiation bit :: field_inverse_zero
@@ -110,4 +159,46 @@
 lemma bit_numeral_odd [simp]: "numeral (Num.Bit1 w) = (1 :: bit)"
   by (simp only: numeral_Bit1 bit_add_self add_0_left)
 
+
+subsection {* Conversion from @{typ bit} *}
+
+context zero_neq_one
+begin
+
+definition of_bit :: "bit \<Rightarrow> 'a"
+where
+  "of_bit b = bit_case 0 1 b" 
+
+lemma of_bit_eq [simp, code]:
+  "of_bit 0 = 0"
+  "of_bit 1 = 1"
+  by (simp_all add: of_bit_def)
+
+lemma of_bit_eq_iff:
+  "of_bit x = of_bit y \<longleftrightarrow> x = y"
+  by (cases x) (cases y, simp_all)+
+
+end  
+
+context semiring_1
+begin
+
+lemma of_nat_of_bit_eq:
+  "of_nat (of_bit b) = of_bit b"
+  by (cases b) simp_all
+
 end
+
+context ring_1
+begin
+
+lemma of_int_of_bit_eq:
+  "of_int (of_bit b) = of_bit b"
+  by (cases b) simp_all
+
+end
+
+hide_const (open) set
+
+end
+
--- a/src/HOL/Library/Code_Binary_Nat.thy	Sun Aug 18 15:59:48 2013 +0200
+++ b/src/HOL/Library/Code_Binary_Nat.thy	Sun Aug 18 17:00:10 2013 +0200
@@ -130,6 +130,15 @@
   "nat_of_num k < nat_of_num l \<longleftrightarrow> k < l"
   by (simp_all add: nat_of_num_numeral)
 
+lemma [code, code del]:
+  "divmod_nat = divmod_nat" ..
+  
+lemma divmod_nat_code [code]:
+  "divmod_nat (nat_of_num k) (nat_of_num l) = divmod k l"
+  "divmod_nat m 0 = (0, m)"
+  "divmod_nat 0 n = (0, 0)"
+  by (simp_all add: prod_eq_iff nat_of_num_numeral del: div_nat_numeral mod_nat_numeral)
+
 
 subsection {* Conversions *}
 
--- a/src/HOL/Library/Code_Target_Int.thy	Sun Aug 18 15:59:48 2013 +0200
+++ b/src/HOL/Library/Code_Target_Int.thy	Sun Aug 18 17:00:10 2013 +0200
@@ -65,9 +65,9 @@
   by simp
 
 lemma [code]:
-  "pdivmod k l = map_pair int_of_integer int_of_integer
+  "Divides.divmod_abs k l = map_pair int_of_integer int_of_integer
     (Code_Numeral.divmod_abs (of_int k) (of_int l))"
-  by (simp add: prod_eq_iff pdivmod_def)
+  by (simp add: prod_eq_iff)
 
 lemma [code]:
   "k div l = int_of_integer (of_int k div of_int l)"
--- a/src/HOL/Num.thy	Sun Aug 18 15:59:48 2013 +0200
+++ b/src/HOL/Num.thy	Sun Aug 18 17:00:10 2013 +0200
@@ -529,6 +529,12 @@
 lemma numeral_times_numeral: "numeral m * numeral n = numeral (m * n)"
   by (rule numeral_mult [symmetric])
 
+lemma mult_2: "2 * z = z + z"
+  unfolding one_add_one [symmetric] distrib_right by simp
+
+lemma mult_2_right: "z * 2 = z + z"
+  unfolding one_add_one [symmetric] distrib_left by simp
+
 end
 
 subsubsection {*
@@ -544,12 +550,6 @@
   by (induct n,
     simp_all only: numeral.simps numeral_class.numeral.simps of_nat_add of_nat_1)
 
-lemma mult_2: "2 * z = z + z"
-  unfolding one_add_one [symmetric] distrib_right by simp
-
-lemma mult_2_right: "z * 2 = z + z"
-  unfolding one_add_one [symmetric] distrib_left by simp
-
 end
 
 lemma nat_of_num_numeral [code_abbrev]:
--- a/src/HOL/Word/Bit_Int.thy	Sun Aug 18 15:59:48 2013 +0200
+++ b/src/HOL/Word/Bit_Int.thy	Sun Aug 18 17:00:10 2013 +0200
@@ -365,7 +365,7 @@
   done
 
 lemmas int_and_le =
-  xtr3 [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or]
+  xtrans(3) [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or]
 
 lemma add_BIT_simps [simp]: (* FIXME: move *)
   "x BIT 0 + y BIT 0 = (x + y) BIT 0"
@@ -553,7 +553,7 @@
     (ALL k. bin_nth b k = (k < n & bin_nth c k))"
   apply (induct n arbitrary: b c)
    apply clarsimp
-  apply (clarsimp simp: Let_def split: ls_splits)
+  apply (clarsimp simp: Let_def split: prod.split_asm)
   apply (case_tac k)
   apply auto
   done
@@ -589,11 +589,11 @@
 
 lemma split_bintrunc: 
   "bin_split n c = (a, b) ==> b = bintrunc n c"
-  by (induct n arbitrary: b c) (auto simp: Let_def split: ls_splits)
+  by (induct n arbitrary: b c) (auto simp: Let_def split: prod.split_asm)
 
 lemma bin_cat_split:
   "bin_split n w = (u, v) ==> w = bin_cat u n v"
-  by (induct n arbitrary: v w) (auto simp: Let_def split: ls_splits)
+  by (induct n arbitrary: v w) (auto simp: Let_def split: prod.split_asm)
 
 lemma bin_split_cat:
   "bin_split n (bin_cat v n w) = (v, bintrunc n w)"
@@ -610,18 +610,18 @@
   "bin_split (min m n) c = (a, b) ==> 
     bin_split n (bintrunc m c) = (bintrunc (m - n) a, b)"
   apply (induct n arbitrary: m b c, clarsimp)
-  apply (simp add: bin_rest_trunc Let_def split: ls_splits)
+  apply (simp add: bin_rest_trunc Let_def split: prod.split_asm)
   apply (case_tac m)
-   apply (auto simp: Let_def split: ls_splits)
+   apply (auto simp: Let_def split: prod.split_asm)
   done
 
 lemma bin_split_trunc1:
   "bin_split n c = (a, b) ==> 
     bin_split n (bintrunc m c) = (bintrunc (m - n) a, bintrunc m b)"
   apply (induct n arbitrary: m b c, clarsimp)
-  apply (simp add: bin_rest_trunc Let_def split: ls_splits)
+  apply (simp add: bin_rest_trunc Let_def split: prod.split_asm)
   apply (case_tac m)
-   apply (auto simp: Let_def split: ls_splits)
+   apply (auto simp: Let_def split: prod.split_asm)
   done
 
 lemma bin_cat_num:
@@ -658,3 +658,4 @@
   by auto
 
 end
+
--- a/src/HOL/Word/Bit_Operations.thy	Sun Aug 18 15:59:48 2013 +0200
+++ b/src/HOL/Word/Bit_Operations.thy	Sun Aug 18 17:00:10 2013 +0200
@@ -8,6 +8,8 @@
 imports "~~/src/HOL/Library/Bit"
 begin
 
+subsection {* Abstract syntactic bit operations *}
+
 class bit =
   fixes bitNOT :: "'a \<Rightarrow> 'a"       ("NOT _" [70] 71)
     and bitAND :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "AND" 64)
@@ -35,7 +37,7 @@
 class bitss = bits +
   fixes msb      :: "'a \<Rightarrow> bool"
 
-
+  
 subsection {* Bitwise operations on @{typ bit} *}
 
 instantiation bit :: bit
--- a/src/HOL/Word/Bit_Representation.thy	Sun Aug 18 15:59:48 2013 +0200
+++ b/src/HOL/Word/Bit_Representation.thy	Sun Aug 18 17:00:10 2013 +0200
@@ -1,16 +1,14 @@
 (* 
   Author: Jeremy Dawson, NICTA
-
-  Basic definitions to do with integers, expressed using Pls, Min, BIT.
 *) 
 
-header {* Basic Definitions for Binary Integers *}
+header {* Integers as implict bit strings *}
 
 theory Bit_Representation
-imports Misc_Numeric "~~/src/HOL/Library/Bit"
+imports "~~/src/HOL/Library/Bit" Misc_Numeric
 begin
 
-subsection {* Further properties of numerals *}
+subsection {* Constructors and destructors for binary integers *}
 
 definition bitval :: "bit \<Rightarrow> 'a\<Colon>zero_neq_one" where
   "bitval = bit_case 0 1"
@@ -23,6 +21,20 @@
 definition Bit :: "int \<Rightarrow> bit \<Rightarrow> int" (infixl "BIT" 90) where
   "k BIT b = bitval b + k + k"
 
+lemma Bit_B0:
+  "k BIT (0::bit) = k + k"
+   by (unfold Bit_def) simp
+
+lemma Bit_B1:
+  "k BIT (1::bit) = k + k + 1"
+   by (unfold Bit_def) simp
+  
+lemma Bit_B0_2t: "k BIT (0::bit) = 2 * k"
+  by (rule trans, rule Bit_B0) simp
+
+lemma Bit_B1_2t: "k BIT (1::bit) = 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))"
 
@@ -99,19 +111,28 @@
   "(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)
 
-lemma Bit_B0:
-  "k BIT (0::bit) = k + k"
-   by (unfold Bit_def) simp
+lemma pred_BIT_simps [simp]:
+  "x BIT 0 - 1 = (x - 1) BIT 1"
+  "x BIT 1 - 1 = x BIT 0"
+  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"
+  by (simp_all add: Bit_B0_2t Bit_B1_2t)
 
-lemma Bit_B1:
-  "k BIT (1::bit) = k + k + 1"
-   by (unfold Bit_def) simp
-  
-lemma Bit_B0_2t: "k BIT (0::bit) = 2 * k"
-  by (rule trans, rule Bit_B0) simp
+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"
+  by (simp_all add: Bit_B0_2t Bit_B1_2t)
 
-lemma Bit_B1_2t: "k BIT (1::bit) = 2 * k + 1"
-  by (rule trans, rule Bit_B1) simp
+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"
+  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"
@@ -140,9 +161,6 @@
   apply force
   done
 
-
-subsection {* Destructors for binary integers *}
-
 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"
@@ -575,7 +593,7 @@
   "sbintrunc (numeral k) 1 = 1"
   by (simp_all add: sbintrunc_numeral)
 
-lemma no_bintr_alt1: "bintrunc n = (%w. w mod 2 ^ n :: int)"
+lemma no_bintr_alt1: "bintrunc n = (\<lambda>w. w mod 2 ^ n :: int)"
   by (rule ext) (rule bintrunc_mod2p)
 
 lemma range_bintrunc: "range (bintrunc n) = {i. 0 <= i & i < 2 ^ n}"
@@ -674,7 +692,7 @@
   "(bin_rest (bintrunc n bin)) = bintrunc (n - 1) (bin_rest bin)"
   by (induct n arbitrary: bin) auto
 
-lemma bin_rest_power_trunc [rule_format] :
+lemma bin_rest_power_trunc:
   "(bin_rest ^^ k) (bintrunc n bin) = 
     bintrunc (n - k) ((bin_rest ^^ k) bin)"
   by (induct k) (auto simp: bin_rest_trunc)
@@ -720,17 +738,12 @@
   apply simp
   done
 
-lemma rco_alt: "(f o g) ^^ n o f = f o (g o f) ^^ n"
-  apply (rule ext)
-  apply (induct n)
-   apply (simp_all add: o_def)
-  done
-
 lemmas rco_bintr = bintrunc_rest' 
   [THEN rco_lem [THEN fun_cong], unfolded o_def]
 lemmas rco_sbintr = sbintrunc_rest' 
   [THEN rco_lem [THEN fun_cong], unfolded o_def]
 
+  
 subsection {* Splitting and concatenation *}
 
 primrec bin_split :: "nat \<Rightarrow> int \<Rightarrow> int \<times> int" where
@@ -747,59 +760,5 @@
   Z: "bin_cat w 0 v = w"
   | Suc: "bin_cat w (Suc n) v = bin_cat w n (bin_rest v) BIT bin_last v"
 
-subsection {* Miscellaneous lemmas *}
-
-lemma funpow_minus_simp:
-  "0 < n \<Longrightarrow> f ^^ n = f \<circ> f ^^ (n - 1)"
-  by (cases n) simp_all
-
-lemma funpow_numeral [simp]:
-  "f ^^ numeral k = f \<circ> f ^^ (pred_numeral k)"
-  by (simp add: numeral_eq_Suc)
-
-lemma replicate_numeral [simp]: (* TODO: move to List.thy *)
-  "replicate (numeral k) x = x # replicate (pred_numeral k) x"
-  by (simp add: numeral_eq_Suc)
-
-lemmas power_minus_simp = 
-  trans [OF gen_minus [where f = "power f"] power_Suc] for f
-
-lemma list_exhaust_size_gt0:
-  assumes y: "\<And>a list. y = a # list \<Longrightarrow> P"
-  shows "0 < length y \<Longrightarrow> P"
-  apply (cases y, simp)
-  apply (rule y)
-  apply fastforce
-  done
+end
 
-lemma list_exhaust_size_eq0:
-  assumes y: "y = [] \<Longrightarrow> P"
-  shows "length y = 0 \<Longrightarrow> P"
-  apply (cases y)
-   apply (rule y, simp)
-  apply simp
-  done
-
-lemma size_Cons_lem_eq:
-  "y = xa # list ==> size y = Suc k ==> size list = k"
-  by auto
-
-lemmas ls_splits = prod.split prod.split_asm split_if_asm
-
-lemma not_B1_is_B0: "y \<noteq> (1::bit) \<Longrightarrow> y = (0::bit)"
-  by (cases y) auto
-
-lemma B1_ass_B0: 
-  assumes y: "y = (0::bit) \<Longrightarrow> y = (1::bit)"
-  shows "y = (1::bit)"
-  apply (rule classical)
-  apply (drule not_B1_is_B0)
-  apply (erule y)
-  done
-
--- "simplifications for specific word lengths"
-lemmas n2s_ths [THEN eq_reflection] = add_2_eq_Suc add_2_eq_Suc'
-
-lemmas s2n_ths = n2s_ths [symmetric]
-
-end
--- a/src/HOL/Word/Bool_List_Representation.thy	Sun Aug 18 15:59:48 2013 +0200
+++ b/src/HOL/Word/Bool_List_Representation.thy	Sun Aug 18 17:00:10 2013 +0200
@@ -12,6 +12,23 @@
 imports Bit_Int
 begin
 
+definition map2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list"
+where
+  "map2 f as bs = map (split f) (zip as bs)"
+
+lemma map2_Nil [simp, code]:
+  "map2 f [] ys = []"
+  unfolding map2_def by auto
+
+lemma map2_Nil2 [simp, code]:
+  "map2 f xs [] = []"
+  unfolding map2_def by auto
+
+lemma map2_Cons [simp, code]:
+  "map2 f (x # xs) (y # ys) = f x y # map2 f xs ys"
+  unfolding map2_def by auto
+
+
 subsection {* Operations on lists of booleans *}
 
 primrec bl_to_bin_aux :: "bool list \<Rightarrow> int \<Rightarrow> int" where
@@ -40,19 +57,6 @@
       case xs of [] => fill # takefill fill n xs
         | y # ys => y # takefill fill n ys)"
 
-definition map2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list" where
-  "map2 f as bs = map (split f) (zip as bs)"
-
-lemma map2_Nil [simp]: "map2 f [] ys = []"
-  unfolding map2_def by auto
-
-lemma map2_Nil2 [simp]: "map2 f xs [] = []"
-  unfolding map2_def by auto
-
-lemma map2_Cons [simp]:
-  "map2 f (x # xs) (y # ys) = f x y # map2 f xs ys"
-  unfolding map2_def by auto
-
 
 subsection "Arithmetic in terms of bool lists"
 
@@ -314,12 +318,12 @@
    apply clarsimp
   apply clarsimp
   apply safe
-  apply (drule meta_spec, erule xtr8 [rotated], simp add: Bit_def)+
+  apply (drule meta_spec, erule xtrans(8) [rotated], simp add: Bit_def)+
   done
 
 lemma bl_to_bin_lt2p: "bl_to_bin bs < (2 ^ length bs)"
   apply (unfold bl_to_bin_def)
-  apply (rule xtr1)
+  apply (rule xtrans(1))
    prefer 2
    apply (rule bl_to_bin_lt2p_aux)
   apply simp
@@ -337,7 +341,7 @@
 
 lemma bl_to_bin_ge0: "bl_to_bin bs >= 0"
   apply (unfold bl_to_bin_def)
-  apply (rule xtr4)
+  apply (rule xtrans(4))
    apply (rule bl_to_bin_ge2p_aux)
   apply simp
   done
@@ -534,7 +538,7 @@
     bin_to_bl m a = take m (bin_to_bl (m + n) c)"
   apply (induct n arbitrary: b c)
    apply clarsimp
-  apply (clarsimp simp: Let_def split: ls_splits)
+  apply (clarsimp simp: Let_def split: prod.split_asm)
   apply (simp add: bin_to_bl_def)
   apply (simp add: take_bin2bl_lem)
   done
@@ -748,11 +752,6 @@
 lemmas rbl_Nils =
   rbl_pred.Nil rbl_succ.Nil rbl_add.Nil rbl_mult.Nil
 
-lemma pred_BIT_simps [simp]:
-  "x BIT 0 - 1 = (x - 1) BIT 1"
-  "x BIT 1 - 1 = x BIT 0"
-  by (simp_all add: Bit_B0_2t Bit_B1_2t)
-
 lemma rbl_pred:
   "rbl_pred (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin - 1))"
   apply (induct n arbitrary: bin, simp)
@@ -763,11 +762,6 @@
    apply (clarsimp simp: bin_to_bl_aux_alt)+
   done
 
-lemma succ_BIT_simps [simp]:
-  "x BIT 0 + 1 = x BIT 1"
-  "x BIT 1 + 1 = (x + 1) BIT 0"
-  by (simp_all add: Bit_B0_2t Bit_B1_2t)
-
 lemma rbl_succ: 
   "rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin + 1))"
   apply (induct n arbitrary: bin, simp)
@@ -778,13 +772,6 @@
    apply (clarsimp simp: bin_to_bl_aux_alt)+
   done
 
-lemma add_BIT_simps [simp]: (* FIXME: move *)
-  "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"
-  by (simp_all add: Bit_B0_2t Bit_B1_2t)
-
 lemma rbl_add: 
   "!!bina binb. rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = 
     rev (bin_to_bl n (bina + binb))"
@@ -870,12 +857,6 @@
   apply (simp add: bin_to_bl_aux_alt)
   done
 
-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"
-  by (simp_all add: Bit_B0_2t Bit_B1_2t algebra_simps)
-
 lemma rbl_mult: "!!bina binb. 
     rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = 
     rev (bin_to_bl n (bina * binb))"
@@ -907,56 +888,6 @@
     (y --> P (rbl_add ws xs)) & (~ y -->  P ws))"
   by (clarsimp simp add : Let_def)
   
-lemma and_len: "xs = ys ==> xs = ys & length xs = length ys"
-  by auto
-
-lemma size_if: "size (if p then xs else ys) = (if p then size xs else size ys)"
-  by auto
-
-lemma tl_if: "tl (if p then xs else ys) = (if p then tl xs else tl ys)"
-  by auto
-
-lemma hd_if: "hd (if p then xs else ys) = (if p then hd xs else hd ys)"
-  by auto
-
-lemma if_Not_x: "(if p then ~ x else x) = (p = (~ x))"
-  by auto
-
-lemma if_x_Not: "(if p then x else ~ x) = (p = x)"
-  by auto
-
-lemma if_same_and: "(If p x y & If p u v) = (if p then x & u else y & v)"
-  by auto
-
-lemma if_same_eq: "(If p x y  = (If p u v)) = (if p then x = (u) else y = (v))"
-  by auto
-
-lemma if_same_eq_not:
-  "(If p x y  = (~ If p u v)) = (if p then x = (~u) else y = (~v))"
-  by auto
-
-(* note - if_Cons can cause blowup in the size, if p is complex,
-  so make a simproc *)
-lemma if_Cons: "(if p then x # xs else y # ys) = If p x y # If p xs ys"
-  by auto
-
-lemma if_single:
-  "(if xc then [xab] else [an]) = [if xc then xab else an]"
-  by auto
-
-lemma if_bool_simps:
-  "If p True y = (p | y) & If p False y = (~p & y) & 
-    If p y True = (p --> y) & If p y False = (p & y)"
-  by auto
-
-lemmas if_simps = if_x_Not if_Not_x if_cancel if_True if_False if_bool_simps
-
-lemmas seqr = eq_reflection [where x = "size w"] for w (* FIXME: delete *)
-
-(* TODO: move name bindings to List.thy *)
-lemmas tl_Nil = tl.simps (1)
-lemmas tl_Cons = tl.simps (2)
-
 
 subsection "Repeated splitting or concatenation"
 
@@ -1008,7 +939,7 @@
   apply (induct n m c bs rule: bin_rsplit_aux.induct)
   apply (subst bin_rsplit_aux.simps)
   apply (subst bin_rsplit_aux.simps)
-  apply (clarsimp split: ls_splits)
+  apply (clarsimp split: prod.split)
   apply auto
   done
 
@@ -1017,7 +948,7 @@
   apply (induct n m c bs rule: bin_rsplitl_aux.induct)
   apply (subst bin_rsplitl_aux.simps)
   apply (subst bin_rsplitl_aux.simps)
-  apply (clarsimp split: ls_splits)
+  apply (clarsimp split: prod.split)
   apply auto
   done
 
@@ -1065,7 +996,7 @@
    apply clarsimp
   apply clarsimp
   apply (drule bthrs)
-  apply (simp (no_asm_use) add: Let_def split: ls_splits)
+  apply (simp (no_asm_use) add: Let_def split: prod.split_asm split_if_asm)
   apply clarify
   apply (drule split_bintrunc)
   apply simp
@@ -1081,7 +1012,7 @@
    apply clarsimp
   apply clarsimp
   apply (drule bthrs)
-  apply (simp (no_asm_use) add: Let_def split: ls_splits)
+  apply (simp (no_asm_use) add: Let_def split: prod.split_asm split_if_asm)
   apply clarify
   apply (erule allE, erule impE, erule exI)
   apply (case_tac k)
@@ -1097,7 +1028,7 @@
 lemma bin_rsplit_all:
   "0 < nw ==> nw <= n ==> bin_rsplit n (nw, w) = [bintrunc n w]"
   unfolding bin_rsplit_def
-  by (clarsimp dest!: split_bintrunc simp: rsplit_aux_simp2ls split: ls_splits)
+  by (clarsimp dest!: split_bintrunc simp: rsplit_aux_simp2ls split: prod.split)
 
 lemma bin_rsplit_l [rule_format] :
   "ALL bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, bintrunc m bin)"
@@ -1106,7 +1037,7 @@
   apply (rule allI)
   apply (subst bin_rsplitl_aux.simps)
   apply (subst bin_rsplit_aux.simps)
-  apply (clarsimp simp: Let_def split: ls_splits)
+  apply (clarsimp simp: Let_def split: prod.split)
   apply (drule bin_split_trunc)
   apply (drule sym [THEN trans], assumption)
   apply (subst rsplit_aux_alts(1))
@@ -1132,7 +1063,7 @@
     length ws \<le> m \<longleftrightarrow> nw + length bs * n \<le> m * n"
   apply (induct n nw w bs rule: bin_rsplit_aux.induct)
   apply (subst bin_rsplit_aux.simps)
-  apply (simp add: lrlem Let_def split: ls_splits)
+  apply (simp add: lrlem Let_def split: prod.split)
   done
 
 lemma bin_rsplit_len_le: 
@@ -1144,7 +1075,7 @@
     (nw + n - 1) div n + length cs"
   apply (induct n nw w cs rule: bin_rsplit_aux.induct)
   apply (subst bin_rsplit_aux.simps)
-  apply (clarsimp simp: Let_def split: ls_splits)
+  apply (clarsimp simp: Let_def split: prod.split)
   apply (erule thin_rl)
   apply (case_tac m)
   apply simp
@@ -1172,7 +1103,7 @@
     by auto
     show ?thesis using `length bs = length cs` `n \<noteq> 0`
       by (auto simp add: bin_rsplit_aux_simp_alt Let_def bin_rsplit_len
-        split: ls_splits)
+        split: prod.split)
   qed
 qed
 
--- a/src/HOL/Word/Misc_Numeric.thy	Sun Aug 18 15:59:48 2013 +0200
+++ b/src/HOL/Word/Misc_Numeric.thy	Sun Aug 18 17:00:10 2013 +0200
@@ -8,38 +8,26 @@
 imports Main Parity
 begin
 
-lemma the_elemI: "y = {x} ==> the_elem y = x" 
-  by simp
-
-lemma nonemptyE: "S ~= {} ==> (!!x. x : S ==> R) ==> R" by auto
-
-lemma gt_or_eq_0: "0 < y \<or> 0 = (y::nat)" by arith 
+lemma zmod_zsub_self [simp]: (* FIXME move to Divides.thy *) 
+  "((b :: int) - a) mod a = b mod a"
+  by (simp add: mod_diff_right_eq)
 
 declare iszero_0 [iff]
 
-lemmas xtr1 = xtrans(1)
-lemmas xtr2 = xtrans(2)
-lemmas xtr3 = xtrans(3)
-lemmas xtr4 = xtrans(4)
-lemmas xtr5 = xtrans(5)
-lemmas xtr6 = xtrans(6)
-lemmas xtr7 = xtrans(7)
-lemmas xtr8 = xtrans(8)
+lemma min_pm [simp]: "min a b + (a - b) = (a :: nat)" by arith
+  
+lemmas min_pm1 [simp] = trans [OF add_commute min_pm]
 
-lemmas nat_simps = diff_add_inverse2 diff_add_inverse
-lemmas nat_iffs = le_add1 le_add2
+lemma rev_min_pm [simp]: "min b a + (a - b) = (a::nat)" by arith
 
-lemma sum_imp_diff: "j = k + i ==> j - i = (k :: nat)" by arith
-
-lemma zless2: "0 < (2 :: int)" by arith
+lemmas rev_min_pm1 [simp] = trans [OF add_commute rev_min_pm]
 
-lemmas zless2p = zless2 [THEN zero_less_power]
-lemmas zle2p = zless2p [THEN order_less_imp_le]
+lemma min_minus [simp] : "min m (m - k) = (m - k :: nat)" by arith
+  
+lemmas min_minus' [simp] = trans [OF min_max.inf_commute min_minus]
 
-lemmas pos_mod_sign2 = zless2 [THEN pos_mod_sign [where b = "2::int"]]
-lemmas pos_mod_bound2 = zless2 [THEN pos_mod_bound [where b = "2::int"]]
-
-lemma nmod2: "n mod (2::int) = 0 | n mod 2 = 1" by arith
+lemma z1pmod2:
+  "(2 * b + 1) mod 2 = (1::int)" by arith
 
 lemma emep1:
   "even n ==> even d ==> 0 <= d ==> (n + 1) mod (d :: int) = (n mod d) + 1"
@@ -50,114 +38,6 @@
   apply (simp add: mod_mult_mult1)
  done
 
-lemmas eme1p = emep1 [simplified add_commute]
-
-lemma le_diff_eq': "(a \<le> c - b) = (b + a \<le> (c::int))" by arith
-
-lemma less_diff_eq': "(a < c - b) = (b + a < (c::int))" by arith
-
-lemma diff_le_eq': "(a - b \<le> c) = (a \<le> b + (c::int))" by arith
-
-lemma diff_less_eq': "(a - b < c) = (a < b + (c::int))" by arith
-
-lemmas m1mod2k = zless2p [THEN zmod_minus1]
-lemmas m1mod22k = mult_pos_pos [OF zless2 zless2p, THEN zmod_minus1]
-lemmas p1mod22k' = zless2p [THEN order_less_imp_le, THEN pos_zmod_mult_2]
-lemmas z1pmod2' = zero_le_one [THEN pos_zmod_mult_2, simplified]
-lemmas z1pdiv2' = zero_le_one [THEN pos_zdiv_mult_2, simplified]
-
-lemma p1mod22k:
-  "(2 * b + 1) mod (2 * 2 ^ n) = 2 * (b mod 2 ^ n) + (1::int)"
-  by (simp add: p1mod22k' add_commute)
-
-lemma z1pmod2:
-  "(2 * b + 1) mod 2 = (1::int)" by arith
-  
-lemma z1pdiv2:
-  "(2 * b + 1) div 2 = (b::int)" by arith
-
-lemmas zdiv_le_dividend = xtr3 [OF div_by_1 [symmetric] zdiv_mono2,
-  simplified int_one_le_iff_zero_less, simplified]
-  
-lemma axxbyy:
-  "a + m + m = b + n + n ==> (a = 0 | a = 1) ==> (b = 0 | b = 1) ==>  
-   a = b & m = (n :: int)" by arith
-
-lemma axxmod2:
-  "(1 + x + x) mod 2 = (1 :: int) & (0 + x + x) mod 2 = (0 :: int)" by arith
-
-lemma axxdiv2:
-  "(1 + x + x) div 2 = (x :: int) & (0 + x + x) div 2 = (x :: int)"  by arith
-
-lemmas iszero_minus = trans [THEN trans,
-  OF iszero_def neg_equal_0_iff_equal iszero_def [symmetric]]
-
-lemmas zadd_diff_inverse = trans [OF diff_add_cancel [symmetric] add_commute]
-
-lemmas add_diff_cancel2 = add_commute [THEN diff_eq_eq [THEN iffD2]]
-
-lemma zmod_zsub_self [simp]: 
-  "((b :: int) - a) mod a = b mod a"
-  by (simp add: mod_diff_right_eq)
-
-lemmas rdmods [symmetric] = mod_minus_eq
-  mod_diff_left_eq mod_diff_right_eq mod_add_left_eq
-  mod_add_right_eq mod_mult_right_eq mod_mult_left_eq
-
-lemma mod_plus_right:
-  "((a + x) mod m = (b + x) mod m) = (a mod m = b mod (m :: nat))"
-  apply (induct x)
-   apply (simp_all add: mod_Suc)
-  apply arith
-  done
-
-lemma nat_minus_mod: "(n - n mod m) mod m = (0 :: nat)"
-  by (induct n) (simp_all add : mod_Suc)
-
-lemmas nat_minus_mod_plus_right = trans [OF nat_minus_mod mod_0 [symmetric],
-  THEN mod_plus_right [THEN iffD2], simplified]
-
-lemmas push_mods' = mod_add_eq
-  mod_mult_eq mod_diff_eq
-  mod_minus_eq
-
-lemmas push_mods = push_mods' [THEN eq_reflection]
-lemmas pull_mods = push_mods [symmetric] rdmods [THEN eq_reflection]
-lemmas mod_simps = 
-  mod_mult_self2_is_0 [THEN eq_reflection]
-  mod_mult_self1_is_0 [THEN eq_reflection]
-  mod_mod_trivial [THEN eq_reflection]
-
-lemma nat_mod_eq:
-  "!!b. b < n ==> a mod n = b mod n ==> a mod n = (b :: nat)" 
-  by (induct a) auto
-
-lemmas nat_mod_eq' = refl [THEN [2] nat_mod_eq]
-
-lemma nat_mod_lem: 
-  "(0 :: nat) < n ==> b < n = (b mod n = b)"
-  apply safe
-   apply (erule nat_mod_eq')
-  apply (erule subst)
-  apply (erule mod_less_divisor)
-  done
-
-lemma mod_nat_add: 
-  "(x :: nat) < z ==> y < z ==> 
-   (x + y) mod z = (if x + y < z then x + y else x + y - z)"
-  apply (rule nat_mod_eq)
-   apply auto
-  apply (rule trans)
-   apply (rule le_mod_geq)
-   apply simp
-  apply (rule nat_mod_eq')
-  apply arith
-  done
-
-lemma mod_nat_sub: 
-  "(x :: nat) < z ==> (x - y) mod z = x - y"
-  by (rule nat_mod_eq') arith
-
 lemma int_mod_lem: 
   "(0 :: int) < n ==> (0 <= b & b < n) = (b mod n = b)"
   apply safe
@@ -166,18 +46,6 @@
    apply auto
   done
 
-lemma int_mod_eq:
-  "(0 :: int) <= b ==> b < n ==> a mod n = b mod n ==> a mod n = b"
-  by clarsimp (rule mod_pos_pos_trivial)
-
-lemmas int_mod_eq' = refl [THEN [3] int_mod_eq]
-
-lemma int_mod_le: "(0::int) <= a ==> a mod n <= a"
-  by (fact zmod_le_nonneg_dividend) (* FIXME: delete *)
-
-lemma int_mod_le': "(0::int) <= b - n ==> b mod n <= b - n"
-  using zmod_le_nonneg_dividend [of "b - n" "n"] by simp
-
 lemma int_mod_ge: "a < n ==> 0 < (n :: int) ==> a <= a mod n"
   apply (cases "0 <= a")
    apply (drule (1) mod_pos_pos_trivial)
@@ -190,121 +58,36 @@
 lemma int_mod_ge': "b < 0 ==> 0 < (n :: int) ==> b + n <= b mod n"
   by (rule int_mod_ge [where a = "b + n" and n = n, simplified])
 
-lemma mod_add_if_z:
-  "(x :: int) < z ==> y < z ==> 0 <= y ==> 0 <= x ==> 0 <= z ==> 
-   (x + y) mod z = (if x + y < z then x + y else x + y - z)"
-  by (auto intro: int_mod_eq)
-
-lemma mod_sub_if_z:
-  "(x :: int) < z ==> y < z ==> 0 <= y ==> 0 <= x ==> 0 <= z ==> 
-   (x - y) mod z = (if y <= x then x - y else x - y + z)"
-  by (auto intro: int_mod_eq)
-
-lemmas zmde = zmod_zdiv_equality [THEN diff_eq_eq [THEN iffD2], symmetric]
-lemmas mcl = mult_cancel_left [THEN iffD1, THEN make_pos_rule]
-
-(* already have this for naturals, div_mult_self1/2, but not for ints *)
-lemma zdiv_mult_self: "m ~= (0 :: int) ==> (a + m * n) div m = a div m + n"
-  apply (rule mcl)
-   prefer 2
-   apply (erule asm_rl)
-  apply (simp add: zmde ring_distribs)
-  done
+lemma zless2:
+  "0 < (2 :: int)"
+  by arith
 
-lemma mod_power_lem:
-  "a > 1 ==> a ^ n mod a ^ m = (if m <= n then 0 else (a :: int) ^ n)"
-  apply clarsimp
-  apply safe
-   apply (simp add: dvd_eq_mod_eq_0 [symmetric])
-   apply (drule le_iff_add [THEN iffD1])
-   apply (force simp: power_add)
-  apply (rule mod_pos_pos_trivial)
-   apply (simp)
-  apply (rule power_strict_increasing)
-   apply auto
-  done
-
-lemma min_pm [simp]: "min a b + (a - b) = (a :: nat)" by arith
-  
-lemmas min_pm1 [simp] = trans [OF add_commute min_pm]
+lemma zless2p:
+  "0 < (2 ^ n :: int)"
+  by arith
 
-lemma rev_min_pm [simp]: "min b a + (a - b) = (a::nat)" by arith
-
-lemmas rev_min_pm1 [simp] = trans [OF add_commute rev_min_pm]
-
-lemma pl_pl_rels: 
-  "a + b = c + d ==> 
-   a >= c & b <= d | a <= c & b >= (d :: nat)" by arith
-
-lemmas pl_pl_rels' = add_commute [THEN [2] trans, THEN pl_pl_rels]
+lemma zle2p:
+  "0 \<le> (2 ^ n :: int)"
+  by arith
 
-lemma minus_eq: "(m - k = m) = (k = 0 | m = (0 :: nat))"  by arith
-
-lemma pl_pl_mm: "(a :: nat) + b = c + d ==> a - c = d - b"  by arith
-
-lemmas pl_pl_mm' = add_commute [THEN [2] trans, THEN pl_pl_mm]
- 
-lemma min_minus [simp] : "min m (m - k) = (m - k :: nat)" by arith
-  
-lemmas min_minus' [simp] = trans [OF min_max.inf_commute min_minus]
+lemma int_mod_le': "(0::int) <= b - n ==> b mod n <= b - n"
+  using zmod_le_nonneg_dividend [of "b - n" "n"] by simp
 
-lemmas dme = box_equals [OF div_mod_equality add_0_right add_0_right]
-lemmas dtle = xtr3 [OF dme [symmetric] le_add1]
-lemmas th2 = order_trans [OF order_refl [THEN [2] mult_le_mono] dtle]
+lemma diff_le_eq': "(a - b \<le> c) = (a \<le> b + (c::int))" by arith
 
-lemma td_gal: 
-  "0 < c ==> (a >= b * c) = (a div c >= (b :: nat))"
-  apply safe
-   apply (erule (1) xtr4 [OF div_le_mono div_mult_self_is_m])
-  apply (erule th2)
-  done
-  
-lemmas td_gal_lt = td_gal [simplified not_less [symmetric], simplified]
-
-lemma div_mult_le: "(a :: nat) div b * b <= a"
-  by (fact dtle)
-
-lemmas sdl = split_div_lemma [THEN iffD1, symmetric]
-
-lemma given_quot: "f > (0 :: nat) ==> (f * l + (f - 1)) div f = l"
-  by (rule sdl, assumption) (simp (no_asm))
+lemma m1mod2k:
+  "-1 mod 2 ^ n = (2 ^ n - 1 :: int)"
+  using zless2p by (rule zmod_minus1)
 
-lemma given_quot_alt: "f > (0 :: nat) ==> (l * f + f - Suc 0) div f = l"
-  apply (frule given_quot)
-  apply (rule trans)
-   prefer 2
-   apply (erule asm_rl)
-  apply (rule_tac f="%n. n div f" in arg_cong)
-  apply (simp add : mult_ac)
-  done
-    
-lemma diff_mod_le: "(a::nat) < d ==> b dvd d ==> a - a mod b <= d - b"
-  apply (unfold dvd_def)
-  apply clarify
-  apply (case_tac k)
-   apply clarsimp
-  apply clarify
-  apply (cases "b > 0")
-   apply (drule mult_commute [THEN xtr1])
-   apply (frule (1) td_gal_lt [THEN iffD1])
-   apply (clarsimp simp: le_simps)
-   apply (rule mult_div_cancel [THEN [2] xtr4])
-   apply (rule mult_mono)
-      apply auto
-  done
+lemma p1mod22k':
+  fixes b :: int
+  shows "(1 + 2 * b) mod (2 * 2 ^ n) = 1 + 2 * (b mod 2 ^ n)"
+  using zle2p by (rule pos_zmod_mult_2) 
 
-lemma less_le_mult':
-  "w * c < b * c ==> 0 \<le> c ==> (w + 1) * c \<le> b * (c::int)"
-  apply (rule mult_right_mono)
-   apply (rule zless_imp_add1_zle)
-   apply (erule (1) mult_right_less_imp_less)
-  apply assumption
-  done
-
-lemmas less_le_mult = less_le_mult' [simplified distrib_right, simplified]
-
-lemmas less_le_mult_minus = iffD2 [OF le_diff_eq less_le_mult, 
-  simplified left_diff_distrib]
+lemma p1mod22k:
+  fixes b :: int
+  shows "(2 * b + 1) mod (2 * 2 ^ n) = 2 * (b mod 2 ^ n) + 1"
+  by (simp add: p1mod22k' add_commute)
 
 lemma lrlem':
   assumes d: "(i::nat) \<le> j \<or> m < j'"
@@ -328,15 +111,5 @@
   apply simp
   done
 
-lemma gen_minus: "0 < n ==> f n = f (Suc (n - 1))"
-  by auto
-
-lemma mpl_lem: "j <= (i :: nat) ==> k < j ==> i - j + k < i" by arith
+end
 
-lemma nonneg_mod_div:
-  "0 <= a ==> 0 <= b ==> 0 <= (a mod b :: int) & 0 <= a div b"
-  apply (cases "b = 0", clarsimp)
-  apply (auto intro: pos_imp_zdiv_nonneg_iff [THEN iffD2])
-  done
-
-end
--- a/src/HOL/Word/Misc_Typedef.thy	Sun Aug 18 15:59:48 2013 +0200
+++ b/src/HOL/Word/Misc_Typedef.thy	Sun Aug 18 17:00:10 2013 +0200
@@ -1,7 +1,7 @@
 (* 
   Author:     Jeremy Dawson and Gerwin Klein, NICTA
 
-  Consequences of type definition theorems, and of extended type definition. theorems
+  Consequences of type definition theorems, and of extended type definition.
 *)
 
 header {* Type Definition Theorems *}
--- a/src/HOL/Word/Word.thy	Sun Aug 18 15:59:48 2013 +0200
+++ b/src/HOL/Word/Word.thy	Sun Aug 18 17:00:10 2013 +0200
@@ -7,9 +7,10 @@
 theory Word
 imports
   Type_Length
-  Misc_Typedef
   "~~/src/HOL/Library/Boolean_Algebra"
   Bool_List_Representation
+  Misc_Typedef
+  Word_Miscellaneous
 begin
 
 text {* see @{text "Examples/WordExamples.thy"} for examples *}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Word/Word_Miscellaneous.thy	Sun Aug 18 17:00:10 2013 +0200
@@ -0,0 +1,389 @@
+(*  Title:      HOL/Word/Word_Miscellaneous.thy
+    Author:     Miscellaneous
+*)
+
+header {* Miscellaneous lemmas, of at least doubtful value *}
+
+theory Word_Miscellaneous
+imports Main Parity "~~/src/HOL/Library/Bit" Misc_Numeric
+begin
+
+lemma power_minus_simp:
+  "0 < n \<Longrightarrow> a ^ n = a * a ^ (n - 1)"
+  by (auto dest: gr0_implies_Suc)
+
+lemma funpow_minus_simp:
+  "0 < n \<Longrightarrow> f ^^ n = f \<circ> f ^^ (n - 1)"
+  by (auto dest: gr0_implies_Suc)
+
+lemma power_numeral:
+  "a ^ numeral k = a * a ^ (pred_numeral k)"
+  by (simp add: numeral_eq_Suc)
+
+lemma funpow_numeral [simp]:
+  "f ^^ numeral k = f \<circ> f ^^ (pred_numeral k)"
+  by (simp add: numeral_eq_Suc)
+
+lemma replicate_numeral [simp]:
+  "replicate (numeral k) x = x # replicate (pred_numeral k) x"
+  by (simp add: numeral_eq_Suc)
+
+lemma rco_alt: "(f o g) ^^ n o f = f o (g o f) ^^ n"
+  apply (rule ext)
+  apply (induct n)
+   apply (simp_all add: o_def)
+  done
+
+lemma list_exhaust_size_gt0:
+  assumes y: "\<And>a list. y = a # list \<Longrightarrow> P"
+  shows "0 < length y \<Longrightarrow> P"
+  apply (cases y, simp)
+  apply (rule y)
+  apply fastforce
+  done
+
+lemma list_exhaust_size_eq0:
+  assumes y: "y = [] \<Longrightarrow> P"
+  shows "length y = 0 \<Longrightarrow> P"
+  apply (cases y)
+   apply (rule y, simp)
+  apply simp
+  done
+
+lemma size_Cons_lem_eq:
+  "y = xa # list ==> size y = Suc k ==> size list = k"
+  by auto
+
+lemmas ls_splits = prod.split prod.split_asm split_if_asm
+
+lemma not_B1_is_B0: "y \<noteq> (1::bit) \<Longrightarrow> y = (0::bit)"
+  by (cases y) auto
+
+lemma B1_ass_B0: 
+  assumes y: "y = (0::bit) \<Longrightarrow> y = (1::bit)"
+  shows "y = (1::bit)"
+  apply (rule classical)
+  apply (drule not_B1_is_B0)
+  apply (erule y)
+  done
+
+-- "simplifications for specific word lengths"
+lemmas n2s_ths [THEN eq_reflection] = add_2_eq_Suc add_2_eq_Suc'
+
+lemmas s2n_ths = n2s_ths [symmetric]
+
+lemma and_len: "xs = ys ==> xs = ys & length xs = length ys"
+  by auto
+
+lemma size_if: "size (if p then xs else ys) = (if p then size xs else size ys)"
+  by auto
+
+lemma tl_if: "tl (if p then xs else ys) = (if p then tl xs else tl ys)"
+  by auto
+
+lemma hd_if: "hd (if p then xs else ys) = (if p then hd xs else hd ys)"
+  by auto
+
+lemma if_Not_x: "(if p then ~ x else x) = (p = (~ x))"
+  by auto
+
+lemma if_x_Not: "(if p then x else ~ x) = (p = x)"
+  by auto
+
+lemma if_same_and: "(If p x y & If p u v) = (if p then x & u else y & v)"
+  by auto
+
+lemma if_same_eq: "(If p x y  = (If p u v)) = (if p then x = (u) else y = (v))"
+  by auto
+
+lemma if_same_eq_not:
+  "(If p x y  = (~ If p u v)) = (if p then x = (~u) else y = (~v))"
+  by auto
+
+(* note - if_Cons can cause blowup in the size, if p is complex,
+  so make a simproc *)
+lemma if_Cons: "(if p then x # xs else y # ys) = If p x y # If p xs ys"
+  by auto
+
+lemma if_single:
+  "(if xc then [xab] else [an]) = [if xc then xab else an]"
+  by auto
+
+lemma if_bool_simps:
+  "If p True y = (p | y) & If p False y = (~p & y) & 
+    If p y True = (p --> y) & If p y False = (p & y)"
+  by auto
+
+lemmas if_simps = if_x_Not if_Not_x if_cancel if_True if_False if_bool_simps
+
+lemmas seqr = eq_reflection [where x = "size w"] for w (* FIXME: delete *)
+
+(* TODO: move name bindings to List.thy *)
+lemmas tl_Nil = tl.simps (1)
+lemmas tl_Cons = tl.simps (2)
+
+lemma the_elemI: "y = {x} ==> the_elem y = x" 
+  by simp
+
+lemma nonemptyE: "S ~= {} ==> (!!x. x : S ==> R) ==> R" by auto
+
+lemma gt_or_eq_0: "0 < y \<or> 0 = (y::nat)" by arith 
+
+lemmas xtr1 = xtrans(1)
+lemmas xtr2 = xtrans(2)
+lemmas xtr3 = xtrans(3)
+lemmas xtr4 = xtrans(4)
+lemmas xtr5 = xtrans(5)
+lemmas xtr6 = xtrans(6)
+lemmas xtr7 = xtrans(7)
+lemmas xtr8 = xtrans(8)
+
+lemmas nat_simps = diff_add_inverse2 diff_add_inverse
+lemmas nat_iffs = le_add1 le_add2
+
+lemma sum_imp_diff: "j = k + i ==> j - i = (k :: nat)" by arith
+
+lemmas pos_mod_sign2 = zless2 [THEN pos_mod_sign [where b = "2::int"]]
+lemmas pos_mod_bound2 = zless2 [THEN pos_mod_bound [where b = "2::int"]]
+
+lemma nmod2: "n mod (2::int) = 0 | n mod 2 = 1" by arith
+
+lemma emep1:
+  "even n ==> even d ==> 0 <= d ==> (n + 1) mod (d :: int) = (n mod d) + 1"
+  apply (simp add: add_commute)
+  apply (safe dest!: even_equiv_def [THEN iffD1])
+  apply (subst pos_zmod_mult_2)
+   apply arith
+  apply (simp add: mod_mult_mult1)
+ done
+
+lemmas eme1p = emep1 [simplified add_commute]
+
+lemma le_diff_eq': "(a \<le> c - b) = (b + a \<le> (c::int))" by arith
+
+lemma less_diff_eq': "(a < c - b) = (b + a < (c::int))" by arith
+
+lemma diff_less_eq': "(a - b < c) = (a < b + (c::int))" by arith
+
+lemmas m1mod22k = mult_pos_pos [OF zless2 zless2p, THEN zmod_minus1]
+lemmas z1pmod2' = zero_le_one [THEN pos_zmod_mult_2, simplified]
+lemmas z1pdiv2' = zero_le_one [THEN pos_zdiv_mult_2, simplified]
+
+lemma z1pmod2:
+  "(2 * b + 1) mod 2 = (1::int)" by arith
+  
+lemma z1pdiv2:
+  "(2 * b + 1) div 2 = (b::int)" by arith
+
+lemmas zdiv_le_dividend = xtr3 [OF div_by_1 [symmetric] zdiv_mono2,
+  simplified int_one_le_iff_zero_less, simplified]
+  
+lemma axxbyy:
+  "a + m + m = b + n + n ==> (a = 0 | a = 1) ==> (b = 0 | b = 1) ==>  
+   a = b & m = (n :: int)" by arith
+
+lemma axxmod2:
+  "(1 + x + x) mod 2 = (1 :: int) & (0 + x + x) mod 2 = (0 :: int)" by arith
+
+lemma axxdiv2:
+  "(1 + x + x) div 2 = (x :: int) & (0 + x + x) div 2 = (x :: int)"  by arith
+
+lemmas iszero_minus = trans [THEN trans,
+  OF iszero_def neg_equal_0_iff_equal iszero_def [symmetric]]
+
+lemmas zadd_diff_inverse = trans [OF diff_add_cancel [symmetric] add_commute]
+
+lemmas add_diff_cancel2 = add_commute [THEN diff_eq_eq [THEN iffD2]]
+
+lemmas rdmods [symmetric] = mod_minus_eq
+  mod_diff_left_eq mod_diff_right_eq mod_add_left_eq
+  mod_add_right_eq mod_mult_right_eq mod_mult_left_eq
+
+lemma mod_plus_right:
+  "((a + x) mod m = (b + x) mod m) = (a mod m = b mod (m :: nat))"
+  apply (induct x)
+   apply (simp_all add: mod_Suc)
+  apply arith
+  done
+
+lemma nat_minus_mod: "(n - n mod m) mod m = (0 :: nat)"
+  by (induct n) (simp_all add : mod_Suc)
+
+lemmas nat_minus_mod_plus_right = trans [OF nat_minus_mod mod_0 [symmetric],
+  THEN mod_plus_right [THEN iffD2], simplified]
+
+lemmas push_mods' = mod_add_eq
+  mod_mult_eq mod_diff_eq
+  mod_minus_eq
+
+lemmas push_mods = push_mods' [THEN eq_reflection]
+lemmas pull_mods = push_mods [symmetric] rdmods [THEN eq_reflection]
+lemmas mod_simps = 
+  mod_mult_self2_is_0 [THEN eq_reflection]
+  mod_mult_self1_is_0 [THEN eq_reflection]
+  mod_mod_trivial [THEN eq_reflection]
+
+lemma nat_mod_eq:
+  "!!b. b < n ==> a mod n = b mod n ==> a mod n = (b :: nat)" 
+  by (induct a) auto
+
+lemmas nat_mod_eq' = refl [THEN [2] nat_mod_eq]
+
+lemma nat_mod_lem: 
+  "(0 :: nat) < n ==> b < n = (b mod n = b)"
+  apply safe
+   apply (erule nat_mod_eq')
+  apply (erule subst)
+  apply (erule mod_less_divisor)
+  done
+
+lemma mod_nat_add: 
+  "(x :: nat) < z ==> y < z ==> 
+   (x + y) mod z = (if x + y < z then x + y else x + y - z)"
+  apply (rule nat_mod_eq)
+   apply auto
+  apply (rule trans)
+   apply (rule le_mod_geq)
+   apply simp
+  apply (rule nat_mod_eq')
+  apply arith
+  done
+
+lemma mod_nat_sub: 
+  "(x :: nat) < z ==> (x - y) mod z = x - y"
+  by (rule nat_mod_eq') arith
+
+lemma int_mod_lem: 
+  "(0 :: int) < n ==> (0 <= b & b < n) = (b mod n = b)"
+  apply safe
+    apply (erule (1) mod_pos_pos_trivial)
+   apply (erule_tac [!] subst)
+   apply auto
+  done
+
+lemma int_mod_eq:
+  "(0 :: int) <= b ==> b < n ==> a mod n = b mod n ==> a mod n = b"
+  by clarsimp (rule mod_pos_pos_trivial)
+
+lemmas int_mod_eq' = refl [THEN [3] int_mod_eq]
+
+lemma int_mod_le: "(0::int) <= a ==> a mod n <= a"
+  by (fact zmod_le_nonneg_dividend) (* FIXME: delete *)
+
+lemma mod_add_if_z:
+  "(x :: int) < z ==> y < z ==> 0 <= y ==> 0 <= x ==> 0 <= z ==> 
+   (x + y) mod z = (if x + y < z then x + y else x + y - z)"
+  by (auto intro: int_mod_eq)
+
+lemma mod_sub_if_z:
+  "(x :: int) < z ==> y < z ==> 0 <= y ==> 0 <= x ==> 0 <= z ==> 
+   (x - y) mod z = (if y <= x then x - y else x - y + z)"
+  by (auto intro: int_mod_eq)
+
+lemmas zmde = zmod_zdiv_equality [THEN diff_eq_eq [THEN iffD2], symmetric]
+lemmas mcl = mult_cancel_left [THEN iffD1, THEN make_pos_rule]
+
+(* already have this for naturals, div_mult_self1/2, but not for ints *)
+lemma zdiv_mult_self: "m ~= (0 :: int) ==> (a + m * n) div m = a div m + n"
+  apply (rule mcl)
+   prefer 2
+   apply (erule asm_rl)
+  apply (simp add: zmde ring_distribs)
+  done
+
+lemma mod_power_lem:
+  "a > 1 ==> a ^ n mod a ^ m = (if m <= n then 0 else (a :: int) ^ n)"
+  apply clarsimp
+  apply safe
+   apply (simp add: dvd_eq_mod_eq_0 [symmetric])
+   apply (drule le_iff_add [THEN iffD1])
+   apply (force simp: power_add)
+  apply (rule mod_pos_pos_trivial)
+   apply (simp)
+  apply (rule power_strict_increasing)
+   apply auto
+  done
+
+lemma pl_pl_rels: 
+  "a + b = c + d ==> 
+   a >= c & b <= d | a <= c & b >= (d :: nat)" by arith
+
+lemmas pl_pl_rels' = add_commute [THEN [2] trans, THEN pl_pl_rels]
+
+lemma minus_eq: "(m - k = m) = (k = 0 | m = (0 :: nat))"  by arith
+
+lemma pl_pl_mm: "(a :: nat) + b = c + d ==> a - c = d - b"  by arith
+
+lemmas pl_pl_mm' = add_commute [THEN [2] trans, THEN pl_pl_mm]
+
+lemmas dme = box_equals [OF div_mod_equality add_0_right add_0_right]
+lemmas dtle = xtr3 [OF dme [symmetric] le_add1]
+lemmas th2 = order_trans [OF order_refl [THEN [2] mult_le_mono] dtle]
+
+lemma td_gal: 
+  "0 < c ==> (a >= b * c) = (a div c >= (b :: nat))"
+  apply safe
+   apply (erule (1) xtr4 [OF div_le_mono div_mult_self_is_m])
+  apply (erule th2)
+  done
+  
+lemmas td_gal_lt = td_gal [simplified not_less [symmetric], simplified]
+
+lemma div_mult_le: "(a :: nat) div b * b <= a"
+  by (fact dtle)
+
+lemmas sdl = split_div_lemma [THEN iffD1, symmetric]
+
+lemma given_quot: "f > (0 :: nat) ==> (f * l + (f - 1)) div f = l"
+  by (rule sdl, assumption) (simp (no_asm))
+
+lemma given_quot_alt: "f > (0 :: nat) ==> (l * f + f - Suc 0) div f = l"
+  apply (frule given_quot)
+  apply (rule trans)
+   prefer 2
+   apply (erule asm_rl)
+  apply (rule_tac f="%n. n div f" in arg_cong)
+  apply (simp add : mult_ac)
+  done
+    
+lemma diff_mod_le: "(a::nat) < d ==> b dvd d ==> a - a mod b <= d - b"
+  apply (unfold dvd_def)
+  apply clarify
+  apply (case_tac k)
+   apply clarsimp
+  apply clarify
+  apply (cases "b > 0")
+   apply (drule mult_commute [THEN xtr1])
+   apply (frule (1) td_gal_lt [THEN iffD1])
+   apply (clarsimp simp: le_simps)
+   apply (rule mult_div_cancel [THEN [2] xtr4])
+   apply (rule mult_mono)
+      apply auto
+  done
+
+lemma less_le_mult':
+  "w * c < b * c ==> 0 \<le> c ==> (w + 1) * c \<le> b * (c::int)"
+  apply (rule mult_right_mono)
+   apply (rule zless_imp_add1_zle)
+   apply (erule (1) mult_right_less_imp_less)
+  apply assumption
+  done
+
+lemmas less_le_mult = less_le_mult' [simplified distrib_right, simplified]
+
+lemmas less_le_mult_minus = iffD2 [OF le_diff_eq less_le_mult, 
+  simplified left_diff_distrib]
+
+lemma gen_minus: "0 < n ==> f n = f (Suc (n - 1))"
+  by auto
+
+lemma mpl_lem: "j <= (i :: nat) ==> k < j ==> i - j + k < i" by arith
+
+lemma nonneg_mod_div:
+  "0 <= a ==> 0 <= b ==> 0 <= (a mod b :: int) & 0 <= a div b"
+  apply (cases "b = 0", clarsimp)
+  apply (auto intro: pos_imp_zdiv_nonneg_iff [THEN iffD2])
+  done
+
+end
+